From 8b90d201f4b897d2cf14b5106e062d7389bcf498 Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 23 Oct 2020 14:46:46 +0200 Subject: [PATCH 001/167] migration of Isabelle theories up to NatOrderings --- isabelle/FixedPoints.thy | 59 +- isabelle/Functions.thy | 228 ++++--- isabelle/NatOrderings.thy | 217 +++--- isabelle/Peano.thy | 79 ++- isabelle/PredicateLogic.thy | 1218 ++++++++++++++++------------------ isabelle/ROOT | 7 + isabelle/SetTheory.thy | 696 +++++++++---------- isabelle/document/root.tex | 4 +- isabelle/simplifier_setup.ML | 150 +++-- 9 files changed, 1325 insertions(+), 1333 deletions(-) create mode 100644 isabelle/ROOT diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index 1af421e9..e4a086b3 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -1,28 +1,27 @@ (* Title: TLA+/FixedPoints.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:18 merz> + Version: Isabelle2020 *) -header {* Fixed points for set-theoretical constructions *} +section \Fixed points for set-theoretical constructions\ theory FixedPoints imports SetTheory begin -text {* +text \ As a test for the encoding of \tlaplus{} set theory, we develop the Knaster-Tarski theorems for least and greatest fixed points in the subset lattice. Again, the proofs essentially follow Paulson's developments for Isabelle/ZF. -*} +\ -subsection {* Monotonic operators *} +subsection \ Monotonic operators \ -definition Monotonic :: "[c, c \ c] \ c" -- {* monotonic operator on a domain *} +definition Monotonic :: "[c, c \ c] \ c" \ \monotonic operator on a domain\ where "Monotonic(D,f) \ f(D) \ D \ (\S,T \ SUBSET D : S \ T \ f(S) \ f(T))" lemma monotonicDomain: @@ -37,10 +36,10 @@ lemma monotonicSubsetDomain: "\ Monotonic(D,f); S \ D\ \ f(S) \ D" by (unfold Monotonic_def, blast) -lemma monotonicCup: +lemma monotonicUnion: assumes mono: "Monotonic(D,f)" and s: "S \ D" and t: "T \ D" shows "f(S) \ f(T) \ f(S \ T)" -proof (rule cupLUB) +proof (rule unionLUB) from s t show "f(S) \ f(S \ T)" by (intro monotonicSubset[OF mono], blast+) next @@ -48,10 +47,10 @@ next by (intro monotonicSubset[OF mono], blast+) qed -lemma monotonicCap: +lemma monotonicInter: assumes mono: "Monotonic(D,f)" and s: "S \ D" and t: "T \ D" shows "f(S \ T) \ f(S) \ f(T)" -proof (rule capGLB) +proof (rule interGLB) from s t show "f(S \ T) \ f(S)" by (intro monotonicSubset[OF mono], blast+) from s t show "f(S \ T) \ f(T)" @@ -59,29 +58,29 @@ proof (rule capGLB) qed -subsection {* Least fixed point *} +subsection \ Least fixed point \ -text {* +text \ The least fixed point is defined as the greatest lower bound of the set of all pre-fixed points, and the Knaster-Tarski theorem is shown for monotonic operators. -*} +\ -definition lfp :: "[c, c \ c] \ c" -- {* least fixed point as GLB of pre-fp's *} +definition lfp :: "[c, c \ c] \ c" \ \least fixed point as GLB of pre-fp's\ where "lfp(D,f) \ INTER {S \ SUBSET D : f(S) \ S}" -lemma lfpLB: -- {* The lfp is contained in each pre-fixed point. *} +lemma lfpLB: \ \The lfp is contained in each pre-fixed point.\ "\ f(S) \ S; S \ D \ \ lfp(D,f) \ S" by (auto simp: lfp_def) -lemma lfpGLB: -- {* \ldots{} and it is the GLB of all such sets *} +lemma lfpGLB: \ \ \ldots{} and it is the GLB of all such sets \ "\f(D) \ D; \S. \f(S) \ S; S \ D\ \ A \ S\ \ A \ lfp(D,f)" by (force simp: lfp_def) lemma lfpSubsetDomain: "lfp(D,f) \ D" (* monotonicity not required *) by (auto simp: lfp_def) -lemma lfpPreFP: -- {* @{text lfp} is a pre-fixed point \ldots *} +lemma lfpPreFP: \ \@{text lfp} is a pre-fixed point \ldots \ assumes mono: "Monotonic(D,f)" shows "f(lfp(D,f)) \ lfp(D,f)" proof (rule lfpGLB) @@ -95,7 +94,7 @@ next with pfp show "f(?mu) \ S" by blast qed -lemma lfpPostFP: -- {* \ldots{} and a post-fixed point *} +lemma lfpPostFP: \ \ \ldots{} and a post-fixed point \ assumes mono: "Monotonic(D,f)" shows "lfp(D,f) \ f(lfp(D,f))" proof - @@ -120,7 +119,7 @@ lemma lfpLeastFixedPoint: shows "lfp(D,f) \ S" using assms by (intro lfpLB, auto) -lemma lfpMono: -- {* monotonicity of the @{text lfp} operator *} +lemma lfpMono: \ \ monotonicity of the @{text lfp} operator \ assumes g: "g(D) \ D" and f: "\S. S \ D \ f(S) \ g(S)" shows "lfp(D,f) \ lfp(D,g)" using g @@ -132,29 +131,29 @@ proof (rule lfpGLB) qed -subsection {* Greatest fixed point *} +subsection \ Greatest fixed point \ -text {* +text \ Dually, the least fixed point is defined as the least upper bound of the set of all post-fixed points, and the Knaster-Tarski theorem is again established. -*} +\ -definition gfp :: "[c, c \ c] \ c" -- {* greatest fixed point as LUB of post-fp's *} +definition gfp :: "[c, c \ c] \ c" \ \ greatest fixed point as LUB of post-fp's \ where "gfp(D,f) \ UNION {S \ SUBSET D : S \ f(S)}" -lemma gfpUB: -- {* The gfp contains each post-fixed point \ldots *} +lemma gfpUB: \ \ The gfp contains each post-fixed point \ldots \ "\S \ f(S); S \ D\ \ S \ gfp(D,f)" by (auto simp: gfp_def) -lemma gfpLUB: -- {* \ldots{} and it is the LUB of all such sets. *} +lemma gfpLUB: \ \ \ldots{} and it is the LUB of all such sets. \ "\f(D) \ D; \S. \S \ f(S); S \ D\ \ S \ A\ \ gfp(D,f) \ A" by (auto simp: gfp_def) lemma gfpSubsetDomain: "gfp(D,f) \ D" by (auto simp: gfp_def) -lemma gfpPostFP: -- {* @text{gfp} is a post-fixed point \ldots *} +lemma gfpPostFP: \ \ @text{gfp} is a post-fixed point \ldots \ assumes mono: "Monotonic(D,f)" shows "gfp(D,f) \ f(gfp(D,f))" proof (rule gfpLUB) @@ -168,7 +167,7 @@ next with pfp show "S \ f(?nu)" by blast qed -lemma gfpPreFP: -- {* \ldots{} and a pre-fixed point *} +lemma gfpPreFP: \ \ \ldots{} and a pre-fixed point \ assumes mono: "Monotonic(D,f)" shows "f(gfp(D,f)) \ gfp(D,f)" proof - @@ -193,7 +192,7 @@ lemma gfpGreatestFixedPoint: shows "S \ gfp(D,f)" using assms by (intro gfpUB, auto) -lemma gfpMono: -- {* monotonicity of the @{text gfp} operator *} +lemma gfpMono: \ \ monotonicity of the @{text gfp} operator \ assumes g: "g(D) \ D" and f: "\S. S \ D \ f(S) \ g(S)" shows "gfp(D,f) \ gfp(D,g)" proof (rule gfpLUB) diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index dade3389..922271f0 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -1,67 +1,80 @@ (* Title: TLA+/Functions.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:56 merz> + Version: Isabelle2020 *) -header {* \tlaplus{} Functions *} +section \ \tlaplus{} Functions \ theory Functions imports SetTheory begin -subsection {* Syntax and axioms for functions *} +subsection \ Syntax and axioms for functions \ -text {* +text \ Functions in \tlaplus{} are not defined (e.g., as sets of pairs), but axiomatized, and in fact, pairs and tuples will be defined as special functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined - as functions over products. - - We follow the development of functions given in Section 16.1 of - ``Specifying Systems''. In particular, we define the predicate - @{text IsAFcn} that is true precisely of functional values. -*} + as functions over products. We follow the development of functions given + in Section 16.1 of ``Specifying Systems''. +\ consts - isAFcn :: "c \ c" -- {* characteristic predicate *} - Fcn :: "[c, c \ c] \ c" -- {* function constructor *} - DOMAIN :: "c \ c" ("(DOMAIN _)" [100]90) -- {* domain of a function *} - fapply :: "[c, c] \ c" ("(_[_])" [89,0]90) -- {* function application *} - FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) -- {* function space *} + Fcn :: "[c, c \ c] \ c" \ \ function constructor \ + DOMAIN :: "c \ c" ("(DOMAIN _)" [100]90) \ \ domain of a function \ + fapply :: "[c, c] \ c" ("(_[_])" [89,0]90) \ \ function application \ syntax - "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \\in _ |-> _])" 900) -syntax (xsymbols) - FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) - "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \ _ \ _])" 900) -syntax (HTML output) - FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \ _ \ _])" 900) +syntax (ASCII) + "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \\in _ |-> _])" 900) translations "[x \ S \ e]" \ "CONST Fcn(S, \x. e)" axiomatization where - fcnIsAFcn [intro!,simp]: "isAFcn(Fcn(S,e))" and - isAFcn_def: "isAFcn(f) \ f = [x \ DOMAIN f \ f[x]]" and DOMAIN [simp]: "DOMAIN Fcn(S,e) = S" and - fapply [simp]: "v \ S \ Fcn(S,e)[v] = e(v)" and - fcnEqual[elim!]: "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ - \ f = g" and + fapply [simp]: "v \ S \ Fcn(S,e)[v] = e(v)" + +text \The predicate @{text isAFcn} characterizes functional values.\ +definition isAFcn :: "c \ c" + where "isAFcn(f) \ f = Fcn(DOMAIN f, \x. f[x])" + +axiomatization where + fcnIsAFcn [intro!, simp]: "isAFcn(Fcn(S,e))" + +text \ + Two functions are equal if they have the same domain and agree on all + arguments within the domain. +\ +axiomatization where fcnEqual[elim!]: + "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ + \ f = g" + +text \ + @{text "FuncSet(S,T)"} represents the set of functions with domain @{text S} + and co-domain @{text T}. It cannot be defined as a set comprehension because + we do not have a bounding set for its elements. +\ +consts + FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) \ \ function space \ +syntax (ASCII) + FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) + +axiomatization where FuncSet: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T)" -lemmas -- {* establish set equality for domains and function spaces *} - setEqualI [where A = "DOMAIN f", standard, intro!] - setEqualI [where B = "DOMAIN f", standard, intro!] - setEqualI [where A = "[S \ T]", standard, intro!] - setEqualI [where B = "[S \ T]", standard, intro!] +lemmas \ \ establish set equality for domains and function spaces \ + setEqualI [where A = "DOMAIN f" for f, intro!] + setEqualI [where B = "DOMAIN f" for f, intro!] + setEqualI [where A = "[S \ T]" for S T, intro!] + setEqualI [where B = "[S \ T]" for S T, intro!] -definition except :: "[c,c,c] \ c" -- {* function override *} +definition except :: "[c,c,c] \ c" \ \ function override \ where "except(f,v,e) \ [x \ DOMAIN f \ (IF x=v THEN e ELSE f[x])]" nonterminal @@ -74,11 +87,11 @@ translations "_except(f, _xcpts(v,e, xcs))" \ "_except(CONST except(f,v,e), xcs)" "[f EXCEPT ![v] = e]" \ "CONST except(f,v,e)" -text {* +text \ The following operators are useful for representing functions with finite domains by enumeration. They are not part of basic \tlaplus{}, but they are defined in the TLC module of the standard library. -*} +\ definition oneArg :: "[c,c] \ c" (infixl ":>" 75) where "d :> e \ [x \ {d} \ e]" @@ -88,7 +101,7 @@ where "f @@ g \ [x \ (DOMAIN f) \ (DOMAIN g) \ IF x \ DOMAIN f THEN f[x] ELSE g[x]]" -subsection {* @{text isAFcn}: identifying functional values *} +subsection \ @{text isAFcn}: identifying functional values \ lemma boolifyIsAFcn [simp]: "boolify(isAFcn(f)) = isAFcn(f)" by (simp add: isAFcn_def) @@ -106,17 +119,17 @@ by auto lemma [intro!,simp]: "isAFcn([f EXCEPT ![c] = e])" by (simp add: except_def) -text {* +text \ We derive instances of axiom @{text fcnEqual} that help in automating proofs about equality of functions. -*} +\ lemma fcnEqual2[elim!]: "\isAFcn(g); isAFcn(f); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" by (rule fcnEqual) --- {* possibly useful as a simplification rule, but cannot be active by default *} +text \ possibly useful as a simplification rule, but cannot be active by default \ lemma fcnEqualIff: assumes "isAFcn(f)" and "isAFcn(g)" shows "(f = g) = (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : f[x] = g[x]))" @@ -137,7 +150,7 @@ lemma [intro!]: by (auto simp: except_def) -subsection {* Theorems about functions *} +subsection \ Theorems about functions \ lemma fcnCong (*[cong] -- not active by default (??) *): assumes "S = T" and "\x. x \ T \ e(x) = f(x)" @@ -167,11 +180,11 @@ lemma exceptEqual [simp]: shows "([f EXCEPT ![v] = e] = f) = (v \ DOMAIN f \ f[v] = e)" using assms by (auto simp: fcnEqualIff) -text {* +text \ A function can be defined from a predicate. Using the @{text CHOOSE} operator, the definition does not require the predicate to be functional. -*} +\ lemma fcnConstruct: assumes hyp: "\x \ S : \y : P(x,y)" @@ -190,7 +203,7 @@ proof - qed -subsection {* Function spaces *} +subsection \ Function spaces \ lemma inFuncSetIff: "(f \ [S \ T]) = (isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T))" @@ -267,12 +280,12 @@ next from assms show "\x \ S : ?exc[x] \ T" by auto qed (simp) -text {* +text \ The following special case is useful for invariant proofs where one proves type correctness. The additional hypotheses make the type of $f$ available and are useful, for example, when the expression $e$ is of the form $f[u]$ for some $u \in S$. -*} +\ lemma exceptInFuncSetSame: assumes "f \ [S \ T]" @@ -281,7 +294,7 @@ lemma exceptInFuncSetSame: using assms by auto -subsection {* Finite functions and extension *} +subsection \ Finite functions and extension \ lemma oneArgIsAFcn [simp, intro!]: "isAFcn(d :> e)" by (simp add: oneArg_def) @@ -321,10 +334,10 @@ lemma oneArgEqualIff [simp]: "isAFcn(f) \ ((d :> e) = f) = ((DOMAIN f = {d}) \ f[d] = e)" by auto --- {* infer equalities @{text "f = g @@ h"} *} +text \ infer equations @{text "f = g @@ h"} \ lemmas - fcnEqual[where f = "f @@ h", standard, intro!] - fcnEqual[where g = "g @@ h", standard, intro!] + fcnEqual[where f = "f @@ h" for f h, intro!] + fcnEqual[where g = "g @@ h" for g h, intro!] lemma extendEqualIff [simp]: "isAFcn(f) \ (f = g @@ h) = @@ -352,25 +365,25 @@ by auto **) -subsection {* Notions about functions *} +subsection \ Notions about functions \ -subsubsection {* Image and Range *} +subsubsection \ Image and Range \ -text {* +text \ Image of a set under a function, and range of a function. Because the application of a function to an argument outside of its domain usually leads to silliness, we restrict to the domain when defining the image. -*} +\ definition Image where "Image(f,A) \ { f[x] : x \ A \ DOMAIN f }" -text {* - The range of a function, introduced as an abbreviation (macro). +text \ + The range of a function, introduced as an abbreviation. To reason about the range, apply the theorems about @{text Image}, or simply rewrite with @{text Image_def}. -*} +\ abbreviation Range where "Range(f) \ Image(f, DOMAIN f)" @@ -390,7 +403,7 @@ lemma imageI_exEq [intro]: shows "y \ Image(f,A)" using assms by (auto intro: imageI_eq) -lemma rangeI: -- {* useful special case *} +lemma rangeI: \ \ useful special case \ assumes "\x \ DOMAIN f: y = f[x]" shows "y \ Range(f)" using assms by auto @@ -416,20 +429,20 @@ lemma imageEmpty [simp]: "({} = Image(f,A)) = (A \ DOMAIN f = {})" by auto -subsubsection {* Injective functions *} +subsubsection \ Injective functions \ definition InjectiveOn where "InjectiveOn(f,A) \ \x,y \ A \ DOMAIN f : f[x] = f[y] \ x = y" -abbreviation Injective -- {* special case: injective function *} +abbreviation Injective \ \ special case: injective function \ where "Injective(f) \ InjectiveOn(f, DOMAIN f)" definition Injections where "Injections(S,T) \ { f \ [S \ T] : Injective(f) }" lemmas - setEqualI [where A = "Injections(S,T)", standard, intro!] - setEqualI [where B = "Injections(S,T)", standard, intro!] + setEqualI [where A = "Injections(S,T)" for S T, intro!] + setEqualI [where B = "Injections(S,T)" for S T, intro!] lemma injectiveOnIsBool [intro!,simp]: "isBool(InjectiveOn(f,A))" @@ -439,9 +452,7 @@ lemma boolifyInjectiveOn [simp]: "boolify(InjectiveOn(f,A)) = InjectiveOn(f,A)" by auto -text {* - For the moment, no support by default for automatic reasoning. -*} +text \ For the moment, no support by default for automatic reasoning. \ lemma injectiveOnI: assumes "\x y. \ x \ A; x \ DOMAIN f; y \ A; y \ DOMAIN f; f[x] = f[y] \ \ x = y" @@ -460,7 +471,7 @@ lemma injectiveOnE: shows "P" using assms by (auto simp: InjectiveOn_def) -lemma injectiveOnIff: -- {* useful for simplification *} +lemma injectiveOnIff: \ \ useful for simplification \ assumes "InjectiveOn(f,A)" and "x \ A \ DOMAIN f" and "y \ A \ DOMAIN f" shows "(f[x] = f[y]) = (x = y)" using assms injectiveOnD by auto @@ -470,14 +481,12 @@ lemma injectiveOnSubset: shows "InjectiveOn(f,B)" using assms by (auto simp: InjectiveOn_def) -lemma injectiveOnDifference: +lemma injectiveOnSetminus: assumes "InjectiveOn(f,A)" - shows "InjectiveOn(f, A \\ B)" + shows "InjectiveOn(f, A \ B)" using assms by (auto simp: InjectiveOn_def) -text {* - The existence of an inverse function implies injectivity. -*} +text \The existence of an inverse function implies injectivity.\ lemma inverseThenInjective: assumes inv: "\x. \ x \ A; x \ DOMAIN f \ \ g[f[x]] = x" @@ -492,7 +501,7 @@ proof (rule injectiveOnI) from x1 y1 eq show "x = y" by simp qed -text {* Trivial cases. *} +text \ Trivial cases. \ lemma injectiveOnEmpty [intro!,simp]: "InjectiveOn(f, {})" @@ -507,7 +516,7 @@ lemma "Injective(d :> e)" by simp **) -text {* Injectivity for function extensions. *} +text \ Injectivity and EXCEPT. \ lemma injectiveOnExcept: assumes 1: "InjectiveOn(f, A \\ {v})" and 2: "isAFcn(f)" @@ -541,7 +550,7 @@ proof (rule injectiveOnI) with fx eq show "FALSE" by simp qed with x1 x2 have x3: "x \ (A \\ {v})" "x \ DOMAIN f" by auto - have "y \ v" -- {* symmetrical reasoning *} + have "y \ v" \ \symmetrical reasoning\ proof assume contr: "y = v" with True have fy: "?exc[y] = e" by auto @@ -557,7 +566,7 @@ proof (rule injectiveOnI) qed lemma injectiveOnExtend: - assumes f: "InjectiveOn(f,A)" and g: "InjectiveOn(g,A \\ DOMAIN f)" + assumes f: "InjectiveOn(f,A)" and g: "InjectiveOn(g, A \\ DOMAIN f)" and disj: "Image(f, A) \ Image(g, A \\ DOMAIN f) = {}" shows "InjectiveOn(f @@ g, A)" proof (rule injectiveOnI) @@ -596,25 +605,25 @@ proof (rule injectiveOnI) qed qed -lemma injectiveExtend: -- {* special case *} - assumes 1: "Injective(f)" and 2: "InjectiveOn(g, DOMAIN g \\ DOMAIN f)" - and 3: "Range(f) \ Image(g, DOMAIN g \\ DOMAIN f) = {}" +lemma injectiveExtend: \ \ special case \ + assumes 1: "Injective(f)" and 2: "InjectiveOn(g, DOMAIN g \ DOMAIN f)" + and 3: "Range(f) \ Image(g, DOMAIN g \ DOMAIN f) = {}" shows "Injective(f @@ g)" proof (rule injectiveOnExtend) from 1 show "InjectiveOn(f, DOMAIN (f @@ g))" by (auto simp: InjectiveOn_def) next - from 2 show "InjectiveOn(g, DOMAIN (f @@ g) \\ DOMAIN f)" + from 2 show "InjectiveOn(g, DOMAIN (f @@ g) \ DOMAIN f)" by (auto simp: InjectiveOn_def) next - show "Image(f, DOMAIN (f @@ g)) \ Image(g, DOMAIN (f @@ g) \\ DOMAIN f) = {}" + show "Image(f, DOMAIN (f @@ g)) \ Image(g, DOMAIN (f @@ g) \ DOMAIN f) = {}" proof (clarify) fix x assume xf: "x \ Image(f, DOMAIN (f @@ g))" - and xg: "x \ Image(g, DOMAIN (f @@ g) \\ DOMAIN f)" + and xg: "x \ Image(g, DOMAIN (f @@ g) \ DOMAIN f)" from xf have "x \ Range(f)" by auto moreover - from xg have "x \ Image(g, DOMAIN g \\ DOMAIN f)" by auto + from xg have "x \ Image(g, DOMAIN g \ DOMAIN f)" by auto moreover note 3 ultimately show FALSE by blast @@ -626,9 +635,9 @@ lemma injectiveOnImageInter: shows "Image(f, B \ C) = Image(f,B) \ Image(f,C)" using assms by (auto simp: InjectiveOn_def Image_def) -lemma injectiveOnImageDifference: +lemma injectiveOnImageSetminus: assumes "InjectiveOn(f,A)" and "B \ A" and "C \ A" - shows "Image(f, B \\ C) = Image(f,B) \\ Image(f,C)" + shows "Image(f, B \ C) = Image(f,B) \ Image(f,C)" using assms by (auto simp: InjectiveOn_def Image_def) lemma injectiveImageMember: @@ -639,14 +648,17 @@ using assms by (auto simp: InjectiveOn_def Image_def) lemma injectiveImageSubset: assumes f: "Injective(f)" shows "(Image(f,A) \ Image(f,B)) = (A \ DOMAIN f \ B \ DOMAIN f)" -proof (auto) -- {* the inclusion ``$\supseteq$'' is solved automatically *} - fix x - assume ab: "Image(f,A) \ Image(f,B)" and x: "x \ A" "x \ DOMAIN f" - from x have "f[x] \ Image(f,A)" by (rule imageI) - with ab have "f[x] \ Image(f,B)" .. - then obtain z where z: "z \ B \ DOMAIN f" "f[x] = f[z]" by auto - from f z x have "x = z" by (auto elim: injectiveOnD) - with z show "x \ B" by simp +proof - + { + fix x + assume ab: "Image(f,A) \ Image(f,B)" and x: "x \ A" "x \ DOMAIN f" + from x have "f[x] \ Image(f,A)" by (rule imageI) + with ab have "f[x] \ Image(f,B)" .. + then obtain z where z: "z \ B \ DOMAIN f" "f[x] = f[z]" by auto + from f z x have "x = z" by (auto elim: injectiveOnD) + with z have "x \ B" by simp + } + thus ?thesis by blast \ \the reverse inclusion is proved automatically\ qed lemma injectiveImageEqual: @@ -674,7 +686,7 @@ lemma injectionsE: using assms unfolding Injections_def InjectiveOn_def by blast -subsubsection {* Surjective functions *} +subsubsection \ Surjective functions \ definition Surjective where "Surjective(f,A) \ A \ Range(f)" @@ -683,8 +695,8 @@ definition Surjections where "Surjections(S,T) \ { f \ [S \ T] : Surjective(f,T) }" lemmas - setEqualI [where A = "Surjections(S,T)", standard, intro!] - setEqualI [where B = "Surjections(S,T)", standard, intro!] + setEqualI [where A = "Surjections(S,T)" for S T, intro!] + setEqualI [where B = "Surjections(S,T)" for S T, intro!] lemma surjectiveIsBool [intro!,simp]: "isBool(Surjective(f,A))" @@ -704,10 +716,10 @@ lemma surjectiveD: shows "\x \ DOMAIN f : y = f[x]" using assms by (auto simp: Surjective_def Image_def) -text {* +text \ @{text "\ Surjective(f,A); y \ A; \x. \ x \ DOMAIN f; y = f[x] \ \ P \ \ P"} -*} -lemmas surjectiveE = surjectiveD[THEN bExE, standard] +\ +lemmas surjectiveE = surjectiveD[THEN bExE] lemma surjectiveRange: shows "Surjective(f, Range(f))" @@ -749,20 +761,20 @@ lemma surjectionsRange: using assms by (rule surjectionsE, auto) -subsubsection {* Bijective functions *} +subsubsection \ Bijective functions \ -text {* +text \ Here we do not define a predicate @{text Bijective} because it would require a set parameter for the codomain and would therefore be curiously asymmetrical. -*} +\ definition Bijections where "Bijections(S,T) \ Injections(S,T) \ Surjections(S,T)" lemmas - setEqualI [where A = "Bijections(S,T)", standard, intro!] - setEqualI [where B = "Bijections(S,T)", standard, intro!] + setEqualI [where A = "Bijections(S,T)" for S T, intro!] + setEqualI [where B = "Bijections(S,T)" for S T, intro!] lemma bijectionsI [intro!]: assumes "f \ [S \ T]" and "Injective(f)" and "Surjective(f,T)" @@ -786,7 +798,7 @@ lemma bijectionsE: using 1 by (intro 2, auto simp: Bijections_def Injections_def Surjections_def) -subsubsection {* Inverse of a function *} +subsubsection \ Inverse of a function \ definition Inverse where "Inverse(f) \ [ y \ Range(f) \ CHOOSE x \ DOMAIN f : f[x] = y ]" @@ -842,9 +854,9 @@ proof (intro injectiveOnI, auto) ultimately show "f[x] = f[x']" by simp qed -text {* +text \ For injective functions, @{text Inverse} really inverts the function. -*} +\ lemma injectiveInverse: assumes f: "Injective(f)" and x: "x \ DOMAIN f" @@ -879,9 +891,7 @@ proof (rule surjectiveI) with x show "\y \ DOMAIN Inverse(f) : x = Inverse(f)[y]" by auto qed -text {* - The inverse of a bijection is a bijection. -*} +text \ The inverse of a bijection is a bijection. \ lemma inverseBijections: assumes f: "f \ Bijections(S,T)" diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f8ddcd09..f558c286 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -1,46 +1,41 @@ (* Title: TLA+/NatOrderings.thy Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:28 merz> + Version: Isabelle2020 *) -header {* Orders on natural numbers *} +section \ Orders on natural numbers \ theory NatOrderings imports Peano begin -text {* +text \ Using the sets @{text upto} we can now define the standard ordering on natural numbers. The constant @{text "\"} is defined over the naturals by the axiom (conditional definition) @{text "nat_leq_def"} below; it should be defined over other domains as appropriate later on. We generally define the constant @{text "<"} such that @{text "ab \ a\b"}, over any domain. -*} + iff @{text "a\b \ a\b"}, for any values. +\ -definition leq :: "[c,c] \ c" (infixl "<=" 50) -(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m <= n) \ (m \ upto(n))"*) -where nat_leq_def: "(m <= n) \ (m \ upto(n))" +definition leq :: "[c,c] \ c" (infixl "\" 50) +(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m \ n) \ (m \ upto(n))"*) +where nat_leq_def: "(m \ n) \ (m \ upto(n))" abbreviation (input) - geq :: "[c,c] \ c" (infixl ">=" 50) -where "x >= y \ y <= x" + geq :: "[c,c] \ c" (infixl "\" 50) +where "x \ y \ y \ x" -notation (xsymbols) - leq (infixl "\" 50) and - geq (infixl "\" 50) +notation (ASCII) + leq (infixl "<=" 50) and + geq (infixl ">=" 50) -notation (HTML output) - leq (infixl "\" 50) and - geq (infixl "\" 50) - -subsection {* Operator definitions and generic facts about @{text "<"} *} +subsection \ Operator definitions and generic facts about @{text "<"} \ definition less :: "[c,c] \ c" (infixl "<" 50) where "a < b \ a \ b \ a \ b" @@ -81,7 +76,7 @@ declare leq_neq_trans[simplified,trans] lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" by auto -subsection {* Facts about @{text "\"} over @{text Nat} *} +subsection \ Facts about @{text "\"} over @{text Nat} \ lemma nat_boolify_leq [simp]: "boolify(m \ n) = (m \ n)" by (simp add: nat_leq_def) @@ -92,7 +87,7 @@ by (simp add: nat_leq_def) lemma nat_leq_refl [intro,simp]: "n \ Nat \ n \ n" unfolding nat_leq_def by (rule uptoRefl) -lemma eq_leq_bothE: -- {* reduce equality over integers to double inequality *} +lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ assumes "m \ Nat" and "n \ Nat" and "m = n" and "\ m \ n; n \ m \ \ P" shows "P" using assms by simp @@ -103,55 +98,64 @@ unfolding nat_leq_def by (rule zeroInUpto) lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" by (simp add: nat_leq_def uptoZero) -lemma nat_leq_SuccI [elim!,simp]: (** FIXME: why simp ? **) - assumes "m \ n" and "m \ Nat" and "n \ Nat" +lemma nat_leq_SuccI [elim!(*,simp*)]: + assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" shows "m \ Succ[n]" using assms by (auto simp: nat_leq_def uptoSucc) -lemma nat_leq_Succ: (** FIXME: make this simp ? **) - assumes "m \ Nat" and "n \ Nat" +lemma nat_leq_Succ: + assumes (*"m \ Nat" and*) "n \ Nat" shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" using assms by (auto simp: nat_leq_def uptoSucc) lemma nat_leq_SuccE [elim]: - assumes "m \ Succ[n]" and "m \ Nat" and "n \ Nat" + assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" and "m \ n \ P" and "m = Succ[n] \ P" shows "P" using assms by (auto simp: nat_leq_Succ) lemma nat_leq_limit: - assumes "m \ n" and "\(Succ[m] \ n)" and "m \ Nat" and "n \ Nat" + assumes "m \ n" and "\(Succ[m] \ n)" and (*"m \ Nat" and*) "n \ Nat" shows "m=n" using assms by (auto simp: nat_leq_def intro: uptoLimit) lemma nat_leq_trans [trans]: - assumes "k \ m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" + assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" shows "k \ n" using assms by (auto simp: nat_leq_def elim: uptoTrans) lemma nat_leq_antisym: - assumes "m \ n" and "n \ m" and "m \ Nat" and "n \ Nat" + assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) shows "m = n" using assms by (auto simp add: nat_leq_def elim: uptoAntisym) lemma nat_Succ_not_leq_self [simp]: assumes n: "n \ Nat" shows "(Succ[n] \ n) = FALSE" -using n by (auto dest: nat_leq_antisym) +using n by (auto simp: nat_leq_Succ dest: nat_leq_antisym) -lemma nat_Succ_leqD: +lemma nat_Succ_leqI: + assumes "m \ n" and "m \ Nat" and "n \ Nat" and "m \ n" + shows "Succ[m] \ n" +using assms by (auto dest: nat_leq_limit) + +lemma nat_Succ_leqD [elim]: assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" shows "m \ n" proof - - from m have "m \ Succ[m]" by simp + from m have "m \ Succ[m]" by (simp add: nat_leq_Succ) with leq m n show ?thesis by (elim nat_leq_trans, auto) qed -lemma nat_Succ_leq_Succ: - (* needn't be added to simp: consequence of nat_Succ_leq_iff_less and nat_less_Succ_iff_leq *) +lemma nat_Succ_leq_Succ [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(Succ[m] \ Succ[n]) = (m \ n)" -using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit elim: nat_Succ_leqD) +using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit) + +lemma nat_Succ_leq: + assumes "m \ Nat" and "n \ Nat" + shows "(Succ[m] \ n) = (m \ n \ m \ n)" + using assms by (auto intro: nat_Succ_leqI) lemma nat_leq_linear: "\m \ Nat; n \ Nat\ \ m \ n \ n \ m" unfolding nat_leq_def using uptoLinear . @@ -176,13 +180,13 @@ next qed qed -lemma nat_leq_induct: -- {* sometimes called ``complete induction'' *} +lemma nat_leq_induct: \ \sometimes called ``complete induction''\ assumes "P(0)" and "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(Succ[n])" shows "\n\Nat : P(n)" proof - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp: nat_leq_Succ) + by (intro natInduct, auto simp add: nat_leq_Succ) thus ?thesis by (blast dest: nat_leq_refl) qed @@ -193,39 +197,51 @@ lemma nat_leq_inductE: using assms by (blast dest: nat_leq_induct) -subsection {* Facts about @{text "<"} over @{text Nat} *} +subsection \ Facts about @{text "<"} over @{text Nat} \ -lemma nat_Succ_leq_iff_less [simp]: +lemma nat_less_iff_Succ_leq [simp]: assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] \ n) = (m < n)" + shows "(m < n) = (Succ[m] \ n)" using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) --- {* alternative definition of @{text "<"} over @{text Nat} *} -lemmas nat_less_iff_Succ_leq = sym[OF nat_Succ_leq_iff_less, standard] - -text {* Reduce @{text "\"} to @{text "<"}. *} +text \ Reduce @{text "\"} to @{text "<"}. \ -lemma nat_leq_less: -- {* premises needed for @{text "isBool(m\n)"} and reflexivity *} - assumes "m \ Nat" and "n \ Nat" +lemma nat_leq_less: + assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" -using assms by (auto simp: less_def) + using assms by (auto simp: less_def) -lemma nat_less_Succ_iff_leq [simp]: +lemma nat_leqE: + assumes "m \ n" and "n \ Nat" and 1: "m P" and 2: "m=n \ P" + shows "P" +proof - + from \m \ n\ \n \ Nat\ have "m < n \ m=n" + by (auto simp add: less_def) + with 1 2 show ?thesis by blast +qed + +lemma nat_less_Succ_iff_leq (*[simp]*): assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m \ n)" -using assms -by (simp del: nat_Succ_leq_iff_less add: nat_less_iff_Succ_leq nat_Succ_leq_Succ) + using assms by simp + +lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq] -lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq, standard] +lemma not_leq_Succ_leq [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "((m \ n) = FALSE) = (Succ[n] \ m)" +proof - + from assms nat_leq_linear[of m n] have "((m \ n) = FALSE) = (n < m)" + by (auto simp: less_def dest: nat_leq_antisym) + with assms show ?thesis by simp +qed lemma nat_not_leq_one: assumes "n \ Nat" shows "(\(1 \ n)) = (n = 0)" -using assms by (cases, auto) - -declare nat_not_leq_one[simplified,simp] +using assms by simp -text {* @{text "<"} and @{text "Succ"}. *} +text \ @{text "<"} and @{text "Succ"}. \ lemma nat_Succ_less_mono: assumes "m \ Nat" and "n \ Nat" @@ -240,7 +256,7 @@ using assms by simp lemma nat_not_less0 [simp]: assumes "n \ Nat" shows "(n < 0) = FALSE" -using assms by (auto simp: less_def) +using assms by auto lemma nat_less0E (*[elim!]*): assumes "n < 0" and "n \ Nat" @@ -255,14 +271,14 @@ using assms by auto lemma nat_Succ_lessD: assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" shows "m < n" -using 1[unfolded less_def] 2 3 by simp +using assms by auto lemma nat_less_leq_not_leq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m < n) = (m \ n \ \ n \ m)" using assms by (auto simp: less_def dest: nat_leq_antisym) -text {* Transitivity. *} +text \ Transitivity. \ lemma nat_less_trans (*[trans]*): assumes "k < m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" @@ -289,21 +305,22 @@ lemma nat_less_leq_trans [trans]: shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) -text {* Asymmetry. *} +text \ Asymmetry. \ lemma nat_less_not_sym: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "(n < m) = FALSE" -using assms by (simp add: nat_less_leq_not_leq) - +using assms by auto + lemma nat_less_asym: assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" shows "P" proof (rule contradiction) - assume "\P" with assms show "FALSE" by (auto dest: nat_less_not_sym) + assume "\P" with nat_less_not_sym[OF \m \m \ Nat\ \n \ Nat\] \\P \ n < m\ + show "FALSE" by simp qed -text {* Linearity (totality). *} +text \ Linearity (totality). \ lemma nat_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" @@ -338,9 +355,9 @@ unfolding nat_not_less[OF m n] using assms by simp lemma nat_not_leq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m \ n) = (n < m)" -using assms by (simp add: sym[OF nat_not_less]) +using assms by simp --- {* often useful, but not active by default *} +text \often useful, but not active by default\ lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq lemma nat_not_leq_eq: @@ -366,7 +383,7 @@ using assms by (auto simp: nat_leq_less) lemma nat_antisym_conv2: assumes "m \ n" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_antisym_conv1) +using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) lemma nat_antisym_conv3: assumes "\ n < m" and "m \ Nat" and "n \ Nat" @@ -388,9 +405,9 @@ lemma nat_gt0_not0 (*[simp]*): shows "(0 < n) = (n \ 0)" using assms by (auto simp: nat_neq_iff[simplified]) -lemmas nat_neq0_conv = sym[OF nat_gt0_not0, standard] +lemmas nat_neq0_conv = sym[OF nat_gt0_not0] -text {* Introduction properties *} +text \ Introduction properties \ lemma nat_less_Succ_self (*[iff]*): assumes "n \ Nat" @@ -402,37 +419,37 @@ lemma nat_zero_less_Succ (*[iff]*): shows "0 < Succ[n]" using assms by simp -text {* Elimination properties. *} +text \ Elimination properties \ lemma nat_less_Succ: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m < n \ m = n)" -using assms by (simp add: nat_leq_less) +using assms by (simp add: nat_leq_less[of n m]) lemma nat_less_SuccE: assumes "m < Succ[n]" and "m \ Nat" and "n \ Nat" and "m < n \ P" and "m = n \ P" shows P -using assms by (auto simp: nat_leq_less) +using assms by (auto simp: nat_leq_less[of n m]) lemma nat_less_one (*[simp]*): assumes "n \ Nat" shows "(n < 1) = (n = 0)" using assms by simp -text {* "Less than" is antisymmetric, sort of. *} +text \ "Less than" is antisymmetric, sort of \ lemma nat_less_antisym: assumes m: "m \ Nat" and n: "n \ Nat" shows "\ \(n < m); n < Succ[m] \ \ m = n" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -text {* Lifting @{text "<"} monotonicity to @{text "\"} monotonicity. *} +text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ lemma less_mono_imp_leq_mono: assumes i: "i \ Nat" and j: "j \ Nat" and f: "\n \ Nat : f(n) \ Nat" and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" shows "f(i) \ f(j)" -using ij proof (auto simp: nat_leq_less[OF i j]) +using ij proof (auto simp: nat_leq_less[OF j]) assume "i < j" with i j have "f(i) < f(j)" by (rule mono) thus "f(i) \ f(j)" by (simp add: less_imp_leq) @@ -440,7 +457,7 @@ next from j f show "f(j) \ f(j)" by auto qed -text {* Inductive (?) properties. *} +text \ Inductive (?) properties. \ lemma nat_Succ_lessI: assumes "m \ Nat" and "n \ Nat" and "m < n" and "Succ[m] \ n" @@ -458,9 +475,8 @@ proof - fix n assume n: "n \ Nat" and 1: "i < Succ[n]" and ih: "i < n \ \j\Nat : i \ j \ n = Succ[j]" - from i n 1 have "i < n \ i = n" by (simp add: nat_leq_less) - thus "\j\Nat : i \ j \ Succ[n] = Succ[j]" - proof + from 1 i n show "\j\Nat : i \ j \ Succ[n] = Succ[j]" + proof (rule nat_less_SuccE) assume "i < n" then obtain j where "j \ Nat" and "i \ j" and "n = Succ[j]" by (blast dest: ih) @@ -478,31 +494,48 @@ qed lemma nat_Succ_lessE: assumes major: "Succ[i] < k" and i: "i \ Nat" and k: "k \ Nat" obtains j where "j \ Nat" and "i < j" and "k = Succ[j]" -using assms by (auto elim: nat_lessE) +proof - + from i have "Succ[i] \ Nat" .. + from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" + by (rule nat_lessE) + with i that show ?thesis by simp +qed +(* lemma nat_gt0_implies_Succ: assumes 1: "0 < n" and 2: "n \ Nat" shows "\m : m \ Nat \ n = Succ[m]" using 2 1 by (cases, auto) +*) +lemma nat_ge1_implies_Succ: + assumes 1: "1 \ n" and 2: "n \ Nat" + shows "\m : m \ Nat \ n = Succ[m]" + using 2 1 by (cases, auto) lemma nat_gt0_iff_Succ: assumes n: "n \ Nat" shows "(0 < n) = (\m \ Nat: n = Succ[m])" -using n by (auto dest: nat_gt0_implies_Succ) +using n by (auto dest: nat_ge1_implies_Succ) +(* lemma nat_less_Succ_eq_0_disj: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" using assms by (induct m, auto) +*) +lemma nat_le_eq_0_disj: + assumes "m \ Nat" and "n \ Nat" + shows "(m \ n) = (m = 0 \ (\j\Nat : m = Succ[j] \ j < n))" + using assms by (induct m, auto) lemma nat_less_antisym_false: "\m < n; m \ Nat; n \ Nat\ \ n < m = FALSE" - unfolding less_def using nat_leq_antisym by auto + by auto lemma nat_less_antisym_leq_false: "\m < n; m \ Nat; n \ Nat\ \ n \ m = FALSE" - unfolding less_def using nat_leq_antisym[of m n] by auto + by auto -subsection {* Intervals of natural numbers *} +subsection \Intervals of natural numbers\ definition natInterval :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) where "m .. n \ { k \ Nat : m \ k \ k \ n }" @@ -518,11 +551,11 @@ lemma inNatIntervalE [elim]: using 1 by (intro 2, auto simp add: natInterval_def) lemma inNatInterval_iff: "(k \ m .. n) = (k \ Nat \ m \ k \ k \ n)" -using assms by auto +by auto lemmas - setEqualI [where A = "m .. n", standard, intro] - setEqualI [where B = "m .. n", standard, intro] + setEqualI [where A = "m .. n" for m n, intro] + setEqualI [where B = "m .. n" for m n, intro] lemma lowerInNatInterval [iff]: assumes "m \ n" and "m \ Nat" @@ -538,8 +571,8 @@ lemma gtNotinNatInterval: assumes gt: "m > n" and k: "k \ m .. n" and m: "m \ Nat" and n: "n \ Nat" shows "P" proof - - from k have 1: "m \ k" and 2: "k \ n" and 3: "k \ Nat" by auto - from 1 2 m 3 n have "m \ n" by (rule nat_leq_trans) + from k have 1: "m \ k" and 2: "k \ n" by auto + from 1 2 n have "m \ n" by (rule nat_leq_trans) with m n have "\(n < m)" by (simp add: nat_not_order_simps) from this gt show ?thesis .. qed @@ -576,7 +609,7 @@ using assms by (auto simp: natInterval_def) lemma succNatInterval: assumes "m \ Nat" and "n \ Nat" shows "Succ[m] .. n = (m .. n \ {m})" -using assms by (auto simp: natInterval_def) +using assms by (auto simp: natInterval_def nat_Succ_leq) lemma natIntervalEqual_iff: assumes k: "k \ Nat" and l: "l \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -599,10 +632,10 @@ proof - from False eq m n have 12: "m \ n" by (simp only: natIntervalEmpty_iff nat_not_less) from 11 k eq have 13: "m \ k" by auto from 12 m eq have 14: "k \ m" by auto - from 14 13 k m have 15: "k = m" by (rule nat_leq_antisym) + from 14 13 have 15: "k = m" by (rule nat_leq_antisym) from 11 l eq have 16: "l \ n" by auto from 12 n eq have 17: "n \ l" by auto - from 16 17 l n have "l = n" by (rule nat_leq_antisym) + from 16 17 have "l = n" by (rule nat_leq_antisym) with 15 show ?rhs by blast qed qed diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index ba06d431..6b848505 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -1,30 +1,29 @@ (* Title: TLA+/Peano.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:03 merz> + Version: Isabelle2020 *) -header {* Peano's axioms and natural numbers *} +section \ Peano's axioms and natural numbers \ theory Peano imports FixedPoints Functions begin -text {* +text \ As a preparation for the definition of numbers and arithmetic in \tlaplus{}, we state Peano's axioms for natural numbers and prove the existence of a structure satisfying them. The presentation of the axioms is somewhat simplified compared to the \tlaplus{} book. (Moreover, the existence of such a structure is assumed, but not proven in the book.) -*} +\ -subsection {* The Peano Axioms *} +subsection \ The Peano Axioms \ definition PeanoAxioms :: "[c,c,c] \ c" where - -- {* parameters: the set of natural numbers, zero, and succ function *} + \ \parameters: the set of natural numbers, zero, and successor function\ "PeanoAxioms(N,Z,Sc) \ Z \ N \ Sc \ [N \ N] @@ -32,7 +31,7 @@ definition PeanoAxioms :: "[c,c,c] \ c" where \ (\m,n \ N: Sc[m] = Sc[n] \ m = n) \ (\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S)" -text {* +text \ The existence of a structure satisfying Peano's axioms is proven following the standard ZF construction where @{text "{}"} is zero, @{text "i \ {i}"} is taken as the successor of any natural @@ -42,15 +41,15 @@ text {* set theory). In \tlaplus{}, natural numbers are defined by a sequence of @{text CHOOSE}'s below, so there is no commitment to that particular structure. -*} +\ theorem peanoExists: "\N,Z,Sc : PeanoAxioms(N,Z,Sc)" proof - - let ?sc = "\n. addElt(n,n)" -- {* successor \emph{operator} *} - def expand \ "\S. {{}} \ { ?sc(n) : n \ S}" - def N \ "lfp(infinity, expand)" - def Z \ "{}" - def Sc \ "[n \ N \ ?sc(n)]" -- {* successor \emph{function} *} + let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ + define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" + define N where "N \ lfp(infinity, expand)" + define Z where "Z \ {}" + define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ have mono: "Monotonic(infinity, expand)" using infinity by (auto simp: Monotonic_def expand_def) hence expandN: "expand(N) \ N" @@ -139,12 +138,12 @@ proof - qed -subsection {* Natural numbers: definition and elementary theorems *} +subsection \ Natural numbers: definition and elementary theorems \ -text {* +text \ The structure of natural numbers is now defined to be some set, zero, and successor satisfying Peano's axioms. -*} +\ definition Succ :: "c" where "Succ \ CHOOSE Sc : \N,Z : PeanoAxioms(N,Z,Sc)" @@ -203,8 +202,8 @@ proof - qed lemmas - setEqualI [where A = "Nat", standard, intro!] - setEqualI [where B = "Nat", standard, intro!] + setEqualI [where A = "Nat", intro!] + setEqualI [where B = "Nat", intro!] lemma zeroIsNat [intro!,simp]: "0 \ Nat" using peanoNatZeroSucc by (simp add: PeanoAxioms_def) @@ -215,9 +214,9 @@ using peanoNatZeroSucc by (simp add: PeanoAxioms_def) lemma succIsAFcn [intro!,simp]: "isAFcn(Succ)" using succInNatNat by blast --- {* @{text "DOMAIN Succ = Nat"} *} +\ \@{text "DOMAIN Succ = Nat"}\ lemmas domainSucc [intro!,simp] = funcSetDomain[OF succInNatNat] --- {* @{text "n \ Nat \ Succ[n] \ Nat"} *} +\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ lemmas succIsNat [intro!,simp] = funcSetValue[OF succInNatNat] lemma oneIsNat [intro!,simp]: "1 \ Nat" @@ -236,7 +235,7 @@ lemma [simp]: shows "(0 = Succ[n]) = FALSE" using assms by (auto dest: sym) -lemma succNotZero (*[elim] -- yields "ignoring weak elimination rule" *): +lemma succNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): "\Succ[n] = 0; n \ Nat\ \ P" "\0 = Succ[n]; n \ Nat\ \ P" by (simp+) @@ -262,7 +261,7 @@ proof - thus ?thesis by auto qed --- {* version of above suitable for the inductive reasoning package *} +\ \version of above suitable for the inductive reasoning package\ lemma natInductE [case_names 0 Succ, induct set: Nat]: assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(Succ[n])" shows "P(n)" @@ -313,7 +312,7 @@ lemma succIrrefl_iff [simp]: by (auto dest: succIrrefl) --- {* Induction over two parameters along the ``diagonal''. *} +text \Induction over two parameters along the ``diagonal''.\ lemma diffInduction: assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, Succ[n])" and step: "\m,n\Nat : P(m,n) \ P(Succ[m], Succ[n])" @@ -351,23 +350,23 @@ lemma not0_implies_Suc: "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" by(rule natCases, auto) -subsection {* Initial intervals of natural numbers and ``less than'' *} +subsection \ Initial intervals of natural numbers and ``less than'' \ -text {* +text \ The set of natural numbers up to (and including) a given $n$ is inductively defined as the smallest set of natural numbers that contains $n$ and that is closed under predecessor. NB: ``less than'' is not first-order definable from the Peano axioms, a set-theoretic definition such as the following seems to be unavoidable. -*} +\ definition upto :: "c \ c" where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : Succ[k] \ S })" lemmas - setEqualI [where A = "upto(n)", standard, intro!] - setEqualI [where B = "upto(n)", standard, intro!] + setEqualI [where A = "upto(n)" for n, intro!] + setEqualI [where B = "upto(n)" for n, intro!] lemma uptoNat: "upto(n) \ Nat" unfolding upto_def by (rule lfpSubsetDomain) @@ -399,20 +398,20 @@ lemma uptoSucc: assumes n: "n \ Nat" shows "upto(Succ[n]) = upto(n) \ {Succ[n]}" (is "?lhs = ?rhs") proof - - let "?preds(S)" = "{k \ Nat : Succ[k] \ S}" - let "?f(S,k)" = "{k} \ ?preds(S)" + let ?preds = "\S. {k \ Nat : Succ[k] \ S}" + let ?f = "\S k. {k} \ ?preds(S)" have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" by (auto simp: Monotonic_def) - -- {* ``$\subseteq$'' *} + \ \``$\subseteq$''\ from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto also have "\ \ upto(n)" by (unfold upto_def, rule lfpPreFP, rule mono, rule n) finally have "?f(?rhs, Succ[n]) \ ?rhs" by auto moreover from n have "?rhs \ Nat" - by (intro cupLUB, auto elim: uptoNat[THEN subsetD]) + by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) ultimately have 1: "?lhs \ ?rhs" by (unfold upto_def[where n="Succ[n]"], rule lfpLB) - -- {* ``$\supseteq$'' *} + \ \``$\supseteq$''\ from n mono have 2: "?f(?lhs, Succ[n]) \ ?lhs" unfolding upto_def by (intro lfpPreFP, blast) with n have "?f(?lhs, n) \ ?lhs" by auto @@ -496,8 +495,8 @@ proof - fix n assume "Succ[m] \ upto(n)" and "n \ upto(m)" from this m have "Succ[m] \ upto(m)" by (rule uptoTrans) - with m show "Succ[m] = n" -- {* contradiction *} - by (blast dest: succNotinUpto) + with m show "Succ[m] = n" \ \contradiction\ + by (blast dest: succNotinUpto) qed qed with m n mn nm show ?thesis by blast @@ -540,13 +539,13 @@ next qed -subsection {* Primitive Recursive Functions *} +subsection \ Primitive Recursive Functions \ -text {* +text \ We axiomatize a primitive recursive scheme for functions with one argument and domain on natural numbers. Later, we use it to define addition, multiplication and difference. -*} +\ (** FIXME: replace axiom with proper fixpoint definition **) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index a8e713dd..80d69b78 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -2,79 +2,82 @@ Author: Stephan Merz, LORIA Copyright (C) 2008-2011 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:37 merz> + Version: Isabelle2017 *) -header {* First-order logic for TLA+ *} +section \First-Order Logic for TLA+\ theory PredicateLogic -imports Pure -uses - "~~/src/Tools/misc_legacy.ML" - "~~/src/Tools/intuitionistic.ML" - "~~/src/Provers/splitter.ML" - "~~/src/Provers/hypsubst.ML" - "~~/src/Tools/atomize_elim.ML" - "~~/src/Provers/classical.ML" - "~~/src/Provers/blast.ML" - "~~/src/Provers/quantifier1.ML" - "~~/src/Provers/clasimp.ML" - "~~/src/Tools/IsaPlanner/isand.ML" "~~/src/Tools/IsaPlanner/rw_tools.ML" "~~/src/Tools/IsaPlanner/rw_inst.ML" "~~/src/Tools/IsaPlanner/zipper.ML" "~~/src/Tools/eqsubst.ML" - "~~/src/Tools/induct.ML" - ("simplifier_setup.ML") + imports Pure + keywords + "print_claset" "print_induct_rules" :: diag begin +ML_file "~~/src/Tools/misc_legacy.ML" +ML_file "~~/src/Tools/intuitionistic.ML" +ML_file "~~/src/Tools/IsaPlanner/isand.ML" +ML_file "~~/src/Tools/IsaPlanner/rw_inst.ML" +ML_file "~~/src/Tools/IsaPlanner/zipper.ML" +ML_file "~~/src/Provers/splitter.ML" +ML_file "~~/src/Provers/hypsubst.ML" +ML_file "~~/src/Provers/classical.ML" +ML_file "~~/src/Provers/blast.ML" +ML_file "~~/src/Tools/atomize_elim.ML" +ML_file "~~/src/Provers/quantifier1.ML" +ML_file "~~/src/Provers/clasimp.ML" +ML_file "~~/src/Tools/eqsubst.ML" +ML_file "~~/src/Tools/induct.ML" declare [[ eta_contract = false ]] (* eta contraction doesn't make much sense for first-order logics *) -text {* +text \ We define classical first-order logic as a basis for an encoding of \tlaplus{}. \tlaplus{} is untyped, to the extent that it does not even distinguish between terms and formulas. We therefore - declare a single type $c$ that represents the universe of - ``constants'' rather than introducing the traditional types $i$ and - $o$ of first-order logic that, for example, underly Isabelle/ZF. -*} + declare a single type @{text \c\} that represents the universe of + ``constants'' rather than introducing the types @{text \i\} + and @{text \o\} representing non-Boolean and Boolean values that + are traditional in first-order logic. +\ (* First-order syntax for applications: f(x,y) instead of (f x y) *) setup Pure_Thy.old_appl_syntax_setup -setup {* Intuitionistic.method_setup @{binding iprover} *} +setup \ Intuitionistic.method_setup @{binding iprover} \ typedecl c -text {* +text \ The following (implicit) lifting from the object to the Isabelle meta level is always needed when formalizing a logic. It corresponds - to judgments $\vdash F$ or $\models F$ in standard presentations, + to judgments @{text \\ F\} or @{text \\ F\} in textbook presentations, asserting that a formula is considered true (either because it is an assumption or because it is a theorem). -*} +\ judgment Trueprop :: "c \ prop" ("(_)" 5) -subsection {* Equality *} +subsection \ Equality \ -text {* +text \ The axioms for equality are reflexivity and a rule that asserts that equal terms are interchangeable at the meta level (this is essential for setting up Isabelle's rewriting machinery). In particular, we can derive a substitution rule. -*} +\ axiomatization "eq" :: "c \ c \ c" (infixl "=" 50) where - refl [intro!]: "a = a" + refl [Pure.intro!]: "a = a" and eq_reflection: "t = u \ t \ u" -text {* +text \ Left and right hand sides of definitions are equal. This is the converse of axiom @{thm eq_reflection}. -*} +\ theorem meta_to_obj_eq: assumes df: "t \ u" @@ -97,19 +100,20 @@ proof (elim subst) show "a=a" .. qed +text \@{term "b=a \ P(a) \ P(b)"}\ +lemmas ssubst = sym[THEN subst] + theorem trans [trans]: "a=b \ b=c \ a=c" by (rule subst) -theorems ssubst = sym[THEN subst] - -text {* +text \ \textsc{let}-expressions in \tlaplus{} expressions. \emph{Limitation:} bindings cannot be unfolded selectively. Rewrite with @{text Let_def} in order to expand \emph{all} bindings within an expression or a context. -*} +\ definition Let :: "'b \ ('b \ 'a) \ 'a" @@ -135,95 +139,90 @@ translations "LET x \ a IN e" \ "CONST Let(a, (%x. e))" -subsection {* Propositional logic *} +subsection \ Propositional logic \ -text {* +text \ Propositional logic is introduced in a rather non-standard way by declaring constants @{text TRUE} and @{text FALSE} as well as - conditional expressions. The Boolean connectives - are defined such that they always return either @{text TRUE} or - @{text FALSE}, irrespectively of their arguments. This allows us - to prove many equational laws of propositional logic, which is - useful for automatic reasoning based on rewriting. Note that we - have equivalence as well as equality. The two relations agree over + conditional expressions. The ordinary Boolean connectives + are defined in terms of these basic operators in such a way + that they always return either @{text TRUE} or @{text FALSE}, + irrespectively of their arguments. We also require that equality + returns a Boolean value. In this way, we can prove standard + equational laws of propositional logic, which is useful for + automatic reasoning based on rewriting. Note that we have + equivalence as well as equality. The two relations agree over Boolean values, but equivalence may be stricter weaker than equality over non-Booleans. -*} - -consts - TRUE :: "c" - FALSE :: "c" - cond :: "c \ c \ c \ c" ("(IF (_)/ THEN (_)/ ELSE (_))" 10) -consts - Not :: "c \ c" ("~ _" [40] 40) - conj :: "c \ c \ c" (infixr "&" 35) - disj :: "c \ c \ c" (infixr "|" 30) - imp :: "c \ c \ c" (infixr "=>" 25) - iff :: "c \ c \ c" (infixr "<=>" 25) - -abbreviation not_equal :: "c \ c \ c" (infixl "~=" 50) -where "x ~= y \ ~ (x = y)" - -notation (xsymbols) - Not ("\ _" [40] 40) -and conj (infixr "\" 35) -and disj (infixr "\" 30) -and imp (infixr "\" 25) -and iff (infixr "\" 25) -and not_equal (infix "\" 50) - -notation (HTML output) - Not ("\ _" [40] 40) -and conj (infixr "\" 35) -and disj (infixr "\" 30) -and imp (infixr "\" 25) -and iff (infixr "\" 25) -and not_equal (infix "\" 50) - -defs - not_def: "\ A \ IF A THEN FALSE ELSE TRUE" - conj_def: "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE FALSE" - disj_def: "A \ B \ IF A THEN TRUE ELSE IF B THEN TRUE ELSE FALSE" - imp_def: "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE TRUE" - iff_def: "A \ B \ (A \ B) \ (B \ A)" - -text {* - We adopt the following axioms of propositional logic: + The following axioms govern propositional logic: \begin{enumerate} - \item $A$ is a theorem if and only if it equals @{term TRUE}. - \item @{term FALSE} (more precisely, @{text "\ TRUE"}) implies anything. - \item Conditionals are reasoned about by case distinction. + \item @{term TRUE} is valid, and @{term FALSE} implies anything}, + \item @{text A} is a theorem if and only if it equals @{term TRUE}, + \item conditionals are reasoned about by case distinction, + \item equations evaluate to @{term TRUE} or @{term FALSE}. \end{enumerate} - We also assert that the equality predicate returns either @{term TRUE} - or @{term FALSE}. -*} +\ -axiomatization where - trueI [intro!]: "TRUE" +axiomatization + TRUE :: "c" and - eqTrueI: "A \ A = TRUE" + FALSE :: "c" and - notTrueE [elim!]: "\ TRUE \ A" + cond :: "c \ c \ c \ c" ("(IF (_)/ THEN (_)/ ELSE (_))" 10) +where + trueI [Pure.intro!]: "TRUE" +and + eqTrueI: "A \ A = TRUE" +and + falseE [Pure.elim!]: "FALSE \ A" and - condI: "\ A \ P(t); \ A \ P(e) \ \ P(IF A THEN t ELSE e)" + condI: "\ A \ P(t); (A \ FALSE) \ P(e) \ \ P(IF A THEN t ELSE e)" and - eqBoolean : "x \ y \ (x=y) = FALSE" + eqBoolean : "(x = y \ FALSE) \ (x=y) = FALSE" + + +definition not :: "c \ c" ("\ _" [40] 40) + where "\ A \ IF A THEN FALSE ELSE TRUE" + +definition conj :: "c \ c \ c" (infixr "\" 35) + where "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE FALSE" -text {* - We now derive the standard proof rules of propositional logic. +definition disj :: "c \ c \ c" (infixr "\" 30) + where "A \ B \ IF A THEN TRUE ELSE IF B THEN TRUE ELSE FALSE" + +definition imp :: "c \ c \ c" (infixr "\" 25) + where "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE TRUE" + +definition iff :: "c \ c \ c" (infixr "\" 25) + where "A \ B \ (A \ B) \ (B \ A)" + +abbreviation not_equal :: "c \ c \ c" (infixl "\" 50) + where "x \ y \ \(x = y)" + +notation (ASCII) + not ("~ _" [40] 40) +and conj (infixr "&" 35) +and disj (infixr "|" 30) +and imp (infixr "=>" 25) +and iff (infixr "<=>" 25) +and not_equal (infix "~=" 50) + +text \ + We now derive the standard laws of propositional logic. The first lemmas are about @{term TRUE}, @{term FALSE}, and conditional expressions. -*} +\ -lemma eqTrueD: -- {* converse of eqTrueI *} +text \ Converse of eqTrueI. \ +lemma eqTrueD: assumes a: "A = TRUE" shows "A" by (unfold a[THEN eq_reflection], rule trueI) -text {* +text \ Assumption @{term TRUE} is useless and can be deleted. -*} +\ lemma TrueAssumption: "(TRUE \ PROP P) == PROP P" proof @@ -237,23 +236,16 @@ lemma condT: "(IF TRUE THEN t ELSE e) = t" proof (rule condI) show "t = t" .. next - assume "\ TRUE" thus "e = t" .. + assume "TRUE \ FALSE" + from this trueI have "FALSE" . + thus "e = t" .. qed lemma notTrue: "(\ TRUE) = FALSE" by (unfold not_def, rule condT) -theorem falseE [elim!]: - assumes f: "FALSE" - shows "A" -proof (rule notTrueE) - have "FALSE = (\ TRUE)" - by (rule sym[OF notTrue]) - hence r: "FALSE \ \ TRUE" - by (rule eq_reflection) - from f show "\ TRUE" - by (unfold r) -qed +lemma notTrueE: "\ TRUE \ A" + unfolding notTrue[THEN eq_reflection] by (rule falseE) lemma condF: "(IF FALSE THEN t ELSE e) = e" proof (rule condI) @@ -277,7 +269,7 @@ proof - by (unfold r, rule condT) qed -lemma condD1 [elim 2]: +lemma condD1 [Pure.elim 2]: assumes c: "IF A THEN P ELSE Q" (is ?if) and a: "A" shows "P" proof (rule eqTrueD) @@ -287,52 +279,71 @@ proof (rule eqTrueD) finally show "P = TRUE" by (rule sym) qed +theorem notI [Pure.intro!]: + assumes hyp: "A \ FALSE" + shows "\ A" +proof (unfold not_def, rule condI) + assume "A" thus "FALSE" by (rule hyp) +next + show "TRUE" .. +qed + +text \ + We do not have @{text "\ A \ A = FALSE"}: this + is true only for Booleans, not for arbitrary values. However, + we have the following theorem, which is just the ordinary + elimination rule for negation. +\ + +lemma notE: + assumes notA: "\ A" and a: "A" + shows "B" +proof - + from notA[unfolded not_def] a have "FALSE" by (rule condD1) + thus ?thesis .. +qed + lemma condElse: assumes na: "\ A" shows "(IF A THEN t ELSE e) = e" proof (rule condI) - assume "A" hence "A = TRUE" - by (rule eqTrueI) - hence r: "A \ TRUE" - by (rule eq_reflection) - from na have "\ TRUE" - by (unfold r) - thus "t = e" .. + assume "A" with na show "t = e" by (rule notE) next show "e = e" .. qed -lemma condD2 [elim 2]: +lemma condD2 [Pure.elim 2]: assumes c: "IF A THEN P ELSE Q" (is ?if) and a: "\ A" shows "Q" -proof (rule eqTrueD) - from c have "?if = TRUE" by (rule eqTrueI) - hence "TRUE = ?if" by (rule sym) - also from a have "?if = Q" by (rule condElse) - finally show "Q = TRUE" by (rule sym) +proof - + from a have "?if = Q" by (rule condElse) + hence "Q = ?if" by (rule sym) + also from c have "?if = TRUE" by (rule eqTrueI) + finally show ?thesis by (rule eqTrueD) qed -text {* - The following theorem shows that we have a classical logic. -*} +text \ + The following theorem confirms that we have a classical logic. +\ lemma cond_id: "(IF A THEN t ELSE t) = t" -proof (rule condI) - show "t=t" .. - show "t=t" .. -qed + by (rule condI) (rule refl)+ -theorem case_split: - assumes p: "P \ Q" - and np: "\ P \ Q" +theorem case_split [case_names True False]: + assumes p: "P \ Q" and np: "\ P \ Q" shows "Q" proof - - from p np have "IF P THEN Q ELSE Q" by (rule condI) + from p np have "IF P THEN Q ELSE Q" + proof (rule condI) + assume "P" thus "P" . + next + assume "P \ FALSE" thus "\P" .. + qed thus "Q" by (unfold eq_reflection[OF cond_id]) qed +text \ Use conditionals in hypotheses. \ theorem condE: - -- {* use conditionals in hypotheses *} assumes p: "P(IF A THEN t ELSE e)" and pos: "\ A; P(t) \ \ B" and neg: "\ \ A; P(e) \ \ B" @@ -341,87 +352,73 @@ proof (rule case_split[of "A"]) assume a: "A" hence r: "IF A THEN t ELSE e \ t" by (unfold eq_reflection[OF condThen]) - with p have "P(t)" - by (unfold r) - with a show "B" - by (rule pos) + with p have "P(t)" by (unfold r) + with a show "B" by (rule pos) next assume a: "\ A" hence r: "IF A THEN t ELSE e \ e" by (unfold eq_reflection[OF condElse]) - with p have "P(e)" - by (unfold r) - with a show "B" - by (rule neg) + with p have "P(e)" by (unfold r) + with a show "B" by (rule neg) qed -text {* +text \ Theorems @{text "condI"} and @{text "condE"} require higher-order unification and are therefore unsuitable for automatic tactics (in particular the \texttt{blast} method). We now derive some special cases that can be given to these methods. -*} +\ --- {* @{text "\A \ t; \A \ e\ \ IF A THEN t ELSE e"} *} -lemmas cond_boolI [intro!] = condI[where P="\ Q. Q"] +lemma cond_boolI [Pure.intro!]: + assumes a: "A \ P" and na: "\A \ Q" + shows "IF A THEN P ELSE Q" +proof (rule condI) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "Q" by (rule na) +qed (rule a) -lemma cond_eqLI [intro!]: +lemma cond_eqLI [Pure.intro!]: assumes 1: "A \ t = v" and 2: "\A \ u = v" shows "(IF A THEN t ELSE u) = v" proof (rule condI) - show "A \ t=v" by (rule 1) + show "A \ t = v" by (rule 1) next - show "\A \ u=v" by (rule 2) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "u = v" by (rule 2) qed -lemma cond_eqRI [intro!]: +lemma cond_eqRI [Pure.intro!]: assumes 1: "A \ v = t" and 2: "\A \ v = u" shows "v = (IF A THEN t ELSE u)" proof (rule condI) show "A \ v = t" by (rule 1) next - show "\A \ v = u" by (rule 2) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "v = u" by (rule 2) qed --- {* @{text "\IF A THEN t ELSE e; \A;t\ \ B; \\A;e\ \ B\ \ B"} *} +text \ @{text "\IF A THEN P ELSE Q; \A;P\ \ B; \\A;Q\ \ B\ \ B"} \ lemmas cond_boolE [elim!] = condE[where P="\ Q. Q"] -lemma cond_eqLE [elim!]: - assumes maj: "(IF A THEN t ELSE e) = u" - and 1: "\ A; t=u \ \ B" and 2: "\ \A; e=u \ \ B" +lemma cond_eqLE [Pure.elim!]: + assumes "(IF A THEN t ELSE e) = u" + and "\ A; t=u \ \ B" and "\ \A; e=u \ \ B" shows "B" -using maj -proof (rule condE) - show "\ A; t=u \ \ B" by (rule 1) -next - show "\ \A; e=u \ \ B" by (rule 2) -qed + using assms by (rule condE) -lemma cond_eqRE [elim!]: - assumes maj: "u = (IF A THEN t ELSE e)" - and 1: "\ A; u=t \ \ B" and 2: "\ \A; u=e \ \ B" +lemma cond_eqRE [Pure.elim!]: + assumes "u = (IF A THEN t ELSE e)" + and "\ A; u=t \ \ B" and "\ \A; u=e \ \ B" shows "B" -using maj -proof (rule condE) - show "\ A; u=t \ \ B" by (rule 1) -next - show "\ \A; u=e \ \ B" by (rule 2) -qed + using assms by (rule condE) -text {* +text \ Derive standard propositional proof rules, based on the operator definitions in terms of @{text "IF _ THEN _ ELSE _"}. -*} - -theorem notI [intro!]: - assumes hyp: "A \ FALSE" - shows "\ A" -proof (unfold not_def, rule condI) - assume "A" thus "FALSE" - by (rule hyp) -next - show "TRUE" .. -qed +\ lemma false_neq_true: "FALSE \ TRUE" proof @@ -432,10 +429,7 @@ qed lemma false_eq_trueE: assumes ft: "FALSE = TRUE" shows "B" -proof (rule falseE) - from ft show "FALSE" - by (rule eqTrueD) -qed + using false_neq_true ft by (rule notE) lemmas true_eq_falseE = sym[THEN false_eq_trueE] @@ -445,42 +439,22 @@ by iprover lemma A_then_notA_false: assumes a: "A" shows "(\ A) = FALSE" -using a -by (unfold not_def, rule condThen) + using a by (unfold not_def, rule condThen) -text {* +text \ The following is an alternative introduction rule for negation that is useful when we know that @{text A} is Boolean. -*} +\ lemma eq_false_not: assumes a: "A = FALSE" shows "\ A" -proof (rule eqTrueD) - show "(\ A) = TRUE" by (unfold eq_reflection[OF a], rule notFalse) -qed - -text {* - Note that we do not have @{text "\ A \ A = FALSE"}: this - is true only for Booleans, not for arbitrary values. However, - we have the following theorem, which is just the ordinary - elimination rule for negation. -*} - -theorem notE: - assumes notA: "\ A" and a: "A" - shows "B" -proof (rule false_eq_trueE) - from a have "(\ A) = FALSE" - by (rule A_then_notA_false) - hence "FALSE = (\ A)" - by (rule sym) - also from notA have "(\ A) = TRUE" - by (rule eqTrueI) - finally show "FALSE = TRUE" . +proof + assume "A" + thus "FALSE" by (unfold a[THEN eq_reflection]) qed -theorem notE' [elim 2]: +theorem notE' [Pure.elim 2]: assumes notA: "\ A" and r: "\ A \ A" shows "B" using notA @@ -505,18 +479,17 @@ proof with hyp show "FALSE" .. qed -text {* +text \ Some derived proof rules of classical logic. -*} +\ theorem contradiction: assumes hyp: "\ A \ FALSE" shows "A" proof (rule case_split[of "A"]) - assume "\ A" hence "FALSE" - by (rule hyp) + assume "\ A" hence "FALSE" by (rule hyp) thus "A" .. -qed -- {* the other case is trivial *} +qed text \ the other case is trivial \ theorem classical: assumes c: "\ A \ A" @@ -535,17 +508,13 @@ proof (rule contradiction) with a show "FALSE" .. qed -theorem notnotD [dest]: - assumes nn: "\\ A" shows "A" -proof (rule contradiction) - assume "\ A" - with nn show "FALSE" .. -qed +theorem notnotD [Pure.dest]: "\\A \ A" + by (rule swap) -text {* +text \ Note again: @{text A} and @{text "\\ A"} are inter-derivable - (and hence equivalent), but not equal! -*} + (and hence provably equivalent), but not equal! +\ lemma contrapos: assumes b: "\ B" and ab: "A \ B" @@ -569,14 +538,14 @@ proof - thus "A" .. qed -theorem conjI [intro!]: +theorem conjI [Pure.intro!]: assumes a: "A" and b: "B" shows "A \ B" -proof (rule eqTrueD) +proof - from a have "(A \ B) = (IF B THEN TRUE ELSE FALSE)" by (unfold conj_def, rule condThen) also from b have "\ = TRUE" by (rule condThen) - finally show "(A \ B) = TRUE" . + finally show ?thesis by (rule eqTrueD) qed theorem conjD1: @@ -584,8 +553,7 @@ theorem conjD1: shows "A" proof (rule contradiction) assume "\ A" - with ab show "FALSE" - by (unfold conj_def, elim condD2) + with ab show "FALSE" by (unfold conj_def, elim condD2) qed theorem conjD2: @@ -594,12 +562,11 @@ theorem conjD2: proof (rule contradiction) assume b: "\ B" from ab have "A" by (rule conjD1) - with ab have "IF B THEN TRUE ELSE FALSE" - by (unfold conj_def, elim condD1) + with ab have "IF B THEN TRUE ELSE FALSE" by (unfold conj_def, elim condD1) with b show "FALSE" by (elim condD2) qed -theorem conjE [elim!]: +theorem conjE [Pure.elim!]: assumes ab: "A \ B" and c: "A \ B \ C" shows "C" proof (rule c) @@ -607,9 +574,9 @@ proof (rule c) from ab show "B" by (rule conjD2) qed -text {* Disjunction *} +text \ Disjunction \ -theorem disjI1 [elim]: +theorem disjI1 [Pure.elim]: assumes a: "A" shows "A \ B" proof (unfold disj_def, rule) @@ -619,39 +586,28 @@ next from this a show "IF B THEN TRUE ELSE FALSE" .. qed -theorem disjI2 [elim]: +theorem disjI2 [Pure.elim]: assumes b: "B" shows "A \ B" proof (unfold disj_def, rule) show "TRUE" .. next - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE" .. - qed + from b show "IF B THEN TRUE ELSE FALSE" by iprover qed -theorem disjI [intro!]: -- {* classical introduction rule *} +text \ Classical introduction rule for disjunction. \ +theorem disjI [Pure.intro!]: assumes ab: "\ A \ B" shows "A \ B" proof (unfold disj_def, rule) show "TRUE" .. next assume "\ A" - hence b: "B" by (rule ab) - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE".. - qed + hence "B" by (rule ab) + thus "IF B THEN TRUE ELSE FALSE" by iprover qed -theorem disjE [elim!]: +theorem disjE [Pure.elim!]: assumes ab: "A \ B" and ac: "A \ C" and bc: "B \ C" shows "C" proof (rule case_split[where P=A]) @@ -669,30 +625,22 @@ next qed theorem excluded_middle: "A \ \ A" -proof - assume "\ A" thus "\ A" . -qed + by (rule disjI) -text {* Implication *} +text \ Implication \ -theorem impI [intro!]: +theorem impI [Pure.intro!]: assumes ab: "A \ B" shows "A \ B" proof (unfold imp_def, rule) assume "A" - hence b: "B" by (rule ab) - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE" .. - qed + hence "B" by (rule ab) + thus "IF B THEN TRUE ELSE FALSE" by iprover next show "TRUE" .. qed -theorem mp (*[elim 2]*): +theorem mp (*[Pure.elim 2]*): assumes ab: "A \ B" and a: "A" shows "B" proof (rule contradiction) @@ -702,12 +650,12 @@ proof (rule contradiction) from this notb show "FALSE" by (rule condD2) qed -theorem rev_mp (*[elim 2]*): +theorem rev_mp (*[Pure.elim 2]*): assumes a: "A" and ab: "A \ B" shows "B" using ab a by (rule mp) -theorem impE: +theorem impE [Pure.elim?]: assumes ab: "A \ B" and a: "A" and bc: "B \ C" shows "C" proof - @@ -715,33 +663,31 @@ proof - thus "C" by (rule bc) qed -theorem impCE [elim]: +theorem impCE [Pure.elim]: assumes ab: "A \ B" and b: "B \ P" and a: "\ A \ P" shows "P" -proof (rule classical) - assume contra: "\ P" - have "A" - proof (rule contradiction) - assume "\ A" hence "P" by (rule a) - with contra show "FALSE" .. - qed +proof (rule case_split[of "A"]) + assume "A" with ab have "B" by (rule mp) thus "P" by (rule b) +next + assume "\A" + thus "P" by (rule a) qed theorem impCE': (* used for classical prover *) - assumes ab: "A \ B" and a: "\C \ A" and b: "B \ C" - shows "C" + assumes ab: "A \ B" and a: "\P \ A" and b: "B \ P" + shows "P" proof (rule classical) - assume "\C" + assume "\P" hence "A" by (rule a) with ab have "B" by (rule mp) - thus "C" by (rule b) + thus "P" by (rule b) qed -text {* Equivalence *} +text \ Equivalence \ -theorem iffI [intro!]: +theorem iffI [Pure.intro!]: assumes ab: "A \ B" and ba: "B \ A" shows "A \ B" proof (unfold iff_def, rule) @@ -763,7 +709,7 @@ proof - thus ?thesis by (rule meta_eq_to_iff) qed -theorem iffD1 [elim 2]: +theorem iffD1 [Pure.elim?]: assumes ab: "A \ B" and a: "A" shows "B" using ab @@ -772,7 +718,7 @@ proof (unfold iff_def, elim conjE) from this a show "B" by (rule mp) qed -theorem iffD2 [elim 2]: +theorem iffD2 [Pure.elim?]: assumes ab: "A \ B" and b: "B" shows "A" using ab @@ -789,7 +735,7 @@ proof (rule r) from ab show "B \ A" by (unfold iff_def, elim conjE) qed -theorem iffCE [elim!]: +theorem iffCE [Pure.elim!]: assumes ab: "A \ B" and pos: "\A; B\ \ C" and neg: "\\ A; \ B\ \ C" shows "C" @@ -822,10 +768,10 @@ next qed -subsection {* Predicate Logic *} +subsection \ Predicate Logic \ -text {* - We take Hilbert's $\varepsilon$ as the basic binder and define the +text \ + We take Hilbert's @{text \} as the basic binder and define the other quantifiers as derived connectives. Again, we make sure that quantified formulas evaluate to @{term TRUE} or @{term FALSE}. @@ -836,43 +782,18 @@ text {* and automatic provers such as \texttt{blast} rely on reflection to the object level for reasoning about meta-connectives, which would not be possible with purely first-order quantification. -*} +\ consts Choice :: "('a \ c) \ 'a" - Ex :: "('a \ c) \ c" - All :: "('a \ c) \ c" - -text {* Concrete syntax: several variables as in @{text "\x,y : P(x,y)"}. *} - -nonterminal cidts (* comma-separated idt's *) - -syntax - "" :: "idt \ cidts" ("_" [100] 100) - "@cidts" :: "[idt, cidts] \ cidts" ("_,/ _" [100,100] 100) syntax (* BEWARE: choosing several values in sequence is not equivalent to choosing a tuple of values simultaneously. To avoid confusion, we do not allow several variables in a CHOOSE. *) "@Choice" :: "[idt, c] \ c" ("(3CHOOSE _ :/ _)" [100,10] 10) - "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100,10] 10) - "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100,10] 10) -syntax (xsymbols) - "@Ex" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100, 10] 10) - "@All" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100, 10] 10) translations "CHOOSE x: P" \ "CONST Choice(\x. P)" - (* separate parse and print translations for nested quantifiers because - they operate in opposite directions *) - "\x,xs : P" \ "CONST Ex(\x. \xs : P)" -(** FIXME: print translation doesn't work correctly: - will print \x : \y: P as \x, y : P -- why ??? **) -(* "\x,xs : P" \ "\x : Ex(\xs. P)" *) - "\x : P" \ "CONST Ex(\x. P)" - "\x,xs : P" \ "CONST All(\x. \xs : P)" -(* "\x,xs : P" \ "\x : All(\xs. P)" *) - "\x : P" \ "CONST All(\x. P)" (* Declare the two following axioms separately, otherwise the type checker infers the same type for P, and the type of P in @@ -880,20 +801,46 @@ translations in the first one. *) axiomatization where - chooseI: "P(t) \ P(CHOOSE x: P(x))" + chooseI: "P(t) \ P(CHOOSE x: P(x))" for t::'a axiomatization where choose_det : "(\x. P(x) \ Q(x)) \ (CHOOSE x: P(x)) = (CHOOSE x: Q(x))" -defs - Ex_def: "Ex(P) \ P(CHOOSE x : P(x) = TRUE) = TRUE" - All_def: "All(P) \ \(\x : \ P(x))" +definition Ex :: "('a \ c) \ c" + where "Ex(P) \ P(CHOOSE x : P(x) = TRUE) = TRUE" + +definition All :: "('a \ c) \ c" + where "All(P) \ \ Ex(\x. \ P(x))" + +text \ Concrete syntax: several variables as in @{text "\x,y : P(x,y)"}. \ -text {* +nonterminal cidts (* comma-separated idt's *) + +syntax + "" :: "idt \ cidts" ("_" [100] 100) + "@cidts" :: "[idt, cidts] \ cidts" ("_,/ _" [100,100] 100) + +syntax + "@Ex" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100,10] 10) + "@All" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100,10] 10) +syntax (ASCII) + "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) + "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) +translations + (* separate parse and print translations for nested quantifiers because + they operate in opposite directions *) + "\x,xs : P" \ "CONST Ex(\x. \xs : P)" + "\x,xs : P" \ "\x : CONST Ex(\xs. P)" + "\x : P" \ "CONST Ex(\x. P)" + "\x,xs : P" \ "CONST All(\x. \xs : P)" + "\x,xs : P" \ "\x : CONST All(\xs. P)" + "\x : P" \ "CONST All(\x. P)" + +text \ We introduce two constants @{text arbitrary} and @{text default} - that correspond to unconstrained and overconstrained choice, + that correspond to unconstrained and over-constrained choice, respectively. -*} +\ definition arbitrary::c where "arbitrary \ CHOOSE x : TRUE" @@ -901,7 +848,7 @@ definition arbitrary::c where definition default::c where "default \ CHOOSE x : FALSE" -theorem exI [intro]: +theorem exI [Pure.intro]: assumes hyp: "P(t)" shows "\x : P(x)" proof - @@ -909,7 +856,7 @@ proof - thus ?thesis by (unfold Ex_def, rule chooseI) qed -theorem exE [elim!]: +theorem exE [Pure.elim!]: assumes hyp: "\x : P(x)" and r: "\x. P(x) \ Q" shows "Q" proof - @@ -918,7 +865,7 @@ proof - thus "Q" by (rule r) qed -theorem allI [intro!]: +theorem allI [Pure.intro!]: assumes hyp: "\x. P(x)" shows "\x : P(x)" proof (unfold All_def, rule) @@ -936,7 +883,7 @@ proof (rule contradiction) with hyp show "FALSE" by (unfold All_def, elim notE) qed -theorem allE [elim]: +theorem allE [Pure.elim 2]: assumes hyp: "\x : P(x)" and r: "P(x) \ Q" shows "Q" proof (rule r) @@ -967,18 +914,18 @@ proof (rule q) from p show "P(CHOOSE x : P(x))" by (rule chooseI_ex) qed -lemma choose_equality [intro]: +lemma choose_equality [Pure.intro]: assumes "P(t)" and "\x. P(x) \ x=a" shows "(CHOOSE x : P(x)) = a" -using assms by (rule chooseI2[where Q="\x. x=a"]) + using assms by (rule chooseI2) -lemmas choose_equality' = sym[OF choose_equality, standard, intro] +lemmas choose_equality' = sym[OF choose_equality, Pure.intro] -text {* +text \ Skolemization rule: note that the existential quantifier in the conclusion introduces an operator (of type @{text "c \ c"}), not a value; second-order quantification is necessary here. -*} +\ lemma skolemI: assumes h: "\x : \y: P(x,y)" shows "\f : \x : P(x, f(x))" @@ -1000,16 +947,16 @@ next qed -subsection {* Setting up the automatic proof methods *} +subsection \ Setting up the automatic proof methods \ -subsubsection {* Reflection of meta-level to object-level *} +subsubsection \ Reflection of meta-level to object-level \ -text {* +text \ Our next goal is to getting Isabelle's automated tactics to work for \tlaplus{}. We follow the setup chosen for Isabelle/HOL as far as it applies to \tlaplus{}. The following lemmas, when used as rewrite rules, replace meta- by object-level connectives. -*} +\ lemma atomize_all [atomize]: "(\x. P(x)) \ Trueprop (\x : P(x))" proof @@ -1058,7 +1005,6 @@ next qed qed - lemmas [symmetric,rulify] = atomize_all atomize_imp and [symmetric,defn] = atomize_all atomize_imp atomize_eq @@ -1070,19 +1016,8 @@ lemma tmp2: "(\P x. P(x)) \ (\a::c. Q(a))" . lemma tmp3: "\a. (\x. P(x)) \ P(a)" . -ML {* - let - val thm = @{thm tmp2} - val rew = Object_Logic.atomize (cprop_of thm) - in - Raw_Simplifier.rewrite_rule [rew] thm - end -*} - ** end test cases **) -setup AtomizeElim.setup - lemma atomize_exL[atomize_elim]: "(\x. P(x) \ Q) \ ((\x : P(x)) \ Q)" by rule iprover+ @@ -1095,49 +1030,28 @@ by rule iprover+ lemma atomize_elimL[atomize_elim]: "(\B. (A \ B) \ B) \ Trueprop(A)" .. -subsubsection {* Setting up the classical reasoner *} +subsubsection \ Setting up the classical reasoner \ -text {* +text \ We now instantiate Isabelle's classical reasoner for \tlaplus{}. This includes methods such as \texttt{fast} and \texttt{blast}. -*} +\ lemma thin_refl: "\x=x; PROP W\ \ PROP W" . -ML {* +ML \ (* functions to take apart judgments and formulas, see Isabelle reference manual, section 9.3 *) - fun dest_Trueprop (Const(@{const_name Trueprop}, _) $ P) = P + fun dest_Trueprop (Const(@{const_name \Trueprop\}, _) $ P) = P | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); - fun dest_eq (Const(@{const_name eq}, _) $ t $ u) = (t,u) + fun dest_eq (Const(@{const_name \eq\}, _) $ t $ u) = (t,u) | dest_eq t = raise TERM ("dest_eq", [t]); - fun dest_imp (Const(@{const_name imp}, _) $ A $ B) = (A, B) + fun dest_imp (Const(@{const_name \imp\}, _) $ A $ B) = (A, B) | dest_imp t = raise TERM ("dest_imp", [t]); -(** - structure Hypsubst_Data = - struct - structure Simplifier = Simplifier - val dest_Trueprop = dest_Trueprop - val dest_eq = dest_eq - val dest_imp = dest_imp - val eq_reflection = @{thm eq_reflection} - val rev_eq_reflection = @{thm meta_to_obj_eq} - val imp_intr = @{thm impI} - val rev_mp = @{thm rev_mp} - val subst = @{thm subst} - val sym = @{thm sym} - val thin_refl = @{thm thin_refl} - val prop_subst = @{lemma "PROP P(t) ==> PROP prop (x = t ==> PROP P(x))" - by (unfold prop_def) (drule eq_reflection, unfold)} - end; - structure Hypsubst = HypsubstFun(Hypsubst_Data); - open Hypsubst; -**) - structure Hypsubst = Hypsubst( val dest_Trueprop = dest_Trueprop val dest_eq = dest_eq @@ -1152,19 +1066,6 @@ ML {* ); open Hypsubst; -(** - structure Classical_Data = - struct - val imp_elim = @{thm impCE'} - val not_elim = @{thm notE} - val swap = @{thm swap} - val classical = @{thm classical} - val sizef = Drule.size_of_thm - val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] - end; - structure Classical = ClassicalFun(Classical_Data); -**) - structure Classical = Classical( val imp_elim = @{thm impCE'} val not_elim = @{thm notE} @@ -1176,10 +1077,12 @@ ML {* structure BasicClassical : BASIC_CLASSICAL = Classical; open BasicClassical; -*} +\ +(* setup hypsubst_setup setup Classical.setup +*) declare refl [intro!] and trueI [intro!] @@ -1208,38 +1111,20 @@ declare refl [intro!] and choose_equality [intro] and sym[OF choose_equality, intro] -ML {* -(** - structure Blast = Blast - ( struct - val thy = @{theory} - type claset = Classical.claset - val equality_name = @{const_name PredicateLogic.eq} - val not_name = @{const_name PredicateLogic.Not} - val notE = @{thm notE} - val ccontr = @{thm contradiction} - val contr_tac = Classical.contr_tac - val dup_intr = Classical.dup_intr - val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac - val rep_cs = Classical.rep_cs - val cla_modifiers = Classical.cla_modifiers; - val cla_meth' = Classical.cla_meth' - end ); -**) - +ML \ structure Blast = Blast ( structure Classical = Classical - val Trueprop_const = dest_Const @{const Trueprop} + val Trueprop_const = dest_Const @{const \Trueprop\} val equality_name = @{const_name PredicateLogic.eq} - val not_name = @{const_name PredicateLogic.Not} + val not_name = @{const_name PredicateLogic.not} val notE = @{thm notE} val ccontr = @{thm contradiction} val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac ); -*} +\ -setup Blast.setup +(* setup Blast.setup *) (*** TEST DATA *** @@ -1248,7 +1133,7 @@ apply hypsubst apply (rule refl) done -lemma "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)" +lemma "\x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)" by hypsubst lemma "x = x" @@ -1303,35 +1188,44 @@ by blast *** END TEST DATA ***) -subsubsection {* Setting up the simplifier *} +subsubsection \ Setting up the simplifier \ -text {* +text \ We instantiate the simplifier, Isabelle's generic rewriting machinery. Equational laws for predicate logic will be proven below; they automate much of the purely logical reasoning. -*} +\ + +text \first use of axiom @{text eqBoolean}\ +lemma neq_conv_eq_False: + assumes "x \ y" shows "(x = y) = FALSE" + by (rule eqBoolean) (rule notE[OF assms]) -lemma if_bool_eq_conj: +lemma not_refl_False: + assumes "(x \ x)" shows "A" + using assms by (rule notE) (rule refl) + +lemma if_bool_equiv_conj: "(IF A THEN B ELSE C) \ ((A \ B) \ (\A \ C))" by fast -text {* +text \ A copy of Isabelle's meta-level implication is introduced, which is used internally by the simplifier for fine-tuning congruence rules by simplifying their premises. -*} +\ definition simp_implies :: "[prop, prop] \ prop" (infixr "=simp=>" 1) where - "simp_implies \ op \" + "simp_implies \ (\)" lemma simp_impliesI: - assumes PQ: "(PROP P ==> PROP Q)" + assumes PQ: "(PROP P \ PROP Q)" shows "PROP P =simp=> PROP Q" unfolding simp_implies_def by (rule PQ) lemma simp_impliesE: assumes PQ: "PROP P =simp=> PROP Q" - and P: "PROP P" and QR: "PROP Q ==> PROP R" + and P: "PROP P" and QR: "PROP Q \ PROP R" shows "PROP R" proof - from P have "PROP Q" by (rule PQ [unfolded simp_implies_def]) @@ -1339,32 +1233,29 @@ proof - qed lemma simp_implies_cong: - assumes PP' :"PROP P == PROP P'" - and P'QQ': "PROP P' ==> (PROP Q == PROP Q')" - shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')" + assumes PP' :"PROP P \ PROP P'" + and P'QQ': "PROP P' \ (PROP Q \ PROP Q')" + shows "(PROP P =simp=> PROP Q) \ (PROP P' =simp=> PROP Q')" unfolding simp_implies_def proof (rule equal_intr_rule) - assume PQ: "PROP P ==> PROP Q" and P': "PROP P'" + assume PQ: "PROP P \ PROP Q" and P': "PROP P'" from PP' [symmetric] and P' have "PROP P" by (rule equal_elim_rule1) hence "PROP Q" by (rule PQ) with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) next - assume P'Q': "PROP P' ==> PROP Q'" and P: "PROP P" + assume P'Q': "PROP P' \ PROP Q'" and P: "PROP P" from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) hence "PROP Q'" by (rule P'Q') with P'QQ' [OF P', symmetric] show "PROP Q" by (rule equal_elim_rule1) qed -use "simplifier_setup.ML" - -setup {* - Simplifier.map_simpset_global (K Simpdata.PL_basic_ss) - #> Simplifier.method_setup Splitter.split_modifiers - #> Splitter.setup - #> Simpdata.clasimp_setup - #> EqSubst.setup -*} +ML_file "simplifier_setup.ML" +ML \open Simpdata\ +setup \ + map_theory_simpset (put_simpset PL_basic_ss) #> + Simplifier.method_setup Splitter.split_modifiers +\ (*** TEST DATA -- note: only basic simplification, no theorems set up *** @@ -1393,10 +1284,13 @@ lemma "x=y \ y=z \ x=z \ f(y) = f(x)" by auto lemma "x=y \ y=z \ x=z" -by force + by force lemma "(\x : P(x)) \ P(a)" -by auto + by auto + +lemma "x \ y \ y = x \ A" + by simp lemma assumes ind: "P(z) \ (\n. P(n) \ P(s(n))) \ (\n. P(n))" @@ -1410,13 +1304,11 @@ lemma shows "Q(n)" using assms by auto -(* without converting all connectives to meta-level (via attribute - [rule_format]), the following loops *) lemma assumes "\Q. (\n: (\i: L(i,n) \ Q(i)) \ Q(n)) \ (\n: Q(n))" and "\n: (\i: L(i,n) \ Q(i)) \ Q(n)" shows "Q(n)" -using assms[rule_format] by auto +using assms(*[rule_format]*) by auto *** END TEST DATA ***) @@ -1438,38 +1330,49 @@ next thus "TRUE = A" by (rule sym) qed +lemma implies_true_equals: "(PROP P \ TRUE) \ Trueprop(TRUE)" + by standard (intro trueI) + +lemma false_implies_equals: "(FALSE \ P) \ Trueprop(TRUE)" + by standard simp_all + +lemma cond_cancel: "(IF x=y THEN y ELSE x) = x" + by fast + lemmas [simp] = triv_forall_equality - TrueAssumption trueprop_eq_true trueprop_true_eq - refl[THEN eqTrueI] -- {* @{text "(x = x) \ TRUE"} *} + TrueAssumption implies_true_equals false_implies_equals + refl[THEN eqTrueI] (* (x = x) = TRUE *) condT (*condThen*) notTrue condF (*condElse*) notFalse - cond_id - false_neq_true[THEN eqBoolean] - not_sym[OF false_neq_true, THEN eqBoolean] + cond_id cond_cancel + false_neq_true[THEN neq_conv_eq_False] (* (FALSE = TRUE) = FALSE *) + not_sym[OF false_neq_true, THEN neq_conv_eq_False] (* (TRUE = FALSE) = FALSE *) iff_refl lemmas [cong] = simp_implies_cong -subsubsection {* Reasoning by cases *} +subsubsection \ Reasoning by cases \ -text {* +text \ The next bit of code sets up reasoning by cases as a proper Isar method, so we can write ``proof cases'' etc. Following the development of FOL, we introduce a set of ``shadow connectives'' that will only be used for this purpose. -*} +\ -theorems cases = case_split [case_names True False] +(* lemmas cases = case_split [case_names True False] *) +context +begin -definition cases_equal where "cases_equal \ eq" -definition cases_implies where "cases_implies \ imp" -definition cases_conj where "cases_conj \ conj" -definition cases_forall where "cases_forall(P) \ \x: P(x)" -definition cases_true where "cases_true \ TRUE" -definition cases_false where "cases_false \ FALSE" +qualified definition cases_equal where "cases_equal \ eq" +qualified definition cases_implies where "cases_implies \ imp" +qualified definition cases_conj where "cases_conj \ conj" +qualified definition cases_forall where "cases_forall(P) \ \x: P(x)" +qualified definition cases_true where "cases_true \ TRUE" +qualified definition cases_false where "cases_false \ FALSE" lemma cases_equal_eq: "(x \ y) \ Trueprop(cases_equal(x, y))" unfolding atomize_eq cases_equal_def . @@ -1488,8 +1391,8 @@ unfolding cases_true_def .. lemmas cases_atomize' = cases_implies_eq cases_conj_eq cases_forall_eq lemmas cases_atomize = cases_atomize' cases_equal_eq -lemmas cases_rulify' [symmetric, standard] = cases_atomize' -lemmas cases_rulify [symmetric, standard] = cases_atomize +lemmas cases_rulify' [symmetric] = cases_atomize' +lemmas cases_rulify [symmetric] = cases_atomize lemmas cases_rulify_fallback = cases_equal_def cases_implies_def cases_conj_def cases_forall_def cases_true_def cases_false_def @@ -1515,48 +1418,50 @@ qed lemmas cases_conj = cases_forall_conj cases_implies_conj cases_conj_curry -ML {* +ML \ structure Induct = Induct ( - val cases_default = @{thm cases} + val cases_default = @{thm case_split} val atomize = @{thms cases_atomize} val rulify = @{thms cases_rulify'} val rulify_fallback = @{thms cases_rulify_fallback} val equal_def = @{thm cases_equal_def} fun dest_def (Const (@{const_name cases_equal}, _) $ t $ u) = SOME (t, u) | dest_def _ = NONE - val trivial_tac = match_tac @{thms cases_trueI} + fun trivial_tac ctxt = match_tac ctxt @{thms cases_trueI} ) -*} +\ -setup {* - Induct.setup #> - Context.theory_map (Induct.map_simpset (fn ss => ss - setmksimps (fn ss => Simpdata.mksimps Simpdata.mksimps_pairs ss #> - map (Simplifier.rewrite_rule (map Thm.symmetric @{thms cases_rulify_fallback}))) +declaration \ + fn _ => Induct.map_simpset (fn ss => ss addsimprocs - [Simplifier.simproc_global @{theory} "swap_cases_false" - ["cases_false ==> PROP P ==> PROP Q"] - (fn _ => fn _ => - (fn _ $ (P as _ $ @{const cases_false}) $ (_ $ Q $ _) => - if P <> Q then SOME Drule.swap_prems_eq else NONE - | _ => NONE)), - Simplifier.simproc_global @{theory} "cases_equal_conj_curry" - ["cases_conj(P, Q) ==> PROP R"] - (fn _ => fn _ => - (fn _ $ (_ $ P) $ _ => - let - fun is_conj (@{const cases_conj} $ P $ Q) = - is_conj P andalso is_conj Q - | is_conj (Const (@{const_name cases_equal}, _) $ _ $ _) = true - | is_conj @{const cases_true} = true - | is_conj @{const cases_false} = true - | is_conj _ = false - in if is_conj P then SOME @{thm cases_conj_curry} else NONE end - | _ => NONE))])) -*} - -text {* Pre-simplification of induction and cases rules *} + [Simplifier.make_simproc @{context} "swap_cases_false" + {lhss = [@{term "cases_false \ PROP P \ PROP Q"}], + proc = fn _ => fn _ => fn ct => + (case Thm.term_of ct of + _ $ (P as _ $ @{const cases_false}) $ (_ $ Q $ _) => + if P <> Q then SOME Drule.swap_prems_eq else NONE + | _ => NONE)}, + Simplifier.make_simproc @{context} "cases_equal_conj_curry" + {lhss = [@{term "cases_conj(P, Q) \ PROP R"}], + proc = fn _ => fn _ => fn ct => + (case Thm.term_of ct of + _ $ (_ $ P) $ _ => + let + fun is_conj (@{const cases_conj} $ P $ Q) = + is_conj P andalso is_conj Q + | is_conj (Const (@{const_name cases_equal}, _) $ _ $ _) = true + | is_conj @{const cases_true} = true + | is_conj @{const cases_false} = true + | is_conj _ = false + in if is_conj P then SOME @{thm cases_conj_curry} else NONE end + | _ => NONE)}] + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> + map (rewrite_rule ctxt (map Thm.symmetric @{thms cases_rulify_fallback})))) +\ + +text \ Pre-simplification of induction and cases rules \ lemma [induct_simp]: "(\x. cases_equal(x, t) \ PROP P(x)) \ PROP P(t)" unfolding cases_equal_def @@ -1593,7 +1498,7 @@ qed lemma [induct_simp]: "(PROP P \ cases_true) \ Trueprop(cases_true)" unfolding cases_true_def by (iprover intro: equal_intr_rule) -lemma [induct_simp]: "(\x. cases_true) \ Trueprop(cases_true)" +lemma [induct_simp]: "(\x::'a::{}. cases_true) \ Trueprop(cases_true)" unfolding cases_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "Trueprop(cases_implies(cases_true, P)) \ Trueprop(P)" @@ -1602,7 +1507,7 @@ lemma [induct_simp]: "Trueprop(cases_implies(cases_true, P)) \ Trueprop(P lemma [induct_simp]: "(x = x) = TRUE" by simp -hide_const cases_forall cases_implies cases_equal cases_conj cases_true cases_false +end (*** BASIC TEST DATA *** @@ -1616,18 +1521,17 @@ qed *** END TEST DATA ***) +subsection \ Propositional simplification \ -subsection {* Propositional simplification *} +subsubsection \ Conversion to Boolean values \ -subsubsection {* Conversion to Boolean values *} - -text {* +text \ Because \tlaplus{} is untyped, equivalence is different from equality, and one has to be careful about stating the laws of propositional logic. For example, although the equivalence @{text "(TRUE \ A) \ A"} is valid, we cannot state the law @{text "(TRUE \ A) = A"} because we have no way of knowing the value of, e.g., @{text "TRUE \ 3"}. - These equalities are valid only if + Such equations are valid only if the connectives are applied to Boolean operands. For automatic reasoning, we are interested in equations that can be used by Isabelle's simplifier. We therefore introduce an auxiliary @@ -1637,7 +1541,7 @@ text {* We will prove below that propositional formulas return a Boolean value when applied to arbitrary arguments. -*} +\ definition boolify :: "c \ c" where "boolify(x) \ IF x THEN TRUE ELSE FALSE" @@ -1645,10 +1549,10 @@ definition boolify :: "c \ c" where definition isBool :: "c \ c" where "isBool(x) \ boolify(x) = x" -text {* +text \ The formulas @{text "P"} and @{text "boolify(P)"} are inter-derivable (but need of course not be equal, unless @{text P} is a Boolean). -*} +\ lemma boolifyI [intro!]: "P \ boolify(P)" unfolding boolify_def by fast @@ -1665,11 +1569,11 @@ lemma boolify_cases: (** too general to be used as a default rule **) shows "P(boolify(x))" unfolding boolify_def using assms by (fast intro: condI) -text {* +text \ @{text boolify} can be defined as @{text "x = TRUE"}. For automatic reasoning, we rewrite the latter to the former, and derive calculational rules for @{text boolify}. -*} +\ lemma [simp]: "(x = TRUE) = boolify(x)" proof (cases "x") @@ -1704,14 +1608,14 @@ by (auto simp add: boolify_def) lemma falseIsBool [intro!,simp]: "isBool(FALSE)" by (unfold isBool_def, rule boolifyFalse) -text {* +text \ The following lemma is used to turn hypotheses @{text "\A"} into - rewrite rules @{text "A = FALSE"}. -*} + rewrite rules @{text "A = FALSE"} when we know that @{text A} is Boolean. +\ lemma boolifyFalseI [intro]: "\ A \ boolify(A) = FALSE" by (auto simp add: boolify_def) -text {* idempotence of @{text boolify} *} +text \ idempotence of @{text boolify} \ lemma boolifyBoolify [simp]: "boolify(boolify(x)) = boolify(x)" by (auto simp add: boolify_def) @@ -1734,18 +1638,16 @@ proof - show ?thesis by (unfold bx) qed -lemmas isBoolE [elim!] = isBoolTrueFalse[THEN disjE, standard] +lemmas isBoolE [elim!] = isBoolTrueFalse[THEN disjE] -lemma boolifyEq [simp]: "boolify(t=u) = (t=u)" (* first use of axiom eqBoolean *) +lemma boolifyEq [simp]: "boolify(t=u) = (t=u)" proof (cases "t=u") case True - hence "(t=u) = TRUE" by (rule eqTrueI) - hence tu: "(t=u) \ TRUE" by (rule eq_reflection) + hence tu: "(t=u) \ TRUE" by (rule eqTrueI[THEN eq_reflection]) show ?thesis by (unfold tu, rule boolifyTrue) next case False - hence "(t=u) = FALSE" by (rule eqBoolean) - hence tu: "(t=u) \ FALSE" by (rule eq_reflection) + hence tu: "(t=u) \ FALSE" by (rule neq_conv_eq_False[THEN eq_reflection]) show "boolify(t=u) = (t=u)" by (unfold tu, rule boolifyFalse) qed @@ -1795,18 +1697,18 @@ by (simp add: iff_def) lemma iffIsBool [intro!,simp]: "isBool(A \ B)" unfolding isBool_def by (rule boolifyIff) -text {* +text \ We can now rewrite equivalences to equations between ``boolified'' arguments, and this gives rise to a technique for proving equations between formulas. -*} +\ lemma boolEqual: assumes "P \ Q" and "isBool(P)" and "isBool(Q)" shows "P = Q" using assms by auto -text {* +text \ The following variant converts equivalences to equations. It might be useful as a (non-conditional) simplification rule, but I suspect that for goals it is more useful to use the standard introduction @@ -1814,12 +1716,10 @@ text {* For assumptions we can use lemma @{text boolEqual} for turning equivalences into conditional rewrites. -*} +\ lemma iffIsBoolifyEqual: "(A \ B) = (boolify(A) = boolify(B))" -proof (rule boolEqual) - show "(A \ B) \ (boolify(A) = boolify(B))" by (auto simp: boolifyFalseI) -qed (simp_all) + by (rule boolEqual) (auto simp: boolifyFalseI) lemma iffThenBoolifyEqual: assumes "A \ B" shows "boolify(A) = boolify(B)" @@ -1834,21 +1734,23 @@ lemma boolEqualIff: shows "(P = Q) = (P \ Q)" using assms by (auto intro: boolEqual) -ML {* +ML \ structure Simpdata = struct open Simpdata; - val mksimps_pairs = [(@{const_name Not}, (@{thms boolifyFalseI}, true)), + val mksimps_pairs = [(@{const_name not}, (@{thms boolifyFalseI}, true)), (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] @ mksimps_pairs; end; open Simpdata; -*} +\ -declaration {* fn _ => - Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) -*} +declaration \ + fn _ => Simplifier.map_ss (fn ss => ss + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +\ (** TEST DATA ** @@ -1864,39 +1766,39 @@ using assms by simp ** END TEST DATA **) -text {* +text \ The following code rewrites @{text "x = y"} to @{text "FALSE"} in the presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. The simplifier is set up so that @{text "y = x"} is already simplified to @{text "FALSE"}, this code adds symmetry of disequality to simplification. -*} +\ lemma symEqLeft: "(x = y) = b \ (y = x) = b" by (auto simp: boolEqualIff) -simproc_setup neq ("x = y") = {* fn _ => -let - val neq_to_EQ_False = @{thm not_sym} RS @{thm eqBoolean} RS @{thm eq_reflection}; - val symEqLeft_to_symEQLeft = @{thm symEqLeft} RS @{thm eq_reflection}; - fun is_neq lhs rhs thm = - (case Thm.prop_of thm of - _ $ (Not' $ (eq' $ l' $ r')) => - Not' = @{const "Not"} andalso eq' = @{const "eq"} andalso - r' aconv lhs andalso l' aconv rhs +simproc_setup neq ("x = y") = \fn _ => + let + val neq_to_EQ_False = @{thm not_sym} RS @{thm neq_conv_eq_False} RS @{thm eq_reflection}; + val symEqLeft_to_symEQLeft = @{thm symEqLeft} RS @{thm eq_reflection}; + fun is_neq eq lhs rhs thm = + (case Thm.prop_of thm of + _ $ (Not $ (eq' $ l' $ r')) => + Not = @{const \not\} andalso eq' = eq andalso + r' aconv lhs andalso l' aconv rhs | _ $ (eq' $ (eq'' $ l' $ r') $ f') => - eq' = @{const "eq"} andalso eq'' = @{const "eq"} andalso - f' = @{const "FALSE"} andalso r' aconv lhs andalso l' aconv rhs - | _ => false); - fun proc ss ct = - (case Thm.term_of ct of - eq $ lhs $ rhs => - (case find_first (is_neq lhs rhs) (Simplifier.prems_of ss) of - SOME thm => SOME ((thm RS symEqLeft_to_symEQLeft) - handle _ => thm RS neq_to_EQ_False) - | NONE => NONE) - | _ => NONE); -in proc end; -*} + eq' = @{const \eq\} andalso eq'' = @{const \eq\} andalso + f' = @{const \FALSE\} andalso r' aconv lhs andalso l' aconv rhs + | _ => false); + fun proc ss ct = + (case Thm.term_of ct of + eq $ lhs $ rhs => + (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of + SOME thm => SOME ((thm RS symEqLeft_to_symEQLeft) + handle _ => thm RS neq_to_EQ_False) + | NONE => NONE) + | _ => NONE); + in proc end +\ (*** TEST *** lemma "a \ b \ (IF b = a THEN t ELSE u) = u" @@ -1951,80 +1853,76 @@ by auto Check what needs to be done to derive the necessary chains of equivalences. **** - -text {* +text \ The following lemmas are used by the simplifier to rearrange quantifiers. -*} +\ lemma uncurry: assumes "P \ Q \ R" shows "P \ Q \ R" using assms by blast +lemma bool_iffI: + assumes "P \ Q" and "Q \ P" + shows "boolify(P) = boolify(Q)" + sorry + lemma iff_allI: assumes "\x. P(x) = Q(x)" shows "(\x : P(x)) = (\x : Q(x))" - using assms by blast + unfolding assms[THEN eq_reflection] by (rule refl) lemma iff_exI: - assumes "!!x. P x = Q x" - shows "(\x. P x) = (\x. Q x)" - using assms by blast + assumes "\x. P(x) = Q(x)" + shows "(\x : P(x)) = (\x : Q(x))" + unfolding assms[THEN eq_reflection] by (rule refl) -lemma all_comm: - "(\x y. P x y) = (\y x. P x y)" +lemma all_comm: "(\x,y : P(x,y)) = (\y,x : P(x,y))" by blast -lemma ex_comm: - "(\x y. P x y) = (\y x. P x y)" +lemma ex_comm: "(\x,y : P(x,y)) = (\y,x : P(x,y))" by blast -ML {* - use "~~/src/Provers/quantifier1.ML"; - local - val uncurry = - prove_goal (the_context()) "P => Q => R ==> P & Q => R" - (fn prems => [cut_facts_tac prems 1, Blast_tac 1]); - val iff_allI = - prove_goal (the_context()) "(!!x. P(x) <=> Q(x)) ==> (ALL x. P(x)) = (ALL x. Q(x))" - (fn prems => [cut_facts_tac prems 1, - auto_tac (claset() addIs [boolEqual], simpset())]); - val iff_exI = - prove_goal (the_context()) "(!!x. P(x) <=> Q(x)) ==> (EX x. P(x)) = (EX x. Q(x))" - (fn prems => [cut_facts_tac prems 1, - auto_tac (claset() addIs [boolEqual], simpset())]); - val all_comm = - prove_goal (the_context()) "(ALL x y. P(x,y)) = (ALL y x. P(x,y))" - (fn _ => [auto_tac (claset() addIs [boolEqual], simpset())]); - val ex_comm = - prove_goal (the_context()) "(EX x y. P(x,y)) = (EX y x. P(x,y))" - (fn _ => [auto_tac (claset() addIs [boolEqual], simpset())]); - val c_type = Type("c", []); - structure Quantifier1 = Quantifier1Fun( - struct - (* syntax *) - fun dest_eq((c as Const("eq",_)) $ s $ t) = SOME(c,s,t) - | dest_eq _ = NONE; - fun dest_conj((c as Const("conj",_)) $ s $ t) = SOME(c,s,t) - | dest_conj _ = NONE; - fun dest_imp((c as Const("imp",_)) $ s $ t) = SOME(c,s,t) - | dest_imp _ = NONE; - val conj = @{term conj}; - val imp = @{term imp}; - (* rules *) - val iff_reflection = thm "eq_reflection"; - val iffI = @{thm iffI}; - val iff_trans = @{thm trans}; - end); - in - -*} +ML \ + structure Quantifier1 = Quantifier1( + (* syntax *) + fun dest_eq(Const(@{const_name \eq\},_) $ s $ t) = SOME(s,t) + | dest_eq _ = NONE; + fun dest_conj(Const(@{const_name \conj\},_) $ s $ t) = SOME(s,t) + | dest_conj _ = NONE; + fun dest_imp(Const(@{const_name \imp\},_) $ s $ t) = SOME(s,t) + | dest_imp _ = NONE; + val conj = @{term conj}; + val imp = @{term imp}; + (* rules *) + val iff_reflection = @{thm eq_reflection}; + val iffI = @{thm bool_iffI}; + val iff_trans = @{thm trans}; + val conjI= @{thm conjI} + val conjE= @{thm conjE} + val impI = @{thm impI} + val mp = @{thm mp} + val uncurry = @{thm uncurry} + val exI = @{thm exI} + val exE = @{thm exE} + val iff_allI = @{thm iff_allI} + val iff_exI = @{thm iff_exI} + val all_comm = @{thm all_comm} + val ex_comm = @{thm ex_comm} + ); +\ + +simproc_setup defined_Ex ("\x : P(x)") = \fn _ => Quantifier1.rearrange_ex\ +simproc_setup defined_All ("\x : P(x)") = \fn _ => Quantifier1.rearrange_all\ + +lemma "\x,y : y \x \ x=y \ x = FALSE" + apply simp ***) -text {* +text \ Orient equations with Boolean constants such that the constant appears on the right-hand side. -*} +\ lemma boolConstEqual [simp]: "(TRUE = P) = (P = TRUE)" @@ -2032,7 +1930,43 @@ lemma boolConstEqual [simp]: by blast+ -subsubsection {* Simplification laws for conditionals *} +subsubsection \Reorienting equations\ + +lemma eq_commute: "(a = b) = (b = a)" + by (auto intro: boolEqual) + +ML \ +signature REORIENT_PROC = +sig + val add : (term -> bool) -> theory -> theory + val proc : morphism -> Proof.context -> cterm -> thm option +end; + +structure Reorient_Proc : REORIENT_PROC = +struct + structure Data = Theory_Data + ( + type T = ((term -> bool) * stamp) list; + val empty = []; + val extend = I; + fun merge data : T = Library.merge (eq_snd (op =)) data; + ); + fun add m = Data.map (cons (m, stamp ())); + fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); + + val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; + fun proc _ ctxt ct = + let + val thy = Proof_Context.theory_of ctxt; + in + case Thm.term_of ct of + (_ $ _ $ u) => if matches thy u then NONE else SOME meta_reorient + | _ => NONE + end; +end; +\ + +subsubsection \ Simplification laws for conditionals \ lemma splitCond [split]: assumes q: "\x. isBool(Q(x))" @@ -2045,7 +1979,8 @@ next thus ?thesis by (auto intro: q) qed -lemma splitCondAsm: -- {* useful with conditionals in hypotheses *} +text \useful with conditionals in hypotheses\ +lemma splitCondAsm: assumes "\x. isBool(Q(x))" shows "Q(IF P THEN t ELSE e) = (\((P \ \Q(t)) \ (\P \ \Q(e))))" using assms by (simp only: splitCond, blast) @@ -2054,7 +1989,8 @@ lemma condCong (*[cong]*): (* strangely, seems to simplify less when active ?! "P = Q \ (IF P THEN t ELSE e) = (IF Q THEN t ELSE e)" by simp -lemma condFullCong: -- {* not active by default, because too expensive *} +text \not active by default, because too expensive\ +lemma condFullCong: "\P = Q; Q \ t=t'; \Q \ e=e'\ \ (IF P THEN t ELSE e) = (IF Q THEN t' ELSE e')" by auto @@ -2065,7 +2001,6 @@ lemma substCond [intro]: using assms by auto lemma cond_simps [simp]: - "(IF x = y THEN y ELSE x) = x" "(IF (IF A THEN B ELSE C) THEN t ELSE e) = (IF (A \ B) \ (\A \ C) THEN t ELSE e)" "(IF A THEN (IF B THEN t ELSE u) ELSE v) = @@ -2073,7 +2008,7 @@ lemma cond_simps [simp]: by auto -subsubsection {* Simplification laws for conjunction *} +subsubsection \ Simplification laws for conjunction \ lemma conj_simps [simp]: "(P \ TRUE) = boolify(P)" @@ -2089,17 +2024,17 @@ lemma conj_simps [simp]: "((P \ Q) \ R) = (P \ Q \ R)" by auto -text {* +text \ The congruence rule for conjunction is occasionally useful, but not active by default. -*} +\ lemma conj_cong: assumes "P = P'" and "P' \ Q = Q'" shows "(P \ Q) = (P' \ Q')" using assms by auto -text {* Commutativity laws are not active by default *} +text \Commutativity laws are not active by default\ lemma conj_comms: "(P \ Q) = (Q \ P)" @@ -2107,7 +2042,7 @@ lemma conj_comms: by auto -subsubsection {* Simplification laws for disjunction *} +subsubsection \Simplification laws for disjunction\ lemma disj_simps [simp]: "(P \ TRUE) = TRUE" @@ -2123,14 +2058,14 @@ lemma disj_simps [simp]: "((P \ Q) \ R) = (P \ Q \ R)" by auto -text {* Congruence rule, not active by default *} +text \Congruence rule, not active by default\ lemma disj_cong: assumes "P = P'" and "\P' \ Q = Q'" shows "(P \ Q) = (P' \ Q')" using assms by auto -text {* Commutativity laws are not active by default *} +text \Commutativity laws are not active by default\ lemma disj_comms: "(P \ Q) = (Q \ P)" @@ -2138,13 +2073,13 @@ lemma disj_comms: by auto -subsubsection {* Simplification laws for negation *} +subsubsection \Simplification laws for negation\ -text {* +text \ Negated formulas create simplifications of the form @{text "A = FALSE"}, we therefore prove two versions of the - following lemmas to complete critical pairs. -*} + following lemmas in order to complete critical pairs. +\ lemma not_simps [simp]: "(\(P \ Q)) = (\P \ \Q)" @@ -2156,14 +2091,14 @@ lemma not_simps [simp]: "(x \ x) = FALSE" "\P. (\(\x : P(x))) = (\x : \P(x))" "\P. (\(\x : P(x))) = (\x : \P(x))" -by (auto simp del: notBoolifyFalse) + by (auto simp del: notBoolifyFalse) declare not_simps [simplified,simp] lemma eqFalse_eqFalse [simp]: "isBool(P) \ ((P = FALSE) = FALSE) = P" by auto -subsubsection {* Simplification laws for implication *} +subsubsection \Simplification laws for implication\ lemma imp_simps [simp]: "(P \ FALSE) = (\P)" @@ -2179,7 +2114,7 @@ lemma imp_cong [cong]: by auto -subsubsection {* Simplification laws for equivalence *} +subsubsection \Simplification laws for equivalence\ lemma iff_simps [simp]: "(TRUE \ P) = boolify(P)" @@ -2194,7 +2129,7 @@ lemma iff_cong [cong]: by auto -subsubsection {* Simplification laws for quantifiers *} +subsubsection \Simplification laws for quantifiers\ (*** subsumed by following simplifications lemma [simp]: "\x : TRUE" @@ -2212,8 +2147,7 @@ lemma quant_simps [simp]: "\ P. (\x : t=x \ P(x)) = boolify(P(t))" by auto - -text {* Miniscoping of quantifiers. *} +text \Miniscoping of quantifiers\ lemma miniscope_ex [simp]: "\P Q. (\x : P(x) \ Q) = ((\x : P(x)) \ Q)" @@ -2238,15 +2172,15 @@ by (rule chooseI, rule refl) declare choose_det [cong] -text {* +text \ A @{text CHOOSE} expression evaluates to @{text default} if the only possible value satisfying the predicate equals @{text default}. Note that the reverse implication is not necessarily true: there could be several values satisfying @{text "P(x)"}, including @{text default}, - and @{text CHOOSE} may return @{text default}. This rule can be useful - for reasoning about \textsc{case} expressions where none of the guards - is true. -*} + and @{text CHOOSE} may return @{text default}. The following rule + can be useful for reasoning about \textsc{case} expressions where + none of the guards is true. +\ lemma equal_default [intro!]: assumes p: "\x : P(x) \ x = default" @@ -2264,11 +2198,9 @@ next unfolding default_def by (blast intro: choose_det) qed -lemmas [intro!] = sym[OF equal_default, standard] +lemmas [intro!] = sym[OF equal_default] -text {* - Similar lemma for @{text arbitrary}. -*} +text \Similar lemma for @{text arbitrary}.\ lemma equal_arbitrary: assumes p: "\x : P(x)" @@ -2279,9 +2211,9 @@ unfolding arbitrary_def proof (rule choose_det) qed -subsubsection {* Distributive laws *} +subsubsection \Distributive laws\ -text {* Not active by default. *} +text \Not active by default\ lemma prop_distrib: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" @@ -2299,7 +2231,7 @@ lemma quant_distrib: by auto -subsubsection {* Further calculational laws *} +subsubsection \Further calculational laws\ lemma cases_simp (*[simp]*): "((P \ Q) \ (\P \ Q)) = boolify(Q)" by auto diff --git a/isabelle/ROOT b/isabelle/ROOT new file mode 100644 index 00000000..53c1d9d9 --- /dev/null +++ b/isabelle/ROOT @@ -0,0 +1,7 @@ +(* build the session using "isabelle build -D ." *) +session "TLA+" = "Pure" + + options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] + theories + Peano + document_files + "root.tex" diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index fa227b21..e37ff62f 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -1,107 +1,92 @@ (* Title: TLA+/SetTheory.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:48 merz> + Version: Isabelle2020 *) -header {* \tlaplus{} Set Theory *} +section \\tlaplus{} Set Theory\ theory SetTheory imports PredicateLogic begin -text {* - This theory defines the version of Zermelo-Fr\"ankel set theory +text \ + This theory defines the version of Zermelo-Fraenkel set theory that underlies \tlaplus{}. -*} +\ -subsection {* Basic syntax and axiomatic basis of set theory. *} +subsection \Basic syntax and axiomatic basis of set theory\ -text {* - We take the set-theoretic constructs of \tlaplus{}, but add generalized - intersection for symmetry and convenience. (Note that @{text "INTER {} = {}"}.) -*} +text \ + The following constants are at the basis of our set theory. +\ consts - emptyset :: "c" ("{}" 100) -- {* empty set *} - upair :: "[c, c] \ c" -- {* unordered pairs *} - addElt :: "[c, c] \ c" -- {* add element to set *} - infinity :: "c" -- {* infinity set *} - SUBSET :: "c \ c" ("SUBSET _" [100]90) -- {* power set *} - UNION :: "c \ c" ("UNION _" [100]90) -- {* generalized union *} - INTER :: "c \ c" ("INTER _" [100]90) -- {* generalized intersection *} - "cup" :: "[c, c] \ c" (infixl "\\cup" 65) -- {* binary union *} - "cap" :: "[c, c] \ c" (infixl "\\cap" 70) -- {* binary intersection *} - "setminus":: "[c, c] \ c" (infixl "\\" 65) -- {* binary set difference *} - "in" :: "[c, c] \ c" (infixl "\\in" 50) -- {* membership relation *} - "subseteq":: "[c, c] \ c" (infixl "\\subseteq" 50) -- {* subset relation *} - subsetOf :: "[c, c \ c] \ c" -- {*@{text "subsetOf(S,p)"} = $\{x \in S : p\}$*} - setOfAll :: "[c, c \ c] \ c" -- {*@{text "setOfAll(S,e)"} = $\{e : x \in S\}$*} - bChoice :: "[c, c \ c] \ c" -- {* bounded choice *} - bAll :: "[c, c \ c] \ c" -- {* bounded universal quantifier *} - bEx :: "[c, c \ c] \ c" -- {* bounded existential quantifier *} - -notation (xsymbols) - "cup" (infixl "\" 65) and - "cap" (infixl "\" 70) and - "setminus" (infixl "\" 65) and - "in" (infixl "\" 50) and - "subseteq" (infixl "\" 50) - -notation (HTML output) - "cup" (infixl "\" 65) and - "cap" (infixl "\" 70) and - "setminus" (infixl "\" 65) and - "in" (infixl "\" 50) and - "subseteq" (infixl "\" 50) - -abbreviation "notin(a, S) \ \(a \ S)" -- {* negated membership *} + "in" :: "[c, c] \ c" (infixl "\" 50) (* set membership *) + emptyset :: "c" ("{}" 100) (* empty set *) + SUBSET :: "c \ c" ("SUBSET _" [100]90) (* power set *) + UNION :: "c \ c" ("UNION _" [100]90) (* generalized union *) + subsetOf :: "[c, c \ c] \ c" (* subsetOf(S,p) = {x \in S : p} *) + setOfAll :: "[c, c \ c] \ c" (* setOfAll(S,e) = {e : x \in S} *) + infinity :: "c" (* infinity set *) + +notation (ASCII) + "in" (infixl "\\in" 50) + +text \Negated set membership\ + +abbreviation "notin(a, S) \ \(a \ S)" notation - "notin" (infixl "\\notin" 50) -notation (xsymbols) - "notin" (infixl "\" 50) -notation (HTML output) "notin" (infixl "\" 50) +notation (ASCII) + "notin" (infixl "\\notin" 50) -abbreviation (input) "supseteq(S, T) \ T \ S" -notation - "supseteq" (infixl "\\supseteq" 50) -notation (xsymbols) - "supseteq" (infixl "\" 50) -notation (HTML output) - "supseteq" (infixl "\" 50) +text \ + We introduce some more notation and operators before stating the + axioms of set theory, in order to make them more readable. +\ -text {* Concrete syntax: proper sub and superset *} +text \ Concrete syntax: set comprehension \ -definition "psubset" :: "[c, c] \ c" (infixl "\\subset" 50) -where "S \\subset T \ S \ T \ S \ T" +(*** + We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} + because logical constants (such as setOfAll) must have fixed arity. + Multiple bindings will be introduced as shorthands involving tuples + when tuples are available -- fortunately we don't need them earlier! -notation (xsymbols) - "psubset" (infixl "\" 50) -notation (HTML output) - "psubset" (infixl "\" 50) + Note that the expression "{x \ S : y \ T}" is ambiguous + (as mentioned in the TLA+ book). In such cases, use logical syntax, e.g. + subsetOf(S, \x. y \ T). See the definition cap_def below. +***) +syntax + "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \ _})") + "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \ _ : _})") +syntax (ASCII) + "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \\in _})") + "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \\in _ : _})") +translations + "{e : x \ S}" \ "CONST setOfAll(S, \x. e)" + "{x \ S : P}" \ "CONST subsetOf(S, \x. P)" -abbreviation (input) "psupset(S,T) \ T \ S" -notation - "psupset" (infix "\\supset" 50) -notation (xsymbols) - "psupset" (infix "\" 50) -notation (HTML output) - "psupset" (infix "\" 50) +text \Unordered pairs\ -lemma psubset_intro [intro!]: - "\ S \ T ; S \ T \ \ S \ T" -unfolding psubset_def by safe +definition upair :: "[c, c] \ c" + where "upair(a,b) \ { IF x={} THEN a ELSE b : x \ SUBSET (SUBSET {}) }" -lemma psubset_elim [elim!]: - "\ S \ T ; \ S \ T ; S \ T \ \ C \ \ C" -unfolding psubset_def by safe +text \Binary set union\ + +definition union :: "[c, c] \ c" (infixl "\" 65) + where "A \ B \ UNION upair(A,B)" +notation (ASCII) + union (infixl "\\union" 65) -text {* Concrete syntax: set enumerations *} +definition addElt :: "[c, c] \ c" + where "addElt(a, A) \ upair(a,a) \ A" + +text \ Concrete syntax: set enumerations \ nonterminal cs @@ -117,7 +102,18 @@ translations abbreviation BOOLEAN :: "c" where "BOOLEAN \ {TRUE, FALSE}" -text {* Concrete syntax: bounded quantification *} +text \Bounded quantification\ + +definition bChoose :: "[c, c \ c] \ c" + where "bChoose(A,P) \ CHOOSE x : x \ A \ P(x)" + +definition bEx :: "[c, c \ c] \ c" + where "bEx(A, P) \ \x : x \ A \ P(x)" + +definition bAll :: "[c, c \ c] \ c" + where "bAll(A, P) \ \x : x \ A \ P(x)" + +text \ Concrete syntax: bounded quantification \ (*** FIXME: allow multiple bindings as in "ALL x:S, y:T. P(x,y)". The following was a valiant attempt to define appropriate syntax @@ -127,23 +123,23 @@ text {* Concrete syntax: bounded quantification *} nonterminal sbinds syntax - "@sbind" :: "[idt,c] \ sbinds" ("_ \\in _") - "@sbinds" :: "[idt,c,sbinds] \ sbinds" ("_ \\in _,/ _") + "@sbind" :: "[idt,c] \ sbinds" ("_ \\in _" [100, 10] 100) + "@sbinds" :: "[idt,c,sbinds] \ sbinds" ("_ \\in _,/ _" [100, 10, 100] 100) syntax (xsymbols) - "@sbind" :: "[idt,c] \ sbinds" ("_ \ _") - "@sbinds" :: "[idt,c, sbinds] \ sbinds" ("_ \ _,/ _") + "@sbind" :: "[idt,c] \ sbinds" ("_ \ _" [100, 10] 100) + "@sbinds" :: "[idt,c, sbinds] \ sbinds" ("_ \ _,/ _" [100, 10, 100] 100) syntax "@bChoice" :: "[sbinds, c] \ c" ("(3CHOOSE _ :/ _)" 10) - "@bEx" :: "[sbinds, c] \ c" ("(3\\E _ :/ _)" 10) - "@bAll" :: "[sbinds, c] \ c" ("(3\\A _ :/ _)" 10) + "@bEx" :: "[sbinds, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) + "@bAll" :: "[sbinds, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) syntax (xsymbols) - "@bEx" :: "[sbinds, c] \ c" ("(3\_ :/ _)" 10) - "@bAll" :: "[sbinds, c] \ c" ("(3\_ :/ _)" 10) + "@bEx" :: "[sbinds, c] \ c" ("(3\_ :/ _)" [100,10] 10) + "@bAll" :: "[sbinds, c] \ c" ("(3\_ :/ _)" [100,10] 10) translations - "ALL x:S, sbs. P" \ "bAll(S, (ALL sbs. (\x. P)))" - "ALL x:S. P" \ "bAll(S, \x. P)" + "\x \ S, bds : P" \ "CONST bAll(S, \x. \bds : P)" + "\x \ S : P" \ "CONST bAll(S, \x. P)" ***) (*** TEMPORARY FIX: @@ -152,16 +148,16 @@ translations ***) syntax - (* Again, only single variable for bounded choice. *) - "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \\in _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) -syntax (xsymbols) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) +syntax (ASCII) + (* Again, only single variable for bounded choice. *) + "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) translations - "CHOOSE x \ S : P" \ "CONST bChoice(S, \x. P)" + "CHOOSE x \ S : P" \ "CONST bChoose(S, \x. P)" (* the following cannot be a print macro because the RHS is non-linear -- need a print translation for display *) "\x, xs \ S : P" \ "CONST bEx(S, \x. \xs \ S : P)" @@ -169,7 +165,7 @@ translations "\x, xs \ S : P" \ "CONST bAll(S, \x. \xs \ S : P)" "\x \ S : P" \ "CONST bAll(S, \x. P)" -print_translation {* +print_translation \ let fun bEx_tr' [S, Abs(x, T, P as (Const (@{const_syntax "bEx"},_) $ S' $ Q))] = (* bEx(S, bEx(S', Q)) => \\E x,y \\in S : Q if S = S' *) @@ -201,9 +197,10 @@ print_translation {* in (Syntax.const "@bAll") $ x' $ S $ P' end | bAll_tr' _ = raise Match; - in [(@{const_syntax "bEx"}, bEx_tr'), (@{const_syntax "bAll"}, bAll_tr')] + in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), + (@{const_syntax "bAll"}, (fn _ => bAll_tr'))] end -*} +\ (*** TEST DATA for parse and print translations *** lemma "\x,y,z \ S : x = y \ y = z" @@ -222,90 +219,86 @@ lemma "\x,y \ S : \z \ T : x=y \ y \ z" oops *** END TEST DATA ***) -text {* Concrete syntax: set comprehension *} - -(*** - We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} - because logical constants (such as setOfAll) must have fixed arity. - Multiple bindings will be introduced as shorthands involving tuples - when tuples are available -- fortunately we don't need them earlier! - - Note that the expression "{x \ S : y \ T}" is ambiguous - (as mentioned in the TLA+ book). In such cases, use logical syntax, e.g. - subsetOf(S, \x. y \ T). See the definition cap_def below. -***) -syntax - "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \\in _})") - "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \\in _ : _})") -syntax (xsymbols) - "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \ _})") - "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \ _ : _})") -translations - "{e : x \ S}" \ "CONST setOfAll(S, \x. e)" - "{x \ S : P}" \ "CONST subsetOf(S, \x. P)" - -text {* - The following definitions make the axioms of set theory more readable. - Observe that @{text \} is treated as an uninterpreted predicate symbol. -*} - -defs - bChoose_def: "bChoice(A,P) \ CHOOSE x : x \ A \ P(x)" - bEx_def: "bEx(A, P) \ \x : x \ A \ P(x)" - bAll_def: "bAll(A, P) \ \x : x \ A \ P(x)" - subset_def: "A \ B \ \x \ A : x \ B" - -text {* - We now state a first batch of axioms of set theory: extensionality +text \ + We now state the axioms of set theory: extensionality and the definitions of @{text UNION}, @{text SUBSET}, and @{text setOfAll}. Membership is also asserted to be produce Boolean values---in traditional presentations of ZF set theory this is ensured by distinguishing sorts of - terms and formulas. -*} - + terms and formulas. @{term infinity} is some infinite set, but it is not + uniquely defined. The foundation axiom rules out sets that are ``too big''. +\ axiomatization where inIsBool [intro!,simp]: "isBool(x \ A)" and - extension: "(A = B) \ (A \ B) \ (B \ A)" + extension: "(A = B) \ (\x : x \ A \ x \ B)" +and + UNION: "(A \ UNION S) \ (\B \ S : A \ B)" and - UNION: "(A \ UNION S) \ (\B\S : A \ B)" + SUBSET: "(A \ SUBSET S) \ (\x \ A : x \ S)" and - SUBSET: "(A \ SUBSET S) \ (A \ S)" + setOfAll: "(y \ setOfAll(S, e)) \ (\x : x \ S \ y = e(x))" and - setOfAll: "(y \ { e(x) : x \ S }) \ (\x\S : y = e(x))" + subsetOf: "(y \ subsetOf(S, P)) \ (y \ S \ P(y))" and - subsetOf: "(y \ { x \ S : P(x) }) \ (y \ S \ P(y))" + infinity: "({} \ infinity) \ (\x \ infinity : {x} \ x \ infinity)" +and + foundation: "(A = {}) \ (\x \ A : \y \ x : y \ A)" -text {* - Armed with this understanding, we can now define the remaining operators - of set theory. -*} -defs - upair_def: "upair(a,b) \ { IF x={} THEN a ELSE b : x \ SUBSET (SUBSET {}) }" - cup_def: "A \ B \ UNION upair(A,B)" - addElt_def: "addElt(a, A) \ upair(a,a) \ A" - cap_def: "A \ B \ subsetOf(A, \x. x \ B)" - diff_def: "A \ B \ {x \ A : x \ B}" - INTER_def: "INTER A \ { x \ UNION A : \B \ A : x \ B }" +text \More set operations\ -text {* - The following two axioms complete our presentation of set theory. -*} +definition subset :: "[c, c] \ c" (infixl "\" 50) + where "A \ B \ \x \ A : x \ B" -axiomatization where - -- {* @{term infinity} is some infinite set, but it is not uniquely defined. *} - infinity: "({} \ infinity) \ (\x \ infinity : {x} \ x \ infinity)" -and - -- {* The foundation axiom rules out sets that are ``too big''. *} - foundation: "(A = {}) \ (\x \ A : \y \ x : y \ A)" +definition inter :: "[c, c] \ c" (infixl "\" 70) (* binary intersection *) + where "A \ B \ subsetOf(A, \x. x \ B)" + +definition setminus :: "[c, c] \ c" (infixl "\" 65) (* set difference *) + where "A \ B \ {x \ A : x \ B}" + +definition INTER :: "c \ c" ("INTER _" [100]90) (* generalized intersection *) + where "INTER A \ { x \ UNION A : \B \ A : x \ B }" + +text \Note that @{text "INTER {}"} is @{text "{}"}.}\ + +notation (ASCII) + "inter" (infixl "\\inter" 70) and + "setminus" (infixl "\\" 65) and + "subset" (infixl "\\subseteq" 50) + +abbreviation (input) "supset(S, T) \ T \ S" +notation + "supset" (infixl "\" 50) +notation (ASCII) + "supset" (infixl "\\supseteq" 50) + +definition "psubset" :: "[c, c] \ c" (infixl "\" 50) +where "S \ T \ S \ T \ S \ T" + +notation (ASCII) + "psubset" (infixl "\\subset" 50) -subsection {* Boolean operators *} +abbreviation (input) "psupset(S,T) \ T \ S" +notation + "psupset" (infix "\" 50) +notation (ASCII) + "psupset" (infix "\\supset" 50) + +lemma psubset_intro [intro!]: + "\ S \ T ; S \ T \ \ S \ T" +unfolding psubset_def by safe -text {* +lemma psubset_elim [elim!]: + "\ S \ T ; \ S \ T ; S \ T \ \ C \ \ C" +unfolding psubset_def by safe + + +subsection \ Boolean operators \ + +text \ The following lemmas assert that certain operators always return Boolean values; these are helpful for the automated reasoning methods. -*} +\ lemma boolifyIn [simp]: "boolify(x \ A) = (x \ A)" by (rule inIsBool[unfolded isBool_def]) @@ -331,6 +324,12 @@ by (simp add: subset_def) lemma subsetIsBool [intro!,simp]: "isBool(A \ B)" by (unfold isBool_def, rule boolifySubset) +lemma boolifypSubset [simp]: "boolify(A \ B) = (A \ B)" +by (simp add: psubset_def) + +lemma psubsetIsBool [intro!,simp]: "isBool(A \ B)" +by (unfold isBool_def, rule boolifypSubset) + lemma [intro!]: "\isBool(P); x\S \ P\ \ (x\S) = P" "\isBool(P); P \ x\S\ \ P = (x\S)" @@ -343,7 +342,7 @@ lemma [intro!]: by auto -subsection {* Substitution rules *} +subsection \ Substitution rules \ lemma subst_elem [trans]: assumes "b \ A" and "a=b" @@ -366,7 +365,7 @@ lemma subst_set_rev [trans]: using assms by simp -subsection {* Bounded quantification *} +subsection \ Bounded quantification \ lemma bAllI [intro!]: assumes "\x. x \ A \ P(x)" @@ -379,8 +378,8 @@ lemma bspec [dest?]: using assms unfolding bAll_def by blast (*** currently inactive -- causes several obscure warnings about renamed - variables and seems to slow down large examples such as AtomicBakery -ML {* + variables and seems to slow down large examples such as AtomicBakery ***) +ML \ structure Simpdata = struct @@ -391,12 +390,15 @@ ML {* end; open Simpdata; -*} +\ -declaration {* fn _ => - Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) -*} -***) +declaration \ + fn _ => Simplifier.map_ss (fn ss => ss + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +\ + +(***) lemma bAllE [elim]: assumes "\x\A : P(x)" and "x \ A \ Q" and "P(x) \ Q" @@ -416,7 +418,7 @@ lemma bExI [intro]: shows "\x\A : P(x)" using assms unfolding bEx_def by blast -lemma bExCI: -- {* implicit proof by contradiction *} +lemma bExCI: (* proof by contradiction *) assumes "(\x\A : \P(x)) \ P(a)" and "a \ A" shows "\x\A : P(x)" using assms by blast @@ -499,9 +501,9 @@ qed TODO: investigate the issue and find out what is really needed about type checking. -subsubsection {* Setting up type checking for \tlaplus{}. *} +subsubsection \ Setting up type checking for \tlaplus{}. \ -text {* +text \ Type checking is intended to solve simple goals, for example showing that conjunctions are Boolean, and that the result of a conditional expression is a natural number if both branches yield naturals. Unlike @@ -512,7 +514,7 @@ text {* unknowns. The code is essentially taken from Isabelle/ZF. -*} +\ use "typecheck.ML" setup TypeCheck.setup @@ -538,7 +540,7 @@ by typecheck ***) -subsection {* Simplification of conditional expressions *} +subsection \ Simplification of conditional expressions \ lemma inCond [simp]: "(a \ (IF P THEN S ELSE T)) = ((P \ a \ S) \ (\P \ a \ T))" by (force intro: condI elim: condE) @@ -583,7 +585,7 @@ by (force intro: condI elim: condE) **) -subsection {* Rules for subsets and set equality *} +subsection \ Rules for subsets and set equality \ lemma subsetI [intro!]: assumes "\x. x \ A \ x \ B" @@ -599,7 +601,7 @@ lemma rev_subsetD [trans]: "\ c \ A; A \ B \ \ c \ B" by (rule subsetD) -lemma subsetCE [elim]: -- {* elimination rule for classical logic *} +lemma subsetCE [elim]: assumes "A \ B" and "c \ A \ P" and "c \ B \ P" shows "P" using assms unfolding subset_def by blast @@ -630,7 +632,7 @@ lemma notSubsetE (*[elim!]*): shows "P" using assms by blast -text {* The subset relation is a partial order. *} +text \ The subset relation is a partial order. \ lemma subsetRefl [simp,intro!]: "A \ A" by blast @@ -650,45 +652,48 @@ lemma setEqualI: shows "A = B" by (rule setEqual, (blast intro: assms)+) -text {* +text \ The rule @{text setEqualI} is too general for use as a default introduction rule: we don't want to apply it for Booleans, for example. However, instances where at least one term results from a set constructor are useful. -*} +\ lemmas (** the instances for the empty set are superseded by rules below - setEqualI [where A = "{}", standard, intro!] - setEqualI [where B = "{}", standard, intro!] + setEqualI [where A = "{}", intro!] + setEqualI [where B = "{}", intro!] **) - setEqualI [where A = "addElt(a,C)", standard, intro!] - setEqualI [where B = "addElt(a,C)", standard, intro!] - setEqualI [where A = "SUBSET C", standard, intro!] - setEqualI [where B = "SUBSET C", standard, intro!] - setEqualI [where A = "UNION C", standard, intro!] - setEqualI [where B = "UNION C", standard, intro!] - setEqualI [where A = "INTER C", standard, intro!] - setEqualI [where B = "INTER C", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "subsetOf(S,P)", standard, intro!] - setEqualI [where B = "subsetOf(S,P)", standard, intro!] - setEqualI [where A = "setOfAll(S,e)", standard, intro!] - setEqualI [where B = "setOfAll(S,e)", standard, intro!] - -lemmas setEqualD1 = extension[THEN iffD1, THEN conjD1, standard] -- {* @{text "A = B \ A \ B"} *} -lemmas setEqualD2 = extension[THEN iffD1, THEN conjD2, standard] -- {* @{text "A = B \ B \ A"} *} - -text {* - We declare the elimination rule for set equalities as an unsafe + setEqualI [where A = "addElt(a,C)" for a C, intro!] + setEqualI [where B = "addElt(a,C)" for a C, intro!] + setEqualI [where A = "SUBSET C" for C, intro!] + setEqualI [where B = "SUBSET C" for C, intro!] + setEqualI [where A = "UNION C" for C, intro!] + setEqualI [where B = "UNION C" for C, intro!] + setEqualI [where A = "INTER C" for C, intro!] + setEqualI [where B = "INTER C" for C, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "subsetOf(S,P)" for S P, intro!] + setEqualI [where B = "subsetOf(S,P)" for S P, intro!] + setEqualI [where A = "setOfAll(S,e)" for S e, intro!] + setEqualI [where B = "setOfAll(S,e)" for S e, intro!] + +lemma setEqualD1: "A = B \ A \ B" + by blast + +lemma setEqualD2: "A = B \ B \ A" + by blast + +text \ + We declare the elimination rule for set equations as an unsafe rule to use with the classical reasoner, so it will be tried if the more obvious uses of equality fail. -*} +\ lemma setEqualE [elim]: assumes "A = B" and "\ c \ A; c \ B \ \ P" and "\ c \ A; c \ B \ \ P" @@ -696,10 +701,10 @@ lemma setEqualE [elim]: using assms by (blast dest: setEqualD1 setEqualD2) lemma setEqual_iff: "(A = B) = (\x : x \ A \ x \ B)" -by (blast intro: setEqualI (*elim: setEqualE*)) +by (blast intro: setEqualI) -subsection {* Set comprehension: @{text setOfAll} and @{text subsetOf} *} +subsection \ Set comprehension: @{text setOfAll} and @{text subsetOf} \ (*** lemma setOfAllI: @@ -735,10 +740,10 @@ lemma setOfAll_cong (*[cong]*): shows "{ e(x) : x \ S } = { f(y) : y \ T }" using assms by auto -text {* +text \ The following rule for showing equality of sets defined by comprehension is probably too general to use by default with the automatic proof methods. -*} +\ lemma setOfAllEqual: assumes "\x. x \ S \ \y\T : e(x) = f(y)" @@ -784,7 +789,7 @@ lemma subsetOfEqual: by (safe elim!: assms) -subsection {* @{text UNION} -- basic rules for generalized union *} +subsection \ @{text UNION} -- basic rules for generalized union \ lemma UNIONI [intro]: assumes "B \ C" and "a \ B" @@ -801,14 +806,14 @@ lemma UNION_iff [simp]: by blast -subsection {* The empty set *} +subsection \ The empty set \ -text {* +text \ Proving that the empty set has no elements is a bit tricky. We first show that the set @{term "{x \ {} : FALSE}"} is empty and then use the foundation axiom to show that it equals the empty set. -*} +\ lemma emptysetEmpty: "a \ {}" proof @@ -820,8 +825,8 @@ proof thus "FALSE" by blast qed -text {* @{text "a \ {} \ P"} *} -lemmas emptyE [elim!] = emptysetEmpty[THEN notE, standard] +text \ @{text "a \ {} \ P"} \ +lemmas emptyE [elim!] = emptysetEmpty[THEN notE] lemma [simp]: "(a \ {}) = FALSE" by blast @@ -844,10 +849,10 @@ lemma nonEmpty [simp]: "({} \ A) = (\x : x \ A)" by (blast+) -text {* Complete critical pairs *} +text \ Complete critical pairs \ lemmas nonEmpty' [simp] = nonEmpty[simplified] --- {* @{text "((A = {}) = FALSE) = (\x : x \ A)"}, - @{text "(({} = A) = FALSE) = (\x : x \ A)"} *} +\ \ @{text "((A = {}) = FALSE) = (\x : x \ A)"}, + @{text "(({} = A) = FALSE) = (\x : x \ A)"} \ lemma emptysetD (*[dest]*): assumes "A = {}" @@ -878,7 +883,7 @@ lemma [simp]: by (blast+) -subsection {* @{text SUBSET} -- the powerset operator *} +subsection \ @{text SUBSET} -- the powerset operator \ lemma SUBSETI [intro!]: assumes "A \ B" @@ -894,18 +899,18 @@ lemma SUBSET_iff [simp]: "(A \ SUBSET B) = (A \ B)" by blast -lemmas emptySUBSET = emptySubset[THEN SUBSETI, standard] -- {* @{term "{} \ SUBSET A"} *} -lemmas selfSUBSET = subsetRefl[THEN SUBSETI, standard] -- {* @{term "A \ SUBSET A"} *} +lemmas emptySUBSET = emptySubset[THEN SUBSETI] \ \ @{term "{} \ SUBSET A"} \ +lemmas selfSUBSET = subsetRefl[THEN SUBSETI] \ \ @{term "A \ SUBSET A"} \ -subsection {* @{text INTER} -- basic rules for generalized intersection *} +subsection \ @{text INTER} -- basic rules for generalized intersection \ -text {* +text \ Generalized intersection is not officially part of \tlaplus{} but can easily be defined as above. Observe that the rules are not exactly dual to those for @{text UNION} because the intersection of the empty set is defined to be the empty set. -*} +\ lemma INTERI [intro]: assumes "\B. B \ C \ a \ B" and "\B : B \ C" @@ -922,62 +927,49 @@ lemma INTER_iff [simp]: by blast -subsection {* Binary union, intersection, and difference: basic rules *} +subsection \ Binary union, intersection, and difference: basic rules \ -text {* +text \ We begin by proving some lemmas about the auxiliary pairing operator @{text upair}. None of these theorems is active by default, as the operator is not part of \tlaplus{} and should not occur in actual reasoning. The dependencies between these operators are quite tricky, therefore the order of the first few lemmas in this section is tightly constrained. -*} +\ lemma upairE: assumes "x \ upair(a,b)" and "x=a \ P" and "x=b \ P" shows "P" using assms by (auto simp: upair_def) -lemma cupE [elim!]: +lemma unionE [elim!]: assumes "x \ A \ B" and "x \ A \ P" and "x \ B \ P" shows "P" -using assms by (auto simp: cup_def elim: upairE) +using assms by (auto simp: union_def elim: upairE) lemma upairI1: "a \ upair(a,b)" by (auto simp: upair_def) -lemma singleton_iff [simp]: "(a \ {b}) = (a = b)" -proof auto - assume a: "a \ {b}" - thus "a = b" - by (auto simp: addElt_def elim: upairE) -next - show "b \ {b}" - by (auto simp: addElt_def cup_def intro: upairI1) -qed - -lemma singletonI (*[intro!]*): "a \ {a}" -by simp - lemma upairI2: "b \ upair(a,b)" by (auto simp: upair_def) lemma upair_iff: "(c \ upair(a,b)) = (c=a \ c=b)" by (blast intro: upairI1 upairI2 elim: upairE) -lemma cupI1: +lemma unionI1: assumes "a \ A" shows "a \ A \ B" -using assms by (auto simp: cup_def upair_iff) +using assms by (auto simp: union_def upair_iff) -lemma cupI2: +lemma unionI2: assumes "a \ B" shows "a \ A \ B" -using assms by (auto simp: cup_def upair_iff) +using assms by (auto simp: union_def upair_iff) -lemma cup_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (auto simp: cup_def upair_iff) +lemma union_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (auto simp: union_def upair_iff) -lemma cupCI [intro!]: +lemma unionCI [intro!]: assumes "c \ B \ c \ A" shows "c \ A \ B" using assms by auto @@ -996,6 +988,12 @@ lemma addEltE [elim!]: shows "P" using assms by auto +lemma singleton_iff (*[simp]*): "(a \ {b}) = (a = b)" + by auto + +lemma singletonI (*[intro!]*): "a \ {a}" +by simp + lemma addEltSubsetI: assumes "a \ B" and "A \ B" shows "addElt(a,A) \ B" @@ -1009,7 +1007,7 @@ using assms by blast lemma addEltSubset_iff: "(addElt(a,A) \ B) = (a \ B \ A \ B)" by blast --- {* Adding the two following lemmas to the simpset breaks proofs. *} +text \ Adding the two following lemmas to the simpset breaks proofs. \ lemma addEltEqual_iff: "(addElt(a,A) = S) = (a \ S \ A \ S \ S \ addElt(a,A))" by blast @@ -1021,48 +1019,48 @@ lemma addEltEqualAddElt: (a \ addElt(b,B) \ A \ addElt(b,B) \ b \ addElt(a,A) \ B \ addElt(a,A))" by (auto simp: addEltEqual_iff) -lemma cap_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (simp add: cap_def) +lemma inter_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (simp add: inter_def) -lemma capI [intro!]: +lemma interI [intro!]: assumes "c \ A" and "c \ B" shows "c \ A \ B" using assms by simp -lemma capD1: +lemma interD1: assumes "c \ A \ B" shows "c \ A" using assms by simp -lemma capD2: +lemma interD2: assumes "c \ A \ B" shows "c \ B" using assms by simp -lemma capE [elim!]: +lemma interE [elim!]: assumes "c \ A \ B" and "\ c \ A; c \ B \ \ P" shows "P" using assms by simp -lemma diff_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (simp add: diff_def) +lemma setminus_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (simp add: setminus_def) -lemma diffI [intro!]: +lemma setminusI [intro!]: assumes "c \ A" and "c \ B" shows "c \ A \ B" using assms by simp -lemma diffD1: +lemma setminusD1: assumes "c \ A \ B" shows "c \ A" using assms by simp -lemma diffD2: +lemma setminusD2: assumes "c \ A \ B" shows "c \ B" using assms by simp -lemma diffE [elim!]: +lemma setminusE [elim!]: assumes "c \ A \ B" and "\ c \ A; c \ B \ \ P" shows "P" using assms by simp @@ -1097,7 +1095,7 @@ lemma subsetAddEltE [elim]: using assms by (auto simp: subsetAddElt_iff) -subsection {* Consequences of the foundation axiom *} +subsection \ Consequences of the foundation axiom \ lemma inAsym: assumes hyps: "a \ b" "\P \ b \ a" @@ -1123,7 +1121,7 @@ lemma equalNotIn: using assms by (blast elim: inIrrefl) -subsection {* Miniscoping of bounded quantifiers *} +subsection \ Miniscoping of bounded quantifiers \ lemma miniscope_bAll [simp]: "\P Q. (\x\A : P(x) \ Q) = ((\x\A : P(x)) \ (A = {} \ Q))" @@ -1139,11 +1137,11 @@ lemma miniscope_bAll [simp]: "\P Q. (\x \ {y \ A : P(y)} : Q(x)) = (\y\A: P(y) \ Q(y))" by (blast+) -lemma bAllCup [simp]: +lemma bAllUnion [simp]: "(\x \ A \ B : P(x)) = ((\x\A : P(x)) \ (\x\B : P(x)))" by blast -lemma bAllCap [simp]: +lemma bAllInter [simp]: "(\x \ A \ B : P(x)) = (\x\A : x \ B \ P(x))" by blast @@ -1161,21 +1159,21 @@ lemma miniscope_bEx [simp]: "\P Q. (\x \ {y \ S : P(y)} : Q(x)) = (\y\S : P(y) \ Q(y))" by (blast+) --- {* completing critical pairs for negated assumption *} +text \ completing critical pairs for negated assumption \ lemma notbQuant' [simp]: "\P. ((\x\A : P(x)) = FALSE) = (\x\A : \P(x))" "\P. ((\x\A : P(x)) = FALSE) = (\x\A : \P(x))" by (auto simp: miniscope_bAll[simplified] miniscope_bEx[simplified]) -lemma bExistsCup [simp]: +lemma bExistsUnion [simp]: "(\x \ A \ B : P(x)) = ((\x\A : P(x)) \ (\x\B : P(x)))" by blast -lemma bExistsCap [simp]: +lemma bExistsInter [simp]: "(\x \ A \ B : P(x)) = (\x\A : x \ B \ P(x))" by blast -lemma bQuant_distribs: -- {* not active by default *} +lemma bQuant_distribs: \ \ not active by default \ "(\x\A : P(x) \ Q(x)) = ((\x\A : P(x)) \ (\x\A : Q(x)))" "(\x\A : P(x) \ Q(x)) = ((\x\A : P(x)) \ (\x\A : Q(x)))" by (blast+) @@ -1190,7 +1188,7 @@ lemma bQuantOnePoint [simp]: by (blast+) -subsection {* Simplification of set comprehensions *} +subsection \ Simplification of set comprehensions \ lemma comprehensionSimps [simp]: "\e. setOfAll({}, e) = {}" @@ -1206,24 +1204,24 @@ lemma comprehensionSimps [simp]: "\A e. setOfAll(A, \x. e) = (IF A = {} THEN {} ELSE {e})" by auto -text {* The following are not active by default. *} +text \ The following are not active by default. \ lemma comprehensionDistribs: "\e. { e(x) : x \ A \ B } = {e(x) : x\A} \ {e(x) : x\B}" "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" - -- {* @{text setOfAll} and intersection or difference do not distribute *} + \ \ @{text setOfAll} and intersection or difference do not distribute \ "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" by (blast+) -subsection {* Binary union, intersection, and difference: inclusions and equalities *} +subsection \ Binary union, intersection, and difference: inclusions and equations \ -text {* +text \ The following list contains many simple facts about set theory. Only the most trivial of these are included in the default set of rewriting rules. -*} +\ lemma addEltCommute: "addElt(a, addElt(b, C)) = addElt(b, addElt(a, C))" by blast @@ -1234,59 +1232,59 @@ by blast lemma addEltTwice: "addElt(a, addElt(a, A)) = addElt(a, A)" by blast -lemma capAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN addElt(a, B \ C) ELSE B \ C)" +lemma interAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN addElt(a, B \ C) ELSE B \ C)" by (blast intro: condI elim: condE) -lemma capAddEltRight: "C \ addElt(a,B) = (IF a \ C THEN addElt(a, C \ B) ELSE C \ B)" +lemma interAddEltRight: "C \ addElt(a,B) = (IF a \ C THEN addElt(a, C \ B) ELSE C \ B)" by (blast intro: condI elim: condE) -lemma addEltCap: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" +lemma addEltInter: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" by blast -lemma diffAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN B \ C ELSE addElt(a, B \ C))" +lemma setminusAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN B \ C ELSE addElt(a, B \ C))" by (blast intro: condI elim: condE) -lemma capSubset: "(C \ A \ B) = (C \ A \ C \ B)" +lemma interSubset: "(C \ A \ B) = (C \ A \ C \ B)" by blast -lemma capLB1: "A \ B \ A" +lemma interLB1: "A \ B \ A" by blast -lemma capLB2: "A \ B \ B" +lemma interLB2: "A \ B \ B" by blast -lemma capGLB: +lemma interGLB: assumes "C \ A" and "C \ B" shows "C \ A \ B" using assms by blast -lemma capEmpty [simp]: +lemma interEmpty [simp]: "A \ {} = {}" "{} \ A = {}" by blast+ -lemma capAbsorb [simp]: "A \ A = A" +lemma interAbsorb [simp]: "A \ A = A" by blast -lemma capLeftAbsorb: "A \ (A \ B) = A \ B" +lemma interLeftAbsorb: "A \ (A \ B) = A \ B" by blast -lemma capCommute: "A \ B = B \ A" +lemma interCommute: "A \ B = B \ A" by blast -lemma capLeftCommute: "A \ (B \ C) = B \ (A \ C)" +lemma interLeftCommute: "A \ (B \ C) = B \ (A \ C)" by blast -lemma capAssoc: "(A \ B) \ C = A \ (B \ C)" +lemma interAssoc: "(A \ B) \ C = A \ (B \ C)" by blast -text {* Intersection is an AC operator: can be added to simp where appropriate *} -lemmas capAC = capAssoc capCommute capLeftCommute capLeftAbsorb +text \ Intersection is an AC operator: can be added to simp where appropriate \ +lemmas interAC = interAssoc interCommute interLeftCommute interLeftAbsorb -lemma subsetOfCap: "{x\A : P(x)} \ B = {x \ A \ B : P(x)}" +lemma subsetOfInter: "{x\A : P(x)} \ B = {x \ A \ B : P(x)}" by blast -lemma capSubsetOf: +lemma interSubsetOf: "B \ {x\A : P(x)} = {x \ B\A : P(x)}" by blast @@ -1298,113 +1296,115 @@ lemma subsetOfConj: "{x\A : P(x) \ Q(x)} = {x\A : P(x)} \ {x\A : Q(x)}" by blast -lemma subsetCup: "(A \ B \ C) = (A \ C \ B \ C)" +lemma subsetUnion: "(A \ B \ C) = (A \ C \ B \ C)" by blast -lemma cupUB1: "A \ A \ B" +lemma unionUB1: "A \ A \ B" by blast -lemma cupUB2: "B \ A \ B" +lemma unionUB2: "B \ A \ B" by blast -lemma cupLUB: +lemma unionLUB: assumes "A \ C" and "B \ C" shows "A \ B \ C" using assms by blast -lemma cupEmpty [simp]: +lemma unionEmpty [simp]: "A \ {} = A" "{} \ A = A" by blast+ -lemma cupAddEltLeft: "addElt(a,B) \ C = addElt(a, B \ C)" +lemma unionAddEltLeft: "addElt(a,B) \ C = addElt(a, B \ C)" by blast -lemma cupAddEltRight: "C \ addElt(a,B) = addElt(a, C \ B)" +lemma unionAddEltRight: "C \ addElt(a,B) = addElt(a, C \ B)" by blast -lemma addEltCup: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" +lemma addEltUnion: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" by blast -lemma cupAbsorb [simp]: "A \ A = A" +lemma unionAbsorb [simp]: "A \ A = A" by blast -lemma cupLeftAbsorb: "A \ (A \ B) = A \ B" +lemma unionLeftAbsorb: "A \ (A \ B) = A \ B" by blast -lemma cupCommute: "A \ B = B \ A" +lemma unionCommute: "A \ B = B \ A" by blast -lemma cupLeftCommute: "A \ (B \ C) = B \ (A \ C)" +lemma unionLeftCommute: "A \ (B \ C) = B \ (A \ C)" by blast -lemma cupAssoc: "(A \ B) \ C = A \ (B \ C)" +lemma unionAssoc: "(A \ B) \ C = A \ (B \ C)" by blast -text {* Union is an AC operator: can be added to simp where appropriate *} -lemmas cupAC = cupAssoc cupCommute cupLeftCommute cupLeftAbsorb +text \ Union is an AC operator: can be added to simp where appropriate \ +lemmas unionAC = unionAssoc unionCommute unionLeftCommute unionLeftAbsorb -text {* Lemmas useful for simplifying enumerated sets are active by default *} +text \ Lemmas useful for simplifying enumerated sets are active by default \ lemmas enumeratedSetSimps [simp] = addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice - capAddEltLeft capAddEltRight cupAddEltLeft cupAddEltRight diffAddEltLeft + interAddEltLeft interAddEltRight unionAddEltLeft unionAddEltRight setminusAddEltLeft -lemma cupEqualEmpty [simp]: "(A \ B = {}) = (A = {} \ B = {})" +(* included below +lemma unionEqualEmpty [simp]: "(A \ B = {}) = (A = {} \ B = {})" by blast +*) -lemma capCupDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma interUnionDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma cupCapDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma unionInterDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma diffSubset: "A \ B \ A" +lemma setminusSubset: "A \ B \ A" by blast -lemma subsetDiff: +lemma subsetSetminus: assumes "C \ A" and "C \ B = {}" shows "C \ A \ B" using assms by blast -lemma diffSelf [simp]: "A \ A = {}" +lemma setminusSelf [simp]: "A \ A = {}" by blast -lemma diffDisjoint: +lemma setminusDisjoint: assumes "A \ B = {}" shows "A \ B = A" using assms by blast -lemma emptyDiff [simp]: "{} \ A = {}" +lemma emptySetminus [simp]: "{} \ A = {}" by blast -lemma diffAddElt: "A \ addElt(a,B) = (A \ B) \ {a}" +lemma setminusAddElt: "A \ addElt(a,B) = (A \ B) \ {a}" by blast -lemma cupDiffSelf: "A \ (B \ A) = A \ B" +lemma unionSetminusSelf: "A \ (B \ A) = A \ B" by blast -lemma diffCupLeft: "(A \ B) \ C = (A \ C) \ (B \ C)" +lemma setminusUnionLeft: "(A \ B) \ C = (A \ C) \ (B \ C)" by blast -lemma diffCupRight: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma setminusUnionRight: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma cupDiffLeft: "(A \ B) \ C = (A \ C) \ (A \ (B \ C))" +lemma unionSetminusLeft: "(A \ B) \ C = (A \ C) \ (A \ (B \ C))" by blast -lemma cupDiffRight: "C \ (A \ B) = (C \ A) \ (A \ (B \ C))" +lemma unionSetminusRight: "C \ (A \ B) = (C \ A) \ (A \ (B \ C))" by blast -lemma diffCapLeft: "(A \ B) \ C = A \ (B \ C)" +lemma setminusInterLeft: "(A \ B) \ C = A \ (B \ C)" by blast -lemma diffCapRight: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma setminusInterRight: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma capDiffLeft: "(A \ B) \ C = (A \ C) \ B" +lemma interSetminusLeft: "(A \ B) \ C = (A \ C) \ B" by blast -lemma capDiffRight: "C \ (A \ B) = (C \ A) \ B" +lemma interSetminusRight: "C \ (A \ B) = (C \ A) \ B" by blast lemma isEmptySimps [simp]: @@ -1425,7 +1425,7 @@ lemma isEmptySimps [simp]: by (blast+) -subsection {* Generalized union: inclusions and equalities *} +subsection \ Generalized union: inclusions and equalities \ lemma UNIONSimps [simp]: "UNION {} = {}" @@ -1446,19 +1446,19 @@ lemma UNION_LUB: shows "UNION A \ C" using assms by blast -lemma UNIONCupDistrib: "UNION (A \ B) = UNION A \ UNION B" +lemma UNIONunionDistrib: "UNION (A \ B) = UNION A \ UNION B" by blast -lemma UNIONCap: "UNION (A \ B) \ UNION A \ UNION B" +lemma UNIONinter: "UNION (A \ B) \ UNION A \ UNION B" by blast lemma UNIONDisjoint: "((UNION A) \ C = {}) = (\B\A : B \ C = {})" by blast -lemma capUNION: "(UNION A) \ B = UNION { C \ B : C \ A }" +lemma interUNION: "(UNION A) \ B = UNION { C \ B : C \ A }" by blast -lemma diffUNIONLeft: "(UNION A) \ B = UNION { C \ B : C \ A }" +lemma setminusUNIONLeft: "(UNION A) \ B = UNION { C \ B : C \ A }" by blast lemma UNION_mono: @@ -1467,7 +1467,7 @@ lemma UNION_mono: using assms by blast -subsection {* Generalized intersection: inclusions and equalities *} +subsection \ Generalized intersection: inclusions and equations \ lemma INTERSimps [simp]: @@ -1491,23 +1491,23 @@ lemma INTER_GLB: shows "C \ INTER A" using assms by blast -lemma INTERCupDistrib: +lemma INTERunionDistrib: assumes "A \ {}" and "B \ {}" shows "INTER (A \ B) = INTER A \ INTER B" using assms by auto -lemma capINTER: "(INTER A) \ B \ INTER {C \ B : C \ A}" +lemma interINTER: "(INTER A) \ B \ INTER {C \ B : C \ A}" by blast -lemma cupINTER: +lemma unionINTER: assumes "A \ {}" shows "(INTER A) \ B = INTER {C \ B : C \ A}" using assms by auto -lemma diffINTERLeft: "(INTER A) \ B = INTER {C \ B : C \ A}" +lemma setminusINTERLeft: "(INTER A) \ B = INTER {C \ B : C \ A}" by auto -lemma diffINTERRight: +lemma setminusINTERRight: assumes "A \ {}" shows "B \ (INTER A) = UNION {B \ C : C \ A}" using assms by auto @@ -1523,16 +1523,16 @@ lemma INTER_anti_mono: using assms by (auto simp: INTER_def) -subsection {* Powerset: inclusions and equalities *} +subsection \ Powerset: inclusions and equations \ lemma SUBSETEmpty [simp]: "SUBSET {} = { {} }" by blast -lemma SUBSETAddElt: +lemma SUBSETaddElt: "SUBSET addElt(a,A) = SUBSET A \ {addElt(a,X) : X \ SUBSET A}" by (rule setEqualI, auto) -lemma cupSUBSET: "(SUBSET A) \ (SUBSET B) \ SUBSET (A \ B)" +lemma unionSUBSET: "(SUBSET A) \ (SUBSET B) \ SUBSET (A \ B)" by blast lemma UNION_SUBSET [simp]: "UNION (SUBSET A) = A" @@ -1544,7 +1544,7 @@ by blast lemma UNION_in_SUBSET: "(UNION A \ SUBSET B) = (A \ SUBSET (SUBSET B))" by blast -lemma SUBSETcap: "SUBSET (A \ B) = SUBSET A \ SUBSET B" +lemma SUBSETinter: "SUBSET (A \ B) = SUBSET A \ SUBSET B" by blast end diff --git a/isabelle/document/root.tex b/isabelle/document/root.tex index 25b77b15..4e2c13c4 100644 --- a/isabelle/document/root.tex +++ b/isabelle/document/root.tex @@ -31,8 +31,8 @@ \newcommand{\tlaplus}{\textrm{\upshape TLA\textsuperscript{+}}} \newcommand{\deq}{\stackrel{\Delta}{=}} -\renewcommand{\isacharbackslash}{\isamath{\backslash}} -\newcommand{\isactrlconst}{} +% \renewcommand{\isacharbackslash}{\isamath{\backslash}} +% \newcommand{\isactrlconst}{} \begin{document} diff --git a/isabelle/simplifier_setup.ML b/isabelle/simplifier_setup.ML index de942192..6957e6a5 100644 --- a/isabelle/simplifier_setup.ML +++ b/isabelle/simplifier_setup.ML @@ -1,9 +1,8 @@ (* Title: TLA+/simplifier_setup.ML - Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2009 - Time-stamp: <2011-10-11 14:59:15 merz> + Version: Isabelle2020 Code for setting up the simplifier for TLA+. The code is essentially copied from FOL and HOL. @@ -19,55 +18,54 @@ struct Negation and equivalence will be handled later. The implicit operator below is Trueprop. *) - fun mk_eq th = case concl_of th of - Const(@{const_name "=="},_) $ _ $ _ => th - | _ $ (Const(@{const_name "eq"},_) $ _ $ _) => th RS @{thm eq_reflection} - | _ $ (Const(@{const_name "Not"},_) $ (Const(@{const_name "eq"},_) $ _ $ _)) => - th RS (@{thm eqBoolean} RS @{thm eq_reflection}) - | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection}; + fun mk_eq th = case Thm.concl_of th of + Const(@{const_name \Pure.eq\},_) $ _ $ _ => th + | _ $ (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _) => th RS @{thm eq_reflection} + | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ + (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => + th RS (@{thm neq_conv_eq_False} RS @{thm eq_reflection}) + | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection} (* turn P into P == TRUE *) - fun mk_eq_True (_: simpset) th = + fun mk_eq_True (_: Proof.context) th = SOME ((th RS @{thm eqTrueI}) RS @{thm eq_reflection}) handle Thm.THM _ => NONE (* Produce theorems of the form (P1 =simp=> ... =simp=> Pn ==> x == y) ==> (P1 =simp=> ... =simp=> Pn ==> x = y) *) - fun lift_meta_to_obj_eq i st = + fun lift_meta_to_obj_eq ctxt i st = let - fun count_imp (Const (@{const_name PredicateLogic.simp_implies}, _) $ _ $ Q) = + fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = 1 + count_imp Q | count_imp _ = 0; - val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) + val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) in if j = 0 then @{thm meta_to_obj_eq} else let val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); fun mk_simp_implies Q = fold_rev (fn R => fn S => - Const (@{const_name PredicateLogic.simp_implies}, propT --> propT --> propT) $ R $ S) Ps Q - val cT = Type ("c", []); - val x = Free ("x", cT); - val y = Free ("y", cT); - val x_eq_y = Const (@{const_name Trueprop}, cT --> propT) $ (Const (@{const_name "eq"}, [cT, cT] ---> cT) $ x $ y) - in Goal.prove_global (Thm.theory_of_thm st) [] - [mk_simp_implies (Logic.mk_equals (x, y))] - (mk_simp_implies x_eq_y) - (fn {prems, ...} => EVERY - [rewrite_goals_tac @{thms simp_implies_def}, - REPEAT (ares_tac (@{thm meta_to_obj_eq} :: - map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) + Const (@{const_name \PredicateLogic.simp_implies\}, propT --> propT --> propT) $ R $ S) Ps Q + in Goal.prove_global (Proof_Context.theory_of ctxt) [] + [mk_simp_implies @{prop \(x::c) == y\}] + (mk_simp_implies @{prop \(x::c) = y\}) + (fn {context = ctxt, prems} => EVERY + [rewrite_goals_tac ctxt @{thms simp_implies_def}, + REPEAT (assume_tac ctxt 1 ORELSE + resolve_tac ctxt + (@{thm meta_to_obj_eq} :: + map (rewrite_rule ctxt @{thms simp_implies_def}) prems) 1)]) end - end; + end (* Congruence rules expecting = instead of == *) - fun mk_meta_cong (_: simpset) rl = zero_var_indexes - (let val rl' = Seq.hd (TRYALL (fn i => fn st => - rtac (lift_meta_to_obj_eq i st) i st) rl) - in (rl' RS @{thm eq_reflection}) + fun mk_meta_cong ctxt rl = + let val rl' = Seq.hd (TRYALL (fn i => fn st => + resolve_tac ctxt [lift_meta_to_obj_eq ctxt i st] i st) rl) + in (rl' RS @{thm eq_reflection}) handle THM _ => - if can Logic.dest_equals (concl_of rl') then rl' + if can Logic.dest_equals (Thm.concl_of rl') then rl' else error "Conclusion of congruence rules must use =-equality" - end); + end |> zero_var_indexes (* (string * (thm list * bool)) list The following list associates a list of theorems and a flag with @@ -78,71 +76,84 @@ struct allows us to discard "type-correctness" hypotheses in conditional rewrite rules that are added later on. *) val mksimps_pairs = - [(@{const_name imp}, ([@{thm mp}], false)), - (@{const_name conj}, ([@{thm conjD1}, @{thm conjD2}], false)), - (@{const_name All}, ([@{thm spec}], false)), - (@{const_name TRUE}, ([], false)), - (@{const_name FALSE}, ([], false)), - (@{const_name cond}, ([@{thm if_bool_eq_conj} RS @{thm iffD1}], false)) + [(@{const_name \imp\}, ([@{thm mp}], false)), + (@{const_name \conj\}, ([@{thm conjD1}, @{thm conjD2}], false)), + (@{const_name \All\}, ([@{thm spec}], false)), + (@{const_name \TRUE\}, ([], false)), + (@{const_name \FALSE\}, ([], false)), + (@{const_name \cond\}, ([@{thm if_bool_equiv_conj} RS @{thm iffD1}], false)) ]; (* val mk_atomize: (string * (thm list * bool)) list -> thm -> thm list *) - fun mk_atomize pairs = + fun mk_atomize ctxt pairs = let fun atoms thm = let - val simps = global_simpset_of (theory_of_thm thm); fun res th = map (fn rl => th RS rl); (* may raise exception THM *) + val thm_ctxt = Variable.declare_thm thm ctxt fun res_fixed rls = if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls - else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.global_thm_context thm) [thm]; + else Variable.trade (K (fn [thm'] => res thm' rls)) thm_ctxt [thm]; in - case concl_of thm of - Const(@{const_name Trueprop},_) $ p => + case Thm.concl_of thm of + Const(@{const_name \Trueprop\},_) $ p => (case head_of p of Const(a,_) => (case AList.lookup (op =) pairs a of SOME (rls, b) => let - val rls = maps atoms (res_fixed rls) handle THM _ => [thm] + val rrls = maps atoms (res_fixed rls) handle THM _ => [thm] in - if b then map (simplify simps) rls - else rls + if b then map (simplify ctxt) rrls + else rrls end | NONE => [thm]) | _ => [thm]) | _ => [thm] end; - in atoms end; + in atoms end + + fun mksimps pairs ctxt = + map_filter (try mk_eq) o mk_atomize ctxt pairs o Variable.gen_all ctxt + + fun unsafe_solver_tac ctxt = + let + val sol_thms = + reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI}:: Simplifier.prems_of ctxt; + fun sol_tac i = + FIRST + [resolve_tac ctxt sol_thms i, + assume_tac ctxt i, + eresolve_tac ctxt [@{thm falseE}, @{thm not_refl_False}] i] ORELSE + (match_tac ctxt [@{thm conjI}] + THEN_ALL_NEW sol_tac) i + in + (fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' sol_tac + end; - fun mksimps pairs (_: simpset) = - map_filter (try mk_eq) o mk_atomize pairs o gen_all - - fun unsafe_solver_tac ss = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [resolve_tac (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ss), - atac, etac @{thm falseE}]; - (*the above may instantiate variables*) val unsafe_solver = mk_solver "PredicateLogic unsafe" unsafe_solver_tac; - fun safe_solver_tac ss = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [match_tac (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ss), - eq_assume_tac, ematch_tac @{thms falseE}]; + (*No premature instantiation of variables during simplification*) + fun safe_solver_tac ctxt = + (fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' + FIRST' [match_tac ctxt (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ctxt), + eq_assume_tac, ematch_tac ctxt [@{thm falseE}, @{thm not_refl_False}]]; val safe_solver = mk_solver "PredicateLogic safe" safe_solver_tac; (* Basic infrastructure for simplification *) - val PL_basic_ss = Simplifier.global_context @{theory} empty_ss - setsubgoaler asm_simp_tac + val PL_basic_ss = + empty_simpset @{context} setSSolver safe_solver setSolver unsafe_solver - setmksimps (mksimps mksimps_pairs) - setmkeqTrue mk_eq_True - setmkcong mk_meta_cong; + |> Simplifier.set_subgoaler asm_simp_tac + |> Simplifier.set_mksimps (mksimps mksimps_pairs) + |> Simplifier.set_mkeqTrue mk_eq_True + |> Simplifier.set_mkcong mk_meta_cong + |> simpset_of (* Combination of classical reasoning and simplifier. *) structure Splitter = Splitter ( - val thy = @{theory} + val context = @{context} val mk_eq = mk_eq val meta_eq_to_iff = @{thm meta_eq_to_iff} val iffD = @{thm iffD2} @@ -152,13 +163,11 @@ struct val contrapos = @{thm contrapos} val contrapos2 = @{thm contrapos2} val notnotD = @{thm notnotD} + val safe_tac = Classical.safe_tac ); val split_tac = Splitter.split_tac; val split_inside_tac = Splitter.split_inside_tac; - (* val split_asm_tac = Splitter.split_asm_tac; *) - val op addsplits = Splitter.addsplits; - val op delsplits = Splitter.delsplits; structure Clasimp = Clasimp ( structure Simplifier = Simplifier @@ -171,6 +180,9 @@ struct ); open Clasimp; + fun unfold_tac ctxt ths = + ALLGOALS (full_simp_tac (clear_simpset (put_simpset PL_basic_ss ctxt) addsimps ths)) + end; (* structure Simpdata *) structure Splitter = Simpdata.Splitter; From 95fc3d278eedf806b0c024e53d2dffff8898bf19 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 20 Dec 2020 16:48:15 +0200 Subject: [PATCH 002/167] STY: rm trailing whitespace --- isabelle/CaseExpressions.thy | 20 ++-- isabelle/FixedPoints.thy | 2 +- isabelle/Functions.thy | 38 +++---- isabelle/Integers.thy | 102 +++++++++--------- isabelle/NatArith.thy | 194 +++++++++++++++++------------------ isabelle/NatDivision.thy | 94 ++++++++--------- isabelle/NatOrderings.thy | 48 ++++----- isabelle/Peano.thy | 12 +-- isabelle/PredicateLogic.thy | 50 ++++----- isabelle/README.html | 2 +- isabelle/SMT.thy | 146 +++++++++++++------------- isabelle/SetTheory.thy | 24 ++--- isabelle/Strings.thy | 12 +-- isabelle/Tuples.thy | 74 ++++++------- isabelle/Zenon.thy | 8 +- isabelle/simplifier_setup.ML | 22 ++-- 16 files changed, 424 insertions(+), 424 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 10dd8c6f..c9fc0a31 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -41,7 +41,7 @@ definition Case where "Case(ps, es) \ CHOOSE x : x \ (UNION { CaseArm(ps[i], es[i]) : i \ DOMAIN ps })" definition CaseOther where - "CaseOther(ps, es, oth) \ + "CaseOther(ps, es, oth) \ CHOOSE x : x \ (UNION { CaseArm(ps[i], es[i]) : i \ DOMAIN ps }) \ CaseArm((\i \ DOMAIN ps : \ps[i]), oth)" @@ -73,14 +73,14 @@ parse_ast_translation {* (* get_case_constituents extracts the lists of predicates, terms, and default value from the arms of a case expression. The order of the ASTs is reversed. *) - fun get_case_constituents (Ast.Appl[Ast.Constant "_other", t]) = + fun get_case_constituents (Ast.Appl[Ast.Constant "_other", t]) = (* 1st case: single "OTHER" arm *) ([], [], SOME t) - | get_case_constituents (Ast.Appl[Ast.Constant "_case1", p, t]) = + | get_case_constituents (Ast.Appl[Ast.Constant "_case1", p, t]) = (* 2nd case: a single arm, no "OTHER" branch *) ([p], [t], NONE) - | get_case_constituents (Ast.Appl[Ast.Constant "_case2", - Ast.Appl[Ast.Constant "_case1", p, t], + | get_case_constituents (Ast.Appl[Ast.Constant "_case2", + Ast.Appl[Ast.Constant "_case1", p, t], arms]) = (* 3rd case: one arm, followed by remaining arms *) let val (ps, ts, oth) = get_case_constituents arms @@ -125,8 +125,8 @@ print_ast_translation {* | list_from_tp t = [t] in list_from_tp tp end - (* make_case_arms constructs an AST representing the arms of the - CASE expression. The result is an AST of "type" case_arms. + (* make_case_arms constructs an AST representing the arms of the + CASE expression. The result is an AST of "type" case_arms. The lists of predicates and terms are of equal length, oth is optional. The lists can be empty only if "oth" is present, corresponding to a degenerated expression "CASE OTHER -> e". *) @@ -138,7 +138,7 @@ print_ast_translation {* in (* last arm: check if OTHER defaults *) if oth = NONE then arm - else Ast.Appl[Ast.Constant @{syntax_const "_case2"}, arm, + else Ast.Appl[Ast.Constant @{syntax_const "_case2"}, arm, Ast.Appl[Ast.Constant @{syntax_const "_other"}, the oth]] end | make_case_arms (p::ps) (t::ts) oth = @@ -164,13 +164,13 @@ print_ast_translation {* val trms = list_from_tuple tTuple in (* make sure that tuples are of equal length, otherwise give up *) if length prds = length trms - then Ast.Appl[Ast.Constant @{syntax_const "_case_syntax"}, + then Ast.Appl[Ast.Constant @{syntax_const "_case_syntax"}, make_case_arms prds trms (SOME oth)] else Ast.Appl[Ast.Constant "CaseOther", pTuple, tTuple, oth] end | caseother_tr' _ = raise Match in - [(@{const_syntax "Case"}, case_syntax_tr'), + [(@{const_syntax "Case"}, case_syntax_tr'), (@{const_syntax "CaseOther"}, caseother_tr')] end *} diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index e4a086b3..7ddf3970 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -14,7 +14,7 @@ begin text \ As a test for the encoding of \tlaplus{} set theory, we develop the Knaster-Tarski theorems for least and greatest fixed points in the - subset lattice. Again, the proofs essentially follow Paulson's + subset lattice. Again, the proofs essentially follow Paulson's developments for Isabelle/ZF. \ diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index 922271f0..b68d19a2 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -14,12 +14,12 @@ begin subsection \ Syntax and axioms for functions \ text \ - Functions in \tlaplus{} are not defined (e.g., as sets of pairs), + Functions in \tlaplus{} are not defined (e.g., as sets of pairs), but axiomatized, and in fact, pairs and tuples will be defined as special functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined - as functions over products. We follow the development of functions given + as functions over products. We follow the development of functions given in Section 16.1 of ``Specifying Systems''. \ @@ -50,7 +50,7 @@ text \ Two functions are equal if they have the same domain and agree on all arguments within the domain. \ -axiomatization where fcnEqual[elim!]: +axiomatization where fcnEqual[elim!]: "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" @@ -62,7 +62,7 @@ text \ consts FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) \ \ function space \ syntax (ASCII) - FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) + FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) axiomatization where FuncSet: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T)" @@ -124,13 +124,13 @@ text \ proofs about equality of functions. \ -lemma fcnEqual2[elim!]: +lemma fcnEqual2[elim!]: "\isAFcn(g); isAFcn(f); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" by (rule fcnEqual) text \ possibly useful as a simplification rule, but cannot be active by default \ -lemma fcnEqualIff: +lemma fcnEqualIff: assumes "isAFcn(f)" and "isAFcn(g)" shows "(f = g) = (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : f[x] = g[x]))" using assms by auto @@ -141,7 +141,7 @@ lemma [intro!]: by auto lemma [intro!]: - "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; + "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; \y \ DOMAIN g : y \ v \ f[y] = g[y]\ \ [g EXCEPT ![v] = w] = f" "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; @@ -271,7 +271,7 @@ lemma functionInFuncSet: using assms by auto lemma exceptInFuncSet[elim!]: - assumes 1: "f \ [S \ U]" and 2: "U \ T" + assumes 1: "f \ [S \ U]" and 2: "U \ T" and 3: "\v \ S; isAFcn(f); DOMAIN f = S; \x \ S : f[x] \ U\ \ e \ T" shows "[f EXCEPT ![v]=e] \ [S \ T]" (is "?exc \ [S \ T]") proof @@ -288,7 +288,7 @@ text \ \ lemma exceptInFuncSetSame: - assumes "f \ [S \ T]" + assumes "f \ [S \ T]" and "\v \ S; isAFcn(f); DOMAIN f = S; \x \ S : f[x] \ T\ \ e \ T" shows "[f EXCEPT ![v]=e] \ [S \ T]" using assms by auto @@ -340,13 +340,13 @@ lemmas fcnEqual[where g = "g @@ h" for g h, intro!] lemma extendEqualIff [simp]: - "isAFcn(f) \ (f = g @@ h) = - (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ - (\x \ DOMAIN g : f[x] = g[x]) \ + "isAFcn(f) \ (f = g @@ h) = + (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ + (\x \ DOMAIN g : f[x] = g[x]) \ (\x \ DOMAIN h \\ DOMAIN g : f[x] = h[x]))" - "isAFcn(f) \ (g @@ h = f) = - (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ - (\x \ DOMAIN g : g[x] = f[x]) \ + "isAFcn(f) \ (g @@ h = f) = + (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ + (\x \ DOMAIN g : g[x] = f[x]) \ (\x \ DOMAIN h \\ DOMAIN g : h[x] = f[x]))" by auto @@ -431,7 +431,7 @@ by auto subsubsection \ Injective functions \ -definition InjectiveOn +definition InjectiveOn where "InjectiveOn(f,A) \ \x,y \ A \ DOMAIN f : f[x] = f[y] \ x = y" abbreviation Injective \ \ special case: injective function \ @@ -460,7 +460,7 @@ lemma injectiveOnI: using assms by (auto simp: InjectiveOn_def) lemma injectiveOnD: - assumes "f[x] = f[y]" and "InjectiveOn(f,A)" + assumes "f[x] = f[y]" and "InjectiveOn(f,A)" and "x \ A" and "x \ DOMAIN f" and "y \ A" and "y \ DOMAIN f" shows "x = y" using assms by (auto simp: InjectiveOn_def) @@ -493,7 +493,7 @@ lemma inverseThenInjective: shows "InjectiveOn(f,A)" proof (rule injectiveOnI) fix x y - assume x: "x \ A" "x \ DOMAIN f" + assume x: "x \ A" "x \ DOMAIN f" and y: "y \ A" "y \ DOMAIN f" and eq: "f[x] = f[y]" from x have x1: "x = g[f[x]]" by (rule sym[OF inv]) @@ -673,7 +673,7 @@ proof - by auto finally show ?thesis . qed - + lemma injectionsI: assumes "f \ [S \ T]" and "\x y. \ x \ S; y \ S; f[x] = f[y] \ \ x = y" shows "f \ Injections(S,T)" diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 3fa326e5..3502ce19 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -30,17 +30,17 @@ translations (* lemma eq_imp_negeq: "n = m \ -.n = -.m" by simp *) axiomatization where - neg0 [simp]: "-.0 = 0" + neg0 [simp]: "-.0 = 0" and neg_neg [simp]: "-.-.n = n" and negNotInNat [simp]: "-.(Succ[n]) \ Nat" -lemma negNat_noteq_Nat [simp]: +lemma negNat_noteq_Nat [simp]: "\m \ Nat; n \ Nat\ \ (-. Succ[m] = Succ[n]) = FALSE" proof (rule contradiction) assume "(-. Succ[m] = Succ[n]) \ FALSE" - and m: "m \ Nat" and n: "n \ Nat" + and m: "m \ Nat" and n: "n \ Nat" hence "-. Succ[m] = Succ[n]" by auto hence "-. Succ[m] \ Nat" using n by auto with negNotInNat[of m] show FALSE by simp @@ -58,7 +58,7 @@ qed lemma nat_not_eq_inv: "n \ Nat \ n = 0 \ -.n \ n" using not0_implies_Suc[of n] by auto -lemma minusInj [dest]: +lemma minusInj [dest]: assumes hyp: "-.n = -.m" shows "n = m" proof - @@ -66,8 +66,8 @@ proof - thus ?thesis by simp qed -lemma minusInj_iff [simp]: - "-.x = -.y = (x = y)" +lemma minusInj_iff [simp]: + "-.x = -.y = (x = y)" by auto lemma neg0_imp_0 [simp]: "-.n = 0 = (n = 0)" @@ -77,7 +77,7 @@ proof auto thus "n = 0" by simp qed -lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" +lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" by simp lemma notneg0_imp_not0 [dest]: "-.n \ 0 \ n \ 0" @@ -113,7 +113,7 @@ by auto lemma negNat_exists: "-.n \ Nat \ \k \ Nat: n = -.k" by force -lemma nat_eq_negnat_is_0 [simp]: +lemma nat_eq_negnat_is_0 [simp]: assumes "n \ Nat" shows "(n = -.n) = (n = 0)" using assms by (cases "n", auto) @@ -147,10 +147,10 @@ using assms unfolding Int_def by auto -- {* Integer cases over two parameters *} lemma intCases2: assumes m: "m \ Int" and n: "n \ Int" - and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" - and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" - and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" - and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" + and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" + and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" + and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" + and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" shows "P(m,n)" using m proof (cases "m") assume "m \ Nat" @@ -163,14 +163,14 @@ qed lemma intCases3: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" - and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" - and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" - and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" - and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" - and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" - and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" - and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" + and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" + and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" + and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" + and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" + and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" + and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" + and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" + and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" shows "P(m,n,p)" proof (rule intCases2[OF m n]) fix m n @@ -235,11 +235,11 @@ by(simp add: isPos_def) lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \ n \ 0)" unfolding isPos_def using not0_implies_Suc[of n] by auto -(*lemma isNeg_eq_inNegNat: +(*lemma isNeg_eq_inNegNat: "isNeg(n) = (n \ {-.n : n \ Nat } \\ {0})" unfolding isNeg_def by force*) -lemma isNeg_eq_inNegNat: +lemma isNeg_eq_inNegNat: "isNeg(n) = (n \ {-.n : n \ Nat } \ n \ 0)" unfolding isNeg_def by force @@ -250,7 +250,7 @@ lemma negNotNat_isNat: assumes n: "n \ Int" shows "(-.n \ Nat) = FALSE \ n \ Nat" using n by (cases, auto) -lemma noNatisNeg [simp]: +lemma noNatisNeg [simp]: "n \ Nat \ isNeg(n) = FALSE" -- {* No natural number is negative *} unfolding isNeg_def using negNotInNat by blast @@ -270,11 +270,11 @@ lemma isPos_isNeg_false [simp]: "n \ Int \ isPos(n) \ Int \\ {0} \ \ isPos(n) \ isPos(-.n)" unfolding isPos_eq_inNat1 using intNotNatIsNegNat by auto*) -lemma isPos_neg_isNeg [simp]: +lemma isPos_neg_isNeg [simp]: assumes n: "n \ Int" shows "isPos(-.n) = isNeg(n)" by (auto simp: minus_sym isPos_def isNeg_def) -lemma notIsNeg0_isPos: +lemma notIsNeg0_isPos: assumes n: "n \ Int" shows "\\ isNeg(n); n \ 0\ \ isPos(n)" using n by (cases, auto simp: isPos_eq_inNat1 dest: negNat_isNeg) @@ -317,7 +317,7 @@ by (auto simp: sgn_def) lemma sgn1_imp_pos (*[simp]*): "sgn(n) = 1 \ n \ Nat \ n \ 0" unfolding sgn_def isPos_eq_inNat1 by auto -(*lemma neg_imp_negSuc (*[simp]*): +(*lemma neg_imp_negSuc (*[simp]*): assumes h:"n \ {-.n : n \ Nat } \\ {0}" shows "isNeg(n)" unfolding isNeg_def proof - @@ -328,7 +328,7 @@ proof - with 3 show "\k \ Nat : n = -.(Succ[k])" by auto qed*) -lemma sgnm1_imp_neg: +lemma sgnm1_imp_neg: assumes n:"n \ Int" shows "sgn(n) = -.1 \ isNeg(n)" unfolding sgn_def using intThenPosZeroNeg[OF n] by auto @@ -347,7 +347,7 @@ lemma sgnNat_not1: "\n \ Nat; sgn(n) \ 1\ \ n = 0" using sgnNat_is_0or1[of n] by auto -lemma sgnNat_not_neg [simp]: +lemma sgnNat_not_neg [simp]: "n \ Nat \ sgn(n) = -.1 = FALSE" unfolding sgn_def isPos_eq_inNat1 by auto @@ -378,7 +378,7 @@ by (auto simp: sgn_eq_neg1_is_not_nat) lemma sgn_neg_eq_1_false: "\sgn(-.m) = 1; m \ Nat\ \ P" unfolding sgn_def by auto -lemma sgn_minus [simp]: +lemma sgn_minus [simp]: assumes n: "n \ Int" shows "sgn(-.n) = -.sgn(n)" unfolding sgn_def using n by (cases, auto) @@ -400,22 +400,22 @@ lemma absn0 [simp]: "abs(-.0) = 0" by simp **) -lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" +lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" unfolding abs_def by (auto dest: sgnNat_not1) -lemma abs_neg [simp]: +lemma abs_neg [simp]: assumes n: "n \ Int" shows "abs(-.n) = abs(n)" unfolding abs_def using n by (auto dest: sgnNat_not1) -subsection {* Orders on integers *} +subsection {* Orders on integers *} -text {* +text {* We distinguish four cases, depending on the arguments being in Nat or negative. *} -lemmas int_leq_pp_def = nat_leq_def +lemmas int_leq_pp_def = nat_leq_def -- {* 'positive-positive' case, ie: both arguments are naturals *} axiomatization where @@ -427,11 +427,11 @@ and (* theorems int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) -lemma int_boolify_leq [simp]: +lemma int_boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" by(rule intCases2[of a b], simp_all) -lemma int_leq_isBool [intro!,simp]: +lemma int_leq_isBool [intro!,simp]: "\a \ Int; b \ Int\ \ isBool(a \ b)" unfolding isBool_def by auto @@ -450,7 +450,7 @@ by(rule intCases2[of m n], simp_all) subsection {* Addition of integers *} -text {* +text {* Again, we distinguish four cases in the definition of @{text "a + b"}, according to each argument being positive or negative. *} @@ -483,14 +483,14 @@ by (auto simp: int_add_pn_def dest: nat_leq_antisym) text {* Closure *} -lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" +lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" by (rule intCases2[of m n], auto simp: int_add_def) text {* Neutral element *} -lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" +lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" by(rule intCases, auto simp add: int_add_np_def) -lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" +lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" by(rule intCases, auto simp add: int_add_pn_def) text {* Additive inverse element *} @@ -509,7 +509,7 @@ by (rule intCases, auto simp: int_add_def) text {* Commutativity *} -lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" +lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" by(simp add: int_add_def) lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" @@ -548,7 +548,7 @@ by(rule adiff_adiff_left_nat[OF m n p]) declare leq_neq_iff_less [simplified,simp] -lemma int_add_assoc1: +lemma int_add_assoc1: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + -.p) = (m + n) + -.p" apply(rule nat_leq_cases[OF p n]) @@ -601,7 +601,7 @@ apply(rule nat_leq_cases[of "n + p" m]) using adiff_add_assoc2[OF _ p n m, symmetric] apply (simp add: adiff_is_0_eq') apply(rule nat_leq_cases[OF n m], simp_all) apply(rule nat_leq_cases[of p "m -- n"], simp_all) - using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp + using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp using adiff_add_assoc3[symmetric] add_commute_nat[OF n p] apply simp using adiff_add_assoc2[symmetric] add_commute_nat[OF n p] apply simp done @@ -667,7 +667,7 @@ lemma int_add_assoc6: shows "-.m + (-.n + p) = -.(m + n) + p" using assms add_commute_int[of "-.n" p] - add_commute_int[of "-.m" "p + -.n"] + add_commute_int[of "-.m" "p + -.n"] add_commute_int[of "-.(m + n)" p] apply simp apply(rule nat_leq_cases[OF n p], simp_all) apply(rule nat_leq_cases[of m "p -- n"], simp+) @@ -687,11 +687,11 @@ apply(rule nat_leq_cases[OF n p], simp_all) apply(rule adiff_add_assoc[symmetric], simp+) done -lemma add_assoc_int: +lemma add_assoc_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m + (n + p) = (m + n) + p" -using m n p -by (rule intCases3, +using m n p +by (rule intCases3, auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 int_add_assoc4 int_add_assoc5 int_add_assoc6) @@ -785,10 +785,10 @@ apply(rule nat_leq_cases[OF p n]) apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) done -lemma distrib_left_int: +lemma distrib_left_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m * (n + p) = (m * n + m * p)" -apply(rule intCases3[OF m n p], +apply(rule intCases3[OF m n p], simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) apply(rule add_mult_distrib_left_nat, assumption+) apply(rule ppn_distrib_left_nat, assumption+) @@ -828,7 +828,7 @@ done lemma distrib_right_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "(m + n) * p = (m * p + n * p)" -apply(rule intCases3[OF m n p], +apply(rule intCases3[OF m n p], simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) apply(rule add_mult_distrib_right_nat, assumption+) apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) @@ -866,7 +866,7 @@ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" lemma diffIsInt [simp]: -- {* Closure *} - "\m \ Int; n \ Int\ \ m - n \ Int" + "\m \ Int; n \ Int\ \ m - n \ Int" by (simp add: int_diff_def) lemma diff_neg_is_add [simp]: "\m \ Int; n \ Int\ \ m - -.n = m + n" diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index fc2618f9..6b17dc37 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -20,7 +20,7 @@ structure AlgebraSimps = setup AlgebraSimps.setup -text{* +text{* The rewrites accumulated in @{text algebra_simps} deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and @@ -46,17 +46,17 @@ lemma addnatType: assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" using assms unfolding addnat_def by (rule bprimrecType_nat, auto) -lemma addIsNat [intro!,simp]: +lemma addIsNat [intro!,simp]: assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" unfolding nat_add_def[OF assms] using assms addnatType by blast text {* Base case and Inductive case *} -lemma addnat_0: +lemma addnat_0: assumes "m \ Nat" shows "addnat(m)[0] = m" using assms unfolding addnat_def by (rule primrec_natE, auto) -lemma add_0_nat [simp]: +lemma add_0_nat [simp]: assumes "m \ Nat" shows "m + 0 = m" by (simp add: nat_add_def[OF assms] addnat_0[OF assms]) @@ -71,24 +71,24 @@ next with n show "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" by blast qed (auto) -lemma add_Succ_nat [simp]: +lemma add_Succ_nat [simp]: assumes "m \ Nat" and "n \ Nat" shows "m + Succ[n] = Succ[m + n]" using assms by (simp add: nat_add_def addnat_Succ[OF assms]) -lemma add_0_left_nat [simp]: - assumes n: "n \ Nat" +lemma add_0_left_nat [simp]: + assumes n: "n \ Nat" shows "0 + n = n" using n by (induct, auto) -lemma add_Succ_left_nat [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" +lemma add_Succ_left_nat [simp]: + assumes n: "n \ Nat" and m: "m \ Nat" shows "Succ[m] + n = Succ[m + n]" using n apply induct using m by auto -lemma add_Succ_shift_nat: - assumes n: "n \ Nat" and m: "m \ Nat" +lemma add_Succ_shift_nat: + assumes n: "n \ Nat" and m: "m \ Nat" shows "Succ[m] + n = m + Succ[n]" using assms by simp @@ -102,7 +102,7 @@ using assms by auto text {* Associativity *} -lemma add_assoc_nat [algebra_simps]: +lemma add_assoc_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + p) = (m + n) + p" using assms by (induct, simp_all) @@ -120,8 +120,8 @@ lemma add_right_cancel_nat [simp]: using assms by (induct, simp_all) -lemma add_left_commute_nat [algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" +lemma add_left_commute_nat [algebra_simps]: + assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a + (b + c) = b + (a + c)" using assms by(simp only: add_assoc_nat add_commute_nat) @@ -172,7 +172,7 @@ lemma multnatType: unfolding multnat_def by (rule bprimrecType_nat, auto simp: assms) lemma multIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" + assumes m: "m \ Nat" and n: "n \ Nat" shows "m * n \ Nat" unfolding nat_mult_def[OF assms] using assms multnatType by blast @@ -198,17 +198,17 @@ next qed (auto simp: m) lemma mult_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" + assumes "m \ Nat" and "n \ Nat" shows "m * Succ[n] = m * n + m" using assms by (simp add: nat_mult_def multnat_Succ[OF assms]) lemma mult_0_left_nat [simp]: - assumes n: "n \ Nat" + assumes n: "n \ Nat" shows "0 * n = 0" using n by (induct, simp_all) -lemma mult_Succ_left_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" +lemma mult_Succ_left_nat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] * n = m * n + n" using n apply induct using m by (simp_all add: add_ac_nat) @@ -265,7 +265,7 @@ lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp text {* Associativity *} -lemma mult_assoc_nat[algebra_simps]: +lemma mult_assoc_nat[algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m * (n * p) = (m * n) * p" using m apply induct @@ -273,19 +273,19 @@ using assms by (auto simp add: add_mult_distrib_right_nat) text {* Reasoning about @{text "m * n = 0"}, etc. *} -lemma mult_is_0_nat [simp]: +lemma mult_is_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = 0) = (m = 0 \ n = 0)" using m apply induct using n by auto -lemma mult_eq_1_iff_nat [simp]: +lemma mult_eq_1_iff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = 1) = (m = 1 \ n = 1)" using m apply induct using n by (induct, auto)+ -lemma one_eq_mult_iff_nat [simp]: +lemma one_eq_mult_iff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 = m * n) = (m = 1 \ n = 1)" proof - @@ -296,7 +296,7 @@ qed text {* Cancellation rules *} -lemma mult_cancel1_nat [simp]: +lemma mult_cancel1_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k * m = k * n) = (m = n \ (k = 0))" proof - @@ -329,18 +329,18 @@ proof - with k m n show ?thesis by auto qed -lemma mult_cancel2_nat [simp]: +lemma mult_cancel2_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k = n * k) = (m = n \ k = 0)" using assms by (simp add: mult_commute_nat) -lemma Suc_mult_cancel1_nat: +lemma Suc_mult_cancel1_nat: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(Succ[k] * m = Succ[k] * n) = (m = n)" using k m n mult_cancel1_nat[of "Succ[k]" m n] by simp -lemma mult_left_commute_nat[algebra_simps]: +lemma mult_left_commute_nat[algebra_simps]: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * (b * c) = b * (a * c)" using assms by(simp only: mult_commute_nat mult_assoc_nat) @@ -360,7 +360,7 @@ by (simp add: Pred_def) lemma Pred_Succ_nat [simp]: "n \ Nat \ Pred[Succ[n]] = n" unfolding Pred_def by (auto intro: bChooseI2) -lemma Succ_Pred_nat [simp]: +lemma Succ_Pred_nat [simp]: assumes "n \ Nat" and "n \ 0" shows "Succ[Pred[n]] = n" using assms unfolding Pred_def by (cases "n", auto intro: bChooseI2) @@ -372,7 +372,7 @@ using assms by (cases "n", auto) subsection {* Difference of natural numbers *} text {* - We define a form of difference @{text "--"} of natural numbers that cuts off + We define a form of difference @{text "--"} of natural numbers that cuts off at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes called ``arithmetic difference''. *} @@ -395,11 +395,11 @@ unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast text {* Neutral element and Inductive step *} -lemma adiffnat_0: +lemma adiffnat_0: assumes "m \ Nat" shows "adiffnat(m)[0] = m" using assms unfolding adiffnat_def by (rule primrec_natE, auto) -lemma adiff_0_nat [simp]: +lemma adiff_0_nat [simp]: assumes "m \ Nat" shows "m -- 0 = m" by (simp add: nat_adiff_def[OF assms] adiffnat_0[OF assms]) @@ -426,40 +426,40 @@ using n apply induct by (simp_all add: adiff_Succ_nat) text {* Reasoning about @{text "m -- m = 0"}, etc. *} -lemma adiff_Succ_Succ_nat [simp]: +lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] -- Succ[n] = m -- n" using n apply induct using assms by (auto simp add: adiff_Succ_nat) -lemma adiff_self_eq_0_nat [simp]: +lemma adiff_self_eq_0_nat [simp]: assumes m: "m \ Nat" shows "m -- m = 0" using m apply induct by auto text {* Associativity *} -lemma adiff_adiff_left_nat: +lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "(i -- j) -- k = i -- (j + k)" using i j apply (rule diffInduct) using k by auto -lemma Succ_adiff_adiff_nat [simp]: +lemma Succ_adiff_adiff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" using assms by (simp add: adiff_adiff_left_nat) text {* Commutativity *} -lemma adiff_commute_nat: +lemma adiff_commute_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i -- j -- k = i -- k -- j" using assms by (simp add: adiff_adiff_left_nat add_commute_nat) text {* Cancellation rules *} -lemma adiff_add_inverse_nat [simp]: +lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(n + m) -- n = m" using n apply induct @@ -489,20 +489,20 @@ using assms by simp_all text {* Difference distributes over multiplication *} -lemma adiff_mult_distrib_nat [algebra_simps]: +lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m -- n) * k = (m * k) -- (n * k)" using m n apply(rule diffInduct) using k by simp_all -lemma adiff_mult_distrib2_nat [algebra_simps]: +lemma adiff_mult_distrib2_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) -- {* NOT added as rewrites, since sometimes they are used from right-to-left *} lemmas nat_distrib = - add_mult_distrib_right_nat add_mult_distrib_left_nat + add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat @@ -520,25 +520,25 @@ lemma nat_add_left_cancel_leq [simp]: shows "(k + m \ k + n) = (m \ n)" using assms by (induct k) simp_all -lemma nat_add_left_cancel_less [simp]: +lemma nat_add_left_cancel_less [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k + m < k + n) = (m < n)" using k apply induct using assms by simp_all -lemma nat_add_right_cancel_less [simp]: +lemma nat_add_right_cancel_less [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m + k < n + k) = (m < n)" using k apply induct using assms by simp_all -lemma nat_add_right_cancel_leq [simp]: +lemma nat_add_right_cancel_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m + k \ n + k) = (m \ n)" using k apply induct using assms by simp_all -lemma add_gr_0 [simp]: +lemma add_gr_0 [simp]: assumes "n \ Nat" and "m \ Nat" shows "(m + n > 0) = (m > 0 \ n > 0)" using assms by (auto dest: nat_gt0_implies_Succ nat_not_lessD) @@ -576,7 +576,7 @@ subsubsection {* (Partially) Ordered Groups *} -- {* The two following lemmas are just ``one half'' of @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} proved above. *} -lemma add_leq_left_mono: +lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" using assms by simp @@ -591,7 +591,7 @@ lemma add_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c \ d \ a + c \ b + d" using assms - add_leq_right_mono[OF a b c] + add_leq_right_mono[OF a b c] add_leq_left_mono[OF c d b] nat_leq_trans[of "a + c" "b + c" "b + d"] by simp @@ -602,7 +602,7 @@ lemma add_less_left_mono: shows "a < b \ c + a < c + b" using assms by simp -lemma add_less_right_mono: +lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" using assms by simp @@ -612,7 +612,7 @@ lemma add_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c < d \ a + c < b + d" using assms - add_less_right_mono[OF a b c] + add_less_right_mono[OF a b c] add_less_left_mono[OF c d b] nat_less_trans[of "a + c" "b + c" "b + d"] by simp @@ -621,7 +621,7 @@ lemma add_less_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c \ d \ a + c < b + d" using assms - add_less_right_mono[OF a b c] + add_less_right_mono[OF a b c] add_leq_left_mono[OF c d b] nat_less_leq_trans[of "a + c" "b + c" "b + d"] by blast @@ -630,58 +630,58 @@ lemma add_leq_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c < d \ a + c < b + d" using assms - add_leq_right_mono[OF a b c] + add_leq_right_mono[OF a b c] add_less_left_mono[OF c d b] nat_leq_less_trans[of "a + c" "b + c" "b + d"] by blast lemma leq_add1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" + assumes m: "m \ Nat" and n: "n \ Nat" shows "n \ n + m" using assms add_leq_left_mono[of 0 m n] by simp lemma leq_add2 [simp]: - assumes "m \ Nat" and "n \ Nat" + assumes "m \ Nat" and "n \ Nat" shows "n \ m + n" using assms add_leq_right_mono [of 0 m n] by simp lemma less_add_Succ1: - assumes "i \ Nat" and "m \ Nat" + assumes "i \ Nat" and "m \ Nat" shows "i < Succ[i + m]" using assms by simp -lemma less_add_Succ2: - assumes "i \ Nat" and "m \ Nat" +lemma less_add_Succ2: + assumes "i \ Nat" and "m \ Nat" shows "i < Succ[m + i]" using assms by simp -lemma less_iff_Succ_add: - assumes "m \ Nat" and "n \ Nat" +lemma less_iff_Succ_add: + assumes "m \ Nat" and "n \ Nat" shows "(m < n) = (\k \ Nat: n = Succ[m + k])" using assms by (auto intro!: less_imp_Succ_add) -lemma trans_leq_add1: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_leq_add1: + assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i \ j + m" using assms by (auto elim: nat_leq_trans) -lemma trans_leq_add2: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_leq_add2: + assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i \ m + j" using assms by (auto elim: nat_leq_trans) -lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_less_add1: + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < j + m" using assms by (auto elim: nat_less_leq_trans) -lemma trans_less_add2: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_less_add2: + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" using assms by (auto elim: nat_less_leq_trans) -lemma add_lessD1: - assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" +lemma add_lessD1: + assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "i < k" using assms by (intro nat_leq_less_trans[of "i" "i+j" "k"], simp+) @@ -695,17 +695,17 @@ lemma not_add_less2 [simp]: shows "(j + i < i) = FALSE" using assms by (simp add: add_commute_nat) -lemma add_leqD1: +lemma add_leqD1: assumes "m + k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "m \ n" using assms by (intro nat_leq_trans[of "m" "m+k" "n"], simp+) -lemma add_leqD2: +lemma add_leqD2: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k \ n" using assms by (intro nat_leq_trans[of "k" "m+k" "n"], simp+) -lemma add_leqE: +lemma add_leqE: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "(m \ n \ k \ n \ R) \ R" using assms by (blast dest: add_leqD1 add_leqD2) @@ -720,13 +720,13 @@ subsubsection {* More results about arithmetic difference *} text {* Addition is the inverse of subtraction: if @{term "n \ m"} then @{term "n + (m -- n) = m"}. *} -lemma add_adiff_inverse: +lemma add_adiff_inverse: assumes "m \ Nat" and "n \ Nat" shows "\(m < n) \ n + (m -- n) = m" apply (induct m n rule: diffInduct) using assms by simp_all -lemma le_add_adiff_inverse [simp]: +lemma le_add_adiff_inverse [simp]: assumes "m \ Nat" and "n \ Nat" shows "n \ m \ n + (m -- n) = m" using assms by (simp add: add_adiff_inverse nat_not_order_simps) @@ -736,7 +736,7 @@ lemma le_add_adiff_inverse2 [simp]: shows "n \ m \ (m -- n) + n = m" using assms by (simp add: add_commute_nat) -lemma Succ_adiff_leq: +lemma Succ_adiff_leq: assumes "m \ Nat" and "n \ Nat" shows "n \ m \ Succ[m] -- n = Succ[m -- n]" apply (induct m n rule: diffInduct) @@ -748,13 +748,13 @@ lemma adiff_less_Succ: apply (induct m n rule: diffInduct) using assms by (auto simp: nat_less_Succ) -lemma adiff_leq_self [simp]: +lemma adiff_leq_self [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ m" apply (induct m n rule: diffInduct) using assms by simp_all -lemma leq_iff_add: +lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") proof - @@ -768,7 +768,7 @@ proof - from 1 2 assms show ?thesis by blast qed -lemma less_imp_adiff_less: +lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" proof - @@ -776,18 +776,18 @@ proof - with j k jk show ?thesis by (auto elim: nat_leq_less_trans) qed -lemma adiff_Succ_less (*[simp]*): +lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" shows "0 < n \ n -- Succ[i] < n" using n apply cases using i by auto -lemma adiff_add_assoc: +lemma adiff_add_assoc: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(i + j) -- k = i + (j -- k)" using assms by (induct j k rule: diffInduct, simp+) -lemma adiff_add_assoc2: +lemma adiff_add_assoc2: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(j + i) -- k = (j -- k) + i" using assms by (simp add: add_commute_nat adiff_add_assoc) @@ -802,11 +802,11 @@ lemma adiff_add_assoc4: and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "p -- (m -- n) = (p -- m) + n" using assms - adiff_add_assoc2[OF _ n p m, symmetric] + adiff_add_assoc2[OF _ n p m, symmetric] adiff_add_assoc3[OF _ _ p n m] apply simp using trans_leq_add1[OF _ m p n] by simp -lemma le_imp_adiff_is_add: +lemma le_imp_adiff_is_add: assumes "i \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(j -- i = k) = (j = k + i)" using assms by auto @@ -821,7 +821,7 @@ lemma adiff_is_0_eq' (*[simp]*): shows "m -- n = 0" using assms by simp -lemma zero_less_adiff [simp]: +lemma zero_less_adiff [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(0 < n -- m) = (m < n)" by (induct m n rule: diffInduct, simp_all add: assms) @@ -859,7 +859,7 @@ lemma leq_adiff_right_false [simp]: shows "m \ m -- n = FALSE" using assms by (simp add: leq_adiff_right_add_left[OF _ m m n]) -lemma leq_adiff_right_imp_0: +lemma leq_adiff_right_imp_0: assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" shows "p = 0" using p h apply (induct) @@ -870,13 +870,13 @@ subsubsection {* Monotonicity of Multiplication *} (* from HOL/Ring_and_Field.thy *) -lemma mult_leq_left_mono: +lemma mult_leq_left_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "c * a \ c * b" using c apply induct using 1 a b by (simp_all add: add_leq_mono) -lemma mult_leq_right_mono: +lemma mult_leq_right_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * c \ b * c" using c apply induct @@ -884,8 +884,8 @@ using 1 a b by (simp_all add: add_leq_mono add_commute_nat) text {* @{text "\"} monotonicity, BOTH arguments *} -lemma mult_leq_mono: - assumes 1: "i \ j" "k \ l" +lemma mult_leq_mono: + assumes 1: "i \ j" "k \ l" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" and l: "l \ Nat" shows "i * k \ j * l" using assms @@ -896,7 +896,7 @@ by simp text {* strict, in 1st argument *} -lemma mult_less_left_mono: +lemma mult_less_left_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "k * i < k * j" using 1 @@ -907,7 +907,7 @@ proof (auto simp: nat_gt0_iff_Succ[OF k]) by (induct m, simp_all add: add_less_mono) qed -lemma mult_less_right_mono: +lemma mult_less_right_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i * k < j * k" using 1 @@ -918,7 +918,7 @@ proof (auto simp: nat_gt0_iff_Succ[OF k]) by (induct m, simp_all add: add_less_mono) qed -lemma nat_0_less_mult_iff [simp]: +lemma nat_0_less_mult_iff [simp]: assumes i: "i \ Nat" and j: "j \ Nat" shows "(0 < i * j) = (0 < i \ 0 < j)" using i apply induct @@ -931,7 +931,7 @@ lemma one_leq_mult_iff (*[simp]*): shows "(1 \ m * n) = (1 \ m \ 1 \ n)" using assms by simp -lemma mult_less_cancel_left [simp]: +lemma mult_less_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m < k * n) = (0 < k \ m < n)" proof (auto intro!: mult_less_left_mono[OF _ _ m n k]) @@ -948,7 +948,7 @@ next qed qed -lemma mult_less_cancel_right [simp]: +lemma mult_less_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k < n * k) = (0 < k \ m < n)" proof (auto intro!: mult_less_right_mono[OF _ _ m n k]) @@ -975,7 +975,7 @@ lemma mult_less_self_right [dest]: shows "k=0" using k assms by (cases, auto) -lemma mult_leq_cancel_left [simp]: +lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) @@ -989,7 +989,7 @@ using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) qed qed -lemma mult_leq_cancel_right [simp]: +lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) @@ -1003,17 +1003,17 @@ using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) qed qed -lemma Suc_mult_less_cancel1: +lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" using assms by (simp del: mult_Succ_left_nat) -lemma Suc_mult_leq_cancel1: +lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m \ Succ[k] * n) = (m \ n)" using assms by (simp del: mult_Succ_left_nat) -lemma nat_leq_square: +lemma nat_leq_square: assumes m: "m \ Nat" shows "m \ m * m" using m by (cases, auto) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 81591110..abef1aed 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -32,18 +32,18 @@ lemma [intro!]: "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" by auto -lemma dvdI [intro]: +lemma dvdI [intro]: assumes a: "a \ Nat" and b: "b \ Nat" and k: "k \ Nat" shows "a = b * k \ b dvd a" unfolding dvd_def[OF a b] using k by blast -lemma dvdE [elim]: +lemma dvdE [elim]: assumes "b dvd a" and "a \ Nat" and "b \ Nat" shows "(\k. \k \ Nat; a = b * k\ \ P) \ P" using assms by (auto simp add: dvd_def) -lemma dvd_refl [iff]: - assumes a: "a \ Nat" +lemma dvd_refl [iff]: + assumes a: "a \ Nat" shows "a dvd a" proof - from a have "a = a*1" by simp @@ -64,21 +64,21 @@ proof - qed lemma dvd_0_left_iff [simp]: - assumes "a \ Nat" + assumes "a \ Nat" shows "(0 dvd a) = (a = 0)" using assms by force -lemma dvd_0_right [iff]: +lemma dvd_0_right [iff]: assumes a: "a \ Nat" shows "a dvd 0" using assms by force -lemma one_dvd [iff]: - assumes a: "a \ Nat" +lemma one_dvd [iff]: + assumes a: "a \ Nat" shows "1 dvd a" using assms by force lemma dvd_mult (*[simp]*): - assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a dvd (b * c)" proof - from dvd a c obtain k where k: "k \ Nat" and "c = a*k" by blast @@ -87,16 +87,16 @@ proof - qed lemma dvd_mult2 (*[simp]*): - assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a dvd (b * c)" using mult_commute_nat[OF b c] dvd_mult[OF dvd a c b] by simp -lemma dvd_triv_right [iff]: +lemma dvd_triv_right [iff]: assumes a: "a \ Nat" and b: "b \ Nat" shows "a dvd (b * a)" using assms by (intro dvd_mult, simp+) -lemma dvd_triv_left [iff]: +lemma dvd_triv_left [iff]: assumes a: "a \ Nat" and b: "b \ Nat" shows "a dvd a * b" using assms by (intro dvd_mult2, simp+) @@ -116,17 +116,17 @@ proof - with a c b' d' show ?thesis by blast qed -lemma dvd_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" +lemma dvd_mult_left: + assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a * b dvd c" shows "a dvd c" proof - - from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" + from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" by (auto simp add: mult_assoc_nat) with a b c show ?thesis by blast qed -lemma dvd_mult_right: +lemma dvd_mult_right: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a*b dvd c" shows "b dvd c" proof - @@ -134,14 +134,14 @@ proof - with b a c show ?thesis by (rule dvd_mult_left) qed -lemma dvd_0_left: +lemma dvd_0_left: assumes "a \ Nat" shows "0 dvd a \ a = 0" using assms by simp lemma dvd_add [iff]: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "a dvd c" + and 1: "a dvd b" and 2: "a dvd c" shows "a dvd (b + c)" proof - from a b 1 obtain b' where b': "b' \ Nat" "b = a * b'" by blast @@ -165,7 +165,7 @@ text {* *} definition divmod_rel where - "divmod_rel(m,n,q,r) \ m = q * n + r + "divmod_rel(m,n,q,r) \ m = q * n + r \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" text {* @{const divmod_rel} is total if $n$ is non-zero. *} @@ -183,7 +183,7 @@ proof - fix m' assume m': "m' \ Nat" and ih: "\q, r \ Nat : m' = q*n + r \ r < n" from ih obtain q' r' - where h1: "m' = q' * n + r'" and h2: "r' < n" + where h1: "m' = q' * n + r'" and h2: "r' < n" and q': "q' \ Nat" and r': "r' \ Nat" by blast show "\q, r \ Nat : Succ[m'] = q * n + r \ r < n" proof (cases "Succ[r'] < n") @@ -204,7 +204,7 @@ qed text {* @{const divmod_rel} has unique solutions in the natural numbers. *} lemma divmod_rel_unique_div: assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" - and m: "m \ Nat" and n: "n \ Nat" + and m: "m \ Nat" and n: "n \ Nat" and q: "q \ Nat" and r: "r \ Nat" and q': "q' \ Nat" and r': "r' \ Nat" shows "q = q'" proof - @@ -270,7 +270,7 @@ lemma divmodNat_divmod_rel [rule_format]: unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]], auto) lemma divmodNat_unique: - assumes h: "divmod_rel(m,n,q,r)" + assumes h: "divmod_rel(m,n,q,r)" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" shows "divmodNat(m,n) = \q,r\" @@ -297,7 +297,7 @@ lemma divIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m div n \ Nat" using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) - + lemma modIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n \ Nat" @@ -431,23 +431,23 @@ lemma mod_div_nat_equality2 [simp]: shows "n * (m div n) + m mod n = m" using assms mult_commute_nat[of "n" "m div n"] by simp -lemma mod_div_nat_equality3 [simp]: +lemma mod_div_nat_equality3 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" shows "m mod n + (m div n) * n = m" using assms add_commute_nat[of "m mod n"] by simp -lemma mod_div_nat_equality4 [simp]: +lemma mod_div_nat_equality4 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" shows "m mod n + n * (m div n) = m" using assms mult_commute_nat[of "n" "m div n"] by simp (** immediate consequence of above -lemma div_mod_nat_equality: +lemma div_mod_nat_equality: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "((m div n) * n + m mod n) + k = m + k" using assms by simp -lemma div_mod_nat_equality2: +lemma div_mod_nat_equality2: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * (m div n) + m mod n) + k = m + k" using assms by simp @@ -487,12 +487,12 @@ lemma div_nat_0: shows "0 div n = 0" using assms by simp -lemma mod_0: +lemma mod_0: assumes "n \ Nat" and "0 < n" shows "0 mod n = 0" using assms by simp -lemma mod_nat_mult_self1 [simp]: +lemma mod_nat_mult_self1 [simp]: assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "(q + m * n) mod n = q mod n" proof - @@ -511,57 +511,57 @@ proof - with assms show ?thesis by (simp del: mod_div_nat_equality) qed -lemma mod_nat_mult_self2 [simp]: +lemma mod_nat_mult_self2 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(q + n * m) mod n = q mod n" using assms by (simp add: mult_commute_nat) -lemma mod_nat_mult_self3 [simp]: +lemma mod_nat_mult_self3 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n + q) mod n = q mod n" using assms by (simp add: add_commute_nat) -lemma mod_nat_mult_self4 [simp]: +lemma mod_nat_mult_self4 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m + q) mod n = q mod n" using assms by (simp add: add_commute_nat) -lemma div_nat_mult_self1_is_id [simp]: +lemma div_nat_mult_self1_is_id [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n) div n = m" using assms div_nat_mult_self1 [of 0 m n] by simp -lemma div_nat_mult_self2_is_id [simp]: +lemma div_nat_mult_self2_is_id [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m) div n = m" using assms div_nat_mult_self2[of 0 n m] by simp -lemma mod_nat_mult_self1_is_0 [simp]: +lemma mod_nat_mult_self1_is_0 [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n) mod n = 0" using assms mod_nat_mult_self1 [of 0 m n] by simp -lemma mod_nat_mult_self2_is_0 [simp]: +lemma mod_nat_mult_self2_is_0 [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m) mod n = 0" using assms mod_nat_mult_self2 [of 0 m n] by simp -lemma div_nat_by_1 [simp]: +lemma div_nat_by_1 [simp]: assumes "m \ Nat" shows "m div 1 = m" using assms div_nat_mult_self1_is_id [of m 1] by simp -lemma mod_nat_by_1 [simp]: +lemma mod_nat_by_1 [simp]: assumes "m \ Nat" shows "m mod 1 = 0" using assms mod_nat_mult_self1_is_0[of m 1] by simp -lemma mod_nat_self [simp]: +lemma mod_nat_self [simp]: assumes "n \ Nat" and "0 < n" shows "n mod n = 0" using assms mod_nat_mult_self1_is_0[of 1] by simp -lemma div_nat_self [simp]: +lemma div_nat_self [simp]: assumes "n \ Nat" and "0 < n" shows "n div n = 1" using assms div_nat_mult_self1_is_id [of 1 n] by simp @@ -610,7 +610,7 @@ proof - from 1 2 assms show ?thesis by blast qed -lemma mod_div_nat_trivial [simp]: +lemma mod_div_nat_trivial [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m mod n) div n = 0" proof - @@ -622,7 +622,7 @@ proof - using assms by simp qed -lemma mod_mod_nat_trivial [simp]: +lemma mod_mod_nat_trivial [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m mod n) mod n = m mod n" proof - @@ -632,12 +632,12 @@ proof - finally show ?thesis . qed -lemma dvd_nat_imp_mod_0: +lemma dvd_nat_imp_mod_0: assumes "n dvd m" and "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n = 0" using assms by (simp add: dvd_nat_eq_mod_eq_0) -lemma dvd_nat_div_mult_self: +lemma dvd_nat_div_mult_self: assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "(m div n) * n = m" using assms @@ -645,7 +645,7 @@ using assms mod_div_nat_equality[OF m n pos] by simp -lemma dvd_nat_div_mult: +lemma dvd_nat_div_mult: assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and k: "k \ Nat" shows "(m div n) * k = (m * k) div n" @@ -680,7 +680,7 @@ next with 3 4 show "(b div a) dvd (c div a)" by simp qed (auto simp: assms) -lemma dvd_mod_nat_imp_dvd: +lemma dvd_mod_nat_imp_dvd: assumes 1: "k dvd (m mod n)" and 2: "k dvd n" and k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "k dvd m" @@ -690,7 +690,7 @@ proof - with m n pos show ?thesis by simp qed -(*** TODO +(*** TODO text {* Addition respects modular equivalence. *} text {* Multiplication respects modular equivalence. *} diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f558c286..5009aa4e 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -13,7 +13,7 @@ imports Peano begin text \ - Using the sets @{text upto} we can now define the standard ordering on + Using the sets @{text upto} we can now define the standard ordering on natural numbers. The constant @{text "\"} is defined over the naturals by the axiom (conditional definition) @{text "nat_leq_def"} below; it should be defined over other domains as appropriate later on. @@ -119,12 +119,12 @@ lemma nat_leq_limit: shows "m=n" using assms by (auto simp: nat_leq_def intro: uptoLimit) -lemma nat_leq_trans [trans]: +lemma nat_leq_trans [trans]: assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" shows "k \ n" using assms by (auto simp: nat_leq_def elim: uptoTrans) -lemma nat_leq_antisym: +lemma nat_leq_antisym: assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) shows "m = n" using assms by (auto simp add: nat_leq_def elim: uptoAntisym) @@ -139,7 +139,7 @@ lemma nat_Succ_leqI: shows "Succ[m] \ n" using assms by (auto dest: nat_leq_limit) -lemma nat_Succ_leqD [elim]: +lemma nat_Succ_leqD [elim]: assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" shows "m \ n" proof - @@ -253,7 +253,7 @@ lemma nat_Succ_less_SuccE: shows "P" using assms by simp -lemma nat_not_less0 [simp]: +lemma nat_not_less0 [simp]: assumes "n \ Nat" shows "(n < 0) = FALSE" using assms by auto @@ -263,12 +263,12 @@ lemma nat_less0E (*[elim!]*): shows "P" using assms by simp -lemma nat_less_SuccI: +lemma nat_less_SuccI: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "m < Succ[n]" using assms by auto -lemma nat_Succ_lessD: +lemma nat_Succ_lessD: assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" shows "m < n" using assms by auto @@ -300,19 +300,19 @@ lemma nat_leq_less_trans [trans]: shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) -lemma nat_less_leq_trans [trans]: +lemma nat_less_leq_trans [trans]: assumes "k < m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) text \ Asymmetry. \ -lemma nat_less_not_sym: +lemma nat_less_not_sym: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "(n < m) = FALSE" using assms by auto - -lemma nat_less_asym: + +lemma nat_less_asym: assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" shows "P" proof (rule contradiction) @@ -322,12 +322,12 @@ qed text \ Linearity (totality). \ -lemma nat_less_linear: +lemma nat_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" shows "m < n \ m = n \ n < m" unfolding less_def using nat_leq_linear[OF m n] by blast -lemma nat_leq_less_linear: +lemma nat_leq_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n \ n < m" using assms nat_less_linear[OF m n] by (auto simp: less_def) @@ -347,7 +347,7 @@ lemma nat_not_less_iff_gr_or_eq: shows "(\ m < n) = (m > n \ m = n)" unfolding nat_not_less[OF m n] using assms by (auto simp: less_def) -lemma nat_not_less_eq: +lemma nat_not_less_eq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m < n) = (n < Succ[m])" unfolding nat_not_less[OF m n] using assms by simp @@ -360,12 +360,12 @@ using assms by simp text \often useful, but not active by default\ lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq -lemma nat_not_leq_eq: +lemma nat_not_leq_eq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m \ n) = (Succ[n] \ m)" unfolding nat_not_leq[OF m n] using assms by simp -lemma nat_neq_iff: +lemma nat_neq_iff: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (m < n \ n < m)" using assms nat_less_linear[OF m n] by auto @@ -375,22 +375,22 @@ lemma nat_neq_lessE: shows "(m < n \ R) \ (n < m \ R) \ R" using assms by (auto simp: nat_neq_iff[simplified]) -lemma nat_antisym_conv1: +lemma nat_antisym_conv1: assumes "\(m < n)" and "m \ Nat" and "n \ Nat" shows "(m \ n) = (m = n)" using assms by (auto simp: nat_leq_less) -lemma nat_antisym_conv2: +lemma nat_antisym_conv2: assumes "m \ n" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -lemma nat_antisym_conv3: +lemma nat_antisym_conv3: assumes "\ n < m" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -lemma nat_not_lessD: +lemma nat_not_lessD: assumes "\(m < n)" and "m \ Nat" and "n \ Nat" shows "n \ m" using assms by (simp add: nat_not_order_simps) @@ -439,7 +439,7 @@ using assms by simp text \ "Less than" is antisymmetric, sort of \ -lemma nat_less_antisym: +lemma nat_less_antisym: assumes m: "m \ Nat" and n: "n \ Nat" shows "\ \(n < m); n < Succ[m] \ \ m = n" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) @@ -450,7 +450,7 @@ lemma less_mono_imp_leq_mono: and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" shows "f(i) \ f(j)" using ij proof (auto simp: nat_leq_less[OF j]) - assume "i < j" + assume "i < j" with i j have "f(i) < f(j)" by (rule mono) thus "f(i) \ f(j)" by (simp add: less_imp_leq) next @@ -497,7 +497,7 @@ lemma nat_Succ_lessE: proof - from i have "Succ[i] \ Nat" .. from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" - by (rule nat_lessE) + by (rule nat_lessE) with i that show ?thesis by simp qed @@ -518,7 +518,7 @@ lemma nat_gt0_iff_Succ: using n by (auto dest: nat_ge1_implies_Succ) (* -lemma nat_less_Succ_eq_0_disj: +lemma nat_less_Succ_eq_0_disj: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" using assms by (induct m, auto) diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 6b848505..8bbc8cac 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -295,7 +295,7 @@ proof - thus "P" by (rule sc) qed qed - + lemma succIrrefl: assumes n: "n \ Nat" shows "Succ[n] \ n" @@ -346,7 +346,7 @@ proof - with n m show ?thesis by blast qed -lemma not0_implies_Suc: +lemma not0_implies_Suc: "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" by(rule natCases, auto) @@ -542,8 +542,8 @@ qed subsection \ Primitive Recursive Functions \ text \ - We axiomatize a primitive recursive scheme for functions - with one argument and domain on natural numbers. Later, we + We axiomatize a primitive recursive scheme for functions + with one argument and domain on natural numbers. Later, we use it to define addition, multiplication and difference. \ @@ -553,7 +553,7 @@ axiomatization where primrec_nat: "\f : isAFcn(f) \ DOMAIN f = Nat \ f[0] = e \ (\n \ Nat : f[Succ[n]] = h(n,f[n]))" -lemma bprimrec_nat: +lemma bprimrec_nat: assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" shows "\f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))" proof - @@ -589,7 +589,7 @@ qed lemma bprimrecType_nat: assumes "e \ S" and "\n \ Nat : \x \ S : h(n,x) \ S" - shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ + shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))) \ [Nat \ S]" by (rule primrec_natE[OF assms], auto) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index 80d69b78..2a23721c 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -48,7 +48,7 @@ setup \ Intuitionistic.method_setup @{binding iprover} \ typedecl c text \ - The following (implicit) lifting from the object to the Isabelle + The following (implicit) lifting from the object to the Isabelle meta level is always needed when formalizing a logic. It corresponds to judgments @{text \\ F\} or @{text \\ F\} in textbook presentations, asserting that a formula is considered true (either because @@ -145,8 +145,8 @@ text \ Propositional logic is introduced in a rather non-standard way by declaring constants @{text TRUE} and @{text FALSE} as well as conditional expressions. The ordinary Boolean connectives - are defined in terms of these basic operators in such a way - that they always return either @{text TRUE} or @{text FALSE}, + are defined in terms of these basic operators in such a way + that they always return either @{text TRUE} or @{text FALSE}, irrespectively of their arguments. We also require that equality returns a Boolean value. In this way, we can prove standard equational laws of propositional logic, which is useful for @@ -736,7 +736,7 @@ proof (rule r) qed theorem iffCE [Pure.elim!]: - assumes ab: "A \ B" + assumes ab: "A \ B" and pos: "\A; B\ \ C" and neg: "\\ A; \ B\ \ C" shows "C" proof (rule case_split[of A]) @@ -827,7 +827,7 @@ syntax (ASCII) "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) translations - (* separate parse and print translations for nested quantifiers because + (* separate parse and print translations for nested quantifiers because they operate in opposite directions *) "\x,xs : P" \ "CONST Ex(\x. \xs : P)" "\x,xs : P" \ "\x : CONST Ex(\xs. P)" @@ -900,14 +900,14 @@ qed (rule hyp) lemma chooseI_ex: "\x : P(x) \ P(CHOOSE x : P(x))" by (elim exE chooseI) -lemma chooseI2: +lemma chooseI2: assumes p: "P(a)" and q:"\x. P(x) \ Q(x)" shows "Q(CHOOSE x : P(x))" proof (rule q) from p show "P(CHOOSE x : P(x))" by (rule chooseI) qed -lemma chooseI2_ex: +lemma chooseI2_ex: assumes p: "\x : P(x)" and q:"\x. P(x) \ Q(x)" shows "Q(CHOOSE x : P(x))" proof (rule q) @@ -997,7 +997,7 @@ proof from conj show "B" by (rule conjunctionD2) qed next - assume conj: "A \ B" + assume conj: "A \ B" show "A &&& B" proof - from conj show A .. @@ -1177,7 +1177,7 @@ lemma shows "Q(n)" using assms by blast -lemma +lemma assumes "\P. P(n)" shows "Q(n)" using assms by fast @@ -1321,7 +1321,7 @@ qed lemma trueprop_true_eq: "Trueprop(TRUE = A) \ Trueprop(A)" proof - assume "TRUE = A" + assume "TRUE = A" hence "A = TRUE" by (rule sym) thus "A" by (rule eqTrueD) next @@ -1344,7 +1344,7 @@ lemmas [simp] = trueprop_eq_true trueprop_true_eq TrueAssumption implies_true_equals false_implies_equals refl[THEN eqTrueI] (* (x = x) = TRUE *) - condT (*condThen*) notTrue + condT (*condThen*) notTrue condF (*condElse*) notFalse cond_id cond_cancel false_neq_true[THEN neq_conv_eq_False] (* (FALSE = TRUE) = FALSE *) @@ -1358,13 +1358,13 @@ subsubsection \ Reasoning by cases \ text \ The next bit of code sets up reasoning by cases as a proper Isar - method, so we can write ``proof cases'' etc. Following the + method, so we can write ``proof cases'' etc. Following the development of FOL, we introduce a set of ``shadow connectives'' that will only be used for this purpose. \ (* lemmas cases = case_split [case_names True False] *) -context +context begin qualified definition cases_equal where "cases_equal \ eq" @@ -1393,7 +1393,7 @@ lemmas cases_atomize' = cases_implies_eq cases_conj_eq cases_forall_eq lemmas cases_atomize = cases_atomize' cases_equal_eq lemmas cases_rulify' [symmetric] = cases_atomize' lemmas cases_rulify [symmetric] = cases_atomize -lemmas cases_rulify_fallback = +lemmas cases_rulify_fallback = cases_equal_def cases_implies_def cases_conj_def cases_forall_def cases_true_def cases_false_def @@ -1534,7 +1534,7 @@ text \ Such equations are valid only if the connectives are applied to Boolean operands. For automatic reasoning, we are interested in equations that can be used by - Isabelle's simplifier. We therefore introduce an auxiliary + Isabelle's simplifier. We therefore introduce an auxiliary predicate that is true precisely of Boolean arguments, and an operation that converts arbitrary arguments to an equivalent (in the sense of @{text "\"}) Boolean expression. @@ -1634,7 +1634,7 @@ lemma isBoolTrueFalse: proof - from hyp have "boolify(x) = x" by (unfold isBool_def) hence bx: "boolify(x) \ x" by (rule eq_reflection) - from boolifyTrueFalse[of x] + from boolifyTrueFalse[of x] show ?thesis by (unfold bx) qed @@ -1739,7 +1739,7 @@ ML \ struct open Simpdata; val mksimps_pairs = [(@{const_name not}, (@{thms boolifyFalseI}, true)), - (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] + (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] @ mksimps_pairs; end; @@ -1768,8 +1768,8 @@ using assms by simp text \ The following code rewrites @{text "x = y"} to @{text "FALSE"} in the - presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. - The simplifier is set up so that @{text "y = x"} is already simplified + presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. + The simplifier is set up so that @{text "y = x"} is already simplified to @{text "FALSE"}, this code adds symmetry of disequality to simplification. \ @@ -1785,7 +1785,7 @@ simproc_setup neq ("x = y") = \fn _ => _ $ (Not $ (eq' $ l' $ r')) => Not = @{const \not\} andalso eq' = eq andalso r' aconv lhs andalso l' aconv rhs - | _ $ (eq' $ (eq'' $ l' $ r') $ f') => + | _ $ (eq' $ (eq'' $ l' $ r') $ f') => eq' = @{const \eq\} andalso eq'' = @{const \eq\} andalso f' = @{const \FALSE\} andalso r' aconv lhs andalso l' aconv rhs | _ => false); @@ -1980,7 +1980,7 @@ next qed text \useful with conditionals in hypotheses\ -lemma splitCondAsm: +lemma splitCondAsm: assumes "\x. isBool(Q(x))" shows "Q(IF P THEN t ELSE e) = (\((P \ \Q(t)) \ (\P \ \Q(e))))" using assms by (simp only: splitCond, blast) @@ -1990,7 +1990,7 @@ lemma condCong (*[cong]*): (* strangely, seems to simplify less when active ?! by simp text \not active by default, because too expensive\ -lemma condFullCong: +lemma condFullCong: "\P = Q; Q \ t=t'; \Q \ e=e'\ \ (IF P THEN t ELSE e) = (IF Q THEN t' ELSE e')" by auto @@ -2036,7 +2036,7 @@ using assms by auto text \Commutativity laws are not active by default\ -lemma conj_comms: +lemma conj_comms: "(P \ Q) = (Q \ P)" "(P \ Q \ R) = (Q \ P \ R)" by auto @@ -2067,7 +2067,7 @@ using assms by auto text \Commutativity laws are not active by default\ -lemma disj_comms: +lemma disj_comms: "(P \ Q) = (Q \ P)" "(P \ Q \ R) = (Q \ P \ R)" by auto @@ -2109,7 +2109,7 @@ lemma imp_simps [simp]: "(P \ \P) = (\P)" by auto -lemma imp_cong [cong]: +lemma imp_cong [cong]: "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) = (P' \ Q'))" by auto diff --git a/isabelle/README.html b/isabelle/README.html index ea6896e3..ab8e434b 100644 --- a/isabelle/README.html +++ b/isabelle/README.html @@ -81,7 +81,7 @@

TLA+: Lamport's Temporal Logic of Actions specification language

{x,y} such that f[x] \in S and f[y] \in T.

- + Finally, this theory introduces standard notions for binary relations, such as orderings, equivalence relations and so on. diff --git a/isabelle/SMT.thy b/isabelle/SMT.thy index 1a07942b..1a0f6f87 100644 --- a/isabelle/SMT.thy +++ b/isabelle/SMT.thy @@ -4,7 +4,7 @@ theory SMT imports NatDivision CaseExpressions Strings Integers begin -text {* We can declare the rewriting rules used for normalization +text {* We can declare the rewriting rules used for normalization with the attribute "norm". *} ML {* @@ -43,15 +43,15 @@ theorem [norm]: "{} \\ S = {}" by simp theorem [norm]: "{a1,a2} \ {b1,b2} = {a1,a2,b1,b2}" by simp theorem [norm]: "{a,b} \ S \ a \ S \ b \ S" by simp -theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" +theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" by (simp add:isBool_def) -theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" +theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" by (simp add:isBool_def) theorem [norm]: "\\x. isBool(P(x)) ; \x \ S \\ T : P(x)\ \ \x \ S : x \ T \ P(x)" - by (simp add: bAll_def) + by (simp add: bAll_def) theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} \\ S : P(x)) \ \x \ {a,b} : x \ S \ P(x)" - by (auto simp only: bAll_def) + by (auto simp only: bAll_def) text {* Sets *} @@ -72,7 +72,7 @@ theorem [norm]: "x \ {y \ S : P(y)} \ x \ S \ P(x)" by s theorem [norm]: "S \ T \ \ x \ S : x \ T" by auto theorem [norm]: "S = {} = (\x \ S : FALSE)" by auto -text {* Of the following two rules, the first is applied when S \ {} +text {* Of the following two rules, the first is applied when S \ {} is a hypothesis ; the second, when it is a conclusion. *} theorem [norm]: "S \ {} \ (\ x \ S : TRUE)" by simp theorem [norm]: "S \ {} \ (S = {} \ FALSE)" by simp @@ -81,7 +81,7 @@ text {* For the moment, the interval of numbers ".." is defined only for natural theorem [norm]: "x \ a..b \ x \ Nat \ a \ x \ x \ b" by (simp add: natInterval_def) text {* \b \ Nat; a \ Nat; c \ Nat\ \ a .. b = a .. c = (b < a \ c < a \ b = c) *} -theorems +theorems natIntervalEqual_iff [of a b a c, simplified, standard, norm] natIntervalEqual_iff [of a b c b, simplified, standard, norm] @@ -93,26 +93,26 @@ theorem [norm]: "isAFcn([f EXCEPT ![a] = b])" by simp theorem [norm]: "f = [x \ S \ e(x)] \ isAFcn(f) \ DOMAIN f = S \ (\ x \ S : f[x] = e(x))" by auto theorem [norm]: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\ x \ S : f[x] \ T)" by auto theorem [norm]: "a \ S \ [x \ S \ e(x)][a] = e(a)" by simp -theorem [norm]: +theorem [norm]: assumes "a \ DOMAIN f" - and "g = [f EXCEPT ![a] = b]" - shows "isAFcn(g)" - and "DOMAIN g = DOMAIN f" - and "g[a] = b \ (\ x \ DOMAIN f \\ {a} : g[x] = f[x])" + and "g = [f EXCEPT ![a] = b]" + shows "isAFcn(g)" + and "DOMAIN g = DOMAIN f" + and "g[a] = b \ (\ x \ DOMAIN f \\ {a} : g[x] = f[x])" using assms by auto -theorem [norm]: - "a \ DOMAIN f \ [f EXCEPT ![x] = y][a] = (IF x = a THEN y ELSE f[a])" +theorem [norm]: + "a \ DOMAIN f \ [f EXCEPT ![x] = y][a] = (IF x = a THEN y ELSE f[a])" by (auto simp: except_def) -theorem [norm]: +theorem [norm]: "DOMAIN [x \ S \ e(x)] = S" by simp -theorem [norm]: +theorem [norm]: "DOMAIN [f EXCEPT ![a] = b] = DOMAIN f" by simp -theorem [norm]: - assumes "a \ DOMAIN f" - and "b \ T" - and "[f EXCEPT ![a] = b] \ [S \ T]" - shows "DOMAIN f = S" and "a \ S" and "b \ T" +theorem [norm]: + assumes "a \ DOMAIN f" + and "b \ T" + and "[f EXCEPT ![a] = b] \ [S \ T]" + shows "DOMAIN f = S" and "a \ S" and "b \ T" and "\ x \ S \\ {a} : f[x] \ T" using assms by (simp_all only:FuncSet except_def, auto) @@ -122,64 +122,64 @@ theorem [norm]: theorem [norm]: "[f EXCEPT ![a] = b] = [g EXCEPT ![c] = d] = - (DOMAIN f = DOMAIN g \ + (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : (IF x = a THEN b ELSE f[x]) = (IF x = c THEN d ELSE g[x])))" by (force simp add: fcnEqualIff) text {* IFs *} -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(IF P THEN t ELSE e) \ IF P THEN Q(t) ELSE Q(e)" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(x, IF P THEN t ELSE e) \ IF P THEN Q(x,t) ELSE Q(x,e)" "isBool(P) \ Q(IF P THEN t ELSE e, x) \ IF P THEN Q(t,x) ELSE Q(e,x)" by (auto simp add:isBool_def) (* SM: should the above rather be the following? *) -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(x, IF P THEN t ELSE e) = (IF P THEN Q(x,t) ELSE Q(x,e))" "isBool(P) \ Q(IF P THEN t ELSE e, x) = (IF P THEN Q(t,x) ELSE Q(e,x))" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "isBool(P) \ (IF P THEN t ELSE e)[x] = (IF P THEN t[x] ELSE e[x])" by (auto simp add:isBool_def) -theorem [norm]: - "isBool(P) \ [(IF P THEN t ELSE e) EXCEPT ![x] = y] = +theorem [norm]: + "isBool(P) \ [(IF P THEN t ELSE e) EXCEPT ![x] = y] = (IF P THEN [t EXCEPT ![x] = y] ELSE [e EXCEPT ![x] = y])" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "x = (IF P THEN t ELSE e) \ IF P THEN x = t ELSE x = e" "(IF P THEN t ELSE e) = x \ IF P THEN t = x ELSE e = x" by auto (* SM: should the above rather be the following? *) -theorem [norm]: +theorem [norm]: "(x = (IF P THEN t ELSE e)) = (IF P THEN x = t ELSE x = e)" "((IF P THEN t ELSE e) = x) = (IF P THEN t = x ELSE e = x)" by auto -theorem [norm]: +theorem [norm]: "x \ (IF P THEN t ELSE e) \ IF P THEN x \ t ELSE x \ e" "(IF P THEN t ELSE e) \ S \ IF P THEN t \ S ELSE e \ S" by auto (* SM: similar *) -theorem [norm]: +theorem [norm]: "x \ (IF P THEN t ELSE e) = (IF P THEN x \ t ELSE x \ e)" "(IF P THEN t ELSE e) \ S = (IF P THEN t \ S ELSE e \ S)" by auto -theorem [norm]: +theorem [norm]: "isBool(P) \ x \ (IF P THEN t ELSE e) \ IF P THEN x \ t ELSE x \ e" "isBool(P) \ (IF P THEN t ELSE e) \ x \ IF P THEN t \ x ELSE e \ x" by (auto simp add:isBool_def) (* SM: similar *) -theorem [norm]: +theorem [norm]: "isBool(P) \ x \ (IF P THEN t ELSE e) = (IF P THEN x \ t ELSE x \ e)" "isBool(P) \ (IF P THEN t ELSE e) \ x = (IF P THEN t \ x ELSE e \ x)" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "\isBool(P); P = Q\ \ (IF P THEN t ELSE e) \ (IF Q THEN u ELSE v) \ IF P THEN t \ u ELSE e \ v" by (auto simp add:isBool_def) (* SM: similar *) -theorem [norm]: +theorem [norm]: "\isBool(P); P = Q\ \ (IF P THEN t ELSE e) \ (IF Q THEN u ELSE v) = (IF P THEN t \ u ELSE e \ v)" by (auto simp add:isBool_def) @@ -202,13 +202,13 @@ theorem [norm]: "DOMAIN <<>> = {}" by simp theorem [norm]: "<>[1] = e1" by simp theorem [norm]: "<>[2] = e2" by simp -theorem [norm]: "t = <> \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] = e1 \ t[2] = e2" +theorem [norm]: "t = <> \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] = e1 \ t[2] = e2" by (intro iffI, simp, rule seqEqualI, simp_all, intro LenI, simp_all) theorem [norm]: "t \ S \ T \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] \ S \ t[2] \ T" by (intro iffI, force, auto simp: prod_def Product_def) theorem [norm]: "<> = <> \ a1 = b1 \ a2 = b2" by simp -theorem [norm]: +theorem [norm]: assumes "[i \ 1..2 \ e(i)] = <>" shows "\i \ 1..2 : e(i) = <>[i]" proof - @@ -225,9 +225,9 @@ text {* Records *} theorem [norm]: "r = (''h1'' :> e1 @@ ''h2'' :> e2) \ r[''h1''] = e1 \ r[''h2''] = e2" by simp theorem [norm]: "r \ [''h1'' : e1, ''h2'' : e2] \ r[''h1''] \ e1 \ r[''h2''] \ e2" by auto -theorem [norm]: +theorem [norm]: assumes "h \ DOMAIN r" and "s = [r EXCEPT ![h] = e]" - shows "\f \ DOMAIN r : IF f = h THEN s[f] = e ELSE s[f] = r[f]" + shows "\f \ DOMAIN r : IF f = h THEN s[f] = e ELSE s[f] = r[f]" using assms by auto text {* Arithmetic *} @@ -251,7 +251,7 @@ theorem "\x \ Int; y \ Int; z \ Int\ \x \ Int; y \ Int; z \ Int\ \ (y + x \ z + x) = (y \ z)" sorry theorem "\x \ Int; y \ Int; z \ Int\ \ (x \ y - z) = (x + z \ y)" - apply (rule intCases3[of x y z],simp+) + apply (rule intCases3[of x y z],simp+) apply (simp add: int_diff_def) thm nat_add_right_cancel_leq thm leq_adiff_right_add_left @@ -287,7 +287,7 @@ lemma (*[simp]:*) "n \ Nat \ 1..n = {} \ n = 0" sorry lemma (*[simp]:*) "n \ Nat \ {} = 1..n \ n = 0" sorry lemma OneToNegative (*[simp]*): "n \ Int \ n \ 0 \ 1..n = 1..0" sorry (* -theorem range_empty [simp]: "b < a \ a..b = {}" +theorem range_empty [simp]: "b < a \ a..b = {}" sorry SM: shouldn't this rather be the following *) @@ -298,9 +298,9 @@ lemma (*[simp]:*) "\n \ Int; {1} = 1..n\ \ n lemma (*[simp]:*) "\n \ Int; {1,2} = 1..n\ \ n = 2" sorry lemma (*[simp]:*) "\n \ Int; {1,2,3} = 1..n\ \ n = 3" sorry -(* Seq(S) is defined over intervals of naturals but it should work +(* Seq(S) is defined over intervals of naturals but it should work as well for intervals of integers. *) -lemma SeqInt: "UNION {[1 .. n \ S] : n \ Int} = Seq(S)" +lemma SeqInt: "UNION {[1 .. n \ S] : n \ Int} = Seq(S)" proof - have "(UNION {[1 .. n \ S] : n \ Int}) = (UNION {[1 .. n \ S] : n \ Nat})" (is "?uint = ?unat") @@ -326,7 +326,7 @@ text {* Sequences *} (* NB: "\\circ" is already -- mistakenly -- used for relation composition in Tuples.thy !!*) definition "Concat" :: "[c, c] \ c" (infixl "\\o" 50) - where "Concat(s,t) \ [i \ 1 .. (Len(s) + Len(t)) \ + where "Concat(s,t) \ [i \ 1 .. (Len(s) + Len(t)) \ IF i \ Len(s) THEN s[i] ELSE t[i - Len(s)]]" (* SM: changed definition of Tail to use "--" instead of "-", which shouldn't change the semantics *) @@ -363,7 +363,7 @@ theorem [norm,simp]: "isASeq(s) \ s \\o \\ = s" theorem [norm]: "Append(\\,e) = \e\" .. theorem [norm]: "Len(\\) = 0" by simp theorem [norm,simp]: "Tail(\\) = \\" by (auto simp: Tail_def OneToNegative) -theorem [norm]: "SubSeq(s,m,n) = \\ \ s = \\ \ 1..(1 + n - m) = {}" +theorem [norm]: "SubSeq(s,m,n) = \\ \ s = \\ \ 1..(1 + n - m) = {}" by (simp add: SubSeq_def) (* Rules about Len *) @@ -372,7 +372,7 @@ theorem [norm]: "Len(\e1,e2\) = 2" by simp theorem SeqLenGeqZero [norm]: "isASeq(s) \ 0 \ Len(s)" by simp -theorem [simp]: +theorem [simp]: assumes 1: "isASeq(s)" and 2: "s \ \\" shows "0 < Len(s)" proof - @@ -387,21 +387,21 @@ theorem LenNonEmptySeq (*[simp]*): shows "Len(s) - 1 \ Nat" sorry -theorem [norm,simp]: "n \ Nat \ Len([x \ 1..n \ e]) = n" +theorem [norm,simp]: "n \ Nat \ Len([x \ 1..n \ e]) = n" by (intro LenI, auto) -theorem [norm,simp]: "n \ Int \ n \ 0 \ Len([x \ 1..n \ e]) = 0" +theorem [norm,simp]: "n \ Int \ n \ 0 \ Len([x \ 1..n \ e]) = 0" by (intro LenI, simp_all add: OneToNegative) theorem [norm]: "\isASeq(s); s \ \\\ \ Len(Tail(s)) = Len(s) - 1" by (force simp: Tail_def LenNonEmptySeq) -theorem [norm]: "\isASeq(s); isASeq(t)\ \ Len(s \\o t) = Len(s) + Len(t)" +theorem [norm]: "\isASeq(s); isASeq(t)\ \ Len(s \\o t) = Len(s) + Len(t)" by (force simp: Concat_def) theorem [norm]: "isASeq(s) \ Len(Append(s,e)) = Len(s) + 1" by (force simp: Append_def) theorem [norm]: "isASeq(s) \ Len([s EXCEPT ![x] = y]) = Len(s)" by force -theorem [simp]: +theorem [simp]: assumes 1: "isASeq(s)" and 2: "isASeq(t)" and 3: "s \\o t = \\" shows "s = \\ \ t = \\" proof - @@ -430,7 +430,7 @@ next thus ?thesis by (auto simp: Tail_def) qed -theorem SubSeqIsASeq [norm]: +theorem SubSeqIsASeq [norm]: assumes s: "isASeq(s)" and m: "m \ Int" and n: "n \ Int" shows "isASeq(SubSeq(s,m,n))" proof - @@ -447,12 +447,12 @@ qed theorem [norm]: "\isASeq(s)\ \ isASeq([s EXCEPT ![a] = b])" by (simp add: except_def) (* Rules of the form seq[i] *) -theorem [norm]: "\isASeq(s); i \ 1..Len(s)\ \ Append(s,e)[i] = s[i]" +theorem [norm]: "\isASeq(s); i \ 1..Len(s)\ \ Append(s,e)[i] = s[i]" by (auto simp: inNatInterval_iff) theorem [norm]: "isASeq(s) \ Append(s,e)[Len(s) + 1] = e" by auto theorem TailElt [norm]: "\isASeq(s); i \ 1..(Len(s) - 1)\ \ Tail(s)[i] = s[i + 1]" by (simp add: Tail_def) theorem [norm]: "i \ 1..(1 + n - m) \ SubSeq(s,m,n)[i] = s[i + m - 1]" by (simp add: SubSeq_def) -theorem ConcatEltFirst [norm]: +theorem ConcatEltFirst [norm]: assumes s: "isASeq(s)" and t: "isASeq(t)" and i: "i \ 1..Len(s)" shows "(s \\o t)[i] = s[i]" proof - @@ -462,7 +462,7 @@ proof - ultimately show ?thesis by (simp add: Concat_def) qed -theorem ConcatEltSecond [norm]: +theorem ConcatEltSecond [norm]: assumes s: "isASeq(s)" and t: "isASeq(t)" and i: "i \ (Len(s)+1) .. (Len(s)+Len(t))" shows "(s \\o t)[i] = t[i - Len(s)]" proof - @@ -488,12 +488,12 @@ theorem [norm]: "e1 \ S \ e2 \ S \ \e1,e2\< theorem [norm]: "\e1,e2\ \ Seq(S) \ e1 \ S \ e2 \ S" by force theorem [norm]: "n \ Nat \ [i \ 1 .. n \ e(i)] \ Seq(S) = (\i \ 1..n : e(i) \ S)" by (simp add:Seq_def FuncSet) -theorem [norm]: +theorem [norm]: assumes n: "n \ Int" shows "[i \ 1 .. n \ e(i)] \ Seq(S) = (\i \ 1..n : e(i) \ S)" apply (simp add: Seq_def FuncSet) - using n apply (rule intCases, simp_all) + using n apply (rule intCases, simp_all) by force -theorem AppendInSeq [norm]: +theorem AppendInSeq [norm]: assumes s: "isASeq(s)" shows "(Append(s,e) \ Seq(S)) \ (e \ S \ s \ Seq(S))" proof @@ -519,7 +519,7 @@ qed theorem [norm]: "\e \ S; s \ Seq(S)\ \ Append(s,e) \ Seq(S)" by auto -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" shows "Tail(s) \ Seq(S)" proof (cases "s = \\") @@ -546,7 +546,7 @@ next qed qed -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" and t: "t \ Seq(S)" shows "s \\o t \ Seq(S)" proof @@ -566,7 +566,7 @@ next qed qed -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" and m: "m \ Int" and n: "n \ Int" and mn: "(n < m) \ (m \ 1 .. Len(s) \ n \ m .. Len(s))" shows "SubSeq(s,m,n) \ Seq(S)" @@ -592,7 +592,7 @@ qed (* Rules of the form x = seq *) -theorem [norm]: +theorem [norm]: assumes m: "m \ Int" and n: "n \ Int" and t: "t = SubSeq(s,m,n)" and i: "i \ m .. n" shows "t[1 + i - m] = s[i]" proof - @@ -601,7 +601,7 @@ proof - with t show ?thesis by (simp add: SubSeq_def) qed -theorem [norm]: +theorem [norm]: assumes n: "n \ 1..Len(s)" shows "SubSeq(s,n,n) = \s[n]\" proof - @@ -612,32 +612,32 @@ qed theorem [norm]: "x = Len(s) \ isASeq(s) \ x \ Nat \ DOMAIN s = 1..x" by simp -theorem [norm]: +theorem [norm]: assumes "isASeq(s)" and "x = Append(s,e)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) + 1)" - and "\i \ 1 .. Len(s): x[i] = s[i]" + and "\i \ 1 .. Len(s): x[i] = s[i]" and "x[Len(s) + 1] = e" using assms by (simp_all, auto simp: Append_def) - -theorem [norm]: - assumes "isASeq(s)" + +theorem [norm]: + assumes "isASeq(s)" and "isASeq(t)" and "x = (s \\o t)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) + Len(t))" - and "\i \ 1 .. Len(s): x[i] = s[i]" + and "\i \ 1 .. Len(s): x[i] = s[i]" and "\i \ (Len(s) + 1) .. (Len(s) + Len(t)): x[i] = t[i - Len(s)]" using assms by (auto simp: ConcatEltFirst ConcatEltSecond) -theorem [norm]: +theorem [norm]: assumes "isASeq(s)" and "s \ \\" and "x = Tail(s)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) - 1)" - and "\i \ 1 .. (Len(s) - 1): x[i] = s[i + 1]" + and "\i \ 1 .. (Len(s) - 1): x[i] = s[i + 1]" using assms by (simp_all add: TailIsASeq TailElt) text {* Setting up the smt tactic. *} @@ -645,7 +645,7 @@ text {* Setting up the smt tactic. *} thm norm ML {* -fun smt_tac ctxt = +fun smt_tac ctxt = (* let val my_ctxt = ctxt |> Simplifier.map_simpset @@ -662,4 +662,4 @@ method_setup smt = {* (fn _ => fn ctxt => SIMPLE_METHOD (CHANGED (smt_tac ctxt))) *} -end \ No newline at end of file +end diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index e37ff62f..ca9e6bb4 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -50,7 +50,7 @@ text \ text \ Concrete syntax: set comprehension \ -(*** +(*** We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} because logical constants (such as setOfAll) must have fixed arity. Multiple bindings will be introduced as shorthands involving tuples @@ -197,7 +197,7 @@ print_translation \ in (Syntax.const "@bAll") $ x' $ S $ P' end | bAll_tr' _ = raise Match; - in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), + in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), (@{const_syntax "bAll"}, (fn _ => bAll_tr'))] end \ @@ -443,7 +443,7 @@ proof - let ?ch = "CHOOSE x \ A : P(x)" from 1 2 have "t \ A \ P(t)" .. hence "?ch \ A \ P(?ch)" - by (unfold bChoose_def, rule chooseI) + by (unfold bChoose_def, rule chooseI) thus ?thesis .. qed @@ -454,7 +454,7 @@ proof - let ?ch = "CHOOSE x \ A : P(x)" from 1 2 have "t \ A \ P(t)" .. hence "?ch \ A \ P(?ch)" - by (unfold bChoose_def, rule chooseI) + by (unfold bChoose_def, rule chooseI) thus ?thesis .. qed @@ -548,13 +548,13 @@ by (force intro: condI elim: condE) lemma condIn [simp]: "((IF P THEN a ELSE b) \ S) = ((P \ a \ S) \ (\P \ b \ S))" by (force intro: condI elim: condE) -lemma inCondI (*[TC]*): +lemma inCondI (*[TC]*): (* adding this as (unsafe) introduction rule cripples blast & co -- why? *) assumes "P \ c \ A" and "\P \ c \ B" shows "c \ (IF P THEN A ELSE B)" using assms by auto -lemma condInI (*[TC]*): +lemma condInI (*[TC]*): (* too general to add as introduction rule? *) assumes "P \ a \ S" and "\P \ b \ S" shows "(IF P THEN a ELSE b) \ S" @@ -572,7 +572,7 @@ lemma condInE: (* too general to add as elimination rule? *) shows "Q" using assms by auto -lemma subsetCond [simp]: +lemma subsetCond [simp]: "(A \ (IF P THEN S ELSE T)) = ((P \ A \ S) \ (\P \ A \ T))" by (blast intro: condI elim: condE) @@ -956,7 +956,7 @@ by (auto simp: upair_def) lemma upair_iff: "(c \ upair(a,b)) = (c=a \ c=b)" by (blast intro: upairI1 upairI2 elim: upairE) -lemma unionI1: +lemma unionI1: assumes "a \ A" shows "a \ A \ B" using assms by (auto simp: union_def upair_iff) @@ -1015,7 +1015,7 @@ lemma equalAddElt_iff: "(S = addElt(a,A)) = (a \ S \ A \ S \< by blast lemma addEltEqualAddElt: - "(addElt(a,A) = addElt(b,B)) = + "(addElt(a,A) = addElt(b,B)) = (a \ addElt(b,B) \ A \ addElt(b,B) \ b \ addElt(a,A) \ B \ addElt(a,A))" by (auto simp: addEltEqual_iff) @@ -1073,7 +1073,7 @@ proof - proof assume 1: "?lhs" show ?rhs proof (cases "a \ B") - case True + case True from 1 have "B \ {a} \ SUBSET A" by blast moreover from True 1 have "B = addElt(a, B \ {a})" by blast @@ -1343,8 +1343,8 @@ text \ Union is an AC operator: can be added to simp where appropriate \ Lemmas useful for simplifying enumerated sets are active by default \ -lemmas enumeratedSetSimps [simp] = - addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice +lemmas enumeratedSetSimps [simp] = + addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice interAddEltLeft interAddEltRight unionAddEltLeft unionAddEltRight setminusAddEltLeft (* included below diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index fdeb66ce..e951d12c 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -51,7 +51,7 @@ syntax text {* The following parse and print translations convert between the internal - and external representations of strings. Strings are written using + and external representations of strings. Strings are written using two single quotes in Isabelle, such as \verb|''abc''|. Note that the empty string is just the empty sequence in \tlaplus{}, so \verb|''''| gets printed as @{term "\\"}. Single characters are printed in the form @@ -70,7 +70,7 @@ parse_ast_translation {* (* convert an ML character to a TLA+ Char *) fun mkChar c = if Symbol.is_ascii c - then Ast.Appl [Ast.Constant "Strings.char", + then Ast.Appl [Ast.Constant "Strings.char", mkNibble (ord c div 16), mkNibble (ord c mod 16)] else error ("Non-ASCII symbol: " ^ quote c); @@ -87,7 +87,7 @@ parse_ast_translation {* | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); (* parse AST translation for strings *) - fun string_ast_tr [Ast.Variable xstr] = + fun string_ast_tr [Ast.Variable xstr] = list2TupleReverse (rev (Lexicon.explode_xstr xstr)) | string_ast_tr asts = raise Ast.AST ("string_ast_tr", asts); in @@ -110,7 +110,7 @@ oops print_ast_translation {* let - (* convert a nibble to an ML integer -- because translation macros have + (* convert a nibble to an ML integer -- because translation macros have already been applied, we see constants "0" through "15", not Succ[...] terms! *) fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 @@ -134,7 +134,7 @@ print_ast_translation {* fun destNbls nb1 nb2 = let val specials = raw_explode "\"\\`'" val c = chr (destNibble nb1 * 16 + destNibble nb2) - in if not (member (op =) specials c) andalso Symbol.is_ascii c + in if not (member (op =) specials c) andalso Symbol.is_ascii c andalso Symbol.is_printable c then c else raise Match end; @@ -155,7 +155,7 @@ print_ast_translation {* (* print AST translation for single characters that do not occur in a string *) fun char_ast_tr' [nb1, nb2] = - Ast.Appl [Ast.Constant @{syntax_const "_Char"}, + Ast.Appl [Ast.Constant @{syntax_const "_Char"}, list2String [destNbls nb1 nb2]] | char_ast_tr' _ = raise Match; diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 06b5d145..4dfb7d6b 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -100,7 +100,7 @@ using assms by auto lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] lemma seqEqualI: - assumes "isASeq(s)" and "isASeq(t)" + assumes "isASeq(s)" and "isASeq(t)" and "Len(s) = Len(t)" and "\k \ 1 .. Len(t) : s[k] = t[k]" shows "s = t" using assms by (intro fcnEqual[of s t], auto) @@ -160,7 +160,7 @@ by (auto elim: seqInSeqRange) subsection {* Sequences via @{text emptySeq} and @{text Append} *} text {* - Sequences can be built from the constructors @{text emptySeq} + Sequences can be built from the constructors @{text emptySeq} (written @{text "\\"}) and Append. *} @@ -264,7 +264,7 @@ lemma appendElt2 (*[simp]*): assumes "isASeq(s)" shows "Append(s,e)[Succ[Len(s)]] = e" using assms by (auto simp: Append_def) - + lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2, standard] lemma isAppend [intro!]: @@ -343,7 +343,7 @@ by (simp add: emptySeq_def) lemma imageAppend [simp]: assumes s: "isASeq(s)" - shows "Image(Append(s,e), A) = + shows "Image(Append(s,e), A) = (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" unfolding appendExtend[OF s] using assms by (auto elim!: inNatIntervalE, force+) @@ -451,7 +451,7 @@ lemma "\0,1\ \ \1,2,3\" by simp lemma "(\a,b\ = \c,d\) = (a=c \ b=d)" by simp ***) -text {* +text {* \tlaplus{} has a form of quantification over tuples written $\exists \langle x,y,z \rangle \in S: P(x,y,z)$. We cannot give a generic definition of this form for arbitrary tuples, but introduce input syntax @@ -496,7 +496,7 @@ text {* @{text x}, @{text y}) such that @{text "f[x] \ S"} and @{text "f[y] \ T"}. Typically, elements of such a function set will be constructed as @{text "(x :> s)@@(y :> t)"}. - This notation for sets of finite functions generalizes similar + This notation for sets of finite functions generalizes similar \tlaplus{} notation for records. Internally, the set is represented as @{text "EnumFuncSet(\x,y\, \S,T\)"}, @@ -528,7 +528,7 @@ qed lemma EnumFuncSetE [elim!]: assumes "f \ EnumFuncSet(doms, rngs)" - and "\f \ [Range(doms) \ UNION Range(rngs)]; + and "\f \ [Range(doms) \ UNION Range(rngs)]; \i \ DOMAIN doms : f[doms[i]] \ rngs[i] \ \ P" shows "P" using assms by (auto simp: EnumFuncSet_def) @@ -550,11 +550,11 @@ parse_ast_translation {* (* get_doms_ranges extracts the lists of arguments and ranges from the arms of a "domrngs" expression. The order of the ASTs is reversed. *) - fun get_doms_ranges (Ast.Appl[Ast.Constant "@domrng", d, r]) = + fun get_doms_ranges (Ast.Appl[Ast.Constant "@domrng", d, r]) = (* base case: one domain, one range *) ([d], [r]) - | get_doms_ranges (Ast.Appl[Ast.Constant "@domrngs", - Ast.Appl[Ast.Constant "@domrng", d, r], + | get_doms_ranges (Ast.Appl[Ast.Constant "@domrngs", + Ast.Appl[Ast.Constant "@domrng", d, r], pairs]) = (* one domrng, followed by remaining doms and ranges *) let val (ds, rs) = get_doms_ranges pairs @@ -645,7 +645,7 @@ text {* *} definition Product -where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : +where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : \i \ 1 .. Len(s) : f[i] \ s[i] }" lemma inProductI [intro!]: @@ -666,7 +666,7 @@ using assms by (auto simp add: Product_def) lemma inProductE [elim!]: assumes "p \ Product(s)" and "isASeq(s)" - and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; + and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; \k \ 1 .. Len(s) : p[k] \ s[k] \ \ P" shows "P" using assms by (auto simp add: Product_def) @@ -836,7 +836,7 @@ lemma inSetOfPairsI [intro]: using assms by (auto simp: setOfPairs_def) lemma inSetOfPairsE [elim!]: -- {* converse true only if $R$ is a relation *} - assumes 1: "z \ setOfPairs(R, e)" + assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" proof - @@ -846,12 +846,12 @@ proof - with pR pz show "P" by (intro 3, auto) qed -lemmas setOfPairsEqualI = +lemmas setOfPairsEqualI = setEqualI [where A = "setOfPairs(R,f)", standard,intro!] setEqualI [where B = "setOfPairs(R,f)", standard,intro!] -lemma setOfPairs_triv [simp]: - assumes s: "R \ A \ B" +lemma setOfPairs_triv [simp]: + assumes s: "R \ A \ B" shows "{ \x,y\ : \x,y\ \ R } = R" using assms by auto @@ -886,7 +886,7 @@ definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" definition rel_comp :: "[c,c] => c" (infixr "\" 75) -- {* binary relation composition *} -where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : +where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) @@ -1018,20 +1018,20 @@ by force subsubsection {* Converse relation *} -lemmas converseEqualI = +lemmas converseEqualI = setEqualI [where A = "r^-1", standard, intro!] setEqualI [where B = "r^-1", standard, intro!] -lemma converse_iff [iff]: +lemma converse_iff [iff]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" using r prodProj by (auto simp: converse_def) -lemma converseI [intro!]: +lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" unfolding converse_def by auto -lemma converseD [sym]: +lemma converseD [sym]: assumes r: "r \ A \ B" shows "\a,b\ \ r^-1 \ \b,a\ \ r" using converse_iff[OF r] by simp @@ -1053,7 +1053,7 @@ proof - show P by simp qed -lemma converse_converse [simp]: +lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" using assms prodProj by (auto elim!: converseE) @@ -1081,18 +1081,18 @@ lemma converse_mono_2: using assms prodProj by auto lemma converse_mono: - assumes r:"r \ A \ B" and s:"s \ A \ B" + assumes r:"r \ A \ B" and s:"s \ A \ B" shows "r^-1 \ s^-1 = (r \ s)" using converse_mono_1[OF r s] converse_mono_2[OF r s] by blast (* from HOL *) -lemma reflexive_converse [simp]: +lemma reflexive_converse [simp]: "r \ A \ B \ reflexive(A, r^-1) = reflexive(A,r)" unfolding reflexive_def by auto -lemma symmetric_converse [simp]: +lemma symmetric_converse [simp]: "r \ A \ B \ symmetric(r^-1) = symmetric(r)" unfolding symmetric_def by auto @@ -1100,11 +1100,11 @@ lemma antisymmetric_converse [simp]: "r \ A \ B \ antisymmetric(r^-1) = antisymmetric(r)" unfolding antisymmetric_def by auto -lemma transitive_converse [simp]: +lemma transitive_converse [simp]: "r \ A \ B \ transitive(r^-1) = transitive(r)" unfolding transitive_def by auto -lemma symmetric_iff_converse_eq: +lemma symmetric_iff_converse_eq: assumes r:"r \ A \ B" shows "symmetric(r) = (r^-1 = r)" proof auto @@ -1125,7 +1125,7 @@ qed subsubsection {* Identity relation over a set *} -lemmas idEqualI = +lemmas idEqualI = setEqualI [where A = "Id(S)", standard, intro!] setEqualI [where B = "Id(S)", standard, intro!] @@ -1175,11 +1175,11 @@ unfolding rel_range_def Id_def by auto subsubsection {* Composition of relations *} -lemmas compEqualI = +lemmas compEqualI = setEqualI [where A = "r \ s", standard, intro!] setEqualI [where B = "r \ s", standard, intro!] -lemma compI [intro]: +lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "\ \a,b\ \ s; \b,c\ \ r \ \ \a,c\ \ r \ s" using assms unfolding rel_comp_def by auto @@ -1189,7 +1189,7 @@ lemma compE [elim!]: shows "(\x y z. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r \ P) \ P" using assms unfolding rel_comp_def by auto -lemma compEpair: +lemma compEpair: assumes "\a,c\ \ r \ s" and "r \ B \ C" and s: "s \ A \ B" shows "\\b. \ \a,b\ \ s; \b,c\ \ r \ \ P \ \ P" using assms by auto @@ -1204,7 +1204,7 @@ lemma rel_comp_in_prodE (*[elim]*): shows "p \ A \ C" using assms by force -lemma converse_comp: +lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "((r \ s)^-1) = (s^-1 \ r^-1)" (is "?lhs = ?rhs") proof @@ -1227,7 +1227,7 @@ next with r s show "x \ ?lhs" by (auto dest: converseSubset) qed -lemma R_comp_Id [simp]: +lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" using r proof auto @@ -1239,7 +1239,7 @@ using r proof auto with 1 show "p \ R \ Id(B)" by simp qed -lemma Id_comp_R [simp]: +lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" using r proof auto @@ -1257,7 +1257,7 @@ unfolding rel_comp_def by auto lemma rel_comp_empty2 [simp]: "R \ {} = {}" unfolding rel_comp_def by auto -lemma comp_assoc: +lemma comp_assoc: assumes t: "T \ A \ B" and s: "S \ B \ C" and r: "R \ C \ D" shows "(R \ S) \ T = R \ (S \ T)" proof @@ -1296,7 +1296,7 @@ next qed qed -lemma rel_comp_mono: +lemma rel_comp_mono: assumes hr': "r' \ r" and hs': "s' \ s" shows "(r' \ s') \ (r \ s)" unfolding rel_comp_def using subrel_dom[OF hs'] subrel_ran[OF hr'] @@ -1358,7 +1358,7 @@ by (blast intro: symmetricI dest: symmetricE) text {* Antisymmetry *} -lemma antisymmetricI [intro]: +lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" unfolding antisymmetric_def by blast diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index d8a797ea..46a9796f 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -347,7 +347,7 @@ by blast lemma zenon_notin_cup_0 : "x \\notin A \\cup B ==> x \\notin A" by blast - + lemma zenon_notin_cup_1 : "x \\notin A \\cup B ==> x \\notin B" by blast @@ -542,7 +542,7 @@ proof - have h10: "x = v | x ~= v" by (rule excluded_middle) show "except (f, v, e)[x] \\in B" using h4 h9 h5 except_def by auto - qed + qed show "except (f, v, e) \\in FuncSet (A, B)" using h6 h7 h8 FuncSet by blast qed @@ -2689,7 +2689,7 @@ proof - using h0 h1 by auto from hf - have hh0: "?gxx" + have hh0: "?gxx" by (rule zenon_disjE1 [OF _ hg1x]) have hi: "cas \\in UNION {CaseArm (<<>>[i], <<>>[i]) : i \\in DOMAIN <<>>}" @@ -2795,7 +2795,7 @@ proof - using h0 h1 by auto from hf - have hh0: "?gxx" + have hh0: "?gxx" by (rule zenon_disjE1 [OF _ hg1x]) have hi: "cas \\in UNION {CaseArm (<<>>[i], <<>>[i]) : i \\in DOMAIN <<>>}" diff --git a/isabelle/simplifier_setup.ML b/isabelle/simplifier_setup.ML index 6957e6a5..8df22c2a 100644 --- a/isabelle/simplifier_setup.ML +++ b/isabelle/simplifier_setup.ML @@ -21,8 +21,8 @@ struct fun mk_eq th = case Thm.concl_of th of Const(@{const_name \Pure.eq\},_) $ _ $ _ => th | _ $ (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _) => th RS @{thm eq_reflection} - | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ - (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => + | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ + (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => th RS (@{thm neq_conv_eq_False} RS @{thm eq_reflection}) | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection} @@ -35,7 +35,7 @@ struct *) fun lift_meta_to_obj_eq ctxt i st = let - fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = + fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = 1 + count_imp Q | count_imp _ = 0; val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) @@ -59,7 +59,7 @@ struct (* Congruence rules expecting = instead of == *) fun mk_meta_cong ctxt rl = - let val rl' = Seq.hd (TRYALL (fn i => fn st => + let val rl' = Seq.hd (TRYALL (fn i => fn st => resolve_tac ctxt [lift_meta_to_obj_eq ctxt i st] i st) rl) in (rl' RS @{thm eq_reflection}) handle THM _ => @@ -76,10 +76,10 @@ struct allows us to discard "type-correctness" hypotheses in conditional rewrite rules that are added later on. *) val mksimps_pairs = - [(@{const_name \imp\}, ([@{thm mp}], false)), + [(@{const_name \imp\}, ([@{thm mp}], false)), (@{const_name \conj\}, ([@{thm conjD1}, @{thm conjD2}], false)), - (@{const_name \All\}, ([@{thm spec}], false)), - (@{const_name \TRUE\}, ([], false)), + (@{const_name \All\}, ([@{thm spec}], false)), + (@{const_name \TRUE\}, ([], false)), (@{const_name \FALSE\}, ([], false)), (@{const_name \cond\}, ([@{thm if_bool_equiv_conj} RS @{thm iffD1}], false)) ]; @@ -100,7 +100,7 @@ struct (case head_of p of Const(a,_) => (case AList.lookup (op =) pairs a of - SOME (rls, b) => + SOME (rls, b) => let val rrls = maps atoms (res_fixed rls) handle THM _ => [thm] in @@ -113,10 +113,10 @@ struct end; in atoms end - fun mksimps pairs ctxt = + fun mksimps pairs ctxt = map_filter (try mk_eq) o mk_atomize ctxt pairs o Variable.gen_all ctxt - fun unsafe_solver_tac ctxt = + fun unsafe_solver_tac ctxt = let val sol_thms = reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI}:: Simplifier.prems_of ctxt; @@ -141,7 +141,7 @@ struct val safe_solver = mk_solver "PredicateLogic safe" safe_solver_tac; (* Basic infrastructure for simplification *) - val PL_basic_ss = + val PL_basic_ss = empty_simpset @{context} setSSolver safe_solver setSolver unsafe_solver From 06290b7ea721e83e055c439623d5d84a5f6ac434 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 21 Dec 2020 20:57:04 +0200 Subject: [PATCH 003/167] MAI: update `NatArith` to Isabelle2020 The addition of `fixes` statements to theorems in the theory `NatOrderings` is for making explicit the order of theorem parameters, used when instantiating a theorem with explicit substitutions. --- isabelle/NatArith.thy | 572 ++++++++++++++++++++++++++++++-------- isabelle/NatOrderings.thy | 5 + 2 files changed, 462 insertions(+), 115 deletions(-) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 6b17dc37..179a0799 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -2,25 +2,29 @@ Author: Hernan Vanzetto, LORIA Copyright (C) 2009-2011 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 + Version: Isabelle2020 Time-stamp: <2011-10-11 17:39:37 merz> *) -header {* Arithmetic (except division) over natural numbers *} +section \ Arithmetic (except division) over natural numbers \ theory NatArith imports NatOrderings begin -ML{* +named_theorems algebra_simps "algebra simplification rules" + +(* +ML \ structure AlgebraSimps = Named_Thms(val name = "algebra_simps" val description = "algebra simplification rules"); -*} +\ setup AlgebraSimps.setup +*) -text{* +text \ The rewrites accumulated in @{text algebra_simps} deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and @@ -29,10 +33,10 @@ text{* Of course it also works for fields, but it knows nothing about multiplicative inverses or division. This should be catered for by @{text field_simps}. -*} +\ -subsection {* Addition of natural numbers *} +subsection \ Addition of natural numbers \ definition addnat where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]])" @@ -40,7 +44,7 @@ where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ c" (infixl "+" 65) where nat_add_def: "\m \ Nat; n \ Nat\ \ (m + n) \ addnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma addnatType: assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" @@ -50,7 +54,7 @@ lemma addIsNat [intro!,simp]: assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" unfolding nat_add_def[OF assms] using assms addnatType by blast -text {* Base case and Inductive case *} +text \ Base case and Inductive case \ lemma addnat_0: assumes "m \ Nat" shows "addnat(m)[0] = m" @@ -92,7 +96,7 @@ lemma add_Succ_shift_nat: shows "Succ[m] + n = m + Succ[n]" using assms by simp -text {* Commutativity *} +text \ Commutativity \ lemma add_commute_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -100,14 +104,14 @@ lemma add_commute_nat [algebra_simps]: using n apply induct using assms by auto -text {* Associativity *} +text \ Associativity \ lemma add_assoc_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + p) = (m + n) + p" using assms by (induct, simp_all) -text {* Cancellation rules *} +text \ Cancellation rules \ lemma add_left_cancel_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -126,9 +130,9 @@ lemma add_left_commute_nat [algebra_simps]: using assms by(simp only: add_assoc_nat add_commute_nat) (* from HOL/OrderedGroups.thy *) -theorems add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat +lemmas add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat -text {* Reasoning about @{text "m + n = 0"}, etc. *} +text \ Reasoning about @{text "m + n = 0"}, etc. \ lemma add_is_0_nat [iff]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -157,7 +161,7 @@ using n apply (rule natCases) done -subsection {* Multiplication of natural numbers *} +subsection \ Multiplication of natural numbers \ definition multnat where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m)" @@ -165,7 +169,7 @@ where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ c" (infixl "*" 70) where nat_mult_def: "\m \ Nat; n \ Nat\ \ m * n \ multnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma multnatType: assumes "m \ Nat" shows "multnat(m) \ [Nat \ Nat]" @@ -176,16 +180,17 @@ lemma multIsNat [intro!, simp]: shows "m * n \ Nat" unfolding nat_mult_def[OF assms] using assms multnatType by blast -text {* Base case and Inductive step *} +text \ Base case and Inductive step \ lemma multnat_0: assumes "m \ Nat" shows "multnat(m)[0] = 0" unfolding multnat_def by (rule primrec_natE, auto simp: assms) -lemma mult_0_nat [simp]: -- {* neutral element *} +lemma mult_0_nat [simp]: (* -- "neutral element" *) assumes n: "n \ Nat" shows "n * 0 = 0" by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) + lemma multnat_Succ: assumes m: "m \ Nat" and n: "n \ Nat" shows "multnat(m)[Succ[n]] = multnat(m)[n] + m" @@ -213,14 +218,14 @@ lemma mult_Succ_left_nat [simp]: using n apply induct using m by (simp_all add: add_ac_nat) -text {* Commutativity *} +text \ Commutativity \ lemma mult_commute_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m * n = n * m" using assms by (induct, simp_all) -text {* Distributivity *} +text \ Distributivity \ lemma add_mult_distrib_left_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -257,13 +262,13 @@ proof - by simp qed -text {* Identity element *} +text \ Identity element \ (* used for arithmetic simplification setup *) lemma mult_1_right_nat: "a \ Nat \ a * 1 = a" by simp lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp -text {* Associativity *} +text \ Associativity \ lemma mult_assoc_nat[algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -271,7 +276,7 @@ lemma mult_assoc_nat[algebra_simps]: using m apply induct using assms by (auto simp add: add_mult_distrib_right_nat) -text {* Reasoning about @{text "m * n = 0"}, etc. *} +text \ Reasoning about @{text "m * n = 0"}, etc. \ lemma mult_is_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -294,7 +299,7 @@ proof - finally show ?thesis . qed -text {* Cancellation rules *} +text \ Cancellation rules \ lemma mult_cancel1_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -346,10 +351,10 @@ lemma mult_left_commute_nat[algebra_simps]: using assms by(simp only: mult_commute_nat mult_assoc_nat) (* from HOL/OrderedGroups.thy *) -theorems mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat +lemmas mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat -subsection {* Predecessor *} +subsection \ Predecessor \ definition Pred where "Pred \ [n \ Nat \ IF n=0 THEN 0 ELSE CHOOSE x \ Nat : n=Succ[x]]" @@ -369,13 +374,13 @@ lemma Pred_in_nat [intro!, simp]: assumes "n \ Nat" shows "Pred[n] \ Nat" using assms by (cases "n", auto) -subsection {* Difference of natural numbers *} +subsection \ Difference of natural numbers \ -text {* +text \ We define a form of difference @{text "--"} of natural numbers that cuts off at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes called ``arithmetic difference''. -*} +\ definition adiffnat where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]])" @@ -383,7 +388,7 @@ where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \< definition adiff (infixl "--" 65) where nat_adiff_def: "\m \ Nat; n \ Nat\ \ (m -- n) \ adiffnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma adiffnatType: assumes "m \ Nat" shows "adiffnat(m) \ [Nat \ Nat]" @@ -393,7 +398,7 @@ lemma adiffIsNat [intro!, simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ Nat" unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast -text {* Neutral element and Inductive step *} +text \ Neutral element and Inductive step \ lemma adiffnat_0: assumes "m \ Nat" shows "adiffnat(m)[0] = m" @@ -424,7 +429,7 @@ lemma adiff_0_eq_0_nat [simp]: shows "0 -- n = 0" using n apply induct by (simp_all add: adiff_Succ_nat) -text {* Reasoning about @{text "m -- m = 0"}, etc. *} +text \ Reasoning about @{text "m -- m = 0"}, etc. \ lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -437,7 +442,7 @@ lemma adiff_self_eq_0_nat [simp]: shows "m -- m = 0" using m apply induct by auto -text {* Associativity *} +text \ Associativity \ lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" @@ -450,14 +455,14 @@ lemma Succ_adiff_adiff_nat [simp]: shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" using assms by (simp add: adiff_adiff_left_nat) -text {* Commutativity *} +text \ Commutativity \ lemma adiff_commute_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i -- j -- k = i -- k -- j" using assms by (simp add: adiff_adiff_left_nat add_commute_nat) -text {* Cancellation rules *} +text \ Cancellation rules \ lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -487,7 +492,7 @@ lemma adiff_add_0_nat [simp]: using n apply induct using assms by simp_all -text {* Difference distributes over multiplication *} +text \ Difference distributes over multiplication \ lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -500,15 +505,15 @@ lemma adiff_mult_distrib2_nat [algebra_simps]: shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) --- {* NOT added as rewrites, since sometimes they are used from right-to-left *} +(* -- "NOT added as rewrites, since sometimes they are used from right-to-left" *) lemmas nat_distrib = add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat -subsection {* Additional arithmetic theorems *} +subsection \ Additional arithmetic theorems \ -subsubsection {* Monotonicity of Addition *} +subsubsection \ Monotonicity of Addition \ lemma Succ_pred [simp]: assumes n: "n \ Nat" @@ -539,9 +544,118 @@ using k apply induct using assms by simp_all lemma add_gr_0 [simp]: - assumes "n \ Nat" and "m \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" -using assms by (auto dest: nat_gt0_implies_Succ nat_not_lessD) + assumes ndom: "n \ Nat" and mdom: "m \ Nat" + shows "(m + n > 0) = (m > 0 \ n > 0)" + proof - + have s1_1: "m + n > 0 ==> m > 0 \ n > 0" + proof - + assume s1_1_asm: "m + n > 0" + have s2_1: "m \ 0 \ n \ 0" + proof - + have s3_1: "m = 0 \ n = 0 ==> FALSE" + proof - + assume s3_1_asm: "m = 0 \ n = 0" + have s4_1: "m + n = 0" + using s3_1_asm ndom mdom add_is_0_nat by auto + have s4_2: "m + n \ 0" + unfolding less_def using s1_1_asm by auto + show "FALSE" + using s4_1 s4_2 by auto + qed + show ?thesis + using s3_1 by auto + qed + show "m > 0 \ n > 0" + using s2_1 ndom mdom nat_gt0_not0 by auto + qed + have s1_2: "m > 0 \ n > 0 ==> m + n > 0" + proof - + assume s1_2_asm: "m > 0 \ n > 0" + have s2_1: "m > 0 ==> m + n > 0" + proof - + assume s2_1_asm: "m > 0" + have s3_1: "m + n > 0 + n" + proof - + have s4_1: "0 < m" + using s2_1_asm by auto + have s4_2: "(0 + n < m + n) = (0 < m)" + using ndom mdom zeroIsNat nat_add_right_cancel_less + by blast + have s4_3: "0 + n < m + n" + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + have s3_2: "0 + n = n" + using ndom add_0_nat add_commute_nat zeroIsNat + by auto + have s3_3: "n >= 0" + using ndom nat_zero_leq by auto + have s3_4: "m + n >= 0 + n" + using s3_1 by (auto simp: less_def) + have s3_5: "m + n >= n" + using s3_4 s3_2 by auto + have s3_6: "m + n >= 0" + using s3_5 s3_3 nat_leq_trans + zeroIsNat ndom mdom addIsNat by auto + have s3_7: "m + n = 0 ==> FALSE" + proof - + assume s3_7_asm: "m + n = 0" + have s4_1: "m = 0 \ n = 0" + using s3_7_asm mdom ndom add_is_0_nat + by auto + show "FALSE" + using s4_1 s2_1_asm by (auto simp: less_def) + qed + show "m + n > 0" + unfolding less_def + using s3_6 s3_7 by auto + qed + have s2_2: "n > 0 ==> m + n > 0" + proof - + assume s2_2_asm: "n > 0" + have s3_1: "m + n > m + 0" + proof - + have s4_1: "0 < n" + using s2_2_asm by auto + have s4_2: "(m + 0 < m + n) = (0 < n)" + using ndom mdom zeroIsNat nat_add_left_cancel_less + by blast + have s4_3: "m + 0 < m + n" + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + have s3_2: "m + 0 = m" + using mdom add_0_nat by auto + have s3_3: "m >= 0" + using mdom nat_zero_leq by auto + have s3_4: "m + n >= m + 0" + using s3_1 by (auto simp: less_def) + have s3_5: "m + n >= m" + using s3_4 s3_2 by auto + have s3_6: "m + n >= 0" + using s3_5 s3_3 nat_leq_trans + zeroIsNat ndom mdom addIsNat by auto + have s3_7: "m + n = 0 ==> FALSE" + proof - + assume s3_7_asm: "m + n = 0" + have s4_1: "m = 0 \ n = 0" + using s3_7_asm mdom ndom add_is_0_nat + by auto + show "FALSE" + using s4_1 s2_2_asm by (auto simp: less_def) + qed + show "m + n > 0" + unfolding less_def + using s3_6 s3_7 by auto + qed + show "m + n > 0" + using s1_2_asm s2_1 s2_2 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed lemma less_imp_Succ_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -571,11 +685,11 @@ apply (induct n p rule: diffInduct) using assms by simp_all -subsubsection {* (Partially) Ordered Groups *} +subsubsection \ (Partially) Ordered Groups \ --- {* The two following lemmas are just ``one half'' of +(* -- "The two following lemmas are just ``one half'' of @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above. *} + proved above." *) lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" @@ -586,7 +700,7 @@ lemma add_leq_right_mono: shows "a \ b \ a + c \ b + c" using assms by simp -text {* non-strict, in both arguments *} +text \ non-strict, in both arguments \ lemma add_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c \ d \ a + c \ b + d" @@ -596,18 +710,43 @@ using assms nat_leq_trans[of "a + c" "b + c" "b + d"] by simp --- {* Similar for strict less than. *} +(* -- "Similar for strict less than." *) lemma add_less_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ c + a < c + b" -using assms by simp + proof - + have s1_1: "a <= b ==> c + a <= c + b" + using a b c add_leq_left_mono by auto + have s1_2: "a \ b ==> (c + a) \ (c + b)" + using a b c add_left_cancel_nat by auto + have s1_3: "a < b ==> c + a < c + b" + proof - + assume s1_3_asm: "a < b" + have s2_1: "a <= b" + using s1_3_asm by (auto simp: less_def) + have s2_2: "a \ b" + using s1_3_asm by (auto simp: less_def) + have s2_3: "c + a <= c + b" + using s1_1 s2_1 by auto + have s2_4: "(c + a) \ (c + b)" + using s1_2 s2_2 by auto + show "c + a < c + b" + using s2_3 s2_4 by (auto simp: less_def) + qed + show "a < b \ c + a < c + b" + using a b c s1_3 by auto + qed +(* +using assms add_leq_left_mono[OF "a" "b" "c"] add_left_cancel_nat[OF "c" "b" "a"] + by (auto simp: nat_less_Succ_iff_leq nat_gt0_not0) +*) lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" -using assms by simp +using assms add_less_left_mono add_commute_nat by auto -text{* Strict monotonicity in both arguments *} +text \ Strict monotonicity in both arguments \ lemma add_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c < d \ a + c < b + d" @@ -671,14 +810,36 @@ lemma trans_leq_add2: using assms by (auto elim: nat_leq_trans) lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" + assumes ij: "i < j" and idom: "i \ Nat" and + jdom: "j \ Nat" and mdom: "m \ Nat" shows "i < j + m" -using assms by (auto elim: nat_less_leq_trans) + proof - + have s1_1: "i <= i + m" + proof - + have s2_1: "0 <= m" + using mdom nat_zero_leq by auto + have s2_2: "i + 0 <= i + m" + using idom zeroIsNat mdom s2_1 nat_add_left_cancel_leq by auto + have s2_3: "i + 0 = i" + using idom add_0_nat by auto + show ?thesis + using s2_2 s2_3 by auto + qed + have s1_2: "i + m < j + m" + using ij idom jdom mdom add_less_right_mono by auto + have s1_3: "i + m \ Nat" + using idom mdom addIsNat by auto + have s1_4: "j + m \ Nat" + using jdom mdom addIsNat by auto + show ?thesis + using s1_1 s1_2 idom s1_3 s1_4 nat_leq_less_trans + by auto + qed lemma trans_less_add2: assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" -using assms by (auto elim: nat_less_leq_trans) +using assms trans_less_add1 add_commute_nat by auto lemma add_lessD1: assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" @@ -711,15 +872,32 @@ lemma add_leqE: using assms by (blast dest: add_leqD1 add_leqD2) lemma leq_add_left_false [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "n \ 0" + assumes mdom: "m \ Nat" and ndom: "n \ Nat" and n0: "n \ 0" shows "m + n \ m = FALSE" -using assms nat_leq_less[of "m + n" m] add_eq_self_zero_nat[OF m n] by auto + proof - + have s1_1: "m + n \ m = (m + n < m \ m + n = m)" + proof - + define p where "p \ m + n" + have s2_1: "p \ m = (p < m \ p = m)" + using mdom nat_leq_less by auto + show ?thesis + using s2_1 by (auto simp: p_def) + qed + have s1_2: "(m + n = m) = (n = 0)" + using mdom ndom add_eq_self_zero_nat[of "m" "n"] by auto + have s1_3: "(m + n <= m) = (m + n < m)" + using n0 s1_1 s1_2 by auto + have s1_4: "(m + n < m) = FALSE" + using mdom ndom not_add_less1[of "m" "n"] by auto + show ?thesis + using s1_3 s1_4 by auto + qed -subsubsection {* More results about arithmetic difference *} +subsubsection \ More results about arithmetic difference \ -text {* Addition is the inverse of subtraction: - if @{term "n \ m"} then @{term "n + (m -- n) = m"}. *} +text \ Addition is the inverse of subtraction: + if @{term "n \ m"} then @{term "n + (m -- n) = m"}. \ lemma add_adiff_inverse: assumes "m \ Nat" and "n \ Nat" shows "\(m < n) \ n + (m -- n) = m" @@ -752,7 +930,7 @@ lemma adiff_leq_self [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ m" apply (induct m n rule: diffInduct) -using assms by simp_all +using assms by auto lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -771,10 +949,14 @@ qed lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" -proof - - from j n have "j -- n \ Nat" "j -- n \ j" by simp+ - with j k jk show ?thesis by (auto elim: nat_leq_less_trans) -qed + proof - + have s1_1: "j -- n \ Nat" + using j n by simp+ + have s1_2: "j -- n \ j" + using j n by simp+ + show ?thesis + using j k jk s1_1 s1_2 nat_leq_less_trans by auto + qed lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" @@ -827,15 +1009,17 @@ lemma zero_less_adiff [simp]: by (induct m n rule: diffInduct, simp_all add: assms) lemma less_imp_add_positive: - assumes "i < j" and i: "i \ Nat" and j: "j \ Nat" + assumes ij: "i < j" and i: "i \ Nat" and j: "j \ Nat" shows "\k \ Nat: 0 < k \ i + k = j" proof - - from assms have "i \ j" by auto - with i j have "i + (j -- i) = j" by simp - moreover - from assms have "j -- i \ Nat" "0 < j -- i" by simp+ - ultimately - show ?thesis by blast + have s1_1: "i + (j -- i) = j" + using i j ij le_add_adiff_inverse by (auto simp: less_def) + have s1_2: "j -- i \ Nat" + using assms by simp+ + have s1_3: "0 < j -- i" + using assms zero_less_adiff by auto + show ?thesis + using s1_1 s1_2 s1_3 by blast qed lemma leq_adiff_right_add_left: @@ -866,7 +1050,7 @@ using p h apply (induct) using n by auto -subsubsection {* Monotonicity of Multiplication *} +subsubsection \ Monotonicity of Multiplication \ (* from HOL/Ring_and_Field.thy *) @@ -882,7 +1066,7 @@ lemma mult_leq_right_mono: using c apply induct using 1 a b by (simp_all add: add_leq_mono add_commute_nat) -text {* @{text "\"} monotonicity, BOTH arguments *} +text \ @{text "\"} monotonicity, BOTH arguments \ lemma mult_leq_mono: assumes 1: "i \ j" "k \ l" @@ -894,29 +1078,95 @@ using assms nat_leq_trans[of "i * k" "j * k" "j * l"] by simp -text {* strict, in 1st argument *} +text \ strict, in 1st argument \ lemma mult_less_left_mono: - assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" + assumes ij: "i < j" and k0: "0 < k" and + idom: "i \ Nat" and jdom: "j \ Nat" and kdom: "k \ Nat" shows "k * i < k * j" -using 1 -proof (auto simp: nat_gt0_iff_Succ[OF k]) - fix m - assume m: "m \ Nat" and "i < j" - with m i j show "Succ[m] * i < Succ[m] * j" - by (induct m, simp_all add: add_less_mono) -qed + proof - + have s1_1: "Succ[0] * i < Succ[0] * j" + using mult_1_left_nat ij idom jdom by auto + have s1_2: "\ n. \ + n \ Nat; + Succ[n] * i < Succ[n] * j + \ \ + Succ[Succ[n]] * i < Succ[Succ[n]] * j" + proof - + fix "n" :: "c" + assume s1_2_ndom: "n \ Nat" + assume s1_2_induct: "Succ[n] * i < Succ[n] * j" + have s2_1: "Succ[n] * i + i < Succ[n] * j + j" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "Succ[n] * i \ Nat" + using s3_1 idom multIsNat by auto + have s3_3: "Succ[n] * j \ Nat" + using s3_1 jdom multIsNat by auto + show ?thesis + using s3_2 s3_3 idom jdom s1_2_induct ij + add_less_mono by auto + qed + have s2_2: "i * Succ[n] + i < j * Succ[n] + j" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "Succ[n] * i = i * Succ[n]" + using s3_1 idom mult_commute_nat by auto + have s3_3: "Succ[n] * j = j * Succ[n]" + using s3_1 jdom mult_commute_nat by auto + show ?thesis + using s2_1 s3_2 s3_3 by auto + qed + have s2_3: "i * Succ[Succ[n]] < j * Succ[Succ[n]]" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "i * Succ[Succ[n]] = i * Succ[n] + i" + using idom s3_1 multnat_Succ by (auto simp: nat_mult_def) + have s3_3: "j * Succ[Succ[n]] = j * Succ[n] + j" + using jdom s3_1 multnat_Succ by (auto simp: nat_mult_def) + show ?thesis + using s2_2 s3_2 s3_3 by auto + qed + have s2_4: "i * Succ[Succ[n]] = Succ[Succ[n]] * i" + proof - + have s3_1: "Succ[Succ[n]] \ Nat" + using s1_2_ndom succIsNat by auto + show ?thesis + using s3_1 idom mult_commute_nat by auto + qed + have s2_5: "j * Succ[Succ[n]] = Succ[Succ[n]] * j" + proof - + have s3_1: "Succ[Succ[n]] \ Nat" + using s1_2_ndom succIsNat by auto + show ?thesis + using s3_1 jdom mult_commute_nat by auto + qed + show "Succ[Succ[n]] * i < Succ[Succ[n]] * j" + using s2_3 s2_4 s2_5 by auto + qed + have s1_3: "\ n \ Nat: Succ[n] * i < Succ[n] * j" + using s1_1 s1_2 + natInduct[of "\ q. Succ[q] * i < Succ[q] * j"] by auto + show ?thesis + using s1_3 nat_gt0_iff_Succ[of "k"] k0 kdom by auto + qed lemma mult_less_right_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i * k < j * k" -using 1 -proof (auto simp: nat_gt0_iff_Succ[OF k]) - fix m - assume m: "m \ Nat" and "i < j" - with m i j show "i * Succ[m] < j * Succ[m]" - by (induct m, simp_all add: add_less_mono) -qed + proof - + have s1_1: "k * i < k * j" + using 1 i j k mult_less_left_mono by auto + have s1_2: "k * i = i * k" + using i k mult_commute_nat by auto + have s1_3: "k * j = j * k" + using j k mult_commute_nat by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed lemma nat_0_less_mult_iff [simp]: assumes i: "i \ Nat" and j: "j \ Nat" @@ -929,7 +1179,7 @@ done lemma one_leq_mult_iff (*[simp]*): assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 \ m * n) = (1 \ m \ 1 \ n)" -using assms by simp +using assms nat_gt0_not0 by simp lemma mult_less_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" @@ -966,47 +1216,140 @@ next qed lemma mult_less_self_left [dest]: - assumes less: "n*k < n" and n: "n \ Nat" and k: "k \ Nat" + assumes less: "n * k < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" -using k assms by (cases, auto) + proof - + have s1_1: "1 \ k ==> FALSE" + proof - + assume s1_1_asm: "1 \ k" + have s2_1: "n * 1 \ n * k" + using n k oneIsNat s1_1_asm mult_leq_left_mono[of "1" "k" "n"] + by auto + have s2_2: "n \ n * k" + proof - + have s3_1: "n * 1 = n" + using n mult_1_right_nat by auto + show ?thesis + using s2_1 s3_1 by auto + qed + define p where "p \ n" + define q where "q \ n * k" + have s2_3: "q < p" + unfolding p_def q_def using less by auto + have s2_4: "n < p" + proof - + have s3_1: "q \ Nat" + unfolding q_def using n k multIsNat by auto + have s3_2: "p \ Nat" + unfolding p_def using n by auto + have s3_3: "n \ q" + unfolding q_def using s2_2 by auto + show ?thesis + using s3_1 s3_2 n s3_3 s2_3 + nat_leq_less_trans[of "n" "q" "p"] + by auto + qed + have s2_5: "n \ n" + using s2_4 by (auto simp: less_def p_def) + show "FALSE" + using s2_5 by auto + qed + have s1_2: "\ (1 \ k)" + using s1_1 by auto + show ?thesis + using s1_2 k nat_not_leq_one[of "k"] by auto + qed lemma mult_less_self_right [dest]: assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" -using k assms by (cases, auto) + proof - + have s1_1: "n * k < n" + proof - + have s2_1: "k * n = n * k" + using n k mult_commute_nat by auto + show ?thesis + using less s2_1 by auto + qed + show ?thesis + using s1_1 n k mult_less_self_left by auto + qed lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" -using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) - assume 1: "k*m \ k*n" and 2: "0 < k" - show "m \ n" - proof (rule contradiction) - assume "\(m \ n)" - with 2 m n k have "k*n < k*m" by (simp add: nat_not_order_simps mult_less_left_mono) - with m n k have "\(k*m \ k*n)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed + proof - + have s1_1: "(k * m \ k * n) ==> (k = 0 \ m \ n)" + proof - + assume s1_1_asm: "k * m \ k * n" + have s2_1: "\ (k = 0) ==> m \ n" + proof - + assume s2_1_asm: "\ (k = 0)" + have s3_1: "k > 0" + using k nat_gt0_not0 s2_1_asm by auto + have s3_2: "\(m \ n) ==> FALSE" + proof - + assume s3_2_asm: "\(m \ n)" + have s2_1: "n < m" + using s3_2_asm m n nat_not_leq[of "m" "n"] by auto + have s2_2: "k * n < k * m" + using s3_1 s2_1 m n k mult_less_left_mono[of "n" "m" "k"] by auto + with m n k have s2_3: "\ (k * m \ k * n)" + by (simp add: nat_not_order_simps) + show "FALSE" + using s2_2 s2_3 s1_1_asm by auto + qed + show "m \ n" + using s3_2 by auto + qed + show "k = 0 \ m \ n" + using s2_1 by auto + qed + have s1_2: "(k = 0 \ m \ n) ==> (k * m \ k * n)" + proof - + assume s1_2_asm: "k = 0 \ m \ n" + have s2_1: "k = 0 ==> k * m \ k * n" + proof - + assume s2_1_asm: "k = 0" + have s3_1: "k * m = 0" + using s2_1_asm m mult_0_left_nat by auto + have s3_2: "k * n = 0" + using s2_1_asm n mult_0_left_nat by auto + show "k * m \ k * n" + using s3_1 s3_2 by auto + qed + have s2_2: "m \ n ==> k * m \ k * n" + proof - + assume s2_2_asm: "m \ n" + show "k * m \ k * n" + using s2_2_asm m n k mult_leq_left_mono by auto + qed + show "k * m \ k * n" + using s1_2_asm s2_1 s2_2 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" -using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) - assume 1: "m*k \ n*k" and 2: "0 < k" - show "m \ n" - proof (rule contradiction) - assume "\(m \ n)" - with 2 m n k have "n*k < m*k" by (simp add: nat_not_order_simps mult_less_right_mono) - with m n k have "\(m*k \ n*k)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed + proof - + have s1_1: "(k * m \ k * n) = (k = 0 \ m \ n)" + using m n k mult_leq_cancel_left by auto + have s1_2: "k * m = m * k" + using k m mult_commute_nat by auto + have s1_3: "k * n = n * k" + using k n mult_commute_nat by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" -using assms by (simp del: mult_Succ_left_nat) +using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] + nat_zero_less_Succ[of "k"] by auto lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" @@ -1023,7 +1366,7 @@ lemma nat_leq_cube: shows "m \ m * (m * m)" using m by (cases, auto) -text {* Lemma for @{text gcd} *} +text \ Lemma for @{text gcd} \ lemma mult_eq_self_implies_10: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") @@ -1033,5 +1376,4 @@ proof - finally show ?thesis . qed - end diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index 5009aa4e..75b44e09 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -207,6 +207,8 @@ using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) text \ Reduce @{text "\"} to @{text "<"}. \ lemma nat_leq_less: + fixes "m" :: "c" + fixes "n" :: "c" assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" using assms by (auto simp: less_def) @@ -296,6 +298,9 @@ proof - qed lemma nat_leq_less_trans [trans]: + fixes "k" :: "c" + fixes "m" :: "c" + fixes "n" :: "c" assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) From bf2edac9c1e78e779ad17e0e53767990c8df56f4 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:03:55 +0200 Subject: [PATCH 004/167] MAI: port sectioning and text of theory `NatDivision` to Isabelle2020 --- isabelle/NatDivision.thy | 36 ++++++++++++++++++------------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index abef1aed..fcc8a7f2 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:39:56 merz> *) -header {* The division operators div and mod on Naturals *} +section \ The division operators div and mod on Naturals \ theory NatDivision imports NatArith Tuples begin -subsection {* The divisibility relation *} +subsection \ The divisibility relation \ definition dvd (infixl "dvd" 50) where "a \ Nat \ b \ Nat \ b dvd a \ (\k \ Nat : a = b * k)" @@ -151,9 +151,9 @@ proof - with a b' c' show ?thesis by blast qed -subsection {* Division on @{const Nat} *} +subsection \ Division on @{const Nat} \ -text {* +text \ We define division and modulo over @{const Nat} by means of a characteristic relation with two input arguments @{term "m"}, @{term "n"} and two output arguments @@ -162,13 +162,13 @@ text {* The following definition works for natural numbers, but also for possibly negative integers. Obviously, the second disjunct cannot hold for natural numbers. -*} +\ definition divmod_rel where "divmod_rel(m,n,q,r) \ m = q * n + r \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" -text {* @{const divmod_rel} is total if $n$ is non-zero. *} +text \ @{const divmod_rel} is total if $n$ is non-zero. \ lemma divmod_rel_ex: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" @@ -201,7 +201,7 @@ proof - with pos that show ?thesis by (auto simp: divmod_rel_def) qed -text {* @{const divmod_rel} has unique solutions in the natural numbers. *} +text \ @{const divmod_rel} has unique solutions in the natural numbers. \ lemma divmod_rel_unique_div: assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" and m: "m \ Nat" and n: "n \ Nat" @@ -242,10 +242,10 @@ proof - with assms show ?thesis by (auto simp: divmod_rel_def) qed -text {* +text \ We instantiate divisibility on the natural numbers by means of @{const divmod_rel}: -*} +\ definition divmodNat where "divmodNat(m,n) \ CHOOSE z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" @@ -282,9 +282,9 @@ proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) qed -text {* +text \ We now define division and modulus over natural numbers. -*} +\ definition div (infixr "div" 70) where div_nat_def: "\m \ Nat; n \ Nat\ \ m div n \ divmodNat(m,n)[1]" @@ -331,7 +331,7 @@ lemma mod_nat_less_divisor: shows "m mod n < n" using assms divmod_rel_div_mod_nat[OF assms] by (simp add: divmod_rel_def) -text {* ``Recursive'' computation of @{text div} and @{text mod}. *} +text \ ``Recursive'' computation of @{text div} and @{text mod}. \ lemma divmodNat_base: assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" @@ -386,7 +386,7 @@ proof - qed -text {* The ''recursion'' equations for @{text div} and @{text mod} *} +text \ The ''recursion'' equations for @{text div} and @{text mod} \ lemma div_nat_less [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" @@ -419,7 +419,7 @@ using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] by simp -subsection {* Facts about @{const div} and @{const mod} *} +subsection \ Facts about @{const div} and @{const mod} \ lemma mod_div_nat_equality [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" @@ -691,13 +691,13 @@ proof - qed (*** TODO -text {* Addition respects modular equivalence. *} +text \ Addition respects modular equivalence. \ -text {* Multiplication respects modular equivalence. *} +text \ Multiplication respects modular equivalence. \ -text {* Negation respects modular equivalence. *} +text \ Negation respects modular equivalence. \ -text {* Subtraction respects modular equivalence. *} +text \ Subtraction respects modular equivalence. \ *) end From 25b04f0d11d2ff4863ff39c908a82a99f25af998 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:04:27 +0200 Subject: [PATCH 005/167] API: change numerals to definitions in theory `Peano` because abbreviations do not work well with Isabelle (nested applications of `Succ`). --- isabelle/Peano.thy | 36 ++++++++++++++++++++++++++++++++++-- 1 file changed, 34 insertions(+), 2 deletions(-) diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 8bbc8cac..1f19fe45 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -154,6 +154,38 @@ where "Nat \ DOMAIN Succ" definition zero :: "c" ("0") where "0 \ CHOOSE Z : PeanoAxioms(Nat, Z, Succ)" +abbreviation "one \ Succ[0]" + notation "one" ("1") +definition two :: "c" ("2") + where "2 \ Succ[1]" +definition three :: "c" ("3") + where "3 \ Succ[2]" +definition four :: "c" ("4") + where "4 \ Succ[3]" +definition five :: "c" ("5") + where "five \ Succ[4]" +definition six :: "c" ("6") + where "six \ Succ[5]" +definition seven :: "c" ("7") + where "seven \ Succ[6]" +definition eight :: "c" ("8") + where "eight \ Succ[7]" +definition nine :: "c" ("9") + where "nine \ Succ[8]" +definition ten :: "c" ("10") + where "ten \ Succ[9]" +definition eleven :: "c" ("11") + where "eleven \ Succ[10]" +definition twelve :: "c" ("12") + where "twelve \ Succ[11]" +definition thirteen :: "c" ("13") + where "thirteen \ Succ[12]" +definition fourteen :: "c" ("14") + where "fourteen \ Succ[13]" +definition fifteen :: "c" ("15") + where "fifteen \ Succ[14]" + +(* abbreviation "one \ Succ[0]" notation "one" ("1") abbreviation "two \ Succ[1]" @@ -184,7 +216,7 @@ abbreviation "fourteen \ Succ[13]" notation "fourteen" ("14") abbreviation "fifteen \ Succ[14]" notation "fifteen" ("15") - +*) lemma peanoNatZeroSucc: "PeanoAxioms(Nat, 0, Succ)" proof - @@ -223,7 +255,7 @@ lemma oneIsNat [intro!,simp]: "1 \ Nat" by simp lemma twoIsNat [intro!,simp]: "2 \ Nat" -by simp +unfolding two_def by simp lemma [simp]: assumes "n \ Nat" From 9fc6c089433850e0540f8769691ae75a533772e2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:20:09 +0200 Subject: [PATCH 006/167] ENH: add theory `NewIntegers` --- isabelle/NewIntegers.thy | 12578 +++++++++++++++++++++++++++++++++++++ 1 file changed, 12578 insertions(+) create mode 100644 isabelle/NewIntegers.thy diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy new file mode 100644 index 00000000..6c7ae08c --- /dev/null +++ b/isabelle/NewIntegers.thy @@ -0,0 +1,12578 @@ +(* An attempt at defining the integers. + +The steps in Isabelle/Isar are named `si_j` where `s` stands for "step", +`i` is the proof level, and `j` the step number. + +Author: Ioannis Filippidis +All rights reserved. Licensed under 2-clause BSD. +*) +theory NewIntegers +imports + NatArith + (*NatDivision + Tuples*) +begin + +(* +Properties to prove: + typeness, + commutativity, associativity, distributivity + transitivity, symmetry, antisymmetry, reflexivity, + monotonicity +*) + +(*----------------------------------------------------------------------------*) +(* Overriding notation. *) + +no_notation NatOrderings.leq (infixl "<=" 50) +no_notation NatOrderings.geq (infixl ">=" 50) +no_notation NatOrderings.less (infixl "<" 50) +no_notation NatOrderings.greater (infixl ">" 50) +no_notation NatOrderings.natInterval ("(_ .. _)" [90,90] 70) + +no_notation NatArith.arith_add (infixl "+" 65) +(* no_notation NatArith.adiff (infixl "--" 65) *) +no_notation NatArith.mult (infixl "*" 70) + +(* no_notation NatDivision.dvd (infixl "dvd" 50) *) + + +(*----------------------------------------------------------------------------*) +(* The minus operator. *) + + +(* +The minus operator can be defined, instead of axiomatized, as follows. + +minus(x) == + LET singletons == {<>: u \in Nat} + IN + IF x \in Nat THEN + <> + ELSE IF x \in Singletons THEN + CHOOSE u \in Nat: x = <> + ELSE + CHOOSE u \notin (Nat \cup Singletons): TRUE +*) +consts + "minus" :: "[c] \ c" ("-._" [75] 75) + +axiomatization where + neg0 [simp]: "-.0 = 0" +and + neg_neg: "n \ Nat \ -.n \ Nat ==> + -.-.n = n" +and + neg_not_in_nat: + "n \ (Nat \ {0}) ==> + -.n \ Nat" + + +(*----------------------------------------------------------------------------*) +(* The set of integers. *) + +definition Int +where "Int \ + Nat \ {-.n: n \ Nat}" + + +definition negNat :: "c" +where "negNat \ {-.n: n \ Nat}" + + +(*----------------------------------------------------------------------------*) +(* Successor and predecessor on integers. *) + +definition iSucc :: "c" +where "iSucc \ + [i \ Int \ IF i \ Nat THEN Succ[i] + ELSE -.Pred[-.i]]" + + +definition iPred :: "c" +where "iPred \ + [i \ Int \ IF i \ Nat \ {0} + THEN Pred[i] + ELSE -.Succ[-.i]]" + + +(*----------------------------------------------------------------------------*) +(* Recursive definitions on the integers. *) + +(* TODO: replace axiom with proof of existence. *) +axiomatization where + primrec_int: " + \ f: + isAFcn(f) \ + DOMAIN f = Int \ + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n])) + " + +(*----------------------------------------------------------------------------*) +(* Arithmetic operators. *) + + +(* Incrementing and decrementing function. *) +definition addint :: "[c] \ c" +where "addint(m) \ + CHOOSE g \ [Int -> Int]: + g[0] = m \ + (\ n \ Nat: + g[Succ[n]] = iSucc[g[n]] \ + g[-.Succ[n]] = iPred[g[-.n]])" + + +(* Addition on integers. *) +definition add :: "[c, c] \ c" (infixl "+" 65) +where "add(m, n) \ addint(m)[n]" + + +(* Subtraction on integers. *) +definition diff :: "[c, c] \ c" (infixl "-" 65) +where "diff(m, n) \ add(m, -.n)" + + +(* Adding and subtracting function. *) +definition multint :: "[c] \ c" +where "multint(m) \ + CHOOSE g \ [Int -> Int]: + g[0] = 0 \ + (\ n \ Nat: + g[Succ[n]] = add(g[n], m) \ + g[-.Succ[n]] = add(g[-.n], -.m))" + + +(* Multiplication on integers. *) +definition mult :: "[c,c] \ c" (infixl "*" 70) +where "mult(m, n) \ multint(m)[n]" + + +(* Divisibility on the integers. *) +definition dvd :: "[c, c] \ c" (infixl "dvd" 50) +where "b dvd a \ + \ k \ Int: a = mult(b, k)" + + +(* Order on the integers. +This definition is motivated by proofs, where equality is easier to +reason about (adding on both sides, multiplying both sides). +*) +definition leq :: "[c,c] \ c" (infixl "<=" 50) +where "leq(m, n) \ + \ k \ Nat: + add(m, k) = n" + + +abbreviation (input) + geq :: "[c, c] \ c" (infixl ">=" 50) + where "geq(x, y) \ leq(y, x)" + + +definition less :: "[c, c] \ c" (infixl "<" 50) +where "less(a, b) \ leq(a, b) + \ (a \ b)" + + +abbreviation (input) + greater :: "[c, c] \ c" (infixl ">" 50) + where "greater(x, y) \ less(y, x)" + + +definition interval :: "[c, c] \ c" ("(_ .. _)" [90,90] 70) +where "m .. n \ + {k \\in Int: m <= k \ k <= n}" + +(*----------------------------------------------------------------------------*) + + +theorem minus_involutive[simp]: + assumes "n \ Int" + shows "-.-.n = n" + proof - + have s1_1: "n \ Nat \ (\ k \ Nat: n = -.k)" + using assms by (auto simp: Int_def) + have s1_2: "n \ Nat ==> -.-.n = n" + using neg_neg by auto + have s1_3: "n \ Nat ==> -.-.n = n" + proof - + { + assume s2_1: "n \ Nat" + have s2_2: "\ k \ Nat: n = -.k" + using s1_1 s2_1 by auto + have s2_3: "\ k \ Nat: -.n = -.-.k" + using s2_2 by auto + have s2_4: "\ k \ Nat: k = -.n" + using s2_3 neg_neg by auto + have s2_5: "-.n \ Nat" + using s2_4 by auto + have s2_6: "-.-.n = n" + using s2_5 neg_neg by auto + } + from this show "n \ Nat ==> -.-.n = n" by auto + qed + show "-.-.n = n" using s1_2 s1_3 by iprover + qed + + +theorem minus_injective[dest]: + assumes "-.n = -.m" and "n \ Int \ m \ Int" + shows "n = m" + proof - + have "<1>1": "-.-.n = -.-.m" + using assms by simp + have "<1>2": "-.-.n = n" + proof - + have "<2>1": "n \ Int ==> -.-.n = n" + using minus_involutive by auto + have "<2>2": "n \ Int" + using assms by auto + show "-.-.n = n" + using "<2>1" "<2>2" by auto + qed + have "<1>3": "-.-.m = m" + proof - + have "<2>1": "m \ Int ==> -.-.m = m" + using minus_involutive by auto + have "<2>2": "m \ Int" + using assms by auto + show "-.-.m = m" + using "<2>1" "<2>2" by auto + qed + show "n = m" + using "<1>1" "<1>2" "<1>3" by auto + qed + + +theorem minus_eq: + assumes "x = y" + shows "-.x = -.y" + using assms by auto + + +theorem neg_succ_not_in_nat[simp]: + "n \ Nat ==> + -.(Succ[n]) \ Nat" + proof - + { + fix "n" + assume "<2>1": "n \ Nat" + have "<2>2": "Succ[n] \ (Nat \ {0})" + using "<2>1" by auto + have "-.(Succ[n]) \ Nat" + using "<2>2" neg_not_in_nat by auto + } + from this show "n \ Nat ==> + -.(Succ[n]) \ Nat" + by auto + qed + + +theorem neg0_iff_eq0[simp]: + "(-.n = 0) = (n = 0)" + proof auto + assume "<1>1": "-.n = 0" + have "<1>2": "-.n \ Nat" + using "<1>1" by auto + have "<1>3": "-.-.n = n" + using "<1>2" neg_neg by auto + have "<1>4": "-.-.n = 0" + using "<1>1" by auto + show "n = 0" + using "<1>3" "<1>4" by simp + qed + + +theorem neg0_imp_eq0[dest]: + "(-.n = 0) ==> (n = 0)" + by simp + + +theorem not_neg0_imp_not0[dest]: + "(-.n \ 0) ==> (n \ 0)" + by simp + + +(*----------------------------------------------------------------------------*) +(* Properties of the set `Int`. *) + + +theorem nat_in_int[simp]: + "n \ Nat ==> n \ Int" + by (simp add: Int_def) + + +theorem minus_nat_in_int: + "n \ Nat ==> -.n \ Int" + by (auto simp: Int_def) + + +theorem minus_nat_in_neg_nat: + "n \ Nat ==> -.n \ negNat" + by (auto simp: negNat_def) + + +theorem neg_nat_in_int: + "n \ negNat ==> n \ Int" + by (auto simp: negNat_def Int_def) + + +theorem int_disj: + "n \ Int ==> + n \ Nat + \ n \ {-.n: n \ Nat}" + by (auto simp: Int_def) + + +theorem int_disj_nat_negnat: + "n \ Int ==> + n \ Nat + \ n \ negNat" + unfolding negNat_def + using int_disj by auto + + +theorem int_union_nat_negnat: + "Int = Nat \ negNat" + using int_disj_nat_negnat nat_in_int + neg_nat_in_int + by auto + + +theorem int_union_nat_0_negnat: + "Int = (Nat \ {0}) \ negNat" + proof - + have s1_1: "Int = Nat \ negNat" + using int_union_nat_negnat by auto + have s1_2: "0 \ negNat" + unfolding negNat_def + using zeroIsNat neg0 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem int_union_nat_negnat_0: + "Int = Nat \ (negNat \ {0})" + proof - + have s1_1: "Int = Nat \ negNat" + using int_union_nat_negnat by auto + have s1_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem neg_int_eq_int[simp]: + "n \ Int ==> -.n \ Int" + unfolding Int_def by auto + + +theorem minus_negnat_in_int: + "n \ negNat ==> -.n \ Int" + using neg_nat_in_int neg_int_eq_int by auto + + +theorem minus_neg_int_in_nat: + "n \ Int \ n \ Nat ==> + -.n \ Nat \ {0}" + proof - + assume "<1>1": "n \ Int \ n \ Nat" + have "<1>2": "n \ {-.k: k \ Nat}" + using "<1>1" int_disj by iprover + have "<1>3": "\ k \ Nat: n = -.k" + using "<1>2" by auto + have "<1>4": "\ k \ Nat: k = -.n" + using "<1>3" minus_involutive by auto + have "<1>5": "-.n \ Nat" + using "<1>4" by auto + have "<1>6": "-.n \ 0" + proof - + { + assume "<2>1": "-.n = 0" + have "<2>2": "n = 0" + using neg0 "<2>1" by auto + have "<2>3": "n \ Nat" + using "<2>2" by auto + have "<2>4": "n \ Nat" + using "<1>1" by auto + have "<2>5": "FALSE" + using "<2>3" "<2>4" by auto + } + from this show "-.n \ 0" + by auto + qed + show "-.n \ Nat \ {0}" + using "<1>5" "<1>6" by auto + qed + + +theorem neg_negative_in_nat: + assumes hyp: "n \ Int \ n \ Nat" + shows "-. n \ Nat" + proof - + have int: "n \ Int" + using hyp by auto + from int int_disj have disj: + "n \ Nat \ n \ {-. m: m \ Nat}" + by auto + from disj hyp have neg: "n \ {-. m: m \ Nat}" + by auto + from neg have ex1: "\ m \ Nat: n = -. m" + by auto + from ex1 have ex2: "\ m \ Nat: -. n = -. -. m" + by auto + from ex2 neg_neg have ex3: "\ m \ Nat: -. n = m" + by auto + from ex3 show "-. n \ Nat" + by auto + qed + + +theorem minus_nat_0_or_not_in_nat: + "n \ Nat ==> + -.n = 0 \ -.n \ Nat" + using neg0 neg_not_in_nat by auto + + +theorem int_eq_neg_int_is_0[simp]: + "n \ Int ==> (n = -.n) = (n = 0)" + proof - + have "<1>1": "n \ Nat \ {0} ==> + -.n \ Nat" + proof - + assume "<2>1": "n \ Nat \ {0}" + show "-.n \ Nat" + using neg_not_in_nat "<2>1" by auto + qed + have "<1>2": "n \ Nat \ {0} ==> + n \ -.n" + using "<1>1" by force + have "<1>3": "n \ Int \ n \ Nat + ==> -.n \ Nat" + using minus_neg_int_in_nat by auto + have "<1>4": "n \ Int \ n \ Nat + ==> n \ -.n" + using "<1>3" by force + have "<1>5": "n \ Int \ {0} ==> + (n \ Nat \ {0}) \ + (n \ {-.k: k \ Nat})" + using int_disj by auto + have "<1>6": "n \ Int \ {0} ==> + (n \ -.n)" + using "<1>2" "<1>4" "<1>5" by auto + have "<1>7": "n \ Int \ (n \ 0) ==> + (n \ -.n)" + using "<1>6" by auto + have "<1>8": "n \ Int \ (n = -.n) ==> + (n = 0)" + using "<1>7" by auto + have "<1>9": "n \ Int \ (n = 0) ==> + (n = -.n)" + using neg0 by auto + show "n \ Int ==> (n = -.n) = (n = 0)" + using "<1>8" "<1>9" by auto + qed + + +theorem minus_neg_nat_in_nat: + assumes hyp: "n \ negNat" + shows "-.n \ Nat" + proof - + have s1_1: "\ k \ Nat: -.k = n" + using hyp by (auto simp: negNat_def) + have s1_2: "\ k \ Nat: -.-.k = -.n" + using s1_1 by auto + have s1_3: "\ k \ Nat: k = -.n" + using s1_2 neg_neg by auto + show "-.n \ Nat" + using s1_3 by auto + qed + + +theorem minus_neg_nat_0_in_nat_0: + assumes hyp: "n \ negNat \ {0}" + shows "-.n \ Nat \ {0}" + proof - + have s1_1: "-.n \ Nat" + using hyp minus_neg_nat_in_nat by auto + have s1_2: "-.n \ 0" + proof - + { + assume s2_1: "-.n = 0" + have s2_2: "-.-.n = -.0" + using s2_1 by auto + have s2_3: "-.-.n = 0" + using s2_2 neg0 by auto + have s2_4: "n = 0" + proof - + have s3_1: "n \ Int" + using hyp neg_nat_in_int by auto + have s3_2: "-.-.n = n" + using s3_1 minus_involutive + by auto + show ?thesis + using s3_2 s2_3 by auto + qed + have s2_5: "n \ 0" + using hyp by auto + have "FALSE" + using s2_4 s2_5 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem neg_nat_0_not_in_nat: + assumes hyp: "n \ negNat \ {0}" + shows "n \ Nat" + proof auto + assume s1_1: "n \ Nat" + have s1_2: "n \ 0" + using hyp by auto + have s1_3: "n \ Nat \ {0}" + using s1_1 s1_2 by auto + have s1_4: "-.n \ Nat" + using s1_3 neg_not_in_nat by auto + have s1_5: "-.n \ Nat" + using hyp minus_neg_nat_in_nat by auto + show "FALSE" + using s1_4 s1_5 by auto + qed + + +theorem neg_nat_not_in_nat_setminus_0: + assumes hyp: "n \ negNat" + shows "n \ Nat \ {0}" + proof - + have s1_1: "n \ negNat ==> + n \ negNat \ {0} + \ n = 0" + by auto + have s1_2: "n \ negNat \ {0} + ==> n \ Nat \ {0}" + using neg_nat_0_not_in_nat by auto + have s1_3: "n = 0 ==> + n \ Nat \ {0}" + by auto + show "n \ Nat \ {0}" + using hyp s1_1 s1_2 s1_3 by iprover + qed + + +theorem minus_nat_0_in_negnat_0: + assumes hyp: "n \ Nat \ {0}" + shows "-.n \ negNat \ {0}" + proof - + have s1_1: "n \ Nat" + using hyp by auto + have s1_2: "-.n \ negNat" + using s1_1 minus_nat_in_neg_nat by auto + have s1_3: "-.n \ 0" + proof - + { + assume s2_1: "-.n = 0" + have s2_2: "-.-.n = -.0" + using s2_1 by auto + have s2_3: "-.-.n = 0" + using s2_2 by auto + have s2_4: "n = 0" + using s1_1 neg_neg s2_3 by auto + have s2_5: "n \ 0" + using hyp by auto + have "FALSE" + using s2_4 s2_5 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_2 s1_3 by auto + qed + + +theorem minus_one_in_negnat: + "-.1 \\in negNat" + unfolding negNat_def + using oneIsNat by auto + + +theorem minus_one_in_int: + "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + + +theorem minus_one_neq_0: + "-.1 \ 0" + proof - + { + assume s1_1: "-.1 = 0" + have s1_2: "-.-.1 = -.0" + using s1_1 by auto + have s1_3: "-.-.1 = 0" + using s1_2 neg0 by auto + have s1_4: "0 \ Nat" + by auto + have s1_5: "1 = 0" + using s1_3 s1_4 by auto + have "FALSE" + using s1_5 succIrrefl + by auto + } + from this show "-.1 \ 0" + by auto + qed + + +theorem minus_minus_one: + "-.-.1 = 1" + proof - + have s1_1: "1 \ Nat" + using oneIsNat by auto + show "-.-.1 = 1" + using neg_neg s1_1 by auto + qed + + +theorem succ_zero_equality: + assumes hyp: "x = Succ[0]" + shows "x = 1" + using hyp by auto + + +theorem pred_1: + "Pred[1] = 0" + proof - + have s1_1: "1 = Succ[0]" + by auto + have s1_2: "\ x \ Nat: (1 = Succ[x]) => (x = 0)" + using s1_1 succInjIff by auto + show ?thesis + using s1_1 s1_2 choose_equality by auto + qed + + +theorem pred_2: + "Pred[2] = 1" + proof - + have s1_1: "1 \\in Nat \ Succ[1] = 2" + unfolding two_def + using oneIsNat by auto + have s1_2: "\ x. x \\in Nat \ (Succ[x] = 2) ==> (x = 1)" + using oneIsNat succInjIff by (auto simp: two_def) + have s1_4: "Pred[2] = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" + proof - + have s2_1: "Pred[2] = + (IF 2 = 0 THEN 0 ELSE CHOOSE x \ Nat: 2 = Succ[x])" + unfolding Pred_def + using twoIsNat by auto + also have s2_2: "... = (CHOOSE x \ Nat: 2 = Succ[x])" + proof - + have s3_1: "2 \ 0" + using succNotZero by (auto simp: two_def) + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = (CHOOSE x: x \\in Nat \ 2 = Succ[x])" + unfolding bChoose_def + by auto + also have s2_4: "... = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" + proof - + have s3_1: "\ x. (2 = Succ[x]) = (Succ[x] = 2)" + by auto + show ?thesis + using s3_1 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s1_1 s1_2 s1_4 choose_equality[ + of "\ x. x \\in Nat \ Succ[x] = 2" + "1" "1"] + by auto + qed + + +theorem zero_in_int: + "0 \ Int" + using zeroIsNat nat_in_int by auto + + +theorem ipred_0: + shows "iPred[0] = -.1" + proof - + have s1_1: "iPred[0] = -.Succ[-.0]" + proof - + have s2_1: "0 \ Int" + using zero_in_int by auto + have s2_2: "0 \ Nat \ {0}" + by auto + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_2: "... = -.Succ[0]" + using neg0 by auto + also have s1_3: "... = -.1" + by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed + + +theorem isucc_minus_1: + "iSucc[-.1] = 0" + unfolding iSucc_def + using oneIsNat neg_neg pred_1 neg0 by auto + + +theorem negnat_succ_in_nat: + assumes ndom: "n \\in negNat" + shows "Succ[-.n] \\in Nat" + using ndom minus_neg_nat_in_nat succIsNat by auto + + +theorem minus_negnat_succ_in_negnat: + assumes ndom: "n \\in negNat" + shows "-.Succ[-.n] \\in negNat" + using ndom negnat_succ_in_nat minus_nat_in_neg_nat by auto + + +theorem minus_succ_minus_negnat_in_int: + assumes ndom: "n \\in negNat" + shows "-.Succ[-.n] \\in Int" + proof - + have s1_1: "-.n \\in Nat" + using ndom minus_neg_nat_in_nat by auto + have s1_2: "Succ[-.n] \\in Nat" + using s1_1 succIsNat by auto + show ?thesis + using s1_2 minus_nat_in_int by auto + qed + +(*----------------------------------------------------------------------------*) +(* `iSucc` and `iPred` properties. *) + +(* Typeness *) +theorem iSucc_type: + assumes idom: "i \ Int" + shows "iSucc[i] \ Int" + proof - + have s1_1: "i \ Nat ==> + iSucc[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s2_1 by auto + have s2_3: "Succ[i] \ Nat" + using s2_1 succIsNat by auto + have s2_4: "Succ[i] \ Int" + using s2_3 nat_in_int by auto + show "iSucc[i] \ Int" + using s2_2 s2_4 by auto + qed + have s1_2: "i \ Nat ==> + iSucc[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using s2_1 idom by auto + have s2_3: "Pred[-.i] \ Nat" + proof - + have s3_1: "-.i \ Nat" + using idom s2_1 minus_neg_int_in_nat + by auto + show "Pred[-.i] \ Nat" + using s3_1 Pred_in_nat by auto + qed + have s2_4: "-.Pred[-.i] \ Int" + using s2_3 minus_nat_in_int by auto + show "iSucc[i] \ Int" + using s2_2 s2_4 by auto + qed + show "iSucc[i] \ Int" + using s1_1 s1_2 by auto + qed + + +theorem iPred_type: + assumes idom: "i \ Int" + shows "iPred[i] \ Int" + proof - + have s1_1: "i \ Nat \ {0} + ==> iPred[i] \ Int" + proof - + assume s2_1: "i \ Nat \ {0}" + have s2_2: "iPred[i] = Pred[i]" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "Pred[i] \ Nat" + using s2_1 Pred_in_nat by auto + have s2_4: "Pred[i] \ Int" + using s2_3 nat_in_int by auto + show "iPred[i] \ Int" + using s2_2 s2_4 by auto + qed + have s1_2: "i = 0 ==> + iPred[i] \ Int" + proof - + assume s2_1: "i = 0" + have s2_2: "iPred[i] = -.1" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "-.1 \ Int" + using oneIsNat by (auto simp: Int_def) + show "iPred[i] \ Int" + using s2_2 s2_3 by auto + qed + have s1_3: "i \ Nat ==> + iPred[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iPred[i] = -.Succ[-.i]" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "-.Succ[-.i] \ Int" + proof - + have s3_1: "-.i \ Nat" + using idom s2_1 minus_neg_int_in_nat + by auto + have s3_2: "Succ[-.i] \ Nat" + using s3_1 succIsNat by auto + show "-.Succ[-.i] \ Int" + using s3_2 minus_nat_in_int by auto + qed + show "iPred[i] \ Int" + using s2_2 s2_3 by auto + qed + show "iPred[i] \ Int" + using s1_1 s1_2 s1_3 by auto + qed + + +(* +THEOREM iSucciPredCommute == + \A i \in Int: iSucc[iPred[i]] = iPred[iSucc[i]] + PROOF + <1>1. SUFFICES ASSUME NEW i \in Int + PROVE iSucc[iPred[i]] = iPred[iSucc[i]] + BY <1>1 + <1>2. CASE i \in Nat \ {0} + <2>1. iSucc[iPred[i]] = i + <3>1. iPred[i] = Pred[i] + <4>1. i \in Nat \ {0} + BY <1>2 + <4> QED + BY <4>1 DEF iPred + <3>2. iSucc[iPred[i]] = Succ[Pred[i]] + <4>1. Pred[i] \in Nat + <5>1. i \in Nat + BY <1>2 + <5> QED + BY <5>1, Pred_in_nat + <4> QED + BY <3>1, <4>1 DEF iSucc + <3>3. Succ[Pred[i]] = i + <4>1. i \in Nat \ {0} + BY <1>2 + <4> QED + BY <4>1, Succ_Pred_nat + <3> QED + BY <3>2, <3>3 + <2>2. iPred[iSucc[i]] = i + <3>1. iSucc[i] = Succ[i] + <4>1. i \in Nat + BY <1>2 + <4> QED + BY <4>1 DEF iSucc + <3>2. iPred[iSucc[i]] = Pred[Succ[i]] + <4>1. Succ[i] \in Nat + <5>1. i \in Nat + BY <1>2 + <5> QED + BY <5>1, succIsNat + <4>2. Succ[i] # 0 + BY <1>2, succNotZero + <4>3. Succ[i] \in Nat \ {0} + BY <4>1, <4>2 + <4>4. iPred[Succ[i]] = Pred[Succ[i]] + BY <4>3 DEF iPred + <4> QED + BY <4>4, <3>1 + <3>3. Pred[Succ[i]] = i + <4>1. i \in Nat + BY <1>2 + <4> QED + BY <4>1, Pred_Succ_nat + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2 + <1>3. CASE i = 0 + <2>1. iSucc[iPred[i]] = 0 + <3>1. iSucc[iPred[i]] = iSucc[-.1] + <4>1. iPred[i] = -.1 + BY <1>3 DEF iPred + <4> QED + BY <4>1 + <3>2. iSucc[-.1] = -.Pred[-.-.1] + <4>1. -.1 \in Int + <5>1. -.1 \in negNat + BY oneIsNat DEF negNat + <5> QED + BY <5>1, neg_nat_in_int + <4>2. -.1 \notin Nat + <5>1. -.1 \in negNat \ {0} + <6>1. -.1 # 0 + BY minus_one_neq_0 + <6>2. -.1 \in negNat + BY oneIsNat DEF negNat + <6> QED + BY <6>1, <6>2 + <5> QED + BY <5>1, neg_nat_0_not_in_nat + <4> QED + BY <4>1, <4>2 DEF iSucc + <3>3. -.Pred[-.-.1] = 0 + <4>1. -.-.1 = 1 + BY minus_minus_one + <4>2. Pred[1] = 0 + BY pred_1 + <4>3. -.0 = 0 + BY neg0 + <4> QED + BY <4>1, <4>2, <4>3 + <3> QED + BY <3>1, <3>2, <3>3 + <2>2. iPred[iSucc[i]] = 0 + <3>1. iSucc[i] = 1 + <4>1. i \in Nat + BY <1>3, zeroIsNat + <4>2. iSucc[i] = Succ[i] + BY <4>1 DEF iSucc + <4>3. Succ[i] = Succ[0] + BY <1>3 + <4>4. Succ[0] = 1 + OBVIOUS + <4> QED + BY <4>2, <4>3, <4>4 + <3>2. iPred[1] = 0 + <4>1. 1 \in Nat \ {0} + BY oneIsNat, succIrrefl + <4>2. iPred[1] = Pred[1] + BY <4>1 DEF iPred + <4>3. Pred[1] = 0 + BY pred_1 + <4> QED + BY <4>2, <4>3 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 + <1>4. CASE i \in negNat + <2>1. iSucc[iPred[i]] = i + <3>1. iPred[i] = -.Succ[-.i] + <4>1. i \in Int + BY <1>1 + <4>2. i \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3>2. -.Succ[-.i] \in negNat \ {0} + <4>1. Succ[-.i] \in Nat \ {0} + <5>1. -.i \ Nat + BY <1>4, minus_neg_nat_in_nat + <5>2. Succ[-.i] \in Nat + BY <5>1, succIsNat + <5>3. Succ[-.i] \ 0 + BY <5>1, succNotZero + <5> QED + BY <5>2, <5>3 + <4> QED + BY <4>1, minus_nat_0_in_negnat_0 + <3>3. iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]] + <4>1. iPred[i] \in negNat \ {0} + BY <3>1, <3>2 + <4>2. iPred[i] \in Int + BY <4>1, neg_nat_in_int + <4>3. iPred[i] \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4>4. iSucc[iPred[i]] = -.Pred[-.iPred[i]] + BY <4>2, <4>3 DEF iSucc + <4> QED + BY <4>4, <3>1 + <3>4. @ = -.Pred[Succ[-.i]] + <4>1. -.i \in Nat + BY <1>4, minus_neg_nat_in_nat + <4>2. Succ[-.i] \in Int + BY <4>1, succIsNat, nat_in_int + <4>3. -.-.Succ[-.i] = Succ[-.i] + BY minus_involutive + <4> QED + BY <4>3 + <3>5. @ = -.-.i + <4>1. -.i \in Nat + BY <1>4, minus_neg_nat_in_nat + <4>2. Pred[Succ[-.i]] = -.i + BY <4>1, Pred_Succ_nat + <4> QED + BY <4>2 + <3>6. @ = i + BY <1>4, minus_involutive + <3> QED + BY <3>3, <3>4, <3>5, <3>6 + <2>2. iPred[iSucc[i]] = i + <3>1. CASE i = 0 + <4>1. iSucc[i] = 1 + <5>1. 0 \in Nat + BY zeroIsNat + <5>2. 0 \in Int + BY <5>1, nat_in_int + <5>3. iSucc[i] = Succ[0] + BY <5>1, <5>2, <3>1 DEF iSucc + <5>4. Succ[0] = 1 + OBVIOUS + <5> QED + BY <5>3, <5>4 + <4>2. iPred[1] = 0 + <5>1. 1 \in Nat \ {0} + BY oneIsNat, succNotZero + <5>2. 1 \in Int + BY <5>1, nat_in_int + <5>3. iPred[1] = Pred[1] + BY <5>1, <5>2 DEF iPred + <5>4. Pred[1] = 0 + BY pred_1 + <5> QED + BY <5>3, <5>4 + <4> QED + BY <4>1, <4>2, <3>1 + <3>2. CASE i # 0 + <4>1. i \in negNat \ {0} + BY <1>4, <3>1 + <4>2. i \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4>3. i \in Int + BY <4>1, neg_nat_in_int + <4>4. iSucc[i] = -.Pred[-.i] + BY <4>2, <4>3 DEF iSucc + <4>5. -.Pred[-.i] \in negNat + <5>1. -.i \in Nat + BY <4>1, minus_neg_nat_in_nat + <5>2. Pred[-.i] \in Nat + BY <5>1, Pred_in_nat + <5> QED + BY <5>2, minus_nat_in_neg_nat + <4>6. -.Pred[-.i] \in Int + BY <4>5, neg_nat_in_int + <4>7. -.Pred[-.i] \notin Nat \ {0} + BY <4>5, neg_nat_not_in_nat_setminus_0 + <4>8. iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]] + BY <4>6, <4>7 DEF iPred + <4>9. -.Succ[-.-.Pred[-.i]] = i + <5>1. -.-.Pred[-.i] = Pred[-.i] + <6>1. -.i \in Nat + BY <4>1, minus_neg_nat_in_nat + <6>2. Pred[-.i] \in Nat + BY <6>1, Pred_in_nat + <6>3. Pred[-.i] \in Int + BY <6>2, nat_in_int + <6> QED + BY <6>3, minus_involutive + <5>2. Succ[Pred[-.i]] = -.i + <6>1. -.i \in Nat \ {0} + BY <4>1, minus_neg_nat_0_in_nat_0 + <6> QED + BY <6>1, Succ_Pred_nat + <5>3. -.-.i = i + BY <4>3, minus_involutive + <5> QED + BY <5>1, <5>2, <5>3 + <4> QED + BY <4>4, <4>8, <4>9 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 + <1> QED + BY <1>1, <1>2, <1>3, <1>4 +*) +theorem iSucciPredCommute: + "\ i \ Int: + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + { + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "i \ Nat \ {0} ==> + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + assume idom: "i \ Nat \ {0}" + have s2_1: "iSucc[iPred[i]] = i" + proof - + have s3_1: "iPred[i] = Pred[i]" + unfolding iPred_def + using s1_1 idom by auto + have s3_2: "iSucc[iPred[i]] = Succ[Pred[i]]" + proof - + have s4_1: "Pred[i] \ Nat" + proof - + have s5_1: "i \ Nat" + using idom by auto + show ?thesis + using s5_1 Pred_in_nat by auto + qed + show ?thesis + unfolding iSucc_def + using s4_1 s3_1 by auto + qed + have s3_3: "Succ[Pred[i]] = i" + using idom Succ_Pred_nat + by auto + show ?thesis + using s3_2 s3_3 by auto + qed + have s2_2: "iPred[iSucc[i]] = i" + proof - + have s3_1: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s1_1 idom by auto + have s3_2: "iPred[iSucc[i]] = Pred[Succ[i]]" + proof - + have s4_1: "Succ[i] \ Nat" + proof - + have s5_1: "i \ Nat" + using idom by auto + show ?thesis + using s5_1 succIsNat by auto + qed + have s4_2: "Succ[i] \ 0" + using idom succNotZero by auto + have s4_3: "Succ[i] \ Nat \ {0}" + using s4_1 s4_2 by auto + have s4_4: "Succ[i] \ Int" + using s4_1 nat_in_int by auto + have s4_5: "iPred[Succ[i]] = Pred[Succ[i]]" + unfolding iPred_def + using s4_3 s4_4 by auto + show ?thesis + using s4_5 s3_1 by auto + qed + have s3_3: "Pred[Succ[i]] = i" + proof - + have s4_1: "i \ Nat" + using idom by auto + show ?thesis + using s4_1 Pred_Succ_nat + by auto + qed + show ?thesis + using s3_2 s3_3 by auto + qed + show ?thesis + using s2_1 s2_2 + by auto + qed + have s1_3: "i \ negNat ==> + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + assume idom: "i \ negNat" + have s2_1: "iSucc[iPred[i]] = i" + proof - + have s3_1: "iPred[i] = -.Succ[-.i]" + proof - + have s4_1: "i \ Int" + using s1_1 by auto + have s4_2: "i \ Nat \ {0}" + using idom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + have s3_2: "-.Succ[-.i] \ negNat \ {0}" + proof - + have s4_1: "Succ[-.i] \ Nat \ {0}" + proof - + have s5_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat + by auto + have s5_2: "Succ[-.i] \ Nat" + using s5_1 succIsNat by auto + have s5_3: "Succ[-.i] \ 0" + using s5_1 succNotZero by auto + show ?thesis + using s5_2 s5_3 by auto + qed + show ?thesis + using s4_1 minus_nat_0_in_negnat_0 + by auto + qed + have s3_3: "iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]]" + proof - + have s4_1: "iPred[i] \ negNat \ {0}" + using s3_1 s3_2 by auto + have s4_2: "iPred[i] \ Int" + using s4_1 neg_nat_in_int by auto + have s4_3: "iPred[i] \ Nat" + using s4_1 neg_nat_0_not_in_nat by auto + have s4_4: "iSucc[iPred[i]] = -.Pred[-.iPred[i]]" + unfolding iSucc_def + using s4_2 s4_3 by auto + show ?thesis + using s4_4 s3_1 by auto + qed + have s3_4: "-.Pred[-.-.Succ[-.i]] = -.Pred[Succ[-.i]]" + proof - + have s4_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat by auto + have s4_2: "Succ[-.i] \ Int" + using s4_1 succIsNat nat_in_int by auto + have s4_3: "-.-.Succ[-.i] = Succ[-.i]" + using s4_2 minus_involutive by auto + show ?thesis + using s4_3 by auto + qed + have s3_5: "-.Pred[Succ[-.i]] = -.-.i" + proof - + have s4_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat by auto + have s4_2: "Pred[Succ[-.i]] = -.i" + using s4_1 Pred_Succ_nat by auto + show ?thesis + using s4_2 by auto + qed + have s3_6: "-.-.i = i" + using s1_1 minus_involutive by auto + show ?thesis + using s3_3 s3_4 s3_5 s3_6 by auto + qed + have s2_2: "iPred[iSucc[i]] = i" + proof - + have s3_1: "i = 0 ==> + iPred[iSucc[i]] = i" + proof - + assume i0: "i = 0" + have s4_1: "iSucc[i] = 1" + proof - + have s5_1: "0 \ Nat" + using zeroIsNat by auto + have s5_2: "0 \ Int" + using s5_1 nat_in_int by auto + have s5_3: "iSucc[i] = Succ[0]" + unfolding iSucc_def + using s5_1 s5_2 i0 + by auto + have s5_4: "Succ[0] = 1" + by auto + show ?thesis + using s5_3 s5_4 by auto + qed + have s4_2: "iPred[1] = 0" + proof - + have s5_1: "1 \ Nat \ {0}" + using oneIsNat succNotZero + by auto + have s5_2: "1 \ Int" + using s5_1 nat_in_int by auto + have s5_3: "iPred[1] = Pred[1]" + unfolding iPred_def + using s5_1 s5_2 + by auto + have s5_4: "Pred[1] = 0" + using pred_1 by auto + show ?thesis + using s5_3 s5_4 by auto + qed + show "iPred[iSucc[i]] = i" + using s4_1 s4_2 i0 by auto + qed + have s3_2: "i \ 0 ==> + iPred[iSucc[i]] = i" + proof - + assume inot0: "i \ 0" + have s4_1: "i \ negNat \ {0}" + using idom inot0 by auto + have s4_2: "i \ Nat" + using s4_1 neg_nat_0_not_in_nat + by auto + have s4_3: "i \ Int" + using s4_1 neg_nat_in_int + by auto + have s4_4: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using s4_2 s4_3 by auto + have s4_5: "-.Pred[-.i] \ negNat" + proof - + have s5_1: "-.i \ Nat" + using s4_1 minus_neg_nat_in_nat + by auto + have s5_2: "Pred[-.i] \ Nat" + using s5_1 Pred_in_nat + by auto + show ?thesis + using s5_2 minus_nat_in_neg_nat + by auto + qed + have s4_6: "-.Pred[-.i] \ Int" + using s4_5 neg_nat_in_int + by auto + have s4_7: "-.Pred[-.i] \ Nat \ {0}" + using s4_5 neg_nat_not_in_nat_setminus_0 + by blast + have s4_8: "iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]]" + unfolding iPred_def + using s4_6 s4_7 by auto + have s4_9: "-.Succ[-.-.Pred[-.i]] = i" + proof - + have s5_1: "-.-.Pred[-.i] = Pred[-.i]" + proof - + have s6_1: "-.i \ Nat" + using s4_1 minus_neg_nat_in_nat + by auto + have s6_2: "Pred[-.i] \ Nat" + using s6_1 Pred_in_nat + by auto + have s6_3: "Pred[-.i] \ Int" + using s6_2 nat_in_int + by auto + show ?thesis + using s6_3 minus_involutive + by auto + qed + have s5_2: "Succ[Pred[-.i]] = -.i" + proof - + have s6_1: "-.i \ Nat \ {0}" + using s4_1 minus_neg_nat_0_in_nat_0 + by auto + show ?thesis + using s6_1 Succ_Pred_nat + by auto + qed + have s5_3: "-.-.i = i" + using s4_3 minus_involutive by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + show "iPred[iSucc[i]] = i" + using s4_4 s4_8 s4_9 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + show "iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i" + using s2_1 s2_2 by auto + qed + have s1_4: "Int = (Nat \ {0}) \ negNat" + using int_union_nat_0_negnat by auto + have "iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i" + using s1_1 s1_2 s1_3 s1_4 by auto + } + from this show ?thesis + using allI by auto + qed + + +(* +THEOREM iSuccMinusFlip == + \A i \in Int: -.iSucc[-.i] = iPred[i] +PROOF +<1>1. SUFFICES + ASSUME NEW i \in Int + PROVE -.iSucc[-.i] = iPred[i] + BY <1>1 +<1>2. CASE i \in Nat \ {0} + <2>1. i \in Int + BY <1>2, nat_in_int + <2>2. -.iSucc[-.i] = Pred[i] + <3>1. -.iSucc[-.i] = -.-.Pred[-.-.i] + <4>1. -.i \in negNat \ {0} + BY <1>2, minus_nat_0_in_negnat_0 + <4>2. -.i \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4> QED + BY <2>1, <4>2 DEF iSucc + <3>2. @ = -.-.Pred[i] + <4>1. -.-.i = i + BY <2>1, minus_involutive + <4> QED + BY <4>1 + <3>3. @ = Pred[i] + <4>1. Pred[i] \in Nat + BY <1>2, Pred_in_nat + <4>2. Pred[i] \in Int + BY <4>1, nat_in_int + <4> QED + BY <4>2, minus_involutive + <3> QED + BY <3>1, <3>2, <3>3 + <2>3. iPred[i] = Pred[i] + BY <2>1, <1>2 DEF iPred + <2> QED + BY <2>2, <2>3 +<1>3. CASE i \in negNat + <2>1. i \in Int + BY <1>3, neg_nat_in_int + <2>2. i \notin Nat \ {0} + BY <1>3, neg_nat_not_in_nat_setminus_0 + <2>3. Pred[i] = -.Succ[-.i] + BY <2>1, <2>2 DEF iPred + <2>4. -.iSucc[-.i] = -.Succ[-.i] + <3>1. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3>2. -.i \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>1, <3>2 DEF iSucc + <2> QED + BY <2>3, <2>4 +<1> QED + BY <1>1, <1>2, <1>3, int_union_nat_0_negnat +*) +theorem iSuccMinusFlip: + "\ i \ Int: + -.iSucc[-.i] = iPred[i]" + proof auto + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "i \ Nat \ {0} ==> + -.iSucc[-.i] = iPred[i]" + proof - + assume s1_2_asm: "i \ Nat \ {0}" + have s2_1: "i \ Int" + using s1_2_asm nat_in_int + by auto + have s2_2: "-.iSucc[-.i] = Pred[i]" + proof - + have s3_1: "-.iSucc[-.i] = -.-.Pred[-.-.i]" + proof - + have s4_1: "-.i \ negNat \ {0}" + using s1_2_asm minus_nat_0_in_negnat_0 + by auto + have s4_2: "-.i \ Nat" + using s4_1 neg_nat_0_not_in_nat + by auto + show ?thesis + unfolding iSucc_def + using s2_1 s4_2 by auto + qed + have s3_2: "-.-.Pred[-.-.i] = -.-.Pred[i]" + proof - + have s4_1: "-.-.i = i" + using s2_1 minus_involutive + by auto + show ?thesis + using s4_1 by auto + qed + have s3_3: "-.-.Pred[i] = Pred[i]" + proof - + have s4_1: "Pred[i] \ Nat" + using s1_2_asm Pred_in_nat by auto + have s4_2: "Pred[i] \ Int" + using s4_1 nat_in_int by auto + show ?thesis + using s4_2 minus_involutive by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 by auto + qed + have s2_3: "iPred[i] = Pred[i]" + unfolding iPred_def + using s2_1 s1_2_asm by auto + show "-.iSucc[-.i] = iPred[i]" + using s2_2 s2_3 by auto + qed + have s1_3: "i \ negNat ==> + -.iSucc[-.i] = iPred[i]" + proof - + assume s1_3_asm: "i \ negNat" + have s2_1: "i \ Int" + using s1_3_asm neg_nat_in_int + by auto + have s2_2: "i \ Nat \ {0}" + using s1_3_asm neg_nat_not_in_nat_setminus_0 + by auto + have s2_3: "-.Succ[-.i] = iPred[i]" + unfolding iPred_def + using s2_1 s2_2 by auto + have s2_4: "-.iSucc[-.i] = -.Succ[-.i]" + proof - + have s3_1: "-.i \ Nat" + using s1_3_asm minus_neg_nat_in_nat by auto + have s3_2: "-.i \ Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + show "-.iSucc[-.i] = iPred[i]" + using s2_3 s2_4 by auto + qed + show "-.iSucc[-.i] = iPred[i]" + using s1_1 s1_2 s1_3 int_union_nat_0_negnat + by auto + qed + + +(* +THEOREM iPredMinusFlip == + \A i \in Int: -.iPred[-.i] = iSucc[i] +PROOF +<1>1. SUFFICES + ASSUME NEW i \in Int + PROVE -.iPred[-.i] = iSucc[i] + BY <1>1 +<1>2. -.iSucc[-.-.i] = iPred[-.i] + <2>1. -.i \in Int + BY <1>1, neg_int_eq_int + <2> QED + BY <2>1, iSuccMinusFlip +<1>3. -.iSucc[i] = iPred[-.i] + <2>1. -.-.i = i + BY <1>1, minus_involutive + <2> QED + BY <1>2, <2>1 +<1>4. -.-.iSucc[i] = -.iPred[-.i] + BY <1>3 +<1>5. -.-.iSucc[i] = iSucc[i] + <2>1. iSucc[i] \in Int + BY <1>1, iSucc_type + <2> QED + BY <2>1, minus_involutive +<1> QED + BY <1>4, <1>5 +*) +theorem iPredMinusFlip: + "\ i \ Int: -.iPred[-.i] = iSucc[i]" + proof auto + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "-.iSucc[-.-.i] = iPred[-.i]" + proof - + have s2_1: "-.i \ Int" + using s1_1 neg_int_eq_int by auto + show ?thesis + using s2_1 iSuccMinusFlip by auto + qed + have s1_3: "-.iSucc[i] = iPred[-.i]" + proof - + have s2_1: "-.-.i = i" + using s1_1 minus_involutive by auto + show ?thesis + using s1_2 s2_1 by auto + qed + have s1_4: "-.-.iSucc[i] = -.iPred[-.i]" + using s1_3 by auto + have s1_5: "-.-.iSucc[i] = iSucc[i]" + proof - + have s2_1: "iSucc[i] \ Int" + using s1_1 iSucc_type by auto + show ?thesis + using s2_1 minus_involutive by auto + qed + show "-.iPred[-.i] = iSucc[i]" + using s1_4 s1_5 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* `neg_nat_induction` *) + + +(* +negNat == {-.n: n \in Nat} + + +THEOREM negNatInduction == + ASSUME + NEW P(n), P(0), + \A i \in negNat: P(i) => P(-.Succ[-.i]) + PROVE + \A i \in negNat: P(i) + PROOF + <1> DEFINE Q(n) == P(-.n) + <1>1. \A n \in Nat: Q(n) + <2>1. Q(0) + BY P(0), -.0 = 0 DEF Q + <2>2. ASSUME NEW n \in Nat, Q(n) + PROVE Q(Succ[n]) + <3>1. P(-.n) + BY <2>2 DEF Q + <3>2. -.n \in negNat + BY <2>2, n \in Nat DEF negNat + <3>3. \A i \in negNat: P(i) => P(-.Succ[-.i]) + OBVIOUS + <3>4. P(-.n) => P(-.Succ[-.-.n]) + BY <3>2, <3>3 + <3>5. P(-.Succ[-.-.n]) + BY <3>1, <3>4 + <3>6. P(-.Succ[n]) + <4>1. -.-.n = n + BY <2>2, minus_involutive, nat_in_int + <4> QED + BY <3>5, <4>1 + <3> QED + BY <3>6 DEF Q + <2> QED + BY <2>1, <2>2, NatInduction + <1>2. ASSUME NEW i \in negNat + PROVE P(i) + <2>1. \E n \in Nat: n = -.i + <3>1. i \in negNat + BY <1>3 + <3>2. \E n \in Nat: i = -.n + BY <3>1 DEF negNat + <3>3. \E n \in Nat: -.i = -.-.n + BY <3>2 + <3>4. \A n \in Nat: -.-.n = n + BY minus_involutive, nat_in_int + <3> QED + BY <3>3, <3>4 + <2>2. \E n \in Nat: Q(-.i) + <3>1. \E n \in Nat: n = -.i /\ Q(n) + BY <1>1, <2>1 + <3>2. \E n \in Nat: n = -.i /\ Q(-.i) + BY <3>1 + <3> QED + BY <3>2 + <2>3. Q(-.i) + BY <2>2 + <2>4. P(-.-.i) + BY <2>3 DEF Q + <2>5. -.-.i = i + <3>1. i \in Int + BY <1>2 DEF negNat, Int + <3> QED + BY <3>1, minus_involutive + <2> QED + BY <2>4, <2>5 + <1> QED + BY <1>2 +*) +theorem neg_nat_induction: + assumes z: "P(0)" and + sc: "\ i. + \ + i \\in negNat; + P(i) + \ \ + P(-.Succ[-.i])" + shows "\ i \ negNat: P(i)" + proof - + { + fix "Q" :: "[c] \ c" + assume q_def: "\ n: Q(n) = P(-.n)" + have s1_1: "\ n \ Nat: Q(n)" + proof - + have s2_1: "Q(0)" + proof - + have s3_1: "P(-.0)" + using z neg0 by auto + show "Q(0)" + using s3_1 q_def allE by auto + qed + have s2_2: "\ n. + \ + n \ Nat; Q(n) + \ \ + Q(Succ[n])" + proof - + fix "n" :: "c" + assume ndom: "n \ Nat" + assume Qn: "Q(n)" + have s3_1: "P(-.n)" + using Qn q_def by auto + have s3_2: "-.n \ negNat" + unfolding negNat_def + using ndom by auto + have s3_3: "\ i. \ i \ negNat; + P(i)\ \ + P(-.Succ[-.i])" + using sc by auto + have s3_4: "P(-.n) ==> + P(-.Succ[-.-.n])" + using s3_2 s3_3 by auto + have s3_5: "P(-.Succ[-.-.n])" + using s3_1 s3_4 by auto + have s3_6: "P(-.Succ[n])" + proof - + have s4_1: "-.-.n = n" + using ndom minus_involutive nat_in_int + by auto + show "P(-.Succ[n])" + using s3_5 s4_1 by auto + qed + show "Q(Succ[n])" + using s3_6 q_def by auto + qed + show "\ n \ Nat: Q(n)" + using s2_1 s2_2 natInduct + by auto + qed + have s1_2: "\ i. i \\in negNat ==> P(i)" + proof - + fix "i" :: "c" + assume idom: "i \ negNat" + have s2_1: "\ n \ Nat: n = -.i" + proof - + have s3_1: "i \ negNat" + using idom by auto + have s3_2: "\ n \ Nat: i = -.n" + using s3_1 by (auto simp: negNat_def) + have s3_3: "\ n \ Nat: -.i = -.-.n" + using s3_2 by auto + have s3_4: "\ n \ Nat: -.-.n = n" + using minus_involutive nat_in_int + by auto + show "\ n \ Nat: n = -.i" + using s3_3 s3_4 by auto + qed + have s2_2: "\ n \ Nat: Q(-.i)" + proof - + have s3_1: "\ n \ Nat: + n = -.i \ Q(n)" + using s1_1 s2_1 by auto + have s3_2: "\ n \ Nat: + n = -.i \ Q(-.i)" + using s3_1 by auto + show "\ n \ Nat: Q(-.i)" + using s3_2 by auto + qed + have s2_3: "Q(-.i)" + using s2_2 by auto + have s2_4: "P(-.-.i)" + using s2_3 q_def by auto + have s2_5: "-.-.i = i" + proof - + have s3_1: "i \ Int" + using idom + by (auto simp: negNat_def Int_def) + show "-.-.i = i" + using s3_1 minus_involutive + by auto + qed + show "P(i)" + using s2_4 s2_5 by auto + qed + have "\ i \ negNat: P(i)" + proof - + have s2_1: + "\ i. i \ negNat => P(i)" + using s1_2 impI by auto + have s2_2: + "\ i: i \ negNat => P(i)" + using s2_1 allI by auto + show "\ i \ negNat: P(i)" + using s2_2 by auto + (* by (rule allI, rule impI, rule s1_2) *) + qed + } + from this have "\ Q: ( + (\ n: Q(n) = P(-.n)) + => + (\ i \ negNat: P(i)) + )" + by auto + from this have " + (\ n: P(-.n) = P(-.n)) + => + (\ i \ negNat: P(i))" + by auto + thus " + (\ i \ negNat: P(i))" + by auto + qed + +(*----------------------------------------------------------------------------*) +(* Primitive recursion in two directions: plus and minus infinity. *) + + +theorem bprimrec_int: + assumes + e: "e \ S" and + succ_h: "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + succ_g: "\ n \ Nat: + \ x \ S: + g(n, x) \ S" + shows "\ f \ [Int -> S]: + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + )" + proof - + from primrec_int[of e h g] obtain f where + 1: "isAFcn(f)" and + 2: "DOMAIN f = Int" and + 3: "f[0] = e" and + 4: "\ n \ Nat: + f[Succ[n]] = h(n, f[n]) + \ f[-.Succ[n]] = g(n, f[-.n])" + by blast + have s1_2: "\ n \ Nat: f[n] \ S" + proof (rule natInduct) + from 3 e show "f[0] \ S" by simp + next + fix n + assume "n \ Nat" and "f[n] \ S" + with succ_h 4 show "f[Succ[n]] \ S" + by force + qed + have s1_3: "\ i \ negNat: f[i] \ S" + proof (rule neg_nat_induction) + from 3 e show "f[0] \ S" by simp + next + fix i + assume ndom: "i \ negNat" and + fidom: "f[i] \ S" + have s2_1: "-.i \ Nat" + using ndom minus_neg_nat_in_nat by auto + have s2_2: "f[-.Succ[-.i]] = g(-.i, f[-.-.i])" + using 4 s2_1 by auto + have s2_3: "g(-.i, f[-.-.i]) = g(-.i, f[i])" + using ndom neg_nat_in_int minus_involutive + by auto + have s2_4: "g(-.i, f[i]) \ S" + using s2_1 fidom succ_g by auto + show "f[-.Succ[-.i]] \ S" + using s2_2 s2_3 s2_4 by auto + qed + have 5: "\ i \ Int: f[i] \ S" + using s1_2 s1_3 int_union_nat_negnat + by auto + with 1 2 3 4 5 show ?thesis + by blast + qed + + +theorem primrec_intE: + assumes e: "e \ S" and + succ_h: "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + succ_g: "\ n \ Nat: + \ x \ S: + g(n, x) \ S" and + f: "f = (CHOOSE r \ [Int -> S]: + r[0] = e \ + (\ n \ Nat: + r[Succ[n]] = h(n, r[n]) \ + r[-.Succ[n]] = g(n, r[-.n]) + ))" (is "f = ?r") and + maj: "\ + f \ [Int -> S]; + f[0] = e; + \ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + \ + \ P" + shows "P" + proof - + from e succ_h succ_g have " + \ r \ [Int -> S]: + r[0] = e \ + (\ n \ Nat: + r[Succ[n]] = h(n, r[n]) \ + r[-.Succ[n]] = g(n, r[-.n]) + )" + by (rule bprimrec_int) + hence "?r \ [Int -> S] \ + ?r[0] = e \ + (\ n \ Nat: + ?r[Succ[n]] = h(n, ?r[n]) \ + ?r[-.Succ[n]] = g(n, ?r[-.n]) + )" + by (rule bChooseI2, auto) + with f maj show ?thesis by auto + qed + + +theorem bprimrecType_int: + assumes "e \ S" and + "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + "\ n \ Nat: + \ x \ S: + g(n, x) \ S" + shows "(CHOOSE f \ [Int -> S]: + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + )) \ [Int -> S]" + by (rule primrec_intE[OF assms], auto) + + +(*----------------------------------------------------------------------------*) +(* Typeness of addition and recursive properties. *) + + +theorem addintType: + assumes mdom: "m \ Int" + shows "addint(m) \ [Int -> Int]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: "(CHOOSE f \ [Int -> Int]: + f[0] = m \ + (\ n \ Nat: + f[Succ[n]] = iSucc[f[n]] \ + f[-.Succ[n]] = iPred[f[-.n]] + )) \ [Int -> Int]" + using mdom s1_1 s1_2 + by (rule bprimrecType_int) + show "addint(m) \ [Int -> Int]" + unfolding addint_def + using s1_3 by auto + qed + + +theorem addIsInt: + assumes "m \ Int" and "n \ Int" + shows "add(m, n) \ Int" + unfolding add_def + using assms addintType by blast + + +theorem diffIsInt: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "diff(m, n) \\in Int" + proof - + have s1_1: "diff(m, n) = add(m, -.n)" + unfolding diff_def by auto + have s1_2: "-.n \\in Int" + using ndom neg_int_eq_int by auto + have s1_3: "add(m, -.n) \\in Int" + using mdom s1_2 addIsInt by auto + show ?thesis + using s1_1 s1_3 by auto + qed + + +theorem addint_0: + assumes mdom: "m \ Int" + shows "addint(m)[0] = m" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: " + addint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = m \ + (\ n \ Nat: + r[Succ[n]] = iSucc[r[n]] \ + r[-.Succ[n]] = iPred[r[-.n]]) + )" + unfolding addint_def + by auto + have s1_4: " + \ + addint(m) \ [Int -> Int]; + addint(m)[0] = m; + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + \ + \ addint(m)[0] = m" + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "m" "Int" + "\ n x. iSucc[x]" + "\ n x. iPred[x]" + "addint(m)"] + by auto + qed + + +theorem add_0: + assumes "m \ Int" + shows "m + 0 = m" + unfolding add_def + using assms addint_0 by auto + + +theorem addint_succ: + assumes mdom: "m \ Int" + shows "\ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: " + addint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = m \ + (\ n \ Nat: + r[Succ[n]] = iSucc[r[n]] \ + r[-.Succ[n]] = iPred[r[-.n]]) + )" + unfolding addint_def + by auto + have s1_4: " + \ + addint(m) \ [Int -> Int]; + addint(m)[0] = m; + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + \ + \ + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + " + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "m" "Int" + "\ n x. iSucc[x]" + "\ n x. iPred[x]" + "addint(m)"] + by auto + qed + + +theorem addint_succ_nat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m + Succ[n] = iSucc[m + n]" + unfolding add_def + using mdom ndom addint_succ[of "m"] spec by auto + + +theorem addint_succ_negnat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m + -.Succ[n] = iPred[m + -.n]" + unfolding add_def + using mdom ndom addint_succ[of "m"] spec by auto + + +theorem nat_add_1: + assumes ndom: "n \\in Nat" + shows "Succ[n] = n + 1" + proof - + have s1_1: "Succ[n] = iSucc[n]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "n \\in Nat" + using ndom by auto + show ?thesis + unfolding iSucc_def + using s2_1 s2_2 by auto + qed + also have s1_2: "... = iSucc[n + 0]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "n + 0 = n" + using s2_1 add_0 by auto + have s2_3: "n = n + 0" + using s2_2[symmetric] by auto + show ?thesis + using s2_3 by auto + qed + also have s1_3: "... = n + Succ[0]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + using s2_1 s2_2 addint_succ_nat by auto + qed + also have s1_4: "... = n + 1" + by auto (* 1 is an abbreviation *) + from calculation show ?thesis . + qed + + +theorem nat_add_in_nat: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" + shows "m + n \\in Nat" + proof - + have s1_1: "m + 0 \\in Nat" + proof - + have s2_1: "m + 0 = m" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + show ?thesis + using s3_1 add_0 by auto + qed + have s2_2: "m \\in Nat" + using mdom by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_2: "\ k. \ + k \\in Nat; + m + k \\in Nat + \ \ + m + Succ[k] \\in Nat" + proof - + fix "k" :: "c" + assume s1_2_kdom: "k \\in Nat" + assume s1_2_induct: "m + k \\in Nat" + have s2_1: "m + Succ[k] = iSucc[m + k]" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + have s3_2: "k \\in Nat" + using s1_2_kdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + have s2_2: "iSucc[m + k] = Succ[m + k]" + proof - + have s3_1: "m + k \\in Nat" + using s1_2_induct by auto + have s3_2: "m + k \\in Int" + using s1_2_induct nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_3: "Succ[m + k] \\in Nat" + using s1_2_induct succIsNat by auto + show "m + Succ[k] \\in Nat" + using s2_1 s2_2 s2_3 by auto + qed + have s1_3: "\ k \ Nat: m + k \\in Nat" + using s1_1 s1_2 natInduct by auto + show ?thesis + using s1_3 ndom spec by auto + qed + + +theorem nat_0_succ: + "\ n \ Nat: + n=0 \ + (\ m \ Nat: n = Succ[m])" + by (rule natInduct, auto) + + +theorem nat_add_0: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" and + mn: "m + n = 0" + shows "m = 0 \ n = 0" + proof - + have s1_1: "n \ 0 ==> FALSE" + proof - + assume s2_1: "n \ 0" + have s2_2: "\ k: k \\in Nat \ n = Succ[k]" + using ndom s2_1 nat_0_succ by auto + def P \ "\ x. x \\in Nat \ n = Succ[x]" + def r \ "CHOOSE x: P(x)" + have s2_3: "r \\in Nat \ n = Succ[r]" + proof - + have s3_1: "\ x: P(x)" + using s2_2 by (auto simp: P_def) + have s3_2: "P(r)" + using s3_1 chooseI_ex by (auto simp: r_def) + show ?thesis + using s3_2 by (auto simp: P_def) + qed + have s2_4: "m + n = m + Succ[r]" + using s2_3 by auto + have s2_5: "m + Succ[r] = iSucc[m + r]" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + have s3_2: "r \\in Nat" + using s2_3 by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + have s2_6: "iSucc[m + r] = Succ[m + r]" + proof - + have s3_1: "m + r \\in Nat" + using mdom s2_3 nat_add_in_nat by auto + have s3_2: "m + r \\in Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_7: "Succ[m + r] \ 0" + proof - + have s3_1: "m + r \\in Nat" + using mdom s2_3 nat_add_in_nat by auto + show ?thesis + using s3_1 succNotZero by auto + qed + have s2_7: "m + n \ 0" + using s2_4 s2_5 s2_6 s2_7 by auto + show "FALSE" + using s2_7 mn by auto + qed + have s1_2: "n = 0" + using s1_1 by auto + have s1_3: "m + 0 = 0" + using mn s1_2 by auto + have s1_4: "m = 0" + using s1_3 mdom nat_in_int add_0 by auto + show ?thesis + using s1_2 s1_4 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Typeness of multiplication and recursive properties. *) + + +theorem multType: + assumes mdom: "m \ Int" + shows "multint(m) \ [Int -> Int]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: "(CHOOSE f \ [Int -> Int]: + f[0] = 0 \ + (\ n \ Nat: + f[Succ[n]] = f[n] + m \ + f[-.Succ[n]] = f[-.n] + -.m + )) \ [Int -> Int]" + using zero_in_int s1_1 s1_2 + by (rule bprimrecType_int[of "0"]) + show "multint(m) \ [Int -> Int]" + unfolding multint_def + using s1_3 by auto + qed + + +theorem multIsInt: + assumes "m \ Int" and "n \ Int" + shows "mult(m, n) \ Int" + unfolding mult_def + using assms multType by blast + + +theorem multint_0: + assumes mdom: "m \ Int" + shows "multint(m)[0] = 0" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: " + multint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = 0 \ + (\ n \ Nat: + r[Succ[n]] = r[n] + m \ + r[-.Succ[n]] = r[-.n] + -.m) + )" + unfolding multint_def + by auto + have s1_4: " + \ + multint(m) \ [Int -> Int]; + multint(m)[0] = 0; + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + \ + \ + multint(m)[0] = 0" + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "0" "Int" + "\ n x. x + m" + "\ n x. x + -.m" + "multint(m)"] + by auto + qed + + +theorem mult_0: + assumes "m \ Int" + shows "m * 0 = 0" + unfolding mult_def + using assms multint_0 by auto + + +theorem multint_succ: + assumes mdom: "m \ Int" + shows "\ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: " + multint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = 0 \ + (\ n \ Nat: + r[Succ[n]] = r[n] + m \ + r[-.Succ[n]] = r[-.n] + -.m) + )" + unfolding multint_def + by auto + have s1_4: " + \ + multint(m) \ [Int -> Int]; + multint(m)[0] = 0; + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + \ + \ + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + " + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "0" "Int" + "\ n x. x + m" + "\ n x. x + -.m" + "multint(m)"] + by auto + qed + + +theorem multint_succ_nat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m * Succ[n] = m * n + m" + unfolding mult_def + using mdom ndom multint_succ[of "m"] spec by auto + + +theorem multint_succ_negnat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m * -.Succ[n] = m * -.n + -.m" + unfolding mult_def + using mdom ndom multint_succ[of "m"] spec by auto + + +theorem nat_mult_in_nat: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" + shows "m * n \\in Nat" + proof - + have s1_1: "m * 0 \\in Nat" + proof - + have s2_1: "m * 0 = 0" + using mdom nat_in_int mult_0 by auto + have s2_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_2: "\ k. \ + k \\in Nat; + m * k \\in Nat + \ \ + m * Succ[k] \\in Nat" + proof - + fix "k" :: "c" + assume s1_2_kdom: "k \\in Nat" + assume s1_2_induct: "m * k \\in Nat" + have s2_1: "m * Succ[k] = m * k + m" + using mdom s1_2_kdom multint_succ_nat by auto + have s2_2: "m * k + m \\in Nat" + using s1_2_induct mdom nat_add_in_nat by auto + show "m * Succ[k] \\in Nat" + using s2_1 s2_2 by auto + qed + have s1_3: "\ k \ Nat: m * k \\in Nat" + using s1_1 s1_2 natInduct by auto + show ?thesis + using s1_3 ndom spec by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Commutativity of addition. *) + + +(* +THEOREM pred_irrefl == + ASSUME NEW n \in Nat \ {0} + PROVE Pred[n] # n +PROOF +<1>1. Pred[n] \in Nat + BY n \in Nat, Pred_in_nat +<1>2. Succ[Pred[n]] # Pred[n] + BY <1>1, succIrrefl +<1>3. Succ[Pred[n]] = n + BY n \in Nat \ {0}, Succ_Pred_nat +<1> QED + BY <1>2, <1>3 +*) +theorem pred_irrefl: + assumes ndom: "n \\in Nat \ {0}" + shows "Pred[n] \ n" + proof - + have s1_1: "Pred[n] \\in Nat" + using ndom Pred_in_nat by auto + have s1_2: "Succ[Pred[n]] \ Pred[n]" + using s1_1 succIrrefl by auto + have s1_3: "Succ[Pred[n]] = n" + using ndom Succ_Pred_nat by auto + show ?thesis + using s1_2 s1_3 by auto + qed + + +theorem ipred_plus_1: + assumes hyp: "a \ Int" + shows "iPred[a + 1] = a" + proof - + have s1_1: "iPred[a + 1] = iPred[a + Succ[0]]" + by auto + have s1_2: "... = iPred[iSucc[a + 0]]" + proof - + have s2_1: "a \ Int" + using hyp by auto + have s2_2: "0 \ Nat" + using zeroIsNat by auto + have s2_3: "a + Succ[0] = iSucc[a + 0]" + using s2_1 s2_2 addint_succ_nat by auto + show ?thesis + using s2_3 by auto + qed + have s1_3: "... = iPred[iSucc[a]]" + using hyp add_0 by auto + have s1_4: "... = a" + using hyp iSucciPredCommute spec by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +(* +THEOREM AddCommutativeNatNat == + \A j \in Nat: \A i \in Nat: + j + i = i + j +PROOF +<1>1. \A i \in Nat: 0 + i = i + 0 + <2>1. 0 + 0 = 0 + 0 + <3>1. 0 + 0 = addint(0)[0] + BY DEF + + <3>2. @ = 0 + BY DEF addint + <3> QED + BY <3>1, <3>2 + + <2>2. ASSUME NEW i \in Nat, 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + + <3>1. 0 + Succ[i] = iSucc[0 + i] + BY DEF + + <3>2. @ = iSucc[i + 0] + BY <2>2 + <3>3. @ = iSucc[i] + BY DEF + + <3>4. @ = Succ[i] + BY <2>2, i \in Nat + <3>5. @ = Succ[i] + 0 + BY DEF + + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5 + <2> QED + BY <2>1, <2>2, NatInduction + +<1>2. ASSUME NEW j \in Nat, + \A i \in Nat: j + i = i + j + PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] + + <2>1. Succ[j] + 0 = 0 + Succ[j] + <3>1. Succ[j] + 0 = Succ[j] + BY DEF + + <3>2. @ = Succ[j + 0] + BY DEF + + <3>3. @ = Succ[0 + j] + BY <1>2 + <3>4. @ = 0 + Succ[j] + BY DEF + + <3> QED + BY <3>1, <3>2, <3>3, <3>4 + + <2>2. ASSUME NEW i \in Nat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] + + <3>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] + <4>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. Succ[i] \in Nat + BY <2>2, succIsNat + <5> QED + BY <5>1, <5>2, addint_succ + <4>2. @ = iSucc[i + Succ[j]] + BY <2>2 + <4>3. @ = iSucc[iSucc[i + j]] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, addint_succ + <4>4. @ = iSucc[iSucc[j + i]] + BY <1>2, <2>2 + <4> QED + BY <4>1, <4>2, <4>3, <4>4 + + <3>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] + <4>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. Succ[j] \in Nat + BY <1>2, succIsNat + <5> QED + BY <5>1, <5>2, addint_succ + <4>2. @ = iSucc[j + Succ[i]] + <5>1. Succ[i] \in Nat + BY <2>2, succIsNat + <5> QED + BY <1>2, <5>1 + <4>3. @ = iSucc[iSucc[j + i]] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, addint_succ + <4> QED + BY <4>1, <4>2, <4>3 + + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, NatInduction +<1> QED + BY <1>1, <1>2, NatInduction + + +THEOREM AddCommutative == + \A j \in Int: \A i \in Int: + j + i = i + j +PROOF +<1>1. \A j \in Nat: \A i \in Int: + j + i = i + j + <2>1. \A j \in Nat: \A i \in Nat: + j + i = i + j + \* The proof of step <2>1 is the same with the + \* proof of the theorem `AddCommutativeNatNat`. + <3>1. \A i \in Nat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in Nat, + 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + <5>1. 0 + Succ[i] = iSucc[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Nat + BY <4>2 + <6> QED + BY <6>1, <6>2, addint_succ + DEF addd + <5>2. @ = iSucc[i + 0] + BY <4>2 + <5>3. @ = iSucc[i] + <6>1. i \in Int + BY <4>2, nat_in_int + <6> QED + BY <6>1 addint_0 + <5>4. @ = Succ[i] + BY <4>2 DEF iSucc + <5>5. @ = Succ[i] + 0 + <6>1. Succ[i] \in Nat + BY <4>2, succIsNat + <6>2. Succ[i] \in Int + BY <6>1, nat_in_int + <6> QED + BY <6>2, addint_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, NatInduction + <3>2. ASSUME NEW j \in Nat, + \A i \in Nat: j + i = i + j + PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] + <4>1. Succ[j] + 0 = 0 + Succ[j] + <5>1. Succ[j] + 0 = Succ[j] + <6>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <6> QED + BY <6>1, addint_0 + <5>2. @ = iSucc[j] + BY <3>2, nat_in_int DEF iSucc + <5>3. @ = iSucc[j + 0] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iSucc[0 + j] + <6>1. j + 0 = 0 + j + BY <3>2, zeroIsNat + <6> QED + By <6>1 + <5>5. @ = 0 + Succ[j] + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2, addint_succ + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4>2. ASSUME NEW i \in Nat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] + <5>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] + <6>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] + <7>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6>2. @ = iSucc[i + Succ[j]] + BY <4>2 + <6>3. @ = iSucc[iSucc[i + j]] + <7>1. i \in Int + BY <4>2, nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6>4. @ = iSucc[iSucc[j + i]] + BY <3>2, <4>2 + <6> QED + By <6>1, <6>2, <6>3, <6>4 + <5>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] + <6>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] + <7>1. Succ[i] \in Int + BY <4>2, succIsNat, nat_in_int + <7>2. j \in Nat + BY <3>2, succIsNat + <7> QED + BY <7>1, <7>2, addint_succ + <6>2. @ = iSucc[j + Succ[i]] + <7>1. Succ[i] \in Nat + BY <4>2, succIsNat + <7> QED + BY <3>2, <7>1 + <6>3. @ = iSucc[iSucc[j + i]] + <7>1. j \in Int + BY <3>2, nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6> QED + BY <6>1, <6>2, <6>3 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, NatInduction + <3> QED + BY <3>1, <3>2, NatInduction + <2>2. \A j \in Nat: \A i \in negNat: + j + i = i + j + <3>1. \A i \in negNat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in negNat, + 0 + i = i + 0 + PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 + <5>1. 0 + -.Succ[-.i] = iPred[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <6>3. 0 + -.Succ[-.i] = iPred[0 + -.-.i] + BY <6>1, <6>2, addint_succ_negnat + <6>4. -.-.i = i + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7> QED + BY <7>1, minus_involutive + <6> QED + BY <6>3, <6>4 + <5>2. @ = iPred[i + 0] + BY <4>2 + <5>3. @ = iPred[i] + <6>1. i \in Int + BY <4>2, neg_nat_in_int + <6>2. i + 0 = i + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = -.Succ[-.i] + <6>1. i \in Int + BY <4>2, neg_nat_in_int + <6>2. i \notin Nat \ {0} + BY <4>2, neg_nat_not_in_nat_setminus_0 + <6> QED + BY <6>1, <6>2 DEF iPred + <5>5. @ = -.Succ[-.i] + 0 + <6>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <6>2. Succ[-.i] \in Nat + BY <6>1, succIsNat + <6>3. -.Succ[-.i] \in Int + BY <6>2, minus_nat_in_int + <6> QED + BY <6>3, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3>2. ASSUME NEW j \in Nat, + \A i \in negNat: j + i = i + j + PROVE \A i \in negNat: Succ[j] + i = i + Succ[j] + <4>1. Succ[j] + 0 = 0 + Succ[j] + <5>1. Succ[j] + 0 = Succ[j] + <6>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <6> QED + BY <6>1, add_0 + <5>2. @ = iSucc[j] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2 DEF iSucc + <5>3. @ = iSucc[j + 0] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iSucc[0 + j] + <6>1. 0 \in negNat + BY neg0 DEF negNat + <6> QED + BY <3>2, <6>1 + <5>5. @ = 0 + Succ[j] + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2, addint_succ_nat + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4>2. ASSUME NEW i \in negNat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j] + <5>1. Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]] + <6>1. Succ[j] + -.Succ[-.i] = iPred[Succ[j] + i] + <7>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7>3. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7>4. Succ[j] + -.Succ[-.i] + = iPred[Succ[j] + -.-.i] + BY <7>1, <7>2, addint_succ_negnat + <7> QED + BY <7>3, <7>4 + <6>2. @ = iPred[i + Succ[j]] + BY <4>2 + <6>3. @ = iPred[iSucc[i + j]] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>4. @ = iPred[iSucc[j + i]] + <7>1. i \in negNat + BY <4>2 + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + i = i + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. -.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]] + <6>1. -.Succ[-.i] + Succ[j] = iSucc[-.Succ[-.i] + j] + <7>1. -.Succ[-.i] \in Int + BY <4>2, minus_neg_nat_in_nat, succIsNat, + minus_nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>2. @ = iSucc[j + -.Succ[-.i]] + <7>1. -.Succ[-.i] \in negNat + BY <4>2, minus_neg_nat_in_nat, succIsNat, + minus_nat_in_neg_nat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6>3. @ = iSucc[iPred[j + i]] + <7>1. j \in Int + BY <3>2, nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7>3. j + -.Succ[-.i] = iPred[j + -.-.i] + BY <7>1, <7>2, addint_succ_negnat + <7>4. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>3, <7>4 + <6>4. @ = iPred[iSucc[j + i]] + <7>1. j + i \in Int + <8>1. j \in Int + BY <3>2, nat_in_int + <8>2. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, <8>2, addIsInt + <7> QED + BY <7>1, iSucciPredCommute + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3> QED + BY <3>1, <3>2, NatInduction + <2> QED + BY <2>1, <2>2, int_union_nat_negnat +<1>2. \A j \in negNat: \A i \in Int: + j + i = i + j + <2>1. \A j \in negNat: \A i \in Nat: + j + i = i + j + <3>1. \A i \in Nat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME + NEW i \in Nat, + 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + \* The level 5 proof is the same as + \* for step <1>1!<2>1!<3>1!<4>2 . + <5>1. 0 + Succ[i] = iSucc[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Nat + BY <4>2 + <6> QED + BY <6>1, <6>2 DEF addint_succ_nat + <5>2. @ = iSucc[i + 0] + BY <4>2 + <5>3. @ = iSucc[i] + <6>1. i \in Int + BY <4>2, nat_in_int + <6>2. i + 0 = i + BY <6>1, addint_0 + <6> QED + BY <6>2 + <5>4. @ = Succ[i] + <6>1. i \in Nat + BY <4>2 + <6>2. i \in Int + BY <4>2, nat_in_int + <6> QED + BY <6>1, <6>2 DEF iSucc + <5>5. @ = Succ[i] + 0 + <6>1. Succ[i] \in Nat + BY <4>2, succIsNat + <6>2. Succ[i] \in Int + BY <6>1, nat_in_int + <6> QED + BY <6>2, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, NatInduction + <3>2. ASSUME + NEW j \in negNat, + \A i \in Nat: j + i = i + j + PROVE + \A i \in Nat: + -.Succ[-.j] + i = i + -.Succ[-.j] + <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] + <5>1. -.Succ[-.j] + 0 = -.Succ[-.j] + <6>1. -.Succ[-.j] \in Int + <7>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>2. Succ[-.j] \in Nat + BY <7>1, succIsNat + <7> QED + BY <7>2, minus_nat_in_int + <6> QED + BY <6>1, add_0 + <5>2. @ = iPred[j] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. j \notin Nat \ {0} + BY <3>2, neg_nat_not_in_nat_setminus_0 + <6> QED + BY <6>1, <6>2 DEF iPred + <5>3. @ = iPred[j + 0] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iPred[0 + j] + <6>1. 0 \in Nat + BY zeroIsNat + <6>2. \A k \in Nat: j + k = k + j + BY <3>2 + <6>3. j + 0 = 0 + j + BY <6>1, <6>2 + <6> QED + BY <6>3 + <5>5. @ = iPred[0 + -.-.j] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. -.-.j = j + BY <6>1, minus_involutive + <6> QED + BY <6>2 + <5>6. @ = 0 + -.Succ[-.j] + <6>1. 0 \in Int + BY zeroIsNat, nat_in_int + <6>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <6> QED + BY <6>1, <6>2, addint_succ_negnat + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5, <5>6 + <4>2. ASSUME + NEW i \in Nat, + -.Succ[-.j] + i = i + -.Succ[-.j] + PROVE + -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j] + <5>1. -.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]] + <6>1. -.Succ[-.j] + Succ[i] = + iSucc[-.Succ[-.j] + i] + <7>1. -.Succ[-.j] \in Int + <8>1. j \in negNat + BY <3>2 + <8>2. -.j \in Nat + BY <8>1, minus_neg_nat_in_nat + <8>3. Succ[-.j] \in Nat + BY <8>2, succIsNat + <8> QED + BY <8>3, minus_nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>2. @ = iSucc[i + -.Succ[-.j]] + BY <4>2 + <6>3. @ = iSucc[iPred[i + j]] + <7>1. i \in Int + BY <4>2, nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>3. -.-.j = j + BY <3>2, neg_nat_in_int, minus_involutive + <7> QED + BY <7>1, <7>2, <7>3, addint_succ_negnat + <6>4. @ = iSucc[iPred[j + i]] + <7>1. i \in Nat + BY <4>2 + <7>2. j + i = i + j + BY <3>2, <7>1 + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]] + <6>1. Succ[i] + -.Succ[-.j] = + iPred[Succ[i] + -.-.j] + <7>1. Succ[i] \in Int + BY <4>2, succIsNat, nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[Succ[i] + j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>2 + <6>3. @ = iPred[j + Succ[i]] + <7>1. Succ[i] \in Nat + BY <4>2, succIsNat + <7>2. \A k \in Nat: j + k = k + j + BY <3>2 + <7>3. j + Succ[i] = Succ[i] + j + BY <7>1, <7>2 + <7> QED + BY <7>3 + <6>4. @ = iPred[iSucc[j + i]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>5. @ = iSucc[iPred[j + i]] + <7>1. j + i \in Int + <8>1. j \in Int + BY <3>2, neg_nat_in_int + <8>2. i \in Int + BY <4>2, nat_in_int + <8> QED + BY <8>1, <8>2, addIsInt + <7> QED + BY <7>1, iSucciPredCommute + <6> QED + BY <6>1, <6>2, <6>3, <6>4, <6>5 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, NatInduction + <3> QED + BY <3>1, <3>2, neg_nat_induction + <2>2. \A j \in negNat: \A i \in negNat: + j + i = i + j + <3>1. \A i \in negNat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in negNat, 0 + i = i + 0 + PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 + <5>1. 0 + -.Succ[-.i] = iPred[i] + <6>1. 0 + -.Succ[-.i] = iPred[0 + -.-.i] + <7>1. 0 \in Int + BY zero_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[0 + i] + <7>1. -.-.i = i + BY <4>2, neg_nat_in_int, minus_involutive + <7> QED + BY <7>1 + <6>3. @ = iPred[i + 0] + BY <4>2 + <6>4. @ = iPred[i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. i + 0 = i + BY <7>1, add_0 + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. -.Succ[-.i] + 0 = iPred[i] + <6>1. -.Succ[-.i] + 0 = -.Succ[-.i] + <7>1. -.Succ[-.i] \in Int + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7> QED + BY <7>1, add_0 + <6>2. @ = iPred[i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. i \notin Nat \ {0} + BY <4>2, neg_nat_not_in_nat_setminus_0 + <7> QED + BY <7>1, <7>2 DEF iPred + <6> QED + BY <6>1, <6>2 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3>2. ASSUME NEW j \in negNat, + \A i \in negNat: j + i = i + j + PROVE \A i \in negNat: + -.Succ[-.j] + i = i + -.Succ[-.j] + <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] + <5>1. -.Succ[-.j] + 0 = iPred[j] + <6>1. -.Succ[-.j] + 0 = -.Succ[-.j] + <7>1. -.Succ[-.j] \in Int + <8>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <8>2. Succ[-.j] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7> QED + BY <7>1, add_0 + <6>2. @ = iPred[j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. j \notin Nat \ {0} + BY <3>2, neg_nat_not_in_nat_setminus_0 + <7> QED + BY <7>1, <7>2 DEF iPred + <6> QED + BY <6>1, <6>2 + <5>2. 0 + -.Succ[-.j] = iPred[j] + <6>1. 0 + -.Succ[-.j] = iPred[0 + -.-.j] + <7>1. 0 \in Int + BY zero_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[0 + j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>1, <7>2 + <6>3. @ = iPred[j + 0] + <7>1. 0 \in negNat + BY neg0 DEF negNat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + 0 = 0 + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6>4. @ = iPred[j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7> QED + BY <7>1, add_0 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5> QED + BY <5>1, <5>2 + <4>2. ASSUME NEW i \in negNat, + -.Succ[-.j] + i = i + -.Succ[-.j] + PROVE -.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j] + <5>1. -.Succ[-.j] + -.Succ[-.i] = iPred[iPred[i + j]] + <6>1. -.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i] + <7>1. -.Succ[-.j] \in Int + <8>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <8>2. Succ[-.j] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[-.Succ[-.j] + i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. -.-.i = i + BY <7>1, minus_involutive + <7> QED + BY <7>1, <7>2 + <6>3. @ = iPred[i + -.Succ[-.j]] + BY <4>2 + <6>4. @ = iPred[iPred[i + -.-.j]] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>3. i + -.Succ[-.j] = iPred[i + -.-.j] + BY <7>1, <7>2, addint_succ_negnat + <7> QED + BY <7>3 + <6>5. @ = iPred[iPred[i + j]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4, <6>5 + <5>2. -.Succ[-.i] + -.Succ[-.j] = iPred[iPred[i + j]] + <6>1. -.Succ[-.i] + -.Succ[-.j] = + iPred[-.Succ[-.i] + -.-.j] + <7>1. -.Succ[-.i] \in Int + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[-.Succ[-.i] + j] + <7>1. -.-.j = j + <8>1. j \in Int + BY <3>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>1 + <6>3. @ = iPred[j + -.Succ[-.i]] + <7>1. -.Succ[-.i] \in negNat + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_neg_nat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j + BY <7>1, <7>2 + <7> QED + BY <7>3 + <6>4. @ = iPred[iPred[j + -.-.i]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>5. @ = iPred[iPred[j + i]] + <7>1. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>1 + <6>6. @ = iPred[iPred[i + j]] + <7>1. i \in negNat + BY <4>2 + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + i = i + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6> QED + BY <6>1, <6>2, <6>3, <6>4, + <6>5, <6>6 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3> QED + BY <3>1, <3>2, neg_nat_induction + <2> QED + BY <2>1, <2>2, int_union_nat_negnat +<1> QED + BY <1>1, <1>2 + +*) +theorem AddCommutative: + "\ j \ Int: \ i \ Int: + j + i = i + j" + proof - + have s1_1: "\ j \ Nat: \ i \ Int: + j + i = i + j" + proof - + have s2_1: "\ j \ Nat: + \ i \ Nat: + j + i = i + j" + proof - + have s3_1: "\ i \ Nat: + 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ Nat; + 0 + i = i + 0 + \ + \ + 0 + Succ[i] = Succ[i] + 0" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + Succ[i] = iSucc[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + unfolding add_def + using s6_1 s6_2 + addint_succ[of "0"] + by auto + qed + have s5_2: "iSucc[0 + i] = iSucc[i + 0]" + using s4_2_asm2 by auto + have s5_3: "iSucc[i + 0] = iSucc[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 nat_in_int + by auto + show ?thesis + unfolding add_def + using s6_1 addint_0[of "i"] + by auto + qed + have s5_4: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s4_2_asm1 by auto + have s5_5: "Succ[i] = Succ[i] + 0" + proof - + have s6_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + have s6_2: "Succ[i] \ Int" + using s6_1 nat_in_int by auto + show ?thesis + unfolding add_def + using s6_2 addint_0 by auto + qed + show "0 + Succ[i] = Succ[i] + 0" + using s5_1 s5_2 s5_3 s5_4 s5_5 + by auto + qed + show ?thesis + using s4_1 s4_2 + natInduct[of "\ i. 0 + i = i + 0"] + by auto + qed + have s3_2: "\ j. \ + j \ Nat; + \ i \ Nat: j + i = i + j + \ + \ + \ i \ Nat: + Succ[j] + i = i + Succ[j]" + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ Nat" + assume s3_2_asm2: "\ i \ Nat: + j + i = i + j" + have s4_1: "Succ[j] + 0 = 0 + Succ[j]" + proof - + have s5_1: "Succ[j] + 0 = Succ[j]" + proof - + have s6_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat + nat_in_int by auto + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iSucc[j]" + unfolding iSucc_def + using s3_2_asm1 nat_in_int + by auto + also have s5_3: "... = iSucc[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iSucc[0 + j]" + proof - + have s6_1: "j + 0 = 0 + j" + using s3_2_asm2 zeroIsNat by auto + show ?thesis + using s6_1 by auto + qed + also have s5_5: "... = 0 + Succ[j]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + have s6_3: "0 \ Int ==> + j \ Nat ==> + 0 + Succ[j] = iSucc[0 + j]" + using addint_succ_nat + by auto + have s6_4: " + 0 + Succ[j] = iSucc[0 + j]" + using s6_1 s6_2 s6_3 by auto + have s6_5: "iSucc[0 + j] = 0 + Succ[j]" + using s6_4 by auto + show ?thesis + using s6_5 by auto + qed + finally show ?thesis . + qed + have s4_2: "\ i. \ + i \ Nat; + Succ[j] + i = i + Succ[j] + \ \ + Succ[j] + Succ[i] = Succ[i] + Succ[j]" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" + have s5_1: "Succ[j] + Succ[i] = iSucc[iSucc[j + i]]" + proof - + have s6_1: "Succ[j] + Succ[i] = iSucc[Succ[j] + i]" + proof - + have s7_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "i \ Nat" + using s4_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + also have s6_2: "... = iSucc[i + Succ[j]]" + using s4_2_asm2 by auto + also have s6_3: "... = iSucc[iSucc[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + also have s6_4: "... = iSucc[iSucc[j + i]]" + using s3_2_asm2 s4_2_asm1 spec by auto + finally show ?thesis + by auto + qed + have s5_2: "Succ[i] + Succ[j] = iSucc[iSucc[j + i]]" + proof - + have s6_1: "Succ[i] + Succ[j] = iSucc[Succ[i] + j]" + proof - + have s7_1: "Succ[i] \ Int" + using s4_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "j \ Nat" + using s3_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[j + Succ[i]]" + proof - + have s7_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s3_2_asm2 by auto + qed + also have s6_3: "... = iSucc[iSucc[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 nat_in_int + by auto + have s7_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + finally show ?thesis + by auto + qed + show "Succ[j] + Succ[i] = Succ[i] + Succ[j]" + using s5_1 s5_2 by auto + qed + show "\ i \ Nat: Succ[j] + i = i + Succ[j]" + using s4_1 s4_2 natInduct + by auto + qed + show ?thesis + using s3_1 s3_2 natInduct by auto + qed + have s2_2: "\ j \ Nat: + \ i \ negNat: j + i = i + j" + proof - + have s3_1: "\ i \ negNat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ negNat; + 0 + i = i + 0 + \ \ + 0 + -.Succ[-.i] = -.Succ[-.i] + 0" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ negNat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + -.Succ[-.i] = iPred[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat + by auto + have s6_3: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" + using s6_1 s6_2 addint_succ_negnat + by auto + have s6_4: "-.-.i = i" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s7_1 minus_involutive by auto + qed + show ?thesis + using s6_3 s6_4 by auto + qed + also have s5_2: "... = iPred[i + 0]" + using s4_2_asm2 by auto + also have s5_3: "... = iPred[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s6_2: "i + 0 = i" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = -.Succ[-.i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s6_2: "i \ Nat \ {0}" + using s4_2_asm1 neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s6_1 s6_2 by auto + qed + also have s5_5: "... = -.Succ[-.i] + 0" + proof - + have s6_1: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat by auto + have s6_2: "Succ[-.i] \ Nat" + using s6_1 succIsNat by auto + have s6_3: "-.Succ[-.i] \ Int" + using s6_2 minus_nat_in_int by auto + show ?thesis + using s6_3 add_0 by auto + qed + finally show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" + by auto + qed + show ?thesis + using s4_1 s4_2 neg_nat_induction by auto + qed + have s3_2: "\ j. \ + j \ Nat; + \ i \ negNat: j + i = i + j + \ \ + \ i \ negNat: Succ[j] + i = i + Succ[j]" + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ Nat" + assume s3_2_asm2: "\ i \ negNat: + j + i = i + j" + have s4_1: "Succ[j] + 0 = 0 + Succ[j]" + proof - + have s5_1: "Succ[j] + 0 = Succ[j]" + proof - + have s6_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int by auto + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iSucc[j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + unfolding iSucc_def + using s6_1 s6_2 by auto + qed + also have s5_3: "... = iSucc[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iSucc[0 + j]" + proof - + have s6_1: "0 \ negNat" + unfolding negNat_def + using neg0 by auto + show ?thesis + using s3_2_asm2 s6_1 by auto + qed + also have s5_5: "... = 0 + Succ[j]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s6_1 s6_2 addint_succ_nat by auto + qed + finally show ?thesis + by auto + qed + have s4_2: "\ i. \ + i \ negNat; + Succ[j] + i = i + Succ[j] + \ \ + Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ negNat" + assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" + have s5_1: "Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]]" + proof - + have s6_1: "Succ[j] + -.Succ[-.i] = + iPred[Succ[j] + i]" + proof - + have s7_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat + by auto + have s7_3: "-.-.i = i" + using s4_2_asm1 neg_nat_in_int + minus_involutive by auto + have s7_4: "Succ[j] + -.Succ[-.i] = + iPred[Succ[j] + -.-.i]" + using s7_1 s7_2 addint_succ_negnat + by auto + show ?thesis + using s7_3 s7_4 by auto + qed + also have s6_2: "... = iPred[i + Succ[j]]" + using s4_2_asm2 by auto + also have s6_3: "... = iPred[iSucc[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "i \ negNat" + using s4_2_asm1 by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + i = i + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + finally show ?thesis + by auto + qed + have s5_2: "-.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]]" + proof - + have s6_1: "-.Succ[-.i] + Succ[j] = + iSucc[-.Succ[-.i] + j]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + using s4_2_asm1 minus_neg_nat_in_nat succIsNat + minus_nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[j + -.Succ[-.i]]" + proof - + have s7_1: "-.Succ[-.i] \ negNat" + using s4_2_asm1 minus_neg_nat_in_nat succIsNat + minus_nat_in_neg_nat by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + also have s6_3: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat by auto + have s7_3: "j + -.Succ[-.i] = iPred[j + -.-.i]" + using s7_1 s7_2 addint_succ_negnat by auto + have s7_4: "-.-.i = i" + proof - + have s8_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_3 s7_4 by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "j + i \ Int" + proof - + have s8_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s8_2: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s8_1 s8_2 addIsInt by auto + qed + show ?thesis + using s7_1 iSucciPredCommute by auto + qed + finally show ?thesis + by auto + qed + show "Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" + using s5_1 s5_2 by auto + qed + show "\ i \ negNat: Succ[j] + i = i + Succ[j]" + using s4_1 s4_2 neg_nat_induction + by auto + qed + show ?thesis + using s3_1 s3_2 natInduct by auto + qed + show ?thesis + proof - + have s3_1: "\ j. \ i. + j \ Nat => ( + i \ Int => ( + j + i = i + j))" + proof auto + fix "j" :: "c" + fix "i" :: "c" + assume s4_1: "j \ Nat" + assume s4_2: "i \ Int" + have s4_3: "i \ Nat ==> j + i = i + j" + proof - + assume s5_1: "i \ Nat" + show "j + i = i + j" + using s4_1 s5_1 s2_1 by auto + qed + have s4_4: "i \ negNat ==> j + i = i + j" + proof - + assume s5_1: "i \ negNat" + show "j + i = i + j" + using s4_1 s5_1 s2_2 by auto + qed + show "j + i = i + j" + using s4_2 s4_3 s4_4 int_union_nat_negnat by auto + qed + show ?thesis + using s3_1 by auto + qed + qed + have s1_2: "\ j \ negNat: \ i \ Int: + j + i = i + j" + proof - + have s2_1: "\ j \ negNat: \ i \ Nat: + j + i = i + j" + proof - + have s3_1: "\ i \ Nat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ Nat; + 0 + i = i + 0 + \ \ + 0 + Succ[i] = Succ[i] + 0 + " + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + Succ[i] = iSucc[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + using s6_1 s6_2 addint_succ_nat by auto + qed + also have s5_2: "... = iSucc[i + 0]" + using s4_2_asm2 by force + also have s5_3: "... = iSucc[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 nat_in_int by auto + have s6_2: "i + 0 = i" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by force + qed + also have s5_4: "... = Succ[i]" + proof - + have s6_1: "i \ Nat" + using s4_2_asm1 by auto + have s6_2: "i \ Int" + using s4_2_asm1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s6_1 s6_2 by auto + qed + also have s5_5: "... = Succ[i] + 0" + proof - + have s6_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + have s6_2: "Succ[i] \ Int" + using s6_1 nat_in_int by auto + show ?thesis + using s6_2 add_0 by auto + qed + finally show "0 + Succ[i] = Succ[i] + 0" + (* . did not work *) + using s5_1 s5_2 s5_3 s5_4 s5_5 by auto + qed + show ?thesis + using s4_1 s4_2 natInduct by auto + qed + have s3_2: "\ j. \ + j \ negNat; + \ i \ Nat: j + i = i + j + \ \ + \ i \ Nat: + -.Succ[-.j] + i = i + -.Succ[-.j] + " + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ negNat" + assume s3_2_asm2: "\ i \ Nat: + j + i = i + j" + have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" + proof - + have s6_1: "-.Succ[-.j] \ Int" + proof - + have s7_1: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat + by auto + have s7_2: "Succ[-.j] \ Nat" + using s7_1 succIsNat by auto + show ?thesis + using s7_2 minus_nat_in_int by auto + qed + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iPred[j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "j \ Nat \ {0}" + using s3_2_asm1 neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s6_1 s6_2 by auto + qed + also have s5_3: "... = iPred[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iPred[0 + j]" + proof - + have s6_1: "0 \ Nat" + using zeroIsNat by auto + have s6_2: "\ k \ Nat: + j + k = k + j" + using s3_2_asm2 by auto + have s6_3: "j + 0 = 0 + j" + using s6_1 s6_2 by auto + show ?thesis + using s6_3 by auto + qed + also have s5_5: "... = iPred[0 + -.-.j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "-.-.j = j" + using s6_1 minus_involutive by auto + show ?thesis + using s6_2 by auto + qed + also have s5_6: "... = 0 + -.Succ[-.j]" + proof - + have s6_1: "0 \ Int" + using zeroIsNat nat_in_int by auto + have s6_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + show ?thesis + using s6_1 s6_2 addint_succ_negnat by auto + qed + finally show ?thesis . + qed + have s4_2: "\ i. \ + i \ Nat; + -.Succ[-.j] + i = i + -.Succ[-.j] + \ \ + -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ Nat" + assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" + have s5_1: "-.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]]" + proof - + have s6_1: "-.Succ[-.j] + Succ[i] = + iSucc[-.Succ[-.j] + i]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "j \ negNat" + using s3_2_asm1 by auto + have s8_2: "-.j \ Nat" + using s8_1 minus_neg_nat_in_nat by auto + have s8_3: "Succ[-.j] \ Nat" + using s8_2 succIsNat by auto + show ?thesis + using s8_3 minus_nat_in_int by auto + qed + have s7_2: "i \ Nat" + using s4_2_idom by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[i + -.Succ[-.j]]" + using s4_2_induct by auto + also have s6_3: "... = iSucc[iPred[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_idom nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + have s7_3: "-.-.j = j" + using s3_2_asm1 neg_nat_in_int + minus_involutive by auto + show ?thesis + using s7_1 s7_2 s7_3 addint_succ_negnat + by auto + qed + also have s6_4: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "i \ Nat" + using s4_2_idom by auto + have s7_2: "j + i = i + j" + using s3_2_asm2 s7_1 spec by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]]" + proof - + have s6_1: "Succ[i] + -.Succ[-.j] = + iPred[Succ[i] + -.-.j]" + proof - + have s7_1: "Succ[i] \ Int" + using s4_2_idom succIsNat nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[Succ[i] + j]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_2 by auto + qed + also have s6_3: "... = iPred[j + Succ[i]]" + proof - + have s7_1: "Succ[i] \ Nat" + using s4_2_idom succIsNat by auto + have s7_2: "\ k \ Nat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + Succ[i] = Succ[i] + j" + using s7_1 s7_2 by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s7_2: "i \ Nat" + using s4_2_idom by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_5: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "j + i \ Int" + proof - + have s8_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s8_2: "i \ Int" + using s4_2_idom nat_in_int by auto + show ?thesis + using s8_1 s8_2 addIsInt by auto + qed + show ?thesis + using s7_1 iSucciPredCommute by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" + using s5_1 s5_2 by auto + qed + show "\ i \ Nat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + using s4_1 s4_2 natInduct by auto + qed + show ?thesis + using s3_1 s3_2 neg_nat_induction by auto + qed + have s2_2: "\ j \ negNat: \ i \ negNat: + j + i = i + j" + proof - + have s3_1: "\ i \ negNat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ negNat; + 0 + i = i + 0 + \ \ + 0 + -.Succ[-.i] = -.Succ[-.i] + 0" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ negNat" + assume s4_2_induct: "0 + i = i + 0" + have s5_1: "0 + -.Succ[-.i] = iPred[i]" + proof - + have s6_1: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" + proof - + have s7_1: "0 \ Int" + using zero_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[0 + i]" + proof - + have s7_1: "-.-.i = i" + using s4_2_idom neg_nat_in_int + minus_involutive by auto + show ?thesis + using s7_1 by auto + qed + also have s6_3: "... = iPred[i + 0]" + using s4_2_induct by auto + also have s6_4: "... = iPred[i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "i + 0 = i" + using s7_1 add_0 by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "-.Succ[-.i] + 0 = iPred[i]" + proof - + have s6_1: "-.Succ[-.i] + 0 = -.Succ[-.i]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + show ?thesis + using s7_1 add_0 by auto + qed + also have s6_2: "... = iPred[i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "i \ Nat \ {0}" + using s4_2_idom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s7_1 s7_2 by auto + qed + finally show ?thesis . + qed + show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" + using s5_1 s5_2 by auto + qed + show ?thesis + using s4_1 s4_2 neg_nat_induction by auto + qed + have s3_2: "\ j. \ + j \ negNat; + \ i \ negNat: j + i = i + j + \ \ + \ i \ negNat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + proof - + fix "j" :: "c" + assume s3_2_jdom: "j \ negNat" + assume s3_2_induct: "\ i \ negNat: + j + i = i + j" + have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] + 0 = iPred[j]" + proof - + have s6_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.j] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + show ?thesis + using s7_1 add_0 by auto + qed + also have s6_2: "... = iPred[j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "j \ Nat \ {0}" + using s3_2_jdom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s7_1 s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "0 + -.Succ[-.j] = iPred[j]" + proof - + have s6_1: "0 + -.Succ[-.j] = iPred[0 + -.-.j]" + proof - + have s7_1: "0 \ Int" + using zero_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[0 + j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_1 s7_2 by auto + qed + also have s6_3: "... = iPred[j + 0]" + proof - + have s7_1: "0 \ negNat" + unfolding negNat_def + using neg0 by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + 0 = 0 + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + show ?thesis + using s7_1 add_0 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s5_1 s5_2 by auto + qed + have s4_2: "\ i. \ + i \ negNat; + -.Succ[-.j] + i = i + -.Succ[-.j] + \ \ + -.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j]" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ negNat" + assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" + have s5_1: "-.Succ[-.j] + -.Succ[-.i] = + iPred[iPred[i + j]]" + proof - + have s6_1: "-.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.j] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s7_3: "-.Succ[-.j] \ Int + ==> + -.i \ Nat + ==> + -.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i]" + using addint_succ_negnat[ + of "-.Succ[-.j]" "-.i"] + by auto + show ?thesis + by (rule s7_3, rule s7_1, rule s7_2) + qed + also have s6_2: "... = iPred[-.Succ[-.j] + i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "-.-.i = i" + using s7_1 minus_involutive by auto + show ?thesis + using s7_1 s7_2 by auto + qed + also have s6_3: "... = iPred[i + -.Succ[-.j]]" + using s4_2_induct by auto + also have s6_4: "... = iPred[iPred[i + -.-.j]]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat by auto + have s7_3: "i + -.Succ[-.j] = iPred[i + -.-.j]" + using s7_1 s7_2 addint_succ_negnat by auto + show ?thesis + using s7_3 by auto + qed + also have s6_5: "... = iPred[iPred[i + j]]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "-.Succ[-.i] + -.Succ[-.j] = + iPred[iPred[i + j]]" + proof - + have s6_1: "-.Succ[-.i] + -.Succ[-.j] = + iPred[-.Succ[-.i] + -.-.j]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[-.Succ[-.i] + j]" + proof - + have s7_1: "-.-.j = j" + proof - + have s8_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_1 by auto + qed + also have s6_3: "... = iPred[j + -.Succ[-.i]]" + proof - + have s7_1: "-.Succ[-.i] \ negNat" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_neg_nat by auto + qed + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" + using s7_1 s7_2 by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[iPred[j + -.-.i]]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_5: "... = iPred[iPred[j + i]]" + proof - + have s7_1: "-.-.i = i" + proof - + have s8_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_1 by auto + qed + also have s6_6: "... = iPred[iPred[i + j]]" + proof - + have s7_1: "i \ negNat" + using s4_2_idom by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + i = i + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j]" + using s5_1 s5_2 by auto + qed + show "\ i \ negNat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + using s4_1 s4_2 neg_nat_induction by auto + qed + show ?thesis + using s3_1 s3_2 neg_nat_induction by auto + qed + show ?thesis + proof - + have s3_1: "\ j. \ i. + j \ negNat => ( + i \ Int => ( + j + i = i + j))" + proof auto + fix "j" :: "c" + fix "i" :: "c" + assume s4_1: "j \ negNat" + assume s4_2: "i \ Int" + have s4_3: "i \ Nat ==> j + i = i + j" + proof - + assume s5_1: "i \ Nat" + show "j + i = i + j" + using s4_1 s5_1 s2_1 by auto + qed + have s4_4: "i \ negNat ==> j + i = i + j" + proof - + assume s5_1: "i \ negNat" + show "j + i = i + j" + using s4_1 s5_1 s2_2 by auto + qed + show "j + i = i + j" + using s4_2 s4_3 s4_4 int_union_nat_negnat by auto + qed + show ?thesis + using s3_1 by auto + qed + qed + show ?thesis + using s1_1 s1_2 int_union_nat_negnat by auto + qed + + +theorem AddCommutative_sequent: + assumes mdom: "m \ Int" and ndom: "n \ Int" + shows "m + n = n + m" + using mdom ndom AddCommutative by auto + + +theorem add_0_left: + assumes mdom: "m \ Int" + shows "0 + m = m" + proof - + have s1_1: "m + 0 = m" + using mdom add_0 by auto + have s1_2: "m + 0 = 0 + m" + using mdom zero_in_int AddCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + + +(*----------------------------------------------------------------------------*) +(* Associativity of addition. *) + + +(* +THEOREM iSuccRightDistributesAdd == + ASSUME NEW a \in Int + PROVE \A b \in Int: iSucc[a + b] = a + iSucc[b] +<1>1. ASSUME NEW b \in Nat + PROVE iSucc[a + b] = a + iSucc[b] + <2>1. iSucc[a + b] = a + Succ[b] + <3>1. a \in Int + OBVIOUS + <3>2. b \in Nat + BY <1>1 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>2. @ = a + iSucc[b] + <3>1. b \in Nat + BY <1>1 + <3>2. b \in Int + BY <1>1, nat_in_int + <3>3. iSucc[b] = Succ[b] + BY <3>1, <3>2 DEF iSucc + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2 +<1>2. \A b \in negNat: iSucc[a + b] = a + iSucc[b] + <2>1. iSucc[a + 0] = a + iSucc[0] + <3>1. iSucc[a + 0] = a + Succ[0] + <4>1. a \in Int + OBVIOUS + <4>2. 0 \in Nat + BY zeroIsNat + <4> QED + BY <4>1, <4>2, addint_succ_nat + <3>2. @ = a + iSucc[0] + <4>1. 0 \in Nat + BY zeroIsNat + <4>2. 0 \in Int + BY zeroIsNat, nat_in_int + <4>3. iSucc[0] = Succ[0] + BY <4>1, <4>2 DEF iSucc + <4> QED + BY <4>3 + <3> QED + BY <3>1, <3>2 + <2>2. ASSUME NEW b \in negNat, + iSucc[a + b] = a + iSucc[b] + PROVE iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]] + <3>1. iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]] + <4>1. a \in Int + OBVIOUS + <4>2. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <4>3. a + -.Succ[-.b] = iPred[a + -.-.b] + BY <4>1, <4>2, addint_succ_negnat + <4>4. -.-.b = b + <5>1. b \in Int + BY <2>2, neg_nat_in_int + <5> QED + BY <5>1, minus_involutive + <4>5. a + -.Succ[-.b] = iPred[a + b] + BY <4>3, <4>4 + <4> QED + BY <4>5 + <3>2. @ = iPred[iSucc[a + b]] + <4>1. a + b \in Int + <5>1. a \in Int + OBVIOUS + <5>2. b \in Int + BY <2>2, neg_nat_in_int + <5> QED + BY <5>1, <5>2, addIsInt + <4> QED + BY <4>1, iSucciPredCommute + <3>3. CASE b = 0 + <4>1. iPred[iSucc[a + b]] = a + <5>1. iPred[iSucc[a + b]] = iPred[iSucc[a + 0]] + BY <3>3 + <5>2. @ = iPred[iSucc[a]] + <6>1. a + 0 = a + BY a \in Int, add_0 + <6> QED + BY <6>1 + <5>3. @ = a + <6>1. a \in Int + OBVIOUS + <6> QED + BY <6>1, iSucciPredCommute, spec + <5> QED + BY <5>1, <5>2, <5>3 + <4>2. a + iSucc[-.Succ[-.b]] = a + <5>1. a + iSucc[-.Succ[-.b]] = a + iSucc[-.Succ[0]] + BY <3>3 + <5>2. @ = a + iSucc[-.1] + OBVIOUS \* 1 is an abbreviation + <5>3. @ = a + -.Pred[-.-.1] + <6>1. -.1 \in Int + BY oneIsNat, minus_nat_in_int + <6>2. -.1 \notin Nat + <7>1. 1 \in Nat \ {0} + BY oneIsNat succNotZero + <7> QED + BY <7>1, minus_nat_0_in_negnat_0, + neg_nat_0_not_in_nat + <6>3. iSucc[-.1] = -.Pred[-.-.1] + BY <6>1, <6>2 DEF iSucc + <6> QED + BY <6>3 + <5>4. @ = a + -.Pred[1] + <6>1. -.-.1 = 1 + <7>1. 1 \in Int + BY oneIsNat nat_in_int + <7> QED + BY <7>1, minus_involutive + <6> QED + BY <6>1 + <5>5. @ = a + -.0 + BY pred_1 + <5>6. @ = a + 0 + BY neg0 + <5>7. @ = a + BY a \in Int, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, + <5>5, <5>6, <5>7 + <4> QED + BY <3>1, <3>2, <4>1, <4>2 + <3>4. CASE b # 0 + <4>1. iPred[iSucc[a + b]] = + iPred[a + iSucc[b]] + BY <2>2 + <4>2. @ = a + -.Succ[-.iSucc[b]] + <5>1. a \in Int + OBVIOUS + <5>2. /\ -.iSucc[b] \in Nat + /\ iSucc[b] \in negNat + <6>1. b \in negNat \ {0} + BY <2>2, <3>4 + <6>2. iSucc[b] = -.Pred[-.b] + <7>1. b \in Int + BY <6>1, neg_nat_in_int + <7>2. b \notin Nat + BY <6>1, neg_nat_0_not_in_nat + <7> QED + BY <7>1, <7>2 DEF iSucc + <6>3. Pred[-.b] \in Nat + <7>1. -.b \in Nat \ {0} + BY <6>1, minus_neg_nat_0_in_nat_0 + <7> QED + BY <7>1, Pred_in_nat + <6>4. -.iSucc[b] \in Nat + <7>1. -.iSucc[b] = Pred[-.b] + <8>1. -.iSucc[b] = -.-.Pred[-.b] + BY <6>2 + <8>2. Pred[-.b] \in Int + BY <6>3, nat_in_int + <8>3. -.-.Pred[-.b] = Pred[-.b] + BY <8>2, minus_involutive + <8> QED + BY <8>1, <8>3 + <7> QED + BY <6>3, <7>1 + <6>5. iSucc[b] \in negNat + BY <6>2, <6>3, minus_nat_in_neg_nat + <6> QED + BY <6>4, <6>5 + <5>3. a + -.Succ[-.iSucc[b]] = + iPred[a + -.-.iSucc[b]] + BY <5>1, <5>2, addint_succ_negnat + <5>4. -.-.iSucc[b] = iSucc[b] + <6>1. iSucc[b] \in Int + BY <5>2, neg_nat_in_int + <6> QED + BY <6>1, minus_involutive + <5> QED + BY <5>3, <5>4 + <4>3. @ = a + -.Succ[Pred[-.b]] + <5>1. b \in Int + BY <2>2, neg_nat_in_int + <5>2. b \notin Nat + <6>1. b \in negNat \ {0} + BY <2>2, <3>4 + <6> QED + BY <6>1, neg_nat_0_not_in_nat + <5>3. iSucc[b] = -.Pred[-.b] + BY <5>1, <5>2 DEF iSucc + <5>4. -.iSucc[b] = -.-.Pred[-.b] + BY <5>3 + <5>5. -.-.Pred[-.b] = Pred[-.b] + <6>1. -.b \in Nat + BY <2>2, <3>4, minus_neg_nat_0_in_nat_0 + <6>2. Pred[-.b] \in Nat + BY <6>1, Pred_in_nat + <6>3. Pred[-.b] \in Int + BY <6>2, nat_in_int + <6> QED + BY <6>3, minus_involutive + <5>6. -.iSucc[b] = Pred[-.b] + BY <5>4, <5>5 + <5> QED + BY <5>6 + <4>4. @ = a + -.Pred[Succ[-.b]] + <5>1. b \in negNat \ {0} + BY <2>2, <3>4 + <5>2. -.b \in Nat \ {0} + BY <5>1, minus_neg_nat_0_in_nat_0 + <5>3. Succ[Pred[-.b]] = Pred[Succ[-.b]] + BY <5>2, Succ_Pred_nat, Pred_Succ_nat + <5> QED + BY <5>3 + <4>5. @ = a + -.Pred[-.-.Succ[-.b]] + <5>1. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <5>2. Succ[-.b] \in Nat + BY <5>1, succIsNat + <5>3. Succ[-.b] \in Int + BY <5>2, nat_in_int + <5>4. -.-.Succ[-.b] = Succ[-.b] + BY <5>3, minus_involutive + <5> QED + BY <5>4 + <4>6. @ = a + iSucc[-.Succ[-.b]] + <5>1. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <5>2. Succ[-.b] \in Nat \ {0} + BY <5>1, succIsNat, succNotZero + <5>3. -.Succ[-.b] \in negNat \ {0} + BY <5>2, minus_nat_0_in_negnat_0 + <5>4. -.Succ[-.b] \in Int + BY <5>3, neg_nat_in_int + <5>5. -.Succ[-.b] \notin Nat + BY <5>3, neg_nat_0_not_in_nat + <5> QED + BY <5>4, <5>5 DEF iSucc + <4> QED + BY <3>1, <3>2, <4>1, <4>2, <4>3, + <4>4, <4>5, <4>6 + <3> QED + BY <3>3, <3>4 + <2> QED + BY <2>1, <2>2, neg_nat_induction +<1> QED + BY <1>1, <1>2 +*) +theorem iSuccRightDistributesAdd: + "\ a \ Int: \ b \ Int: + iSucc[a + b] = a + iSucc[b]" + proof - + { + fix "a" :: "c" + assume adom: "a \ Int" + have s1_1: "\ b \ Nat: iSucc[a + b] = a + iSucc[b]" + proof auto + fix "b" :: "c" + assume bdom: "b \ Nat" + have s2_1: "iSucc[a + b] = a + Succ[b]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b \ Nat" + using bdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + also have s2_2: "... = a + iSucc[b]" + proof - + have s3_1: "b \ Nat" + using bdom by auto + have s3_2: "b \ Int" + using bdom nat_in_int by auto + have s3_3: "iSucc[b] = Succ[b]" + unfolding iSucc_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "iSucc[a + b] = a + iSucc[b]" . + qed + have s1_2: "\ b \ negNat: iSucc[a + b] = a + iSucc[b]" + proof - + have s2_1: "iSucc[a + 0] = a + iSucc[0]" + proof - + have s3_1: "iSucc[a + 0] = a + Succ[0]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s4_1 s4_2 addint_succ_nat by auto + qed + have s3_2: "... = a + iSucc[0]" + proof - + have s4_1: "0 \ Nat" + using zeroIsNat by auto + have s4_2: "0 \ Int" + using zeroIsNat nat_in_int by auto + have s4_3: "iSucc[0] = Succ[0]" + unfolding iSucc_def + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + have s2_2: "\ b. \ + b \ negNat; + iSucc[a + b] = a + iSucc[b] + \ \ + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + fix "b" :: "c" + assume s2_2_bdom: "b \ negNat" + assume s2_2_induct: "iSucc[a + b] = a + iSucc[b]" + have s3_1: "iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s4_3: "a + -.Succ[-.b] = iPred[a + -.-.b]" + using s4_1 s4_2 addint_succ_negnat by auto + have s4_4: "-.-.b = b" + proof - + have s5_1: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + show ?thesis + using s5_1 minus_involutive by auto + qed + have s4_5: "a + -.Succ[-.b] = iPred[a + b]" + using s4_3 s4_4 by auto + show ?thesis + using s4_5 by auto + qed + have s3_2: "... = iPred[iSucc[a + b]]" + proof - + have s4_1: "a + b \ Int" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + show ?thesis + using s5_1 s5_2 addIsInt by auto + qed + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + have s3_3: "b = 0 ==> + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + assume s3_3_bdom: "b = 0" + have s4_1: "iPred[iSucc[a + b]] = a" + proof - + have s5_1: "iPred[iSucc[a + b]] = iPred[iSucc[a + 0]]" + using s3_3_bdom by auto + also have s5_2: "... = iPred[iSucc[a]]" + proof - + have s6_1: "a + 0 = a" + using adom add_0 by auto + show ?thesis + using s6_1 by auto + qed + also have s5_3: "... = a" + using adom iSucciPredCommute spec by auto + finally show ?thesis . + qed + have s4_2: "a + iSucc[-.Succ[-.b]] = a" + proof - + have s5_1: "a + iSucc[-.Succ[-.b]] = + a + iSucc[-.Succ[0]]" + using s3_3_bdom by auto + have s5_2: "... = a + iSucc[-.1]" + by auto + have s5_3: "... = a + -.Pred[-.-.1]" + (* `also` raises error: + empty result sequence + *) + proof - + have s6_1: "-.1 \ Int" + using oneIsNat minus_nat_in_int by auto + have s6_2: "-.1 \ Nat" + proof - + have s7_1: "1 \ Nat \ {0}" + using oneIsNat succNotZero by auto + show ?thesis + using s7_1 minus_nat_0_in_negnat_0 + neg_nat_0_not_in_nat by auto + qed + have s6_3: "iSucc[-.1] = -.Pred[-.-.1]" + unfolding iSucc_def + using s6_1 s6_2 by auto + show ?thesis + using s6_3 by auto + qed + have s5_4: "... = a + -.Pred[1]" + proof - + have s6_1: "-.-.1 = 1" + proof - + have s7_1: "1 \ Int" + using oneIsNat nat_in_int by auto + show ?thesis + using s7_1 minus_involutive by auto + qed + show ?thesis + using s6_1 by auto + qed + have s5_5: "... = a + -.0" + using pred_1 by auto + have s5_6: "... = a + 0" + using neg0 by auto + have s5_7: "... = a" + using adom add_0 by auto + show ?thesis + using s5_1 s5_2 s5_3 s5_4 s5_5 s5_6 s5_7 + by auto + qed + show ?thesis + using s3_1 s3_2 s4_1 s4_2 by auto + qed + have s3_4: "b \ 0 ==> + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + assume s3_4_bdom: "b \ 0" + have s4_1: "iPred[iSucc[a + b]] = + iPred[a + iSucc[b]]" + using s2_2_induct by auto + also have s4_2: "... = a + -.Succ[-.iSucc[b]]" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "-.iSucc[b] \ Nat \ + iSucc[b] \ negNat" + proof - + have s6_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + have s6_2: "iSucc[b] = -.Pred[-.b]" + proof - + have s7_1: "b \ Int" + using s6_1 neg_nat_in_int by auto + have s7_2: "b \ Nat" + using s6_1 neg_nat_0_not_in_nat by auto + show ?thesis + unfolding iSucc_def + using s7_1 s7_2 by auto + qed + have s6_3: "Pred[-.b] \ Nat" + proof - + have s7_1: "-.b \ Nat \ {0}" + using s6_1 minus_neg_nat_0_in_nat_0 by auto + show ?thesis + using s7_1 Pred_in_nat by auto + qed + have s6_4: "-.iSucc[b] \ Nat" + proof - + have s7_1: "-.iSucc[b] = Pred[-.b]" + proof - + have s8_1: "-.iSucc[b] = -.-.Pred[-.b]" + using s6_2 by auto + have s8_2: "Pred[-.b] \ Int" + using s6_3 nat_in_int by auto + have s8_3: "-.-.Pred[-.b] = Pred[-.b]" + using s8_2 minus_involutive by auto + show ?thesis + using s8_1 s8_3 by auto + qed + show ?thesis + using s6_3 s7_1 by auto + qed + have s6_5: "iSucc[b] \ negNat" + using s6_2 s6_3 minus_nat_in_neg_nat by auto + show ?thesis + using s6_4 s6_5 by auto + qed + have s5_3: "a + -.Succ[-.iSucc[b]] = + iPred[a + -.-.iSucc[b]]" + using s5_1 s5_2 addint_succ_negnat by auto + have s5_4: "-.-.iSucc[b] = iSucc[b]" + proof - + have s6_1: "iSucc[b] \ Int" + using s5_2 neg_nat_in_int by auto + show ?thesis + using s6_1 minus_involutive by auto + qed + show ?thesis + using s5_3 s5_4 by auto + qed + also have s4_3: "... = a + -.Succ[Pred[-.b]]" + proof - + have s5_1: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + have s5_2: "b \ Nat" + proof - + have s6_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + show ?thesis + using s6_1 neg_nat_0_not_in_nat by auto + qed + have s5_3: "iSucc[b] = -.Pred[-.b]" + unfolding iSucc_def + using s5_1 s5_2 by auto + have s5_4: "-.iSucc[b] = -.-.Pred[-.b]" + using s5_3 by auto + have s5_5: "-.-.Pred[-.b] = Pred[-.b]" + proof - + have s6_1: "-.b \ Nat" + using s2_2_bdom s3_4_bdom + minus_neg_nat_0_in_nat_0 + by auto + have s6_2: "Pred[-.b] \ Nat" + using s6_1 Pred_in_nat by auto + have s6_3: "Pred[-.b] \ Int" + using s6_2 nat_in_int by auto + show ?thesis + using s6_3 minus_involutive by auto + qed + have s5_6: "-.iSucc[b] = Pred[-.b]" + using s5_4 s5_5 by auto + show ?thesis + using s5_6 by auto + qed + also have s4_4: "... = a + -.Pred[Succ[-.b]]" + proof - + have s5_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + have s5_2: "-.b \ Nat \ {0}" + using s5_1 minus_neg_nat_0_in_nat_0 by auto + have s5_3: "Succ[Pred[-.b]] = Pred[Succ[-.b]]" + using s5_2 Succ_Pred_nat Pred_Succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = a + -.Pred[-.-.Succ[-.b]]" + proof - + have s5_1: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s5_2: "Succ[-.b] \ Nat" + using s5_1 succIsNat by auto + have s5_3: "Succ[-.b] \ Int" + using s5_2 nat_in_int by auto + have s5_4: "-.-.Succ[-.b] = Succ[-.b]" + using s5_3 minus_involutive by auto + show ?thesis + using s5_4 by auto + qed + also have s4_6: "... = a + iSucc[-.Succ[-.b]]" + proof - + have s5_1: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s5_2: "Succ[-.b] \ Nat \ {0}" + using s5_1 succIsNat succNotZero by auto + have s5_3: "-.Succ[-.b] \ negNat \ {0}" + using s5_2 minus_nat_0_in_negnat_0 by auto + have s5_4: "-.Succ[-.b] \ Int" + using s5_3 neg_nat_in_int by auto + have s5_5: "-.Succ[-.b] \ Nat" + using s5_3 neg_nat_0_not_in_nat by auto + show ?thesis + unfolding iSucc_def + using s5_4 s5_5 by auto + qed + finally show "iSucc[a + -.Succ[-.b]] = + a + iSucc[-.Succ[-.b]]" + using s3_1 s3_2 by auto + qed + show "iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + using s3_3 s3_4 by blast + qed + show ?thesis + using s2_1 s2_2 neg_nat_induction by auto + qed + have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" + using s1_1 s1_2 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem iSuccRightDistributesAdd_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "iSucc[a + b] = a + iSucc[b]" + using adom bdom iSuccRightDistributesAdd by auto + + +(* +THEOREM iPredRightDistributesAdd == + ASSUME NEW a \in Int + PROVE \A b \in Int: iPred[a + b] = a + iPred[b] +PROOF +<1>1. iPred[a + 0] = a + iPred[0] + <2>1. iPred[a + 0] = iPred[a + -.0] + BY neg0 + <2>2. @ = a + -.Succ[0] + <3>1. a \in Int + OBVIOUS + <3>2. 0 \in Nat + BY zeroIsNat + <3> QED + BY <3>1, <3>2, addint_succ_negnat + <2>3. @ = a + -.1 + OBVIOUS \* 1 is an abbreviation + <2>4. @ = a + iPred[0] + BY ipred_0 + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>2. \A b \in Nat: iPred[a + b] = a + iPred[b] + <2>1. ASSUME NEW b \in Nat, + iPred[a + b] = a + iPred[b] + PROVE iPred[a + Succ[b]] = a + iPred[Succ[b]] + <3>1. iPred[a + Succ[b]] = iPred[iSucc[a + b]] + <4>1. a \in Int + OBVIOUS + <4>2. b \in Nat + BY <2>1 + <4>3. a + Succ[b] = iSucc[a + b] + BY <4>1, <4>2, addint_succ_nat + <4> QED + BY <4>3 + <3>2. @ = iSucc[iPred[a + b]] + <4>1. a + b \in Int + <5>1. a \in Int + OBVIOUS + <5>2. b \in Int + BY <2>2, nat_in_int + <5> QED + BY <5>1, <5>2, addIsInt + <4> QED + BY <4>1, iSucciPredCommute + <3>3. @ = iSucc[a + iPred[b]] + BY <2>1 + <3>4. @ = a + iSucc[iPred[b]] + <4>1. b \in Int + BY <2>1, nat_in_int + <4>2. iPred[b] \in Int + BY <4>1, iPred_type + <4>3. a \in Int + OBVIOUS + <4> QED + BY <4>2, <4>3, iSuccRightDistributesAdd + <3>5. @ = a + iPred[iSucc[b]] + <4>1. b \in Int + BY <2>1, nat_in_int + <4> QED + BY <4>1, iSucciPredCommute + <3>6. @ = a + iPred[Succ[b]] + <4>1. b \in Nat + BY <2>1 + <4>2. b \in Int + BY <2>1, nat_in_int + <4>3. iSucc[b] = Succ[b] + BY <4>1, <4>2, nat_in_int DEF iSucc + <4> QED + BY <4>3 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 + <2> QED + BY <2>1, <1>1, NatInduction +<1>3. \A b \in negNat \ {0}: iPred[a + b] = a + iPred[b] + <2>1. ASSUME NEW b \in negNat \ {0} + PROVE iPred[a + b] = a + iPred[b] + <3>1. iPred[a + b] = iPred[a + -.-.b] + <4>1. b \in Int + BY <2>1, neg_nat_in_int + <4>2. -.-.b = b + BY <4>1, minus_involutive + <4> QED + BY <4>2 + <3>2. @ = a + -.Succ[-.b] + <4>1. a \in Int + OBVIOUS + <4>2. -.b \in Nat + BY <2>1, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, addint_succ_negnat + <3>3. @ = a + iPred[b] + <4>1. b \in Int + BY <2>1, neg_nat_in_int + <4>2. b \notin Nat \ {0} + BY <2>1, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3> QED + BY <3>1, <3>2, <3>3 + <2> QED + BY <2>1 +<1> QED + BY <1>2, <1>3 +*) +theorem iPredRightDistributesAdd: + "\ a \ Int: \ b \ Int: + iPred[a + b] = a + iPred[b]" + proof - + { + fix "a" :: "c" + assume adom: "a \ Int" + have s1_1: "iPred[a + 0] = a + iPred[0]" + proof - + have s2_1: "iPred[a + 0] = iPred[a + -.0]" + using neg0 by auto + also have s2_2: "... = a + -.Succ[0]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s3_1 s3_2 addint_succ_negnat by auto + qed + also have s2_3: "... = a + -.1" + by auto + have s2_4: "... = a + iPred[0]" + using ipred_0 by auto + (* `also` raises an error here *) + finally show ?thesis + using s2_4 by auto + qed + have s1_2: "\ b \ Nat: + iPred[a + b] = a + iPred[b]" + proof - + have s2_1: "\ b. \ + b \ Nat; + iPred[a + b] = a + iPred[b] + \ \ + iPred[a + Succ[b]] = a + iPred[Succ[b]]" + proof - + fix "b" :: "c" + assume s2_1_bdom: "b \ Nat" + assume s2_1_induct: "iPred[a + b] = a + iPred[b]" + have s3_1: "iPred[a + Succ[b]] = iPred[iSucc[a + b]]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "b \ Nat" + using s2_1_bdom by auto + have s4_3: "a + Succ[b] = iSucc[a + b]" + using s4_1 s4_2 addint_succ_nat by auto + show ?thesis + using s4_3 by auto + qed + also have s3_2: "... = iSucc[iPred[a + b]]" + proof - + have s4_1: "a + b \ Int" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "b \ Int" + using s2_1_bdom nat_in_int by auto + show ?thesis + using s5_1 s5_2 addIsInt by auto + qed + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + also have s3_3: "... = iSucc[a + iPred[b]]" + using s2_1_induct by auto + also have s3_4: "... = a + iSucc[iPred[b]]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom nat_in_int by auto + have s4_2: "iPred[b] \ Int" + using s4_1 iPred_type by auto + have s4_3: "a \ Int" + using adom by auto + show ?thesis + using s4_2 s4_3 iSuccRightDistributesAdd + by auto + qed + also have s3_5: "... = a + iPred[iSucc[b]]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom nat_in_int by auto + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + also have s3_6: "... = a + iPred[Succ[b]]" + proof - + have s4_1: "b \ Nat" + using s2_1_bdom by auto + have s4_2: "b \ Int" + using s2_1_bdom nat_in_int by auto + have s4_3: "iSucc[b] = Succ[b]" + unfolding iSucc_def + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . + qed + show ?thesis + using s1_1 s2_1 natInduct by auto + qed + have s1_3: "\ b \ negNat \ {0}: + iPred[a + b] = a + iPred[b]" + proof - + { + fix "b" :: "c" + assume s2_1_bdom: "b \ negNat \ {0}" + { + have s3_1: "iPred[a + b] = iPred[a + -.-.b]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom neg_nat_in_int by auto + have s4_2: "-.-.b = b" + using s4_1 minus_involutive by auto + show ?thesis + using s4_2 by auto + qed + also have s3_2: "... = a + -.Succ[-.b]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "-.b \ Nat" + using s2_1_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 addint_succ_negnat by auto + qed + also have s3_3: "... = a + iPred[b]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom neg_nat_in_int by auto + have s4_2: "b \ Nat \ {0}" + proof - + have s5_1: "b \ negNat" + using s2_1_bdom by auto + show ?thesis + using s5_1 neg_nat_not_in_nat_setminus_0 + by auto + qed + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + finally have "iPred[a + b] = a + iPred[b]" . + } + from this have s2_2: "iPred[a + b] = a + iPred[b]" + by auto + } + from this show ?thesis + by auto + qed + have s1_4: "\ b \ Int: + iPred[a + b] = a + iPred[b]" + using s1_2 s1_3 int_union_nat_negnat_0 spec allI by auto + } + from this show ?thesis + by auto + qed + + +theorem iPredRightDistributesAdd_sequent: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "iPred[m + n] = m + iPred[n]" + using mdom ndom iPredRightDistributesAdd by auto + + +theorem isucc_eq_add_1: + assumes idom: "i \\in Int" + shows "iSucc[i] = i + 1" + proof - + have s1_1: "iSucc[i] = iSucc[i + 0]" + using idom add_0 by auto + also have s1_2: "... = i + iSucc[0]" + proof - + have s2_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + show ?thesis + using idom s2_1 iSuccRightDistributesAdd_sequent + by auto + qed + also have s1_3: "... = i + 1" + proof - + have s2_1: "iSucc[0] = Succ[0]" + proof - + have s3_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s3_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_2: "Succ[0] = 1" + by auto + have s2_3: "iSucc[0] = 1" + using s2_1 s2_2 by auto + show ?thesis + using s2_3 s1_2 by auto + qed + finally show ?thesis . + qed + + +(* +THEOREM AddAssociative == + \A a \in Int: \A b \in Int: \A c \in Int: + (a + b) + c = a + (b + c) +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE \A c \in Int: + (a + b) + c = a + (b + c) + BY <1>1 +<1>2. (a + b) + 0 = a + (b + 0) + <2>1. (a + b) + 0 = a + b + <3>1. a + b \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, add_0 + <2>2. @ = a + (b + 0) + <3>1. b \in Int + BY <1>1 + <3>2. b + 0 = b + BY <3>1, add_0 + <3> QED + BY <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW c \in Nat, + (a + b) + c = a + (b + c) + PROVE (a + b) + Succ[c] = a + (b + Succ[c]) + <2>1. (a + b) + Succ[c] = iSucc[(a + b) + c] + <3>1. a + b \in Int + BY <1>1, addIsInt + <3>2. c \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>2. @ = iSucc[a + (b + c)] + BY <1>3 \* induction hypothesis + <2>3. @ = a + iSucc[b + c] + <3>1. a \in Int + BY <1>1 + <3>2. b + c \in Int + BY <1>1, <1>3, nat_in_int, addIsInt + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>4. @ = a + (b + iSucc[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>3, nat_in_int + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>5. @ = a + (b + Succ[c]) + <3>1. c \in Nat + BY <1>3 + <3>2. c \in Int + BY <3>1, nat_in_int + <3>3. iSucc[c] = Succ[c] + BY <3>1, <3>2 DEF iSucc + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5 +<1>4. ASSUME NEW c \in negNat, + (a + b) + c = a + (b + c) + PROVE (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c]) + <2>1. (a + b) + -.Succ[-.c] = (a + b) + iPred[c] + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. c \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <3>3. iPred[c] = -.Succ[-.c] + BY <3>1, <3>2 DEF iPred + <3> QED + BY <3>3 + <2>2. @ = iPred[(a + b) + c] + <3>1. a + b \in Int + BY <1>1, addIsInt + <3>2. c \in Int + BY <1>4, neg_nat_in_int + <3>3 QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>3. @ = iPred[a + (b + c)] + BY <1>4 \* induction hypothesis + <2>4. @ = a + iPred[b + c] + <3>1. a \in Int + BY <1>1 + <3>2. b + c \in Int + BY <1>1, <1>4, neg_nat_in_int, addIsInt + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>5. @ = a + (b + iPred[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>4, neg_nat_in_int + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>6. @ = a + (b + -.Succ[-.c]) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. c \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <3>3. iPred[c] = -.Succ[-.c] + BY <3>1, <3>2 DEF iPred + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 +<1>5. \A c \in Nat: (a + b) + c = a + (b + c) + BY <1>2, <1>3, NatInduction +<1>6. \A c \in negNat: (a + b) + c = a + (b + c) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6 +*) +theorem AddAssociative: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + (a + b) + c = a + (b + c)" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \ Int" + assume bdom: "b \ Int" + have s1_2: "(a + b) + 0 = a + (b + 0)" + proof - + have s2_1: "(a + b) + 0 = a + b" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + show ?thesis + using s3_1 add_0 by auto + qed + have s2_2: "... = a + (b + 0)" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "b + 0 = b" + using s3_1 add_0 by auto + show ?thesis + using s3_2 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ c. \ + c \ Nat; + (a + b) + c = a + (b + c) + \ \ + (a + b) + Succ[c] = a + (b + Succ[c])" + proof - + fix "c" :: "c" + assume s1_3_cdom: "c \ Nat" + assume s1_3_induct: "(a + b) + c = a + (b + c)" + have s2_1: "(a + b) + Succ[c] = iSucc[(a + b) + c]" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + have s3_2: "c \ Nat" + using s1_3_cdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + also have s2_2: "... = iSucc[a + (b + c)]" + using s1_3_induct by auto + also have s2_3: "... = a + iSucc[b + c]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b + c \ Int" + using bdom s1_3_cdom nat_in_int addIsInt by auto + show ?thesis + using s3_1 s3_2 iSuccRightDistributesAdd by auto + qed + also have s2_4: "... = a + (b + iSucc[c])" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "c \ Int" + using s1_3_cdom nat_in_int by auto + have s3_3: "iSucc[b + c] = b + iSucc[c]" + using s3_1 s3_2 iSuccRightDistributesAdd by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = a + (b + Succ[c])" + proof - + have s3_1: "c \ Nat" + using s1_3_cdom by auto + have s3_2: "c \ Int" + using s1_3_cdom nat_in_int by auto + have s3_3: "iSucc[c] = Succ[c]" + unfolding iSucc_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a + b) + Succ[c] = a + (b + Succ[c])" . + qed + have s1_4: "\ c. \ + c \ negNat; + (a + b) + c = a + (b + c) + \ \ + (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" + proof - + fix "c" :: "c" + assume s1_4_cdom: "c \ negNat" + assume s1_4_induct: "(a + b) + c = a + (b + c)" + have s2_1: "(a + b) + -.Succ[-.c] = (a + b) + iPred[c]" + proof - + have s3_1: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "c \ Nat \ {0}" + using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto + have s3_3: "iPred[c] = -.Succ[-.c]" + unfolding iPred_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + also have s2_2: "... = iPred[(a + b) + c]" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + have s3_2: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by force + qed + also have s2_3: "... = iPred[a + (b + c)]" + using s1_4_induct by auto + also have s2_4: "... = a + iPred[b + c]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b + c \ Int" + using bdom s1_4_cdom neg_nat_in_int addIsInt by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by auto + qed + also have s2_5: "... = a + (b + iPred[c])" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_3: "iPred[b + c] = b + iPred[c]" + using s3_1 s3_2 iPredRightDistributesAdd by force + show ?thesis + using s3_3 by auto + qed + also have s2_6: "... = a + (b + -.Succ[-.c])" + proof - + have s3_1: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "c \ Nat \ {0}" + using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto + have s3_3: "iPred[c] = -.Succ[-.c]" + unfolding iPred_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" . + qed + have s1_5: "\ c \ Nat: + (a + b) + c = a + (b + c)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ c \ negNat: + (a + b) + c = a + (b + c)" + using s1_2 s1_4 neg_nat_induction by auto + have "\ c \ Int: + (a + b) + c = a + (b + c)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem AddAssociative_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + and cdom: "c \\in Int" + shows "(a + b) + c = a + (b + c)" + using adom bdom cdom AddAssociative by auto + + +theorem AddLeftCommutativity: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + and cdom: "c \\in Int" + shows "b + (a + c) = a + (b + c)" + proof - + have s1_1: "(b + a) + c = (a + b) + c" + using adom bdom by (simp only: AddCommutative_sequent) + show ?thesis + using adom bdom cdom s1_1 by (simp only: AddAssociative_sequent) + qed + + +theorems add_ac_int[algebra_simps] = + AddCommutative_sequent AddAssociative_sequent + AddLeftCommutativity add_0 add_0_left + + +(* A test that the simplification rules work. *) +theorem + assumes "a \\in Int" and "b \\in Int" and "c \\in Int" + shows "a + (b + c) = (a + b) + c" + using assms by (simp add: algebra_simps) + + +(* +THEOREM MinusSuccIsMinusOne == + ASSUME NEW n \in Nat + PROVE -.Succ[n] = -.n + -.1 +<1>1. -.Succ[n] = -.Succ[-.-.n] + <2>1. -.-.n = n + <3>1. n \in Nat + OBVIOUS + <3>2. n \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>2, minus_involutive + <2> QED + BY <2>1 +<1>2. @ = iPred[-.n] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. -.n \notin Nat \ {0} + <3>1. n \in Nat + OBVIOUS + <3>2. -.n \in negNat + BY <3>1, minus_nat_in_neg_nat + <3> QED + BY <3>2, neg_nat_not_in_nat_setminus_0 + <2> QED + BY <2>1, <2>2 DEF iPred +<1>3. @ = iPred[-.n + 0] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. -.n + 0 = -.n + BY <2>1, add_0 + <2> QED + BY <2>2 +<1>4. @ = -.n + iPred[0] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. 0 \in Int + BY zero_in_int + <2> QED + BY <2>1, <2>2, iPredRightDistributesAdd +<1>5. @ = -.n + -.Succ[-.0] + <2>1. 0 \in Int + BY zero_in_int + <2>2. 0 \notin Nat \ {0} + OBVIOUS + <2> QED + BY <2>1, <2>2 DEF iPred +<1>6. @ = -.n + -.Succ[0] + BY neg0 +<1>7. @ = -.n + -.1 + OBVIOUS \* 1 is an abbreviation +<1> QED + BY <1>1, <1>2, <1>3, <1>4, + <1>5, <1>6, <1>7 +*) +theorem MinusSuccIsMinusOne: + assumes ndom: "n \ Nat" + shows "-.Succ[n] = -.n + -.1" + proof - + have s1_1: "-.Succ[n] = -.Succ[-.-.n]" + proof - + have s2_1: "-.-.n = n" + proof - + have s3_1: "n \ Int" + using ndom nat_in_int by auto + show ?thesis + using s3_1 minus_involutive by auto + qed + show ?thesis + using s2_1 by auto + qed + also have s1_2: "... = iPred[-.n]" + proof - + have s2_1: "-.n \ Int" + using ndom minus_nat_in_int by auto + have s2_2: "-.n \ Nat \ {0}" + proof - + have s3_1: "-.n \ negNat" + using ndom minus_nat_in_neg_nat by auto + show ?thesis + using s3_1 neg_nat_not_in_nat_setminus_0 by fast + qed + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_3: "... = iPred[-.n + 0]" + proof - + have s2_1: "-.n \ Int" + proof - + have s3_1: "n \ Nat" + using ndom by auto + show ?thesis + using s3_1 minus_nat_in_int by auto + qed + have s2_2: "-.n + 0 = -.n" + using s2_1 add_0 by auto + show ?thesis + using s2_2 by auto + qed + also have s1_4: "... = -.n + iPred[0]" + proof - + have s2_1: "-.n \ Int" + proof - + have s3_1: "n \ Nat" + using ndom by auto + show ?thesis + using s3_1 minus_nat_in_int by auto + qed + have s2_2: "0 \ Int" + using zero_in_int by auto + show ?thesis + using s2_1 s2_2 iPredRightDistributesAdd by fast + qed + also have s1_5: "... = -.n + -.Succ[-.0]" + proof - + have s2_1: "0 \ Int" + using zero_in_int by auto + have s2_2: "0 \ Nat \ {0}" + by auto + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_6: "... = -.n + -.Succ[0]" + using neg0 by auto + also have s1_7: "... = -.n + -.1" + by auto + from calculation show ?thesis + by auto + qed + + +theorem SuccMinusIsPlusOne: + assumes idom: "i \\in negNat" + shows "Succ[-.i] = -.i + 1" + proof - + have s1_1: "-.i \\in Nat" + using idom minus_neg_nat_in_nat by auto + have s1_2: "-.i \\in Int" + using s1_1 nat_in_int by auto + have s1_3: "Succ[-.i] = iSucc[-.i]" + unfolding iSucc_def + using s1_1 s1_2 by auto + also have s1_4: "... = iSucc[-.i + 0]" + using s1_2 add_0 by auto + also have s1_5: "... = -.i + Succ[0]" + using s1_2 zeroIsNat addint_succ_nat by auto + also have s1_6: "... = -.i + 1" + by auto (* 1 is an abbreviation *) + from calculation show ?thesis . + qed + + +(*----------------------------------------------------------------------------*) +(* Properties of addition and difference. *) + + +(* +THEOREM AddNegCancels == + \A i \in Int: i + -.i = 0 +PROOF +<1>1. 0 + -.0 = 0 + <2>1. 0 + -.0 = 0 + 0 + BY neg0 + <2>2. @ = 0 + BY zero_in_int, add_0 + <2> QED + BY <2>1, <2>2 +<1>2. ASSUME NEW i \in Nat, i + -.i = 0 + PROVE Succ[i] + -.Succ[i] = 0 + <2>1. Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i] + <3>1. Succ[i] \in Int + BY <1>2, succIsNat, nat_in_int + <3>2. -.Succ[i] \in Int + BY <3>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, AddCommutative + <2>2. @ = iSucc[-.Succ[i] + i] + <3>1. -.Succ[i] \in Int + BY <1>2, succIsNat, minus_nat_in_int + <3>2. i \in Nat + BY <1>2 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>3. @ = iSucc[i + -.Succ[i]] + <3>1. i \in Int + BY <1>2, nat_in_int + <3>2. -.Succ[i] \in Int + BY <1>2, succIsNat, minus_nat_in_int + <3>3. -.Succ[i] + i = i + -.Succ[i] + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>4. @ = iSucc[iPred[i + -.i]] + <3>1. i \in Int + BY <1>2, nat_in_int + <3>2. i \in Nat + BY <1>2 + <3> QED + BY <3>1, <3>2, addint_succ_negnat + <2>5. @ = iSucc[iPred[0]] + BY <1>2 + <2>6. @ = iSucc[-.1] + BY ipred_0 + <2>7. @ = 0 + BY isucc_minus_1 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6, <2>7 +<1>3. ASSUME NEW i \in negNat, i + -.i = 0 + PROVE -.Succ[-.i] + -.(-.Succ[-.i]) = 0 + <2>1. -.Succ[-.i] + -.(-.Succ[-.i]) + = -.Succ[-.i] + Succ[-.i] + <3>1. Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, nat_in_int + <3> QED + BY <3>1, minus_involutive + <2>2. @ = iSucc[-.Succ[-.i] + -.i] + <3>1. -.Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, minus_nat_in_int + <3>2. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>3. @ = iSucc[-.i + -.Succ[-.i]] + <3>1. -.Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, minus_nat_in_int + <3>2. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>3. -.Succ[-.i] + -.i = -.i + -.Succ[-.i] + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>4. @ = iSucc[iPred[-.i + -.-.i]] + <3>1. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>2. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3>3. -.i + -.Succ[-.i] = iPred[-.i + -.-.i] + BY <3>1, <3>2, addint_succ_negnat + <3> QED + BY <3>3 + <2>5. @ = iSucc[iPred[-.i + i]] + <3>1. -.-.i = i + BY <1>3, neg_nat_in_int, minus_involutive + <3> QED + BY <3>1 + <2>6. @ = iSucc[iPred[i + -.i]] + <3>1. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>2. i \in Int + BY <1>3, neg_nat_in_int + <3>3. -.i + i = i + -.i + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>7. @ = iSucc[iPred[0]] + BY <1>3 \* induction hypothesis + <2>8. @ = iSucc[-.1] + BY ipred_0 + <2>9. @ = 0 + BY isucc_minus_1 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, + <2>6, <2>7, <2>8, <2>9 +<1>4. \A i \in Nat: i + -.i = 0 + BY <1>1, <1>2, NatInduction +<1>5. \A i \in negNat: i + -.i = 0 + BY <1>1, <1>3, neg_nat_induction +<1> QED + BY <1>4, <1>5, int_union_nat_negnat +*) +theorem AddNegCancels: + "\ i \ Int: i + -.i = 0" + proof - + have s1_1: "0 + -.0 = 0" + proof - + have s2_1: "0 + -.0 = 0 + 0" + using neg0 by auto + also have s2_2: "... = 0" + using zero_in_int add_0 by auto + finally show ?thesis . + qed + have s1_2: "\ i. \ + i \\in Nat; + i + -.i = 0 + \ \ + Succ[i] + -.Succ[i] = 0" + proof - + fix "i" :: "c" + assume s1_2_idom: "i \\in Nat" + assume s1_2_induct: "i + -.i = 0" + have s2_1: "Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i]" + proof - + have s3_1: "Succ[i] \\in Int" + using s1_2_idom succIsNat nat_in_int by auto + have s3_2: "-.Succ[i] \\in Int" + using s3_1 neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 AddCommutative by fast + qed + also have s2_2: "... = iSucc[-.Succ[i] + i]" + proof - + have s3_1: "-.Succ[i] \\in Int" + using s1_2_idom succIsNat minus_nat_in_int by auto + have s3_2: "i \\in Nat" + using s1_2_idom by auto + show ?thesis + using s3_1 s3_2 by (rule addint_succ_nat) + qed + also have s2_3: "... = iSucc[i + -.Succ[i]]" + proof - + have s3_1: "i \\in Int" + using s1_2_idom nat_in_int by auto + have s3_2: "-.Succ[i] \\in Int" + using s1_2_idom succIsNat minus_nat_in_int + by auto + have s3_3: "-.Succ[i] + i = i + -.Succ[i]" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = iSucc[iPred[i + -.i]]" + proof - + have s3_1: "i \\in Int" + using s1_2_idom nat_in_int by auto + have s3_2: "i \\in Nat" + using s1_2_idom by auto + show ?thesis + using s3_1 s3_2 addint_succ_negnat by auto + qed + also have s2_5: "... = iSucc[iPred[0]]" + using s1_2_induct by auto + also have s2_6: "... = iSucc[-.1]" + using ipred_0 by auto + also have s2_7: "... = 0" + using isucc_minus_1 by auto + finally show "Succ[i] + -.Succ[i] = 0" . + qed + have s1_3: "\ i. \ + i \\in negNat; + i + -.i = 0 + \ \ + -.Succ[-.i] + -.-.Succ[-.i] = 0" + proof - + fix "i" :: "c" + assume s1_3_idom: "i \\in negNat" + assume s1_3_induct: "i + -.i = 0" + have s2_1: "-.Succ[-.i] + -.(-.Succ[-.i]) + = -.Succ[-.i] + Succ[-.i]" + proof - + have s3_1: "Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat nat_in_int by auto + show ?thesis + using s3_1 minus_involutive by auto + qed + also have s2_2: "... = iSucc[-.Succ[-.i] + -.i]" + proof - + have s3_1: "-.Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat minus_nat_in_int by auto + have s3_2: "-.i \\in Nat" + using s1_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 by (rule addint_succ_nat) + qed + also have s2_3: "... = iSucc[-.i + -.Succ[-.i]]" + proof - + have s3_1: "-.Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat minus_nat_in_int by auto + have s3_2: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_3: "-.Succ[-.i] + -.i = -.i + -.Succ[-.i]" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = iSucc[iPred[-.i + -.-.i]]" + proof - + have s3_1: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_2: "-.i \\in Nat" + using s1_3_idom minus_neg_nat_in_nat by auto + have s3_3: "-.i + -.Succ[-.i] = iPred[-.i + -.-.i]" + using s3_1 s3_2 addint_succ_negnat by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = iSucc[iPred[-.i + i]]" + proof - + have s3_1: "-.-.i = i" + using s1_3_idom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s3_1 by auto + qed + also have s2_6: "... = iSucc[iPred[i + -.i]]" + proof - + have s3_1: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_2: "i \\in Int" + using s1_3_idom neg_nat_in_int by auto + have s3_3: "-.i + i = i + -.i" + using s3_1 s3_2 AddCommutative by fast + show ?thesis + using s3_3 by auto + qed + also have s2_7: "... = iSucc[iPred[0]]" + using s1_3_induct by auto + also have s2_8: "... = iSucc[-.1]" + using ipred_0 by auto + also have s2_9: "... = 0" + using isucc_minus_1 by auto + finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . + qed + have s1_4: "\ i \ Nat: i + -.i = 0" + using s1_1 s1_2 natInduct by auto + have s1_5: "\ i \ negNat: i + -.i = 0" + using s1_1 s1_3 neg_nat_induction by auto + show ?thesis + using s1_4 s1_5 int_union_nat_negnat by auto + qed + + +theorem AddNegCancels_sequent: + assumes idom: "i \\in Int" + shows "i + -.i = 0" + using idom AddNegCancels spec by auto + + +theorem AddNegCancels_left: + assumes idom: "i \\in Int" + shows "-.i + i = 0" + proof - + have s1_1: "i + -.i = 0" + using idom AddNegCancels_sequent by auto + have s1_2: "i + -.i = -.i + i" + proof - + have s2_1: "-.i \\in Int" + using idom neg_int_eq_int by auto + show ?thesis + using idom s2_1 AddCommutative_sequent by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM a_plus_b_minus_a_eq_b == + ASSUME NEW a \in Int, NEW b \in Int + PROVE a + (b - a) = b +PROOF +<1>1. a + (b - a) = a + (b + -.a) + BY DEF diff +<1>2. @ = a + (-.a + b) + <2>1. b \in Int + OBVIOUS + <2>2. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>3. b + -.a = -.a + b + BY <2>1, <2>2, AddCommutative + <2> QED + BY <2>3 +<1>3. @ = (a + -.a) + b + <2>1. a \in Int + OBVIOUS + <2>2. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>3. b \in Int + OBVIOUS + <2> QED + BY <2>1, <2>2, <2>3, AddAssociative +<1>4. @ = 0 + b + <2>1. a \in Int + OBVIOUS + <2>2. a + -.a = 0 + BY <2>1, AddNegCancels + <2> QED + BY <2>2 +<1>5. @ = b + BY add_0_left +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5 +*) +theorem a_plus_b_minus_a_eq_b: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "a + (b - a) = b" + proof - + have s1_1: "a + (b - a) = a + (b + -.a)" + unfolding diff_def by auto + also have s1_2: "... = a + (-.a + b)" + proof - + have s2_1: "b \\in Int" + using bdom by auto + have s2_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_3: "b + -.a = -.a + b" + using s2_1 s2_2 AddCommutative by auto + show ?thesis + using s2_3 by auto + qed + also have s1_3: "... = (a + -.a) + b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_3: "b \\in Int" + using bdom by auto + have s2_4: "a \\in Int ==> -.a \\in Int ==> b \\in Int + ==> (a + -.a) + b = a + (-.a + b)" + using AddAssociative_sequent by blast + have s2_5: "(a + -.a) + b = a + (-.a + b)" + using s2_1 s2_2 s2_3 s2_4 by auto + show ?thesis + by (rule s2_5[symmetric]) + qed + also have s1_4: "... = 0 + b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "a + -.a = 0" + using s2_1 AddNegCancels by auto + show ?thesis + using s2_2 by auto + qed + also have s1_5: "... = b" + using bdom add_0_left by auto + finally show ?thesis . + qed + + +(* +THEOREM MinusDistributesAdd == + \A a \in Int: \A b \in Int: + -.(a + b) = (-.a) + (-.b) +PROOF +<1>1. SUFFICES ASSUME NEW a \in Int + PROVE \A b \in Int: + -.(a + b) = (-.a) + (-.b) + BY <1>1 +<1>2. -.(a + 0) = (-.a) + (-.0) + <2>1. -.(a + 0) = -.a + <3>1. a + 0 = a + BY <1>1, add_0 + <3> QED + BY <3>1 + <2>2. -.a + -.0 = -.a + <3>1. -.a + -.0 = -.a + 0 + BY neg0 + <3>2. @ = -.a + <4>1. -.a \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, add_0 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW b \in Nat, + -.(a + b) = (-.a) + (-.b) + PROVE + -.(a + Succ[b]) = (-.a) + (-.Succ[b]) + <2>1. -.(a + Succ[b]) = -.iSucc[a + b] + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3>3. a + Succ[b] = iSucc[a + b] + BY <3>1, <3>2, addint_succ_nat + <3> QED + BY <3>3 + <2>2. @ = -. iSucc[-.((-.a) + (-.b))] + <3>1. a + b = -.(-.a + -.b) + <4>1. -.-.(a + b) = -.(-.a + -.b) + BY <1>3 + <4>2. -.-.(a + b) = a + b + <5>1. a + b \in Int + <6>1. a \in Int + BY <1>1 + <6>2. b \in Int + BY <1>3, nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, minus_involutive + <4> QED + BY <4>1, <4>2 + <3> QED + BY <3>1 + <2>3. @ = iPred[-.a + -.b] + <3>1. -.a + -.b \in Int + <4>1. -.a \in Int + BY <1>1, neg_int_eq_int + <4>2. -.b \in Int + BY <1>3, minus_nat_in_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, iSuccMinusFlip + <2>4. @ = -.a + iPred[-.b] + <3>1. -.a \in Int + BY <1>1, neg_int_eq_int + <3>2. -.b \in Int + BY <1>3, minus_nat_in_int + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>5. @ = -.a + -.Succ[b] + <3>1. iPred[-.b] = -.Succ[-.-.b] + <4>1. -.b \in Int + BY <1>3, minus_nat_in_int + <4>2. -.b \notin Nat \ {0} + <5>1. -.b \in negNat + BY <1>3, minus_nat_in_neg_nat + <5> QED + BY <5>1, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3>2. @ = -.Succ[b] + <4>1. -.-.b = b + BY <1>3, nat_in_int, minus_involutive + <4> QED + BY <4>1 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5 +<1>4. ASSUME NEW b \in negNat, + -.(a + b) = (-.a) + (-.b) + PROVE + -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b]) + <2>1. -.(a + -.Succ[-.b]) = -. iPred[a + b] + <3>1. a + -.Succ[-.b] = iPred[a + -.-.b] + <4>1. a \in Int + BY <1>1 + <4>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, addint_succ_negnat + <3>2. a + -.Succ[-.b] = iPred[a + b] + <4>1. -.-.b = b + BY <1>4, neg_nat_in_int, minus_involutive + <4> QED + BY <3>1, <4>1 + <3> QED + BY <3>2 + <2>2. @ = -. iPred[-.((-.a) + (-.b))] + <3>1. a + b = -.(-.a + -.b) + <4>1. -.(a + b) = (-.a) + (-.b) + BY <1>4 + <4>2. -.-.(a + b) = -.(-.a + -.b) + BY <4>1 + <4>3. -.-.(a + b) = a + b + <5>1. a + b \in Int + <6>1. a \in Int + BY <1>1 + <6>2. b \in Int + BY <1>4, neg_nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, minus_involutive + <4> QED + BY <4>2, <4>3 + <3> QED + BY <3>1 + <2>3. @ = iSucc[-.a + -.b] + <3>1. -.a + -.b \in Int + <4>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <4>2. -.b \in Int + BY <1>4, minus_negnat_in_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, iPredMinusFlip + <2>4. @ = -.a + iSucc[-.b] + <3>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <3>2. -.b \in Int + BY <1>4, minus_negnat_in_int + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>5. @ = -.a + Succ[-.b] + <3>1. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>2. -.b \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>1, <3>2 DEF iSucc + <2>6. @ = -.a + (-.-.Succ[-.b]) + <3>1. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>2. Succ[-.b] \in Nat + BY <3>1, succIsNat + <3>3. Succ[-.b] \in Int + BY <3>2, nat_in_int + <3>4. -.-.Succ[-.b] = Succ[-.b] + BY <3>3, minus_involutive + <3> QED + BY <3>4 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 +<1>5. \A b \in Nat: -.(a + b) = (-.a) + (-.b) + BY <1>2, <1>3, NatInduction +<1>6. \A b \in negNat: -.(a + b) = (-.a) + (-.b) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MinusDistributesAdd: + "\ a \ Int: \ b \ Int: + -.(a + b) = (-.a) + (-.b)" + proof - + { + fix "a" :: "c" + assume adom: "a \\in Int" + have s1_2: "-.(a + 0) = (-.a) + (-.0)" + proof - + have s2_1: "-.(a + 0) = -.a" + proof - + have s3_1: "a + 0 = a" + using adom add_0 by auto + show ?thesis + using s3_1 by auto + qed + have s2_2: "-.a + -.0 = -.a" + proof - + have s3_1: "-.a + -.0 = -.a + 0" + using neg0 by auto + also have s3_2: "... = -.a" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + show ?thesis + using s4_1 add_0 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ b. \ + b \\in Nat; + -.(a + b) = -.a + -.b + \ \ + -.(a + Succ[b]) = (-.a) + (-.Succ[b])" + proof - + fix "b" :: "c" + assume s1_3_bdom: "b \\in Nat" + assume s1_3_induct: "-.(a + b) = -.a + -.b" + have s2_1: "-.(a + Succ[b]) = -.iSucc[a + b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + have s3_3: "a + Succ[b] = iSucc[a + b]" + using s3_1 s3_2 addint_succ_nat by auto + show ?thesis + using s3_3 by auto + qed + also have s2_2: "... = -. iSucc[-.((-.a) + (-.b))]" + proof - + have s3_1: "a + b = -.(-.a + -.b)" + proof - + have s4_1: "-.-.(a + b) = -.(-.a + -.b)" + using s1_3_induct by auto + have s4_2: "-.-.(a + b) = a + b" + proof - + have s5_1: "a + b \\in Int" + proof - + have s6_1: "a \\in Int" + using adom by auto + have s6_2: "b \\in Int" + using s1_3_bdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 minus_involutive by auto + qed + show ?thesis + using s4_1 s4_2 by auto + qed + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = iPred[-.a + -.b]" + proof - + have s3_1: "-.a + -.b \\in Int" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s4_2: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 iSuccMinusFlip by auto + qed + also have s2_4: "... = -.a + iPred[-.b]" + proof - + have s3_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s3_2: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by auto + qed + also have s2_5: "... = -.a + -.Succ[b]" + proof - + have s3_1: "iPred[-.b] = -.Succ[-.-.b]" + proof - + have s4_1: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + have s4_2: "-.b \ Nat \ {0}" + proof - + have s5_1: "-.b \\in negNat" + using s1_3_bdom minus_nat_in_neg_nat by auto + show ?thesis + using s5_1 neg_nat_not_in_nat_setminus_0 by fast + qed + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + also have s3_2: "... = -.Succ[b]" + proof - + have s4_1: "-.-.b = b" + using s1_3_bdom nat_in_int minus_involutive by auto + show ?thesis + using s4_1 by auto + qed + finally have s3_3: "iPred[-.b] = -.Succ[b]" . + show ?thesis + using s3_3 by auto + qed + finally show "-.(a + Succ[b]) = (-.a) + (-.Succ[b])" . + qed + have s1_4: "\ b. \ + b \\in negNat; + -.(a + b) = -.a + -.b + \ \ + -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" + proof - + fix "b" :: "c" + assume s1_4_bdom: "b \\in negNat" + assume s1_4_induct: "-.(a + b) = -.a + -.b" + have s2_1: "-.(a + -.Succ[-.b]) = -. iPred[a + b]" + proof - + have s3_1: "a + -.Succ[-.b] = iPred[a + -.-.b]" + proof - + have s4_1: "a \\in Int" + using adom by auto + have s4_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 addint_succ_negnat by auto + qed + have s3_2: "a + -.Succ[-.b] = iPred[a + b]" + proof - + have s4_1: "-.-.b = b" + using s1_4_bdom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s3_1 s4_1 by auto + qed + show ?thesis + using s3_2 by auto + qed + also have s2_2: "... = -. iPred[-.((-.a) + (-.b))]" + proof - + have s3_1: "a + b = -.(-.a + -.b)" + proof - + have s4_1: "-.(a + b) = (-.a) + (-.b)" + using s1_4_induct by auto + have s4_2: "-.-.(a + b) = -.(-.a + -.b)" + using s4_1 by auto + have s4_3: "-.-.(a + b) = a + b" + proof - + have s5_1: "a + b \\in Int" + proof - + have s6_1: "a \\in Int" + using adom by auto + have s6_2: "b \\in Int" + using s1_4_bdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 minus_involutive by auto + qed + show ?thesis + using s4_2 s4_3 by auto + qed + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = iSucc[-.a + -.b]" + proof - + have s3_1: "-.a + -.b \\in Int" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s4_2: "-.b \\in Int" + using s1_4_bdom minus_negnat_in_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 iPredMinusFlip by auto + qed + also have s2_4: "... = -.a + iSucc[-.b]" + proof - + have s3_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s3_2: "-.b \\in Int" + using s1_4_bdom minus_negnat_in_int by auto + show ?thesis + using s3_1 s3_2 iSuccRightDistributesAdd by auto + qed + also have s2_5: "... = -.a + Succ[-.b]" + proof - + have s3_1: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + have s3_2: "-.b \\in Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + also have s2_6: "... = -.a + (-.-.Succ[-.b])" + proof - + have s3_1: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + have s3_2: "Succ[-.b] \\in Nat" + using s3_1 succIsNat by auto + have s3_3: "Succ[-.b] \\in Int" + using s3_2 nat_in_int by auto + have s3_4: "-.-.Succ[-.b] = Succ[-.b]" + using s3_3 minus_involutive by auto + show ?thesis + using s3_4 by auto + qed + finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . + qed + have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem MinusDistributesAdd_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "-.(a + b) = (-.a) + (-.b)" + using adom bdom MinusDistributesAdd by auto + + +(* +THEOREM MinusDiff == + ASSUME NEW a \in Int, NEW b \in Int + PROVE -.(a - b) = b - a +PROOF +<1>1. -.(a - b) = -.(a + -.b) + BY DEF diff +<1>2. @ = (-.a) + (-.-.b) + <2>1. a \in Int + OBVIOUS + <2>2. -.b \in Int + BY b \in Int, neg_int_eq_int + <2> QED + BY <2>1, <2>2, MinusDistributesAdd +<1>3. @ = (-.a) + b + BY b \in Int, minus_involutive +<1>4. @ = b + (-.a) + <2>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>2. b \in Int + OBVIOUS + <2> QED + BY <2>1, <2>2, AddCommutative +<1>5. @ = b - a + BY DEF diff +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5 +*) +theorem MinusDiff: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "-.(a - b) = b - a" + proof - + have s1_1: "-.(a - b) = -.(a + -.b)" + unfolding diff_def by auto + also have s1_2: "... = -.a + -.-.b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + show ?thesis + using s2_1 s2_2 MinusDistributesAdd by auto + qed + also have s1_3: "... = -.a + b" + using bdom minus_involutive by auto + also have s1_4: "... = b + -.a" + proof - + have s2_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_2: "b \\in Int" + using bdom by auto + show ?thesis + using s2_1 s2_2 AddCommutative by auto + qed + also have s1_5: "... = b - a" + unfolding diff_def by auto + finally show ?thesis . + qed + + +(*----------------------------------------------------------------------------*) +(* Commutativity of multiplication. *) + +(* +THEOREM MultCommutative == + \A i, j \in Int: i * j = j * i +PROOF +<1>1. \A i \in Int: i * 0 = 0 * i + <2>1. 0 * 0 = 0 * 0 + OBVIOUS + <2>2. ASSUME NEW i \in Nat, i * 0 = 0 * i + PROVE Succ[i] * 0 = 0 * Succ[i] + <3>1. 0 * Succ[i] = 0 * i + 0 + <4>1. 0 \in Int + BY zero_in_int + <4>2. i \in Nat + BY <2>2 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3>2. @ = 0 * i + <4>1. 0 * i \in Int + <5>1 0 \in Int + BY zero_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5> QED + BY <5>1, <5>2, multIsInt + <4> QED + BY <4>1, add_0 + <3>3. @ = i * 0 + BY <2>2 + <3>4. @ = 0 + <4>1. i \in Int + BY <2>2, nat_in_int + <4> QED + BY <4>1, mult_0 + <3>5. @ = Succ[i] * 0 + <4>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5 + <2>3. ASSUME NEW i \in negNat, i * 0 = 0 * i + PROVE -.Succ[-.i] * 0 = 0 * -.Succ[-.i] + <3>1. -.Succ[-.i] * 0 = 0 + <4>1. -.Succ[-.i] \in Int + BY <2>3, minus_succ_minus_negnat_in_int + <4> QED + BY <4>1, mult_0 + <3>2. 0 * -.Succ[-.i] = 0 + <4>1. 0 * -.Succ[-.i] = 0 * -.-.i + -.0 + <5>1. 0 \in Int + BY zero_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = 0 * i + -.0 + <5>1. -.-.i = i + BY <2>3, neg_nat_in_int, minus_involutive + <5> QED + BY <5>1 + <4>3. @ = 0 * i + 0 + BY neg0 + <4>4. @ = 0 * i + <5>1. 0 * i \in Int + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Int + BY <2>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5> QED + BY <5>1, add_0 + <4>5. @ = i * 0 + BY <2>3 \* induction hypothesis + <4>6. @ = 0 + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5> QED + BY <5>1, mult_0 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, + <4>5, <4>6 + <3> QED + BY <3>1, <3>2 + <2>4. \A i \in Nat: i * 0 = 0 * i + BY <2>1, <2>2, NatInduction + <2>5. \A i \in negNat: i * 0 = 0 * i + BY <2>1, <2>3, neg_nat_induction + <2> QED + BY <2>4, <2>5, int_union_nat_negnat +<1>2. ASSUME NEW j \in Nat, + \A i \in Int: i * j = j * i + PROVE \A i \in Int: i * Succ[j] = Succ[j] * i + <2>1. 0 * Succ[j] = Succ[j] * 0 + <3>1. 0 * Succ[j] = 0 + <4>1. 0 * Succ[j] = 0 * j + 0 + <5>1. 0 \in Int + BY zero_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = 0 * j + <5>1. 0 * j \in Int + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Int + BY <1>2, nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5> QED + BY <5>1, add_0 + <4>3. @ = j * 0 + BY <1>2, zero_in_int, spec + <4>4. @ = 0 + <5>1. j \in Int + BY <1>2, nat_in_int + <5> QED + BY <5>1, mult_0 + <4> QED + BY <4>1, <4>2, <4>3, <4>4 + <3>2. Succ[j] * 0 = 0 + <4>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2 + <2>2. ASSUME NEW i \in Nat, + i * Succ[j] = Succ[j] * i + PROVE Succ[i] * Succ[j] = Succ[j] * Succ[i] + <3>1. Succ[i] * Succ[j] = i * j + ((i + j) + 1) + <4>1. Succ[i] * Succ[j] = Succ[i] * j + Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = j * Succ[i] + Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. Succ[i] * j = j * Succ[i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>3. @ = (j * i + j) + Succ[i] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5>3. j * Succ[i] = j * i + j + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>4. @ = (i * j + j) + Succ[i] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>5. @ = (i * j + j) + (i + 1) + <5>1. i \in Nat + BY <2>2 + <5>2. Succ[i] = i + 1 + BY <5>1, nat_add_1 + <5> QED + BY <5>2 + <4>6. @ = i * j + (j + (i + 1)) + <5>1. i * j \in Int + BY <1>2, <2>2, nat_in_int, multIsInt + <5>2. j \in Int + BY <1>2, nat_in_int + <5>3. i + 1 \in Int + BY <2>2, oneIsNat, nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + ((j + i) + 1) + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. j + (i + 1) = (j + i) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>8. @ = i * j + ((i + j) + 1) + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. j + i = i + j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8 + <3>2. Succ[j] * Succ[i] = i * j + ((i + j) + 1) + <4>1. Succ[j] * Succ[i] = Succ[j] * i + Succ[j] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = i * Succ[j] + Succ[j] + BY <2>2 \* induction hypothesis + <4>3. @ = (i * j + i) + Succ[j] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5>3. i * Succ[j] = i * j + i + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>4. @ = (i * j + i) + (j + 1) + <5>1. j \in Nat + BY <1>2 + <5>2. Succ[j] = j + 1 + BY <5>1, nat_add_1 + <5> QED + BY <5>2 + <4>5. @ = i * j + (i + (j + 1)) + <5>1. i * j \in Int + BY <1>2, <2>2, nat_in_int, multIsInt + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. j + 1 \in Int + BY <1>2, oneIsNat, nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>6. @ = i * j + ((i + j) + 1) + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Int + BY <1>2, nat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. i + (j + 1) = (i + j) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, <4>6 + <3> QED + BY <3>1, <3>2 + <2>3. ASSUME NEW i \in negNat, + i * Succ[j] = Succ[j] * i + PROVE -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i] + <3>1. -.Succ[-.i] * Succ[j] = + i * j + ((i + -.j) + -.1) + <4>1. -.Succ[-.i] * Succ[j] = + -.Succ[-.i] * j + -.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = j * -.Succ[-.i] + -.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>3. @ = (j * -.-.i + -.j) + -.Succ[-.i] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>3. j * -.Succ[-.i] = j * -.-.i + -.j + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>4. @ = (j * i + -.j) + -.Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>5. @ = (i * j + -.j) + -.Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>6. @ = (i * j + -.j) + (i + -.1) + <5>1. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>2. -.Succ[-.i] = -.-.i + -.1 + BY <5>1, MinusSuccIsMinusOne + <5>3. -.-.i = i + BY <2>3, neg_nat_in_int, minus_involutive + <5>4. -.Succ[-.i] = i + -.1 + BY <5>2, <5>3 + <5> QED + BY <5>4 + <4>7. @ = i * j + (-.j + (i + -.1)) + <5>1. i * j \in Int + BY <1>2, nat_in_int, <2>3, neg_nat_in_int, + multIsInt + <5>2. -.j \in Int + BY <1>2, minus_nat_in_int + <5>3. i + -.1 \in Int + BY <2>3, minus_one_in_negnat, + neg_nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + ((-.j + i) + -.1) + <5>1. -.j \in Int + BY <1>2, minus_nat_in_int + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <5>4. -.j + (i + -.1) = (-.j + i) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>9. @ = i * j + ((i + -.j) + -.1) + <5>1. -.j \in Int + BY <1>2, minus_nat_in_int + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.j + i = i + -.j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, + <4>5, <4>6, <4>7, <4>8, <4>9 + <3>2. Succ[j] * -.Succ[-.i] = + i * j + ((i + -.j) + -.1) + <4>1. Succ[j] * -.Succ[-.i] = + Succ[j] * -.-.i + -.Succ[j] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = Succ[j] * i + -.Succ[j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = i * Succ[j] + -.Succ[j] + BY <2>3 \* induction hypothesis + <4>4. @ = (i * j + i) + -.Succ[j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. j \in Nat + BY <1>2 + <5>3. i * Succ[j] = i * j + i + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>5. @ = (i * j + i) + (-.j + -.1) + <5>1. j \in Nat + BY <1>2 + <5>2. -.Succ[j] = -.j + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>6. @ = i * j + (i + (-.j + -.1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>2, nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.j + -.1 \in Int + <6>1. -.j \in Int + BY <1>2, minus_nat_in_int + <6>2. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + ((i + -.j) + -.1) + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.j \in Int + BY <1>2, minus_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <5>4. i + (-.j + -.1) = (I + -.j) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, NatInduction, neg_nat_induction, + int_union_nat_negnat +<1>3. ASSUME NEW j \in negNat, + \A i \in Int: i * j = j * i + PROVE \A i \in Int: i * -.Succ[-.j] = -.Succ[-.j] * i + <2>1. 0 * -.Succ[-.j] = -.Succ[-.j] * 0 + <3>1. 0 * -.Succ[-.j] = 0 * -.-.j + -.0 + <4>1. 0 \in Int + BY zeroIsNat nat_in_int + <4>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3>2. @ = 0 * -.-.j + <4>1. 0 * -.-.j \in Int + <5>1. 0 \in Int + BY zeroIsNat nat_in_int + <5>2. -.-.j \in Int + BY <1>3, neg_nat_in_int, minus_involutive + <5> QED + BY <5>1, <5>2, multIsInt + <4>2. 0 * -.-.j + 0 = 0 * -.-.j + BY <4>1, add_0 + <4> QED + BY <4>2, neg0 + <3>3. @ = 0 * j + <4>1. j \in Int + BY <1>3, neg_nat_in_int + <4> QED + BY <4>1, minus_involutive + <3>4. @ = j * 0 + <4>1. 0 \in Int + BY zeroIsNat, nat_in_int + <4>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <4> QED + BY <4>1, <4>2, spec + <3>5. @ = 0 + <4>1. j \in Int + BY <1>3, neg_nat_in_int + <4> QED + BY <4>1, mult_0 + <3>6. @ = -.Succ[-.j] * 0 + <4>1. -.Succ[-.j] \in Int + BY <1>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 + <2>2. ASSUME NEW i \in Nat, + i * -.Succ[-.j] = -.Succ[-.j] * i + PROVE Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i] + <3>1. Succ[i] * -.Succ[-.j] = i * j + ((-.i + j) + -.1) + <4>1. Succ[i] * -.Succ[-.j] = + Succ[i] * -.-.j + -.Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = Succ[i] * j + -.Succ[i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = j * Succ[i] + -.Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. Succ[i] * j = j * Succ[i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>4. @ = (j * i + j) + -.Succ[i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. i \in Nat + BY <2>2 + <5>3. j * Succ[i] = j * i + j + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>5. @ = (i * j + j) + -.Succ[i] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>6. @ = i * j + (j + -.Succ[i]) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>2, nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. j \in Int + BY <1>3, neg_nat_in_int + <5>3. -.Succ[i] \in Int + BY <2>2, succIsNat, minus_nat_in_int + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + (j + (-.i + -.1)) + <5>1. i \in Nat + BY <2>2 + <5>2. -.Succ[i] = -.i + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>8. @ = i * j + ((j + -.i) + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_int + <5>4. j + (-.i + -.1) = (j + -.i) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>9. @ = i * j + ((-.i + j) + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. j + -.i = -.i + j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9 + <3>2. -.Succ[-.j] * Succ[i] = i * j + ((-.i + j) + -.1) + <4>1. -.Succ[-.j] * Succ[i] = + -.Succ[-.j] * i + -.Succ[-.j] + <5>1. -.Succ[-.j] \in Int + BY <1>3, minus_negnat_succ_in_negnat, + neg_nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = i * -.Succ[-.j] + -.Succ[-.j] + BY <2>2 \* induction hypothesis + <4>3. @ = (i * -.-.j + -.i) + -.Succ[-.j] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>3. i * -.Succ[-.j] = i * -.-.j + -.i + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>4. @ = (i * j + -.i) + -.Succ[-.j] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>5. @ = (i * j + -.i) + (-.-.j + -.1) + <5>1. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>2. -.Succ[-.j] = -.-.j + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>6. @ = (i * j + -.i) + (j + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = i * j + (-.i + (j + -.1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>2, nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. j + -.1 \in Int + <6>1. j \in Int + BY <1>3, neg_nat_in_int + <6>2. -.1 \in Int + BY minus_one_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + ((-.i + j) + -.1) + <5>1. -.i \in Int + BY <2>2, minus_nat_in_int + <5>2. j \in Int + BY <1>3, neg_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_int + <5>4. -.i + (j + -.1) = (-.i + j) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8 + <3> QED + BY <3>1, <3>2 + <2>3. ASSUME NEW i \in negNat, + i * -.Succ[-.j] = -.Succ[-.j] * i + PROVE -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i] + <3>1. -.Succ[-.i] * -.Succ[-.j] = + i * j + ((-.i + -.j) + 1) + <4>1. -.Succ[-.i] * -.Succ[-.j] = + -.Succ[-.i] * -.-.j + -.-.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, + neg_nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = -.Succ[-.i] * j + -.-.Succ[-.i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = -.Succ[-.i] * j + Succ[-.i] + <5>1. Succ[-.i] \in Int + BY <2>3, negnat_succ_in_nat, nat_in_int + <5>2. -.-.Succ[-.i] = Succ[-.i] + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>4. @ = j * -.Succ[-.i] + Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>5. @ = (j * -.-.i + -.j) + Succ[-.i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>3. j * -.Succ[-.i] = j * -.-.i + -.j + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>6. @ = (j * i + -.j) + Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = (i * j + -.j) + Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>8. @ = (i * j + -.j) + (-.i + 1) + <5>1. i \in negNat + BY <2>3 + <5>2. Succ[-.i] = -.i + 1 + BY <5>1, SuccMinusIsPlusOne + <5> QED + BY <5>2 + <4>9. @ = i * j + (-.j + (-.i + 1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>3. -.i + 1 \in Int + <6>1. -.i \in Int + BY <2>3, minus_negnat_in_int + <6>2. 1 \in Int + BY oneIsNat, nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>10. @ = i * j + ((-.j + -.i) + 1) + <5>1. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. -.j + (-.i + 1) = (-.j + -.i) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>11. @ = i * j + ((-.i + -.j) + 1) + <5>1. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. -.j + -.i = -.i + -.j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9, + <4>10, <4>11 + <3>2. -.Succ[-.j] * -.Succ[-.i] = + i * j + ((-.i + -.j) + 1) + <4>1. -.Succ[-.j] * -.Succ[-.i] = + -.Succ[-.j] * -.-.i + -.-.Succ[-.j] + <5>1. -.Succ[-.j] \in Int + BY <1>3, minus_succ_minus_negnat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = -.Succ[-.j] * i + -.-.Succ[-.j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = -.Succ[-.j] * i + Succ[-.j] + <5>1. Succ[-.j] \in Int + BY <1>3, negnat_succ_in_nat, nat_in_int + <5>2. -.-.Succ[-.j] = Succ[-.j] + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>4. @ = i * -.Succ[-.j] + Succ[-.j] + BY <2>3 \* induction hypothesis + <4>5. @ = (i * -.-.j + -.i) + Succ[-.j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>3. i * -.Succ[-.j] = i * -.-.j + -.i + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>6. @ = (i * j + -.i) + Succ[-.j] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = i * j + (-.i + Succ[-.j]) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. Succ[-.j] \in Int + BY <1>3, negnat_succ_in_nat, nat_in_int + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + (-.i + (-.j + 1)) + <5>1. Succ[-.j] = -.j + 1 + BY <1>3, SuccMinusIsPlusOne + <5> QED + BY <5>1 + <4>9. @ = i * j + ((-.i + -.j) + 1) + <5>1. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>2. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. -.i + (-.j + 1) = (-.i + -.j) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, NatInduction, neg_nat_induction, + int_union_nat_negnat +<1> QED + BY <1>1, <1>2, <1>3, NatInduction, neg_nat_induction, + int_union_nat_negnat +*) +theorem MultCommutative: + "\ i \ Int: \ j \ Int: + i * j = j * i" + proof - + have s1_1: "\ i \ Int: i * 0 = 0 * i" + proof - + have s2_1: "0 * 0 = 0 * 0" + by auto + have s2_2: "\ i. \ + i \\in Nat; + i * 0 = 0 * i + \ \ + Succ[i] * 0 = 0 * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * 0 = 0 * i" + have s3_1: "0 * Succ[i] = 0 * i + 0" + proof - + have s4_1: "0 \\in Int" + using zero_in_int by auto + have s4_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + also have s3_2: "... = 0 * i" + proof - + have s4_1: "0 * i \\in Int" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + show ?thesis + using s4_1 add_0 by auto + qed + also have s3_3: "... = i * 0" + using s2_2_induct by auto + also have s3_4: "... = 0" + proof - + have s4_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + also have s3_5: "... = Succ[i] * 0" + proof - + have s4_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + finally show "Succ[i] * 0 = 0 * Succ[i]" by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * 0 = 0 * i + \ \ + -.Succ[-.i] * 0 = 0 * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * 0 = 0 * i" + have s3_1: "-.Succ[-.i] * 0 = 0" + proof - + have s4_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_succ_minus_negnat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + have s3_2: "0 * -.Succ[-.i] = 0" + proof - + have s4_1: "0 * -.Succ[-.i] = 0 * -.-.i + -.0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = 0 * i + -.0" + proof - + have s5_1: "-.-.i = i" + using s2_3_idom neg_nat_in_int + minus_involutive by auto + show ?thesis + using s5_1 by auto + qed + also have s4_3: "... = 0 * i + 0" + using neg0 by auto + also have s4_4: "... = 0 * i" + proof - + have s5_1: "0 * i \\in Int" + proof - + have s6_1: "0 \\in Int" + using zero_in_int by auto + have s6_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + show ?thesis + using s5_1 add_0 by auto + qed + also have s4_5: "... = i * 0" + using s2_3_induct by auto + also have s4_6: "... = 0" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + show ?thesis + using s5_1 mult_0 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * 0 = 0 * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: i * 0 = 0 * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: i * 0 = 0 * i" + using s2_1 s2_3 neg_nat_induction by auto + show ?thesis + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_2: "\ j. \ + j \\in Nat; + \ i \ Int: i * j = j * i + \ \ + \ i \ Int: i * Succ[j] = Succ[j] * i" + proof - + fix "j" :: "c" + assume s1_2_jdom: "j \\in Nat" + assume s1_2_induct: "\ i \ Int: i * j = j * i" + have s2_1: "0 * Succ[j] = Succ[j] * 0" + proof - + have s3_1: "0 * Succ[j] = 0" + proof - + have s4_1: "0 * Succ[j] = 0 * j + 0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = 0 * j" + proof - + have s5_1: "0 * j \\in Int" + proof - + have s6_1: "0 \\in Int" + using zero_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + show ?thesis + using s5_1 add_0 by auto + qed + also have s4_3: "... = j * 0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + show ?thesis + using s5_1 s5_2 spec by fast + qed + also have s4_4: "... = 0" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s5_1 mult_0 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * 0 = 0" + proof - + have s4_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + have s2_2: "\ i. \ + i \\in Nat; + i * Succ[j] = Succ[j] * i + \ \ + Succ[i] * Succ[j] = Succ[j] * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * Succ[j] = Succ[j] * i" + have s3_1: "Succ[i] * Succ[j] = i * j + ((i + j) + 1)" + proof - + have s4_1: "Succ[i] * Succ[j] = Succ[i] * j + Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = j * Succ[i] + Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_2_induct by auto + have s5_3: "Succ[i] * j = j * Succ[i]" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_3: "... = (j * i + j) + Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + have s5_3: "j * Succ[i] = j * i + j" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + j) + Succ[i]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_2_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + j) + (i + 1)" + proof - + have s5_1: "i \\in Nat" + using s2_2_idom by auto + have s5_2: "Succ[i] = i + 1" + using s5_1 nat_add_1 by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = i * j + (j + (i + 1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_3: "i + 1 \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "1 \\in Nat" + using oneIsNat nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + ((j + i) + 1)" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "j + (i + 1) = (j + i) + 1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_8: "... = i * j + ((i + j) + 1)" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "j + i = i + j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * Succ[i] = i * j + ((i + j) + 1)" + proof - + have s4_1: "Succ[j] * Succ[i] = Succ[j] * i + Succ[j]" + proof - + have s5_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = i * Succ[j] + Succ[j]" + using s2_2_induct by auto + also have s4_3: "... = (i * j + i) + Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + have s5_3: "i * Succ[j] = i * j + i" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + i) + (j + 1)" + proof - + have s5_1: "j \\in Nat" + using s1_2_jdom by auto + have s5_2: "Succ[j] = j + 1" + using s5_1 nat_add_1 by fast + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = i * j + (i + (j + 1))" + proof - + have s5_1: "i * j \\in Int" + using s1_2_jdom s2_2_idom nat_in_int multIsInt by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "j + 1 \\in Int" + using s1_2_jdom oneIsNat nat_in_int addIsInt by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_6: "... = i * j + ((i + j) + 1)" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "i + (j + 1) = (i + j) + 1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "Succ[i] * Succ[j] = Succ[j] * Succ[i]" + using s3_1 s3_2 by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * Succ[j] = Succ[j] * i + \ \ + -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * Succ[j] = Succ[j] * i" + have s3_1: "-.Succ[-.i] * Succ[j] = + i * j + ((i + -.j) + -.1)" + proof - + have s4_1: "-.Succ[-.i] * Succ[j] = + -.Succ[-.i] * j + -.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = j * -.Succ[-.i] + -.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_3: "... = (j * -.-.i + -.j) + -.Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (j * i + -.j) + -.Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = (i * j + -.j) + -.Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (i * j + -.j) + (i + -.1)" + proof - + have s5_1: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_2: "-.Succ[-.i] = -.-.i + -.1" + using s5_1 MinusSuccIsMinusOne by fast + have s5_3: "-.-.i = i" + using s2_3_idom neg_nat_in_int minus_involutive + by auto + have s5_4: "-.Succ[-.i] = i + -.1" + using s5_2 s5_3 by auto + show ?thesis + using s5_4 by auto + qed + also have s4_7: "... = i * j + (-.j + (i + -.1))" + proof - + have s5_1: "i * j \\in Int" + using s1_2_jdom nat_in_int s2_3_idom + neg_nat_in_int multIsInt by auto + have s5_2: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_3: "i + -.1 \\in Int" + using s2_3_idom minus_one_in_negnat + neg_nat_in_int addIsInt by auto + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + ((-.j + i) + -.1)" + proof - + have s5_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + have s5_4: " -.j + (i + -.1) = (-.j + i) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_9: "... = i * j + ((i + -.j) + -.1)" + proof - + have s5_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.j + i = i + -.j" + using s5_1 s5_2 AddCommutative_sequent + by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * -.Succ[-.i] = + i * j + ((i + -.j) + -.1)" + proof - + have s4_1: "Succ[j] * -.Succ[-.i] = + Succ[j] * -.-.i + -.Succ[j]" + proof - + have s5_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = Succ[j] * i + -.Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = i * Succ[j] + -.Succ[j]" + using s2_3_induct by auto + also have s4_4: "... = (i * j + i) + -.Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + have s5_3: "i * Succ[j] = i * j + i" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + i) + (-.j + -.1)" + proof - + have s5_1: "j \\in Nat" + using s1_2_jdom by auto + have s5_2: "-.Succ[j] = -.j + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = i * j + (i + (-.j + -.1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.j + -.1 \\in Int" + proof - + have s6_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s6_2: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + ((i + -.j) + -.1)" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + have s5_4: "i + (-.j + -.1) = (i + -.j) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" + using s2_1 s2_3 neg_nat_induction by auto + show "\ i \ Int: i * Succ[j] = Succ[j] * i" + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_3: "\ j. \ + j \\in negNat; + \ i \ Int: i * j = j * i + \ \ + \ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" + proof - + fix "j" :: "c" + assume s1_3_jdom: "j \\in negNat" + assume s1_3_induct: "\ i \ Int: i * j = j * i" + have s2_1: "0 * -.Succ[-.j] = -.Succ[-.j] * 0" + proof - + have s3_1: "0 * -.Succ[-.j] = 0 * -.-.j + -.0" + proof - + have s4_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s4_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + also have s3_2: "... = 0 * -.-.j" + proof - + have s4_1: "0 * -.-.j \\in Int" + proof - + have s5_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s5_2: "-.-.j \\in Int" + using s1_3_jdom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + have s4_2: "0 * -.-.j + 0 = 0 * -.-.j" + using s4_1 add_0 by auto + show ?thesis + using s4_2 neg0 by auto + qed + also have s3_3: "... = 0 * j" + proof - + have s4_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s4_1 minus_involutive by auto + qed + also have s3_4: "... = j * 0" + proof - + have s4_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s4_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + show ?thesis + using s4_1 s4_2 spec by fast + qed + also have s3_5: "... = 0" + proof - + have s4_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + also have s3_6: "... = -.Succ[-.j] * 0" + proof - + have s4_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + finally show ?thesis . + qed + have s2_2: "\ i. \ + i \\in Nat; + i * -.Succ[-.j] = -.Succ[-.j] * i + \ \ + Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" + have s3_1: "Succ[i] * -.Succ[-.j] = + i * j + ((-.i + j) + -.1)" + proof - + have s4_1: "Succ[i] * -.Succ[-.j] = + Succ[i] * -.-.j + -.Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = Succ[i] * j + -.Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = j * Succ[i] + -.Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + have s5_3: "Succ[i] * j = j * Succ[i]" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (j * i + j) + -.Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + have s5_3: "j * Succ[i] = j * i + j" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + j) + -.Succ[i]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = i * j + (j + -.Succ[i])" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_3: "-.Succ[i] \\in Int" + using s2_2_idom succIsNat minus_nat_in_int by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + (j + (-.i + -.1))" + proof - + have s5_1: "i \\in Nat" + using s2_2_idom by auto + have s5_2: "-.Succ[i] = -.i + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_8: "... = i * j + ((j + -.i) + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_int by auto + have s5_4: "j + (-.i + -.1) = (j + -.i) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_9: "... = i * j + ((-.i + j) + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "j + -.i = -.i + j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "-.Succ[-.j] * Succ[i] = + i * j + ((-.i + j) + -.1)" + proof - + have s4_1: "-.Succ[-.j] * Succ[i] = + -.Succ[-.j] * i + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = i * -.Succ[-.j] + -.Succ[-.j]" + using s2_2_induct by auto + also have s4_3: "... = (i * -.-.j + -.i) + -.Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + -.i) + -.Succ[-.j]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = (i * j + -.i) + (-.-.j + -.1)" + proof - + have s5_1: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_2: "-.Succ[-.j] = -.-.j + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = (i * j + -.i) + (j + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = i * j + (-.i + (j + -.1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "j + -.1 \\in Int" + proof - + have s6_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s6_2: "-.1 \\in Int" + using minus_one_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + ((-.i + j) + -.1)" + proof - + have s5_1: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_int by auto + have s5_4: "-.i + (j + -.1) = (-.i + j) + -.1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" + using s3_1 s3_2 by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * -.Succ[-.j] = -.Succ[-.j] * i + \ \ + -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" + have s3_1: "-.Succ[-.i] * -.Succ[-.j] = + i * j + ((-.i + -.j) + 1)" + proof - + have s4_1: "-.Succ[-.i] * -.Succ[-.j] = + -.Succ[-.i] * -.-.j + -.-.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = -.Succ[-.i] * j + -.-.Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = -.Succ[-.i] * j + Succ[-.i]" + proof - + have s5_1: "Succ[-.i] \\in Int" + using s2_3_idom negnat_succ_in_nat nat_in_int by auto + have s5_2: "-.-.Succ[-.i] = Succ[-.i]" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_4: "... = j * -.Succ[-.i] + Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_3_induct by auto + have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (j * -.-.i + -.j) + Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (j * i + -.j) + Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = (i * j + -.j) + Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_3_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_8: "... = (i * j + -.j) + (-.i + 1)" + proof - + have s5_1: "i \\in negNat" + using s2_3_idom by auto + have s5_2: "Succ[-.i] = -.i + 1" + using s5_1 SuccMinusIsPlusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_9: "... = i * j + (-.j + (-.i + 1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_3: "-.i + 1 \\in Int" + proof - + have s6_1: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s6_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_10: "... = i * j + ((-.j + -.i) + 1)" + proof - + have s5_1: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "-.j + (-.i + 1) = (-.j + -.i) + 1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_11: "... = i * j + ((-.i + -.j) + 1)" + proof - + have s5_1: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "-.j + -.i = -.i + -.j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "-.Succ[-.j] * -.Succ[-.i] = + i * j + ((-.i + -.j) + 1)" + proof - + have s4_1: "-.Succ[-.j] * -.Succ[-.i] = + -.Succ[-.j] * -.-.i + -.-.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_succ_minus_negnat_in_int + by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = -.Succ[-.j] * i + -.-.Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = -.Succ[-.j] * i + Succ[-.j]" + proof - + have s5_1: "Succ[-.j] \\in Int" + using s1_3_jdom negnat_succ_in_nat + nat_in_int by auto + have s5_2: "-.-.Succ[-.j] = Succ[-.j]" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_4: "... = i * -.Succ[-.j] + Succ[-.j]" + using s2_3_induct by auto + also have s4_5: "... = (i * -.-.j + -.i) + Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (i * j + -.i) + Succ[-.j]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = i * j + (-.i + Succ[-.j])" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "Succ[-.j] \\in Int" + using s1_3_jdom negnat_succ_in_nat nat_in_int by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + (-.i + (-.j + 1))" + proof - + have s5_1: "Succ[-.j] = -.j + 1" + using s1_3_jdom SuccMinusIsPlusOne by fast + show ?thesis + using s5_1 by auto + qed + also have s4_9: "... = i * j + ((-.i + -.j) + 1)" + proof - + have s5_1: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_2: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "-.i + (-.j + 1) = (-.i + -.j) + 1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: + i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: + i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_1 s2_3 neg_nat_induction by auto + show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_4: "\ j \ Nat: \ i \ Int: + i * j = j * i" + using s1_1 s1_2 natInduct by auto + have s1_5: "\ j \ negNat: \ i \ Int: + i * j = j * i" + using s1_1 s1_3 neg_nat_induction by auto + show ?thesis + using s1_4 s1_5 int_union_nat_negnat by auto + qed + + +theorem MultCommutative_sequent: + assumes idom: "i \\in Int" and jdom: "j \\in Int" + shows "i * j = j * i" + using idom jdom MultCommutative spec by auto + + +(* +THEOREM mult_0_left == + ASSUME m \in Int + PROVE 0 * m = 0 + BY mult_0, MultCommutative +*) +theorem mult_0_left: + assumes ndom: "n \\in Int" + shows "0 * n = 0" + proof - + have s1_1: "n * 0 = 0" + using ndom mult_0 by auto + have s1_2: "n * 0 = 0 * n" + using ndom zero_in_int MultCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Associativity of multiplication. *) + + +(* +THEOREM MultLeftDistributesAdd == + \A a, b, c \in Int: + a * (b + c) = a * b + a * c +PROOF +<1>1. SUFFICES + ASSUME NEW b \in Int, NEW c \in Int + PROVE \A a \in Int: a * (b + c) = a * b + a * c + BY <1>1 +<1>2. 0 * (b + c) = 0 * b + 0 * c + <2>1. 0 * (b + c) = 0 + <3>1. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, mult_0_left + <2>2. 0 * b + 0 * c = 0 + <3>1. 0 * b = 0 + BY <1>1, mult_0_left + <3>2. 0 * c = 0 + BY <1>1, mult_0_left + <3> QED + BY <3>1, <3>2, zero_in_int, add_0 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW a \in Nat, + a * (b + c) = a * b + a * c + PROVE Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c + <2>1. Succ[a] * (b + c) = (b + c) * Succ[a] + <3>1. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, MultCommutative + <2>2. @ = (b + c) * a + (b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. a \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>3. @ = a * (b + c) + (b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. a \in Int + BY <1>3, nat_in_int + <3>3. (b + c) * a = a * (b + c) + BY <3>1, <3>2, MultCommutative + <3> QED + BY <3>3 + <2>4. @ = (a * b + a * c) + (b + c) + BY <1>3 \* induction hypothesis + <2>5. @ = a * b + (a * c + (b + c)) + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>6. @ = a * b + ((a * c + b) + c) + <3>1. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. c \in Int + BY <1>1 + <3>4. a * c + (b + c) = (a * c + b) + c + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>7. @ = a * b + ((b + a * c) + c) + <3>1. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. a * c + b = b + a * c + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>8. @ = a * b + (b + (a * c + c)) + <3>1. b \in Int + BY <1>1 + <3>2. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. c \in Int + BY <1>1 + <3>4. (b + a * c) + c = b + (a * c + c) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>9. @ = (a * b + b) + (a * c + c) + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. (a * c + c) \in Int + BY <1>1, <1>3, nat_in_int, addIsInt, multIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>10. @ = (b * a + b) + (c * a + c) + <3>1. a * b = b * a + <4>1. a \in Int + BY <1>3, nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. a * c = c * a + <4>1. a \in Int + BY <1>3, nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2>11. @ = b * Succ[a] + c * Succ[a] + <3>1. b * Succ[a] = b * a + b + <4>1. b \in Int + BY <1>1 + <4>2. a \in Nat + BY <1>3 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3>2. c * Succ[a] = c * a + c + <4>1. c \in Int + BY <1>1 + <4>2. a \in Nat + BY <1>3 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3> QED + BY <3>1, <3>2 + <2>12. @ = Succ[a] * b + Succ[a] * c + <3>1. b * Succ[a] = Succ[a] * b + <4>1. b \in Int + BY <1>1 + <4>2. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. c * Succ[a] = Succ[a] * c + <4>1. c \in Int + BY <1>1 + <4>2. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, + <2>5, <2>6, <2>7, <2>8, <2>9, + <2>10, <2>11, <2>12 +<1>4. ASSUME NEW a \in negNat, + a * (b + c) = a * b + a * c + PROVE -.Succ[-.a] * (b + c) = -.Succ[-.a] * b + -.Succ[-.a] * c + <2>1. -.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a] + <3>1. -.Succ[-.a] \in Int + BY <1>4, minus_succ_minus_negnat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, MultCommutative + <2>2. @ = (b + c) * -.-.a + -.(b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>3. @ = (b + c) * a + -.(b + c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>4. @ = a * (b + c) + -.(b + c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3>3. (b + c) * a = a * (b + c) + BY <3>1, <3>2, MultCommutative + <3> QED + BY <3>3 + <2>5. @ = (a * b + a * c) + -.(b + c) + BY <1>4 \* induction hypothesis + <2>6. @ = (a * b + a * c) + (-.b + -.c) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>1 + <3>3. -.(b + c) = -.b + -.c + BY <3>1, <3>2, MinusDistributesAdd, spec + <3> QED + BY <3>3 + <2>7. @ = a * b + (a * c + (-.b + -.c)) + <3>1. a * b \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.b + -.c \in Int + <4>1. -.b \in Int + BY <1>1, neg_int_eq_int + <4>2. -.c \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>8. @ = a * b + ((a * c + -.b) + -.c) + <3>1. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. -.c \in Int + BY <1>1, neg_int_eq_int + <3>4. a * c + (-.b + -.c) = + (a * c + -.b) + -.c + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>9. @ = a * b + ((-.b + a * c) + -.c) + <3>1. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. a * c + -.b = -.b + a * c + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>10. @ = a * b + (-.b + (a * c + -.c)) + <3>1. -.b \in Int + BY <1>1, neg_int_eq_int + <3>2. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.c \in Int + BY <1>1, neg_int_eq_int + <3>4. (-.b + a * c) + -.c = + -.b + (a * c + -.c) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>11. @ = (a * b + -.b) + (a * c + -.c) + <3>1. a * b \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. a * c + -.c \in Int + <4>1. a * c \in Int + <5>1. a \in Int + BY <1>4, neg_nat_in_int + <5>2. c \in Int + BY <1>1 + <5> QED + BY <5>1, <5>2, multIsInt + <4>2. -.c \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>12. @ = (b * a + -.b) + (c * a + -.c) + <3>1. a * b = b * a + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. a * c = c * a + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2>13. @ = (b * -.-.a + -.b) + (c * -.-.a + -.c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>14. @ = b * -.Succ[-.a] + c * -.Succ[-.a] + <3>1. b * -.Succ[-.a] = b * -.-.a + -.b + <4>1. b \in Int + BY <1>1 + <4>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3>2. c * -.Succ[-.a] = c * -.-.a + -.c + <4>1. c \in Int + BY <1>1 + <4>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3> QED + BY <3>1, <3>2 + <2>15. @ = -.Succ[-.a] * b + -.Succ[-.a] * c + <3>1. -.Succ[-.a] \in Int + BY <1>4, minus_succ_minus_negnat_in_int + <3>2. b * -.Succ[-.a] = -.Succ[-.a] * b + <4>1. b \in Int + BY <1>1 + <4> QED + BY <3>1, <4>1, MultCommutative + <3>3. c * -.Succ[-.a] = -.Succ[-.a] * c + <4>1. c \in Int + BY <1>1 + <4> QED + BY <3>1, <4>1, MultCommutative + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, 2<6>, + <2>7, <2>8, <2>9, <2>10, <2>11, + <2>12, <2>13, <2>14, <2>15 +<1>5. \A a \in Nat: a * (b + c) = a * b + a * c + BY <1>2, <1>3, NatInduction +<1>6. \A a \in negNat: a * (b + c) = a * b + a * c + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MultLeftDistributesAdd_forall: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + a * (b + c) = a * b + a * c" + proof - + { + fix "b" :: "c" + fix "c" :: "c" + assume bdom: "b \\in Int" + assume cdom: "c \\in Int" + have s1_2: "0 * (b + c) = 0 * b + 0 * c" + proof - + have s2_1: "0 * (b + c) = 0" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 mult_0_left by auto + qed + have s2_2: "0 * b + 0 * c = 0" + proof - + have s3_1: "0 * b = 0" + using bdom mult_0_left by auto + have s3_2: "0 * c = 0" + using cdom mult_0_left by auto + show ?thesis + using s3_1 s3_2 zero_in_int add_0 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ a. \ + a \\in Nat; + a * (b + c) = a * b + a * c + \ \ + Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" + proof - + fix "a" :: "c" + assume s1_3_adom: "a \\in Nat" + assume s1_3_induct: "a * (b + c) = a * b + a * c" + have s2_1: "Succ[a] * (b + c) = (b + c) * Succ[a]" + proof - + have s3_1: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 MultCommutative_sequent by auto + qed + also have s2_2: "... = (b + c) * a + (b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_3: "... = a * (b + c) + (b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s3_3: "(b + c) * a = a * (b + c)" + using s3_1 s3_2 MultCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = (a * b + a * c) + (b + c)" + using s1_3_induct by auto + also have s2_5: "... = a * b + (a * c + (b + c))" + proof - + have s3_1: "a * b \\in Int" + using bdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_3: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_6: "... = a * b + ((a * c + b) + c)" + proof - + have s3_1: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "c \\in Int" + using cdom by auto + have s3_4: "a * c + (b + c) = (a * c + b) + c" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_7: "... = a * b + ((b + a * c) + c)" + proof - + have s3_1: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "a * c + b = b + a * c" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_8: "... = a * b + (b + (a * c + c))" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_3: "c \\in Int" + using cdom by auto + have s3_4: "(b + a * c) + c = b + (a * c + c)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_9: "... = (a * b + b) + (a * c + c)" + proof - + have s3_1: "a * b \\in Int" + using bdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "a * c + c \\in Int" + using cdom s1_3_adom nat_in_int addIsInt multIsInt by auto + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_10: "... = (b * a + b) + (c * a + c)" + proof - + have s3_1: "a * b = b * a" + proof - + have s4_1: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + have s3_2: "a * c = c * a" + proof - + have s4_1: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_11: "... = b * Succ[a] + c * Succ[a]" + proof - + have s3_1: "b * Succ[a] = b * a + b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + have s3_2: "c * Succ[a] = c * a + c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_12: "... = Succ[a] * b + Succ[a] * c" + proof - + have s3_1: "b * Succ[a] = Succ[a] * b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + show ?thesis + using s4_1 s4_2 MultCommutative by auto + qed + have s3_2: "c * Succ[a] = Succ[a] * c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + show ?thesis + using s4_1 s4_2 MultCommutative by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + finally show "Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" . + qed + have s1_4: "\ a. \ + a \\in negNat; + a * (b + c) = a * b + a * c + \ \ + -.Succ[-.a] * (b + c) = + -.Succ[-.a] * b + -.Succ[-.a] * c" + proof - + fix "a" :: "c" + assume s1_4_adom: "a \\in negNat" + assume s1_4_induct: "a * (b + c) = a * b + a * c" + have s2_1: "-.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a]" + proof - + have s3_1: "-.Succ[-.a] \\in Int" + using s1_4_adom minus_succ_minus_negnat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 MultCommutative by auto + qed + also have s2_2: "... = (b + c) * -.-.a + -.(b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_3: "... = (b + c) * a + -.(b + c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_4: "... = a * (b + c) + -.(b + c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_3: "(b + c) * a = a * (b + c)" + using s3_1 s3_2 MultCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = (a * b + a * c) + -.(b + c)" + using s1_4_induct by auto + also have s2_6: "... = (a * b + a * c) + (-.b + -.c)" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "c \\in Int" + using cdom by auto + have s3_3: "-.(b + c) = -.b + -.c" + using s3_1 s3_2 MinusDistributesAdd spec by auto + show ?thesis + using s3_3 by auto + qed + also have s2_7: "... = a * b + (a * c + (-.b + -.c))" + proof - + have s3_1: "a * b \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.b + -.c \\in Int" + proof - + have s4_1: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s4_2: "-.c \\in Int" + using cdom neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_8: "... = a * b + ((a * c + -.b) + -.c)" + proof - + have s3_1: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "-.c \\in Int" + using cdom neg_int_eq_int by auto + have s3_4: "a * c + (-.b + -.c) = + (a * c + -.b) + -.c" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_9: "... = a * b + ((-.b + a * c) + -.c)" + proof - + have s3_1: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "a * c + -.b = -.b + a * c" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_10: "... = a * b + (-.b + (a * c + -.c))" + proof - + have s3_1: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_2: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.c \\in Int" + using cdom neg_int_eq_int by auto + have s3_4: "(-.b + a * c) + -.c = + -.b + (a * c + -.c)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_11: "... = (a * b + -.b) + (a * c + -.c)" + proof - + have s3_1: "a * b \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "a * c + -.c \\in Int" + proof - + have s4_1: "a * c \\in Int" + proof - + have s5_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s5_2: "c \\in Int" + using cdom by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + have s4_2: "-.c \\in Int" + using cdom neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_12: "... = (b * a + -.b) + (c * a + -.c)" + proof - + have s3_1: "a * b = b * a" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + have s3_2: "a * c = c * a" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_13: "... = (b * -.-.a + -.b) + (c * -.-.a + -.c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_14: "... = b * -.Succ[-.a] + c * -.Succ[-.a]" + proof - + have s3_1: "b * -.Succ[-.a] = b * -.-.a + -.b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + have s3_2: "c * -.Succ[-.a] = c * -.-.a + -.c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_15: "... = -.Succ[-.a] * b + -.Succ[-.a] * c" + proof - + have s3_1: "-.Succ[-.a] \\in Int" + using s1_4_adom minus_succ_minus_negnat_in_int by auto + have s3_2: "b * -.Succ[-.a] = -.Succ[-.a] * b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + show ?thesis + using s3_1 s4_1 MultCommutative_sequent by auto + qed + have s3_3: "c * -.Succ[-.a] = -.Succ[-.a] * c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + show ?thesis + using s3_1 s4_1 MultCommutative_sequent by auto + qed + show ?thesis + using s3_2 s3_3 by auto + qed + finally show "-.Succ[-.a] * (b + c) = + -.Succ[-.a] * b + -.Succ[-.a] * c" . + qed + have s1_5: "\ a \ Nat: + a * (b + c) = a * b + a * c" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ a \ negNat: + a * (b + c) = a * b + a * c" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ a \ Int: + a * (b + c) = a * b + a * c" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem MultLeftDistributesAdd: + assumes adom: "a \\in Int" and bdom: "b \\in Int" and + cdom: "c \\in Int" + shows "a * (b + c) = a * b + a * c" + using adom bdom cdom MultLeftDistributesAdd_forall by auto + + +theorem MultRightDistributesAdd: + assumes adom: "a \\in Int" and bdom: "b \\in Int" and + cdom: "c \\in Int" + shows "(b + c) * a = b * a + c * a" + proof - + have s1_1: "a * (b + c) = a * b + a * c" + using adom bdom cdom MultLeftDistributesAdd by auto + have s1_2: "a * (b + c) = (b + c) * a" + using adom bdom cdom addIsInt MultCommutative by auto + have s1_3: "a * b = b * a" + using adom bdom MultCommutative by auto + have s1_4: "a * c = c * a" + using adom cdom MultCommutative by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorems mult_distributes[algebra_simps] = + MultLeftDistributesAdd MultRightDistributesAdd + + +(* +THEOREM MinusCommutesRightMult == + \A a, b \in Int: -.(a * b) = a * -.b +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int + PROVE \A b \in Int: -.(a * b) = a * -.b + BY <1>1 +<1>2. -.(a * 0) = a * -.0 + <2>1. -.(a * 0) = 0 + BY <1>1, mult_0, neg0 + <2>2. a * -.0 = 0 + BY <1>1, neg0, mult_0 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW b \in Nat, + -.(a * b) = a * -.b + PROVE -.(a * Succ[b]) = a * -.Succ[b] + <2>1. -.(a * Succ[b]) = -.(a * b + a) + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>2. @ = -.(a * b) + -.a + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. a \in Int + BY <1>1 + <3> QED + BY <3>1, <3>2, MinusDistributesAdd + <2>3. @ = a * -.b + -.a + BY <1>3 \* induction hypothesis + <2>4. @ = a * -.Succ[b] + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>4. ASSUME NEW b \in negNat, + -.(a * b) = a * -.b + PROVE -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b] + <2>1. -.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a) + <3>1. a \in Int + BY <1>1 + <3>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>2. @ = -.(a * -.-.b) + -.-.a + <3>1. a * -.-.b \in Int + <4>1. a \in Int + BY <1>1 + <4>2. -.-.b \in Int + BY <1>4, minus_negnat_in_int, + neg_int_eq_int + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.a \in Int + BY <1>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, MinusDistributesAdd + <2>3. @ = -.(a * -.-.b) + a + <3>1. a \in Int + BY <1>1 + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>4. @ = -.(a * b) + a + <3>1. b \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.b = b + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>5. @ = a * -.b + a + BY <1>4 \* induction hypothesis + <2>6. @ = a * Succ[-.b] + <3>1. a \in Int + BY <1>1 + <3>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>7. @ = a * -.-.Succ[-.b] + <3>1. Succ[-.b] \in Int + BY <1>4, negnat_succ_in_nat, nat_in_int + <3>2. -.-.Succ[-.b] = Succ[-.b] + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2> QED + BY <2>1, <2>2, <2>3, ,2>4, + <2>5, <2>6, <2>7 +<1>5. \A b \in Nat: -.(a * b) = a * -.b + BY <1>2, <1>3, NatInduction +<1>6. \A b \in negNat: -.(a * b) = a * -.b + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MinusCommutesRightMult: + "\ a \ Int: \ b \ Int: + -.(a * b) = a * -.b" + proof - + { + fix "a" :: "c" + assume adom: "a \\in Int" + have s1_2: "-.(a * 0) = a * -.0" + proof - + have s2_1: "-.(a * 0) = 0" + using adom mult_0 neg0 by auto + have s2_2: "a * -.0 = 0" + using adom neg0 mult_0 by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ b. \ + b \\in Nat; + -.(a * b) = a * -.b + \ \ + -.(a * Succ[b]) = a * -.Succ[b]" + proof - + fix "b" :: "c" + assume s1_3_bdom: "b \\in Nat" + assume s1_3_induct: "-.(a * b) = a * -.b" + have s2_1: "-.(a * Succ[b]) = -.(a * b + a)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_2: "... = -.(a * b) + -.a" + proof - + have s3_1: "a * b \\in Int" + using adom s1_3_bdom nat_in_int multIsInt by auto + have s3_2: "a \\in Int" + using adom by auto + show ?thesis + using s3_1 s3_2 MinusDistributesAdd by auto + qed + also have s2_3: "... = a * -.b + -.a" + using s1_3_induct by auto + also have s2_4: "... = a * -.Succ[b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + finally show "-.(a * Succ[b]) = a * -.Succ[b]" . + qed + have s1_4: "\ b. \ + b \\in negNat; + -.(a * b) = a * -.b + \ \ + -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" + proof - + fix "b" :: "c" + assume s1_4_bdom: "b \\in negNat" + assume s1_4_induct: "-.(a * b) = a * -.b" + have s2_1: "-.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_2: "... = -.(a * -.-.b) + -.-.a" + proof - + have s3_1: "a * -.-.b \\in Int" + proof - + have s4_1: "a \\in Int" + using adom by auto + have s4_2: "-.-.b \\in Int" + using s1_4_bdom minus_negnat_in_int + neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 MinusDistributesAdd by auto + qed + also have s2_3: "... = -.(a * -.-.b) + a" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_4: "... = -.(a * b) + a" + proof - + have s3_1: "b \\in Int" + using s1_4_bdom neg_nat_in_int by auto + have s3_2: "-.-.b = b" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_5: "... = a * -.b + a" + using s1_4_induct by auto + also have s2_6: "... = a * Succ[-.b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_7: "... = a * -.-.Succ[-.b]" + proof - + have s3_1: "Succ[-.b] \\in Int" + using s1_4_bdom negnat_succ_in_nat nat_in_int by auto + have s3_2: "-.-.Succ[-.b] = Succ[-.b]" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + finally show "-.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" . + qed + have s1_5: "\ b \ Nat: + -.(a * b) = a * -.b" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ b \ negNat: + -.(a * b) = a * -.b" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ b \ Int: + -.(a * b) = a * -.b" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +(* +THEOREM MinusCommutesLeftMult == + \A a, b \in Int: -.(a * b) = -.a * b +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE -.(a * b) = -.a * b + BY <1>1 +<1>2. -.(a * b) = -.(b * a) + BY <1>1, MultCommutative +<1>3. @ = b * -.a + BY <1>1, MinusCommutesRightMult +<1>4. @ = -.a * b + BY <1>1, neg_int_eq_int, MultCommutative +<1> QED + BY <1>2, <1>3, <1>4 +*) +theorem MinusCommutesLeftMult: + "\ a \ Int: \ b \ Int: + -.(a * b) = -.a * b" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \\in Int" + assume bdom: "b \\in Int" + have s1_2: "-.(a * b) = -.(b * a)" + using adom bdom MultCommutative_sequent by auto + also have s1_3: "... = b * -.a" + using adom bdom MinusCommutesRightMult by auto + also have s1_4: "... = -.a * b" + using adom bdom neg_int_eq_int + MultCommutative_sequent by auto + finally have "-.(a * b) = -.a * b" + by auto + } + from this show ?thesis + by auto + qed + + +(* +THEOREM MultAssociative == + \A a \in Int: \A b \in Int: \A c \in Int: + (a * b) * c = a * (b * c) +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE \A c \in Int: (a * b) * c = a * (b * c) + BY <1>1, allI +<1>2. (a * b) * 0 = a * (b * 0) + <2>1. (a * b) * 0 = 0 + <3>1. a * b \in Int + BY <1>1, multIsInt + <3> QED + BY <3>1, mult_0 + <2>2. a * (b * 0) = 0 + <3>1. b * 0 = 0 + BY <1>1, mult_0 + <3>2. a * 0 = 0 + BY <1>1, mult_0 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW c \in Nat, + (a * b) * c = a * (b * c) + PROVE (a * b) * Succ[c] = a * (b * Succ[c]) + <2>1. (a * b) * Succ[c] = (a * b) * c + (a * b) + <3>1. a * b \in Int + BY <1>1, multIsInt + <3>2. c \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>2. @ = a * (b * c) + (a * b) + BY <1>3 \* induction hypothesis + <2>3. @ = a * (b * c + b) + <3>1. a \in Int + BY <1>1 + <3>2. b * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. b \in Int + BY <1>1 + <3> QED + BY <3>1, <3>2, <3>3, MultLeftDistributesAdd + <2>4. @ = a * (b * Succ[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Nat + BY <1>3 + <3>3. b * Succ[c] = b * c + b + BY <3>1, <3>2, multint_succ_nat + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>4. ASSUME NEW c \in negNat, + (a * b) * c = a * (b * c) + PROVE (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c]) + <2>1. (a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b) + <3>1. a * b \in Int + BY <1>1, multIsInt + <3>2. -.c \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>2. @ = (a * b) * c + -.(a * b) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.c = c + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>3. @ = a * (b * c) + -.(a * b) + BY <1>4 \* induction hypothesis + <2>4. @ = a * (b * c) + a * -.b + <3>1. a \in Int + BY <1>1 + <3>2. b \in Int + BY <1>1 + <3>3. -.(a * b) = a * -.b + BY <3>1, <3>2, MinusCommutesRightMult + <3> QED + BY <3>3 + <2>5. @ = a * (b * c + -.b) + <3>1. a \in Int + BY <1>1 + <3>2. b * c \in Int + BY <1>1 <1>4, neg_nat_in_int, multIsInt + <3>3. -.b \in Int + BY <1>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, <3>3, MultLeftDistributesAdd + <2>6. @ = a * (b * -.-.c + -.b) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.c = c + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>7. @ = a * (b * -.Succ[-.c]) + <3>1. b \in Int + BY <1>1 + <3>2. -.c \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>3. b * -.Succ[-.c] = b * -.-.c + -.b + BY <3>1, <3>2, multint_succ_negnat + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, + <2>5, <2>6, <2>7 +<1>5. \A c \in Nat: (a * b) * c = a * (b * c) + BY <1>2, <1>3, NatInduction +<1>6. \A c \in negNat: (a * b) * c = a * (b * c) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MultAssociative: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + (a * b) * c = a * (b * c)" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \\in Int" + assume bdom: "b \\in Int" + have s1_2: "(a * b) * 0 = a * (b * 0)" + proof - + have s2_1: "(a * b) * 0 = 0" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + show ?thesis + using s3_1 mult_0 by auto + qed + have s2_2: "a * (b * 0) = 0" + proof - + have s3_1: "b * 0 = 0" + using bdom mult_0 by auto + have s3_2: "a * 0 = 0" + using adom mult_0 by auto + show ?thesis + using s3_1 s3_2 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ c. \ + c \\in Nat; + (a * b) * c = a * (b * c) + \ \ + (a * b) * Succ[c] = a * (b * Succ[c])" + proof - + fix "c" :: "c" + assume s1_3_cdom: "c \\in Nat" + assume s1_3_induct: "(a * b) * c = a * (b * c)" + have s2_1: "(a * b) * Succ[c] = (a * b) * c + (a * b)" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + have s3_2: "c \\in Nat" + using s1_3_cdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_2: "... = a * (b * c) + (a * b)" + using s1_3_induct by auto + also have s2_3: "... = a * (b * c + b)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b * c \\in Int" + using bdom s1_3_cdom nat_in_int multIsInt by auto + have s3_3: "b \\in Int" + using bdom by auto + show ?thesis + using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] by auto + qed + also have s2_4: "... = a * (b * Succ[c])" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "c \\in Nat" + using s1_3_cdom by auto + have s3_3: "b * Succ[c] = b * c + b" + using s3_1 s3_2 multint_succ_nat by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a * b) * Succ[c] = a * (b * Succ[c])" . + qed + have s1_4: "\ c. \ + c \\in negNat; + (a * b) * c = a * (b * c) + \ \ + (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" + proof - + fix "c" :: "c" + assume s1_4_cdom: "c \\in negNat" + assume s1_4_induct: "(a * b) * c = a * (b * c)" + have s2_1: "(a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b)" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + have s3_2: "-.c \\in Nat" + using s1_4_cdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_2: "... = (a * b) * c + -.(a * b)" + proof - + have s3_1: "c \\in Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "-.-.c = c" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_3: "... = a * (b * c) + -.(a * b)" + using s1_4_induct by auto + also have s2_4: "... = a * (b * c) + a * -.b" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "-.(a * b) = a * -.b" + using s3_1 s3_2 MinusCommutesRightMult by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = a * (b * c + -.b)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b * c \\in Int" + using bdom s1_4_cdom neg_nat_in_int multIsInt by auto + have s3_3: "-.b \\in Int" + using bdom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] + by auto + qed + also have s2_6: "... = a * (b * -.-.c + -.b)" + proof - + have s3_1: "c \\in Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "-.-.c = c" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_7: "... = a * (b * -.Succ[-.c])" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "-.c \\in Nat" + using s1_4_cdom minus_neg_nat_in_nat by auto + have s3_3: "b * -.Succ[-.c] = b * -.-.c + -.b" + using s3_1 s3_2 multint_succ_negnat by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" . + qed + have s1_5: "\ c \ Nat: + (a * b) * c = a * (b * c)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ c \ negNat: + (a * b) * c = a * (b * c)" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ c \ Int: + (a * b) * c = a * (b * c)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + + +(*----------------------------------------------------------------------------*) +(* Properties of `leq` and `add` and `diff`. *) + + +theorem LeqIsBool: + "(m <= n) \\in BOOLEAN" + unfolding leq_def by auto + + +theorem LeqReflexive: + assumes ndom: "n \\in Int" + shows "n <= n" + proof - + have s1_1: "n + 0 = n" + using ndom add_0 by auto + have s1_2: "0 \\in Nat" + using zeroIsNat by auto + have s1_3: "\ r \ Nat: n + r = n" + using s1_1 s1_2 by auto + show ?thesis + unfolding leq_def + using s1_3 by auto + qed + + +theorem LeqTransitive: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m <= n" and nk: "n <= k" + shows "m <= k" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + (* PICK t \in Nat: n + t = k *) + def Q \ "\ x. x \\in Nat \ n + x = k" + def t \ "CHOOSE x: Q(x)" + have s1_2: "t \\in Nat \ n + t = k" + proof - + have s2_1: "\\E x: Q(x)" + using nk by (auto simp: leq_def Q_def) + have s2_2: "Q(t)" + unfolding t_def + using s2_1 chooseI_ex[of Q] by auto + show ?thesis + using s2_2 by (auto simp: Q_def) + qed + have s1_3: "(m + r) + t = k" + using s1_1 s1_2 by auto + have s1_4: "m + (r + t) = k" + proof - + have s2_1: "m \\in Int" + using mdom by auto + have s2_2: "r \\in Int" + using s1_1 nat_in_int by auto + have s2_3: "t \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "(m + r) + t = m + (r + t)" + using s2_1 s2_2 s2_3 AddAssociative_sequent by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "r + t \\in Nat" + using s1_1 s1_2 nat_add_in_nat by auto + have s1_6: "\ q \ Nat: m + q = k" + using s1_4 s1_5 by auto + show ?thesis + unfolding leq_def + using s1_6 by auto + qed + + +theorem leq_0: + assumes ndom: "n \\in Int" + shows "(n \\in Nat) = (0 <= n)" + proof - + have s1_1: "n \\in Nat ==> 0 <= n" + proof - + assume s1_1_ndom: "n \\in Nat" + have s2_1: "0 + n = n" + using ndom add_0_left by auto + have s2_2: "\ r \ Nat: 0 + r = n" + using s2_1 s1_1_ndom by auto + show "0 <= n" + unfolding leq_def + using s2_2 by auto + qed + have s1_2: "0 <= n ==> n \\in Nat" + proof - + assume s1_2_leq: "0 <= n" + def P \ "\ x. x \\in Nat \ 0 + x = n" + def r \ "CHOOSE x: P(x)" + have s2_1: "r \\in Nat \ 0 + r = n" + proof - + have s3_1: "\\E x: P(x)" + using s1_2_leq by (auto simp: leq_def P_def) + have s3_2: "P(r)" + unfolding r_def + using s3_1 chooseI_ex[of P] by auto + show ?thesis + using s3_2 by (auto simp: P_def) + qed + have s2_2: "0 + r = r" + proof - + have s3_1: "r \\in Int" + using s2_1 nat_in_int by auto + show ?thesis + using s3_1 add_0_left by auto + qed + have s2_3: "r = n" + using s2_1 s2_2 by auto + show "n \\in Nat" + using s2_1 s2_3 by auto + qed + have s1_3: "(0 <= n) \\in BOOLEAN" + using LeqIsBool by auto + have s1_4: "(n \\in Nat) \\in BOOLEAN" + using inIsBool isBoolTrueFalse by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorem eq_leq_both: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m = n" + shows "m <= n \ n <= m" + proof - + have s1_1: "m <= n" + proof - + have s2_1: "m + 0 = m" + using mdom add_0 by auto + have s2_2: "m + 0 = n" + using s2_1 mn by auto + have s2_3: "0 \\in Nat" + using zeroIsNat by auto + have s2_4: "\ r \ Nat: m + r = n" + using s2_2 s2_3 by auto + show ?thesis + unfolding leq_def + using s2_4 by auto + qed + have s1_2: "n <= m" + proof - + have s2_1: "n + 0 = n" + using ndom add_0 by auto + have s2_2: "n + 0 = m" + using s2_1 mn by auto + have s2_3: "0 \\in Nat" + using zeroIsNat by auto + have s2_4: "\ r \ Nat: n + r = m" + using s2_2 s2_3 by auto + show ?thesis + unfolding leq_def + using s2_4 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* The inverse of the theorem `eq_leq_both`. *) +theorem leq_both_eq: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m <= n" and nm: "n <= m" + shows "m = n" + proof - + def P \ "\ x. x \\in Nat \ m + x = n" + def p \ "CHOOSE x: P(x)" + def Q \ "\ x. x \\in Nat \ n + x = m" + def q \ "CHOOSE x: Q(x)" + have s1_1: "p \\in Nat \ m + p = n" + proof - + have s2_1: "\ x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(p)" + unfolding p_def using s2_1 chooseI_ex[of P] + by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "q \\in Nat \ n + q = m" + proof - + have s2_1: "\ x: Q(x)" + using nm by (auto simp: leq_def Q_def) + have s2_2: "Q(q)" + unfolding q_def using s2_1 chooseI_ex[of Q] + by auto + show ?thesis + using s2_2 by (auto simp: Q_def) + qed + have s1_3: "(m + p) + (n + q) = n + m" + using s1_1 s1_2 by auto + have s1_4: "p + q = 0" + proof - + have s2_1: "(m + p) + (n + q) = n + m" + using s1_3 by auto + have s2_2: "m + (p + (n + q)) = n + m" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "p \\in Int" + using s1_1 nat_in_int by auto + have s3_3: "n + q \\in Int" + using ndom s1_2 nat_in_int addIsInt by auto + have s3_4: "(m + p) + (n + q) = + m + (p + (n + q))" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + have s2_3: "m + (p + (q + n)) = n + m" + proof - + have s3_1: "n \\in Int" + using ndom by auto + have s3_2: "q \\in Int" + using s1_2 nat_in_int by auto + have s3_3: "n + q = q + n" + using s3_1 s3_2 AddCommutative_sequent + by auto + show ?thesis + using s2_2 s3_3 by auto + qed + have s2_4: "m + ((p + q) + n) = n + m" + proof - + have s3_1: "p \\in Int" + using s1_1 nat_in_int by auto + have s3_2: "q \\in Int" + using s1_2 nat_in_int by auto + have s3_3: "n \\in Int" + using ndom by auto + have s3_4: "p + (q + n) = (p + q) + n" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_3 s3_4 by auto + qed + have s2_5: "m + (n + (p + q)) = n + m" + proof - + have s3_1: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "(p + q) + n = n + (p + q)" + using s3_1 s3_2 + AddCommutative_sequent by auto + show ?thesis + using s2_4 s3_3 by auto + qed + have s2_6: "(m + n) + (p + q) = n + m" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_4: "m + (n + (p + q)) = (m + n) + (p + q)" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_5 s3_4 by auto + qed + have s2_7: "-.(m + n) + ((m + n) + (p + q)) = + -.(m + n) + (n + m)" + using s2_6 by auto + have s2_8: "(-.(m + n) + (m + n)) + (p + q) = + -.(m + n) + (n + m)" + proof - + have s3_1: "m + n \\in Int" + using mdom ndom addIsInt by auto + have s3_2: "-.(m + n) \\in Int" + using s3_1 neg_int_eq_int by auto + have s3_3: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_4: "-.(m + n) + ((m + n) + (p + q)) = + (-.(m + n) + (m + n)) + (p + q)" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_7 s3_4 by auto + qed + have s2_9: "(-.(m + n) + (m + n)) + (p + q) = + -.(m + n) + (m + n)" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "n + m = m + n" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s2_8 s3_3 by auto + qed + have s2_10: "0 + (p + q) = 0" + proof - + have s3_1: "-.(m + n) + (m + n) = 0" + proof - + have s4_1: "m + n \\in Int" + using mdom ndom addIsInt by auto + show ?thesis + using s4_1 AddNegCancels_left by auto + qed + show ?thesis + using s2_9 s3_1 by auto + qed + have s2_11: "p + q = 0" + proof - + have s3_1: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + show ?thesis + using s2_10 s3_1 add_0_left by auto + qed + show ?thesis + using s2_11 by auto + qed + have s1_5: "p = 0 \ q = 0" + using s1_1 s1_2 s1_4 nat_add_0 by auto + have s1_6: "m + 0 = n" + using s1_1 s1_5 by auto + show ?thesis + using s1_6 mdom add_0 by auto + qed + + +theorem eq_iff_leq_both: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "(m = n) = ((m <= n) \ (n <= m))" + using mdom ndom eq_leq_both leq_both_eq by auto + + +(* +THEOREM dichotomy_leq == + ASSUME NEW a \in Int, NEW b \in Int + PROVE (a <= b) \ (b <= a) +PROOF +<1>1. CASE (a - b) \in Nat + <2>1. b + (a - b) = a + BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b + <2>2. \E n \in Nat: b + n = a + BY <1>1, <2>1 + <2>3. b <= a + BY <2>2 DEF leq + <2> QED + BY <2>3 +<1>2. CASE (a - b) \notin Nat + <2>1. -.(a - b) \in Nat + <3>1. (a - b) \in Int + <4>1. a - b = a + -.b + BY DEF diff + <4>2. a \in Int + OBVIOUS + <4>3. -.b \in Int + BY b \in Int, neg_int_eq_int + <4>4. a + -.b \in Int + BY <4>2, <4>3, addIsInt + <4> QED + BY <4>1, <4>4 + <3> QED + BY <3>1, <1>2, neg_negative_in_nat + <2>2. (b - a) \in Nat + <3>1. -.(a - b) = b - a + BY a \in Int, b \in Int, MinusDiff + <3> QED + BY <2>1, <3>1 + <2>3. a + (b - a) = b + BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b + <2>4. \E n \in Nat: a + n = b + BY <2>2, <2>3 + <2>5. a <= b + BY <2>4 DEF leq + <2> QED + BY <2>5 +<1> QED + BY <1>1, <1>2 +*) +theorem dichotomy_leq: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(a <= b) \ (b <= a)" + proof - + have s1_1: "a - b \\in Nat ==> + (a <= b) \ (b <= a)" + proof - + assume s1_1_asm: "a - b \\in Nat" + have s2_1: "b + (a - b) = a" + using adom bdom a_plus_b_minus_a_eq_b by auto + have s2_2: "\ n \ Nat: b + n = a" + using s1_1_asm s2_1 by auto + have s2_3: "b <= a" + unfolding leq_def + using s2_2 by auto + show "(a <= b) \ (b <= a)" + using s2_3 by auto + qed + have s1_2: "a - b \\notin Nat ==> + (a <= b) \ (b <= a)" + proof - + assume s1_2_asm: "a - b \\notin Nat" + have s2_1: "-.(a - b) \\in Nat" + proof - + have s3_1: "(a - b) \\in Int" + proof - + have s4_1: "a - b = a + -.b" + unfolding diff_def by auto + have s4_2: "a \\in Int" + using adom by auto + have s4_3: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s4_4: "a + -.b \\in Int" + using s4_2 s4_3 addIsInt by auto + show ?thesis + using s4_1 s4_4 by auto + qed + show ?thesis + using s3_1 s1_2_asm neg_negative_in_nat by auto + qed + have s2_2: "(b - a) \\in Nat" + proof - + have s3_1: "-.(a - b) = b - a" + using adom bdom MinusDiff by auto + show ?thesis + using s2_1 s3_1 by auto + qed + have s2_3: "a + (b - a) = b" + using adom bdom a_plus_b_minus_a_eq_b by auto + have s2_4: "\ n \ Nat: a + n = b" + using s2_2 s2_3 by auto + have s2_5: "a <= b" + unfolding leq_def + using s2_4 by auto + show "(a <= b) \ (b <= a)" + using s2_5 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem trichotomy_less: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(a < b) \ (a = b) \ (b < a)" + using adom bdom dichotomy_leq by (auto simp: less_def) + + +theorem trichotomy_less_0: + assumes ndom: "n \\in Int" + shows "(n < 0) \ (n = 0) \ (0 < n)" + using ndom zero_in_int trichotomy_less by auto + + +(* +THEOREM leq_mult_monotonic == + \A m \in Int: \A n \in Int: + (m <= n) => (\A k \in Nat: k * m <= k * n) +PROOF +<1>1. SUFFICES + ASSUME NEW m \in Int, NEW n \in Int, NEW k \in Nat, + m <= n + PROVE k * m <= k * n + BY <1>1 +<1>2. PICK r \in Nat: m + r = n + BY <1>1 DEF leq +<1>3. k * (m + r) = k * n + BY <1>2 +<1>4. k * m + k * r = k * n + <2>1. k \in Int + BY <1>1, nat_in_int + <2>2. m \in Int + BY <1>1 + <2>3. r \in Int + BY <1>2, nat_in_int + <2>4. k * (m + r) = k * m + k * r + BY <2>1, <2>2, <2>3, MultLeftDistributesAdd + <2> QED + BY <1>3, <2>4 +<1>5. k * r \in Nat + BY <1>1, <1>2, nat_mult_in_nat +<1>6. \E x \in Nat: k * m + x = k * n + BY <1>4, <1>5 +<1> QED + BY <1>6 DEF leq +*) +theorem leq_mult_monotonic: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Nat" and mn: "m <= n" + shows "k * m <= k * n" + proof - + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_2: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_3: "k * (m + r) = k * n" + using s1_2 by auto + have s1_4: "k * m + k * r = k * n" + proof - + have s2_1: "k \\in Int" + using kdom nat_in_int by auto + have s2_2: "m \\in Int" + using mdom by auto + have s2_3: "r \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "k * (m + r) = k * m + k * r" + using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "k * r \\in Nat" + using kdom s1_2 nat_mult_in_nat by auto + have s1_6: "\ x \ Nat: k * m + x = k * n" + using s1_4 s1_5 by auto + show ?thesis + unfolding leq_def + using s1_6 by auto + qed + + +(* +THEOREM leq_mult_monotonic_neg: + \A m \in Int: \A n \in Int: + (m <= n) => (\A k \in negNat: k * n <= k * m) +PROOF +<1>1. SUFFICES + ASSUME NEW m \in Int, NEW n \in Int, NEW k \in negNat, + m <= n + PROVE k * n <= k * m + BY <1>1 +<1>2. PICK r \in Nat: m + r = n + BY <1>1 DEF leq +<1>3. k * (m + r) = k * n + BY <1>2 +<1>4. k * m + k * r = k * n + <2>1. k \in Int + BY <1>1, neg_nat_in_int + <2>2. m \in Int + BY <1>1 + <2>3. r \in Int + BY <1>2, nat_in_int + <2>4. k * (m + r) = k * m + k * r + BY <2>1, <2>2, <2>3, MultLeftDistributesAdd + <2> QED + BY <1>3, <2>4 +<1>5. k * m = k * n + -.(k * r) + <2>1. (k * m + k * r) + -.(k * r) = k * n + -.(k * r) + BY <1>4 + <2>2. k * m + (k * r + -.(k * r)) = k * n + -.(k * r) + <3>1. k * m \in Int + BY <1>1, neg_nat_in_int, multIsInt + <3>2. k * r \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. r \in Int + BY <1>2, nat_in_int + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.(k * r) \in Int + BY <3>2, neg_int_eq_int + <3>4. (k * m + k * r) + -.(k * r) = + k * m + (k * r + -.(k * r)) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <2>1, <3>4 + <2>3. k * r + -.(k * r) = 0 + <3>1. k * r \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. r \in Int + BY <1>2, nat_in_int + <4> QED + BY <4>1, <4>2, multIsInt + <3> QED + BY <3>1, AddNegCancels + <2>4. k * m + 0 = k * n + -.(k * r) + BY <2>2, <2>3 + <2>5. k * m + 0 = k * m + <3>1. k * m \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. m \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3> QED + BY <3>1, add_0 + <2> QED + BY <2>4, <2>5 +<1>6. k * n + (-.k * r) = k * m + <2>1. k * n + -.(k * r) = k * m + BY <1>5 \* symmetric + <2>2. -.(k * r) = -.k * r + <3>1. k \in Int + BY <1>1, neg_nat_in_int + <3>2. r \in Int + BY <1>2, nat_in_int + <3> QED + BY <3>1, <3>2, MinusCommutesLeftMult + <2> QED + BY <2>1, <2>2 +<1>7. -.k * r \in Nat + <2>1. -.k \in Nat + BY <1>1, minus_neg_nat_in_nat + <2>2. r \in Nat + BY <1>2 + <2> QED + BY <2>1, <2>2, nat_mult_in_nat +<1>8. \E x \in Nat: k * n + x = k * m + BY <1>6, <1>7 +<1> QED + BY <1>8 DEF leq +*) +theorem leq_mult_monotonic_neg: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in negNat" and mn: "m <= n" + shows "k * n <= k * m" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_2: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_3: "k * (m + r) = k * n" + using s1_2 by auto + have s1_4: "k * m + k * r = k * n" + proof - + have s2_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s2_2: "m \\in Int" + using mdom by auto + have s2_3: "r \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "k * (m + r) = k * m + k * r" + using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "k * m = k * n + -.(k * r)" + proof - + have s2_1: "(k * m + k * r) + -.(k * r) = k * n + -.(k * r)" + using s1_4 by auto + have s2_2: "k * m + (k * r + -.(k * r)) = k * n + -.(k * r)" + proof - + have s3_1: "k * m \\in Int" + using mdom kdom neg_nat_in_int multIsInt by auto + have s3_2: "k * r \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.(k * r) \\in Int" + using s3_2 neg_int_eq_int by auto + have s3_4: "(k * m + k * r) + -.(k * r) = + k * m + (k * r + -.(k * r))" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + have s2_3: "k * r + -.(k * r) = 0" + proof - + have s3_1: "k * r \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + show ?thesis + using s3_1 AddNegCancels by auto + qed + have s2_4: "k * m + 0 = k * n + -.(k * r)" + using s2_2 s2_3 by auto + have s2_5: "k * m + 0 = k * m" + proof - + have s3_1: "k * m \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "m \\in Int" + using mdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + show ?thesis + using s3_1 add_0 by auto + qed + show ?thesis + using s2_4 s2_5 by auto + qed + have s1_6: "k * n + (-.k * r) = k * m" + proof - + have s2_1: "k * n + -.(k * r) = k * m" + using s1_5 by auto + have s2_2: "-.(k * r) = -.k * r" + proof - + have s3_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s3_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s3_1 s3_2 MinusCommutesLeftMult by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_7: "-.k * r \\in Nat" + proof - + have s2_1: "-.k \\in Nat" + using kdom minus_neg_nat_in_nat by auto + have s2_2: "r \\in Nat" + using s1_2 by auto + show ?thesis + using s2_1 s2_2 nat_mult_in_nat by auto + qed + have s1_8: "\ x \ Nat: k * n + x = k * m" + using s1_6 s1_7 by auto + show ?thesis + unfolding leq_def + using s1_8 by auto + qed + + +(*----------------------------------------------------------------------------*) + + +(* Monotonicity of addition with respect to order. *) +theorem leq_add_monotonic_right: + assumes mdom: "m \ Int" and ndom: "n \ Int" and + kdom: "k \ Int" and mn: "m <= n" + shows "m + k <= n + k" + proof - + have s1_1: "\ r: r \ Nat \ m + r = n" + using mn by (auto simp: leq_def) + def r \ "CHOOSE r: r \ Nat \ m + r = n" + have s1_2: "r \ Nat \ m + r = n" + unfolding r_def + by (rule chooseI_ex, rule s1_1) + have s1_3: "(m + r) + k = n + k" + using s1_2 by auto + have s1_4: "(m + k) + r = n + k" + proof - + have s2_1: "(r + m) + k = n + k" + proof - + have s3_1: "m \ Int" + using mdom by auto + have s3_2: "r \ Int" + using s1_2 nat_in_int by auto + have s3_3: "m + r = r + m" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s1_3 s3_3 by auto + qed + have s2_2: "r + (m + k) = n + k" + proof - + have s3_1: "r \ Int" + using s1_2 nat_in_int by auto + have s3_2: "m \ Int" + using mdom by auto + have s3_3: "k \ Int" + using kdom by auto + have s3_4: "(r + m) + k = r + (m + k)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + show ?thesis + proof - + have s3_1: "r \ Int" + using s1_2 nat_in_int by auto + have s3_2: "m + k \ Int" + using mdom kdom addIsInt by auto + have s3_3: "r + (m + k) = (m + k) + r" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s2_2 s3_3 by auto + qed + qed + have s1_5: "\ t \ Nat: + (m + k) + t = n + k" + using s1_2 s1_4 by auto + show "m + k <= n + k" + using s1_5 by (auto simp: leq_def) + qed + + +theorem eq_add_inj_right: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m + k = n + k" + shows "m = n" + proof - + have s1_1: "(m + k) + -.k = (n + k) + -.k" + using mn by auto + have s1_2: "m + (k + -.k) = n + (k + -.k)" + proof - + have minus_kdom: "-.k \\in Int" + using kdom neg_int_eq_int by auto + have s2_1: "(m + k) + -.k = m + (k + -.k)" + using mdom kdom minus_kdom AddAssociative_sequent + by auto + have s2_2: "(n + k) + -.k = n + (k + -.k)" + using ndom kdom minus_kdom AddAssociative_sequent + by auto + show ?thesis + using s1_1 s2_1 s2_2 by auto + qed + have s1_3: "m + 0 = n + 0" + proof - + have s2_1: "k + -.k = 0" + using kdom AddNegCancels_sequent by auto + show ?thesis + using s1_2 s2_1 by auto + qed + have s1_4: "m + 0 = m" + using mdom add_0 by auto + have s1_5: "n + 0 = n" + using ndom add_0 by auto + show ?thesis + using s1_3 s1_4 s1_5 by auto + qed + + +theorem less_add_monotonic_right: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m < n" + shows "m + k < n + k" + proof - + have s1_1: "m + k <= n + k" + proof - + have s2_1: "m <= n" + using mn by (auto simp: less_def) + show ?thesis + using mdom ndom kdom s2_1 leq_add_monotonic_right by auto + qed + have s1_2: "m + k \ n + k" + proof - + have s2_1: "m + k = n + k ==> FALSE" + proof - + assume s2_1_asm: "m + k = n + k" + have s3_1: "m = n" + using s2_1_asm mdom ndom kdom eq_add_inj_right + by auto + have s3_2: "m \ n" + using mn by (auto simp: less_def) + show "FALSE" + using s3_1 s3_2 by auto + qed + show ?thesis + using s2_1 by auto + qed + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_add_monotonic_left: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m < n" + shows "k + m < k + n" + proof - + have s1_1: "m + k < n + k" + using mdom ndom kdom mn less_add_monotonic_right by auto + have s1_2: "m + k = k + m" + using mdom kdom AddCommutative_sequent by auto + have s1_3: "n + k = k + n" + using ndom kdom AddCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed + + +theorem minus_less: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m < n" + shows "-.n < -.m" + proof - + have s1_1: "m + -.m < n + -.m" + proof - + have s2_1: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using mdom s2_1 ndom mn less_add_monotonic_right by auto + qed + have s1_2: "0 < n + -.m" + proof - + have s2_1: "m + -.m = 0" + using mdom AddNegCancels_sequent by auto + show ?thesis + using s1_1 s2_1 by auto + qed + have s1_3: "-.n + 0 < -.n + (n + -.m)" + proof - + have s2_1: "-.n \\in Int" + using ndom neg_int_eq_int by auto + have s2_2: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s2_3: "n + -.m \\in Int" + proof - + have s3_1: "n \\in Int" + using ndom by auto + have s3_2: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 addIsInt by auto + qed + show ?thesis + using s1_2 s2_1 s2_2 s2_3 less_add_monotonic_left + by auto + qed + have s1_4: "-.n < -.n + (n + -.m)" + proof - + have s2_1: "-.n + 0 = -.n" + using ndom neg_int_eq_int add_0 by auto + show ?thesis + using s1_3 s2_1 by auto + qed + have s1_5: "-.n + (n + -.m) = -.m" + proof - + have s2_1: "-.n + (n + -.m) = (-.n + n) + -.m" + using ndom mdom neg_int_eq_int AddAssociative_sequent + by auto + have s2_2: "-.n + n = 0" + using ndom AddNegCancels_left by auto + have s2_3: "(-.n + n) + -.m = 0 + -.m" + using s2_2 by auto + have s2_4: "0 + -.m = -.m" + proof - + have s3_1: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s3_1 add_0_left by auto + qed + show ?thesis + using s2_1 s2_3 s2_4 by auto + qed + show ?thesis + using s1_4 s1_5 by auto + qed + + +theorem leq_diff_add: + assumes ndom: "n \ Int" and mdom: "m \ Int" and + kdom: "k \ Int" and nmk: "n - m <= k" + shows "n <= k + m" + proof - + def P \ "\ x. x \\in Nat \ (n - m) + x = k" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ (n - m) + r = k" + proof - + have s2_1: "\\E x: P(x)" + using nmk by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "m + ((n - m) + r) = m + k" + using s1_1 by auto + have s1_3: "(m + -.m) + (n + r) = k + m" + proof - + have s2_1: "m + ((n - m) + r) = (m + -.m) + (n + r)" + proof - + have s3_1: "m + ((n - m) + r) = + m + ((n + -.m) + r)" + unfolding diff_def by auto + also have s3_2: "... = m + ((-.m + n) + r)" + using mdom neg_int_eq_int ndom + AddCommutative_sequent by auto + also have s3_3: "... = (m + (-.m + n)) + r" + using mdom neg_int_eq_int ndom s1_1 nat_in_int + addIsInt AddAssociative_sequent by auto + also have s3_4: "... = ((m + -.m) + n) + r" + using mdom neg_int_eq_int ndom + AddAssociative_sequent by auto + also have s3_5: "... = (m + -.m) + (n + r)" + using mdom neg_int_eq_int addIsInt ndom + s1_1 nat_in_int AddAssociative_sequent by auto + finally show ?thesis . + qed + have s2_2: "m + k = k + m" + using mdom kdom AddCommutative_sequent by auto + show ?thesis + using s1_2 s2_1 s2_2 by auto + qed + have s1_4: "n + r = k + m" + proof - + have s2_1: "m + -.m = 0" + using mdom AddNegCancels_sequent by auto + have s2_2: "(m + -.m) + (n + r) = n + r" + proof - + have s3_1: "(m + -.m) + (n + r) = 0 + (n + r)" + using s2_1 by auto + have s3_2: "0 + (n + r) = n + r" + proof - + have s4_1: "n + r \\in Int" + using ndom s1_1 nat_in_int addIsInt by auto + show ?thesis + using s4_1 add_0_left by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + show ?thesis + using s1_3 s2_2 by auto + qed + have s1_5: "\ x \ Nat: n + x = k + m" + using s1_1 s1_4 by auto + show ?thesis + unfolding leq_def + using s1_5 by auto + qed + + +theorem leq_diff_nat: + assumes mdom: "m \\in Int" and ndom: "n \\in Nat" + shows "m - n <= m" + proof - + have s1_1: "(m - n) + n = m" + proof - + have s2_1: "-.n + n = 0" + proof - + have s3_1: "n \\in Int" + using ndom nat_in_int by auto + show ?thesis + using s3_1 AddNegCancels_left by auto + qed + have s2_2: "m + (-.n + n) = m + 0" + using s2_1 by auto + have s2_3: "(m + -.n) + n = m" + proof - + have s3_1: "m + (-.n + n) = (m + -.n) + n" + proof - + have s4_1: "m \\in Int" + using mdom by auto + have s4_2: "-.n \\in Int" + using ndom minus_nat_in_int by auto + have s4_3: "n \\in Int" + using ndom nat_in_int by auto + show ?thesis + using s4_1 s4_2 s4_3 AddAssociative_sequent by auto + qed + have s3_2: "m + 0 = m" + using mdom add_0 by auto + show ?thesis + using s2_2 s3_1 s3_2 by auto + qed + show ?thesis + unfolding diff_def + using s2_3 by auto + qed + have s1_2: "\ r \ Nat: (m - n) + r = m" + using ndom s1_1 by auto + show ?thesis + unfolding leq_def + using s1_2 by auto + qed + + +(* See also the theorem `less_leq_trans`. *) +theorem leq_less_trans: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + and kdom: "k \\in Int" and + mn: "m <= n" and nk: "n < k" + shows "m < k" + proof - + have s1_1: "(n <= k) \ (n \ k)" + using nk by (auto simp: less_def) + have s1_2: "m <= k" + proof - + have s2_2: "n <= k" + using s1_1 by auto + have s2_3: "m <= n" + using mn by auto + show ?thesis + using s2_2 s2_3 mdom ndom kdom LeqTransitive + by iprover + qed + have s1_3: "m \ k" + proof - + { + assume s3_1: "m = k" + have s3_2: "k <= n" + using mn s3_1 by auto + have s3_3: "n <= k" + using nk by (auto simp: less_def) + have s3_4: "n = k" + using s3_2 s3_3 ndom kdom leq_both_eq + by auto + have "FALSE" + using s3_4 s1_1 by auto + } + from this show "m \ k" + by auto + qed + show ?thesis + unfolding less_def + using s1_2 s1_3 by auto + qed + + +(* See also the theorem `leq_less_trans`. *) +theorem less_leq_trans: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + and kdom: "k \\in Int" and + mn: "m < n" and nk: "n <= k" + shows "m < k" + proof - + have s1_1: "(m <= n) \ (m \ n)" + using mn by (auto simp: less_def) + have s1_2: "m <= k" + proof - + have s2_1: "m <= n" + using s1_1 by auto + have s2_2: "n <= k" + using nk by auto + show ?thesis + using s2_1 s2_2 mdom ndom kdom LeqTransitive + by iprover + qed + have s1_3: "m \ k" + proof - + have s2_1: "m = k ==> FALSE" + proof - + assume s2_1_asm: "m = k" + have s3_1: "n <= m" + using nk s2_1_asm by auto + have s3_2: "m <= n" + using mn by (auto simp: less_def) + have s3_3: "m = n" + using s3_1 s3_2 mdom ndom leq_both_eq + by auto + show "FALSE" + using s3_3 s1_1 by auto + qed + show ?thesis + using s2_1 by auto + qed + show ?thesis + unfolding less_def + using s1_2 s1_3 by auto + qed + + +theorem less_not_leq: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "(m < n) = (~ (n <= m))" + proof - + have s1_1: "m < n ==> ~ (n <= m)" + proof - + assume mn: "m < n" + have s2_1: "(m <= n) \ (m \ n)" + using mn by (auto simp: less_def) + have s2_2: "n <= m ==> FALSE" + proof - + assume nm: "n <= m" + have s3_1: "m = n" + using mdom ndom s2_1 nm leq_both_eq by blast + show "FALSE" + using s2_1 s3_1 by auto + qed + show "~ (n <= m)" + using s2_2 by auto + qed + have s1_2: "(~ (n <= m)) ==> m < n" + proof - + assume nm: "~ (n <= m)" + have s2_1: "(m <= n) \ (n <= m)" + using mdom ndom dichotomy_leq by auto + have s2_2: "m <= n" + using nm s2_1 by auto + have s2_3: "m = n ==> FALSE" + proof - + assume eq: "m = n" + have s3_1: "n <= n" + using ndom LeqReflexive by auto + have s3_2: "n <= m" + using eq s3_1 by auto + show "FALSE" + using nm s3_2 by auto + qed + have s2_4: "m \ n" + using s2_3 by auto + show "m < n" + unfolding less_def + using s2_2 s2_4 by auto + qed + have s1_3: "(~ (n <= m)) \\in BOOLEAN" + by auto + have s1_4: "(m < n) \\in BOOLEAN" + unfolding less_def by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorem leq_linear: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(~ (a <= b)) \ (a = b) \ (~ (b <= a))" + using adom bdom trichotomy_less[of "a" "b"] + less_not_leq[of "a" "b"] + less_not_leq[of "b" "a"] by auto + + +theorem leq_geq_linear: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(~ (b >= a)) \ (a = b) \ (~ (b <= a))" + using adom bdom leq_linear by auto + + +(* +THEOREM less_is_leq_plus_one == + ASSUME NEW m \in Int, NEW n \in Int, m < n + PROVE m + 1 <= n +PROOF +<1>1. PICK r \in Nat: m + r = n + BY m < n DEF less, leq +<1>2. PICK k \in Nat: r = Succ[k] + <2>1. r = 0 \/ \E k \in Nat: r = Succ[k] + BY <1>1, nat_0_succ + <2>2. ASSUME r = 0 + PROVE FALSE + <3>1. m + 0 = n + BY <1>1, <2>2 + <3>2. m = n + BY <3>1, m \in Int, add_0 + <3>3. m # n + BY m < n DEF less + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2 +<1>3. m + Succ[k] = n + BY <1>1, <1>2 +<1>4. m + (k + 1) = n + <2>1. Succ[k] = k + 1 + BY <1>2, nat_add_1 + <2> QED + BY <1>3, <2>1 +<1>5. m + (1 + k) = n + <2>1. k \in Int + BY <1>2, nat_in_int + <2>2. 1 \in Int + BY oneIsNat, nat_in_int + <2>3. k + 1 = 1 + k + BY <2>1, <2>2, AddCommutative + <2> QED + BY <1>4, <2>3 +<1>6. (m + 1) + k = n + <2>1. m \in Int + OBVIOUS + <2>2. 1 \in Int + BY oneIsNat nat_in_int + <2>3. k \in Int + BY <1>2, nat_in_int + <2>4. m + (1 + k) = (m + 1) + k + BY <2>1, <2>2, <2>3, AddAssociative + <2> QED + BY <1>5, <2>4 +<1>7. \E q \in Nat: (m + 1) + q = n + BY <1>2, <1>6 +<1> QED + BY <1>7 DEF leq +*) +theorem less_is_leq_plus_one: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m < n" + shows "m + 1 <= n" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: less_def leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + (* PICK k \in Nat: r = Succ[k] *) + def Q \ "\ x. x \\in Nat \ r = Succ[x]" + def k \ "CHOOSE x: Q(x)" + have s1_2: "k \\in Nat \ r = Succ[k]" + proof - + have s2_1: "r = 0 \ + (\ x \ Nat: r = Succ[x])" + using s1_1 nat_0_succ by auto + have s2_2: "r = 0 ==> FALSE" + proof - + assume s2_2_asm: "r = 0" + have s3_1: "m + 0 = n" + using s1_1 s2_2_asm by auto + have s3_2: "m = n" + using s3_1 mdom add_0 by auto + have s3_3: "m \ n" + using mn by (auto simp: less_def) + show ?thesis + using s3_2 s3_3 by auto + qed + have s2_3: "\ x \ Nat: r = Succ[x]" + using s2_1 s2_2 by auto + have s2_4: "\\E x: Q(x)" + unfolding Q_def using s2_3 by auto + have s2_5: "Q(k)" + unfolding k_def + using s2_4 chooseI_ex[of Q] by auto + show ?thesis + using s2_5 by (auto simp: Q_def) + qed + have s1_3: "m + Succ[k] = n" + using s1_1 s1_2 by auto + have s1_4: "m + (k + 1) = n" + proof - + have s2_1: "Succ[k] = k + 1" + using s1_2 nat_add_1 by fast + show ?thesis + using s1_3 s2_1 by auto + qed + have s1_5: "m + (1 + k) = n" + proof - + have s2_1: "k \\in Int" + using s1_2 nat_in_int by auto + have s2_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s2_3: "k + 1 = 1 + k" + using s2_1 s2_2 AddCommutative_sequent by auto + show ?thesis + using s1_4 s2_3 by auto + qed + have s1_6: "(m + 1) + k = n" + proof - + have s2_1: "m \\in Int" + using mdom by auto + have s2_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s2_3: "k \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "m + (1 + k) = (m + 1) + k" + using s2_1 s2_2 s2_3 AddAssociative_sequent by auto + show ?thesis + using s1_5 s2_4 by auto + qed + have s1_7: "\ q \ Nat: (m + 1) + q = n" + using s1_2 s1_6 by auto + show ?thesis + unfolding leq_def using s1_7 by auto + qed + + +(* +THEOREM less_succ_nat == + ASSUME NEW n \in Nat + PROVE n < Succ[n] +PROOF +<1>1. n <= Succ[n] + <2>1. Succ[n] = n + 1 + BY n \in Nat, nat_add_1 + <2>2. 1 \in Nat + BY oneIsNat + <2>3. \E r \in Nat: n + r = Succ[n] + BY <2>1, <2>2 + <2> QED + BY <2>3 DEF leq +<1>2. n # Succ[n] + BY n \in Nat, succIrrefl +<1> QED + BY <1>1, <1>2 DEF less +*) +theorem less_succ_nat: + assumes ndom: "n \\in Nat" + shows "n < Succ[n]" + proof - + have s1_1: "n <= Succ[n]" + proof - + have s2_1: "Succ[n] = n + 1" + using ndom nat_add_1 by fast + have s2_2: "1 \\in Nat" + using oneIsNat by auto + have s2_3: "\\E r \\in Nat: n + r = Succ[n]" + using s2_1 s2_2 by auto + show ?thesis + unfolding leq_def + using s2_3 by auto + qed + have s1_2: "n \ Succ[n]" + using ndom succIrrefl by auto + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_pred_nat: + assumes ndom: "n \\in Nat \ {0}" + shows "Pred[n] < n" + proof - + have s1_1: "Pred[n] <= n" + proof - + have s2_1: "Pred[n] \\in Nat" + using ndom Pred_in_nat by auto + have s2_2: "Succ[Pred[n]] = n" + using ndom Succ_Pred_nat by auto + have s2_3: "Pred[n] + 1 = n" + proof - + have s3_1: "Succ[Pred[n]] = Pred[n] + 1" + using s2_1 nat_add_1 by fast + show ?thesis + using s2_2 s3_1 by auto + qed + have s2_4: "1 \\in Nat" + using oneIsNat by auto + have s2_5: "\\E x \\in Nat: Pred[n] + x = n" + using s2_3 s2_4 by auto + show ?thesis + unfolding leq_def + using s2_5 by auto + qed + have s1_2: "Pred[n] \ n" + using ndom pred_irrefl by auto + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_isucc: + assumes idom: "i \\in Int" + shows "i < iSucc[i]" + proof - + have s1_1: "i \\in Nat ==> i < iSucc[i]" + proof - + assume s1_1_asm: "i \\in Nat" + have s2_1: "i < Succ[i]" + using s1_1_asm less_succ_nat by auto + have s2_2: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using idom s1_1_asm by auto + show "i < iSucc[i]" + using s2_1 s2_2 by auto + qed + have s1_2: "i \\notin Nat ==> i < iSucc[i]" + proof - + assume s1_2_asm: "i \\notin Nat" + have s2_1: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using idom s1_2_asm by auto + have s2_2: "-.i \\in Nat \ {0}" + using idom s1_2_asm minus_neg_int_in_nat by auto + have s2_3: "Pred[-.i] < -.i" + using s2_2 less_pred_nat by auto + have s2_4: "-.-.i < -.Pred[-.i]" + proof - + have s3_1: "Pred[-.i] \\in Int" + using s2_2 Pred_in_nat nat_in_int by auto + have s3_2: "-.i \\in Int" + using s2_2 nat_in_int by auto + show ?thesis + using s2_3 s3_1 s3_2 minus_less by auto + qed + have s2_5: "-.-.i = i" + using idom minus_involutive by auto + show "i < iSucc[i]" + using s2_1 s2_4 s2_5 by auto + qed + show ?thesis + using s1_1 s1_2 by iprover + qed + + +theorem less_i_add_1: + assumes idom: "i \\in Int" + shows "i < i + 1" + proof - + have s1_1: "i < iSucc[i]" + using idom less_isucc by auto + have s1_2: "iSucc[i] = i + 1" + using idom isucc_eq_add_1 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM i_less_j_nat == + ASSUME NEW i \in Int, NEW j \in Nat, i <= j + PROVE i < Succ[j] +PROOF +<1>1. j < Succ[j] + BY j \in Nat, less_succ_nat +<1>2. j \in Int + BY j \in Nat, nat_in_int +<1>3. Succ[j] \in Int + BY j \in Nat, succIsNat, nat_in_int +<1> QED + BY i \in Int, <1>2, <1>3, i <= j, <1>1, + leq_less_trans +*) +theorem i_less_succ_j_nat: + assumes idom: "i \\in Int" and jdom: "j \\in Nat" and + ij: "i <= j" + shows "i < Succ[j]" + proof - + have s1_1: "j < Succ[j]" + using jdom less_succ_nat by auto + have s1_2: "j \\in Int" + using jdom nat_in_int by auto + have s1_3: "Succ[j] \\in Int" + using jdom succIsNat nat_in_int by auto + show ?thesis + using idom s1_2 s1_3 ij s1_1 leq_less_trans by auto + qed + + +theorem less_add_1: + assumes idom: "i \\in Int" and jdom: "j \\in Int" and + ij: "i <= j" + shows "i < j + 1" + proof - + have s1_1: "j < j + 1" + using jdom less_i_add_1 by auto + have s1_2: "j + 1 \\in Int" + proof - + have s2_1: "1 \\in Int" + using oneIsNat nat_in_int by auto + show ?thesis + using jdom s2_1 addIsInt by auto + qed + show ?thesis + using idom jdom s1_2 ij s1_1 leq_less_trans by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Intervals of integers. *) + + +(* +THEOREM SplitNat0 == + Nat = {0} \cup {n + 1: n \in Nat} +PROOF +<1>1. ASSUME NEW n \in Nat + PROVE n \in {0} \cup {k + 1: k \in Nat} + <2>1. n = 0 \/ \E m \in Nat: n = Succ[m] + BY <1>1, nat_0_succ + <2>2. (\E m \in Nat: n = Succ[m]) <=> + n \in {Succ[m]: m \in Nat} + OBVIOUS + <2>3. ASSUME NEW m \in Nat + PROVE Succ[m] = m + 1 + BY <2>3, nat_add_1 + <2>4. {Succ[m]: m \in Nat} = + {m + 1: m \in Nat} + BY <2>3 + <2>5. n \in {0} \/ n \in {m + 1: m \in Nat} + BY <2>1, <2>2, <2>4 + <2> QED + BY <2>5 +<1>2. ASSUME NEW n \in {0} \cup {k + 1: k \in Nat} + PROVE n \in Nat + <2>1. CASE n = 0 + BY <2>1, zeroIsNat + <2>2. CASE n \in {k + 1: k \in Nat} + <3>1. PICK k \in Nat: n = k + 1 + BY <2>2 + <3>2. k + 1 \in Nat + BY <3>1, oneIsNat, nat_add_in_nat + <3> QED + BY <3>1, <3>2 + <2> QED + BY <1>2, <2>1, <2>2 +<1> QED + BY <1>1, <1>2 +*) +theorem SplitNat0: + "Nat = {0} \ {n + 1: n \\in Nat}" + proof - + have s1_1: "\\A n \\in Nat: n \\in {0} \ {k + 1: k \\in Nat}" + proof - + { + fix "n" :: "c" + assume s1_1_asm: "n \\in Nat" + have s2_1: "n = 0 \ (\\E m \\in Nat: n = Succ[m])" + using s1_1_asm nat_0_succ by auto + have s2_2: "\\E m \\in Nat: n = Succ[m] ==> + n \\in {Succ[m]: m \\in Nat}" + by auto + have s2_3: "{Succ[m]: m \\in Nat} = {m + 1: m \\in Nat}" + proof - + have s3_1: "\\A y \\in {Succ[m]: m \\in Nat}: + y \\in {m + 1: m \\in Nat}" + proof - + { + fix "y" :: "c" + assume s4_1: "y \\in {Succ[m]: m \\in Nat}" + have s4_2: "\\E m \\in Nat: y = Succ[m]" + using s4_1 setOfAll by auto + have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" + using s4_2 by (auto simp: bEx_def) + (* PICK r \in Nat: y = Succ[r] *) + def P \ "\ x. x \\in Nat \ y = Succ[x]" + def r \ "CHOOSE x: P(x)" + have s4_4: "r \\in Nat \ y = Succ[r]" + proof - + have s5_1: "\\E x: P(x)" + using s4_3 by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_5: "r \\in Nat \ y = r + 1" + proof - + have s5_1: "r \\in Nat" + using s4_4 by auto + have s5_2: "Succ[r] = r + 1" + using s5_1 nat_add_1 by fast + have s5_3: "y = Succ[r]" + using s4_4 by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + have "y \\in {m + 1: m \\in Nat}" + using s4_5 setOfAll_eqI by auto + } + from this show ?thesis + by auto + qed + have s3_2: "\\A y \\in {m + 1: m \\in Nat}: + y \\in {Succ[m]: m \\in Nat}" + proof - + { + fix "y" :: "c" + assume s4_1: "y \\in {m + 1: m \\in Nat}" + have s4_2: "\\E m \\in Nat: y = m + 1" + using s4_1 setOfAll by auto + have s4_3: "\\E m: m \\in Nat \ y = m + 1" + using s4_2 by (auto simp: bEx_def) + (* PICK r \in Nat: y = r + 1 *) + def P \ "\ x. x \\in Nat \ y = x + 1" + def r \ "CHOOSE x: P(x)" + have s4_4: "r \\in Nat \ y = r + 1" + proof - + have s5_1: "\\E x: P(x)" + using s4_3 by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_5: "r \\in Nat \ y = Succ[r]" + proof - + have s5_1: "r \\in Nat" + using s4_4 by auto + have s5_2: "Succ[r] = r + 1" + using s5_1 nat_add_1 by fast + have s5_3: "y = r + 1" + using s4_4 by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + have "y \\in {Succ[m]: m \\in Nat}" + using s4_5 setOfAll_eqI by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s3_1 s3_2 extension + by (auto simp: subset_def) + qed + have s2_4: "n \\in {0} \ n \\in {m + 1: m \\in Nat}" + using s2_1 s2_2 s2_3 by auto + have s2_5: "n \\in {0} \ {k + 1: k \\in Nat}" + using s2_4 by auto + } + from this show ?thesis + by auto + qed + have s1_2: "\\A n \\in {0} \ {k + 1: k \\in Nat}: n \\in Nat" + proof - + { + fix "n" :: "c" + assume s1_2_asm: "n \\in {0} \ {k + 1: k \\in Nat}" + have s2_1: "n = 0 ==> n \\in Nat" + proof - + assume s2_1_asm: "n = 0" + show "n \\in Nat" + using s2_1_asm zeroIsNat by auto + qed + have s2_2: "n \\in {k + 1: k \\in Nat} ==> n \\in Nat" + proof - + assume s2_2_asm: "n \\in {k + 1: k \\in Nat}" + have s3_1: "\\E k \\in Nat: n = k + 1" + using s2_2_asm by auto + have s3_2: "\\E k: k \\in Nat \ n = k + 1" + using s3_1 by auto + (* PICK k \in Nat: n = k + 1 *) + def P \ "\ x. x \\in Nat \ n = x + 1" + def k \ "CHOOSE x: P(x)" + have s3_3: "k \\in Nat \ n = k + 1" + proof - + have s4_1: "\\E x: P(x)" + using s3_2 by (auto simp: P_def) + have s4_2: "P(k)" + unfolding k_def + using s4_1 chooseI_ex[of P] by auto + show ?thesis + using s4_2 by (auto simp: P_def) + qed + have s3_4: "k + 1 \\in Nat" + using s3_3 oneIsNat nat_add_in_nat by auto + show "n \\in Nat" + using s3_3 s3_4 by auto + qed + have s2_3: "n \\in Nat" + using s1_2_asm s2_1 s2_2 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_1 s1_2 extension + by (auto simp: subset_def) + qed + + +(* +THEOREM NatIsAdd0 == + Nat = {k + 0: k \in Nat} +PROOF +<1>1. Nat = {k: k \in Nat} + OBVIOUS +<1>2. ASSUME NEW k \in Nat + PROVE k + 0 = k + BY <1>2, nat_in_int, add_0 +<1>3. {k: k \in Nat} = + {k + 0: k \in Nat} + BY <1>2 +<1> QED + BY <1>1, <1>3 +*) +theorem NatIsAdd0: + "Nat = {k + 0: k \\in Nat}" + proof - + have s1_1: "Nat = {k: k \\in Nat}" + by auto + have s1_2: "{k: k \\in Nat} = {k + 0: k \\in Nat}" + using nat_in_int add_0 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM SplitAddi == + ASSUME NEW i \in Nat + PROVE {k + i: k \in Nat} = + {i} \cup {k + Succ[i]: k \in Nat} +PROOF +<1>1. n \in {k + i: k \in Nat} + = \E k \in Nat: n = k + i + OBVIOUS +<1>2. @ = \E k \in {0} \cup {q + 1: q \in Nat}: + n = k + i + BY SplitNat0 +<1>3. @ = \/ \E k \in {0}: n = k + i + \/ \E k \in {q + 1: q \in Nat}: n = k + i + OBVIOUS +<1>4. @ = \/ n = 0 + i + \/ \E k: /\ k \in {q + 1: q \in Nat} + /\ n = k + i + OBVIOUS +<1>5. @ = \/ n = i + \/ \E k: /\ \E q \in Nat: k = q + 1 + /\ n = k + i + BY i \in Nat, nat_in_int, add_0_left +<1>6. @ = \/ n = i + \/ \E k: \E q \in Nat: /\ k = q + 1 + /\ n = k + i + OBVIOUS +<1>7. @ = \/ n = i + \/ \E q \in Nat: n = (q + 1) + i + OBVIOUS +<1>8. @ = \/ n \in {i} + \/ n \in {q + (i + 1): q \in Nat} + BY AddCommutative, AddAssociative, + oneIsNat, nat_in_int +<1>9. @ = \/ n \in {i} + \/ n \in {q + Succ[i]: q \in Nat} + BY nat_add_1 +<1>10. @ = n \in {i} \cup {q + Succ[i]: q \in Nat} + OBVIOUS +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5, + <1>6, <1>7, <1>8, <1>9, <1>10 +*) +theorem SplitAddi: + assumes idom: "i \\in Nat" + shows "{k + i: k \\in Nat} = + ({i} \ {k + Succ[i]: k \\in Nat})" + proof - + { + fix "n" :: "c" + have s1_1: "n \\in {k + i: k \\in Nat} + = (\\E k \\in Nat: n = k + i)" + by auto + also have s1_2: "... = ( + \\E k \\in ({0} \ {q + 1: q \\in Nat}): + n = k + i)" + using SplitNat0 by auto + also have s1_3: "... = ( + (\\E k \\in {0}: n = k + i) \ + (\\E k \\in {q + 1: q \\in Nat}: n = k + i))" + by auto + also have s1_4: "... = ( + (n = 0 + i) \ + (\\E k: (k \\in {q + 1: q \\in Nat}) \ n = k + i))" + by auto + also have s1_5: "... = ( + (n = i) \ + (\\E k: (\\E q \\in Nat: k = q + 1) \ n = k + i))" + using idom nat_in_int add_0_left by auto + also have s1_6: "... = ( + (n = i) \ + (\\E k: \\E q \\in Nat: k = q + 1 \ n = k + i))" + by auto + also have s1_7: "... = ( + (n = i) \ + (\\E q \\in Nat: n = (q + 1) + i))" + by auto + also have s1_8: "... = ( + (n = i) \ + (\\E q \\in Nat: n = q + (i + 1)))" + using AddCommutative_sequent AddAssociative_sequent + oneIsNat idom nat_in_int by auto + also have s1_9: "... = ( + (n = i) \ + (\\E q \\in Nat: n = q + Succ[i]))" + using idom nat_add_1[symmetric] by force + also have s1_10: "... = ( + (n \\in {i}) \ + (n \\in {q + Succ[i]: q \\in Nat}))" + by auto + also have s1_11: "... = ( + n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + by auto + finally have s1_12: " + (n \\in {k + i: k \\in Nat}) + = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" . + } + from this have s1_13: "\ n. (n \\in {k + i: k \\in Nat}) + = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + by auto + have s1_14: "\ n. (n \\in {k + i: k \\in Nat}) + ==> (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + using s1_13 by auto + have s1_15: "\ n. (n \\in ({i} \ {k + Succ[i]: k \\in Nat})) + ==> (n \\in {k + i: k \\in Nat})" + using s1_13 by auto + show ?thesis + using s1_14 s1_15 + by (rule setEqualI, blast) + qed + + +(* +THEOREM LeqNatOpenInterval == + ASSUME NEW i \in Nat, + NEW j \in {k + Succ[i]: k \in Nat} + PROVE ~ (j <= i) +PROOF +<1>1. PICK k \in Nat: j = k + Succ[i] + OBVIOUS +<1>2. Succ[i] <= j + BY <1>1 DEF leq +<1>3. i < Succ[i] + BY i \in Nat, less_succ_nat +<1>4. i < j + <2>1. i \in Int + BY i \in Nat, nat_in_int + <2>2. j \in Int + BY <1>1, i \in Nat, succIsNat, nat_in_int, + addIsInt + <2>3. Succ[i] \in Int + BY i \in Nat, succIsNat, nat_in_int + <2> QED + BY <2>1, <2>2, <2>3, <1>3, <1>2, + leq_less_trans +<1>5. i \in Int + BY i \in Nat, nat_in_int +<1>6. j \in Int + BY <1>1, i \in Nat, succIsNat, nat_in_int, + addIsInt +<1> QED + BY <1>4, <1>5, <1>6, less_not_leq +*) +theorem LeqNatOpenInterval: + assumes idom: "i \\in Nat" and + jdom: "j \\in {k + Succ[i]: k \\in Nat}" + shows "~ (j <= i)" + proof - + def P \ "\ x. x \\in Nat \ j = x + Succ[i]" + def k \ "CHOOSE x: P(x)" + have s1_1: "k \\in Nat \ j = k + Succ[i]" + proof - + have s2_1: "\\E x: P(x)" + using jdom by (auto simp: P_def) + have s2_2: "P(k)" + unfolding k_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "Succ[i] <= j" + proof - + have s2_1: "k + Succ[i] = j" + using s1_1 by auto + have s2_2: "Succ[i] + k = j" + using s2_1 s1_1 idom succIsNat nat_in_int + AddCommutative_sequent by auto + have s2_3: "k \\in Nat" + using s1_1 by auto + show ?thesis + unfolding leq_def + using s2_2 s2_3 by auto + qed + have s1_3: "i < Succ[i]" + using idom less_succ_nat by auto + have s1_4: "i < j" + proof - + have s2_1: "i \\in Int" + using idom nat_in_int by auto + have s2_2: "j \\in Int" + proof - + have s3_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s3_2: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s3_3: "j = k + Succ[i]" + using s1_1 by auto + show ?thesis + using s3_1 s3_2 s3_3 addIsInt by auto + qed + have s2_3: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + show ?thesis + using s2_1 s2_2 s2_3 s1_3 s1_2 + less_leq_trans by iprover + qed + have s1_5: "i \\in Int" + using idom nat_in_int by auto + have s1_6: "j \\in Int" + proof - + have s2_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s2_2: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s2_3: "j = k + Succ[i]" + using s1_1 by auto + show ?thesis + using s2_1 s2_2 s2_3 addIsInt by auto + qed + show ?thesis + using s1_4 s1_5 s1_6 less_not_leq by auto + qed + + +(* +negNatIsAdd0 == + negNat = {-.k + -.0: k \in Nat} +PROOF +<1>1. negNat = {-.k: k \in Nat} + BY DEF negNat +<1>2. ASSUME NEW k \in Nat + PROVE -.k + -.0 = -.k + <2>1. -.k + -.0 = -.k + 0 + BY neg0 + <2>2. @ = -.k + <3>1. -.k \in Int + BY <1>2, minus_nat_in_int + <3> QED + BY <3>1, add_0 + <2> QED + BY <2>1, <2>2 +<1>3. {-.k: k \in Nat} = + (-.k + -.0: k \in Nat} + BY <1>2 +<1> QED + BY <1>1, <1>3 +*) +theorem negNatIsAdd0: + "negNat = {-.k + -.0: k \\in Nat}" + proof - + have s1_1: "negNat = {-.k: k \\in Nat}" + unfolding negNat_def by auto + have s1_2: "\ k. k \\in Nat ==> -.k + -.0 = -.k" + proof - + fix "k" :: "c" + assume s1_2_asm: "k \\in Nat" + have s2_1: "-.k + -.0 = -.k + 0" + using neg0 by auto + also have s2_2: "... = -.k" + proof - + have s3_1: "-.k \\in Int" + using s1_2_asm minus_nat_in_int by auto + show ?thesis + using s3_1 add_0 by auto + qed + finally show "-.k + -.0 = -.k" . + qed + have s1_3: "{-.k: k \\in Nat} = + {-.k + -.0: k \\in Nat}" + using s1_2 by auto + show ?thesis + using s1_1 s1_3 by auto + qed + + +theorem negNatSplitAddi: + assumes idom: "i \\in Nat" + shows "{-.k + -.i: k \\in Nat} = + ({-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + proof - + have s1_1: "\ n. (n \\in {-.k + -.i: k \\in Nat}) + = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + proof - + fix "n" :: "c" + have s2_1: "n \\in {-.k + -.i: k \\in Nat} + = (\\E k \\in Nat: n = -.k + -.i)" + by auto + also have s2_2: "... = ( + \\E k \\in ({0} \ {q + 1: q \\in Nat}): + n = -.k + -.i)" + using SplitNat0 by auto + also have s2_3: "... = ( + (\\E k \\in {0}: n = -.k + -.i) \ + (\\E k \\in {q + 1: q \\in Nat}: n = -.k + -.i) + )" + by auto + also have s2_4: "... = ( + (n = -.0 + -.i) \ + (\\E k: (k \\in {q + 1: q \\in Nat}) \ (n = -.k + -.i)))" + by auto + also have s2_5: "... = ( + (n = -.i) \ + (\\E k: (\\E q \\in Nat: k = q + 1) \ n = -.k + -.i))" + using idom minus_nat_in_int add_0_left by auto + also have s2_6: "... = ( + (n = -.i) \ + (\\E k: \\E q \\in Nat: k = q + 1 \ (n = -.k + -.i)))" + by auto + also have s2_7: "... = ( + (n = -.i) \ + (\\E q \\in Nat: n = -.(q + 1) + -.i))" + by auto + also have s2_8: "... = ( + (n = -.i) \ + (\\E q \\in Nat: n = -.q + -.Succ[i]))" + proof - + have s3_1: "\\E q \\in Nat: n = -.(q + 1) + -.i ==> + \\E q \\in Nat: n = -.q + -.Succ[i]" + proof - + assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" + def P \ "\ x. x \\in Nat \ + n = -.(x + 1) + -.i" + def r \ "CHOOSE x: P(x)" + have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" + proof - + have s5_1: "\\E x: P(x)" + using s3_1_asm by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_2: "n = (-.r + -.1) + -.i" + using s4_1 oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_3: "n = -.r + (-.i + -.1)" + using s4_2 s4_1 idom oneIsNat nat_in_int minus_nat_in_int + AddAssociative_sequent AddCommutative_sequent by auto + have s4_4: "n = -.r + -.(i + 1)" + using s4_3 idom oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_5: "n = -.r + -.Succ[i]" + proof - + have s5_1: "Succ[i] = i + 1" + using idom nat_add_1 by fast + show ?thesis + using s5_1 s4_4 by auto + qed + show "\\E q \\in Nat: n = -.q + -.Succ[i]" + using s4_1 s4_5 by auto + qed + have s3_2: "\\E q \\in Nat: n = -.q + -.Succ[i] ==> + \\E q \\in Nat: n = -.(q + 1) + -.i" + proof - + assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" + def P \ "\ x. x \\in Nat \ + n = -.x + -.Succ[i]" + def r \ "CHOOSE x: P(x)" + have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" + proof - + have s5_1: "\\E x: P(x)" + using s3_2_asm by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_2: "n = -.r + -.(i + 1)" + proof - + have s5_1: "Succ[i] = i + 1" + using idom nat_add_1 by fast + show ?thesis + using s4_1 s5_1 by auto + qed + have s4_3: "n = -.r + (-.i + -.1)" + using s4_2 idom oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_4: "n = (-.r + -.1) + -.i" + using s4_3 s4_1 idom oneIsNat minus_nat_in_int + AddAssociative_sequent AddCommutative_sequent + by auto + have s4_5: "n = -.(r + 1) + -.i" + using s4_4 s4_1 oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + show "\\E q \\in Nat: n = -.(q + 1) + -.i" + using s4_1 s4_5 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_9: "... = ( + (n \\in {-.i}) \ + (n \\in {-.k + -.Succ[i]: k \\in Nat}))" + by auto + also have s2_10: "... = ( + n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + by auto + finally show "(n \\in {-.k + -.i: k \\in Nat}) + = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" . + qed + have s1_2: "\ n. (n \\in {-.k + -.i: k \\in Nat}) + ==> (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + using s1_1 by auto + have s1_3: "\ n. (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat}) + ==> (n \\in {-.k + -.i: k \\in Nat})" + using s1_1 by auto + show ?thesis + using s1_2 s1_3 by (rule setEqualI, blast) + qed + + +theorem LeqnegNatOpenInterval: + assumes idom: "i \\in Nat" and + jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" + shows "~ (-.i <= j)" + proof - + def P \ "\ x. x \\in Nat \ j = -.x + -.Succ[i]" + def k \ "CHOOSE x: P(x)" + have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" + proof - + have s2_1: "\\E x: P(x)" + using jdom by (auto simp: P_def) + have s2_2: "P(k)" + unfolding k_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_types: "j \\in Int \ -.i \\in Int \ + -.j \\in Int \ Succ[i] \\in Int \ + -.k \\in Int \ -.Succ[i] \\in Int \ + k \\in Int \ i \\in Int" + proof - + have s2_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s2_2: "-.k \\in Int" + using s2_1 neg_int_eq_int by auto + have s2_3: "i \\in Int" + using idom nat_in_int by auto + have s2_4: "-.i \\in Int" + using s2_3 neg_int_eq_int by auto + have s2_5: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s2_6: "-.Succ[i] \\in Int" + using s2_5 neg_int_eq_int by auto + have s2_7: "j \\in Int" + proof - + have s2_1: "j = -.k + -.Succ[i]" + using s1_1 by auto + show ?thesis + using s2_1 s2_2 s2_6 addIsInt by auto + qed + have s2_8: "-.j \\in Int" + using s2_7 neg_int_eq_int by auto + show ?thesis + using s2_1 s2_2 s2_3 s2_4 s2_5 + s2_6 s2_7 s2_8 by auto + qed + have s1_2: "Succ[i] <= -.j" + proof - + have s2_1: "Succ[i] + k = -.j" + proof - + have s3_1: "j = -.k + -.Succ[i]" + using s1_1 by auto + have s3_2: "-.j + j = -.j + (-.k + -.Succ[i])" + using s3_1 by auto + have s3_3: "0 = -.j + (-.k + -.Succ[i])" + using s3_2 s1_types AddNegCancels_left by auto + have s3_4: "0 + Succ[i] = (-.j + (-.k + -.Succ[i])) + Succ[i]" + using s3_3 by auto + have s3_5: "0 + Succ[i] = ((-.j + -.k) + -.Succ[i]) + Succ[i]" + using s3_4 s1_types AddAssociative_sequent by auto + have s3_6: "0 + Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" + using s3_5 s1_types addIsInt AddAssociative_sequent by auto + have s3_7: "Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" + proof - + have s4_1: "Succ[i] \\in Int" + using s1_types by auto + have s4_2: "0 + Succ[i] = Succ[i]" + using s4_1 add_0_left by fast + show ?thesis + using s3_6 s4_2 by auto + qed + have s3_8: "Succ[i] = (-.j + -.k) + 0" + using s3_7 s1_types AddNegCancels_left by auto + have s3_9: "Succ[i] = -.j + -.k" + proof - + have s4_1: "-.j + -.k \\in Int" + using s1_types addIsInt by auto + show ?thesis + using s3_8 s4_1 add_0 by auto + qed + have s3_10: "Succ[i] + k = (-.j + -.k) + k" + using s3_9 by auto + have s3_11: "Succ[i] + k = -.j + (-.k + k)" + using s3_10 s1_types AddAssociative_sequent by auto + have s3_12: "Succ[i] + k = -.j + 0" + using s3_11 s1_types AddNegCancels_left by auto + show ?thesis + using s3_12 s1_types add_0 by auto + qed + have s2_2: "\\E r \\in Nat: Succ[i] + r = -.j" + using s1_1 s2_1 by auto + show ?thesis + unfolding leq_def + using s2_2 by auto + qed + have s1_3: "i < Succ[i]" + using idom less_succ_nat by fast + have s1_4: "i < -.j" + using s1_types s1_2 s1_3 less_leq_trans by auto + have s1_5: "j < -.i" + proof - + have s2_1: "-.-.j < -.i" + proof - + have s3_1: "i \\in Int" + using s1_types by auto + have s3_2: "-.j \\in Int" + using s1_types by auto + show ?thesis + using s3_1 s3_2 s1_4 minus_less by auto + qed + have s2_2: "-.-.j = j" + using s1_types minus_involutive by auto + show ?thesis + using s2_1 s2_2 by auto + qed + show ?thesis + using s1_5 s1_types less_not_leq by auto + qed + + +theorem split_interval_nat_negnat: + shows "{k \\in Int: a <= k \ k <= b} = + {k \\in negNat: a <= k \ k <= b} \ + {k \\in Nat: a <= k \ k <= b}" + using int_union_nat_negnat by auto + + +end From 9992e76b439d3deca061127058bd75e79668233d Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:14:12 +0200 Subject: [PATCH 007/167] API: add TLA+ syntax for bounded quantifiers --- isabelle/SetTheory.thy | 2 ++ 1 file changed, 2 insertions(+) diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index ca9e6bb4..8729d7e5 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -151,6 +151,8 @@ syntax "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) syntax (ASCII) (* Again, only single variable for bounded choice. *) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) From 42b197944a805e57e6ccb14a6989448cc0e91be0 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:26:25 +0200 Subject: [PATCH 008/167] MAI: update theory `NewIntegers` to Isabelle2020 --- isabelle/NewIntegers.thy | 234 ++++++++++++++++++++++++++------------- 1 file changed, 156 insertions(+), 78 deletions(-) diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy index 6c7ae08c..103b7221 100644 --- a/isabelle/NewIntegers.thy +++ b/isabelle/NewIntegers.thy @@ -2190,7 +2190,7 @@ theorem nat_add_in_nat: using s2_1 s2_2 s2_3 by auto qed have s1_3: "\ k \ Nat: m + k \\in Nat" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ k. m + k \\in Nat"] by auto show ?thesis using s1_3 ndom spec by auto qed @@ -2213,8 +2213,9 @@ theorem nat_add_0: assume s2_1: "n \ 0" have s2_2: "\ k: k \\in Nat \ n = Succ[k]" using ndom s2_1 nat_0_succ by auto - def P \ "\ x. x \\in Nat \ n = Succ[x]" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ n = Succ[x]" + define r where "r \ CHOOSE x: P(x)" have s2_3: "r \\in Nat \ n = Succ[r]" proof - have s3_1: "\ x: P(x)" @@ -2468,7 +2469,7 @@ theorem nat_mult_in_nat: using s2_1 s2_2 by auto qed have s1_3: "\ k \ Nat: m * k \\in Nat" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ k. m * k \\in Nat"] by auto show ?thesis using s1_3 ndom spec by auto qed @@ -3443,7 +3444,7 @@ theorem AddCommutative: show ?thesis using s4_1 s4_2 natInduct[of "\ i. 0 + i = i + 0"] - by auto + by blast qed have s3_2: "\ j. \ j \ Nat; @@ -3584,11 +3585,15 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show "\ i \ Nat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 natInduct - by auto + using s4_1 s4_2 + natInduct[of "\ i. Succ[j] + i = i + Succ[j]"] + by auto qed show ?thesis - using s3_1 s3_2 natInduct by auto + using s3_1 s3_2 + natInduct[of "\ j. + \ i \ Nat: j + i = i + j"] + by auto qed have s2_2: "\ j \ Nat: \ i \ negNat: j + i = i + j" @@ -3663,7 +3668,9 @@ theorem AddCommutative: by auto qed show ?thesis - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ Nat; @@ -3844,11 +3851,16 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show "\ i \ negNat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 neg_nat_induction - by auto + using s4_1 s4_2 + neg_nat_induction[of + "\ i. Succ[j] + i = i + Succ[j]"] + by auto qed show ?thesis - using s3_1 s3_2 natInduct by auto + using s3_1 s3_2 + natInduct[of "\ j. + \ i \ negNat: j + i = i + j"] + by auto qed show ?thesis proof - @@ -3944,7 +3956,9 @@ theorem AddCommutative: using s5_1 s5_2 s5_3 s5_4 s5_5 by auto qed show ?thesis - using s4_1 s4_2 natInduct by auto + using s4_1 s4_2 + natInduct[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ negNat; @@ -4146,10 +4160,16 @@ theorem AddCommutative: qed show "\ i \ Nat: -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 natInduct by auto + using s4_1 s4_2 + natInduct[of "\ i. + -.Succ[-.j] + i = i + -.Succ[-.j]"] + by auto qed show ?thesis - using s3_1 s3_2 neg_nat_induction by auto + using s3_1 s3_2 + neg_nat_induction[of "\ j. + \ i \ Nat: j + i = i + j"] + by auto qed have s2_2: "\ j \ negNat: \ i \ negNat: j + i = i + j" @@ -4233,7 +4253,9 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show ?thesis - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ negNat; @@ -4492,10 +4514,16 @@ theorem AddCommutative: qed show "\ i \ negNat: -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. + -.Succ[-.j] + i = i + -.Succ[-.j]"] + by auto qed show ?thesis - using s3_1 s3_2 neg_nat_induction by auto + using s3_1 s3_2 + neg_nat_induction[of "\ j. + \ i \ negNat: j + i = i + j"] + by auto qed show ?thesis proof - @@ -5125,7 +5153,10 @@ theorem iSuccRightDistributesAdd: using s3_3 s3_4 by blast qed show ?thesis - using s2_1 s2_2 neg_nat_induction by auto + using s2_1 s2_2 + neg_nat_induction[of "\ b. + iSucc[a + b] = a + iSucc[b]"] + by auto qed have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" using s1_1 s1_2 int_union_nat_negnat by auto @@ -5346,7 +5377,10 @@ theorem iPredRightDistributesAdd: finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . qed show ?thesis - using s1_1 s2_1 natInduct by auto + using s1_1 s2_1 + natInduct[of "\ b. + iPred[a + b] = a + iPred[b]"] + by auto qed have s1_3: "\ b \ negNat \ {0}: iPred[a + b] = a + iPred[b]" @@ -5718,10 +5752,14 @@ theorem AddAssociative: qed have s1_5: "\ c \ Nat: (a + b) + c = a + (b + c)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ c. + (a + b) + c = a + (b + c)"] + by auto have s1_6: "\ c \ negNat: (a + b) + c = a + (b + c)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ c. + (a + b) + c = a + (b + c)"] + by auto have "\ c \ Int: (a + b) + c = a + (b + c)" using s1_5 s1_6 int_union_nat_negnat by auto @@ -5750,7 +5788,7 @@ theorem AddLeftCommutativity: qed -theorems add_ac_int[algebra_simps] = +lemmas add_ac_int[algebra_simps] = AddCommutative_sequent AddAssociative_sequent AddLeftCommutativity add_0 add_0_left @@ -6192,9 +6230,13 @@ theorem AddNegCancels: finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . qed have s1_4: "\ i \ Nat: i + -.i = 0" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ i. + i + -.i = 0"] + by auto have s1_5: "\ i \ negNat: i + -.i = 0" - using s1_1 s1_3 neg_nat_induction by auto + using s1_1 s1_3 neg_nat_induction[of "\ i. + i + -.i = 0"] + by auto show ?thesis using s1_4 s1_5 int_union_nat_negnat by auto qed @@ -6726,9 +6768,13 @@ theorem MinusDistributesAdd: finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . qed have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ b. + -.(a + b) = (-.a) + (-.b)"] + by auto have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ b. + -.(a + b) = (-.a) + (-.b)"] + by auto have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" using s1_5 s1_6 int_union_nat_negnat by auto } @@ -7760,9 +7806,13 @@ theorem MultCommutative: using s3_1 s3_2 by auto qed have s2_4: "\ i \ Nat: i * 0 = 0 * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * 0 = 0 * i"] + by blast have s2_5: "\ i \ negNat: i * 0 = 0 * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * 0 = 0 * i"] + by blast show ?thesis using s2_4 s2_5 int_union_nat_negnat by auto qed @@ -8222,9 +8272,13 @@ theorem MultCommutative: using s3_1 s3_2 by auto qed have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * Succ[j] = Succ[j] * i"] + by auto have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * Succ[j] = Succ[j] * i"] + by auto show "\ i \ Int: i * Succ[j] = Succ[j] * i" using s2_4 s2_5 int_union_nat_negnat by auto qed @@ -8759,19 +8813,27 @@ theorem MultCommutative: qed have s2_4: "\ i \ Nat: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * -.Succ[-.j] = -.Succ[-.j] * i"] + by auto have s2_5: "\ i \ negNat: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * -.Succ[-.j] = -.Succ[-.j] * i"] + by auto show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" using s2_4 s2_5 int_union_nat_negnat by auto qed have s1_4: "\ j \ Nat: \ i \ Int: i * j = j * i" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ j. + \ i \ Int: i * j = j * i"] + by blast have s1_5: "\ j \ negNat: \ i \ Int: i * j = j * i" - using s1_1 s1_3 neg_nat_induction by auto + using s1_1 s1_3 neg_nat_induction[of "\ j. + \ i \ Int: i * j = j * i"] + by blast show ?thesis using s1_4 s1_5 int_union_nat_negnat by auto qed @@ -9637,10 +9699,14 @@ theorem MultLeftDistributesAdd_forall: qed have s1_5: "\ a \ Nat: a * (b + c) = a * b + a * c" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ a. + a * (b + c) = a * b + a * c"] + by auto have s1_6: "\ a \ negNat: a * (b + c) = a * b + a * c" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ a. + a * (b + c) = a * b + a * c"] + by auto have s1_7: "\ a \ Int: a * (b + c) = a * b + a * c" using s1_5 s1_6 int_union_nat_negnat by auto @@ -9675,7 +9741,7 @@ theorem MultRightDistributesAdd: qed -theorems mult_distributes[algebra_simps] = +lemmas mult_distributes[algebra_simps] = MultLeftDistributesAdd MultRightDistributesAdd @@ -9918,10 +9984,14 @@ theorem MinusCommutesRightMult: qed have s1_5: "\ b \ Nat: -.(a * b) = a * -.b" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ b. + -.(a * b) = a * -.b"] + by auto have s1_6: "\ b \ negNat: -.(a * b) = a * -.b" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ b. + -.(a * b) = a * -.b"] + by auto have s1_7: "\ b \ Int: -.(a * b) = a * -.b" using s1_5 s1_6 int_union_nat_negnat by auto @@ -10243,10 +10313,14 @@ theorem MultAssociative: qed have s1_5: "\ c \ Nat: (a * b) * c = a * (b * c)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ c. + (a * b) * c = a * (b * c)"] + by auto have s1_6: "\ c \ negNat: (a * b) * c = a * (b * c)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ c. + (a * b) * c = a * (b * c)"] + by auto have s1_7: "\ c \ Int: (a * b) * c = a * (b * c)" using s1_5 s1_6 int_union_nat_negnat by auto @@ -10288,8 +10362,8 @@ theorem LeqTransitive: shows "m <= k" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -10301,8 +10375,8 @@ theorem LeqTransitive: using s2_2 by (auto simp: P_def) qed (* PICK t \in Nat: n + t = k *) - def Q \ "\ x. x \\in Nat \ n + x = k" - def t \ "CHOOSE x: Q(x)" + define Q where "Q \ \ x. x \\in Nat \ n + x = k" + define t where "t \ CHOOSE x: Q(x)" have s1_2: "t \\in Nat \ n + t = k" proof - have s2_1: "\\E x: Q(x)" @@ -10356,8 +10430,8 @@ theorem leq_0: have s1_2: "0 <= n ==> n \\in Nat" proof - assume s1_2_leq: "0 <= n" - def P \ "\ x. x \\in Nat \ 0 + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ 0 + x = n" + define r where "r \ CHOOSE x: P(x)" have s2_1: "r \\in Nat \ 0 + r = n" proof - have s3_1: "\\E x: P(x)" @@ -10433,10 +10507,10 @@ theorem leq_both_eq: mn: "m <= n" and nm: "n <= m" shows "m = n" proof - - def P \ "\ x. x \\in Nat \ m + x = n" - def p \ "CHOOSE x: P(x)" - def Q \ "\ x. x \\in Nat \ n + x = m" - def q \ "CHOOSE x: Q(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define p where "p \ CHOOSE x: P(x)" + define Q where "Q \ \ x. x \\in Nat \ n + x = m" + define q where "q \ CHOOSE x: Q(x)" have s1_1: "p \\in Nat \ m + p = n" proof - have s2_1: "\ x: P(x)" @@ -10754,8 +10828,8 @@ theorem leq_mult_monotonic: kdom: "k \\in Nat" and mn: "m <= n" shows "k * m <= k * n" proof - - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_2: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -10890,8 +10964,8 @@ theorem leq_mult_monotonic_neg: shows "k * n <= k * m" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_2: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -11019,7 +11093,7 @@ theorem leq_add_monotonic_right: proof - have s1_1: "\ r: r \ Nat \ m + r = n" using mn by (auto simp: leq_def) - def r \ "CHOOSE r: r \ Nat \ m + r = n" + define r where "r \ CHOOSE r: r \ Nat \ m + r = n" have s1_2: "r \ Nat \ m + r = n" unfolding r_def by (rule chooseI_ex, rule s1_1) @@ -11231,8 +11305,8 @@ theorem leq_diff_add: kdom: "k \ Int" and nmk: "n - m <= k" shows "n <= k + m" proof - - def P \ "\ x. x \\in Nat \ (n - m) + x = k" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ (n - m) + x = k" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ (n - m) + r = k" proof - have s2_1: "\\E x: P(x)" @@ -11556,8 +11630,8 @@ theorem less_is_leq_plus_one: shows "m + 1 <= n" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -11569,8 +11643,8 @@ theorem less_is_leq_plus_one: using s2_2 by (auto simp: P_def) qed (* PICK k \in Nat: r = Succ[k] *) - def Q \ "\ x. x \\in Nat \ r = Succ[x]" - def k \ "CHOOSE x: Q(x)" + define Q where "Q \ \ x. x \\in Nat \ r = Succ[x]" + define k where "k \ CHOOSE x: Q(x)" have s1_2: "k \\in Nat \ r = Succ[k]" proof - have s2_1: "r = 0 \ @@ -11888,8 +11962,9 @@ theorem SplitNat0: have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" using s4_2 by (auto simp: bEx_def) (* PICK r \in Nat: y = Succ[r] *) - def P \ "\ x. x \\in Nat \ y = Succ[x]" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ y = Succ[x]" + define r where "r \ CHOOSE x: P(x)" have s4_4: "r \\in Nat \ y = Succ[r]" proof - have s5_1: "\\E x: P(x)" @@ -11928,8 +12003,9 @@ theorem SplitNat0: have s4_3: "\\E m: m \\in Nat \ y = m + 1" using s4_2 by (auto simp: bEx_def) (* PICK r \in Nat: y = r + 1 *) - def P \ "\ x. x \\in Nat \ y = x + 1" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ y = x + 1" + define r where "r \ CHOOSE x: P(x)" have s4_4: "r \\in Nat \ y = r + 1" proof - have s5_1: "\\E x: P(x)" @@ -11988,8 +12064,9 @@ theorem SplitNat0: have s3_2: "\\E k: k \\in Nat \ n = k + 1" using s3_1 by auto (* PICK k \in Nat: n = k + 1 *) - def P \ "\ x. x \\in Nat \ n = x + 1" - def k \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ n = x + 1" + define k where "k \ CHOOSE x: P(x)" have s3_3: "k \\in Nat \ n = k + 1" proof - have s4_1: "\\E x: P(x)" @@ -12013,7 +12090,7 @@ theorem SplitNat0: qed show ?thesis using s1_1 s1_2 extension - by (auto simp: subset_def) + by blast qed @@ -12192,8 +12269,8 @@ theorem LeqNatOpenInterval: jdom: "j \\in {k + Succ[i]: k \\in Nat}" shows "~ (j <= i)" proof - - def P \ "\ x. x \\in Nat \ j = x + Succ[i]" - def k \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ j = x + Succ[i]" + define k where "k \ CHOOSE x: P(x)" have s1_1: "k \\in Nat \ j = k + Succ[i]" proof - have s2_1: "\\E x: P(x)" @@ -12354,9 +12431,9 @@ theorem negNatSplitAddi: \\E q \\in Nat: n = -.q + -.Succ[i]" proof - assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" - def P \ "\ x. x \\in Nat \ + define P where "P \ \ x. x \\in Nat \ n = -.(x + 1) + -.i" - def r \ "CHOOSE x: P(x)" + define r where "r \ CHOOSE x: P(x)" have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" proof - have s5_1: "\\E x: P(x)" @@ -12390,9 +12467,9 @@ theorem negNatSplitAddi: \\E q \\in Nat: n = -.(q + 1) + -.i" proof - assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" - def P \ "\ x. x \\in Nat \ + define P where "P \ \ x. x \\in Nat \ n = -.x + -.Succ[i]" - def r \ "CHOOSE x: P(x)" + define r where "r \ CHOOSE x: P(x)" have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" proof - have s5_1: "\\E x: P(x)" @@ -12452,8 +12529,9 @@ theorem LeqnegNatOpenInterval: jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" shows "~ (-.i <= j)" proof - - def P \ "\ x. x \\in Nat \ j = -.x + -.Succ[i]" - def k \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ j = -.x + -.Succ[i]" + define k where "k \ CHOOSE x: P(x)" have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" proof - have s2_1: "\\E x: P(x)" From bc17079dee05c384b949a34154b99bd3b8368e96 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:35:26 +0200 Subject: [PATCH 009/167] MAI: update sectioning and text of theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 84 ++++++++++++++++++++++----------------------- 1 file changed, 42 insertions(+), 42 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 4dfb7d6b..d23d908a 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:39:46 merz> *) -header {* Tuples and Relations in \tlaplus{} *} +section \ Tuples and Relations in \tlaplus{} \ theory Tuples imports NatOrderings begin -text {* +text \ We develop a theory of tuples and relations in \tlaplus{}. Tuples are functions whose domains are intervals of the form $1 .. n$, for some natural number $n$, and relations are sets of tuples. In particular, @@ -21,15 +21,15 @@ text {* when defining functions recursively, as we have a fixed point theorem on sets but not on functions.) We also introduce standard notions for binary relations, such as orderings, equivalence relations and so on. -*} +\ -subsection {* Sequences and Tuples *} +subsection \ Sequences and Tuples \ -text {* +text \ Tuples and sequences are the same mathematical objects in \tlaplus{}, so we give elementary definitions for sequences here. Further operations on sequences require arithmetic and will be introduced in a separate theory. -*} +\ definition Seq -- {* set of finite sequences with elements from $S$ *} where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" @@ -157,12 +157,12 @@ lemma isASeqInSeq: "isASeq(s) = (\S: s \ Seq(S))" by (auto elim: seqInSeqRange) -subsection {* Sequences via @{text emptySeq} and @{text Append} *} +subsection \ Sequences via @{text emptySeq} and @{text Append} \ -text {* +text \ Sequences can be built from the constructors @{text emptySeq} (written @{text "\\"}) and Append. -*} +\ definition emptySeq ("(<< >>)") where "<< >> \ [x \ 1 .. 0 \ {}]" @@ -328,10 +328,10 @@ lemma appendEqualIff [simp]: shows "(Append(s,e) = Append(t,f)) = (s = t \ e = f)" using appendD1[OF s t] appendD2[OF s t] by auto -text {* +text \ The following lemma gives a possible alternative definition of @{text Append}. -*} +\ lemma appendExtend: assumes "isASeq(s)" @@ -348,10 +348,10 @@ lemma imageAppend [simp]: unfolding appendExtend[OF s] using assms by (auto elim!: inNatIntervalE, force+) -text {* +text \ Inductive reasoning about sequences, based on @{term "\\"} and @{text Append}. -*} +\ lemma seqInduct [case_names empty append, induct set: Seq]: assumes s: "s \ Seq(S)" @@ -411,12 +411,12 @@ lemma seqCases [case_names Empty Append, cases set: Seq]: using assms by (auto dest: seqEmptyOrAppend) -subsection {* Enumerated sequences *} +subsection \ Enumerated sequences \ -text {* +text \ We introduce the conventional syntax @{text "\a,b,c\"} for tuples and enumerated sequences, based on the above constructors. -*} +\ nonterminal tpl @@ -451,12 +451,12 @@ lemma "\0,1\ \ \1,2,3\" by simp lemma "(\a,b\ = \c,d\) = (a=c \ b=d)" by simp ***) -text {* +text \ \tlaplus{} has a form of quantification over tuples written $\exists \langle x,y,z \rangle \in S: P(x,y,z)$. We cannot give a generic definition of this form for arbitrary tuples, but introduce input syntax for tuples of length up to $5$. -*} +\ syntax "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) @@ -488,9 +488,9 @@ translations "\\x,y,z,u\ \ S : P" \ "\x,y,z,u : \x,y,z,u\ \ S \ P" "\\x,y,z,u,v\ \ S : P" \ "\x,y,z,u,v : \x,y,z,u,v\ \ S \ P" -subsection {* Sets of finite functions *} +subsection \ Sets of finite functions \ -text {* +text \ We introduce notation such as @{text "[x: S, y: T]"} to designate the set of finite functions @{text f} with domain @{text "{x,y}"} (for constants @{text x}, @{text y}) such that @{text "f[x] \ S"} and @{text "f[y] \ T"}. @@ -502,7 +502,7 @@ text {* Internally, the set is represented as @{text "EnumFuncSet(\x,y\, \S,T\)"}, using appropriate translation functions between the internal and external representations. -*} +\ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : @@ -635,14 +635,14 @@ by auto ***) -subsection {* Set product *} +subsection \ Set product \ -text {* +text \ The cartesian product of two sets $A$ and $B$ is the set of pairs whose first component is in $A$ and whose second component is in $B$. We generalize the definition of products to an arbitrary number of sets: $Product(\langle A_1,\ldots,A_n \rangle) = A_1 \times\cdots\times A_n$. -*} +\ definition Product where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : @@ -707,7 +707,7 @@ lemma "\a\ \ Product(\A,B\)" by auto ***) -text {* Special case: binary product *} +text \ Special case: binary product \ definition prod :: "c \ c \ c" (infixr "\\X" 100) where @@ -813,7 +813,7 @@ unfolding prod_def Product_def using assms by auto -subsection {* Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} *} +subsection \ Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} \ definition setOfPairs :: "[c, [c,c]\c] \ c" where "setOfPairs(R, f) \ { f(p[1], p[2]) : p \ R }" @@ -874,7 +874,7 @@ lemma setOfPairsEqual: using assms by (auto, blast+) -subsection {* Basic notions about binary relations *} +subsection \ Basic notions about binary relations \ definition rel_domain :: "c => c" where "rel_domain(r) \ { p[1] : p \ r }" @@ -896,7 +896,7 @@ definition Id :: "c \ c" -- {* diagonal: identity over a set *} where "Id(A) \ { \x,x\ : x \ A }" -text {* Properties of relations *} +text \ Properties of relations \ definition reflexive -- {* reflexivity over a set *} where "reflexive(A,r) \ \x \ A: \x,x\ \ r" @@ -953,7 +953,7 @@ lemma equivalenceIsBool[intro!,simp]: "isBool(equivalence(A,r))" unfolding isBool_def by (rule boolifyEquivalence) -subsubsection {* Domain and Range *} +subsubsection \ Domain and Range \ lemma prod_in_dom_x_ran: assumes "r \ A \ B" and "p \ r" @@ -1016,7 +1016,7 @@ lemma in_ran_imp_in_B: "r \ A \ B \ p \ rel by force -subsubsection {* Converse relation *} +subsubsection \ Converse relation \ lemmas converseEqualI = setEqualI [where A = "r^-1", standard, intro!] @@ -1123,7 +1123,7 @@ next qed -subsubsection {* Identity relation over a set *} +subsubsection \ Identity relation over a set \ lemmas idEqualI = setEqualI [where A = "Id(S)", standard, intro!] @@ -1173,7 +1173,7 @@ lemma ran_Id [simp]: "rel_range(Id(A)) = A" unfolding rel_range_def Id_def by auto -subsubsection {* Composition of relations *} +subsubsection \ Composition of relations \ lemmas compEqualI = setEqualI [where A = "r \ s", standard, intro!] @@ -1330,9 +1330,9 @@ unfolding rel_comp_def proof auto qed -subsubsection {* Properties of relations *} +subsubsection \ Properties of relations \ -text {* Reflexivity *} +text \ Reflexivity \ lemma reflI [intro]: "(\x. x \ A \ \x,x\ \ r) \ reflexive(A,r)" unfolding reflexive_def by blast @@ -1344,7 +1344,7 @@ lemma reflexive_empty (*[simp]*): "reflexive({}, {})" by auto -text {* Symmetry *} +text \ Symmetry \ lemma symmetricI: "\ \x y. \x,y\ \ r \ \y,x\ \ r \ \ symmetric(r)" unfolding symmetric_def by blast @@ -1356,7 +1356,7 @@ lemma symmetric_Int: "\ symmetric(r); symmetric(s) \ \ Antisymmetry \ lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" @@ -1372,7 +1372,7 @@ lemma antisym_empty (*[simp]*): "antisymmetric({})" by blast -text {* Transitivity *} +text \ Transitivity \ lemma transitiveI [intro]: "(\x y z. \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r) \ transitive(r)" @@ -1388,7 +1388,7 @@ lemma transitive_iff_comp_subset: "transitive(r) = (r \ r \ r)" unfolding transitive_def rel_comp_def by (auto elim!: subsetD) -text {* Irreflexivity *} +text \ Irreflexivity \ lemma irreflexiveI [intro]: "\ \x. x \ A \ \x,x\ \ r \ \ irreflexive(A,r)" unfolding irreflexive_def by blast @@ -1397,7 +1397,7 @@ lemma irreflexiveE [dest]: "\ irreflexive(A,r); x \ A \ \ Equivalence Relations \ (**************** NOT USED ANYWHERE ************************************ definition @@ -1422,11 +1422,11 @@ abbreviation --{* Abbreviation for the common case where the relations are identical *} ***************************************************************************) -text{* @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} *} +text \ @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} \ (* from Suppes, Theorem 70 *) -text {* First half: ``only if'' part *} +text \ First half: ``only if'' part \ lemma sym_trans_comp_subset: assumes "r \ A \ A" and "symmetric(r)" and "transitive(r)" @@ -1460,7 +1460,7 @@ using eq sym_trans_comp_subset[OF r] refl_comp_subset[OF r] unfolding equivalence_def by (intro setEqual, simp+) -text {* Second half: ``if'' part, needs totality of relation $r$ *} +text \ Second half: ``if'' part, needs totality of relation $r$ \ lemma comp_equivI: assumes dom: "rel_domain(r) = A" and r: "r \ A \ A" and comp: "r^-1 \ r = r" From b7727e65e758b7cf5467af5490c0a202c774afde Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:38:38 +0200 Subject: [PATCH 010/167] MAI: port comments `--` in theory `Tuples` to Isabelle2020 The `-- {* ... *}` have been changed to comments, not knowing another equivalent syntax in Isabelle2020. --- isabelle/Tuples.thy | 50 ++++++++++++++++++++++----------------------- 1 file changed, 25 insertions(+), 25 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index d23d908a..64cb6464 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -31,13 +31,13 @@ text \ on sequences require arithmetic and will be introduced in a separate theory. \ -definition Seq -- {* set of finite sequences with elements from $S$ *} +definition Seq (* -- \ set of finite sequences with elements from $S$ \ *) where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" -definition isASeq -- {* characteristic predicate for sequences or tuples *} +definition isASeq (* -- \ characteristic predicate for sequences or tuples \close> *) where "isASeq(s) \ isAFcn(s) \ (\n \ Nat : DOMAIN s = 1 .. n)" -definition Len -- {* length of a sequence *} +definition Len (* -- \ length of a sequence \ *) where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: @@ -80,7 +80,7 @@ lemma SeqIsAFcn (*[elim!]*): shows "isAFcn(s)" using assms by auto --- {* @{text "s \ Seq(S) \ isAFcn(s)"} *} +(* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn, standard] lemma LenInNat [simp]: @@ -88,7 +88,7 @@ lemma LenInNat [simp]: shows "Len(s) \ Nat" using assms by auto --- {* @{text "s \ Seq(S) \ Len(s) \ Nat"} *} +(* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat, standard] lemma DomainSeqLen [simp]: @@ -96,7 +96,7 @@ lemma DomainSeqLen [simp]: shows "DOMAIN s = 1 .. Len(s)" using assms by auto --- {* @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} *} +(* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] lemma seqEqualI: @@ -121,7 +121,7 @@ lemma SeqI [intro!]: shows "s \ Seq(S)" using assms by (auto simp: Seq_def) -lemma SeqI': -- {* closer to the definition but probably less useful *} +lemma SeqI': (* -- \ closer to the definition but probably less useful \ *) assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" using assms by (auto simp: Seq_def) @@ -179,7 +179,7 @@ where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" by (auto simp: emptySeq_def isASeq_def) --- {* @{text "isAFcn(\\)"} *} +(* -- \ @{text "isAFcn(\\)"} \ *) lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] lemma lenEmptySeq [simp]: "Len(\\) = 0" @@ -199,7 +199,7 @@ lemma appendIsASeq [simp,intro!]: using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) --- {* @{text "isASeq(s) \ isAFcn(Append(s,e))"} *} +(* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn, standard] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" @@ -213,11 +213,11 @@ lemma isEmptySeq [intro!]: "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" by auto --- {* immediate consequence of @{text isEmptySeq} *} +(* -- \ immediate consequence of @{text isEmptySeq} \ *) lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto --- {* Symmetric equation could be a useful rewrite rule (it is applied by TLC) *} +(* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn, standard] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" @@ -243,7 +243,7 @@ lemma lenAppend [simp]: shows "Len(Append(s,e)) = Succ[Len(s)]" using assms by (intro LenI, auto simp: Append_def) --- {* @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} *} +(* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend, standard] lemma appendElt [simp]: @@ -398,7 +398,7 @@ proof - with s show ?thesis unfolding Seq_def by auto qed --- {* example of an inductive proof about sequences *} +(* -- \ example of an inductive proof about sequences \ *) lemma seqEmptyOrAppend: assumes "s \ Seq(S)" shows "s = \\ \ (\s' \ Seq(S): \e \ S : s = Append(s',e))" @@ -508,7 +508,7 @@ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" -lemmas -- {* establish set equality for sets of enumerated functions *} +lemmas (* -- \ establish set equality for sets of enumerated functions \ *) setEqualI [where A = "EnumFuncSet(doms, rngs)", standard, intro!] setEqualI [where B = "EnumFuncSet(doms, rngs)", standard, intro!] @@ -835,7 +835,7 @@ lemma inSetOfPairsI [intro]: shows "a \ setOfPairs(R, e)" using assms by (auto simp: setOfPairs_def) -lemma inSetOfPairsE [elim!]: -- {* converse true only if $R$ is a relation *} +lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" @@ -885,20 +885,20 @@ where "rel_range(r) \ { p[2] : p \ r }" definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" -definition rel_comp :: "[c,c] => c" (infixr "\" 75) -- {* binary relation composition *} +definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" -definition Id :: "c \ c" -- {* diagonal: identity over a set *} +definition Id :: "c \ c" (* -- \ diagonal: identity over a set \ *) where "Id(A) \ { \x,x\ : x \ A }" text \ Properties of relations \ -definition reflexive -- {* reflexivity over a set *} +definition reflexive (* -- \ reflexivity over a set \ *) where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" @@ -907,7 +907,7 @@ unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" unfolding isBool_def by (rule boolifyReflexive) -definition symmetric -- {* symmetric relation *} +definition symmetric (* -- \ symmetric relation \ *) where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" @@ -916,7 +916,7 @@ unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" unfolding isBool_def by (rule boolifySymmetric) -definition antisymmetric -- {* antisymmetric relation *} +definition antisymmetric (* -- \ antisymmetric relation \ *) where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" @@ -925,7 +925,7 @@ unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" unfolding isBool_def by (rule boolifyAntisymmetric) -definition transitive -- {* transitivity predicate *} +definition transitive (* -- \ transitivity predicate \ *) where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" @@ -934,7 +934,7 @@ unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" unfolding isBool_def by (rule boolifyTransitive) -definition irreflexive -- {* irreflexivity predicate *} +definition irreflexive (* -- \ irreflexivity predicate \ *) where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" @@ -943,7 +943,7 @@ unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" unfolding isBool_def by (rule boolifyIrreflexive) -definition equivalence :: "[c,c] \ c" -- {* (partial) equivalence relation *} +definition equivalence :: "[c,c] \ c" (* -- \ (partial) equivalence relation \ *) where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" @@ -1043,7 +1043,7 @@ lemma converseE [elim]: (** consider [elim!] ?? **) assumes yx: "yx \ r^-1" and r: "r \ A \ B" and p: "\x y. yx = \y,x\ \ \x,y\ \ r \ P" shows "P" - -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *} + (* -- \ More general than @{text converseD}, as it ``splits'' the member of the relation. \ *) proof - from prodProj[OF subsetD[OF converseSubset[OF r] yx]] have 2: "yx = \yx[1], yx[2]\" . with yx have 3: "\yx[2], yx[1]\ \ r" @@ -1419,7 +1419,7 @@ abbreviation abbreviation RESPECTS2 ::"[c \ c \ c, c] \ c" (infixr "respects2 " 80) where "f respects2 r \ congruent2(r,r,f)" - --{* Abbreviation for the common case where the relations are identical *} + (* -- \ Abbreviation for the common case where the relations are identical \ *) ***************************************************************************) text \ @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} \ From 17da587b3a575cdb45d4b10300ba305c5f49d476 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:44:10 +0200 Subject: [PATCH 011/167] MAI: rm attribute `standard` from `lemmas` because in Isabelle2020 this attribute is not needed, and unsupported. This conclusion is based on the file `NEWS` in the Isabelle `hg` repository. --- isabelle/Tuples.thy | 40 ++++++++++++++++++++-------------------- 1 file changed, 20 insertions(+), 20 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 64cb6464..e939c1bb 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -81,7 +81,7 @@ lemma SeqIsAFcn (*[elim!]*): using assms by auto (* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) -lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn, standard] +lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" @@ -89,7 +89,7 @@ lemma LenInNat [simp]: using assms by auto (* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) -lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat, standard] +lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] lemma DomainSeqLen [simp]: assumes "isASeq(s)" @@ -97,7 +97,7 @@ lemma DomainSeqLen [simp]: using assms by auto (* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) -lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] +lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: assumes "isASeq(s)" and "isASeq(t)" @@ -200,7 +200,7 @@ using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) (* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) -lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn, standard] +lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" by (simp add: emptySeq_def) @@ -218,7 +218,7 @@ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto (* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) -lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn, standard] +lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" by auto @@ -244,28 +244,28 @@ lemma lenAppend [simp]: using assms by (intro LenI, auto simp: Append_def) (* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) -lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend, standard] +lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] lemma appendElt [simp]: assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Succ[Len(s)]" shows "Append(s,e)[k] = (IF k = Succ[Len(s)] THEN e ELSE s[k])" using assms by (auto simp: Append_def) -lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt, standard] +lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt] lemma appendElt1 (*[simp]*): assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Len(s)" shows "Append(s,e)[k] = s[k]" using assms by (auto simp: Append_def) -lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1, standard] +lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1] lemma appendElt2 (*[simp]*): assumes "isASeq(s)" shows "Append(s,e)[Succ[Len(s)]] = e" using assms by (auto simp: Append_def) -lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2, standard] +lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2] lemma isAppend [intro!]: assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. Succ[Len(s)]" and s: "isASeq(s)" @@ -280,7 +280,7 @@ next by (auto simp: Append_def) qed -lemmas isAppend' [intro!] = isAppend[symmetric, standard] +lemmas isAppend' [intro!] = isAppend[symmetric] lemma appendInSeq [simp]: assumes s: "s \ Seq(S)" and e: "e \ S" @@ -509,8 +509,8 @@ where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNI \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" lemmas (* -- \ establish set equality for sets of enumerated functions \ *) - setEqualI [where A = "EnumFuncSet(doms, rngs)", standard, intro!] - setEqualI [where B = "EnumFuncSet(doms, rngs)", standard, intro!] + setEqualI [where A = "EnumFuncSet(doms, rngs)", intro!] + setEqualI [where B = "EnumFuncSet(doms, rngs)", intro!] lemma EnumFuncSetI [intro!,simp]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" @@ -847,8 +847,8 @@ proof - qed lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)", standard,intro!] - setEqualI [where B = "setOfPairs(R,f)", standard,intro!] + setEqualI [where A = "setOfPairs(R,f)",intro!] + setEqualI [where B = "setOfPairs(R,f)",intro!] lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" @@ -1019,8 +1019,8 @@ by force subsubsection \ Converse relation \ lemmas converseEqualI = - setEqualI [where A = "r^-1", standard, intro!] - setEqualI [where B = "r^-1", standard, intro!] + setEqualI [where A = "r^-1", intro!] + setEqualI [where B = "r^-1", intro!] lemma converse_iff [iff]: assumes r: "r \ A \ B" @@ -1126,8 +1126,8 @@ qed subsubsection \ Identity relation over a set \ lemmas idEqualI = - setEqualI [where A = "Id(S)", standard, intro!] - setEqualI [where B = "Id(S)", standard, intro!] + setEqualI [where A = "Id(S)", intro!] + setEqualI [where B = "Id(S)", intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" unfolding Id_def by auto @@ -1176,8 +1176,8 @@ unfolding rel_range_def Id_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s", standard, intro!] - setEqualI [where B = "r \ s", standard, intro!] + setEqualI [where A = "r \ s", intro!] + setEqualI [where B = "r \ s", intro!] lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From c8bafaefc26fdbe6863151c63737d8e76c1e3c34 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:44:47 +0200 Subject: [PATCH 012/167] MAI: update `def` to `define ... where` for Isabelle2020 in the theory `Tuples`. --- isabelle/Tuples.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index e939c1bb..e0884db3 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -369,8 +369,8 @@ proof - proof fix sn assume sn: "sn \ [1 .. Succ[n] \ S]" - def so \ "[k \ 1 .. n \ sn[k]]" - def lst \ "sn[Succ[n]]" + define so where "so \ [k \ 1 .. n \ sn[k]]" + define lst where "lst \ sn[Succ[n]]" have 1: "sn = Append(so, lst)" proof from sn show "isAFcn(sn)" by simp From 0ea80153134a6052234441b071bfd304d6ec1599 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:46:10 +0200 Subject: [PATCH 013/167] MAI: update ML section delimiters in theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index e0884db3..5bd08f33 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -541,7 +541,7 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") -parse_ast_translation {* +parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. The order of elements is reversed. *) @@ -571,9 +571,9 @@ parse_ast_translation {* in [("@EnumFuncSet", enum_funcset_tr)] end -*} +\ -print_ast_translation {* +print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] | list_from_tuple (Ast.Appl[Ast.Constant "@tuple", tp]) = @@ -602,8 +602,7 @@ print_ast_translation {* in [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] end -*} - +\ (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" From 7c573ff62e076d0c1e3f343737394b815476ad44 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:47:13 +0200 Subject: [PATCH 014/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. --- isabelle/Tuples.thy | 2 ++ 1 file changed, 2 insertions(+) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5bd08f33..96c35178 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -541,6 +541,7 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") +(* parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -603,6 +604,7 @@ print_ast_translation \ [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] end \ +*) (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" From 72eebdda5ce279a006dcb75878922253e3c8722c Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:48:28 +0200 Subject: [PATCH 015/167] MAI: update `lemmas` of theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 96c35178..0ca96fe4 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -848,8 +848,8 @@ proof - qed lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)",intro!] - setEqualI [where B = "setOfPairs(R,f)",intro!] + setEqualI [where A = "setOfPairs(R,f)" for R f, intro!] + setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" @@ -1020,8 +1020,8 @@ by force subsubsection \ Converse relation \ lemmas converseEqualI = - setEqualI [where A = "r^-1", intro!] - setEqualI [where B = "r^-1", intro!] + setEqualI [where A = "r^-1" for r, intro!] + setEqualI [where B = "r^-1" for r, intro!] lemma converse_iff [iff]: assumes r: "r \ A \ B" @@ -1127,8 +1127,8 @@ qed subsubsection \ Identity relation over a set \ lemmas idEqualI = - setEqualI [where A = "Id(S)", intro!] - setEqualI [where B = "Id(S)", intro!] + setEqualI [where A = "Id(S)" for S, intro!] + setEqualI [where B = "Id(S)" for S, intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" unfolding Id_def by auto @@ -1177,8 +1177,8 @@ unfolding rel_range_def Id_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s", intro!] - setEqualI [where B = "r \ s", intro!] + setEqualI [where A = "r \ s" for r s, intro!] + setEqualI [where B = "r \ s" for r s, intro!] lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From f6253d615a4c0db1b2f79e3f6a96f33149b832d2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 25 Dec 2020 14:03:52 +0200 Subject: [PATCH 016/167] MAI: update theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 398 +++++++++++++++++++++++++++++++++++++------- 1 file changed, 336 insertions(+), 62 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 0ca96fe4..47d556ce 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -336,7 +336,28 @@ text \ lemma appendExtend: assumes "isASeq(s)" shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" -using assms by force + proof - + have s1_1: "Append(s, e) = + [k \ 1 .. Succ[Len(s)] \ + IF k = Succ[Len(s)] THEN e ELSE s[k]]" + unfolding Append_def by auto + have s1_2: "(Succ[Len(s)] :> e) = + [x \ {Succ[Len(s)]} \ e]" + unfolding oneArg_def by auto + define p where "p \ s @@ (Succ[Len(s)] :> e)" + have s1_3: "p = + [x \ (DOMAIN s) \ {Succ[Len(s)]} + \ IF x \ DOMAIN s THEN s[x] ELSE e]" + unfolding p_def extend_def using s1_2 by auto + have s1_4: "(DOMAIN s) \ {Succ[Len(s)]} + = 1 .. Succ[Len(s)]" + using assms by auto + have s1_5: "p = [x \ 1..Succ[Len(s)] \ + IF x \ DOMAIN s THEN s[x] ELSE e]" + using s1_3 s1_4 by auto + show ?thesis + using s1_1 s1_5 assms by (auto simp: p_def) + qed lemma imageEmptySeq [simp]: "Image(\\, A) = {}" by (simp add: emptySeq_def) @@ -725,6 +746,7 @@ by (auto simp add: prod_def) lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" +unfolding two_def using assms by (auto simp add: prod_def) lemma inProd': @@ -750,7 +772,7 @@ proof next show "\a,b\[1] = a" by simp next - show "\a,b\[2] = b" by simp + show "\a,b\[2] = b" unfolding two_def by simp qed lemma inProdE [elim]: @@ -829,12 +851,13 @@ translations lemma inSetOfPairsI_ex: assumes "\\x,y\ \ R : a = e(x,y)" shows "a \ { e(x,y) : \x,y\ \ R }" -using assms by (auto simp: setOfPairs_def) +using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsI [intro]: assumes "a = e(x,y)" and "\x,y\ \ R" shows "a \ setOfPairs(R, e)" -using assms by (auto simp: setOfPairs_def) +unfolding two_def +using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) assumes 1: "z \ setOfPairs(R, e)" @@ -853,27 +876,187 @@ lemmas setOfPairsEqualI = lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" - shows "{ \x,y\ : \x,y\ \ R } = R" -using assms by auto + shows "{ \x, y\ : \x, y\ \ R } = R" + proof - + define Q where "Q \ + { \x, y\ : \x, y\ \ R }" + have s1_1: "Q = {<>: p \ R}" + unfolding Q_def setOfPairs_def by auto + have s1_2: "\ p. p \ R ==> p \ A \ B" + using s by auto + have s1_3: "\ p. p \ R ==> p = <>" + proof - + fix "p" :: "c" + assume s1_3_asm: "p \\in R" + have s2_1: "p \ A \ B" + using s1_3_asm s by auto + show "p = <>" + using s2_1 prodProj by auto + qed + define e :: "[c] => c" where "e \ \ x. x" + define f :: "[c] => c" where "f \ \ x. <>" + have s1_4: "\ x. x \ R ==> e(x) = f(x)" + unfolding e_def f_def using s1_3 by auto + have s1_5: "{e(x): x \ R} = {f(y): y \ R}" + apply (rule setOfAll_cong[of "R" "R" "e" "f"]) + apply (simp) + by (rule s1_4) + show ?thesis + using s1_1 s1_5 by (auto simp: Q_def e_def f_def) + qed lemma setOfPairs_cong (*[cong]*): - assumes 1: "R = S" and 2: "S \ A \ B" and 3: "\x y. \x,y\ \ S \ e(x,y) = f(x,y)" - shows "{ e(x,y) : \x,y\ \ R } = { f(u,v) : \u,v\ \ S }" -using assms proof (auto) - fix u v - let ?p = "\u,v\" - assume uv: "?p \ S" - with 3 have "f(u,v) = e(?p[1], ?p[2])" by simp - with uv show "f(u,v) \ setOfPairs(S, e)" by auto -qed + assumes 1: "R = S" and + 2: "S \ A \ B" and + 3: "\x y. \x, y\ \ S + \ + e(x, y) = f(x, y)" + shows "{ e(x, y) : \x, y\ \ R } = + { f(u, v) : \u, v\ \ S }" + proof - + have s1_1: "{ e(x, y) : \x, y\ \ R } = + { e(x, y) : \x, y\ \ S }" + using 1 by auto + have s1_2: "{ e(x, y) : \x, y\ \ S } = + { e(p[1], p[2]) : p \ S }" + by (auto simp: setOfPairs_def) + have s1_3: "{ f(u, v) : \u, v\ \ S } = + { f(p[1], p[2]) : p \ S }" + by (auto simp: setOfPairs_def) + define P where "P \ { e(p[1], p[2]): p \ S }" + define Q where "Q \ { f(p[1], p[2]): p \ S }" + have s1_4: "P = Q" + proof - + have s2_1: "\ x. x \ P <=> x \ Q" + proof - + fix "x" :: "c" + have s3_1: "x \\in P ==> x \\in Q" + proof - + assume s3_1_asm: "x \\in P" + have s4_1: "x \ { e(p[1], p[2]): p \ S }" + using s3_1_asm by (auto simp: P_def) + have s4_2: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_1 setOfAll by auto + have s4_3: "\\E p: p \\in S \ x = f(p[1], p[2])" + using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) + have s4_4: "x \ { f(p[1], p[2]): p \ S }" + using s4_3 setOfAll by auto + show "x \\in Q" + unfolding Q_def using s4_4 by auto + qed + have s3_2: "x \\in Q ==> x \\in P" + proof - + assume s3_2_asm: "x \\in Q" + have s4_1: "x \ { f(p[1], p[2]): p \ S}" + using s3_2_asm by (auto simp: Q_def) + have s4_2: "\\E p: p \\in S \ x = f(p[1], p[2])" + using s4_1 setOfAll by auto + have s4_3: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) + have s4_4: "x \ { e(p[1], p[2]): p \ S }" + using s4_3 setOfAll by auto + show "x \\in P" + unfolding P_def using s4_4 by auto + qed + show "x \ P <=> x \ Q" + using s3_1 s3_2 by auto + qed + have s2_2: "\ x: x \ P <=> x \ Q" + using s2_1 allI by auto + show ?thesis + using s2_2 extension by auto + qed + show ?thesis + using s1_1 s1_2 s1_3 s1_4 + by (auto simp: P_def Q_def) + qed lemma setOfPairsEqual: - assumes 1: "\x y. \x,y\ \ S \ \\x',y'\ \ T : e(x,y) = f(x',y')" - and 2: "\x' y'. \x',y'\ \ T \ \\x,y\\S : f(x',y') = e(x,y)" - and 3: "S \ A \ B" and 4: "T \ C \ D" - shows "{ e(x,y) : \x,y\ \ S } = { f(x,y) : \x,y\ \ T }" -using assms by (auto, blast+) - + assumes 1: "\x y. + \x, y\ \ S + \ \\x', y'\ \ T : + e(x, y) = f(x', y')" + and 2: "\x' y'. + \x', y'\ \ T + \ \\x, y\ \ S : + f(x', y') = e(x, y)" + and 3: "S \ A \ B" and + 4: "T \ C \ D" + shows "{ e(x, y) : \x, y\ \ S } = + { f(x, y) : \x, y\ \ T }" + proof - + have s1_1: "{ e(x, y) : \x, y\ \ S } = + { e(p[1], p[2]): p \ S }" + by (auto simp: setOfPairs_def) + have s1_2: "{ f(x, y) : \x, y\ \ T } = + { f(p[1], p[2]): p \ T }" + by (auto simp: setOfPairs_def) + define P where "P \ { e(p[1], p[2]): p \ S }" + define Q where "Q \ { f(p[1], p[2]): p \ T }" + have s1_3: "P = Q" + proof - + have s2_1: "\ x. x \ P <=> x \ Q" + proof - + fix "x" :: "c" + have s3_1: "x \\in P ==> x \\in Q" + proof - + assume s3_1_asm: "x \\in P" + have s4_1: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s3_1_asm setOfAll by (auto simp: P_def) + have s4_2: "\\E p: p \\in S \ + \p[1], p[2]\ \ S \ + x = e(p[1], p[2])" + using s4_1 pairProj_in_rel 3 by auto + have s4_3: "\\E p: p \\in S \ + x = e(p[1], p[2]) \ + (\\u, v\ \ T: + e(p[1], p[2]) = f(u, v))" + using s4_2 1 by auto + have s4_4: "\\u, v\ \ T: + x = f(u, v)" + using s4_3 by auto + have s4_5: "\\E u, v: <> \\in T \ x = f(u, v)" + using s4_4 by auto + have s4_6: "\\E p: p \\in T \ x = f(p[1], p[2])" + using s4_5 4 pairProj_in_rel by (auto simp: two_def) + show "x \\in Q" + using s4_6 setOfAll by (auto simp: Q_def) + qed + have s3_2: "x \\in Q ==> x \\in P" + proof - + assume s3_2_asm: "x \\in Q" + have s4_1: "\\E p: p \\in T \ x = f(p[1], p[2])" + using s3_2_asm setOfAll by (auto simp: Q_def) + have s4_2: "\\E p: p \\in T \ + <> \\in T \ + x = f(p[1], p[2])" + using s4_1 pairProj_in_rel 4 by auto + have s4_3: "\\E p: p \\in T \ + x = f(p[1], p[2]) \ + (\\u, v\ \ S: + f(p[1], p[2]) = e(u, v))" + using s4_2 2 by auto + have s4_4: "\\u, v\ \ S: + x = e(u, v)" + using s4_3 by auto + have s4_5: "\\E u, v: <> \\in S \ x = e(u, v)" + using s4_4 by auto + have s4_6: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_5 3 pairProj_in_rel by (auto simp: two_def) + show "x \\in P" + using s4_6 setOfAll by (auto simp: P_def) + qed + show "x \\in P <=> x \\in Q" + using s3_1 s3_2 by auto + qed + have s2_2: "\ x: x \ P <=> x \ Q" + using s2_1 allI by auto + show ?thesis + using s2_2 extension by auto + qed + show ?thesis + using s1_1 s1_2 s1_3 by (auto simp: P_def Q_def) + qed subsection \ Basic notions about binary relations \ @@ -986,7 +1169,7 @@ unfolding rel_range_def using pairProj_in_prod by auto lemma in_rel_rangeI [iff]: assumes "\x,y\ \ r" shows "y \ rel_range(r)" -unfolding rel_range_def using assms by auto +unfolding rel_range_def two_def using assms by auto lemma in_rel_rangeE [elim]: assumes y: "y \ rel_range(r)" and r: "r \ A \ B" and p: "\x. \x,y\ \ r \ P" @@ -1026,11 +1209,11 @@ lemmas converseEqualI = lemma converse_iff [iff]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" -using r prodProj by (auto simp: converse_def) +using r prodProj by (auto simp: converse_def two_def) lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" -unfolding converse_def by auto +unfolding converse_def two_def by auto lemma converseD [sym]: assumes r: "r \ A \ B" @@ -1057,10 +1240,52 @@ qed lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" -using assms prodProj by (auto elim!: converseE) + proof - + have s1_1: "\ x. x \\in A \ B ==> x = <>" + using prodProj by (auto simp: two_def) + have s1_2: "\ x. x \\in r ==> x \\in A \ B" + using s1_1 r by auto + have s1_3: "\ x. x \\in r ==> x = <>" + using s1_1 s1_2 by auto + show ?thesis + using s1_3 assms prodProj pairProj_in_rel + by (auto simp: converse_def two_def) + qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" -using prodProj by auto + proof - + have s1_1: "(A \ B)^-1 = + {<>: p \ A \ B}" + unfolding converse_def by auto + have s1_2: "B \ A = + {<>: p \ B \ A}" + using setOfPairs_triv prodProj by auto + have s1_3: "\ x. x \\in B \\X A ==> + \\E p \\in A \\X B: x = <>" + proof - + fix "x" :: "c" + assume s1_5_asm: "x \\in B \\X A" + have s2_1: "<> \\in A \\X B" + using s1_5_asm by (auto simp: two_def) + define q where "q \ <>" + have s2_2: "q \\in A \\X B \ q = <>" + using s2_1 by (auto simp: q_def) + have s2_3: "x = <>" + using s1_5_asm prodProj by auto + have s2_4: "q[1] = x[2] \ q[2] = x[1]" + using s2_2 by (auto simp: two_def) + have s2_5: "x = <>" + using s2_3 s2_4 by auto + show "\\E p \\in A \\X B: x = <>" + using s2_2 s2_5 by auto + qed + have s1_4: "\ p. p \\in A \\X B ==> + p[1] \\in A \ p[2] \\in B" + using pairProj_in_rel inProd prodProj + by (auto simp: two_def) + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed lemma converse_empty [simp]: "converse({}) = {}" by auto @@ -1106,23 +1331,44 @@ lemma transitive_converse [simp]: unfolding transitive_def by auto lemma symmetric_iff_converse_eq: - assumes r:"r \ A \ B" + assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" -proof auto - fix p - assume "symmetric(r)" and "p \ r^-1" - with r show "p \ r" by (auto elim!: converseE simp add: symmetric_def) -next - fix p - assume 1: "symmetric(r)" and 2: "p \ r" - from r 2 have 3: "p = \p[1],p[2]\" by (intro prodProj, auto) - with 1 2 have "\p[2],p[1]\ \ r" by (force simp add: symmetric_def) - with 3 show "p \ r^-1" by (auto dest: converseI) -next - assume "r^-1 = r" thus "symmetric(r)" - by (auto simp add: symmetric_def) -qed - + proof - + have s1_1: "symmetric(r) ==> r^-1 = r" + proof - + assume s1_1_asm: "symmetric(r)" + have s2_1: "\ p. p \\in r^-1 ==> p \\in r" + proof - + fix "p" :: "c" + assume s2_1_asm: "p \\in r^-1" + show "p \\in r" + using r s1_1_asm s2_1_asm + by (auto elim!: converseE simp: symmetric_def) + qed + have s2_2: "\ p. p \\in r ==> p \\in r^-1" + proof - + fix "p" :: "c" + assume s2_2_asm: "p \\in r" + have s3_1: "p = <>" + using s2_2_asm r by (intro prodProj, auto) + have s3_2: "<> \\in r" + using s1_1_asm s2_2_asm s3_1 by (auto simp: symmetric_def) + show "p \\in r^-1" + using s3_1 s3_2 by (auto dest: converseI) + qed + show "r^-1 = r" + using s2_1 s2_2 setEqualI[of "r^-1" "r"] + by blast + qed + have s1_2: "r^-1 = r ==> symmetric(r)" + proof - + assume s1_2_asm: "r^-1 = r" + show "symmetric(r)" + using s1_2_asm by (auto simp: symmetric_def) + qed + show ?thesis + using s1_1 s1_2 by auto + qed subsubsection \ Identity relation over a set \ @@ -1165,13 +1411,14 @@ lemma Id_eqI: "a = b \ a \ A \ \a,b\ by simp lemma converse_Id [simp]: "Id(A)^-1 = Id(A)" -by auto +by (auto simp: Id_def two_def) + lemma dom_Id [simp]: "rel_domain(Id(A)) = A" unfolding rel_domain_def Id_def by auto lemma ran_Id [simp]: "rel_range(Id(A)) = A" -unfolding rel_range_def Id_def by auto +unfolding rel_range_def Id_def two_def by auto subsubsection \ Composition of relations \ @@ -1208,7 +1455,7 @@ using assms by force lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "((r \ s)^-1) = (s^-1 \ r^-1)" (is "?lhs = ?rhs") -proof +proof (rule setEqualI) fix x assume x: "x \ ?lhs" from s r have "r \ s \ A \ C" by (rule rel_comp_in_prod) @@ -1231,26 +1478,53 @@ qed lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" -using r proof auto - fix p - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) - with 1 show "p \ R \ Id(B)" by simp -qed + using r + proof - + have s1_1: "\ p. p \\in R ==> p \\in R \ Id(B)" + proof - + fix "p" :: "c" + assume p: "p \ R" + with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) + from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) + with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) + with 1 show "p \ R \ Id(B)" by simp + qed + have s1_2: "\ p. p \\in R \ Id(B) ==> p \\in R" + proof - + fix "p" :: "c" + assume s1_2_asm: "p \\in R \ Id(B)" + show "p \\in R" + using r s1_2_asm by auto + qed + show ?thesis + using s1_1 s1_2 setEqualI[of "R \ Id(B)" "R"] by blast + qed + lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" -using r proof auto - fix p - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) - with 1 show "p \ Id(B) \ R" by simp -qed + using r + proof - + have s1_1: "\ p. p \\in R ==> p \\in Id(B) \ R" + proof - + fix "p" :: "c" + assume p: "p \ R" + with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) + from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) + with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) + with 1 show "p \ Id(B) \ R" by simp + qed + have s1_2: "\ p. p \\in Id(B) \ R ==> p \\in R" + proof - + fix "p" :: "c" + assume s1_2_asm: "p \\in Id(B) \ R" + show "p \\in R" + using r s1_2_asm by auto + qed + show ?thesis + using s1_1 s1_2 setEqualI[of "Id(B) \ R" "R"] by blast + qed lemma rel_comp_empty1 [simp]: "{} \ R = {}" unfolding rel_comp_def by auto @@ -1261,7 +1535,7 @@ unfolding rel_comp_def by auto lemma comp_assoc: assumes t: "T \ A \ B" and s: "S \ B \ C" and r: "R \ C \ D" shows "(R \ S) \ T = R \ (S \ T)" -proof +proof (rule setEqualI) fix p assume p: "p \ (R \ S) \ T" from r s have "R \ S \ B \ D" by simp From 232fc196fcbe3fb6c0f516b8dee546f975267f20 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 25 Dec 2020 14:04:28 +0200 Subject: [PATCH 017/167] STY: shorter lines --- isabelle/Tuples.thy | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 47d556ce..7c7bfc97 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1070,8 +1070,13 @@ definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) -where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : - \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" +where "r \ s \ + { p \ rel_domain(s) \ rel_range(r) : + \ x, z : + p = \x, z\ \ + (\ y: + \x, y\ \ s \ + \y, z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" From a763bf205f2a9332e69cd9eef0e6e14481ce05ae Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 13:32:13 +0200 Subject: [PATCH 018/167] DRAFT: comment attribute `intro!` in `lemmas compEqualI` in theory `Tuples`, because it gives rise to non-terminating behavior in Isabelle2020. --- isabelle/Tuples.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 7c7bfc97..d644e8b5 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1429,8 +1429,8 @@ unfolding rel_range_def Id_def two_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s" for r s, intro!] - setEqualI [where B = "r \ s" for r s, intro!] + setEqualI [where A = "r \ s" for r s] (* , intro!] *) + setEqualI [where B = "r \ s" for r s] (* , intro!] *) lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From 174fc6d4c4b29dbc9bd5c2690ba393d8d6442fc8 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 13:51:04 +0200 Subject: [PATCH 019/167] MAI: update theory `NatDivision` after change of numerals to definitions --- isabelle/NatDivision.thy | 19 ++++++++++++------- 1 file changed, 12 insertions(+), 7 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index fcc8a7f2..1d2869aa 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -256,7 +256,7 @@ lemma divmodNatPairEx: proof - from assms obtain q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" by (rule divmod_rel_ex) - thus ?thesis by force + thus ?thesis unfolding two_def by force qed lemma divmodNatInNatNat: @@ -279,6 +279,7 @@ proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) fix z assume "z \ Nat \ Nat" and "divmod_rel(m,n,z[1],z[2])" with m n q r h show "z = \q,r\" + unfolding two_def by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) qed @@ -301,18 +302,20 @@ using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) lemma modIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def) +using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def two_def) lemma divmodNat_div_mod: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "divmodNat(m,n) = \m div n, m mod n\" -unfolding div_nat_def[OF m n] mod_nat_def[OF m n] using divmodNatInNatNat[OF assms] +unfolding div_nat_def[OF m n] mod_nat_def[OF m n] two_def +using divmodNatInNatNat[OF assms] by force lemma divmod_rel_div_mod_nat: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "divmod_rel(m, n, m div n, m mod n)" -using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] by simp +using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] +by (auto simp: two_def) lemma div_nat_unique: assumes h: "divmod_rel(m,n,q,r)" @@ -324,7 +327,7 @@ lemma mod_nat_unique: assumes h: "divmod_rel(m,n,q,r)" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" shows "m mod n = r" -unfolding mod_nat_def[OF m n] using divmodNat_unique[OF assms] by simp +unfolding mod_nat_def[OF m n] two_def using divmodNat_unique[OF assms] by simp lemma mod_nat_less_divisor: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" @@ -342,10 +345,12 @@ proof - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" by (simp add: divmodNat_divmod_rel) from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - with 1 2 less n have "?dm[1] * n < n" by (auto simp: divmod_rel_def elim!: add_lessD1) + with 1 2 less n have "?dm[1] * n < n" + unfolding two_def + by (auto simp: divmod_rel_def elim!: add_lessD1) with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) - with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def) + with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) with 3 prodProj[OF 2] show ?thesis by simp qed From 4070b971ad2cf916e571dc9c3b119e6bde0ecde0 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 16:56:43 +0200 Subject: [PATCH 020/167] MAI: update theory `NatDivision` to Isabelle2020 --- isabelle/NatDivision.thy | 38 +++++++++++++++++++++++++++++--------- 1 file changed, 29 insertions(+), 9 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 1d2869aa..94eefffd 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -221,12 +221,18 @@ proof - assume "\(x \ x')" with nat have "x' < x" by (simp add: nat_not_leq[simplified]) with nat obtain k where k: "k \ Nat" "x = Succ[x'+k]" - by (auto simp: less_iff_Succ_add) + using less_iff_Succ_add by auto with eq nat n have "x'*n + (k*n + n + y) = x'*n + y'" by (simp add: add_mult_distrib_right_nat add_assoc_nat) with nat k n have "k*n + n + y = y'" by simp - with less k n nat have "(k*n + y) + n < n" by (simp add: add_ac_nat) - with k n nat show "FALSE" by simp + with less k n nat have s3_1: "(k*n + y) + n < n" by (simp add: add_ac_nat) + define j where "j \ k * n + y" + have s3_2: "j + n < n" + unfolding j_def using s3_1 by auto + have s3_3: "j \\in Nat" + unfolding j_def using multIsNat addIsNat k n nat by auto + show "FALSE" + using s3_2 s3_3 n not_add_less2[of "n" "j"] by auto qed } from this[OF q r q' r'] this[OF q' r' q r] q q' mqr mqr' rn rn' @@ -345,9 +351,19 @@ proof - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" by (simp add: divmodNat_divmod_rel) from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - with 1 2 less n have "?dm[1] * n < n" - unfolding two_def - by (auto simp: divmod_rel_def elim!: add_lessD1) + have "?dm[1] * n < n" + proof - + have s2_1: "m = ?dm[1] * n + ?dm[2]" + using 1 by (auto simp: divmod_rel_def) + have s2_2: "?dm[1] * n + ?dm[2] < n" + using less s2_1 by auto + have s2_3: "?dm[1] \\in Nat" + using 2 by auto + have s2_4: "?dm[2] \\in Nat" + using 2 by (auto simp: two_def) + show ?thesis + using s2_2 s2_3 s2_4 multIsNat n add_lessD1 by auto + qed with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) @@ -364,8 +380,11 @@ proof - have 2: "m div n \ 0" proof assume "m div n = 0" - with 1 m n pos have "m < n" by (auto simp: divmod_rel_def) - with geq m n show "FALSE" by (auto simp: nat_less_leq_not_leq) + with 1 m n pos have s2_1: "m < n" by (auto simp: divmod_rel_def) + have s2_2: "\ (n \ m)" + using s2_1 n m nat_not_leq by auto + show "FALSE" + using geq s2_2 by auto qed with m n pos obtain k where k1: "k \ Nat" and k2: "m div n = Succ[k]" using not0_implies_Suc[of "m div n"] by auto @@ -621,7 +640,8 @@ lemma mod_div_nat_trivial [simp]: proof - from assms have "m div n + (m mod n) div n = (m mod n + (m div n) * n) div n" - by (simp add: mod_nat_less_divisor) + using div_nat_mult_self1[of "m mod n" "m div n" "n"] + by (auto simp: mod_nat_less_divisor) also from assms have "... = m div n + 0" by simp finally show ?thesis using assms by simp From c3667176cd94d3693eb95e926c56b294bceb76f2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 16:56:58 +0200 Subject: [PATCH 021/167] STY: shorter lines --- isabelle/NatDivision.thy | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 94eefffd..cdfa1eb4 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -165,8 +165,11 @@ text \ \ definition divmod_rel where - "divmod_rel(m,n,q,r) \ m = q * n + r - \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" + "divmod_rel(m,n,q,r) \ + m = q * n + r + \ + ((0 < n \ 0 \ r \ r < n) \ + (n < 0 \ r \ 0 \ n < r))" text \ @{const divmod_rel} is total if $n$ is non-zero. \ From ccc8a877bd2e072ad8b645d3ff215f10f8ef5478 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:03:24 +0200 Subject: [PATCH 022/167] MAI: update delimiters in theory `CaseExpressions` to Isabelle2020 --- isabelle/CaseExpressions.thy | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index c9fc0a31..d7fcf341 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:40:22 merz> *) -header {* Case expressions *} +section \ Case expressions \ theory CaseExpressions imports Tuples begin -text {* +text \ A \textsc{case} expression in \tlaplus{} has the form \[ \textsc{case } p_1 \rightarrow e_1 @@ -32,7 +32,7 @@ text {* guard of the \textsc{other}-branch, when present, is the conjunction of the negated guards of all other branches, so every guard appears twice (and will be simplified twice) in a @{text CaseOther} expression. -*} +\ definition CaseArm -- {* preliminary construct to convert case arm into set *} where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" @@ -64,7 +64,7 @@ syntax (HTML output) "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") -parse_ast_translation {* +parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. The order of elements is reversed. *) @@ -99,10 +99,10 @@ parse_ast_translation {* in [("_case_syntax", case_syntax_tr)] end -*} +\ (** debugging ** -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" oops @@ -113,10 +113,10 @@ oops lemma "CASE OTHER \ g" (* degenerated case *) oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) **) -print_ast_translation {* +print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] | list_from_tuple (Ast.Appl[Ast.Constant "@tuple", tp]) = @@ -173,10 +173,10 @@ print_ast_translation {* [(@{const_syntax "Case"}, case_syntax_tr'), (@{const_syntax "CaseOther"}, caseother_tr')] end -*} +\ (** debugging ** -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" oops @@ -187,7 +187,7 @@ oops lemma "CASE OTHER \ g" (* degenerated case *) oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) **) lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def From c251804370335c73fc23eee4493e4ea6d627f062 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:04:05 +0200 Subject: [PATCH 023/167] MAI: update `--` comments in theory `CaseExpressions` to Isabelle2020 --- isabelle/CaseExpressions.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index d7fcf341..bcb75e57 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -34,7 +34,7 @@ text \ twice (and will be simplified twice) in a @{text CaseOther} expression. \ -definition CaseArm -- {* preliminary construct to convert case arm into set *} +definition CaseArm (* -- preliminary construct to convert case arm into set *) where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" definition Case where From 1cce60a78579c496b1aec13e14a037f82cc9b0a9 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:06:12 +0200 Subject: [PATCH 024/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. Mark them as "TODO". --- isabelle/CaseExpressions.thy | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index bcb75e57..09c2ca77 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -64,6 +64,7 @@ syntax (HTML output) "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") +(* TODO parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -100,6 +101,7 @@ parse_ast_translation \ [("_case_syntax", case_syntax_tr)] end \ +*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) @@ -116,6 +118,7 @@ oops (*ML \ reset Syntax.trace_ast; \*) **) +(* TODO print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] @@ -174,6 +177,7 @@ print_ast_translation \ (@{const_syntax "CaseOther"}, caseother_tr')] end \ +*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) From f4774bd93e0a0ba8ecc7d29779da297af86be2b9 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:15:13 +0200 Subject: [PATCH 025/167] MAI: update delimiters in theory `Strings` to Isabelle2020 --- isabelle/Strings.thy | 36 ++++++++++++++++++------------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index e951d12c..b274d2cf 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -6,18 +6,18 @@ Time-stamp: <2011-10-11 17:40:15 merz> *) -header {* Characters and strings *} +section \ Characters and strings \ theory Strings imports Tuples begin -subsection {* Characters *} +subsection \ Characters \ -text {* +text \ Characters are represented as pairs of hexadecimal digits (also called \emph{nibbles}). -*} +\ definition Nibble where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" @@ -40,7 +40,7 @@ lemma isChar [simp]: "(c \ Char) = (\a,b \ Nibble : c = char(a,b unfolding Char_def by auto -subsection {* Strings *} +subsection \ Strings \ definition String where "String \ Seq(Char)" @@ -49,7 +49,7 @@ syntax "_Char" :: "xstr \ c" ("CHAR _") "_String" :: "xstr \ c" ("_") -text {* +text \ The following parse and print translations convert between the internal and external representations of strings. Strings are written using two single quotes in Isabelle, such as \verb|''abc''|. Note that the @@ -57,9 +57,9 @@ text {* printed as @{term "\\"}. Single characters are printed in the form \verb|CHAR ''a''|: Isabelle doesn't provide single characters in its lexicon. -*} +\ -parse_ast_translation {* +parse_ast_translation \ let (* convert an ML integer to a nibble *) fun mkNibble n = @@ -93,11 +93,11 @@ parse_ast_translation {* in [("_Char", char_ast_tr), ("_String", string_ast_tr)] end -*} +\ (** debug **) -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "''a''" oops @@ -105,10 +105,10 @@ oops lemma "CHAR ''a''" oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) (**) -print_ast_translation {* +print_ast_translation \ let (* convert a nibble to an ML integer -- because translation macros have already been applied, we see constants "0" through "15", not Succ[...] terms! *) @@ -167,11 +167,11 @@ print_ast_translation {* in [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] end -*} +\ (*** examples ** (* not sure if the following is the way to set tracing? *) -(* ML {* val _ = Config.put Ast.trace true @{context} *} *) +(* ML \ val _ = Config.put Ast.trace true @{context} \ *) lemma "CHAR ''a'' \ CHAR ''b''" by simp @@ -191,13 +191,13 @@ by simp lemma "''abc''[2] \ CHAR ''a''" by simp -(* ML {* val _ = Config.put Ast.trace false @{context} *} *) +(* ML \ val _ = Config.put Ast.trace false @{context} \ *) **) -subsection {* Records and sets of records *} +subsection \ Records and sets of records \ -text {* +text \ Records are simply represented as enumerated functions with string arguments, such as @{text "(''foo'' :> 1) @ (''bar'' :> TRUE)"}. Similarly, there is no specific @{text "EXCEPT"} construct for records; use the standard one for @@ -207,7 +207,7 @@ text {* record syntax in Isabelle seems difficult, because the Isabelle lexer distinguishes between identifiers and strings: the latter must be surrounded by two single quotes. -*} +\ (** Examples ** From ceb17186c0428a45298fe16f76131f88b6b7bf5a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:15:59 +0200 Subject: [PATCH 026/167] MAI: update `--` comments in theory `Strings` to Isabelle2020 --- isabelle/Strings.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index b274d2cf..41b761f5 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -22,7 +22,7 @@ text \ definition Nibble where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" -definition char -- {* @{text char} is intended to be applied to nibbles *} +definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" lemma charInj [simp]: "(char(a,b) = char(c,d)) = (a=c \ b=d)" From a088bc573182236351da681cd1c426a1f4061db7 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:22:05 +0200 Subject: [PATCH 027/167] MAI: update `xstr` to `str_token` for Isabelle2020 based on the file `NEWS` of the Isabelle `hg` repository. --- isabelle/Strings.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 41b761f5..63c91b52 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -46,8 +46,8 @@ definition String where "String \ Seq(Char)" syntax - "_Char" :: "xstr \ c" ("CHAR _") - "_String" :: "xstr \ c" ("_") + "_Char" :: "str_token \ c" ("CHAR _") + "_String" :: "str_token \ c" ("_") text \ The following parse and print translations convert between the internal From c79245fbb1c6f6fa653616050ebded506e3a798a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:22:31 +0200 Subject: [PATCH 028/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. Mark them as "TODO". --- isabelle/Strings.thy | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 63c91b52..162b00f2 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -59,6 +59,7 @@ text \ lexicon. \ +(* TODO parse_ast_translation \ let (* convert an ML integer to a nibble *) @@ -94,6 +95,7 @@ parse_ast_translation \ [("_Char", char_ast_tr), ("_String", string_ast_tr)] end \ +*) (** debug **) @@ -108,6 +110,7 @@ oops (*ML \ reset Syntax.trace_ast; \*) (**) +(* TODO print_ast_translation \ let (* convert a nibble to an ML integer -- because translation macros have @@ -168,6 +171,7 @@ print_ast_translation \ [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] end \ +*) (*** examples ** (* not sure if the following is the way to set tracing? *) From 88b6431d25349c4b7a2b1452a239374ce30680aa Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:24:06 +0200 Subject: [PATCH 029/167] DRAFT: comment two `lemma`s that are for debugging because the first lemma raises an error about undefined constant `"_String"`. --- isabelle/Strings.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 162b00f2..cabb62c8 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -98,7 +98,7 @@ parse_ast_translation \ *) -(** debug **) +(** debug ** (*ML \ set Syntax.trace_ast; \*) lemma "''a''" @@ -108,7 +108,7 @@ lemma "CHAR ''a''" oops (*ML \ reset Syntax.trace_ast; \*) -(**) +**) (* TODO print_ast_translation \ From ab4ef5721479017a8d7321344dd00f4e120a42cc Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:27:27 +0200 Subject: [PATCH 030/167] MAI: update delimiters in theory `Constant` to Isabelle2020 --- isabelle/Constant.thy | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index 23c186be..f7aff81d 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -6,15 +6,15 @@ Time-stamp: <2011-10-11 17:40:28 merz> *) -header {* Main theory for constant-level Isabelle/\tlaplus{} *} +section \ Main theory for constant-level Isabelle/\tlaplus{} \ theory Constant imports NatDivision CaseExpressions Strings Integers begin -text {* +text \ This is just an umbrella for the component theories. -*} +\ end From c4c53ff9c74b42394b1492b807bb36e72a861631 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:37:48 +0200 Subject: [PATCH 031/167] MAI: update delimiters in theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 61 ++++++++++++++++++++++--------------------- 1 file changed, 31 insertions(+), 30 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 3502ce19..166ea6c3 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -7,13 +7,13 @@ *) -header {* The Integers as a superset of natural numbers *} +section \ The Integers as a superset of natural numbers \ theory Integers imports Tuples NatArith begin -subsection {* The minus sign *} +subsection \ The minus sign \ consts "minus" :: "c \ c" ("-._" [75] 75) @@ -123,7 +123,7 @@ lemma (*[simp]*) "\x \ Nat : -.1 = -.x" by auto lemma (*[simp]*) "x \ Nat \ (1 = -.x) = FALSE" by (auto simp: sym[OF minus_sym]) -subsection {* The set of Integers *} +subsection \ The set of Integers \ definition Int where "Int \ Nat \ {-.n : n \ Nat}" @@ -200,7 +200,7 @@ lemma intNotNatIsNegNat: "\n \ Nat; n \ Int\ \ Predicates ''is positive'' and 'is negative' \ definition isPos -- {* Predicate ''is positive'' *} where "isPos(n) \ \k \ Nat: n = Succ[k]" @@ -288,7 +288,7 @@ lemma intThenPosZeroNeg: by (auto elim: notIsNeg0_isPos[OF n]) -subsection {* Signum function and absolute value *} +subsection \ Signum function and absolute value \ definition sgn -- {* signum function *} where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" @@ -383,7 +383,7 @@ lemma sgn_minus [simp]: shows "sgn(-.n) = -.sgn(n)" unfolding sgn_def using n by (cases, auto) -text {* Absolute value *} +text \ Absolute value \ lemma absIsNat [simp]: assumes n: "n \ Int" shows "abs(n) \ Nat" @@ -408,12 +408,12 @@ lemma abs_neg [simp]: unfolding abs_def using n by (auto dest: sgnNat_not1) -subsection {* Orders on integers *} +subsection \ Orders on integers \ -text {* +text \ We distinguish four cases, depending on the arguments being in Nat or negative. -*} +\ lemmas int_leq_pp_def = nat_leq_def -- {* 'positive-positive' case, ie: both arguments are naturals *} @@ -448,18 +448,19 @@ lemma neg_le_iff_le [simp]: by(rule intCases2[of m n], simp_all) -subsection {* Addition of integers *} +subsection \ Addition of integers \ -text {* +text \ Again, we distinguish four cases in the definition of @{text "a + b"}, according to each argument being positive or negative. -*} +\ +(* cf. NatArith *) (** The following is rejected by Isabelle because the two definitions are not distinguishable by argument types. defs (unchecked overloaded) int_add_def: "\a \ Int; b \ Int \ \ a + b \ - IF a \ Nat \ b \ Nat THEN addnat(a)[b] (* cf. NatArith *) + IF a \ Nat \ b \ Nat THEN addnat(a)[b] ELSE IF isNeg(a) \ isNeg(b) THEN -.(addnat(-.a)[-.b]) ELSE IF isNeg(a) THEN IF -.a \ b THEN b -- a ELSE -.(a -- b) ELSE IF a \ -.b THEN -.(b -- a) ELSE a -- b" @@ -481,19 +482,19 @@ theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" by (auto simp: int_add_pn_def dest: nat_leq_antisym) -text {* Closure *} +text \ Closure \ lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" by (rule intCases2[of m n], auto simp: int_add_def) -text {* Neutral element *} +text \ Neutral element \ lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" by(rule intCases, auto simp add: int_add_np_def) lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" by(rule intCases, auto simp add: int_add_pn_def) -text {* Additive inverse element *} +text \ Additive inverse element \ lemma add_inverse_nat [simp]: "n \ Nat \ n + -.n = 0" by(simp add: int_add_pn_def) @@ -507,7 +508,7 @@ by (rule intCases, auto simp: int_add_def) lemma add_inverse2_int [simp]: "n \ Int \ -.n + n = 0" by (rule intCases, auto simp: int_add_def) -text {* Commutativity *} +text \ Commutativity \ lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" by(simp add: int_add_def) @@ -515,7 +516,7 @@ by(simp add: int_add_def) lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" by(rule intCases2[of m n], auto simp add: int_add_def add_commute_nat) -text {* Associativity *} +text \ Associativity \ lemma add_pn_eq_adiff [simp]: "\m \ n; m \ Nat; n \ Nat\ \ m + -.n = -.(n -- m)" @@ -695,7 +696,7 @@ by (rule intCases3, auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 int_add_assoc4 int_add_assoc5 int_add_assoc6) -text {* Minus sign distributes over addition *} +text \ Minus sign distributes over addition \ lemma minus_distrib_pn_int [simp]: "m \ Nat \ n \ Nat \ -.(m + -.n) = -.m + n" @@ -713,7 +714,7 @@ lemma int_add_minus_distrib [simp]: by (rule intCases2[OF m n], simp_all) -subsection {* Multiplication of integers *} +subsection \ Multiplication of integers \ axiomatization where int_mult_pn_def: "\a \ Nat; b \ Nat\ \ a * -.b = -.(a * b)" @@ -724,12 +725,12 @@ and theorems int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) -text {* Closure *} +text \ Closure \ lemma multIsInt [simp]: "\a \ Int; b \ Int\ \ a * b \ Int" by (rule intCases2[of a b], simp_all add: int_mult_def) -text {* Neutral element *} +text \ Neutral element \ lemma mult_0_right_int [simp]: "a \ Int \ a * 0 = 0" by (rule intCases[of a], simp_all add: int_mult_np_def) @@ -737,12 +738,12 @@ by (rule intCases[of a], simp_all add: int_mult_np_def) lemma mult_0_left_int [simp]: "a \ Int \ 0 * a = 0" by (rule intCases[of a], simp_all add: int_mult_pn_def) -text {* Commutativity *} +text \ Commutativity \ lemma mult_commute_int: "\a \ Int; b \ Int\ \ a * b = b * a" by (rule intCases2[of a b], simp_all add: int_mult_def mult_commute_nat) -text {* Identity element *} +text \ Identity element \ lemma mult_1_right_int [simp]: "a \ Int \ a * 1 = a" by (rule intCases[of a], simp_all add: int_mult_def) @@ -750,14 +751,14 @@ by (rule intCases[of a], simp_all add: int_mult_def) lemma mult_1_left_int [simp]: "a \ Int \ 1 * a = a" by (rule intCases[of a], simp_all add: int_mult_def) -text {* Associativity *} +text \ Associativity \ lemma mult_assoc_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m * (n * p) = (m * n) * p" by(rule intCases3[OF m n p], simp_all add: mult_assoc_nat int_mult_def) -text {* Distributivity *} +text \ Distributivity \ lemma ppn_distrib_left_nat: (* ppn stands for m=positive, n=positive, p=negative *) assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -840,7 +841,7 @@ apply(rule intCases3[OF m n p], apply(simp only: add_mult_distrib_right_nat) done -text {* Minus sign distributes over multiplication *} +text \ Minus sign distributes over multiplication \ lemma minus_mult_left_int: assumes m: "m \ Int" and n: "n \ Int" @@ -853,14 +854,14 @@ lemma minus_mult_right_int: by (rule intCases2[OF m n], simp_all add: int_mult_def) -subsection {* Difference of integers *} +subsection \ Difference of integers \ -text {* +text \ Difference over integers is simply defined as addition of the complement. Note that this difference, noted @{text "-"}, is different from the difference over natural numbers, noted @{text "--"}, even for two natural numbers, because the latter cuts off at $0$. -*} +\ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" From f496b754982cb13ffcc1fd38fb7eb9fa222a2aa8 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:40:15 +0200 Subject: [PATCH 032/167] MAI: update `--` comments in theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 166ea6c3..f899a391 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -18,7 +18,7 @@ subsection \ The minus sign \ consts "minus" :: "c \ c" ("-._" [75] 75) -syntax -- {* syntax for negative naturals *} +syntax (* -- syntax for negative naturals *) "-.0" :: "c" ("-.0") "-.1" :: "c" ("-.1") "-.2" :: "c" ("-.2") @@ -144,7 +144,7 @@ lemma intCases [case_names Positive Negative, cases set: Int]: shows "P" using assms unfolding Int_def by auto --- {* Integer cases over two parameters *} +(* -- Integer cases over two parameters *) lemma intCases2: assumes m: "m \ Int" and n: "n \ Int" and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" @@ -202,10 +202,10 @@ unfolding Int_def by auto subsection \ Predicates ''is positive'' and 'is negative' \ -definition isPos -- {* Predicate ''is positive'' *} +definition isPos (* -- Predicate ''is positive'' *) where "isPos(n) \ \k \ Nat: n = Succ[k]" -definition isNeg -- {* Predicate ''is negative'' *} +definition isNeg (* -- Predicate ''is negative'' *) where "isNeg(n) \ \k \ Nat: n = -.Succ[k]" lemma boolify_isPos [simp]: "boolify(isPos(n)) = (isPos(n))" @@ -251,7 +251,7 @@ lemma negNotNat_isNat: using n by (cases, auto) lemma noNatisNeg [simp]: - "n \ Nat \ isNeg(n) = FALSE" -- {* No natural number is negative *} + "n \ Nat \ isNeg(n) = FALSE" (* -- No natural number is negative *) unfolding isNeg_def using negNotInNat by blast lemma negNat_isNeg [intro]: "\m \ Nat; m \ 0\ \ isNeg(-.m)" @@ -290,10 +290,10 @@ by (auto elim: notIsNeg0_isPos[OF n]) subsection \ Signum function and absolute value \ -definition sgn -- {* signum function *} +definition sgn (* -- signum function *) where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" -definition abs -- {* absolute value *} +definition abs (* -- absolute value *) where "abs(n) \ IF sgn(n) = -.1 THEN -.n ELSE n" lemma sgnInInt [simp]: "n \ Int \ sgn(n) \ Int" @@ -416,7 +416,7 @@ text \ \ lemmas int_leq_pp_def = nat_leq_def - -- {* 'positive-positive' case, ie: both arguments are naturals *} + (* -- 'positive-positive' case, ie: both arguments are naturals *) axiomatization where int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = FALSE" @@ -438,7 +438,7 @@ unfolding isBool_def by auto lemma int_leq_refl [iff]: "n \ Int \ n \ n" by(rule intCases, auto) -lemma eq_leq_bothE: -- {* reduce equality over integers to double inequality *} +lemma eq_leq_bothE: (* -- reduce equality over integers to double inequality *) assumes "m \ Int" and "n \ Int" and "m = n" and "\m \ n; n \ m\ \ P" shows "P" using assms by simp @@ -467,7 +467,7 @@ defs (unchecked overloaded) **) lemmas - int_add_pp_def = nat_add_def -- {* both numbers are positive, ie. naturals *} + int_add_pp_def = nat_add_def (* -- both numbers are positive, ie. naturals *) axiomatization where int_add_pn_def: "\a \ Nat; b \ Nat\ \ a + (-.b) \ IF a \ b THEN -.(b -- a) ELSE a -- b" @@ -477,7 +477,7 @@ and int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) - --{* When we use these definitions, we don't want to unfold the 'pp' case *} + (* -- When we use these definitions, we don't want to unfold the 'pp' case *) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" by (auto simp: int_add_pn_def dest: nat_leq_antisym) @@ -866,7 +866,7 @@ text \ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" -lemma diffIsInt [simp]: -- {* Closure *} +lemma diffIsInt [simp]: (* -- Closure *) "\m \ Int; n \ Int\ \ m - n \ Int" by (simp add: int_diff_def) From 5f5927c1e50cabb43acb8f35475742a865929230 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 18:52:02 +0200 Subject: [PATCH 033/167] MAI: update to `lemmas` for Isabelle2020 in theory `Integers` --- isabelle/Integers.thy | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index f899a391..650c63ea 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -425,7 +425,7 @@ and and int_leq_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ -.b = (b \ a)" -(* theorems int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) +(* lemmas int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) lemma int_boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" @@ -476,7 +476,7 @@ int_add_np_def: "\a \ Nat; b \ Nat\ \ (- and int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" -theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) +lemmas int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) (* -- When we use these definitions, we don't want to unfold the 'pp' case *) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" @@ -723,7 +723,7 @@ and and int_mult_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a * -.b = a * b" -theorems int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) +lemmas int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) text \ Closure \ From 86de33850a45d69ce73284e3eb9577816c11bcfe Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 18:52:42 +0200 Subject: [PATCH 034/167] MAI: rm nonexistent `assms` (raises error in Isabelle2020) --- isabelle/Integers.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 650c63ea..e9159290 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -883,7 +883,7 @@ lemma diff_self_eq_0_int [simp]: "m \ Int \ m - m = 0" by (simp add: int_diff_def) lemma neg_diff_is_diff [simp]: "\m \ Int; n \ Int\ \ -.(m - n) = n - m" -using assms by (simp add: int_diff_def add_commute_int) +by (simp add: int_diff_def add_commute_int) lemma diff_nat_is_add_neg: "\m \ Nat; n \ Nat\ \ m - n = m + -.n" by (simp add: int_diff_def) From c3729500b06f7582a53e14200ccdd85962060d8e Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:17:37 +0200 Subject: [PATCH 035/167] ENH: add theorem `leq_adiff_left_add_right_equiv` to the theory `NatArith`, to be used for proving the theorem `int_add_assoc1` in the theory `Integers`. --- isabelle/NatArith.thy | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 179a0799..912320e2 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -1027,6 +1027,11 @@ lemma leq_adiff_right_add_left: shows "m \ n -- k = (m + k \ n)" using assms by (induct n k rule: diffInduct, simp+) +lemma leq_adiff_left_add_right_equiv: + assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" + shows "(n -- k \ m) = (n \ m + k)" +using assms by (induct n k rule: diffInduct, simp+) + lemma leq_adiff_left_add_right: assumes 1: "n -- p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "n \ m + p" From 397644c57dc4d188f01d4e130417a8b3a463bf0b Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:20:07 +0200 Subject: [PATCH 036/167] MAI: partially update theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 182 +++++++++++++++++++++++++++++++++++++++--- 1 file changed, 169 insertions(+), 13 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index e9159290..d314b323 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -552,19 +552,175 @@ declare leq_neq_iff_less [simplified,simp] lemma int_add_assoc1: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + -.p) = (m + n) + -.p" -apply(rule nat_leq_cases[OF p n]) - using assms apply simp_all - apply(rule nat_leq_cases[of p "m + n"], simp_all) - apply(simp add: adiff_add_assoc[OF _ m n p]) - apply(rule nat_leq_cases[of p "m + n"], simp+) - apply(rule nat_leq_cases[of "p -- n" m], simp+) - apply(rule adiff_add_assoc3, simp+) - apply(rule adiff_add_assoc5, simp+) - apply(rule nat_leq_cases[of "p -- n" m], simp_all) - apply(rule adiff_add_assoc6, simp_all) - apply(simp only: add_commute_nat[of m n]) - apply(rule adiff_adiff_left_nat, simp+) -done +proof - +have s1_1: "n \ p ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s1_1_asm: "n \ p" + have s2_1: "n + -.p = -.(p -- n)" + using s1_1_asm n p int_add_pn_def[of "n" "p"] by auto + have s2_2: "p \ m + n ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s2_2_asm: "p \ m + n" + have s3_1: "m + (n + -.p) = m -- (p -- n)" + proof - + have s4_1: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_2: "m + -.(p -- n) = m -- (p -- n)" + proof - + have s5_1: "p -- n \ m" + using s1_1_asm s2_2_asm + leq_adiff_left_add_right_equiv[of "n" "p" "m"] + p n m by auto + show ?thesis + using s5_1 p n m adiffIsNat + int_add_pn_def by auto + qed + show ?thesis + using s4_1 s4_2 by auto + qed + have s3_2: "(m + n) + -.p = (m + n) -- p" + using s2_2_asm p n m addIsNat int_add_pn_def by auto + have s3_3: "m -- (p -- n) = (m + n) -- p" + using s1_1_asm s2_2_asm + adiff_add_assoc3[of "n" "p" "m"] m n p by auto + show "m + (n + -.p) = (m + n) + -.p" + using s3_1 s3_2 s3_3 by auto + qed + have s2_3: "\ (p \ m + n) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s2_3_asm: "\ (p \ m + n)" + have s3_1: "m + n < p" + using s2_3_asm p m n addIsNat nat_not_leq by auto + have s3_2: "m + n \ p" + using s3_1 by (auto simp: less_def) + have s3_3: "(m + n) + -.p = -.(p -- (m + n))" + using s3_2 m n p addIsNat int_add_pn_def by auto + have s3_4: "p -- n \ m ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s3_4_asm: "p -- n \ m" + have s4_1: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_2: "m + -.(p -- n) = m -- (p -- n)" + using s3_4_asm p n m adiffIsNat int_add_pn_def + by auto + have s4_3: "m + (n + -.p) = m -- (p -- n)" + using s4_1 s4_2 by auto + have s4_4: "m -- (p -- n) = -.(p -- (m + n))" + using m n p s1_1_asm s3_2 s3_4_asm + adiff_add_assoc6 by auto + show "m + (n + -.p) = (m + n) + -.p" + using s3_3 s4_3 s4_4[symmetric] by auto + qed + have s3_5: "\ (p -- n \ m) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s3_5_asm: "\ (p -- n \ m)" + have s4_1: "m \ p -- n" + proof - + have s5_1: "m < p -- n" + using p n m adiffIsNat nat_not_leq s3_5_asm by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_3: "m + -.(p -- n) = -.((p -- n) -- m)" + using s4_1 s3_5_asm p n m adiffIsNat + int_add_pn_def[of "m" "p -- n"] + by auto + have s4_4: "m + (n + -.p) = -.((p -- n) -- m)" + using s4_2 s4_3 by auto + have s4_5: "(p -- n) -- m = p -- (n + m)" + using p n m adiff_adiff_left_nat by auto + have s4_6: "(p -- n) -- m = p -- (m + n)" + using n m add_commute_nat s4_5 by auto + have s4_7: "-.((p -- n) -- m) = -.(p -- (m + n))" + using s4_6 by auto + show "m + (n + -.p) = (m + n) + -.p" + using s4_7 s4_4 s3_3 by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s3_4 s3_5 nat_leq_isBool[of "p -- n" "m"] + isBoolTrueFalse[of "p -- n \ m"] + by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s2_2 s2_3 + nat_leq_isBool[of "p" "m + n"] + isBoolTrueFalse[of "p \ m + n"] + by auto + qed +have s1_2: "\ (n \ p) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s1_2_asm: "\ (n \ p)" + have s2_1: "n + -.p = n -- p" + proof - + have s3_1: "\n \ Nat; p \ Nat\ + \ + n + (-.p) \ + IF n \ p THEN -.(p -- n) ELSE n -- p" + using int_add_pn_def[of "n" "p"] by auto + have s3_2: "n + (-.p) \ n -- p" + using s3_1 s1_2_asm n p by auto + show ?thesis + using s3_2 by auto + qed + have s2_2: "\ ((m + n) \ p)" + proof - + have s3_1: "p < n" + using p n s1_2_asm nat_not_leq by auto + have s3_2: "p < m + n" + using m p n s3_1 trans_less_add2 by auto + show ?thesis + using s3_2 p m n addIsNat nat_not_leq by auto + qed + have s2_3: "(m + n) + -.p = (m + n) -- p" + proof - + have s3_1: "m + n \\in Nat" + using m n addIsNat by auto + have s3_2: "\(m + n) \\in Nat; p \\in Nat\ + \ + (m + n) + (-.p) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE ((m + n) -- p)" + using int_add_pn_def[of "m + n" "p"] by auto + have s3_3: "(m + n) + (-.p) \ (m + n) -- p" + using s3_1 p s3_2 s2_2 by auto + show ?thesis + using s3_3 by auto + qed + have s2_4: "(m + (n + -.p) = (m + n) + -.p) = + (m + (n -- p) = (m + n) -- p)" + using s2_1 s2_3 by auto + have s2_5: "(m + n) -- p = m + (n -- p)" + proof - + have s3_1: "p \ n" + proof - + have s4_1: "p < n" + using s1_2_asm p n nat_not_leq by auto + show ?thesis + using s4_1 by (auto simp: less_def) + qed + show ?thesis + using m n p s3_1 adiff_add_assoc[of "p" "n" "m"] by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s2_4 s2_5[symmetric] by auto + qed +show ?thesis + proof - + have s2_1: "((n \ p) = TRUE) \ ((n \ p) = FALSE)" + using nat_leq_isBool[of "n" "p"] isBoolTrueFalse[of "n \ p"] + by auto + show ?thesis + using s1_1 s1_2 s2_1 by auto + qed +qed lemma int_add_assoc2: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" From fbfb3cf2ffb5e5a1cd1442509ea690284c741a35 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:20:27 +0200 Subject: [PATCH 037/167] STY: shorter lines --- isabelle/NatOrderings.thy | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index 75b44e09..f0407ee6 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -162,7 +162,10 @@ unfolding nat_leq_def using uptoLinear . lemma nat_leq_cases: assumes m: "m \ Nat" and n: "n \ Nat" - and leq: "m \ n \ P" and geq: "\n \ m; n \ m\ \ P" + and leq: "m \ n \ P" and + geq: "\n \ m; + n \ m + \ \ P" shows "P" proof (cases "m \ n") case True thus "P" by (rule leq) From eb6e90978a1ddb2441dc2ea24800e0dde6eac133 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:23:48 +0200 Subject: [PATCH 038/167] DEV: ignore directory `isabelle/output/` which contains the document generated by Isabelle for the TLA+ theories. --- .gitignore | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.gitignore b/.gitignore index 4e62973a..60b1bb25 100644 --- a/.gitignore +++ b/.gitignore @@ -19,3 +19,5 @@ /tools/installer/tlaps-release.sh /tools/installer/tlaps-source-release.sh tlapm + +/isabelle/output/ From 1279bdd12111ac70912b0275f550c7d49239606e Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 20:20:48 +0200 Subject: [PATCH 039/167] MAI: update theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 292 +++++++++++++++++++++++++++++++++++++++--- 1 file changed, 274 insertions(+), 18 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index d314b323..c4398167 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -725,24 +725,280 @@ qed lemma int_add_assoc2: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (-.p + n) = (m + -.p) + n" -using assms apply (simp add: add_commute_int[of "-.p" n]) - using int_add_assoc1[OF m n p] apply simp - apply(rule nat_leq_cases[of p "m + n"], simp_all) - apply(rule nat_leq_cases[OF p m], simp_all) - apply(rule adiff_add_assoc2, simp_all) - apply (simp add: add_commute_int[of "-.(p -- m)" n]) - apply(simp only: add_commute_nat[OF m n]) - apply(rule nat_leq_cases[of "p -- m" n], simp_all) - apply(rule adiff_add_assoc3[symmetric], simp+) - apply(rule adiff_add_assoc5[symmetric], simp+) - apply(rule nat_leq_cases[OF p m], simp_all) - apply (simp add: add_commute_nat[OF m n]) - apply (simp add: add_commute_int[of "-.(p -- m)" n]) - apply(rule nat_leq_cases[of "p -- m" n], simp_all) - apply(simp add: add_commute_nat[OF m n]) - apply(rule adiff_add_assoc6[symmetric], simp+) - apply(rule adiff_adiff_left_nat[symmetric], simp+) -done +proof - +have s1_1: "(m + n) + -.p = (m + -.p) + n" + proof - + have s2_1: "m + n \ p ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s2_1_asm: "m + n \ p" + have s3_1: "(m + n) + -.p = -.(p -- (m + n))" + proof - + have s4_1: "\m + n \ Nat; p \ Nat\ + \ ((m + n) + (-.p)) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE + (m + n) -- p" + using int_add_pn_def[of "m + n" "p"] by auto + have s4_2: "(m + n) + -.p \ -.(p -- (m + n))" + using s4_1 s2_1_asm m n p addIsNat by auto + show ?thesis + using s4_2 by auto + qed + have s3_2: "m \ p" + using s2_1_asm add_leqD1 m n p by auto + have s3_3: "(m + -.p) + n = -.(p -- m) + n" + proof - + have s4_1: "m + -.p = -.(p -- m)" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ -.(p -- m)" + using s5_1 m p s3_2 by auto + show ?thesis + using s5_2 by auto + qed + show ?thesis + using s4_1 by auto + qed + have s3_4: "n \ p -- m ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_4_asm: "n \ p -- m" + have s4_1: "(m + -.p) + n = -.((p -- m) -- n)" + proof - + have s5_1: "-.(p -- m) + n = n + -.(p -- m)" + using p n m adiffIsNat add_commute_pn_nat by auto + have s5_2: "n + -.(p -- m) = -.((p -- m) -- n)" + proof - + have s6_1: "\n \ Nat; p -- m \ Nat\ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s6_2: "n + -.(p -- m) \ -.((p -- m) -- n)" + using s6_1 n m p adiffIsNat s3_4_asm by auto + show ?thesis + using s6_2 by auto + qed + show ?thesis + using s3_3 s5_1 s5_2 by auto + qed + have s4_2: "(p -- m) -- n = p -- (m + n)" + using p n m adiff_adiff_left_nat by auto + show "(m + n) + -.p = (m + -.p) + n" + using s3_1 s4_1 s4_2 by auto + qed + have s3_5: "\ (n \ p -- m) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_5_asm: "\ (n \ p -- m)" + have s4_1: "p -- m \ n" + proof - + have s5_1: "p -- m < n" + using s3_5_asm nat_not_leq p n m adiffIsNat by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "-.(p -- m) + n = n -- (p -- m)" + proof - + have s5_1: "-.(p -- m) + n = n + -.(p -- m)" + using p n m add_commute_pn_nat adiffIsNat by auto + have s5_2: "n + -.(p -- m) = n -- (p -- m)" + proof - + have s6_1: "\n \ Nat; p -- m \ Nat\ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s6_2: "n + -.(p -- m) \ n -- (p -- m)" + using s6_1 s3_5_asm p n m adiffIsNat by auto + show ?thesis + using s6_2 by auto + qed + show ?thesis + using s5_1 s5_2 by auto + qed + have s4_3: "n -- (p -- m) = -.(p -- (m + n))" + using adiff_add_assoc6 m n p s3_2 s2_1_asm + s4_1 add_commute_nat by auto + show "(m + n) + -.p = (m + -.p) + n" + using s3_1 s3_3 s4_2 s4_3[symmetric] by auto + qed + show ?thesis + using s3_4 s3_5 nat_leq_isBool[of "n" "p -- m"] + isBoolTrueFalse[of "n \ p -- m"] by auto + qed + have s2_2: "\ (m + n \ p) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s2_2_asm: "\ (m + n \ p)" + have s3_1: "(m + n) + -.p = (m + n) -- p" + proof - + have s4_1: "\(m + n) \ Nat; p \ Nat\ + \ ((m + n) + (-.p)) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE + (m + n) -- p" + using int_add_pn_def[of "m + n" "p"] by auto + have s4_2: "((m + n) + (-.p)) \ ((m + n) -- p)" + proof - + have s5_1: "m + n \\in Nat" + using m n addIsNat by auto + show "((m + n) + (-.p)) \ ((m + n) -- p)" + using s2_2_asm s4_1[OF s5_1 p] by auto + qed + show ?thesis + using s4_2 by auto + qed + have s3_2: "m \ p ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_2_asm: "m \ p" + have s4_1: "m + -.p = -.(p -- m)" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ -.(p -- m)" + using s5_1 m p s3_2_asm by auto + show ?thesis + using s5_2 by auto + qed + have s4_4: "p \ m + n" + proof - + have s5_1: "p < m + n" + using s2_2_asm p m n addIsNat nat_not_leq by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "n \ p -- m ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s4_2_asm: "n \ p -- m" + have s5_1: "(m + -.p) + n = -.((p -- m) -- n)" + proof - + have s6_1: "(m + -.p) + n = -.(p -- m) + n" + using s4_1 by auto + have s6_2: "-.(p -- m) + n = n + -.(p -- m)" + using p m n adiffIsNat add_commute_pn_nat by auto + have s6_3: "n + -.(p -- m) = -.((p -- m) -- n)" + proof - + have s7_1: "\n \ Nat; p -- m \ Nat + \ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s7_2: "n + -.(p -- m) \ -.((p -- m) -- n)" + using s7_1 p n m adiffIsNat s4_2_asm by auto + show ?thesis + using s7_2 by auto + qed + show ?thesis + using s6_1 s6_2 s6_3 by auto + qed + have s5_2: "-.((p -- m) -- n) = (m + n) -- p" + using adiff_add_assoc5 s4_4 s4_2_asm s3_2_asm + add_commute_nat + m n p by auto + show "(m + n) + -.p = (m + -.p) + n" + using s5_1 s5_2 s3_1 by auto + qed + have s4_3: "\ (n \ p -- m) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s4_3_asm: "\ (n \ p -- m)" + have s5_1: "(m + n) -- p = n -- (p -- m)" + using m n p adiff_add_assoc3[of "m" "p" "n"] + s3_2_asm s4_4 add_commute_nat + by auto + have s5_2: "(m + -.p) + n = n -- (p -- m)" + proof - + have s6_1: "(m + -.p) + n = -.(p -- m) + n" + using s4_1 by auto + have s6_2: "-.(p -- m) + n = n + -.(p -- m)" + using s6_1 n p m adiffIsNat + add_commute_pn_nat by auto + have s6_3: "n + -.(p -- m) = n -- (p -- m)" + proof - + have s7_1: "\n \ Nat; p -- m \ Nat + \ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s7_2: "n + -.(p -- m) \ n -- (p -- m)" + using s7_1 s4_3_asm n p m adiffIsNat + by auto + show ?thesis + using s7_2 by auto + qed + show ?thesis + using s6_1 s6_2 s6_3 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s5_1 s3_1 s5_2 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s4_2 s4_3 nat_leq_isBool[of "n" "p -- m"] + isBoolTrueFalse[of "n \ p -- m"] by auto + qed + have s3_3: "\ (m \ p) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_3_asm: "\ (m \ p)" + have s4_1: "p \ m" + proof - + have s5_1: "p < m" + using s3_3_asm p m nat_not_leq by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "(m + n) -- p = (m -- p) + n" + using adiff_add_assoc2 s4_1 p n m by auto + have s4_3: "m + -.p = m -- p" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ m -- p" + using s5_1 m p s3_3_asm by auto + show ?thesis + using s5_2 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s4_2 s3_1 s4_3 by auto + qed + show ?thesis + using s3_2 s3_3 nat_leq_isBool[of "m" "p"] + isBoolTrueFalse[of "m \ p"] by auto + qed + show ?thesis + using s2_1 s2_2 nat_leq_isBool[of "m + n" "p"] + isBoolTrueFalse[of "m + n \ p"] by auto + qed +have s1_2: "m + (n + -.p) = (m + -.p) + n" + using s1_1 int_add_assoc1[OF m n p] by auto +have s1_3: "n + -.p = -.p + n" + using n p add_commute_pn_nat[of "n" "p"] by auto +show ?thesis + using s1_2 s1_3 by auto +qed declare leq_neq_iff_less [simplified,simp del] From b6ee3052da957bcceaf68e77aa65ab685003182d Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:35:12 +0200 Subject: [PATCH 040/167] MAI: update delimiters in theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 46a9796f..133039c9 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -11,7 +11,7 @@ theory Zenon imports Constant begin -text {* The following lemmas make a cleaner meta-object reification *} +text \ The following lemmas make a cleaner meta-object reification \ lemma atomize_meta_bAll [atomize]: "(\x. (x \ S \ P(x))) \ Trueprop (\ x \ S : P(x))" From 6d12905408b70766588b242b3f13d51363ad0b70 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:36:51 +0200 Subject: [PATCH 041/167] DRAFT: define abbreviations `\\cup` and `\\cap` in theory `Zenon` This syntax was present in the theory `SetTheory` in its version for Isabelle2011-1. The abbreviation in the theory `Zenon` is for use within the theory `Zenon`, and possibly by proofs generated by `zenon`. --- isabelle/Zenon.thy | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 133039c9..f90e1d6f 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -336,6 +336,14 @@ by blast (* INTER -- needed ? *) +abbreviation + cup :: "[c, c] \ c" (infixl "\\cup" 65) + where "a \\cup b \ a \ b" + +abbreviation + cap :: "[c, c] \ c" (infixl "\\cap" 70) + where "a \\cap b \ a \ b" + lemma zenon_in_cup : "x \\in A \\cup B ==> (x \\in A ==> FALSE) ==> (x \\in B ==> FALSE) ==> FALSE" by blast From 71fc7ec24cc71b6c1350859afd21941b774c340d Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:37:39 +0200 Subject: [PATCH 042/167] MAI: partially update theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index f90e1d6f..702cbe57 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -944,7 +944,8 @@ proof - assume h2: "~c ==> P(b) ==> FALSE" have h2x: "~c ==> ~P(b)" using h2 by auto have h3: "~P (IF c THEN a ELSE b)" - by (rule condI [of c "\ x . ~P (x)", OF h1x h2x]) + using atomize_not condI[of c "\ x . ~P (x)", OF h1x h2x] + by (auto) show FALSE by (rule notE [OF h3 main]) qed @@ -1211,7 +1212,7 @@ lemma zenon_tuple_acc_1 : lemma zenon_tuple_acc_2 : "isASeq (r) ==> k \\in Nat ==> 0 < k ==> k \ Len (r) ==> x = r[k] ==> x = Append (r, e) [k]" -by auto +using appendElt1 by auto definition zenon_ss :: "c \ c" where "zenon_ss (n) \ IF n \\in Nat THEN Succ[n] ELSE 1" From 0309414e4db60fffec555305be4f546cdd893f11 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:40:29 +0200 Subject: [PATCH 043/167] DRAFT: use `EnumFuncSet` instead of `[... : ...]` syntax because the required `parse_ast_translation` and `print_ast_translation` in the theory `Tuples` are currently commented, until they are updated to Isabelle2020. --- isabelle/Zenon.thy | 36 +++++++++++++++++++++++++++++------- 1 file changed, 29 insertions(+), 7 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 702cbe57..460cc898 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1263,7 +1263,9 @@ lemma zenon_inrecordsetI1 : assumes hfcn: "isAFcn (r)" assumes hdom: "DOMAIN r = {l1x}" assumes h1: "r[l1x] \\in s1x" - shows "r \\in [l1x : s1x]" + (* TODO + shows "r \\in [l1x : s1x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x}" @@ -1324,7 +1326,9 @@ lemma zenon_inrecordsetI2 : assumes hdom: "DOMAIN r = {l1x, l2x}" assumes h1: "r[l1x] \\in s1x" assumes h2: "r[l2x] \\in s2x" - shows "r \\in [l1x : s1x, l2x : s2x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x}" @@ -1404,7 +1408,9 @@ lemma zenon_inrecordsetI3 : assumes h1: "r[l1x] \\in s1x" assumes h2: "r[l2x] \\in s2x" assumes h3: "r[l3x] \\in s3x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x}" @@ -1503,7 +1509,11 @@ lemma zenon_inrecordsetI4 : assumes h2: "r[l2x] \\in s2x" assumes h3: "r[l3x] \\in s3x" assumes h4: "r[l4x] \\in s4x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x}" @@ -1621,7 +1631,11 @@ lemma zenon_inrecordsetI5 : assumes h3: "r[l3x] \\in s3x" assumes h4: "r[l4x] \\in s4x" assumes h5: "r[l5x] \\in s5x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x}" @@ -1758,7 +1772,11 @@ lemma zenon_inrecordsetI6 : assumes h4: "r[l4x] \\in s4x" assumes h5: "r[l5x] \\in s5x" assumes h6: "r[l6x] \\in s6x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x, l6x}" @@ -1914,7 +1932,11 @@ lemma zenon_inrecordsetI7 : assumes h5: "r[l5x] \\in s5x" assumes h6: "r[l6x] \\in s6x" assumes h7: "r[l7x] \\in s7x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x, l7x : s7x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x, l7x : s7x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x, l6x, l7x}" From 2daf47461e1f573e1b9c310e228b44c8b5050dca Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:57:07 +0200 Subject: [PATCH 044/167] MAI: update definition syntax in theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 179 ++++++++++++++++++++++++++------------------- 1 file changed, 103 insertions(+), 76 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 460cc898..49cbfd67 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -2108,7 +2108,7 @@ lemma zenon_caseother0 : proof - fix P e0 assume h: "P (CASE OTHER -> e0)" - def cas == "CASE OTHER -> e0" + define cas where "cas \ CASE OTHER -> e0" have hh: "P (cas)" using h by (fold cas_def) assume hoth: "P (e0) ==> FALSE" have hb: "(\ i \ DOMAIN <<>> : ~<<>>[i]) = TRUE" by auto @@ -2666,13 +2666,14 @@ lemma zenon_case1 : assumes hoth: "~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c1x" (is "?dcs") + define dcs where "dcs \ c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2682,10 +2683,10 @@ proof - using hoth hh1 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq ( + define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -2740,26 +2741,28 @@ lemma zenon_caseother1 : assumes hoth: "~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq ( + define dcs where "dcs \ c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -2845,13 +2848,14 @@ lemma zenon_case2 : assumes hoth: "~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c2x | c1x" (is "?dcs") + define dcs where "dcs \ c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2861,10 +2865,10 @@ proof - using hoth hh1 hh2 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq ( + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -2923,26 +2927,28 @@ lemma zenon_caseother2 : assumes hoth: "~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3033,13 +3039,14 @@ lemma zenon_case3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3050,10 +3057,12 @@ proof - using hoth hh1 hh2 hh3 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3118,26 +3127,29 @@ lemma zenon_caseother3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3235,13 +3247,14 @@ lemma zenon_case4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3253,10 +3266,12 @@ proof - using hoth hh1 hh2 hh3 hh4 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3327,26 +3342,30 @@ lemma zenon_caseother4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c4x | c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3451,13 +3470,14 @@ lemma zenon_case5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c5x | c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3470,10 +3490,12 @@ proof - using hoth hh1 hh2 hh3 hh4 hh5 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3550,26 +3572,31 @@ lemma zenon_caseother5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ + c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) From 0c872e0f7f4d64e9b369feed9e0c8938e8583b1b Mon Sep 17 00:00:00 2001 From: merz Date: Sun, 3 Jan 2021 09:38:00 +0100 Subject: [PATCH 045/167] pass (mainly) over NatArith --- isabelle/FixedPoints.thy | 2 +- isabelle/Functions.thy | 2 +- isabelle/NatArith.thy | 826 ++++++++++++++---------------------- isabelle/NatOrderings.thy | 30 +- isabelle/Peano.thy | 2 +- isabelle/PredicateLogic.thy | 13 +- isabelle/SetTheory.thy | 8 +- 7 files changed, 338 insertions(+), 545 deletions(-) diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index 7ddf3970..28162aea 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -1,6 +1,6 @@ (* Title: TLA+/FixedPoints.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index b68d19a2..ca831784 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -1,6 +1,6 @@ (* Title: TLA+/Functions.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 912320e2..1e58c382 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -1,9 +1,8 @@ (* Title: TLA+/NatArith.thy Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 - Time-stamp: <2011-10-11 17:39:37 merz> *) section \ Arithmetic (except division) over natural numbers \ @@ -152,7 +151,7 @@ lemma one_is_add_nat: using m apply (rule natCases) using n by (induct, auto)+ -lemma add_eq_self_zero_nat [simp]: +lemma add_eq_self_zero_nat1 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m + n = m) = (n = 0)" using n apply (rule natCases) @@ -160,6 +159,11 @@ using n apply (rule natCases) using m apply induct apply auto done +lemma add_eq_self_zero_nat2 [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(n + m = m) = (n = 0)" +using assms by (simp add: add_commute_nat) + subsection \ Multiplication of natural numbers \ @@ -186,7 +190,7 @@ lemma multnat_0: assumes "m \ Nat" shows "multnat(m)[0] = 0" unfolding multnat_def by (rule primrec_natE, auto simp: assms) -lemma mult_0_nat [simp]: (* -- "neutral element" *) +lemma mult_0_nat [simp]: \ \ neutral element \ assumes n: "n \ Nat" shows "n * 0 = 0" by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) @@ -427,28 +431,26 @@ using assms by (simp add: nat_adiff_def adiffnat_Succ[OF assms]) lemma adiff_0_eq_0_nat [simp]: assumes n: "n \ Nat" shows "0 -- n = 0" -using n apply induct by (simp_all add: adiff_Succ_nat) +using assms by (induct) (simp_all add: adiff_Succ_nat) text \ Reasoning about @{text "m -- m = 0"}, etc. \ lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] -- Succ[n] = m -- n" -using n apply induct -using assms by (auto simp add: adiff_Succ_nat) + using n by induct (auto simp add: m adiff_Succ_nat) lemma adiff_self_eq_0_nat [simp]: assumes m: "m \ Nat" shows "m -- m = 0" -using m apply induct by auto +using m by induct auto text \ Associativity \ lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "(i -- j) -- k = i -- (j + k)" -using i j apply (rule diffInduct) -using k by auto +using i j by (rule diffInduct) (auto simp: k) lemma Succ_adiff_adiff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" @@ -467,8 +469,7 @@ text \ Cancellation rules \ lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(n + m) -- n = m" -using n apply induct -using assms by auto +using n by induct (simp_all add: m) lemma adiff_add_inverse2_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -478,8 +479,7 @@ using assms by (simp add: add_commute_nat [of m n]) lemma adiff_cancel_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k + m) -- (k + n) = m -- n" -using k apply induct -using assms by simp_all +using k by induct (simp_all add: m n) lemma adiff_cancel2_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -489,23 +489,21 @@ using assms by (simp add: add_commute_nat) lemma adiff_add_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "n -- (n + m) = 0" -using n apply induct -using assms by simp_all +using n by induct (simp_all add: m) text \ Difference distributes over multiplication \ lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m -- n) * k = (m * k) -- (n * k)" -using m n apply(rule diffInduct) -using k by simp_all +using m n by (rule diffInduct) (simp_all add: k) lemma adiff_mult_distrib2_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) -(* -- "NOT added as rewrites, since sometimes they are used from right-to-left" *) +\ \ NOT added as rewrites, since sometimes they are used from right to left \ lemmas nat_distrib = add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat @@ -525,137 +523,31 @@ lemma nat_add_left_cancel_leq [simp]: shows "(k + m \ k + n) = (m \ n)" using assms by (induct k) simp_all -lemma nat_add_left_cancel_less [simp]: +lemma nat_add_right_cancel_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m < k + n) = (m < n)" -using k apply induct -using assms by simp_all + shows "(m + k \ n + k) = (m \ n)" +using assms by (induct k) simp_all -lemma nat_add_right_cancel_less [simp]: +lemma nat_add_left_cancel_Succ_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k < n + k) = (m < n)" -using k apply induct -using assms by simp_all + shows "(Succ[k + m] \ k + n) = (Succ[m] \ n)" +using assms by (induct k) simp_all -lemma nat_add_right_cancel_leq [simp]: +(* immediate corollary *) +lemma nat_add_left_cancel_less: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k \ n + k) = (m \ n)" -using k apply induct -using assms by simp_all + shows "(k + m < k + n) = (m < n)" +using assms by simp -lemma add_gr_0 [simp]: - assumes ndom: "n \ Nat" and mdom: "m \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" - proof - - have s1_1: "m + n > 0 ==> m > 0 \ n > 0" - proof - - assume s1_1_asm: "m + n > 0" - have s2_1: "m \ 0 \ n \ 0" - proof - - have s3_1: "m = 0 \ n = 0 ==> FALSE" - proof - - assume s3_1_asm: "m = 0 \ n = 0" - have s4_1: "m + n = 0" - using s3_1_asm ndom mdom add_is_0_nat by auto - have s4_2: "m + n \ 0" - unfolding less_def using s1_1_asm by auto - show "FALSE" - using s4_1 s4_2 by auto - qed - show ?thesis - using s3_1 by auto - qed - show "m > 0 \ n > 0" - using s2_1 ndom mdom nat_gt0_not0 by auto - qed - have s1_2: "m > 0 \ n > 0 ==> m + n > 0" - proof - - assume s1_2_asm: "m > 0 \ n > 0" - have s2_1: "m > 0 ==> m + n > 0" - proof - - assume s2_1_asm: "m > 0" - have s3_1: "m + n > 0 + n" - proof - - have s4_1: "0 < m" - using s2_1_asm by auto - have s4_2: "(0 + n < m + n) = (0 < m)" - using ndom mdom zeroIsNat nat_add_right_cancel_less - by blast - have s4_3: "0 + n < m + n" - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - have s3_2: "0 + n = n" - using ndom add_0_nat add_commute_nat zeroIsNat - by auto - have s3_3: "n >= 0" - using ndom nat_zero_leq by auto - have s3_4: "m + n >= 0 + n" - using s3_1 by (auto simp: less_def) - have s3_5: "m + n >= n" - using s3_4 s3_2 by auto - have s3_6: "m + n >= 0" - using s3_5 s3_3 nat_leq_trans - zeroIsNat ndom mdom addIsNat by auto - have s3_7: "m + n = 0 ==> FALSE" - proof - - assume s3_7_asm: "m + n = 0" - have s4_1: "m = 0 \ n = 0" - using s3_7_asm mdom ndom add_is_0_nat - by auto - show "FALSE" - using s4_1 s2_1_asm by (auto simp: less_def) - qed - show "m + n > 0" - unfolding less_def - using s3_6 s3_7 by auto - qed - have s2_2: "n > 0 ==> m + n > 0" - proof - - assume s2_2_asm: "n > 0" - have s3_1: "m + n > m + 0" - proof - - have s4_1: "0 < n" - using s2_2_asm by auto - have s4_2: "(m + 0 < m + n) = (0 < n)" - using ndom mdom zeroIsNat nat_add_left_cancel_less - by blast - have s4_3: "m + 0 < m + n" - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - have s3_2: "m + 0 = m" - using mdom add_0_nat by auto - have s3_3: "m >= 0" - using mdom nat_zero_leq by auto - have s3_4: "m + n >= m + 0" - using s3_1 by (auto simp: less_def) - have s3_5: "m + n >= m" - using s3_4 s3_2 by auto - have s3_6: "m + n >= 0" - using s3_5 s3_3 nat_leq_trans - zeroIsNat ndom mdom addIsNat by auto - have s3_7: "m + n = 0 ==> FALSE" - proof - - assume s3_7_asm: "m + n = 0" - have s4_1: "m = 0 \ n = 0" - using s3_7_asm mdom ndom add_is_0_nat - by auto - show "FALSE" - using s4_1 s2_2_asm by (auto simp: less_def) - qed - show "m + n > 0" - unfolding less_def - using s3_6 s3_7 by auto - qed - show "m + n > 0" - using s1_2_asm s2_1 s2_2 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed +lemma nat_add_right_cancel_Succ_less [simp]: + assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" + shows "(Succ[m + k] \ n + k) = (Succ[m] \ n)" +using assms by (induct k) simp_all + +lemma nat_add_right_cancel_less: + assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" + shows "(m + k < n + k) = (m < n)" +using assms by simp lemma less_imp_Succ_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -687,9 +579,9 @@ using assms by simp_all subsubsection \ (Partially) Ordered Groups \ -(* -- "The two following lemmas are just ``one half'' of - @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above." *) +\ \ The two following lemmas are just ``one half'' of + @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} + proved above." \ lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" @@ -710,41 +602,16 @@ using assms nat_leq_trans[of "a + c" "b + c" "b + d"] by simp -(* -- "Similar for strict less than." *) +\ \ Similar for strict less than. \ lemma add_less_left_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes "a \ Nat" and "b \ Nat" and "c \ Nat" shows "a < b \ c + a < c + b" - proof - - have s1_1: "a <= b ==> c + a <= c + b" - using a b c add_leq_left_mono by auto - have s1_2: "a \ b ==> (c + a) \ (c + b)" - using a b c add_left_cancel_nat by auto - have s1_3: "a < b ==> c + a < c + b" - proof - - assume s1_3_asm: "a < b" - have s2_1: "a <= b" - using s1_3_asm by (auto simp: less_def) - have s2_2: "a \ b" - using s1_3_asm by (auto simp: less_def) - have s2_3: "c + a <= c + b" - using s1_1 s2_1 by auto - have s2_4: "(c + a) \ (c + b)" - using s1_2 s2_2 by auto - show "c + a < c + b" - using s2_3 s2_4 by (auto simp: less_def) - qed - show "a < b \ c + a < c + b" - using a b c s1_3 by auto - qed -(* -using assms add_leq_left_mono[OF "a" "b" "c"] add_left_cancel_nat[OF "c" "b" "a"] - by (auto simp: nat_less_Succ_iff_leq nat_gt0_not0) -*) + using assms by (simp add: add_leq_left_mono less_def) lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" -using assms add_less_left_mono add_commute_nat by auto +using assms by (auto simp: add_less_left_mono add_commute_nat) text \ Strict monotonicity in both arguments \ lemma add_less_mono: @@ -784,6 +651,36 @@ lemma leq_add2 [simp]: shows "n \ m + n" using assms add_leq_right_mono [of 0 m n] by simp +lemma trans_leq_add1 [elim]: + assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m+k" +using assms by (simp add: nat_leq_trans) + +lemma trans_leq_add2 [elim]: + assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ k+m" +using assms by (simp add: nat_leq_trans) + +lemma add_leqD1 [elim]: + assumes "n+k \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m" +using assms by (simp add: nat_leq_trans[of "n" "n+k" "m"]) + +lemma add_leqD2 [elim]: + assumes "k+n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m" +using assms by (simp add: nat_leq_trans[of "n" "k+n" "m"]) + +lemma add_Succ_leqD1 [elim]: + assumes "Succ[n+k] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "Succ[n] \ m" +using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[n+k]" "m"]) + +lemma add_Succ_leqD2 [elim]: + assumes "Succ[k+n] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "Succ[n] \ m" +using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[k+n]" "m"]) + lemma less_add_Succ1: assumes "i \ Nat" and "m \ Nat" shows "i < Succ[i + m]" @@ -799,99 +696,88 @@ lemma less_iff_Succ_add: shows "(m < n) = (\k \ Nat: n = Succ[m + k])" using assms by (auto intro!: less_imp_Succ_add) -lemma trans_leq_add1: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i \ j + m" -using assms by (auto elim: nat_leq_trans) - -lemma trans_leq_add2: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i \ m + j" -using assms by (auto elim: nat_leq_trans) - lemma trans_less_add1: - assumes ij: "i < j" and idom: "i \ Nat" and - jdom: "j \ Nat" and mdom: "m \ Nat" + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < j + m" - proof - - have s1_1: "i <= i + m" - proof - - have s2_1: "0 <= m" - using mdom nat_zero_leq by auto - have s2_2: "i + 0 <= i + m" - using idom zeroIsNat mdom s2_1 nat_add_left_cancel_leq by auto - have s2_3: "i + 0 = i" - using idom add_0_nat by auto - show ?thesis - using s2_2 s2_3 by auto - qed - have s1_2: "i + m < j + m" - using ij idom jdom mdom add_less_right_mono by auto - have s1_3: "i + m \ Nat" - using idom mdom addIsNat by auto - have s1_4: "j + m \ Nat" - using jdom mdom addIsNat by auto - show ?thesis - using s1_1 s1_2 idom s1_3 s1_4 nat_leq_less_trans - by auto - qed +using assms by auto lemma trans_less_add2: assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" -using assms trans_less_add1 add_commute_nat by auto +using assms by auto lemma add_lessD1: assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "i < k" -using assms by (intro nat_leq_less_trans[of "i" "i+j" "k"], simp+) +using assms by auto -lemma not_add_less1 [simp]: +lemma add_lessD2: + assumes "j+i < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" + shows "i < k" +using assms by auto + +lemma not_Succ_add_le1: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(Succ[i + j] \ i) = FALSE" +using assms by simp + +lemma not_Succ_add_le2: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(Succ[j + i] \ i) = FALSE" +using assms by simp + +lemma not_add_less1: assumes i: "i \ Nat" and j: "j \ Nat" shows "(i + j < i) = FALSE" -by (auto dest: add_lessD1[OF _ i j i]) +using assms by simp -lemma not_add_less2 [simp]: +lemma not_add_less2: assumes "i \ Nat" and "j \ Nat" shows "(j + i < i) = FALSE" -using assms by (simp add: add_commute_nat) - -lemma add_leqD1: - assumes "m + k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "m \ n" -using assms by (intro nat_leq_trans[of "m" "m+k" "n"], simp+) - -lemma add_leqD2: - assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k \ n" -using assms by (intro nat_leq_trans[of "k" "m+k" "n"], simp+) +using assms by simp lemma add_leqE: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "(m \ n \ k \ n \ R) \ R" using assms by (blast dest: add_leqD1 add_leqD2) -lemma leq_add_left_false [simp]: - assumes mdom: "m \ Nat" and ndom: "n \ Nat" and n0: "n \ 0" - shows "m + n \ m = FALSE" - proof - - have s1_1: "m + n \ m = (m + n < m \ m + n = m)" - proof - - define p where "p \ m + n" - have s2_1: "p \ m = (p < m \ p = m)" - using mdom nat_leq_less by auto - show ?thesis - using s2_1 by (auto simp: p_def) - qed - have s1_2: "(m + n = m) = (n = 0)" - using mdom ndom add_eq_self_zero_nat[of "m" "n"] by auto - have s1_3: "(m + n <= m) = (m + n < m)" - using n0 s1_1 s1_2 by auto - have s1_4: "(m + n < m) = FALSE" - using mdom ndom not_add_less1[of "m" "n"] by auto - show ?thesis - using s1_3 s1_4 by auto - qed +lemma leq_add_self1 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(m+n \ m) = (n=0)" +proof - + { + assume "m+n \ m" + with assms have "m+n = m" by (intro nat_leq_antisym) simp_all + } + with assms show ?thesis by auto +qed + +lemma leq_add_self2 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(n+m \ m) = (n=0)" +proof - + { + assume "n+m \ m" + with assms have "n+m = m" by (intro nat_leq_antisym) simp_all + } + with assms show ?thesis by auto +qed + +lemma add_gt_0: + assumes "m \ Nat" and "n \ Nat" + shows "(m + n > 0) = (m > 0 \ n > 0)" +proof - + from assms have "m + n \ Nat" by simp + hence "(m + n > 0) = (m + n \ 0)" by (rule nat_gt0_not0) + also from assms have "\ = (m \ 0 \ n \ 0)" by simp + finally show ?thesis using assms by (simp only: nat_gt0_not0) +qed + +(* The above lemma follows from the following simplification rule. *) +lemma add_ge_1 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(m + n \ 1) = (m \ 1 \ n \ 1)" + using assms by (auto simp add: add_gt_0) subsubsection \ More results about arithmetic difference \ @@ -936,33 +822,31 @@ lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") proof - - have 1: "?lhs \ ?rhs" + have "?lhs \ ?rhs" proof assume mn: "m \ n" with m n have "n = m + (n -- m)" by simp with m n show "?rhs" by blast qed - from assms have 2: "?rhs \ ?lhs" by auto - from 1 2 assms show ?thesis by blast + moreover + from assms have "?rhs \ ?lhs" by auto + ultimately show ?thesis by blast qed lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" - proof - - have s1_1: "j -- n \ Nat" - using j n by simp+ - have s1_2: "j -- n \ j" - using j n by simp+ - show ?thesis - using j k jk s1_1 s1_2 nat_leq_less_trans by auto - qed +proof - + from j n have "j -- n \ Nat" by simp + moreover + from j n have s1_2: "j -- n \ j" by simp + ultimately show ?thesis using j k jk nat_leq_less_trans by auto +qed lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" shows "0 < n \ n -- Succ[i] < n" -using n apply cases -using i by auto +using n by cases (auto simp: i) lemma adiff_add_assoc: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" @@ -1003,23 +887,24 @@ lemma adiff_is_0_eq' (*[simp]*): shows "m -- n = 0" using assms by simp -lemma zero_less_adiff [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(0 < n -- m) = (m < n)" +lemma one_leq_adiff [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(1 \ n -- m) = (Succ[m] \ n)" by (induct m n rule: diffInduct, simp_all add: assms) +lemma zero_less_adiff: + assumes "m \ Nat" and "n \ Nat" + shows "(0 < n -- m) = (m < n)" +using assms by simp + lemma less_imp_add_positive: - assumes ij: "i < j" and i: "i \ Nat" and j: "j \ Nat" + assumes "i < j" and "i \ Nat" and "j \ Nat" shows "\k \ Nat: 0 < k \ i + k = j" proof - - have s1_1: "i + (j -- i) = j" - using i j ij le_add_adiff_inverse by (auto simp: less_def) - have s1_2: "j -- i \ Nat" - using assms by simp+ - have s1_3: "0 < j -- i" - using assms zero_less_adiff by auto - show ?thesis - using s1_1 s1_2 s1_3 by blast + thm leq_iff_add + from assms obtain k where k: "k \ Nat" "j = Succ[i]+k" + by (auto simp: leq_iff_add) + with assms show ?thesis by force qed lemma leq_adiff_right_add_left: @@ -1043,16 +928,10 @@ lemma leq_adiff_trans: apply(rule nat_leq_trans[of "p -- n" p m]) using assms adiff_leq_self[OF p n] by simp_all -lemma leq_adiff_right_false [simp]: - assumes "n \ 0" "n \ m" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ m -- n = FALSE" -using assms by (simp add: leq_adiff_right_add_left[OF _ m m n]) - -lemma leq_adiff_right_imp_0: - assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" - shows "p = 0" -using p h apply (induct) -using n by auto +lemma leq_adiff_right [simp]: + assumes "n \ m" and "m \ Nat" and "n \ Nat" + shows "(m \ m -- n) = (n = 0)" + using assms by (simp add: leq_adiff_right_add_left) subsubsection \ Monotonicity of Multiplication \ @@ -1062,14 +941,12 @@ subsubsection \ Monotonicity of Multiplication \ lemma mult_leq_left_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "c * a \ c * b" -using c apply induct -using 1 a b by (simp_all add: add_leq_mono) +using c by induct (simp_all add: add_leq_mono 1 a b) lemma mult_leq_right_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * c \ b * c" -using c apply induct -using 1 a b by (simp_all add: add_leq_mono add_commute_nat) +using c by induct (simp_all add: add_leq_mono add_commute_nat 1 a b) text \ @{text "\"} monotonicity, BOTH arguments \ @@ -1081,280 +958,193 @@ using assms mult_leq_right_mono[OF _ i j k] mult_leq_left_mono[OF _ k l j] nat_leq_trans[of "i * k" "j * k" "j * l"] -by simp + by simp + +lemma self_leq_mult_left: + assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" + shows "a \ a * b" +proof - + have "1 \ Nat" by simp + from 1 this b a have "a * 1 \ a * b" by (rule mult_leq_left_mono) + with a b show ?thesis by simp +qed -text \ strict, in 1st argument \ +lemma self_leq_mult_right: + assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" + shows "a \ b * a" +using assms by (simp add: self_leq_mult_left mult_commute_nat) -lemma mult_less_left_mono: - assumes ij: "i < j" and k0: "0 < k" and - idom: "i \ Nat" and jdom: "j \ Nat" and kdom: "k \ Nat" - shows "k * i < k * j" - proof - - have s1_1: "Succ[0] * i < Succ[0] * j" - using mult_1_left_nat ij idom jdom by auto - have s1_2: "\ n. \ - n \ Nat; - Succ[n] * i < Succ[n] * j - \ \ - Succ[Succ[n]] * i < Succ[Succ[n]] * j" - proof - - fix "n" :: "c" - assume s1_2_ndom: "n \ Nat" - assume s1_2_induct: "Succ[n] * i < Succ[n] * j" - have s2_1: "Succ[n] * i + i < Succ[n] * j + j" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "Succ[n] * i \ Nat" - using s3_1 idom multIsNat by auto - have s3_3: "Succ[n] * j \ Nat" - using s3_1 jdom multIsNat by auto - show ?thesis - using s3_2 s3_3 idom jdom s1_2_induct ij - add_less_mono by auto - qed - have s2_2: "i * Succ[n] + i < j * Succ[n] + j" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "Succ[n] * i = i * Succ[n]" - using s3_1 idom mult_commute_nat by auto - have s3_3: "Succ[n] * j = j * Succ[n]" - using s3_1 jdom mult_commute_nat by auto - show ?thesis - using s2_1 s3_2 s3_3 by auto - qed - have s2_3: "i * Succ[Succ[n]] < j * Succ[Succ[n]]" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "i * Succ[Succ[n]] = i * Succ[n] + i" - using idom s3_1 multnat_Succ by (auto simp: nat_mult_def) - have s3_3: "j * Succ[Succ[n]] = j * Succ[n] + j" - using jdom s3_1 multnat_Succ by (auto simp: nat_mult_def) - show ?thesis - using s2_2 s3_2 s3_3 by auto - qed - have s2_4: "i * Succ[Succ[n]] = Succ[Succ[n]] * i" - proof - - have s3_1: "Succ[Succ[n]] \ Nat" - using s1_2_ndom succIsNat by auto - show ?thesis - using s3_1 idom mult_commute_nat by auto - qed - have s2_5: "j * Succ[Succ[n]] = Succ[Succ[n]] * j" - proof - - have s3_1: "Succ[Succ[n]] \ Nat" - using s1_2_ndom succIsNat by auto - show ?thesis - using s3_1 jdom mult_commute_nat by auto - qed - show "Succ[Succ[n]] * i < Succ[Succ[n]] * j" - using s2_3 s2_4 s2_5 by auto - qed - have s1_3: "\ n \ Nat: Succ[n] * i < Succ[n] * j" - using s1_1 s1_2 - natInduct[of "\ q. Succ[q] * i < Succ[q] * j"] by auto - show ?thesis - using s1_3 nat_gt0_iff_Succ[of "k"] k0 kdom by auto - qed +text \ Similar lemmas for @{text "<"} \ + +lemma mult_Succ_leq_right_mono1: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + shows "Succ[a * Succ[c]] \ b * Succ[c]" +using c proof induct + case 0 + with assms show ?case by simp +next + fix n + assume n: "n \ Nat" and ih: "Succ[a * Succ[n]] \ b * Succ[n]" + from ab a b have 1: "a \ b" + by (simp add: nat_leq_trans[of "a" "Succ[a]" "b"]) + from a n have "Succ[a * Succ[Succ[n]]] = Succ[a * Succ[n]] + a" + by simp + also from a b n ih have "\ \ b* Succ[n] + a" + by simp + also from a b n 1 have "\ \ b * Succ[n] + b" + by simp + finally show "Succ[a * Succ[Succ[n]]] \ b * Succ[Succ[n]]" + using b n by simp +qed + +lemma mult_Succ_leq_right_mono2: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" + shows "Succ[a * c] \ b * c" +proof - + from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" + by (blast dest: nat_ge1_implies_Succ) + with ab a b show ?thesis by (auto dest: mult_Succ_leq_right_mono1[of a b k]) +qed lemma mult_less_right_mono: - assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "i * k < j * k" - proof - - have s1_1: "k * i < k * j" - using 1 i j k mult_less_left_mono by auto - have s1_2: "k * i = i * k" - using i k mult_commute_nat by auto - have s1_3: "k * j = j * k" - using j k mult_commute_nat by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed + assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" + shows "a * c < b * c" +using assms by (simp add: mult_Succ_leq_right_mono2) -lemma nat_0_less_mult_iff [simp]: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(0 < i * j) = (0 < i \ 0 < j)" -using i apply induct - using j apply simp - using j apply(induct, simp_all) -done +lemma mult_Succ_leq_left_mono1: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + shows "Succ[Succ[c] * a] \ Succ[c] * b" +proof - + have "Succ[Succ[c] * a] = Succ[a * Succ[c]]" + using a c by (simp add: mult_commute_nat) + also have "\ \ b * Succ[c]" + using assms by (rule mult_Succ_leq_right_mono1) + also have "\ = Succ[c] * b" + using b c by (simp add: mult_commute_nat) + finally show ?thesis . +qed -lemma one_leq_mult_iff (*[simp]*): +lemma mult_Succ_leq_left_mono2: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" + shows "Succ[c * a] \ c * b" +proof - + from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" + by (blast dest: nat_ge1_implies_Succ) + with ab a b show ?thesis by (auto dest: mult_Succ_leq_left_mono1[of a b k]) +qed + +lemma mult_less_left_mono: + assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" + shows "c * a < c * b" +using assms by (simp add: mult_Succ_leq_left_mono2) + +lemma one_leq_mult_iff[simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 \ m * n) = (1 \ m \ 1 \ n)" using assms nat_gt0_not0 by simp -lemma mult_less_cancel_left [simp]: +lemma nat_0_less_mult_iff: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(0 < i * j) = (0 < i \ 0 < j)" +using assms by simp + +lemma mult_Succ_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m < k * n) = (0 < k \ m < n)" -proof (auto intro!: mult_less_left_mono[OF _ _ m n k]) - assume "k*m < k*n" - from k m n this show "0 < k" by (cases k, simp_all) + shows "(Succ[k * m] \ k * n) = (1 \ k \ Succ[m] \ n)" +proof (auto intro: mult_Succ_leq_left_mono2[OF _ m n k _]) + assume "Succ[k * m] \ k * n" + from k m n this show "1 \ k" by (cases k) simp_all next - assume 1: "k*m < k*n" - show "m < n" + assume 1: "Succ[k * m] \ k * n" + show "Succ[m] \ n" proof (rule contradiction) - assume "\(m < n)" + assume "\(Succ[m] \ n)" with m n k have "k*n \ k*m" by (simp add: nat_not_order_simps mult_leq_left_mono) - with m n k have "\(k*m < k*n)" by (simp add: nat_not_order_simps) + with m n k have "\(Succ[k*m] \ k*n)" by (simp add: nat_not_order_simps) with 1 show "FALSE" by simp qed qed -lemma mult_less_cancel_right [simp]: +lemma mult_less_cancel_left: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k < n * k) = (0 < k \ m < n)" -proof (auto intro!: mult_less_right_mono[OF _ _ m n k]) - assume "m*k < n*k" - from k m n this show "0 < k" by (cases k, simp_all) + shows "(k * m < k * n) = (0 < k \ m < n)" +using assms by simp + +lemma mult_Succ_leq_cancel_right [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" + shows "(Succ[m * k] \ n * k) = (1 \ k \ Succ[m] \ n)" +proof (auto intro: mult_Succ_leq_right_mono2[OF _ m n k _]) + assume "Succ[m * k] \ n * k" + from k m n this show "1 \ k" by (cases k) simp_all next - assume 1: "m*k < n*k" - show "m < n" + assume 1: "Succ[m * k] \ n * k" + show "Succ[m] \ n" proof (rule contradiction) - assume "\(m < n)" + assume "\(Succ[m] \ n)" with m n k have "n*k \ m*k" by (simp add: nat_not_order_simps mult_leq_right_mono) - with m n k have "\(m*k < n*k)" by (simp add: nat_not_order_simps) + with m n k have "\(Succ[m*k] \ n*k)" by (simp add: nat_not_order_simps) with 1 show "FALSE" by simp qed qed -lemma mult_less_self_left [dest]: - assumes less: "n * k < n" and n: "n \ Nat" and k: "k \ Nat" +lemma mult_less_cancel_right: + assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" + shows "(m * k < n * k) = (0 < k \ m < n)" +using assms by simp + +lemma mult_Succ_leq_self_left [dest]: + assumes nk: "Succ[n*k] \ n" and n: "n \ Nat" and k: "k \ Nat" + shows "k = 0" +proof (rule contradiction) + assume "k \ 0" + with k nat_neq0_conv have "1 \ k" by simp + with n k have "n \ n * k" by (rule self_leq_mult_left) + with n k have "\(Succ[n*k] \ n)" by (simp add: nat_not_order_simps) + from this nk show "FALSE" .. +qed + +lemma mult_less_self_left: + assumes "n * k < n" and "n \ Nat" and "k \ Nat" shows "k=0" - proof - - have s1_1: "1 \ k ==> FALSE" - proof - - assume s1_1_asm: "1 \ k" - have s2_1: "n * 1 \ n * k" - using n k oneIsNat s1_1_asm mult_leq_left_mono[of "1" "k" "n"] - by auto - have s2_2: "n \ n * k" - proof - - have s3_1: "n * 1 = n" - using n mult_1_right_nat by auto - show ?thesis - using s2_1 s3_1 by auto - qed - define p where "p \ n" - define q where "q \ n * k" - have s2_3: "q < p" - unfolding p_def q_def using less by auto - have s2_4: "n < p" - proof - - have s3_1: "q \ Nat" - unfolding q_def using n k multIsNat by auto - have s3_2: "p \ Nat" - unfolding p_def using n by auto - have s3_3: "n \ q" - unfolding q_def using s2_2 by auto - show ?thesis - using s3_1 s3_2 n s3_3 s2_3 - nat_leq_less_trans[of "n" "q" "p"] - by auto - qed - have s2_5: "n \ n" - using s2_4 by (auto simp: less_def p_def) - show "FALSE" - using s2_5 by auto - qed - have s1_2: "\ (1 \ k)" - using s1_1 by auto - show ?thesis - using s1_2 k nat_not_leq_one[of "k"] by auto - qed +using assms by auto -lemma mult_less_self_right [dest]: +lemma mult_Succ_leq_self_right [dest]: + assumes less: "Succ[k*n] \ n" and n: "n \ Nat" and k: "k \ Nat" + shows "k=0" +using assms by (auto simp: mult_commute_nat[OF k n]) + +lemma mult_less_self_right: assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" - proof - - have s1_1: "n * k < n" - proof - - have s2_1: "k * n = n * k" - using n k mult_commute_nat by auto - show ?thesis - using less s2_1 by auto - qed - show ?thesis - using s1_1 n k mult_less_self_left by auto - qed +using assms by auto lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" - proof - - have s1_1: "(k * m \ k * n) ==> (k = 0 \ m \ n)" - proof - - assume s1_1_asm: "k * m \ k * n" - have s2_1: "\ (k = 0) ==> m \ n" - proof - - assume s2_1_asm: "\ (k = 0)" - have s3_1: "k > 0" - using k nat_gt0_not0 s2_1_asm by auto - have s3_2: "\(m \ n) ==> FALSE" - proof - - assume s3_2_asm: "\(m \ n)" - have s2_1: "n < m" - using s3_2_asm m n nat_not_leq[of "m" "n"] by auto - have s2_2: "k * n < k * m" - using s3_1 s2_1 m n k mult_less_left_mono[of "n" "m" "k"] by auto - with m n k have s2_3: "\ (k * m \ k * n)" - by (simp add: nat_not_order_simps) - show "FALSE" - using s2_2 s2_3 s1_1_asm by auto - qed - show "m \ n" - using s3_2 by auto - qed - show "k = 0 \ m \ n" - using s2_1 by auto - qed - have s1_2: "(k = 0 \ m \ n) ==> (k * m \ k * n)" - proof - - assume s1_2_asm: "k = 0 \ m \ n" - have s2_1: "k = 0 ==> k * m \ k * n" - proof - - assume s2_1_asm: "k = 0" - have s3_1: "k * m = 0" - using s2_1_asm m mult_0_left_nat by auto - have s3_2: "k * n = 0" - using s2_1_asm n mult_0_left_nat by auto - show "k * m \ k * n" - using s3_1 s3_2 by auto - qed - have s2_2: "m \ n ==> k * m \ k * n" - proof - - assume s2_2_asm: "m \ n" - show "k * m \ k * n" - using s2_2_asm m n k mult_leq_left_mono by auto - qed - show "k * m \ k * n" - using s1_2_asm s2_1 s2_2 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed +proof - + { + assume 1: "k*m \ k*n" and 2: "k \ 0" and 3: "\(m \ n)" + from k 2 nat_gt0_not0 have "k > 0" by simp + with 3 k m n have "\(k * m \ k * n)" by (simp add: nat_not_order_simps) + from this 1 have "FALSE" .. + } + moreover have "k=0 \ k*m \ k*n" + using assms by simp + moreover have "m \ n \ k*m \ k*n" + using assms by (simp add: mult_leq_left_mono) + ultimately show ?thesis by blast +qed lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" - proof - - have s1_1: "(k * m \ k * n) = (k = 0 \ m \ n)" - using m n k mult_leq_cancel_left by auto - have s1_2: "k * m = m * k" - using k m mult_commute_nat by auto - have s1_3: "k * n = n * k" - using k n mult_commute_nat by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed +using assms by (simp add: mult_commute_nat) lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] - nat_zero_less_Succ[of "k"] by auto + nat_zero_less_Succ[of "k"] by auto lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" @@ -1377,7 +1167,7 @@ lemma mult_eq_self_implies_10: shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") proof - from assms have "(m*n = m) = (m*n = m*1)" by simp - also have "... = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) + also have "\ = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) finally show ?thesis . qed diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f0407ee6..97e45a2f 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -1,6 +1,6 @@ (* Title: TLA+/NatOrderings.thy Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) @@ -34,6 +34,11 @@ notation (ASCII) leq (infixl "<=" 50) and geq (infixl ">=" 50) +lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" +by (rule ssubst) + +lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" +by (rule subst) subsection \ Operator definitions and generic facts about @{text "<"} \ @@ -50,6 +55,12 @@ by (simp add: less_def) lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" +by (rule ssubst) + +lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" +by (rule subst) + lemma less_imp_leq [elim!]: "a < b \ a \ b" unfolding less_def by simp @@ -98,16 +109,22 @@ unfolding nat_leq_def by (rule zeroInUpto) lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" by (simp add: nat_leq_def uptoZero) -lemma nat_leq_SuccI [elim!(*,simp*)]: +lemma nat_leq_SuccI [(*elim!,*)intro]: assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" shows "m \ Succ[n]" using assms by (auto simp: nat_leq_def uptoSucc) +(* Do not make this "simp": interferes with other rules, e.g. simplifying "Succ[m] \ Succ[n]" *) lemma nat_leq_Succ: assumes (*"m \ Nat" and*) "n \ Nat" shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" using assms by (auto simp: nat_leq_def uptoSucc) +lemma nat_is_leq_Succ [simp]: + assumes "n \ Nat" + shows "n \ Succ[n]" +using assms by (simp add: nat_leq_Succ) + lemma nat_leq_SuccE [elim]: assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" and "m \ n \ P" and "m = Succ[n] \ P" @@ -189,7 +206,7 @@ lemma nat_leq_induct: \ \sometimes called ``complete induction'' shows "\n\Nat : P(n)" proof - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp add: nat_leq_Succ) + by (intro natInduct, auto simp: nat_leq_Succ) thus ?thesis by (blast dest: nat_leq_refl) qed @@ -210,8 +227,6 @@ using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) text \ Reduce @{text "\"} to @{text "<"}. \ lemma nat_leq_less: - fixes "m" :: "c" - fixes "n" :: "c" assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" using assms by (auto simp: less_def) @@ -301,9 +316,6 @@ proof - qed lemma nat_leq_less_trans [trans]: - fixes "k" :: "c" - fixes "m" :: "c" - fixes "n" :: "c" assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) @@ -522,7 +534,7 @@ lemma nat_ge1_implies_Succ: lemma nat_gt0_iff_Succ: assumes n: "n \ Nat" - shows "(0 < n) = (\m \ Nat: n = Succ[m])" + shows "(1 \ n) = (\m \ Nat: n = Succ[m])" using n by (auto dest: nat_ge1_implies_Succ) (* diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 1f19fe45..90fb821e 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -1,6 +1,6 @@ (* Title: TLA+/Peano.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index 2a23721c..afade712 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -1,8 +1,8 @@ (* Title: TLA+/PredicateLogic.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2017 + Version: Isabelle2020 *) section \First-Order Logic for TLA+\ @@ -1079,11 +1079,6 @@ ML \ \ -(* -setup hypsubst_setup -setup Classical.setup -*) - declare refl [intro!] and trueI [intro!] and conjI [intro!] @@ -1124,8 +1119,6 @@ ML \ ); \ -(* setup Blast.setup *) - (*** TEST DATA *** lemma "x = y \ y = x" @@ -1615,7 +1608,7 @@ text \ lemma boolifyFalseI [intro]: "\ A \ boolify(A) = FALSE" by (auto simp add: boolify_def) -text \ idempotence of @{text boolify} \ +\ \ idempotency of @{text boolify} \ lemma boolifyBoolify [simp]: "boolify(boolify(x)) = boolify(x)" by (auto simp add: boolify_def) diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index 8729d7e5..707897f4 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -1,6 +1,6 @@ (* Title: TLA+/SetTheory.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) @@ -151,13 +151,11 @@ syntax "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) syntax (ASCII) (* Again, only single variable for bounded choice. *) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) translations "CHOOSE x \ S : P" \ "CONST bChoose(S, \x. P)" (* the following cannot be a print macro because the RHS is non-linear From ac7cd269cddc4a3dbf293bffba7d948e2e468fc7 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 22 Jan 2021 12:15:32 +0200 Subject: [PATCH 046/167] API: update to keywords of Isabelle 2021 The script `isabelle_keywords_get.sh` is outdated, because the file it gets the keywords from is not part of recent Isabelle distributions. Instead, a Python script `isabelle_keywords_update.py` has been added, which relies on the Scala script `outer_keywords.scala`, which gets the Isabelle 2021 keywords for the logic `Pure`. --- src/isabelle_keywords.ml | 546 +++++++++++++------------------- src/isabelle_keywords_get.sh | 21 -- src/isabelle_keywords_update.py | 22 ++ 3 files changed, 237 insertions(+), 352 deletions(-) delete mode 100755 src/isabelle_keywords_get.sh create mode 100644 src/isabelle_keywords_update.py diff --git a/src/isabelle_keywords.ml b/src/isabelle_keywords.ml index 908f5c0b..4a3f13fc 100644 --- a/src/isabelle_keywords.ml +++ b/src/isabelle_keywords.ml @@ -1,334 +1,218 @@ -(* AUTOMATICALLY GENERATED by isabelle_keywords_get.sh - DO NOT EDIT *) +(* AUTOMATICALLY GENERATED by `update_isabelle_keywords.py` - DO NOT EDIT *) Revision.f "$Rev$";; let v = [ - "."; - ".."; - "Isabelle.command"; - "ML"; - "ML_command"; - "ML_prf"; - "ML_val"; - "ProofGeneral.inform_file_processed"; - "ProofGeneral.inform_file_retracted"; - "ProofGeneral.kill_proof"; - "ProofGeneral.pr"; - "ProofGeneral.process_pgip"; - "ProofGeneral.restart"; - "ProofGeneral.undo"; - "abbreviation"; - "advanced"; - "also"; - "and"; - "apply"; - "apply_end"; - "arities"; - "assume"; - "assumes"; - "atom_decl"; - "attach"; - "attribute_setup"; - "avoids"; - "ax_specification"; - "axiomatization"; - "axioms"; - "back"; - "begin"; - "binder"; - "boogie_end"; - "boogie_open"; - "boogie_status"; - "boogie_vc"; - "by"; - "cannot_undo"; - "case"; - "case_eqns"; - "cd"; - "chapter"; - "checking"; - "class"; - "class_deps"; - "classes"; - "classrel"; - "codatatype"; - "code_abort"; - "code_class"; - "code_const"; - "code_datatype"; - "code_deps"; - "code_include"; - "code_instance"; - "code_library"; - "code_module"; - "code_modulename"; - "code_monad"; - "code_pred"; - "code_reflect"; - "code_reserved"; - "code_thms"; - "code_type"; - "coinductive"; - "coinductive_set"; - "commit"; - "con_defs"; - "congs"; - "constrains"; - "consts"; - "consts_code"; - "contains"; - "context"; - "corollary"; - "cpodef"; - "datatype"; - "datatypes"; - "declaration"; - "declare"; - "def"; - "default_sort"; - "defer"; - "defer_recdef"; - "defines"; - "definition"; - "defs"; - "disable_pr"; - "display_drafts"; - "domain"; - "domain_isomorphism"; - "domaindef"; - "domains"; - "done"; - "elimination"; - "enable_pr"; - "end"; - "enriched_type"; - "equivariance"; - "exit"; - "export_code"; - "extract"; - "extract_type"; - "file"; - "finalconsts"; - "finally"; - "find_consts"; - "find_theorems"; - "fix"; - "fixes"; - "fixrec"; - "for"; - "from"; - "full_prf"; - "fun"; - "function"; - "functions"; - "guess"; - "have"; - "header"; - "help"; - "hence"; - "hide_class"; - "hide_const"; - "hide_fact"; - "hide_type"; - "hints"; - "identifier"; - "if"; - "imports"; - "in"; - "induction"; - "inductive"; - "inductive_cases"; - "inductive_set"; - "inductive_simps"; - "infix"; - "infixl"; - "infixr"; - "init_toplevel"; - "instance"; - "instantiation"; - "interpret"; - "interpretation"; - "intros"; - "is"; - "judgment"; - "kill"; - "kill_thy"; - "lazy"; - "lemma"; - "lemmas"; - "let"; - "linear_undo"; - "local_setup"; - "locale"; - "method_setup"; - "module_name"; - "monos"; - "moreover"; - "morphisms"; - "next"; - "nitpick"; - "nitpick_params"; - "no_notation"; - "no_syntax"; - "no_translations"; - "no_type_notation"; - "nominal_datatype"; - "nominal_inductive"; - "nominal_inductive2"; - "nominal_primrec"; - "nonterminal"; - "notation"; - "note"; - "notepad"; - "notes"; - "obtain"; - "obtains"; - "oops"; - "open"; - "oracle"; - "output"; - "overloaded"; - "overloading"; - "parse_ast_translation"; - "parse_translation"; - "partial_function"; - "pcpodef"; - "permissive"; - "pervasive"; - "pr"; - "prefer"; - "presume"; - "pretty_setmargin"; - "prf"; - "primrec"; - "print_abbrevs"; - "print_antiquotations"; - "print_ast_translation"; - "print_attributes"; - "print_binds"; - "print_cases"; - "print_claset"; - "print_classes"; - "print_codeproc"; - "print_codesetup"; - "print_coercion_maps"; - "print_coercions"; - "print_commands"; - "print_configs"; - "print_context"; - "print_dependencies"; - "print_drafts"; - "print_facts"; - "print_induct_rules"; - "print_interps"; - "print_locale"; - "print_locales"; - "print_methods"; - "print_orders"; - "print_quotconsts"; - "print_quotients"; - "print_quotmaps"; - "print_rules"; - "print_simpset"; - "print_statement"; - "print_syntax"; - "print_tcset"; - "print_theorems"; - "print_theory"; - "print_trans_rules"; - "print_translation"; - "proof"; - "prop"; - "pwd"; - "qed"; - "quickcheck"; - "quickcheck_params"; - "quit"; - "quotient_definition"; - "quotient_type"; - "realizability"; - "realizers"; - "recdef"; - "recdef_tc"; - "record"; - "recursor_eqns"; - "refute"; - "refute_params"; - "remove_thy"; - "rep_datatype"; - "schematic_corollary"; - "schematic_lemma"; - "schematic_theorem"; - "sect"; - "section"; - "setup"; - "show"; - "shows"; - "simproc_setup"; - "sledgehammer"; - "sledgehammer_params"; - "smt_status"; - "solve_direct"; - "sorry"; - "spark_end"; - "spark_open"; - "spark_proof_functions"; - "spark_status"; - "spark_types"; - "spark_vc"; - "specification"; - "statespace"; - "structure"; - "subclass"; - "sublocale"; - "subsect"; - "subsection"; - "subsubsect"; - "subsubsection"; - "syntax"; - "syntax_declaration"; - "term"; - "termination"; - "text"; - "text_raw"; - "then"; - "theorem"; - "theorems"; - "theory"; - "thm"; - "thm_deps"; - "thus"; - "thy_deps"; - "translations"; - "try"; - "try_methods"; - "txt"; - "txt_raw"; - "typ"; - "type_elims"; - "type_intros"; - "type_notation"; - "type_synonym"; - "typed_print_translation"; - "typedecl"; - "typedef"; - "types"; - "types_code"; - "ultimately"; - "unchecked"; - "undo"; - "undos_proof"; - "unfolding"; - "unsafe"; - "unused_thms"; - "use"; - "use_thy"; - "uses"; - "using"; - "value"; - "values"; - "welcome"; - "where"; - "with"; - "write"; - "{"; - "}"; + "ML"; + "ML_command"; + "ML_export"; + "ML_file"; + "ML_file_debug"; + "ML_file_no_debug"; + "ML_prf"; + "ML_val"; + "ROOTS_file"; + "SML_export"; + "SML_file"; + "SML_file_debug"; + "SML_file_no_debug"; + "SML_import"; + "abbreviation"; + "abbrevs"; + "alias"; + "also"; + "and"; + "apply"; + "apply_end"; + "assume"; + "assumes"; + "attribute_setup"; + "axiomatization"; + "back"; + "begin"; + "bibtex_file"; + "binder"; + "bundle"; + "by"; + "case"; + "chapter"; + "class"; + "class_deps"; + "code_datatype"; + "compile_generated_files"; + "consider"; + "constrains"; + "consts"; + "context"; + "corollary"; + "declaration"; + "declare"; + "default_sort"; + "defer"; + "define"; + "defines"; + "definition"; + "done"; + "end"; + "experiment"; + "export_files"; + "export_generated_files"; + "export_prefix"; + "external_file"; + "external_files"; + "extract"; + "extract_type"; + "finally"; + "find_consts"; + "find_theorems"; + "fix"; + "fixes"; + "for"; + "from"; + "full_prf"; + "generate_file"; + "global_interpretation"; + "guess"; + "have"; + "help"; + "hence"; + "hide_class"; + "hide_const"; + "hide_fact"; + "hide_type"; + "if"; + "imports"; + "in"; + "include"; + "includes"; + "including"; + "infix"; + "infixl"; + "infixr"; + "instance"; + "instantiation"; + "interpret"; + "interpretation"; + "is"; + "judgment"; + "keywords"; + "lemma"; + "lemmas"; + "let"; + "local_setup"; + "locale"; + "locale_deps"; + "method_setup"; + "moreover"; + "named_theorems"; + "next"; + "no_notation"; + "no_syntax"; + "no_translations"; + "no_type_notation"; + "nonterminal"; + "notation"; + "note"; + "notepad"; + "notes"; + "obtain"; + "obtains"; + "oops"; + "open"; + "opening"; + "oracle"; + "output"; + "overloaded"; + "overloading"; + "paragraph"; + "parse_ast_translation"; + "parse_translation"; + "pervasive"; + "prefer"; + "premises"; + "presume"; + "prf"; + "print_ML_antiquotations"; + "print_abbrevs"; + "print_antiquotations"; + "print_ast_translation"; + "print_attributes"; + "print_bundles"; + "print_cases"; + "print_classes"; + "print_codesetup"; + "print_commands"; + "print_context"; + "print_definitions"; + "print_defn_rules"; + "print_facts"; + "print_interps"; + "print_locale"; + "print_locales"; + "print_methods"; + "print_options"; + "print_rules"; + "print_simpset"; + "print_state"; + "print_statement"; + "print_syntax"; + "print_term_bindings"; + "print_theorems"; + "print_theory"; + "print_trans_rules"; + "print_translation"; + "private"; + "proof"; + "prop"; + "proposition"; + "qed"; + "qualified"; + "realizability"; + "realizers"; + "rewrites"; + "schematic_goal"; + "section"; + "setup"; + "show"; + "shows"; + "simproc_setup"; + "sorry"; + "structure"; + "subclass"; + "subgoal"; + "sublocale"; + "subparagraph"; + "subsection"; + "subsubsection"; + "supply"; + "syntax"; + "syntax_declaration"; + "term"; + "text"; + "text_raw"; + "then"; + "theorem"; + "theory"; + "thm"; + "thm_deps"; + "thm_oracles"; + "thus"; + "thy_deps"; + "translations"; + "txt"; + "typ"; + "type_alias"; + "type_notation"; + "type_synonym"; + "typed_print_translation"; + "typedecl"; + "ultimately"; + "unbundle"; + "unchecked"; + "unfolding"; + "unused_thms"; + "using"; + "welcome"; + "when"; + "where"; + "with"; + "write"; ];; diff --git a/src/isabelle_keywords_get.sh b/src/isabelle_keywords_get.sh deleted file mode 100755 index a2779387..00000000 --- a/src/isabelle_keywords_get.sh +++ /dev/null @@ -1,21 +0,0 @@ -#!/bin/sh - -# Copyright (C) 2012 INRIA and Microsoft Corporation - -# This script generates the "isabelle_keywords.ml" file. -# It assumes a working installation of Isabelle is in the PATH. - -eval `isabelle getenv ISABELLE_HOME` - -{ - echo '(* AUTOMATICALLY GENERATED by isabelle_keywords_get.sh - DO NOT EDIT *)' - echo 'Revision.f "$Rev: 29196 $";;' - echo 'let v = [' - - sed -n -e '/^.*"\([^"]*\)".*$/s//\1/p' $ISABELLE_HOME/etc/isar-keywords*.el \ - | sed -e 's/\\\\\././g' \ - | sort -u \ - | sed -e 's/.*/ "&";/' - - echo '];;' -} > isabelle_keywords.ml diff --git a/src/isabelle_keywords_update.py b/src/isabelle_keywords_update.py new file mode 100644 index 00000000..ee7ad720 --- /dev/null +++ b/src/isabelle_keywords_update.py @@ -0,0 +1,22 @@ +"""Generate OCaml list of Isabelle keywords. + +This script reads the output of the invocation: +./Isabelle2021-RC2.app/Isabelle/bin/isabelle outer_keywords Pure +which is contained in the file with name equal to `infile`. +""" +infile = 'isabelle_keywords.txt' + + +def main(): + with open(infile, 'r') as f: + text = f.read() + for line in text.split('\n'): + s = '"{line}";'.format(line=line) + if not s: + continue + print(s) + + +if __name__ == '__main__': + main() + From a3b7e9b32ef57b6ade725e3c0d39df9f68c5ad01 Mon Sep 17 00:00:00 2001 From: Isabelle Group <> Date: Sat, 23 Jan 2021 12:40:46 +0200 Subject: [PATCH 047/167] ENH: add script for Isabelle 2021 logic keywords Kindly provided by Makarius Wenzel. --- src/outer_keywords.scala | 51 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 51 insertions(+) create mode 100644 src/outer_keywords.scala diff --git a/src/outer_keywords.scala b/src/outer_keywords.scala new file mode 100644 index 00000000..b51073ca --- /dev/null +++ b/src/outer_keywords.scala @@ -0,0 +1,51 @@ +// DESCRIPTION: print outer syntax keywords of specified theory + +object Tool extends isabelle.Isabelle_Tool.Body { + import isabelle._ + + def outer_keywords( + options: Options, + theory_name: String, + dirs: List[Path] = Nil, + progress: Progress = new Progress) + { + val sessions_structure = Sessions.load_structure(options, dirs = dirs) + val session_name = sessions_structure.bootstrap.theory_qualifier(theory_name) + val deps = Sessions.deps(sessions_structure.selection(session_name)) + + deps(session_name).loaded_theory_syntax(theory_name) match { + case None => error("Bad theory " + quote(theory_name)) + case Some(syntax) => + val keywords = syntax.keywords.kinds.keySet.toList.sorted + keywords.foreach(progress.echo) + } + } + + def apply(args: List[String]) + { + var dirs: List[Path] = Nil + + val getopts = Getopts(""" +Usage: isabelle outer_keywords [OPTIONS] THEORY + + Options are: + -d DIR include session directory + + Print outer syntax keywords of specific theory (long name), e.g. + "Pure", "Main", "HOL-Library.Multiset". +""", + "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg)))) + + val more_args = getopts(args) + val theory_name = + more_args match { + case List(name) => name + case _ => getopts.usage() + } + + val options = Options.init() + val progress = new Console_Progress() + + outer_keywords(options, theory_name, dirs = dirs, progress = progress) + } +} From e1f3bdf653cf4fd5ac5eb5ad30687c183341ed10 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 12:53:02 +0200 Subject: [PATCH 048/167] DOC: how to install and use `outer_keywords.scala` --- src/isabelle_keywords_update.py | 6 +++--- src/outer_keywords.scala | 13 +++++++++++++ 2 files changed, 16 insertions(+), 3 deletions(-) diff --git a/src/isabelle_keywords_update.py b/src/isabelle_keywords_update.py index ee7ad720..6315e8ec 100644 --- a/src/isabelle_keywords_update.py +++ b/src/isabelle_keywords_update.py @@ -1,8 +1,9 @@ """Generate OCaml list of Isabelle keywords. This script reads the output of the invocation: -./Isabelle2021-RC2.app/Isabelle/bin/isabelle outer_keywords Pure -which is contained in the file with name equal to `infile`. + isabelle outer_keywords Pure > isabelle_keywords.txt +The script `outer_keywords.scala` is present in this directory, +and contains its installation and usage documentation. """ infile = 'isabelle_keywords.txt' @@ -19,4 +20,3 @@ def main(): if __name__ == '__main__': main() - diff --git a/src/outer_keywords.scala b/src/outer_keywords.scala index b51073ca..72c61e3a 100644 --- a/src/outer_keywords.scala +++ b/src/outer_keywords.scala @@ -1,4 +1,17 @@ // DESCRIPTION: print outer syntax keywords of specified theory +// +// INSTALLATION: +// Place this Scala script in the directory $ISABELLE_HOME_USER/Tools +// and to the file $ISABELLE_HOME_USER/etc/settings the line +// ISABELLE_TOOLS="$ISABELLE_TOOLS:$ISABELLE_HOME_USER/Tools" +// To find the path ISABELLE_HOME_USER the tool `isabelle getenv` can be used +// by invoking: +// isabelle getenv ISABELLE_HOME_USER +// For example, when this script was first used for `tlapm`, this path was +// $HOME/.isabelle/Isabelle2021-RC2 +// +// Invoke the script with: +// isabelle outer_keywords Pure object Tool extends isabelle.Isabelle_Tool.Body { import isabelle._ From f66dd307678ebdfb291f58dd6872e6466d032a1c Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:04:55 +0200 Subject: [PATCH 049/167] MAI: use theories `Constant` and `Zenon` in `ROOT` The attribute `global` is included to enable directly importing these theories within user theories. --- isabelle/ROOT | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/isabelle/ROOT b/isabelle/ROOT index 53c1d9d9..b84420fd 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -2,6 +2,7 @@ session "TLA+" = "Pure" + options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] theories - Peano + Constant (global) + Zenon (global) document_files "root.tex" From 8fe2b12ccfcaf8f990d52a55332876fa17beddf8 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:05:32 +0200 Subject: [PATCH 050/167] DRAFT: comment document generation for Isabelle theories because it currently raises errors. --- isabelle/ROOT | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/ROOT b/isabelle/ROOT index b84420fd..84a6f8b6 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,8 +1,12 @@ (* build the session using "isabelle build -D ." *) session "TLA+" = "Pure" + + (* options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] + *) theories Constant (global) Zenon (global) + (* document_files "root.tex" + *) From 33c4da70f30a3a47c8d15c5bcde90ca838ef7b06 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:06:29 +0200 Subject: [PATCH 051/167] MAI: add theorem `leq_adiff_right_imp_0` to theory `Integers` because this theorem is used within the theory `Integers`. --- isabelle/Integers.thy | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index c4398167..776ddd84 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -1044,6 +1044,12 @@ apply(rule nat_leq_cases[of m "n + p" ], simp_all) apply(rule adiff_adiff_left_nat[symmetric], simp+) done +lemma leq_adiff_right_imp_0: + assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" + shows "p = 0" +using p h apply (induct) +using n by auto + lemma int_add_assoc5: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "-.m + (n + -.p) = -.m + n + -.p" From bcd431fc7d6bb0446cc810a69c2a0bc9f527edc3 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:07:01 +0200 Subject: [PATCH 052/167] REL: update `isabelle/Makefile` for Isabelle2020 --- isabelle/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Makefile b/isabelle/Makefile index 93e3091d..cd752015 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -10,4 +10,4 @@ rebuild: isabelle usedir -b -i true -d pdf Pure TLA+ heap-only: - isabelle usedir -b Pure TLA+ + isabelle build -v -d . -b TLA+ From 27f6ae6f794204f887add6696ab5f64c802d6a54 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:08:10 +0200 Subject: [PATCH 053/167] MAI: update theory `Zenon` to use `\\A i \\in` syntax following changes to bounded quantification syntax in the theory `SetTheory`. --- isabelle/Zenon.thy | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 49cbfd67..c00ac8c5 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1047,7 +1047,7 @@ proof (rule inProductE [OF a b]) assume h1: "isASeq(p)" assume h2: "Len(p) = Len(s)" assume h3: "p \\in [1 .. Len (s) -> UNION Range (s)]" - assume h4: "ALL k in 1 .. Len (s) : p[k] \\in s[k]" + assume h4: "\\A k \\in 1 .. Len (s) : p[k] \\in s[k]" show "p[i] \\in s[i]" (is "?c") proof (rule bAllE [where x = i, OF h4]) assume h5: "i \\notin 1 .. Len(s)" @@ -1137,7 +1137,7 @@ lemma zenon_in_recordset_field : shows "r[doms[i]] \\in rngs[i]" proof (rule EnumFuncSetE [OF a]) assume h1: "r \\in [Range(doms) -> UNION Range(rngs)]" - assume h2: "ALL i in DOMAIN doms : r[doms[i]] \\in rngs[i]" + assume h2: "\\A i \\in DOMAIN doms : r[doms[i]] \\in rngs[i]" show "r[doms[i]] \\in rngs[i]" (is "?c") proof (rule bAllE [where x = i, OF h2]) assume h3: "i \\notin DOMAIN doms" @@ -1198,12 +1198,12 @@ lemma zenon_tuple_dom_3 : "isASeq (s) & isASeq (t) & DOMAIN s = DOMAIN t ==> DOMAIN s = DOMAIN t" by auto -lemma zenon_all_rec_1 : "ALL i in {} : r[doms[i]] \\in rngs[i]" by auto +lemma zenon_all_rec_1 : "\\A i \\in {} : r[doms[i]] \\in rngs[i]" by auto lemma zenon_all_rec_2 : - "ALL i in s : r[doms[i]] \\in rngs[i] ==> + "\\A i \\in s : r[doms[i]] \\in rngs[i] ==> r[doms[j]] \\in rngs[j] ==> - ALL i in addElt (j, s) : r[doms[i]] \\in rngs[i]" + \\A i \\in addElt (j, s) : r[doms[i]] \\in rngs[i]" by auto lemma zenon_tuple_acc_1 : @@ -1296,7 +1296,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1360,7 +1360,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1443,7 +1443,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1547,7 +1547,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1670,7 +1670,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1812,7 +1812,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1973,7 +1973,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" From 2bf74426a01b0e82d607fa3c748f7c414d49baec Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:09:13 +0200 Subject: [PATCH 054/167] DRAFT: comment auto-generated proof rules for `CASE` expressions in the theory `Zenon`, because currently these raise errors, due to an update that is needed in the theory `CaseExpressions`. --- isabelle/Zenon.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index c00ac8c5..4c5d5c2d 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -2100,7 +2100,7 @@ proof - qed (*** END AUTO_GENERATED STUFF ***) - +(* (* ------------- Proof rules for CASE expressions ------ *) lemma zenon_caseother0 : @@ -3694,7 +3694,7 @@ proof - by (rule zenon_case_empty_union [OF hi]) qed qed - +*) (*** END AUTO_GENERATED STUFF ***) end From b21f4dfd93de8e5fa943c3a1d3caf493bdc2eb49 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:15:02 +0200 Subject: [PATCH 055/167] MAI: update `tlapm` to call `isabelle process` of Isabelle2020 As part of this change, the Isabelle/TLA+ theories are installed under the directory `${prefix}/lib/isabelle_tla/`. where `${prefix}` the value of the corresponding variable in the file `Makefile`. --- src/params.ml | 16 ++++++++++++---- tools/Makefile.in | 6 ++++++ 2 files changed, 18 insertions(+), 4 deletions(-) diff --git a/src/params.ml b/src/params.ml index f98398f8..2fb31a3b 100644 --- a/src/params.ml +++ b/src/params.ml @@ -62,6 +62,13 @@ let library_path = let d = Filename.concat d "tlaps" in d +let isabelle_tla_path = + let d = Sys.executable_name in + let d = Filename.dirname (Filename.dirname d) in + let d = Filename.concat d "lib" in + let d = Filename.concat d "isabelle_tla" in + d + type executable = | Unchecked of string * string * string (* exec, command, version_command *) | User of string (* command *) @@ -86,7 +93,7 @@ let get_exec e = let r = 10000 + (Random.int 10000) in let stat = sprintf "echo path prefix: \"%s\" ; %s stat `which %s` >/tmp/stat-%d-%d.txt" path_prefix path_prefix exec tod r in - + ignore (Sys.command stat); let check = sprintf "%s type %s >/tmp/check-%d-%d.txt" path_prefix exec tod r in *) @@ -130,11 +137,12 @@ let isabelle_success_string = "((TLAPS SUCCESS))" let isabelle = let cmd = - Printf.sprintf "isabelle-process -r -q -e \"(use_thy \\\"$file\\\"; \ - writeln \\\"%s\\\");\" TLA+" + Printf.sprintf "isabelle process -e \"(use_thy \\\"$file\\\"; \ + writeln \\\"%s\\\");\" -d %s -l TLA+" isabelle_success_string + isabelle_tla_path in - make_exec "isabelle-process" cmd "isabelle version" + make_exec "isabelle process" cmd "isabelle version" ;; let set_fast_isabelle () = diff --git a/tools/Makefile.in b/tools/Makefile.in index 613a1438..d4ccacf5 100644 --- a/tools/Makefile.in +++ b/tools/Makefile.in @@ -10,6 +10,7 @@ VPATH = @srcdir@ bindir = @bindir@ LIBPATH = @libdir@/tlaps +ISATLAPATH = @libdir@/isabelle_tla all: check-configure tlapm@EXE@ @@ -101,6 +102,8 @@ install: mkdir -p -m 0755 ${LIBPATH} ${INSTALL} -m 0644 library/*.tla ${LIBPATH} @TAR@ -cf - examples | ( cd ${LIBPATH} && @TAR@ -xf - ) + mkdir -p -m 0755 ${ISATLAPATH} + ${INSTALL} -m 0644 isabelle/*.thy isabelle/*.ML isabelle/ROOT ${ISATLAPATH} .PHONY: uninstall uninstall: @@ -108,6 +111,9 @@ uninstall: if test -n "`ls -A ${LIBPATH}`"; \ then rm -rf ${LIBPATH} ; \ fi + if test -n "`ls -A ${ISATLAPATH}`"; \ + then rm -rf ${ISATLAPATH} ; \ + fi .PHONY: newversion newversion: From f95915fdd7be75fc52227ebecadfa56aba36e29b Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:17:06 +0200 Subject: [PATCH 056/167] MAI: include `using assms` only if there exist assumptions in the generated Isabelle theories. This change is needed because Isabelle2020 raises an error if `using assms` appears without any `assumes` keywords preceding it. --- src/backend/isabelle.ml | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/src/backend/isabelle.ml b/src/backend/isabelle.ml index d6ae628c..e83d3676 100644 --- a/src/backend/isabelle.ml +++ b/src/backend/isabelle.ml @@ -29,6 +29,7 @@ let bump : ctx -> ctx = Ctx.bump let length : ctx -> int = Ctx.length +let emitted_assumptions = ref false let cook x = "is" ^ pickle x @@ -484,6 +485,7 @@ and pp_print_sequent_outer cx ff sq = match Deque.front sq.context with pp_print_sequent_outer ncx ff { sq with context = hs } | Fresh (nm, _, _, Bounded (b, Visible)) -> let (ncx, nm) = adj cx nm in + emitted_assumptions := true; fprintf ff "fixes %s@,assumes %s_in : @[\"(%s \\ %a)\"@]@," nm nm nm (pp_print_expr false cx) b ; pp_print_sequent_outer ncx ff { sq with context = hs } @@ -496,6 +498,7 @@ and pp_print_sequent_outer cx ff sq = match Deque.front sq.context with begin try let null = make_formatter (fun _ _ _ -> ()) (fun () -> ()) in pp_print_expr false cx null e; (* may raise Unsupported *) + emitted_assumptions := true; fprintf ff "assumes v'%d: \"%a\"@," (Ctx.length cx) (pp_print_expr false cx) e with Unsupported msg -> @@ -644,6 +647,7 @@ let thy_temp ob tac tempname thyout = Printf.fprintf thyout "lemma ob'%d: (* %s *)\n" obid obfp; let ff = Format.formatter_of_out_channel thyout in begin try + emitted_assumptions := false; pp_print_expr true Ctx.dot ff (Sequent ob.obl.core @@ nowhere); with Unsupported msg -> failwith ("isabelle : unsupported operator " ^ msg) @@ -652,7 +656,10 @@ let thy_temp ob tac tempname thyout = Printf.fprintf thyout "(is \"PROP ?ob'%d\")\n" obid; Printf.fprintf thyout "proof -\n"; Printf.fprintf thyout "show \"PROP ?ob'%d\"\n" obid; - Printf.fprintf thyout "using assms by %s\n" tac; + if !emitted_assumptions then + Printf.fprintf thyout "using assms by %s\n" tac + else + Printf.fprintf thyout "by %s\n" tac; Printf.fprintf thyout "qed\n"; Printf.fprintf thyout "end\n"; ;; From 9ef3683ffa8d293b728934016ee6b01ea15d30dd Mon Sep 17 00:00:00 2001 From: merz Date: Sat, 23 Jan 2021 14:37:28 +0100 Subject: [PATCH 057/167] fixed syntax translations for Isa2020 --- isabelle/CaseExpressions.thy | 45 ++- isabelle/Tuples.thy | 525 ++++++++--------------------------- 2 files changed, 136 insertions(+), 434 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 09c2ca77..6507100a 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -1,9 +1,8 @@ (* Title: TLA+/CaseExpressions.thy Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:22 merz> + Version: Isabelle2020 *) section \ Case expressions \ @@ -34,7 +33,7 @@ text \ twice (and will be simplified twice) in a @{text CaseOther} expression. \ -definition CaseArm (* -- preliminary construct to convert case arm into set *) +definition CaseArm \ \ preliminary construct to convert case arm into set \ where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" definition Case where @@ -49,22 +48,16 @@ nonterminal case_arm and case_arms syntax "_case_syntax":: "case_arms \ c" ("(CASE _ )" 10) - "_case1" :: "[c, c] \ case_arm" ("(2_ ->/ _)" 10) - "" :: "case_arm \ case_arms" ("_") - "_other" :: "c \ case_arms" ("OTHER -> _") - "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ [] _") - -syntax (xsymbols) "_case1" :: "[c, c] \ case_arm" ("(2_ \/ _)" 10) + "" :: "case_arm \ case_arms" ("_") "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") -syntax (HTML output) - "_case1" :: "[c, c] \ case_arm" ("(2_ \/ _)" 10) - "_other" :: "c \ case_arms" ("OTHER \ _") - "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") +syntax (ASCII) + "_case1" :: "[c, c] \ case_arm" ("(2_ ->/ _)" 10) + "_other" :: "c \ case_arms" ("OTHER -> _") + "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ [] _") -(* TODO parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -98,15 +91,15 @@ parse_ast_translation \ end | case_syntax_tr _ = raise Match; in - [("_case_syntax", case_syntax_tr)] + [("_case_syntax", K case_syntax_tr)] end \ -*) + (** debugging ** (*ML \ set Syntax.trace_ast; \*) -lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" +lemma "CASE a \ b \ c \ d \ e \ f \ OTHER \ g" oops lemma "CASE a \ b" @@ -118,7 +111,6 @@ oops (*ML \ reset Syntax.trace_ast; \*) **) -(* TODO print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] @@ -173,16 +165,15 @@ print_ast_translation \ end | caseother_tr' _ = raise Match in - [(@{const_syntax "Case"}, case_syntax_tr'), - (@{const_syntax "CaseOther"}, caseother_tr')] + [(@{const_syntax "Case"}, K case_syntax_tr'), + (@{const_syntax "CaseOther"}, K caseother_tr')] end \ -*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) -lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" +lemma "CASE a \ b \ c \ d \ e \ f \ OTHER \ g" oops lemma "CASE a \ b" @@ -199,15 +190,15 @@ lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def (** test cases ** lemma - "i=1 \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1) = 0" -by simp + "i=1 \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = 0" +by (simp add: two_def) lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1) = default" + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" by auto lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1 [] OTHER \ a) = a" + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" by auto **) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index d644e8b5..5ecc293f 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Tuples.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:46 merz> + Version: Isabelle2020 *) section \ Tuples and Relations in \tlaplus{} \ @@ -31,13 +30,13 @@ text \ on sequences require arithmetic and will be introduced in a separate theory. \ -definition Seq (* -- \ set of finite sequences with elements from $S$ \ *) +definition Seq \ \ set of finite sequences with elements from $S$ \ where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" -definition isASeq (* -- \ characteristic predicate for sequences or tuples \close> *) +definition isASeq \ \ characteristic predicate for sequences or tuples \ where "isASeq(s) \ isAFcn(s) \ (\n \ Nat : DOMAIN s = 1 .. n)" -definition Len (* -- \ length of a sequence \ *) +definition Len \ \ length of a sequence \ where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: @@ -80,15 +79,15 @@ lemma SeqIsAFcn (*[elim!]*): shows "isAFcn(s)" using assms by auto -(* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) +\ \ @{text "s \ Seq(S) \ isAFcn(s)"} \ lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" shows "Len(s) \ Nat" -using assms by auto +using assms by (rule isASeqE) -(* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) +\ \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] lemma DomainSeqLen [simp]: @@ -96,7 +95,7 @@ lemma DomainSeqLen [simp]: shows "DOMAIN s = 1 .. Len(s)" using assms by auto -(* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) +\ \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: @@ -121,7 +120,7 @@ lemma SeqI [intro!]: shows "s \ Seq(S)" using assms by (auto simp: Seq_def) -lemma SeqI': (* -- \ closer to the definition but probably less useful \ *) +lemma SeqI': \ \ closer to the definition but probably less useful \ assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" using assms by (auto simp: Seq_def) @@ -141,7 +140,7 @@ qed lemma seqFuncSet: assumes "s \ Seq(S)" shows "s \ [1 .. Len(s) \ S]" -using assms by auto +using assms by (rule SeqE) lemma seqElt [elim!]: assumes "s \ Seq(S)" and "n \ Nat" and "1 \ n" and "n \ Len(s)" @@ -164,14 +163,11 @@ text \ (written @{text "\\"}) and Append. \ -definition emptySeq ("(<< >>)") -where "<< >> \ [x \ 1 .. 0 \ {}]" +definition emptySeq ("(\\)") +where "\\ \ [x \ 1 .. 0 \ {}]" -notation (xsymbols) - emptySeq ("(\\)") - -notation (HTML output) - emptySeq ("(\\)") +notation (ASCII) + emptySeq ("(<< >>)") definition Append :: "[c,c] \ c" where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len(s)] THEN e ELSE s[k]]" @@ -179,7 +175,7 @@ where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" by (auto simp: emptySeq_def isASeq_def) -(* -- \ @{text "isAFcn(\\)"} \ *) +\ \ @{text "isAFcn(\\)"} \ lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] lemma lenEmptySeq [simp]: "Len(\\) = 0" @@ -199,7 +195,7 @@ lemma appendIsASeq [simp,intro!]: using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) -(* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) +\ \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" @@ -213,11 +209,11 @@ lemma isEmptySeq [intro!]: "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" by auto -(* -- \ immediate consequence of @{text isEmptySeq} \ *) +\ \ immediate consequence of @{text isEmptySeq} \ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto -(* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) +\ \ The reverse equation could be a useful rewrite rule (it is applied by TLC) \ lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" @@ -243,7 +239,7 @@ lemma lenAppend [simp]: shows "Len(Append(s,e)) = Succ[Len(s)]" using assms by (intro LenI, auto simp: Append_def) -(* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) +\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] lemma appendElt [simp]: @@ -336,28 +332,7 @@ text \ lemma appendExtend: assumes "isASeq(s)" shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" - proof - - have s1_1: "Append(s, e) = - [k \ 1 .. Succ[Len(s)] \ - IF k = Succ[Len(s)] THEN e ELSE s[k]]" - unfolding Append_def by auto - have s1_2: "(Succ[Len(s)] :> e) = - [x \ {Succ[Len(s)]} \ e]" - unfolding oneArg_def by auto - define p where "p \ s @@ (Succ[Len(s)] :> e)" - have s1_3: "p = - [x \ (DOMAIN s) \ {Succ[Len(s)]} - \ IF x \ DOMAIN s THEN s[x] ELSE e]" - unfolding p_def extend_def using s1_2 by auto - have s1_4: "(DOMAIN s) \ {Succ[Len(s)]} - = 1 .. Succ[Len(s)]" - using assms by auto - have s1_5: "p = [x \ 1..Succ[Len(s)] \ - IF x \ DOMAIN s THEN s[x] ELSE e]" - using s1_3 s1_4 by auto - show ?thesis - using s1_1 s1_5 assms by (auto simp: p_def) - qed +using assms by (intro isAppend') auto lemma imageEmptySeq [simp]: "Image(\\, A) = {}" by (simp add: emptySeq_def) @@ -367,7 +342,7 @@ lemma imageAppend [simp]: shows "Image(Append(s,e), A) = (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" unfolding appendExtend[OF s] -using assms by (auto elim!: inNatIntervalE, force+) +using assms by auto (force+) text \ Inductive reasoning about sequences, based on @{term "\\"} and @@ -419,7 +394,7 @@ proof - with s show ?thesis unfolding Seq_def by auto qed -(* -- \ example of an inductive proof about sequences \ *) +\ \ example of an inductive proof about sequences \ lemma seqEmptyOrAppend: assumes "s \ Seq(S)" shows "s = \\ \ (\s' \ Seq(S): \e \ S : s = Append(s',e))" @@ -444,13 +419,10 @@ nonterminal tpl syntax "" :: "c \ tpl" ("_") "@app" :: "[tpl, c] \ tpl" ("_,/ _") - "@tuple" :: "tpl \ c" ("<<(_)>>") - -syntax (xsymbols) "@tuple" :: "tpl \ c" ("\(_)\") -syntax (HTML output) - "@tuple" :: "tpl \ c" ("\(_)\") +syntax (ASCII) + "@tuple" :: "tpl \ c" ("<<(_)>>") translations "\ tp, x \" \ "CONST Append(\tp\, x)" @@ -458,13 +430,13 @@ translations (*** examples *** -lemma "Len(\a,b,c\) = 3" by simp +lemma "Len(\a,b,c\) = 3" by (simp add: three_def two_def) lemma "Append(\a,b\, c) = \a,b,c\" .. lemma "\a,b,c\[1] = a" by simp -lemma "\a,b,c\[2] = b" by simp -lemma "\a,b,c\[3] = c" by simp +lemma "\a,b,c\[2] = b" by (simp add: two_def) +lemma "\a,b,c\[3] = c" by (simp add: two_def three_def) lemma "\0,1\ \ \\" by simp lemma "\0,1\ \ \1\" by simp lemma "\0,1\ \ \1,2\" by simp @@ -480,16 +452,6 @@ text \ \ syntax - "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) - "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) - "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) - "@bEx5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) - "@bAll2" :: "[idt,idt,c,c] \ c" ("(3ALL <<_,_>> in _ :/ _)" [100,100,0,0] 10) - "@bAll3" :: "[idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) - "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) - "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) - -syntax (xsymbols) "@bEx2" :: "[idt,idt,c,c] \ c" ("(3\\_,_\ \ _ :/ _)" [100,100,0,0] 10) "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3\\_,_,_\ \ _ :/ _)" [100,100,100,0,0] 10) "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_\ \ _ :/ _)" [100,100,100,100,0,0] 10) @@ -499,6 +461,16 @@ syntax (xsymbols) "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_\ \ _ :/ _)" [100,100,100,100,0,0] 10) "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_,_\ \ _ :/ _)" [100,100,100,100,100,0,0] 10) +syntax (ASCII) + "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) + "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) + "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) + "@bEx5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) + "@bAll2" :: "[idt,idt,c,c] \ c" ("(3ALL <<_,_>> in _ :/ _)" [100,100,0,0] 10) + "@bAll3" :: "[idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) + "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) + "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) + translations "\\x,y\ \ S : P" \ "\x,y : \x,y\ \ S \ P" "\\x,y,z\ \ S : P" \ "\x,y,z : \x,y,z\ \ S \ P" @@ -529,9 +501,9 @@ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" -lemmas (* -- \ establish set equality for sets of enumerated functions \ *) - setEqualI [where A = "EnumFuncSet(doms, rngs)", intro!] - setEqualI [where B = "EnumFuncSet(doms, rngs)", intro!] +lemmas \ \ establish set equality for sets of enumerated functions \ + setEqualI [where A = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] + setEqualI [where B = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] lemma EnumFuncSetI [intro!,simp]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" @@ -562,7 +534,6 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") -(* parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -591,7 +562,7 @@ parse_ast_translation \ end | enum_funcset_tr _ = raise Match; in - [("@EnumFuncSet", enum_funcset_tr)] + [("@EnumFuncSet", K enum_funcset_tr)] end \ @@ -622,35 +593,35 @@ print_ast_translation \ end | enum_funcset_tr' _ = raise Match in - [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] + [(@{const_syntax "EnumFuncSet"}, K enum_funcset_tr')] end \ -*) + (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" by auto lemma "(1 :> TRUE) @@ (5 :> 2) \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE @@ 5 :> 2) = (5 :> 2 @@ 1 :> TRUE)" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (2 :> {}) \ [1 : BOOLEAN, 2: SUBSET Nat, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![1] = FALSE] \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![5] = {}] \ [1 : BOOLEAN, 5 : SUBSET Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (1 :> {}) \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[1 : BOOLEAN, 5 : Nat] = [5 : Nat, 1 : BOOLEAN]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "\f \ [1 : BOOLEAN, 2: Nat]; g \ [2: Nat, 1: BOOLEAN]; f[1] = g[1]; f[2] = g[2]\ \ f = g" by auto @@ -732,12 +703,10 @@ lemma "\a\ \ Product(\A,B\)" by auto text \ Special case: binary product \ definition - prod :: "c \ c \ c" (infixr "\\X" 100) where - "A \\X B \ Product(\A,B\)" -notation (xsymbols) - prod (infixr "\" 100) -notation (HTML output) - prod (infixr "\" 100) + prod :: "c \ c \ c" (infixr "\" 100) where + "A \ B \ Product(\A,B\)" +notation (ASCII) + prod (infixr "\\X" 100) lemma inProd [simp]: "(\a,b\ \ A \ B) = (a \ A \ b \ B)" @@ -746,8 +715,7 @@ by (auto simp add: prod_def) lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" -unfolding two_def -using assms by (auto simp add: prod_def) +using assms by (auto simp: two_def prod_def) lemma inProd': "(p \ A \ B) = (\a \ A : \ b \ B : p = \a,b\)" @@ -821,13 +789,12 @@ using assms by (auto dest: pairProj_in_prod) lemma relProj2 [elim]: (** consider [elim!] ?? **) assumes "\a,b\ \ r" and "r \ A \ B" shows "b \ B" -using assms by (auto dest: pairProj_in_prod) + using assms by (auto dest: pairProj_in_prod) lemma setOfAllPairs_eq_r (*[simp]*): assumes r: "r \ A \ B" shows "{\p[1], p[2]\ : p \ r} = r" -apply auto -using subsetD[OF r, THEN prodProj[of _ A B]] by simp_all +using subsetD[OF r, THEN prodProj] by force lemma subsetsInProd: assumes "A' \ A" and "B' \ B" @@ -842,9 +809,9 @@ definition setOfPairs :: "[c, [c,c]\c] \ c" where "setOfPairs(R, f) \ { f(p[1], p[2]) : p \ R }" syntax - "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : <<_,_>> \\in _})") -syntax (xsymbols) "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : \_,_\ \ _})") +syntax (ASCII) + "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : <<_,_>> \\in _})") translations "{e : \x,y\ \ R}" \ "CONST setOfPairs(R, \x y. e)" @@ -859,7 +826,7 @@ lemma inSetOfPairsI [intro]: unfolding two_def using assms by (auto simp: setOfPairs_def two_def) -lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) +lemma inSetOfPairsE [elim!]: \ \ the converse is true only if $R$ is a relation \ assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" @@ -875,188 +842,27 @@ lemmas setOfPairsEqualI = setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] lemma setOfPairs_triv [simp]: - assumes s: "R \ A \ B" + assumes "R \ A \ B" shows "{ \x, y\ : \x, y\ \ R } = R" - proof - - define Q where "Q \ - { \x, y\ : \x, y\ \ R }" - have s1_1: "Q = {<>: p \ R}" - unfolding Q_def setOfPairs_def by auto - have s1_2: "\ p. p \ R ==> p \ A \ B" - using s by auto - have s1_3: "\ p. p \ R ==> p = <>" - proof - - fix "p" :: "c" - assume s1_3_asm: "p \\in R" - have s2_1: "p \ A \ B" - using s1_3_asm s by auto - show "p = <>" - using s2_1 prodProj by auto - qed - define e :: "[c] => c" where "e \ \ x. x" - define f :: "[c] => c" where "f \ \ x. <>" - have s1_4: "\ x. x \ R ==> e(x) = f(x)" - unfolding e_def f_def using s1_3 by auto - have s1_5: "{e(x): x \ R} = {f(y): y \ R}" - apply (rule setOfAll_cong[of "R" "R" "e" "f"]) - apply (simp) - by (rule s1_4) - show ?thesis - using s1_1 s1_5 by (auto simp: Q_def e_def f_def) - qed +using assms by auto lemma setOfPairs_cong (*[cong]*): assumes 1: "R = S" and 2: "S \ A \ B" and - 3: "\x y. \x, y\ \ S - \ - e(x, y) = f(x, y)" + 3: "\x y. \x, y\ \ S \ e(x, y) = f(x, y)" shows "{ e(x, y) : \x, y\ \ R } = { f(u, v) : \u, v\ \ S }" - proof - - have s1_1: "{ e(x, y) : \x, y\ \ R } = - { e(x, y) : \x, y\ \ S }" - using 1 by auto - have s1_2: "{ e(x, y) : \x, y\ \ S } = - { e(p[1], p[2]) : p \ S }" - by (auto simp: setOfPairs_def) - have s1_3: "{ f(u, v) : \u, v\ \ S } = - { f(p[1], p[2]) : p \ S }" - by (auto simp: setOfPairs_def) - define P where "P \ { e(p[1], p[2]): p \ S }" - define Q where "Q \ { f(p[1], p[2]): p \ S }" - have s1_4: "P = Q" - proof - - have s2_1: "\ x. x \ P <=> x \ Q" - proof - - fix "x" :: "c" - have s3_1: "x \\in P ==> x \\in Q" - proof - - assume s3_1_asm: "x \\in P" - have s4_1: "x \ { e(p[1], p[2]): p \ S }" - using s3_1_asm by (auto simp: P_def) - have s4_2: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_1 setOfAll by auto - have s4_3: "\\E p: p \\in S \ x = f(p[1], p[2])" - using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) - have s4_4: "x \ { f(p[1], p[2]): p \ S }" - using s4_3 setOfAll by auto - show "x \\in Q" - unfolding Q_def using s4_4 by auto - qed - have s3_2: "x \\in Q ==> x \\in P" - proof - - assume s3_2_asm: "x \\in Q" - have s4_1: "x \ { f(p[1], p[2]): p \ S}" - using s3_2_asm by (auto simp: Q_def) - have s4_2: "\\E p: p \\in S \ x = f(p[1], p[2])" - using s4_1 setOfAll by auto - have s4_3: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) - have s4_4: "x \ { e(p[1], p[2]): p \ S }" - using s4_3 setOfAll by auto - show "x \\in P" - unfolding P_def using s4_4 by auto - qed - show "x \ P <=> x \ Q" - using s3_1 s3_2 by auto - qed - have s2_2: "\ x: x \ P <=> x \ Q" - using s2_1 allI by auto - show ?thesis - using s2_2 extension by auto - qed - show ?thesis - using s1_1 s1_2 s1_3 s1_4 - by (auto simp: P_def Q_def) - qed +using assms by (force intro!: inSetOfPairsI_ex) lemma setOfPairsEqual: - assumes 1: "\x y. - \x, y\ \ S - \ \\x', y'\ \ T : - e(x, y) = f(x', y')" - and 2: "\x' y'. - \x', y'\ \ T - \ \\x, y\ \ S : - f(x', y') = e(x, y)" - and 3: "S \ A \ B" and - 4: "T \ C \ D" + assumes 1: "\x y. \x, y\ \ S \ \\x', y'\ \ T : e(x, y) = f(x', y')" + and 2: "\x' y'. \x', y'\ \ T \ \\x, y\ \ S : f(x', y') = e(x, y)" + and 3: "S \ A \ B" + and 4: "T \ C \ D" shows "{ e(x, y) : \x, y\ \ S } = { f(x, y) : \x, y\ \ T }" - proof - - have s1_1: "{ e(x, y) : \x, y\ \ S } = - { e(p[1], p[2]): p \ S }" - by (auto simp: setOfPairs_def) - have s1_2: "{ f(x, y) : \x, y\ \ T } = - { f(p[1], p[2]): p \ T }" - by (auto simp: setOfPairs_def) - define P where "P \ { e(p[1], p[2]): p \ S }" - define Q where "Q \ { f(p[1], p[2]): p \ T }" - have s1_3: "P = Q" - proof - - have s2_1: "\ x. x \ P <=> x \ Q" - proof - - fix "x" :: "c" - have s3_1: "x \\in P ==> x \\in Q" - proof - - assume s3_1_asm: "x \\in P" - have s4_1: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s3_1_asm setOfAll by (auto simp: P_def) - have s4_2: "\\E p: p \\in S \ - \p[1], p[2]\ \ S \ - x = e(p[1], p[2])" - using s4_1 pairProj_in_rel 3 by auto - have s4_3: "\\E p: p \\in S \ - x = e(p[1], p[2]) \ - (\\u, v\ \ T: - e(p[1], p[2]) = f(u, v))" - using s4_2 1 by auto - have s4_4: "\\u, v\ \ T: - x = f(u, v)" - using s4_3 by auto - have s4_5: "\\E u, v: <> \\in T \ x = f(u, v)" - using s4_4 by auto - have s4_6: "\\E p: p \\in T \ x = f(p[1], p[2])" - using s4_5 4 pairProj_in_rel by (auto simp: two_def) - show "x \\in Q" - using s4_6 setOfAll by (auto simp: Q_def) - qed - have s3_2: "x \\in Q ==> x \\in P" - proof - - assume s3_2_asm: "x \\in Q" - have s4_1: "\\E p: p \\in T \ x = f(p[1], p[2])" - using s3_2_asm setOfAll by (auto simp: Q_def) - have s4_2: "\\E p: p \\in T \ - <> \\in T \ - x = f(p[1], p[2])" - using s4_1 pairProj_in_rel 4 by auto - have s4_3: "\\E p: p \\in T \ - x = f(p[1], p[2]) \ - (\\u, v\ \ S: - f(p[1], p[2]) = e(u, v))" - using s4_2 2 by auto - have s4_4: "\\u, v\ \ S: - x = e(u, v)" - using s4_3 by auto - have s4_5: "\\E u, v: <> \\in S \ x = e(u, v)" - using s4_4 by auto - have s4_6: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_5 3 pairProj_in_rel by (auto simp: two_def) - show "x \\in P" - using s4_6 setOfAll by (auto simp: P_def) - qed - show "x \\in P <=> x \\in Q" - using s3_1 s3_2 by auto - qed - have s2_2: "\ x: x \ P <=> x \ Q" - using s2_1 allI by auto - show ?thesis - using s2_2 extension by auto - qed - show ?thesis - using s1_1 s1_2 s1_3 by (auto simp: P_def Q_def) - qed +using assms by (force intro!: inSetOfPairsI_ex) + subsection \ Basic notions about binary relations \ @@ -1069,25 +875,22 @@ where "rel_range(r) \ { p[2] : p \ r }" definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" -definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) +\ \ composition of binary relations \ +definition rel_comp :: "[c,c] => c" (infixr "\" 75) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : - \ x, z : - p = \x, z\ \ - (\ y: - \x, y\ \ s \ - \y, z\ \ r) }" + \ x, z : p = \x, z\ \ (\ y: \x, y\ \ s \ \y, z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" -definition Id :: "c \ c" (* -- \ diagonal: identity over a set \ *) +definition Id :: "c \ c" \ \ diagonal: identity over a set \ where "Id(A) \ { \x,x\ : x \ A }" text \ Properties of relations \ -definition reflexive (* -- \ reflexivity over a set \ *) +definition reflexive \ \ reflexivity over a set \ where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" @@ -1096,7 +899,7 @@ unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" unfolding isBool_def by (rule boolifyReflexive) -definition symmetric (* -- \ symmetric relation \ *) +definition symmetric \ \ symmetric relation \ where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" @@ -1105,7 +908,7 @@ unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" unfolding isBool_def by (rule boolifySymmetric) -definition antisymmetric (* -- \ antisymmetric relation \ *) +definition antisymmetric \ \ antisymmetric relation \ where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" @@ -1114,7 +917,7 @@ unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" unfolding isBool_def by (rule boolifyAntisymmetric) -definition transitive (* -- \ transitivity predicate \ *) +definition transitive \ \ transitivity predicate \ where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" @@ -1123,7 +926,7 @@ unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" unfolding isBool_def by (rule boolifyTransitive) -definition irreflexive (* -- \ irreflexivity predicate \ *) +definition irreflexive \ \ irreflexivity predicate \ where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" @@ -1132,7 +935,7 @@ unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" unfolding isBool_def by (rule boolifyIrreflexive) -definition equivalence :: "[c,c] \ c" (* -- \ (partial) equivalence relation \ *) +definition equivalence :: "[c,c] \ c" \ \ (partial) equivalence relation \ where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" @@ -1211,7 +1014,7 @@ lemmas converseEqualI = setEqualI [where A = "r^-1" for r, intro!] setEqualI [where B = "r^-1" for r, intro!] -lemma converse_iff [iff]: +lemma converse_iff [simp]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" using r prodProj by (auto simp: converse_def two_def) @@ -1236,8 +1039,7 @@ lemma converseE [elim]: (** consider [elim!] ?? **) proof - from prodProj[OF subsetD[OF converseSubset[OF r] yx]] have 2: "yx = \yx[1], yx[2]\" . with yx have 3: "\yx[2], yx[1]\ \ r" - unfolding converse_def apply auto - using r prodProj by auto + unfolding converse_def using r prodProj by auto from p[of "yx[1]" "yx[2]"] 2 3 show P by simp qed @@ -1245,52 +1047,18 @@ qed lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" - proof - - have s1_1: "\ x. x \\in A \ B ==> x = <>" - using prodProj by (auto simp: two_def) - have s1_2: "\ x. x \\in r ==> x \\in A \ B" - using s1_1 r by auto - have s1_3: "\ x. x \\in r ==> x = <>" - using s1_1 s1_2 by auto - show ?thesis - using s1_3 assms prodProj pairProj_in_rel - by (auto simp: converse_def two_def) - qed +proof - + { + fix x + assume x: "x \ r" + have "x = \x[1], x[2]\" + using x r by (intro prodProj) auto + } + then show ?thesis by (force simp: converse_def two_def) +qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" - proof - - have s1_1: "(A \ B)^-1 = - {<>: p \ A \ B}" - unfolding converse_def by auto - have s1_2: "B \ A = - {<>: p \ B \ A}" - using setOfPairs_triv prodProj by auto - have s1_3: "\ x. x \\in B \\X A ==> - \\E p \\in A \\X B: x = <>" - proof - - fix "x" :: "c" - assume s1_5_asm: "x \\in B \\X A" - have s2_1: "<> \\in A \\X B" - using s1_5_asm by (auto simp: two_def) - define q where "q \ <>" - have s2_2: "q \\in A \\X B \ q = <>" - using s2_1 by (auto simp: q_def) - have s2_3: "x = <>" - using s1_5_asm prodProj by auto - have s2_4: "q[1] = x[2] \ q[2] = x[1]" - using s2_2 by (auto simp: two_def) - have s2_5: "x = <>" - using s2_3 s2_4 by auto - show "\\E p \\in A \\X B: x = <>" - using s2_2 s2_5 by auto - qed - have s1_4: "\ p. p \\in A \\X B ==> - p[1] \\in A \ p[2] \\in B" - using pairProj_in_rel inProd prodProj - by (auto simp: two_def) - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed +by auto lemma converse_empty [simp]: "converse({}) = {}" by auto @@ -1338,42 +1106,8 @@ lemma transitive_converse [simp]: lemma symmetric_iff_converse_eq: assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" - proof - - have s1_1: "symmetric(r) ==> r^-1 = r" - proof - - assume s1_1_asm: "symmetric(r)" - have s2_1: "\ p. p \\in r^-1 ==> p \\in r" - proof - - fix "p" :: "c" - assume s2_1_asm: "p \\in r^-1" - show "p \\in r" - using r s1_1_asm s2_1_asm - by (auto elim!: converseE simp: symmetric_def) - qed - have s2_2: "\ p. p \\in r ==> p \\in r^-1" - proof - - fix "p" :: "c" - assume s2_2_asm: "p \\in r" - have s3_1: "p = <>" - using s2_2_asm r by (intro prodProj, auto) - have s3_2: "<> \\in r" - using s1_1_asm s2_2_asm s3_1 by (auto simp: symmetric_def) - show "p \\in r^-1" - using s3_1 s3_2 by (auto dest: converseI) - qed - show "r^-1 = r" - using s2_1 s2_2 setEqualI[of "r^-1" "r"] - by blast - qed - have s1_2: "r^-1 = r ==> symmetric(r)" - proof - - assume s1_2_asm: "r^-1 = r" - show "symmetric(r)" - using s1_2_asm by (auto simp: symmetric_def) - qed - show ?thesis - using s1_1 s1_2 by auto - qed +using assms by (auto simp: symmetric_def) + subsubsection \ Identity relation over a set \ @@ -1483,53 +1217,30 @@ qed lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" - using r - proof - - have s1_1: "\ p. p \\in R ==> p \\in R \ Id(B)" - proof - - fix "p" :: "c" - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) - with 1 show "p \ R \ Id(B)" by simp - qed - have s1_2: "\ p. p \\in R \ Id(B) ==> p \\in R" - proof - - fix "p" :: "c" - assume s1_2_asm: "p \\in R \ Id(B)" - show "p \\in R" - using r s1_2_asm by auto - qed - show ?thesis - using s1_1 s1_2 setEqualI[of "R \ Id(B)" "R"] by blast - qed - +proof (rule setEqualI) + fix x + assume "x \ R \ Id(B)" + with r show "x \ R" by auto +next + fix x + assume x: "x \ R" + with r obtain b c where "b \ B" "c \ C" "x = \b,c\" by force + with x r show "x \ R \ Id(B)" by auto +qed lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" - using r - proof - - have s1_1: "\ p. p \\in R ==> p \\in Id(B) \ R" - proof - - fix "p" :: "c" - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) - with 1 show "p \ Id(B) \ R" by simp - qed - have s1_2: "\ p. p \\in Id(B) \ R ==> p \\in R" - proof - - fix "p" :: "c" - assume s1_2_asm: "p \\in Id(B) \ R" - show "p \\in R" - using r s1_2_asm by auto - qed - show ?thesis - using s1_1 s1_2 setEqualI[of "Id(B) \ R" "R"] by blast - qed +proof (rule setEqualI) + fix x + assume "x \ Id(B) \ R" + with r show "x \ R" by auto +next + fix x + assume x: "x \ R" + with r obtain a b where "a \ A" "b \ B" "x = \a,b\" by force + with x r show "x \ Id(B) \ R" by auto +qed lemma rel_comp_empty1 [simp]: "{} \ R = {}" unfolding rel_comp_def by auto @@ -1553,7 +1264,7 @@ proof (rule setEqualI) fix u assume "\y,u\ \ S" and "\u,z\ \ R" with r s t 2 have "\x,z\ \ R \ (S \ T)" - by (intro compI, auto elim!: relProj1 relProj2) + by (intro compI) (auto elim!: relProj1 relProj2) with 1 show ?thesis by simp qed qed @@ -1570,7 +1281,7 @@ next fix u assume "\x,u\ \ T" and "\u,y\ \ S" with r s t 3 have "\x,z\ \ (R \ S) \ T" - by (intro compI, auto elim!: relProj1 relProj2) + by (intro compI) (auto elim!: relProj1 relProj2) with 1 show ?thesis by simp qed qed @@ -1759,7 +1470,7 @@ proof - have sym: "symmetric(r)" proof - from comp have "r^-1 = (r^-1 \ r)^-1" by simp - also from r r1 have "... = r^-1 \ r" by (simp add: converse_comp) + also from r r1 have "\ = r^-1 \ r" by (simp add: converse_comp) finally have "r^-1 = r" by (simp add: comp) with r show ?thesis by (simp add: symmetric_iff_converse_eq) qed From 8c82ca5fa7a930ad0c45919bfde33d3e759df2f7 Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 29 Jan 2021 18:54:56 +0100 Subject: [PATCH 058/167] remaining syntax translations for Isabelle-2020 --- isabelle/Constant.thy | 5 +-- isabelle/Strings.thy | 101 +++++++++++++++++++++++++++--------------- 2 files changed, 68 insertions(+), 38 deletions(-) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index f7aff81d..dba29b43 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Constant.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:28 merz> + Version: Isabelle2020 *) section \ Main theory for constant-level Isabelle/\tlaplus{} \ diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index cabb62c8..42152c84 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Strings.thy Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:15 merz> + Version: Isabelle2020 *) section \ Characters and strings \ @@ -20,7 +19,23 @@ text \ \ definition Nibble -where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" +(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) + where "Nibble \ {0, + 1, + Succ[1], + Succ[Succ[1]], + Succ[Succ[Succ[1]]], + Succ[Succ[Succ[Succ[1]]]], + Succ[Succ[Succ[Succ[Succ[1]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]]]}" definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" @@ -30,7 +45,8 @@ by (simp add: char_def) (*** example ** lemma "char(3,13) \ char(14,2)" -by simp + by (simp add: two_def three_def four_def five_def six_def seven_def eight_def + nine_def ten_def eleven_def twelve_def thirteen_def fourteen_def) ***) definition Char @@ -59,7 +75,6 @@ text \ lexicon. \ -(* TODO parse_ast_translation \ let (* convert an ML integer to a nibble *) @@ -75,46 +90,64 @@ parse_ast_translation \ mkNibble (ord c div 16), mkNibble (ord c mod 16)] else error ("Non-ASCII symbol: " ^ quote c); + (* remove two leading quotes from a list of characters *) + fun trim_quotes ("'" :: ( "'" :: cs)) = cs + | trim_quotes cs = cs + (* convert a list of ML characters into a TLA+ string, in reverse order *) fun list2TupleReverse [] = Ast.Constant "Tuples.emptySeq" | list2TupleReverse (c :: cs) = Ast.Appl [Ast.Constant "Tuples.Append", list2TupleReverse cs, mkChar c]; (* parse AST translation for characters *) - fun char_ast_tr [Ast.Variable xstr] = - (case Lexicon.explode_xstr xstr of + fun char_ast_tr [Ast.Variable str] = + let val trimmed_str = (trim_quotes (rev (trim_quotes (Symbol.explode str)))) + in (case trimmed_str of [c] => mkChar c - | _ => error ("Expected single character, not " ^ xstr)) - | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); + | _ => error ("Expected single character, not " ^ str)) + end + | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); (* parse AST translation for strings *) - fun string_ast_tr [Ast.Variable xstr] = - list2TupleReverse (rev (Lexicon.explode_xstr xstr)) + fun string_ast_tr [Ast.Variable str] = + list2TupleReverse (trim_quotes (rev ((trim_quotes (Symbol.explode str))))) | string_ast_tr asts = raise Ast.AST ("string_ast_tr", asts); in - [("_Char", char_ast_tr), ("_String", string_ast_tr)] + [("_Char", K char_ast_tr), ("_String", K string_ast_tr)] end \ -*) (** debug ** -(*ML \ set Syntax.trace_ast; \*) +(*declare [[syntax_ast_trace = true]]*) lemma "''a''" oops +lemma "''''" +oops + +lemma "''abc''" +oops + lemma "CHAR ''a''" oops -(*ML \ reset Syntax.trace_ast; \*) +(* declare [[syntax_ast_trace = false]] *) **) -(* TODO print_ast_translation \ let - (* convert a nibble to an ML integer -- because translation macros have - already been applied, we see constants "0" through "15", not Succ[...] terms! *) + (* convert a nibble to an ML integer *) + fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 + | destNibble (Ast.Constant @{const_syntax "one"}) = 1 + | destNibble (Ast.Appl [Ast.Constant @{const_syntax "Functions.fapply"}, + Ast.Constant @{const_syntax "Peano.Succ"}, nb]) + = (destNibble nb) + 1 + | destNibble _ = raise Match + + (* the following version should be used when 2 .. 15 are abbreviations, not definitions *) +(* fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 | destNibble (Ast.Constant @{const_syntax "two"}) = 2 @@ -132,7 +165,7 @@ print_ast_translation \ | destNibble (Ast.Constant @{const_syntax "fourteen"}) = 14 | destNibble (Ast.Constant @{const_syntax "fifteen"}) = 15 | destNibble _ = raise Match; - +*) (* convert a pair of nibbles to an ML character *) fun destNbls nb1 nb2 = let val specials = raw_explode "\"\\`'" @@ -154,13 +187,12 @@ print_ast_translation \ (* convert a list of TLA+ characters to the output representation of a TLA+ string *) fun list2String cs = Ast.Appl [Ast.Constant "_inner_string", - Ast.Variable (Lexicon.implode_xstr cs)]; + Ast.Variable (Lexicon.implode_str cs)]; (* print AST translation for single characters that do not occur in a string *) fun char_ast_tr' [nb1, nb2] = - Ast.Appl [Ast.Constant @{syntax_const "_Char"}, - list2String [destNbls nb1 nb2]] - | char_ast_tr' _ = raise Match; + Ast.Appl [Ast.Constant @{syntax_const "_Char"}, list2String [destNbls nb1 nb2]] + | char_ast_tr' _ = raise Match (* print AST translation for non-empty literal strings, fails (by raising exception Match) @@ -168,14 +200,12 @@ print_ast_translation \ fun string_ast_tr' [args] = list2String (map destChar (tuple2List args)) | string_ast_tr' _ = raise Match; in - [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] + [(@{const_syntax "char"}, K char_ast_tr'), ("@tuple", K string_ast_tr')] end \ -*) (*** examples ** -(* not sure if the following is the way to set tracing? *) -(* ML \ val _ = Config.put Ast.trace true @{context} \ *) +(*declare [[syntax_ast_trace = true]]*) lemma "CHAR ''a'' \ CHAR ''b''" by simp @@ -193,9 +223,10 @@ lemma "''ab''[1] = CHAR ''a''" by simp lemma "''abc''[2] \ CHAR ''a''" -by simp +by (simp add: two_def) + +(* declare [[syntax_ast_trace = false]] *) -(* ML \ val _ = Config.put Ast.trace false @{context} \ *) **) @@ -220,7 +251,7 @@ by auto lemma "r \ [''bar'' : BOOLEAN, ''foo'' : Nat] \ [r EXCEPT ![''foo''] = 3] \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by force (* "by auto" also works, but is slow *) +by (force simp: two_def three_def) (* "by auto" also works, but is slow *) lemma "(''a'' :> 1) \ (''b'' :> 1)" by simp @@ -229,19 +260,19 @@ lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1)" by simp lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1 @@ ''b'' :> 3)" -by simp +by (simp add: three_def) lemma "(''a'' :> 1 @@ ''b'' :> 2) = (''b'' :> 2 @@ ''a'' :> 1)" by simp lemma "''ab'' = [i \ {1,2} \ ''ab''[i]]" -by auto +by (auto simp: two_def) lemma "''ab'' = (1 :> CHAR ''a'') @@ (2 :> CHAR ''b'')" -by auto +by (auto simp: two_def) lemma "Len(''ab'') = 2" -by simp +by (simp add: two_def) **) From 2c9211db1f31830fa0a4a6c780dabd7e95e87a04 Mon Sep 17 00:00:00 2001 From: merz Date: Sun, 7 Feb 2021 09:42:39 +0100 Subject: [PATCH 059/167] restored document preparation --- isabelle/PredicateLogic.thy | 6 +++--- isabelle/ROOT | 11 +++++------ isabelle/ROOT.ML | 15 --------------- isabelle/SetTheory.thy | 2 +- 4 files changed, 9 insertions(+), 25 deletions(-) delete mode 100644 isabelle/ROOT.ML diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index afade712..cd9411a9 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -157,10 +157,10 @@ text \ The following axioms govern propositional logic: \begin{enumerate} - \item @{term TRUE} is valid, and @{term FALSE} implies anything}, - \item @{text A} is a theorem if and only if it equals @{term TRUE}, + \item @{text TRUE} is valid, and @{text FALSE} implies anything, + \item @{text A} is a theorem if and only if it equals @{text TRUE}, \item conditionals are reasoned about by case distinction, - \item equations evaluate to @{term TRUE} or @{term FALSE}. + \item equations evaluate to @{text TRUE} or @{text FALSE}. \end{enumerate} \ diff --git a/isabelle/ROOT b/isabelle/ROOT index 84a6f8b6..2bd32182 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,12 +1,11 @@ -(* build the session using "isabelle build -D ." *) +(* build the session using "isabelle build -b -D ." *) +chapter "TLA+" session "TLA+" = "Pure" + - (* - options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] - *) + options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + theories Constant (global) Zenon (global) - (* + document_files "root.tex" - *) diff --git a/isabelle/ROOT.ML b/isabelle/ROOT.ML deleted file mode 100644 index f174a865..00000000 --- a/isabelle/ROOT.ML +++ /dev/null @@ -1,15 +0,0 @@ -(* Title: TLA+/ROOT.ML - Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2011 - Time-stamp: <2011-07-26 16:53:42 merz> - -Encoding of TLA+ constant level as an object logic in Isabelle. -*) - -val banner = "TLA+"; -writeln banner; - -use_thy "Constant"; -use_thy "Zenon"; diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index 707897f4..bbb6fbe0 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -259,7 +259,7 @@ definition setminus :: "[c, c] \ c" (infixl "\" 65) ( definition INTER :: "c \ c" ("INTER _" [100]90) (* generalized intersection *) where "INTER A \ { x \ UNION A : \B \ A : x \ B }" -text \Note that @{text "INTER {}"} is @{text "{}"}.}\ +text \Note that @{text "INTER {}"} is @{text "{}"}.\ notation (ASCII) "inter" (infixl "\\inter" 70) and From 4a5bf2d5c94844e64fa53db86f5d2c8f54a1e805 Mon Sep 17 00:00:00 2001 From: merz Date: Thu, 23 Dec 2021 18:49:50 +0100 Subject: [PATCH 060/167] lemmas for SMT axioms --- isabelle/Integers.thy | 4 +- isabelle/NewSMT.thy | 281 ++++++++++++++++++++++++++++++++++++++++++ isabelle/ROOT | 1 + 3 files changed, 284 insertions(+), 2 deletions(-) create mode 100644 isabelle/NewSMT.thy diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 776ddd84..fdcc141f 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -419,7 +419,7 @@ lemmas int_leq_pp_def = nat_leq_def (* -- 'positive-positive' case, ie: both arguments are naturals *) axiomatization where - int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = FALSE" + int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = (a = 0 \ b = 0)" and int_leq_np_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ b = TRUE" and @@ -445,7 +445,7 @@ using assms by simp lemma neg_le_iff_le [simp]: "\m \ Int; n \ Int \ \ -.n \ -.m = (m \ n)" -by(rule intCases2[of m n], simp_all) + by (rule intCases2[of m n]) auto subsection \ Addition of integers \ diff --git a/isabelle/NewSMT.thy b/isabelle/NewSMT.thy new file mode 100644 index 00000000..766ae093 --- /dev/null +++ b/isabelle/NewSMT.thy @@ -0,0 +1,281 @@ +section \Axioms used by the new SMT backend\ + +theory NewSMT + imports Constant +begin + +subsection \Set Theory\ + +lemma ChooseDef: "P(x) \ P(Choice(P))" + by (rule chooseI) + +lemma SetExtensionality: "(\z : z \ a \ z \ b) \ a = b" + by (simp add: extension) + +lemma SubseteqIntro: "(\z : z \ a \ z \ b) \ a \ b" + by blast + +lemma SubseteqElim: "a \ b \ x \ a \ x \ b" + by blast + +text \ + The two following axioms are stated for $n$-ary enumerations, + we prove their instances for $n=2$. +\ +lemma EnumIntro: "a \ {a,b} \ b \ {a,b}" + by blast + +lemma EnumElim: "x \ {a,b} \ x=a \ x=b" + by blast + +lemma SubsetDef: "x \ SUBSET a \ x \ a" + by blast + +lemma UnionIntro: "x \ y \ y \ a \ x \ UNION a" + by blast + +lemma UnionElim: "x \ UNION a \ (\y : x \ y \ y \ a)" + by blast + +lemma SetstDef: "x \ subsetOf(a,P) \ x \ a \ P(x)" + by (rule subsetOf) + +text \ + The two following axioms are stated for an $n$-ary + operator $f$ but Isabelle only provides the construct + for unary operators. +\ +lemma SetofIntro: "y \ a \ f(y) \ setOfAll(a,f)" + by blast + +lemma SetofElim: "x \ setOfAll(a,f) \ (\y : y \ a \ x = f(y))" + by blast + +lemma CupDef: "x \ a \ b \ x \ a \ x \ b" + by blast + +lemma CapDef: "x \ a \ b \ x \ a \ x \ b" + by blast + +lemma SetminusDef: "x \ a \ b \ x \ a \ x \ b" + by blast + + +subsection \Functions\ + +lemma FcnExtensionality: + assumes "isAFcn(f)" "isAFcn(g)" "DOMAIN f = DOMAIN g" + "\x \ DOMAIN f : f[x] = g[x]" + shows "f = g" + using assms by auto + +lemma FcnIsafcn: "isAFcn(Fcn(a,F))" + by (rule fcnIsAFcn) + +lemma FcnDom: "DOMAIN Fcn(a,F) = a" + by (rule DOMAIN) + +lemma FcnApp: "x \ a \ Fcn(a,F)[x] = F(x)" + by auto + +lemma ArrowIntro: + assumes "isAFcn(f)" "DOMAIN f = a" "\x \ a : f[x] \ b" + shows "f \ [a \ b]" + using assms by (rule inFuncSet) + +lemma ArrowElim1: "f \ [a \ b] \ isAFcn(f) \ DOMAIN f = a" + by blast + +lemma ArrowElim2: "f \ [a \ b] \ x \ a \ f[x] \ b" + by blast + +lemma ExceptIsafcn: "isAFcn(except(f,x,y))" + by blast + +lemma ExceptDom: "DOMAIN except(f,x,y) = DOMAIN f" + by (rule domainExcept) + +lemma ExceptApp1: "x \ DOMAIN f \ except(f,x,y)[x] = y" + by auto + +lemma ExceptApp2: "z \ DOMAIN f \ z \ x \ except(f,x,y)[z] = f[z]" + by auto + + +subsection \Integers\ + +lemma PlusTyping: "x \ Int \ y \ Int \ x+y \ Int" + by simp + +lemma UminusTyping: "x \ Int \ -.x \ Int" + by simp + +lemma MinusTyping: "x \ Int \ y \ Int \ x-y \ Int" + by simp + +lemma TimesTyping: "x \ Int \ y \ Int \ x*y \ Int" + by simp + +(* Isabelle currently provides integer division only for natural + numbers, so we cannot prove the following. + +lemma DivTyping: "x \ Int \ y \ Int \ y \ 0 \ x \\div y \ Int" +*) + +lemma NatDef: "x \ Nat \ x \ Int \ 0 \ x" +proof + assume "x \ Nat" + then show "x \ Int \ 0 \ x" by simp +next + assume "x \ Int \ 0 \ x" + hence "x \ Int" "0 \ x" by blast+ + from \x \ Int\ show "x \ Nat" + proof (rule intCases) + fix m + assume "m \ Nat" "x = -.m" + with \0 \ x\ show ?thesis + by (cases m) auto + qed +qed + +(* Isabelle currently does not support integer intervals, + so we cannot prove the following. + +lemma RangeDef: "x \ a .. b \ x \ Int \ a \ x \ x \ b" +*) + + +subsection \Booleans\ + +text \ + The following lemmas do not exactly correspond to the SMT + axioms because Isabelle/TLA+ is untyped and does not have + cast operators. +\ + +lemma BooleanIntro: "TRUE \ BOOLEAN \ FALSE \ BOOLEAN" + by blast + +lemma BooleanElim: "x \ BOOLEAN \ x = TRUE \ x = FALSE" + by blast + + +subsection \Tuples and Records\ + +text \ + The following SMT axioms are for tuples or records + of arbitrary length, we prove instances for pairs. +\ + +lemma TupIsafcn: "isAFcn(\x,y\)" + by blast + +lemma TupDom: "DOMAIN \x,y\ = {1,2}" + by (auto simp: two_def) + +lemma TupApp: "\x,y\[1] = x \ \x,y\[2] = y" + by (auto simp: two_def) + +lemma ProductIntro: "x \ a \ y \ b \ \x,y\ \ a \ b" + by auto + +lemma ProductElim: "x \ a \ b \ x = \x[1], x[2]\" + by (rule prodProj) + +lemma RecordIsafcn: "isAFcn(''foo'' :> x @@ ''bar'' :> y)" + by auto + +lemma RecordDom: "DOMAIN (''foo'' :> x @@ ''bar'' :> y) = {''foo'', ''bar''}" + by auto + +lemma RecordApp: + "(''foo'' :> x @@ ''bar'' :> y)[''foo''] = x" + "(''foo'' :> x @@ ''bar'' :> y)[''bar''] = y" + by auto + +lemma RectIntro: + assumes "x \ a" "y \ b" + shows "(''foo'' :> x @@ ''bar'' :> y) \ [''foo'' : a, ''bar'' : b]" + using assms by auto + +lemma RectElim: + assumes "x \ [''foo'' : a, ''bar'' : b]" + shows "x = (''foo'' :> x[''foo''] @@ ''bar'' :> x[''bar''])" + using assms by force + + +subsection \Sequences\ + +text \ + Somewhat rewritten, in particular the original hypothesis + @{text "\i: i \ DOMAIN s \ 1 \ i \ i \ Len(s)"} + doesn't imply that @{text "DOMAIN s \ Nat"}. + Could instead assume + @{text "DOMAIN s = 1 .. Len(s)"} and + @{text "\i \ DOMAIN s : s[i] \ a"}. +\ +lemma SeqIntro: + assumes "isAFcn(s)" "Len(s) \ Nat" + "\i : i \ DOMAIN s \ i \ Nat \ 1 \ i \ i \ Len(s)" + "\i \ Nat : 1 \ i \ i \ Len(s) \ s[i] \ a" + shows "s \ Seq(a)" + using assms by auto + +lemma SeqElim1: + "s \ Seq(a) \ isAFcn(s) \ Len(s) \ Nat \ DOMAIN s = 1 .. Len(s)" + by auto + +text \ + Similarly rewritten as theorem @{text SeqIntro} above. +\ +lemma SeqElim2: + "s \ Seq(a) \ i \ Nat \ 1 \ i \ i \ Len(s) \ s[i] \ a" + by auto + +lemma LenDef: + "x \ Nat \ DOMAIN s = 1 .. x \ Len(s) = x" + by auto + +text \ + Isabelle/TLA+ currently doesn't know about concatenation of sequences, + so we cannot prove lemmas corresponding to the SMT axioms about concatenation. +\ + +lemma AppendTyping: "s \ Seq(a) \ x \ a \ Append(s,x) \ Seq(a)" + by auto + +lemma AppendLen: + assumes "Len(s) \ Nat" + shows "Len(Append(s,x)) = Len(s)+1" +proof + from assms show "DOMAIN Append(s,x) = 1 .. (Len(s)+1)" by auto +next + from assms show "Len(s)+1 \ Nat" by auto +qed + +text \ + Some changes w.r.t. the original formulations of the axioms about @{text Append}. +\ +lemma AppendApp1: + assumes "Len(s) \ Nat" (* missing from original formulation *) + "i \ Nat" (* changed from Int because Isabelle doesn't support Int intervals *) + "1 \ i" "i \ Len(s)" + shows "Append(s,x)[i] = s[i]" + using assms unfolding Append_def by auto + +lemma AppendApp2: + "Len(s) \ Nat \ Append(s,x)[Len(s)+1] = x" + unfolding Append_def by auto + +text \ + Isabelle/TLA+ doesn't know about the remaining operations on sequences. +\ + +lemma TupSeqTyping: "x \ a \ y \ a \ \x,y\ \ Seq(a)" + by auto + +lemma TupSeqLen: "Len(\x,y\) = 2" + by (auto simp: two_def) + + +end diff --git a/isabelle/ROOT b/isabelle/ROOT index 2bd32182..30960a90 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -6,6 +6,7 @@ session "TLA+" = "Pure" + theories Constant (global) Zenon (global) +(* NewSMT (global) *) document_files "root.tex" From d15857b1e1b1188b03fc3e9bf222cd7bd0f4961b Mon Sep 17 00:00:00 2001 From: merz Date: Wed, 16 Mar 2022 16:19:52 +0100 Subject: [PATCH 061/167] refactoring of integer theories, update to Isabelle 2021-1 --- isabelle/CaseExpressions.thy | 16 +- isabelle/Constant.thy | 6 +- isabelle/IntegerArithmetic.thy | 2140 ++++++++++++++++++++++++++ isabelle/IntegerDivision.thy | 713 +++++++++ isabelle/Integers.thy | 2579 +++++++++++++++++--------------- isabelle/README.html | 61 +- isabelle/Strings.thy | 48 +- isabelle/Tuples.thy | 511 ++++--- isabelle/Zenon.thy | 498 ++---- 9 files changed, 4716 insertions(+), 1856 deletions(-) create mode 100644 isabelle/IntegerArithmetic.thy create mode 100644 isabelle/IntegerDivision.thy diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 6507100a..f9c6cbeb 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -1,8 +1,8 @@ (* Title: TLA+/CaseExpressions.thy - Author: Stephan Merz, LORIA - Copyright (C) 2009-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Case expressions \ @@ -191,15 +191,15 @@ lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def lemma "i=1 \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = 0" -by (simp add: two_def) + by simp lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" -by auto + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" + by auto lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" -by auto + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" + by auto **) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index dba29b43..d085827c 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -1,14 +1,14 @@ (* Title: TLA+/Constant.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Copyright (C) 2008-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Main theory for constant-level Isabelle/\tlaplus{} \ theory Constant -imports NatDivision CaseExpressions Strings Integers +imports IntegerDivision CaseExpressions Strings begin text \ diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy new file mode 100644 index 00000000..07aaf7a6 --- /dev/null +++ b/isabelle/IntegerArithmetic.thy @@ -0,0 +1,2140 @@ +(* Title: TLA+/IntegerArithmetic.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation + License: BSD + Version: Isabelle2021-1 +*) + +section \ Arithmetic (except division) for the integers \ + +theory IntegerArithmetic +imports Integers +begin + +named_theorems algebra_simps "algebra simplification rules" + +text \ + The rewrites accumulated in @{text algebra_simps} deal with the + classical algebraic structures of groups, rings and family. They simplify + terms by multiplying everything out (in case of a ring) and bringing sums and + products into a canonical form (by ordered rewriting). As a result these + rewrites decide the equational theory of groups and rings but also help + with inequalities. +\ + + +subsection \ Addition and difference of integers \ + +text \ + We define integer addition as a recursive function. + Subtraction is just addition of the complement, and + is introduced as an abbreviation. +\ + +definition pr_addint where +"pr_addint(a) \ int_primrec(Int, a, \n x y. succ[x], \n x y. pred[y])" + +definition addint :: "[c,c] \ c" (infixl "+" 65) +where "\a \ Int; b \ Int\ \ (a + b) \ pr_addint(a)[b]" + +abbreviation subint (infixl "-" 65) +where "a - b \ a + (-b)" + +text \ Closure \ + +lemma pr_addintType: + assumes "a \ Int" shows "pr_addint(a) \ [Int \ Int]" + using assms unfolding pr_addint_def by (rule int_primrecType) auto + +lemma addintType [iff]: "\a \ Int; b \ Int\ \ a+b \ Int" + unfolding addint_def by (blast dest: pr_addintType) + +text \ Base case and successor cases \ + +lemma pr_addint_0: "a \ Int \ pr_addint(a)[0] = a" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma add_0_int [simp]: "a \ Int \ a + 0 = a" + by (simp add: addint_def pr_addint_0) + +lemma pr_addint_SuccNat: + "\a \ Int; n \ Nat\ \ pr_addint(a)[succ[n]] = succ[pr_addint(a)[n]]" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma pr_addint_uminusSuccNat: + "\a \ Int; n \ Nat\ \ pr_addint(a)[-succ[n]] = pred[pr_addint(a)[-n]]" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma addint_succ [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + succ[b] = succ[a+b]" + using b a + by (cases b) + (auto dest: pr_addintType + simp: addint_def pr_addint_SuccNat pr_addint_uminusSuccNat) + +lemma addint_pred [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + pred[b] = pred[a+b]" + using b a + by (cases b) + (auto dest: pr_addintType + simp: addint_def pr_addint_SuccNat pr_addint_uminusSuccNat pred0 sym[OF uminusSucc]) + +lemma subint_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "a - succ[b] = pred[a-b]" + using assms by (simp add: uminusSucc) + +lemma subint_pred [simp]: + assumes "a \ Int" and "b \ Int" + shows "a - pred[b] = succ[a-b]" + using assms by (simp add: uminusPred) + +lemma addintNat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m + n \ Nat" + using n m by induct auto + +lemma addint_0_left [simp]: + assumes "a \ Int" + shows "0 + a = a" + using assms + by induct (auto simp: uminusSucc) + +lemma add_succ_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[a] + b = succ[a + b]" + using b a by (induct b) auto + +lemma add_pred_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[a] + b = pred[a + b]" + using b a by (induct b) auto + +lemma add_uminus_succ_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "-succ[a] + b = pred[-a + b]" + using assms by (simp add: uminusSucc) + +lemma add_uminus_pred_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "-pred[a] + b = succ[-a + b]" + using assms by (simp add: uminusPred) + +lemma succ_pred_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[pred[a] + b] = a + b" + using b a by induct auto + +lemma pred_succ_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[succ[a] + b] = a + b" + using b a by induct auto + +lemma subint_self [simp]: + assumes "a \ Int" + shows "a - a = 0" + using assms by induct auto + +lemma addint_uminus_self [simp]: + assumes "a \ Int" + shows "-a + a = 0" + using assms by induct auto + +lemma uminus_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "-(a + b) = -a - b" + using b a by induct (auto simp: uminusSucc) + +text \ Commutativity \ + +lemma add_commute [algebra_simps]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + b = b + a" + using b a by (induct b) (auto simp: uminusSucc) + +text \ Associativity \ + +lemma add_assoc [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + (b + c) = (a + b) + c" + using assms by induct simp+ + +lemma add_left_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + (b + c) = b + (a + c)" + using assms by (simp only: add_assoc add_commute) + +lemma add_right_commute: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + b + c = a + c + b" +proof - + from assms have "a + b + c = a + (b + c)" + by (rule sym[OF add_assoc]) + also from assms have "\ = a + (c + b)" + by (simp add: add_commute) + finally show ?thesis + using assms by (simp add: add_assoc) +qed + +lemmas add_ac = add_assoc add_commute add_left_commute + + +text \ Cancellation rules \ + +lemma add_left_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b = a + c) = (b = c)" + using assms by induct simp+ + +lemma add_right_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(b + a = c + a) = (b = c)" + using assms by induct simp+ + +lemma add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b = a) = (b = 0)" + using assms by induct (simp add: uminusSucc)+ + +lemma add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (a = a + b) = (b = 0)" + "\a \ Int; b \ Int\ \ (b + a = a) = (b = 0)" + "\a \ Int; b \ Int\ \ (a = b + a) = (b = 0)" + by (auto simp: add_commute dest: sym) + +lemma succ_add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a+b] = a) = (b = -1)" +proof - + { + assume "succ[a+b] = a" + with assms have "a + succ[b] = a" by simp + with assms add_eq_self1[of "a" "succ[b]"] + have "succ[b] = 0" by simp + hence "pred[succ[b]] = pred[0]" by simp + with assms have "b = -1" by (simp add: pred0) + } + with assms show ?thesis by auto +qed + +lemma succ_add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (succ[b+a] = a) = (b = -1)" + "\a \ Int; b \ Int\ \ (a = succ[a+b]) = (b = -1)" + "\a \ Int; b \ Int\ \ (a = succ[b+a]) = (b = -1)" + by (auto simp: add_commute dest: sym) + +lemma pred_add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a+b] = a) = (b = 1)" +proof - + { + assume "pred[a+b] = a" + with assms have "a + pred[b] = a" by simp + with assms add_eq_self1[of "a" "pred[b]"] + have "pred[b] = 0" by simp + hence "succ[pred[b]] = succ[0]" by simp + with assms have "b = 1" by simp + } + with assms show ?thesis by auto +qed + +lemma pred_add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (pred[b+a] = a) = (b = 1)" + "\a \ Int; b \ Int\ \ (a = pred[a+b]) = (b = 1)" + "\a \ Int; b \ Int\ \ (a = pred[b+a]) = (b = 1)" + by (auto simp: add_commute dest: sym) + + +text \ Rules about integer difference \ + +lemma diff_diff_left: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(a - b) - c = a - (b + c)" + using assms by (simp add: add_assoc) + +lemma add_diff_same_1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b) - a = b" + using assms by (simp add: add_ac) + +lemma add_diff_same_2 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b) - b = a" +proof - + from assms have "(a + b) - b = a + (b - b)" + by (intro sym[OF add_assoc]) auto + with assms show ?thesis by simp +qed + +lemma diff_add_same_1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a - b) + b = a" + using assms add_diff_same_2[of "a" "-b"] by simp + +lemma diff_add_same_2 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a + b) + a = b" + using assms by (simp add: add_ac) + +lemma add_diff_cancel_left [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b) - (a + c) = b - c" + using assms by (simp add: add_ac) + +lemma add_diff_cancel_right [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b) - (c + b) = a - c" (is "?lhs = ?rhs") +proof - + from assms have "?lhs = a + (b - (c + b))" + by (intro sym[OF add_assoc]) simp+ + also have "\ = ?rhs" + using assms by (simp add: add_ac) + finally show ?thesis . +qed + +lemma add_left_diff_right: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b = c) = (a = c - b)" +proof - + { + assume "a + b = c" + hence "(a + b) - b = c - b" by simp + with assms have "a = c - b" by simp + } + moreover + { + assume "a = c - b" + hence "a + b = (c - b) + b" by simp + with assms have "a + b = c" by simp + } + ultimately show ?thesis by auto +qed + +lemma diff_0_eq [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a - b = 0) = (a = b)" +proof - + { + assume "a - b = 0" + hence "(a - b) + b = 0 + b" by simp + with assms have "a = b" by simp + } + with assms show ?thesis by auto +qed + +lemma add_right_diff_left: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a = b + c) = (c = a - b)" + using assms by (auto simp: add_left_diff_right add_assoc) + +lemma diff_0_eqs [simp]: + "\a \ Int; b \ Int\ \ (-a + b = 0) = (a = b)" + "\a \ Int; b \ Int\ \ (0 = a - b) = (a = b)" + "\a \ Int; b \ Int\ \ (0 = -a + b) = (a = b)" + by (auto simp: add_commute dest: sym) + +text \ Reasoning about @{text "m + n = 0"}, etc. \ + +lemma add_is_0_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(m + n = 0) = (m = 0 \ n = 0)" + using m apply cases + using n by (induct, auto) + +lemma zero_is_add_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(0 = m + n) = (m = 0 \ n = 0)" + using m apply cases + using n by (induct, auto) + +lemma add_is_1_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(m + n = 1) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" + using m apply cases + using n by (induct, auto simp: add_is_0_nat) + +lemma one_is_add_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(1 = m + n) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" + using m apply (rule natCases) + using n by (induct, auto simp: zero_is_add_nat) + + +subsection \ Multiplication of natural numbers \ + +definition pr_multint where +"pr_multint(a) \ int_primrec(Int, 0, \n x y. x+a, \n x y. y-a)" + +definition multint :: "[c,c] \ c" (infixl "*" 70) +where "\a \ Int; b \ Int\ \ (a * b) \ pr_multint(a)[b]" + +text \ Closure \ + +lemma pr_multintType: + assumes "a \ Int" shows "pr_multint(a) \ [Int \ Int]" + using assms unfolding pr_multint_def + by (intro int_primrecType) auto + +lemma multintType [iff]: "\a \ Int; b \ Int\ \ a*b \ Int" + unfolding multint_def by (blast dest: pr_multintType) + +text \ Base case and successor cases \ + +lemma pr_multint_0: "a \ Int \ pr_multint(a)[0] = 0" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma times_0 [simp]: "a \ Int \ a * 0 = 0" + by (simp add: multint_def pr_multint_0) + +lemma pr_multint_SuccNat: + "\a \ Int; n \ Nat\ \ pr_multint(a)[succ[n]] = pr_multint(a)[n] + a" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma pr_multint_uminusSuccNat: + "\a \ Int; n \ Nat\ \ pr_multint(a)[-succ[n]] = pr_multint(a)[-n] - a" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma times_succ [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * succ[b] = a*b + a" + using b a + by (cases b) + (simp add: multint_def pr_multint_SuccNat pr_multint_uminusSuccNat + pr_multintType[THEN funcSetValue])+ + +lemma times_pred [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * pred[b] = a*b - a" + using b a + by (cases b) + (simp add: multint_def pr_multint_SuccNat pr_multint_uminusSuccNat + pr_multintType[THEN funcSetValue] pred0 sym[OF uminusSucc])+ + +lemma times_uminus_right [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * (-b) = -(a * b)" +using b proof induct + case 0 + with a show ?case by simp +next + case (pos n) + with a show ?case by (simp add: uminusSucc) +next + case (neg n) + with a show ?case + by simp (simp add: uminusSucc) +qed + +lemma times_nat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m * n \ Nat" + using n m by induct simp+ + +lemma times_0_left [simp]: + assumes "a \ Int" + shows "0 * a = 0" + using assms by induct simp+ + +lemma times_succ_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[a] * b = a * b + b" +using b proof induct + case 0 + with a show ?case by simp +next + case (pos n) + with a show ?case by (simp add: add_ac) +next + case (neg n) + with a show ?case + (* the two applications of simp cannot be combined ... *) + by (simp add: add_right_commute) (simp add: uminusSucc) +qed + +lemma times_pred_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[a] * b = a*b - b" + using b a by induct (simp add: add_right_commute)+ + +text \ Commutativity \ + +lemma times_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" + shows "a * b = b * a" +using assms by induct (simp add: uminusSucc)+ + +lemma times_uminus_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a) * b = -(a * b)" + using assms by (simp add: times_commute) + +text \ Distributivity \ + +lemma add_times_distrib_left [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b + c) = a * b + a * c" + using assms by induct (simp add: uminusSucc add_right_commute add_assoc)+ + +lemma add_times_distrib_right [algebra_simps]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(b + c) * a = b * a + c * a" + using assms by (simp add: times_commute add_times_distrib_left) + +\ \ NOT added as rewrites, since sometimes they are used from right to left \ +lemmas int_distrib = add_times_distrib_left add_times_distrib_right + +(* +text \ Identity element \ + +(no longer?) used for arithmetic simplification setup +lemma times_1_right: "a \ Int \ a * 1 = a" by simp +lemma times_1_left: "a \ Int \ 1 * a = a" by simp +lemma times_uminus_1_right: "a \ Int \ a * (-1) = -a" by simp +lemma times_uminus_1_left: "a \ Int \ (-1) * a = -a" by simp +*) + +text \ Associativity \ + +lemma times_assoc [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b * c) = (a * b) * c" + using assms by induct (auto simp add: add_times_distrib_right) + +lemma times_left_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b * c) = b * (a * c)" +using assms by (simp only: times_commute times_assoc) + +(* from HOL/OrderedGroups.thy *) +lemmas times_ac = times_assoc times_commute times_left_commute + +lemma times_right_commute: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * b * c = a * c * b" +proof - + from assms have "a * b * c = a * (b * c)" + by (rule sym[OF times_assoc]) + also from assms have "\ = a * (c * b)" + by (simp add: times_commute) + finally show ?thesis + using assms by (simp add: times_assoc) +qed + +text \ Reasoning about @{text "m * n = 0"}, etc. \ + +lemma times_is_0_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = 0) = (a = 0 \ b = 0)" +using a proof cases + case 0 + with b show ?thesis by simp +next + case (pos n) + with b show ?thesis + by cases auto +next + case na: (neg n) + with b show ?thesis + proof cases + case 0 + with na show ?thesis by simp + next + case (pos m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 0" + by (simp add: sym[OF uminusSucc]) + with na pos show ?thesis by simp + next + case (neg m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 0" + by (simp add: sym[OF uminusSucc]) + with na neg show ?thesis by simp + qed +qed + +lemma zero_istimes_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(0 = a * b) = (a = 0 \ b = 0)" + using assms by (auto dest: sym) + +lemma times_is_1_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = 1) = ((a = 1 \ b = 1) \ (a = -1 \ b = -1))" +using a proof induct + case 0 + with b show ?case by simp +next + case (pos n) + with b show ?case by cases (auto simp: add_is_0_nat) +next + case na: (neg n) + with b show ?case + proof cases + case 0 + with na show ?thesis by simp + next + case (pos m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 1" + by (simp add: sym[OF uminusSucc]) + with na pos show ?thesis by simp + next + case (neg m) + with na show ?thesis by (auto simp: zero_is_add_nat) + qed +qed + +lemma one_is_times_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(1 = a * b) = ((a = 1 \ b = 1) \ (a = -1 \ b = -1))" + using assms by (auto dest: sym) + +text \ Cancellation rules \ + +lemma times_cancel1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(a * b = a * c) = (b = c \ (a = 0))" +proof - + { + assume "a \ 0" "a*b = a*c" + from b c \a*b = a*c\ have "b = c" + proof (induct arbitrary: c) + case 0 + with a \a \ 0\ show ?case by simp + next + fix n c + assume n: "n \ Nat" and c: "c \ Int" and eq: "a * succ[n] = a * c" + and ih: "\d. \d \ Int; a*n = a *d\ \ n = d" + from a n have "a * n = (a * succ[n]) - a" by simp + also have "\ = a*c - a" using eq by simp + finally have "a * n = a * pred[c]" using a c by simp + with c ih have "n = pred[c]" by blast + with n c show "succ[n] = c" by simp + next + fix n c + assume n: "n \ Nat" and c: "c \ Int" and eq: "a * -succ[n] = a * c" + and ih: "\d. \d \ Int; a*(-n) = a *d\ \ -n = d" + from a n have "a * (-n) = (a * -succ[n]) + a" by simp + also have "\ = a * c + a" using eq by simp + finally have "a * (-n) = a * succ[c]" using a c by simp + with c ih have "-n = succ[c]" by blast + with n c show "-succ[n] = c" by (simp add: uminusSucc) + qed + } + with assms show ?thesis by auto +qed + +lemma times_cancel2 [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a * c = b * c) = (a = b \ c = 0)" +proof - + from assms have "a*c = c*a" "b*c = c*b" + by (simp add: times_commute)+ + with assms show ?thesis + by simp +qed + +lemma mult_eq_self1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = a) = (b = 1 \ a = 0)" (is "?lhs = ?rhs") +proof - + from assms have "?lhs = (a*b = a*1)" by simp + also from assms have "\ = ?rhs" by (intro times_cancel1) simp+ + finally show ?thesis . +qed + +lemma mult_eq_self2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(b * a = a) = (b = 1 \ a = 0)" + using assms by (simp add: times_commute) + + +subsection \Order on integer numbers\ + +text \ + We define the standard order @{text "\"} on integers using a conditional + definition, and we define the constant @{text "<"} such that @{text "ab \ a\b"}, for any values. +\ + +definition leq :: "[c,c] \ c" (infixl "\" 50) +where int_leq_def: "\a \ Int; b \ Int\ \ a \ b \ b - a \ Nat" + +abbreviation (input) + geq :: "[c,c] \ c" (infixl "\" 50) +where "x \ y \ y \ x" + +notation (ASCII) + leq (infixl "<=" 50) and + geq (infixl ">=" 50) + +lemma boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" +by (simp add: int_leq_def) + +lemma leq_isBool [intro!,simp]: "\a \ Int; b \ Int\ \ isBool(a \ b)" +by (simp add: int_leq_def) + +lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" +by (rule ssubst) + +lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" +by (rule subst) + +subsection \ Operator definitions and generic facts about @{text "<"} \ + +definition less :: "[c,c] \ c" (infixl "<" 50) +where "a < b \ a \ b \ a \ b" + +abbreviation (input) + greater :: "[c,c] \ c" (infixl ">" 50) +where "x > y \ y < x" + +lemma boolify_less [simp]: "boolify(a < b) = (a < b)" +by (simp add: less_def) + +lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" +by (rule ssubst) + +lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" +by (rule subst) + +lemma less_imp_leq [elim!]: "a < b \ a \ b" +unfolding less_def by simp + +lemma less_irrefl [simp]: "(a < a) = FALSE" +unfolding less_def by simp + +lemma less_irreflE [elim!]: "a < a \ R" +by simp + +lemma less_not_refl: "a < b \ a \ b" +by auto + +lemma neq_leq_trans [trans]: "a \ b \ a \ b \ a < b" +by (simp add: less_def) + +declare neq_leq_trans[simplified,trans] + +lemma leq_neq_trans [trans,elim!]: "a \ b \ a \ b \ a < b" +by (simp add: less_def) + +declare leq_neq_trans[simplified,trans] + +(* Don't add to [simp]: will be tried on all disequalities! *) +lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" +by auto + +subsection \ Facts about @{text "\"} over @{text Int} \ + +lemma int_leq_refl [intro,simp]: "a \ Int \ a \ a" + by (simp add: int_leq_def) + +lemma int_leq_antisym [elim]: + assumes ab: "a \ b" and ba: "b \ a" and a: "a \ Int" and b: "b \ Int" + shows "a = b" +proof - + from a b ab have "b - a \ Nat" + by (simp add: int_leq_def) + moreover + from a b ba have "-(b - a) \ Nat" + by (simp add: int_leq_def add_commute) + ultimately have "b - a = 0" + by simp + with a b show ?thesis + by simp +qed + +declare int_leq_antisym[THEN sym, elim] + +lemma int_leq_trans [trans]: + assumes "a \ b" and "b \ c" + and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a \ c" +proof - + from a b c have "c - a = (c - b) + (b - a)" + by (simp add: add_assoc) + with assms show ?thesis + by (simp add: int_leq_def) +qed + +lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ + assumes "a \ Int" and "b \ Int" and "a = b" and "\ a \ b; b \ a \ \ P" + shows "P" + using assms by simp + +lemma int_leq_linear: + assumes "a \ Int" and "b \ Int" + shows "a \ b \ b \ a" +proof - + from assms have "b - a \ Int" by simp + hence "b - a \ Nat \ -(b - a) \ Nat" by (auto elim: intElim) + with assms show ?thesis by (simp add: int_leq_def add_commute) +qed + +lemma int_leq_uminus [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a \ -b) = (b \ a)" + using assms by (simp add: int_leq_def add_commute) + +lemma uminus_leq_zero [simp]: + "a \ Int \ (-a \ 0) = (0 \ a)" + "a \ Int \ (0 \ -a) = (a \ 0)" + by (simp add: int_leq_def)+ + +lemma int_leq_succ [iff]: + "a \ Int \ a \ succ[a]" + by (simp add: int_leq_def) + +lemma int_pred_leq [iff]: + "a \ Int \ pred[a] \ a" + by (simp add: int_leq_def) + +lemma int_leq_iff_zero_leq_diff: + (* loops, not suitable as a rewrite rule *) + assumes "a \ Int" and "b \ Int" + shows "(a \ b) = (0 \ b-a)" + using assms by (simp add: int_leq_def) + +lemma int_leqI: + "\a \ Int; b \ Int; 0 \ b-a\ \ a \ b" + by (simp add: int_leq_def) + +lemma int_leqE: + "\a \ b; a \ Int; b \ Int; 0 \ b-a \ P\ \ P" + by (simp add: int_leq_def) + +text \ + From now on, we reduce the set @{text "Nat"} to the set of + non-negative integers. +\ +lemma nat_iff_int_geq0 [simp]: "n \ Nat = (n \ Int \ 0 \ n)" + by (auto simp: int_leq_def) + +declare natIsInt [simp del] +declare succInNat [simp del] +declare predInNat [simp del] + +lemma succ_geq_0 [simp]: + assumes "a \ Int" + shows "(0 \ succ[a]) = (0 \ a \ a = -1)" + using assms succInNat[OF assms] by auto + +lemma pred_geq_0 [simp]: + assumes "a \ Int" + shows "(0 \ pred[a]) = (0 \ a \ a \ 0)" + using assms predInNat[OF assms] by auto + +lemma succ_leq_0 [simp]: + assumes "a \ Int" + shows "(succ[a] \ 0) = (a \ 0 \ a \ 0)" +proof - + from assms have "(succ[a] \ 0) = (-0 \ -succ[a])" + by (intro sym[OF int_leq_uminus]) simp+ + with assms show ?thesis by (simp add: uminusSucc) +qed + +lemma pred_leq_0 [simp]: + assumes "a \ Int" + shows "(pred[a] \ 0) = (a \ 0 \ a = 1)" +proof - + from assms have "(pred[a] \ 0) = (-0 \ -pred[a])" + by (intro sym[OF int_leq_uminus]) simp+ + with assms show ?thesis by (simp add: uminusPred) +qed + +lemma nat_zero_leq: "n \ Nat \ 0 \ n" + by simp + +lemma nat_leq_zero [simp]: + "\n \ Int; 0 \ n\ \ (n \ 0) = (n = 0)" + by auto + +lemma uminus_leq_nat: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "-m \ n" + using assms by (simp add: int_leq_trans[of "-m" "0" "n"]) + +lemma int_succ_leq_succ [simp]: + "\a \ Int; b \ Int\ \ (succ[a] \ succ[b]) = (a \ b)" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "succ[a]" "succ[b]"] + by simp + +lemma int_succ_leq_succI [intro!]: + "\a \ Int; b \ Int; a \ b\ \ succ[a] \ succ[b]" + by simp + +lemma int_succ_leq_succE [elim!]: + "\succ[a] \ succ[b]; a \ Int; b \ Int; a \ b \ P\ \ P" + by simp + +lemma int_pred_leq_pred [simp]: + "\a \ Int; b \ Int\ \ (pred[a] \ pred[b]) = (a \ b)" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "pred[a]" "pred[b]"] + by simp + +lemma int_pred_leq_predI [intro!]: + "\a \ Int; b \ Int; a \ b\ \ pred[a] \ pred[b]" + by simp + +lemma int_succ_leq_predE [elim!]: + "\pred[a] \ pred[b]; a \ Int; b \ Int; a \ b \ P\ \ P" + by simp + +lemma int_leq_succI [intro]: + "\a \ b; a \ Int; b \ Int\ \ a \ succ[b]" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "a" "succ[b]"] + by simp + +lemma int_pred_leqI [intro]: + "\a \ b; a \ Int; b \ Int\ \ pred[a] \ b" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "pred[a]" "b"] + by simp + +lemma int_pred_leq_is_leq_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] \ b) = (a \ succ[b])" +proof - + from assms have "(pred[a] \ b) = (succ[pred[a]] \ succ[b])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with assms show ?thesis by simp +qed + +lemma int_leq_pred_is_succ_leq [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ pred[b]) = (succ[a] \ b)" +proof - + from assms have "(a \ pred[b]) = (succ[a] \ succ[pred[b]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with assms show ?thesis by simp +qed + +lemma int_leq_succ_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a \ succ[b]) = (a \ b \ a = succ[b])" +proof - + { + assume "b - a = -1" + with assms have "b = -1 + a" + by (simp add: add_left_diff_right) + with assms have "a = succ[b]" + by simp + } + with assms int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "a" "succ[b]"] + show ?thesis by auto +qed + +lemma int_pred_leq_iff: + assumes a: "a \ Int" and b: "b \ Int" + shows "(pred[a] \ b) = (a \ b \ a = succ[b])" + using assms by simp + +lemma int_leq_succE [elim!]: + assumes "a \ succ[b]" and "a \ Int" and "b \ Int" + and "a \ b \ P" and "a = succ[b] \ P" + shows "P" + using assms by auto + +lemma int_pred_leqE [elim!]: + assumes "pred[a] \ b" and "a \ Int" and "b \ Int" + and "a \ b \ P" and "a = succ[b] \ P" + shows "P" + using assms by auto + +lemma int_succ_leq_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a] \ b) = (a < b)" +proof - + { + assume "succ[a] \ b" + with assms have "a \ b" + by (simp add: int_leq_trans[of "a" "succ[a]" "b"]) + moreover + have "a \ b" + proof + assume "a = b" + with assms int_leq_succ[of "a"] \succ[a] \ b\ have "succ[a] = a" + by (intro int_leq_antisym) simp+ + with assms show "FALSE" + by simp + qed + ultimately have "a < b" + by (simp add: less_def) + } + moreover + { + assume "a < b" + have "succ[a] \ b" + proof (rule int_leqI) + from assms have "b - a \ Int" by simp + thus "0 \ b - succ[a]" + proof (cases) + case 0 + with assms \a < b\ show ?thesis + by (simp add: less_def) + next + case (pos n) + with assms show ?thesis + by simp + next + case (neg n) + with assms \a < b\ show ?thesis + by (simp add: less_def int_leq_iff_zero_leq_diff[of "a" "b"]) + qed + qed (simp add: assms)+ + } + ultimately show ?thesis + using assms by (intro boolEqual) auto +qed + +lemma int_leq_pred_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ pred[b]) = (a < b)" + using assms by simp + +lemma int_leq_uminus_succ_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ -succ[b]) = (a < -b)" + using assms by (simp add: uminusSucc) + +lemma int_uminus_pred_leq_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-pred[a] \ b) = (-a < b)" + using assms by (simp add: uminusPred) + +lemma int_succ_leqI [intro]: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "succ[a] \ b" + using assms by simp + +lemma int_succ_leqE [elim!]: + assumes "succ[a] \ b" and "a \ Int" and "b \ Int" + and "\a < b\ \ P" + shows "P" + using assms by simp + +lemma int_leq_predI [intro]: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "a \ pred[b]" + using assms by simp + +lemma int_leq_predE [elim!]: + assumes "a \ pred[b]" and "a \ Int" and "b \ Int" + and "a < b \ P" + shows "P" + using assms by simp + +lemma int_not_succ_leq [simp]: + "a \ Int \ (succ[a] \ a) = FALSE" + by auto + +lemma int_not_leq_pred [simp]: + "a \ Int \ (a \ pred[a]) = FALSE" + by auto + +lemma int_leq_limit_1: + assumes "a \ b" and "\(succ[a] \ b)" and "a \ Int" and "b \ Int" + shows "a=b" + using assms by (auto simp: less_def) + +lemma int_leq_limit_2: + assumes "a \ b" and "\(a \ pred[b])" and "a \ Int" and "b \ Int" + shows "a=b" + using assms by (auto elim: int_leq_limit_1) + +text \ + Case distinction and induction rules involving @{text "\"}. +\ + +lemma int_leq_cases: + assumes a: "a \ Int" and b: "b \ Int" + and "a \ b \ P" and "\b \ a; b \ a\ \ P" + shows "P" + using assms int_leq_linear[OF a b] by auto + +lemma nat_leq_induct: \ \sometimes called ``complete induction''\ + assumes base: "P(0)" + and step: "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(succ[n])" + shows "\n\Nat : P(n)" +proof - + have "\n\Nat : \m\Nat : m \ n \ P(m)" + proof (rule natInduct) + from base show "\m \ Nat : m \ 0 \ P(m)" by auto + next + fix n + assume "n \ Nat" "\m \ Nat : m \ n \ P(m)" + with step show "\m \ Nat : m \ succ[n] \ P(m)" by auto + qed + thus ?thesis by auto +qed + +lemma nat_leq_inductE: + assumes "n \ Nat" + and "P(0)" and "\n. \n \ Nat; \m\Nat : m \ n \ P(m)\ \ P(succ[n])" + shows "P(n)" + using assms by (blast dest: nat_leq_induct) + +lemma int_leq_induct: \ \analogous induction principle for the integers\ + assumes base: "P(0)" + and step: "\n\Nat : (\i\Int : -n \ i \ i \ n \ P(i)) \ P(succ[n]) \ P(-succ[n])" + shows "\a\Int : P(a)" +proof - + have "\n\Nat : \a\Int : -n \ a \ a \ n \ P(a)" + proof (rule natInduct) + from base show "\a \ Int : -0 \ a \ a \ 0 \ P(a)" + by (auto dest: int_leq_antisym) + next + fix n + assume n: "n \ Nat" and ih: "\a \ Int : -n \ a \ a \ n \ P(a)" + show "\a \ Int : -succ[n] \ a \ a \ succ[n] \ P(a)" + proof (clarify) + fix a + assume m: "a \ Int" "-succ[n] \ a" "a \ succ[n]" + from n ih step have "P(succ[n]) \ P(-succ[n])" by blast + moreover { + assume "-n \ a" "a \ n" + with ih \a \ Int\ have "P(a)" by blast + } + ultimately show "P(a)" + using n m by (auto simp: uminusSucc) + qed + qed + moreover + { + fix n + assume "n \ Nat" + hence "-n \ Int \ n \ Int \ -n \ -n \ -n \ n \ n \ n" + by (simp add: int_leq_trans[of "-n" "0" "n"]) + } + ultimately have "\n \ Nat : P(n) \ P(-n)" by blast + thus ?thesis by (blast elim: intElim) +qed + +lemma int_leq_inductE: + assumes "a \ Int" + and "P(0)" + and "\n. \n \ Nat; \i\Int : -n \ i \ i \ n \ P(i)\ \ P(succ[n])" + and "\n. \n \ Nat; \i\Int : -n \ i \ i \ n \ P(i)\ \ P(-succ[n])" + shows "P(a)" + using assms by (blast dest: int_leq_induct) + + +subsection \ Facts about @{text "<"} over @{text Int} \ + +text \ Reduce @{text "\"} to @{text "<"}. \ + +lemma int_leq_less: + assumes "a \ Int" and "b \ Int" + shows "a \ b = (a < b \ a = b)" + using assms by (auto simp: less_def) + +lemma int_leq_lessE: + assumes "a \ b" and "a \ Int" and "b \ Int" + and "a < b \ P" and "a=b \ P" + shows "P" + using assms by (auto simp: int_leq_less) + +lemma int_less_iff_zero_less_diff: + assumes "a \ Int" and "b \ Int" + shows "(a < b) = (0 < b-a)" + using assms unfolding less_def + by (auto simp: int_leq_iff_zero_leq_diff[of "a" "b"]) + +lemma int_lessI: + "\a \ Int; b \ Int; 0 < b-a\ \ a < b" + by (simp add: int_less_iff_zero_less_diff[of "a" "b"]) + +lemma int_lessE: + "\a < b; a \ Int; b \ Int; 0 < b-a \ P\ \ P" + by (simp add: int_less_iff_zero_less_diff[of "a" "b"]) + +lemma int_leq_not_less: + assumes "a \ Int" and "b \ Int" and "a \ b" + shows "(b < a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_not_leq: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "(b \ a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +text \ + Declaring the above lemmas as simplification rules makes the prover + very slow, but the instances where @{text "a"} is @{text "0"} are useful. +\ +lemmas zero_leq_not_less_zero [simp] = int_leq_not_less[of 0, simplified] +lemmas zero_less_not_leq_zero [simp] = int_less_not_leq[of 0, simplified] + +lemma int_less_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a < succ[b]) = (a \ b)" + using assms by (auto simp: less_def int_leq_succI) + +lemmas int_leq_is_less_succ = sym[OF int_less_succ] + +lemma int_less_succE: + assumes "a < succ[b]" and "a \ Int" and "b \ Int" + and "a < b \ P" and "a = b \ P" + shows P + using assms by (auto elim: int_leq_lessE) + +lemma int_pred_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] < b) = (a \ b)" + using assms by (auto simp: less_def) + +lemmas int_leq_is_pred_less = sym[OF int_pred_less] + +lemma not_leq_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "((a \ b) = FALSE) = (b < a)" + using assms int_leq_linear[of a b] + by (auto simp: less_def dest: int_leq_antisym) + +lemma not_less_leq [simp]: + assumes "a \ Int" and "b \ Int" + shows "((a < b) = FALSE) = (b \ a)" + using assms int_leq_linear[of a b] + by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_uminus [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a < -b) = (b < a)" + using assms by (auto simp: less_def) + +lemma int_uminus_zero [simp]: + "a \ Int \ (-a < 0) = (0 < a)" + "a \ Int \ (0 < -a) = (a < 0)" + by (auto simp: less_def) + +text \ @{text "<"} and @{text "succ / pred"}. \ + +lemma int_succ_less_succ [simp]: + "\a \ Int; b \ Int\ \ (succ[a] < succ[b]) = (a < b)" + by (auto simp: less_def int_leq_succI) + +lemma int_succ_less_succI [intro]: + "\a \ Int; b \ Int; a < b\ \ succ[a] < succ[b]" + by simp + +lemma int_succ_less_succE [elim]: + assumes "succ[a] < succ[b]" and "a \ Int" and "b \ Int" and "a < b \ P" + shows "P" + using assms by simp + +lemma int_pred_less_pred [simp]: + "\a \ Int; b \ Int\ \ (pred[a] < pred[b]) = (a < b)" + by (simp add: less_def) + +lemma int_pred_less_predI [intro]: + "\a \ Int; b \ Int; a < b\ \ pred[a] < pred[b]" + by simp + +lemma int_pred_less_predE [elim]: + assumes "pred[a] < pred[b]" and "a \ Int" and "b \ Int" and "a < b \ P" + shows "P" + using assms by simp + +lemma int_less_succ_same [iff]: + "a \ Int \ a < succ[a]" + by (simp add: less_def) + +lemma int_pred_less_same [iff]: + "a \ Int \ pred[a] < a" + by (simp add: less_def) + +lemma int_pred_less_is_less_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] < b) = (a < succ[b])" + using assms by (auto simp: less_def) + +lemma int_less_pred_is_succ_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a < pred[b]) = (succ[a] < b)" + using assms by (auto simp: less_def) + +lemma succ_nat_not_neg [simp]: + "\n \ Int; 0 \ n\ \ (succ[n] < 0) = FALSE" + by (auto simp: less_def) + +lemma int_less_leq_not_leq: + assumes "a \ Int" and "b \ Int" + shows "(a < b) = (a \ b \ \ b \ a)" + using assms by (auto simp: less_def dest: int_leq_antisym) + +text \ Transitivity. \ + +lemma int_less_trans [trans]: + assumes "a < b" and "b < c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma nat_less_trans_Succ: \ \stronger than the above\ + assumes lt1: "a < b" and lt2: "b < c" + and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "succ[a] < c" + using assms +proof - + from a b lt1 have "succ[succ[a]] \ succ[b]" + by (auto simp: less_def int_leq_succI) + also from b c lt2 have "\ \ c" + by (auto simp: less_def) + finally show ?thesis + using a b c by (auto simp: less_def) +qed + +lemma int_leq_less_trans [trans]: + assumes "a \ b" and "b < c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma int_less_leq_trans [trans]: + assumes "a < b" and "b \ c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma nat_zero_less_succ [iff]: + "\n \ Int; 0 \ n\ \ 0 < succ[n]" + by (simp add: int_leq_less_trans) + +text \ Asymmetry. \ + +lemma int_less_not_sym: + assumes "a < b" and "a \ Int" and "b \ Int" + shows "(b < a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_asym: + assumes 1: "a < b" and 2: "a \ Int" and 3: "b \ Int" and 4: "\P \ b < a" + shows "P" + using int_less_not_sym[OF 1 2 3] 4 by auto + +text \ Linearity (totality). \ + +lemma int_less_linear: + assumes a: "a \ Int" and b: "b \ Int" + shows "a < b \ a = b \ b < a" + unfolding less_def using int_leq_linear[OF a b] by blast + +lemma int_leq_less_linear: + assumes a: "a \ Int" and b: "b \ Int" + shows "a \ b \ b < a" + using assms int_less_linear[OF a b] by (auto simp: less_def) + +lemma int_less_cases [case_names less equal greater]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a < b \ P) \ (a = b \ P) \ (b < a \ P) \ P" + using int_less_linear[OF a b] by blast + +lemma int_neq_iff: + assumes a: "a \ Int" and b: "b \ Int" + shows "((a = b) = FALSE) = (a < b \ b < a)" + using assms int_less_linear[OF a b] by auto + +lemma int_neq_lessE: + assumes "a \ b" and "a \ Int" and "b \ Int" + shows "(a < b \ P) \ (b < a \ P) \ P" + using assms by (auto simp: int_neq_iff) + +lemma int_antisym_conv1: + assumes "\ a < b" and "a \ Int" and "b \ Int" + shows "(a \ b) = (a = b)" + using assms by (auto dest: int_leq_antisym) + +lemma int_antisym_conv2: + assumes "a \ b" and "a \ Int" and "b \ Int" + shows "(\ a < b) = (b = a)" + using assms by (auto dest: int_leq_antisym) + +lemma int_antisym_conv3: + assumes "\ a < b" and "a \ Int" and "b \ Int" + shows "(\ b < a) = (a = b)" + using assms by (auto dest: int_leq_antisym) + +lemma nat_gt0_not0: + assumes "n \ Int" and "0 \ n" + shows "(0 < n) = (n \ 0)" + using assms by auto + +lemmas nat_neq0_conv = sym[OF nat_gt0_not0, simplified, simp] + +lemma nat_less_one (*[simp]*): + assumes "n \ Int" and "0 \ n" + shows "(n < 1) = (n = 0)" + using assms by auto + +text \ "Less than" is antisymmetric, sort of \ + +lemma int_less_antisym: + assumes "a \ Int" and "b \ Int" + shows "\ \(b < a); b < succ[a] \ \ a = b" + using assms by (auto dest: int_leq_antisym) + +text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ +lemma int_less_mono_imp_leq_mono: + assumes ij: "i \ j" and i: "i \ Int" and j: "j \ Int" + and f: "\a \ Int : f(a) \ Int" + and mono: "\i j. \i \ Int; j \ Int; i \ f(i) < f(j)" + shows "f(i) \ f(j)" +proof (rule int_leq_lessE[OF ij i j]) + assume "i < j" + with i j have "f(i) < f(j)" by (rule mono) + thus ?thesis by auto +next + assume "i = j" + with i f show ?thesis by simp +qed + +lemma nat_less_mono_imp_leq_mono: + assumes ij: "i \ j" and i: "i \ Nat" and j: "j \ Nat" + and f: "\a \ Nat : f(a) \ Int" + and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" + shows "f(i) \ f(j)" +proof (rule int_leq_lessE[OF ij]) + assume "i < j" + with i j have "f(i) < f(j)" by (rule mono) + thus ?thesis by auto +next + assume "i = j" + with i f show ?thesis by simp +qed (auto simp: i[simplified] j[simplified]) + +lemma int_succ_lessI: + assumes "a \ Int" and "b \ Int" and "a < b" and "succ[a] \ b" + shows "succ[a] < b" + using assms by (auto simp: less_def) + +lemma int_less_antisym_false: "\a < b; a \ Int; b \ Int\ \ b < a = FALSE" + by auto + +lemma nat_less_antisym_leq_false: "\a < b; a \ Int; b \ Int\ \ b \ a = FALSE" + by auto + + +subsection \ Additional arithmetic theorems \ + +subsubsection \ Monotonicity of Addition \ + +lemma int_add_left_cancel_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(i + j \ i + k) = (j \ k)" + using assms by (induct i) (simp del: int_leq_succ_iff)+ + +lemma int_add_right_cancel_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(j + i \ k + i) = (j \ k)" + using assms by (induct i) (simp del: int_leq_succ_iff)+ + +lemma int_add_left_cancel_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(i + j < i + k) = (j < k)" + using assms by (simp add: less_def) + +lemma int_add_right_cancel_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(j + i < k + i) = (j < k)" + using assms by (simp add: less_def) + +lemma int_add_left_cancel_succ_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[i + j] \ i + k) = (succ[j] \ k)" +using assms proof (induct i) + case (neg n) + then have "(-n + j \ pred[-n + k]) = (succ[-n+j] \ succ[pred[-n + k]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with neg show ?case by simp +qed (simp del: int_leq_succ_iff)+ + +lemma int_add_right_cancel_succ_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[j + i] \ k + i) = (succ[j] \ k)" +using assms proof (induct i) + case (neg n) + then have "(j - n \ pred[k - n]) = (succ[j-n] \ succ[pred[k-n]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with neg show ?case by simp +qed (simp del: int_leq_succ_iff)+ + +lemma int_add_left_cancel_pred_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(pred[i + j] \ i + k) = (pred[j] \ k)" +proof - + from assms have "pred[i+j] = i + pred[j]" by simp + hence "(pred[i + j] \ i + k) = (i + pred[j] \ i + k)" by simp + also have "\ = (pred[j] \ k)" + using assms by (intro int_add_left_cancel_leq) simp+ + finally show ?thesis . +qed + +lemma int_add_right_cancel_pred_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(pred[j + i] \ k + i) = (pred[j] \ k)" +proof - + from assms have "pred[j+i] = pred[j] + i" by simp + hence "(pred[j+i] \ k + i) = (pred[j] + i \ k + i)" by simp + also have "\ = (pred[j] \ k)" + using assms by (intro int_add_right_cancel_leq) simp+ + finally show ?thesis . +qed + +lemma int_leq_imp_add: + assumes i: "i \ Int" and j: "j \ Int" and ij: "i \ j" + shows "\k \ Nat : j = i+k" +proof - + from assms have "j-i \ Nat" + by (simp add: int_leq_iff_zero_leq_diff[of "i" "j"]) + moreover + from i j have "j = i + (j-i)" + by (simp add: add_ac) + ultimately show ?thesis .. +qed + +lemma int_less_imp_succ_add: + assumes ij: "i < j" and i: "i \ Int" and j: "j \ Int" + shows "\k \ Nat: j = succ[i + k]" +proof - + from assms have "succ[i] \ j" + by simp + with i j have "j - succ[i] \ Nat" + by (simp add: int_less_iff_zero_less_diff[of "i" "j"]) + moreover + from i j have "j = succ[i + (j - succ[i])]" + by (simp add: add_ac) + ultimately show ?thesis .. +qed + + +subsubsection \ (Partially) Ordered Groups \ + +\ \ The two following lemmas are just ``one half'' of + @{text int_add_left_cancel_leq} and @{text int_add_right_cancel_leq} + proved above." \ +lemma add_leq_left_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a \ b \ c + a \ c + b" + using assms by simp + +lemma add_leq_right_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a \ b \ a + c \ b + c" + using assms by simp + +text \ non-strict, in both arguments \ +lemma add_leq_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a \ b \ c \ d \ a + c \ b + d" + using assms + add_leq_right_mono[of a b c] + add_leq_left_mono[of c d b] + int_leq_trans[of "a + c" "b + c" "b + d"] + by simp + +\ \ Similar for strict less than. \ +lemma add_less_left_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a < b \ c + a < c + b" + using assms by (simp add: add_leq_left_mono less_def) + +lemma add_less_right_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a < b \ a + c < b + c" + using assms by (simp add: add_leq_right_mono less_def) + +text \ Strict monotonicity in both arguments \ +lemma add_less_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a < b \ c < d \ a + c < b + d" + using assms + add_less_right_mono[of a b c] + add_less_left_mono[of c d b] + int_less_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma add_less_leq_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a < b \ c \ d \ a + c < b + d" + using assms + add_less_right_mono[of a b c] + add_leq_left_mono[of c d b] + int_less_leq_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma add_leq_diff1: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b \ c) = (a \ c - b)" +proof - + from assms have "(a + b \ c) = ((a + b) - b \ c - b)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis by simp +qed + +lemma add_leq_diff2: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b \ c) = (b \ c - a)" +proof - + from assms have "(a + b \ c) = ((a + b) - a \ c - a)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis by simp +qed + +lemma add_leq_less_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a \ b \ c < d \ a + c < b + d" + using assms + add_leq_right_mono[of a b c] + add_less_left_mono[of c d b] + int_leq_less_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma leq_add_nat1 [iff]: + assumes "a \ Int" and "0 \ a" and "b \ Int" + shows "b \ b + a" + using assms add_leq_left_mono[of 0 a b] by simp + +lemma leq_add_nat2 [iff]: + assumes "a \ Int" and "0 \ a" and "b \ Int" + shows "b \ a + b" + using assms add_leq_right_mono [of 0 a b] by simp + +lemma trans_leq_add_nat1 [simp] (* was [elim]*): + assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" + shows "a \ b+m" + using assms by (simp add: int_leq_trans[of "a" "b" "b+m"]) + +lemma trans_leq_add_nat2 [simp] (* was [elim] *): + assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" + shows "a \ m+b" + using assms by (simp add: int_leq_trans[of "a" "b" "m+b"]) + +lemma add_nat_leqD1 [elim]: + assumes "a+k \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "a \ b" + using assms by (simp add: int_leq_trans[of "a" "a+k" "b"]) + +lemma add_nat_leqD2 [elim]: + assumes "k+a \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "a \ b" + using assms by (simp add: int_leq_trans[of "a" "k+a" "b"]) + +lemma add_nat_pred_leqD1 [elim]: + assumes "pred[a+k] \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "pred[a] \ b" +proof - + { + assume 1: "a + k = succ[b]" + have "a \ succ[b]" + proof (rule int_leqI) + from assms 1 have "succ[b] - a = k" + by (simp add: add_left_diff_right add_ac) + with \0 \ k\ show "0 \ succ[b] - a" + by simp + qed (simp add: assms)+ + } + with assms show ?thesis + by (auto simp del: int_leq_succ_iff) +qed + +lemma add_nat_pred_leqD2 [elim]: + assumes "pred[k+a] \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "pred[a] \ b" +proof (rule add_nat_pred_leqD1) + from assms show "pred[a+k] \ b" + by (simp add: add_commute) +qed (simp add: assms)+ + +lemma less_succ_add_nat1 (*[simp]*): + assumes "i \ Int" and "m \ Int" and "0 \ m" + shows "i < succ[i + m]" + using assms by simp + +lemma less_add_Succ2 (*[simp]*): + assumes "i \ Int" and "m \ Int" and "0 \ m" + shows "i < succ[m + i]" + using assms by simp + +lemma diff_nat_leq_self1 [simp]: + assumes "i \ Int" and "n \ Int" and "0 \ n" + shows "i - n \ i" +proof - + from assms have "i-n \ (i-n)+n" + by (intro leq_add_nat1) simp+ + with assms show ?thesis by simp +qed + +lemma diff_nat_leq_self2 [simp]: + assumes "i \ Int" and "n \ Int" and "0 \ n" + shows "-n + i \ i" + using assms by (simp add: add_commute) + +lemma trans_leq_diff_nat1 [simp]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i-m \ j" + using assms by (simp add: int_leq_trans[of "i-m" "i" "j"]) + +lemma trans_leq_diff_nat2 [simp]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "-m + i \ j" + using assms by (simp add: int_leq_trans[of "-m + i" "i" "j"]) + +lemma int_leq_iff_add: + assumes "i \ Int" and "j \ Int" + shows "(i \ j) = (\k \ Nat: j = i + k)" (is "?lhs = ?rhs") + using assms by (auto intro: int_leq_imp_add) + +lemma int_less_iff_succ_add: + assumes "i \ Int" and "j \ Int" + shows "(i < j) = (\k \ Nat: j = succ[i + k])" (is "?lhs = ?rhs") + using assms by (auto intro: int_less_imp_succ_add) + +lemma leq_add_self1 [simp]: + assumes "i \ Int" and "j \ Int" + shows "(i + j \ i) = (j \ 0)" +proof - + from assms have "(i+j \ i) = (i+j-i \ i-i)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis + by simp +qed + +lemma leq_add_self2 [simp]: + assumes "i \ Int" and "j \ Int" + shows "(j + i \ i) = (j \ 0)" + using assms by (simp add: add_commute) + +lemma trans_less_add_nat1 [simp]: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i < j + m" + using assms by (simp add: int_less_leq_trans[of "i" "j" "j+m"]) + +lemma trans_less_add_nat2 [simp]: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i < m + j" + using assms by (simp add: int_less_leq_trans[of "i" "j" "m+j"]) + +lemma trans_less_diff_nat1: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i - m < j" + using assms by (simp add: int_leq_less_trans[of "i-m" "i" "j"]) + +lemma add_nat_lessD1: + assumes "i+k < j" and "i \ Int" and "j \ Int" and "k \ Int" and "0 \ k" + shows "i < j" + using assms by (simp add: int_leq_less_trans[of "i" "i+k" "j"]) + +lemma add_lessD2: + assumes "k+i < j" and "i \ Int" and "j \ Int" and "k \ Int" and "0 \ k" + shows "i < j" + using assms by (simp add: int_leq_less_trans[of "i" "k+i" "j"]) + +lemma nat_add_gt_0: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "(m + n > 0) = (m > 0 \ n > 0)" + using assms by auto + +lemma add_ge_1: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "(m + n \ 1) = (m \ 1 \ n \ 1)" + using assms by auto + +lemma less_imp_add_positive: + assumes "i < j" and "i \ Int" and "j \ Int" + shows "\k \ Nat: 0 < k \ i + k = j" +proof - + from assms obtain k where k: "k \ Nat" "i + succ[k] = j" + by (auto dest: int_less_imp_succ_add) + from k have "0 < succ[k]" "succ[k] \ Nat" by simp+ + with assms k show ?thesis by blast +qed + + +subsubsection \ Monotonicity of Multiplication \ + +lemma times_leq_left_pos_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "c * a \ c * b" + using c by induct (simp add: add_leq_mono 1 a b)+ + +lemma times_leq_left_neg_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "(-c) * b \ (-c) * a" +using c proof (induct) + case 0 + with a b show ?case by simp +next + case (Succ n) + from a b 1 have "-b \ -a" by simp + with a b Succ show ?case by (simp add: add_leq_mono) +qed + +lemma times_leq_right_pos_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "a * c \ b * c" + using c by induct (simp add: add_leq_mono 1 a b)+ + +lemma times_leq_right_neg_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "b * (-c) \ a * (-c)" +using c proof (induct) + case 0 + with a b show ?case by simp +next + case (Succ n) + from a b 1 have "-b \ -a" by simp + with a b Succ show ?case by (simp add: add_leq_mono) +qed + +lemma nat_leq_times_left: + assumes a: "a \ Nat" and b: "b \ Int" and 1: "1 \ b" + shows "a \ a * b" +proof - + have "1 \ Int" by simp + from 1 this b a have "a * 1 \ a * b" by (rule times_leq_left_pos_mono) + with a b show ?thesis by simp +qed + +lemma nat_leq_times_right: + assumes "a \ Nat" and "b \ Int" and "1 \ b" + shows "a \ b * a" + using assms by (simp add: nat_leq_times_left times_commute) + +text \ Similar lemmas for @{text "<"} \ + +lemma times_less_left_pos_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "c * a < c * b" +proof - + from c 0 obtain n where "n \ Nat" "c = succ[n]" + by cases auto + from \n \ Nat\ have "succ[n] * a < succ[n] * b" + by induct (simp add: add_less_mono 1 a b)+ + with \c = succ[n]\ show ?thesis by simp +qed + +lemma times_less_left_neg_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "c * b < c * a" +proof - + from c 0 obtain n where "n \ Nat" "c = -succ[n]" + by cases auto + from a b 1 have "-b < -a" by simp + with \n \ Nat\ have "(-succ[n]) * b < (-succ[n]) * a" + by (induct) (simp add: add_less_mono a b)+ + with \c = -succ[n]\ show ?thesis by simp +qed + +lemma times_less_right_pos_mono: + assumes "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and "0 < c" + shows "a * c < b * c" + using times_less_left_pos_mono[OF assms] a b c by (simp add: times_commute) + +lemma times_less_right_neg_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and "c < 0" + shows "b * c < a * c" + using times_less_left_neg_mono[OF assms] a b c by (simp add: times_commute) + +lemma nat_less_times_left: + assumes a: "a \ Int" and 0: "0 < a" and b: "b \ Int" and 1: "1 < b" + shows "a < a * b" +proof - + have "1 \ Int" by simp + from 1 this b a 0 have "a * 1 < a * b" by (rule times_less_left_pos_mono) + with a b show ?thesis by simp +qed + +lemma nat_less_times_right: + assumes "a \ Int" and 0: "0 < a" and "b \ Int" and "1 < b" + shows "a < b * a" + using assms by (simp add: nat_less_times_left times_commute) + +lemma times_leq_left_pos_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "(c * a \ c * b) = (a \ b)" (is "?lhs = ?rhs") +proof (rule boolEqual) + { + assume lhs: "?lhs" and contr: "\?rhs" + from a b contr have "b < a" by simp + from this b a c 0 have "c * b < c * a" by (rule times_less_left_pos_mono) + with a b c have "\?lhs" by simp + from this lhs have "FALSE" .. + } + moreover + from c 0 have "c \ Nat" + by (simp add: less_def) + ultimately show "?lhs \ ?rhs" + using times_leq_left_pos_mono[of a b c] assms by blast +qed (simp add: a b c)+ + +lemma times_leq_right_pos_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "(a * c \ b * c) = (a \ b)" (is "?lhs = ?rhs") +proof - + from a b c have "(a * c \ b * c) = (c * a \ c * b)" + by (simp add: times_commute) + with assms show ?thesis + by simp +qed + +lemma times_leq_left_neg_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "(c * a \ c * b) = (b \ a)" (is "?lhs = ?rhs") +proof (rule boolEqual) + { + assume lhs: "?lhs" and contr: "\?rhs" + from a b contr have "a < b" by simp + from this a b c 0 have "c * b < c * a" by (rule times_less_left_neg_mono) + with a b c have "\?lhs" by simp + from this lhs have "FALSE" .. + } + moreover + from c 0 obtain d where "d \ Nat" "c = -d" + by (cases c) force+ + ultimately show "?lhs \ ?rhs" + using times_leq_left_neg_mono[of b a d] assms by blast +qed (simp add: a b c)+ + +lemma times_leq_right_neg_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "(a * c \ b * c) = (b \ a)" (is "?lhs = ?rhs") +proof - + from a b c have "(a * c \ b * c) = (c * a \ c * b)" + by (simp add: times_commute) + with assms show ?thesis + by simp +qed + +lemma times_less_left_pos_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(c * a < c * b) = (a < b)" + unfolding less_def using assms by auto + +lemma times_less_left_neg_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "c < 0" + shows "(c * a < c * b) = (b < a)" + unfolding less_def using assms by auto + +lemma times_less_right_pos_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * c < b * c) = (a < b)" + unfolding less_def using assms by auto + +lemma times_less_right_neg_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "c < 0" + shows "(a * c < b * c) = (b < a)" + unfolding less_def using assms by auto + +lemma times_pos [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(0 < a * b) = ((0 < a \ 0 < b) \ (a < 0 \ b < 0))" +using a proof (cases) + case 0 + with b show ?thesis by simp +next + case (pos n) + hence a': "a \ Nat" "0 < a" + by auto + from a have "(0 < a * b) = (a * 0 < a * b)" + by simp + also from a \0 < a\ b have "\ = (0 < b)" + by (intro times_less_left_pos_cancel) simp+ + finally show ?thesis + using a' by simp +next + case (neg n) + hence a': "a \ Int" "a < 0" + by auto + from a have "(0 < a * b) = (a * 0 < a * b)" + by simp + also from a' b have "\ = (b < 0)" + by (intro times_less_left_neg_cancel) simp+ + finally show ?thesis + using a' by (simp add: int_less_not_sym[of a 0]) +qed + +lemma times_neg: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b < 0) = ((0 < a \ b < 0) \ (a < 0 \ 0 < b))" +proof - + from assms have "(a * b < 0) = (0 < (-a) * b)" + by simp + also from assms have "\ = ((0 < -a \ 0 < b) \ (-a < 0 \ b < 0))" + by (intro times_pos) simp+ + finally show ?thesis + using assms by auto +qed + + +subsection \Intervals of integers\ + +definition intvl :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) +where "i .. j \ { k \ Int : i \ k \ k \ j }" + +lemma intvlI [iff]: + assumes "k \ Int" and "i \ k" and "k \ j" + shows "k \ i .. j" + using assms by (simp add: intvl_def) + +lemma intvlE [elim]: + assumes "k \ i .. j" and "\k \ Int; i \ k; k \ j\ \ P" + shows P + using assms by (simp add: intvl_def) + +lemma intvlIff [simp]: "(k \ i .. j) = (k \ Int \ i \ k \ k \ j)" + by auto + +lemmas + setEqualI [where A = "i .. j" for i j, intro] + setEqualI [where B = "i .. j" for i j, intro] + +lemma gtNotinIntvl: + assumes gt: "i > j" and k: "k \ i .. j" and i: "i \ Int" and j: "j \ Int" + shows "P" +proof - + from k have "k \ Int" "i \ k" "k \ j" by auto + with i j have "i \ j" by (auto elim: int_leq_trans) + with i j have "\(j < i)" by simp + from this gt show ?thesis .. +qed + +lemma intvlIsEmpty: + assumes "i \ Int" and "j \ Int" and "i > j" + shows "i .. j = {}" + using assms by (blast dest: gtNotinIntvl) + +lemma intvlEmpty_iff: + assumes i: "i \ Int" and j: "j \ Int" + shows "(i .. j = {}) = (i > j)" +proof - + { + assume mt: "i .. j = {}" + have "i > j" + proof (rule contradiction) + assume "\(i > j)" + with assms have "i \ j" by simp + with i have "i \ i ..j" by simp + with mt show "FALSE" by simp + qed + } + moreover + from i j have "i > j \ i..j = {}" + by (blast dest: intvlIsEmpty) + ultimately show ?thesis + by blast +qed + +lemma intvlSingleton [simp]: + "i \ Int \ i .. i = {i}" + by (auto dest: int_leq_antisym) + +lemma intvlSucc [simp]: + assumes "i \ Int" and "j \ Int" and "i \ succ[j]" + shows "i .. succ[j] = (i..j \ {succ[j]})" + using assms by auto + +lemma succIntvl: + assumes "i \ Int" and "j \ Int" + shows "succ[i] .. j = (i .. j \ {i})" + using assms unfolding intvl_def by auto + +lemma intvlEqual_iff: + assumes i: "i \ Int" and j: "j \ Int" and k: "k \ Int" and l: "l \ Int" + shows "(i .. j = k .. l) = ((i > j \ k > l) \ (i = k \ j = l))" (is "?lhs = ?rhs") +proof (cases "i > j") + case True + with i j have "i .. j = {}" + by (simp add: intvlEmpty_iff) + moreover + from k l have "(k .. l = {}) = (k > l)" + by (simp add: intvlEmpty_iff) + ultimately show ?thesis + using assms True by force +next + case False + with i j have ij: "i \ j" by simp + with i j have "i \ i .. j" "j \ i .. j" by simp+ + { + assume eq: "i..j = k..l" and idx: "i \ k \ j \ l" + from eq \i \ i..j\ intvlIsEmpty[OF k l] have "\(k > l)" + by blast + with k l have "k \ k..l" "l \ k..l" by simp+ + from idx have "FALSE" + proof + assume "i \ k" + from this i k show "FALSE" + proof (rule int_neq_lessE) + assume "i < k" + with i k have "i \ k .. l" by auto + with eq \i \ i .. j\ show "FALSE" by simp + next + assume "k < i" + with i k have "k \ i .. j" by auto + with eq \k \ k..l\ show "FALSE" by simp + qed + next + assume "j \ l" + from this j l show "FALSE" + proof (rule int_neq_lessE) + assume "j < l" + with j l have "l \ i .. j" by auto + with eq \l \ k ..l\ show "FALSE" by simp + next + assume "l < j" + with j l have "j \ k..l" by auto + with eq \j \ i..j\ show "FALSE" by simp + qed + qed + } + thus ?thesis using False by auto +qed + +lemma zerotoInj [simp]: + assumes "l \ Nat" and "m \ Int" and "n \ Int" + shows "(0 .. l = m .. n) = (m=0 \ l=n)" +using assms by (auto simp: intvlEqual_iff) + +lemma zerotoInj' [simp]: + assumes "k \ Int" and "l \ Int" and "n \ Nat" + shows "(k .. l = 0 .. n) = (k=0 \ l=n)" +using assms by (auto simp: intvlEqual_iff) + +lemma zerotoEmpty [simp]: + assumes "m \ Nat" + shows "succ[m] .. 0 = {}" + using assms by (intro intvlIsEmpty) auto + +lemma onetoInj [simp]: + assumes "l \ Nat" and "m \ Int" and "n \ Int" and "l \ 0 \ 1 \ l" + shows "(1 .. l = m .. n) = (m=1 \ l=n)" + using assms by (auto simp: intvlEqual_iff) + +lemma onetoInj' [simp]: + assumes "k \ Int" and "l \ Int" and "n \ Nat" and "n\0 \ 1 \ n" + shows "(k .. l = 1 .. n) = (k=1 \ l=n)" + using assms by (auto simp: intvlEqual_iff) + +lemma SuccInIntvlSucc: + assumes "m \ k" and "k \ Int" and "m \ Int" and "n \ Int" + shows "(succ[k] \ m .. succ[n]) = (k \ m .. n)" + using assms by auto + +lemma SuccInIntvlSuccSucc: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[i] \ succ[j] .. succ[k]) = (i \ j .. k)" + using assms by auto + +end diff --git a/isabelle/IntegerDivision.thy b/isabelle/IntegerDivision.thy new file mode 100644 index 00000000..2da1b4eb --- /dev/null +++ b/isabelle/IntegerDivision.thy @@ -0,0 +1,713 @@ +(* Title: TLA+/IntegerDivision.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation + License: BSD + Version: Isabelle2021-1 +*) + +section \ The division operators div and mod on the integers \ + +theory IntegerDivision +imports IntegerArithmetic +begin + +subsection \ The divisibility relation \ + +definition dvd (infixl "dvd" 50) +where "a \ Int \ b \ Int \ b dvd a \ (\k \ Int : a = b * k)" + +lemma boolify_dvd [simp]: + assumes "a \ Int" and "b \ Int" + shows "boolify(b dvd a) = (b dvd a)" +using assms by (simp add: dvd_def) + +lemma dvdIsBool [intro!,simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "isBool(b dvd a)" +using assms by (simp add: dvd_def) + +lemma [intro!]: + "\isBool(P); isBool(a dvd b); (a dvd b) \ P\ \ (a dvd b) = P" + "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" +by auto + +lemma dvdI [intro]: + assumes a: "a \ Int" and b: "b \ Int" and k: "k \ Int" + shows "a = b * k \ b dvd a" +unfolding dvd_def[OF a b] using k by blast + +lemma dvdE [elim]: + assumes "b dvd a" and "a \ Int" and "b \ Int" + shows "(\k. \k \ Int; a = b * k\ \ P) \ P" +using assms by (auto simp: dvd_def) + +lemma dvd_refl [iff]: + assumes a: "a \ Int" + shows "a dvd a" +proof - + from a have "a = a*1" by simp + with a show ?thesis by blast +qed + +lemma dvd_trans [trans]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and 1: "a dvd b" and 2: "b dvd c" + shows "a dvd c" +proof - + from a b 1 obtain k where k: "k \ Int" and "b = a * k" by blast + moreover + from b c 2 obtain l where l: "l \ Int" and "c = b * l" by blast + ultimately have h: "c = a * (k * l)" + using a b c by (simp add: times_assoc) + thus ?thesis using a c k l by blast +qed + +lemma dvd_0_left_iff [simp]: + assumes "a \ Int" + shows "(0 dvd a) = (a = 0)" +using assms by auto + +lemma dvd_0_right [iff]: + assumes a: "a \ Int" shows "a dvd 0" +using assms by force + +lemma one_dvd [iff]: + assumes a: "a \ Int" + shows "1 dvd a" +using assms by force + +lemma dvd_mult (*[simp]*): + assumes dvd: "a dvd c" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a dvd (b * c)" +proof - + from dvd a c obtain k where k: "k \ Int" and "c = a*k" by blast + with a b c have "b*c = a * (b*k)" by (simp add: times_left_commute) + with a b c k show ?thesis by blast +qed + +lemma dvd_mult2 (*[simp]*): + assumes dvd: "a dvd b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a dvd (b * c)" +using times_commute[OF b c] dvd_mult[OF dvd a c b] by simp + +lemma dvd_triv_right [iff]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a dvd (b * a)" +using assms by (intro dvd_mult) simp+ + +lemma dvd_triv_left [iff]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a dvd a * b" +using assms by (intro dvd_mult2) simp+ + +lemma mult_dvd_mono: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and d: "d \ Int" + and 1: "a dvd b" and 2: "c dvd d" + shows "(a * c) dvd (b * d)" +proof - + from a b 1 obtain b' where b': "b = a * b'" "b' \ Int" by blast + from c d 2 obtain d' where d': "d = c * d'" "d' \ Int" by blast + with b' a b c d + times_assoc[of a b' "c * d'"] + times_left_commute[of b' c d'] + times_assoc[of a c "b' * d'"] + have "b * d = (a * c) * (b' * d')" by simp + with a c b' d' show ?thesis by blast +qed + +lemma dvd_mult_left: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and h: "a * b dvd c" + shows "a dvd c" +proof - + from h a b c obtain k where k: "k \ Int" "c = a*(b*k)" + by (auto simp add: times_assoc) + with a b c show ?thesis by blast +qed + +lemma dvd_mult_right: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and h: "a*b dvd c" + shows "b dvd c" +proof - + from h a b c have "b*a dvd c" by (simp add: times_ac) + with b a c show ?thesis by (rule dvd_mult_left) +qed + +lemma dvd_0_left: + assumes "a \ Int" + shows "0 dvd a \ a = 0" + using assms by simp + +lemma dvd_add [iff]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and 1: "a dvd b" and 2: "a dvd c" + shows "a dvd (b + c)" +proof - + from a b 1 obtain b' where b': "b' \ Int" "b = a * b'" by blast + from a c 2 obtain c' where c': "c' \ Int" "c = a * c'" by blast + from a b c b' c' + have "b + c = a * (b' + c')" by (simp add: add_times_distrib_left) + with a b' c' show ?thesis by blast +qed + +subsection \ Division and modulus on @{const Int} \ + +text \ + \tlaplus{} defines integer division only for the case + where the divisor is positive. Two different extensions + are commonly introduced for negative divisors, and we + leave this case undefined in order to avoid confusion. +\ + +definition div (infixr "div" 70) +where int_div_def: "\a \ Int; b \ Int\ \ + a div b \ CHOOSE q \ Int : \r \ 0 .. (b-1) : a = b * q + r" + +definition mod (infixr "%" 70) +where "a % b \ a - b * (a div b)" + +lemma div_exists: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" + shows "\q \ Int : \r \ 0 .. (b-1) : a = b * q + r" +using a proof induct + case 0 + from b have "0 = b * 0 + 0" "0 \ 0 .. (b-1)" by simp+ + thus ?case by blast +next + case (pos n) + then obtain q r where qr: "q \ Int" "r \ 0 .. (b-1)" "n = b * q + r" + by blast + show ?case + proof (cases "r = b-1") + case True + with \n \ Nat\ b qr + have "succ[n] = b * succ[q] + 0" "0 \ 0 .. (b-1)" + by simp+ + with \q \ Int\ show ?thesis by blast + next + case False + with b qr + have "succ[r] \ 0 .. (b-1)" "succ[n] = b * q + succ[r]" + by (auto simp: less_def) + with \q \ Int\ show ?thesis by blast + qed +next + case (neg n) + then obtain q r where qr: "q \ Int" "r \ 0 .. (b-1)" "-n = b * q + r" + by blast + show ?case + proof (cases "r = 0") + case True + with \n \ Nat\ b qr + have "-succ[n] = b * pred[q] + (b - 1)" "b-1 \ 0 .. (b-1)" + by (simp add: uminusSucc[of n])+ + with \q \ Int\ show ?thesis by blast + next + case False + with b qr \n \ Nat\ + have "pred[r] \ 0 .. (b-1)" "-succ[n] = b * q + pred[r]" + by (auto simp: uminusSucc[of n]) + with \q \ Int\ show ?thesis by blast + qed +qed + +lemma div_mod_unique: + assumes qr: "a = b * q + r" "q \ Int" "r \ 0 .. (b-1)" + and qr': "a = b * q' + r'" "q' \ Int" "r' \ 0 .. (b-1)" + and ab: "a \ Int" "b \ Int" "0 < b" + shows "q' = q" "r' = r" +proof - + { + fix x x' y y' + assume xy: "x \ Int" "x' \ Int" "y \ Nat" "y' \ Nat" + and y': "y' < b" + and eq: "b * x + y = b * x' + y'" + have "x \ x'" + proof (rule contradiction) + assume "\(x \ x')" + with xy obtain k where k: "k \ Nat" "x = succ[x' + k]" + by (auto dest: int_less_imp_succ_add) + with xy ab eq have "b*x' + (b*k + b + y) = b*x' + y'" + by (simp add: add_times_distrib_left add_assoc) + with xy ab k have "b*k + b + y = y'" by simp + with xy ab k y' have "(b*k + y) + b < b" by (simp add: add_ac) + with xy ab k add_less_right_mono[of "(b*k + y) + b" "b" "-b"] + have "b*k + y < 0" by simp + moreover + from xy ab k times_leq_right_pos_mono[of "0" "b" "k"] + have "0 \ b*k + y" by (simp add: less_def) + ultimately show "FALSE" + using xy ab k by simp + qed + } + from this[of q q' r r'] this[of q' q r' r] assms + show 1: "q' = q" by (auto dest: int_leq_antisym) + with add_right_diff_left[of a "b*q" r] add_right_diff_left[of a "b*q'" r'] assms + show "r' = r" by simp +qed + +lemma div_type [iff]: + assumes ab: "a \ Int" "b \ Int" and "0 < b" + shows "a div b \ Int" + using div_exists[OF assms] unfolding int_div_def[OF ab] + by (rule bChooseI2) + +lemma mod_type [iff]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "a % b \ Int" + using assms unfolding mod_def by simp + +lemma mod_range: + assumes ab: "a \ Int" "b \ Int" and "0 < b" + shows "a % b \ 0 .. (b-1)" +proof - + from assms have "a div b \ Int" by simp + from div_exists[OF assms] + have "\r \ 0 .. (b-1) : a = b * (a div b) + r" + unfolding int_div_def[OF ab] by (rule bChooseI2) + then obtain r where "r \ 0 .. (b-1)" "a = b * (a div b) + r" + by blast + with assms show ?thesis + by (simp add: mod_def add_right_diff_left[of "a" "b * (a div b)" r]) +qed + +lemma mod_div_equality: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) + b * (a div b) = a" + using assms unfolding mod_def by simp + +lemma mod_div_equality2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "b * (a div b) + (a % b) = a" + using assms by (simp add: mod_div_equality add_commute) + +lemma mod_div_equality3: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a div b) * b + (a % b) = a" + using assms by (simp add: mod_div_equality2 times_commute) + +lemma mod_div_equality4: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) + (a div b) * b = a" + using assms by (simp add: mod_div_equality times_commute) + +lemma div_equal: + assumes "a = b*q + r" and "a \ Int" and "b \ Int" and "0 < b" + and "q \ Int" and "r \ 0 .. (b-1)" + shows "a div b = q" + using assms mod_div_equality2[of a b] mod_range[of a b] + div_mod_unique[of a b q r "a div b" "a % b"] + by auto + +lemma mod_equal: + assumes "a = b*q + r" and "a \ Int" and "b \ Int" and "0 < b" + and "q \ Int" and "r \ 0 .. (b-1)" + shows "a % b = r" + using assms mod_div_equality2[of a b] mod_range[of a b] + div_mod_unique[of a b q r "a div b" "a % b"] + by auto + + +text \ ``Recursive'' laws about @{text div} and @{text "%"}. \ + +lemma div_mod_nat_less_divisor: + assumes "a \ 0 .. (b-1)" "b \ Int" "0 < b" + shows "a div b = 0" "a % b = a" +proof - + from assms have 1: "a = b*0 + a" by simp + from 1[THEN div_equal] assms show "a div b = 0" + by blast + from 1[THEN mod_equal] assms show "a % b = a" + by blast +qed + +lemma div_is_zero_iff [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a div b = 0) = (a \ 0 .. (b-1))" + using assms mod_div_equality2[of a b] mod_range[of a b] + by (auto simp: div_mod_nat_less_divisor) + +lemma mod_is_self_iff [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b = a) = (a \ 0 .. (b-1))" + using assms mod_div_equality2[of a b] mod_range[of a b] + by (auto simp: div_mod_nat_less_divisor) + +lemma div_mod_minus_less_divisor: + assumes "a \ (-b) .. (-1)" "b \ Int" "0 < b" + shows "a div b = -1" "a % b = b+a" +proof - + from assms have "0 \ b+a" + using add_leq_left_mono[of "-b" "a" "b"] by simp + moreover + from assms have "b+a < b" + using int_leq_pred_iff[of a 0] add_less_left_mono[of a 0 b] + by (simp add: pred0) + ultimately have 1: "b+a \ 0 .. (b-1)" + using assms by simp + from assms have 2: "a = b*(-1) + (b+a)" by (simp add: add_assoc) + from 1 2[THEN div_equal] assms show "a div b = -1" + by blast + from 1 2[THEN mod_equal] assms show "a % b = b+a" + by blast +qed + +lemma div_is_minusone_iff [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a div b = -1) = (a \ (-b) .. (-1))" +proof - + { + assume "a div b = -1" + with assms mod_div_equality2[of a b] + have 1: "a = -b + (a % b)" by simp + from assms mod_range[of a b] have "-b \ -b + (a % b)" + by simp + moreover + from assms mod_range[of a b] have "-b + (a % b) < -b + b" + by (intro add_leq_less_mono) simp+ + ultimately have "a \ (-b) .. (-1)" + using assms 1 by simp + } + moreover + { + assume "a \ (-b) .. (-1)" + with assms div_mod_minus_less_divisor[of a b] have "a div b = -1" + by blast + } + ultimately show ?thesis + using assms by blast +qed + +lemma div_mod_plus_divisor1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a+b) div b = succ[a div b]" "(a+b) % b = a % b" +proof - + from assms have "a = b * (a div b) + (a % b)" + by (simp add: mod_div_equality2) + with assms have "a+b = b * (a div b) + (a % b) + b" + by simp + with assms have 1: "a+b = b * succ[a div b] + (a % b)" + by (simp add: add_ac) + from 1[THEN div_equal] assms mod_range[of a b] + show "(a+b) div b = succ[a div b]" + by blast + from 1[THEN mod_equal] assms mod_range[of a b] + show "(a+b) % b = a % b" + by blast +qed + +lemma div_mod_plus_divisor2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b+a) div b = succ[a div b]" "(b+a) % b = a % b" + using assms by (simp add: add_commute)+ + +lemma div_mod_minus_divisor1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a-b) div b = pred[a div b]" "(a-b) % b = a % b" +proof - + from assms have "a = b * (a div b) + (a % b)" + by (simp add: mod_div_equality2) + with assms have "a-b = b * (a div b) + (a % b) - b" + by simp + with assms have 1: "a-b = b * pred[a div b] + (a % b)" + by (simp add: add_right_commute) + from 1[THEN div_equal] assms mod_range[of a b] + show "(a-b) div b = pred[a div b]" + by blast + from 1[THEN mod_equal] assms mod_range[of a b] + show "(a-b) % b = a % b" + by blast +qed + +lemma div_mod_minus_divisor2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(-b + a) div b = pred[a div b]" "(-b + a) % b = a % b" + using assms by (simp add: add_commute)+ + + +subsection \ Facts about @{text div} and @{text "%"} \ + +lemma div_times_divisor1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + c*b) div b = c + (a div b)" + using c by (induct) (simp add: a b add_assoc)+ + +lemma div_times_divisor2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + b*c) div b = c + (a div b)" + using c by (induct) (simp add: a b add_assoc)+ + +lemma div_times_divisor3 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(b*c + a) div b = c + (a div b)" + using assms by (simp add: add_commute) + +lemma div_times_divisor4 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(c*b + a) div b = c + (a div b)" + using assms by (simp add: add_commute) + +lemma mod_times_divisor1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + c*b) % b = a % b" + using c by (induct) (simp add: a b add_assoc)+ + +lemma mod_times_divisor2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + b*c) % b = a % b" + using c by (induct) (simp add: a b add_assoc)+ + +lemma mod_times_divisor3 [simp]: + assumes "a \ Int" and "b \ Int" "0 < b" and "c \ Int" + shows "(b*c + a) % b = a % b" + using assms by (simp add: add_commute) + +lemma mod_times_divisor4 [simp]: + assumes "a \ Int" and "b \ Int" "0 < b" and "c \ Int" + shows "(c*b + a) % b = a % b" + using assms by (simp add: add_commute) + +lemma times_div_self1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a*b) div b = a" + using assms div_times_divisor1[of 0 b a] by simp + +lemma times_div_self2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b*a) div b = a" + using assms by (simp add: times_commute) + +lemma times_mod_self1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a*b) % b = 0" + using assms mod_times_divisor1[of 0 b a] by simp + +lemma times_mod_self2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b*a) % b = 0" + using assms by (simp add: times_commute) + +lemma div_mod_by_one [simp]: + "a \ Int \ a div 1 = a" + "a \ Int \ a % 1 = 0" + using times_div_self1[of a 1] times_mod_self1[of a 1] by simp+ + +lemma div_mod_by_self [simp]: + "\a \ Int; 0 < a\ \ a div a = 1" + "\a \ Int; 0 < a\ \ a % a = 0" + using times_div_self1[of 1 a] times_mod_self1[of 1 a] by simp+ + +lemma mod_div_same [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) div b = 0" + using assms mod_range[of a b] by simp + +lemma mod_mod_same [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) % b = a % b" + using assms mod_times_divisor1[of "a % b" "b" "a div b"] mod_range[of a b] + by simp + +lemma div_mod_decomposition: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" + obtains q r where "q \ Int" and "r \ 0 .. (b-1)" + and "q = a div b" and "r = a % b" and "a = b*q + r" + using assms mod_div_equality2[of a b, THEN sym] mod_range[of a b] + by blast + + +subsection \Further arithmetic laws for division and modulus\ + +lemma add_div: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a + b) div c = a div c + b div c + (a%c + b%c) div c" +proof - + from assms have "a + b = (c * (a div c) + a%c) + (c * (b div c) + b%c)" + by (simp add: mod_div_equality2) + with assms have "a + b = c * (a div c + b div c) + (a%c + b%c)" + by (simp add: add_times_distrib_left add_commute add_left_commute) + with assms show ?thesis by simp +qed + +lemma add_mod: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a + b) % c = ((a % c) + (b % c)) % c" +proof - + from assms have "(a%c + b%c) % c = (c * (a div c + b div c) + (a%c + b%c)) % c" + by simp + moreover + from assms have "c * (a div c + b div c) + (a%c + b%c) + = (c * (a div c) + a%c) + (c * (b div c) + b%c)" + by (simp add: add_times_distrib_left add_commute add_left_commute) + with assms have "c * (a div c + b div c) + (a%c + b%c) = a+b" + by (simp add: mod_div_equality2) + ultimately show ?thesis by simp +qed + +lemma times_div: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * b) div c = a * (b div c) + (a div c) * (b%c) + ((a%c) * (b%c)) div c" +proof - + from assms have "a * b = (c * (a div c) + a%c) * (c * (b div c) + b%c)" + by (simp add: mod_div_equality2) + also from assms + have "\ = (c * (a div c) + a%c) * (c * (b div c)) + (c * (a div c) + a%c) * (b%c)" + by (simp add: add_times_distrib_left) + also from assms + have "\ = (c * (a div c) * (c * (b div c)) + + (a%c) * (c * (b div c))) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: add_times_distrib_right) + also from assms + have "\ = (c * c * (a div c) * (b div c) + + (a%c) * c * (b div c)) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: times_assoc times_right_commute) + also from assms + have "\ = (c * c * (a div c) + (a%c) * c) * (b div c) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: add_times_distrib_right) + also from assms + have "\ = (c * c * (a div c) + c * (a%c)) * (b div c) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: times_commute) + also from assms + have "\ = (c * (c * (a div c)) + c * (a%c)) * (b div c) + + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + by (simp add: times_assoc) + also from assms + have "\ = (c * (c * (a div c) + (a%c))) * (b div c) + + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + by (simp add: add_times_distrib_left) + finally have "a * b = c * (a * (b div c)) + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + using assms mod_div_equality2[of a c] + by (simp add: times_assoc) + with assms have "(a * b) div c = a * (b div c) + (c * ((a div c) * (b%c)) + (a%c) * (b%c)) div c" + by simp + moreover + from assms + have "(c * ((a div c) * (b%c)) + (a%c) * (b%c)) div c = (a div c) * (b%c) + ((a%c) * (b%c)) div c" + by simp + ultimately + have "(a * b) div c = a * (b div c) + ((a div c) * (b%c) + ((a%c) * (b%c)) div c)" + by simp + with assms show ?thesis by (simp add: add_assoc) +qed + +lemma times_mod: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * b) % c = ((a % c) * (b % c)) % c" +proof - + from assms + have "((a % c) * (b % c)) % c + = (c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c)) % c" + by simp + moreover + from assms + have "c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c) + = (c * (a div c) + a % c) * (c * (b div c) + b % c)" + by (simp add: add_times_distrib_left add_times_distrib_right + times_commute times_left_commute + add_commute add_left_commute) + with assms + have "c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c) + = a * b" + by (simp add: mod_div_equality2) + ultimately show ?thesis by simp +qed + +lemma uminus_div1: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) div b = -(a div b) + (-(a%b)) div b" +proof - + from assms have "-a = -(b * (a div b) + a%b)" + by (simp add: mod_div_equality2) + with assms have "-a = b * (-(a div b)) - (a % b)" + by simp + with assms show ?thesis + by (simp del: times_uminus_right) +qed + +lemma uminus_div2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) div b = -(a div b) + (IF a%b = 0 THEN 0 ELSE -1)" +proof (cases "a % b = 0") + case True + with assms show ?thesis by (simp add: uminus_div1) +next + case False + with assms mod_range[of a b] have "-(a % b) \ (-b) .. (-1)" + by auto + with assms have "(-(a % b)) div b = -1" + by (simp add: div_mod_minus_less_divisor) + with assms False uminus_div1[of a b] + show ?thesis by simp +qed + +lemma uminus_mod1: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) % b = (-(a % b)) % b" +proof - + from assms have "(-a) % b = (-(b * (a div b) + a % b)) % b" + by (simp add: mod_div_equality2) + moreover + from assms have "-(b * (a div b) + a % b) = b * (-(a div b)) - a % b" + by simp + ultimately show ?thesis + using assms mod_times_divisor3[of "-(a % b)" "b" "-(a div b)"] + by simp +qed + +lemma uminus_mod2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) % b = (IF a % b = 0 THEN 0 ELSE b - (a%b))" +proof (cases "a % b = 0") + case True + with assms show ?thesis by (simp add: uminus_mod1) +next + case False + with assms mod_range[of a b] have "-(a % b) \ (-b) .. (-1)" + by auto + with assms have "(-(a % b)) % b = b - (a % b)" + by (simp add: div_mod_minus_less_divisor) + with assms uminus_mod1[of a b] False + show ?thesis by simp +qed + + +subsection \Divisibility and @{text div} / @{text mod}\ + +lemma dvd_is_mod0: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(b dvd a) = (a % b = 0)" + using assms mod_div_equality2[of a b] by auto + +lemma dvd_times_div_self: + assumes "b dvd a" and "a \ Int" and "b \ Int" and "0 < b" + shows "b * (a div b) = a" + using assms dvd_is_mod0[of a b] mod_div_equality2[of a b] by simp + +lemma dvd_times_div_distrib: + assumes dvd: "b dvd a" and a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "c * (a div b) = (c*a) div b" +proof - + from dvd a b obtain d where d: "d \ Int" "a = b*d" by auto + with a b c have "c*a = b * (c*d)" by (simp add: times_left_commute) + with a b c d show ?thesis by simp +qed + +lemma dvd_mod_imp_dvd: + assumes "a dvd (c % b)" and "a dvd b" + and a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "a dvd c" +proof - + from assms have "a dvd ((c div b) * b + c % b)" + by (simp add: dvd_mult) + with a b c show ?thesis + by (simp add: mod_div_equality3) +qed + +end diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index fdcc141f..bfbefc7f 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -1,1314 +1,1537 @@ -(* Title: Integers.thy - Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation +(* Title: TLA+/Integers.thy + Author: Stephan Merz, LORIA + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:02 merz> - + Version: Isabelle2021-1 *) -section \ The Integers as a superset of natural numbers \ +section \ The set of integer numbers \ theory Integers -imports Tuples NatArith +imports FixedPoints Functions begin -subsection \ The minus sign \ - -consts - "minus" :: "c \ c" ("-._" [75] 75) - -syntax (* -- syntax for negative naturals *) - "-.0" :: "c" ("-.0") - "-.1" :: "c" ("-.1") - "-.2" :: "c" ("-.2") -translations - "-.0" \ "-.(0)" - "-.1" \ "-.(1)" - "-.2" \ "-.(2)" - -(* lemma eq_imp_negeq: "n = m \ -.n = -.m" by simp *) - -axiomatization where - neg0 [simp]: "-.0 = 0" -and - neg_neg [simp]: "-.-.n = n" -and - negNotInNat [simp]: "-.(Succ[n]) \ Nat" - -lemma negNat_noteq_Nat [simp]: - "\m \ Nat; n \ Nat\ \ (-. Succ[m] = Succ[n]) = FALSE" -proof (rule contradiction) - assume "(-. Succ[m] = Succ[n]) \ FALSE" - and m: "m \ Nat" and n: "n \ Nat" - hence "-. Succ[m] = Succ[n]" by auto - hence "-. Succ[m] \ Nat" using n by auto - with negNotInNat[of m] show FALSE by simp -qed +subsection \Extended Peano axioms\ -lemma negNat_noteq_Nat2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m] = -. Succ[n]) = FALSE" -proof auto - assume "Succ[m] = -. Succ[n]" - hence "-. Succ[n] = Succ[m]" by simp - with m n show "FALSE" by simp -qed +text \ + We extend the standard Peano axioms by a complement operation + for constructing negative integers, and we prove the existence + of a structure satisfying these axioms. +\ -lemma nat_not_eq_inv: "n \ Nat \ n = 0 \ -.n \ n" - using not0_implies_Suc[of n] by auto +definition ExtendedPeano :: "[c,c,c,c,c] \ c" where + \ \Parameters: the set of integer and natural numbers, zero, + and successor and complement functions.\ + "ExtendedPeano(I,N,Z,sc,cp) \ + Z \ N + \ sc \ [N \ N] + \ (\n \ N : sc[n] \ Z) + \ (\m,n \ N : sc[m] = sc[n] \ m = n) + \ (\S \ SUBSET N : Z \ S \ (\n\S : sc[n] \ S) \ N \ S) + \ cp \ [I \ I] + \ I = N \ {cp[n] : n \ N} + \ cp[Z] = Z + \ (\k \ N : cp[k] \ N \ k = Z) + \ (\k,l \ I : cp[k] = cp[l] \ k = l) + \ (\k \ I : cp[cp[k]] = k)" -lemma minusInj [dest]: - assumes hyp: "-.n = -.m" - shows "n = m" +text \ + We now prove the existence of a structure satisfying the + extended Peano axioms. For @{text N}, we take the standard + ZF construction where @{text "{}"} is zero, where + @{text "i \ {i}"} is the successor of any natural number + @{text i}, and where the set @{text N} is defined as + the least set that contains zero and is closed under successor + (this is a subset of the infinity set asserted to exist in ZF + set theory). We then pick a value @{text e} that does not + occur in the construction of @{text N} and define the + complement function by adding or removing that value to / from + the set representing the argument. Later, integers are defined + using a sequence of @{text CHOOSE}'s, so there is no commitment + to that particular structure. +\ + +theorem extendedPeanoExists: "\I,N,Z,sc,cp : ExtendedPeano(I,N,Z,sc,cp)" proof - - from hyp have "-.-.n = -.-.m" by simp - thus ?thesis by simp -qed + let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ + define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" + define N where "N \ lfp(infinity, expand)" + define Z where "Z \ {}" + define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ + have mono: "Monotonic(infinity, expand)" + using infinity by (auto simp: Monotonic_def expand_def) + hence expandN: "expand(N) \ N" + by (unfold N_def, rule lfpPreFP) + from expandN have 1: "Z \ N" + by (auto simp: expand_def Z_def) + have 2: "Sc \ [N \ N]" + proof (unfold Sc_def, rule functionInFuncSet) + show "\n \ N : ?sc(n) \ N" using expandN by (auto simp: expand_def) + qed + have 3: "\m\N : Sc[m] \ Z" + unfolding Z_def Sc_def by auto + have 4: "\m,n \ N : Sc[m] = Sc[n] \ m = n" + proof (clarify) + fix m n + assume "m \ N" and "n \ N" and "Sc[m] = Sc[n]" + hence eq: "?sc(m) = ?sc(n)" by (simp add: Sc_def) + show "m = n" + proof (rule setEqual) + show "m \ n" + proof (rule subsetI) + fix x + assume x: "x \ m" show "x \ n" + proof (rule contradiction) + assume "x \ n" + with x eq have "n \ m" by auto + moreover + from eq have "m \ ?sc(n)" by auto + ultimately + show "FALSE" by (blast elim: inAsym) + qed + qed + next + show "n \ m" + proof (rule subsetI) + fix x + assume x: "x \ n" show "x \ m" + proof (rule contradiction) + assume "x \ m" + with x eq have "m \ n" by auto + moreover + from eq have "n \ ?sc(m)" by auto + ultimately + show "FALSE" by (blast elim: inAsym) + qed + qed + qed + qed + have 5: "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" + proof (clarify del: subsetI) + fix S + assume sub: "S \ N" and Z: "Z \ S" and Sc: "\n\S : Sc[n] \ S" + show "N \ S" + proof (unfold N_def, rule lfpLB) + show "expand(S) \ S" + proof (auto simp: expand_def) + from Z show "{} \ S" by (simp add: Z_def) + next + fix n + assume n: "n \ S" + with Sc have "Sc[n] \ S" .. + moreover + from n sub have "n \ N" by auto + hence "Sc[n] = ?sc(n)" by (simp add: Sc_def) + ultimately show "?sc(n) \ S" by simp + qed + next + have "N \ infinity" + by (unfold N_def, rule lfpSubsetDomain) + with sub show "S \ infinity" by auto + qed + qed + + define e where "e \ CHOOSE e : e \ UNION {n : n \ N}" + define cp where "cp \ \n. IF n = Z THEN n ELSE IF e \ n THEN n \ {e} ELSE n \ {e}" + define I where "I \ N \ {cp(n) : n \ N}" + let ?cp = "[k \ I \ cp(k)]" + + have "\e : e \ UNION {n : n \ N}" by (blast intro: inIrrefl) + hence "e \ UNION {n : n \ N}" unfolding e_def by (rule chooseI_ex) + hence e: "\n \ N : e \ n" by blast + + have cpZ: "cp(Z) = Z" by (simp add: cp_def) + + have diff: "\k \ I \ {Z} : cp(k) \ k" + proof (clarify) + fix k + assume "k \ I" "k \ Z" "cp(k) = k" + show "FALSE" + proof (cases "k \ N") + case True + with e \k \ Z\ have "e \ k \ cp(k) = k \ {e}" by (simp add: cp_def) + with \cp(k) = k\ show ?thesis by blast + next + case False + with \k \ I\ obtain n where "n \ N" "k = cp(n)" by (auto simp: I_def) + with \k \ Z\ have "n \ Z" by (auto simp: cp_def) + with e \n \ N\ \k = cp(n)\ \k \ Z\ + have "k = n \ {e}" "cp(k) = k \ {e}" by (auto simp: cp_def) + with \cp(k) = k\ show ?thesis by auto + qed + qed + + from e have cpN: "\n \ N : cp(cp(n)) = n" + by (auto simp: Z_def cp_def) + + from e have cpNN: "\k \ N : cp(k) \ N \ k = Z" + by (auto simp: cp_def) + + { + fix k + assume k: "k \ I" "cp(k) = Z" + have "k = Z" + proof (cases "k \ N") + case True + with k 1 cpNN show ?thesis by blast + next + case False + with \k \ I\ obtain n where "n \ N" "k = cp(n)" + by (auto simp: I_def) + with \cp(k) = Z\ cpN have "n = Z" by simp + with \k = cp(n)\ show ?thesis by (simp add: cp_def) + qed + } + with cpZ have cpisZ: "\k \ I: cp(k) = Z \ k = Z" by blast + + from cpN have cpI: "\k \ I : cp(k) \ I" by (auto simp: I_def) + hence 6: "?cp \ [I \ I]" by (rule functionInFuncSet) + + have 7: "I = N \ {?cp[n] : n \ N}" by (force simp: I_def) + + from cpZ 1 have 8: "?cp[Z] = Z" by (simp add: I_def) + + from cpNN have 9: "\n \ N : ?cp[n] \ N \ n = Z" by (simp add: I_def) + + { + fix k l + assume dom: "k \ I" "l \ I" and eq: "cp(k) = cp(l)" and kl: "k \ l" + have "FALSE" + proof (cases "k \ N") + case True + have "k \ Z" + proof + assume "k = Z" + with eq cpZ dom kl cpisZ show "FALSE" by auto + qed + with \k \ N\ e have cpk: "e \ k" "cp(k) = k \ {e}" by (auto simp: cp_def) + have "l \ N" + proof (rule contradiction) + assume "l \ N" + with \l \ I\ obtain n where "n \ N" "l = cp(n)" + by (auto simp: I_def) + with cpk cpN eq have "n = k \ {e}" by simp + with \n \ N\ e show "FALSE" by blast + qed + have "l \ Z" + proof + assume "l = Z" + with eq cpZ dom kl cpisZ show "FALSE" by auto + qed + with \l \ N\ e have "e \ l" "cp(l) = l \ {e}" by (auto simp: cp_def) + with cpk eq have "k \ l" "l \ k" by auto + with kl show ?thesis by (auto dest: setEqual) + next + case False + with dom obtain nk where nk: "nk \ N" "k = cp(nk)" + by (auto simp: I_def) + with cpN eq have cpl: "cp(l) \ N" by simp + have "l \ N" + proof + assume "l \ N" + with cpl cpNN have "l = Z" by simp + with eq dom cpZ cpisZ False 1 show "FALSE" by simp + qed + with dom obtain nl where "nl \ N" "l = cp(nl)" + by (auto simp: I_def) + with nk eq kl cpN show ?thesis by simp + qed + } + hence 10: "\k,l \ I : ?cp[k] = ?cp[l] \ k = l" by auto -lemma minusInj_iff [simp]: - "-.x = -.y = (x = y)" -by auto + from cpI cpN have 11: "\k \ I : ?cp[?cp[k]] = k" + by (auto simp: I_def) -lemma neg0_imp_0 [simp]: "-.n = 0 = (n = 0)" -proof auto - assume "-.n = 0" - hence "-.-.n = 0" by simp - thus "n = 0" by simp + from 1 2 3 4 5 6 7 8 9 10 11 have "ExtendedPeano(I,N,Z,Sc,?cp)" + unfolding ExtendedPeano_def by blast + thus ?thesis by blast qed -lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" -by simp +subsection \The structure of integer numbers\ -lemma notneg0_imp_not0 [dest]: "-.n \ 0 \ n \ 0" -by auto +text \ + The integer numbers are now defined as some structure + satisfying the extended Peano axioms. +\ -lemma not0_imp_notNat [simp]: "n \ Nat \ n \ 0 \ -.n \ Nat" - using not0_implies_Suc[of n] by auto +definition Succ :: "c" where + "Succ \ CHOOSE sc : \cp,I,N,Z : ExtendedPeano(I,N,Z,sc,cp)" -lemma negSuccNotZero [simp]: "n \ Nat \ (-. Succ[n] = 0) = FALSE" -by auto +definition Nat :: "c" where + "Nat \ DOMAIN Succ" -lemma negSuccNotZero2 [simp]: "n \ Nat \ (0 = -. Succ[n]) = FALSE" -proof auto - assume n: "n \ Nat" and 1: "0 = -. Succ[n]" - from 1 have "-. Succ[n] = 0" by simp - with n show FALSE by simp -qed +definition zero :: "c" ("0") where + "0 \ CHOOSE Z : \cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" -lemma negInNat_imp_false [dest]: "-.Succ[n] \ Nat \ FALSE" - using negNotInNat[of n] by simp +definition intCplt :: "c" where + "intCplt \ CHOOSE cp : \I : ExtendedPeano(I,Nat,0,Succ,cp)" -lemma negInNatFalse [simp]: "-.Succ[n] \ Nat = FALSE" - using negNotInNat[of n] by auto +definition Int :: "c" where + "Int \ DOMAIN intCplt" -lemma n_negn_inNat_is0 [simp]: - assumes "n \ Nat" - shows "-.n \ Nat = (n = 0)" -using assms by (cases "n", auto) +definition uminus :: "c \ c" ("-_" [80] 80) where + "-z \ intCplt[z]" -lemma minus_sym: "-.a = b = (a = -.b)" -by auto +lemmas + setEqualI [where A = "Nat", intro!] + setEqualI [where B = "Nat", intro!] + setEqualI [where A = "Int", intro!] + setEqualI [where B = "Int", intro!] -lemma negNat_exists: "-.n \ Nat \ \k \ Nat: n = -.k" -by force +lemma intExtendedPeano: "ExtendedPeano(Int,Nat,0,Succ,intCplt)" +proof - + have "\cp,I,N,Z : ExtendedPeano(I,N,Z,Succ,cp)" + proof (unfold Succ_def, rule chooseI_ex) + from extendedPeanoExists + show "\Sc,cp,I,N,Z : ExtendedPeano(I,N,Z,Sc,cp)" by blast + qed + then obtain N Z where PNZ: "\cp,I : ExtendedPeano(I,N,Z,Succ,cp)" by blast + hence "Succ \ [N \ N]" + by (simp add: ExtendedPeano_def) + hence "N = Nat" + by (simp add: Nat_def) + with PNZ have "\cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" by simp + hence "\cp,I : ExtendedPeano(I,Nat,0,Succ,cp)" + unfolding zero_def by (rule chooseI) + hence "\I : ExtendedPeano(I,Nat,0,Succ,intCplt)" + unfolding intCplt_def by (rule chooseI_ex) + then obtain I where IEP: "ExtendedPeano(I,Nat,0,Succ,intCplt)" + by blast + hence "intCplt \ [I \ I]" + by (auto simp: ExtendedPeano_def) + hence "I = Int" + by (simp add: Int_def) + with IEP show ?thesis by simp +qed -lemma nat_eq_negnat_is_0 [simp]: - assumes "n \ Nat" - shows "(n = -.n) = (n = 0)" -using assms by (cases "n", auto) +lemma natIsInt [simp,intro]: "n \ Nat \ n \ Int" + using intExtendedPeano by (auto simp: ExtendedPeano_def) -(* used for simplification in additions *) -lemma (*[simp]*) "\x \ Nat : -.1 = -.x" by auto -lemma (*[simp]*) "x \ Nat \ (1 = -.x) = FALSE" by (auto simp: sym[OF minus_sym]) +lemma uminusIsInt [simp,intro!]: "n \ Int \ -n \ Int" + using intExtendedPeano by (auto simp: ExtendedPeano_def uminus_def) +text \ + Fundamental lemmas about zero, successor, and complement. +\ -subsection \ The set of Integers \ +lemma zeroIsNat [intro!,simp]: "0 \ Nat" + using intExtendedPeano by (simp add: ExtendedPeano_def) -definition Int -where "Int \ Nat \ {-.n : n \ Nat}" +(* redundant, but will become relevant later *) +lemma zeroIsInt [intro!,simp]: "0 \ Int" + by simp -lemma natInInt [simp]: "n \ Nat \ n \ Int" -by (simp add: Int_def) +lemma SuccInNatNat [intro!,simp]: "Succ \ [Nat \ Nat]" + using intExtendedPeano by (simp add: ExtendedPeano_def) -lemma intDisj: "n \ Int \ n \ Nat \ n \ {-.n : n \ Nat}" -by (auto simp: Int_def) +lemma SuccIsAFcn [intro!,simp]: "isAFcn(Succ)" + using SuccInNatNat by blast -lemma negint_eq_int [simp]: "-.n \ Int = (n \ Int)" -unfolding Int_def by force +\ \@{text "DOMAIN Succ = Nat"}\ +lemmas domainSucc [intro!,simp] = funcSetDomain[OF SuccInNatNat] +\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ +lemmas SuccIsNat [intro!,simp] = funcSetValue[OF SuccInNatNat] -lemma intCases [case_names Positive Negative, cases set: Int]: - assumes n: "n \ Int" - and sc: "n \ Nat \ P" - and nsc: "\m. \m \ Nat; n = -.m\ \ P" - shows "P" -using assms unfolding Int_def by auto - -(* -- Integer cases over two parameters *) -lemma intCases2: - assumes m: "m \ Int" and n: "n \ Int" - and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" - and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" - and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" - and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" - shows "P(m,n)" -using m proof (cases "m") - assume "m \ Nat" - from n this pp pn show "P(m,n)" by (cases "n", auto) -next - fix m' - assume "m' \ Nat" "m = -. m'" - from n this np nn show "P(m,n)" by (cases "n", auto) +lemma [simp]: + assumes "n \ Nat" + shows "(Succ[n] = 0) = FALSE" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def) + +lemma [simp]: + assumes n: "n \ Nat" + shows "(0 = Succ[n]) = FALSE" + using assms by (auto dest: sym) + +lemma SuccNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): + "\Succ[n] = 0; n \ Nat\ \ P" + "\0 = Succ[n]; n \ Nat\ \ P" +by (simp+) + +lemma SuccInj [dest]: + assumes "Succ[m] = Succ[n]" and "m \ Nat" and "n \ Nat" + shows "m=n" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def) + +lemma SuccInjIff [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(Succ[m] = Succ[n]) = (m = n)" +using assms by auto + +\ \The primitive induction rule for natural numbers, will be superseded below.\ +lemma pr_natInduct: + assumes z: "P(0)" + and sc: "\n. \n \ Nat; P(n)\ \ P(Succ[n])" + shows "\n\Nat : P(n)" +proof - + let ?P = "{n \ Nat : P(n)}" + have "?P \ SUBSET Nat" by auto + moreover + from z have "0 \ ?P" by simp + moreover + from sc have "\n \ ?P : Succ[n] \ ?P" by simp + ultimately have "Nat \ ?P" + using intExtendedPeano by (simp add: ExtendedPeano_def) + thus ?thesis by auto qed -lemma intCases3: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" - and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" - and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" - and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" - and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" - and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" - and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" - and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" - shows "P(m,n,p)" -proof (rule intCases2[OF m n]) - fix m n - assume "m \ Nat" and "n \ Nat" - from p this ppp ppn show "P(m,n,p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this pnp pnn show "P(m, -.n, p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this npp npn show "P(-.m, n, p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this nnp nnn show "P(-.m, -.n, p)" by (cases "p", auto) +lemma pr_natCases: + assumes n: "n \ Nat" + and 0: "n = 0 \ P" + and suc: "\m. \m \ Nat; n = Succ[m]\ \ P" + shows "P" +proof - + have "\n \ Nat : n = 0 \ (\m \ Nat : n = Succ[m])" + by (rule pr_natInduct) auto + with assms show ?thesis by blast qed -lemma int_eq_negint_is_0 [simp]: "n \ Int \ n = -.n = (n = 0)" -by(rule intCases, auto) - -lemma intNotNatIsNeg: "\n \ Nat; n \ Int\ \ \k \ Nat: n = -.k" -unfolding Int_def by auto - -lemma intNotNatIsNegNat: "\n \ Nat; n \ Int\ \ -.n \ Nat" -unfolding Int_def by auto - - -subsection \ Predicates ''is positive'' and 'is negative' \ - -definition isPos (* -- Predicate ''is positive'' *) -where "isPos(n) \ \k \ Nat: n = Succ[k]" - -definition isNeg (* -- Predicate ''is negative'' *) -where "isNeg(n) \ \k \ Nat: n = -.Succ[k]" - -lemma boolify_isPos [simp]: "boolify(isPos(n)) = (isPos(n))" -by (simp add: isPos_def) - -lemma isPos_isBool [intro!,simp]: "isBool(isPos(n))" -by (simp add: isPos_def) +lemma [simp]: "-0 = 0" + using intExtendedPeano by (simp add: ExtendedPeano_def uminus_def) -lemma boolify_isNeg [simp]: "boolify(isNeg(n)) = (isNeg(n))" -by (simp add: isNeg_def) - -lemma isNeg_isBool [intro!,simp]: "isBool(isNeg(n))" -by (simp add: isNeg_def) - -lemma zeroNotPos [dest]: "isPos(0) \ FALSE" by (auto simp: isPos_def) -lemma zeroNotNeg [dest]: "isNeg(0) \ FALSE" by (auto simp: isNeg_def) - -lemma natIsPos [simp]: "n \ Nat \ isPos(Succ[n])" by(simp add: isPos_def) -lemma negIsNeg [simp]: "n \ Nat \ isNeg(-.Succ[n])" by(simp add: isNeg_def) - -lemma negIsNotPos [simp]: "n \ Nat \ isPos(-.Succ[n]) = FALSE" -by(simp add: isPos_def) - -(*lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \\ {0})" - unfolding isPos_def using not0_implies_Suc[of n] by auto*) - -lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \ n \ 0)" -unfolding isPos_def using not0_implies_Suc[of n] by auto - -(*lemma isNeg_eq_inNegNat: - "isNeg(n) = (n \ {-.n : n \ Nat } \\ {0})" - unfolding isNeg_def by force*) - -lemma isNeg_eq_inNegNat: - "isNeg(n) = (n \ {-.n : n \ Nat } \ n \ 0)" -unfolding isNeg_def by force - -lemma intIsPos_isNat: "n \ Int \ isPos(n) \ n \ Nat" -by(auto simp: isPos_def) - -lemma negNotNat_isNat: - assumes n: "n \ Int" shows "(-.n \ Nat) = FALSE \ n \ Nat" -using n by (cases, auto) - -lemma noNatisNeg [simp]: - "n \ Nat \ isNeg(n) = FALSE" (* -- No natural number is negative *) -unfolding isNeg_def using negNotInNat by blast - -lemma negNat_isNeg [intro]: "\m \ Nat; m \ 0\ \ isNeg(-.m)" -unfolding isNeg_eq_inNegNat by auto - -lemma nat_is_0_or_pos: "(n = 0 \ isPos(n)) = (n \ Nat)" -unfolding isPos_def by force - -lemma isNeg_dichotomy (*[simp]*): "n \ Int \ isNeg(-.n) \ isNeg(n) = FALSE" -unfolding isNeg_def by auto +lemma uminusNat [simp]: + assumes "n \ Nat" + shows "(-n \ Nat) = (n = 0)" +proof - + { + assume "-n \ Nat" + with assms intExtendedPeano have "n = 0" + by (auto simp: ExtendedPeano_def uminus_def) + } + thus ?thesis by auto +qed -lemma isPos_isNeg_false [simp]: "n \ Int \ isPos(n) \ isNeg(n) = FALSE" - unfolding isPos_def by force +lemma uminusInj [dest]: + assumes "-a = -b" and "a \ Int" and "b \ Int" + shows "a = b" + using assms intExtendedPeano + unfolding ExtendedPeano_def uminus_def by blast -(*lemma notPosIsNegPos: - "n \ Int \\ {0} \ \ isPos(n) \ isPos(-.n)" -unfolding isPos_eq_inNat1 using intNotNatIsNegNat by auto*) +lemma uminusInjIff [simp]: + "\a \ Int; b \ Int\ \ (-a = -b) = (a = b)" + by auto -lemma isPos_neg_isNeg [simp]: - assumes n: "n \ Int" shows "isPos(-.n) = isNeg(n)" -by (auto simp: minus_sym isPos_def isNeg_def) +lemma uminusUminus [simp]: + "a \ Int \ --a = a" + using intExtendedPeano by (simp add: ExtendedPeano_def uminus_def) -lemma notIsNeg0_isPos: - assumes n: "n \ Int" - shows "\\ isNeg(n); n \ 0\ \ isPos(n)" -using n by (cases, auto simp: isPos_eq_inNat1 dest: negNat_isNeg) - -lemma notIsPos_notNat [simp]: "\\ isPos(n); n \ 0\ \ n \ Nat = FALSE" -by (auto simp: isPos_eq_inNat1) +lemma uminusSwap: + assumes "a \ Int" and "b \ Int" + shows "(-a = b) = (a = -b)" +proof - + from assms have "(-a = b) = (--a = -b)" by auto + with assms show ?thesis by simp +qed -lemma intThenPosZeroNeg: +lemma intElim: assumes n: "n \ Int" - shows "isNeg(n) \ n = 0 \ isPos(n)" -by (auto elim: notIsNeg0_isPos[OF n]) - - -subsection \ Signum function and absolute value \ + and pos: "n \ Nat \ P" and neg: "\m. \m \ Nat; n = -m\ \ P" + shows "P" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def uminus_def) -definition sgn (* -- signum function *) -where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" +lemma uminusZero [simp]: + assumes "a \ Int" + shows "(-a = 0) = (a = 0)" + using assms by (simp add: uminusSwap) -definition abs (* -- absolute value *) -where "abs(n) \ IF sgn(n) = -.1 THEN -.n ELSE n" +lemma uminusZero' [simp]: + assumes "a \ Int" + shows "(0 = -a) = (a = 0)" +using assms by (auto dest: sym) -lemma sgnInInt [simp]: "n \ Int \ sgn(n) \ Int" -by (auto simp: sgn_def) +lemma uminusReflZero: + assumes "a \ Int" "-a = a" + shows "a = 0" +proof - + { + fix k + assume "k \ Nat" "-k = k" + hence "-k \ Nat" by simp + with \k \ Nat\ have "k = 0" by simp + } + from assms this show ?thesis + by (elim intElim) (force+) +qed -lemma sgn0 [simp]: "sgn(0) = 0" -by (simp add: sgn_def) +lemma [simp]: + "a \ Int \ (-a = a) = (a = 0)" + "a \ Int \ (a = -a) = (a = 0)" + by (auto dest: uminusReflZero) -lemma sgnPos [simp]: "n \ Nat \ sgn(Succ[n]) = 1" -by (simp add: sgn_def) +lemma pr_intCases: + assumes a: "a \ Int" + and 0: "a = 0 \ P" + and pos: "\n. \ n \ Nat; a = Succ[n] \ \ P" + and neg: "\n. \ n \ Nat; a = -Succ[n] \ \ P" + shows "P" +using a proof (rule intElim) + assume "a \ Nat" + with 0 pos show ?thesis by (auto elim: pr_natCases) +next + fix m + assume "m \ Nat" "a = -m" + with 0 neg show ?thesis by (auto elim: pr_natCases) +qed -lemma sgnNeg [simp]: "n \ Nat \ sgn(-.Succ[n]) = -.1" -by (simp add: sgn_def) -lemma sgn0_imp_0: "sgn(n) = 0 \ n = 0" -by (auto simp: sgn_def) +subsection \Successor and predecessor on integer numbers.\ -lemma sgn0_iff_0 [simp]: "(sgn(n) = 0) = (n = 0)" -by (auto simp: sgn_def) +text \ + We extend the successor function for the set of integer numbers + and define the predecessor function as its inverse. These + functions are denoted @{text "succ"} and @{text "pred"} and will + replace the primitive function @{text "Succ"} for reasoning + about integers. +\ -lemma sgn1_imp_pos (*[simp]*): "sgn(n) = 1 \ n \ Nat \ n \ 0" -unfolding sgn_def isPos_eq_inNat1 by auto +definition succ where + "succ \ [a \ Int \ IF a \ Nat THEN Succ[a] + ELSE -(CHOOSE m \ Nat : a = -Succ[m])]" + +abbreviation "one \ succ[0]" +notation "one" ("1") +(* +abbreviation "two \ succ[1]" +notation "two" ("2") +abbreviation "three \ succ[2]" +notation "three" ("3") +abbreviation "four \ succ[3]" +notation "four" ("4") +abbreviation "five \ succ[4]" +notation "five" ("5") +abbreviation "six \ succ[5]" +notation "six" ("6") +abbreviation "seven \ succ[6]" +notation "seven" ("7") +abbreviation "eight \ succ[7]" +notation "eight" ("8") +abbreviation "nine \ succ[8]" +notation "nine" ("9") +abbreviation "ten \ succ[9]" +notation "ten" ("10") +abbreviation "eleven \ succ[10]" +notation "eleven" ("11") +abbreviation "twelve \ succ[11]" +notation "twelve" ("12") +abbreviation "thirteen \ succ[12]" +notation "thirteen" ("13") +abbreviation "fourteen \ succ[13]" +notation "fourteen" ("14") +abbreviation "fifteen \ succ[14]" +notation "fifteen" ("15") +*) -(*lemma neg_imp_negSuc (*[simp]*): - assumes h:"n \ {-.n : n \ Nat } \\ {0}" shows "isNeg(n)" -unfolding isNeg_def +definition two :: "c" ("2") + where "2 \ succ[1]" +definition three :: "c" ("3") + where "3 \ succ[2]" +definition four :: "c" ("4") + where "4 \ succ[3]" +definition five :: "c" ("5") + where "five \ succ[4]" +definition six :: "c" ("6") + where "six \ succ[5]" +definition seven :: "c" ("7") + where "seven \ succ[6]" +definition eight :: "c" ("8") + where "eight \ succ[7]" +definition nine :: "c" ("9") + where "nine \ succ[8]" +definition ten :: "c" ("10") + where "ten \ succ[9]" +definition eleven :: "c" ("11") + where "eleven \ succ[10]" +definition twelve :: "c" ("12") + where "twelve \ succ[11]" +definition thirteen :: "c" ("13") + where "thirteen \ succ[12]" +definition fourteen :: "c" ("14") + where "fourteen \ succ[13]" +definition fifteen :: "c" ("15") + where "fifteen \ succ[14]" + +lemma succType: "succ \ [Int \ Int]" proof - - from h have 1:"\k \ Nat : n = -.k" and 2:"n \ 0" by auto - from 1 obtain k where k:"k \ Nat" and 3:"n = -.k" by auto - with 2 have "k \ 0" by auto - with k not0_implies_Suc[of k] have "\m \ Nat : k = Succ[m]" by auto - with 3 show "\k \ Nat : n = -.(Succ[k])" by auto -qed*) - -lemma sgnm1_imp_neg: - assumes n:"n \ Int" shows "sgn(n) = -.1 \ isNeg(n)" -unfolding sgn_def using intThenPosZeroNeg[OF n] by auto + { + fix a + assume "a \ Int" + hence "succ[a] \ Int" + by (rule pr_intCases) (auto simp: succ_def intro: bChooseI2) + } + thus ?thesis by (auto simp: succ_def) +qed -lemma isPos_sgn [simp]: "isPos(sgn(n)) = isPos(n)" -unfolding isPos_def sgn_def by force +lemma succIsAFcn [intro!,simp]: "isAFcn(succ)" + using succType by blast + +\ \@{text "DOMAIN succ = Int"}\ +lemmas [intro!,simp] = funcSetDomain[OF succType] +\ \@{text "a \ Int \ succ[a] \ Int"}\ +lemmas [intro!,simp] = funcSetValue[OF succType] + +lemma succIsSucc: "n \ Nat \ succ[n] = Succ[n]" + by (simp add: succ_def) + +lemma succUminusSuccNat: "n \ Nat \ succ[-succ[n]] = -n" + by (auto simp: succ_def intro: bChooseI2) + +lemma succUminusSucc [simp]: + assumes "a \ Int" + shows "succ[-succ[a]] = -a" +using assms proof (rule pr_intCases) + fix n + assume n: "n \ Nat" "a = -Succ[n]" + hence "succ[a] = -n" + by (simp add: sym[OF succIsSucc] succUminusSuccNat) + with n show ?thesis + by (simp add: succIsSucc) +qed (simp add: succUminusSuccNat)+ + +lemma succInNat [simp]: + assumes "a \ Int" + shows "(succ[a] \ Nat) = (a \ Nat \ {-1})" +proof - + { + assume "a \ Nat" hence "succ[a] \ Nat" + by (simp add: succIsSucc) + } + moreover + { + assume "a = -1" hence "succ[a] \ Nat" + by simp + } + moreover + { + assume s: "succ[a] \ Nat" + from assms have "a \ Nat \ {-1}" + proof (rule pr_intCases) + fix n + assume "n \ Nat" "a = -Succ[n]" + with s show ?thesis by (simp add: sym[OF succIsSucc]) + qed (simp+) + } + ultimately show ?thesis by blast +qed -lemma sgnNat_is_0or1 (*[simp]*): - "n \ Nat \ sgn(n) = 0 \ sgn(n) = 1" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma succIsUminus [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(succ[a] = -b) = (a = -succ[b])" +proof - + { + assume "succ[a] = -b" + with b have "succ[b] = succ[-succ[a]]" by simp + also from a have "\ = -a" by simp + finally have "a = -succ[b]" using a by simp + } + with assms show ?thesis by auto +qed -lemma sgnNat_not0: - "\n \ Nat; sgn(n) \ 0\ \ sgn(n) = 1" -using sgnNat_is_0or1[of n] by auto +lemma uminusIsSucc [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(-a = succ[b]) = (b = -succ[a])" + using assms by (auto simp: uminusSwap) -lemma sgnNat_not1: - "\n \ Nat; sgn(n) \ 1\ \ n = 0" -using sgnNat_is_0or1[of n] by auto +lemma succIs0: (* better not add to simp *) + assumes "a \ Int" + shows "(succ[a] = 0) = (a = -1)" +proof - + from assms have "(succ[a] = -0) = (a = -1)" + by (rule succIsUminus) simp + thus ?thesis by simp +qed -lemma sgnNat_not_neg [simp]: - "n \ Nat \ sgn(n) = -.1 = FALSE" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma zeroIsSucc: + assumes "a \ Int" + shows "(0 = succ[a]) = (a = -1)" + by (auto simp: sym[OF succIs0[OF assms]]) + +lemma succNatNotZero (*[elim] -- "ignoring weak elimination rule" *): + "\succ[n] = 0; n \ Nat\ \ P" + "\0 = succ[n]; n \ Nat\ \ P" + by (auto simp: succIsSucc) + +lemma succNatZeroIff [simp]: + "n \ Nat \ (succ[n] = 0) = FALSE" + "n \ Nat \ (0 = succ[n]) = FALSE" + by (auto simp: succIsSucc) + +lemma succInj [dest]: + assumes eq: "succ[a] = succ[b]" and a: "a \ Int" and b: "b \ Int" + shows "a = b" +using a proof (rule pr_intCases) + assume a0: "a = 0" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" with a0 show ?thesis by simp + next + fix n + assume "n \ Nat" "b = Succ[n]" + with a0 eq show ?thesis by (simp add: succIsSucc) + next + fix n + assume "n \ Nat" "b = -Succ[n]" + with a0 eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +next + fix n + assume n: "n \ Nat" "a = Succ[n]" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" + with n eq show ?thesis by (simp add: succIsSucc) + next + fix m + assume "m \ Nat" "b = Succ[m]" + with n eq show ?thesis by (simp add: succIsSucc) + next + fix m + assume "m \ Nat" "b = -Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +next + fix n + assume n: "n \ Nat" "a = -Succ[n]" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + next + fix m + assume "m \ Nat" "b = Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + next + fix m + assume "m \ Nat" "b = -Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +qed -lemma notNat_imp_sgn_neg1 [intro]: "n \ Nat \ sgn(n) = -.1" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma succInjIff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a] = succ[b]) = (a = b)" + using assms by auto + +lemma intIsSucc: + assumes "a \ Int" + shows "\b \ Int : a = succ[b]" +using assms proof (rule pr_intCases) + assume "a = 0" + hence "a = succ[-1]" by simp + moreover have "-1 \ Int" by simp + ultimately show ?thesis by blast +next + fix n + assume n: "n \ Nat" "a = Succ[n]" + hence "a = succ[n]" by (simp add: succIsSucc) + with n show ?thesis by blast +next + fix n + assume n: "n \ Nat" "a = -Succ[n]" + hence "a = succ[-succ[Succ[n]]]" by simp + with n show ?thesis by blast +qed -lemma eqSgnNat_imp_nat: "sgn(m) = sgn(n) \ m \ Nat \ n \ Nat" -unfolding sgn_def isPos_eq_inNat1 by auto +definition pred where + "pred = [a \ Int \ CHOOSE b \ Int : a = succ[b]]" -lemma eqSgn_imp_0_nat [simp]: "n \ Nat \ sgn(n) = sgn(-.n) = (n = 0)" -unfolding sgn_def isPos_def by force +lemma predType: "pred \ [Int \ Int]" + unfolding pred_def by (auto intro: intIsSucc[THEN bChooseI2]) -lemma eqSgn_imp_0_nat2 [simp]: "n \ Nat \ sgn(-.n) = sgn(n) = (n = 0)" -unfolding sgn_def isPos_def by force +lemma predIsAFcn [intro!,simp]: "isAFcn(pred)" + using predType by blast -lemma eqSgn_imp_0 [simp]: "n \ Int \ sgn(n) = sgn(-.n) = (n = 0)" -by(rule intCases, auto) +\ \@{text "DOMAIN pred = Int"}\ +lemmas [intro!,simp] = funcSetDomain[OF predType] +\ \@{text "a \ Int \ pred[a] \ Int"}\ +lemmas [intro!,simp] = funcSetValue[OF predType] -(*lemma notSgnNegs [simp]: "n \ Int \\ {0} \ sgn(n) = sgn(-.n) = FALSE" - by(rule intCases[of n], simp_all) *) +lemma predValue: + assumes "b \ Int" and "a = succ[b]" + shows "b = pred[a]" + using assms unfolding pred_def by (auto intro: bChooseI) -lemma sgn_eq_neg1_is_not_nat (*[simp]*): "(sgn(n) = -.1) = (n \ Nat \ n \ 0)" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma pred0: "pred[0] = -1" + by (rule sym[OF predValue]) auto -lemma sgn_not_neg1_is_nat [simp]: "((sgn(n) = -.1) = FALSE) = (n \ Nat)" -by (auto simp: sgn_eq_neg1_is_not_nat) +lemma predSucc [simp]: + assumes "a \ Int" + shows "pred[succ[a]] = a" + by (rule predValue[OF assms, THEN sym, OF refl]) -lemma sgn_neg_eq_1_false: "\sgn(-.m) = 1; m \ Nat\ \ P" -unfolding sgn_def by auto +lemma succPred [simp]: + assumes "a \ Int" + shows "succ[pred[a]] = a" + using assms unfolding pred_def by (auto intro: intIsSucc[THEN bChooseI2]) -lemma sgn_minus [simp]: - assumes n: "n \ Int" - shows "sgn(-.n) = -.sgn(n)" -unfolding sgn_def using n by (cases, auto) +lemma predInj [dest]: + assumes eq: "pred[a] = pred[b]" and ab: "a \ Int" "b \ Int" + shows "a = b" +proof - + from eq have "succ[pred[a]] = succ[pred[b]]" + by simp + with ab show ?thesis by simp +qed -text \ Absolute value \ +lemma predInjIff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] = pred[b]) = (a = b)" + using assms by auto -lemma absIsNat [simp]: - assumes n: "n \ Int" shows "abs(n) \ Nat" -unfolding abs_def using intNotNatIsNegNat[OF _ n] by auto +lemma uminusSucc (*[simp]*): + assumes "a \ Int" + shows "-succ[a] = pred[-a]" +proof - + from assms have "succ[-succ[a]] = succ[pred[-a]]" + by simp + with assms show ?thesis + by blast +qed -lemma absNat [simp]: "n \ Nat \ abs(n) = n" -unfolding abs_def by auto +lemma uminusPred (*[simp]*): + assumes "a \ Int" + shows "-pred[a] = succ[-a]" + using assms by auto -lemma abs0 [simp]: "abs(0) = 0" -unfolding abs_def by simp +text \ + We restate the induction and case splitting theorems for + natural numbers and integers in terms of @{text "succ"} and + @{text "pred"}. +\ +lemma natInduct: + assumes "P(0)" and "\n. \ n \ Nat; P(n) \ \ P(succ[n])" + shows "\n \ Nat : P(n)" + using assms by (intro pr_natInduct) (auto simp: succIsSucc) -(** -lemma absn0 [simp]: "abs(-.0) = 0" -by simp -**) +\ \version of above suitable for the inductive reasoning package\ +lemma natInductE [case_names 0 Succ, induct set: Nat]: + assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(succ[n])" + shows "P(n)" +using bspec[OF natInduct, where P=P] assms by blast -lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" -unfolding abs_def by (auto dest: sgnNat_not1) +(*** EXAMPLE INDUCTION PROOFS *** -lemma abs_neg [simp]: - assumes n: "n \ Int" shows "abs(-.n) = abs(n)" -unfolding abs_def using n by (auto dest: sgnNat_not1) +lemma "\n\Nat : n=0 \ (\m \ Nat : n = succ[m])" +by (rule natInduct, auto) +lemma + assumes 1: "n \ Nat" + shows "n=0 \ (\m \ Nat : n = succ[m])" +using 1 by (induct, auto) -subsection \ Orders on integers \ +*** END EXAMPLE ***) -text \ - We distinguish four cases, depending on the arguments being in - Nat or negative. -\ +lemma natCases [case_names 0 succ, cases set: Nat]: + assumes n: "n \ Nat" + and z: "n=0 \ P" and sc: "\m. \m \ Nat; n = succ[m]\ \ P" + shows "P" +proof - + from n have "n=0 \ (\m \ Nat : n = succ[m])" + by (induct, auto) + thus ?thesis + proof + assume "n=0" thus "P" by (rule z) + next + assume "\m\Nat : n = succ[m]" + then obtain m where "m \ Nat" and "n = succ[m]" .. + thus "P" by (rule sc) + qed +qed -lemmas int_leq_pp_def = nat_leq_def - (* -- 'positive-positive' case, ie: both arguments are naturals *) +lemma not0_implies_Suc: + "\n \ Nat; n \ 0\ \ \m \ Nat: n = succ[m]" + by (rule natCases) auto + +text \Induction over two parameters along the ``diagonal''.\ +lemma diffInduction: + assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, succ[n])" + and step: "\m,n\Nat : P(m,n) \ P(succ[m], succ[n])" + shows "\m,n\Nat : P(m,n)" +proof (rule natInduct) + show "\n\Nat : P(0,n)" + using b1 b2 by (intro natInduct) auto +next + fix m + assume m: "m \ Nat" and ih: "\n\Nat : P(m,n)" + show "\n\Nat : P(succ[m],n)" + proof (rule bAllI) + fix n + assume "n \ Nat" thus "P(succ[m],n)" + proof (cases) + case 0 with b1 m show ?thesis by auto + next + case succ with step ih m show ?thesis by auto + qed + qed +qed -axiomatization where - int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = (a = 0 \ b = 0)" -and - int_leq_np_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ b = TRUE" -and - int_leq_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ -.b = (b \ a)" +lemma diffInduct: + assumes n: "n \ Nat" and m: "m \ Nat" + and b1: "\m. m\Nat \ P(m,0)" and b2: "\n. n\Nat \ P(0, succ[n])" + and step: "\m n. \m \ Nat; n\Nat; P(m,n) \ \ P(succ[m], succ[n])" + shows "P(m,n)" +proof - + have "\m,n\Nat : P(m,n)" + by (rule diffInduction, auto intro: b1 b2 step) + with n m show ?thesis by blast +qed -(* lemmas int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) +lemma intInduct: + assumes z: "P(0)" + and pos: "\n. \n \ Nat; P(n); P(-n)\ \ P(succ[n])" + and neg: "\n. \n \ Nat; P(n); P(-n)\ \ P(-succ[n])" + shows "\a\Int : P(a)" +proof - + from assms have 1: "\n \ Nat : P(n) \ P(-n)" + by (intro natInduct) auto + show ?thesis + proof + fix a + assume "a \ Int" + from this 1 show "P(a)" + by (elim intElim) auto + qed +qed -lemma int_boolify_leq [simp]: - "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" -by(rule intCases2[of a b], simp_all) +\ \turn the above lemma into an induction method\ +lemma intInductE [case_names 0 pos neg, induct set: Int]: + assumes "a \ Int" and "P(0)" + and "\n. \n \ Nat; P(n); P(-n)\ \ P(succ[n])" + and "\n. \n \ Nat; P(n); P(-n)\ \ P(-succ[n])" + shows "P(a)" + using bspec[OF intInduct, where P=P] assms by blast + +lemma intZeroPosNeg: + assumes "a \ Int" + shows "a = 0 \ (\n \ Nat : a = succ[n]) \ (\n \ Nat : a = -succ[n])" +using assms by (induct a) auto + +lemma intCases [case_names 0 pos neg, cases set: Int]: + assumes "a \ Int" + and "a = 0 \ P" + and "\n. \n \ Nat; a = succ[n]\ \ P" + and "\n. \n \ Nat; a = -succ[n]\ \ P" + shows "P" + using assms by (blast dest: intZeroPosNeg) + +lemma succIrrefl: + assumes "a \ Int" + shows "succ[a] \ a" + using assms by induct auto + +lemma succIrreflE (*[elim] -- don't: "ignoring weak elimination rule" *): + "\succ[a] = a; a \ Int\ \ P" + "\a = succ[a]; a \ Int\ \ P" +by (auto dest: succIrrefl) + +lemma succIrrefl_iff [simp]: + "a \ Int \ (succ[a] = a) = FALSE" + "a \ Int \ (a = succ[a]) = FALSE" +by (auto dest: succIrrefl) + +lemma predIrrefl: + assumes "a \ Int" + shows "pred[a] \ a" +proof + assume "pred[a] = a" + hence "succ[pred[a]] = succ[a]" by simp + with assms show "FALSE" by simp +qed -lemma int_leq_isBool [intro!,simp]: - "\a \ Int; b \ Int\ \ isBool(a \ b)" -unfolding isBool_def by auto +lemma predIrreflE (*[elim]*): + "\pred[a] = a; a \ Int\ \ P" + "\a = pred[a]; a \ Int\ \ P" +by (auto dest: predIrrefl) -lemma int_leq_refl [iff]: "n \ Int \ n \ n" -by(rule intCases, auto) +lemma predIrrefl_iff [simp]: + "a \ Int \ (pred[a] = a) = FALSE" + "a \ Int \ (a = pred[a]) = FALSE" +by (auto dest: predIrrefl) -lemma eq_leq_bothE: (* -- reduce equality over integers to double inequality *) - assumes "m \ Int" and "n \ Int" and "m = n" and "\m \ n; n \ m\ \ P" - shows "P" -using assms by simp +lemma predInNat [simp]: + assumes "a \ Int" + shows "(pred[a] \ Nat) = (a \ Nat \ {0})" + using assms by cases (auto simp: pred0 sym[OF uminusSucc]) -lemma neg_le_iff_le [simp]: - "\m \ Int; n \ Int \ \ -.n \ -.m = (m \ n)" - by (rule intCases2[of m n]) auto +lemma succNatNotNeg [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(succ[m] = -n) = FALSE" "(n = -succ[m]) = FALSE" + "(-n = succ[m]) = FALSE" "(-succ[m] = n) = FALSE" + using assms by auto -subsection \ Addition of integers \ +subsection \ Initial intervals of natural numbers \ text \ - Again, we distinguish four cases in the definition of @{text "a + b"}, - according to each argument being positive or negative. + The set of natural numbers up to (and including) a given $n$ is + inductively defined as the smallest set of natural numbers that + contains $n$ and all numbers whose successor is in the set. + + NB: ``less than'' is not first-order definable from the Peano axioms, + a set-theoretic definition such as the following seems to be unavoidable. \ -(* cf. NatArith *) -(** The following is rejected by Isabelle because the two definitions - are not distinguishable by argument types. -defs (unchecked overloaded) - int_add_def: "\a \ Int; b \ Int \ \ a + b \ - IF a \ Nat \ b \ Nat THEN addnat(a)[b] - ELSE IF isNeg(a) \ isNeg(b) THEN -.(addnat(-.a)[-.b]) - ELSE IF isNeg(a) THEN IF -.a \ b THEN b -- a ELSE -.(a -- b) - ELSE IF a \ -.b THEN -.(b -- a) ELSE a -- b" -**) +definition upto :: "c \ c" +where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : succ[k] \ S })" lemmas - int_add_pp_def = nat_add_def (* -- both numbers are positive, ie. naturals *) + setEqualI [where A = "upto(n)" for n, intro!] + setEqualI [where B = "upto(n)" for n, intro!] -axiomatization where -int_add_pn_def: "\a \ Nat; b \ Nat\ \ a + (-.b) \ IF a \ b THEN -.(b -- a) ELSE a -- b" -and -int_add_np_def: "\a \ Nat; b \ Nat\ \ (-.a) + b \ IF b \ a THEN -.(a -- b) ELSE b -- a" -and -int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" +lemma uptoNat: "upto(n) \ Nat" + unfolding upto_def by (rule lfpSubsetDomain) -lemmas int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) - (* -- When we use these definitions, we don't want to unfold the 'pp' case *) +lemma uptoPred: + assumes suc: "succ[m] \ upto(n)" and m: "m \ Nat" and n: "n \ Nat" + shows "m \ upto(n)" +proof - + let ?f = "\S. {n} \ {k\Nat : succ[k] \ S}" + from n have mono: "Monotonic(Nat, ?f)" + unfolding Monotonic_def by blast + from m suc have 1: "m \ ?f(upto(n))" by auto + from mono have 2: "?f(upto(n)) \ upto(n)" + unfolding upto_def by (rule lfpPreFP) + from 1 2 show ?thesis by blast +qed -lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" -by (auto simp: int_add_pn_def dest: nat_leq_antisym) +lemma uptoZero: "upto(0) = {0}" +proof (rule setEqual) + have "{0} \ { k \ Nat : succ[k] \ {0} } \ {0}" by auto + thus "upto(0) \ {0}" + unfolding upto_def by (rule lfpLB, auto) +next + show "{0} \ upto(0)" + unfolding upto_def by (rule lfpGLB, auto) +qed -text \ Closure \ +lemma uptoSucc: + assumes n: "n \ Nat" + shows "upto(succ[n]) = upto(n) \ {succ[n]}" (is "?lhs = ?rhs") +proof - + let ?preds = "\S. {k \ Nat : succ[k] \ S}" + let ?f = "\S k. {k} \ ?preds(S)" + have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" + by (auto simp: Monotonic_def) + \ \``$\subseteq$''\ + from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto + also have "\ \ upto(n)" + by (unfold upto_def, rule lfpPreFP, rule mono, rule n) + finally have "?f(?rhs, succ[n]) \ ?rhs" by auto + moreover from n have "?rhs \ Nat" + by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) + ultimately have 1: "?lhs \ ?rhs" + by (unfold upto_def[where n="succ[n]"], rule lfpLB) + \ \``$\supseteq$''\ + from n mono have 2: "?f(?lhs, succ[n]) \ ?lhs" + unfolding upto_def by (intro lfpPreFP) auto + with n have "?f(?lhs, n) \ ?lhs" by auto + moreover have "?lhs \ Nat" by (rule uptoNat) + ultimately have 3: "upto(n) \ ?lhs" + unfolding upto_def[where n=n] by (rule lfpLB) + from 2 have 4: "succ[n] \ ?lhs" by auto + from 3 4 have "?rhs \ ?lhs" by auto + with 1 show ?thesis by (rule setEqual) +qed -lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" -by (rule intCases2[of m n], auto simp: int_add_def) +lemma uptoRefl: + assumes n: "n \ Nat" + shows "n \ upto(n)" +using n proof (cases) + case 0 thus ?thesis by (simp add: uptoZero) +next + case (succ m) thus ?thesis by (auto simp: uptoSucc) +qed -text \ Neutral element \ +lemma zeroInUpto: + assumes n: "n \ Nat" + shows "0 \ upto(n)" +using n by (induct, auto simp: uptoZero uptoSucc) -lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" -by(rule intCases, auto simp add: int_add_np_def) -lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" -by(rule intCases, auto simp add: int_add_pn_def) +lemma SuccNotUptoZero: + assumes "n \ Nat" and "succ[n] \ upto(0)" + shows "P" +using assms by (auto simp: uptoZero) -text \ Additive inverse element \ +lemma uptoTrans: + assumes "k \ upto(m)" and "m \ upto(n)" and "n \ Nat" + shows "k \ upto(n)" +proof - + have "\n\Nat : m \ upto(n) \ upto(m) \ upto(n)" + by (rule natInduct, auto simp: uptoZero uptoSucc) + with assms show ?thesis by blast +qed -lemma add_inverse_nat [simp]: "n \ Nat \ n + -.n = 0" -by(simp add: int_add_pn_def) +lemma succNotinUpto: + assumes n: "n \ Nat" + shows "succ[n] \ upto(n)" +using n proof (induct) + show "1 \ upto(0)" by (auto simp: uptoZero) +next + fix n + assume n: "n \ Nat" and ih: "succ[n] \ upto(n)" + show "succ[succ[n]] \ upto(succ[n])" + proof (auto simp: uptoSucc n) + assume "succ[succ[n]] \ upto(n)" + with n have "succ[n] \ upto(n)" + by (auto elim: uptoPred) + with ih show "FALSE" .. + qed +qed -lemma add_inverse2_nat [simp]: "n \ Nat \ -.n + n = 0" -by(simp add: int_add_np_def) +lemma uptoLimit: + assumes m: "m \ upto(n)" and suc: "succ[m] \ upto(n)" and n: "n \ Nat" + shows "m=n" +proof - + from m uptoNat have mNat: "m \ Nat" by blast + from n have "\m\Nat: m \ upto(n) \ succ[m] \ upto(n) \ m=n" (is "?P(n)") + by (induct, auto simp: uptoZero uptoSucc) + with mNat m suc show ?thesis by blast +qed -lemma add_inverse_int [simp]: "n \ Int \ n + -.n = 0" -by (rule intCases, auto simp: int_add_def) +lemma uptoAntisym: + assumes mn: "m \ upto(n)" and nm: "n \ upto(m)" + shows "m=n" +proof - + from mn uptoNat have m: "m \ Nat" by blast + from nm uptoNat have n: "n \ Nat" by blast + have "\m,n\Nat : m \ upto(n) \ n \ upto(m) \ m=n" (is "\m,n\Nat : ?P(m,n)") + proof (rule natInduct) + show "\n\Nat : ?P(0,n)" by (auto simp: uptoZero) + next + fix m + assume m: "m \ Nat" and ih: "\n\Nat : ?P(m,n)" + show "\n\Nat : ?P(succ[m],n)" + proof (auto simp: uptoSucc m) + fix n + assume "succ[m] \ upto(n)" and "n \ upto(m)" + from this m have "succ[m] \ upto(m)" by (rule uptoTrans) + with m show "succ[m] = n" \ \contradiction\ + by (blast dest: succNotinUpto) + qed + qed + with m n mn nm show ?thesis by blast +qed -lemma add_inverse2_int [simp]: "n \ Int \ -.n + n = 0" -by (rule intCases, auto simp: int_add_def) +lemma uptoInj: + assumes n: "n \ Nat" and m: "m \ Nat" + shows "(upto(n) = upto(m)) = (n = m)" +proof (auto) + assume 1: "upto(n) = upto(m)" + from n have "n \ upto(n)" by (rule uptoRefl) + with 1 have "n \ upto(m)" by auto + moreover + from m have "m \ upto(m)" by (rule uptoRefl) + with 1 have "m \ upto(n)" by auto + ultimately + show "n = m" by (rule uptoAntisym) +qed -text \ Commutativity \ +lemma uptoLinear: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m \ upto(n) \ n \ upto(m)" (is "?P(m,n)") +using m proof induct + from n show "?P(0,n)" by (auto simp: zeroInUpto) +next + fix k + assume k: "k \ Nat" and ih: "?P(k,n)" + from k show "?P(succ[k],n)" + proof (auto simp: uptoSucc) + assume kn: "(succ[k] \ upto(n)) = FALSE" + show "n \ upto(k)" + proof (rule contradiction) + assume c: "n \ upto(k)" + with ih have "k \ upto(n)" by simp + from this kn n have "k = n" by (rule uptoLimit[simplified]) + with n have "n \ upto(k)" by (simp add: uptoRefl) + with c show "FALSE" .. + qed + qed +qed -lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" -by(simp add: int_add_def) +lemma uptoInduct: + assumes n: "n \ Nat" + and base: "P(0)" + and step: "\m. \m \ Nat; m \ upto(n) \ P(m); succ[m] \ upto(n)\ \ P(succ[m])" + shows "\m \ upto(n) : P(m)" +proof - + have "\m \ Nat: m \ upto(n) \ P(m)" + by (rule natInduct) (auto simp: base dest: step) + with uptoNat show ?thesis by auto +qed -lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" - by(rule intCases2[of m n], auto simp add: int_add_def add_commute_nat) -text \ Associativity \ +subsection \ Primitive recursive functions over natural numbers \ -lemma add_pn_eq_adiff [simp]: - "\m \ n; m \ Nat; n \ Nat\ \ m + -.n = -.(n -- m)" -by (simp add: int_add_def) +text \ + We justify the definition of function over natural numbers by + primitive recursion. The idea is to construct a sequence of + approximations, and then obtain the function by diagonalization. +\ -lemma adiff_add_assoc5: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ p; p \ m + n; m \ p -- n\ \ -.(p -- n -- m) = m + n -- p" -apply (induct p n rule: diffInduct) -using assms by (auto dest: nat_leq_antisym) +definition primrec_nat_upto where + "primrec_nat_upto(base, step, f, n) \ + isAFcn(f) \ DOMAIN f = Nat + \ f[0] = base + \ (\m \ Nat : succ[m] \ upto(n) \ f[succ[m]] = step(m, f[m]))" + +lemma primrec_nat_upto_deterministic: + assumes n: "n \ Nat" and m: "m \ upto(n)" + and f: "primrec_nat_upto(base, step, f, n)" + and g: "primrec_nat_upto(base, step, g, m)" + and k: "k \ upto(m)" + shows "f[k] = g[k]" +proof - + from m uptoNat have mNat: "m \ Nat" by auto + hence "\k \ upto(m): f[k] = g[k]" + proof (rule uptoInduct) + from f g show "f[0] = g[0]" unfolding primrec_nat_upto_def by simp + next + fix k + assume kNat: "k \ Nat" and ih: "k \ upto(m) \ f[k] = g[k]" + and sk: "succ[k] \ upto(m)" + from sk kNat g have "g[succ[k]] = step(k, g[k])" + unfolding primrec_nat_upto_def by simp + moreover + from sk m n have "succ[k] \ upto(n)" + by (rule uptoTrans) + with kNat f have "f[succ[k]] = step(k, f[k])" + unfolding primrec_nat_upto_def by simp + moreover + from sk kNat mNat have "k \ upto(m)" + by (rule uptoPred) + ultimately show "f[succ[k]] = g[succ[k]]" + using ih by simp + qed + with k show ?thesis by blast +qed -lemma adiff_add_assoc6: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ p; m + n \ p; p -- n \ m\ \ m -- (p -- n) = -.(p -- (m + n))" -apply (induct p n rule: diffInduct) -using assms by (auto dest: nat_leq_antisym) +lemma primrec_nat_upto_exists: + "\n \ Nat : \f : primrec_nat_upto(base, step, f, n)" +proof (rule natInduct) + have "primrec_nat_upto(base, step, [n \ Nat \ base], 0)" + unfolding primrec_nat_upto_def by (auto simp: uptoZero) + thus "\f : primrec_nat_upto(base, step, f, 0)" .. +next + fix n + assume n: "n \ Nat" + and ih: "\f : primrec_nat_upto(base, step, f, n)" + from n have ssn: "succ[succ[n]] \ n" + by induct (simp+) + from ih obtain f where f: "primrec_nat_upto(base, step, f, n)" .. + define g where "g \ [f EXCEPT ![succ[n]] = step(n, f[n])]" + from f n have "isAFcn(g) \ DOMAIN g = Nat \ g[0] = base" + by (simp add: primrec_nat_upto_def g_def) + moreover + { + fix m + assume 1: "m \ Nat" and 2: "succ[m] \ upto(n)" + have "m \ succ[n]" + proof + assume 3: "m = succ[n]" + from n have "n \ upto(succ[succ[n]])" + by (simp add: uptoSucc uptoRefl) + with 2 3 have "succ[succ[n]] = n" + by (auto dest: uptoAntisym) + with ssn show "FALSE" .. + qed + with f n 1 2 have "g[succ[m]] = step(m, g[m])" + by (auto simp: primrec_nat_upto_def g_def) + } + moreover + from f n have "g[succ[n]] = step(n, g[n])" + by (simp add: primrec_nat_upto_def g_def) + ultimately have "primrec_nat_upto(base, step, g, succ[n])" + using n by (auto simp: primrec_nat_upto_def uptoSucc) + thus "\f : primrec_nat_upto(base, step, f, succ[n])" .. +qed -lemma adiff_add_assoc7: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\p + n \ m; m \ n\ \ -.(m -- (p + n)) = n -- m + p" -apply (induct n m rule: diffInduct) -using assms by simp_all +lemma primrec_nat: + "\f : isAFcn(f) \ DOMAIN f = Nat + \ f[0] = base \ (\n \ Nat : f[succ[n]] = step(n,f[n]))" +proof - + from primrec_nat_upto_exists[THEN fcnConstruct] obtain F where + F: "isAFcn(F) \ DOMAIN F = Nat + \ (\n\Nat : primrec_nat_upto(base, step, F[n], n))" + by blast + define g where "g \ [n \ Nat \ F[n][n]]" + have "isAFcn(g)" "DOMAIN g = Nat" + unfolding g_def by (simp+) + moreover + from F have "g[0] = base" + by (simp add: primrec_nat_upto_def g_def) + moreover + have "\n \ Nat : g[succ[n]] = step(n, g[n])" + proof + fix n + assume n: "n \ Nat" + hence 1: "n \ upto(n)" "n \ upto(succ[n])" "succ[n] \ upto(succ[n])" + by (simp add: uptoSucc uptoRefl)+ + moreover + from F n have 2: "primrec_nat_upto(base, step, F[n], n)" + and 3: "primrec_nat_upto(base, step, F[succ[n]], succ[n])" + by auto + ultimately have 4: "F[succ[n]][n] = F[n][n]" + using n by (auto intro: primrec_nat_upto_deterministic[where n="succ[n]"]) + with n 1 3 show "g[succ[n]] = step(n, g[n])" + by (auto simp: g_def primrec_nat_upto_def) + qed + ultimately + show ?thesis by blast +qed -lemma adiff_add_assoc8: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ m; p \ m -- n; p \ m; m -- p \ n\ \ m -- n -- p = -.(n -- (m -- p))" -using adiff_add_assoc6[OF n p m] apply simp -using leq_adiff_right_add_left[OF _ p m n] add_commute_nat[OF p n] apply simp -by(rule adiff_adiff_left_nat[OF m n p]) +lemma bprimrec_nat: + assumes base: "base \ S" and step: "\n \ Nat : \x \ S : step(n,x) \ S" + shows "\f \ [Nat \ S] : f[0] = base \ (\n \ Nat: f[succ[n]] = step(n,f[n]))" +proof - + from primrec_nat[of base step] obtain f where + 1: "isAFcn(f)" and 2: "DOMAIN f = Nat" + and 3: "f[0] = base" and 4: "\n\Nat : f[succ[n]] = step(n,f[n])" + by blast + have "\n\Nat : f[n] \ S" + proof (rule natInduct) + from 3 base show "f[0] \ S" by simp + next + fix n + assume "n \ Nat" and "f[n] \ S" + with step 4 show "f[succ[n]] \ S" by force + qed + with 1 2 3 4 show ?thesis + by blast +qed -declare leq_neq_iff_less [simplified,simp] +definition nat_primrec where + "nat_primrec(S, base, step) \ + CHOOSE f \ [Nat \ S] : f[0] = base + \ (\n \ Nat : f[succ[n]] = step(n, f[n]))" -lemma int_add_assoc1: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (n + -.p) = (m + n) + -.p" +lemma nat_primrecE: + assumes f: "f = nat_primrec(S, base, step)" (is "f = ?g") + and base: "base \ S" + and step: "\n \ Nat : \x \ S : step(n,x) \ S" + and maj: "\ f \ [Nat \ S]; f[0] = base; \n \ Nat : f[succ[n]] = step(n, f[n]) \ \ P" + shows "P" proof - -have s1_1: "n \ p ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s1_1_asm: "n \ p" - have s2_1: "n + -.p = -.(p -- n)" - using s1_1_asm n p int_add_pn_def[of "n" "p"] by auto - have s2_2: "p \ m + n ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s2_2_asm: "p \ m + n" - have s3_1: "m + (n + -.p) = m -- (p -- n)" - proof - - have s4_1: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_2: "m + -.(p -- n) = m -- (p -- n)" - proof - - have s5_1: "p -- n \ m" - using s1_1_asm s2_2_asm - leq_adiff_left_add_right_equiv[of "n" "p" "m"] - p n m by auto - show ?thesis - using s5_1 p n m adiffIsNat - int_add_pn_def by auto - qed - show ?thesis - using s4_1 s4_2 by auto - qed - have s3_2: "(m + n) + -.p = (m + n) -- p" - using s2_2_asm p n m addIsNat int_add_pn_def by auto - have s3_3: "m -- (p -- n) = (m + n) -- p" - using s1_1_asm s2_2_asm - adiff_add_assoc3[of "n" "p" "m"] m n p by auto - show "m + (n + -.p) = (m + n) + -.p" - using s3_1 s3_2 s3_3 by auto - qed - have s2_3: "\ (p \ m + n) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s2_3_asm: "\ (p \ m + n)" - have s3_1: "m + n < p" - using s2_3_asm p m n addIsNat nat_not_leq by auto - have s3_2: "m + n \ p" - using s3_1 by (auto simp: less_def) - have s3_3: "(m + n) + -.p = -.(p -- (m + n))" - using s3_2 m n p addIsNat int_add_pn_def by auto - have s3_4: "p -- n \ m ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s3_4_asm: "p -- n \ m" - have s4_1: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_2: "m + -.(p -- n) = m -- (p -- n)" - using s3_4_asm p n m adiffIsNat int_add_pn_def - by auto - have s4_3: "m + (n + -.p) = m -- (p -- n)" - using s4_1 s4_2 by auto - have s4_4: "m -- (p -- n) = -.(p -- (m + n))" - using m n p s1_1_asm s3_2 s3_4_asm - adiff_add_assoc6 by auto - show "m + (n + -.p) = (m + n) + -.p" - using s3_3 s4_3 s4_4[symmetric] by auto - qed - have s3_5: "\ (p -- n \ m) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s3_5_asm: "\ (p -- n \ m)" - have s4_1: "m \ p -- n" - proof - - have s5_1: "m < p -- n" - using p n m adiffIsNat nat_not_leq s3_5_asm by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_3: "m + -.(p -- n) = -.((p -- n) -- m)" - using s4_1 s3_5_asm p n m adiffIsNat - int_add_pn_def[of "m" "p -- n"] - by auto - have s4_4: "m + (n + -.p) = -.((p -- n) -- m)" - using s4_2 s4_3 by auto - have s4_5: "(p -- n) -- m = p -- (n + m)" - using p n m adiff_adiff_left_nat by auto - have s4_6: "(p -- n) -- m = p -- (m + n)" - using n m add_commute_nat s4_5 by auto - have s4_7: "-.((p -- n) -- m) = -.(p -- (m + n))" - using s4_6 by auto - show "m + (n + -.p) = (m + n) + -.p" - using s4_7 s4_4 s3_3 by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s3_4 s3_5 nat_leq_isBool[of "p -- n" "m"] - isBoolTrueFalse[of "p -- n \ m"] - by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s2_2 s2_3 - nat_leq_isBool[of "p" "m + n"] - isBoolTrueFalse[of "p \ m + n"] - by auto - qed -have s1_2: "\ (n \ p) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s1_2_asm: "\ (n \ p)" - have s2_1: "n + -.p = n -- p" - proof - - have s3_1: "\n \ Nat; p \ Nat\ - \ - n + (-.p) \ - IF n \ p THEN -.(p -- n) ELSE n -- p" - using int_add_pn_def[of "n" "p"] by auto - have s3_2: "n + (-.p) \ n -- p" - using s3_1 s1_2_asm n p by auto - show ?thesis - using s3_2 by auto - qed - have s2_2: "\ ((m + n) \ p)" - proof - - have s3_1: "p < n" - using p n s1_2_asm nat_not_leq by auto - have s3_2: "p < m + n" - using m p n s3_1 trans_less_add2 by auto - show ?thesis - using s3_2 p m n addIsNat nat_not_leq by auto - qed - have s2_3: "(m + n) + -.p = (m + n) -- p" - proof - - have s3_1: "m + n \\in Nat" - using m n addIsNat by auto - have s3_2: "\(m + n) \\in Nat; p \\in Nat\ - \ - (m + n) + (-.p) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE ((m + n) -- p)" - using int_add_pn_def[of "m + n" "p"] by auto - have s3_3: "(m + n) + (-.p) \ (m + n) -- p" - using s3_1 p s3_2 s2_2 by auto - show ?thesis - using s3_3 by auto - qed - have s2_4: "(m + (n + -.p) = (m + n) + -.p) = - (m + (n -- p) = (m + n) -- p)" - using s2_1 s2_3 by auto - have s2_5: "(m + n) -- p = m + (n -- p)" - proof - - have s3_1: "p \ n" - proof - - have s4_1: "p < n" - using s1_2_asm p n nat_not_leq by auto - show ?thesis - using s4_1 by (auto simp: less_def) - qed - show ?thesis - using m n p s3_1 adiff_add_assoc[of "p" "n" "m"] by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s2_4 s2_5[symmetric] by auto - qed -show ?thesis - proof - - have s2_1: "((n \ p) = TRUE) \ ((n \ p) = FALSE)" - using nat_leq_isBool[of "n" "p"] isBoolTrueFalse[of "n \ p"] - by auto - show ?thesis - using s1_1 s1_2 s2_1 by auto - qed + from base step have "\ g \ [Nat \ S] : g[0] = base \ (\n \ Nat: g[succ[n]] = step(n,g[n]))" + by (rule bprimrec_nat) + hence "?g \ [Nat \ S] \ ?g[0] = base \ (\n \ Nat: ?g[succ[n]] = step(n,?g[n]))" + unfolding nat_primrec_def by (rule bChooseI2, auto) + with f maj show ?thesis by blast qed -lemma int_add_assoc2: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (-.p + n) = (m + -.p) + n" +lemma nat_primrecType: + assumes "base \ S" and "\n \ Nat : \x \ S : step(n,x) \ S" + shows "nat_primrec(S, base, step) \ [Nat \ S]" proof - -have s1_1: "(m + n) + -.p = (m + -.p) + n" - proof - - have s2_1: "m + n \ p ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s2_1_asm: "m + n \ p" - have s3_1: "(m + n) + -.p = -.(p -- (m + n))" - proof - - have s4_1: "\m + n \ Nat; p \ Nat\ - \ ((m + n) + (-.p)) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE - (m + n) -- p" - using int_add_pn_def[of "m + n" "p"] by auto - have s4_2: "(m + n) + -.p \ -.(p -- (m + n))" - using s4_1 s2_1_asm m n p addIsNat by auto - show ?thesis - using s4_2 by auto - qed - have s3_2: "m \ p" - using s2_1_asm add_leqD1 m n p by auto - have s3_3: "(m + -.p) + n = -.(p -- m) + n" - proof - - have s4_1: "m + -.p = -.(p -- m)" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ -.(p -- m)" - using s5_1 m p s3_2 by auto - show ?thesis - using s5_2 by auto - qed - show ?thesis - using s4_1 by auto - qed - have s3_4: "n \ p -- m ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_4_asm: "n \ p -- m" - have s4_1: "(m + -.p) + n = -.((p -- m) -- n)" - proof - - have s5_1: "-.(p -- m) + n = n + -.(p -- m)" - using p n m adiffIsNat add_commute_pn_nat by auto - have s5_2: "n + -.(p -- m) = -.((p -- m) -- n)" - proof - - have s6_1: "\n \ Nat; p -- m \ Nat\ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s6_2: "n + -.(p -- m) \ -.((p -- m) -- n)" - using s6_1 n m p adiffIsNat s3_4_asm by auto - show ?thesis - using s6_2 by auto - qed - show ?thesis - using s3_3 s5_1 s5_2 by auto - qed - have s4_2: "(p -- m) -- n = p -- (m + n)" - using p n m adiff_adiff_left_nat by auto - show "(m + n) + -.p = (m + -.p) + n" - using s3_1 s4_1 s4_2 by auto - qed - have s3_5: "\ (n \ p -- m) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_5_asm: "\ (n \ p -- m)" - have s4_1: "p -- m \ n" - proof - - have s5_1: "p -- m < n" - using s3_5_asm nat_not_leq p n m adiffIsNat by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "-.(p -- m) + n = n -- (p -- m)" - proof - - have s5_1: "-.(p -- m) + n = n + -.(p -- m)" - using p n m add_commute_pn_nat adiffIsNat by auto - have s5_2: "n + -.(p -- m) = n -- (p -- m)" - proof - - have s6_1: "\n \ Nat; p -- m \ Nat\ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s6_2: "n + -.(p -- m) \ n -- (p -- m)" - using s6_1 s3_5_asm p n m adiffIsNat by auto - show ?thesis - using s6_2 by auto - qed - show ?thesis - using s5_1 s5_2 by auto - qed - have s4_3: "n -- (p -- m) = -.(p -- (m + n))" - using adiff_add_assoc6 m n p s3_2 s2_1_asm - s4_1 add_commute_nat by auto - show "(m + n) + -.p = (m + -.p) + n" - using s3_1 s3_3 s4_2 s4_3[symmetric] by auto - qed - show ?thesis - using s3_4 s3_5 nat_leq_isBool[of "n" "p -- m"] - isBoolTrueFalse[of "n \ p -- m"] by auto - qed - have s2_2: "\ (m + n \ p) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s2_2_asm: "\ (m + n \ p)" - have s3_1: "(m + n) + -.p = (m + n) -- p" - proof - - have s4_1: "\(m + n) \ Nat; p \ Nat\ - \ ((m + n) + (-.p)) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE - (m + n) -- p" - using int_add_pn_def[of "m + n" "p"] by auto - have s4_2: "((m + n) + (-.p)) \ ((m + n) -- p)" - proof - - have s5_1: "m + n \\in Nat" - using m n addIsNat by auto - show "((m + n) + (-.p)) \ ((m + n) -- p)" - using s2_2_asm s4_1[OF s5_1 p] by auto - qed - show ?thesis - using s4_2 by auto - qed - have s3_2: "m \ p ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_2_asm: "m \ p" - have s4_1: "m + -.p = -.(p -- m)" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ -.(p -- m)" - using s5_1 m p s3_2_asm by auto - show ?thesis - using s5_2 by auto - qed - have s4_4: "p \ m + n" - proof - - have s5_1: "p < m + n" - using s2_2_asm p m n addIsNat nat_not_leq by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "n \ p -- m ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s4_2_asm: "n \ p -- m" - have s5_1: "(m + -.p) + n = -.((p -- m) -- n)" - proof - - have s6_1: "(m + -.p) + n = -.(p -- m) + n" - using s4_1 by auto - have s6_2: "-.(p -- m) + n = n + -.(p -- m)" - using p m n adiffIsNat add_commute_pn_nat by auto - have s6_3: "n + -.(p -- m) = -.((p -- m) -- n)" - proof - - have s7_1: "\n \ Nat; p -- m \ Nat - \ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s7_2: "n + -.(p -- m) \ -.((p -- m) -- n)" - using s7_1 p n m adiffIsNat s4_2_asm by auto - show ?thesis - using s7_2 by auto - qed - show ?thesis - using s6_1 s6_2 s6_3 by auto - qed - have s5_2: "-.((p -- m) -- n) = (m + n) -- p" - using adiff_add_assoc5 s4_4 s4_2_asm s3_2_asm - add_commute_nat - m n p by auto - show "(m + n) + -.p = (m + -.p) + n" - using s5_1 s5_2 s3_1 by auto - qed - have s4_3: "\ (n \ p -- m) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s4_3_asm: "\ (n \ p -- m)" - have s5_1: "(m + n) -- p = n -- (p -- m)" - using m n p adiff_add_assoc3[of "m" "p" "n"] - s3_2_asm s4_4 add_commute_nat - by auto - have s5_2: "(m + -.p) + n = n -- (p -- m)" - proof - - have s6_1: "(m + -.p) + n = -.(p -- m) + n" - using s4_1 by auto - have s6_2: "-.(p -- m) + n = n + -.(p -- m)" - using s6_1 n p m adiffIsNat - add_commute_pn_nat by auto - have s6_3: "n + -.(p -- m) = n -- (p -- m)" - proof - - have s7_1: "\n \ Nat; p -- m \ Nat - \ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s7_2: "n + -.(p -- m) \ n -- (p -- m)" - using s7_1 s4_3_asm n p m adiffIsNat - by auto - show ?thesis - using s7_2 by auto - qed - show ?thesis - using s6_1 s6_2 s6_3 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s5_1 s3_1 s5_2 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s4_2 s4_3 nat_leq_isBool[of "n" "p -- m"] - isBoolTrueFalse[of "n \ p -- m"] by auto - qed - have s3_3: "\ (m \ p) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_3_asm: "\ (m \ p)" - have s4_1: "p \ m" - proof - - have s5_1: "p < m" - using s3_3_asm p m nat_not_leq by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "(m + n) -- p = (m -- p) + n" - using adiff_add_assoc2 s4_1 p n m by auto - have s4_3: "m + -.p = m -- p" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ m -- p" - using s5_1 m p s3_3_asm by auto - show ?thesis - using s5_2 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s4_2 s3_1 s4_3 by auto - qed - show ?thesis - using s3_2 s3_3 nat_leq_isBool[of "m" "p"] - isBoolTrueFalse[of "m \ p"] by auto - qed - show ?thesis - using s2_1 s2_2 nat_leq_isBool[of "m + n" "p"] - isBoolTrueFalse[of "m + n \ p"] by auto - qed -have s1_2: "m + (n + -.p) = (m + -.p) + n" - using s1_1 int_add_assoc1[OF m n p] by auto -have s1_3: "n + -.p = -.p + n" - using n p add_commute_pn_nat[of "n" "p"] by auto -show ?thesis - using s1_2 s1_3 by auto + define f where "f \ nat_primrec(S, base, step)" + from f_def[THEN meta_to_obj_eq] assms show ?thesis + by (rule nat_primrecE) (simp add: f_def) qed -declare leq_neq_iff_less [simplified,simp del] - -lemma int_add_assoc3: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + -.(n + p) = m + -.n + -.p" -apply(rule nat_leq_cases[of "n + p" m]) - using assms apply simp_all - apply(rule nat_leq_cases[OF n m], simp_all) - apply(rule nat_leq_cases[of p "m -- n"], simp_all) - apply(rule adiff_adiff_left_nat[symmetric], simp+) - using adiff_add_assoc6 add_commute_nat[OF n p] apply simp - using adiff_add_assoc2[OF _ p n m, symmetric] apply (simp add: adiff_is_0_eq') - apply(rule nat_leq_cases[OF n m], simp_all) - apply(rule nat_leq_cases[of p "m -- n"], simp_all) - using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp - using adiff_add_assoc3[symmetric] add_commute_nat[OF n p] apply simp - using adiff_add_assoc2[symmetric] add_commute_nat[OF n p] apply simp -done - -lemma int_add_assoc4: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (n + p) = (-.m + n) + p" -using assms add_commute_int[of "-.m" "n + p"] add_commute_int[of "-.m" n] apply simp -apply(rule nat_leq_cases[of m "n + p" ], simp_all) - apply(rule nat_leq_cases[OF m n], simp_all) - apply(rule adiff_add_assoc2, simp+) - apply(simp add: add_commute_int[of "-.(m -- n)" p]) - apply(rule nat_leq_cases[of "m -- n" p], simp_all) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc3[symmetric]) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc5[symmetric]) - apply(rule nat_leq_cases[OF m n], simp_all) - apply(simp only: add_commute_nat[of n p]) - apply(rule adiff_add_assoc7, simp_all) - apply(simp add: add_commute_int[of "-.(m -- n)" p]) - apply(rule nat_leq_cases[of "m -- n" p], simp+) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc6[symmetric]) - apply(simp only: add_commute_nat[of n p]) - apply(simp add: add_commute_nat[of p n]) - apply(rule adiff_adiff_left_nat[symmetric], simp+) -done - -lemma leq_adiff_right_imp_0: - assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" - shows "p = 0" -using p h apply (induct) -using n by auto - -lemma int_add_assoc5: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (n + -.p) = -.m + n + -.p" -using assms -apply(simp add: add_commute_int[of "-.m" "n + -.p"] add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[OF p n], simp_all) - apply(rule nat_leq_cases[of m "n -- p"], simp+) - apply(rule nat_leq_cases[of m n], simp+) - apply(rule nat_leq_cases[of p "n -- m"], simp_all) - apply(rule adiff_commute_nat[OF n p m]) - apply(rule adiff_add_assoc8, simp+) - using nat_leq_trans[of n m "n -- p"] apply simp - using leq_adiff_right_imp_0[OF _ _ n p] nat_leq_antisym[of m n] apply simp - apply(rule nat_leq_cases[OF m n], simp_all) - apply(rule nat_leq_cases[of p "n -- m"], simp_all) - apply(rule adiff_add_assoc8[symmetric], simp_all) - using leq_adiff_left_add_right[OF _ p n m] - add_commute_nat[OF p m] - apply(simp add: adiff_add_assoc3) - apply(simp add: adiff_add_assoc4) - apply(rule nat_leq_cases[of m n], simp_all) - apply(rule nat_leq_cases[of p "n -- m"], simp+) - using nat_leq_trans[of n p "n -- m"] apply simp - using leq_adiff_right_imp_0[OF _ _ n m] apply simp - using nat_leq_antisym[of n p] apply simp - apply(rule minusInj, simp) - apply(rule adiff_add_assoc4[symmetric], simp+) - apply(simp add: adiff_add_assoc2[symmetric]) - apply(simp add: add_commute_nat) -done - -lemma int_add_assoc6: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (-.n + p) = -.(m + n) + p" -using assms - add_commute_int[of "-.n" p] - add_commute_int[of "-.m" "p + -.n"] - add_commute_int[of "-.(m + n)" p] apply simp -apply(rule nat_leq_cases[OF n p], simp_all) - apply(rule nat_leq_cases[of m "p -- n"], simp+) - apply(rule nat_leq_cases[of "m + n" p], simp+) - apply(simp only: add_commute_nat[of m n]) - apply(rule adiff_adiff_left_nat, simp_all) - apply(simp only: minus_sym[symmetric]) - apply(rule adiff_add_assoc5, simp_all) - apply(rule nat_leq_cases[of "m + n" p], simp_all) - apply(simp only: minus_sym) - apply(rule adiff_add_assoc6, simp_all) - apply(rule adiff_add_assoc3, simp_all) - apply(rule nat_leq_cases[of "m + n" p], simp_all) - apply(simp only: minus_sym) - apply(rule adiff_add_assoc7[symmetric], simp_all) - apply(simp add: add_commute_nat[of "n -- p" m]) - apply(rule adiff_add_assoc[symmetric], simp+) -done - -lemma add_assoc_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m + (n + p) = (m + n) + p" -using m n p -by (rule intCases3, - auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 - int_add_assoc4 int_add_assoc5 int_add_assoc6) - -text \ Minus sign distributes over addition \ - -lemma minus_distrib_pn_int [simp]: - "m \ Nat \ n \ Nat \ -.(m + -.n) = -.m + n" -apply(simp add: add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[of n m], simp_all) -done - -lemma minus_distrib_np_int [simp]: - "m \ Nat \ n \ Nat \ -.(-.m + n) = m + -.n" -by(simp add: add_commute_int) - -lemma int_add_minus_distrib [simp]: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m + n) = -.m + -.n" -by (rule intCases2[OF m n], simp_all) - - -subsection \ Multiplication of integers \ - -axiomatization where - int_mult_pn_def: "\a \ Nat; b \ Nat\ \ a * -.b = -.(a * b)" -and - int_mult_np_def: "\a \ Nat; b \ Nat\ \ -.a * b = -.(a * b)" -and - int_mult_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a * -.b = a * b" - -lemmas int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) - -text \ Closure \ - -lemma multIsInt [simp]: "\a \ Int; b \ Int\ \ a * b \ Int" -by (rule intCases2[of a b], simp_all add: int_mult_def) - -text \ Neutral element \ - -lemma mult_0_right_int [simp]: "a \ Int \ a * 0 = 0" -by (rule intCases[of a], simp_all add: int_mult_np_def) - -lemma mult_0_left_int [simp]: "a \ Int \ 0 * a = 0" -by (rule intCases[of a], simp_all add: int_mult_pn_def) - -text \ Commutativity \ - -lemma mult_commute_int: "\a \ Int; b \ Int\ \ a * b = b * a" -by (rule intCases2[of a b], simp_all add: int_mult_def mult_commute_nat) - -text \ Identity element \ - -lemma mult_1_right_int [simp]: "a \ Int \ a * 1 = a" -by (rule intCases[of a], simp_all add: int_mult_def) - -lemma mult_1_left_int [simp]: "a \ Int \ 1 * a = a" -by (rule intCases[of a], simp_all add: int_mult_def) - -text \ Associativity \ - -lemma mult_assoc_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m * (n * p) = (m * n) * p" -by(rule intCases3[OF m n p], simp_all add: mult_assoc_nat int_mult_def) - -text \ Distributivity \ - -lemma ppn_distrib_left_nat: (* ppn stands for m=positive, n=positive, p=negative *) - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n + -.p) = m * n + -.(m * p)" -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"]) - using assms apply(simp_all add: adiff_mult_distrib2_nat int_mult_def) -done - -lemma npn_distrib_left_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m * (n + -.p) = -.(m * n) + m * p" -using assms apply (simp add: add_commute_int[of "-.(m * n)" "m * p"]) -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"], simp_all) - apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) -done - -lemma nnp_distrib_left_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m * (-.n + p) = m * n + -.(m * p)" -using assms apply (simp add: add_commute_int[of "-.n" p]) -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"], simp_all) - apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) -done - -lemma distrib_left_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m * (n + p) = (m * n + m * p)" -apply(rule intCases3[OF m n p], - simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) - apply(rule add_mult_distrib_left_nat, assumption+) - apply(rule ppn_distrib_left_nat, assumption+) - apply(simp add: add_commute_int, rule ppn_distrib_left_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_left_nat)+ - apply(rule npn_distrib_left_nat, assumption+) - apply(rule nnp_distrib_left_nat, assumption+) - apply(simp only: add_mult_distrib_left_nat) -done - -lemma pnp_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + -.n) * p = m * p + -.(n * p)" -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - using assms apply(simp_all add: adiff_mult_distrib_nat int_mult_def) -done - -lemma pnn_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + -.n) * -.p = -.(m * p) + n * p" -using assms apply (simp add: add_commute_int[of "-.(m * p)" "n * p"]) -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - apply (auto simp: adiff_mult_distrib_nat int_mult_def dest: nat_leq_antisym) -done - -lemma npn_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(-.m + n) * -.p = m * p + -.(n * p)" -using assms apply (simp add: add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - apply (auto simp: adiff_mult_distrib_nat int_mult_def dest: nat_leq_antisym) -done - -lemma distrib_right_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "(m + n) * p = (m * p + n * p)" -apply(rule intCases3[OF m n p], - simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) - apply(rule add_mult_distrib_right_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) - apply(rule pnp_distrib_right_nat, assumption+) - apply(rule pnn_distrib_right_nat, assumption+) - apply(simp add: add_commute_int, rule pnp_distrib_right_nat, assumption+) - apply(rule npn_distrib_right_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) - apply(simp only: add_mult_distrib_right_nat) -done - -text \ Minus sign distributes over multiplication \ - -lemma minus_mult_left_int: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m * n) = -.m * n" -by (rule intCases2[OF m n], simp_all add: int_mult_def) - -lemma minus_mult_right_int: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m * n) = m * -.n" -by (rule intCases2[OF m n], simp_all add: int_mult_def) - - -subsection \ Difference of integers \ + +subsection \ Primitive recursive functions over the integers \ text \ - Difference over integers is simply defined as addition of the complement. - Note that this difference, noted @{text "-"}, is different from the - difference over natural numbers, noted @{text "--"}, even for two natural - numbers, because the latter cuts off at $0$. + We introduce a similar construction for functions defined + over the integers. We have two step functions for positive and + negative integers and allow both of them to refer to the value + of the function at the predecessor and its conjugate. \ -definition diff (infixl "-" 65) -where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" - -lemma diffIsInt [simp]: (* -- Closure *) - "\m \ Int; n \ Int\ \ m - n \ Int" -by (simp add: int_diff_def) - -lemma diff_neg_is_add [simp]: "\m \ Int; n \ Int\ \ m - -.n = m + n" -by (simp add: int_diff_def) - -lemma diff_0_right_int [simp]: "m \ Int \ m - 0 = m" -by (simp add: int_diff_def) +definition primrec_int_upto where + "primrec_int_upto(base, pstep, nstep, f, n) \ + isAFcn(f) \ DOMAIN f = Int + \ f[0] = base + \ (\m \ Nat : succ[m] \ upto(n) \ + f[succ[m]] = pstep(m, f[m], f[-m]) + \ f[-succ[m]] = nstep(m, f[m], f[-m]))" + +lemma primrec_int_upto_deterministic: + assumes n: "n \ Nat" and m: "m \ upto(n)" + and f: "primrec_int_upto(base, pstep, nstep, f, n)" + and g: "primrec_int_upto(base, pstep, nstep, g, m)" + and k: "k \ upto(m)" + shows "f[k] = g[k]" "f[-k] = g[-k]" +proof - + from m uptoNat have mNat: "m \ Nat" by auto + hence "\k \ upto(m): f[k] = g[k] \ f[-k] = g[-k]" + proof (rule uptoInduct) + from f g show "f[0] = g[0] \ f[-0] = g[-0]" + unfolding primrec_int_upto_def by simp + next + fix k + assume kNat: "k \ Nat" + and ih: "k \ upto(m) \ f[k] = g[k] \ f[-k] = g[-k]" + and sk: "succ[k] \ upto(m)" + from sk kNat g + have "g[succ[k]] = pstep(k, g[k], g[-k]) + \ g[-succ[k]] = nstep(k, g[k], g[-k])" + unfolding primrec_int_upto_def by simp + moreover + from sk m n have "succ[k] \ upto(n)" + by (rule uptoTrans) + with kNat f + have "f[succ[k]] = pstep(k, f[k], f[-k]) + \ f[-succ[k]] = nstep(k, f[k], f[-k])" + unfolding primrec_int_upto_def by simp + moreover + from sk kNat mNat have "k \ upto(m)" + by (rule uptoPred) + ultimately + show "f[succ[k]] = g[succ[k]] \ f[-succ[k]] = g[-succ[k]]" + using ih by simp + qed + with k show "f[k] = g[k]" "f[-k] = g[-k]" by (blast+) +qed -lemma diff_0_left_int [simp]: "n \ Int \ 0 - n = -.n" -by (simp add: int_diff_def) +lemma primrec_int_upto_exists: + "\n \ Nat : \f : primrec_int_upto(base, pstep, nstep, f, n)" +proof (rule natInduct) + have "primrec_int_upto(base, pstep, nstep, [n \ Int \ base], 0)" + unfolding primrec_int_upto_def by (auto simp: uptoZero) + thus "\f : primrec_int_upto(base, pstep, nstep, f, 0)" .. +next + fix n + assume n: "n \ Nat" + and ih: "\f : primrec_int_upto(base, pstep, nstep, f, n)" + from n have ssn: "succ[succ[n]] \ n" + by induct (simp+) + from ih obtain f where f: "primrec_int_upto(base, pstep, nstep, f, n)" .. + define g where "g \ [m \ Int \ + IF m = succ[n] THEN pstep(n, f[n], f[-n]) + ELSE IF m = -succ[n] THEN nstep(n, f[n], f[-n]) + ELSE f[m]]" + have "isAFcn(g) \ DOMAIN g = Int" + by (simp add: g_def) + moreover + from f n have "g[0] = base" + by (simp add: primrec_int_upto_def g_def) + moreover + { + fix m + assume 1: "m \ Nat" and 2: "succ[m] \ upto(n)" + have m1: "m \ succ[n]" + proof + assume 3: "m = succ[n]" + from n have "n \ upto(succ[succ[n]])" + by (simp add: uptoSucc uptoRefl) + with 2 3 have "succ[succ[n]] = n" + by (auto dest: uptoAntisym) + with ssn show "FALSE" .. + qed + from 1 n have m2: "m \ -succ[n]" by auto + from m1 m2 f n 1 2 + have "g[succ[m]] = pstep(m, g[m], g[-m]) + \ g[-succ[m]] = nstep(m, g[m], g[-m])" + by (auto simp: primrec_int_upto_def g_def) + } + moreover + from f n have "g[succ[n]] = pstep(n, g[n], g[-n])" + by (simp add: g_def) + moreover + from f n have "g[-succ[n]] = nstep(n, g[n], g[-n])" + by (simp add: g_def) + ultimately have "primrec_int_upto(base, pstep, nstep, g, succ[n])" + using n by (auto simp: primrec_int_upto_def uptoSucc) + thus "\f : primrec_int_upto(base, pstep, nstep, f, succ[n])" .. +qed -lemma diff_self_eq_0_int [simp]: "m \ Int \ m - m = 0" -by (simp add: int_diff_def) +lemma primrec_int: + "\f : isAFcn(f) \ DOMAIN f = Int + \ f[0] = base + \ (\n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]))" +proof - + from primrec_int_upto_exists[THEN fcnConstruct] obtain F where + F: "isAFcn(F) \ DOMAIN F = Nat + \ (\n\Nat : primrec_int_upto(base, pstep, nstep, F[n], n))" + by blast + define g where "g \ [i \ Int \ IF i \ Nat THEN F[i][i] ELSE F[-i][i]]" + have "isAFcn(g)" "DOMAIN g = Int" + unfolding g_def by (simp+) + moreover + from F have "g[0] = base" + by (simp add: primrec_int_upto_def g_def) + moreover + have "\n \ Nat : g[succ[n]] = pstep(n, g[n], g[-n]) + \ g[-succ[n]] = nstep(n, g[n], g[-n])" + (is "\n \ Nat : ?P(n)") + proof + fix n + assume n: "n \ Nat" + hence 1: "n \ upto(n)" "n \ upto(succ[n])" "succ[n] \ upto(succ[n])" + by (simp add: uptoSucc uptoRefl)+ + moreover + from F n have 2: "primrec_int_upto(base, pstep, nstep, F[n], n)" + and 3: "primrec_int_upto(base, pstep, nstep, F[succ[n]], succ[n])" + by auto + with 1 have 4: "F[succ[n]][n] = F[n][n]" "F[succ[n]][-n] = F[n][-n]" + using n by (auto intro: primrec_int_upto_deterministic[where n = "succ[n]"]) + moreover + from n have "g[-n] = F[n][-n]" + by (auto simp add: g_def) + ultimately show "?P(n)" + using 3 n by (auto simp: g_def primrec_int_upto_def) + qed + ultimately + show ?thesis by blast +qed -lemma neg_diff_is_diff [simp]: "\m \ Int; n \ Int\ \ -.(m - n) = n - m" -by (simp add: int_diff_def add_commute_int) +lemma bprimrec_int: + assumes base: "base \ S" + and pstep: "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and nstep: "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + shows "\f \ [Int \ S] : + f[0] = base + \ (\n \ Nat: f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat: f[-succ[n]] = nstep(n, f[n], f[-n]))" +proof - + from primrec_int[of base pstep nstep] obtain f where + 1: "isAFcn(f)" and 2: "DOMAIN f = Int" + and 3: "f[0] = base" + and 4: "\n\Nat : f[succ[n]] = pstep(n, f[n], f[-n])" + and 5: "\n\Nat : f[-succ[n]] = nstep(n, f[n], f[-n])" + by blast + from 3 4 5 base pstep nstep have "\n\Int : f[n] \ S" + by (intro intInduct) auto + with 1 2 3 4 5 show ?thesis + by blast +qed -lemma diff_nat_is_add_neg: "\m \ Nat; n \ Nat\ \ m - n = m + -.n" -by (simp add: int_diff_def) +definition int_primrec where + "int_primrec(S, base, pstep, nstep) \ + CHOOSE f \ [Int \ S] : f[0] = base + \ (\n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]))" + +lemma int_primrecE: + assumes f: "f = int_primrec(S, base, pstep, nstep)" (is "f = ?g") + and base: "base \ S" + and pstep: "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and nstep: "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + and maj: "\ f \ [Int \ S]; f[0] = base; + \n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n]); + \n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]) \ \ P" + shows "P" +proof - + from base pstep nstep + have "\g \ [Int \ S] : g[0] = base + \ (\n \ Nat: g[succ[n]] = pstep(n, g[n], g[-n])) + \ (\n \ Nat: g[-succ[n]] = nstep(n, g[n], g[-n]))" + by (rule bprimrec_int) + hence "?g \ [Int \ S] \ ?g[0] = base + \ (\n \ Nat: ?g[succ[n]] = pstep(n, ?g[n], ?g[-n])) + \ (\n \ Nat: ?g[-succ[n]] = nstep(n, ?g[n], ?g[-n]))" + unfolding int_primrec_def by (rule bChooseI2) auto + with f maj show ?thesis by blast +qed -(** -lemma diff_neg_nat_is_add [simp]: "\m \ Nat; n \ Nat\ \ m - -.n = m + n" -by simp -**) +lemma int_primrecType: + assumes "base \ S" + and "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + shows "int_primrec(S, base, pstep, nstep) \ [Int \ S]" +proof - + define f where "f \ int_primrec(S, base, pstep, nstep)" + from f_def[THEN meta_to_obj_eq] assms show ?thesis + by (rule int_primrecE) (simp add: f_def) +qed end diff --git a/isabelle/README.html b/isabelle/README.html index ab8e434b..77e5d740 100644 --- a/isabelle/README.html +++ b/isabelle/README.html @@ -9,9 +9,9 @@ -

TLA+: Lamport's Temporal Logic of Actions specification language

+

TLA+: Lamport's specification language based on the Temporal Logic of Actions

-

TLA +

TLA is a linear-time temporal logic introduced by Leslie Lamport in The Temporal Logic of Actions (ACM TOPLAS 16(3), 1994, 872-923). Unlike other temporal logics, both systems and properties @@ -20,10 +20,14 @@

TLA+: Lamport's Temporal Logic of Actions specification language

structural relations such as refinement, parallel composition, and hiding.

-

TLA+ is a language for specifying and verifying concurrent and distributed systems. It is based on a variant of Zermelo-Fraenkel set theory for describing the data structures manipulated by the algorithms to be verified, and on TLA for describing their dynamic behavior. TLA+ has been applied to numerous case studies. +

TLA+ is a language for specifying and verifying concurrent and distributed +systems. It is based on a variant of Zermelo-Fraenkel set theory for describing +the data structures manipulated by the algorithms to be verified, and on TLA for +describing their dynamic behavior. TLA+ has been applied to numerous case +studies.

-

This directory formalizes TLA+ in Isabelle (version 2009-1), as follows: +

This directory formalizes TLA+ in Isabelle (version 2021-1), as follows:

  • PredicateLogic theory @@ -44,14 +48,14 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

    The FixedPoints theory develops the Knaster-Tarski theorems for least and greatest fixed points in the subset lattice. This development is used for the construction of - natural numbers; it also serves as a test for the encoding of TLA+ + integers; it also serves as a test for the encoding of TLA+ set theory.
  • Functions theory: functions in TLA+ are not defined (e.g., as sets of pairs), - but axiomatized, and in fact, pairs and tuples are later defined as special + but axiomatized, and in fact, pairs and tuples are later defined as functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined @@ -59,19 +63,32 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

  • - Peano theory: Peano's axioms for natural numbers. We prove - the existence of a structure satisfying these axioms and derive the - set of natural numbers. + Integers theory: the structure of integer numbers is introduced + using an extended version of Peano's axioms. We prove the existence of a + structure satisfying these axioms and define the integers as some such + structure (using CHOOSE expressions), together with a subet of + natural numbers. We define the successor and predecessor functions over + integers and prove induction rules for natural numbers and integers. + We also justify the definition of functions over integers by primitive + recursion.
  • - NatOrderings theory: introduces the order <= over - natural numbers. It also defines a < b (for arbitrary - values) as a <= b & a # b. Proves many lemmas over - these orders. Also defines intervals m .. n for natural - numbers m and n. + IntegerArithmetic theory: introduces addition (including + subtraction) and multiplication over integers, as well as the order + <= over integers and a < b (for arbitrary + values) as a <= b & a # b. Proves many lemmas for + integer arithmetic. Also defines intervals m .. n for + integers m and n.
  • +
  • + IntegerDivision theory: introduces integer division and + modulus, as well as the divisibility relation. In accordance with + the definition in the standard library of TLA+, div and + % (modulus) are defined only if the second argument is a + positive integer. +
  • The Tuples theory defines tuples and relations in TLA+. Tuples are functions whose domains are intervals of the form @@ -104,18 +121,6 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

    in theory Tuples.
  • -
  • - NatArith theory: addition, subtraction, and multiplication - over natural numbers. The subtraction (noted --) is defined - such that the result is cut off at 0: this is different from - the standard difference over integers. -
  • - -
  • - NatDivision theory: divisibility relation, division and - modulus for natural numbers. -
  • -

@@ -142,7 +147,5 @@

TLA+: Lamport's Temporal Logic of Actions specification language

Stephan Merz
- -Last modified: Mon Jun 8 11:15:49 CET 2009 - +Last modified: Tue Mar 15 15:05:26 CET 2022 diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 42152c84..da52b7a4 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -19,23 +19,23 @@ text \ \ definition Nibble -(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) +(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) where "Nibble \ {0, 1, - Succ[1], - Succ[Succ[1]], - Succ[Succ[Succ[1]]], - Succ[Succ[Succ[Succ[1]]]], - Succ[Succ[Succ[Succ[Succ[1]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]]]}" + succ[1], + succ[succ[1]], + succ[succ[succ[1]]], + succ[succ[succ[succ[1]]]], + succ[succ[succ[succ[succ[1]]]]], + succ[succ[succ[succ[succ[succ[1]]]]]], + succ[succ[succ[succ[succ[succ[succ[1]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]]]]}" definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" @@ -80,8 +80,8 @@ parse_ast_translation \ (* convert an ML integer to a nibble *) fun mkNibble n = if n = 0 - then Ast.Constant "Peano.zero" - else Ast.Appl [Ast.Constant "Functions.fapply", Ast.Constant "Peano.Succ", mkNibble (n-1)]; + then Ast.Constant "Integers.zero" + else Ast.Appl [Ast.Constant "Functions.fapply", Ast.Constant "Integers.succ", mkNibble (n-1)]; (* convert an ML character to a TLA+ Char *) fun mkChar c = @@ -139,14 +139,14 @@ oops print_ast_translation \ let (* convert a nibble to an ML integer *) + (* version to be used if 2 .. 15 are definitions, not abbreviations *) fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 | destNibble (Ast.Appl [Ast.Constant @{const_syntax "Functions.fapply"}, - Ast.Constant @{const_syntax "Peano.Succ"}, nb]) + Ast.Constant @{const_syntax "Integers.succ"}, nb]) = (destNibble nb) + 1 | destNibble _ = raise Match - - (* the following version should be used when 2 .. 15 are abbreviations, not definitions *) + (* version to be used when 2 .. 15 are abbreviations, not definitions *) (* fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 @@ -247,11 +247,11 @@ text \ (** Examples ** lemma "(''foo'' :> 1 @@ ''bar'' :> TRUE) \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by auto +by auto (* slow *) lemma "r \ [''bar'' : BOOLEAN, ''foo'' : Nat] \ [r EXCEPT ![''foo''] = 3] \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by (force simp: two_def three_def) (* "by auto" also works, but is slow *) +by (force simp: two_def three_def) (* even slower, "auto" also works *) lemma "(''a'' :> 1) \ (''b'' :> 1)" by simp @@ -260,7 +260,7 @@ lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1)" by simp lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1 @@ ''b'' :> 3)" -by (simp add: three_def) +by (simp add: two_def three_def) lemma "(''a'' :> 1 @@ ''b'' :> 2) = (''b'' :> 2 @@ ''a'' :> 1)" by simp @@ -278,7 +278,7 @@ by (simp add: two_def) end -(* NB: Make sure that the following are proved automatically: +(* NB: Make sure that the following is proved automatically once concatenation is defined: THEOREM Thm1 == "ab" = "a" \o "b" OBVIOUS diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5ecc293f..5aa1bd04 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1,14 +1,14 @@ (* Title: TLA+/Tuples.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2022 Inria and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Tuples and Relations in \tlaplus{} \ theory Tuples -imports NatOrderings +imports IntegerArithmetic begin text \ @@ -16,10 +16,8 @@ text \ functions whose domains are intervals of the form $1 .. n$, for some natural number $n$, and relations are sets of tuples. In particular, \tlaplus{} distinguishes between a function and its graph, and we have - functions to convert between the two. (This is useful, for example, - when defining functions recursively, as we have a fixed point theorem - on sets but not on functions.) We also introduce standard notions for - binary relations, such as orderings, equivalence relations and so on. + functions to convert between the two. We also introduce standard notions + for binary relations, such as orderings, equivalence relations and so on. \ subsection \ Sequences and Tuples \ @@ -41,96 +39,104 @@ where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: "isBool(isASeq(s))" -by (simp add: isASeq_def) + by (simp add: isASeq_def) lemma boolifyIsASeq [simp]: "boolify(isASeq(s)) = isASeq(s)" -by auto + by auto lemma isASeqI [intro (*,simp*)]: assumes "isAFcn(s)" and "n \ Nat" and "DOMAIN s = 1 .. n" shows "isASeq(s)" -using assms by (auto simp: isASeq_def) + using assms by (auto simp: isASeq_def) lemma SeqIsASeq [elim!]: assumes "s \ Seq(S)" shows "isASeq(s)" -using assms by (auto simp: Seq_def) + using assms by (auto simp: Seq_def) lemma LenI [intro]: - assumes "DOMAIN s = 1 .. n" and "n \ Nat" + assumes dom: "DOMAIN s = 1 .. n" and n: "n \ Nat" shows "Len(s) = n" -proof (unfold Len_def, rule bChooseI2) - from assms show "\x \ Nat : DOMAIN s = 1 .. x" by blast +unfolding Len_def proof (rule bChooseI2) + from assms show "\m \ Nat : DOMAIN s = 1 .. m" + by blast next fix m assume "m \ Nat" and "DOMAIN s = 1 .. m" - with assms show "m = n" by auto + with dom n show "m = n" + by (auto simp: intvlEqual_iff) qed lemma isASeqE [elim]: - assumes "isASeq(s)" - and "\isAFcn(s); DOMAIN s = 1 .. Len(s); Len(s) \ Nat\ \ P" - shows "P" -using assms by (auto simp: isASeq_def dest: LenI) + assumes s: "isASeq(s)" + and p: "\isAFcn(s); DOMAIN s = 1 .. Len(s); Len(s) \ Nat\ \ P" + shows "P" + using assms unfolding isASeq_def by (blast dest: LenI) +\ \The corresponding lemma for @{text "Seq(S)"} will be proved later.\ lemma SeqIsAFcn (*[elim!]*): assumes "isASeq(s)" shows "isAFcn(s)" -using assms by auto + using assms by (rule isASeqE) \ \ @{text "s \ Seq(S) \ isAFcn(s)"} \ lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" - shows "Len(s) \ Nat" -using assms by (rule isASeqE) + shows "Len(s) \ Int" "0 \ Len(s)" + using assms by (rule isASeqE, simp)+ -\ \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ -lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] +lemma LenInNat' [simp]: + assumes "s \ Seq(S)" + shows "Len(s) \ Int" "0 \ Len(s)" + using assms[THEN SeqIsASeq] by simp+ lemma DomainSeqLen [simp]: assumes "isASeq(s)" shows "DOMAIN s = 1 .. Len(s)" -using assms by auto + using assms by auto \ \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ -lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] +lemmas DomainSeqLen' [simp] = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: assumes "isASeq(s)" and "isASeq(t)" and "Len(s) = Len(t)" and "\k \ 1 .. Len(t) : s[k] = t[k]" shows "s = t" -using assms by (intro fcnEqual[of s t], auto) + using assms by (intro fcnEqual[of s t]) (simp add: SeqIsAFcn)+ lemma seqEqualE: assumes "isASeq(s)" and "isASeq(t)" and "s=t" and "\ Len(s) = Len(t); \k \ 1 .. Len(t) : s[k] = t[k] \ \ P" shows "P" -using assms by auto + using assms by auto lemma seqEqualIff: assumes "isASeq(s)" and "isASeq(t)" shows "(s = t) = (Len(s) = Len(t) \ (\k \ 1 .. Len(t) : s[k] = t[k]))" -by (auto elim: seqEqualI[OF assms] seqEqualE[OF assms]) + by (auto elim: seqEqualI[OF assms] seqEqualE[OF assms]) -lemma SeqI [intro!]: - assumes "isASeq(s)" and "\k. k \ 1 .. Len(s) \ s[k] \ S" +lemma SeqI [intro]: + assumes s: "isASeq(s)" and "\k. k \ 1 .. Len(s) \ s[k] \ S" shows "s \ Seq(S)" -using assms by (auto simp: Seq_def) +proof - + from assms have "s \ [1 .. Len(s) \ S]" by auto + with s show ?thesis unfolding Seq_def by auto +qed lemma SeqI': \ \ closer to the definition but probably less useful \ assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" -using assms by (auto simp: Seq_def) + using assms by (auto simp: Seq_def) lemma SeqE [elim]: assumes s: "s \ Seq(S)" and p: "\s \ [1 .. Len(s) \ S]; Len(s) \ Nat\ \ P" shows "P" proof (rule p) - from s show "Len(s) \ Nat" by (rule LenInNat') + from s show "Len(s) \ Nat" by simp next from s obtain n where "n \ Nat" and "s \ [1 .. n \ S]" by (auto simp: Seq_def) @@ -140,20 +146,20 @@ qed lemma seqFuncSet: assumes "s \ Seq(S)" shows "s \ [1 .. Len(s) \ S]" -using assms by (rule SeqE) + using assms by (rule SeqE) -lemma seqElt [elim!]: - assumes "s \ Seq(S)" and "n \ Nat" and "1 \ n" and "n \ Len(s)" +lemma seqElt [elim]: + assumes "s \ Seq(S)" and "n \ Int" and "0 < n" and "n \ Len(s)" shows "s[n] \ S" -using assms by auto + using assms by auto lemma seqInSeqRange: assumes "isASeq(s)" shows "s \ Seq(Range(s))" -using assms by auto + using assms by auto lemma isASeqInSeq: "isASeq(s) = (\S: s \ Seq(S))" -by (auto elim: seqInSeqRange) + by (auto elim: seqInSeqRange) subsection \ Sequences via @{text emptySeq} and @{text Append} \ @@ -170,15 +176,15 @@ notation (ASCII) emptySeq ("(<< >>)") definition Append :: "[c,c] \ c" -where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len(s)] THEN e ELSE s[k]]" +where "Append(s,e) \ [k \ 1 .. succ[Len(s)] \ IF k = succ[Len(s)] THEN e ELSE s[k]]" -lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" -by (auto simp: emptySeq_def isASeq_def) +lemma emptySeqIsASeq [iff]: "isASeq(\\)" + unfolding emptySeq_def by force \ \ @{text "isAFcn(\\)"} \ -lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] +lemmas emptySeqIsAFcn [iff] = emptySeqIsASeq[THEN SeqIsAFcn] -lemma lenEmptySeq [simp]: "Len(\\) = 0" +lemma lenEmptySeq [iff]: "Len(\\) = 0" by (auto simp: emptySeq_def) lemma emptySeqInSeq (*[simp,intro!]*): "\\ \ Seq(S)" @@ -187,101 +193,105 @@ by auto lemma SeqNotEmpty [simp]: "(Seq(S) = {}) = FALSE" "({} = Seq(S)) = FALSE" -by auto + by (force intro: emptySeqInSeq)+ -lemma appendIsASeq [simp,intro!]: +lemma appendIsASeq [iff]: assumes s: "isASeq(s)" shows "isASeq(Append(s,e))" -using s unfolding Append_def -by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) + using s unfolding Append_def + by (rule isASeqE, intro isASeqI, auto) + +\ \@{text "isASeq(s) \ isAFcn(Append(s,e))"}\ +lemmas appendIsAFcn [iff] = appendIsASeq[THEN SeqIsAFcn] -\ \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ -lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] +\ \@{text "s \ Seq(S) \ isAFcn(Append(s,e))"}\ +lemmas appendIsAFcn' [iff] = SeqIsASeq[THEN appendIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" -by (simp add: emptySeq_def) + by (simp add: emptySeq_def) -lemma domainAppend [simp]: "DOMAIN Append(s,e) = 1 .. Succ[Len(s)]" -by (simp add: Append_def) +lemma domainAppend [simp]: "DOMAIN Append(s,e) = 1 .. succ[Len(s)]" + by (simp add: Append_def) -lemma isEmptySeq [intro!]: +lemma isEmptySeq [intro]: "\isAFcn(f); DOMAIN f = {}\ \ f = \\" "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" -by auto + by auto \ \ immediate consequence of @{text isEmptySeq} \ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" -by auto + by auto \ \ The reverse equation could be a useful rewrite rule (it is applied by TLC) \ lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" -by auto + by auto lemma seqLenZeroIsEmpty (*used to be [simp], but emptySeqIff seems more useful*): assumes "isASeq(s)" shows "(Len(s) = 0) = (s = \\)" -using assms by auto + using assms by auto lemma emptySeqIff [simp]: assumes "isAFcn(s)" shows "(s = \\) = (DOMAIN s = {} \ Len(s) = 0)" -using assms by auto + using assms by auto lemma emptySeqIff' [simp]: assumes "isAFcn(s)" shows "(\\ = s) = (DOMAIN s = {} \ Len(s) = 0)" -using assms by auto + using assms by auto lemma lenAppend [simp]: assumes "isASeq(s)" - shows "Len(Append(s,e)) = Succ[Len(s)]" -using assms by (intro LenI, auto simp: Append_def) + shows "Len(Append(s,e)) = succ[Len(s)]" + using assms unfolding Append_def by (intro LenI) force+ -\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ +\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = succ[Len(s)]"} \ lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] +lemma appendNotEmpty [simp]: + assumes "isASeq(s)" + shows "(Append(s,e) = \\) = FALSE" + using assms by auto + lemma appendElt [simp]: - assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Succ[Len(s)]" - shows "Append(s,e)[k] = (IF k = Succ[Len(s)] THEN e ELSE s[k])" -using assms by (auto simp: Append_def) + assumes "isASeq(s)" and "k \ Int" and "0 < k" and "k \ succ[Len(s)]" + shows "Append(s,e)[k] = (IF k = succ[Len(s)] THEN e ELSE s[k])" + using assms by (auto simp: Append_def) lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt] -lemma appendElt1 (*[simp]*): - assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Len(s)" +lemma appendElt1 [simp]: + assumes "isASeq(s)" and "k \ Int" and "0 < k" and "k \ Len(s)" shows "Append(s,e)[k] = s[k]" -using assms by (auto simp: Append_def) + using assms unfolding Append_def by (force simp: int_leq_succI) -lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1] +lemmas appendElt1' [simp] = SeqIsASeq[THEN appendElt1] -lemma appendElt2 (*[simp]*): +lemma appendElt2 [simp]: assumes "isASeq(s)" - shows "Append(s,e)[Succ[Len(s)]] = e" -using assms by (auto simp: Append_def) + shows "Append(s,e)[succ[Len(s)]] = e" + using assms unfolding Append_def by force -lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2] +lemmas appendElt2' [simp] = SeqIsASeq[THEN appendElt2] -lemma isAppend [intro!]: - assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. Succ[Len(s)]" and s: "isASeq(s)" - and elt1: "\n \ 1 .. Len(s) : f[n] = s[n]" and elt2: "f[Succ[Len(s)]] = e" +lemma isAppend [intro]: + assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. succ[Len(s)]" and s: "isASeq(s)" + and elt1: "\n \ 1 .. Len(s) : f[n] = s[n]" and elt2: "f[succ[Len(s)]] = e" shows "f = Append(s,e)" proof (rule fcnEqual[OF f]) - from s show "isAFcn(Append(s,e))" by simp -next - from dom show "DOMAIN f = DOMAIN Append(s,e)" by simp -next from s elt1 elt2 show "\x \ DOMAIN Append(s, e) : f[x] = Append(s, e)[x]" - by (auto simp: Append_def) -qed + unfolding Append_def by force +qed (simp add: s dom)+ -lemmas isAppend' [intro!] = isAppend[symmetric] +lemmas isAppend' [intro] = isAppend[symmetric] lemma appendInSeq [simp]: assumes s: "s \ Seq(S)" and e: "e \ S" shows "Append(s,e) \ Seq(S)" -using assms by (force simp: nat_leq_Succ) + using assms by (intro SeqI) auto lemma appendD1: assumes s: "isASeq(s)" and t: "isASeq(t)" and app: "Append(s,e) = Append(t,f)" @@ -292,15 +302,15 @@ proof - from s have 1: "isASeq(?s1)" by simp from t have 2: "isASeq(?t1)" by simp from 1 2 app have len: "Len(?s1) = Len(?t1)" and elt: "\k \ 1 .. Len(?t1) : ?s1[k] = ?t1[k]" - by (blast elim: seqEqualE)+ - from s t len have ls: "Len(s) = Len(t)" by simp + by (auto elim: seqEqualE) + from s t len have ls: "Len(s) = Len(t)" by auto thus ?thesis - proof (rule seqEqualI[OF s t], auto) + proof (rule seqEqualI[OF s t], clarify) fix k assume k: "k \ 1 .. Len(t)" - with s ls have "s[k] = ?s1[k]" by (intro sym[OF appendElt1], auto) - also from k elt t have "... = ?t1[k]" by auto - also from t k have "... = t[k]" by (intro appendElt1, auto) + with s ls have "s[k] = ?s1[k]" by (intro sym[OF appendElt1]) auto + also from k elt t have "\ = ?t1[k]" by auto + also from t k have "... = t[k]" by (intro appendElt1) auto finally show "s[k] = t[k]" . qed qed @@ -314,15 +324,17 @@ proof - from s have 1: "isASeq(?s1)" by simp from t have 2: "isASeq(?t1)" by simp from 1 2 app have "Len(?s1) = Len(?t1)" and "\k \ 1 .. Len(?t1) : ?s1[k] = ?t1[k]" - by (blast elim: seqEqualE)+ - with s t have "?s1[Len(?s1)] = ?t1[Len(?t1)]" by auto + by (auto elim: seqEqualE)+ + moreover + from s t app have "Len(?s1) \ 1 .. Len(?t1)" by auto + ultimately have "?s1[Len(?s1)] = ?t1[Len(?t1)]" by auto with s t show ?thesis by simp qed lemma appendEqualIff [simp]: assumes s: "isASeq(s)" and t: "isASeq(t)" shows "(Append(s,e) = Append(t,f)) = (s = t \ e = f)" -using appendD1[OF s t] appendD2[OF s t] by auto + using appendD1[OF s t] appendD2[OF s t] by auto text \ The following lemma gives a possible alternative definition of @@ -331,18 +343,35 @@ text \ lemma appendExtend: assumes "isASeq(s)" - shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" -using assms by (intro isAppend') auto + shows "Append(s,e) = s @@ (succ[Len(s)] :> e)" (is "_ = ?rhs") + using assms by (intro isAppend') force+ lemma imageEmptySeq [simp]: "Image(\\, A) = {}" -by (simp add: emptySeq_def) + by (simp add: emptySeq_def) lemma imageAppend [simp]: assumes s: "isASeq(s)" shows "Image(Append(s,e), A) = - (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" -unfolding appendExtend[OF s] -using assms by auto (force+) + (IF succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" +proof - + { + fix i + assume "i \ A" "i \ Int" "0 < i" "i \ Len(s)" + with s have "\j \ A : j \ Int \ 1 \ j \ j \ Len(s) \ + s[i] = (s @@ succ[Len(s)] :> e)[j]" + by auto + } + moreover + { + fix i + assume "i \ A" "i \ Int" "0 < i" "i \ Len(s)" + with s have "\j \ A : j \ Int \ 1 \ j \ j \ Len(s) \ + s[i] = s[j]" + by auto + } + ultimately show ?thesis + unfolding appendExtend[OF s] using s by auto +qed text \ Inductive reasoning about sequences, based on @{term "\\"} and @@ -357,21 +386,21 @@ lemma seqInduct [case_names empty append, induct set: Seq]: proof - have "\n \ Nat : \s \ [1 .. n \ S] : P(s)" (is "\n \ Nat : ?A(n)") proof (rule natInduct) - from base show "?A(0)" by (auto del: funcSetE') + from base show "?A(0)" by auto next fix n assume n: "n \ Nat" and ih: "?A(n)" - show "?A(Succ[n])" + show "?A(succ[n])" proof fix sn - assume sn: "sn \ [1 .. Succ[n] \ S]" + assume sn: "sn \ [1 .. succ[n] \ S]" define so where "so \ [k \ 1 .. n \ sn[k]]" - define lst where "lst \ sn[Succ[n]]" + define lst where "lst \ sn[succ[n]]" have 1: "sn = Append(so, lst)" proof from sn show "isAFcn(sn)" by simp next - from sn n show "DOMAIN sn = 1 .. Succ[Len(so)]" + from sn n show "DOMAIN sn = 1 .. succ[Len(so)]" by (simp add: so_def LenI) next from n show "isASeq(so)" by (force simp: so_def) @@ -379,11 +408,11 @@ proof - from n show "\k \ 1 .. Len(so) : sn[k] = so[k]" by (auto simp: so_def LenI) next - from n show "sn[Succ[Len(so)]] = lst" + from n show "sn[succ[Len(so)]] = lst" by (simp add: so_def lst_def LenI) qed from sn n have 2: "so \ [1 .. n \ S]" - by (force simp: so_def) + unfolding so_def by force with ih have 3: "P(so)" .. from 2 n have 4: "so \ Seq(S)" unfolding Seq_def by auto @@ -430,7 +459,8 @@ translations (*** examples *** -lemma "Len(\a,b,c\) = 3" by (simp add: three_def two_def) +lemma "Len(\a,b,c\) = 3" + by (simp add: two_def three_def) lemma "Append(\a,b\, c) = \a,b,c\" .. @@ -505,7 +535,7 @@ lemmas \ \ establish set equality for sets of enumerated function setEqualI [where A = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] setEqualI [where B = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] -lemma EnumFuncSetI [intro!,simp]: +lemma EnumFuncSetI [intro!]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" and 3: "DOMAIN rngs = DOMAIN doms" and 4: "\i \ DOMAIN doms: f[doms[i]] \ rngs[i]" @@ -598,76 +628,78 @@ print_ast_translation \ \ (*** Examples *** - lemma "(1 :> TRUE) \ [1 : BOOLEAN]" -by auto + by auto lemma "(1 :> TRUE) @@ (5 :> 2) \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE @@ 5 :> 2) = (5 :> 2 @@ 1 :> TRUE)" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (2 :> {}) \ [1 : BOOLEAN, 2: SUBSET Nat, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![1] = FALSE] \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![5] = {}] \ [1 : BOOLEAN, 5 : SUBSET Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (1 :> {}) \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[1 : BOOLEAN, 5 : Nat] = [5 : Nat, 1 : BOOLEAN]" -by (auto simp: two_def three_def four_def five_def) + by auto lemma "\f \ [1 : BOOLEAN, 2: Nat]; g \ [2: Nat, 1: BOOLEAN]; f[1] = g[1]; f[2] = g[2]\ \ f = g" -by auto - + by auto ***) subsection \ Set product \ text \ - The cartesian product of two sets $A$ and $B$ is the set of pairs + The Cartesian product of two sets $A$ and $B$ is the set of pairs whose first component is in $A$ and whose second component is in $B$. We generalize the definition of products to an arbitrary number of sets: $Product(\langle A_1,\ldots,A_n \rangle) = A_1 \times\cdots\times A_n$. \ definition Product -where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : +where "Product(s) \ { f \ [1 .. Len(s) \ UNION {s[i] : i \ 1 .. Len(s)}] : \i \ 1 .. Len(s) : f[i] \ s[i] }" +lemmas \ \ establish set equality for products \ + setEqualI [where A = "Product(s)" for s, intro!] + setEqualI [where B = "Product(s)" for s, intro!] + lemma inProductI [intro!]: assumes "isASeq(p)" and "isASeq(s)" and "Len(p) = Len(s)" and "\k \ 1 .. Len(s) : p[k] \ s[k]" shows "p \ Product(s)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductIsASeq: assumes "p \ Product(s)" and "isASeq(s)" shows "isASeq(p)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductLen: assumes "p \ Product(s)" and "isASeq(s)" shows "Len(p) = Len(s)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductE [elim!]: assumes "p \ Product(s)" and "isASeq(s)" and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; \k \ 1 .. Len(s) : p[k] \ s[k] \ \ P" shows "P" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto (*** examples *** lemma "Product(\\) = { \\ }" by auto -lemma "\2,1\ \ Product(\Nat, Nat\)" by auto +lemma "\2,1\ \ Product(\Nat, Nat\)" by (auto simp: two_def) lemma assumes "x \ X" and "y \ Y" and "z \ Z" @@ -708,14 +740,18 @@ definition notation (ASCII) prod (infixr "\\X" 100) +lemmas \ \ establish set equality for binary products \ + setEqualI [where A = "prod(X, Y)" for X Y, intro!] + setEqualI [where B = "prod(X, Y)" for X Y, intro!] + lemma inProd [simp]: "(\a,b\ \ A \ B) = (a \ A \ b \ B)" -by (auto simp add: prod_def) + unfolding prod_def by auto lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" -using assms by (auto simp: two_def prod_def) + using assms unfolding prod_def by (auto simp: two_def) lemma inProd': "(p \ A \ B) = (\a \ A : \ b \ B : p = \a,b\)" @@ -730,7 +766,7 @@ qed lemma inProdI [intro]: assumes a: "a \ A" and b: "b \ B" and p: "P(\a,b\)" shows "\p \ A \ B : P(p)" -using assms by (intro bExI[of "\a,b\"], simp+) + using assms by (intro bExI[of "\a,b\"]) simp+ lemma inProdI': assumes a: "a \ A" and b: "b \ B" @@ -740,7 +776,7 @@ proof next show "\a,b\[1] = a" by simp next - show "\a,b\[2] = b" unfolding two_def by simp + show "\a,b\[2] = b" by (simp add: two_def) qed lemma inProdE [elim]: @@ -750,43 +786,41 @@ lemma inProdE [elim]: using assms by (auto simp add: inProd') (*** examples *** -lemma "\2,1\ \ Nat \ Nat" by simp +lemma "\2,1\ \ Nat \ Nat" by (simp add: two_def) ***) lemma prodEmptyIff [simp]: "(A \ B = {}) = ((A = {}) \ (B = {}))" -proof auto - fix a b - assume a: "a \ A" and b: "b \ B" and prod: "A \ B = {}" - from a b have "\a,b\ \ A \ B" by simp - with prod show "FALSE" by blast +proof - + { + fix a b + assume a: "a \ A" and b: "b \ B" and prod: "A \ B = {}" + from a b have "\a,b\ \ A \ B" by simp + with prod have "FALSE" by blast + } + thus ?thesis by auto qed lemma prodEmptyIff' [simp]: "({} = A \ B) = ((A = {}) \ (B = {}))" -proof auto - fix a b - assume a: "a \ A" and b: "b \ B" and prod: "{} = A \ B" - from a b have "\a,b\ \ A \ B" by simp - with prod show "FALSE" by blast -qed + by (auto dest: sym) lemma pairProj_in_rel (*[simp]*): assumes r: "r \ A \ B" and p: "p \ r" shows "\p[1],p[2]\ \ r" -using p prodProj[OF rev_subsetD[OF p r], symmetric] by simp + using p prodProj[OF rev_subsetD[OF p r], symmetric] by simp lemma pairProj_in_prod (*[simp]*): assumes r: "r \ A \ B" and p: "p \ r" shows "\p[1],p[2]\ \ A \ B" -using subsetD[OF r p] prodProj[OF rev_subsetD[OF p r], symmetric] by simp + using subsetD[OF r p] prodProj[OF rev_subsetD[OF p r], symmetric] by simp -lemma relProj1 [elim]: (** consider [elim!] ?? **) +lemma relProj1 [elim]: assumes "\a,b\ \ r" and "r \ A \ B" shows "a \ A" -using assms by (auto dest: pairProj_in_prod) + using assms by (auto dest: pairProj_in_prod) -lemma relProj2 [elim]: (** consider [elim!] ?? **) +lemma relProj2 [elim]: assumes "\a,b\ \ r" and "r \ A \ B" shows "b \ B" using assms by (auto dest: pairProj_in_prod) @@ -794,13 +828,12 @@ lemma relProj2 [elim]: (** consider [elim!] ?? **) lemma setOfAllPairs_eq_r (*[simp]*): assumes r: "r \ A \ B" shows "{\p[1], p[2]\ : p \ r} = r" -using subsetD[OF r, THEN prodProj] by force + using subsetD[OF r, THEN prodProj] by force lemma subsetsInProd: assumes "A' \ A" and "B' \ B" shows "A' \ B' \ A \ B" -unfolding prod_def Product_def -using assms by auto + unfolding prod_def Product_def using assms by auto subsection \ Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} \ @@ -815,16 +848,19 @@ syntax (ASCII) translations "{e : \x,y\ \ R}" \ "CONST setOfPairs(R, \x y. e)" +lemmas \ \ establish set equality for sets of pairs \ + setEqualI [where A = "setOfPairs(R, f)" for R f, intro!] + setEqualI [where B = "setOfPairs(R, f)" for R f, intro!] + lemma inSetOfPairsI_ex: assumes "\\x,y\ \ R : a = e(x,y)" shows "a \ { e(x,y) : \x,y\ \ R }" -using assms by (auto simp: setOfPairs_def two_def) + using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsI [intro]: assumes "a = e(x,y)" and "\x,y\ \ R" shows "a \ setOfPairs(R, e)" -unfolding two_def -using assms by (auto simp: setOfPairs_def two_def) + using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsE [elim!]: \ \ the converse is true only if $R$ is a relation \ assumes 1: "z \ setOfPairs(R, e)" @@ -837,14 +873,10 @@ proof - with pR pz show "P" by (intro 3, auto) qed -lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)" for R f, intro!] - setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] - lemma setOfPairs_triv [simp]: assumes "R \ A \ B" shows "{ \x, y\ : \x, y\ \ R } = R" -using assms by auto + using assms by auto lemma setOfPairs_cong (*[cong]*): assumes 1: "R = S" and @@ -852,7 +884,7 @@ lemma setOfPairs_cong (*[cong]*): 3: "\x y. \x, y\ \ S \ e(x, y) = f(x, y)" shows "{ e(x, y) : \x, y\ \ R } = { f(u, v) : \u, v\ \ S }" -using assms by (force intro!: inSetOfPairsI_ex) + using assms by (force intro!: inSetOfPairsI_ex) lemma setOfPairsEqual: assumes 1: "\x y. \x, y\ \ S \ \\x', y'\ \ T : e(x, y) = f(x', y')" @@ -861,27 +893,27 @@ lemma setOfPairsEqual: and 4: "T \ C \ D" shows "{ e(x, y) : \x, y\ \ S } = { f(x, y) : \x, y\ \ T }" -using assms by (force intro!: inSetOfPairsI_ex) + using assms by (force intro!: inSetOfPairsI_ex) subsection \ Basic notions about binary relations \ -definition rel_domain :: "c => c" +definition rel_domain :: "c \ c" where "rel_domain(r) \ { p[1] : p \ r }" -definition rel_range :: "c => c" +definition rel_range :: "c \ c" where "rel_range(r) \ { p[2] : p \ r }" -definition converse :: "c => c" ("(_^-1)" [1000] 999) +definition converse :: "c \ c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" \ \ composition of binary relations \ -definition rel_comp :: "[c,c] => c" (infixr "\" 75) +definition rel_comp :: "[c,c] \ c" (infixr "\" 75) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \ x, z : p = \x, z\ \ (\ y: \x, y\ \ s \ \y, z\ \ r) }" -definition rel_image :: "[c,c] => c" (infixl "``" 90) +definition rel_image :: "[c,c] \ c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" definition Id :: "c \ c" \ \ diagonal: identity over a set \ @@ -894,55 +926,55 @@ definition reflexive \ \ reflexivity over a set \ where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" -unfolding reflexive_def by simp + unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" -unfolding isBool_def by (rule boolifyReflexive) + unfolding isBool_def by (rule boolifyReflexive) definition symmetric \ \ symmetric relation \ where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" -unfolding symmetric_def by simp + unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" -unfolding isBool_def by (rule boolifySymmetric) + unfolding isBool_def by (rule boolifySymmetric) definition antisymmetric \ \ antisymmetric relation \ -where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" + where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" -unfolding antisymmetric_def by simp + unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" -unfolding isBool_def by (rule boolifyAntisymmetric) + unfolding isBool_def by (rule boolifyAntisymmetric) definition transitive \ \ transitivity predicate \ where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" -unfolding transitive_def by simp + unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" -unfolding isBool_def by (rule boolifyTransitive) + unfolding isBool_def by (rule boolifyTransitive) definition irreflexive \ \ irreflexivity predicate \ where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" -unfolding irreflexive_def by simp + unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" -unfolding isBool_def by (rule boolifyIrreflexive) + unfolding isBool_def by (rule boolifyIrreflexive) definition equivalence :: "[c,c] \ c" \ \ (partial) equivalence relation \ where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" -unfolding equivalence_def by simp + unfolding equivalence_def by simp lemma equivalenceIsBool[intro!,simp]: "isBool(equivalence(A,r))" -unfolding isBool_def by (rule boolifyEquivalence) + unfolding isBool_def by (rule boolifyEquivalence) subsubsection \ Domain and Range \ @@ -950,13 +982,12 @@ subsubsection \ Domain and Range \ lemma prod_in_dom_x_ran: assumes "r \ A \ B" and "p \ r" shows "\p[1],p[2]\ \ rel_domain(r) \ rel_range(r)" -unfolding inProd rel_domain_def rel_range_def -using assms by auto + unfolding inProd rel_domain_def rel_range_def using assms by auto lemma in_rel_domainI [iff]: assumes "\x,y\ \ r" shows "x \ rel_domain(r)" -unfolding rel_domain_def using assms by auto + unfolding rel_domain_def using assms by auto lemma in_rel_domainE [elim]: assumes x: "x \ rel_domain(r)" and r: "r \ A \ B" and p: "\y. \x,y\ \ r \ P" @@ -969,15 +1000,15 @@ proof - qed lemma rel_domain (*[simp]*): "r \ A \ B \ rel_domain(r) \ A" -unfolding rel_domain_def using pairProj_in_prod by auto + unfolding rel_domain_def using pairProj_in_prod by auto lemma rel_range (*[simp]*): "r \ A \ B \ rel_range(r) \ B" -unfolding rel_range_def using pairProj_in_prod by auto + unfolding rel_range_def using pairProj_in_prod by auto lemma in_rel_rangeI [iff]: assumes "\x,y\ \ r" shows "y \ rel_range(r)" -unfolding rel_range_def two_def using assms by auto + unfolding rel_range_def using assms by (auto simp: two_def) lemma in_rel_rangeE [elim]: assumes y: "y \ rel_range(r)" and r: "r \ A \ B" and p: "\x. \x,y\ \ r \ P" @@ -990,22 +1021,22 @@ proof - qed lemma dom_in_A (*[simp]*): "rel_domain ({ p \ A \ B : P(p) }) \ A" -by auto + by auto lemma ran_in_B (*[simp]*): "rel_range ({ p \ A \ B : P(p) }) \ B" -by auto + by auto lemma subrel_dom: "r' \ r \ x \ rel_domain(r') \ x \ rel_domain(r)" -unfolding rel_domain_def by auto + unfolding rel_domain_def by auto lemma subrel_ran: "r' \ r \ x \ rel_range(r') \ x \ rel_range(r)" -unfolding rel_range_def by auto + unfolding rel_range_def by auto lemma in_dom_imp_in_A: "r \ A \ B \ x \ rel_domain(r) \ x \ A" -by force + by force lemma in_ran_imp_in_B: "r \ A \ B \ p \ rel_range(r) \ p \ B" -by force + by force subsubsection \ Converse relation \ @@ -1017,21 +1048,21 @@ lemmas converseEqualI = lemma converse_iff [simp]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" -using r prodProj by (auto simp: converse_def two_def) + using r prodProj by (auto simp: converse_def two_def) lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" -unfolding converse_def two_def by auto + unfolding converse_def by (auto simp: two_def) lemma converseD [sym]: assumes r: "r \ A \ B" shows "\a,b\ \ r^-1 \ \b,a\ \ r" -using converse_iff[OF r] by simp + using converse_iff[OF r] by simp lemma converseSubset: "r \ A \ B \ r^-1 \ B \ A" -unfolding converse_def using pairProj_in_prod by auto + unfolding converse_def using pairProj_in_prod by auto -lemma converseE [elim]: (** consider [elim!] ?? **) +lemma converseE [elim]: assumes yx: "yx \ r^-1" and r: "r \ A \ B" and p: "\x y. yx = \y,x\ \ \x,y\ \ r \ P" shows "P" @@ -1058,10 +1089,10 @@ proof - qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" -by auto + by auto lemma converse_empty [simp]: "converse({}) = {}" -by auto + by auto lemma converse_mono_1: assumes r: "r \ A \ B" and s: "s \ A \ B" and sub: "r^-1 \ s^-1" @@ -1077,13 +1108,12 @@ qed lemma converse_mono_2: assumes "r \ A \ B" and "s \ A \ B" and "r \ s" shows "r^-1 \ s^-1" -using assms prodProj by auto + using assms prodProj by auto lemma converse_mono: assumes r:"r \ A \ B" and s:"s \ A \ B" shows "r^-1 \ s^-1 = (r \ s)" -using converse_mono_1[OF r s] converse_mono_2[OF r s] -by blast + using converse_mono_1[OF r s] converse_mono_2[OF r s] by blast (* from HOL *) @@ -1106,7 +1136,7 @@ lemma transitive_converse [simp]: lemma symmetric_iff_converse_eq: assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" -using assms by (auto simp: symmetric_def) + using assms by (auto simp: symmetric_def) subsubsection \ Identity relation over a set \ @@ -1116,48 +1146,47 @@ lemmas idEqualI = setEqualI [where B = "Id(S)" for S, intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" -unfolding Id_def by auto + unfolding Id_def by auto lemma IdI' [intro]: "x \ S \ p = \x,x\ \ p \ Id(S)" -unfolding Id_def by auto + unfolding Id_def by auto lemma IdE [elim!]: "p \ Id(S) \ (\x. x \ S \ p = \x,x\ \ P) \ P" -unfolding Id_def by auto + unfolding Id_def by auto lemma Id_iff: "(\a,b\ \ Id(S)) = (a = b \ a \ S)" -by auto + by auto lemma Id_subset_Prod [simp]: "Id(S) \ S \ S" -unfolding Id_def by auto + unfolding Id_def by auto lemma reflexive_Id: "reflexive(S,Id(S))" -unfolding reflexive_def by auto + unfolding reflexive_def by auto lemma antisymmetric_Id [simp]: "antisymmetric(Id(S))" -unfolding antisymmetric_def by auto + unfolding antisymmetric_def by auto lemma symmetric_Id [simp]: "symmetric(Id(S))" -unfolding symmetric_def by auto + unfolding symmetric_def by auto lemma transitive_Id [simp]: "transitive(Id(S))" -unfolding transitive_def by auto + unfolding transitive_def by auto lemma Id_empty [simp]: "Id({}) = {}" -unfolding Id_def by simp + unfolding Id_def by simp lemma Id_eqI: "a = b \ a \ A \ \a,b\ \ Id(A)" -by simp + by simp lemma converse_Id [simp]: "Id(A)^-1 = Id(A)" -by (auto simp: Id_def two_def) - + by (auto simp: Id_def) lemma dom_Id [simp]: "rel_domain(Id(A)) = A" -unfolding rel_domain_def Id_def by auto + unfolding rel_domain_def Id_def by auto lemma ran_Id [simp]: "rel_range(Id(A)) = A" -unfolding rel_range_def Id_def two_def by auto + unfolding rel_range_def Id_def by (auto simp: two_def) subsubsection \ Composition of relations \ @@ -1169,27 +1198,27 @@ lemmas compEqualI = lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "\ \a,b\ \ s; \b,c\ \ r \ \ \a,c\ \ r \ s" -using assms unfolding rel_comp_def by auto + using assms unfolding rel_comp_def by auto lemma compE [elim!]: assumes "xz \ r \ s" and "r \ B \ C" and "s \ A \ B" shows "(\x y z. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r \ P) \ P" -using assms unfolding rel_comp_def by auto + using assms unfolding rel_comp_def by auto lemma compEpair: assumes "\a,c\ \ r \ s" and "r \ B \ C" and s: "s \ A \ B" shows "\\b. \ \a,b\ \ s; \b,c\ \ r \ \ P \ \ P" -using assms by auto + using assms by auto lemma rel_comp_in_prod [iff]: assumes s: "s \ A \ B" and r: "r \ B \ C" shows "r \ s \ A \ C" -using assms by force + using assms by force lemma rel_comp_in_prodE (*[elim]*): assumes "p \ r \ s" and "s \ A \ B" and r: "r \ B \ C" shows "p \ A \ C" -using assms by force + using assms by force lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" @@ -1326,66 +1355,66 @@ subsubsection \ Properties of relations \ text \ Reflexivity \ lemma reflI [intro]: "(\x. x \ A \ \x,x\ \ r) \ reflexive(A,r)" -unfolding reflexive_def by blast + unfolding reflexive_def by blast lemma reflexiveD [elim!]: "reflexive(A,r) \ a \ A \ \a,a\ \ r" -unfolding reflexive_def by blast + unfolding reflexive_def by blast lemma reflexive_empty (*[simp]*): "reflexive({}, {})" -by auto + by auto text \ Symmetry \ lemma symmetricI: "\ \x y. \x,y\ \ r \ \y,x\ \ r \ \ symmetric(r)" -unfolding symmetric_def by blast + unfolding symmetric_def by blast lemma symmetricE: "\ symmetric(r); \x,y\ \ r \ \ \y,x\ \ r" -unfolding symmetric_def by blast + unfolding symmetric_def by blast lemma symmetric_Int: "\ symmetric(r); symmetric(s) \ \ symmetric(r \ s)" -by (blast intro: symmetricI dest: symmetricE) + by (blast intro: symmetricI dest: symmetricE) text \ Antisymmetry \ lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisymmetricE [elim]: "\ antisymmetric(r); \x,y\ \ r; \y,x\ \ r \ \ x = y" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisymmetricSubset: "r \ s \ antisymmetric(s) \ antisymmetric(r)" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisym_empty (*[simp]*): "antisymmetric({})" -by blast + by blast text \ Transitivity \ lemma transitiveI [intro]: "(\x y z. \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r) \ transitive(r)" -unfolding transitive_def by blast + unfolding transitive_def by blast lemma transD [elim]: "\ transitive(r); \x,y\ \ r; \y,z\ \ r \ \ \x,z\ \ r" -unfolding transitive_def by blast + unfolding transitive_def by blast lemma trans_Int: "transitive(r) \ transitive(s) \ transitive(r \ s)" -by fast + by fast lemma transitive_iff_comp_subset: "transitive(r) = (r \ r \ r)" -unfolding transitive_def rel_comp_def by (auto elim!: subsetD) + unfolding transitive_def rel_comp_def by (auto elim!: subsetD) text \ Irreflexivity \ lemma irreflexiveI [intro]: "\ \x. x \ A \ \x,x\ \ r \ \ irreflexive(A,r)" -unfolding irreflexive_def by blast + unfolding irreflexive_def by blast lemma irreflexiveE [dest]: "\ irreflexive(A,r); x \ A \ \ \x,x\ \ r" -unfolding irreflexive_def by blast + unfolding irreflexive_def by blast subsubsection \ Equivalence Relations \ @@ -1422,7 +1451,7 @@ text \ First half: ``only if'' part \ lemma sym_trans_comp_subset: assumes "r \ A \ A" and "symmetric(r)" and "transitive(r)" shows "r^-1 \ r \ r" -using assms by (simp add: symmetric_iff_converse_eq transitive_iff_comp_subset) + using assms by (simp add: symmetric_iff_converse_eq transitive_iff_comp_subset) lemma refl_comp_subset: assumes r: "r \ A \ A" and refl: "reflexive(A,r)" diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 4c5d5c2d..dae0051c 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1,8 +1,7 @@ (* Zenon.thy --- Support lemmas for Zenon proofs * Author: Damien Doligez - * Copyright (C) 2008-2011 INRIA and Microsoft Corporation - * Version: Isabelle2011-1 - * Time-stamp: <2011-10-11 17:41:20 merz> + * Copyright (C) 2008-2022 INRIA and Microsoft Corporation + * Version: Isabelle2021-1 *) (* isabelle usedir -b Pure TLA+ *) @@ -1090,7 +1089,7 @@ using zenon_sa_seq by (auto simp add: zenon_sa_def) lemma zenon_sa_diff_0a : "zenon_sa (zenon_sa (s1, e1), e2) ~= zenon_sa (<<>>, f2)" -using zenon_sa_def zenon_sa_seq by auto +using zenon_sa_seq by (auto simp: zenon_sa_def) lemma zenon_sa_diff_0b : "zenon_sa (<<>>, f2) ~= zenon_sa (zenon_sa (s1, e1), e2)" @@ -1117,10 +1116,10 @@ lemma zenon_in_nat_0 : "~(0 \\in Nat) ==> FALSE" by blast lemma zenon_in_nat_1 : "~(1 \\in Nat) ==> FALSE" -by blast +by simp lemma zenon_in_nat_2 : "~(2 \\in Nat) ==> FALSE" -by blast +by (simp add: two_def) lemma zenon_in_nat_succ : "~(Succ[x] \\in Nat) ==> (~(x \\in Nat) ==> FALSE) ==> FALSE" @@ -1175,7 +1174,10 @@ lemma zenon_range_1 : "isASeq (<<>>) & {} = Range (<<>>)" by auto lemma zenon_range_2 : assumes h: "(isASeq (s) & a = Range (s))" shows "(isASeq (Append (s, x)) & addElt (x, a) = Range (Append (s, x)))" -using h by auto +proof - + from assms have "succ[Len(s)] \ DOMAIN Append(s,x)" by auto + with assms show ?thesis by auto +qed lemma zenon_set_rev_1 : "a = {} \\cup c ==> c = a" by auto @@ -1207,7 +1209,8 @@ lemma zenon_all_rec_2 : by auto lemma zenon_tuple_acc_1 : - "isASeq (r) ==> Len (r) = n ==> x = Append (r, x) [Succ[n]]" by auto + "isASeq (r) ==> Len (r) = n ==> x = Append (r, x) [succ[n]]" + by auto lemma zenon_tuple_acc_2 : "isASeq (r) ==> k \\in Nat ==> 0 < k ==> k \ Len (r) ==> @@ -1215,32 +1218,52 @@ lemma zenon_tuple_acc_2 : using appendElt1 by auto definition zenon_ss :: "c \ c" -where "zenon_ss (n) \ IF n \\in Nat THEN Succ[n] ELSE 1" +where "zenon_ss (n) \ IF n \\in Nat THEN succ[n] ELSE 1" -lemma zenon_ss_nat : "zenon_ss(x) \\in Nat" by (auto simp add: zenon_ss_def) +lemma zenon_ss_nat : "zenon_ss(x) \\in Nat" +unfolding zenon_ss_def proof (rule condI) + assume "x \ Nat" thus "succ[x] \ Nat" by simp +next + show "1 \ Nat" by simp +qed -lemma zenon_ss_1 : "Succ[0] = zenon_ss(0)" by (auto simp add: zenon_ss_def) +lemma zenon_ss_1 : "succ[0] = zenon_ss(0)" by (simp add: zenon_ss_def) -lemma zenon_ss_2 : "Succ[zenon_ss(x)] = zenon_ss(zenon_ss(x))" by (auto simp add: zenon_ss_def) +lemma zenon_ss_2 : "succ[zenon_ss(x)] = zenon_ss(zenon_ss(x))" +proof (cases "x \ Nat") + case True + then show ?thesis by (simp add: zenon_ss_def) +next + case False + hence "zenon_ss(x) = 1" by (auto simp: zenon_ss_def) + thus ?thesis by (simp add: zenon_ss_def) +qed lemma zenon_zero_lt : "0 < zenon_ss(x)" - by (simp add: zenon_ss_def, rule disjE [of "x \\in Nat" "x \\notin Nat"], rule excluded_middle, simp+) +proof (cases "x \ Nat") + case True + then show ?thesis by (simp add: zenon_ss_def) +next + case False + hence "zenon_ss(x) = 1" by (auto simp: zenon_ss_def) + then show ?thesis by simp +qed lemma zenon_ss_le_sa_1 : "zenon_ss(0) <= Len (zenon_sa (s, x))" - by (auto simp add: zenon_ss_def zenon_sa_def, rule disjE [of "isASeq (s)" "~isASeq (s)"], rule excluded_middle, simp+) + by (cases "isASeq(s)") (auto simp: zenon_ss_def zenon_sa_def) lemma zenon_ss_le_sa_2 : fixes x y z assumes h0: "zenon_ss (x) <= Len (zenon_sa (s, y))" shows "zenon_ss (zenon_ss (x)) <= Len (zenon_sa (zenon_sa (s, y), z))" proof - - have h1: "Succ [Len (zenon_sa (s, y))] = Len (zenon_sa (zenon_sa (s, y), z))" + have h1: "succ [Len (zenon_sa (s, y))] = Len (zenon_sa (zenon_sa (s, y), z))" using zenon_sa_seq by (auto simp add: zenon_sa_def) have h2: "Len (zenon_sa (s, y)) \\in Nat" - using zenon_sa_seq by (rule LenInNat) - have h3: "Succ [zenon_ss (x)] \ Succ [Len (zenon_sa (s, y))]" - using zenon_ss_nat h2 h0 by (simp only: nat_Succ_leq_Succ) + using zenon_sa_seq by simp + have h3: "succ [zenon_ss (x)] \ succ [Len (zenon_sa (s, y))]" + using zenon_ss_nat h2 h0 by (intro int_succ_leq_succI) simp+ show ?thesis using h3 by (auto simp add: zenon_ss_2 h1) qed @@ -1250,8 +1273,8 @@ by auto lemma zenon_dom_app_2 : assumes h: "isASeq (s) & n = Len (s) & x = 1 .. Len (s)" - shows "isASeq (Append (s, y)) & Succ[n] = Len (Append (s, y)) - & addElt (Succ[n], x) = 1 .. Len (Append (s, y))" (is "?a & ?b & ?c") + shows "isASeq (Append (s, y)) & succ[n] = Len (Append (s, y)) + & addElt (succ[n], x) = 1 .. Len (Append (s, y))" using h by auto (* generic proof rules instantiated for small n *) @@ -1272,7 +1295,7 @@ proof - let ?domsetrev = "{l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" + let ?n1n = "succ[?n0n]" let ?indices = "{?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1335,8 +1358,8 @@ proof - let ?domsetrev = "{l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" let ?indices = "{?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1364,39 +1387,18 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) - qed qed qed @@ -1417,9 +1419,9 @@ proof - let ?domsetrev = "{l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" let ?indices = "{?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1447,53 +1449,23 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) @@ -1520,10 +1492,10 @@ proof - let ?domsetrev = "{l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" let ?indices = "{?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1551,70 +1523,30 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) @@ -1642,11 +1574,11 @@ proof - let ?domsetrev = "{l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" let ?indices = "{?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1674,87 +1606,37 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) @@ -1783,12 +1665,12 @@ proof - let ?domsetrev = "{l6x, l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" - let ?n6n = "Succ[?n5n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" + let ?n6n = "succ[?n5n]" let ?indices = "{?n6n, ?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1816,104 +1698,44 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) next have hn: "l6x = ?doms[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s6x = ?rngs[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n6n]] \\in ?rngs[?n6n]" by (rule subst[OF hs], rule subst[OF hn], rule h6) @@ -1943,13 +1765,13 @@ proof - let ?domsetrev = "{l7x, l6x, l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" - let ?n6n = "Succ[?n5n]" - let ?n7n = "Succ[?n6n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" + let ?n6n = "succ[?n5n]" + let ?n7n = "succ[?n6n]" let ?indices = "{?n7n, ?n6n, ?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1977,121 +1799,51 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) next have hn: "l6x = ?doms[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s6x = ?rngs[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n6n]] \\in ?rngs[?n6n]" by (rule subst[OF hs], rule subst[OF hn], rule h6) next have hn: "l7x = ?doms[?n7n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s7x = ?rngs[?n7n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n7n]] \\in ?rngs[?n7n]" by (rule subst[OF hs], rule subst[OF hn], rule h7) From 4d6cfc220a2f6e1f9596b92b8254475b45ffe57b Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 18 Mar 2022 19:20:01 +0100 Subject: [PATCH 062/167] fixed Zenon.thy and examples, cleanup of old theory files --- isabelle/NatArith.thy | 1174 --- isabelle/NatDivision.thy | 731 -- isabelle/NatOrderings.thy | 709 -- isabelle/NewIntegers.thy | 12656 -------------------------- isabelle/Peano.thy | 629 -- isabelle/Zenon.thy | 184 +- isabelle/examples/Allocator.thy | 69 +- isabelle/examples/AtomicBakeryG.thy | 910 +- isabelle/examples/ROOT | 18 + isabelle/examples/ROOT.ML | 13 - 10 files changed, 521 insertions(+), 16572 deletions(-) delete mode 100644 isabelle/NatArith.thy delete mode 100644 isabelle/NatDivision.thy delete mode 100644 isabelle/NatOrderings.thy delete mode 100644 isabelle/NewIntegers.thy delete mode 100644 isabelle/Peano.thy create mode 100644 isabelle/examples/ROOT delete mode 100644 isabelle/examples/ROOT.ML diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy deleted file mode 100644 index 1e58c382..00000000 --- a/isabelle/NatArith.thy +++ /dev/null @@ -1,1174 +0,0 @@ -(* Title: TLA+/NatArith.thy - Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - -section \ Arithmetic (except division) over natural numbers \ - -theory NatArith -imports NatOrderings -begin - -named_theorems algebra_simps "algebra simplification rules" - -(* -ML \ -structure AlgebraSimps = - Named_Thms(val name = "algebra_simps" - val description = "algebra simplification rules"); -\ - -setup AlgebraSimps.setup -*) - -text \ - The rewrites accumulated in @{text algebra_simps} deal with the - classical algebraic structures of groups, rings and family. They simplify - terms by multiplying everything out (in case of a ring) and bringing sums and - products into a canonical form (by ordered rewriting). As a result these - rewrites decide group and ring equalities but also help with inequalities. - - Of course it also works for fields, but it knows nothing about multiplicative - inverses or division. This should be catered for by @{text field_simps}. -\ - - -subsection \ Addition of natural numbers \ - -definition addnat -where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]])" - -definition arith_add :: "[c,c] \ c" (infixl "+" 65) -where nat_add_def: "\m \ Nat; n \ Nat\ \ (m + n) \ addnat(m)[n]" - -text \ Closure \ - -lemma addnatType: - assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" -using assms unfolding addnat_def by (rule bprimrecType_nat, auto) - -lemma addIsNat [intro!,simp]: - assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" -unfolding nat_add_def[OF assms] using assms addnatType by blast - -text \ Base case and Inductive case \ - -lemma addnat_0: - assumes "m \ Nat" shows "addnat(m)[0] = m" -using assms unfolding addnat_def by (rule primrec_natE, auto) - -lemma add_0_nat [simp]: - assumes "m \ Nat" shows "m + 0 = m" -by (simp add: nat_add_def[OF assms] addnat_0[OF assms]) - -lemma addnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" -proof (rule primrec_natE[OF m]) - show "addnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]]))" - unfolding addnat_def .. -next - assume "\n \ Nat : addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" - with n show "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" by blast -qed (auto) - -lemma add_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "m + Succ[n] = Succ[m + n]" -using assms by (simp add: nat_add_def addnat_Succ[OF assms]) - -lemma add_0_left_nat [simp]: - assumes n: "n \ Nat" - shows "0 + n = n" -using n by (induct, auto) - -lemma add_Succ_left_nat [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "Succ[m] + n = Succ[m + n]" -using n apply induct -using m by auto - -lemma add_Succ_shift_nat: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "Succ[m] + n = m + Succ[n]" -using assms by simp - -text \ Commutativity \ - -lemma add_commute_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m + n = n + m" -using n apply induct -using assms by auto - -text \ Associativity \ - -lemma add_assoc_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (n + p) = (m + n) + p" -using assms by (induct, simp_all) - -text \ Cancellation rules \ - -lemma add_left_cancel_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + n = m + p) = (n = p)" -using assms by (induct, simp_all) - -lemma add_right_cancel_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(n + m = p + m) = (n = p)" -using assms by (induct, simp_all) - - -lemma add_left_commute_nat [algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a + (b + c) = b + (a + c)" -using assms by(simp only: add_assoc_nat add_commute_nat) - -(* from HOL/OrderedGroups.thy *) -lemmas add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat - -text \ Reasoning about @{text "m + n = 0"}, etc. \ - -lemma add_is_0_nat [iff]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = 0) = (m = 0 \ n = 0)" -using m apply (rule natCases) -using n by (induct, auto) - -lemma add_is_1_nat: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = 1) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" -using m apply (rule natCases) -using n by (induct, auto) - -lemma one_is_add_nat: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 = m + n) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" -using m apply (rule natCases) -using n by (induct, auto)+ - -lemma add_eq_self_zero_nat1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = m) = (n = 0)" -using n apply (rule natCases) - using m apply simp - using m apply induct apply auto -done - -lemma add_eq_self_zero_nat2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(n + m = m) = (n = 0)" -using assms by (simp add: add_commute_nat) - - -subsection \ Multiplication of natural numbers \ - -definition multnat -where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m)" - -definition mult :: "[c,c] \ c" (infixl "*" 70) -where nat_mult_def: "\m \ Nat; n \ Nat\ \ m * n \ multnat(m)[n]" - -text \ Closure \ - -lemma multnatType: - assumes "m \ Nat" shows "multnat(m) \ [Nat \ Nat]" -unfolding multnat_def by (rule bprimrecType_nat, auto simp: assms) - -lemma multIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m * n \ Nat" -unfolding nat_mult_def[OF assms] using assms multnatType by blast - -text \ Base case and Inductive step \ - -lemma multnat_0: - assumes "m \ Nat" shows "multnat(m)[0] = 0" -unfolding multnat_def by (rule primrec_natE, auto simp: assms) - -lemma mult_0_nat [simp]: \ \ neutral element \ - assumes n: "n \ Nat" shows "n * 0 = 0" -by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) - - -lemma multnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "multnat(m)[Succ[n]] = multnat(m)[n] + m" -proof (rule primrec_natE) - show "multnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m))" - unfolding multnat_def .. -next - assume "\n \ Nat : multnat(m)[Succ[n]] = multnat(m)[n] + m" - with n show "multnat(m)[Succ[n]] = multnat(m)[n] + m" by blast -qed (auto simp: m) - -lemma mult_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "m * Succ[n] = m * n + m" -using assms by (simp add: nat_mult_def multnat_Succ[OF assms]) - -lemma mult_0_left_nat [simp]: - assumes n: "n \ Nat" - shows "0 * n = 0" -using n by (induct, simp_all) - -lemma mult_Succ_left_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "Succ[m] * n = m * n + n" -using n apply induct -using m by (simp_all add: add_ac_nat) - -text \ Commutativity \ - -lemma mult_commute_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m * n = n * m" -using assms by (induct, simp_all) - -text \ Distributivity \ - -lemma add_mult_distrib_left_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n + p) = m * n + m * p" -using assms apply induct -proof auto - fix m - assume "m \ Nat" "m * (n + p) = m * n + m * p" - with n p - add_assoc_nat[of "m * n + m * p" n p] - add_assoc_nat[of "m * n" "m * p" n] - add_commute_nat[of "m * p" n] - add_assoc_nat[of "m * n" n "m * p"] - add_assoc_nat[of "m * n + n" "m * p" p] - show "m * n + m * p + (n + p) = m * n + n + (m * p + p)" - by simp -qed - -lemma add_mult_distrib_right_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(n + p) * m = n * m + p * m" -using m apply induct -using n p apply auto -proof - - fix m - assume "m \ Nat" "(n + p) * m = n * m + p * m" - with n p - add_assoc_nat[of "n * m + p * m" n p] - add_assoc_nat[of "n * m" "p * m" n] - add_commute_nat[of "p * m" n] - add_assoc_nat[of "n * m" n "p * m"] - add_assoc_nat[of "n * m + n" "p * m" p] - show "n * m + p * m + (n + p) = n * m + n + (p * m + p)" - by simp -qed - -text \ Identity element \ - -(* used for arithmetic simplification setup *) -lemma mult_1_right_nat: "a \ Nat \ a * 1 = a" by simp -lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp - -text \ Associativity \ - -lemma mult_assoc_nat[algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n * p) = (m * n) * p" -using m apply induct -using assms by (auto simp add: add_mult_distrib_right_nat) - -text \ Reasoning about @{text "m * n = 0"}, etc. \ - -lemma mult_is_0_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = 0) = (m = 0 \ n = 0)" -using m apply induct -using n by auto - -lemma mult_eq_1_iff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = 1) = (m = 1 \ n = 1)" -using m apply induct -using n by (induct, auto)+ - -lemma one_eq_mult_iff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 = m * n) = (m = 1 \ n = 1)" -proof - - have "(1 = m * n) = (m * n = 1)" by auto - also from assms have "... = (m = 1 \ n = 1)" by simp - finally show ?thesis . -qed - -text \ Cancellation rules \ - -lemma mult_cancel1_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k * m = k * n) = (m = n \ (k = 0))" -proof - - have "k \ 0 \ k * m = k * n \ m = n" - using n m proof (induct arbitrary: m) - fix m - assume "k \ 0" "k * m = k * 0" "m \ Nat" - with k show "m = 0" by simp - next - fix n m - assume - n': "n \ Nat" and h1:"\m. \k \ 0; k * m = k * n; m \ Nat\ \ m = n" - and k0: "k \ 0" and h2: "k * m = k * Succ[n]" and m':"m \ Nat" - from m' show "m = Succ[n]" - proof (rule natCases) - assume "m = 0" - with k have "k * m = 0" by simp - moreover - from k k0 n' have "k * Succ[n] \ 0" by simp - moreover - note h2 - ultimately show ?thesis by simp - next - fix w - assume w:"w \ Nat" and h3:"m = Succ[w]" - with k n' h2 have "k * w = k * n" by simp - with k k0 w h1[of w] h3 show ?thesis by simp - qed - qed - with k m n show ?thesis by auto -qed - -lemma mult_cancel2_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k = n * k) = (m = n \ k = 0)" -using assms by (simp add: mult_commute_nat) - -lemma Suc_mult_cancel1_nat: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[k] * m = Succ[k] * n) = (m = n)" -using k m n mult_cancel1_nat[of "Succ[k]" m n] by simp - - -lemma mult_left_commute_nat[algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a * (b * c) = b * (a * c)" -using assms by(simp only: mult_commute_nat mult_assoc_nat) - -(* from HOL/OrderedGroups.thy *) -lemmas mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat - - -subsection \ Predecessor \ - -definition Pred -where "Pred \ [n \ Nat \ IF n=0 THEN 0 ELSE CHOOSE x \ Nat : n=Succ[x]]" - -lemma Pred_0_nat [simp]: "Pred[0] = 0" -by (simp add: Pred_def) - -lemma Pred_Succ_nat [simp]: "n \ Nat \ Pred[Succ[n]] = n" -unfolding Pred_def by (auto intro: bChooseI2) - -lemma Succ_Pred_nat [simp]: - assumes "n \ Nat" and "n \ 0" - shows "Succ[Pred[n]] = n" -using assms unfolding Pred_def by (cases "n", auto intro: bChooseI2) - -lemma Pred_in_nat [intro!, simp]: - assumes "n \ Nat" shows "Pred[n] \ Nat" -using assms by (cases "n", auto) - -subsection \ Difference of natural numbers \ - -text \ - We define a form of difference @{text "--"} of natural numbers that cuts off - at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes - called ``arithmetic difference''. -\ - -definition adiffnat -where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]])" - -definition adiff (infixl "--" 65) -where nat_adiff_def: "\m \ Nat; n \ Nat\ \ (m -- n) \ adiffnat(m)[n]" - -text \ Closure \ - -lemma adiffnatType: - assumes "m \ Nat" shows "adiffnat(m) \ [Nat \ Nat]" -using assms unfolding adiffnat_def by (rule bprimrecType_nat, auto) - -lemma adiffIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ Nat" -unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast - -text \ Neutral element and Inductive step \ - -lemma adiffnat_0: - assumes "m \ Nat" shows "adiffnat(m)[0] = m" -using assms unfolding adiffnat_def by (rule primrec_natE, auto) - -lemma adiff_0_nat [simp]: - assumes "m \ Nat" shows "m -- 0 = m" -by (simp add: nat_adiff_def[OF assms] adiffnat_0[OF assms]) - -lemma adiffnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" -proof (rule primrec_natE[OF m]) - show "adiffnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]]))" - unfolding adiffnat_def .. -next - assume "\n \ Nat : adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" - with n show "adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" by blast -qed (auto) - -lemma adiff_Succ_nat: - assumes "m \ Nat" and "n \ Nat" - shows "m -- Succ[n] = Pred[m -- n]" -using assms by (simp add: nat_adiff_def adiffnat_Succ[OF assms]) - -lemma adiff_0_eq_0_nat [simp]: - assumes n: "n \ Nat" - shows "0 -- n = 0" -using assms by (induct) (simp_all add: adiff_Succ_nat) - -text \ Reasoning about @{text "m -- m = 0"}, etc. \ - -lemma adiff_Succ_Succ_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "Succ[m] -- Succ[n] = m -- n" - using n by induct (auto simp add: m adiff_Succ_nat) - -lemma adiff_self_eq_0_nat [simp]: - assumes m: "m \ Nat" - shows "m -- m = 0" -using m by induct auto - -text \ Associativity \ - -lemma adiff_adiff_left_nat: - assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "(i -- j) -- k = i -- (j + k)" -using i j by (rule diffInduct) (auto simp: k) - -lemma Succ_adiff_adiff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" -using assms by (simp add: adiff_adiff_left_nat) - -text \ Commutativity \ - -lemma adiff_commute_nat: - assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "i -- j -- k = i -- k -- j" -using assms by (simp add: adiff_adiff_left_nat add_commute_nat) - -text \ Cancellation rules \ - -lemma adiff_add_inverse_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(n + m) -- n = m" -using n by induct (simp_all add: m) - -lemma adiff_add_inverse2_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n) -- n = m" -using assms by (simp add: add_commute_nat [of m n]) - -lemma adiff_cancel_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m) -- (k + n) = m -- n" -using k by induct (simp_all add: m n) - -lemma adiff_cancel2_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k) -- (n + k) = m -- n" -using assms by (simp add: add_commute_nat) - -lemma adiff_add_0_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "n -- (n + m) = 0" -using n by induct (simp_all add: m) - -text \ Difference distributes over multiplication \ - -lemma adiff_mult_distrib_nat [algebra_simps]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m -- n) * k = (m * k) -- (n * k)" -using m n by (rule diffInduct) (simp_all add: k) - -lemma adiff_mult_distrib2_nat [algebra_simps]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "k * (m -- n) = (k * m) -- (k * n)" -using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) - -\ \ NOT added as rewrites, since sometimes they are used from right to left \ -lemmas nat_distrib = - add_mult_distrib_right_nat add_mult_distrib_left_nat - adiff_mult_distrib_nat adiff_mult_distrib2_nat - - -subsection \ Additional arithmetic theorems \ - -subsubsection \ Monotonicity of Addition \ - -lemma Succ_pred [simp]: - assumes n: "n \ Nat" - shows "n > 0 \ Succ[n -- 1] = n" -using assms by (simp add: adiff_Succ_nat[OF n zeroIsNat] nat_gt0_not0[OF n]) - -lemma nat_add_left_cancel_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m \ k + n) = (m \ n)" -using assms by (induct k) simp_all - -lemma nat_add_right_cancel_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k \ n + k) = (m \ n)" -using assms by (induct k) simp_all - -lemma nat_add_left_cancel_Succ_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[k + m] \ k + n) = (Succ[m] \ n)" -using assms by (induct k) simp_all - -(* immediate corollary *) -lemma nat_add_left_cancel_less: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m < k + n) = (m < n)" -using assms by simp - -lemma nat_add_right_cancel_Succ_less [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m + k] \ n + k) = (Succ[m] \ n)" -using assms by (induct k) simp_all - -lemma nat_add_right_cancel_less: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k < n + k) = (m < n)" -using assms by simp - -lemma less_imp_Succ_add: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m < n \ (\k \ Nat: n = Succ[m + k])" (is "_ \ ?P(n)") -using n proof (induct) - case 0 with m show ?case by simp -next - fix n - assume n: "n \ Nat" and ih: "m ?P(n)" and suc: "m < Succ[n]" - from suc m n show "?P(Succ[n])" - proof (rule nat_less_SuccE) - assume "m Nat" and "n = Succ[m + k]" by (blast dest: ih) - with m n have "Succ[k] \ Nat" and "Succ[n] = Succ[m + Succ[k]]" by auto - thus ?thesis .. - next - assume "m=n" - with n have "Succ[n] = Succ[m + 0]" by simp - thus ?thesis by blast - qed -qed - -lemma nat_leq_trans_add_left_false [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\m + n \ p; p \ n\ \ (m + n < p) = FALSE" -apply (induct n p rule: diffInduct) -using assms by simp_all - - -subsubsection \ (Partially) Ordered Groups \ - -\ \ The two following lemmas are just ``one half'' of - @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above." \ -lemma add_leq_left_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a \ b \ c + a \ c + b" -using assms by simp - -lemma add_leq_right_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a \ b \ a + c \ b + c" -using assms by simp - -text \ non-strict, in both arguments \ -lemma add_leq_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a \ b \ c \ d \ a + c \ b + d" -using assms - add_leq_right_mono[OF a b c] - add_leq_left_mono[OF c d b] - nat_leq_trans[of "a + c" "b + c" "b + d"] -by simp - -\ \ Similar for strict less than. \ -lemma add_less_left_mono: - assumes "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "a < b \ c + a < c + b" - using assms by (simp add: add_leq_left_mono less_def) - -lemma add_less_right_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a < b \ a + c < b + c" -using assms by (auto simp: add_less_left_mono add_commute_nat) - -text \ Strict monotonicity in both arguments \ -lemma add_less_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a < b \ c < d \ a + c < b + d" -using assms - add_less_right_mono[OF a b c] - add_less_left_mono[OF c d b] - nat_less_trans[of "a + c" "b + c" "b + d"] -by simp - -lemma add_less_leq_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a < b \ c \ d \ a + c < b + d" -using assms - add_less_right_mono[OF a b c] - add_leq_left_mono[OF c d b] - nat_less_leq_trans[of "a + c" "b + c" "b + d"] -by blast - -lemma add_leq_less_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a \ b \ c < d \ a + c < b + d" -using assms - add_leq_right_mono[OF a b c] - add_less_left_mono[OF c d b] - nat_leq_less_trans[of "a + c" "b + c" "b + d"] -by blast - -lemma leq_add1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "n \ n + m" -using assms add_leq_left_mono[of 0 m n] by simp - -lemma leq_add2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m + n" -using assms add_leq_right_mono [of 0 m n] by simp - -lemma trans_leq_add1 [elim]: - assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m+k" -using assms by (simp add: nat_leq_trans) - -lemma trans_leq_add2 [elim]: - assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ k+m" -using assms by (simp add: nat_leq_trans) - -lemma add_leqD1 [elim]: - assumes "n+k \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m" -using assms by (simp add: nat_leq_trans[of "n" "n+k" "m"]) - -lemma add_leqD2 [elim]: - assumes "k+n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m" -using assms by (simp add: nat_leq_trans[of "n" "k+n" "m"]) - -lemma add_Succ_leqD1 [elim]: - assumes "Succ[n+k] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "Succ[n] \ m" -using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[n+k]" "m"]) - -lemma add_Succ_leqD2 [elim]: - assumes "Succ[k+n] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "Succ[n] \ m" -using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[k+n]" "m"]) - -lemma less_add_Succ1: - assumes "i \ Nat" and "m \ Nat" - shows "i < Succ[i + m]" -using assms by simp - -lemma less_add_Succ2: - assumes "i \ Nat" and "m \ Nat" - shows "i < Succ[m + i]" -using assms by simp - -lemma less_iff_Succ_add: - assumes "m \ Nat" and "n \ Nat" - shows "(m < n) = (\k \ Nat: n = Succ[m + k])" -using assms by (auto intro!: less_imp_Succ_add) - -lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i < j + m" -using assms by auto - -lemma trans_less_add2: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i < m + j" -using assms by auto - -lemma add_lessD1: - assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "i < k" -using assms by auto - -lemma add_lessD2: - assumes "j+i < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "i < k" -using assms by auto - -lemma not_Succ_add_le1: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(Succ[i + j] \ i) = FALSE" -using assms by simp - -lemma not_Succ_add_le2: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(Succ[j + i] \ i) = FALSE" -using assms by simp - -lemma not_add_less1: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(i + j < i) = FALSE" -using assms by simp - -lemma not_add_less2: - assumes "i \ Nat" and "j \ Nat" - shows "(j + i < i) = FALSE" -using assms by simp - -lemma add_leqE: - assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(m \ n \ k \ n \ R) \ R" -using assms by (blast dest: add_leqD1 add_leqD2) - -lemma leq_add_self1 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m+n \ m) = (n=0)" -proof - - { - assume "m+n \ m" - with assms have "m+n = m" by (intro nat_leq_antisym) simp_all - } - with assms show ?thesis by auto -qed - -lemma leq_add_self2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(n+m \ m) = (n=0)" -proof - - { - assume "n+m \ m" - with assms have "n+m = m" by (intro nat_leq_antisym) simp_all - } - with assms show ?thesis by auto -qed - -lemma add_gt_0: - assumes "m \ Nat" and "n \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" -proof - - from assms have "m + n \ Nat" by simp - hence "(m + n > 0) = (m + n \ 0)" by (rule nat_gt0_not0) - also from assms have "\ = (m \ 0 \ n \ 0)" by simp - finally show ?thesis using assms by (simp only: nat_gt0_not0) -qed - -(* The above lemma follows from the following simplification rule. *) -lemma add_ge_1 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m + n \ 1) = (m \ 1 \ n \ 1)" - using assms by (auto simp add: add_gt_0) - - -subsubsection \ More results about arithmetic difference \ - -text \ Addition is the inverse of subtraction: - if @{term "n \ m"} then @{term "n + (m -- n) = m"}. \ -lemma add_adiff_inverse: - assumes "m \ Nat" and "n \ Nat" - shows "\(m < n) \ n + (m -- n) = m" -apply (induct m n rule: diffInduct) -using assms by simp_all - -lemma le_add_adiff_inverse [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ n + (m -- n) = m" -using assms by (simp add: add_adiff_inverse nat_not_order_simps) - -lemma le_add_adiff_inverse2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ (m -- n) + n = m" -using assms by (simp add: add_commute_nat) - -lemma Succ_adiff_leq: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ Succ[m] -- n = Succ[m -- n]" -apply (induct m n rule: diffInduct) -using assms by simp_all - -lemma adiff_less_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m -- n < Succ[m]" -apply (induct m n rule: diffInduct) -using assms by (auto simp: nat_less_Succ) - -lemma adiff_leq_self [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m -- n \ m" -apply (induct m n rule: diffInduct) -using assms by auto - -lemma leq_iff_add: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") -proof - - have "?lhs \ ?rhs" - proof - assume mn: "m \ n" - with m n have "n = m + (n -- m)" by simp - with m n show "?rhs" by blast - qed - moreover - from assms have "?rhs \ ?lhs" by auto - ultimately show ?thesis by blast -qed - -lemma less_imp_adiff_less: - assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" - shows "j -- n < k" -proof - - from j n have "j -- n \ Nat" by simp - moreover - from j n have s1_2: "j -- n \ j" by simp - ultimately show ?thesis using j k jk nat_leq_less_trans by auto -qed - -lemma adiff_Succ_less (*[simp]*): - assumes i: "i \ Nat" and n: "n \ Nat" - shows "0 < n \ n -- Succ[i] < n" -using n by cases (auto simp: i) - -lemma adiff_add_assoc: - assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(i + j) -- k = i + (j -- k)" -using assms by (induct j k rule: diffInduct, simp+) - -lemma adiff_add_assoc2: - assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(j + i) -- k = (j -- k) + i" -using assms by (simp add: add_commute_nat adiff_add_assoc) - -lemma adiff_add_assoc3: - assumes "n \ p" and "p \ m+n" and "m \ Nat" and "n \ Nat" and "p \ Nat" - shows "m -- (p -- n) = m + n -- p" -using assms by (induct p n rule: diffInduct, simp+) - -lemma adiff_add_assoc4: - assumes 1: "n \ m" and 2: "m -- n \ p" and 3: "m \ p" - and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "p -- (m -- n) = (p -- m) + n" -using assms - adiff_add_assoc2[OF _ n p m, symmetric] - adiff_add_assoc3[OF _ _ p n m] apply simp -using trans_leq_add1[OF _ m p n] by simp - -lemma le_imp_adiff_is_add: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(j -- i = k) = (j = k + i)" -using assms by auto - -lemma adiff_is_0_eq [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m -- n = 0) = (m \ n)" -by (induct m n rule: diffInduct, simp_all add: assms) - -lemma adiff_is_0_eq' (*[simp]*): - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "m -- n = 0" -using assms by simp - -lemma one_leq_adiff [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(1 \ n -- m) = (Succ[m] \ n)" -by (induct m n rule: diffInduct, simp_all add: assms) - -lemma zero_less_adiff: - assumes "m \ Nat" and "n \ Nat" - shows "(0 < n -- m) = (m < n)" -using assms by simp - -lemma less_imp_add_positive: - assumes "i < j" and "i \ Nat" and "j \ Nat" - shows "\k \ Nat: 0 < k \ i + k = j" -proof - - thm leq_iff_add - from assms obtain k where k: "k \ Nat" "j = Succ[i]+k" - by (auto simp: leq_iff_add) - with assms show ?thesis by force -qed - -lemma leq_adiff_right_add_left: - assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "m \ n -- k = (m + k \ n)" -using assms by (induct n k rule: diffInduct, simp+) - -lemma leq_adiff_left_add_right_equiv: - assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(n -- k \ m) = (n \ m + k)" -using assms by (induct n k rule: diffInduct, simp+) - -lemma leq_adiff_left_add_right: - assumes 1: "n -- p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "n \ m + p" -using assms by (induct n p rule: diffInduct, simp+) - -lemma leq_adiff_trans: - assumes "p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "p -- n \ m" -apply(rule nat_leq_trans[of "p -- n" p m]) -using assms adiff_leq_self[OF p n] by simp_all - -lemma leq_adiff_right [simp]: - assumes "n \ m" and "m \ Nat" and "n \ Nat" - shows "(m \ m -- n) = (n = 0)" - using assms by (simp add: leq_adiff_right_add_left) - - -subsubsection \ Monotonicity of Multiplication \ - -(* from HOL/Ring_and_Field.thy *) - -lemma mult_leq_left_mono: - assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "c * a \ c * b" -using c by induct (simp_all add: add_leq_mono 1 a b) - -lemma mult_leq_right_mono: - assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a * c \ b * c" -using c by induct (simp_all add: add_leq_mono add_commute_nat 1 a b) - -text \ @{text "\"} monotonicity, BOTH arguments \ - -lemma mult_leq_mono: - assumes 1: "i \ j" "k \ l" - and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" and l: "l \ Nat" - shows "i * k \ j * l" -using assms - mult_leq_right_mono[OF _ i j k] - mult_leq_left_mono[OF _ k l j] - nat_leq_trans[of "i * k" "j * k" "j * l"] - by simp - -lemma self_leq_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" - shows "a \ a * b" -proof - - have "1 \ Nat" by simp - from 1 this b a have "a * 1 \ a * b" by (rule mult_leq_left_mono) - with a b show ?thesis by simp -qed - -lemma self_leq_mult_right: - assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" - shows "a \ b * a" -using assms by (simp add: self_leq_mult_left mult_commute_nat) - -text \ Similar lemmas for @{text "<"} \ - -lemma mult_Succ_leq_right_mono1: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "Succ[a * Succ[c]] \ b * Succ[c]" -using c proof induct - case 0 - with assms show ?case by simp -next - fix n - assume n: "n \ Nat" and ih: "Succ[a * Succ[n]] \ b * Succ[n]" - from ab a b have 1: "a \ b" - by (simp add: nat_leq_trans[of "a" "Succ[a]" "b"]) - from a n have "Succ[a * Succ[Succ[n]]] = Succ[a * Succ[n]] + a" - by simp - also from a b n ih have "\ \ b* Succ[n] + a" - by simp - also from a b n 1 have "\ \ b * Succ[n] + b" - by simp - finally show "Succ[a * Succ[Succ[n]]] \ b * Succ[Succ[n]]" - using b n by simp -qed - -lemma mult_Succ_leq_right_mono2: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" - shows "Succ[a * c] \ b * c" -proof - - from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" - by (blast dest: nat_ge1_implies_Succ) - with ab a b show ?thesis by (auto dest: mult_Succ_leq_right_mono1[of a b k]) -qed - -lemma mult_less_right_mono: - assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "a * c < b * c" -using assms by (simp add: mult_Succ_leq_right_mono2) - -lemma mult_Succ_leq_left_mono1: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "Succ[Succ[c] * a] \ Succ[c] * b" -proof - - have "Succ[Succ[c] * a] = Succ[a * Succ[c]]" - using a c by (simp add: mult_commute_nat) - also have "\ \ b * Succ[c]" - using assms by (rule mult_Succ_leq_right_mono1) - also have "\ = Succ[c] * b" - using b c by (simp add: mult_commute_nat) - finally show ?thesis . -qed - -lemma mult_Succ_leq_left_mono2: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" - shows "Succ[c * a] \ c * b" -proof - - from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" - by (blast dest: nat_ge1_implies_Succ) - with ab a b show ?thesis by (auto dest: mult_Succ_leq_left_mono1[of a b k]) -qed - -lemma mult_less_left_mono: - assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "c * a < c * b" -using assms by (simp add: mult_Succ_leq_left_mono2) - -lemma one_leq_mult_iff[simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 \ m * n) = (1 \ m \ 1 \ n)" -using assms nat_gt0_not0 by simp - -lemma nat_0_less_mult_iff: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(0 < i * j) = (0 < i \ 0 < j)" -using assms by simp - -lemma mult_Succ_leq_cancel_left [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[k * m] \ k * n) = (1 \ k \ Succ[m] \ n)" -proof (auto intro: mult_Succ_leq_left_mono2[OF _ m n k _]) - assume "Succ[k * m] \ k * n" - from k m n this show "1 \ k" by (cases k) simp_all -next - assume 1: "Succ[k * m] \ k * n" - show "Succ[m] \ n" - proof (rule contradiction) - assume "\(Succ[m] \ n)" - with m n k have "k*n \ k*m" by (simp add: nat_not_order_simps mult_leq_left_mono) - with m n k have "\(Succ[k*m] \ k*n)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed - -lemma mult_less_cancel_left: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m < k * n) = (0 < k \ m < n)" -using assms by simp - -lemma mult_Succ_leq_cancel_right [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[m * k] \ n * k) = (1 \ k \ Succ[m] \ n)" -proof (auto intro: mult_Succ_leq_right_mono2[OF _ m n k _]) - assume "Succ[m * k] \ n * k" - from k m n this show "1 \ k" by (cases k) simp_all -next - assume 1: "Succ[m * k] \ n * k" - show "Succ[m] \ n" - proof (rule contradiction) - assume "\(Succ[m] \ n)" - with m n k have "n*k \ m*k" by (simp add: nat_not_order_simps mult_leq_right_mono) - with m n k have "\(Succ[m*k] \ n*k)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed - -lemma mult_less_cancel_right: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k < n * k) = (0 < k \ m < n)" -using assms by simp - -lemma mult_Succ_leq_self_left [dest]: - assumes nk: "Succ[n*k] \ n" and n: "n \ Nat" and k: "k \ Nat" - shows "k = 0" -proof (rule contradiction) - assume "k \ 0" - with k nat_neq0_conv have "1 \ k" by simp - with n k have "n \ n * k" by (rule self_leq_mult_left) - with n k have "\(Succ[n*k] \ n)" by (simp add: nat_not_order_simps) - from this nk show "FALSE" .. -qed - -lemma mult_less_self_left: - assumes "n * k < n" and "n \ Nat" and "k \ Nat" - shows "k=0" -using assms by auto - -lemma mult_Succ_leq_self_right [dest]: - assumes less: "Succ[k*n] \ n" and n: "n \ Nat" and k: "k \ Nat" - shows "k=0" -using assms by (auto simp: mult_commute_nat[OF k n]) - -lemma mult_less_self_right: - assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" - shows "k=0" -using assms by auto - -lemma mult_leq_cancel_left [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m \ k * n) = (k = 0 \ m \ n)" -proof - - { - assume 1: "k*m \ k*n" and 2: "k \ 0" and 3: "\(m \ n)" - from k 2 nat_gt0_not0 have "k > 0" by simp - with 3 k m n have "\(k * m \ k * n)" by (simp add: nat_not_order_simps) - from this 1 have "FALSE" .. - } - moreover have "k=0 \ k*m \ k*n" - using assms by simp - moreover have "m \ n \ k*m \ k*n" - using assms by (simp add: mult_leq_left_mono) - ultimately show ?thesis by blast -qed - -lemma mult_leq_cancel_right [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k \ n * k) = (k = 0 \ m \ n)" -using assms by (simp add: mult_commute_nat) - -lemma Suc_mult_less_cancel1: - assumes "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(Succ[k] * m < Succ[k] * n) = (m < n)" -using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] - nat_zero_less_Succ[of "k"] by auto - -lemma Suc_mult_leq_cancel1: - assumes "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(Succ[k] * m \ Succ[k] * n) = (m \ n)" -using assms by (simp del: mult_Succ_left_nat) - -lemma nat_leq_square: - assumes m: "m \ Nat" - shows "m \ m * m" -using m by (cases, auto) - -lemma nat_leq_cube: - assumes m: "m \ Nat" - shows "m \ m * (m * m)" -using m by (cases, auto) - -text \ Lemma for @{text gcd} \ -lemma mult_eq_self_implies_10: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") -proof - - from assms have "(m*n = m) = (m*n = m*1)" by simp - also have "\ = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) - finally show ?thesis . -qed - -end diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy deleted file mode 100644 index cdfa1eb4..00000000 --- a/isabelle/NatDivision.thy +++ /dev/null @@ -1,731 +0,0 @@ -(* Title: TLA+/NatDivision.thy - Author: Hernan Vanzetto - Copyright (C) 2009-2011 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:56 merz> -*) - -section \ The division operators div and mod on Naturals \ - -theory NatDivision -imports NatArith Tuples -begin - -subsection \ The divisibility relation \ - -definition dvd (infixl "dvd" 50) -where "a \ Nat \ b \ Nat \ b dvd a \ (\k \ Nat : a = b * k)" - -lemma boolify_dvd [simp]: - assumes "a \ Nat" and "b \ Nat" - shows "boolify(b dvd a) = (b dvd a)" -using assms by (simp add: dvd_def) - -lemma dvdIsBool [intro!,simp]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "isBool(b dvd a)" -using assms by (simp add: dvd_def) - -lemma [intro!]: - "\isBool(P); isBool(a dvd b); (a dvd b) \ P\ \ (a dvd b) = P" - "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" -by auto - -lemma dvdI [intro]: - assumes a: "a \ Nat" and b: "b \ Nat" and k: "k \ Nat" - shows "a = b * k \ b dvd a" -unfolding dvd_def[OF a b] using k by blast - -lemma dvdE [elim]: - assumes "b dvd a" and "a \ Nat" and "b \ Nat" - shows "(\k. \k \ Nat; a = b * k\ \ P) \ P" -using assms by (auto simp add: dvd_def) - -lemma dvd_refl [iff]: - assumes a: "a \ Nat" - shows "a dvd a" -proof - - from a have "a = a*1" by simp - with a show ?thesis by blast -qed - -lemma dvd_trans [trans]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "b dvd c" - shows "a dvd c" -proof - - from a b 1 obtain k where k: "k \ Nat" and "b = a * k" by blast - moreover - from b c 2 obtain l where l: "l \ Nat" and "c = b * l" by blast - ultimately have h:"c = a * (k * l)" - using a b c by (simp add: mult_assoc_nat) - thus ?thesis using a c k l by blast -qed - -lemma dvd_0_left_iff [simp]: - assumes "a \ Nat" - shows "(0 dvd a) = (a = 0)" -using assms by force - -lemma dvd_0_right [iff]: - assumes a: "a \ Nat" shows "a dvd 0" -using assms by force - -lemma one_dvd [iff]: - assumes a: "a \ Nat" - shows "1 dvd a" -using assms by force - -lemma dvd_mult (*[simp]*): - assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a dvd (b * c)" -proof - - from dvd a c obtain k where k: "k \ Nat" and "c = a*k" by blast - with a b c have "b*c = a * (b*k)" by (simp add: mult_left_commute_nat) - with a b c k show ?thesis by blast -qed - -lemma dvd_mult2 (*[simp]*): - assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a dvd (b * c)" -using mult_commute_nat[OF b c] dvd_mult[OF dvd a c b] by simp - -lemma dvd_triv_right [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "a dvd (b * a)" -using assms by (intro dvd_mult, simp+) - -lemma dvd_triv_left [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "a dvd a * b" -using assms by (intro dvd_mult2, simp+) - -lemma mult_dvd_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - and 1: "a dvd b" and 2: "c dvd d" - shows "(a * c) dvd (b * d)" -proof - - from a b 1 obtain b' where b': "b = a * b'" "b' \ Nat" by blast - from c d 2 obtain d' where d': "d = c * d'" "d' \ Nat" by blast - with b' a b c d - mult_assoc_nat[of a b' "c * d'"] - mult_left_commute_nat[of b' c d'] - mult_assoc_nat[of a c "b' * d'"] - have "b * d = (a * c) * (b' * d')" by simp - with a c b' d' show ?thesis by blast -qed - -lemma dvd_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and h: "a * b dvd c" - shows "a dvd c" -proof - - from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" - by (auto simp add: mult_assoc_nat) - with a b c show ?thesis by blast -qed - -lemma dvd_mult_right: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a*b dvd c" - shows "b dvd c" -proof - - from h a b c have "b*a dvd c" by (simp add: mult_ac_nat) - with b a c show ?thesis by (rule dvd_mult_left) -qed - -lemma dvd_0_left: - assumes "a \ Nat" - shows "0 dvd a \ a = 0" -using assms by simp - -lemma dvd_add [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "a dvd c" - shows "a dvd (b + c)" -proof - - from a b 1 obtain b' where b': "b' \ Nat" "b = a * b'" by blast - from a c 2 obtain c' where c': "c' \ Nat" "c = a * c'" by blast - from a b c b' c' - have "b + c = a * (b' + c')" by (simp add: add_mult_distrib_left_nat) - with a b' c' show ?thesis by blast -qed - -subsection \ Division on @{const Nat} \ - -text \ - We define division and modulo over @{const Nat} by means - of a characteristic relation with two input arguments - @{term "m"}, @{term "n"} and two output arguments - @{term "q"}(uotient) and @{term "r"}(emainder). - - The following definition works for natural numbers, but also for - possibly negative integers. Obviously, the second disjunct cannot - hold for natural numbers. -\ - -definition divmod_rel where - "divmod_rel(m,n,q,r) \ - m = q * n + r - \ - ((0 < n \ 0 \ r \ r < n) \ - (n < 0 \ r \ 0 \ n < r))" - -text \ @{const divmod_rel} is total if $n$ is non-zero. \ - -lemma divmod_rel_ex: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - obtains q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" -proof - - have "\q,r \ Nat : m = q*n + r \ r < n" - using m proof (induct) - case 0 - from n pos have "0 = 0 * n + 0 \ 0 < n" by simp - then show ?case by blast - next - fix m' - assume m': "m' \ Nat" and ih: "\q, r \ Nat : m' = q*n + r \ r < n" - from ih obtain q' r' - where h1: "m' = q' * n + r'" and h2: "r' < n" - and q': "q' \ Nat" and r': "r' \ Nat" by blast - show "\q, r \ Nat : Succ[m'] = q * n + r \ r < n" - proof (cases "Succ[r'] < n") - case True - from h1 h2 m' q' n r' have "Succ[m'] = q' * n + Succ[r']" by simp - with True q' r' show ?thesis by blast - next - case False - with n r' have "n \ Succ[r']" by (simp add: nat_not_less[simplified]) - with r' n h2 have "n = Succ[r']" by (intro nat_leq_antisym, simp+) - with h1 m' q' r' have "Succ[m'] = Succ[q'] * n + 0" by (simp add: add_ac_nat) - with pos q' show ?thesis by blast - qed - qed - with pos that show ?thesis by (auto simp: divmod_rel_def) -qed - -text \ @{const divmod_rel} has unique solutions in the natural numbers. \ -lemma divmod_rel_unique_div: - assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" - and m: "m \ Nat" and n: "n \ Nat" - and q: "q \ Nat" and r: "r \ Nat" and q': "q' \ Nat" and r': "r' \ Nat" - shows "q = q'" -proof - - from n 1 have pos: "0 < n" and mqr: "m = q*n+r" and rn: "r < n" - by (auto simp: divmod_rel_def) - from n 2 have mqr': "m = q'*n+r'" and rn': "r' < n" - by (auto simp: divmod_rel_def) - { - fix x y x' y' - assume nat: "x \ Nat" "y \ Nat" "x' \ Nat" "y' \ Nat" - and eq: "x*n + y = x'*n + y'" and less: "y' < n" - have "x \ x'" - proof (rule contradiction) - assume "\(x \ x')" - with nat have "x' < x" by (simp add: nat_not_leq[simplified]) - with nat obtain k where k: "k \ Nat" "x = Succ[x'+k]" - using less_iff_Succ_add by auto - with eq nat n have "x'*n + (k*n + n + y) = x'*n + y'" - by (simp add: add_mult_distrib_right_nat add_assoc_nat) - with nat k n have "k*n + n + y = y'" by simp - with less k n nat have s3_1: "(k*n + y) + n < n" by (simp add: add_ac_nat) - define j where "j \ k * n + y" - have s3_2: "j + n < n" - unfolding j_def using s3_1 by auto - have s3_3: "j \\in Nat" - unfolding j_def using multIsNat addIsNat k n nat by auto - show "FALSE" - using s3_2 s3_3 n not_add_less2[of "n" "j"] by auto - qed - } - from this[OF q r q' r'] this[OF q' r' q r] q q' mqr mqr' rn rn' - show ?thesis by (intro nat_leq_antisym, simp+) -qed - -lemma divmod_rel_unique_mod: - assumes "divmod_rel(m,n,q,r)" and "divmod_rel(m,n,q',r')" - and "m \ Nat" and "n \ Nat" and "q \ Nat" and "r \ Nat" and "q' \ Nat" and "r' \ Nat" - shows "r = r'" -proof - - from assms have "q = q'" by (rule divmod_rel_unique_div) - with assms show ?thesis by (auto simp: divmod_rel_def) -qed - -text \ - We instantiate divisibility on the natural numbers by - means of @{const divmod_rel}: -\ - -definition divmodNat -where "divmodNat(m,n) \ CHOOSE z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" - -lemma divmodNatPairEx: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "\z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" -proof - - from assms obtain q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" - by (rule divmod_rel_ex) - thus ?thesis unfolding two_def by force -qed - -lemma divmodNatInNatNat: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "divmodNat(m,n) \ Nat \ Nat" -unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]]) - -lemma divmodNat_divmod_rel [rule_format]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "z = divmodNat(m,n) \ divmod_rel(m,n,z[1],z[2])" -unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]], auto) - -lemma divmodNat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - and q: "q \ Nat" and r: "r \ Nat" - shows "divmodNat(m,n) = \q,r\" -unfolding divmodNat_def -proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) - fix z - assume "z \ Nat \ Nat" and "divmod_rel(m,n,z[1],z[2])" - with m n q r h show "z = \q,r\" - unfolding two_def - by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) -qed - -text \ - We now define division and modulus over natural numbers. -\ - -definition div (infixr "div" 70) -where div_nat_def: "\m \ Nat; n \ Nat\ \ m div n \ divmodNat(m,n)[1]" - -definition mod (infixr "mod" 70) -where mod_nat_def: "\m \ Nat; n \ Nat\ \ m mod n \ divmodNat(m,n)[2]" - - -lemma divIsNat [iff]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "m div n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) - -lemma modIsNat [iff]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "m mod n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def two_def) - -lemma divmodNat_div_mod: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "divmodNat(m,n) = \m div n, m mod n\" -unfolding div_nat_def[OF m n] mod_nat_def[OF m n] two_def -using divmodNatInNatNat[OF assms] -by force - -lemma divmod_rel_div_mod_nat: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "divmod_rel(m, n, m div n, m mod n)" -using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] -by (auto simp: two_def) - -lemma div_nat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" - shows "m div n = q" -unfolding div_nat_def[OF m n] using divmodNat_unique[OF assms] by simp - -lemma mod_nat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" - shows "m mod n = r" -unfolding mod_nat_def[OF m n] two_def using divmodNat_unique[OF assms] by simp - -lemma mod_nat_less_divisor: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "m mod n < n" -using assms divmod_rel_div_mod_nat[OF assms] by (simp add: divmod_rel_def) - -text \ ``Recursive'' computation of @{text div} and @{text mod}. \ - -lemma divmodNat_base: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "divmodNat(m,n) = \0, m\" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - let ?dm = "divmodNat(m,n)" - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" - by (simp add: divmodNat_divmod_rel) - from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - have "?dm[1] * n < n" - proof - - have s2_1: "m = ?dm[1] * n + ?dm[2]" - using 1 by (auto simp: divmod_rel_def) - have s2_2: "?dm[1] * n + ?dm[2] < n" - using less s2_1 by auto - have s2_3: "?dm[1] \\in Nat" - using 2 by auto - have s2_4: "?dm[2] \\in Nat" - using 2 by (auto simp: two_def) - show ?thesis - using s2_2 s2_3 s2_4 multIsNat n add_lessD1 by auto - qed - with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) - - with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) - - with 3 prodProj[OF 2] show ?thesis by simp -qed - -lemma divmodNat_step: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "divmodNat(m,n) = \Succ[(m -- n) div n], (m -- n) mod n\" -proof - - from m n pos have 1: "divmod_rel(m, n, m div n, m mod n)" - by (rule divmod_rel_div_mod_nat) - have 2: "m div n \ 0" - proof - assume "m div n = 0" - with 1 m n pos have s2_1: "m < n" by (auto simp: divmod_rel_def) - have s2_2: "\ (n \ m)" - using s2_1 n m nat_not_leq by auto - show "FALSE" - using geq s2_2 by auto - qed - with m n pos obtain k where k1: "k \ Nat" and k2: "m div n = Succ[k]" - using not0_implies_Suc[of "m div n"] by auto - with 1 m n pos have "m = n + k*n + m mod n" - by (auto simp: divmod_rel_def add_commute_nat) - moreover - from m n k1 pos geq have "... -- n = k*n + m mod n" - by (simp add: adiff_add_assoc2) - ultimately - have "m -- n = k*n + m mod n" by simp - with pos m n 1 have "divmod_rel(m -- n, n, k, m mod n)" - by (auto simp: divmod_rel_def) - - with k1 m n pos have "divmodNat(m -- n, n) = \k, m mod n\" - by (intro divmodNat_unique, simp+) - moreover - from m n pos have "divmodNat(m -- n, n) = \(m--n) div n, (m--n) mod n\" - by (intro divmodNat_div_mod, simp+) - ultimately - have "m div n = Succ[(m--n) div n]" and "m mod n = (m--n) mod n" - using m n k2 by auto - thus ?thesis by (simp add: divmodNat_div_mod[OF m n pos]) -qed - - -text \ The ''recursion'' equations for @{text div} and @{text mod} \ - -lemma div_nat_less [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "m div n = 0" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - from divmodNat_base[OF m n less] divmodNat_div_mod[OF m n pos] show ?thesis - by simp -qed - -lemma div_nat_geq: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "m div n = Succ[(m -- n) div n]" -using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] -by simp - -lemma mod_nat_less [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "m mod n = m" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - from divmodNat_base[OF m n less] divmodNat_div_mod[OF m n pos] show ?thesis - by simp -qed - -lemma mod_nat_geq: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "m mod n = (m -- n) mod n" -using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] -by simp - - -subsection \ Facts about @{const div} and @{const mod} \ - -lemma mod_div_nat_equality [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m div n) * n + m mod n = m" -using divmod_rel_div_mod_nat [OF assms] by (simp add: divmod_rel_def) - -lemma mod_div_nat_equality2 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "n * (m div n) + m mod n = m" -using assms mult_commute_nat[of "n" "m div n"] by simp - -lemma mod_div_nat_equality3 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" - shows "m mod n + (m div n) * n = m" -using assms add_commute_nat[of "m mod n"] by simp - -lemma mod_div_nat_equality4 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" - shows "m mod n + n * (m div n) = m" -using assms mult_commute_nat[of "n" "m div n"] by simp - -(** immediate consequence of above -lemma div_mod_nat_equality: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "((m div n) * n + m mod n) + k = m + k" -using assms by simp - -lemma div_mod_nat_equality2: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * (m div n) + m mod n) + k = m + k" -using assms by simp -**) - -lemma div_nat_mult_self1 [simp]: - assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(q + m * n) div n = m + (q div n)" (is "?P(m)") -using m proof (induct m) - from assms show "?P(0)" by simp -next - fix k - assume k: "k \ Nat" and ih: "?P(k)" - from n q k have "n \ q + (k*n + n)" by (simp add: add_assoc_nat) - with q k n pos have "(q + (k*n + n)) div n = Succ[(q + k*n) div n]" - by (simp add: div_nat_geq add_assoc_nat) - with ih q m n k pos show "?P(Succ[k])" by simp -qed - -lemma div_nat_mult_self2 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(q + n * m) div n = m + q div n" -using assms by (simp add: mult_commute_nat) - -lemma div_nat_mult_self3 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(m * n + q) div n = m + q div n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_mult_self4 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(n * m + q) div n = m + q div n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_0: - assumes "n \ Nat" and "0 < n" - shows "0 div n = 0" -using assms by simp - -lemma mod_0: - assumes "n \ Nat" and "0 < n" - shows "0 mod n = 0" -using assms by simp - -lemma mod_nat_mult_self1 [simp]: - assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(q + m * n) mod n = q mod n" -proof - - from assms have "m*n + q = q + m*n" - by (simp add: add_commute_nat) - also from assms have "... = ((q + m*n) div n) * n + (q + m*n) mod n" - by (intro sym[OF mod_div_nat_equality], simp+) - also from assms have "... = (m + q div n) * n + (q + m*n) mod n" - by simp - also from assms have "... = m*n + ((q div n) * n + (q + m*n) mod n)" - by (simp add: add_mult_distrib_right_nat add_assoc_nat) - finally have "q = (q div n) * n + (q + m*n) mod n" - using assms by simp - with q n pos have "(q div n) * n + (q + m*n) mod n = (q div n) * n + q mod n" - by simp - with assms show ?thesis by (simp del: mod_div_nat_equality) -qed - -lemma mod_nat_mult_self2 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(q + n * m) mod n = q mod n" -using assms by (simp add: mult_commute_nat) - -lemma mod_nat_mult_self3 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n + q) mod n = q mod n" -using assms by (simp add: add_commute_nat) - -lemma mod_nat_mult_self4 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m + q) mod n = q mod n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_mult_self1_is_id [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n) div n = m" -using assms div_nat_mult_self1 [of 0 m n] by simp - -lemma div_nat_mult_self2_is_id [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m) div n = m" -using assms div_nat_mult_self2[of 0 n m] by simp - -lemma mod_nat_mult_self1_is_0 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n) mod n = 0" -using assms mod_nat_mult_self1 [of 0 m n] by simp - -lemma mod_nat_mult_self2_is_0 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m) mod n = 0" -using assms mod_nat_mult_self2 [of 0 m n] by simp - -lemma div_nat_by_1 [simp]: - assumes "m \ Nat" - shows "m div 1 = m" -using assms div_nat_mult_self1_is_id [of m 1] by simp - -lemma mod_nat_by_1 [simp]: - assumes "m \ Nat" - shows "m mod 1 = 0" -using assms mod_nat_mult_self1_is_0[of m 1] by simp - -lemma mod_nat_self [simp]: - assumes "n \ Nat" and "0 < n" - shows "n mod n = 0" -using assms mod_nat_mult_self1_is_0[of 1] by simp - -lemma div_nat_self [simp]: - assumes "n \ Nat" and "0 < n" - shows "n div n = 1" -using assms div_nat_mult_self1_is_id [of 1 n] by simp - -lemma div_nat_add_self1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m + n) div n = m div n + 1" -using assms div_nat_mult_self1[OF m oneIsNat n pos] by simp - -lemma div_nat_add_self2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(n + m) div n = m div n + 1" -using assms div_nat_mult_self3[OF m n oneIsNat pos] by simp - -lemma mod_nat_add_self1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m + n) mod n = m mod n" -using assms mod_nat_mult_self1[OF m oneIsNat n pos] by simp - -lemma mod_nat_add_self2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(n + m) mod n = m mod n" -using assms mod_nat_mult_self3[OF m oneIsNat n pos] by simp - -lemma div_mod_nat_decomp: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - obtains q r where "q \ Nat" and "r \ Nat" - and "q = m div n" and "r = m mod n" and "m = q * n + r" -proof - - from m n pos have "m = (m div n) * n + (m mod n)" by simp - with assms that show ?thesis by blast -qed - -lemma dvd_nat_eq_mod_eq_0: - assumes "m \ Nat" and "n \ Nat" and "0 < m" - shows "(m dvd n) = (n mod m = 0)" (is "?lhs = ?rhs") -proof - - from assms have 1: "?lhs \ ?rhs" by auto - have 2: "?rhs \ ?lhs" - proof - assume mod: "n mod m = 0" - with assms mod_div_nat_equality[of n m] have "(n div m) * m = n" by simp - with assms have "n = m * (n div m)" by (simp add: mult_commute_nat) - with assms show "m dvd n" by blast - qed - from 1 2 assms show ?thesis by blast -qed - -lemma mod_div_nat_trivial [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m mod n) div n = 0" -proof - - from assms - have "m div n + (m mod n) div n = (m mod n + (m div n) * n) div n" - using div_nat_mult_self1[of "m mod n" "m div n" "n"] - by (auto simp: mod_nat_less_divisor) - also from assms have "... = m div n + 0" by simp - finally show ?thesis - using assms by simp -qed - -lemma mod_mod_nat_trivial [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m mod n) mod n = m mod n" -proof - - from assms mod_nat_mult_self1[of "m mod n" "m div n" "n"] - have "(m mod n) mod n = (m mod n + (m div n) * n) mod n" by simp - also from assms have "... = m mod n" by simp - finally show ?thesis . -qed - -lemma dvd_nat_imp_mod_0: - assumes "n dvd m" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "m mod n = 0" -using assms by (simp add: dvd_nat_eq_mod_eq_0) - -lemma dvd_nat_div_mult_self: - assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m div n) * n = m" -using assms - dvd_nat_imp_mod_0[OF assms] - mod_div_nat_equality[OF m n pos] -by simp - -lemma dvd_nat_div_mult: - assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - and k: "k \ Nat" - shows "(m div n) * k = (m * k) div n" -proof - - from dvd m n obtain l where l: "l \ Nat" "m = n*l" by auto - with m n k have "m * k = n * (l*k)" by (simp add: mult_assoc_nat) - with m n k l pos show ?thesis by simp -qed - -lemma div_nat_dvd_div [simp]: - assumes 1: "a dvd b" and 2: "a dvd c" - and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and pos: "0 < a" - shows "(b div a) dvd (c div a) = (b dvd c)" -proof (auto) - assume lhs: "(b div a) dvd (c div a)" - with a b c pos have "((b div a) * a) dvd ((c div a) * a)" - by (intro mult_dvd_mono, simp+) - moreover - from 1 a b pos have "(b div a) * a = b" by (simp add: dvd_nat_div_mult_self) - moreover - from 2 a c pos have "(c div a) * a = c" by (simp add: dvd_nat_div_mult_self) - ultimately show "b dvd c" by simp -next - assume rhs: "b dvd c" - with b c obtain k where k: "k \ Nat" "c = b*k" by auto - from 1 a b obtain l where l: "l \ Nat" "b = a*l" by auto - with a pos have 3: "b div a = l" by simp - from 2 a c obtain m where m: "m \ Nat" "c = a*m" by auto - with a pos have 4: "c div a = m" by simp - from k l m a pos mult_assoc_nat[of a l k, symmetric] have "m = l*k" by auto - with k l m have "l dvd m" by auto - with 3 4 show "(b div a) dvd (c div a)" by simp -qed (auto simp: assms) - -lemma dvd_mod_nat_imp_dvd: - assumes 1: "k dvd (m mod n)" and 2: "k dvd n" - and k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "k dvd m" -proof - - from assms have "k dvd ((m div n) * n + m mod n)" - by (simp add: dvd_mult del: mod_div_nat_equality) - with m n pos show ?thesis by simp -qed - -(*** TODO -text \ Addition respects modular equivalence. \ - -text \ Multiplication respects modular equivalence. \ - -text \ Negation respects modular equivalence. \ - -text \ Subtraction respects modular equivalence. \ -*) - -end diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy deleted file mode 100644 index 97e45a2f..00000000 --- a/isabelle/NatOrderings.thy +++ /dev/null @@ -1,709 +0,0 @@ -(* Title: TLA+/NatOrderings.thy - Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - - -section \ Orders on natural numbers \ - -theory NatOrderings -imports Peano -begin - -text \ - Using the sets @{text upto} we can now define the standard ordering on - natural numbers. The constant @{text "\"} is defined over the naturals - by the axiom (conditional definition) @{text "nat_leq_def"} below; it - should be defined over other domains as appropriate later on. - - We generally define the constant @{text "<"} such that @{text "ab \ a\b"}, for any values. -\ - -definition leq :: "[c,c] \ c" (infixl "\" 50) -(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m \ n) \ (m \ upto(n))"*) -where nat_leq_def: "(m \ n) \ (m \ upto(n))" - -abbreviation (input) - geq :: "[c,c] \ c" (infixl "\" 50) -where "x \ y \ y \ x" - -notation (ASCII) - leq (infixl "<=" 50) and - geq (infixl ">=" 50) - -lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" -by (rule ssubst) - -lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" -by (rule subst) - -subsection \ Operator definitions and generic facts about @{text "<"} \ - -definition less :: "[c,c] \ c" (infixl "<" 50) -where "a < b \ a \ b \ a \ b" - -abbreviation (input) - greater :: "[c,c] \ c" (infixl ">" 50) -where "x > y \ y < x" - -lemma boolify_less [simp]: "boolify(a < b) = (a < b)" -by (simp add: less_def) - -lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" -by (rule ssubst) - -lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" -by (rule subst) - -lemma less_imp_leq [elim!]: "a < b \ a \ b" -unfolding less_def by simp - -lemma less_irrefl [simp]: "(a < a) = FALSE" -unfolding less_def by simp - -lemma less_irreflE [elim!]: "a < a \ R" -by simp - -lemma less_not_refl: "a < b \ a \ b" -by auto - -lemma neq_leq_trans [trans]: "a \ b \ a \ b \ a < b" -by (simp add: less_def) - -declare neq_leq_trans[simplified,trans] - -lemma leq_neq_trans [trans,elim!]: "a \ b \ a \ b \ a < b" -by (simp add: less_def) - -declare leq_neq_trans[simplified,trans] - -(* Don't add to [simp]: will be tried on all disequalities! *) -lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" -by auto - -subsection \ Facts about @{text "\"} over @{text Nat} \ - -lemma nat_boolify_leq [simp]: "boolify(m \ n) = (m \ n)" -by (simp add: nat_leq_def) - -lemma nat_leq_isBool [intro,simp]: "isBool(m \ n)" -by (simp add: nat_leq_def) - -lemma nat_leq_refl [intro,simp]: "n \ Nat \ n \ n" -unfolding nat_leq_def by (rule uptoRefl) - -lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ - assumes "m \ Nat" and "n \ Nat" and "m = n" and "\ m \ n; n \ m \ \ P" - shows "P" -using assms by simp - -lemma nat_zero_leq [simp]: "n \ Nat \ 0 \ n" -unfolding nat_leq_def by (rule zeroInUpto) - -lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" -by (simp add: nat_leq_def uptoZero) - -lemma nat_leq_SuccI [(*elim!,*)intro]: - assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" - shows "m \ Succ[n]" -using assms by (auto simp: nat_leq_def uptoSucc) - -(* Do not make this "simp": interferes with other rules, e.g. simplifying "Succ[m] \ Succ[n]" *) -lemma nat_leq_Succ: - assumes (*"m \ Nat" and*) "n \ Nat" - shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" -using assms by (auto simp: nat_leq_def uptoSucc) - -lemma nat_is_leq_Succ [simp]: - assumes "n \ Nat" - shows "n \ Succ[n]" -using assms by (simp add: nat_leq_Succ) - -lemma nat_leq_SuccE [elim]: - assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" - and "m \ n \ P" and "m = Succ[n] \ P" - shows "P" -using assms by (auto simp: nat_leq_Succ) - -lemma nat_leq_limit: - assumes "m \ n" and "\(Succ[m] \ n)" and (*"m \ Nat" and*) "n \ Nat" - shows "m=n" -using assms by (auto simp: nat_leq_def intro: uptoLimit) - -lemma nat_leq_trans [trans]: - assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" - shows "k \ n" -using assms by (auto simp: nat_leq_def elim: uptoTrans) - -lemma nat_leq_antisym: - assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) - shows "m = n" -using assms by (auto simp add: nat_leq_def elim: uptoAntisym) - -lemma nat_Succ_not_leq_self [simp]: - assumes n: "n \ Nat" - shows "(Succ[n] \ n) = FALSE" -using n by (auto simp: nat_leq_Succ dest: nat_leq_antisym) - -lemma nat_Succ_leqI: - assumes "m \ n" and "m \ Nat" and "n \ Nat" and "m \ n" - shows "Succ[m] \ n" -using assms by (auto dest: nat_leq_limit) - -lemma nat_Succ_leqD [elim]: - assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ n" -proof - - from m have "m \ Succ[m]" by (simp add: nat_leq_Succ) - with leq m n show ?thesis by (elim nat_leq_trans, auto) -qed - -lemma nat_Succ_leq_Succ [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m] \ Succ[n]) = (m \ n)" -using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit) - -lemma nat_Succ_leq: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] \ n) = (m \ n \ m \ n)" - using assms by (auto intro: nat_Succ_leqI) - -lemma nat_leq_linear: "\m \ Nat; n \ Nat\ \ m \ n \ n \ m" -unfolding nat_leq_def using uptoLinear . - -lemma nat_leq_cases: - assumes m: "m \ Nat" and n: "n \ Nat" - and leq: "m \ n \ P" and - geq: "\n \ m; - n \ m - \ \ P" - shows "P" -proof (cases "m \ n") - case True thus "P" by (rule leq) -next - case False - with m n have nm: "n \ m" by (blast dest: nat_leq_linear) - thus "P" - proof (cases "n=m") - case True - with m have "m \ n" by simp - thus "P" by (rule leq) - next - case False - with nm show "P" by (rule geq) - qed -qed - -lemma nat_leq_induct: \ \sometimes called ``complete induction''\ - assumes "P(0)" - and "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(Succ[n])" - shows "\n\Nat : P(n)" -proof - - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp: nat_leq_Succ) - thus ?thesis by (blast dest: nat_leq_refl) -qed - -lemma nat_leq_inductE: - assumes "n \ Nat" - and "P(0)" and "\n. \n \ Nat; \m\Nat : m \ n \ P(m)\ \ P(Succ[n])" - shows "P(n)" -using assms by (blast dest: nat_leq_induct) - - -subsection \ Facts about @{text "<"} over @{text Nat} \ - -lemma nat_less_iff_Succ_leq [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m < n) = (Succ[m] \ n)" -using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) - -text \ Reduce @{text "\"} to @{text "<"}. \ - -lemma nat_leq_less: - assumes (*"m \ Nat" and*) "n \ Nat" - shows "m \ n = (m < n \ m = n)" - using assms by (auto simp: less_def) - -lemma nat_leqE: - assumes "m \ n" and "n \ Nat" and 1: "m P" and 2: "m=n \ P" - shows "P" -proof - - from \m \ n\ \n \ Nat\ have "m < n \ m=n" - by (auto simp add: less_def) - with 1 2 show ?thesis by blast -qed - -lemma nat_less_Succ_iff_leq (*[simp]*): - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m \ n)" - using assms by simp - -lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq] - -lemma not_leq_Succ_leq [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "((m \ n) = FALSE) = (Succ[n] \ m)" -proof - - from assms nat_leq_linear[of m n] have "((m \ n) = FALSE) = (n < m)" - by (auto simp: less_def dest: nat_leq_antisym) - with assms show ?thesis by simp -qed - -lemma nat_not_leq_one: - assumes "n \ Nat" - shows "(\(1 \ n)) = (n = 0)" -using assms by simp - -text \ @{text "<"} and @{text "Succ"}. \ - -lemma nat_Succ_less_mono: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] < Succ[n]) = (m < n)" -using assms by simp - -lemma nat_Succ_less_SuccE: - assumes "Succ[m] < Succ[n]" and "m \ Nat" and "n \ Nat" and "m < n \ P" - shows "P" -using assms by simp - -lemma nat_not_less0 [simp]: - assumes "n \ Nat" - shows "(n < 0) = FALSE" -using assms by auto - -lemma nat_less0E (*[elim!]*): - assumes "n < 0" and "n \ Nat" - shows "P" -using assms by simp - -lemma nat_less_SuccI: - assumes "m < n" and "m \ Nat" and "n \ Nat" - shows "m < Succ[n]" -using assms by auto - -lemma nat_Succ_lessD: - assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" - shows "m < n" -using assms by auto - -lemma nat_less_leq_not_leq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m < n) = (m \ n \ \ n \ m)" -using assms by (auto simp: less_def dest: nat_leq_antisym) - -text \ Transitivity. \ - -lemma nat_less_trans (*[trans]*): - assumes "k < m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -lemma nat_less_trans_Succ [trans]: - assumes lt1: "i < j" and lt2: "j < k" - and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "Succ[i] < k" -proof - - from i j lt1 have "Succ[Succ[i]] \ Succ[j]" by simp - also from j k lt2 have "Succ[j] \ k" by simp - finally show ?thesis using i j k by simp -qed - -lemma nat_leq_less_trans [trans]: - assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -lemma nat_less_leq_trans [trans]: - assumes "k < m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -text \ Asymmetry. \ - -lemma nat_less_not_sym: - assumes "m < n" and "m \ Nat" and "n \ Nat" - shows "(n < m) = FALSE" -using assms by auto - -lemma nat_less_asym: - assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" - shows "P" -proof (rule contradiction) - assume "\P" with nat_less_not_sym[OF \m \m \ Nat\ \n \ Nat\] \\P \ n < m\ - show "FALSE" by simp -qed - -text \ Linearity (totality). \ - -lemma nat_less_linear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m < n \ m = n \ n < m" -unfolding less_def using nat_leq_linear[OF m n] by blast - -lemma nat_leq_less_linear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n \ n < m" -using assms nat_less_linear[OF m n] by (auto simp: less_def) - -lemma nat_less_cases [case_names less equal greater]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m < n \ P) \ (m = n \ P) \ (n < m \ P) \ P" -using nat_less_linear[OF m n] by blast - -lemma nat_not_less: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (n \ m)" -using assms nat_leq_linear[OF m n] by (auto simp: less_def dest: nat_leq_antisym) - -lemma nat_not_less_iff_gr_or_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (m > n \ m = n)" -unfolding nat_not_less[OF m n] using assms by (auto simp: less_def) - -lemma nat_not_less_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (n < Succ[m])" -unfolding nat_not_less[OF m n] using assms by simp - -lemma nat_not_leq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m \ n) = (n < m)" -using assms by simp - -text \often useful, but not active by default\ -lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq - -lemma nat_not_leq_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m \ n) = (Succ[n] \ m)" -unfolding nat_not_leq[OF m n] using assms by simp - -lemma nat_neq_iff: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n = (m < n \ n < m)" -using assms nat_less_linear[OF m n] by auto - -lemma nat_neq_lessE: - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "(m < n \ R) \ (n < m \ R) \ R" -using assms by (auto simp: nat_neq_iff[simplified]) - -lemma nat_antisym_conv1: - assumes "\(m < n)" and "m \ Nat" and "n \ Nat" - shows "(m \ n) = (m = n)" -using assms by (auto simp: nat_leq_less) - -lemma nat_antisym_conv2: - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -lemma nat_antisym_conv3: - assumes "\ n < m" and "m \ Nat" and "n \ Nat" - shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -lemma nat_not_lessD: - assumes "\(m < n)" and "m \ Nat" and "n \ Nat" - shows "n \ m" -using assms by (simp add: nat_not_order_simps) - -lemma nat_not_lessI: - assumes "n \ m" and "m \ Nat" and "n \ Nat" - shows "\(m < n)" -using assms by (simp add: nat_not_order_simps) - -lemma nat_gt0_not0 (*[simp]*): - assumes "n \ Nat" - shows "(0 < n) = (n \ 0)" -using assms by (auto simp: nat_neq_iff[simplified]) - -lemmas nat_neq0_conv = sym[OF nat_gt0_not0] - -text \ Introduction properties \ - -lemma nat_less_Succ_self (*[iff]*): - assumes "n \ Nat" - shows "n < Succ[n]" -using assms by simp - -lemma nat_zero_less_Succ (*[iff]*): - assumes "n \ Nat" - shows "0 < Succ[n]" -using assms by simp - -text \ Elimination properties \ - -lemma nat_less_Succ: - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m < n \ m = n)" -using assms by (simp add: nat_leq_less[of n m]) - -lemma nat_less_SuccE: - assumes "m < Succ[n]" and "m \ Nat" and "n \ Nat" - and "m < n \ P" and "m = n \ P" - shows P -using assms by (auto simp: nat_leq_less[of n m]) - -lemma nat_less_one (*[simp]*): - assumes "n \ Nat" - shows "(n < 1) = (n = 0)" -using assms by simp - -text \ "Less than" is antisymmetric, sort of \ - -lemma nat_less_antisym: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "\ \(n < m); n < Succ[m] \ \ m = n" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ -lemma less_mono_imp_leq_mono: - assumes i: "i \ Nat" and j: "j \ Nat" and f: "\n \ Nat : f(n) \ Nat" - and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" - shows "f(i) \ f(j)" -using ij proof (auto simp: nat_leq_less[OF j]) - assume "i < j" - with i j have "f(i) < f(j)" by (rule mono) - thus "f(i) \ f(j)" by (simp add: less_imp_leq) -next - from j f show "f(j) \ f(j)" by auto -qed - -text \ Inductive (?) properties. \ - -lemma nat_Succ_lessI: - assumes "m \ Nat" and "n \ Nat" and "m < n" and "Succ[m] \ n" - shows "Succ[m] < n" -using assms by (simp add: leq_neq_iff_less[simplified]) - -lemma nat_lessE: - assumes major: "i < k" and i: "i \ Nat" and k: "k \ Nat" - obtains j where "j \ Nat" and "i \ j" and "k = Succ[j]" -proof - - from k major have "\j\Nat : i \ j \ k = Succ[j]" - proof (induct k) - case 0 with i show ?case by simp - next - fix n - assume n: "n \ Nat" and 1: "i < Succ[n]" - and ih: "i < n \ \j\Nat : i \ j \ n = Succ[j]" - from 1 i n show "\j\Nat : i \ j \ Succ[n] = Succ[j]" - proof (rule nat_less_SuccE) - assume "i < n" - then obtain j where "j \ Nat" and "i \ j" and "n = Succ[j]" - by (blast dest: ih) - with i have "Succ[j] \ Nat" and "i \ Succ[j]" and "Succ[n] = Succ[Succ[j]]" - by auto - thus ?thesis by blast - next - assume "i = n" - with i show ?thesis by blast - qed - qed - with that show ?thesis by blast -qed - -lemma nat_Succ_lessE: - assumes major: "Succ[i] < k" and i: "i \ Nat" and k: "k \ Nat" - obtains j where "j \ Nat" and "i < j" and "k = Succ[j]" -proof - - from i have "Succ[i] \ Nat" .. - from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" - by (rule nat_lessE) - with i that show ?thesis by simp -qed - -(* -lemma nat_gt0_implies_Succ: - assumes 1: "0 < n" and 2: "n \ Nat" - shows "\m : m \ Nat \ n = Succ[m]" -using 2 1 by (cases, auto) -*) -lemma nat_ge1_implies_Succ: - assumes 1: "1 \ n" and 2: "n \ Nat" - shows "\m : m \ Nat \ n = Succ[m]" - using 2 1 by (cases, auto) - -lemma nat_gt0_iff_Succ: - assumes n: "n \ Nat" - shows "(1 \ n) = (\m \ Nat: n = Succ[m])" -using n by (auto dest: nat_ge1_implies_Succ) - -(* -lemma nat_less_Succ_eq_0_disj: - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" -using assms by (induct m, auto) -*) -lemma nat_le_eq_0_disj: - assumes "m \ Nat" and "n \ Nat" - shows "(m \ n) = (m = 0 \ (\j\Nat : m = Succ[j] \ j < n))" - using assms by (induct m, auto) - -lemma nat_less_antisym_false: "\m < n; m \ Nat; n \ Nat\ \ n < m = FALSE" - by auto - -lemma nat_less_antisym_leq_false: "\m < n; m \ Nat; n \ Nat\ \ n \ m = FALSE" - by auto - - -subsection \Intervals of natural numbers\ - -definition natInterval :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) -where "m .. n \ { k \ Nat : m \ k \ k \ n }" - -lemma inNatIntervalI [intro!,simp]: - assumes "k \ Nat" and "m \ k" and "k \ n" - shows "k \ m .. n" -using assms by (simp add: natInterval_def) - -lemma inNatIntervalE [elim]: - assumes 1: "k \ m .. n" and 2: "\k \ Nat; m \ k; k \ n\ \ P" - shows P -using 1 by (intro 2, auto simp add: natInterval_def) - -lemma inNatInterval_iff: "(k \ m .. n) = (k \ Nat \ m \ k \ k \ n)" -by auto - -lemmas - setEqualI [where A = "m .. n" for m n, intro] - setEqualI [where B = "m .. n" for m n, intro] - -lemma lowerInNatInterval [iff]: - assumes "m \ n" and "m \ Nat" - shows "m \ m .. n" -using assms by (simp add: natInterval_def) - -lemma upperInNatInterval [iff]: - assumes "m \ n" and "n \ Nat" - shows "n \ m .. n" -using assms by (simp add: natInterval_def) - -lemma gtNotinNatInterval: - assumes gt: "m > n" and k: "k \ m .. n" and m: "m \ Nat" and n: "n \ Nat" - shows "P" -proof - - from k have 1: "m \ k" and 2: "k \ n" by auto - from 1 2 n have "m \ n" by (rule nat_leq_trans) - with m n have "\(n < m)" by (simp add: nat_not_order_simps) - from this gt show ?thesis .. -qed - -lemma natIntervalIsEmpty: - assumes "m \ Nat" and "n \ Nat" and "m > n" - shows "m .. n = {}" -using assms by (blast dest: gtNotinNatInterval) - -lemma natIntervalEmpty_iff: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m .. n = {}) = (m > n)" -proof (auto dest: natIntervalIsEmpty[OF m n]) - assume mt: "m .. n = {}" - show "n < m" - proof (rule contradiction) - assume "\(n < m)" - with m n have "m \ n" by (simp add: nat_not_order_simps) - from this m have "m \ m .. n" by (rule lowerInNatInterval) - with mt show "FALSE" by blast - qed -qed - -lemma natIntervalSingleton [simp]: - assumes "n \ Nat" - shows "n .. n = {n}" -using assms by (auto dest: nat_leq_antisym) - -lemma natIntervalSucc [simp]: - assumes "m \ Nat" and "n \ Nat" and "m \ Succ[n]" - shows "m .. Succ[n] = addElt(Succ[n], m .. n)" -using assms by (auto simp: natInterval_def) - -lemma succNatInterval: - assumes "m \ Nat" and "n \ Nat" - shows "Succ[m] .. n = (m .. n \ {m})" -using assms by (auto simp: natInterval_def nat_Succ_leq) - -lemma natIntervalEqual_iff: - assumes k: "k \ Nat" and l: "l \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k .. l = m .. n) = ((k > l \ m > n) \ (k = m \ l = n))" (is "?lhs = ?rhs") -proof - - have 1: "?lhs \ ?rhs" - proof - assume eq: "?lhs" - show ?rhs - proof (cases "k .. l = {}") - case True - with k l have "k > l" by (simp only: natIntervalEmpty_iff) - moreover - from True eq m n have "m > n" by (simp only: natIntervalEmpty_iff) - ultimately - show ?rhs by blast - next - case False - with k l have 11: "k \ l" by (simp only: natIntervalEmpty_iff nat_not_less) - from False eq m n have 12: "m \ n" by (simp only: natIntervalEmpty_iff nat_not_less) - from 11 k eq have 13: "m \ k" by auto - from 12 m eq have 14: "k \ m" by auto - from 14 13 have 15: "k = m" by (rule nat_leq_antisym) - from 11 l eq have 16: "l \ n" by auto - from 12 n eq have 17: "n \ l" by auto - from 16 17 have "l = n" by (rule nat_leq_antisym) - with 15 show ?rhs by blast - qed - qed - have 2: "?rhs \ ?lhs" - proof auto - assume lk: "l < k" and nm: "n < m" - from k l lk have "k .. l = {}" by (rule natIntervalIsEmpty) - moreover - from m n nm have "m .. n = {}" by (rule natIntervalIsEmpty) - ultimately - show ?lhs by auto - qed - from 1 2 show ?thesis by blast -qed - -lemma zerotoInj [simp]: - assumes "l \ Nat" and "m \ Nat" and "n \ Nat" - shows "(0 .. l = m .. n) = (m=0 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma zerotoInj' [simp]: - assumes "k \ Nat" and "l \ Nat" and "n \ Nat" - shows "(k .. l = 0 .. n) = (k=0 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma zerotoEmpty [simp]: - assumes "m \ Nat" - shows "Succ[m] .. 0 = {}" -using assms by auto - -lemma onetoInj [simp]: - assumes "l \ Nat" and "m \ Nat" and "n \ Nat" and "l\0 \ m=1" - shows "(1 .. l = m .. n) = (m=1 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma onetoInj' [simp]: - assumes "k \ Nat" and "l \ Nat" and "n \ Nat" and "n\0 \ k=1" - shows "(k .. l = 1 .. n) = (k=1 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma SuccInNatIntervalSucc: - assumes "m \ k" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(Succ[k] \ m .. Succ[n]) = (k \ m .. n)" -using assms by auto - -lemma SuccInNatIntervalSuccSucc: - assumes "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(Succ[k] \ Succ[m] .. Succ[n]) = (k \ m .. n)" -using assms by auto - -end diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy deleted file mode 100644 index 103b7221..00000000 --- a/isabelle/NewIntegers.thy +++ /dev/null @@ -1,12656 +0,0 @@ -(* An attempt at defining the integers. - -The steps in Isabelle/Isar are named `si_j` where `s` stands for "step", -`i` is the proof level, and `j` the step number. - -Author: Ioannis Filippidis -All rights reserved. Licensed under 2-clause BSD. -*) -theory NewIntegers -imports - NatArith - (*NatDivision - Tuples*) -begin - -(* -Properties to prove: - typeness, - commutativity, associativity, distributivity - transitivity, symmetry, antisymmetry, reflexivity, - monotonicity -*) - -(*----------------------------------------------------------------------------*) -(* Overriding notation. *) - -no_notation NatOrderings.leq (infixl "<=" 50) -no_notation NatOrderings.geq (infixl ">=" 50) -no_notation NatOrderings.less (infixl "<" 50) -no_notation NatOrderings.greater (infixl ">" 50) -no_notation NatOrderings.natInterval ("(_ .. _)" [90,90] 70) - -no_notation NatArith.arith_add (infixl "+" 65) -(* no_notation NatArith.adiff (infixl "--" 65) *) -no_notation NatArith.mult (infixl "*" 70) - -(* no_notation NatDivision.dvd (infixl "dvd" 50) *) - - -(*----------------------------------------------------------------------------*) -(* The minus operator. *) - - -(* -The minus operator can be defined, instead of axiomatized, as follows. - -minus(x) == - LET singletons == {<>: u \in Nat} - IN - IF x \in Nat THEN - <> - ELSE IF x \in Singletons THEN - CHOOSE u \in Nat: x = <> - ELSE - CHOOSE u \notin (Nat \cup Singletons): TRUE -*) -consts - "minus" :: "[c] \ c" ("-._" [75] 75) - -axiomatization where - neg0 [simp]: "-.0 = 0" -and - neg_neg: "n \ Nat \ -.n \ Nat ==> - -.-.n = n" -and - neg_not_in_nat: - "n \ (Nat \ {0}) ==> - -.n \ Nat" - - -(*----------------------------------------------------------------------------*) -(* The set of integers. *) - -definition Int -where "Int \ - Nat \ {-.n: n \ Nat}" - - -definition negNat :: "c" -where "negNat \ {-.n: n \ Nat}" - - -(*----------------------------------------------------------------------------*) -(* Successor and predecessor on integers. *) - -definition iSucc :: "c" -where "iSucc \ - [i \ Int \ IF i \ Nat THEN Succ[i] - ELSE -.Pred[-.i]]" - - -definition iPred :: "c" -where "iPred \ - [i \ Int \ IF i \ Nat \ {0} - THEN Pred[i] - ELSE -.Succ[-.i]]" - - -(*----------------------------------------------------------------------------*) -(* Recursive definitions on the integers. *) - -(* TODO: replace axiom with proof of existence. *) -axiomatization where - primrec_int: " - \ f: - isAFcn(f) \ - DOMAIN f = Int \ - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n])) - " - -(*----------------------------------------------------------------------------*) -(* Arithmetic operators. *) - - -(* Incrementing and decrementing function. *) -definition addint :: "[c] \ c" -where "addint(m) \ - CHOOSE g \ [Int -> Int]: - g[0] = m \ - (\ n \ Nat: - g[Succ[n]] = iSucc[g[n]] \ - g[-.Succ[n]] = iPred[g[-.n]])" - - -(* Addition on integers. *) -definition add :: "[c, c] \ c" (infixl "+" 65) -where "add(m, n) \ addint(m)[n]" - - -(* Subtraction on integers. *) -definition diff :: "[c, c] \ c" (infixl "-" 65) -where "diff(m, n) \ add(m, -.n)" - - -(* Adding and subtracting function. *) -definition multint :: "[c] \ c" -where "multint(m) \ - CHOOSE g \ [Int -> Int]: - g[0] = 0 \ - (\ n \ Nat: - g[Succ[n]] = add(g[n], m) \ - g[-.Succ[n]] = add(g[-.n], -.m))" - - -(* Multiplication on integers. *) -definition mult :: "[c,c] \ c" (infixl "*" 70) -where "mult(m, n) \ multint(m)[n]" - - -(* Divisibility on the integers. *) -definition dvd :: "[c, c] \ c" (infixl "dvd" 50) -where "b dvd a \ - \ k \ Int: a = mult(b, k)" - - -(* Order on the integers. -This definition is motivated by proofs, where equality is easier to -reason about (adding on both sides, multiplying both sides). -*) -definition leq :: "[c,c] \ c" (infixl "<=" 50) -where "leq(m, n) \ - \ k \ Nat: - add(m, k) = n" - - -abbreviation (input) - geq :: "[c, c] \ c" (infixl ">=" 50) - where "geq(x, y) \ leq(y, x)" - - -definition less :: "[c, c] \ c" (infixl "<" 50) -where "less(a, b) \ leq(a, b) - \ (a \ b)" - - -abbreviation (input) - greater :: "[c, c] \ c" (infixl ">" 50) - where "greater(x, y) \ less(y, x)" - - -definition interval :: "[c, c] \ c" ("(_ .. _)" [90,90] 70) -where "m .. n \ - {k \\in Int: m <= k \ k <= n}" - -(*----------------------------------------------------------------------------*) - - -theorem minus_involutive[simp]: - assumes "n \ Int" - shows "-.-.n = n" - proof - - have s1_1: "n \ Nat \ (\ k \ Nat: n = -.k)" - using assms by (auto simp: Int_def) - have s1_2: "n \ Nat ==> -.-.n = n" - using neg_neg by auto - have s1_3: "n \ Nat ==> -.-.n = n" - proof - - { - assume s2_1: "n \ Nat" - have s2_2: "\ k \ Nat: n = -.k" - using s1_1 s2_1 by auto - have s2_3: "\ k \ Nat: -.n = -.-.k" - using s2_2 by auto - have s2_4: "\ k \ Nat: k = -.n" - using s2_3 neg_neg by auto - have s2_5: "-.n \ Nat" - using s2_4 by auto - have s2_6: "-.-.n = n" - using s2_5 neg_neg by auto - } - from this show "n \ Nat ==> -.-.n = n" by auto - qed - show "-.-.n = n" using s1_2 s1_3 by iprover - qed - - -theorem minus_injective[dest]: - assumes "-.n = -.m" and "n \ Int \ m \ Int" - shows "n = m" - proof - - have "<1>1": "-.-.n = -.-.m" - using assms by simp - have "<1>2": "-.-.n = n" - proof - - have "<2>1": "n \ Int ==> -.-.n = n" - using minus_involutive by auto - have "<2>2": "n \ Int" - using assms by auto - show "-.-.n = n" - using "<2>1" "<2>2" by auto - qed - have "<1>3": "-.-.m = m" - proof - - have "<2>1": "m \ Int ==> -.-.m = m" - using minus_involutive by auto - have "<2>2": "m \ Int" - using assms by auto - show "-.-.m = m" - using "<2>1" "<2>2" by auto - qed - show "n = m" - using "<1>1" "<1>2" "<1>3" by auto - qed - - -theorem minus_eq: - assumes "x = y" - shows "-.x = -.y" - using assms by auto - - -theorem neg_succ_not_in_nat[simp]: - "n \ Nat ==> - -.(Succ[n]) \ Nat" - proof - - { - fix "n" - assume "<2>1": "n \ Nat" - have "<2>2": "Succ[n] \ (Nat \ {0})" - using "<2>1" by auto - have "-.(Succ[n]) \ Nat" - using "<2>2" neg_not_in_nat by auto - } - from this show "n \ Nat ==> - -.(Succ[n]) \ Nat" - by auto - qed - - -theorem neg0_iff_eq0[simp]: - "(-.n = 0) = (n = 0)" - proof auto - assume "<1>1": "-.n = 0" - have "<1>2": "-.n \ Nat" - using "<1>1" by auto - have "<1>3": "-.-.n = n" - using "<1>2" neg_neg by auto - have "<1>4": "-.-.n = 0" - using "<1>1" by auto - show "n = 0" - using "<1>3" "<1>4" by simp - qed - - -theorem neg0_imp_eq0[dest]: - "(-.n = 0) ==> (n = 0)" - by simp - - -theorem not_neg0_imp_not0[dest]: - "(-.n \ 0) ==> (n \ 0)" - by simp - - -(*----------------------------------------------------------------------------*) -(* Properties of the set `Int`. *) - - -theorem nat_in_int[simp]: - "n \ Nat ==> n \ Int" - by (simp add: Int_def) - - -theorem minus_nat_in_int: - "n \ Nat ==> -.n \ Int" - by (auto simp: Int_def) - - -theorem minus_nat_in_neg_nat: - "n \ Nat ==> -.n \ negNat" - by (auto simp: negNat_def) - - -theorem neg_nat_in_int: - "n \ negNat ==> n \ Int" - by (auto simp: negNat_def Int_def) - - -theorem int_disj: - "n \ Int ==> - n \ Nat - \ n \ {-.n: n \ Nat}" - by (auto simp: Int_def) - - -theorem int_disj_nat_negnat: - "n \ Int ==> - n \ Nat - \ n \ negNat" - unfolding negNat_def - using int_disj by auto - - -theorem int_union_nat_negnat: - "Int = Nat \ negNat" - using int_disj_nat_negnat nat_in_int - neg_nat_in_int - by auto - - -theorem int_union_nat_0_negnat: - "Int = (Nat \ {0}) \ negNat" - proof - - have s1_1: "Int = Nat \ negNat" - using int_union_nat_negnat by auto - have s1_2: "0 \ negNat" - unfolding negNat_def - using zeroIsNat neg0 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem int_union_nat_negnat_0: - "Int = Nat \ (negNat \ {0})" - proof - - have s1_1: "Int = Nat \ negNat" - using int_union_nat_negnat by auto - have s1_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem neg_int_eq_int[simp]: - "n \ Int ==> -.n \ Int" - unfolding Int_def by auto - - -theorem minus_negnat_in_int: - "n \ negNat ==> -.n \ Int" - using neg_nat_in_int neg_int_eq_int by auto - - -theorem minus_neg_int_in_nat: - "n \ Int \ n \ Nat ==> - -.n \ Nat \ {0}" - proof - - assume "<1>1": "n \ Int \ n \ Nat" - have "<1>2": "n \ {-.k: k \ Nat}" - using "<1>1" int_disj by iprover - have "<1>3": "\ k \ Nat: n = -.k" - using "<1>2" by auto - have "<1>4": "\ k \ Nat: k = -.n" - using "<1>3" minus_involutive by auto - have "<1>5": "-.n \ Nat" - using "<1>4" by auto - have "<1>6": "-.n \ 0" - proof - - { - assume "<2>1": "-.n = 0" - have "<2>2": "n = 0" - using neg0 "<2>1" by auto - have "<2>3": "n \ Nat" - using "<2>2" by auto - have "<2>4": "n \ Nat" - using "<1>1" by auto - have "<2>5": "FALSE" - using "<2>3" "<2>4" by auto - } - from this show "-.n \ 0" - by auto - qed - show "-.n \ Nat \ {0}" - using "<1>5" "<1>6" by auto - qed - - -theorem neg_negative_in_nat: - assumes hyp: "n \ Int \ n \ Nat" - shows "-. n \ Nat" - proof - - have int: "n \ Int" - using hyp by auto - from int int_disj have disj: - "n \ Nat \ n \ {-. m: m \ Nat}" - by auto - from disj hyp have neg: "n \ {-. m: m \ Nat}" - by auto - from neg have ex1: "\ m \ Nat: n = -. m" - by auto - from ex1 have ex2: "\ m \ Nat: -. n = -. -. m" - by auto - from ex2 neg_neg have ex3: "\ m \ Nat: -. n = m" - by auto - from ex3 show "-. n \ Nat" - by auto - qed - - -theorem minus_nat_0_or_not_in_nat: - "n \ Nat ==> - -.n = 0 \ -.n \ Nat" - using neg0 neg_not_in_nat by auto - - -theorem int_eq_neg_int_is_0[simp]: - "n \ Int ==> (n = -.n) = (n = 0)" - proof - - have "<1>1": "n \ Nat \ {0} ==> - -.n \ Nat" - proof - - assume "<2>1": "n \ Nat \ {0}" - show "-.n \ Nat" - using neg_not_in_nat "<2>1" by auto - qed - have "<1>2": "n \ Nat \ {0} ==> - n \ -.n" - using "<1>1" by force - have "<1>3": "n \ Int \ n \ Nat - ==> -.n \ Nat" - using minus_neg_int_in_nat by auto - have "<1>4": "n \ Int \ n \ Nat - ==> n \ -.n" - using "<1>3" by force - have "<1>5": "n \ Int \ {0} ==> - (n \ Nat \ {0}) \ - (n \ {-.k: k \ Nat})" - using int_disj by auto - have "<1>6": "n \ Int \ {0} ==> - (n \ -.n)" - using "<1>2" "<1>4" "<1>5" by auto - have "<1>7": "n \ Int \ (n \ 0) ==> - (n \ -.n)" - using "<1>6" by auto - have "<1>8": "n \ Int \ (n = -.n) ==> - (n = 0)" - using "<1>7" by auto - have "<1>9": "n \ Int \ (n = 0) ==> - (n = -.n)" - using neg0 by auto - show "n \ Int ==> (n = -.n) = (n = 0)" - using "<1>8" "<1>9" by auto - qed - - -theorem minus_neg_nat_in_nat: - assumes hyp: "n \ negNat" - shows "-.n \ Nat" - proof - - have s1_1: "\ k \ Nat: -.k = n" - using hyp by (auto simp: negNat_def) - have s1_2: "\ k \ Nat: -.-.k = -.n" - using s1_1 by auto - have s1_3: "\ k \ Nat: k = -.n" - using s1_2 neg_neg by auto - show "-.n \ Nat" - using s1_3 by auto - qed - - -theorem minus_neg_nat_0_in_nat_0: - assumes hyp: "n \ negNat \ {0}" - shows "-.n \ Nat \ {0}" - proof - - have s1_1: "-.n \ Nat" - using hyp minus_neg_nat_in_nat by auto - have s1_2: "-.n \ 0" - proof - - { - assume s2_1: "-.n = 0" - have s2_2: "-.-.n = -.0" - using s2_1 by auto - have s2_3: "-.-.n = 0" - using s2_2 neg0 by auto - have s2_4: "n = 0" - proof - - have s3_1: "n \ Int" - using hyp neg_nat_in_int by auto - have s3_2: "-.-.n = n" - using s3_1 minus_involutive - by auto - show ?thesis - using s3_2 s2_3 by auto - qed - have s2_5: "n \ 0" - using hyp by auto - have "FALSE" - using s2_4 s2_5 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem neg_nat_0_not_in_nat: - assumes hyp: "n \ negNat \ {0}" - shows "n \ Nat" - proof auto - assume s1_1: "n \ Nat" - have s1_2: "n \ 0" - using hyp by auto - have s1_3: "n \ Nat \ {0}" - using s1_1 s1_2 by auto - have s1_4: "-.n \ Nat" - using s1_3 neg_not_in_nat by auto - have s1_5: "-.n \ Nat" - using hyp minus_neg_nat_in_nat by auto - show "FALSE" - using s1_4 s1_5 by auto - qed - - -theorem neg_nat_not_in_nat_setminus_0: - assumes hyp: "n \ negNat" - shows "n \ Nat \ {0}" - proof - - have s1_1: "n \ negNat ==> - n \ negNat \ {0} - \ n = 0" - by auto - have s1_2: "n \ negNat \ {0} - ==> n \ Nat \ {0}" - using neg_nat_0_not_in_nat by auto - have s1_3: "n = 0 ==> - n \ Nat \ {0}" - by auto - show "n \ Nat \ {0}" - using hyp s1_1 s1_2 s1_3 by iprover - qed - - -theorem minus_nat_0_in_negnat_0: - assumes hyp: "n \ Nat \ {0}" - shows "-.n \ negNat \ {0}" - proof - - have s1_1: "n \ Nat" - using hyp by auto - have s1_2: "-.n \ negNat" - using s1_1 minus_nat_in_neg_nat by auto - have s1_3: "-.n \ 0" - proof - - { - assume s2_1: "-.n = 0" - have s2_2: "-.-.n = -.0" - using s2_1 by auto - have s2_3: "-.-.n = 0" - using s2_2 by auto - have s2_4: "n = 0" - using s1_1 neg_neg s2_3 by auto - have s2_5: "n \ 0" - using hyp by auto - have "FALSE" - using s2_4 s2_5 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_2 s1_3 by auto - qed - - -theorem minus_one_in_negnat: - "-.1 \\in negNat" - unfolding negNat_def - using oneIsNat by auto - - -theorem minus_one_in_int: - "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - - -theorem minus_one_neq_0: - "-.1 \ 0" - proof - - { - assume s1_1: "-.1 = 0" - have s1_2: "-.-.1 = -.0" - using s1_1 by auto - have s1_3: "-.-.1 = 0" - using s1_2 neg0 by auto - have s1_4: "0 \ Nat" - by auto - have s1_5: "1 = 0" - using s1_3 s1_4 by auto - have "FALSE" - using s1_5 succIrrefl - by auto - } - from this show "-.1 \ 0" - by auto - qed - - -theorem minus_minus_one: - "-.-.1 = 1" - proof - - have s1_1: "1 \ Nat" - using oneIsNat by auto - show "-.-.1 = 1" - using neg_neg s1_1 by auto - qed - - -theorem succ_zero_equality: - assumes hyp: "x = Succ[0]" - shows "x = 1" - using hyp by auto - - -theorem pred_1: - "Pred[1] = 0" - proof - - have s1_1: "1 = Succ[0]" - by auto - have s1_2: "\ x \ Nat: (1 = Succ[x]) => (x = 0)" - using s1_1 succInjIff by auto - show ?thesis - using s1_1 s1_2 choose_equality by auto - qed - - -theorem pred_2: - "Pred[2] = 1" - proof - - have s1_1: "1 \\in Nat \ Succ[1] = 2" - unfolding two_def - using oneIsNat by auto - have s1_2: "\ x. x \\in Nat \ (Succ[x] = 2) ==> (x = 1)" - using oneIsNat succInjIff by (auto simp: two_def) - have s1_4: "Pred[2] = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" - proof - - have s2_1: "Pred[2] = - (IF 2 = 0 THEN 0 ELSE CHOOSE x \ Nat: 2 = Succ[x])" - unfolding Pred_def - using twoIsNat by auto - also have s2_2: "... = (CHOOSE x \ Nat: 2 = Succ[x])" - proof - - have s3_1: "2 \ 0" - using succNotZero by (auto simp: two_def) - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = (CHOOSE x: x \\in Nat \ 2 = Succ[x])" - unfolding bChoose_def - by auto - also have s2_4: "... = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" - proof - - have s3_1: "\ x. (2 = Succ[x]) = (Succ[x] = 2)" - by auto - show ?thesis - using s3_1 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s1_1 s1_2 s1_4 choose_equality[ - of "\ x. x \\in Nat \ Succ[x] = 2" - "1" "1"] - by auto - qed - - -theorem zero_in_int: - "0 \ Int" - using zeroIsNat nat_in_int by auto - - -theorem ipred_0: - shows "iPred[0] = -.1" - proof - - have s1_1: "iPred[0] = -.Succ[-.0]" - proof - - have s2_1: "0 \ Int" - using zero_in_int by auto - have s2_2: "0 \ Nat \ {0}" - by auto - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_2: "... = -.Succ[0]" - using neg0 by auto - also have s1_3: "... = -.1" - by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed - - -theorem isucc_minus_1: - "iSucc[-.1] = 0" - unfolding iSucc_def - using oneIsNat neg_neg pred_1 neg0 by auto - - -theorem negnat_succ_in_nat: - assumes ndom: "n \\in negNat" - shows "Succ[-.n] \\in Nat" - using ndom minus_neg_nat_in_nat succIsNat by auto - - -theorem minus_negnat_succ_in_negnat: - assumes ndom: "n \\in negNat" - shows "-.Succ[-.n] \\in negNat" - using ndom negnat_succ_in_nat minus_nat_in_neg_nat by auto - - -theorem minus_succ_minus_negnat_in_int: - assumes ndom: "n \\in negNat" - shows "-.Succ[-.n] \\in Int" - proof - - have s1_1: "-.n \\in Nat" - using ndom minus_neg_nat_in_nat by auto - have s1_2: "Succ[-.n] \\in Nat" - using s1_1 succIsNat by auto - show ?thesis - using s1_2 minus_nat_in_int by auto - qed - -(*----------------------------------------------------------------------------*) -(* `iSucc` and `iPred` properties. *) - -(* Typeness *) -theorem iSucc_type: - assumes idom: "i \ Int" - shows "iSucc[i] \ Int" - proof - - have s1_1: "i \ Nat ==> - iSucc[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s2_1 by auto - have s2_3: "Succ[i] \ Nat" - using s2_1 succIsNat by auto - have s2_4: "Succ[i] \ Int" - using s2_3 nat_in_int by auto - show "iSucc[i] \ Int" - using s2_2 s2_4 by auto - qed - have s1_2: "i \ Nat ==> - iSucc[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using s2_1 idom by auto - have s2_3: "Pred[-.i] \ Nat" - proof - - have s3_1: "-.i \ Nat" - using idom s2_1 minus_neg_int_in_nat - by auto - show "Pred[-.i] \ Nat" - using s3_1 Pred_in_nat by auto - qed - have s2_4: "-.Pred[-.i] \ Int" - using s2_3 minus_nat_in_int by auto - show "iSucc[i] \ Int" - using s2_2 s2_4 by auto - qed - show "iSucc[i] \ Int" - using s1_1 s1_2 by auto - qed - - -theorem iPred_type: - assumes idom: "i \ Int" - shows "iPred[i] \ Int" - proof - - have s1_1: "i \ Nat \ {0} - ==> iPred[i] \ Int" - proof - - assume s2_1: "i \ Nat \ {0}" - have s2_2: "iPred[i] = Pred[i]" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "Pred[i] \ Nat" - using s2_1 Pred_in_nat by auto - have s2_4: "Pred[i] \ Int" - using s2_3 nat_in_int by auto - show "iPred[i] \ Int" - using s2_2 s2_4 by auto - qed - have s1_2: "i = 0 ==> - iPred[i] \ Int" - proof - - assume s2_1: "i = 0" - have s2_2: "iPred[i] = -.1" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "-.1 \ Int" - using oneIsNat by (auto simp: Int_def) - show "iPred[i] \ Int" - using s2_2 s2_3 by auto - qed - have s1_3: "i \ Nat ==> - iPred[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iPred[i] = -.Succ[-.i]" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "-.Succ[-.i] \ Int" - proof - - have s3_1: "-.i \ Nat" - using idom s2_1 minus_neg_int_in_nat - by auto - have s3_2: "Succ[-.i] \ Nat" - using s3_1 succIsNat by auto - show "-.Succ[-.i] \ Int" - using s3_2 minus_nat_in_int by auto - qed - show "iPred[i] \ Int" - using s2_2 s2_3 by auto - qed - show "iPred[i] \ Int" - using s1_1 s1_2 s1_3 by auto - qed - - -(* -THEOREM iSucciPredCommute == - \A i \in Int: iSucc[iPred[i]] = iPred[iSucc[i]] - PROOF - <1>1. SUFFICES ASSUME NEW i \in Int - PROVE iSucc[iPred[i]] = iPred[iSucc[i]] - BY <1>1 - <1>2. CASE i \in Nat \ {0} - <2>1. iSucc[iPred[i]] = i - <3>1. iPred[i] = Pred[i] - <4>1. i \in Nat \ {0} - BY <1>2 - <4> QED - BY <4>1 DEF iPred - <3>2. iSucc[iPred[i]] = Succ[Pred[i]] - <4>1. Pred[i] \in Nat - <5>1. i \in Nat - BY <1>2 - <5> QED - BY <5>1, Pred_in_nat - <4> QED - BY <3>1, <4>1 DEF iSucc - <3>3. Succ[Pred[i]] = i - <4>1. i \in Nat \ {0} - BY <1>2 - <4> QED - BY <4>1, Succ_Pred_nat - <3> QED - BY <3>2, <3>3 - <2>2. iPred[iSucc[i]] = i - <3>1. iSucc[i] = Succ[i] - <4>1. i \in Nat - BY <1>2 - <4> QED - BY <4>1 DEF iSucc - <3>2. iPred[iSucc[i]] = Pred[Succ[i]] - <4>1. Succ[i] \in Nat - <5>1. i \in Nat - BY <1>2 - <5> QED - BY <5>1, succIsNat - <4>2. Succ[i] # 0 - BY <1>2, succNotZero - <4>3. Succ[i] \in Nat \ {0} - BY <4>1, <4>2 - <4>4. iPred[Succ[i]] = Pred[Succ[i]] - BY <4>3 DEF iPred - <4> QED - BY <4>4, <3>1 - <3>3. Pred[Succ[i]] = i - <4>1. i \in Nat - BY <1>2 - <4> QED - BY <4>1, Pred_Succ_nat - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2 - <1>3. CASE i = 0 - <2>1. iSucc[iPred[i]] = 0 - <3>1. iSucc[iPred[i]] = iSucc[-.1] - <4>1. iPred[i] = -.1 - BY <1>3 DEF iPred - <4> QED - BY <4>1 - <3>2. iSucc[-.1] = -.Pred[-.-.1] - <4>1. -.1 \in Int - <5>1. -.1 \in negNat - BY oneIsNat DEF negNat - <5> QED - BY <5>1, neg_nat_in_int - <4>2. -.1 \notin Nat - <5>1. -.1 \in negNat \ {0} - <6>1. -.1 # 0 - BY minus_one_neq_0 - <6>2. -.1 \in negNat - BY oneIsNat DEF negNat - <6> QED - BY <6>1, <6>2 - <5> QED - BY <5>1, neg_nat_0_not_in_nat - <4> QED - BY <4>1, <4>2 DEF iSucc - <3>3. -.Pred[-.-.1] = 0 - <4>1. -.-.1 = 1 - BY minus_minus_one - <4>2. Pred[1] = 0 - BY pred_1 - <4>3. -.0 = 0 - BY neg0 - <4> QED - BY <4>1, <4>2, <4>3 - <3> QED - BY <3>1, <3>2, <3>3 - <2>2. iPred[iSucc[i]] = 0 - <3>1. iSucc[i] = 1 - <4>1. i \in Nat - BY <1>3, zeroIsNat - <4>2. iSucc[i] = Succ[i] - BY <4>1 DEF iSucc - <4>3. Succ[i] = Succ[0] - BY <1>3 - <4>4. Succ[0] = 1 - OBVIOUS - <4> QED - BY <4>2, <4>3, <4>4 - <3>2. iPred[1] = 0 - <4>1. 1 \in Nat \ {0} - BY oneIsNat, succIrrefl - <4>2. iPred[1] = Pred[1] - BY <4>1 DEF iPred - <4>3. Pred[1] = 0 - BY pred_1 - <4> QED - BY <4>2, <4>3 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 - <1>4. CASE i \in negNat - <2>1. iSucc[iPred[i]] = i - <3>1. iPred[i] = -.Succ[-.i] - <4>1. i \in Int - BY <1>1 - <4>2. i \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3>2. -.Succ[-.i] \in negNat \ {0} - <4>1. Succ[-.i] \in Nat \ {0} - <5>1. -.i \ Nat - BY <1>4, minus_neg_nat_in_nat - <5>2. Succ[-.i] \in Nat - BY <5>1, succIsNat - <5>3. Succ[-.i] \ 0 - BY <5>1, succNotZero - <5> QED - BY <5>2, <5>3 - <4> QED - BY <4>1, minus_nat_0_in_negnat_0 - <3>3. iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]] - <4>1. iPred[i] \in negNat \ {0} - BY <3>1, <3>2 - <4>2. iPred[i] \in Int - BY <4>1, neg_nat_in_int - <4>3. iPred[i] \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4>4. iSucc[iPred[i]] = -.Pred[-.iPred[i]] - BY <4>2, <4>3 DEF iSucc - <4> QED - BY <4>4, <3>1 - <3>4. @ = -.Pred[Succ[-.i]] - <4>1. -.i \in Nat - BY <1>4, minus_neg_nat_in_nat - <4>2. Succ[-.i] \in Int - BY <4>1, succIsNat, nat_in_int - <4>3. -.-.Succ[-.i] = Succ[-.i] - BY minus_involutive - <4> QED - BY <4>3 - <3>5. @ = -.-.i - <4>1. -.i \in Nat - BY <1>4, minus_neg_nat_in_nat - <4>2. Pred[Succ[-.i]] = -.i - BY <4>1, Pred_Succ_nat - <4> QED - BY <4>2 - <3>6. @ = i - BY <1>4, minus_involutive - <3> QED - BY <3>3, <3>4, <3>5, <3>6 - <2>2. iPred[iSucc[i]] = i - <3>1. CASE i = 0 - <4>1. iSucc[i] = 1 - <5>1. 0 \in Nat - BY zeroIsNat - <5>2. 0 \in Int - BY <5>1, nat_in_int - <5>3. iSucc[i] = Succ[0] - BY <5>1, <5>2, <3>1 DEF iSucc - <5>4. Succ[0] = 1 - OBVIOUS - <5> QED - BY <5>3, <5>4 - <4>2. iPred[1] = 0 - <5>1. 1 \in Nat \ {0} - BY oneIsNat, succNotZero - <5>2. 1 \in Int - BY <5>1, nat_in_int - <5>3. iPred[1] = Pred[1] - BY <5>1, <5>2 DEF iPred - <5>4. Pred[1] = 0 - BY pred_1 - <5> QED - BY <5>3, <5>4 - <4> QED - BY <4>1, <4>2, <3>1 - <3>2. CASE i # 0 - <4>1. i \in negNat \ {0} - BY <1>4, <3>1 - <4>2. i \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4>3. i \in Int - BY <4>1, neg_nat_in_int - <4>4. iSucc[i] = -.Pred[-.i] - BY <4>2, <4>3 DEF iSucc - <4>5. -.Pred[-.i] \in negNat - <5>1. -.i \in Nat - BY <4>1, minus_neg_nat_in_nat - <5>2. Pred[-.i] \in Nat - BY <5>1, Pred_in_nat - <5> QED - BY <5>2, minus_nat_in_neg_nat - <4>6. -.Pred[-.i] \in Int - BY <4>5, neg_nat_in_int - <4>7. -.Pred[-.i] \notin Nat \ {0} - BY <4>5, neg_nat_not_in_nat_setminus_0 - <4>8. iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]] - BY <4>6, <4>7 DEF iPred - <4>9. -.Succ[-.-.Pred[-.i]] = i - <5>1. -.-.Pred[-.i] = Pred[-.i] - <6>1. -.i \in Nat - BY <4>1, minus_neg_nat_in_nat - <6>2. Pred[-.i] \in Nat - BY <6>1, Pred_in_nat - <6>3. Pred[-.i] \in Int - BY <6>2, nat_in_int - <6> QED - BY <6>3, minus_involutive - <5>2. Succ[Pred[-.i]] = -.i - <6>1. -.i \in Nat \ {0} - BY <4>1, minus_neg_nat_0_in_nat_0 - <6> QED - BY <6>1, Succ_Pred_nat - <5>3. -.-.i = i - BY <4>3, minus_involutive - <5> QED - BY <5>1, <5>2, <5>3 - <4> QED - BY <4>4, <4>8, <4>9 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 - <1> QED - BY <1>1, <1>2, <1>3, <1>4 -*) -theorem iSucciPredCommute: - "\ i \ Int: - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - { - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "i \ Nat \ {0} ==> - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - assume idom: "i \ Nat \ {0}" - have s2_1: "iSucc[iPred[i]] = i" - proof - - have s3_1: "iPred[i] = Pred[i]" - unfolding iPred_def - using s1_1 idom by auto - have s3_2: "iSucc[iPred[i]] = Succ[Pred[i]]" - proof - - have s4_1: "Pred[i] \ Nat" - proof - - have s5_1: "i \ Nat" - using idom by auto - show ?thesis - using s5_1 Pred_in_nat by auto - qed - show ?thesis - unfolding iSucc_def - using s4_1 s3_1 by auto - qed - have s3_3: "Succ[Pred[i]] = i" - using idom Succ_Pred_nat - by auto - show ?thesis - using s3_2 s3_3 by auto - qed - have s2_2: "iPred[iSucc[i]] = i" - proof - - have s3_1: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s1_1 idom by auto - have s3_2: "iPred[iSucc[i]] = Pred[Succ[i]]" - proof - - have s4_1: "Succ[i] \ Nat" - proof - - have s5_1: "i \ Nat" - using idom by auto - show ?thesis - using s5_1 succIsNat by auto - qed - have s4_2: "Succ[i] \ 0" - using idom succNotZero by auto - have s4_3: "Succ[i] \ Nat \ {0}" - using s4_1 s4_2 by auto - have s4_4: "Succ[i] \ Int" - using s4_1 nat_in_int by auto - have s4_5: "iPred[Succ[i]] = Pred[Succ[i]]" - unfolding iPred_def - using s4_3 s4_4 by auto - show ?thesis - using s4_5 s3_1 by auto - qed - have s3_3: "Pred[Succ[i]] = i" - proof - - have s4_1: "i \ Nat" - using idom by auto - show ?thesis - using s4_1 Pred_Succ_nat - by auto - qed - show ?thesis - using s3_2 s3_3 by auto - qed - show ?thesis - using s2_1 s2_2 - by auto - qed - have s1_3: "i \ negNat ==> - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - assume idom: "i \ negNat" - have s2_1: "iSucc[iPred[i]] = i" - proof - - have s3_1: "iPred[i] = -.Succ[-.i]" - proof - - have s4_1: "i \ Int" - using s1_1 by auto - have s4_2: "i \ Nat \ {0}" - using idom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - have s3_2: "-.Succ[-.i] \ negNat \ {0}" - proof - - have s4_1: "Succ[-.i] \ Nat \ {0}" - proof - - have s5_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat - by auto - have s5_2: "Succ[-.i] \ Nat" - using s5_1 succIsNat by auto - have s5_3: "Succ[-.i] \ 0" - using s5_1 succNotZero by auto - show ?thesis - using s5_2 s5_3 by auto - qed - show ?thesis - using s4_1 minus_nat_0_in_negnat_0 - by auto - qed - have s3_3: "iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]]" - proof - - have s4_1: "iPred[i] \ negNat \ {0}" - using s3_1 s3_2 by auto - have s4_2: "iPred[i] \ Int" - using s4_1 neg_nat_in_int by auto - have s4_3: "iPred[i] \ Nat" - using s4_1 neg_nat_0_not_in_nat by auto - have s4_4: "iSucc[iPred[i]] = -.Pred[-.iPred[i]]" - unfolding iSucc_def - using s4_2 s4_3 by auto - show ?thesis - using s4_4 s3_1 by auto - qed - have s3_4: "-.Pred[-.-.Succ[-.i]] = -.Pred[Succ[-.i]]" - proof - - have s4_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat by auto - have s4_2: "Succ[-.i] \ Int" - using s4_1 succIsNat nat_in_int by auto - have s4_3: "-.-.Succ[-.i] = Succ[-.i]" - using s4_2 minus_involutive by auto - show ?thesis - using s4_3 by auto - qed - have s3_5: "-.Pred[Succ[-.i]] = -.-.i" - proof - - have s4_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat by auto - have s4_2: "Pred[Succ[-.i]] = -.i" - using s4_1 Pred_Succ_nat by auto - show ?thesis - using s4_2 by auto - qed - have s3_6: "-.-.i = i" - using s1_1 minus_involutive by auto - show ?thesis - using s3_3 s3_4 s3_5 s3_6 by auto - qed - have s2_2: "iPred[iSucc[i]] = i" - proof - - have s3_1: "i = 0 ==> - iPred[iSucc[i]] = i" - proof - - assume i0: "i = 0" - have s4_1: "iSucc[i] = 1" - proof - - have s5_1: "0 \ Nat" - using zeroIsNat by auto - have s5_2: "0 \ Int" - using s5_1 nat_in_int by auto - have s5_3: "iSucc[i] = Succ[0]" - unfolding iSucc_def - using s5_1 s5_2 i0 - by auto - have s5_4: "Succ[0] = 1" - by auto - show ?thesis - using s5_3 s5_4 by auto - qed - have s4_2: "iPred[1] = 0" - proof - - have s5_1: "1 \ Nat \ {0}" - using oneIsNat succNotZero - by auto - have s5_2: "1 \ Int" - using s5_1 nat_in_int by auto - have s5_3: "iPred[1] = Pred[1]" - unfolding iPred_def - using s5_1 s5_2 - by auto - have s5_4: "Pred[1] = 0" - using pred_1 by auto - show ?thesis - using s5_3 s5_4 by auto - qed - show "iPred[iSucc[i]] = i" - using s4_1 s4_2 i0 by auto - qed - have s3_2: "i \ 0 ==> - iPred[iSucc[i]] = i" - proof - - assume inot0: "i \ 0" - have s4_1: "i \ negNat \ {0}" - using idom inot0 by auto - have s4_2: "i \ Nat" - using s4_1 neg_nat_0_not_in_nat - by auto - have s4_3: "i \ Int" - using s4_1 neg_nat_in_int - by auto - have s4_4: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using s4_2 s4_3 by auto - have s4_5: "-.Pred[-.i] \ negNat" - proof - - have s5_1: "-.i \ Nat" - using s4_1 minus_neg_nat_in_nat - by auto - have s5_2: "Pred[-.i] \ Nat" - using s5_1 Pred_in_nat - by auto - show ?thesis - using s5_2 minus_nat_in_neg_nat - by auto - qed - have s4_6: "-.Pred[-.i] \ Int" - using s4_5 neg_nat_in_int - by auto - have s4_7: "-.Pred[-.i] \ Nat \ {0}" - using s4_5 neg_nat_not_in_nat_setminus_0 - by blast - have s4_8: "iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]]" - unfolding iPred_def - using s4_6 s4_7 by auto - have s4_9: "-.Succ[-.-.Pred[-.i]] = i" - proof - - have s5_1: "-.-.Pred[-.i] = Pred[-.i]" - proof - - have s6_1: "-.i \ Nat" - using s4_1 minus_neg_nat_in_nat - by auto - have s6_2: "Pred[-.i] \ Nat" - using s6_1 Pred_in_nat - by auto - have s6_3: "Pred[-.i] \ Int" - using s6_2 nat_in_int - by auto - show ?thesis - using s6_3 minus_involutive - by auto - qed - have s5_2: "Succ[Pred[-.i]] = -.i" - proof - - have s6_1: "-.i \ Nat \ {0}" - using s4_1 minus_neg_nat_0_in_nat_0 - by auto - show ?thesis - using s6_1 Succ_Pred_nat - by auto - qed - have s5_3: "-.-.i = i" - using s4_3 minus_involutive by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - show "iPred[iSucc[i]] = i" - using s4_4 s4_8 s4_9 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - show "iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i" - using s2_1 s2_2 by auto - qed - have s1_4: "Int = (Nat \ {0}) \ negNat" - using int_union_nat_0_negnat by auto - have "iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i" - using s1_1 s1_2 s1_3 s1_4 by auto - } - from this show ?thesis - using allI by auto - qed - - -(* -THEOREM iSuccMinusFlip == - \A i \in Int: -.iSucc[-.i] = iPred[i] -PROOF -<1>1. SUFFICES - ASSUME NEW i \in Int - PROVE -.iSucc[-.i] = iPred[i] - BY <1>1 -<1>2. CASE i \in Nat \ {0} - <2>1. i \in Int - BY <1>2, nat_in_int - <2>2. -.iSucc[-.i] = Pred[i] - <3>1. -.iSucc[-.i] = -.-.Pred[-.-.i] - <4>1. -.i \in negNat \ {0} - BY <1>2, minus_nat_0_in_negnat_0 - <4>2. -.i \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4> QED - BY <2>1, <4>2 DEF iSucc - <3>2. @ = -.-.Pred[i] - <4>1. -.-.i = i - BY <2>1, minus_involutive - <4> QED - BY <4>1 - <3>3. @ = Pred[i] - <4>1. Pred[i] \in Nat - BY <1>2, Pred_in_nat - <4>2. Pred[i] \in Int - BY <4>1, nat_in_int - <4> QED - BY <4>2, minus_involutive - <3> QED - BY <3>1, <3>2, <3>3 - <2>3. iPred[i] = Pred[i] - BY <2>1, <1>2 DEF iPred - <2> QED - BY <2>2, <2>3 -<1>3. CASE i \in negNat - <2>1. i \in Int - BY <1>3, neg_nat_in_int - <2>2. i \notin Nat \ {0} - BY <1>3, neg_nat_not_in_nat_setminus_0 - <2>3. Pred[i] = -.Succ[-.i] - BY <2>1, <2>2 DEF iPred - <2>4. -.iSucc[-.i] = -.Succ[-.i] - <3>1. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3>2. -.i \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>1, <3>2 DEF iSucc - <2> QED - BY <2>3, <2>4 -<1> QED - BY <1>1, <1>2, <1>3, int_union_nat_0_negnat -*) -theorem iSuccMinusFlip: - "\ i \ Int: - -.iSucc[-.i] = iPred[i]" - proof auto - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "i \ Nat \ {0} ==> - -.iSucc[-.i] = iPred[i]" - proof - - assume s1_2_asm: "i \ Nat \ {0}" - have s2_1: "i \ Int" - using s1_2_asm nat_in_int - by auto - have s2_2: "-.iSucc[-.i] = Pred[i]" - proof - - have s3_1: "-.iSucc[-.i] = -.-.Pred[-.-.i]" - proof - - have s4_1: "-.i \ negNat \ {0}" - using s1_2_asm minus_nat_0_in_negnat_0 - by auto - have s4_2: "-.i \ Nat" - using s4_1 neg_nat_0_not_in_nat - by auto - show ?thesis - unfolding iSucc_def - using s2_1 s4_2 by auto - qed - have s3_2: "-.-.Pred[-.-.i] = -.-.Pred[i]" - proof - - have s4_1: "-.-.i = i" - using s2_1 minus_involutive - by auto - show ?thesis - using s4_1 by auto - qed - have s3_3: "-.-.Pred[i] = Pred[i]" - proof - - have s4_1: "Pred[i] \ Nat" - using s1_2_asm Pred_in_nat by auto - have s4_2: "Pred[i] \ Int" - using s4_1 nat_in_int by auto - show ?thesis - using s4_2 minus_involutive by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 by auto - qed - have s2_3: "iPred[i] = Pred[i]" - unfolding iPred_def - using s2_1 s1_2_asm by auto - show "-.iSucc[-.i] = iPred[i]" - using s2_2 s2_3 by auto - qed - have s1_3: "i \ negNat ==> - -.iSucc[-.i] = iPred[i]" - proof - - assume s1_3_asm: "i \ negNat" - have s2_1: "i \ Int" - using s1_3_asm neg_nat_in_int - by auto - have s2_2: "i \ Nat \ {0}" - using s1_3_asm neg_nat_not_in_nat_setminus_0 - by auto - have s2_3: "-.Succ[-.i] = iPred[i]" - unfolding iPred_def - using s2_1 s2_2 by auto - have s2_4: "-.iSucc[-.i] = -.Succ[-.i]" - proof - - have s3_1: "-.i \ Nat" - using s1_3_asm minus_neg_nat_in_nat by auto - have s3_2: "-.i \ Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - show "-.iSucc[-.i] = iPred[i]" - using s2_3 s2_4 by auto - qed - show "-.iSucc[-.i] = iPred[i]" - using s1_1 s1_2 s1_3 int_union_nat_0_negnat - by auto - qed - - -(* -THEOREM iPredMinusFlip == - \A i \in Int: -.iPred[-.i] = iSucc[i] -PROOF -<1>1. SUFFICES - ASSUME NEW i \in Int - PROVE -.iPred[-.i] = iSucc[i] - BY <1>1 -<1>2. -.iSucc[-.-.i] = iPred[-.i] - <2>1. -.i \in Int - BY <1>1, neg_int_eq_int - <2> QED - BY <2>1, iSuccMinusFlip -<1>3. -.iSucc[i] = iPred[-.i] - <2>1. -.-.i = i - BY <1>1, minus_involutive - <2> QED - BY <1>2, <2>1 -<1>4. -.-.iSucc[i] = -.iPred[-.i] - BY <1>3 -<1>5. -.-.iSucc[i] = iSucc[i] - <2>1. iSucc[i] \in Int - BY <1>1, iSucc_type - <2> QED - BY <2>1, minus_involutive -<1> QED - BY <1>4, <1>5 -*) -theorem iPredMinusFlip: - "\ i \ Int: -.iPred[-.i] = iSucc[i]" - proof auto - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "-.iSucc[-.-.i] = iPred[-.i]" - proof - - have s2_1: "-.i \ Int" - using s1_1 neg_int_eq_int by auto - show ?thesis - using s2_1 iSuccMinusFlip by auto - qed - have s1_3: "-.iSucc[i] = iPred[-.i]" - proof - - have s2_1: "-.-.i = i" - using s1_1 minus_involutive by auto - show ?thesis - using s1_2 s2_1 by auto - qed - have s1_4: "-.-.iSucc[i] = -.iPred[-.i]" - using s1_3 by auto - have s1_5: "-.-.iSucc[i] = iSucc[i]" - proof - - have s2_1: "iSucc[i] \ Int" - using s1_1 iSucc_type by auto - show ?thesis - using s2_1 minus_involutive by auto - qed - show "-.iPred[-.i] = iSucc[i]" - using s1_4 s1_5 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* `neg_nat_induction` *) - - -(* -negNat == {-.n: n \in Nat} - - -THEOREM negNatInduction == - ASSUME - NEW P(n), P(0), - \A i \in negNat: P(i) => P(-.Succ[-.i]) - PROVE - \A i \in negNat: P(i) - PROOF - <1> DEFINE Q(n) == P(-.n) - <1>1. \A n \in Nat: Q(n) - <2>1. Q(0) - BY P(0), -.0 = 0 DEF Q - <2>2. ASSUME NEW n \in Nat, Q(n) - PROVE Q(Succ[n]) - <3>1. P(-.n) - BY <2>2 DEF Q - <3>2. -.n \in negNat - BY <2>2, n \in Nat DEF negNat - <3>3. \A i \in negNat: P(i) => P(-.Succ[-.i]) - OBVIOUS - <3>4. P(-.n) => P(-.Succ[-.-.n]) - BY <3>2, <3>3 - <3>5. P(-.Succ[-.-.n]) - BY <3>1, <3>4 - <3>6. P(-.Succ[n]) - <4>1. -.-.n = n - BY <2>2, minus_involutive, nat_in_int - <4> QED - BY <3>5, <4>1 - <3> QED - BY <3>6 DEF Q - <2> QED - BY <2>1, <2>2, NatInduction - <1>2. ASSUME NEW i \in negNat - PROVE P(i) - <2>1. \E n \in Nat: n = -.i - <3>1. i \in negNat - BY <1>3 - <3>2. \E n \in Nat: i = -.n - BY <3>1 DEF negNat - <3>3. \E n \in Nat: -.i = -.-.n - BY <3>2 - <3>4. \A n \in Nat: -.-.n = n - BY minus_involutive, nat_in_int - <3> QED - BY <3>3, <3>4 - <2>2. \E n \in Nat: Q(-.i) - <3>1. \E n \in Nat: n = -.i /\ Q(n) - BY <1>1, <2>1 - <3>2. \E n \in Nat: n = -.i /\ Q(-.i) - BY <3>1 - <3> QED - BY <3>2 - <2>3. Q(-.i) - BY <2>2 - <2>4. P(-.-.i) - BY <2>3 DEF Q - <2>5. -.-.i = i - <3>1. i \in Int - BY <1>2 DEF negNat, Int - <3> QED - BY <3>1, minus_involutive - <2> QED - BY <2>4, <2>5 - <1> QED - BY <1>2 -*) -theorem neg_nat_induction: - assumes z: "P(0)" and - sc: "\ i. - \ - i \\in negNat; - P(i) - \ \ - P(-.Succ[-.i])" - shows "\ i \ negNat: P(i)" - proof - - { - fix "Q" :: "[c] \ c" - assume q_def: "\ n: Q(n) = P(-.n)" - have s1_1: "\ n \ Nat: Q(n)" - proof - - have s2_1: "Q(0)" - proof - - have s3_1: "P(-.0)" - using z neg0 by auto - show "Q(0)" - using s3_1 q_def allE by auto - qed - have s2_2: "\ n. - \ - n \ Nat; Q(n) - \ \ - Q(Succ[n])" - proof - - fix "n" :: "c" - assume ndom: "n \ Nat" - assume Qn: "Q(n)" - have s3_1: "P(-.n)" - using Qn q_def by auto - have s3_2: "-.n \ negNat" - unfolding negNat_def - using ndom by auto - have s3_3: "\ i. \ i \ negNat; - P(i)\ \ - P(-.Succ[-.i])" - using sc by auto - have s3_4: "P(-.n) ==> - P(-.Succ[-.-.n])" - using s3_2 s3_3 by auto - have s3_5: "P(-.Succ[-.-.n])" - using s3_1 s3_4 by auto - have s3_6: "P(-.Succ[n])" - proof - - have s4_1: "-.-.n = n" - using ndom minus_involutive nat_in_int - by auto - show "P(-.Succ[n])" - using s3_5 s4_1 by auto - qed - show "Q(Succ[n])" - using s3_6 q_def by auto - qed - show "\ n \ Nat: Q(n)" - using s2_1 s2_2 natInduct - by auto - qed - have s1_2: "\ i. i \\in negNat ==> P(i)" - proof - - fix "i" :: "c" - assume idom: "i \ negNat" - have s2_1: "\ n \ Nat: n = -.i" - proof - - have s3_1: "i \ negNat" - using idom by auto - have s3_2: "\ n \ Nat: i = -.n" - using s3_1 by (auto simp: negNat_def) - have s3_3: "\ n \ Nat: -.i = -.-.n" - using s3_2 by auto - have s3_4: "\ n \ Nat: -.-.n = n" - using minus_involutive nat_in_int - by auto - show "\ n \ Nat: n = -.i" - using s3_3 s3_4 by auto - qed - have s2_2: "\ n \ Nat: Q(-.i)" - proof - - have s3_1: "\ n \ Nat: - n = -.i \ Q(n)" - using s1_1 s2_1 by auto - have s3_2: "\ n \ Nat: - n = -.i \ Q(-.i)" - using s3_1 by auto - show "\ n \ Nat: Q(-.i)" - using s3_2 by auto - qed - have s2_3: "Q(-.i)" - using s2_2 by auto - have s2_4: "P(-.-.i)" - using s2_3 q_def by auto - have s2_5: "-.-.i = i" - proof - - have s3_1: "i \ Int" - using idom - by (auto simp: negNat_def Int_def) - show "-.-.i = i" - using s3_1 minus_involutive - by auto - qed - show "P(i)" - using s2_4 s2_5 by auto - qed - have "\ i \ negNat: P(i)" - proof - - have s2_1: - "\ i. i \ negNat => P(i)" - using s1_2 impI by auto - have s2_2: - "\ i: i \ negNat => P(i)" - using s2_1 allI by auto - show "\ i \ negNat: P(i)" - using s2_2 by auto - (* by (rule allI, rule impI, rule s1_2) *) - qed - } - from this have "\ Q: ( - (\ n: Q(n) = P(-.n)) - => - (\ i \ negNat: P(i)) - )" - by auto - from this have " - (\ n: P(-.n) = P(-.n)) - => - (\ i \ negNat: P(i))" - by auto - thus " - (\ i \ negNat: P(i))" - by auto - qed - -(*----------------------------------------------------------------------------*) -(* Primitive recursion in two directions: plus and minus infinity. *) - - -theorem bprimrec_int: - assumes - e: "e \ S" and - succ_h: "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - succ_g: "\ n \ Nat: - \ x \ S: - g(n, x) \ S" - shows "\ f \ [Int -> S]: - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - )" - proof - - from primrec_int[of e h g] obtain f where - 1: "isAFcn(f)" and - 2: "DOMAIN f = Int" and - 3: "f[0] = e" and - 4: "\ n \ Nat: - f[Succ[n]] = h(n, f[n]) - \ f[-.Succ[n]] = g(n, f[-.n])" - by blast - have s1_2: "\ n \ Nat: f[n] \ S" - proof (rule natInduct) - from 3 e show "f[0] \ S" by simp - next - fix n - assume "n \ Nat" and "f[n] \ S" - with succ_h 4 show "f[Succ[n]] \ S" - by force - qed - have s1_3: "\ i \ negNat: f[i] \ S" - proof (rule neg_nat_induction) - from 3 e show "f[0] \ S" by simp - next - fix i - assume ndom: "i \ negNat" and - fidom: "f[i] \ S" - have s2_1: "-.i \ Nat" - using ndom minus_neg_nat_in_nat by auto - have s2_2: "f[-.Succ[-.i]] = g(-.i, f[-.-.i])" - using 4 s2_1 by auto - have s2_3: "g(-.i, f[-.-.i]) = g(-.i, f[i])" - using ndom neg_nat_in_int minus_involutive - by auto - have s2_4: "g(-.i, f[i]) \ S" - using s2_1 fidom succ_g by auto - show "f[-.Succ[-.i]] \ S" - using s2_2 s2_3 s2_4 by auto - qed - have 5: "\ i \ Int: f[i] \ S" - using s1_2 s1_3 int_union_nat_negnat - by auto - with 1 2 3 4 5 show ?thesis - by blast - qed - - -theorem primrec_intE: - assumes e: "e \ S" and - succ_h: "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - succ_g: "\ n \ Nat: - \ x \ S: - g(n, x) \ S" and - f: "f = (CHOOSE r \ [Int -> S]: - r[0] = e \ - (\ n \ Nat: - r[Succ[n]] = h(n, r[n]) \ - r[-.Succ[n]] = g(n, r[-.n]) - ))" (is "f = ?r") and - maj: "\ - f \ [Int -> S]; - f[0] = e; - \ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - \ - \ P" - shows "P" - proof - - from e succ_h succ_g have " - \ r \ [Int -> S]: - r[0] = e \ - (\ n \ Nat: - r[Succ[n]] = h(n, r[n]) \ - r[-.Succ[n]] = g(n, r[-.n]) - )" - by (rule bprimrec_int) - hence "?r \ [Int -> S] \ - ?r[0] = e \ - (\ n \ Nat: - ?r[Succ[n]] = h(n, ?r[n]) \ - ?r[-.Succ[n]] = g(n, ?r[-.n]) - )" - by (rule bChooseI2, auto) - with f maj show ?thesis by auto - qed - - -theorem bprimrecType_int: - assumes "e \ S" and - "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - "\ n \ Nat: - \ x \ S: - g(n, x) \ S" - shows "(CHOOSE f \ [Int -> S]: - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - )) \ [Int -> S]" - by (rule primrec_intE[OF assms], auto) - - -(*----------------------------------------------------------------------------*) -(* Typeness of addition and recursive properties. *) - - -theorem addintType: - assumes mdom: "m \ Int" - shows "addint(m) \ [Int -> Int]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: "(CHOOSE f \ [Int -> Int]: - f[0] = m \ - (\ n \ Nat: - f[Succ[n]] = iSucc[f[n]] \ - f[-.Succ[n]] = iPred[f[-.n]] - )) \ [Int -> Int]" - using mdom s1_1 s1_2 - by (rule bprimrecType_int) - show "addint(m) \ [Int -> Int]" - unfolding addint_def - using s1_3 by auto - qed - - -theorem addIsInt: - assumes "m \ Int" and "n \ Int" - shows "add(m, n) \ Int" - unfolding add_def - using assms addintType by blast - - -theorem diffIsInt: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "diff(m, n) \\in Int" - proof - - have s1_1: "diff(m, n) = add(m, -.n)" - unfolding diff_def by auto - have s1_2: "-.n \\in Int" - using ndom neg_int_eq_int by auto - have s1_3: "add(m, -.n) \\in Int" - using mdom s1_2 addIsInt by auto - show ?thesis - using s1_1 s1_3 by auto - qed - - -theorem addint_0: - assumes mdom: "m \ Int" - shows "addint(m)[0] = m" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: " - addint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = m \ - (\ n \ Nat: - r[Succ[n]] = iSucc[r[n]] \ - r[-.Succ[n]] = iPred[r[-.n]]) - )" - unfolding addint_def - by auto - have s1_4: " - \ - addint(m) \ [Int -> Int]; - addint(m)[0] = m; - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - \ - \ addint(m)[0] = m" - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "m" "Int" - "\ n x. iSucc[x]" - "\ n x. iPred[x]" - "addint(m)"] - by auto - qed - - -theorem add_0: - assumes "m \ Int" - shows "m + 0 = m" - unfolding add_def - using assms addint_0 by auto - - -theorem addint_succ: - assumes mdom: "m \ Int" - shows "\ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: " - addint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = m \ - (\ n \ Nat: - r[Succ[n]] = iSucc[r[n]] \ - r[-.Succ[n]] = iPred[r[-.n]]) - )" - unfolding addint_def - by auto - have s1_4: " - \ - addint(m) \ [Int -> Int]; - addint(m)[0] = m; - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - \ - \ - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - " - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "m" "Int" - "\ n x. iSucc[x]" - "\ n x. iPred[x]" - "addint(m)"] - by auto - qed - - -theorem addint_succ_nat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m + Succ[n] = iSucc[m + n]" - unfolding add_def - using mdom ndom addint_succ[of "m"] spec by auto - - -theorem addint_succ_negnat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m + -.Succ[n] = iPred[m + -.n]" - unfolding add_def - using mdom ndom addint_succ[of "m"] spec by auto - - -theorem nat_add_1: - assumes ndom: "n \\in Nat" - shows "Succ[n] = n + 1" - proof - - have s1_1: "Succ[n] = iSucc[n]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "n \\in Nat" - using ndom by auto - show ?thesis - unfolding iSucc_def - using s2_1 s2_2 by auto - qed - also have s1_2: "... = iSucc[n + 0]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "n + 0 = n" - using s2_1 add_0 by auto - have s2_3: "n = n + 0" - using s2_2[symmetric] by auto - show ?thesis - using s2_3 by auto - qed - also have s1_3: "... = n + Succ[0]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - using s2_1 s2_2 addint_succ_nat by auto - qed - also have s1_4: "... = n + 1" - by auto (* 1 is an abbreviation *) - from calculation show ?thesis . - qed - - -theorem nat_add_in_nat: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" - shows "m + n \\in Nat" - proof - - have s1_1: "m + 0 \\in Nat" - proof - - have s2_1: "m + 0 = m" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - show ?thesis - using s3_1 add_0 by auto - qed - have s2_2: "m \\in Nat" - using mdom by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_2: "\ k. \ - k \\in Nat; - m + k \\in Nat - \ \ - m + Succ[k] \\in Nat" - proof - - fix "k" :: "c" - assume s1_2_kdom: "k \\in Nat" - assume s1_2_induct: "m + k \\in Nat" - have s2_1: "m + Succ[k] = iSucc[m + k]" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - have s3_2: "k \\in Nat" - using s1_2_kdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - have s2_2: "iSucc[m + k] = Succ[m + k]" - proof - - have s3_1: "m + k \\in Nat" - using s1_2_induct by auto - have s3_2: "m + k \\in Int" - using s1_2_induct nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_3: "Succ[m + k] \\in Nat" - using s1_2_induct succIsNat by auto - show "m + Succ[k] \\in Nat" - using s2_1 s2_2 s2_3 by auto - qed - have s1_3: "\ k \ Nat: m + k \\in Nat" - using s1_1 s1_2 natInduct[of "\ k. m + k \\in Nat"] by auto - show ?thesis - using s1_3 ndom spec by auto - qed - - -theorem nat_0_succ: - "\ n \ Nat: - n=0 \ - (\ m \ Nat: n = Succ[m])" - by (rule natInduct, auto) - - -theorem nat_add_0: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" and - mn: "m + n = 0" - shows "m = 0 \ n = 0" - proof - - have s1_1: "n \ 0 ==> FALSE" - proof - - assume s2_1: "n \ 0" - have s2_2: "\ k: k \\in Nat \ n = Succ[k]" - using ndom s2_1 nat_0_succ by auto - define P where "P \ - \ x. x \\in Nat \ n = Succ[x]" - define r where "r \ CHOOSE x: P(x)" - have s2_3: "r \\in Nat \ n = Succ[r]" - proof - - have s3_1: "\ x: P(x)" - using s2_2 by (auto simp: P_def) - have s3_2: "P(r)" - using s3_1 chooseI_ex by (auto simp: r_def) - show ?thesis - using s3_2 by (auto simp: P_def) - qed - have s2_4: "m + n = m + Succ[r]" - using s2_3 by auto - have s2_5: "m + Succ[r] = iSucc[m + r]" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - have s3_2: "r \\in Nat" - using s2_3 by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - have s2_6: "iSucc[m + r] = Succ[m + r]" - proof - - have s3_1: "m + r \\in Nat" - using mdom s2_3 nat_add_in_nat by auto - have s3_2: "m + r \\in Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_7: "Succ[m + r] \ 0" - proof - - have s3_1: "m + r \\in Nat" - using mdom s2_3 nat_add_in_nat by auto - show ?thesis - using s3_1 succNotZero by auto - qed - have s2_7: "m + n \ 0" - using s2_4 s2_5 s2_6 s2_7 by auto - show "FALSE" - using s2_7 mn by auto - qed - have s1_2: "n = 0" - using s1_1 by auto - have s1_3: "m + 0 = 0" - using mn s1_2 by auto - have s1_4: "m = 0" - using s1_3 mdom nat_in_int add_0 by auto - show ?thesis - using s1_2 s1_4 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Typeness of multiplication and recursive properties. *) - - -theorem multType: - assumes mdom: "m \ Int" - shows "multint(m) \ [Int -> Int]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: "(CHOOSE f \ [Int -> Int]: - f[0] = 0 \ - (\ n \ Nat: - f[Succ[n]] = f[n] + m \ - f[-.Succ[n]] = f[-.n] + -.m - )) \ [Int -> Int]" - using zero_in_int s1_1 s1_2 - by (rule bprimrecType_int[of "0"]) - show "multint(m) \ [Int -> Int]" - unfolding multint_def - using s1_3 by auto - qed - - -theorem multIsInt: - assumes "m \ Int" and "n \ Int" - shows "mult(m, n) \ Int" - unfolding mult_def - using assms multType by blast - - -theorem multint_0: - assumes mdom: "m \ Int" - shows "multint(m)[0] = 0" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: " - multint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = 0 \ - (\ n \ Nat: - r[Succ[n]] = r[n] + m \ - r[-.Succ[n]] = r[-.n] + -.m) - )" - unfolding multint_def - by auto - have s1_4: " - \ - multint(m) \ [Int -> Int]; - multint(m)[0] = 0; - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - \ - \ - multint(m)[0] = 0" - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "0" "Int" - "\ n x. x + m" - "\ n x. x + -.m" - "multint(m)"] - by auto - qed - - -theorem mult_0: - assumes "m \ Int" - shows "m * 0 = 0" - unfolding mult_def - using assms multint_0 by auto - - -theorem multint_succ: - assumes mdom: "m \ Int" - shows "\ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: " - multint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = 0 \ - (\ n \ Nat: - r[Succ[n]] = r[n] + m \ - r[-.Succ[n]] = r[-.n] + -.m) - )" - unfolding multint_def - by auto - have s1_4: " - \ - multint(m) \ [Int -> Int]; - multint(m)[0] = 0; - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - \ - \ - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - " - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "0" "Int" - "\ n x. x + m" - "\ n x. x + -.m" - "multint(m)"] - by auto - qed - - -theorem multint_succ_nat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m * Succ[n] = m * n + m" - unfolding mult_def - using mdom ndom multint_succ[of "m"] spec by auto - - -theorem multint_succ_negnat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m * -.Succ[n] = m * -.n + -.m" - unfolding mult_def - using mdom ndom multint_succ[of "m"] spec by auto - - -theorem nat_mult_in_nat: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" - shows "m * n \\in Nat" - proof - - have s1_1: "m * 0 \\in Nat" - proof - - have s2_1: "m * 0 = 0" - using mdom nat_in_int mult_0 by auto - have s2_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_2: "\ k. \ - k \\in Nat; - m * k \\in Nat - \ \ - m * Succ[k] \\in Nat" - proof - - fix "k" :: "c" - assume s1_2_kdom: "k \\in Nat" - assume s1_2_induct: "m * k \\in Nat" - have s2_1: "m * Succ[k] = m * k + m" - using mdom s1_2_kdom multint_succ_nat by auto - have s2_2: "m * k + m \\in Nat" - using s1_2_induct mdom nat_add_in_nat by auto - show "m * Succ[k] \\in Nat" - using s2_1 s2_2 by auto - qed - have s1_3: "\ k \ Nat: m * k \\in Nat" - using s1_1 s1_2 natInduct[of "\ k. m * k \\in Nat"] by auto - show ?thesis - using s1_3 ndom spec by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Commutativity of addition. *) - - -(* -THEOREM pred_irrefl == - ASSUME NEW n \in Nat \ {0} - PROVE Pred[n] # n -PROOF -<1>1. Pred[n] \in Nat - BY n \in Nat, Pred_in_nat -<1>2. Succ[Pred[n]] # Pred[n] - BY <1>1, succIrrefl -<1>3. Succ[Pred[n]] = n - BY n \in Nat \ {0}, Succ_Pred_nat -<1> QED - BY <1>2, <1>3 -*) -theorem pred_irrefl: - assumes ndom: "n \\in Nat \ {0}" - shows "Pred[n] \ n" - proof - - have s1_1: "Pred[n] \\in Nat" - using ndom Pred_in_nat by auto - have s1_2: "Succ[Pred[n]] \ Pred[n]" - using s1_1 succIrrefl by auto - have s1_3: "Succ[Pred[n]] = n" - using ndom Succ_Pred_nat by auto - show ?thesis - using s1_2 s1_3 by auto - qed - - -theorem ipred_plus_1: - assumes hyp: "a \ Int" - shows "iPred[a + 1] = a" - proof - - have s1_1: "iPred[a + 1] = iPred[a + Succ[0]]" - by auto - have s1_2: "... = iPred[iSucc[a + 0]]" - proof - - have s2_1: "a \ Int" - using hyp by auto - have s2_2: "0 \ Nat" - using zeroIsNat by auto - have s2_3: "a + Succ[0] = iSucc[a + 0]" - using s2_1 s2_2 addint_succ_nat by auto - show ?thesis - using s2_3 by auto - qed - have s1_3: "... = iPred[iSucc[a]]" - using hyp add_0 by auto - have s1_4: "... = a" - using hyp iSucciPredCommute spec by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -(* -THEOREM AddCommutativeNatNat == - \A j \in Nat: \A i \in Nat: - j + i = i + j -PROOF -<1>1. \A i \in Nat: 0 + i = i + 0 - <2>1. 0 + 0 = 0 + 0 - <3>1. 0 + 0 = addint(0)[0] - BY DEF + - <3>2. @ = 0 - BY DEF addint - <3> QED - BY <3>1, <3>2 - - <2>2. ASSUME NEW i \in Nat, 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - - <3>1. 0 + Succ[i] = iSucc[0 + i] - BY DEF + - <3>2. @ = iSucc[i + 0] - BY <2>2 - <3>3. @ = iSucc[i] - BY DEF + - <3>4. @ = Succ[i] - BY <2>2, i \in Nat - <3>5. @ = Succ[i] + 0 - BY DEF + - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5 - <2> QED - BY <2>1, <2>2, NatInduction - -<1>2. ASSUME NEW j \in Nat, - \A i \in Nat: j + i = i + j - PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] - - <2>1. Succ[j] + 0 = 0 + Succ[j] - <3>1. Succ[j] + 0 = Succ[j] - BY DEF + - <3>2. @ = Succ[j + 0] - BY DEF + - <3>3. @ = Succ[0 + j] - BY <1>2 - <3>4. @ = 0 + Succ[j] - BY DEF + - <3> QED - BY <3>1, <3>2, <3>3, <3>4 - - <2>2. ASSUME NEW i \in Nat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] - - <3>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] - <4>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. Succ[i] \in Nat - BY <2>2, succIsNat - <5> QED - BY <5>1, <5>2, addint_succ - <4>2. @ = iSucc[i + Succ[j]] - BY <2>2 - <4>3. @ = iSucc[iSucc[i + j]] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, addint_succ - <4>4. @ = iSucc[iSucc[j + i]] - BY <1>2, <2>2 - <4> QED - BY <4>1, <4>2, <4>3, <4>4 - - <3>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] - <4>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. Succ[j] \in Nat - BY <1>2, succIsNat - <5> QED - BY <5>1, <5>2, addint_succ - <4>2. @ = iSucc[j + Succ[i]] - <5>1. Succ[i] \in Nat - BY <2>2, succIsNat - <5> QED - BY <1>2, <5>1 - <4>3. @ = iSucc[iSucc[j + i]] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, addint_succ - <4> QED - BY <4>1, <4>2, <4>3 - - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, NatInduction -<1> QED - BY <1>1, <1>2, NatInduction - - -THEOREM AddCommutative == - \A j \in Int: \A i \in Int: - j + i = i + j -PROOF -<1>1. \A j \in Nat: \A i \in Int: - j + i = i + j - <2>1. \A j \in Nat: \A i \in Nat: - j + i = i + j - \* The proof of step <2>1 is the same with the - \* proof of the theorem `AddCommutativeNatNat`. - <3>1. \A i \in Nat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in Nat, - 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - <5>1. 0 + Succ[i] = iSucc[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Nat - BY <4>2 - <6> QED - BY <6>1, <6>2, addint_succ - DEF addd - <5>2. @ = iSucc[i + 0] - BY <4>2 - <5>3. @ = iSucc[i] - <6>1. i \in Int - BY <4>2, nat_in_int - <6> QED - BY <6>1 addint_0 - <5>4. @ = Succ[i] - BY <4>2 DEF iSucc - <5>5. @ = Succ[i] + 0 - <6>1. Succ[i] \in Nat - BY <4>2, succIsNat - <6>2. Succ[i] \in Int - BY <6>1, nat_in_int - <6> QED - BY <6>2, addint_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, NatInduction - <3>2. ASSUME NEW j \in Nat, - \A i \in Nat: j + i = i + j - PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] - <4>1. Succ[j] + 0 = 0 + Succ[j] - <5>1. Succ[j] + 0 = Succ[j] - <6>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <6> QED - BY <6>1, addint_0 - <5>2. @ = iSucc[j] - BY <3>2, nat_in_int DEF iSucc - <5>3. @ = iSucc[j + 0] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iSucc[0 + j] - <6>1. j + 0 = 0 + j - BY <3>2, zeroIsNat - <6> QED - By <6>1 - <5>5. @ = 0 + Succ[j] - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2, addint_succ - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4>2. ASSUME NEW i \in Nat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] - <5>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] - <6>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] - <7>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6>2. @ = iSucc[i + Succ[j]] - BY <4>2 - <6>3. @ = iSucc[iSucc[i + j]] - <7>1. i \in Int - BY <4>2, nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6>4. @ = iSucc[iSucc[j + i]] - BY <3>2, <4>2 - <6> QED - By <6>1, <6>2, <6>3, <6>4 - <5>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] - <6>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] - <7>1. Succ[i] \in Int - BY <4>2, succIsNat, nat_in_int - <7>2. j \in Nat - BY <3>2, succIsNat - <7> QED - BY <7>1, <7>2, addint_succ - <6>2. @ = iSucc[j + Succ[i]] - <7>1. Succ[i] \in Nat - BY <4>2, succIsNat - <7> QED - BY <3>2, <7>1 - <6>3. @ = iSucc[iSucc[j + i]] - <7>1. j \in Int - BY <3>2, nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6> QED - BY <6>1, <6>2, <6>3 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, NatInduction - <3> QED - BY <3>1, <3>2, NatInduction - <2>2. \A j \in Nat: \A i \in negNat: - j + i = i + j - <3>1. \A i \in negNat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in negNat, - 0 + i = i + 0 - PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 - <5>1. 0 + -.Succ[-.i] = iPred[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <6>3. 0 + -.Succ[-.i] = iPred[0 + -.-.i] - BY <6>1, <6>2, addint_succ_negnat - <6>4. -.-.i = i - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7> QED - BY <7>1, minus_involutive - <6> QED - BY <6>3, <6>4 - <5>2. @ = iPred[i + 0] - BY <4>2 - <5>3. @ = iPred[i] - <6>1. i \in Int - BY <4>2, neg_nat_in_int - <6>2. i + 0 = i - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = -.Succ[-.i] - <6>1. i \in Int - BY <4>2, neg_nat_in_int - <6>2. i \notin Nat \ {0} - BY <4>2, neg_nat_not_in_nat_setminus_0 - <6> QED - BY <6>1, <6>2 DEF iPred - <5>5. @ = -.Succ[-.i] + 0 - <6>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <6>2. Succ[-.i] \in Nat - BY <6>1, succIsNat - <6>3. -.Succ[-.i] \in Int - BY <6>2, minus_nat_in_int - <6> QED - BY <6>3, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3>2. ASSUME NEW j \in Nat, - \A i \in negNat: j + i = i + j - PROVE \A i \in negNat: Succ[j] + i = i + Succ[j] - <4>1. Succ[j] + 0 = 0 + Succ[j] - <5>1. Succ[j] + 0 = Succ[j] - <6>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <6> QED - BY <6>1, add_0 - <5>2. @ = iSucc[j] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2 DEF iSucc - <5>3. @ = iSucc[j + 0] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iSucc[0 + j] - <6>1. 0 \in negNat - BY neg0 DEF negNat - <6> QED - BY <3>2, <6>1 - <5>5. @ = 0 + Succ[j] - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2, addint_succ_nat - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4>2. ASSUME NEW i \in negNat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j] - <5>1. Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]] - <6>1. Succ[j] + -.Succ[-.i] = iPred[Succ[j] + i] - <7>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7>3. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7>4. Succ[j] + -.Succ[-.i] - = iPred[Succ[j] + -.-.i] - BY <7>1, <7>2, addint_succ_negnat - <7> QED - BY <7>3, <7>4 - <6>2. @ = iPred[i + Succ[j]] - BY <4>2 - <6>3. @ = iPred[iSucc[i + j]] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>4. @ = iPred[iSucc[j + i]] - <7>1. i \in negNat - BY <4>2 - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + i = i + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. -.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]] - <6>1. -.Succ[-.i] + Succ[j] = iSucc[-.Succ[-.i] + j] - <7>1. -.Succ[-.i] \in Int - BY <4>2, minus_neg_nat_in_nat, succIsNat, - minus_nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>2. @ = iSucc[j + -.Succ[-.i]] - <7>1. -.Succ[-.i] \in negNat - BY <4>2, minus_neg_nat_in_nat, succIsNat, - minus_nat_in_neg_nat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6>3. @ = iSucc[iPred[j + i]] - <7>1. j \in Int - BY <3>2, nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7>3. j + -.Succ[-.i] = iPred[j + -.-.i] - BY <7>1, <7>2, addint_succ_negnat - <7>4. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>3, <7>4 - <6>4. @ = iPred[iSucc[j + i]] - <7>1. j + i \in Int - <8>1. j \in Int - BY <3>2, nat_in_int - <8>2. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, <8>2, addIsInt - <7> QED - BY <7>1, iSucciPredCommute - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3> QED - BY <3>1, <3>2, NatInduction - <2> QED - BY <2>1, <2>2, int_union_nat_negnat -<1>2. \A j \in negNat: \A i \in Int: - j + i = i + j - <2>1. \A j \in negNat: \A i \in Nat: - j + i = i + j - <3>1. \A i \in Nat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME - NEW i \in Nat, - 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - \* The level 5 proof is the same as - \* for step <1>1!<2>1!<3>1!<4>2 . - <5>1. 0 + Succ[i] = iSucc[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Nat - BY <4>2 - <6> QED - BY <6>1, <6>2 DEF addint_succ_nat - <5>2. @ = iSucc[i + 0] - BY <4>2 - <5>3. @ = iSucc[i] - <6>1. i \in Int - BY <4>2, nat_in_int - <6>2. i + 0 = i - BY <6>1, addint_0 - <6> QED - BY <6>2 - <5>4. @ = Succ[i] - <6>1. i \in Nat - BY <4>2 - <6>2. i \in Int - BY <4>2, nat_in_int - <6> QED - BY <6>1, <6>2 DEF iSucc - <5>5. @ = Succ[i] + 0 - <6>1. Succ[i] \in Nat - BY <4>2, succIsNat - <6>2. Succ[i] \in Int - BY <6>1, nat_in_int - <6> QED - BY <6>2, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, NatInduction - <3>2. ASSUME - NEW j \in negNat, - \A i \in Nat: j + i = i + j - PROVE - \A i \in Nat: - -.Succ[-.j] + i = i + -.Succ[-.j] - <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] - <5>1. -.Succ[-.j] + 0 = -.Succ[-.j] - <6>1. -.Succ[-.j] \in Int - <7>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>2. Succ[-.j] \in Nat - BY <7>1, succIsNat - <7> QED - BY <7>2, minus_nat_in_int - <6> QED - BY <6>1, add_0 - <5>2. @ = iPred[j] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. j \notin Nat \ {0} - BY <3>2, neg_nat_not_in_nat_setminus_0 - <6> QED - BY <6>1, <6>2 DEF iPred - <5>3. @ = iPred[j + 0] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iPred[0 + j] - <6>1. 0 \in Nat - BY zeroIsNat - <6>2. \A k \in Nat: j + k = k + j - BY <3>2 - <6>3. j + 0 = 0 + j - BY <6>1, <6>2 - <6> QED - BY <6>3 - <5>5. @ = iPred[0 + -.-.j] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. -.-.j = j - BY <6>1, minus_involutive - <6> QED - BY <6>2 - <5>6. @ = 0 + -.Succ[-.j] - <6>1. 0 \in Int - BY zeroIsNat, nat_in_int - <6>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <6> QED - BY <6>1, <6>2, addint_succ_negnat - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5, <5>6 - <4>2. ASSUME - NEW i \in Nat, - -.Succ[-.j] + i = i + -.Succ[-.j] - PROVE - -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j] - <5>1. -.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]] - <6>1. -.Succ[-.j] + Succ[i] = - iSucc[-.Succ[-.j] + i] - <7>1. -.Succ[-.j] \in Int - <8>1. j \in negNat - BY <3>2 - <8>2. -.j \in Nat - BY <8>1, minus_neg_nat_in_nat - <8>3. Succ[-.j] \in Nat - BY <8>2, succIsNat - <8> QED - BY <8>3, minus_nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>2. @ = iSucc[i + -.Succ[-.j]] - BY <4>2 - <6>3. @ = iSucc[iPred[i + j]] - <7>1. i \in Int - BY <4>2, nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>3. -.-.j = j - BY <3>2, neg_nat_in_int, minus_involutive - <7> QED - BY <7>1, <7>2, <7>3, addint_succ_negnat - <6>4. @ = iSucc[iPred[j + i]] - <7>1. i \in Nat - BY <4>2 - <7>2. j + i = i + j - BY <3>2, <7>1 - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]] - <6>1. Succ[i] + -.Succ[-.j] = - iPred[Succ[i] + -.-.j] - <7>1. Succ[i] \in Int - BY <4>2, succIsNat, nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[Succ[i] + j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>2 - <6>3. @ = iPred[j + Succ[i]] - <7>1. Succ[i] \in Nat - BY <4>2, succIsNat - <7>2. \A k \in Nat: j + k = k + j - BY <3>2 - <7>3. j + Succ[i] = Succ[i] + j - BY <7>1, <7>2 - <7> QED - BY <7>3 - <6>4. @ = iPred[iSucc[j + i]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>5. @ = iSucc[iPred[j + i]] - <7>1. j + i \in Int - <8>1. j \in Int - BY <3>2, neg_nat_in_int - <8>2. i \in Int - BY <4>2, nat_in_int - <8> QED - BY <8>1, <8>2, addIsInt - <7> QED - BY <7>1, iSucciPredCommute - <6> QED - BY <6>1, <6>2, <6>3, <6>4, <6>5 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, NatInduction - <3> QED - BY <3>1, <3>2, neg_nat_induction - <2>2. \A j \in negNat: \A i \in negNat: - j + i = i + j - <3>1. \A i \in negNat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in negNat, 0 + i = i + 0 - PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 - <5>1. 0 + -.Succ[-.i] = iPred[i] - <6>1. 0 + -.Succ[-.i] = iPred[0 + -.-.i] - <7>1. 0 \in Int - BY zero_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[0 + i] - <7>1. -.-.i = i - BY <4>2, neg_nat_in_int, minus_involutive - <7> QED - BY <7>1 - <6>3. @ = iPred[i + 0] - BY <4>2 - <6>4. @ = iPred[i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. i + 0 = i - BY <7>1, add_0 - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. -.Succ[-.i] + 0 = iPred[i] - <6>1. -.Succ[-.i] + 0 = -.Succ[-.i] - <7>1. -.Succ[-.i] \in Int - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7> QED - BY <7>1, add_0 - <6>2. @ = iPred[i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. i \notin Nat \ {0} - BY <4>2, neg_nat_not_in_nat_setminus_0 - <7> QED - BY <7>1, <7>2 DEF iPred - <6> QED - BY <6>1, <6>2 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3>2. ASSUME NEW j \in negNat, - \A i \in negNat: j + i = i + j - PROVE \A i \in negNat: - -.Succ[-.j] + i = i + -.Succ[-.j] - <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] - <5>1. -.Succ[-.j] + 0 = iPred[j] - <6>1. -.Succ[-.j] + 0 = -.Succ[-.j] - <7>1. -.Succ[-.j] \in Int - <8>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <8>2. Succ[-.j] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7> QED - BY <7>1, add_0 - <6>2. @ = iPred[j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. j \notin Nat \ {0} - BY <3>2, neg_nat_not_in_nat_setminus_0 - <7> QED - BY <7>1, <7>2 DEF iPred - <6> QED - BY <6>1, <6>2 - <5>2. 0 + -.Succ[-.j] = iPred[j] - <6>1. 0 + -.Succ[-.j] = iPred[0 + -.-.j] - <7>1. 0 \in Int - BY zero_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[0 + j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>1, <7>2 - <6>3. @ = iPred[j + 0] - <7>1. 0 \in negNat - BY neg0 DEF negNat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + 0 = 0 + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6>4. @ = iPred[j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7> QED - BY <7>1, add_0 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5> QED - BY <5>1, <5>2 - <4>2. ASSUME NEW i \in negNat, - -.Succ[-.j] + i = i + -.Succ[-.j] - PROVE -.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j] - <5>1. -.Succ[-.j] + -.Succ[-.i] = iPred[iPred[i + j]] - <6>1. -.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i] - <7>1. -.Succ[-.j] \in Int - <8>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <8>2. Succ[-.j] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[-.Succ[-.j] + i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. -.-.i = i - BY <7>1, minus_involutive - <7> QED - BY <7>1, <7>2 - <6>3. @ = iPred[i + -.Succ[-.j]] - BY <4>2 - <6>4. @ = iPred[iPred[i + -.-.j]] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>3. i + -.Succ[-.j] = iPred[i + -.-.j] - BY <7>1, <7>2, addint_succ_negnat - <7> QED - BY <7>3 - <6>5. @ = iPred[iPred[i + j]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4, <6>5 - <5>2. -.Succ[-.i] + -.Succ[-.j] = iPred[iPred[i + j]] - <6>1. -.Succ[-.i] + -.Succ[-.j] = - iPred[-.Succ[-.i] + -.-.j] - <7>1. -.Succ[-.i] \in Int - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[-.Succ[-.i] + j] - <7>1. -.-.j = j - <8>1. j \in Int - BY <3>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>1 - <6>3. @ = iPred[j + -.Succ[-.i]] - <7>1. -.Succ[-.i] \in negNat - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_neg_nat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j - BY <7>1, <7>2 - <7> QED - BY <7>3 - <6>4. @ = iPred[iPred[j + -.-.i]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>5. @ = iPred[iPred[j + i]] - <7>1. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>1 - <6>6. @ = iPred[iPred[i + j]] - <7>1. i \in negNat - BY <4>2 - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + i = i + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6> QED - BY <6>1, <6>2, <6>3, <6>4, - <6>5, <6>6 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3> QED - BY <3>1, <3>2, neg_nat_induction - <2> QED - BY <2>1, <2>2, int_union_nat_negnat -<1> QED - BY <1>1, <1>2 - -*) -theorem AddCommutative: - "\ j \ Int: \ i \ Int: - j + i = i + j" - proof - - have s1_1: "\ j \ Nat: \ i \ Int: - j + i = i + j" - proof - - have s2_1: "\ j \ Nat: - \ i \ Nat: - j + i = i + j" - proof - - have s3_1: "\ i \ Nat: - 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ Nat; - 0 + i = i + 0 - \ - \ - 0 + Succ[i] = Succ[i] + 0" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + Succ[i] = iSucc[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - unfolding add_def - using s6_1 s6_2 - addint_succ[of "0"] - by auto - qed - have s5_2: "iSucc[0 + i] = iSucc[i + 0]" - using s4_2_asm2 by auto - have s5_3: "iSucc[i + 0] = iSucc[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 nat_in_int - by auto - show ?thesis - unfolding add_def - using s6_1 addint_0[of "i"] - by auto - qed - have s5_4: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s4_2_asm1 by auto - have s5_5: "Succ[i] = Succ[i] + 0" - proof - - have s6_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - have s6_2: "Succ[i] \ Int" - using s6_1 nat_in_int by auto - show ?thesis - unfolding add_def - using s6_2 addint_0 by auto - qed - show "0 + Succ[i] = Succ[i] + 0" - using s5_1 s5_2 s5_3 s5_4 s5_5 - by auto - qed - show ?thesis - using s4_1 s4_2 - natInduct[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ Nat; - \ i \ Nat: j + i = i + j - \ - \ - \ i \ Nat: - Succ[j] + i = i + Succ[j]" - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ Nat" - assume s3_2_asm2: "\ i \ Nat: - j + i = i + j" - have s4_1: "Succ[j] + 0 = 0 + Succ[j]" - proof - - have s5_1: "Succ[j] + 0 = Succ[j]" - proof - - have s6_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat - nat_in_int by auto - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iSucc[j]" - unfolding iSucc_def - using s3_2_asm1 nat_in_int - by auto - also have s5_3: "... = iSucc[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iSucc[0 + j]" - proof - - have s6_1: "j + 0 = 0 + j" - using s3_2_asm2 zeroIsNat by auto - show ?thesis - using s6_1 by auto - qed - also have s5_5: "... = 0 + Succ[j]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - have s6_3: "0 \ Int ==> - j \ Nat ==> - 0 + Succ[j] = iSucc[0 + j]" - using addint_succ_nat - by auto - have s6_4: " - 0 + Succ[j] = iSucc[0 + j]" - using s6_1 s6_2 s6_3 by auto - have s6_5: "iSucc[0 + j] = 0 + Succ[j]" - using s6_4 by auto - show ?thesis - using s6_5 by auto - qed - finally show ?thesis . - qed - have s4_2: "\ i. \ - i \ Nat; - Succ[j] + i = i + Succ[j] - \ \ - Succ[j] + Succ[i] = Succ[i] + Succ[j]" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" - have s5_1: "Succ[j] + Succ[i] = iSucc[iSucc[j + i]]" - proof - - have s6_1: "Succ[j] + Succ[i] = iSucc[Succ[j] + i]" - proof - - have s7_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "i \ Nat" - using s4_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - also have s6_2: "... = iSucc[i + Succ[j]]" - using s4_2_asm2 by auto - also have s6_3: "... = iSucc[iSucc[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - also have s6_4: "... = iSucc[iSucc[j + i]]" - using s3_2_asm2 s4_2_asm1 spec by auto - finally show ?thesis - by auto - qed - have s5_2: "Succ[i] + Succ[j] = iSucc[iSucc[j + i]]" - proof - - have s6_1: "Succ[i] + Succ[j] = iSucc[Succ[i] + j]" - proof - - have s7_1: "Succ[i] \ Int" - using s4_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "j \ Nat" - using s3_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[j + Succ[i]]" - proof - - have s7_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s3_2_asm2 by auto - qed - also have s6_3: "... = iSucc[iSucc[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 nat_in_int - by auto - have s7_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - finally show ?thesis - by auto - qed - show "Succ[j] + Succ[i] = Succ[i] + Succ[j]" - using s5_1 s5_2 by auto - qed - show "\ i \ Nat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 - natInduct[of "\ i. Succ[j] + i = i + Succ[j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - natInduct[of "\ j. - \ i \ Nat: j + i = i + j"] - by auto - qed - have s2_2: "\ j \ Nat: - \ i \ negNat: j + i = i + j" - proof - - have s3_1: "\ i \ negNat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ negNat; - 0 + i = i + 0 - \ \ - 0 + -.Succ[-.i] = -.Succ[-.i] + 0" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ negNat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + -.Succ[-.i] = iPred[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat - by auto - have s6_3: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" - using s6_1 s6_2 addint_succ_negnat - by auto - have s6_4: "-.-.i = i" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s7_1 minus_involutive by auto - qed - show ?thesis - using s6_3 s6_4 by auto - qed - also have s5_2: "... = iPred[i + 0]" - using s4_2_asm2 by auto - also have s5_3: "... = iPred[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s6_2: "i + 0 = i" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = -.Succ[-.i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s6_2: "i \ Nat \ {0}" - using s4_2_asm1 neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s6_1 s6_2 by auto - qed - also have s5_5: "... = -.Succ[-.i] + 0" - proof - - have s6_1: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat by auto - have s6_2: "Succ[-.i] \ Nat" - using s6_1 succIsNat by auto - have s6_3: "-.Succ[-.i] \ Int" - using s6_2 minus_nat_in_int by auto - show ?thesis - using s6_3 add_0 by auto - qed - finally show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" - by auto - qed - show ?thesis - using s4_1 s4_2 - neg_nat_induction[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ Nat; - \ i \ negNat: j + i = i + j - \ \ - \ i \ negNat: Succ[j] + i = i + Succ[j]" - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ Nat" - assume s3_2_asm2: "\ i \ negNat: - j + i = i + j" - have s4_1: "Succ[j] + 0 = 0 + Succ[j]" - proof - - have s5_1: "Succ[j] + 0 = Succ[j]" - proof - - have s6_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int by auto - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iSucc[j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - unfolding iSucc_def - using s6_1 s6_2 by auto - qed - also have s5_3: "... = iSucc[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iSucc[0 + j]" - proof - - have s6_1: "0 \ negNat" - unfolding negNat_def - using neg0 by auto - show ?thesis - using s3_2_asm2 s6_1 by auto - qed - also have s5_5: "... = 0 + Succ[j]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s6_1 s6_2 addint_succ_nat by auto - qed - finally show ?thesis - by auto - qed - have s4_2: "\ i. \ - i \ negNat; - Succ[j] + i = i + Succ[j] - \ \ - Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ negNat" - assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" - have s5_1: "Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]]" - proof - - have s6_1: "Succ[j] + -.Succ[-.i] = - iPred[Succ[j] + i]" - proof - - have s7_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat - by auto - have s7_3: "-.-.i = i" - using s4_2_asm1 neg_nat_in_int - minus_involutive by auto - have s7_4: "Succ[j] + -.Succ[-.i] = - iPred[Succ[j] + -.-.i]" - using s7_1 s7_2 addint_succ_negnat - by auto - show ?thesis - using s7_3 s7_4 by auto - qed - also have s6_2: "... = iPred[i + Succ[j]]" - using s4_2_asm2 by auto - also have s6_3: "... = iPred[iSucc[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "i \ negNat" - using s4_2_asm1 by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + i = i + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - finally show ?thesis - by auto - qed - have s5_2: "-.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]]" - proof - - have s6_1: "-.Succ[-.i] + Succ[j] = - iSucc[-.Succ[-.i] + j]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - using s4_2_asm1 minus_neg_nat_in_nat succIsNat - minus_nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[j + -.Succ[-.i]]" - proof - - have s7_1: "-.Succ[-.i] \ negNat" - using s4_2_asm1 minus_neg_nat_in_nat succIsNat - minus_nat_in_neg_nat by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - also have s6_3: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat by auto - have s7_3: "j + -.Succ[-.i] = iPred[j + -.-.i]" - using s7_1 s7_2 addint_succ_negnat by auto - have s7_4: "-.-.i = i" - proof - - have s8_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_3 s7_4 by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "j + i \ Int" - proof - - have s8_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s8_2: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s8_1 s8_2 addIsInt by auto - qed - show ?thesis - using s7_1 iSucciPredCommute by auto - qed - finally show ?thesis - by auto - qed - show "Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" - using s5_1 s5_2 by auto - qed - show "\ i \ negNat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 - neg_nat_induction[of - "\ i. Succ[j] + i = i + Succ[j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - natInduct[of "\ j. - \ i \ negNat: j + i = i + j"] - by auto - qed - show ?thesis - proof - - have s3_1: "\ j. \ i. - j \ Nat => ( - i \ Int => ( - j + i = i + j))" - proof auto - fix "j" :: "c" - fix "i" :: "c" - assume s4_1: "j \ Nat" - assume s4_2: "i \ Int" - have s4_3: "i \ Nat ==> j + i = i + j" - proof - - assume s5_1: "i \ Nat" - show "j + i = i + j" - using s4_1 s5_1 s2_1 by auto - qed - have s4_4: "i \ negNat ==> j + i = i + j" - proof - - assume s5_1: "i \ negNat" - show "j + i = i + j" - using s4_1 s5_1 s2_2 by auto - qed - show "j + i = i + j" - using s4_2 s4_3 s4_4 int_union_nat_negnat by auto - qed - show ?thesis - using s3_1 by auto - qed - qed - have s1_2: "\ j \ negNat: \ i \ Int: - j + i = i + j" - proof - - have s2_1: "\ j \ negNat: \ i \ Nat: - j + i = i + j" - proof - - have s3_1: "\ i \ Nat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ Nat; - 0 + i = i + 0 - \ \ - 0 + Succ[i] = Succ[i] + 0 - " - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + Succ[i] = iSucc[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - using s6_1 s6_2 addint_succ_nat by auto - qed - also have s5_2: "... = iSucc[i + 0]" - using s4_2_asm2 by force - also have s5_3: "... = iSucc[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 nat_in_int by auto - have s6_2: "i + 0 = i" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by force - qed - also have s5_4: "... = Succ[i]" - proof - - have s6_1: "i \ Nat" - using s4_2_asm1 by auto - have s6_2: "i \ Int" - using s4_2_asm1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s6_1 s6_2 by auto - qed - also have s5_5: "... = Succ[i] + 0" - proof - - have s6_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - have s6_2: "Succ[i] \ Int" - using s6_1 nat_in_int by auto - show ?thesis - using s6_2 add_0 by auto - qed - finally show "0 + Succ[i] = Succ[i] + 0" - (* . did not work *) - using s5_1 s5_2 s5_3 s5_4 s5_5 by auto - qed - show ?thesis - using s4_1 s4_2 - natInduct[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ negNat; - \ i \ Nat: j + i = i + j - \ \ - \ i \ Nat: - -.Succ[-.j] + i = i + -.Succ[-.j] - " - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ negNat" - assume s3_2_asm2: "\ i \ Nat: - j + i = i + j" - have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" - proof - - have s6_1: "-.Succ[-.j] \ Int" - proof - - have s7_1: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat - by auto - have s7_2: "Succ[-.j] \ Nat" - using s7_1 succIsNat by auto - show ?thesis - using s7_2 minus_nat_in_int by auto - qed - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iPred[j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "j \ Nat \ {0}" - using s3_2_asm1 neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s6_1 s6_2 by auto - qed - also have s5_3: "... = iPred[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iPred[0 + j]" - proof - - have s6_1: "0 \ Nat" - using zeroIsNat by auto - have s6_2: "\ k \ Nat: - j + k = k + j" - using s3_2_asm2 by auto - have s6_3: "j + 0 = 0 + j" - using s6_1 s6_2 by auto - show ?thesis - using s6_3 by auto - qed - also have s5_5: "... = iPred[0 + -.-.j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "-.-.j = j" - using s6_1 minus_involutive by auto - show ?thesis - using s6_2 by auto - qed - also have s5_6: "... = 0 + -.Succ[-.j]" - proof - - have s6_1: "0 \ Int" - using zeroIsNat nat_in_int by auto - have s6_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - show ?thesis - using s6_1 s6_2 addint_succ_negnat by auto - qed - finally show ?thesis . - qed - have s4_2: "\ i. \ - i \ Nat; - -.Succ[-.j] + i = i + -.Succ[-.j] - \ \ - -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ Nat" - assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" - have s5_1: "-.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]]" - proof - - have s6_1: "-.Succ[-.j] + Succ[i] = - iSucc[-.Succ[-.j] + i]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "j \ negNat" - using s3_2_asm1 by auto - have s8_2: "-.j \ Nat" - using s8_1 minus_neg_nat_in_nat by auto - have s8_3: "Succ[-.j] \ Nat" - using s8_2 succIsNat by auto - show ?thesis - using s8_3 minus_nat_in_int by auto - qed - have s7_2: "i \ Nat" - using s4_2_idom by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[i + -.Succ[-.j]]" - using s4_2_induct by auto - also have s6_3: "... = iSucc[iPred[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_idom nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - have s7_3: "-.-.j = j" - using s3_2_asm1 neg_nat_in_int - minus_involutive by auto - show ?thesis - using s7_1 s7_2 s7_3 addint_succ_negnat - by auto - qed - also have s6_4: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "i \ Nat" - using s4_2_idom by auto - have s7_2: "j + i = i + j" - using s3_2_asm2 s7_1 spec by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]]" - proof - - have s6_1: "Succ[i] + -.Succ[-.j] = - iPred[Succ[i] + -.-.j]" - proof - - have s7_1: "Succ[i] \ Int" - using s4_2_idom succIsNat nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[Succ[i] + j]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_2 by auto - qed - also have s6_3: "... = iPred[j + Succ[i]]" - proof - - have s7_1: "Succ[i] \ Nat" - using s4_2_idom succIsNat by auto - have s7_2: "\ k \ Nat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + Succ[i] = Succ[i] + j" - using s7_1 s7_2 by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s7_2: "i \ Nat" - using s4_2_idom by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_5: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "j + i \ Int" - proof - - have s8_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s8_2: "i \ Int" - using s4_2_idom nat_in_int by auto - show ?thesis - using s8_1 s8_2 addIsInt by auto - qed - show ?thesis - using s7_1 iSucciPredCommute by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" - using s5_1 s5_2 by auto - qed - show "\ i \ Nat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 - natInduct[of "\ i. - -.Succ[-.j] + i = i + -.Succ[-.j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - neg_nat_induction[of "\ j. - \ i \ Nat: j + i = i + j"] - by auto - qed - have s2_2: "\ j \ negNat: \ i \ negNat: - j + i = i + j" - proof - - have s3_1: "\ i \ negNat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ negNat; - 0 + i = i + 0 - \ \ - 0 + -.Succ[-.i] = -.Succ[-.i] + 0" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ negNat" - assume s4_2_induct: "0 + i = i + 0" - have s5_1: "0 + -.Succ[-.i] = iPred[i]" - proof - - have s6_1: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" - proof - - have s7_1: "0 \ Int" - using zero_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[0 + i]" - proof - - have s7_1: "-.-.i = i" - using s4_2_idom neg_nat_in_int - minus_involutive by auto - show ?thesis - using s7_1 by auto - qed - also have s6_3: "... = iPred[i + 0]" - using s4_2_induct by auto - also have s6_4: "... = iPred[i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "i + 0 = i" - using s7_1 add_0 by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "-.Succ[-.i] + 0 = iPred[i]" - proof - - have s6_1: "-.Succ[-.i] + 0 = -.Succ[-.i]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - show ?thesis - using s7_1 add_0 by auto - qed - also have s6_2: "... = iPred[i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "i \ Nat \ {0}" - using s4_2_idom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s7_1 s7_2 by auto - qed - finally show ?thesis . - qed - show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" - using s5_1 s5_2 by auto - qed - show ?thesis - using s4_1 s4_2 - neg_nat_induction[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ negNat; - \ i \ negNat: j + i = i + j - \ \ - \ i \ negNat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - proof - - fix "j" :: "c" - assume s3_2_jdom: "j \ negNat" - assume s3_2_induct: "\ i \ negNat: - j + i = i + j" - have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] + 0 = iPred[j]" - proof - - have s6_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.j] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - show ?thesis - using s7_1 add_0 by auto - qed - also have s6_2: "... = iPred[j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "j \ Nat \ {0}" - using s3_2_jdom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s7_1 s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "0 + -.Succ[-.j] = iPred[j]" - proof - - have s6_1: "0 + -.Succ[-.j] = iPred[0 + -.-.j]" - proof - - have s7_1: "0 \ Int" - using zero_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[0 + j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_1 s7_2 by auto - qed - also have s6_3: "... = iPred[j + 0]" - proof - - have s7_1: "0 \ negNat" - unfolding negNat_def - using neg0 by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + 0 = 0 + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - show ?thesis - using s7_1 add_0 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s5_1 s5_2 by auto - qed - have s4_2: "\ i. \ - i \ negNat; - -.Succ[-.j] + i = i + -.Succ[-.j] - \ \ - -.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j]" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ negNat" - assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" - have s5_1: "-.Succ[-.j] + -.Succ[-.i] = - iPred[iPred[i + j]]" - proof - - have s6_1: "-.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.j] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s7_3: "-.Succ[-.j] \ Int - ==> - -.i \ Nat - ==> - -.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i]" - using addint_succ_negnat[ - of "-.Succ[-.j]" "-.i"] - by auto - show ?thesis - by (rule s7_3, rule s7_1, rule s7_2) - qed - also have s6_2: "... = iPred[-.Succ[-.j] + i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "-.-.i = i" - using s7_1 minus_involutive by auto - show ?thesis - using s7_1 s7_2 by auto - qed - also have s6_3: "... = iPred[i + -.Succ[-.j]]" - using s4_2_induct by auto - also have s6_4: "... = iPred[iPred[i + -.-.j]]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat by auto - have s7_3: "i + -.Succ[-.j] = iPred[i + -.-.j]" - using s7_1 s7_2 addint_succ_negnat by auto - show ?thesis - using s7_3 by auto - qed - also have s6_5: "... = iPred[iPred[i + j]]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "-.Succ[-.i] + -.Succ[-.j] = - iPred[iPred[i + j]]" - proof - - have s6_1: "-.Succ[-.i] + -.Succ[-.j] = - iPred[-.Succ[-.i] + -.-.j]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[-.Succ[-.i] + j]" - proof - - have s7_1: "-.-.j = j" - proof - - have s8_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_1 by auto - qed - also have s6_3: "... = iPred[j + -.Succ[-.i]]" - proof - - have s7_1: "-.Succ[-.i] \ negNat" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_neg_nat by auto - qed - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" - using s7_1 s7_2 by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[iPred[j + -.-.i]]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_5: "... = iPred[iPred[j + i]]" - proof - - have s7_1: "-.-.i = i" - proof - - have s8_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_1 by auto - qed - also have s6_6: "... = iPred[iPred[i + j]]" - proof - - have s7_1: "i \ negNat" - using s4_2_idom by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + i = i + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j]" - using s5_1 s5_2 by auto - qed - show "\ i \ negNat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 - neg_nat_induction[of "\ i. - -.Succ[-.j] + i = i + -.Succ[-.j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - neg_nat_induction[of "\ j. - \ i \ negNat: j + i = i + j"] - by auto - qed - show ?thesis - proof - - have s3_1: "\ j. \ i. - j \ negNat => ( - i \ Int => ( - j + i = i + j))" - proof auto - fix "j" :: "c" - fix "i" :: "c" - assume s4_1: "j \ negNat" - assume s4_2: "i \ Int" - have s4_3: "i \ Nat ==> j + i = i + j" - proof - - assume s5_1: "i \ Nat" - show "j + i = i + j" - using s4_1 s5_1 s2_1 by auto - qed - have s4_4: "i \ negNat ==> j + i = i + j" - proof - - assume s5_1: "i \ negNat" - show "j + i = i + j" - using s4_1 s5_1 s2_2 by auto - qed - show "j + i = i + j" - using s4_2 s4_3 s4_4 int_union_nat_negnat by auto - qed - show ?thesis - using s3_1 by auto - qed - qed - show ?thesis - using s1_1 s1_2 int_union_nat_negnat by auto - qed - - -theorem AddCommutative_sequent: - assumes mdom: "m \ Int" and ndom: "n \ Int" - shows "m + n = n + m" - using mdom ndom AddCommutative by auto - - -theorem add_0_left: - assumes mdom: "m \ Int" - shows "0 + m = m" - proof - - have s1_1: "m + 0 = m" - using mdom add_0 by auto - have s1_2: "m + 0 = 0 + m" - using mdom zero_in_int AddCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - - -(*----------------------------------------------------------------------------*) -(* Associativity of addition. *) - - -(* -THEOREM iSuccRightDistributesAdd == - ASSUME NEW a \in Int - PROVE \A b \in Int: iSucc[a + b] = a + iSucc[b] -<1>1. ASSUME NEW b \in Nat - PROVE iSucc[a + b] = a + iSucc[b] - <2>1. iSucc[a + b] = a + Succ[b] - <3>1. a \in Int - OBVIOUS - <3>2. b \in Nat - BY <1>1 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>2. @ = a + iSucc[b] - <3>1. b \in Nat - BY <1>1 - <3>2. b \in Int - BY <1>1, nat_in_int - <3>3. iSucc[b] = Succ[b] - BY <3>1, <3>2 DEF iSucc - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2 -<1>2. \A b \in negNat: iSucc[a + b] = a + iSucc[b] - <2>1. iSucc[a + 0] = a + iSucc[0] - <3>1. iSucc[a + 0] = a + Succ[0] - <4>1. a \in Int - OBVIOUS - <4>2. 0 \in Nat - BY zeroIsNat - <4> QED - BY <4>1, <4>2, addint_succ_nat - <3>2. @ = a + iSucc[0] - <4>1. 0 \in Nat - BY zeroIsNat - <4>2. 0 \in Int - BY zeroIsNat, nat_in_int - <4>3. iSucc[0] = Succ[0] - BY <4>1, <4>2 DEF iSucc - <4> QED - BY <4>3 - <3> QED - BY <3>1, <3>2 - <2>2. ASSUME NEW b \in negNat, - iSucc[a + b] = a + iSucc[b] - PROVE iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]] - <3>1. iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]] - <4>1. a \in Int - OBVIOUS - <4>2. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <4>3. a + -.Succ[-.b] = iPred[a + -.-.b] - BY <4>1, <4>2, addint_succ_negnat - <4>4. -.-.b = b - <5>1. b \in Int - BY <2>2, neg_nat_in_int - <5> QED - BY <5>1, minus_involutive - <4>5. a + -.Succ[-.b] = iPred[a + b] - BY <4>3, <4>4 - <4> QED - BY <4>5 - <3>2. @ = iPred[iSucc[a + b]] - <4>1. a + b \in Int - <5>1. a \in Int - OBVIOUS - <5>2. b \in Int - BY <2>2, neg_nat_in_int - <5> QED - BY <5>1, <5>2, addIsInt - <4> QED - BY <4>1, iSucciPredCommute - <3>3. CASE b = 0 - <4>1. iPred[iSucc[a + b]] = a - <5>1. iPred[iSucc[a + b]] = iPred[iSucc[a + 0]] - BY <3>3 - <5>2. @ = iPred[iSucc[a]] - <6>1. a + 0 = a - BY a \in Int, add_0 - <6> QED - BY <6>1 - <5>3. @ = a - <6>1. a \in Int - OBVIOUS - <6> QED - BY <6>1, iSucciPredCommute, spec - <5> QED - BY <5>1, <5>2, <5>3 - <4>2. a + iSucc[-.Succ[-.b]] = a - <5>1. a + iSucc[-.Succ[-.b]] = a + iSucc[-.Succ[0]] - BY <3>3 - <5>2. @ = a + iSucc[-.1] - OBVIOUS \* 1 is an abbreviation - <5>3. @ = a + -.Pred[-.-.1] - <6>1. -.1 \in Int - BY oneIsNat, minus_nat_in_int - <6>2. -.1 \notin Nat - <7>1. 1 \in Nat \ {0} - BY oneIsNat succNotZero - <7> QED - BY <7>1, minus_nat_0_in_negnat_0, - neg_nat_0_not_in_nat - <6>3. iSucc[-.1] = -.Pred[-.-.1] - BY <6>1, <6>2 DEF iSucc - <6> QED - BY <6>3 - <5>4. @ = a + -.Pred[1] - <6>1. -.-.1 = 1 - <7>1. 1 \in Int - BY oneIsNat nat_in_int - <7> QED - BY <7>1, minus_involutive - <6> QED - BY <6>1 - <5>5. @ = a + -.0 - BY pred_1 - <5>6. @ = a + 0 - BY neg0 - <5>7. @ = a - BY a \in Int, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, - <5>5, <5>6, <5>7 - <4> QED - BY <3>1, <3>2, <4>1, <4>2 - <3>4. CASE b # 0 - <4>1. iPred[iSucc[a + b]] = - iPred[a + iSucc[b]] - BY <2>2 - <4>2. @ = a + -.Succ[-.iSucc[b]] - <5>1. a \in Int - OBVIOUS - <5>2. /\ -.iSucc[b] \in Nat - /\ iSucc[b] \in negNat - <6>1. b \in negNat \ {0} - BY <2>2, <3>4 - <6>2. iSucc[b] = -.Pred[-.b] - <7>1. b \in Int - BY <6>1, neg_nat_in_int - <7>2. b \notin Nat - BY <6>1, neg_nat_0_not_in_nat - <7> QED - BY <7>1, <7>2 DEF iSucc - <6>3. Pred[-.b] \in Nat - <7>1. -.b \in Nat \ {0} - BY <6>1, minus_neg_nat_0_in_nat_0 - <7> QED - BY <7>1, Pred_in_nat - <6>4. -.iSucc[b] \in Nat - <7>1. -.iSucc[b] = Pred[-.b] - <8>1. -.iSucc[b] = -.-.Pred[-.b] - BY <6>2 - <8>2. Pred[-.b] \in Int - BY <6>3, nat_in_int - <8>3. -.-.Pred[-.b] = Pred[-.b] - BY <8>2, minus_involutive - <8> QED - BY <8>1, <8>3 - <7> QED - BY <6>3, <7>1 - <6>5. iSucc[b] \in negNat - BY <6>2, <6>3, minus_nat_in_neg_nat - <6> QED - BY <6>4, <6>5 - <5>3. a + -.Succ[-.iSucc[b]] = - iPred[a + -.-.iSucc[b]] - BY <5>1, <5>2, addint_succ_negnat - <5>4. -.-.iSucc[b] = iSucc[b] - <6>1. iSucc[b] \in Int - BY <5>2, neg_nat_in_int - <6> QED - BY <6>1, minus_involutive - <5> QED - BY <5>3, <5>4 - <4>3. @ = a + -.Succ[Pred[-.b]] - <5>1. b \in Int - BY <2>2, neg_nat_in_int - <5>2. b \notin Nat - <6>1. b \in negNat \ {0} - BY <2>2, <3>4 - <6> QED - BY <6>1, neg_nat_0_not_in_nat - <5>3. iSucc[b] = -.Pred[-.b] - BY <5>1, <5>2 DEF iSucc - <5>4. -.iSucc[b] = -.-.Pred[-.b] - BY <5>3 - <5>5. -.-.Pred[-.b] = Pred[-.b] - <6>1. -.b \in Nat - BY <2>2, <3>4, minus_neg_nat_0_in_nat_0 - <6>2. Pred[-.b] \in Nat - BY <6>1, Pred_in_nat - <6>3. Pred[-.b] \in Int - BY <6>2, nat_in_int - <6> QED - BY <6>3, minus_involutive - <5>6. -.iSucc[b] = Pred[-.b] - BY <5>4, <5>5 - <5> QED - BY <5>6 - <4>4. @ = a + -.Pred[Succ[-.b]] - <5>1. b \in negNat \ {0} - BY <2>2, <3>4 - <5>2. -.b \in Nat \ {0} - BY <5>1, minus_neg_nat_0_in_nat_0 - <5>3. Succ[Pred[-.b]] = Pred[Succ[-.b]] - BY <5>2, Succ_Pred_nat, Pred_Succ_nat - <5> QED - BY <5>3 - <4>5. @ = a + -.Pred[-.-.Succ[-.b]] - <5>1. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <5>2. Succ[-.b] \in Nat - BY <5>1, succIsNat - <5>3. Succ[-.b] \in Int - BY <5>2, nat_in_int - <5>4. -.-.Succ[-.b] = Succ[-.b] - BY <5>3, minus_involutive - <5> QED - BY <5>4 - <4>6. @ = a + iSucc[-.Succ[-.b]] - <5>1. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <5>2. Succ[-.b] \in Nat \ {0} - BY <5>1, succIsNat, succNotZero - <5>3. -.Succ[-.b] \in negNat \ {0} - BY <5>2, minus_nat_0_in_negnat_0 - <5>4. -.Succ[-.b] \in Int - BY <5>3, neg_nat_in_int - <5>5. -.Succ[-.b] \notin Nat - BY <5>3, neg_nat_0_not_in_nat - <5> QED - BY <5>4, <5>5 DEF iSucc - <4> QED - BY <3>1, <3>2, <4>1, <4>2, <4>3, - <4>4, <4>5, <4>6 - <3> QED - BY <3>3, <3>4 - <2> QED - BY <2>1, <2>2, neg_nat_induction -<1> QED - BY <1>1, <1>2 -*) -theorem iSuccRightDistributesAdd: - "\ a \ Int: \ b \ Int: - iSucc[a + b] = a + iSucc[b]" - proof - - { - fix "a" :: "c" - assume adom: "a \ Int" - have s1_1: "\ b \ Nat: iSucc[a + b] = a + iSucc[b]" - proof auto - fix "b" :: "c" - assume bdom: "b \ Nat" - have s2_1: "iSucc[a + b] = a + Succ[b]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b \ Nat" - using bdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - also have s2_2: "... = a + iSucc[b]" - proof - - have s3_1: "b \ Nat" - using bdom by auto - have s3_2: "b \ Int" - using bdom nat_in_int by auto - have s3_3: "iSucc[b] = Succ[b]" - unfolding iSucc_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "iSucc[a + b] = a + iSucc[b]" . - qed - have s1_2: "\ b \ negNat: iSucc[a + b] = a + iSucc[b]" - proof - - have s2_1: "iSucc[a + 0] = a + iSucc[0]" - proof - - have s3_1: "iSucc[a + 0] = a + Succ[0]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s4_1 s4_2 addint_succ_nat by auto - qed - have s3_2: "... = a + iSucc[0]" - proof - - have s4_1: "0 \ Nat" - using zeroIsNat by auto - have s4_2: "0 \ Int" - using zeroIsNat nat_in_int by auto - have s4_3: "iSucc[0] = Succ[0]" - unfolding iSucc_def - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - have s2_2: "\ b. \ - b \ negNat; - iSucc[a + b] = a + iSucc[b] - \ \ - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - fix "b" :: "c" - assume s2_2_bdom: "b \ negNat" - assume s2_2_induct: "iSucc[a + b] = a + iSucc[b]" - have s3_1: "iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s4_3: "a + -.Succ[-.b] = iPred[a + -.-.b]" - using s4_1 s4_2 addint_succ_negnat by auto - have s4_4: "-.-.b = b" - proof - - have s5_1: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - show ?thesis - using s5_1 minus_involutive by auto - qed - have s4_5: "a + -.Succ[-.b] = iPred[a + b]" - using s4_3 s4_4 by auto - show ?thesis - using s4_5 by auto - qed - have s3_2: "... = iPred[iSucc[a + b]]" - proof - - have s4_1: "a + b \ Int" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - show ?thesis - using s5_1 s5_2 addIsInt by auto - qed - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - have s3_3: "b = 0 ==> - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - assume s3_3_bdom: "b = 0" - have s4_1: "iPred[iSucc[a + b]] = a" - proof - - have s5_1: "iPred[iSucc[a + b]] = iPred[iSucc[a + 0]]" - using s3_3_bdom by auto - also have s5_2: "... = iPred[iSucc[a]]" - proof - - have s6_1: "a + 0 = a" - using adom add_0 by auto - show ?thesis - using s6_1 by auto - qed - also have s5_3: "... = a" - using adom iSucciPredCommute spec by auto - finally show ?thesis . - qed - have s4_2: "a + iSucc[-.Succ[-.b]] = a" - proof - - have s5_1: "a + iSucc[-.Succ[-.b]] = - a + iSucc[-.Succ[0]]" - using s3_3_bdom by auto - have s5_2: "... = a + iSucc[-.1]" - by auto - have s5_3: "... = a + -.Pred[-.-.1]" - (* `also` raises error: - empty result sequence - *) - proof - - have s6_1: "-.1 \ Int" - using oneIsNat minus_nat_in_int by auto - have s6_2: "-.1 \ Nat" - proof - - have s7_1: "1 \ Nat \ {0}" - using oneIsNat succNotZero by auto - show ?thesis - using s7_1 minus_nat_0_in_negnat_0 - neg_nat_0_not_in_nat by auto - qed - have s6_3: "iSucc[-.1] = -.Pred[-.-.1]" - unfolding iSucc_def - using s6_1 s6_2 by auto - show ?thesis - using s6_3 by auto - qed - have s5_4: "... = a + -.Pred[1]" - proof - - have s6_1: "-.-.1 = 1" - proof - - have s7_1: "1 \ Int" - using oneIsNat nat_in_int by auto - show ?thesis - using s7_1 minus_involutive by auto - qed - show ?thesis - using s6_1 by auto - qed - have s5_5: "... = a + -.0" - using pred_1 by auto - have s5_6: "... = a + 0" - using neg0 by auto - have s5_7: "... = a" - using adom add_0 by auto - show ?thesis - using s5_1 s5_2 s5_3 s5_4 s5_5 s5_6 s5_7 - by auto - qed - show ?thesis - using s3_1 s3_2 s4_1 s4_2 by auto - qed - have s3_4: "b \ 0 ==> - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - assume s3_4_bdom: "b \ 0" - have s4_1: "iPred[iSucc[a + b]] = - iPred[a + iSucc[b]]" - using s2_2_induct by auto - also have s4_2: "... = a + -.Succ[-.iSucc[b]]" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "-.iSucc[b] \ Nat \ - iSucc[b] \ negNat" - proof - - have s6_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - have s6_2: "iSucc[b] = -.Pred[-.b]" - proof - - have s7_1: "b \ Int" - using s6_1 neg_nat_in_int by auto - have s7_2: "b \ Nat" - using s6_1 neg_nat_0_not_in_nat by auto - show ?thesis - unfolding iSucc_def - using s7_1 s7_2 by auto - qed - have s6_3: "Pred[-.b] \ Nat" - proof - - have s7_1: "-.b \ Nat \ {0}" - using s6_1 minus_neg_nat_0_in_nat_0 by auto - show ?thesis - using s7_1 Pred_in_nat by auto - qed - have s6_4: "-.iSucc[b] \ Nat" - proof - - have s7_1: "-.iSucc[b] = Pred[-.b]" - proof - - have s8_1: "-.iSucc[b] = -.-.Pred[-.b]" - using s6_2 by auto - have s8_2: "Pred[-.b] \ Int" - using s6_3 nat_in_int by auto - have s8_3: "-.-.Pred[-.b] = Pred[-.b]" - using s8_2 minus_involutive by auto - show ?thesis - using s8_1 s8_3 by auto - qed - show ?thesis - using s6_3 s7_1 by auto - qed - have s6_5: "iSucc[b] \ negNat" - using s6_2 s6_3 minus_nat_in_neg_nat by auto - show ?thesis - using s6_4 s6_5 by auto - qed - have s5_3: "a + -.Succ[-.iSucc[b]] = - iPred[a + -.-.iSucc[b]]" - using s5_1 s5_2 addint_succ_negnat by auto - have s5_4: "-.-.iSucc[b] = iSucc[b]" - proof - - have s6_1: "iSucc[b] \ Int" - using s5_2 neg_nat_in_int by auto - show ?thesis - using s6_1 minus_involutive by auto - qed - show ?thesis - using s5_3 s5_4 by auto - qed - also have s4_3: "... = a + -.Succ[Pred[-.b]]" - proof - - have s5_1: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - have s5_2: "b \ Nat" - proof - - have s6_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - show ?thesis - using s6_1 neg_nat_0_not_in_nat by auto - qed - have s5_3: "iSucc[b] = -.Pred[-.b]" - unfolding iSucc_def - using s5_1 s5_2 by auto - have s5_4: "-.iSucc[b] = -.-.Pred[-.b]" - using s5_3 by auto - have s5_5: "-.-.Pred[-.b] = Pred[-.b]" - proof - - have s6_1: "-.b \ Nat" - using s2_2_bdom s3_4_bdom - minus_neg_nat_0_in_nat_0 - by auto - have s6_2: "Pred[-.b] \ Nat" - using s6_1 Pred_in_nat by auto - have s6_3: "Pred[-.b] \ Int" - using s6_2 nat_in_int by auto - show ?thesis - using s6_3 minus_involutive by auto - qed - have s5_6: "-.iSucc[b] = Pred[-.b]" - using s5_4 s5_5 by auto - show ?thesis - using s5_6 by auto - qed - also have s4_4: "... = a + -.Pred[Succ[-.b]]" - proof - - have s5_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - have s5_2: "-.b \ Nat \ {0}" - using s5_1 minus_neg_nat_0_in_nat_0 by auto - have s5_3: "Succ[Pred[-.b]] = Pred[Succ[-.b]]" - using s5_2 Succ_Pred_nat Pred_Succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = a + -.Pred[-.-.Succ[-.b]]" - proof - - have s5_1: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s5_2: "Succ[-.b] \ Nat" - using s5_1 succIsNat by auto - have s5_3: "Succ[-.b] \ Int" - using s5_2 nat_in_int by auto - have s5_4: "-.-.Succ[-.b] = Succ[-.b]" - using s5_3 minus_involutive by auto - show ?thesis - using s5_4 by auto - qed - also have s4_6: "... = a + iSucc[-.Succ[-.b]]" - proof - - have s5_1: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s5_2: "Succ[-.b] \ Nat \ {0}" - using s5_1 succIsNat succNotZero by auto - have s5_3: "-.Succ[-.b] \ negNat \ {0}" - using s5_2 minus_nat_0_in_negnat_0 by auto - have s5_4: "-.Succ[-.b] \ Int" - using s5_3 neg_nat_in_int by auto - have s5_5: "-.Succ[-.b] \ Nat" - using s5_3 neg_nat_0_not_in_nat by auto - show ?thesis - unfolding iSucc_def - using s5_4 s5_5 by auto - qed - finally show "iSucc[a + -.Succ[-.b]] = - a + iSucc[-.Succ[-.b]]" - using s3_1 s3_2 by auto - qed - show "iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - using s3_3 s3_4 by blast - qed - show ?thesis - using s2_1 s2_2 - neg_nat_induction[of "\ b. - iSucc[a + b] = a + iSucc[b]"] - by auto - qed - have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" - using s1_1 s1_2 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem iSuccRightDistributesAdd_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "iSucc[a + b] = a + iSucc[b]" - using adom bdom iSuccRightDistributesAdd by auto - - -(* -THEOREM iPredRightDistributesAdd == - ASSUME NEW a \in Int - PROVE \A b \in Int: iPred[a + b] = a + iPred[b] -PROOF -<1>1. iPred[a + 0] = a + iPred[0] - <2>1. iPred[a + 0] = iPred[a + -.0] - BY neg0 - <2>2. @ = a + -.Succ[0] - <3>1. a \in Int - OBVIOUS - <3>2. 0 \in Nat - BY zeroIsNat - <3> QED - BY <3>1, <3>2, addint_succ_negnat - <2>3. @ = a + -.1 - OBVIOUS \* 1 is an abbreviation - <2>4. @ = a + iPred[0] - BY ipred_0 - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>2. \A b \in Nat: iPred[a + b] = a + iPred[b] - <2>1. ASSUME NEW b \in Nat, - iPred[a + b] = a + iPred[b] - PROVE iPred[a + Succ[b]] = a + iPred[Succ[b]] - <3>1. iPred[a + Succ[b]] = iPred[iSucc[a + b]] - <4>1. a \in Int - OBVIOUS - <4>2. b \in Nat - BY <2>1 - <4>3. a + Succ[b] = iSucc[a + b] - BY <4>1, <4>2, addint_succ_nat - <4> QED - BY <4>3 - <3>2. @ = iSucc[iPred[a + b]] - <4>1. a + b \in Int - <5>1. a \in Int - OBVIOUS - <5>2. b \in Int - BY <2>2, nat_in_int - <5> QED - BY <5>1, <5>2, addIsInt - <4> QED - BY <4>1, iSucciPredCommute - <3>3. @ = iSucc[a + iPred[b]] - BY <2>1 - <3>4. @ = a + iSucc[iPred[b]] - <4>1. b \in Int - BY <2>1, nat_in_int - <4>2. iPred[b] \in Int - BY <4>1, iPred_type - <4>3. a \in Int - OBVIOUS - <4> QED - BY <4>2, <4>3, iSuccRightDistributesAdd - <3>5. @ = a + iPred[iSucc[b]] - <4>1. b \in Int - BY <2>1, nat_in_int - <4> QED - BY <4>1, iSucciPredCommute - <3>6. @ = a + iPred[Succ[b]] - <4>1. b \in Nat - BY <2>1 - <4>2. b \in Int - BY <2>1, nat_in_int - <4>3. iSucc[b] = Succ[b] - BY <4>1, <4>2, nat_in_int DEF iSucc - <4> QED - BY <4>3 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 - <2> QED - BY <2>1, <1>1, NatInduction -<1>3. \A b \in negNat \ {0}: iPred[a + b] = a + iPred[b] - <2>1. ASSUME NEW b \in negNat \ {0} - PROVE iPred[a + b] = a + iPred[b] - <3>1. iPred[a + b] = iPred[a + -.-.b] - <4>1. b \in Int - BY <2>1, neg_nat_in_int - <4>2. -.-.b = b - BY <4>1, minus_involutive - <4> QED - BY <4>2 - <3>2. @ = a + -.Succ[-.b] - <4>1. a \in Int - OBVIOUS - <4>2. -.b \in Nat - BY <2>1, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, addint_succ_negnat - <3>3. @ = a + iPred[b] - <4>1. b \in Int - BY <2>1, neg_nat_in_int - <4>2. b \notin Nat \ {0} - BY <2>1, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3> QED - BY <3>1, <3>2, <3>3 - <2> QED - BY <2>1 -<1> QED - BY <1>2, <1>3 -*) -theorem iPredRightDistributesAdd: - "\ a \ Int: \ b \ Int: - iPred[a + b] = a + iPred[b]" - proof - - { - fix "a" :: "c" - assume adom: "a \ Int" - have s1_1: "iPred[a + 0] = a + iPred[0]" - proof - - have s2_1: "iPred[a + 0] = iPred[a + -.0]" - using neg0 by auto - also have s2_2: "... = a + -.Succ[0]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s3_1 s3_2 addint_succ_negnat by auto - qed - also have s2_3: "... = a + -.1" - by auto - have s2_4: "... = a + iPred[0]" - using ipred_0 by auto - (* `also` raises an error here *) - finally show ?thesis - using s2_4 by auto - qed - have s1_2: "\ b \ Nat: - iPred[a + b] = a + iPred[b]" - proof - - have s2_1: "\ b. \ - b \ Nat; - iPred[a + b] = a + iPred[b] - \ \ - iPred[a + Succ[b]] = a + iPred[Succ[b]]" - proof - - fix "b" :: "c" - assume s2_1_bdom: "b \ Nat" - assume s2_1_induct: "iPred[a + b] = a + iPred[b]" - have s3_1: "iPred[a + Succ[b]] = iPred[iSucc[a + b]]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "b \ Nat" - using s2_1_bdom by auto - have s4_3: "a + Succ[b] = iSucc[a + b]" - using s4_1 s4_2 addint_succ_nat by auto - show ?thesis - using s4_3 by auto - qed - also have s3_2: "... = iSucc[iPred[a + b]]" - proof - - have s4_1: "a + b \ Int" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "b \ Int" - using s2_1_bdom nat_in_int by auto - show ?thesis - using s5_1 s5_2 addIsInt by auto - qed - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - also have s3_3: "... = iSucc[a + iPred[b]]" - using s2_1_induct by auto - also have s3_4: "... = a + iSucc[iPred[b]]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom nat_in_int by auto - have s4_2: "iPred[b] \ Int" - using s4_1 iPred_type by auto - have s4_3: "a \ Int" - using adom by auto - show ?thesis - using s4_2 s4_3 iSuccRightDistributesAdd - by auto - qed - also have s3_5: "... = a + iPred[iSucc[b]]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom nat_in_int by auto - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - also have s3_6: "... = a + iPred[Succ[b]]" - proof - - have s4_1: "b \ Nat" - using s2_1_bdom by auto - have s4_2: "b \ Int" - using s2_1_bdom nat_in_int by auto - have s4_3: "iSucc[b] = Succ[b]" - unfolding iSucc_def - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . - qed - show ?thesis - using s1_1 s2_1 - natInduct[of "\ b. - iPred[a + b] = a + iPred[b]"] - by auto - qed - have s1_3: "\ b \ negNat \ {0}: - iPred[a + b] = a + iPred[b]" - proof - - { - fix "b" :: "c" - assume s2_1_bdom: "b \ negNat \ {0}" - { - have s3_1: "iPred[a + b] = iPred[a + -.-.b]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom neg_nat_in_int by auto - have s4_2: "-.-.b = b" - using s4_1 minus_involutive by auto - show ?thesis - using s4_2 by auto - qed - also have s3_2: "... = a + -.Succ[-.b]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "-.b \ Nat" - using s2_1_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 addint_succ_negnat by auto - qed - also have s3_3: "... = a + iPred[b]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom neg_nat_in_int by auto - have s4_2: "b \ Nat \ {0}" - proof - - have s5_1: "b \ negNat" - using s2_1_bdom by auto - show ?thesis - using s5_1 neg_nat_not_in_nat_setminus_0 - by auto - qed - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - finally have "iPred[a + b] = a + iPred[b]" . - } - from this have s2_2: "iPred[a + b] = a + iPred[b]" - by auto - } - from this show ?thesis - by auto - qed - have s1_4: "\ b \ Int: - iPred[a + b] = a + iPred[b]" - using s1_2 s1_3 int_union_nat_negnat_0 spec allI by auto - } - from this show ?thesis - by auto - qed - - -theorem iPredRightDistributesAdd_sequent: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "iPred[m + n] = m + iPred[n]" - using mdom ndom iPredRightDistributesAdd by auto - - -theorem isucc_eq_add_1: - assumes idom: "i \\in Int" - shows "iSucc[i] = i + 1" - proof - - have s1_1: "iSucc[i] = iSucc[i + 0]" - using idom add_0 by auto - also have s1_2: "... = i + iSucc[0]" - proof - - have s2_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - show ?thesis - using idom s2_1 iSuccRightDistributesAdd_sequent - by auto - qed - also have s1_3: "... = i + 1" - proof - - have s2_1: "iSucc[0] = Succ[0]" - proof - - have s3_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s3_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_2: "Succ[0] = 1" - by auto - have s2_3: "iSucc[0] = 1" - using s2_1 s2_2 by auto - show ?thesis - using s2_3 s1_2 by auto - qed - finally show ?thesis . - qed - - -(* -THEOREM AddAssociative == - \A a \in Int: \A b \in Int: \A c \in Int: - (a + b) + c = a + (b + c) -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE \A c \in Int: - (a + b) + c = a + (b + c) - BY <1>1 -<1>2. (a + b) + 0 = a + (b + 0) - <2>1. (a + b) + 0 = a + b - <3>1. a + b \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, add_0 - <2>2. @ = a + (b + 0) - <3>1. b \in Int - BY <1>1 - <3>2. b + 0 = b - BY <3>1, add_0 - <3> QED - BY <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW c \in Nat, - (a + b) + c = a + (b + c) - PROVE (a + b) + Succ[c] = a + (b + Succ[c]) - <2>1. (a + b) + Succ[c] = iSucc[(a + b) + c] - <3>1. a + b \in Int - BY <1>1, addIsInt - <3>2. c \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>2. @ = iSucc[a + (b + c)] - BY <1>3 \* induction hypothesis - <2>3. @ = a + iSucc[b + c] - <3>1. a \in Int - BY <1>1 - <3>2. b + c \in Int - BY <1>1, <1>3, nat_in_int, addIsInt - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>4. @ = a + (b + iSucc[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>3, nat_in_int - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>5. @ = a + (b + Succ[c]) - <3>1. c \in Nat - BY <1>3 - <3>2. c \in Int - BY <3>1, nat_in_int - <3>3. iSucc[c] = Succ[c] - BY <3>1, <3>2 DEF iSucc - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5 -<1>4. ASSUME NEW c \in negNat, - (a + b) + c = a + (b + c) - PROVE (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c]) - <2>1. (a + b) + -.Succ[-.c] = (a + b) + iPred[c] - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. c \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <3>3. iPred[c] = -.Succ[-.c] - BY <3>1, <3>2 DEF iPred - <3> QED - BY <3>3 - <2>2. @ = iPred[(a + b) + c] - <3>1. a + b \in Int - BY <1>1, addIsInt - <3>2. c \in Int - BY <1>4, neg_nat_in_int - <3>3 QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>3. @ = iPred[a + (b + c)] - BY <1>4 \* induction hypothesis - <2>4. @ = a + iPred[b + c] - <3>1. a \in Int - BY <1>1 - <3>2. b + c \in Int - BY <1>1, <1>4, neg_nat_in_int, addIsInt - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>5. @ = a + (b + iPred[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>4, neg_nat_in_int - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>6. @ = a + (b + -.Succ[-.c]) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. c \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <3>3. iPred[c] = -.Succ[-.c] - BY <3>1, <3>2 DEF iPred - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 -<1>5. \A c \in Nat: (a + b) + c = a + (b + c) - BY <1>2, <1>3, NatInduction -<1>6. \A c \in negNat: (a + b) + c = a + (b + c) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6 -*) -theorem AddAssociative: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - (a + b) + c = a + (b + c)" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \ Int" - assume bdom: "b \ Int" - have s1_2: "(a + b) + 0 = a + (b + 0)" - proof - - have s2_1: "(a + b) + 0 = a + b" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - show ?thesis - using s3_1 add_0 by auto - qed - have s2_2: "... = a + (b + 0)" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "b + 0 = b" - using s3_1 add_0 by auto - show ?thesis - using s3_2 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ c. \ - c \ Nat; - (a + b) + c = a + (b + c) - \ \ - (a + b) + Succ[c] = a + (b + Succ[c])" - proof - - fix "c" :: "c" - assume s1_3_cdom: "c \ Nat" - assume s1_3_induct: "(a + b) + c = a + (b + c)" - have s2_1: "(a + b) + Succ[c] = iSucc[(a + b) + c]" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - have s3_2: "c \ Nat" - using s1_3_cdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - also have s2_2: "... = iSucc[a + (b + c)]" - using s1_3_induct by auto - also have s2_3: "... = a + iSucc[b + c]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b + c \ Int" - using bdom s1_3_cdom nat_in_int addIsInt by auto - show ?thesis - using s3_1 s3_2 iSuccRightDistributesAdd by auto - qed - also have s2_4: "... = a + (b + iSucc[c])" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "c \ Int" - using s1_3_cdom nat_in_int by auto - have s3_3: "iSucc[b + c] = b + iSucc[c]" - using s3_1 s3_2 iSuccRightDistributesAdd by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = a + (b + Succ[c])" - proof - - have s3_1: "c \ Nat" - using s1_3_cdom by auto - have s3_2: "c \ Int" - using s1_3_cdom nat_in_int by auto - have s3_3: "iSucc[c] = Succ[c]" - unfolding iSucc_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a + b) + Succ[c] = a + (b + Succ[c])" . - qed - have s1_4: "\ c. \ - c \ negNat; - (a + b) + c = a + (b + c) - \ \ - (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" - proof - - fix "c" :: "c" - assume s1_4_cdom: "c \ negNat" - assume s1_4_induct: "(a + b) + c = a + (b + c)" - have s2_1: "(a + b) + -.Succ[-.c] = (a + b) + iPred[c]" - proof - - have s3_1: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "c \ Nat \ {0}" - using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto - have s3_3: "iPred[c] = -.Succ[-.c]" - unfolding iPred_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - also have s2_2: "... = iPred[(a + b) + c]" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - have s3_2: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by force - qed - also have s2_3: "... = iPred[a + (b + c)]" - using s1_4_induct by auto - also have s2_4: "... = a + iPred[b + c]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b + c \ Int" - using bdom s1_4_cdom neg_nat_in_int addIsInt by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by auto - qed - also have s2_5: "... = a + (b + iPred[c])" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_3: "iPred[b + c] = b + iPred[c]" - using s3_1 s3_2 iPredRightDistributesAdd by force - show ?thesis - using s3_3 by auto - qed - also have s2_6: "... = a + (b + -.Succ[-.c])" - proof - - have s3_1: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "c \ Nat \ {0}" - using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto - have s3_3: "iPred[c] = -.Succ[-.c]" - unfolding iPred_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" . - qed - have s1_5: "\ c \ Nat: - (a + b) + c = a + (b + c)" - using s1_2 s1_3 natInduct[of "\ c. - (a + b) + c = a + (b + c)"] - by auto - have s1_6: "\ c \ negNat: - (a + b) + c = a + (b + c)" - using s1_2 s1_4 neg_nat_induction[of "\ c. - (a + b) + c = a + (b + c)"] - by auto - have "\ c \ Int: - (a + b) + c = a + (b + c)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem AddAssociative_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - and cdom: "c \\in Int" - shows "(a + b) + c = a + (b + c)" - using adom bdom cdom AddAssociative by auto - - -theorem AddLeftCommutativity: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - and cdom: "c \\in Int" - shows "b + (a + c) = a + (b + c)" - proof - - have s1_1: "(b + a) + c = (a + b) + c" - using adom bdom by (simp only: AddCommutative_sequent) - show ?thesis - using adom bdom cdom s1_1 by (simp only: AddAssociative_sequent) - qed - - -lemmas add_ac_int[algebra_simps] = - AddCommutative_sequent AddAssociative_sequent - AddLeftCommutativity add_0 add_0_left - - -(* A test that the simplification rules work. *) -theorem - assumes "a \\in Int" and "b \\in Int" and "c \\in Int" - shows "a + (b + c) = (a + b) + c" - using assms by (simp add: algebra_simps) - - -(* -THEOREM MinusSuccIsMinusOne == - ASSUME NEW n \in Nat - PROVE -.Succ[n] = -.n + -.1 -<1>1. -.Succ[n] = -.Succ[-.-.n] - <2>1. -.-.n = n - <3>1. n \in Nat - OBVIOUS - <3>2. n \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>2, minus_involutive - <2> QED - BY <2>1 -<1>2. @ = iPred[-.n] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. -.n \notin Nat \ {0} - <3>1. n \in Nat - OBVIOUS - <3>2. -.n \in negNat - BY <3>1, minus_nat_in_neg_nat - <3> QED - BY <3>2, neg_nat_not_in_nat_setminus_0 - <2> QED - BY <2>1, <2>2 DEF iPred -<1>3. @ = iPred[-.n + 0] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. -.n + 0 = -.n - BY <2>1, add_0 - <2> QED - BY <2>2 -<1>4. @ = -.n + iPred[0] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. 0 \in Int - BY zero_in_int - <2> QED - BY <2>1, <2>2, iPredRightDistributesAdd -<1>5. @ = -.n + -.Succ[-.0] - <2>1. 0 \in Int - BY zero_in_int - <2>2. 0 \notin Nat \ {0} - OBVIOUS - <2> QED - BY <2>1, <2>2 DEF iPred -<1>6. @ = -.n + -.Succ[0] - BY neg0 -<1>7. @ = -.n + -.1 - OBVIOUS \* 1 is an abbreviation -<1> QED - BY <1>1, <1>2, <1>3, <1>4, - <1>5, <1>6, <1>7 -*) -theorem MinusSuccIsMinusOne: - assumes ndom: "n \ Nat" - shows "-.Succ[n] = -.n + -.1" - proof - - have s1_1: "-.Succ[n] = -.Succ[-.-.n]" - proof - - have s2_1: "-.-.n = n" - proof - - have s3_1: "n \ Int" - using ndom nat_in_int by auto - show ?thesis - using s3_1 minus_involutive by auto - qed - show ?thesis - using s2_1 by auto - qed - also have s1_2: "... = iPred[-.n]" - proof - - have s2_1: "-.n \ Int" - using ndom minus_nat_in_int by auto - have s2_2: "-.n \ Nat \ {0}" - proof - - have s3_1: "-.n \ negNat" - using ndom minus_nat_in_neg_nat by auto - show ?thesis - using s3_1 neg_nat_not_in_nat_setminus_0 by fast - qed - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_3: "... = iPred[-.n + 0]" - proof - - have s2_1: "-.n \ Int" - proof - - have s3_1: "n \ Nat" - using ndom by auto - show ?thesis - using s3_1 minus_nat_in_int by auto - qed - have s2_2: "-.n + 0 = -.n" - using s2_1 add_0 by auto - show ?thesis - using s2_2 by auto - qed - also have s1_4: "... = -.n + iPred[0]" - proof - - have s2_1: "-.n \ Int" - proof - - have s3_1: "n \ Nat" - using ndom by auto - show ?thesis - using s3_1 minus_nat_in_int by auto - qed - have s2_2: "0 \ Int" - using zero_in_int by auto - show ?thesis - using s2_1 s2_2 iPredRightDistributesAdd by fast - qed - also have s1_5: "... = -.n + -.Succ[-.0]" - proof - - have s2_1: "0 \ Int" - using zero_in_int by auto - have s2_2: "0 \ Nat \ {0}" - by auto - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_6: "... = -.n + -.Succ[0]" - using neg0 by auto - also have s1_7: "... = -.n + -.1" - by auto - from calculation show ?thesis - by auto - qed - - -theorem SuccMinusIsPlusOne: - assumes idom: "i \\in negNat" - shows "Succ[-.i] = -.i + 1" - proof - - have s1_1: "-.i \\in Nat" - using idom minus_neg_nat_in_nat by auto - have s1_2: "-.i \\in Int" - using s1_1 nat_in_int by auto - have s1_3: "Succ[-.i] = iSucc[-.i]" - unfolding iSucc_def - using s1_1 s1_2 by auto - also have s1_4: "... = iSucc[-.i + 0]" - using s1_2 add_0 by auto - also have s1_5: "... = -.i + Succ[0]" - using s1_2 zeroIsNat addint_succ_nat by auto - also have s1_6: "... = -.i + 1" - by auto (* 1 is an abbreviation *) - from calculation show ?thesis . - qed - - -(*----------------------------------------------------------------------------*) -(* Properties of addition and difference. *) - - -(* -THEOREM AddNegCancels == - \A i \in Int: i + -.i = 0 -PROOF -<1>1. 0 + -.0 = 0 - <2>1. 0 + -.0 = 0 + 0 - BY neg0 - <2>2. @ = 0 - BY zero_in_int, add_0 - <2> QED - BY <2>1, <2>2 -<1>2. ASSUME NEW i \in Nat, i + -.i = 0 - PROVE Succ[i] + -.Succ[i] = 0 - <2>1. Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i] - <3>1. Succ[i] \in Int - BY <1>2, succIsNat, nat_in_int - <3>2. -.Succ[i] \in Int - BY <3>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, AddCommutative - <2>2. @ = iSucc[-.Succ[i] + i] - <3>1. -.Succ[i] \in Int - BY <1>2, succIsNat, minus_nat_in_int - <3>2. i \in Nat - BY <1>2 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>3. @ = iSucc[i + -.Succ[i]] - <3>1. i \in Int - BY <1>2, nat_in_int - <3>2. -.Succ[i] \in Int - BY <1>2, succIsNat, minus_nat_in_int - <3>3. -.Succ[i] + i = i + -.Succ[i] - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>4. @ = iSucc[iPred[i + -.i]] - <3>1. i \in Int - BY <1>2, nat_in_int - <3>2. i \in Nat - BY <1>2 - <3> QED - BY <3>1, <3>2, addint_succ_negnat - <2>5. @ = iSucc[iPred[0]] - BY <1>2 - <2>6. @ = iSucc[-.1] - BY ipred_0 - <2>7. @ = 0 - BY isucc_minus_1 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6, <2>7 -<1>3. ASSUME NEW i \in negNat, i + -.i = 0 - PROVE -.Succ[-.i] + -.(-.Succ[-.i]) = 0 - <2>1. -.Succ[-.i] + -.(-.Succ[-.i]) - = -.Succ[-.i] + Succ[-.i] - <3>1. Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, nat_in_int - <3> QED - BY <3>1, minus_involutive - <2>2. @ = iSucc[-.Succ[-.i] + -.i] - <3>1. -.Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, minus_nat_in_int - <3>2. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>3. @ = iSucc[-.i + -.Succ[-.i]] - <3>1. -.Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, minus_nat_in_int - <3>2. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>3. -.Succ[-.i] + -.i = -.i + -.Succ[-.i] - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>4. @ = iSucc[iPred[-.i + -.-.i]] - <3>1. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>2. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3>3. -.i + -.Succ[-.i] = iPred[-.i + -.-.i] - BY <3>1, <3>2, addint_succ_negnat - <3> QED - BY <3>3 - <2>5. @ = iSucc[iPred[-.i + i]] - <3>1. -.-.i = i - BY <1>3, neg_nat_in_int, minus_involutive - <3> QED - BY <3>1 - <2>6. @ = iSucc[iPred[i + -.i]] - <3>1. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>2. i \in Int - BY <1>3, neg_nat_in_int - <3>3. -.i + i = i + -.i - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>7. @ = iSucc[iPred[0]] - BY <1>3 \* induction hypothesis - <2>8. @ = iSucc[-.1] - BY ipred_0 - <2>9. @ = 0 - BY isucc_minus_1 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, - <2>6, <2>7, <2>8, <2>9 -<1>4. \A i \in Nat: i + -.i = 0 - BY <1>1, <1>2, NatInduction -<1>5. \A i \in negNat: i + -.i = 0 - BY <1>1, <1>3, neg_nat_induction -<1> QED - BY <1>4, <1>5, int_union_nat_negnat -*) -theorem AddNegCancels: - "\ i \ Int: i + -.i = 0" - proof - - have s1_1: "0 + -.0 = 0" - proof - - have s2_1: "0 + -.0 = 0 + 0" - using neg0 by auto - also have s2_2: "... = 0" - using zero_in_int add_0 by auto - finally show ?thesis . - qed - have s1_2: "\ i. \ - i \\in Nat; - i + -.i = 0 - \ \ - Succ[i] + -.Succ[i] = 0" - proof - - fix "i" :: "c" - assume s1_2_idom: "i \\in Nat" - assume s1_2_induct: "i + -.i = 0" - have s2_1: "Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i]" - proof - - have s3_1: "Succ[i] \\in Int" - using s1_2_idom succIsNat nat_in_int by auto - have s3_2: "-.Succ[i] \\in Int" - using s3_1 neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 AddCommutative by fast - qed - also have s2_2: "... = iSucc[-.Succ[i] + i]" - proof - - have s3_1: "-.Succ[i] \\in Int" - using s1_2_idom succIsNat minus_nat_in_int by auto - have s3_2: "i \\in Nat" - using s1_2_idom by auto - show ?thesis - using s3_1 s3_2 by (rule addint_succ_nat) - qed - also have s2_3: "... = iSucc[i + -.Succ[i]]" - proof - - have s3_1: "i \\in Int" - using s1_2_idom nat_in_int by auto - have s3_2: "-.Succ[i] \\in Int" - using s1_2_idom succIsNat minus_nat_in_int - by auto - have s3_3: "-.Succ[i] + i = i + -.Succ[i]" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = iSucc[iPred[i + -.i]]" - proof - - have s3_1: "i \\in Int" - using s1_2_idom nat_in_int by auto - have s3_2: "i \\in Nat" - using s1_2_idom by auto - show ?thesis - using s3_1 s3_2 addint_succ_negnat by auto - qed - also have s2_5: "... = iSucc[iPred[0]]" - using s1_2_induct by auto - also have s2_6: "... = iSucc[-.1]" - using ipred_0 by auto - also have s2_7: "... = 0" - using isucc_minus_1 by auto - finally show "Succ[i] + -.Succ[i] = 0" . - qed - have s1_3: "\ i. \ - i \\in negNat; - i + -.i = 0 - \ \ - -.Succ[-.i] + -.-.Succ[-.i] = 0" - proof - - fix "i" :: "c" - assume s1_3_idom: "i \\in negNat" - assume s1_3_induct: "i + -.i = 0" - have s2_1: "-.Succ[-.i] + -.(-.Succ[-.i]) - = -.Succ[-.i] + Succ[-.i]" - proof - - have s3_1: "Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat nat_in_int by auto - show ?thesis - using s3_1 minus_involutive by auto - qed - also have s2_2: "... = iSucc[-.Succ[-.i] + -.i]" - proof - - have s3_1: "-.Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat minus_nat_in_int by auto - have s3_2: "-.i \\in Nat" - using s1_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 by (rule addint_succ_nat) - qed - also have s2_3: "... = iSucc[-.i + -.Succ[-.i]]" - proof - - have s3_1: "-.Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat minus_nat_in_int by auto - have s3_2: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_3: "-.Succ[-.i] + -.i = -.i + -.Succ[-.i]" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = iSucc[iPred[-.i + -.-.i]]" - proof - - have s3_1: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_2: "-.i \\in Nat" - using s1_3_idom minus_neg_nat_in_nat by auto - have s3_3: "-.i + -.Succ[-.i] = iPred[-.i + -.-.i]" - using s3_1 s3_2 addint_succ_negnat by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = iSucc[iPred[-.i + i]]" - proof - - have s3_1: "-.-.i = i" - using s1_3_idom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s3_1 by auto - qed - also have s2_6: "... = iSucc[iPred[i + -.i]]" - proof - - have s3_1: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_2: "i \\in Int" - using s1_3_idom neg_nat_in_int by auto - have s3_3: "-.i + i = i + -.i" - using s3_1 s3_2 AddCommutative by fast - show ?thesis - using s3_3 by auto - qed - also have s2_7: "... = iSucc[iPred[0]]" - using s1_3_induct by auto - also have s2_8: "... = iSucc[-.1]" - using ipred_0 by auto - also have s2_9: "... = 0" - using isucc_minus_1 by auto - finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . - qed - have s1_4: "\ i \ Nat: i + -.i = 0" - using s1_1 s1_2 natInduct[of "\ i. - i + -.i = 0"] - by auto - have s1_5: "\ i \ negNat: i + -.i = 0" - using s1_1 s1_3 neg_nat_induction[of "\ i. - i + -.i = 0"] - by auto - show ?thesis - using s1_4 s1_5 int_union_nat_negnat by auto - qed - - -theorem AddNegCancels_sequent: - assumes idom: "i \\in Int" - shows "i + -.i = 0" - using idom AddNegCancels spec by auto - - -theorem AddNegCancels_left: - assumes idom: "i \\in Int" - shows "-.i + i = 0" - proof - - have s1_1: "i + -.i = 0" - using idom AddNegCancels_sequent by auto - have s1_2: "i + -.i = -.i + i" - proof - - have s2_1: "-.i \\in Int" - using idom neg_int_eq_int by auto - show ?thesis - using idom s2_1 AddCommutative_sequent by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM a_plus_b_minus_a_eq_b == - ASSUME NEW a \in Int, NEW b \in Int - PROVE a + (b - a) = b -PROOF -<1>1. a + (b - a) = a + (b + -.a) - BY DEF diff -<1>2. @ = a + (-.a + b) - <2>1. b \in Int - OBVIOUS - <2>2. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>3. b + -.a = -.a + b - BY <2>1, <2>2, AddCommutative - <2> QED - BY <2>3 -<1>3. @ = (a + -.a) + b - <2>1. a \in Int - OBVIOUS - <2>2. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>3. b \in Int - OBVIOUS - <2> QED - BY <2>1, <2>2, <2>3, AddAssociative -<1>4. @ = 0 + b - <2>1. a \in Int - OBVIOUS - <2>2. a + -.a = 0 - BY <2>1, AddNegCancels - <2> QED - BY <2>2 -<1>5. @ = b - BY add_0_left -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5 -*) -theorem a_plus_b_minus_a_eq_b: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "a + (b - a) = b" - proof - - have s1_1: "a + (b - a) = a + (b + -.a)" - unfolding diff_def by auto - also have s1_2: "... = a + (-.a + b)" - proof - - have s2_1: "b \\in Int" - using bdom by auto - have s2_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_3: "b + -.a = -.a + b" - using s2_1 s2_2 AddCommutative by auto - show ?thesis - using s2_3 by auto - qed - also have s1_3: "... = (a + -.a) + b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_3: "b \\in Int" - using bdom by auto - have s2_4: "a \\in Int ==> -.a \\in Int ==> b \\in Int - ==> (a + -.a) + b = a + (-.a + b)" - using AddAssociative_sequent by blast - have s2_5: "(a + -.a) + b = a + (-.a + b)" - using s2_1 s2_2 s2_3 s2_4 by auto - show ?thesis - by (rule s2_5[symmetric]) - qed - also have s1_4: "... = 0 + b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "a + -.a = 0" - using s2_1 AddNegCancels by auto - show ?thesis - using s2_2 by auto - qed - also have s1_5: "... = b" - using bdom add_0_left by auto - finally show ?thesis . - qed - - -(* -THEOREM MinusDistributesAdd == - \A a \in Int: \A b \in Int: - -.(a + b) = (-.a) + (-.b) -PROOF -<1>1. SUFFICES ASSUME NEW a \in Int - PROVE \A b \in Int: - -.(a + b) = (-.a) + (-.b) - BY <1>1 -<1>2. -.(a + 0) = (-.a) + (-.0) - <2>1. -.(a + 0) = -.a - <3>1. a + 0 = a - BY <1>1, add_0 - <3> QED - BY <3>1 - <2>2. -.a + -.0 = -.a - <3>1. -.a + -.0 = -.a + 0 - BY neg0 - <3>2. @ = -.a - <4>1. -.a \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, add_0 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW b \in Nat, - -.(a + b) = (-.a) + (-.b) - PROVE - -.(a + Succ[b]) = (-.a) + (-.Succ[b]) - <2>1. -.(a + Succ[b]) = -.iSucc[a + b] - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3>3. a + Succ[b] = iSucc[a + b] - BY <3>1, <3>2, addint_succ_nat - <3> QED - BY <3>3 - <2>2. @ = -. iSucc[-.((-.a) + (-.b))] - <3>1. a + b = -.(-.a + -.b) - <4>1. -.-.(a + b) = -.(-.a + -.b) - BY <1>3 - <4>2. -.-.(a + b) = a + b - <5>1. a + b \in Int - <6>1. a \in Int - BY <1>1 - <6>2. b \in Int - BY <1>3, nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, minus_involutive - <4> QED - BY <4>1, <4>2 - <3> QED - BY <3>1 - <2>3. @ = iPred[-.a + -.b] - <3>1. -.a + -.b \in Int - <4>1. -.a \in Int - BY <1>1, neg_int_eq_int - <4>2. -.b \in Int - BY <1>3, minus_nat_in_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, iSuccMinusFlip - <2>4. @ = -.a + iPred[-.b] - <3>1. -.a \in Int - BY <1>1, neg_int_eq_int - <3>2. -.b \in Int - BY <1>3, minus_nat_in_int - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>5. @ = -.a + -.Succ[b] - <3>1. iPred[-.b] = -.Succ[-.-.b] - <4>1. -.b \in Int - BY <1>3, minus_nat_in_int - <4>2. -.b \notin Nat \ {0} - <5>1. -.b \in negNat - BY <1>3, minus_nat_in_neg_nat - <5> QED - BY <5>1, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3>2. @ = -.Succ[b] - <4>1. -.-.b = b - BY <1>3, nat_in_int, minus_involutive - <4> QED - BY <4>1 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5 -<1>4. ASSUME NEW b \in negNat, - -.(a + b) = (-.a) + (-.b) - PROVE - -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b]) - <2>1. -.(a + -.Succ[-.b]) = -. iPred[a + b] - <3>1. a + -.Succ[-.b] = iPred[a + -.-.b] - <4>1. a \in Int - BY <1>1 - <4>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, addint_succ_negnat - <3>2. a + -.Succ[-.b] = iPred[a + b] - <4>1. -.-.b = b - BY <1>4, neg_nat_in_int, minus_involutive - <4> QED - BY <3>1, <4>1 - <3> QED - BY <3>2 - <2>2. @ = -. iPred[-.((-.a) + (-.b))] - <3>1. a + b = -.(-.a + -.b) - <4>1. -.(a + b) = (-.a) + (-.b) - BY <1>4 - <4>2. -.-.(a + b) = -.(-.a + -.b) - BY <4>1 - <4>3. -.-.(a + b) = a + b - <5>1. a + b \in Int - <6>1. a \in Int - BY <1>1 - <6>2. b \in Int - BY <1>4, neg_nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, minus_involutive - <4> QED - BY <4>2, <4>3 - <3> QED - BY <3>1 - <2>3. @ = iSucc[-.a + -.b] - <3>1. -.a + -.b \in Int - <4>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <4>2. -.b \in Int - BY <1>4, minus_negnat_in_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, iPredMinusFlip - <2>4. @ = -.a + iSucc[-.b] - <3>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <3>2. -.b \in Int - BY <1>4, minus_negnat_in_int - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>5. @ = -.a + Succ[-.b] - <3>1. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>2. -.b \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>1, <3>2 DEF iSucc - <2>6. @ = -.a + (-.-.Succ[-.b]) - <3>1. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>2. Succ[-.b] \in Nat - BY <3>1, succIsNat - <3>3. Succ[-.b] \in Int - BY <3>2, nat_in_int - <3>4. -.-.Succ[-.b] = Succ[-.b] - BY <3>3, minus_involutive - <3> QED - BY <3>4 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 -<1>5. \A b \in Nat: -.(a + b) = (-.a) + (-.b) - BY <1>2, <1>3, NatInduction -<1>6. \A b \in negNat: -.(a + b) = (-.a) + (-.b) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MinusDistributesAdd: - "\ a \ Int: \ b \ Int: - -.(a + b) = (-.a) + (-.b)" - proof - - { - fix "a" :: "c" - assume adom: "a \\in Int" - have s1_2: "-.(a + 0) = (-.a) + (-.0)" - proof - - have s2_1: "-.(a + 0) = -.a" - proof - - have s3_1: "a + 0 = a" - using adom add_0 by auto - show ?thesis - using s3_1 by auto - qed - have s2_2: "-.a + -.0 = -.a" - proof - - have s3_1: "-.a + -.0 = -.a + 0" - using neg0 by auto - also have s3_2: "... = -.a" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - show ?thesis - using s4_1 add_0 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ b. \ - b \\in Nat; - -.(a + b) = -.a + -.b - \ \ - -.(a + Succ[b]) = (-.a) + (-.Succ[b])" - proof - - fix "b" :: "c" - assume s1_3_bdom: "b \\in Nat" - assume s1_3_induct: "-.(a + b) = -.a + -.b" - have s2_1: "-.(a + Succ[b]) = -.iSucc[a + b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - have s3_3: "a + Succ[b] = iSucc[a + b]" - using s3_1 s3_2 addint_succ_nat by auto - show ?thesis - using s3_3 by auto - qed - also have s2_2: "... = -. iSucc[-.((-.a) + (-.b))]" - proof - - have s3_1: "a + b = -.(-.a + -.b)" - proof - - have s4_1: "-.-.(a + b) = -.(-.a + -.b)" - using s1_3_induct by auto - have s4_2: "-.-.(a + b) = a + b" - proof - - have s5_1: "a + b \\in Int" - proof - - have s6_1: "a \\in Int" - using adom by auto - have s6_2: "b \\in Int" - using s1_3_bdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 minus_involutive by auto - qed - show ?thesis - using s4_1 s4_2 by auto - qed - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = iPred[-.a + -.b]" - proof - - have s3_1: "-.a + -.b \\in Int" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s4_2: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 iSuccMinusFlip by auto - qed - also have s2_4: "... = -.a + iPred[-.b]" - proof - - have s3_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s3_2: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by auto - qed - also have s2_5: "... = -.a + -.Succ[b]" - proof - - have s3_1: "iPred[-.b] = -.Succ[-.-.b]" - proof - - have s4_1: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - have s4_2: "-.b \ Nat \ {0}" - proof - - have s5_1: "-.b \\in negNat" - using s1_3_bdom minus_nat_in_neg_nat by auto - show ?thesis - using s5_1 neg_nat_not_in_nat_setminus_0 by fast - qed - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - also have s3_2: "... = -.Succ[b]" - proof - - have s4_1: "-.-.b = b" - using s1_3_bdom nat_in_int minus_involutive by auto - show ?thesis - using s4_1 by auto - qed - finally have s3_3: "iPred[-.b] = -.Succ[b]" . - show ?thesis - using s3_3 by auto - qed - finally show "-.(a + Succ[b]) = (-.a) + (-.Succ[b])" . - qed - have s1_4: "\ b. \ - b \\in negNat; - -.(a + b) = -.a + -.b - \ \ - -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" - proof - - fix "b" :: "c" - assume s1_4_bdom: "b \\in negNat" - assume s1_4_induct: "-.(a + b) = -.a + -.b" - have s2_1: "-.(a + -.Succ[-.b]) = -. iPred[a + b]" - proof - - have s3_1: "a + -.Succ[-.b] = iPred[a + -.-.b]" - proof - - have s4_1: "a \\in Int" - using adom by auto - have s4_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 addint_succ_negnat by auto - qed - have s3_2: "a + -.Succ[-.b] = iPred[a + b]" - proof - - have s4_1: "-.-.b = b" - using s1_4_bdom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s3_1 s4_1 by auto - qed - show ?thesis - using s3_2 by auto - qed - also have s2_2: "... = -. iPred[-.((-.a) + (-.b))]" - proof - - have s3_1: "a + b = -.(-.a + -.b)" - proof - - have s4_1: "-.(a + b) = (-.a) + (-.b)" - using s1_4_induct by auto - have s4_2: "-.-.(a + b) = -.(-.a + -.b)" - using s4_1 by auto - have s4_3: "-.-.(a + b) = a + b" - proof - - have s5_1: "a + b \\in Int" - proof - - have s6_1: "a \\in Int" - using adom by auto - have s6_2: "b \\in Int" - using s1_4_bdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 minus_involutive by auto - qed - show ?thesis - using s4_2 s4_3 by auto - qed - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = iSucc[-.a + -.b]" - proof - - have s3_1: "-.a + -.b \\in Int" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s4_2: "-.b \\in Int" - using s1_4_bdom minus_negnat_in_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 iPredMinusFlip by auto - qed - also have s2_4: "... = -.a + iSucc[-.b]" - proof - - have s3_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s3_2: "-.b \\in Int" - using s1_4_bdom minus_negnat_in_int by auto - show ?thesis - using s3_1 s3_2 iSuccRightDistributesAdd by auto - qed - also have s2_5: "... = -.a + Succ[-.b]" - proof - - have s3_1: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - have s3_2: "-.b \\in Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - also have s2_6: "... = -.a + (-.-.Succ[-.b])" - proof - - have s3_1: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - have s3_2: "Succ[-.b] \\in Nat" - using s3_1 succIsNat by auto - have s3_3: "Succ[-.b] \\in Int" - using s3_2 nat_in_int by auto - have s3_4: "-.-.Succ[-.b] = Succ[-.b]" - using s3_3 minus_involutive by auto - show ?thesis - using s3_4 by auto - qed - finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . - qed - have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_3 natInduct[of "\ b. - -.(a + b) = (-.a) + (-.b)"] - by auto - have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_4 neg_nat_induction[of "\ b. - -.(a + b) = (-.a) + (-.b)"] - by auto - have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem MinusDistributesAdd_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "-.(a + b) = (-.a) + (-.b)" - using adom bdom MinusDistributesAdd by auto - - -(* -THEOREM MinusDiff == - ASSUME NEW a \in Int, NEW b \in Int - PROVE -.(a - b) = b - a -PROOF -<1>1. -.(a - b) = -.(a + -.b) - BY DEF diff -<1>2. @ = (-.a) + (-.-.b) - <2>1. a \in Int - OBVIOUS - <2>2. -.b \in Int - BY b \in Int, neg_int_eq_int - <2> QED - BY <2>1, <2>2, MinusDistributesAdd -<1>3. @ = (-.a) + b - BY b \in Int, minus_involutive -<1>4. @ = b + (-.a) - <2>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>2. b \in Int - OBVIOUS - <2> QED - BY <2>1, <2>2, AddCommutative -<1>5. @ = b - a - BY DEF diff -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5 -*) -theorem MinusDiff: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "-.(a - b) = b - a" - proof - - have s1_1: "-.(a - b) = -.(a + -.b)" - unfolding diff_def by auto - also have s1_2: "... = -.a + -.-.b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - show ?thesis - using s2_1 s2_2 MinusDistributesAdd by auto - qed - also have s1_3: "... = -.a + b" - using bdom minus_involutive by auto - also have s1_4: "... = b + -.a" - proof - - have s2_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_2: "b \\in Int" - using bdom by auto - show ?thesis - using s2_1 s2_2 AddCommutative by auto - qed - also have s1_5: "... = b - a" - unfolding diff_def by auto - finally show ?thesis . - qed - - -(*----------------------------------------------------------------------------*) -(* Commutativity of multiplication. *) - -(* -THEOREM MultCommutative == - \A i, j \in Int: i * j = j * i -PROOF -<1>1. \A i \in Int: i * 0 = 0 * i - <2>1. 0 * 0 = 0 * 0 - OBVIOUS - <2>2. ASSUME NEW i \in Nat, i * 0 = 0 * i - PROVE Succ[i] * 0 = 0 * Succ[i] - <3>1. 0 * Succ[i] = 0 * i + 0 - <4>1. 0 \in Int - BY zero_in_int - <4>2. i \in Nat - BY <2>2 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3>2. @ = 0 * i - <4>1. 0 * i \in Int - <5>1 0 \in Int - BY zero_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5> QED - BY <5>1, <5>2, multIsInt - <4> QED - BY <4>1, add_0 - <3>3. @ = i * 0 - BY <2>2 - <3>4. @ = 0 - <4>1. i \in Int - BY <2>2, nat_in_int - <4> QED - BY <4>1, mult_0 - <3>5. @ = Succ[i] * 0 - <4>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5 - <2>3. ASSUME NEW i \in negNat, i * 0 = 0 * i - PROVE -.Succ[-.i] * 0 = 0 * -.Succ[-.i] - <3>1. -.Succ[-.i] * 0 = 0 - <4>1. -.Succ[-.i] \in Int - BY <2>3, minus_succ_minus_negnat_in_int - <4> QED - BY <4>1, mult_0 - <3>2. 0 * -.Succ[-.i] = 0 - <4>1. 0 * -.Succ[-.i] = 0 * -.-.i + -.0 - <5>1. 0 \in Int - BY zero_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = 0 * i + -.0 - <5>1. -.-.i = i - BY <2>3, neg_nat_in_int, minus_involutive - <5> QED - BY <5>1 - <4>3. @ = 0 * i + 0 - BY neg0 - <4>4. @ = 0 * i - <5>1. 0 * i \in Int - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Int - BY <2>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5> QED - BY <5>1, add_0 - <4>5. @ = i * 0 - BY <2>3 \* induction hypothesis - <4>6. @ = 0 - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5> QED - BY <5>1, mult_0 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, - <4>5, <4>6 - <3> QED - BY <3>1, <3>2 - <2>4. \A i \in Nat: i * 0 = 0 * i - BY <2>1, <2>2, NatInduction - <2>5. \A i \in negNat: i * 0 = 0 * i - BY <2>1, <2>3, neg_nat_induction - <2> QED - BY <2>4, <2>5, int_union_nat_negnat -<1>2. ASSUME NEW j \in Nat, - \A i \in Int: i * j = j * i - PROVE \A i \in Int: i * Succ[j] = Succ[j] * i - <2>1. 0 * Succ[j] = Succ[j] * 0 - <3>1. 0 * Succ[j] = 0 - <4>1. 0 * Succ[j] = 0 * j + 0 - <5>1. 0 \in Int - BY zero_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = 0 * j - <5>1. 0 * j \in Int - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Int - BY <1>2, nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5> QED - BY <5>1, add_0 - <4>3. @ = j * 0 - BY <1>2, zero_in_int, spec - <4>4. @ = 0 - <5>1. j \in Int - BY <1>2, nat_in_int - <5> QED - BY <5>1, mult_0 - <4> QED - BY <4>1, <4>2, <4>3, <4>4 - <3>2. Succ[j] * 0 = 0 - <4>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2 - <2>2. ASSUME NEW i \in Nat, - i * Succ[j] = Succ[j] * i - PROVE Succ[i] * Succ[j] = Succ[j] * Succ[i] - <3>1. Succ[i] * Succ[j] = i * j + ((i + j) + 1) - <4>1. Succ[i] * Succ[j] = Succ[i] * j + Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = j * Succ[i] + Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. Succ[i] * j = j * Succ[i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>3. @ = (j * i + j) + Succ[i] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5>3. j * Succ[i] = j * i + j - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>4. @ = (i * j + j) + Succ[i] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>5. @ = (i * j + j) + (i + 1) - <5>1. i \in Nat - BY <2>2 - <5>2. Succ[i] = i + 1 - BY <5>1, nat_add_1 - <5> QED - BY <5>2 - <4>6. @ = i * j + (j + (i + 1)) - <5>1. i * j \in Int - BY <1>2, <2>2, nat_in_int, multIsInt - <5>2. j \in Int - BY <1>2, nat_in_int - <5>3. i + 1 \in Int - BY <2>2, oneIsNat, nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + ((j + i) + 1) - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. j + (i + 1) = (j + i) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>8. @ = i * j + ((i + j) + 1) - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. j + i = i + j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8 - <3>2. Succ[j] * Succ[i] = i * j + ((i + j) + 1) - <4>1. Succ[j] * Succ[i] = Succ[j] * i + Succ[j] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = i * Succ[j] + Succ[j] - BY <2>2 \* induction hypothesis - <4>3. @ = (i * j + i) + Succ[j] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5>3. i * Succ[j] = i * j + i - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>4. @ = (i * j + i) + (j + 1) - <5>1. j \in Nat - BY <1>2 - <5>2. Succ[j] = j + 1 - BY <5>1, nat_add_1 - <5> QED - BY <5>2 - <4>5. @ = i * j + (i + (j + 1)) - <5>1. i * j \in Int - BY <1>2, <2>2, nat_in_int, multIsInt - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. j + 1 \in Int - BY <1>2, oneIsNat, nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>6. @ = i * j + ((i + j) + 1) - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Int - BY <1>2, nat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. i + (j + 1) = (i + j) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, <4>6 - <3> QED - BY <3>1, <3>2 - <2>3. ASSUME NEW i \in negNat, - i * Succ[j] = Succ[j] * i - PROVE -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i] - <3>1. -.Succ[-.i] * Succ[j] = - i * j + ((i + -.j) + -.1) - <4>1. -.Succ[-.i] * Succ[j] = - -.Succ[-.i] * j + -.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = j * -.Succ[-.i] + -.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>3. @ = (j * -.-.i + -.j) + -.Succ[-.i] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>3. j * -.Succ[-.i] = j * -.-.i + -.j - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>4. @ = (j * i + -.j) + -.Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>5. @ = (i * j + -.j) + -.Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>6. @ = (i * j + -.j) + (i + -.1) - <5>1. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>2. -.Succ[-.i] = -.-.i + -.1 - BY <5>1, MinusSuccIsMinusOne - <5>3. -.-.i = i - BY <2>3, neg_nat_in_int, minus_involutive - <5>4. -.Succ[-.i] = i + -.1 - BY <5>2, <5>3 - <5> QED - BY <5>4 - <4>7. @ = i * j + (-.j + (i + -.1)) - <5>1. i * j \in Int - BY <1>2, nat_in_int, <2>3, neg_nat_in_int, - multIsInt - <5>2. -.j \in Int - BY <1>2, minus_nat_in_int - <5>3. i + -.1 \in Int - BY <2>3, minus_one_in_negnat, - neg_nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + ((-.j + i) + -.1) - <5>1. -.j \in Int - BY <1>2, minus_nat_in_int - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <5>4. -.j + (i + -.1) = (-.j + i) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>9. @ = i * j + ((i + -.j) + -.1) - <5>1. -.j \in Int - BY <1>2, minus_nat_in_int - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.j + i = i + -.j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, - <4>5, <4>6, <4>7, <4>8, <4>9 - <3>2. Succ[j] * -.Succ[-.i] = - i * j + ((i + -.j) + -.1) - <4>1. Succ[j] * -.Succ[-.i] = - Succ[j] * -.-.i + -.Succ[j] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = Succ[j] * i + -.Succ[j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = i * Succ[j] + -.Succ[j] - BY <2>3 \* induction hypothesis - <4>4. @ = (i * j + i) + -.Succ[j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. j \in Nat - BY <1>2 - <5>3. i * Succ[j] = i * j + i - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>5. @ = (i * j + i) + (-.j + -.1) - <5>1. j \in Nat - BY <1>2 - <5>2. -.Succ[j] = -.j + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>6. @ = i * j + (i + (-.j + -.1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>2, nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.j + -.1 \in Int - <6>1. -.j \in Int - BY <1>2, minus_nat_in_int - <6>2. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + ((i + -.j) + -.1) - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.j \in Int - BY <1>2, minus_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <5>4. i + (-.j + -.1) = (I + -.j) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, NatInduction, neg_nat_induction, - int_union_nat_negnat -<1>3. ASSUME NEW j \in negNat, - \A i \in Int: i * j = j * i - PROVE \A i \in Int: i * -.Succ[-.j] = -.Succ[-.j] * i - <2>1. 0 * -.Succ[-.j] = -.Succ[-.j] * 0 - <3>1. 0 * -.Succ[-.j] = 0 * -.-.j + -.0 - <4>1. 0 \in Int - BY zeroIsNat nat_in_int - <4>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3>2. @ = 0 * -.-.j - <4>1. 0 * -.-.j \in Int - <5>1. 0 \in Int - BY zeroIsNat nat_in_int - <5>2. -.-.j \in Int - BY <1>3, neg_nat_in_int, minus_involutive - <5> QED - BY <5>1, <5>2, multIsInt - <4>2. 0 * -.-.j + 0 = 0 * -.-.j - BY <4>1, add_0 - <4> QED - BY <4>2, neg0 - <3>3. @ = 0 * j - <4>1. j \in Int - BY <1>3, neg_nat_in_int - <4> QED - BY <4>1, minus_involutive - <3>4. @ = j * 0 - <4>1. 0 \in Int - BY zeroIsNat, nat_in_int - <4>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <4> QED - BY <4>1, <4>2, spec - <3>5. @ = 0 - <4>1. j \in Int - BY <1>3, neg_nat_in_int - <4> QED - BY <4>1, mult_0 - <3>6. @ = -.Succ[-.j] * 0 - <4>1. -.Succ[-.j] \in Int - BY <1>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 - <2>2. ASSUME NEW i \in Nat, - i * -.Succ[-.j] = -.Succ[-.j] * i - PROVE Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i] - <3>1. Succ[i] * -.Succ[-.j] = i * j + ((-.i + j) + -.1) - <4>1. Succ[i] * -.Succ[-.j] = - Succ[i] * -.-.j + -.Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = Succ[i] * j + -.Succ[i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = j * Succ[i] + -.Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. Succ[i] * j = j * Succ[i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>4. @ = (j * i + j) + -.Succ[i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. i \in Nat - BY <2>2 - <5>3. j * Succ[i] = j * i + j - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>5. @ = (i * j + j) + -.Succ[i] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>6. @ = i * j + (j + -.Succ[i]) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>2, nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. j \in Int - BY <1>3, neg_nat_in_int - <5>3. -.Succ[i] \in Int - BY <2>2, succIsNat, minus_nat_in_int - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + (j + (-.i + -.1)) - <5>1. i \in Nat - BY <2>2 - <5>2. -.Succ[i] = -.i + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>8. @ = i * j + ((j + -.i) + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_int - <5>4. j + (-.i + -.1) = (j + -.i) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>9. @ = i * j + ((-.i + j) + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. j + -.i = -.i + j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9 - <3>2. -.Succ[-.j] * Succ[i] = i * j + ((-.i + j) + -.1) - <4>1. -.Succ[-.j] * Succ[i] = - -.Succ[-.j] * i + -.Succ[-.j] - <5>1. -.Succ[-.j] \in Int - BY <1>3, minus_negnat_succ_in_negnat, - neg_nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = i * -.Succ[-.j] + -.Succ[-.j] - BY <2>2 \* induction hypothesis - <4>3. @ = (i * -.-.j + -.i) + -.Succ[-.j] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>3. i * -.Succ[-.j] = i * -.-.j + -.i - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>4. @ = (i * j + -.i) + -.Succ[-.j] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>5. @ = (i * j + -.i) + (-.-.j + -.1) - <5>1. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>2. -.Succ[-.j] = -.-.j + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>6. @ = (i * j + -.i) + (j + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = i * j + (-.i + (j + -.1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>2, nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. j + -.1 \in Int - <6>1. j \in Int - BY <1>3, neg_nat_in_int - <6>2. -.1 \in Int - BY minus_one_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + ((-.i + j) + -.1) - <5>1. -.i \in Int - BY <2>2, minus_nat_in_int - <5>2. j \in Int - BY <1>3, neg_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_int - <5>4. -.i + (j + -.1) = (-.i + j) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8 - <3> QED - BY <3>1, <3>2 - <2>3. ASSUME NEW i \in negNat, - i * -.Succ[-.j] = -.Succ[-.j] * i - PROVE -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i] - <3>1. -.Succ[-.i] * -.Succ[-.j] = - i * j + ((-.i + -.j) + 1) - <4>1. -.Succ[-.i] * -.Succ[-.j] = - -.Succ[-.i] * -.-.j + -.-.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, - neg_nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = -.Succ[-.i] * j + -.-.Succ[-.i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = -.Succ[-.i] * j + Succ[-.i] - <5>1. Succ[-.i] \in Int - BY <2>3, negnat_succ_in_nat, nat_in_int - <5>2. -.-.Succ[-.i] = Succ[-.i] - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>4. @ = j * -.Succ[-.i] + Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>5. @ = (j * -.-.i + -.j) + Succ[-.i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>3. j * -.Succ[-.i] = j * -.-.i + -.j - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>6. @ = (j * i + -.j) + Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = (i * j + -.j) + Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>8. @ = (i * j + -.j) + (-.i + 1) - <5>1. i \in negNat - BY <2>3 - <5>2. Succ[-.i] = -.i + 1 - BY <5>1, SuccMinusIsPlusOne - <5> QED - BY <5>2 - <4>9. @ = i * j + (-.j + (-.i + 1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>3. -.i + 1 \in Int - <6>1. -.i \in Int - BY <2>3, minus_negnat_in_int - <6>2. 1 \in Int - BY oneIsNat, nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>10. @ = i * j + ((-.j + -.i) + 1) - <5>1. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. -.j + (-.i + 1) = (-.j + -.i) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>11. @ = i * j + ((-.i + -.j) + 1) - <5>1. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. -.j + -.i = -.i + -.j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9, - <4>10, <4>11 - <3>2. -.Succ[-.j] * -.Succ[-.i] = - i * j + ((-.i + -.j) + 1) - <4>1. -.Succ[-.j] * -.Succ[-.i] = - -.Succ[-.j] * -.-.i + -.-.Succ[-.j] - <5>1. -.Succ[-.j] \in Int - BY <1>3, minus_succ_minus_negnat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = -.Succ[-.j] * i + -.-.Succ[-.j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = -.Succ[-.j] * i + Succ[-.j] - <5>1. Succ[-.j] \in Int - BY <1>3, negnat_succ_in_nat, nat_in_int - <5>2. -.-.Succ[-.j] = Succ[-.j] - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>4. @ = i * -.Succ[-.j] + Succ[-.j] - BY <2>3 \* induction hypothesis - <4>5. @ = (i * -.-.j + -.i) + Succ[-.j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>3. i * -.Succ[-.j] = i * -.-.j + -.i - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>6. @ = (i * j + -.i) + Succ[-.j] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = i * j + (-.i + Succ[-.j]) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. Succ[-.j] \in Int - BY <1>3, negnat_succ_in_nat, nat_in_int - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + (-.i + (-.j + 1)) - <5>1. Succ[-.j] = -.j + 1 - BY <1>3, SuccMinusIsPlusOne - <5> QED - BY <5>1 - <4>9. @ = i * j + ((-.i + -.j) + 1) - <5>1. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>2. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. -.i + (-.j + 1) = (-.i + -.j) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, NatInduction, neg_nat_induction, - int_union_nat_negnat -<1> QED - BY <1>1, <1>2, <1>3, NatInduction, neg_nat_induction, - int_union_nat_negnat -*) -theorem MultCommutative: - "\ i \ Int: \ j \ Int: - i * j = j * i" - proof - - have s1_1: "\ i \ Int: i * 0 = 0 * i" - proof - - have s2_1: "0 * 0 = 0 * 0" - by auto - have s2_2: "\ i. \ - i \\in Nat; - i * 0 = 0 * i - \ \ - Succ[i] * 0 = 0 * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * 0 = 0 * i" - have s3_1: "0 * Succ[i] = 0 * i + 0" - proof - - have s4_1: "0 \\in Int" - using zero_in_int by auto - have s4_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - also have s3_2: "... = 0 * i" - proof - - have s4_1: "0 * i \\in Int" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - show ?thesis - using s4_1 add_0 by auto - qed - also have s3_3: "... = i * 0" - using s2_2_induct by auto - also have s3_4: "... = 0" - proof - - have s4_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - also have s3_5: "... = Succ[i] * 0" - proof - - have s4_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - finally show "Succ[i] * 0 = 0 * Succ[i]" by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * 0 = 0 * i - \ \ - -.Succ[-.i] * 0 = 0 * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * 0 = 0 * i" - have s3_1: "-.Succ[-.i] * 0 = 0" - proof - - have s4_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_succ_minus_negnat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - have s3_2: "0 * -.Succ[-.i] = 0" - proof - - have s4_1: "0 * -.Succ[-.i] = 0 * -.-.i + -.0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = 0 * i + -.0" - proof - - have s5_1: "-.-.i = i" - using s2_3_idom neg_nat_in_int - minus_involutive by auto - show ?thesis - using s5_1 by auto - qed - also have s4_3: "... = 0 * i + 0" - using neg0 by auto - also have s4_4: "... = 0 * i" - proof - - have s5_1: "0 * i \\in Int" - proof - - have s6_1: "0 \\in Int" - using zero_in_int by auto - have s6_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - show ?thesis - using s5_1 add_0 by auto - qed - also have s4_5: "... = i * 0" - using s2_3_induct by auto - also have s4_6: "... = 0" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - show ?thesis - using s5_1 mult_0 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * 0 = 0 * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: i * 0 = 0 * i" - using s2_1 s2_2 natInduct[of "\ i. - i * 0 = 0 * i"] - by blast - have s2_5: "\ i \ negNat: i * 0 = 0 * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * 0 = 0 * i"] - by blast - show ?thesis - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_2: "\ j. \ - j \\in Nat; - \ i \ Int: i * j = j * i - \ \ - \ i \ Int: i * Succ[j] = Succ[j] * i" - proof - - fix "j" :: "c" - assume s1_2_jdom: "j \\in Nat" - assume s1_2_induct: "\ i \ Int: i * j = j * i" - have s2_1: "0 * Succ[j] = Succ[j] * 0" - proof - - have s3_1: "0 * Succ[j] = 0" - proof - - have s4_1: "0 * Succ[j] = 0 * j + 0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = 0 * j" - proof - - have s5_1: "0 * j \\in Int" - proof - - have s6_1: "0 \\in Int" - using zero_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - show ?thesis - using s5_1 add_0 by auto - qed - also have s4_3: "... = j * 0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - show ?thesis - using s5_1 s5_2 spec by fast - qed - also have s4_4: "... = 0" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s5_1 mult_0 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * 0 = 0" - proof - - have s4_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - have s2_2: "\ i. \ - i \\in Nat; - i * Succ[j] = Succ[j] * i - \ \ - Succ[i] * Succ[j] = Succ[j] * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * Succ[j] = Succ[j] * i" - have s3_1: "Succ[i] * Succ[j] = i * j + ((i + j) + 1)" - proof - - have s4_1: "Succ[i] * Succ[j] = Succ[i] * j + Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = j * Succ[i] + Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_2_induct by auto - have s5_3: "Succ[i] * j = j * Succ[i]" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_3: "... = (j * i + j) + Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - have s5_3: "j * Succ[i] = j * i + j" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + j) + Succ[i]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_2_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + j) + (i + 1)" - proof - - have s5_1: "i \\in Nat" - using s2_2_idom by auto - have s5_2: "Succ[i] = i + 1" - using s5_1 nat_add_1 by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = i * j + (j + (i + 1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_3: "i + 1 \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "1 \\in Nat" - using oneIsNat nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + ((j + i) + 1)" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "j + (i + 1) = (j + i) + 1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_8: "... = i * j + ((i + j) + 1)" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "j + i = i + j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * Succ[i] = i * j + ((i + j) + 1)" - proof - - have s4_1: "Succ[j] * Succ[i] = Succ[j] * i + Succ[j]" - proof - - have s5_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = i * Succ[j] + Succ[j]" - using s2_2_induct by auto - also have s4_3: "... = (i * j + i) + Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - have s5_3: "i * Succ[j] = i * j + i" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + i) + (j + 1)" - proof - - have s5_1: "j \\in Nat" - using s1_2_jdom by auto - have s5_2: "Succ[j] = j + 1" - using s5_1 nat_add_1 by fast - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = i * j + (i + (j + 1))" - proof - - have s5_1: "i * j \\in Int" - using s1_2_jdom s2_2_idom nat_in_int multIsInt by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "j + 1 \\in Int" - using s1_2_jdom oneIsNat nat_in_int addIsInt by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_6: "... = i * j + ((i + j) + 1)" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "i + (j + 1) = (i + j) + 1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "Succ[i] * Succ[j] = Succ[j] * Succ[i]" - using s3_1 s3_2 by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * Succ[j] = Succ[j] * i - \ \ - -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * Succ[j] = Succ[j] * i" - have s3_1: "-.Succ[-.i] * Succ[j] = - i * j + ((i + -.j) + -.1)" - proof - - have s4_1: "-.Succ[-.i] * Succ[j] = - -.Succ[-.i] * j + -.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = j * -.Succ[-.i] + -.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_3: "... = (j * -.-.i + -.j) + -.Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (j * i + -.j) + -.Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = (i * j + -.j) + -.Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (i * j + -.j) + (i + -.1)" - proof - - have s5_1: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_2: "-.Succ[-.i] = -.-.i + -.1" - using s5_1 MinusSuccIsMinusOne by fast - have s5_3: "-.-.i = i" - using s2_3_idom neg_nat_in_int minus_involutive - by auto - have s5_4: "-.Succ[-.i] = i + -.1" - using s5_2 s5_3 by auto - show ?thesis - using s5_4 by auto - qed - also have s4_7: "... = i * j + (-.j + (i + -.1))" - proof - - have s5_1: "i * j \\in Int" - using s1_2_jdom nat_in_int s2_3_idom - neg_nat_in_int multIsInt by auto - have s5_2: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_3: "i + -.1 \\in Int" - using s2_3_idom minus_one_in_negnat - neg_nat_in_int addIsInt by auto - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + ((-.j + i) + -.1)" - proof - - have s5_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - have s5_4: " -.j + (i + -.1) = (-.j + i) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_9: "... = i * j + ((i + -.j) + -.1)" - proof - - have s5_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.j + i = i + -.j" - using s5_1 s5_2 AddCommutative_sequent - by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * -.Succ[-.i] = - i * j + ((i + -.j) + -.1)" - proof - - have s4_1: "Succ[j] * -.Succ[-.i] = - Succ[j] * -.-.i + -.Succ[j]" - proof - - have s5_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = Succ[j] * i + -.Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = i * Succ[j] + -.Succ[j]" - using s2_3_induct by auto - also have s4_4: "... = (i * j + i) + -.Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - have s5_3: "i * Succ[j] = i * j + i" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + i) + (-.j + -.1)" - proof - - have s5_1: "j \\in Nat" - using s1_2_jdom by auto - have s5_2: "-.Succ[j] = -.j + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = i * j + (i + (-.j + -.1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.j + -.1 \\in Int" - proof - - have s6_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s6_2: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + ((i + -.j) + -.1)" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - have s5_4: "i + (-.j + -.1) = (i + -.j) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_2 natInduct[of "\ i. - i * Succ[j] = Succ[j] * i"] - by auto - have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * Succ[j] = Succ[j] * i"] - by auto - show "\ i \ Int: i * Succ[j] = Succ[j] * i" - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_3: "\ j. \ - j \\in negNat; - \ i \ Int: i * j = j * i - \ \ - \ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" - proof - - fix "j" :: "c" - assume s1_3_jdom: "j \\in negNat" - assume s1_3_induct: "\ i \ Int: i * j = j * i" - have s2_1: "0 * -.Succ[-.j] = -.Succ[-.j] * 0" - proof - - have s3_1: "0 * -.Succ[-.j] = 0 * -.-.j + -.0" - proof - - have s4_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s4_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - also have s3_2: "... = 0 * -.-.j" - proof - - have s4_1: "0 * -.-.j \\in Int" - proof - - have s5_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s5_2: "-.-.j \\in Int" - using s1_3_jdom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - have s4_2: "0 * -.-.j + 0 = 0 * -.-.j" - using s4_1 add_0 by auto - show ?thesis - using s4_2 neg0 by auto - qed - also have s3_3: "... = 0 * j" - proof - - have s4_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s4_1 minus_involutive by auto - qed - also have s3_4: "... = j * 0" - proof - - have s4_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s4_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - show ?thesis - using s4_1 s4_2 spec by fast - qed - also have s3_5: "... = 0" - proof - - have s4_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - also have s3_6: "... = -.Succ[-.j] * 0" - proof - - have s4_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - finally show ?thesis . - qed - have s2_2: "\ i. \ - i \\in Nat; - i * -.Succ[-.j] = -.Succ[-.j] * i - \ \ - Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" - have s3_1: "Succ[i] * -.Succ[-.j] = - i * j + ((-.i + j) + -.1)" - proof - - have s4_1: "Succ[i] * -.Succ[-.j] = - Succ[i] * -.-.j + -.Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = Succ[i] * j + -.Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = j * Succ[i] + -.Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - have s5_3: "Succ[i] * j = j * Succ[i]" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (j * i + j) + -.Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - have s5_3: "j * Succ[i] = j * i + j" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + j) + -.Succ[i]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = i * j + (j + -.Succ[i])" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_3: "-.Succ[i] \\in Int" - using s2_2_idom succIsNat minus_nat_in_int by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + (j + (-.i + -.1))" - proof - - have s5_1: "i \\in Nat" - using s2_2_idom by auto - have s5_2: "-.Succ[i] = -.i + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_8: "... = i * j + ((j + -.i) + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_int by auto - have s5_4: "j + (-.i + -.1) = (j + -.i) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_9: "... = i * j + ((-.i + j) + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "j + -.i = -.i + j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "-.Succ[-.j] * Succ[i] = - i * j + ((-.i + j) + -.1)" - proof - - have s4_1: "-.Succ[-.j] * Succ[i] = - -.Succ[-.j] * i + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = i * -.Succ[-.j] + -.Succ[-.j]" - using s2_2_induct by auto - also have s4_3: "... = (i * -.-.j + -.i) + -.Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + -.i) + -.Succ[-.j]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = (i * j + -.i) + (-.-.j + -.1)" - proof - - have s5_1: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_2: "-.Succ[-.j] = -.-.j + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = (i * j + -.i) + (j + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = i * j + (-.i + (j + -.1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "j + -.1 \\in Int" - proof - - have s6_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s6_2: "-.1 \\in Int" - using minus_one_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + ((-.i + j) + -.1)" - proof - - have s5_1: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_int by auto - have s5_4: "-.i + (j + -.1) = (-.i + j) + -.1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" - using s3_1 s3_2 by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * -.Succ[-.j] = -.Succ[-.j] * i - \ \ - -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" - have s3_1: "-.Succ[-.i] * -.Succ[-.j] = - i * j + ((-.i + -.j) + 1)" - proof - - have s4_1: "-.Succ[-.i] * -.Succ[-.j] = - -.Succ[-.i] * -.-.j + -.-.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = -.Succ[-.i] * j + -.-.Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = -.Succ[-.i] * j + Succ[-.i]" - proof - - have s5_1: "Succ[-.i] \\in Int" - using s2_3_idom negnat_succ_in_nat nat_in_int by auto - have s5_2: "-.-.Succ[-.i] = Succ[-.i]" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_4: "... = j * -.Succ[-.i] + Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_3_induct by auto - have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (j * -.-.i + -.j) + Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (j * i + -.j) + Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = (i * j + -.j) + Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_3_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_8: "... = (i * j + -.j) + (-.i + 1)" - proof - - have s5_1: "i \\in negNat" - using s2_3_idom by auto - have s5_2: "Succ[-.i] = -.i + 1" - using s5_1 SuccMinusIsPlusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_9: "... = i * j + (-.j + (-.i + 1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_3: "-.i + 1 \\in Int" - proof - - have s6_1: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s6_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_10: "... = i * j + ((-.j + -.i) + 1)" - proof - - have s5_1: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "-.j + (-.i + 1) = (-.j + -.i) + 1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_11: "... = i * j + ((-.i + -.j) + 1)" - proof - - have s5_1: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "-.j + -.i = -.i + -.j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "-.Succ[-.j] * -.Succ[-.i] = - i * j + ((-.i + -.j) + 1)" - proof - - have s4_1: "-.Succ[-.j] * -.Succ[-.i] = - -.Succ[-.j] * -.-.i + -.-.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_succ_minus_negnat_in_int - by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = -.Succ[-.j] * i + -.-.Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = -.Succ[-.j] * i + Succ[-.j]" - proof - - have s5_1: "Succ[-.j] \\in Int" - using s1_3_jdom negnat_succ_in_nat - nat_in_int by auto - have s5_2: "-.-.Succ[-.j] = Succ[-.j]" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_4: "... = i * -.Succ[-.j] + Succ[-.j]" - using s2_3_induct by auto - also have s4_5: "... = (i * -.-.j + -.i) + Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (i * j + -.i) + Succ[-.j]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = i * j + (-.i + Succ[-.j])" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "Succ[-.j] \\in Int" - using s1_3_jdom negnat_succ_in_nat nat_in_int by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + (-.i + (-.j + 1))" - proof - - have s5_1: "Succ[-.j] = -.j + 1" - using s1_3_jdom SuccMinusIsPlusOne by fast - show ?thesis - using s5_1 by auto - qed - also have s4_9: "... = i * j + ((-.i + -.j) + 1)" - proof - - have s5_1: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_2: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "-.i + (-.j + 1) = (-.i + -.j) + 1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: - i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_2 natInduct[of "\ i. - i * -.Succ[-.j] = -.Succ[-.j] * i"] - by auto - have s2_5: "\ i \ negNat: - i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * -.Succ[-.j] = -.Succ[-.j] * i"] - by auto - show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_4: "\ j \ Nat: \ i \ Int: - i * j = j * i" - using s1_1 s1_2 natInduct[of "\ j. - \ i \ Int: i * j = j * i"] - by blast - have s1_5: "\ j \ negNat: \ i \ Int: - i * j = j * i" - using s1_1 s1_3 neg_nat_induction[of "\ j. - \ i \ Int: i * j = j * i"] - by blast - show ?thesis - using s1_4 s1_5 int_union_nat_negnat by auto - qed - - -theorem MultCommutative_sequent: - assumes idom: "i \\in Int" and jdom: "j \\in Int" - shows "i * j = j * i" - using idom jdom MultCommutative spec by auto - - -(* -THEOREM mult_0_left == - ASSUME m \in Int - PROVE 0 * m = 0 - BY mult_0, MultCommutative -*) -theorem mult_0_left: - assumes ndom: "n \\in Int" - shows "0 * n = 0" - proof - - have s1_1: "n * 0 = 0" - using ndom mult_0 by auto - have s1_2: "n * 0 = 0 * n" - using ndom zero_in_int MultCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Associativity of multiplication. *) - - -(* -THEOREM MultLeftDistributesAdd == - \A a, b, c \in Int: - a * (b + c) = a * b + a * c -PROOF -<1>1. SUFFICES - ASSUME NEW b \in Int, NEW c \in Int - PROVE \A a \in Int: a * (b + c) = a * b + a * c - BY <1>1 -<1>2. 0 * (b + c) = 0 * b + 0 * c - <2>1. 0 * (b + c) = 0 - <3>1. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, mult_0_left - <2>2. 0 * b + 0 * c = 0 - <3>1. 0 * b = 0 - BY <1>1, mult_0_left - <3>2. 0 * c = 0 - BY <1>1, mult_0_left - <3> QED - BY <3>1, <3>2, zero_in_int, add_0 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW a \in Nat, - a * (b + c) = a * b + a * c - PROVE Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c - <2>1. Succ[a] * (b + c) = (b + c) * Succ[a] - <3>1. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, MultCommutative - <2>2. @ = (b + c) * a + (b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. a \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>3. @ = a * (b + c) + (b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. a \in Int - BY <1>3, nat_in_int - <3>3. (b + c) * a = a * (b + c) - BY <3>1, <3>2, MultCommutative - <3> QED - BY <3>3 - <2>4. @ = (a * b + a * c) + (b + c) - BY <1>3 \* induction hypothesis - <2>5. @ = a * b + (a * c + (b + c)) - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>6. @ = a * b + ((a * c + b) + c) - <3>1. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. c \in Int - BY <1>1 - <3>4. a * c + (b + c) = (a * c + b) + c - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>7. @ = a * b + ((b + a * c) + c) - <3>1. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. a * c + b = b + a * c - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>8. @ = a * b + (b + (a * c + c)) - <3>1. b \in Int - BY <1>1 - <3>2. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. c \in Int - BY <1>1 - <3>4. (b + a * c) + c = b + (a * c + c) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>9. @ = (a * b + b) + (a * c + c) - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. (a * c + c) \in Int - BY <1>1, <1>3, nat_in_int, addIsInt, multIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>10. @ = (b * a + b) + (c * a + c) - <3>1. a * b = b * a - <4>1. a \in Int - BY <1>3, nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. a * c = c * a - <4>1. a \in Int - BY <1>3, nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2>11. @ = b * Succ[a] + c * Succ[a] - <3>1. b * Succ[a] = b * a + b - <4>1. b \in Int - BY <1>1 - <4>2. a \in Nat - BY <1>3 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3>2. c * Succ[a] = c * a + c - <4>1. c \in Int - BY <1>1 - <4>2. a \in Nat - BY <1>3 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3> QED - BY <3>1, <3>2 - <2>12. @ = Succ[a] * b + Succ[a] * c - <3>1. b * Succ[a] = Succ[a] * b - <4>1. b \in Int - BY <1>1 - <4>2. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. c * Succ[a] = Succ[a] * c - <4>1. c \in Int - BY <1>1 - <4>2. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, - <2>5, <2>6, <2>7, <2>8, <2>9, - <2>10, <2>11, <2>12 -<1>4. ASSUME NEW a \in negNat, - a * (b + c) = a * b + a * c - PROVE -.Succ[-.a] * (b + c) = -.Succ[-.a] * b + -.Succ[-.a] * c - <2>1. -.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a] - <3>1. -.Succ[-.a] \in Int - BY <1>4, minus_succ_minus_negnat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, MultCommutative - <2>2. @ = (b + c) * -.-.a + -.(b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>3. @ = (b + c) * a + -.(b + c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>4. @ = a * (b + c) + -.(b + c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3>3. (b + c) * a = a * (b + c) - BY <3>1, <3>2, MultCommutative - <3> QED - BY <3>3 - <2>5. @ = (a * b + a * c) + -.(b + c) - BY <1>4 \* induction hypothesis - <2>6. @ = (a * b + a * c) + (-.b + -.c) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>1 - <3>3. -.(b + c) = -.b + -.c - BY <3>1, <3>2, MinusDistributesAdd, spec - <3> QED - BY <3>3 - <2>7. @ = a * b + (a * c + (-.b + -.c)) - <3>1. a * b \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.b + -.c \in Int - <4>1. -.b \in Int - BY <1>1, neg_int_eq_int - <4>2. -.c \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>8. @ = a * b + ((a * c + -.b) + -.c) - <3>1. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. -.c \in Int - BY <1>1, neg_int_eq_int - <3>4. a * c + (-.b + -.c) = - (a * c + -.b) + -.c - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>9. @ = a * b + ((-.b + a * c) + -.c) - <3>1. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. a * c + -.b = -.b + a * c - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>10. @ = a * b + (-.b + (a * c + -.c)) - <3>1. -.b \in Int - BY <1>1, neg_int_eq_int - <3>2. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.c \in Int - BY <1>1, neg_int_eq_int - <3>4. (-.b + a * c) + -.c = - -.b + (a * c + -.c) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>11. @ = (a * b + -.b) + (a * c + -.c) - <3>1. a * b \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. a * c + -.c \in Int - <4>1. a * c \in Int - <5>1. a \in Int - BY <1>4, neg_nat_in_int - <5>2. c \in Int - BY <1>1 - <5> QED - BY <5>1, <5>2, multIsInt - <4>2. -.c \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>12. @ = (b * a + -.b) + (c * a + -.c) - <3>1. a * b = b * a - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. a * c = c * a - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2>13. @ = (b * -.-.a + -.b) + (c * -.-.a + -.c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>14. @ = b * -.Succ[-.a] + c * -.Succ[-.a] - <3>1. b * -.Succ[-.a] = b * -.-.a + -.b - <4>1. b \in Int - BY <1>1 - <4>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3>2. c * -.Succ[-.a] = c * -.-.a + -.c - <4>1. c \in Int - BY <1>1 - <4>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3> QED - BY <3>1, <3>2 - <2>15. @ = -.Succ[-.a] * b + -.Succ[-.a] * c - <3>1. -.Succ[-.a] \in Int - BY <1>4, minus_succ_minus_negnat_in_int - <3>2. b * -.Succ[-.a] = -.Succ[-.a] * b - <4>1. b \in Int - BY <1>1 - <4> QED - BY <3>1, <4>1, MultCommutative - <3>3. c * -.Succ[-.a] = -.Succ[-.a] * c - <4>1. c \in Int - BY <1>1 - <4> QED - BY <3>1, <4>1, MultCommutative - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, 2<6>, - <2>7, <2>8, <2>9, <2>10, <2>11, - <2>12, <2>13, <2>14, <2>15 -<1>5. \A a \in Nat: a * (b + c) = a * b + a * c - BY <1>2, <1>3, NatInduction -<1>6. \A a \in negNat: a * (b + c) = a * b + a * c - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MultLeftDistributesAdd_forall: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - a * (b + c) = a * b + a * c" - proof - - { - fix "b" :: "c" - fix "c" :: "c" - assume bdom: "b \\in Int" - assume cdom: "c \\in Int" - have s1_2: "0 * (b + c) = 0 * b + 0 * c" - proof - - have s2_1: "0 * (b + c) = 0" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 mult_0_left by auto - qed - have s2_2: "0 * b + 0 * c = 0" - proof - - have s3_1: "0 * b = 0" - using bdom mult_0_left by auto - have s3_2: "0 * c = 0" - using cdom mult_0_left by auto - show ?thesis - using s3_1 s3_2 zero_in_int add_0 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ a. \ - a \\in Nat; - a * (b + c) = a * b + a * c - \ \ - Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" - proof - - fix "a" :: "c" - assume s1_3_adom: "a \\in Nat" - assume s1_3_induct: "a * (b + c) = a * b + a * c" - have s2_1: "Succ[a] * (b + c) = (b + c) * Succ[a]" - proof - - have s3_1: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 MultCommutative_sequent by auto - qed - also have s2_2: "... = (b + c) * a + (b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_3: "... = a * (b + c) + (b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s3_3: "(b + c) * a = a * (b + c)" - using s3_1 s3_2 MultCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = (a * b + a * c) + (b + c)" - using s1_3_induct by auto - also have s2_5: "... = a * b + (a * c + (b + c))" - proof - - have s3_1: "a * b \\in Int" - using bdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_3: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_6: "... = a * b + ((a * c + b) + c)" - proof - - have s3_1: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "c \\in Int" - using cdom by auto - have s3_4: "a * c + (b + c) = (a * c + b) + c" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_7: "... = a * b + ((b + a * c) + c)" - proof - - have s3_1: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "a * c + b = b + a * c" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_8: "... = a * b + (b + (a * c + c))" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_3: "c \\in Int" - using cdom by auto - have s3_4: "(b + a * c) + c = b + (a * c + c)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_9: "... = (a * b + b) + (a * c + c)" - proof - - have s3_1: "a * b \\in Int" - using bdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "a * c + c \\in Int" - using cdom s1_3_adom nat_in_int addIsInt multIsInt by auto - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_10: "... = (b * a + b) + (c * a + c)" - proof - - have s3_1: "a * b = b * a" - proof - - have s4_1: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - have s3_2: "a * c = c * a" - proof - - have s4_1: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_11: "... = b * Succ[a] + c * Succ[a]" - proof - - have s3_1: "b * Succ[a] = b * a + b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - have s3_2: "c * Succ[a] = c * a + c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_12: "... = Succ[a] * b + Succ[a] * c" - proof - - have s3_1: "b * Succ[a] = Succ[a] * b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - show ?thesis - using s4_1 s4_2 MultCommutative by auto - qed - have s3_2: "c * Succ[a] = Succ[a] * c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - show ?thesis - using s4_1 s4_2 MultCommutative by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - finally show "Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" . - qed - have s1_4: "\ a. \ - a \\in negNat; - a * (b + c) = a * b + a * c - \ \ - -.Succ[-.a] * (b + c) = - -.Succ[-.a] * b + -.Succ[-.a] * c" - proof - - fix "a" :: "c" - assume s1_4_adom: "a \\in negNat" - assume s1_4_induct: "a * (b + c) = a * b + a * c" - have s2_1: "-.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a]" - proof - - have s3_1: "-.Succ[-.a] \\in Int" - using s1_4_adom minus_succ_minus_negnat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 MultCommutative by auto - qed - also have s2_2: "... = (b + c) * -.-.a + -.(b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_3: "... = (b + c) * a + -.(b + c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_4: "... = a * (b + c) + -.(b + c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_3: "(b + c) * a = a * (b + c)" - using s3_1 s3_2 MultCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = (a * b + a * c) + -.(b + c)" - using s1_4_induct by auto - also have s2_6: "... = (a * b + a * c) + (-.b + -.c)" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "c \\in Int" - using cdom by auto - have s3_3: "-.(b + c) = -.b + -.c" - using s3_1 s3_2 MinusDistributesAdd spec by auto - show ?thesis - using s3_3 by auto - qed - also have s2_7: "... = a * b + (a * c + (-.b + -.c))" - proof - - have s3_1: "a * b \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.b + -.c \\in Int" - proof - - have s4_1: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s4_2: "-.c \\in Int" - using cdom neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_8: "... = a * b + ((a * c + -.b) + -.c)" - proof - - have s3_1: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "-.c \\in Int" - using cdom neg_int_eq_int by auto - have s3_4: "a * c + (-.b + -.c) = - (a * c + -.b) + -.c" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_9: "... = a * b + ((-.b + a * c) + -.c)" - proof - - have s3_1: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "a * c + -.b = -.b + a * c" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_10: "... = a * b + (-.b + (a * c + -.c))" - proof - - have s3_1: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_2: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.c \\in Int" - using cdom neg_int_eq_int by auto - have s3_4: "(-.b + a * c) + -.c = - -.b + (a * c + -.c)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_11: "... = (a * b + -.b) + (a * c + -.c)" - proof - - have s3_1: "a * b \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "a * c + -.c \\in Int" - proof - - have s4_1: "a * c \\in Int" - proof - - have s5_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s5_2: "c \\in Int" - using cdom by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - have s4_2: "-.c \\in Int" - using cdom neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_12: "... = (b * a + -.b) + (c * a + -.c)" - proof - - have s3_1: "a * b = b * a" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - have s3_2: "a * c = c * a" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_13: "... = (b * -.-.a + -.b) + (c * -.-.a + -.c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_14: "... = b * -.Succ[-.a] + c * -.Succ[-.a]" - proof - - have s3_1: "b * -.Succ[-.a] = b * -.-.a + -.b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - have s3_2: "c * -.Succ[-.a] = c * -.-.a + -.c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_15: "... = -.Succ[-.a] * b + -.Succ[-.a] * c" - proof - - have s3_1: "-.Succ[-.a] \\in Int" - using s1_4_adom minus_succ_minus_negnat_in_int by auto - have s3_2: "b * -.Succ[-.a] = -.Succ[-.a] * b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - show ?thesis - using s3_1 s4_1 MultCommutative_sequent by auto - qed - have s3_3: "c * -.Succ[-.a] = -.Succ[-.a] * c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - show ?thesis - using s3_1 s4_1 MultCommutative_sequent by auto - qed - show ?thesis - using s3_2 s3_3 by auto - qed - finally show "-.Succ[-.a] * (b + c) = - -.Succ[-.a] * b + -.Succ[-.a] * c" . - qed - have s1_5: "\ a \ Nat: - a * (b + c) = a * b + a * c" - using s1_2 s1_3 natInduct[of "\ a. - a * (b + c) = a * b + a * c"] - by auto - have s1_6: "\ a \ negNat: - a * (b + c) = a * b + a * c" - using s1_2 s1_4 neg_nat_induction[of "\ a. - a * (b + c) = a * b + a * c"] - by auto - have s1_7: "\ a \ Int: - a * (b + c) = a * b + a * c" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem MultLeftDistributesAdd: - assumes adom: "a \\in Int" and bdom: "b \\in Int" and - cdom: "c \\in Int" - shows "a * (b + c) = a * b + a * c" - using adom bdom cdom MultLeftDistributesAdd_forall by auto - - -theorem MultRightDistributesAdd: - assumes adom: "a \\in Int" and bdom: "b \\in Int" and - cdom: "c \\in Int" - shows "(b + c) * a = b * a + c * a" - proof - - have s1_1: "a * (b + c) = a * b + a * c" - using adom bdom cdom MultLeftDistributesAdd by auto - have s1_2: "a * (b + c) = (b + c) * a" - using adom bdom cdom addIsInt MultCommutative by auto - have s1_3: "a * b = b * a" - using adom bdom MultCommutative by auto - have s1_4: "a * c = c * a" - using adom cdom MultCommutative by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -lemmas mult_distributes[algebra_simps] = - MultLeftDistributesAdd MultRightDistributesAdd - - -(* -THEOREM MinusCommutesRightMult == - \A a, b \in Int: -.(a * b) = a * -.b -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int - PROVE \A b \in Int: -.(a * b) = a * -.b - BY <1>1 -<1>2. -.(a * 0) = a * -.0 - <2>1. -.(a * 0) = 0 - BY <1>1, mult_0, neg0 - <2>2. a * -.0 = 0 - BY <1>1, neg0, mult_0 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW b \in Nat, - -.(a * b) = a * -.b - PROVE -.(a * Succ[b]) = a * -.Succ[b] - <2>1. -.(a * Succ[b]) = -.(a * b + a) - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>2. @ = -.(a * b) + -.a - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. a \in Int - BY <1>1 - <3> QED - BY <3>1, <3>2, MinusDistributesAdd - <2>3. @ = a * -.b + -.a - BY <1>3 \* induction hypothesis - <2>4. @ = a * -.Succ[b] - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>4. ASSUME NEW b \in negNat, - -.(a * b) = a * -.b - PROVE -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b] - <2>1. -.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a) - <3>1. a \in Int - BY <1>1 - <3>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>2. @ = -.(a * -.-.b) + -.-.a - <3>1. a * -.-.b \in Int - <4>1. a \in Int - BY <1>1 - <4>2. -.-.b \in Int - BY <1>4, minus_negnat_in_int, - neg_int_eq_int - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.a \in Int - BY <1>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, MinusDistributesAdd - <2>3. @ = -.(a * -.-.b) + a - <3>1. a \in Int - BY <1>1 - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>4. @ = -.(a * b) + a - <3>1. b \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.b = b - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>5. @ = a * -.b + a - BY <1>4 \* induction hypothesis - <2>6. @ = a * Succ[-.b] - <3>1. a \in Int - BY <1>1 - <3>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>7. @ = a * -.-.Succ[-.b] - <3>1. Succ[-.b] \in Int - BY <1>4, negnat_succ_in_nat, nat_in_int - <3>2. -.-.Succ[-.b] = Succ[-.b] - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2> QED - BY <2>1, <2>2, <2>3, ,2>4, - <2>5, <2>6, <2>7 -<1>5. \A b \in Nat: -.(a * b) = a * -.b - BY <1>2, <1>3, NatInduction -<1>6. \A b \in negNat: -.(a * b) = a * -.b - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MinusCommutesRightMult: - "\ a \ Int: \ b \ Int: - -.(a * b) = a * -.b" - proof - - { - fix "a" :: "c" - assume adom: "a \\in Int" - have s1_2: "-.(a * 0) = a * -.0" - proof - - have s2_1: "-.(a * 0) = 0" - using adom mult_0 neg0 by auto - have s2_2: "a * -.0 = 0" - using adom neg0 mult_0 by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ b. \ - b \\in Nat; - -.(a * b) = a * -.b - \ \ - -.(a * Succ[b]) = a * -.Succ[b]" - proof - - fix "b" :: "c" - assume s1_3_bdom: "b \\in Nat" - assume s1_3_induct: "-.(a * b) = a * -.b" - have s2_1: "-.(a * Succ[b]) = -.(a * b + a)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_2: "... = -.(a * b) + -.a" - proof - - have s3_1: "a * b \\in Int" - using adom s1_3_bdom nat_in_int multIsInt by auto - have s3_2: "a \\in Int" - using adom by auto - show ?thesis - using s3_1 s3_2 MinusDistributesAdd by auto - qed - also have s2_3: "... = a * -.b + -.a" - using s1_3_induct by auto - also have s2_4: "... = a * -.Succ[b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - finally show "-.(a * Succ[b]) = a * -.Succ[b]" . - qed - have s1_4: "\ b. \ - b \\in negNat; - -.(a * b) = a * -.b - \ \ - -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" - proof - - fix "b" :: "c" - assume s1_4_bdom: "b \\in negNat" - assume s1_4_induct: "-.(a * b) = a * -.b" - have s2_1: "-.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_2: "... = -.(a * -.-.b) + -.-.a" - proof - - have s3_1: "a * -.-.b \\in Int" - proof - - have s4_1: "a \\in Int" - using adom by auto - have s4_2: "-.-.b \\in Int" - using s1_4_bdom minus_negnat_in_int - neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 MinusDistributesAdd by auto - qed - also have s2_3: "... = -.(a * -.-.b) + a" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_4: "... = -.(a * b) + a" - proof - - have s3_1: "b \\in Int" - using s1_4_bdom neg_nat_in_int by auto - have s3_2: "-.-.b = b" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_5: "... = a * -.b + a" - using s1_4_induct by auto - also have s2_6: "... = a * Succ[-.b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_7: "... = a * -.-.Succ[-.b]" - proof - - have s3_1: "Succ[-.b] \\in Int" - using s1_4_bdom negnat_succ_in_nat nat_in_int by auto - have s3_2: "-.-.Succ[-.b] = Succ[-.b]" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - finally show "-.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" . - qed - have s1_5: "\ b \ Nat: - -.(a * b) = a * -.b" - using s1_2 s1_3 natInduct[of "\ b. - -.(a * b) = a * -.b"] - by auto - have s1_6: "\ b \ negNat: - -.(a * b) = a * -.b" - using s1_2 s1_4 neg_nat_induction[of "\ b. - -.(a * b) = a * -.b"] - by auto - have s1_7: "\ b \ Int: - -.(a * b) = a * -.b" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -(* -THEOREM MinusCommutesLeftMult == - \A a, b \in Int: -.(a * b) = -.a * b -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE -.(a * b) = -.a * b - BY <1>1 -<1>2. -.(a * b) = -.(b * a) - BY <1>1, MultCommutative -<1>3. @ = b * -.a - BY <1>1, MinusCommutesRightMult -<1>4. @ = -.a * b - BY <1>1, neg_int_eq_int, MultCommutative -<1> QED - BY <1>2, <1>3, <1>4 -*) -theorem MinusCommutesLeftMult: - "\ a \ Int: \ b \ Int: - -.(a * b) = -.a * b" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \\in Int" - assume bdom: "b \\in Int" - have s1_2: "-.(a * b) = -.(b * a)" - using adom bdom MultCommutative_sequent by auto - also have s1_3: "... = b * -.a" - using adom bdom MinusCommutesRightMult by auto - also have s1_4: "... = -.a * b" - using adom bdom neg_int_eq_int - MultCommutative_sequent by auto - finally have "-.(a * b) = -.a * b" - by auto - } - from this show ?thesis - by auto - qed - - -(* -THEOREM MultAssociative == - \A a \in Int: \A b \in Int: \A c \in Int: - (a * b) * c = a * (b * c) -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE \A c \in Int: (a * b) * c = a * (b * c) - BY <1>1, allI -<1>2. (a * b) * 0 = a * (b * 0) - <2>1. (a * b) * 0 = 0 - <3>1. a * b \in Int - BY <1>1, multIsInt - <3> QED - BY <3>1, mult_0 - <2>2. a * (b * 0) = 0 - <3>1. b * 0 = 0 - BY <1>1, mult_0 - <3>2. a * 0 = 0 - BY <1>1, mult_0 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW c \in Nat, - (a * b) * c = a * (b * c) - PROVE (a * b) * Succ[c] = a * (b * Succ[c]) - <2>1. (a * b) * Succ[c] = (a * b) * c + (a * b) - <3>1. a * b \in Int - BY <1>1, multIsInt - <3>2. c \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>2. @ = a * (b * c) + (a * b) - BY <1>3 \* induction hypothesis - <2>3. @ = a * (b * c + b) - <3>1. a \in Int - BY <1>1 - <3>2. b * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. b \in Int - BY <1>1 - <3> QED - BY <3>1, <3>2, <3>3, MultLeftDistributesAdd - <2>4. @ = a * (b * Succ[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Nat - BY <1>3 - <3>3. b * Succ[c] = b * c + b - BY <3>1, <3>2, multint_succ_nat - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>4. ASSUME NEW c \in negNat, - (a * b) * c = a * (b * c) - PROVE (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c]) - <2>1. (a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b) - <3>1. a * b \in Int - BY <1>1, multIsInt - <3>2. -.c \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>2. @ = (a * b) * c + -.(a * b) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.c = c - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>3. @ = a * (b * c) + -.(a * b) - BY <1>4 \* induction hypothesis - <2>4. @ = a * (b * c) + a * -.b - <3>1. a \in Int - BY <1>1 - <3>2. b \in Int - BY <1>1 - <3>3. -.(a * b) = a * -.b - BY <3>1, <3>2, MinusCommutesRightMult - <3> QED - BY <3>3 - <2>5. @ = a * (b * c + -.b) - <3>1. a \in Int - BY <1>1 - <3>2. b * c \in Int - BY <1>1 <1>4, neg_nat_in_int, multIsInt - <3>3. -.b \in Int - BY <1>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, <3>3, MultLeftDistributesAdd - <2>6. @ = a * (b * -.-.c + -.b) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.c = c - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>7. @ = a * (b * -.Succ[-.c]) - <3>1. b \in Int - BY <1>1 - <3>2. -.c \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>3. b * -.Succ[-.c] = b * -.-.c + -.b - BY <3>1, <3>2, multint_succ_negnat - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, - <2>5, <2>6, <2>7 -<1>5. \A c \in Nat: (a * b) * c = a * (b * c) - BY <1>2, <1>3, NatInduction -<1>6. \A c \in negNat: (a * b) * c = a * (b * c) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MultAssociative: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - (a * b) * c = a * (b * c)" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \\in Int" - assume bdom: "b \\in Int" - have s1_2: "(a * b) * 0 = a * (b * 0)" - proof - - have s2_1: "(a * b) * 0 = 0" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - show ?thesis - using s3_1 mult_0 by auto - qed - have s2_2: "a * (b * 0) = 0" - proof - - have s3_1: "b * 0 = 0" - using bdom mult_0 by auto - have s3_2: "a * 0 = 0" - using adom mult_0 by auto - show ?thesis - using s3_1 s3_2 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ c. \ - c \\in Nat; - (a * b) * c = a * (b * c) - \ \ - (a * b) * Succ[c] = a * (b * Succ[c])" - proof - - fix "c" :: "c" - assume s1_3_cdom: "c \\in Nat" - assume s1_3_induct: "(a * b) * c = a * (b * c)" - have s2_1: "(a * b) * Succ[c] = (a * b) * c + (a * b)" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - have s3_2: "c \\in Nat" - using s1_3_cdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_2: "... = a * (b * c) + (a * b)" - using s1_3_induct by auto - also have s2_3: "... = a * (b * c + b)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b * c \\in Int" - using bdom s1_3_cdom nat_in_int multIsInt by auto - have s3_3: "b \\in Int" - using bdom by auto - show ?thesis - using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] by auto - qed - also have s2_4: "... = a * (b * Succ[c])" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "c \\in Nat" - using s1_3_cdom by auto - have s3_3: "b * Succ[c] = b * c + b" - using s3_1 s3_2 multint_succ_nat by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a * b) * Succ[c] = a * (b * Succ[c])" . - qed - have s1_4: "\ c. \ - c \\in negNat; - (a * b) * c = a * (b * c) - \ \ - (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" - proof - - fix "c" :: "c" - assume s1_4_cdom: "c \\in negNat" - assume s1_4_induct: "(a * b) * c = a * (b * c)" - have s2_1: "(a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b)" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - have s3_2: "-.c \\in Nat" - using s1_4_cdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_2: "... = (a * b) * c + -.(a * b)" - proof - - have s3_1: "c \\in Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "-.-.c = c" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_3: "... = a * (b * c) + -.(a * b)" - using s1_4_induct by auto - also have s2_4: "... = a * (b * c) + a * -.b" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "-.(a * b) = a * -.b" - using s3_1 s3_2 MinusCommutesRightMult by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = a * (b * c + -.b)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b * c \\in Int" - using bdom s1_4_cdom neg_nat_in_int multIsInt by auto - have s3_3: "-.b \\in Int" - using bdom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] - by auto - qed - also have s2_6: "... = a * (b * -.-.c + -.b)" - proof - - have s3_1: "c \\in Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "-.-.c = c" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_7: "... = a * (b * -.Succ[-.c])" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "-.c \\in Nat" - using s1_4_cdom minus_neg_nat_in_nat by auto - have s3_3: "b * -.Succ[-.c] = b * -.-.c + -.b" - using s3_1 s3_2 multint_succ_negnat by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" . - qed - have s1_5: "\ c \ Nat: - (a * b) * c = a * (b * c)" - using s1_2 s1_3 natInduct[of "\ c. - (a * b) * c = a * (b * c)"] - by auto - have s1_6: "\ c \ negNat: - (a * b) * c = a * (b * c)" - using s1_2 s1_4 neg_nat_induction[of "\ c. - (a * b) * c = a * (b * c)"] - by auto - have s1_7: "\ c \ Int: - (a * b) * c = a * (b * c)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - - -(*----------------------------------------------------------------------------*) -(* Properties of `leq` and `add` and `diff`. *) - - -theorem LeqIsBool: - "(m <= n) \\in BOOLEAN" - unfolding leq_def by auto - - -theorem LeqReflexive: - assumes ndom: "n \\in Int" - shows "n <= n" - proof - - have s1_1: "n + 0 = n" - using ndom add_0 by auto - have s1_2: "0 \\in Nat" - using zeroIsNat by auto - have s1_3: "\ r \ Nat: n + r = n" - using s1_1 s1_2 by auto - show ?thesis - unfolding leq_def - using s1_3 by auto - qed - - -theorem LeqTransitive: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m <= n" and nk: "n <= k" - shows "m <= k" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - (* PICK t \in Nat: n + t = k *) - define Q where "Q \ \ x. x \\in Nat \ n + x = k" - define t where "t \ CHOOSE x: Q(x)" - have s1_2: "t \\in Nat \ n + t = k" - proof - - have s2_1: "\\E x: Q(x)" - using nk by (auto simp: leq_def Q_def) - have s2_2: "Q(t)" - unfolding t_def - using s2_1 chooseI_ex[of Q] by auto - show ?thesis - using s2_2 by (auto simp: Q_def) - qed - have s1_3: "(m + r) + t = k" - using s1_1 s1_2 by auto - have s1_4: "m + (r + t) = k" - proof - - have s2_1: "m \\in Int" - using mdom by auto - have s2_2: "r \\in Int" - using s1_1 nat_in_int by auto - have s2_3: "t \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "(m + r) + t = m + (r + t)" - using s2_1 s2_2 s2_3 AddAssociative_sequent by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "r + t \\in Nat" - using s1_1 s1_2 nat_add_in_nat by auto - have s1_6: "\ q \ Nat: m + q = k" - using s1_4 s1_5 by auto - show ?thesis - unfolding leq_def - using s1_6 by auto - qed - - -theorem leq_0: - assumes ndom: "n \\in Int" - shows "(n \\in Nat) = (0 <= n)" - proof - - have s1_1: "n \\in Nat ==> 0 <= n" - proof - - assume s1_1_ndom: "n \\in Nat" - have s2_1: "0 + n = n" - using ndom add_0_left by auto - have s2_2: "\ r \ Nat: 0 + r = n" - using s2_1 s1_1_ndom by auto - show "0 <= n" - unfolding leq_def - using s2_2 by auto - qed - have s1_2: "0 <= n ==> n \\in Nat" - proof - - assume s1_2_leq: "0 <= n" - define P where "P \ \ x. x \\in Nat \ 0 + x = n" - define r where "r \ CHOOSE x: P(x)" - have s2_1: "r \\in Nat \ 0 + r = n" - proof - - have s3_1: "\\E x: P(x)" - using s1_2_leq by (auto simp: leq_def P_def) - have s3_2: "P(r)" - unfolding r_def - using s3_1 chooseI_ex[of P] by auto - show ?thesis - using s3_2 by (auto simp: P_def) - qed - have s2_2: "0 + r = r" - proof - - have s3_1: "r \\in Int" - using s2_1 nat_in_int by auto - show ?thesis - using s3_1 add_0_left by auto - qed - have s2_3: "r = n" - using s2_1 s2_2 by auto - show "n \\in Nat" - using s2_1 s2_3 by auto - qed - have s1_3: "(0 <= n) \\in BOOLEAN" - using LeqIsBool by auto - have s1_4: "(n \\in Nat) \\in BOOLEAN" - using inIsBool isBoolTrueFalse by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -theorem eq_leq_both: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m = n" - shows "m <= n \ n <= m" - proof - - have s1_1: "m <= n" - proof - - have s2_1: "m + 0 = m" - using mdom add_0 by auto - have s2_2: "m + 0 = n" - using s2_1 mn by auto - have s2_3: "0 \\in Nat" - using zeroIsNat by auto - have s2_4: "\ r \ Nat: m + r = n" - using s2_2 s2_3 by auto - show ?thesis - unfolding leq_def - using s2_4 by auto - qed - have s1_2: "n <= m" - proof - - have s2_1: "n + 0 = n" - using ndom add_0 by auto - have s2_2: "n + 0 = m" - using s2_1 mn by auto - have s2_3: "0 \\in Nat" - using zeroIsNat by auto - have s2_4: "\ r \ Nat: n + r = m" - using s2_2 s2_3 by auto - show ?thesis - unfolding leq_def - using s2_4 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* The inverse of the theorem `eq_leq_both`. *) -theorem leq_both_eq: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m <= n" and nm: "n <= m" - shows "m = n" - proof - - define P where "P \ \ x. x \\in Nat \ m + x = n" - define p where "p \ CHOOSE x: P(x)" - define Q where "Q \ \ x. x \\in Nat \ n + x = m" - define q where "q \ CHOOSE x: Q(x)" - have s1_1: "p \\in Nat \ m + p = n" - proof - - have s2_1: "\ x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(p)" - unfolding p_def using s2_1 chooseI_ex[of P] - by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "q \\in Nat \ n + q = m" - proof - - have s2_1: "\ x: Q(x)" - using nm by (auto simp: leq_def Q_def) - have s2_2: "Q(q)" - unfolding q_def using s2_1 chooseI_ex[of Q] - by auto - show ?thesis - using s2_2 by (auto simp: Q_def) - qed - have s1_3: "(m + p) + (n + q) = n + m" - using s1_1 s1_2 by auto - have s1_4: "p + q = 0" - proof - - have s2_1: "(m + p) + (n + q) = n + m" - using s1_3 by auto - have s2_2: "m + (p + (n + q)) = n + m" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "p \\in Int" - using s1_1 nat_in_int by auto - have s3_3: "n + q \\in Int" - using ndom s1_2 nat_in_int addIsInt by auto - have s3_4: "(m + p) + (n + q) = - m + (p + (n + q))" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - have s2_3: "m + (p + (q + n)) = n + m" - proof - - have s3_1: "n \\in Int" - using ndom by auto - have s3_2: "q \\in Int" - using s1_2 nat_in_int by auto - have s3_3: "n + q = q + n" - using s3_1 s3_2 AddCommutative_sequent - by auto - show ?thesis - using s2_2 s3_3 by auto - qed - have s2_4: "m + ((p + q) + n) = n + m" - proof - - have s3_1: "p \\in Int" - using s1_1 nat_in_int by auto - have s3_2: "q \\in Int" - using s1_2 nat_in_int by auto - have s3_3: "n \\in Int" - using ndom by auto - have s3_4: "p + (q + n) = (p + q) + n" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_3 s3_4 by auto - qed - have s2_5: "m + (n + (p + q)) = n + m" - proof - - have s3_1: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "(p + q) + n = n + (p + q)" - using s3_1 s3_2 - AddCommutative_sequent by auto - show ?thesis - using s2_4 s3_3 by auto - qed - have s2_6: "(m + n) + (p + q) = n + m" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_4: "m + (n + (p + q)) = (m + n) + (p + q)" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_5 s3_4 by auto - qed - have s2_7: "-.(m + n) + ((m + n) + (p + q)) = - -.(m + n) + (n + m)" - using s2_6 by auto - have s2_8: "(-.(m + n) + (m + n)) + (p + q) = - -.(m + n) + (n + m)" - proof - - have s3_1: "m + n \\in Int" - using mdom ndom addIsInt by auto - have s3_2: "-.(m + n) \\in Int" - using s3_1 neg_int_eq_int by auto - have s3_3: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_4: "-.(m + n) + ((m + n) + (p + q)) = - (-.(m + n) + (m + n)) + (p + q)" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_7 s3_4 by auto - qed - have s2_9: "(-.(m + n) + (m + n)) + (p + q) = - -.(m + n) + (m + n)" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "n + m = m + n" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s2_8 s3_3 by auto - qed - have s2_10: "0 + (p + q) = 0" - proof - - have s3_1: "-.(m + n) + (m + n) = 0" - proof - - have s4_1: "m + n \\in Int" - using mdom ndom addIsInt by auto - show ?thesis - using s4_1 AddNegCancels_left by auto - qed - show ?thesis - using s2_9 s3_1 by auto - qed - have s2_11: "p + q = 0" - proof - - have s3_1: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - show ?thesis - using s2_10 s3_1 add_0_left by auto - qed - show ?thesis - using s2_11 by auto - qed - have s1_5: "p = 0 \ q = 0" - using s1_1 s1_2 s1_4 nat_add_0 by auto - have s1_6: "m + 0 = n" - using s1_1 s1_5 by auto - show ?thesis - using s1_6 mdom add_0 by auto - qed - - -theorem eq_iff_leq_both: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "(m = n) = ((m <= n) \ (n <= m))" - using mdom ndom eq_leq_both leq_both_eq by auto - - -(* -THEOREM dichotomy_leq == - ASSUME NEW a \in Int, NEW b \in Int - PROVE (a <= b) \ (b <= a) -PROOF -<1>1. CASE (a - b) \in Nat - <2>1. b + (a - b) = a - BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b - <2>2. \E n \in Nat: b + n = a - BY <1>1, <2>1 - <2>3. b <= a - BY <2>2 DEF leq - <2> QED - BY <2>3 -<1>2. CASE (a - b) \notin Nat - <2>1. -.(a - b) \in Nat - <3>1. (a - b) \in Int - <4>1. a - b = a + -.b - BY DEF diff - <4>2. a \in Int - OBVIOUS - <4>3. -.b \in Int - BY b \in Int, neg_int_eq_int - <4>4. a + -.b \in Int - BY <4>2, <4>3, addIsInt - <4> QED - BY <4>1, <4>4 - <3> QED - BY <3>1, <1>2, neg_negative_in_nat - <2>2. (b - a) \in Nat - <3>1. -.(a - b) = b - a - BY a \in Int, b \in Int, MinusDiff - <3> QED - BY <2>1, <3>1 - <2>3. a + (b - a) = b - BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b - <2>4. \E n \in Nat: a + n = b - BY <2>2, <2>3 - <2>5. a <= b - BY <2>4 DEF leq - <2> QED - BY <2>5 -<1> QED - BY <1>1, <1>2 -*) -theorem dichotomy_leq: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(a <= b) \ (b <= a)" - proof - - have s1_1: "a - b \\in Nat ==> - (a <= b) \ (b <= a)" - proof - - assume s1_1_asm: "a - b \\in Nat" - have s2_1: "b + (a - b) = a" - using adom bdom a_plus_b_minus_a_eq_b by auto - have s2_2: "\ n \ Nat: b + n = a" - using s1_1_asm s2_1 by auto - have s2_3: "b <= a" - unfolding leq_def - using s2_2 by auto - show "(a <= b) \ (b <= a)" - using s2_3 by auto - qed - have s1_2: "a - b \\notin Nat ==> - (a <= b) \ (b <= a)" - proof - - assume s1_2_asm: "a - b \\notin Nat" - have s2_1: "-.(a - b) \\in Nat" - proof - - have s3_1: "(a - b) \\in Int" - proof - - have s4_1: "a - b = a + -.b" - unfolding diff_def by auto - have s4_2: "a \\in Int" - using adom by auto - have s4_3: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s4_4: "a + -.b \\in Int" - using s4_2 s4_3 addIsInt by auto - show ?thesis - using s4_1 s4_4 by auto - qed - show ?thesis - using s3_1 s1_2_asm neg_negative_in_nat by auto - qed - have s2_2: "(b - a) \\in Nat" - proof - - have s3_1: "-.(a - b) = b - a" - using adom bdom MinusDiff by auto - show ?thesis - using s2_1 s3_1 by auto - qed - have s2_3: "a + (b - a) = b" - using adom bdom a_plus_b_minus_a_eq_b by auto - have s2_4: "\ n \ Nat: a + n = b" - using s2_2 s2_3 by auto - have s2_5: "a <= b" - unfolding leq_def - using s2_4 by auto - show "(a <= b) \ (b <= a)" - using s2_5 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem trichotomy_less: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(a < b) \ (a = b) \ (b < a)" - using adom bdom dichotomy_leq by (auto simp: less_def) - - -theorem trichotomy_less_0: - assumes ndom: "n \\in Int" - shows "(n < 0) \ (n = 0) \ (0 < n)" - using ndom zero_in_int trichotomy_less by auto - - -(* -THEOREM leq_mult_monotonic == - \A m \in Int: \A n \in Int: - (m <= n) => (\A k \in Nat: k * m <= k * n) -PROOF -<1>1. SUFFICES - ASSUME NEW m \in Int, NEW n \in Int, NEW k \in Nat, - m <= n - PROVE k * m <= k * n - BY <1>1 -<1>2. PICK r \in Nat: m + r = n - BY <1>1 DEF leq -<1>3. k * (m + r) = k * n - BY <1>2 -<1>4. k * m + k * r = k * n - <2>1. k \in Int - BY <1>1, nat_in_int - <2>2. m \in Int - BY <1>1 - <2>3. r \in Int - BY <1>2, nat_in_int - <2>4. k * (m + r) = k * m + k * r - BY <2>1, <2>2, <2>3, MultLeftDistributesAdd - <2> QED - BY <1>3, <2>4 -<1>5. k * r \in Nat - BY <1>1, <1>2, nat_mult_in_nat -<1>6. \E x \in Nat: k * m + x = k * n - BY <1>4, <1>5 -<1> QED - BY <1>6 DEF leq -*) -theorem leq_mult_monotonic: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Nat" and mn: "m <= n" - shows "k * m <= k * n" - proof - - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_2: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_3: "k * (m + r) = k * n" - using s1_2 by auto - have s1_4: "k * m + k * r = k * n" - proof - - have s2_1: "k \\in Int" - using kdom nat_in_int by auto - have s2_2: "m \\in Int" - using mdom by auto - have s2_3: "r \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "k * (m + r) = k * m + k * r" - using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "k * r \\in Nat" - using kdom s1_2 nat_mult_in_nat by auto - have s1_6: "\ x \ Nat: k * m + x = k * n" - using s1_4 s1_5 by auto - show ?thesis - unfolding leq_def - using s1_6 by auto - qed - - -(* -THEOREM leq_mult_monotonic_neg: - \A m \in Int: \A n \in Int: - (m <= n) => (\A k \in negNat: k * n <= k * m) -PROOF -<1>1. SUFFICES - ASSUME NEW m \in Int, NEW n \in Int, NEW k \in negNat, - m <= n - PROVE k * n <= k * m - BY <1>1 -<1>2. PICK r \in Nat: m + r = n - BY <1>1 DEF leq -<1>3. k * (m + r) = k * n - BY <1>2 -<1>4. k * m + k * r = k * n - <2>1. k \in Int - BY <1>1, neg_nat_in_int - <2>2. m \in Int - BY <1>1 - <2>3. r \in Int - BY <1>2, nat_in_int - <2>4. k * (m + r) = k * m + k * r - BY <2>1, <2>2, <2>3, MultLeftDistributesAdd - <2> QED - BY <1>3, <2>4 -<1>5. k * m = k * n + -.(k * r) - <2>1. (k * m + k * r) + -.(k * r) = k * n + -.(k * r) - BY <1>4 - <2>2. k * m + (k * r + -.(k * r)) = k * n + -.(k * r) - <3>1. k * m \in Int - BY <1>1, neg_nat_in_int, multIsInt - <3>2. k * r \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. r \in Int - BY <1>2, nat_in_int - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.(k * r) \in Int - BY <3>2, neg_int_eq_int - <3>4. (k * m + k * r) + -.(k * r) = - k * m + (k * r + -.(k * r)) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <2>1, <3>4 - <2>3. k * r + -.(k * r) = 0 - <3>1. k * r \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. r \in Int - BY <1>2, nat_in_int - <4> QED - BY <4>1, <4>2, multIsInt - <3> QED - BY <3>1, AddNegCancels - <2>4. k * m + 0 = k * n + -.(k * r) - BY <2>2, <2>3 - <2>5. k * m + 0 = k * m - <3>1. k * m \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. m \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3> QED - BY <3>1, add_0 - <2> QED - BY <2>4, <2>5 -<1>6. k * n + (-.k * r) = k * m - <2>1. k * n + -.(k * r) = k * m - BY <1>5 \* symmetric - <2>2. -.(k * r) = -.k * r - <3>1. k \in Int - BY <1>1, neg_nat_in_int - <3>2. r \in Int - BY <1>2, nat_in_int - <3> QED - BY <3>1, <3>2, MinusCommutesLeftMult - <2> QED - BY <2>1, <2>2 -<1>7. -.k * r \in Nat - <2>1. -.k \in Nat - BY <1>1, minus_neg_nat_in_nat - <2>2. r \in Nat - BY <1>2 - <2> QED - BY <2>1, <2>2, nat_mult_in_nat -<1>8. \E x \in Nat: k * n + x = k * m - BY <1>6, <1>7 -<1> QED - BY <1>8 DEF leq -*) -theorem leq_mult_monotonic_neg: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in negNat" and mn: "m <= n" - shows "k * n <= k * m" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_2: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_3: "k * (m + r) = k * n" - using s1_2 by auto - have s1_4: "k * m + k * r = k * n" - proof - - have s2_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s2_2: "m \\in Int" - using mdom by auto - have s2_3: "r \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "k * (m + r) = k * m + k * r" - using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "k * m = k * n + -.(k * r)" - proof - - have s2_1: "(k * m + k * r) + -.(k * r) = k * n + -.(k * r)" - using s1_4 by auto - have s2_2: "k * m + (k * r + -.(k * r)) = k * n + -.(k * r)" - proof - - have s3_1: "k * m \\in Int" - using mdom kdom neg_nat_in_int multIsInt by auto - have s3_2: "k * r \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.(k * r) \\in Int" - using s3_2 neg_int_eq_int by auto - have s3_4: "(k * m + k * r) + -.(k * r) = - k * m + (k * r + -.(k * r))" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - have s2_3: "k * r + -.(k * r) = 0" - proof - - have s3_1: "k * r \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - show ?thesis - using s3_1 AddNegCancels by auto - qed - have s2_4: "k * m + 0 = k * n + -.(k * r)" - using s2_2 s2_3 by auto - have s2_5: "k * m + 0 = k * m" - proof - - have s3_1: "k * m \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "m \\in Int" - using mdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - show ?thesis - using s3_1 add_0 by auto - qed - show ?thesis - using s2_4 s2_5 by auto - qed - have s1_6: "k * n + (-.k * r) = k * m" - proof - - have s2_1: "k * n + -.(k * r) = k * m" - using s1_5 by auto - have s2_2: "-.(k * r) = -.k * r" - proof - - have s3_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s3_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s3_1 s3_2 MinusCommutesLeftMult by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_7: "-.k * r \\in Nat" - proof - - have s2_1: "-.k \\in Nat" - using kdom minus_neg_nat_in_nat by auto - have s2_2: "r \\in Nat" - using s1_2 by auto - show ?thesis - using s2_1 s2_2 nat_mult_in_nat by auto - qed - have s1_8: "\ x \ Nat: k * n + x = k * m" - using s1_6 s1_7 by auto - show ?thesis - unfolding leq_def - using s1_8 by auto - qed - - -(*----------------------------------------------------------------------------*) - - -(* Monotonicity of addition with respect to order. *) -theorem leq_add_monotonic_right: - assumes mdom: "m \ Int" and ndom: "n \ Int" and - kdom: "k \ Int" and mn: "m <= n" - shows "m + k <= n + k" - proof - - have s1_1: "\ r: r \ Nat \ m + r = n" - using mn by (auto simp: leq_def) - define r where "r \ CHOOSE r: r \ Nat \ m + r = n" - have s1_2: "r \ Nat \ m + r = n" - unfolding r_def - by (rule chooseI_ex, rule s1_1) - have s1_3: "(m + r) + k = n + k" - using s1_2 by auto - have s1_4: "(m + k) + r = n + k" - proof - - have s2_1: "(r + m) + k = n + k" - proof - - have s3_1: "m \ Int" - using mdom by auto - have s3_2: "r \ Int" - using s1_2 nat_in_int by auto - have s3_3: "m + r = r + m" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s1_3 s3_3 by auto - qed - have s2_2: "r + (m + k) = n + k" - proof - - have s3_1: "r \ Int" - using s1_2 nat_in_int by auto - have s3_2: "m \ Int" - using mdom by auto - have s3_3: "k \ Int" - using kdom by auto - have s3_4: "(r + m) + k = r + (m + k)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - show ?thesis - proof - - have s3_1: "r \ Int" - using s1_2 nat_in_int by auto - have s3_2: "m + k \ Int" - using mdom kdom addIsInt by auto - have s3_3: "r + (m + k) = (m + k) + r" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s2_2 s3_3 by auto - qed - qed - have s1_5: "\ t \ Nat: - (m + k) + t = n + k" - using s1_2 s1_4 by auto - show "m + k <= n + k" - using s1_5 by (auto simp: leq_def) - qed - - -theorem eq_add_inj_right: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m + k = n + k" - shows "m = n" - proof - - have s1_1: "(m + k) + -.k = (n + k) + -.k" - using mn by auto - have s1_2: "m + (k + -.k) = n + (k + -.k)" - proof - - have minus_kdom: "-.k \\in Int" - using kdom neg_int_eq_int by auto - have s2_1: "(m + k) + -.k = m + (k + -.k)" - using mdom kdom minus_kdom AddAssociative_sequent - by auto - have s2_2: "(n + k) + -.k = n + (k + -.k)" - using ndom kdom minus_kdom AddAssociative_sequent - by auto - show ?thesis - using s1_1 s2_1 s2_2 by auto - qed - have s1_3: "m + 0 = n + 0" - proof - - have s2_1: "k + -.k = 0" - using kdom AddNegCancels_sequent by auto - show ?thesis - using s1_2 s2_1 by auto - qed - have s1_4: "m + 0 = m" - using mdom add_0 by auto - have s1_5: "n + 0 = n" - using ndom add_0 by auto - show ?thesis - using s1_3 s1_4 s1_5 by auto - qed - - -theorem less_add_monotonic_right: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m < n" - shows "m + k < n + k" - proof - - have s1_1: "m + k <= n + k" - proof - - have s2_1: "m <= n" - using mn by (auto simp: less_def) - show ?thesis - using mdom ndom kdom s2_1 leq_add_monotonic_right by auto - qed - have s1_2: "m + k \ n + k" - proof - - have s2_1: "m + k = n + k ==> FALSE" - proof - - assume s2_1_asm: "m + k = n + k" - have s3_1: "m = n" - using s2_1_asm mdom ndom kdom eq_add_inj_right - by auto - have s3_2: "m \ n" - using mn by (auto simp: less_def) - show "FALSE" - using s3_1 s3_2 by auto - qed - show ?thesis - using s2_1 by auto - qed - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_add_monotonic_left: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m < n" - shows "k + m < k + n" - proof - - have s1_1: "m + k < n + k" - using mdom ndom kdom mn less_add_monotonic_right by auto - have s1_2: "m + k = k + m" - using mdom kdom AddCommutative_sequent by auto - have s1_3: "n + k = k + n" - using ndom kdom AddCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed - - -theorem minus_less: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m < n" - shows "-.n < -.m" - proof - - have s1_1: "m + -.m < n + -.m" - proof - - have s2_1: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using mdom s2_1 ndom mn less_add_monotonic_right by auto - qed - have s1_2: "0 < n + -.m" - proof - - have s2_1: "m + -.m = 0" - using mdom AddNegCancels_sequent by auto - show ?thesis - using s1_1 s2_1 by auto - qed - have s1_3: "-.n + 0 < -.n + (n + -.m)" - proof - - have s2_1: "-.n \\in Int" - using ndom neg_int_eq_int by auto - have s2_2: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s2_3: "n + -.m \\in Int" - proof - - have s3_1: "n \\in Int" - using ndom by auto - have s3_2: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 addIsInt by auto - qed - show ?thesis - using s1_2 s2_1 s2_2 s2_3 less_add_monotonic_left - by auto - qed - have s1_4: "-.n < -.n + (n + -.m)" - proof - - have s2_1: "-.n + 0 = -.n" - using ndom neg_int_eq_int add_0 by auto - show ?thesis - using s1_3 s2_1 by auto - qed - have s1_5: "-.n + (n + -.m) = -.m" - proof - - have s2_1: "-.n + (n + -.m) = (-.n + n) + -.m" - using ndom mdom neg_int_eq_int AddAssociative_sequent - by auto - have s2_2: "-.n + n = 0" - using ndom AddNegCancels_left by auto - have s2_3: "(-.n + n) + -.m = 0 + -.m" - using s2_2 by auto - have s2_4: "0 + -.m = -.m" - proof - - have s3_1: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s3_1 add_0_left by auto - qed - show ?thesis - using s2_1 s2_3 s2_4 by auto - qed - show ?thesis - using s1_4 s1_5 by auto - qed - - -theorem leq_diff_add: - assumes ndom: "n \ Int" and mdom: "m \ Int" and - kdom: "k \ Int" and nmk: "n - m <= k" - shows "n <= k + m" - proof - - define P where "P \ \ x. x \\in Nat \ (n - m) + x = k" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ (n - m) + r = k" - proof - - have s2_1: "\\E x: P(x)" - using nmk by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "m + ((n - m) + r) = m + k" - using s1_1 by auto - have s1_3: "(m + -.m) + (n + r) = k + m" - proof - - have s2_1: "m + ((n - m) + r) = (m + -.m) + (n + r)" - proof - - have s3_1: "m + ((n - m) + r) = - m + ((n + -.m) + r)" - unfolding diff_def by auto - also have s3_2: "... = m + ((-.m + n) + r)" - using mdom neg_int_eq_int ndom - AddCommutative_sequent by auto - also have s3_3: "... = (m + (-.m + n)) + r" - using mdom neg_int_eq_int ndom s1_1 nat_in_int - addIsInt AddAssociative_sequent by auto - also have s3_4: "... = ((m + -.m) + n) + r" - using mdom neg_int_eq_int ndom - AddAssociative_sequent by auto - also have s3_5: "... = (m + -.m) + (n + r)" - using mdom neg_int_eq_int addIsInt ndom - s1_1 nat_in_int AddAssociative_sequent by auto - finally show ?thesis . - qed - have s2_2: "m + k = k + m" - using mdom kdom AddCommutative_sequent by auto - show ?thesis - using s1_2 s2_1 s2_2 by auto - qed - have s1_4: "n + r = k + m" - proof - - have s2_1: "m + -.m = 0" - using mdom AddNegCancels_sequent by auto - have s2_2: "(m + -.m) + (n + r) = n + r" - proof - - have s3_1: "(m + -.m) + (n + r) = 0 + (n + r)" - using s2_1 by auto - have s3_2: "0 + (n + r) = n + r" - proof - - have s4_1: "n + r \\in Int" - using ndom s1_1 nat_in_int addIsInt by auto - show ?thesis - using s4_1 add_0_left by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - show ?thesis - using s1_3 s2_2 by auto - qed - have s1_5: "\ x \ Nat: n + x = k + m" - using s1_1 s1_4 by auto - show ?thesis - unfolding leq_def - using s1_5 by auto - qed - - -theorem leq_diff_nat: - assumes mdom: "m \\in Int" and ndom: "n \\in Nat" - shows "m - n <= m" - proof - - have s1_1: "(m - n) + n = m" - proof - - have s2_1: "-.n + n = 0" - proof - - have s3_1: "n \\in Int" - using ndom nat_in_int by auto - show ?thesis - using s3_1 AddNegCancels_left by auto - qed - have s2_2: "m + (-.n + n) = m + 0" - using s2_1 by auto - have s2_3: "(m + -.n) + n = m" - proof - - have s3_1: "m + (-.n + n) = (m + -.n) + n" - proof - - have s4_1: "m \\in Int" - using mdom by auto - have s4_2: "-.n \\in Int" - using ndom minus_nat_in_int by auto - have s4_3: "n \\in Int" - using ndom nat_in_int by auto - show ?thesis - using s4_1 s4_2 s4_3 AddAssociative_sequent by auto - qed - have s3_2: "m + 0 = m" - using mdom add_0 by auto - show ?thesis - using s2_2 s3_1 s3_2 by auto - qed - show ?thesis - unfolding diff_def - using s2_3 by auto - qed - have s1_2: "\ r \ Nat: (m - n) + r = m" - using ndom s1_1 by auto - show ?thesis - unfolding leq_def - using s1_2 by auto - qed - - -(* See also the theorem `less_leq_trans`. *) -theorem leq_less_trans: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - and kdom: "k \\in Int" and - mn: "m <= n" and nk: "n < k" - shows "m < k" - proof - - have s1_1: "(n <= k) \ (n \ k)" - using nk by (auto simp: less_def) - have s1_2: "m <= k" - proof - - have s2_2: "n <= k" - using s1_1 by auto - have s2_3: "m <= n" - using mn by auto - show ?thesis - using s2_2 s2_3 mdom ndom kdom LeqTransitive - by iprover - qed - have s1_3: "m \ k" - proof - - { - assume s3_1: "m = k" - have s3_2: "k <= n" - using mn s3_1 by auto - have s3_3: "n <= k" - using nk by (auto simp: less_def) - have s3_4: "n = k" - using s3_2 s3_3 ndom kdom leq_both_eq - by auto - have "FALSE" - using s3_4 s1_1 by auto - } - from this show "m \ k" - by auto - qed - show ?thesis - unfolding less_def - using s1_2 s1_3 by auto - qed - - -(* See also the theorem `leq_less_trans`. *) -theorem less_leq_trans: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - and kdom: "k \\in Int" and - mn: "m < n" and nk: "n <= k" - shows "m < k" - proof - - have s1_1: "(m <= n) \ (m \ n)" - using mn by (auto simp: less_def) - have s1_2: "m <= k" - proof - - have s2_1: "m <= n" - using s1_1 by auto - have s2_2: "n <= k" - using nk by auto - show ?thesis - using s2_1 s2_2 mdom ndom kdom LeqTransitive - by iprover - qed - have s1_3: "m \ k" - proof - - have s2_1: "m = k ==> FALSE" - proof - - assume s2_1_asm: "m = k" - have s3_1: "n <= m" - using nk s2_1_asm by auto - have s3_2: "m <= n" - using mn by (auto simp: less_def) - have s3_3: "m = n" - using s3_1 s3_2 mdom ndom leq_both_eq - by auto - show "FALSE" - using s3_3 s1_1 by auto - qed - show ?thesis - using s2_1 by auto - qed - show ?thesis - unfolding less_def - using s1_2 s1_3 by auto - qed - - -theorem less_not_leq: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "(m < n) = (~ (n <= m))" - proof - - have s1_1: "m < n ==> ~ (n <= m)" - proof - - assume mn: "m < n" - have s2_1: "(m <= n) \ (m \ n)" - using mn by (auto simp: less_def) - have s2_2: "n <= m ==> FALSE" - proof - - assume nm: "n <= m" - have s3_1: "m = n" - using mdom ndom s2_1 nm leq_both_eq by blast - show "FALSE" - using s2_1 s3_1 by auto - qed - show "~ (n <= m)" - using s2_2 by auto - qed - have s1_2: "(~ (n <= m)) ==> m < n" - proof - - assume nm: "~ (n <= m)" - have s2_1: "(m <= n) \ (n <= m)" - using mdom ndom dichotomy_leq by auto - have s2_2: "m <= n" - using nm s2_1 by auto - have s2_3: "m = n ==> FALSE" - proof - - assume eq: "m = n" - have s3_1: "n <= n" - using ndom LeqReflexive by auto - have s3_2: "n <= m" - using eq s3_1 by auto - show "FALSE" - using nm s3_2 by auto - qed - have s2_4: "m \ n" - using s2_3 by auto - show "m < n" - unfolding less_def - using s2_2 s2_4 by auto - qed - have s1_3: "(~ (n <= m)) \\in BOOLEAN" - by auto - have s1_4: "(m < n) \\in BOOLEAN" - unfolding less_def by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -theorem leq_linear: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(~ (a <= b)) \ (a = b) \ (~ (b <= a))" - using adom bdom trichotomy_less[of "a" "b"] - less_not_leq[of "a" "b"] - less_not_leq[of "b" "a"] by auto - - -theorem leq_geq_linear: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(~ (b >= a)) \ (a = b) \ (~ (b <= a))" - using adom bdom leq_linear by auto - - -(* -THEOREM less_is_leq_plus_one == - ASSUME NEW m \in Int, NEW n \in Int, m < n - PROVE m + 1 <= n -PROOF -<1>1. PICK r \in Nat: m + r = n - BY m < n DEF less, leq -<1>2. PICK k \in Nat: r = Succ[k] - <2>1. r = 0 \/ \E k \in Nat: r = Succ[k] - BY <1>1, nat_0_succ - <2>2. ASSUME r = 0 - PROVE FALSE - <3>1. m + 0 = n - BY <1>1, <2>2 - <3>2. m = n - BY <3>1, m \in Int, add_0 - <3>3. m # n - BY m < n DEF less - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2 -<1>3. m + Succ[k] = n - BY <1>1, <1>2 -<1>4. m + (k + 1) = n - <2>1. Succ[k] = k + 1 - BY <1>2, nat_add_1 - <2> QED - BY <1>3, <2>1 -<1>5. m + (1 + k) = n - <2>1. k \in Int - BY <1>2, nat_in_int - <2>2. 1 \in Int - BY oneIsNat, nat_in_int - <2>3. k + 1 = 1 + k - BY <2>1, <2>2, AddCommutative - <2> QED - BY <1>4, <2>3 -<1>6. (m + 1) + k = n - <2>1. m \in Int - OBVIOUS - <2>2. 1 \in Int - BY oneIsNat nat_in_int - <2>3. k \in Int - BY <1>2, nat_in_int - <2>4. m + (1 + k) = (m + 1) + k - BY <2>1, <2>2, <2>3, AddAssociative - <2> QED - BY <1>5, <2>4 -<1>7. \E q \in Nat: (m + 1) + q = n - BY <1>2, <1>6 -<1> QED - BY <1>7 DEF leq -*) -theorem less_is_leq_plus_one: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m < n" - shows "m + 1 <= n" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: less_def leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - (* PICK k \in Nat: r = Succ[k] *) - define Q where "Q \ \ x. x \\in Nat \ r = Succ[x]" - define k where "k \ CHOOSE x: Q(x)" - have s1_2: "k \\in Nat \ r = Succ[k]" - proof - - have s2_1: "r = 0 \ - (\ x \ Nat: r = Succ[x])" - using s1_1 nat_0_succ by auto - have s2_2: "r = 0 ==> FALSE" - proof - - assume s2_2_asm: "r = 0" - have s3_1: "m + 0 = n" - using s1_1 s2_2_asm by auto - have s3_2: "m = n" - using s3_1 mdom add_0 by auto - have s3_3: "m \ n" - using mn by (auto simp: less_def) - show ?thesis - using s3_2 s3_3 by auto - qed - have s2_3: "\ x \ Nat: r = Succ[x]" - using s2_1 s2_2 by auto - have s2_4: "\\E x: Q(x)" - unfolding Q_def using s2_3 by auto - have s2_5: "Q(k)" - unfolding k_def - using s2_4 chooseI_ex[of Q] by auto - show ?thesis - using s2_5 by (auto simp: Q_def) - qed - have s1_3: "m + Succ[k] = n" - using s1_1 s1_2 by auto - have s1_4: "m + (k + 1) = n" - proof - - have s2_1: "Succ[k] = k + 1" - using s1_2 nat_add_1 by fast - show ?thesis - using s1_3 s2_1 by auto - qed - have s1_5: "m + (1 + k) = n" - proof - - have s2_1: "k \\in Int" - using s1_2 nat_in_int by auto - have s2_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s2_3: "k + 1 = 1 + k" - using s2_1 s2_2 AddCommutative_sequent by auto - show ?thesis - using s1_4 s2_3 by auto - qed - have s1_6: "(m + 1) + k = n" - proof - - have s2_1: "m \\in Int" - using mdom by auto - have s2_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s2_3: "k \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "m + (1 + k) = (m + 1) + k" - using s2_1 s2_2 s2_3 AddAssociative_sequent by auto - show ?thesis - using s1_5 s2_4 by auto - qed - have s1_7: "\ q \ Nat: (m + 1) + q = n" - using s1_2 s1_6 by auto - show ?thesis - unfolding leq_def using s1_7 by auto - qed - - -(* -THEOREM less_succ_nat == - ASSUME NEW n \in Nat - PROVE n < Succ[n] -PROOF -<1>1. n <= Succ[n] - <2>1. Succ[n] = n + 1 - BY n \in Nat, nat_add_1 - <2>2. 1 \in Nat - BY oneIsNat - <2>3. \E r \in Nat: n + r = Succ[n] - BY <2>1, <2>2 - <2> QED - BY <2>3 DEF leq -<1>2. n # Succ[n] - BY n \in Nat, succIrrefl -<1> QED - BY <1>1, <1>2 DEF less -*) -theorem less_succ_nat: - assumes ndom: "n \\in Nat" - shows "n < Succ[n]" - proof - - have s1_1: "n <= Succ[n]" - proof - - have s2_1: "Succ[n] = n + 1" - using ndom nat_add_1 by fast - have s2_2: "1 \\in Nat" - using oneIsNat by auto - have s2_3: "\\E r \\in Nat: n + r = Succ[n]" - using s2_1 s2_2 by auto - show ?thesis - unfolding leq_def - using s2_3 by auto - qed - have s1_2: "n \ Succ[n]" - using ndom succIrrefl by auto - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_pred_nat: - assumes ndom: "n \\in Nat \ {0}" - shows "Pred[n] < n" - proof - - have s1_1: "Pred[n] <= n" - proof - - have s2_1: "Pred[n] \\in Nat" - using ndom Pred_in_nat by auto - have s2_2: "Succ[Pred[n]] = n" - using ndom Succ_Pred_nat by auto - have s2_3: "Pred[n] + 1 = n" - proof - - have s3_1: "Succ[Pred[n]] = Pred[n] + 1" - using s2_1 nat_add_1 by fast - show ?thesis - using s2_2 s3_1 by auto - qed - have s2_4: "1 \\in Nat" - using oneIsNat by auto - have s2_5: "\\E x \\in Nat: Pred[n] + x = n" - using s2_3 s2_4 by auto - show ?thesis - unfolding leq_def - using s2_5 by auto - qed - have s1_2: "Pred[n] \ n" - using ndom pred_irrefl by auto - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_isucc: - assumes idom: "i \\in Int" - shows "i < iSucc[i]" - proof - - have s1_1: "i \\in Nat ==> i < iSucc[i]" - proof - - assume s1_1_asm: "i \\in Nat" - have s2_1: "i < Succ[i]" - using s1_1_asm less_succ_nat by auto - have s2_2: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using idom s1_1_asm by auto - show "i < iSucc[i]" - using s2_1 s2_2 by auto - qed - have s1_2: "i \\notin Nat ==> i < iSucc[i]" - proof - - assume s1_2_asm: "i \\notin Nat" - have s2_1: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using idom s1_2_asm by auto - have s2_2: "-.i \\in Nat \ {0}" - using idom s1_2_asm minus_neg_int_in_nat by auto - have s2_3: "Pred[-.i] < -.i" - using s2_2 less_pred_nat by auto - have s2_4: "-.-.i < -.Pred[-.i]" - proof - - have s3_1: "Pred[-.i] \\in Int" - using s2_2 Pred_in_nat nat_in_int by auto - have s3_2: "-.i \\in Int" - using s2_2 nat_in_int by auto - show ?thesis - using s2_3 s3_1 s3_2 minus_less by auto - qed - have s2_5: "-.-.i = i" - using idom minus_involutive by auto - show "i < iSucc[i]" - using s2_1 s2_4 s2_5 by auto - qed - show ?thesis - using s1_1 s1_2 by iprover - qed - - -theorem less_i_add_1: - assumes idom: "i \\in Int" - shows "i < i + 1" - proof - - have s1_1: "i < iSucc[i]" - using idom less_isucc by auto - have s1_2: "iSucc[i] = i + 1" - using idom isucc_eq_add_1 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM i_less_j_nat == - ASSUME NEW i \in Int, NEW j \in Nat, i <= j - PROVE i < Succ[j] -PROOF -<1>1. j < Succ[j] - BY j \in Nat, less_succ_nat -<1>2. j \in Int - BY j \in Nat, nat_in_int -<1>3. Succ[j] \in Int - BY j \in Nat, succIsNat, nat_in_int -<1> QED - BY i \in Int, <1>2, <1>3, i <= j, <1>1, - leq_less_trans -*) -theorem i_less_succ_j_nat: - assumes idom: "i \\in Int" and jdom: "j \\in Nat" and - ij: "i <= j" - shows "i < Succ[j]" - proof - - have s1_1: "j < Succ[j]" - using jdom less_succ_nat by auto - have s1_2: "j \\in Int" - using jdom nat_in_int by auto - have s1_3: "Succ[j] \\in Int" - using jdom succIsNat nat_in_int by auto - show ?thesis - using idom s1_2 s1_3 ij s1_1 leq_less_trans by auto - qed - - -theorem less_add_1: - assumes idom: "i \\in Int" and jdom: "j \\in Int" and - ij: "i <= j" - shows "i < j + 1" - proof - - have s1_1: "j < j + 1" - using jdom less_i_add_1 by auto - have s1_2: "j + 1 \\in Int" - proof - - have s2_1: "1 \\in Int" - using oneIsNat nat_in_int by auto - show ?thesis - using jdom s2_1 addIsInt by auto - qed - show ?thesis - using idom jdom s1_2 ij s1_1 leq_less_trans by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Intervals of integers. *) - - -(* -THEOREM SplitNat0 == - Nat = {0} \cup {n + 1: n \in Nat} -PROOF -<1>1. ASSUME NEW n \in Nat - PROVE n \in {0} \cup {k + 1: k \in Nat} - <2>1. n = 0 \/ \E m \in Nat: n = Succ[m] - BY <1>1, nat_0_succ - <2>2. (\E m \in Nat: n = Succ[m]) <=> - n \in {Succ[m]: m \in Nat} - OBVIOUS - <2>3. ASSUME NEW m \in Nat - PROVE Succ[m] = m + 1 - BY <2>3, nat_add_1 - <2>4. {Succ[m]: m \in Nat} = - {m + 1: m \in Nat} - BY <2>3 - <2>5. n \in {0} \/ n \in {m + 1: m \in Nat} - BY <2>1, <2>2, <2>4 - <2> QED - BY <2>5 -<1>2. ASSUME NEW n \in {0} \cup {k + 1: k \in Nat} - PROVE n \in Nat - <2>1. CASE n = 0 - BY <2>1, zeroIsNat - <2>2. CASE n \in {k + 1: k \in Nat} - <3>1. PICK k \in Nat: n = k + 1 - BY <2>2 - <3>2. k + 1 \in Nat - BY <3>1, oneIsNat, nat_add_in_nat - <3> QED - BY <3>1, <3>2 - <2> QED - BY <1>2, <2>1, <2>2 -<1> QED - BY <1>1, <1>2 -*) -theorem SplitNat0: - "Nat = {0} \ {n + 1: n \\in Nat}" - proof - - have s1_1: "\\A n \\in Nat: n \\in {0} \ {k + 1: k \\in Nat}" - proof - - { - fix "n" :: "c" - assume s1_1_asm: "n \\in Nat" - have s2_1: "n = 0 \ (\\E m \\in Nat: n = Succ[m])" - using s1_1_asm nat_0_succ by auto - have s2_2: "\\E m \\in Nat: n = Succ[m] ==> - n \\in {Succ[m]: m \\in Nat}" - by auto - have s2_3: "{Succ[m]: m \\in Nat} = {m + 1: m \\in Nat}" - proof - - have s3_1: "\\A y \\in {Succ[m]: m \\in Nat}: - y \\in {m + 1: m \\in Nat}" - proof - - { - fix "y" :: "c" - assume s4_1: "y \\in {Succ[m]: m \\in Nat}" - have s4_2: "\\E m \\in Nat: y = Succ[m]" - using s4_1 setOfAll by auto - have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" - using s4_2 by (auto simp: bEx_def) - (* PICK r \in Nat: y = Succ[r] *) - define P where "P \ - \ x. x \\in Nat \ y = Succ[x]" - define r where "r \ CHOOSE x: P(x)" - have s4_4: "r \\in Nat \ y = Succ[r]" - proof - - have s5_1: "\\E x: P(x)" - using s4_3 by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_5: "r \\in Nat \ y = r + 1" - proof - - have s5_1: "r \\in Nat" - using s4_4 by auto - have s5_2: "Succ[r] = r + 1" - using s5_1 nat_add_1 by fast - have s5_3: "y = Succ[r]" - using s4_4 by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - have "y \\in {m + 1: m \\in Nat}" - using s4_5 setOfAll_eqI by auto - } - from this show ?thesis - by auto - qed - have s3_2: "\\A y \\in {m + 1: m \\in Nat}: - y \\in {Succ[m]: m \\in Nat}" - proof - - { - fix "y" :: "c" - assume s4_1: "y \\in {m + 1: m \\in Nat}" - have s4_2: "\\E m \\in Nat: y = m + 1" - using s4_1 setOfAll by auto - have s4_3: "\\E m: m \\in Nat \ y = m + 1" - using s4_2 by (auto simp: bEx_def) - (* PICK r \in Nat: y = r + 1 *) - define P where "P \ - \ x. x \\in Nat \ y = x + 1" - define r where "r \ CHOOSE x: P(x)" - have s4_4: "r \\in Nat \ y = r + 1" - proof - - have s5_1: "\\E x: P(x)" - using s4_3 by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_5: "r \\in Nat \ y = Succ[r]" - proof - - have s5_1: "r \\in Nat" - using s4_4 by auto - have s5_2: "Succ[r] = r + 1" - using s5_1 nat_add_1 by fast - have s5_3: "y = r + 1" - using s4_4 by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - have "y \\in {Succ[m]: m \\in Nat}" - using s4_5 setOfAll_eqI by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s3_1 s3_2 extension - by (auto simp: subset_def) - qed - have s2_4: "n \\in {0} \ n \\in {m + 1: m \\in Nat}" - using s2_1 s2_2 s2_3 by auto - have s2_5: "n \\in {0} \ {k + 1: k \\in Nat}" - using s2_4 by auto - } - from this show ?thesis - by auto - qed - have s1_2: "\\A n \\in {0} \ {k + 1: k \\in Nat}: n \\in Nat" - proof - - { - fix "n" :: "c" - assume s1_2_asm: "n \\in {0} \ {k + 1: k \\in Nat}" - have s2_1: "n = 0 ==> n \\in Nat" - proof - - assume s2_1_asm: "n = 0" - show "n \\in Nat" - using s2_1_asm zeroIsNat by auto - qed - have s2_2: "n \\in {k + 1: k \\in Nat} ==> n \\in Nat" - proof - - assume s2_2_asm: "n \\in {k + 1: k \\in Nat}" - have s3_1: "\\E k \\in Nat: n = k + 1" - using s2_2_asm by auto - have s3_2: "\\E k: k \\in Nat \ n = k + 1" - using s3_1 by auto - (* PICK k \in Nat: n = k + 1 *) - define P where "P \ - \ x. x \\in Nat \ n = x + 1" - define k where "k \ CHOOSE x: P(x)" - have s3_3: "k \\in Nat \ n = k + 1" - proof - - have s4_1: "\\E x: P(x)" - using s3_2 by (auto simp: P_def) - have s4_2: "P(k)" - unfolding k_def - using s4_1 chooseI_ex[of P] by auto - show ?thesis - using s4_2 by (auto simp: P_def) - qed - have s3_4: "k + 1 \\in Nat" - using s3_3 oneIsNat nat_add_in_nat by auto - show "n \\in Nat" - using s3_3 s3_4 by auto - qed - have s2_3: "n \\in Nat" - using s1_2_asm s2_1 s2_2 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_1 s1_2 extension - by blast - qed - - -(* -THEOREM NatIsAdd0 == - Nat = {k + 0: k \in Nat} -PROOF -<1>1. Nat = {k: k \in Nat} - OBVIOUS -<1>2. ASSUME NEW k \in Nat - PROVE k + 0 = k - BY <1>2, nat_in_int, add_0 -<1>3. {k: k \in Nat} = - {k + 0: k \in Nat} - BY <1>2 -<1> QED - BY <1>1, <1>3 -*) -theorem NatIsAdd0: - "Nat = {k + 0: k \\in Nat}" - proof - - have s1_1: "Nat = {k: k \\in Nat}" - by auto - have s1_2: "{k: k \\in Nat} = {k + 0: k \\in Nat}" - using nat_in_int add_0 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM SplitAddi == - ASSUME NEW i \in Nat - PROVE {k + i: k \in Nat} = - {i} \cup {k + Succ[i]: k \in Nat} -PROOF -<1>1. n \in {k + i: k \in Nat} - = \E k \in Nat: n = k + i - OBVIOUS -<1>2. @ = \E k \in {0} \cup {q + 1: q \in Nat}: - n = k + i - BY SplitNat0 -<1>3. @ = \/ \E k \in {0}: n = k + i - \/ \E k \in {q + 1: q \in Nat}: n = k + i - OBVIOUS -<1>4. @ = \/ n = 0 + i - \/ \E k: /\ k \in {q + 1: q \in Nat} - /\ n = k + i - OBVIOUS -<1>5. @ = \/ n = i - \/ \E k: /\ \E q \in Nat: k = q + 1 - /\ n = k + i - BY i \in Nat, nat_in_int, add_0_left -<1>6. @ = \/ n = i - \/ \E k: \E q \in Nat: /\ k = q + 1 - /\ n = k + i - OBVIOUS -<1>7. @ = \/ n = i - \/ \E q \in Nat: n = (q + 1) + i - OBVIOUS -<1>8. @ = \/ n \in {i} - \/ n \in {q + (i + 1): q \in Nat} - BY AddCommutative, AddAssociative, - oneIsNat, nat_in_int -<1>9. @ = \/ n \in {i} - \/ n \in {q + Succ[i]: q \in Nat} - BY nat_add_1 -<1>10. @ = n \in {i} \cup {q + Succ[i]: q \in Nat} - OBVIOUS -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5, - <1>6, <1>7, <1>8, <1>9, <1>10 -*) -theorem SplitAddi: - assumes idom: "i \\in Nat" - shows "{k + i: k \\in Nat} = - ({i} \ {k + Succ[i]: k \\in Nat})" - proof - - { - fix "n" :: "c" - have s1_1: "n \\in {k + i: k \\in Nat} - = (\\E k \\in Nat: n = k + i)" - by auto - also have s1_2: "... = ( - \\E k \\in ({0} \ {q + 1: q \\in Nat}): - n = k + i)" - using SplitNat0 by auto - also have s1_3: "... = ( - (\\E k \\in {0}: n = k + i) \ - (\\E k \\in {q + 1: q \\in Nat}: n = k + i))" - by auto - also have s1_4: "... = ( - (n = 0 + i) \ - (\\E k: (k \\in {q + 1: q \\in Nat}) \ n = k + i))" - by auto - also have s1_5: "... = ( - (n = i) \ - (\\E k: (\\E q \\in Nat: k = q + 1) \ n = k + i))" - using idom nat_in_int add_0_left by auto - also have s1_6: "... = ( - (n = i) \ - (\\E k: \\E q \\in Nat: k = q + 1 \ n = k + i))" - by auto - also have s1_7: "... = ( - (n = i) \ - (\\E q \\in Nat: n = (q + 1) + i))" - by auto - also have s1_8: "... = ( - (n = i) \ - (\\E q \\in Nat: n = q + (i + 1)))" - using AddCommutative_sequent AddAssociative_sequent - oneIsNat idom nat_in_int by auto - also have s1_9: "... = ( - (n = i) \ - (\\E q \\in Nat: n = q + Succ[i]))" - using idom nat_add_1[symmetric] by force - also have s1_10: "... = ( - (n \\in {i}) \ - (n \\in {q + Succ[i]: q \\in Nat}))" - by auto - also have s1_11: "... = ( - n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - by auto - finally have s1_12: " - (n \\in {k + i: k \\in Nat}) - = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" . - } - from this have s1_13: "\ n. (n \\in {k + i: k \\in Nat}) - = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - by auto - have s1_14: "\ n. (n \\in {k + i: k \\in Nat}) - ==> (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - using s1_13 by auto - have s1_15: "\ n. (n \\in ({i} \ {k + Succ[i]: k \\in Nat})) - ==> (n \\in {k + i: k \\in Nat})" - using s1_13 by auto - show ?thesis - using s1_14 s1_15 - by (rule setEqualI, blast) - qed - - -(* -THEOREM LeqNatOpenInterval == - ASSUME NEW i \in Nat, - NEW j \in {k + Succ[i]: k \in Nat} - PROVE ~ (j <= i) -PROOF -<1>1. PICK k \in Nat: j = k + Succ[i] - OBVIOUS -<1>2. Succ[i] <= j - BY <1>1 DEF leq -<1>3. i < Succ[i] - BY i \in Nat, less_succ_nat -<1>4. i < j - <2>1. i \in Int - BY i \in Nat, nat_in_int - <2>2. j \in Int - BY <1>1, i \in Nat, succIsNat, nat_in_int, - addIsInt - <2>3. Succ[i] \in Int - BY i \in Nat, succIsNat, nat_in_int - <2> QED - BY <2>1, <2>2, <2>3, <1>3, <1>2, - leq_less_trans -<1>5. i \in Int - BY i \in Nat, nat_in_int -<1>6. j \in Int - BY <1>1, i \in Nat, succIsNat, nat_in_int, - addIsInt -<1> QED - BY <1>4, <1>5, <1>6, less_not_leq -*) -theorem LeqNatOpenInterval: - assumes idom: "i \\in Nat" and - jdom: "j \\in {k + Succ[i]: k \\in Nat}" - shows "~ (j <= i)" - proof - - define P where "P \ \ x. x \\in Nat \ j = x + Succ[i]" - define k where "k \ CHOOSE x: P(x)" - have s1_1: "k \\in Nat \ j = k + Succ[i]" - proof - - have s2_1: "\\E x: P(x)" - using jdom by (auto simp: P_def) - have s2_2: "P(k)" - unfolding k_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "Succ[i] <= j" - proof - - have s2_1: "k + Succ[i] = j" - using s1_1 by auto - have s2_2: "Succ[i] + k = j" - using s2_1 s1_1 idom succIsNat nat_in_int - AddCommutative_sequent by auto - have s2_3: "k \\in Nat" - using s1_1 by auto - show ?thesis - unfolding leq_def - using s2_2 s2_3 by auto - qed - have s1_3: "i < Succ[i]" - using idom less_succ_nat by auto - have s1_4: "i < j" - proof - - have s2_1: "i \\in Int" - using idom nat_in_int by auto - have s2_2: "j \\in Int" - proof - - have s3_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s3_2: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s3_3: "j = k + Succ[i]" - using s1_1 by auto - show ?thesis - using s3_1 s3_2 s3_3 addIsInt by auto - qed - have s2_3: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - show ?thesis - using s2_1 s2_2 s2_3 s1_3 s1_2 - less_leq_trans by iprover - qed - have s1_5: "i \\in Int" - using idom nat_in_int by auto - have s1_6: "j \\in Int" - proof - - have s2_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s2_2: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s2_3: "j = k + Succ[i]" - using s1_1 by auto - show ?thesis - using s2_1 s2_2 s2_3 addIsInt by auto - qed - show ?thesis - using s1_4 s1_5 s1_6 less_not_leq by auto - qed - - -(* -negNatIsAdd0 == - negNat = {-.k + -.0: k \in Nat} -PROOF -<1>1. negNat = {-.k: k \in Nat} - BY DEF negNat -<1>2. ASSUME NEW k \in Nat - PROVE -.k + -.0 = -.k - <2>1. -.k + -.0 = -.k + 0 - BY neg0 - <2>2. @ = -.k - <3>1. -.k \in Int - BY <1>2, minus_nat_in_int - <3> QED - BY <3>1, add_0 - <2> QED - BY <2>1, <2>2 -<1>3. {-.k: k \in Nat} = - (-.k + -.0: k \in Nat} - BY <1>2 -<1> QED - BY <1>1, <1>3 -*) -theorem negNatIsAdd0: - "negNat = {-.k + -.0: k \\in Nat}" - proof - - have s1_1: "negNat = {-.k: k \\in Nat}" - unfolding negNat_def by auto - have s1_2: "\ k. k \\in Nat ==> -.k + -.0 = -.k" - proof - - fix "k" :: "c" - assume s1_2_asm: "k \\in Nat" - have s2_1: "-.k + -.0 = -.k + 0" - using neg0 by auto - also have s2_2: "... = -.k" - proof - - have s3_1: "-.k \\in Int" - using s1_2_asm minus_nat_in_int by auto - show ?thesis - using s3_1 add_0 by auto - qed - finally show "-.k + -.0 = -.k" . - qed - have s1_3: "{-.k: k \\in Nat} = - {-.k + -.0: k \\in Nat}" - using s1_2 by auto - show ?thesis - using s1_1 s1_3 by auto - qed - - -theorem negNatSplitAddi: - assumes idom: "i \\in Nat" - shows "{-.k + -.i: k \\in Nat} = - ({-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - proof - - have s1_1: "\ n. (n \\in {-.k + -.i: k \\in Nat}) - = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - proof - - fix "n" :: "c" - have s2_1: "n \\in {-.k + -.i: k \\in Nat} - = (\\E k \\in Nat: n = -.k + -.i)" - by auto - also have s2_2: "... = ( - \\E k \\in ({0} \ {q + 1: q \\in Nat}): - n = -.k + -.i)" - using SplitNat0 by auto - also have s2_3: "... = ( - (\\E k \\in {0}: n = -.k + -.i) \ - (\\E k \\in {q + 1: q \\in Nat}: n = -.k + -.i) - )" - by auto - also have s2_4: "... = ( - (n = -.0 + -.i) \ - (\\E k: (k \\in {q + 1: q \\in Nat}) \ (n = -.k + -.i)))" - by auto - also have s2_5: "... = ( - (n = -.i) \ - (\\E k: (\\E q \\in Nat: k = q + 1) \ n = -.k + -.i))" - using idom minus_nat_in_int add_0_left by auto - also have s2_6: "... = ( - (n = -.i) \ - (\\E k: \\E q \\in Nat: k = q + 1 \ (n = -.k + -.i)))" - by auto - also have s2_7: "... = ( - (n = -.i) \ - (\\E q \\in Nat: n = -.(q + 1) + -.i))" - by auto - also have s2_8: "... = ( - (n = -.i) \ - (\\E q \\in Nat: n = -.q + -.Succ[i]))" - proof - - have s3_1: "\\E q \\in Nat: n = -.(q + 1) + -.i ==> - \\E q \\in Nat: n = -.q + -.Succ[i]" - proof - - assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" - define P where "P \ \ x. x \\in Nat \ - n = -.(x + 1) + -.i" - define r where "r \ CHOOSE x: P(x)" - have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" - proof - - have s5_1: "\\E x: P(x)" - using s3_1_asm by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_2: "n = (-.r + -.1) + -.i" - using s4_1 oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_3: "n = -.r + (-.i + -.1)" - using s4_2 s4_1 idom oneIsNat nat_in_int minus_nat_in_int - AddAssociative_sequent AddCommutative_sequent by auto - have s4_4: "n = -.r + -.(i + 1)" - using s4_3 idom oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_5: "n = -.r + -.Succ[i]" - proof - - have s5_1: "Succ[i] = i + 1" - using idom nat_add_1 by fast - show ?thesis - using s5_1 s4_4 by auto - qed - show "\\E q \\in Nat: n = -.q + -.Succ[i]" - using s4_1 s4_5 by auto - qed - have s3_2: "\\E q \\in Nat: n = -.q + -.Succ[i] ==> - \\E q \\in Nat: n = -.(q + 1) + -.i" - proof - - assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" - define P where "P \ \ x. x \\in Nat \ - n = -.x + -.Succ[i]" - define r where "r \ CHOOSE x: P(x)" - have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" - proof - - have s5_1: "\\E x: P(x)" - using s3_2_asm by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_2: "n = -.r + -.(i + 1)" - proof - - have s5_1: "Succ[i] = i + 1" - using idom nat_add_1 by fast - show ?thesis - using s4_1 s5_1 by auto - qed - have s4_3: "n = -.r + (-.i + -.1)" - using s4_2 idom oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_4: "n = (-.r + -.1) + -.i" - using s4_3 s4_1 idom oneIsNat minus_nat_in_int - AddAssociative_sequent AddCommutative_sequent - by auto - have s4_5: "n = -.(r + 1) + -.i" - using s4_4 s4_1 oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - show "\\E q \\in Nat: n = -.(q + 1) + -.i" - using s4_1 s4_5 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_9: "... = ( - (n \\in {-.i}) \ - (n \\in {-.k + -.Succ[i]: k \\in Nat}))" - by auto - also have s2_10: "... = ( - n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - by auto - finally show "(n \\in {-.k + -.i: k \\in Nat}) - = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" . - qed - have s1_2: "\ n. (n \\in {-.k + -.i: k \\in Nat}) - ==> (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - using s1_1 by auto - have s1_3: "\ n. (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat}) - ==> (n \\in {-.k + -.i: k \\in Nat})" - using s1_1 by auto - show ?thesis - using s1_2 s1_3 by (rule setEqualI, blast) - qed - - -theorem LeqnegNatOpenInterval: - assumes idom: "i \\in Nat" and - jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" - shows "~ (-.i <= j)" - proof - - define P where "P \ - \ x. x \\in Nat \ j = -.x + -.Succ[i]" - define k where "k \ CHOOSE x: P(x)" - have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" - proof - - have s2_1: "\\E x: P(x)" - using jdom by (auto simp: P_def) - have s2_2: "P(k)" - unfolding k_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_types: "j \\in Int \ -.i \\in Int \ - -.j \\in Int \ Succ[i] \\in Int \ - -.k \\in Int \ -.Succ[i] \\in Int \ - k \\in Int \ i \\in Int" - proof - - have s2_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s2_2: "-.k \\in Int" - using s2_1 neg_int_eq_int by auto - have s2_3: "i \\in Int" - using idom nat_in_int by auto - have s2_4: "-.i \\in Int" - using s2_3 neg_int_eq_int by auto - have s2_5: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s2_6: "-.Succ[i] \\in Int" - using s2_5 neg_int_eq_int by auto - have s2_7: "j \\in Int" - proof - - have s2_1: "j = -.k + -.Succ[i]" - using s1_1 by auto - show ?thesis - using s2_1 s2_2 s2_6 addIsInt by auto - qed - have s2_8: "-.j \\in Int" - using s2_7 neg_int_eq_int by auto - show ?thesis - using s2_1 s2_2 s2_3 s2_4 s2_5 - s2_6 s2_7 s2_8 by auto - qed - have s1_2: "Succ[i] <= -.j" - proof - - have s2_1: "Succ[i] + k = -.j" - proof - - have s3_1: "j = -.k + -.Succ[i]" - using s1_1 by auto - have s3_2: "-.j + j = -.j + (-.k + -.Succ[i])" - using s3_1 by auto - have s3_3: "0 = -.j + (-.k + -.Succ[i])" - using s3_2 s1_types AddNegCancels_left by auto - have s3_4: "0 + Succ[i] = (-.j + (-.k + -.Succ[i])) + Succ[i]" - using s3_3 by auto - have s3_5: "0 + Succ[i] = ((-.j + -.k) + -.Succ[i]) + Succ[i]" - using s3_4 s1_types AddAssociative_sequent by auto - have s3_6: "0 + Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" - using s3_5 s1_types addIsInt AddAssociative_sequent by auto - have s3_7: "Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" - proof - - have s4_1: "Succ[i] \\in Int" - using s1_types by auto - have s4_2: "0 + Succ[i] = Succ[i]" - using s4_1 add_0_left by fast - show ?thesis - using s3_6 s4_2 by auto - qed - have s3_8: "Succ[i] = (-.j + -.k) + 0" - using s3_7 s1_types AddNegCancels_left by auto - have s3_9: "Succ[i] = -.j + -.k" - proof - - have s4_1: "-.j + -.k \\in Int" - using s1_types addIsInt by auto - show ?thesis - using s3_8 s4_1 add_0 by auto - qed - have s3_10: "Succ[i] + k = (-.j + -.k) + k" - using s3_9 by auto - have s3_11: "Succ[i] + k = -.j + (-.k + k)" - using s3_10 s1_types AddAssociative_sequent by auto - have s3_12: "Succ[i] + k = -.j + 0" - using s3_11 s1_types AddNegCancels_left by auto - show ?thesis - using s3_12 s1_types add_0 by auto - qed - have s2_2: "\\E r \\in Nat: Succ[i] + r = -.j" - using s1_1 s2_1 by auto - show ?thesis - unfolding leq_def - using s2_2 by auto - qed - have s1_3: "i < Succ[i]" - using idom less_succ_nat by fast - have s1_4: "i < -.j" - using s1_types s1_2 s1_3 less_leq_trans by auto - have s1_5: "j < -.i" - proof - - have s2_1: "-.-.j < -.i" - proof - - have s3_1: "i \\in Int" - using s1_types by auto - have s3_2: "-.j \\in Int" - using s1_types by auto - show ?thesis - using s3_1 s3_2 s1_4 minus_less by auto - qed - have s2_2: "-.-.j = j" - using s1_types minus_involutive by auto - show ?thesis - using s2_1 s2_2 by auto - qed - show ?thesis - using s1_5 s1_types less_not_leq by auto - qed - - -theorem split_interval_nat_negnat: - shows "{k \\in Int: a <= k \ k <= b} = - {k \\in negNat: a <= k \ k <= b} \ - {k \\in Nat: a <= k \ k <= b}" - using int_union_nat_negnat by auto - - -end diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy deleted file mode 100644 index 90fb821e..00000000 --- a/isabelle/Peano.thy +++ /dev/null @@ -1,629 +0,0 @@ -(* Title: TLA+/Peano.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - -section \ Peano's axioms and natural numbers \ - -theory Peano -imports FixedPoints Functions -begin - -text \ - As a preparation for the definition of numbers and arithmetic - in \tlaplus{}, we state Peano's axioms for natural numbers and - prove the existence of a structure satisfying them. The presentation - of the axioms is somewhat simplified compared to the \tlaplus{} book. - (Moreover, the existence of such a structure is assumed, but not - proven in the book.) -\ - -subsection \ The Peano Axioms \ - -definition PeanoAxioms :: "[c,c,c] \ c" where - \ \parameters: the set of natural numbers, zero, and successor function\ - "PeanoAxioms(N,Z,Sc) \ - Z \ N - \ Sc \ [N \ N] - \ (\n \ N : Sc[n] \ Z) - \ (\m,n \ N: Sc[m] = Sc[n] \ m = n) - \ (\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S)" - -text \ - The existence of a structure satisfying Peano's axioms is proven - following the standard ZF construction where @{text "{}"} is zero, - @{text "i \ {i}"} is taken as the successor of any natural - number @{text i}, and the set of natural numbers is defined as - the least set that contains zero and is closed under successor - (this is a subset of the infinity set asserted to exist in ZF - set theory). In \tlaplus{}, natural numbers are defined by a sequence - of @{text CHOOSE}'s below, so there is no commitment to that - particular structure. -\ - -theorem peanoExists: "\N,Z,Sc : PeanoAxioms(N,Z,Sc)" -proof - - let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ - define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" - define N where "N \ lfp(infinity, expand)" - define Z where "Z \ {}" - define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ - have mono: "Monotonic(infinity, expand)" - using infinity by (auto simp: Monotonic_def expand_def) - hence expandN: "expand(N) \ N" - by (unfold N_def, rule lfpPreFP) - from expandN have 1: "Z \ N" - by (auto simp: expand_def Z_def) - have 2: "Sc \ [N \ N]" - proof (unfold Sc_def, rule functionInFuncSet) - show "\n \ N : ?sc(n) \ N" using expandN by (auto simp: expand_def) - qed - have 3: "\m\N : Sc[m] \ Z" - unfolding Z_def Sc_def by auto - have 4: "\m,n \ N : Sc[m] = Sc[n] \ m = n" - proof (clarify) - fix m n - assume "m \ N" and "n \ N" and "Sc[m] = Sc[n]" - hence eq: "?sc(m) = ?sc(n)" by (simp add: Sc_def) - show "m = n" - proof (rule setEqual) - show "m \ n" - proof (rule subsetI) - fix x - assume x: "x \ m" show "x \ n" - proof (rule contradiction) - assume "x \ n" - with x eq have "n \ m" by auto - moreover - from eq have "m \ ?sc(n)" by auto - ultimately - show "FALSE" by (blast elim: inAsym) - qed - qed - next - show "n \ m" - proof (rule subsetI) - fix x - assume x: "x \ n" show "x \ m" - proof (rule contradiction) - assume "x \ m" - with x eq have "m \ n" by auto - moreover - from eq have "n \ ?sc(m)" by auto - ultimately - show "FALSE" by (blast elim: inAsym) - qed - qed - qed - qed - have 5: "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" - proof (clarify del: subsetI) - fix S - assume sub: "S \ N" and Z: "Z \ S" and Sc: "\n\S : Sc[n] \ S" - show "N \ S" - proof (unfold N_def, rule lfpLB) - show "expand(S) \ S" - proof (auto simp: expand_def) - from Z show "{} \ S" by (simp add: Z_def) - next - fix n - assume n: "n \ S" - with Sc have "Sc[n] \ S" .. - moreover - from n sub have "n \ N" by auto - hence "Sc[n] = ?sc(n)" by (simp add: Sc_def) - ultimately show "?sc(n) \ S" by simp - qed - next - have "N \ infinity" - by (unfold N_def, rule lfpSubsetDomain) - with sub show "S \ infinity" by auto - qed - qed - from 1 2 3 4 5 have "PeanoAxioms(N,Z,Sc)" - unfolding PeanoAxioms_def by blast - thus ?thesis by blast -qed - -lemma peanoInduct: - assumes pa: "PeanoAxioms(N,Z,Sc)" - and "S \ N" and "Z \ S" and "\n. n \ S \ Sc[n] \ S" - shows "N \ S" -proof - - from pa have "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" - unfolding PeanoAxioms_def by blast - with assms show ?thesis by blast -qed - - -subsection \ Natural numbers: definition and elementary theorems \ - -text \ - The structure of natural numbers is now defined to be some set, - zero, and successor satisfying Peano's axioms. -\ - -definition Succ :: "c" -where "Succ \ CHOOSE Sc : \N,Z : PeanoAxioms(N,Z,Sc)" - -definition Nat :: "c" -where "Nat \ DOMAIN Succ" - -definition zero :: "c" ("0") -where "0 \ CHOOSE Z : PeanoAxioms(Nat, Z, Succ)" - -abbreviation "one \ Succ[0]" - notation "one" ("1") -definition two :: "c" ("2") - where "2 \ Succ[1]" -definition three :: "c" ("3") - where "3 \ Succ[2]" -definition four :: "c" ("4") - where "4 \ Succ[3]" -definition five :: "c" ("5") - where "five \ Succ[4]" -definition six :: "c" ("6") - where "six \ Succ[5]" -definition seven :: "c" ("7") - where "seven \ Succ[6]" -definition eight :: "c" ("8") - where "eight \ Succ[7]" -definition nine :: "c" ("9") - where "nine \ Succ[8]" -definition ten :: "c" ("10") - where "ten \ Succ[9]" -definition eleven :: "c" ("11") - where "eleven \ Succ[10]" -definition twelve :: "c" ("12") - where "twelve \ Succ[11]" -definition thirteen :: "c" ("13") - where "thirteen \ Succ[12]" -definition fourteen :: "c" ("14") - where "fourteen \ Succ[13]" -definition fifteen :: "c" ("15") - where "fifteen \ Succ[14]" - -(* -abbreviation "one \ Succ[0]" -notation "one" ("1") -abbreviation "two \ Succ[1]" -notation "two" ("2") -abbreviation "three \ Succ[2]" -notation "three" ("3") -abbreviation "four \ Succ[3]" -notation "four" ("4") -abbreviation "five \ Succ[4]" -notation "five" ("5") -abbreviation "six \ Succ[5]" -notation "six" ("6") -abbreviation "seven \ Succ[6]" -notation "seven" ("7") -abbreviation "eight \ Succ[7]" -notation "eight" ("8") -abbreviation "nine \ Succ[8]" -notation "nine" ("9") -abbreviation "ten \ Succ[9]" -notation "ten" ("10") -abbreviation "eleven \ Succ[10]" -notation "eleven" ("11") -abbreviation "twelve \ Succ[11]" -notation "twelve" ("12") -abbreviation "thirteen \ Succ[12]" -notation "thirteen" ("13") -abbreviation "fourteen \ Succ[13]" -notation "fourteen" ("14") -abbreviation "fifteen \ Succ[14]" -notation "fifteen" ("15") -*) - -lemma peanoNatZeroSucc: "PeanoAxioms(Nat, 0, Succ)" -proof - - have "\N,Z : PeanoAxioms(N,Z,Succ)" - proof (unfold Succ_def, rule chooseI_ex) - from peanoExists show "\Sc,N,Z : PeanoAxioms(N,Z,Sc)" by blast - qed - then obtain N Z where PNZ: "PeanoAxioms(N,Z,Succ)" by blast - hence "Succ \ [N \ N]" - by (simp add: PeanoAxioms_def) - hence "N = Nat" - by (simp add: Nat_def) - with PNZ have "PeanoAxioms(Nat, Z, Succ)" by simp - thus ?thesis by (unfold zero_def, rule chooseI) -qed - -lemmas - setEqualI [where A = "Nat", intro!] - setEqualI [where B = "Nat", intro!] - -lemma zeroIsNat [intro!,simp]: "0 \ Nat" -using peanoNatZeroSucc by (simp add: PeanoAxioms_def) - -lemma succInNatNat [intro!,simp]: "Succ \ [Nat \ Nat]" -using peanoNatZeroSucc by (simp add: PeanoAxioms_def) - -lemma succIsAFcn [intro!,simp]: "isAFcn(Succ)" -using succInNatNat by blast - -\ \@{text "DOMAIN Succ = Nat"}\ -lemmas domainSucc [intro!,simp] = funcSetDomain[OF succInNatNat] -\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ -lemmas succIsNat [intro!,simp] = funcSetValue[OF succInNatNat] - -lemma oneIsNat [intro!,simp]: "1 \ Nat" -by simp - -lemma twoIsNat [intro!,simp]: "2 \ Nat" -unfolding two_def by simp - -lemma [simp]: - assumes "n \ Nat" - shows "(Succ[n] = 0) = FALSE" -using assms peanoNatZeroSucc by (auto simp: PeanoAxioms_def) - -lemma [simp]: - assumes n: "n \ Nat" - shows "(0 = Succ[n]) = FALSE" -using assms by (auto dest: sym) - -lemma succNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): - "\Succ[n] = 0; n \ Nat\ \ P" - "\0 = Succ[n]; n \ Nat\ \ P" -by (simp+) - -lemma succInj [dest]: - assumes "Succ[m] = Succ[n]" and "m \ Nat" and "n \ Nat" - shows "m=n" -using peanoNatZeroSucc assms by (auto simp: PeanoAxioms_def) - -lemma succInjIff [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] = Succ[n]) = (m = n)" -using assms by auto - -lemma natInduct: - assumes z: "P(0)" - and sc: "\n. \n \ Nat; P(n)\ \ P(Succ[n])" - shows "\n\Nat : P(n)" -proof - - let ?P = "{n \ Nat : P(n)}" - from peanoNatZeroSucc have "Nat \ ?P" - by (rule peanoInduct, auto simp: z sc) - thus ?thesis by auto -qed - -\ \version of above suitable for the inductive reasoning package\ -lemma natInductE [case_names 0 Succ, induct set: Nat]: - assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(Succ[n])" - shows "P(n)" -using bspec[OF natInduct, where P=P] assms by blast - -(*** EXAMPLE INDUCTION PROOFS *** - -lemma "\n\Nat : n=0 \ (\m \ Nat : n = Succ[m])" -by (rule natInduct, auto) - -lemma - assumes 1: "n \ Nat" - shows "n=0 \ (\m \ Nat : n = Succ[m])" -using 1 by (induct, auto) - -*** END EXAMPLE ***) - -lemma natCases [case_names 0 Succ, cases set: Nat]: - assumes n: "n \ Nat" - and z: "n=0 \ P" and sc: "\m. \m \ Nat; n = Succ[m]\ \ P" - shows "P" -proof - - from n have "n=0 \ (\m \ Nat : n = Succ[m])" - by (induct, auto) - thus ?thesis - proof - assume "n=0" thus "P" by (rule z) - next - assume "\m\Nat : n = Succ[m]" - then obtain m where "m \ Nat" and "n = Succ[m]" .. - thus "P" by (rule sc) - qed -qed - -lemma succIrrefl: - assumes n: "n \ Nat" - shows "Succ[n] \ n" -using n by (induct, auto) - -lemma succIrreflE (*[elim] -- don't: "ignoring weak elimination rule" *): - "\Succ[n] = n; n \ Nat\ \ P" - "\n = Succ[n]; n \ Nat\ \ P" -by (auto dest: succIrrefl) - -lemma succIrrefl_iff [simp]: - "n \ Nat \ (Succ[n] = n) = FALSE" - "n \ Nat \ (n = Succ[n]) = FALSE" -by (auto dest: succIrrefl) - - -text \Induction over two parameters along the ``diagonal''.\ -lemma diffInduction: - assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, Succ[n])" - and step: "\m,n\Nat : P(m,n) \ P(Succ[m], Succ[n])" - shows "\m,n\Nat : P(m,n)" -proof (rule natInduct) - show "\n\Nat : P(0,n)" - using b1 b2 by (intro natInduct, auto) -next - fix m - assume m: "m \ Nat" and ih: "\n\Nat : P(m,n)" - show "\n\Nat : P(Succ[m],n)" - proof (rule bAllI) - fix n - assume "n \ Nat" thus "P(Succ[m],n)" - proof (cases) - case 0 with b1 m show ?thesis by auto - next - case Succ with step ih m show ?thesis by auto - qed - qed -qed - -lemma diffInduct: - assumes n: "n \ Nat" and m: "m \ Nat" - and b1: "\m. m\Nat \ P(m,0)" and b2: "\n. n\Nat \ P(0, Succ[n])" - and step: "\m n. \m \ Nat; n\Nat; P(m,n) \ \ P(Succ[m], Succ[n])" - shows "P(m,n)" -proof - - have "\m,n\Nat : P(m,n)" - by (rule diffInduction, auto intro: b1 b2 step) - with n m show ?thesis by blast -qed - -lemma not0_implies_Suc: - "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" -by(rule natCases, auto) - -subsection \ Initial intervals of natural numbers and ``less than'' \ - -text \ - The set of natural numbers up to (and including) a given $n$ is - inductively defined as the smallest set of natural numbers that - contains $n$ and that is closed under predecessor. - - NB: ``less than'' is not first-order definable from the Peano axioms, - a set-theoretic definition such as the following seems to be unavoidable. -\ - -definition upto :: "c \ c" -where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : Succ[k] \ S })" - -lemmas - setEqualI [where A = "upto(n)" for n, intro!] - setEqualI [where B = "upto(n)" for n, intro!] - -lemma uptoNat: "upto(n) \ Nat" - unfolding upto_def by (rule lfpSubsetDomain) - -lemma uptoPred: - assumes Suc: "Succ[m] \ upto(n)" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ upto(n)" -proof - - let ?f = "\S. {n} \ {k\Nat : Succ[k] \ S}" - from n have mono: "Monotonic(Nat, ?f)" - unfolding Monotonic_def by blast - from m Suc have 1: "m \ ?f(upto(n))" by auto - from mono have 2: "?f(upto(n)) \ upto(n)" - unfolding upto_def by (rule lfpPreFP) - from 1 2 show ?thesis by blast -qed - -lemma uptoZero: "upto(0) = {0}" -proof (rule setEqual) - have "{0} \ { k \ Nat : Succ[k] \ {0} } \ {0}" by auto - thus "upto(0) \ {0}" - unfolding upto_def by (rule lfpLB, auto) -next - show "{0} \ upto(0)" - unfolding upto_def by (rule lfpGLB, auto) -qed - -lemma uptoSucc: - assumes n: "n \ Nat" - shows "upto(Succ[n]) = upto(n) \ {Succ[n]}" (is "?lhs = ?rhs") -proof - - let ?preds = "\S. {k \ Nat : Succ[k] \ S}" - let ?f = "\S k. {k} \ ?preds(S)" - have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" - by (auto simp: Monotonic_def) - \ \``$\subseteq$''\ - from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto - also have "\ \ upto(n)" - by (unfold upto_def, rule lfpPreFP, rule mono, rule n) - finally have "?f(?rhs, Succ[n]) \ ?rhs" by auto - moreover from n have "?rhs \ Nat" - by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) - ultimately have 1: "?lhs \ ?rhs" - by (unfold upto_def[where n="Succ[n]"], rule lfpLB) - \ \``$\supseteq$''\ - from n mono have 2: "?f(?lhs, Succ[n]) \ ?lhs" - unfolding upto_def by (intro lfpPreFP, blast) - with n have "?f(?lhs, n) \ ?lhs" by auto - moreover have "?lhs \ Nat" by (rule uptoNat) - ultimately have 3: "upto(n) \ ?lhs" - unfolding upto_def[where n=n] by (rule lfpLB) - from 2 have 4: "Succ[n] \ ?lhs" by auto - from 3 4 have "?rhs \ ?lhs" by auto - with 1 show ?thesis by (rule setEqual) -qed - -lemma uptoRefl: - assumes n: "n \ Nat" - shows "n \ upto(n)" -using n proof (cases) - case 0 thus ?thesis by (simp add: uptoZero) -next - case Succ thus ?thesis by (auto simp: uptoSucc) -qed - -lemma zeroInUpto: - assumes n: "n \ Nat" - shows "0 \ upto(n)" -using n by (induct, auto simp: uptoZero uptoSucc) - -lemma SuccNotUptoZero: - assumes "n \ Nat" and "Succ[n] \ upto(0)" - shows "P" -using assms by (auto simp: uptoZero) - -lemma uptoTrans: - assumes "k \ upto(m)" and "m \ upto(n)" and "n \ Nat" - shows "k \ upto(n)" -proof - - have "\n\Nat : m \ upto(n) \ upto(m) \ upto(n)" - by (rule natInduct, auto simp: uptoZero uptoSucc) - with assms show ?thesis by blast -qed - -lemma succNotinUpto: - assumes n: "n \ Nat" - shows "Succ[n] \ upto(n)" -using n proof (induct) - show "1 \ upto(0)" by (auto simp: uptoZero) -next - fix n - assume n: "n \ Nat" and ih: "Succ[n] \ upto(n)" - show "Succ[Succ[n]] \ upto(Succ[n])" - proof (auto simp: uptoSucc n) - assume "Succ[Succ[n]] \ upto(n)" - with n have "Succ[n] \ upto(n)" - by (auto elim: uptoPred) - with ih show "FALSE" .. - qed -qed - -lemma uptoLimit: - assumes m: "m \ upto(n)" and suc: "Succ[m] \ upto(n)" and n: "n \ Nat" - shows "m=n" -proof - - from m uptoNat have mNat: "m \ Nat" by blast - from n have "\m\Nat: m \ upto(n) \ Succ[m] \ upto(n) \ m=n" (is "?P(n)") - by (induct, auto simp: uptoZero uptoSucc) - with mNat m suc show ?thesis by blast -qed - -lemma uptoAntisym: - assumes mn: "m \ upto(n)" and nm: "n \ upto(m)" - shows "m=n" -proof - - from mn uptoNat have m: "m \ Nat" by blast - from nm uptoNat have n: "n \ Nat" by blast - have "\m,n\Nat : m \ upto(n) \ n \ upto(m) \ m=n" (is "\m,n\Nat : ?P(m,n)") - proof (rule natInduct) - show "\n\Nat : ?P(0,n)" by (auto simp: uptoZero) - next - fix m - assume m: "m \ Nat" and ih: "\n\Nat : ?P(m,n)" - show "\n\Nat : ?P(Succ[m],n)" - proof (auto simp: uptoSucc m) - fix n - assume "Succ[m] \ upto(n)" and "n \ upto(m)" - from this m have "Succ[m] \ upto(m)" by (rule uptoTrans) - with m show "Succ[m] = n" \ \contradiction\ - by (blast dest: succNotinUpto) - qed - qed - with m n mn nm show ?thesis by blast -qed - -lemma uptoInj [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "(upto(n) = upto(m)) = (n = m)" -proof (auto) - assume 1: "upto(n) = upto(m)" - from n have "n \ upto(n)" by (rule uptoRefl) - with 1 have "n \ upto(m)" by auto - moreover - from m have "m \ upto(m)" by (rule uptoRefl) - with 1 have "m \ upto(n)" by auto - ultimately - show "n = m" by (rule uptoAntisym) -qed - -lemma uptoLinear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ upto(n) \ n \ upto(m)" (is "?P(m,n)") -using m proof induct - from n show "?P(0,n)" by (auto simp: zeroInUpto) -next - fix k - assume k: "k \ Nat" and ih: "?P(k,n)" - from k show "?P(Succ[k],n)" - proof (auto simp: uptoSucc) - assume kn: "(Succ[k] \ upto(n)) = FALSE" - show "n \ upto(k)" - proof (rule contradiction) - assume c: "n \ upto(k)" - with ih have "k \ upto(n)" by simp - from this kn n have "k = n" by (rule uptoLimit[simplified]) - with n have "n \ upto(k)" by (simp add: uptoRefl) - with c show "FALSE" .. - qed - qed -qed - - -subsection \ Primitive Recursive Functions \ - -text \ - We axiomatize a primitive recursive scheme for functions - with one argument and domain on natural numbers. Later, we - use it to define addition, multiplication and difference. -\ - -(** FIXME: replace axiom with proper fixpoint definition **) - -axiomatization where - primrec_nat: "\f : isAFcn(f) \ DOMAIN f = Nat - \ f[0] = e \ (\n \ Nat : f[Succ[n]] = h(n,f[n]))" - -lemma bprimrec_nat: - assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" - shows "\f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))" -proof - - from primrec_nat[of e h] obtain f where - 1: "isAFcn(f)" and 2: "DOMAIN f = Nat" - and 3: "f[0] = e" and 4: "\n\Nat : f[Succ[n]] = h(n,f[n])" - by blast - have "\n\Nat : f[n] \ S" - proof (rule natInduct) - from 3 e show "f[0] \ S" by simp - next - fix n - assume "n \ Nat" and "f[n] \ S" - with suc 4 show "f[Succ[n]] \ S" by force - qed - with 1 2 3 4 show ?thesis - by blast -qed - -lemma primrec_natE: - assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" - and f: "f = (CHOOSE g \ [Nat \ S] : g[0] = e \ (\n \ Nat: g[Succ[n]] = h(n,g[n])))" - (is "f = ?g") - and maj: "\ f \ [Nat \ S]; f[0] = e; \n \ Nat : f[Succ[n]] = h(n, f[n]) \ \ P" - shows "P" -proof - - from e suc have "\ g \ [Nat \ S] : g[0] = e \ (\n \ Nat: g[Succ[n]] = h(n,g[n]))" - by (rule bprimrec_nat) - hence "?g \ [Nat \ S] \ ?g[0] = e \ (\n \ Nat: ?g[Succ[n]] = h(n,?g[n]))" - by (rule bChooseI2, auto) - with f maj show ?thesis by blast -qed - -lemma bprimrecType_nat: - assumes "e \ S" and "\n \ Nat : \x \ S : h(n,x) \ S" - shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ - (\n \ Nat: f[Succ[n]] = h(n,f[n]))) - \ [Nat \ S]" -by (rule primrec_natE[OF assms], auto) - -end diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index dae0051c..91cf9114 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1852,7 +1852,7 @@ proof - qed (*** END AUTO_GENERATED STUFF ***) -(* + (* ------------- Proof rules for CASE expressions ------ *) lemma zenon_caseother0 : @@ -1920,7 +1920,7 @@ proof (rule boolEqual, simp only: zenon_seqify_appseq, rule iffI) assume h4: "~ c" with h1 h6 obtain i where i: "i \ DOMAIN zenon_seqify(cs)" "Append(zenon_seqify(cs), c)[i]" - by auto + by force with h6 have "zenon_seqify(cs)[i]" by (auto simp: leq_neq_iff_less[simplified]) with i show "\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i]" by blast qed @@ -1929,11 +1929,13 @@ next assume h1: "c | (\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i])" show "\ i \ DOMAIN Append (zenon_seqify (cs), c) - : Append (zenon_seqify (cs), c)[i]" (is "?g") + : Append (zenon_seqify (cs), c)[i]" + (is "?g" is "\i \ DOMAIN ?app : Append(?seq,c)[i]") proof (rule disjE [OF h1]) assume h2: c - show "?g" - using h2 h0 by auto + with h0 have "succ[Len(?seq)] \ DOMAIN ?app" "?app[succ[Len(?seq)]]" + by auto + thus "?g" by blast next assume h2: "\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i]" @@ -1944,15 +1946,15 @@ next have h4: "i \\in DOMAIN Append (zenon_seqify (cs), c)" using h0 h3 by auto assume h5: "zenon_seqify (cs)[i]" - have h6: "i ~= Succ[Len (zenon_seqify (cs))]" + have h6: "i ~= succ[Len (zenon_seqify (cs))]" using h0 h3 by auto have h7: "Append (zenon_seqify (cs), c)[i]" using h6 h5 h3 h0 by force show "?g" - using h4 h7 by auto + using h4 h7 by blast qed qed -qed (simp_all) +qed (simp+) lemma zenon_case_seq_empty : assumes h0: "\ i \ DOMAIN zenon_seqify (<<>>) @@ -2009,14 +2011,14 @@ proof proof (rule bExE [OF h5]) fix i assume h6: "i \\in DOMAIN Append (zenon_seqify(cs), c)" - have h7: "i = Succ [Len (zenon_seqify (cs))] + have h7: "i = succ [Len (zenon_seqify (cs))] | i \\in DOMAIN (zenon_seqify (cs))" using h3c h6 by auto assume h8: "x \\in CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i])" have h9: "i \\in DOMAIN (zenon_seqify (cs))" (is "?g") proof (rule disjE [OF h7]) - assume h10: "i = Succ [Len (zenon_seqify (cs))]" + assume h10: "i = succ [Len (zenon_seqify (cs))]" have h11: FALSE using h8 h10 h4 h3c h3e h2 by auto show "?g" using h11 by auto next @@ -2101,10 +2103,16 @@ next using h4 zenon_seqifyIsASeq by auto have h5: "x \\in CaseArm (c, e)" using h2 by blast + have "succ[Len(zenon_seqify(cs))] \\in DOMAIN Append (zenon_seqify (cs), c)" + "CaseArm (Append (zenon_seqify (cs), c)[succ[Len(zenon_seqify(cs))]], + Append (zenon_seqify (es), e)[succ[Len(zenon_seqify(cs))]]) + = CaseArm (c, e)" + using h3 zenon_seqifyIsASeq by auto + with h5 have h6: "x \\in UNION {CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) : i \\in DOMAIN Append (zenon_seqify (cs), c)}" - using h5 zenon_seqifyIsASeq appendElt2 h3 by auto + by blast show "?g" using h4 h6 by auto next @@ -2151,12 +2159,12 @@ next \\in {CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) : i \\in DOMAIN Append (zenon_seqify (cs), c)}" - using h10 by auto + using h10 by blast qed qed show "?g" using h4 h7 by blast qed -qed (simp_all) +qed (simp+) lemma zenon_case_len_simpl : fixes cs c es e @@ -2173,7 +2181,7 @@ next show "Len (Append (zenon_seqify (cs), c)) = Len (Append (zenon_seqify (es), e))" using h1 zenon_seqifyIsASeq by auto -qed (simp_all) +qed (simp+) lemma zenon_case_empty_union : fixes x @@ -2209,9 +2217,17 @@ next fix i assume h7: "i \\in DOMAIN Append (zenon_seqify (cs), c)" show "~Append (zenon_seqify (cs), c)[i]" - using g0 h7 h6 zenon_seqifyIsASeq by (unfold Append_def, auto) + proof (cases "i \\in DOMAIN zenon_seqify(cs)") + case True + with h6 zenon_seqifyIsASeq show ?thesis by force + next + case False + with h7 zenon_seqifyIsASeq[of cs] + have "i = succ[Len(zenon_seqify(cs))]" by (auto simp: int_less_not_leq) + with g0 zenon_seqifyIsASeq show ?thesis by auto + qed qed -qed (simp_all) +qed (simp+) lemma zenon_case_oth_simpl_l2 : fixes cs c es e @@ -2244,9 +2260,9 @@ proof assume h9: "i \\in DOMAIN Append (zenon_seqify (cs), c)" assume h10: "CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) = B" - have h11: "i = Succ[Len (zenon_seqify (cs))] ==> FALSE" + have h11: "i = succ[Len (zenon_seqify (cs))] ==> FALSE" proof - - assume h12: "i = Succ[Len (zenon_seqify (cs))]" + assume h12: "i = succ[Len (zenon_seqify (cs))]" have h13: "B = CaseArm (c, e)" using h10 h12 g0 appendElt2 zenon_seqifyIsASeq by auto show FALSE @@ -2418,14 +2434,14 @@ lemma zenon_case1 : assumes hoth: "~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c1x" (is "?dcs") + define dcs where "dcs \ c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2437,10 +2453,10 @@ proof - assume ha: "?dcs" define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2493,29 +2509,29 @@ lemma zenon_caseother1 : assumes hoth: "~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2600,14 +2616,14 @@ lemma zenon_case2 : assumes hoth: "~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c2x | c1x" (is "?dcs") + define dcs where "dcs \ c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2619,10 +2635,10 @@ proof - assume ha: "?dcs" define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2679,29 +2695,29 @@ lemma zenon_caseother2 : assumes hoth: "~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2791,14 +2807,14 @@ lemma zenon_case3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2812,11 +2828,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2879,30 +2895,30 @@ lemma zenon_caseother3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2999,14 +3015,14 @@ lemma zenon_case4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3021,11 +3037,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -3094,31 +3110,31 @@ lemma zenon_caseother4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3222,14 +3238,14 @@ lemma zenon_case5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3245,11 +3261,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -3324,32 +3340,32 @@ lemma zenon_caseother5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) define dcs where "dcs \ - c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + c5x | c4x | c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3446,7 +3462,7 @@ proof - by (rule zenon_case_empty_union [OF hi]) qed qed -*) + (*** END AUTO_GENERATED STUFF ***) end diff --git a/isabelle/examples/Allocator.thy b/isabelle/examples/Allocator.thy index 2c2647f9..30542f61 100644 --- a/isabelle/examples/Allocator.thy +++ b/isabelle/examples/Allocator.thy @@ -1,18 +1,29 @@ -(* Title: TLA+/AtomicBakery.thy - Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation +(* Title: TLA+/Allocator.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:22 merz> + Version: Isabelle2021-1 *) +section \A Simple Resource Allocator\ + theory Allocator imports Constant begin +text \ + We define and prove the safety of a simple-minded resource allocator + where clients can request and return resources. The allocator can + hand out resources that are available. Since the allocator does + not keep track of the fact if clients will be able to satisfy all + their requests, it may get into a situation where clients hold + part of their requested resources but will wait forever for the + outstanding resources, held by other clients. +\ + declare funcSetValue [dest] -consts -- {* constant parameters of TLA+ SimpleAllocator module *} +consts \ \ constant parameters of TLA+ SimpleAllocator module \ Client :: c Resource :: c @@ -28,14 +39,14 @@ lemma InitTypeInvariant: "Init(unsat,alloc) \ TypeInvariant(unsat,alloc)" by (auto simp: Init_def TypeInvariant_def) --- {* Alternative formulation as proof rule (meta-level implication) *} +\ \ Alternative formulation as proof rule (meta-level implication) \ lemma InitTypeInvariant': assumes "Init(unsat,alloc)" shows "TypeInvariant(unsat,alloc)" using assms by (auto simp: Init_def TypeInvariant_def) definition available where - "available(alloc) \ Resource \\ (UNION {alloc[c] : c \ Client})" + "available(alloc) \ Resource \ (UNION {alloc[c] : c \ Client})" definition Request where "Request(c,S,unsat,alloc,unsat',alloc') \ @@ -49,7 +60,7 @@ lemma RequestTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Request_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma RequestTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Request(c,S,unsat,alloc,unsat',alloc')" @@ -60,7 +71,7 @@ definition Allocate where "Allocate(c,S,unsat,alloc,unsat',alloc') \ S \ {} \ S \ available(alloc) \ unsat[c] \ alloc' = [alloc EXCEPT ![c] = alloc[c] \ S] - \ unsat' = [unsat EXCEPT ![c] = alloc[c] \\ S]" + \ unsat' = [unsat EXCEPT ![c] = alloc[c] \ S]" lemma AllocateTypeInvariant: "TypeInvariant(unsat,alloc) @@ -68,7 +79,7 @@ lemma AllocateTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Allocate_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma AllocateTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Allocate(c,S,unsat,alloc,unsat',alloc')" @@ -78,7 +89,7 @@ using assms unfolding TypeInvariant_def Allocate_def by auto definition Return where "Return(c,S,unsat,alloc,unsat',alloc') \ S \ {} \ S \ alloc[c] - \ alloc' = [alloc EXCEPT ![c] = alloc[c] \\ S] + \ alloc' = [alloc EXCEPT ![c] = alloc[c] \ S] \ unsat' = unsat" lemma ReturnTypeInvariant: @@ -87,7 +98,7 @@ lemma ReturnTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Return_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma ReturnTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Return(c,S,unsat,alloc,unsat',alloc')" @@ -109,14 +120,14 @@ by (blast intro: RequestTypeInvariant [rule_format] AllocateTypeInvariant [rule_format] ReturnTypeInvariant [rule_format]) --- {* Alternative formulation, using the rule formats of the individual lemmas *} +\ \ Alternative formulation, using the rule formats of the individual lemmas \ lemma NextTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "Next(unsat,alloc,unsat',alloc')" shows "TypeInvariant(unsat',alloc')" using assms unfolding Next_def by (blast intro: RequestTypeInvariant' AllocateTypeInvariant' ReturnTypeInvariant') --- {* Direct proof, without use of lemmas for subactions *} +\ \ Direct proof, without use of lemmas for subactions \ lemma NextTypeInvariant'': "TypeInvariant(unsat,alloc) \ Next(unsat,alloc,unsat',alloc') \ TypeInvariant(unsat',alloc')" @@ -124,7 +135,7 @@ unfolding TypeInvariant_def Next_def Request_def Allocate_def Return_def by auto -text {* Proof of mutual exclusion: no two processes ever hold a common resource *} +text \ Proof of mutual exclusion: no two processes ever hold a common resource \ definition Mutex where "Mutex(alloc) \ @@ -133,10 +144,10 @@ definition Mutex where lemma InitMutex: "Init(unsat,alloc) \ Mutex(alloc)" unfolding Init_def Mutex_def by auto -text {* +text \ The @{text Request} action trivially preserves mutual exclusion because it leaves the value of @{text alloc} unchanged. -*} +\ lemma RequestMutex: "Mutex(alloc) @@ -144,10 +155,10 @@ lemma RequestMutex: \ Mutex(alloc')" unfolding Request_def by auto -text {* +text \ The @{text Return} action decreases the set of allocated resources, so preservation of mutual exclusion follows easily. Note the use of the type invariant. -*} +\ lemma ReturnMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -174,10 +185,10 @@ proof (clarify) qed qed -text {* +text \ In fact, the proof is easy for Isabelle's automatic proof methods. Note the use of @{text clarsimp} to simplify the goal before the heavy lifting. -*} +\ lemma "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -185,10 +196,10 @@ lemma \ Mutex(alloc')" by (clarsimp simp: Mutex_def TypeInvariant_def Return_def) auto -text {* +text \ The proof for the @{text Allocate} action requires some case analysis. Unfortunately, we need to prove two symmetric assertions. -*} +\ lemma AllocateMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -206,13 +217,13 @@ proof (clarify) and rc: "r \ alloc'[c]" and rc': "r \ alloc'[c']" show "c' = c" proof (cases "c=clt \ c'=clt") - case False -- {* the simple case first *} + case False \ \ the simple case first \ with all c c' tpg have "alloc'[c] = alloc[c] \ alloc'[c'] = alloc[c']" by (auto simp: Allocate_def TypeInvariant_def) with mux rc rc' c c' r show "c' = c" by (auto simp: Mutex_def) next - case True -- {* one of @{term c} or @{term c'} is @{term clt} *} + case True \ \ one of @{term c} or @{term c'} is @{term clt} \ thus "c' = c" proof assume c1: "c = clt" @@ -233,7 +244,7 @@ proof (clarify) using all c1 c tpg by (auto simp: Allocate_def TypeInvariant_def) with rc show "FALSE" by auto qed - next -- {* symmetric argument *} + next \ \ symmetric argument \ assume c1: "c' = clt" show "c' = c" proof (rule contradiction) @@ -258,10 +269,10 @@ proof (clarify) qed -text {* +text \ Using the above lemmas, it is straightforward to derive that the next-state action preserves mutual exclusion. -*} +\ lemma NextMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) \ Next(unsat,alloc,unsat',alloc') diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index f04f5428..8e0f0a91 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -1,15 +1,23 @@ (* Title: TLA+/AtomicBakery.thy Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:37:50 merz> + Version: Isabelle2021-1 *) +section \Safety Proof of the Atomic Version of the Bakery Algorithm\ + theory AtomicBakeryG imports Constant begin +text \ + This is the version of the atomic Bakery algorithm used for + presenting TLAPS, verified using only Isabelle/TLA+. The + proof is quite tedious, compared to the proof that makes use + of the different back-end provers available in TLAPS. +\ + (************************************************** \algorithm AtomicBakery { @@ -80,7 +88,7 @@ definition Init where definition p1 where "p1(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p1'' - \ unread' = [unread EXCEPT ![self] = P \\ {self}] + \ unread' = [unread EXCEPT ![self] = P \ {self}] \ max' = [max EXCEPT ![self] = 0] \ flag' = [flag EXCEPT ![self] = TRUE] \ pc' = [pc EXCEPT ![self] = ''p2''] @@ -91,7 +99,7 @@ definition p2 where pc[self] = ''p2'' \ (IF unread[self] \ {} THEN \i \ unread[self] : - ( unread' = [unread EXCEPT ![self] = unread[self] \\ {i}] + ( unread' = [unread EXCEPT ![self] = unread[self] \ {i}] \ (IF num[i] > max[self] THEN max' = [max EXCEPT ![self] = num[i]] ELSE (max' = max))) @@ -111,7 +119,7 @@ definition p4 where "p4(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p4'' \ flag' = [flag EXCEPT ![self] = FALSE] - \ unread' = [unread EXCEPT ![self] = P \\ {self}] + \ unread' = [unread EXCEPT ![self] = P \ {self}] \ pc' = [pc EXCEPT ![self] = ''p5''] \ num' = num \ max' = max \ nxt' = nxt" @@ -133,11 +141,11 @@ definition p6 where \ (IF self > nxt[self] THEN num[nxt[self]] > num[self] ELSE num[nxt[self]] \ num[self])) - \ unread' = [unread EXCEPT ![self] = unread[self] \\ {nxt[self]}] + \ unread' = [unread EXCEPT ![self] = unread[self] \ {nxt[self]}] \ pc' = [pc EXCEPT ![self] = ''p5''] \ num' = num \ flag' = flag \ max' = max \ nxt' = nxt" -definition p7 where -- {* Critical section *} +definition p7 where \ \ Critical section \ "p7(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p7'' \ TRUE @@ -193,12 +201,12 @@ definition TypeOK where \ max \ [P \ Nat \ {defaultInitValue}] \ (\j \ P : (pc[j] \ {''p2'', ''p3''}) \ max[j] \ Nat) \ nxt \ [P \ P \ {defaultInitValue}] - \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ P \\ {i}) + \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ P \ {i}) \ pc \ [P \ {''p1'',''p2'',''p3'',''p4'',''p5'',''p6'',''p7'',''p8''}]" ---{* The type invariant in p6 should be - @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \\ {i})"} +\ \ The type invariant in p6 should be + @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \ {i})"} but it works anyway as it is. -*} +\ (** definition TypeOK where (** version for the alternative initial condition **) @@ -212,10 +220,10 @@ definition TypeOK where (** version for the alternative initial condition **) \ nxt \ [P \ P] \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ i) \ pc \ [P \ {''p1'',''p2'',''p3'',''p4'',''p5'',''p6'',''p7'',''p8''}]" ---{* The type invariant in p6 should be - @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \\ {i})"} +\ \ The type invariant in p6 should be + @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \ {i})"} but it works anyway as it is. -*} +\ **) definition GG where @@ -239,12 +247,12 @@ definition IInv where ((num[i] = 0) = (pc[i] \ {''p1'', ''p2'', ''p3''})) \ (flag[i] = (pc[i] \ {''p2'', ''p3'', ''p4''})) \ (pc[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread[i]) \\ {i} : After(j,i,unread,max,flag,pc,num,nxt))) + (\j \ (P \ unread[i]) \ {i} : After(j,i,unread,max,flag,pc,num,nxt))) \ ( pc[i] = ''p6'' \ ( (pc[nxt[i]] = ''p2'') \ i \ unread[nxt[i]] \ (pc[nxt[i]] = ''p3'')) \ max[nxt[i]] \ num[i]) - \ ((pc[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread,max,flag,pc,num,nxt)))" + \ ((pc[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread,max,flag,pc,num,nxt)))" definition Inv where "Inv(unread,max,flag,pc,num,nxt) \ @@ -261,8 +269,8 @@ theorem GGIrreflexive: assumes i: "i \ P" and j: "j \ P" and 1: "i \ j" and 2: "num[i] \ Nat \\ {0}" and 3: "num[j] \ Nat \\ {0}" shows "\(GG(i,j,num) \ GG(j,i,num))" -unfolding GG_def using assms nat_less_linear[of i j] -by (auto dest: procInNat nat_less_not_sym nat_less_leq_trans) +unfolding GG_def using assms int_less_linear[of i j] +by (auto dest: procInNat int_less_not_sym int_less_leq_trans) theorem InitImpliesTypeOK: "Init(unread,max,flag,pc,num,nxt) \ TypeOK(unread,max,flag,pc,num,nxt)" @@ -313,15 +321,27 @@ next with i j pc type show "unread'[j] \ P \ j \ unread'[j]" by (auto simp: TypeOK_def) qed - from u type i have n: "num[i] \ Nat" - by (auto simp: TypeOK_def) + from u \i \ unread[self]\ have "i \ ProcSet" by blast + with type have n: "num[i] \ Nat" by (auto simp: TypeOK_def) with i type have m': "max' \ [P \ Nat \ {defaultInitValue}]" by (auto simp: TypeOK_def) - from n i type have m'': "\j \ P : pc'[j] \ {''p2'',''p3''} \ max'[j] \ Nat" - (* FIXME: - - why doesn't Isabelle do the case distinction by itself? - - why do we need condElse? *) - by (cases "max[self] < num[i]", auto simp: TypeOK_def (*condThen*) condElse) + have m'': "\j \ P : pc'[j] \ {''p2'',''p3''} \ max'[j] \ Nat" + proof (clarify) + fix j + assume "j \ ProcSet" "pc'[j] \ {''p2'', ''p3''}" + with type i have mj: "max[j] \ Nat" + unfolding TypeOK_def by auto + show "max'[j] \ Nat" + proof (cases "j = self") + case True + with i mj type \j \ ProcSet\ n show ?thesis + by (cases "max[self] < num[i]") (auto simp: TypeOK_def) + next + case False + with i mj type \j \ ProcSet\ show ?thesis + by (auto simp: TypeOK_def) + qed + qed from triv u' u'' m' m'' show ?thesis by (auto simp add: TypeOK_def) qed @@ -390,7 +410,7 @@ qed lemma leq_neq_trans' (*[dest!]*): "a \ b \ b \ a \ a < b" - by (drule not_sym) (rule leq_neq_trans) + by (simp add: less_def) theorem InvInvariant: assumes inv: "Inv(unread,max,flag,pc,num,nxt)" @@ -407,26 +427,29 @@ proof auto assume type: "TypeOK(unread, max, flag, pc, num, nxt)" and iinv: "\j \ ProcSet: IInv(j, unread, max, flag, pc, num, nxt)" and i: "i \ ProcSet" - -- {* auxiliary definition that is used in several places of the proof below *} - def after \ "\k. pc[k] = ''p1'' \ + \ \ auxiliary definition that is used in several places of the proof below \ + define after where + "after \ \k. pc[k] = ''p1'' \ pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ pc[k] = ''p3'' \ num[i] \ max[k] \ (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. - -- {* iinv3 and iinv5: particular parts of the invariant, taken to the meta-level for then being instantiated with the proper variables, to ease the work of the classical reasoner. *} + \ \ iinv3 and iinv5: particular parts of the invariant, + taken to the meta-level for then being instantiated with the proper variables, + to ease the work of the classical reasoner. \ from iinv have iinv3: "\i j. - \pc[i] \ {''p5'', ''p6''}; i \ ProcSet; j \ ProcSet \\ unread[i] \\ {i}\ + \pc[i] \ {''p5'', ''p6''}; i \ ProcSet; j \ ProcSet \ unread[i] \ {i}\ \ After(j, i, unread, max, flag, pc, num, nxt)" proof - fix i j assume pci: "pc[i] \ {''p5'', ''p6''}" - and i: "i \ ProcSet" and j:"j \ ProcSet \\ unread[i] \\ {i}" + and i: "i \ ProcSet" and j:"j \ ProcSet \ unread[i] \ {i}" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. hence "pc[i] \ {''p5'', ''p6''} \ - (\j \ ProcSet \\ unread[i] \\ {i} : + (\j \ ProcSet \ unread[i] \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" unfolding IInv_def by auto with pci i j @@ -436,7 +459,7 @@ proof auto from iinv have iinv5: "\i j. - \pc[i] \ {''p7'', ''p8''}; i \ ProcSet; j \ ProcSet \\ {i}\ + \pc[i] \ {''p7'', ''p8''}; i \ ProcSet; j \ ProcSet \ {i}\ \ After(j, i, unread, max, flag, pc, num, nxt)" proof - fix i j @@ -444,7 +467,7 @@ proof auto and i: "i \ ProcSet" and j:"j \ ProcSet \\ {i}" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. hence "pc[i] \ {''p7'', ''p8''} \ - (\j \ ProcSet \\ {i} : + (\j \ ProcSet \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" unfolding IInv_def by auto with pci i j @@ -452,7 +475,7 @@ proof auto by auto qed - -- {* This also is an instantiation of a type invariant, since auto can't resolve it in a reasonable time. *} + \ \ This also is an instantiation of a type invariant, since auto can't resolve it in a reasonable time. \ from type i have nxti: "pc[i] = ''p6'' \ nxt[i] \ ProcSet \ nxt[i] \ i" by (auto simp: TypeOK_def) @@ -474,210 +497,89 @@ proof auto have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p1_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True with selfi p1 type iinvi self i pc' show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: "proof(..., auto)" doesn't finish in reasonable time **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - thus "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE" - by auto - qed - next - case False with selfi p1 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: "proof(..., auto/force) doesn't finish in reasonable time - and auto / force / blast cannot handle the following in one fell swoop - so we have to resort to very low-level reasoning. **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" by auto - assume "\k \ ProcSet \\ unread[i] \\ {i}: - pc[k] = ''p1'' \ - pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ - pc[k] = ''p3'' \ num[i] \ max[k] \ - (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ - GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" - hence aft: "\k \ ProcSet \\ unread[i] \\ {i}: after(k)" unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - hence "pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - unfolding after_def . - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by simp + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + next + case False with selfi p1 type iinvi self i pc' j show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto qed qed - qed - have 4: "pc'[i] = ''p6'' - \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] - \ (pc'[nxt'[i]] = ''p3'')) - \ max'[nxt'[i]] \ num'[i]" - proof (cases "self = nxt[i]") - case True - with selfi p1 type iinvi self i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p1_def) - next - case False - with type self i have "pc[i] = ''p6'' \ nxt[i] \ ProcSet" - by (auto simp: TypeOK_def) - with False selfi p1 type iinvi self i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p1_def) - qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" - proof (rule+) - fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" - show "After(j,i,unread',max',flag',pc',num',nxt')" - proof (cases "self = j") - case True with selfi p1 type iinvi self i pc' show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: same problems as in 3 above **) - assume "(j = i) = FALSE" "pc[i] = ''p7'' \ pc[i] = ''p8''" - thus "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE" - by auto - qed + have 4: "pc'[i] = ''p6'' + \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] + \ (pc'[nxt'[i]] = ''p3'')) + \ max'[nxt'[i]] \ num'[i]" + proof (cases "self = nxt[i]") + case True + with selfi p1 type iinvi self i show ?thesis + by (clarsimp simp: TypeOK_def IInv_def p1_def) next - case False with selfi p1 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: again, similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" by auto - assume "\k \ ProcSet \\ {i} : - pc[k] = ''p1'' \ - pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ - pc[k] = ''p3'' \ num[i] \ max[k] \ - (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ - GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" - hence aft: "\k \ ProcSet \\ {i} : after(k)" unfolding after_def . - with j have "after(j)" by blast - hence "pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - unfolding after_def . - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by simp + case False + with type self i have "pc[i] = ''p6'' \ nxt[i] \ ProcSet" + by (auto simp: TypeOK_def) + with False selfi p1 type iinvi self i show ?thesis + by (clarsimp simp: TypeOK_def IInv_def p1_def) + qed + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + proof (rule+) + fix j + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + show "After(j,i,unread',max',flag',pc',num',nxt')" + proof (cases "self = j") + case True with selfi p1 type iinvi self i pc' show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + next + case False with selfi p1 type iinvi self i pc' j show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto qed qed + from 1 2 3 4 5 show ?thesis + unfolding IInv_def by blast qed - from 1 2 3 4 5 show ?thesis - unfolding IInv_def by blast - qed - next + next assume p2: "p2(self, unread, max, flag, pc, num, nxt, unread', max', flag', pc', num', nxt')" show ?thesis proof (cases "self = i") assume selfi: "self = i" - show ?thesis - proof (cases "unread[self] = {}") - case True - with selfi p2 type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p2_def) - next - case False - from type have "isAFcn(pc)" - (* FIXME: needed for proving that pc'=pc, but why doesn't "simp ..." suffice *) - by (force simp add: TypeOK_def) - with False p2 obtain k where - k: "pc[self] = ''p2''" "k \ unread[self]" - "unread' = [unread EXCEPT ![self] = unread[self] \\ {k}]" - "max' = (IF num[k] > max[self] THEN [max EXCEPT ![self] = num[k]] ELSE max)" - "pc' = pc" "num' = num" "flag' = flag" "nxt' = nxt" - by (auto simp: p2_def) - with selfi type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def) - qed + with selfi p2 type iinvi i show ?thesis + by (auto simp: p2_def TypeOK_def IInv_def) next assume selfi: "self \ i" show ?thesis proof (cases "unread[self] = {}") assume empty: "unread[self] = {}" - with selfi p2 type iinvi self i + from i selfi self p2 type empty + have unch: "pc'[i] = pc[i]" "unread' = unread" "max' = max" "num' = num" "flag' = flag" + by (auto simp: TypeOK_def p2_def) + from empty selfi p2 type iinvi self i have 1: "(num'[i] = 0) = (pc'[i] \ {''p1'', ''p2'', ''p3''})" by (clarsimp simp: TypeOK_def IInv_def p2_def) from empty selfi p2 type iinvi self i have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p2_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" - from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" - by (auto simp: TypeOK_def) + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" + from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" + unfolding IInv_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") - case True with empty selfi p2 type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: again, some manual help necessary **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" - "(j \ unread[i]) = FALSE" "pc[j] = ''p2''" - "unread[j] = {}" - "\j \ ProcSet \\ unread[i] \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ num[i] \ max[j]" - by simp - qed + case True + with empty unch after p2 type j show ?thesis + by (auto simp: p2_def TypeOK_def After_def) next - case False with empty selfi p2 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i}: after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" "(j \ unread[i]) = FALSE" - with aft have "after(j)" by auto - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed + case False + with self j p2 type have "pc'[j] = pc[j]" + by (auto simp: TypeOK_def p2_def) + with after unch show ?thesis + by (auto simp: After_def GG_def) qed qed have 4: "pc'[i] = ''p6'' @@ -695,79 +597,39 @@ proof auto with False empty selfi p2 type iinvi self i show ?thesis by (clarsimp simp: TypeOK_def IInv_def p2_def) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" - from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" - by (auto simp: TypeOK_def) + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" + unfolding IInv_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True - with empty selfi p2 type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "unread[j] = {}" - "\j \ ProcSet \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ num[i] \ max[j]" - by simp - qed + with empty unch after p2 type j show ?thesis + by (auto simp: p2_def TypeOK_def After_def) next case False - with empty selfi p2 type iinvi i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i}: after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 - show "(pc[i] = ''p1'') = FALSE \ - (pc[i] = ''p2'') = FALSE \ - (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed + with self j p2 type have "pc'[j] = pc[j]" + by (auto simp: TypeOK_def p2_def) + with after unch show ?thesis + by (auto simp: After_def GG_def) qed qed from 1 2 3 4 5 show ?thesis unfolding IInv_def by blast next assume nempty: "unread[self] \ {}" - from type have "isAFcn(pc)" - (* FIXME: needed for proving that pc'=pc, but why doesn't "simp ..." suffice *) - by (force simp: TypeOK_def) - with nempty p2 obtain k where + with p2 type obtain k where k: "pc[self] = ''p2''" "k \ unread[self]" - "unread' = [unread EXCEPT ![self] = unread[self] \\ {k}]" + "unread' = [unread EXCEPT ![self] = unread[self] \ {k}]" "max' = (IF num[k] > max[self] THEN [max EXCEPT ![self] = num[k]] ELSE max)" "pc' = pc" "num' = num" "flag' = flag" "nxt' = nxt" - by (auto simp: p2_def) + by (auto simp: p2_def TypeOK_def) with type self have kproc: "k \ ProcSet" by (auto simp: TypeOK_def) + with type have numk: "num[k] \ Nat" + by (auto simp: TypeOK_def) from k selfi type iinvi self i have 1: "(num'[i] = 0) = (pc'[i] \ {''p1'', ''p2'', ''p3''})" by (clarsimp simp: TypeOK_def IInv_def) @@ -775,10 +637,12 @@ proof auto have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" + from pc' i k 1 type have numi: "0 < num'[i]" + by (auto simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -790,38 +654,28 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj less type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[k] = ''p5'' \ pc[k] = ''p6''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj less type iinvi self i pc' j mx numi numk show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj less type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj less type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "pc[j] = ''p2''" "j \ ProcSet" + "pc[j] = ''p2''" "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" + "(i \ unread[j]) = FALSE" "\j \ ProcSet \\ unread[i] \\ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence ii2: "i \ unread[j] \ num[i] \ max[j]" + hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" "i \ ProcSet" - "max[j] \ Nat" - with ii2 kproc have "i \ unread[j] \ num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) - with ii1 - show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ num[k])" - by simp + assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + with ii2 kproc mx i \pc[j] = ''p2''\ + show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed qed next @@ -830,94 +684,37 @@ proof auto proof (cases "i=k") case True with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse nat_not_less[simplified]) - assume "pc[k] = ''p5'' \ pc[k] = ''p6''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + by (auto simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse) + with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" + "j \ ProcSet" "(j \ unread[i]) = FALSE" + "(i \ unread[j]) = FALSE" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ unread[i] \\ {i} : + "\j \ ProcSet \ unread[i] \ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "i \ unread[j] \ num[i] \ max[j]" + thus "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ max[j])" - by simp qed qed qed next assume selfj: "self \ j" - show ?thesis - proof (cases "max[self] < num[k]") - case True - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ - (pc[i] = ''p2'') = FALSE \ - (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - next - case False - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - qed + with selfi i j self type k + have unr: "unread'[i] = unread[i]" "unread'[j] = unread[j]" + "max'[j] = max[j]" + by (auto simp: TypeOK_def) + with iinvi \pc' = pc\ pc' j + have "After(j,i,unread,max,flag,pc,num,nxt)" + by (auto simp: IInv_def) + with unr \num' = num\ \pc' = pc\ show ?thesis + by (auto simp: After_def GG_def) qed qed have 4: "pc'[i] = ''p6'' @@ -926,21 +723,37 @@ proof auto \ max'[nxt'[i]] \ num'[i]" proof (cases "self = nxt[i]") assume nxt: "self = nxt[i]" - from type k self have mx: "max[self] \ Nat" + from type k kproc self have mx: "max[self] \ Nat" "num[k] \ Nat" by (auto simp: TypeOK_def) show ?thesis proof (cases "max[self] < num[k]") case True with k selfi type iinvi self i nxt mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def, cases "i=k", simp, simp) - assume "num \ [ProcSet \ Nat]" "num[i] \ max[nxt[i]]" "max[nxt[i]] < num[k]" - with mx kproc i nxt show "num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) + proof (clarsimp simp: TypeOK_def IInv_def, cases "i \ unread[nxt[i]]", simp) + assume "i \ unread[nxt[i]]" "num \ [ProcSet \ Nat]" + "i \ unread[nxt[i]] = FALSE \ num[i] \ max[nxt[i]]" + "max[nxt[i]] < num[k]" + with mx nxt kproc i show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed next case False with k selfi type iinvi self i nxt mx show ?thesis - by (clarsimp simp: TypeOK_def IInv_def nat_not_less[simplified]) + proof (clarsimp simp: TypeOK_def IInv_def) + from False mx nxt type + have mxexc: "(IF max[nxt[i]] < num[k] + THEN [max EXCEPT ![nxt[i]] = num[k]] + ELSE max) = max" + by (auto simp: TypeOK_def) + assume "i \ unread[nxt[i]] = FALSE \ num[i] \ max[nxt[i]]" + "i \ unread[nxt[i]] = FALSE \ i = k" + "k \ unread[nxt[i]]" + with mxexc mx False nxt + show "num[i] \ (IF max[nxt[i]] < num[k] + THEN [max EXCEPT ![nxt[i]] = num[k]] + ELSE max)[nxt[i]]" + by auto + qed qed next assume nxt: "self \ nxt[i]" @@ -957,10 +770,12 @@ proof auto by (clarsimp simp: TypeOK_def IInv_def condElse) qed qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + from pc' i k 1 type have numi: "0 < num'[i]" + by (auto simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -972,38 +787,28 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj less type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[k] = ''p7'' \ pc[k] = ''p8''" - thus "(pc[k] = ''p1'') = FALSE - \ (pc[k] = ''p2'') = FALSE - \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj less type iinvi self i pc' j mx numi numk show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj less type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj less type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence ii2: "i \ unread[j] \ num[i] \ max[j]" + "pc[j] = ''p2''" "j \ ProcSet" + "(j = i) = FALSE" + "(i \ unread[j]) = FALSE" + "\j \ ProcSet \ {i} : + pc[j] = ''p1'' \ + pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ + pc[j] = ''p3'' \ num[i] \ max[j] \ + (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ + GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" + hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" "max[j] \ Nat" - with i kproc ii2 have "i \ unread[j] \ num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ num[k])" - by simp + assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + with ii2 kproc mx i \pc[j] = ''p2''\ + show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed qed next @@ -1011,92 +816,38 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj nless type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 nat_not_less[simplified]) - assume "pc[k] = ''p7'' \ pc[k] = ''p8''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj nless type iinvi self i pc' j mx show ?thesis + by (auto simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj nless type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ {i} : + "j \ ProcSet" + "(i \ unread[j]) = FALSE" + "(j = i) = FALSE" "pc[j] = ''p2''" + "\j \ ProcSet \ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "i \ unread[j] \ num[i] \ max[j]" + thus "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ max[j])" - by simp qed qed qed next assume selfj: "self \ j" - show ?thesis - proof (cases "max[self] < num[k]") - case True - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - next - case False - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - qed + with selfi i j self type k + have unr: "unread'[i] = unread[i]" "unread'[j] = unread[j]" + "max'[j] = max[j]" + by (auto simp: TypeOK_def) + with iinvi \pc' = pc\ pc' j + have "After(j,i,unread,max,flag,pc,num,nxt)" + by (auto simp: IInv_def) + with unr \num' = num\ \pc' = pc\ show ?thesis + by (auto simp: After_def GG_def) qed qed from 1 2 3 4 5 show ?thesis @@ -1109,7 +860,7 @@ proof auto proof (cases "self = i") case True with p3 type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p3_def) auto + by (clarsimp simp: TypeOK_def IInv_def p3_def) next assume selfi: "self \ i" from selfi p3 type iinvi self i @@ -1117,7 +868,7 @@ proof auto and 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p3_def)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" @@ -1125,7 +876,6 @@ proof auto proof (cases "self = j") case True with selfi type p3 self i j iinv3[of i j] pc' show ?thesis by (clarsimp simp: TypeOK_def p3_def After_def GG_def) auto - (** a bit slow, but it's not necessary to do case analysis on IF **) next case False with selfi p3 type self i j iinv3[of i j] pc' show ?thesis @@ -1177,10 +927,10 @@ proof auto and 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p4_def)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" with selfi type p4 self i iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" unfolding TypeOK_def p4_def After_def GG_def @@ -1199,10 +949,10 @@ proof auto with selfi p4 type iinvi self i nxti show ?thesis by (clarsimp simp: TypeOK_def IInv_def p4_def) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" with selfi type p4 self i j pc' iinv5[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" by (cases "self = j") (clarsimp simp: TypeOK_def p4_def After_def)+ @@ -1224,10 +974,10 @@ proof auto unfolding TypeOK_def IInv_def p5_def by(cases "self = i", clarsimp+) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" with empty type p5 self i j iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" unfolding TypeOK_def After_def GG_def p5_def @@ -1245,19 +995,16 @@ proof auto by (cases "self = nxt[i]", (clarsimp simp: nxti)+) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True with empty type p5 self i j iinv3[of i j] (** IInv part 3 is used, not part 5! **) show ?thesis - apply (clarsimp simp: IInv_def TypeOK_def p5_def After_def) - apply (cases "pc[j] = ''p2''", clarsimp) - apply (cases "pc[j] = ''p3''", clarsimp+) - done + by (clarsimp simp: IInv_def TypeOK_def p5_def After_def) auto next from empty have unreadj: "j \ unread[self]" by auto @@ -1282,18 +1029,28 @@ proof auto unfolding TypeOK_def IInv_def p5_def by(cases "self = i") clarsimp+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" - with empty type p5 self i j iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" - unfolding TypeOK_def After_def GG_def p5_def - apply(cases "self = j", clarsimp) - apply (cases "self = i", clarsimp) - apply (cases "pc[j] = ''p2''", clarsimp) - apply (cases "pc[j] = ''p3''", clarsimp+) - done + proof (cases "self=j") + case True + with empty type p5 self i j pc' iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) auto + next + case False + show ?thesis + proof (cases "self = i") + case True + with empty type p5 self i j pc' iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) auto + next + case False + with empty type p5 self i j pc' \self \ j\ iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) + qed + qed qed have 4: "pc'[i] = ''p6'' \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] @@ -1326,10 +1083,10 @@ proof auto by (cases "self = nxt[i]") clarsimp+ qed qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) from j iinv have iinvj: "IInv(j, unread, max, flag, pc, num, nxt)" by auto @@ -1349,10 +1106,10 @@ proof auto unfolding TypeOK_def IInv_def p6_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof(rule+) fix j - assume pc': "pc'[i] \ {''p5'', ''p6''}" and j: "j \ ProcSet \\ unread'[i] \\ {i}" + assume pc': "pc'[i] \ {''p5'', ''p6''}" and j: "j \ ProcSet \ unread'[i] \ {i}" from iinv j have iinvj: "IInv(j, unread, max, flag, pc, num, nxt)" by auto from type j @@ -1361,16 +1118,17 @@ proof auto show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = i") assume selfi: "self = i" + with iinvi p6 type i have numi: "0 < num[i]" "nxt[i] \ ProcSet" "nxt[i] \ i" + by (auto simp: IInv_def TypeOK_def p6_def) show ?thesis proof (cases "j \ unread[self]") assume junread: "j \ unread[self]" show ?thesis proof(cases "num[j] = 0") case True - with j selfi i pc' p6 type self iinvi junread iinvj (* FIXME: iinvj *) + with j selfi i pc' p6 type self iinvi junread iinvj show ?thesis - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def nat_gt0_not0) - by (cases "pc[nxt[i]] = ''p2''", clarsimp+) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def) auto next assume numj: "num[j] \ 0" show ?thesis @@ -1381,30 +1139,49 @@ proof auto case True with j selfi i pc' p6 type self iinvi junread numj pcj show ?thesis - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) auto next assume notunread: "i \ unread[j]" show ?thesis proof (cases "j < i") case True - from True j selfi i pc' p6 type self iinvi junread numj pcj iinv3[of j i] nxti notunread + with j selfi i pc' p6 type self iinvi junread numj numi pcj + iinv3[of j i] nxti notunread (*procInNat*) show ?thesis - using procInNat nat_less_antisym_false[of j i] - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) (force simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "nxt[i] < i" "j = nxt[i]" + with i j have nxti: "nxt[i] \ ProcSet" "\(i < nxt[i])" + by (auto dest: procInNat elim!: int_less_asym) + assume "num \ [ProcSet \ Nat]" + "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num[i] < num[nxt[i]]" + with i nxti + show "(pc[nxt[i]] = ''p5'') = FALSE + \ (pc[nxt[i]] = ''p6'') = FALSE" + by (auto simp: int_leq_not_less) + qed next case False - with j selfi i pc' p6 type self iinvi junread numj pcj iinv3[of j i] nxti notunread + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinv3[of j i] nxti notunread show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - proof (clarsimp simp: nat_gt0_not0 nat_not_less[simplified] nat_less_not_sym procInNat - dest!: leq_neq_trans'[of "i" "nxt[i]", simplified]) - assume - "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ num[nxt[i]] < num[i]" - "num[i] \ num[nxt[i]]" "nxt[i] \ ProcSet" "i \ ProcSet" - "num \ [ProcSet \ Nat]" - thus "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" - by (auto simp: nat_less_antisym_leq_false) - qed + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "(nxt[i] < i) = FALSE" + with i numi int_less_linear[of "i" "nxt[i]"] have "i < nxt[i]" + by (auto dest: procInNat) + assume + "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num[i] \ num[nxt[i]]" + "num \ [ProcSet \ Nat]" + with i numi \i < nxt[i]\ + show "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" + by (auto dest: procInNat simp: int_leq_not_less) + qed qed qed next @@ -1415,25 +1192,50 @@ proof auto show ?thesis proof (cases "j < i") case True - with j selfi i pc' p6 type self iinvi junread numj pcj iinvj nxti iinv5[of "nxt[i]" i] + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinvj nxti iinv5[of "nxt[i]" i] show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - using nat_less_antisym_false[of j i] - apply (clarsimp simp: nat_gt0_not0 procInNat) - by (clarsimp simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume nxti: "nxt[i] < i" + hence "i \ nxt[i]" by (auto simp: less_def) + assume + "i = nxt[i] = FALSE + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] + THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num \ [ProcSet \ Nat]" + with \i \ nxt[i]\ have "num[nxt[i]] \ num[i]" + by (auto dest: procInNat elim!: int_less_asym[OF nxti]) + assume "num[i] < num[nxt[i]]" + with i numi \num \ [ProcSet \ Nat]\ \num[nxt[i]] \ num[i]\ + show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') + \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" + by (auto simp: int_less_not_leq) + qed next case False - with j selfi i pc' p6 type self iinvi junread numj pcj iinv5[of j i] nxti + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinv5[of j i] nxti + int_less_linear[of "nxt[i]" "i"] show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - by (clarsimp simp: nat_gt0_not0 nat_not_less[simplified] procInNat) - (clarsimp dest!: leq_neq_trans'[simplified] simp: nat_less_antisym_leq_false nat_less_not_sym) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "(nxt[i] < i) = FALSE" + with i numi int_less_linear[of "nxt[i]" "i"] have "i < nxt[i]" + by (auto dest: procInNat) + assume "num[i] \ num[nxt[i]]" + "IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i])" + "num \ [ProcSet \ Nat]" + with i numi \i < nxt[i]\ + show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') + \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" + by (auto simp: int_less_not_leq) + qed qed next assume pcj2: "pc[j] \ {''p7'', ''p8''}" with j selfi i pc' p6 type self iinvi junread numj pcj iinvj nxti show ?thesis - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) + apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) by (erule funcSetE[where f="pc" and x="nxt[i]"], simp+) qed qed @@ -1453,18 +1255,32 @@ proof auto show ?thesis proof (cases "j < nxt[j]") case True - with j pc' selfi selfj i p6 type self iinv3[of i j] + with j pc' selfi selfj i p6 type self iinv3[of i j] nxtj show ?thesis - apply (clarsimp simp: p6_def TypeOK_def After_def GG_def nat_gt0_not0) - by (auto dest: nat_less_trans nat_less_leq_trans) - (** slow **) + by (clarsimp simp: p6_def TypeOK_def After_def GG_def) + (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] + elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False with j pc' selfi selfj i p6 type self iinv3[of i j] show ?thesis - using nxtj nat_not_less procInNat - by (clarsimp simp: p6_def TypeOK_def After_def GG_def nat_gt0_not0 leq_neq_iff_less[simplified]) - (auto simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def After_def GG_def) + assume + "pc[j] = ''p6''" + "(j < nxt[j]) = FALSE" + with i j nxtj int_less_linear[of "j" "nxt[j]"] have "nxt[j] < j" + by (force dest: procInNat) + + assume + "num[nxt[j]] = 0 + \ (IF nxt[j] < j THEN num[j] < num[nxt[j]] ELSE num[j] \ num[nxt[j]])" + "0 < num[i]" + "IF j < i THEN num[i] < num[j] ELSE num[i] \ num[j]" + "num \ [ProcSet \ Nat]" + with i j \nxt[j] < j\ show "(i = nxt[j]) = FALSE" + by (auto simp: int_less_not_leq[of "num[j]" "num[nxt[j]]"] + elim!: int_less_asym[of "num[nxt[j]]" "num[j]"]) + qed qed next case False @@ -1488,10 +1304,10 @@ proof auto apply (clarsimp simp: TypeOK_def IInv_def p6_def) by (cases "self = i", simp+) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True @@ -1501,24 +1317,24 @@ proof auto case True with type p6 i pc' j self selfj iinv5[of i j] show ?thesis - using nat_less_antisym_false[of j i] procInNat - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) - by (cases "i = nxt[j]") (auto simp: nat_less_antisym_leq_false) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] + elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False from i j type have num: "num[i] \ Nat" "num[j] \ Nat" by (auto simp: TypeOK_def) from False type p6 i pc' j self selfj iinv5[of i j] show ?thesis - using nat_less_linear[of i j] procInNat - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) - (auto dest!: nat_leq_less_trans) + using int_less_linear[of i j] procInNat + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + (auto dest!: int_leq_less_trans) qed next case False with type iinv5[of i j] p6 i pc' j show ?thesis apply (clarsimp simp: TypeOK_def IInv_def After_def p6_def) - by (cases "self = i", simp, simp) + by (cases "self = i", simp+) qed qed from 1 2 3 4 5 show ?thesis @@ -1532,10 +1348,10 @@ proof auto unfolding TypeOK_def IInv_def p7_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \\ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True @@ -1564,10 +1380,10 @@ proof auto unfolding TypeOK_def IInv_def p7_def by (cases "self = nxt[i]") clarsimp+ qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = j") case True @@ -1590,10 +1406,10 @@ proof auto unfolding p8_def TypeOK_def IInv_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True @@ -1622,10 +1438,10 @@ proof auto unfolding p8_def TypeOK_def IInv_def by (cases "self = nxt[i]") clarsimp+ qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = i") case True diff --git a/isabelle/examples/ROOT b/isabelle/examples/ROOT new file mode 100644 index 00000000..cbcd3fa6 --- /dev/null +++ b/isabelle/examples/ROOT @@ -0,0 +1,18 @@ +(* Title: TLA+/examples/ROOT + Author: Stephan Merz, Inria Nancy + License: BSD + +Examples of using Isabelle/TLA+ +*) + +(* build the session (without a heap) using "isabelle build -D ." *) +chapter "TLA+ Examples" +session "TLA-Examples" = "TLA+" + + options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + + theories + Allocator + AtomicBakeryG + + document_files + "root.tex" diff --git a/isabelle/examples/ROOT.ML b/isabelle/examples/ROOT.ML deleted file mode 100644 index 6812d8d3..00000000 --- a/isabelle/examples/ROOT.ML +++ /dev/null @@ -1,13 +0,0 @@ -(* Title: TLA+/examples/ROOT.ML - Author: Stephan Merz, LORIA - License: BSD - -Examples of using Isabelle/TLA+ -*) - - -val banner = "TLA+ Examples"; -writeln banner; - -use_thy "Allocator"; -use_thy "AtomicBakeryG"; From 3556e1d4ecf8536fa5664c78ad1c528187d8ea2a Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 23 Oct 2020 14:46:46 +0200 Subject: [PATCH 063/167] migration of Isabelle theories up to NatOrderings --- isabelle/FixedPoints.thy | 59 +- isabelle/Functions.thy | 228 ++++--- isabelle/NatOrderings.thy | 217 +++--- isabelle/Peano.thy | 79 ++- isabelle/PredicateLogic.thy | 1218 ++++++++++++++++------------------ isabelle/ROOT | 7 + isabelle/SetTheory.thy | 696 +++++++++---------- isabelle/document/root.tex | 4 +- isabelle/simplifier_setup.ML | 150 +++-- 9 files changed, 1325 insertions(+), 1333 deletions(-) create mode 100644 isabelle/ROOT diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index 1af421e9..e4a086b3 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -1,28 +1,27 @@ (* Title: TLA+/FixedPoints.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:18 merz> + Version: Isabelle2020 *) -header {* Fixed points for set-theoretical constructions *} +section \Fixed points for set-theoretical constructions\ theory FixedPoints imports SetTheory begin -text {* +text \ As a test for the encoding of \tlaplus{} set theory, we develop the Knaster-Tarski theorems for least and greatest fixed points in the subset lattice. Again, the proofs essentially follow Paulson's developments for Isabelle/ZF. -*} +\ -subsection {* Monotonic operators *} +subsection \ Monotonic operators \ -definition Monotonic :: "[c, c \ c] \ c" -- {* monotonic operator on a domain *} +definition Monotonic :: "[c, c \ c] \ c" \ \monotonic operator on a domain\ where "Monotonic(D,f) \ f(D) \ D \ (\S,T \ SUBSET D : S \ T \ f(S) \ f(T))" lemma monotonicDomain: @@ -37,10 +36,10 @@ lemma monotonicSubsetDomain: "\ Monotonic(D,f); S \ D\ \ f(S) \ D" by (unfold Monotonic_def, blast) -lemma monotonicCup: +lemma monotonicUnion: assumes mono: "Monotonic(D,f)" and s: "S \ D" and t: "T \ D" shows "f(S) \ f(T) \ f(S \ T)" -proof (rule cupLUB) +proof (rule unionLUB) from s t show "f(S) \ f(S \ T)" by (intro monotonicSubset[OF mono], blast+) next @@ -48,10 +47,10 @@ next by (intro monotonicSubset[OF mono], blast+) qed -lemma monotonicCap: +lemma monotonicInter: assumes mono: "Monotonic(D,f)" and s: "S \ D" and t: "T \ D" shows "f(S \ T) \ f(S) \ f(T)" -proof (rule capGLB) +proof (rule interGLB) from s t show "f(S \ T) \ f(S)" by (intro monotonicSubset[OF mono], blast+) from s t show "f(S \ T) \ f(T)" @@ -59,29 +58,29 @@ proof (rule capGLB) qed -subsection {* Least fixed point *} +subsection \ Least fixed point \ -text {* +text \ The least fixed point is defined as the greatest lower bound of the set of all pre-fixed points, and the Knaster-Tarski theorem is shown for monotonic operators. -*} +\ -definition lfp :: "[c, c \ c] \ c" -- {* least fixed point as GLB of pre-fp's *} +definition lfp :: "[c, c \ c] \ c" \ \least fixed point as GLB of pre-fp's\ where "lfp(D,f) \ INTER {S \ SUBSET D : f(S) \ S}" -lemma lfpLB: -- {* The lfp is contained in each pre-fixed point. *} +lemma lfpLB: \ \The lfp is contained in each pre-fixed point.\ "\ f(S) \ S; S \ D \ \ lfp(D,f) \ S" by (auto simp: lfp_def) -lemma lfpGLB: -- {* \ldots{} and it is the GLB of all such sets *} +lemma lfpGLB: \ \ \ldots{} and it is the GLB of all such sets \ "\f(D) \ D; \S. \f(S) \ S; S \ D\ \ A \ S\ \ A \ lfp(D,f)" by (force simp: lfp_def) lemma lfpSubsetDomain: "lfp(D,f) \ D" (* monotonicity not required *) by (auto simp: lfp_def) -lemma lfpPreFP: -- {* @{text lfp} is a pre-fixed point \ldots *} +lemma lfpPreFP: \ \@{text lfp} is a pre-fixed point \ldots \ assumes mono: "Monotonic(D,f)" shows "f(lfp(D,f)) \ lfp(D,f)" proof (rule lfpGLB) @@ -95,7 +94,7 @@ next with pfp show "f(?mu) \ S" by blast qed -lemma lfpPostFP: -- {* \ldots{} and a post-fixed point *} +lemma lfpPostFP: \ \ \ldots{} and a post-fixed point \ assumes mono: "Monotonic(D,f)" shows "lfp(D,f) \ f(lfp(D,f))" proof - @@ -120,7 +119,7 @@ lemma lfpLeastFixedPoint: shows "lfp(D,f) \ S" using assms by (intro lfpLB, auto) -lemma lfpMono: -- {* monotonicity of the @{text lfp} operator *} +lemma lfpMono: \ \ monotonicity of the @{text lfp} operator \ assumes g: "g(D) \ D" and f: "\S. S \ D \ f(S) \ g(S)" shows "lfp(D,f) \ lfp(D,g)" using g @@ -132,29 +131,29 @@ proof (rule lfpGLB) qed -subsection {* Greatest fixed point *} +subsection \ Greatest fixed point \ -text {* +text \ Dually, the least fixed point is defined as the least upper bound of the set of all post-fixed points, and the Knaster-Tarski theorem is again established. -*} +\ -definition gfp :: "[c, c \ c] \ c" -- {* greatest fixed point as LUB of post-fp's *} +definition gfp :: "[c, c \ c] \ c" \ \ greatest fixed point as LUB of post-fp's \ where "gfp(D,f) \ UNION {S \ SUBSET D : S \ f(S)}" -lemma gfpUB: -- {* The gfp contains each post-fixed point \ldots *} +lemma gfpUB: \ \ The gfp contains each post-fixed point \ldots \ "\S \ f(S); S \ D\ \ S \ gfp(D,f)" by (auto simp: gfp_def) -lemma gfpLUB: -- {* \ldots{} and it is the LUB of all such sets. *} +lemma gfpLUB: \ \ \ldots{} and it is the LUB of all such sets. \ "\f(D) \ D; \S. \S \ f(S); S \ D\ \ S \ A\ \ gfp(D,f) \ A" by (auto simp: gfp_def) lemma gfpSubsetDomain: "gfp(D,f) \ D" by (auto simp: gfp_def) -lemma gfpPostFP: -- {* @text{gfp} is a post-fixed point \ldots *} +lemma gfpPostFP: \ \ @text{gfp} is a post-fixed point \ldots \ assumes mono: "Monotonic(D,f)" shows "gfp(D,f) \ f(gfp(D,f))" proof (rule gfpLUB) @@ -168,7 +167,7 @@ next with pfp show "S \ f(?nu)" by blast qed -lemma gfpPreFP: -- {* \ldots{} and a pre-fixed point *} +lemma gfpPreFP: \ \ \ldots{} and a pre-fixed point \ assumes mono: "Monotonic(D,f)" shows "f(gfp(D,f)) \ gfp(D,f)" proof - @@ -193,7 +192,7 @@ lemma gfpGreatestFixedPoint: shows "S \ gfp(D,f)" using assms by (intro gfpUB, auto) -lemma gfpMono: -- {* monotonicity of the @{text gfp} operator *} +lemma gfpMono: \ \ monotonicity of the @{text gfp} operator \ assumes g: "g(D) \ D" and f: "\S. S \ D \ f(S) \ g(S)" shows "gfp(D,f) \ gfp(D,g)" proof (rule gfpLUB) diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index dade3389..922271f0 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -1,67 +1,80 @@ (* Title: TLA+/Functions.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:56 merz> + Version: Isabelle2020 *) -header {* \tlaplus{} Functions *} +section \ \tlaplus{} Functions \ theory Functions imports SetTheory begin -subsection {* Syntax and axioms for functions *} +subsection \ Syntax and axioms for functions \ -text {* +text \ Functions in \tlaplus{} are not defined (e.g., as sets of pairs), but axiomatized, and in fact, pairs and tuples will be defined as special functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined - as functions over products. - - We follow the development of functions given in Section 16.1 of - ``Specifying Systems''. In particular, we define the predicate - @{text IsAFcn} that is true precisely of functional values. -*} + as functions over products. We follow the development of functions given + in Section 16.1 of ``Specifying Systems''. +\ consts - isAFcn :: "c \ c" -- {* characteristic predicate *} - Fcn :: "[c, c \ c] \ c" -- {* function constructor *} - DOMAIN :: "c \ c" ("(DOMAIN _)" [100]90) -- {* domain of a function *} - fapply :: "[c, c] \ c" ("(_[_])" [89,0]90) -- {* function application *} - FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) -- {* function space *} + Fcn :: "[c, c \ c] \ c" \ \ function constructor \ + DOMAIN :: "c \ c" ("(DOMAIN _)" [100]90) \ \ domain of a function \ + fapply :: "[c, c] \ c" ("(_[_])" [89,0]90) \ \ function application \ syntax - "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \\in _ |-> _])" 900) -syntax (xsymbols) - FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) - "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \ _ \ _])" 900) -syntax (HTML output) - FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \ _ \ _])" 900) +syntax (ASCII) + "@Fcn" :: "[idt,c,c] \ c" ("(1[_ \\in _ |-> _])" 900) translations "[x \ S \ e]" \ "CONST Fcn(S, \x. e)" axiomatization where - fcnIsAFcn [intro!,simp]: "isAFcn(Fcn(S,e))" and - isAFcn_def: "isAFcn(f) \ f = [x \ DOMAIN f \ f[x]]" and DOMAIN [simp]: "DOMAIN Fcn(S,e) = S" and - fapply [simp]: "v \ S \ Fcn(S,e)[v] = e(v)" and - fcnEqual[elim!]: "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ - \ f = g" and + fapply [simp]: "v \ S \ Fcn(S,e)[v] = e(v)" + +text \The predicate @{text isAFcn} characterizes functional values.\ +definition isAFcn :: "c \ c" + where "isAFcn(f) \ f = Fcn(DOMAIN f, \x. f[x])" + +axiomatization where + fcnIsAFcn [intro!, simp]: "isAFcn(Fcn(S,e))" + +text \ + Two functions are equal if they have the same domain and agree on all + arguments within the domain. +\ +axiomatization where fcnEqual[elim!]: + "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ + \ f = g" + +text \ + @{text "FuncSet(S,T)"} represents the set of functions with domain @{text S} + and co-domain @{text T}. It cannot be defined as a set comprehension because + we do not have a bounding set for its elements. +\ +consts + FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) \ \ function space \ +syntax (ASCII) + FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) + +axiomatization where FuncSet: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T)" -lemmas -- {* establish set equality for domains and function spaces *} - setEqualI [where A = "DOMAIN f", standard, intro!] - setEqualI [where B = "DOMAIN f", standard, intro!] - setEqualI [where A = "[S \ T]", standard, intro!] - setEqualI [where B = "[S \ T]", standard, intro!] +lemmas \ \ establish set equality for domains and function spaces \ + setEqualI [where A = "DOMAIN f" for f, intro!] + setEqualI [where B = "DOMAIN f" for f, intro!] + setEqualI [where A = "[S \ T]" for S T, intro!] + setEqualI [where B = "[S \ T]" for S T, intro!] -definition except :: "[c,c,c] \ c" -- {* function override *} +definition except :: "[c,c,c] \ c" \ \ function override \ where "except(f,v,e) \ [x \ DOMAIN f \ (IF x=v THEN e ELSE f[x])]" nonterminal @@ -74,11 +87,11 @@ translations "_except(f, _xcpts(v,e, xcs))" \ "_except(CONST except(f,v,e), xcs)" "[f EXCEPT ![v] = e]" \ "CONST except(f,v,e)" -text {* +text \ The following operators are useful for representing functions with finite domains by enumeration. They are not part of basic \tlaplus{}, but they are defined in the TLC module of the standard library. -*} +\ definition oneArg :: "[c,c] \ c" (infixl ":>" 75) where "d :> e \ [x \ {d} \ e]" @@ -88,7 +101,7 @@ where "f @@ g \ [x \ (DOMAIN f) \ (DOMAIN g) \ IF x \ DOMAIN f THEN f[x] ELSE g[x]]" -subsection {* @{text isAFcn}: identifying functional values *} +subsection \ @{text isAFcn}: identifying functional values \ lemma boolifyIsAFcn [simp]: "boolify(isAFcn(f)) = isAFcn(f)" by (simp add: isAFcn_def) @@ -106,17 +119,17 @@ by auto lemma [intro!,simp]: "isAFcn([f EXCEPT ![c] = e])" by (simp add: except_def) -text {* +text \ We derive instances of axiom @{text fcnEqual} that help in automating proofs about equality of functions. -*} +\ lemma fcnEqual2[elim!]: "\isAFcn(g); isAFcn(f); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" by (rule fcnEqual) --- {* possibly useful as a simplification rule, but cannot be active by default *} +text \ possibly useful as a simplification rule, but cannot be active by default \ lemma fcnEqualIff: assumes "isAFcn(f)" and "isAFcn(g)" shows "(f = g) = (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : f[x] = g[x]))" @@ -137,7 +150,7 @@ lemma [intro!]: by (auto simp: except_def) -subsection {* Theorems about functions *} +subsection \ Theorems about functions \ lemma fcnCong (*[cong] -- not active by default (??) *): assumes "S = T" and "\x. x \ T \ e(x) = f(x)" @@ -167,11 +180,11 @@ lemma exceptEqual [simp]: shows "([f EXCEPT ![v] = e] = f) = (v \ DOMAIN f \ f[v] = e)" using assms by (auto simp: fcnEqualIff) -text {* +text \ A function can be defined from a predicate. Using the @{text CHOOSE} operator, the definition does not require the predicate to be functional. -*} +\ lemma fcnConstruct: assumes hyp: "\x \ S : \y : P(x,y)" @@ -190,7 +203,7 @@ proof - qed -subsection {* Function spaces *} +subsection \ Function spaces \ lemma inFuncSetIff: "(f \ [S \ T]) = (isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T))" @@ -267,12 +280,12 @@ next from assms show "\x \ S : ?exc[x] \ T" by auto qed (simp) -text {* +text \ The following special case is useful for invariant proofs where one proves type correctness. The additional hypotheses make the type of $f$ available and are useful, for example, when the expression $e$ is of the form $f[u]$ for some $u \in S$. -*} +\ lemma exceptInFuncSetSame: assumes "f \ [S \ T]" @@ -281,7 +294,7 @@ lemma exceptInFuncSetSame: using assms by auto -subsection {* Finite functions and extension *} +subsection \ Finite functions and extension \ lemma oneArgIsAFcn [simp, intro!]: "isAFcn(d :> e)" by (simp add: oneArg_def) @@ -321,10 +334,10 @@ lemma oneArgEqualIff [simp]: "isAFcn(f) \ ((d :> e) = f) = ((DOMAIN f = {d}) \ f[d] = e)" by auto --- {* infer equalities @{text "f = g @@ h"} *} +text \ infer equations @{text "f = g @@ h"} \ lemmas - fcnEqual[where f = "f @@ h", standard, intro!] - fcnEqual[where g = "g @@ h", standard, intro!] + fcnEqual[where f = "f @@ h" for f h, intro!] + fcnEqual[where g = "g @@ h" for g h, intro!] lemma extendEqualIff [simp]: "isAFcn(f) \ (f = g @@ h) = @@ -352,25 +365,25 @@ by auto **) -subsection {* Notions about functions *} +subsection \ Notions about functions \ -subsubsection {* Image and Range *} +subsubsection \ Image and Range \ -text {* +text \ Image of a set under a function, and range of a function. Because the application of a function to an argument outside of its domain usually leads to silliness, we restrict to the domain when defining the image. -*} +\ definition Image where "Image(f,A) \ { f[x] : x \ A \ DOMAIN f }" -text {* - The range of a function, introduced as an abbreviation (macro). +text \ + The range of a function, introduced as an abbreviation. To reason about the range, apply the theorems about @{text Image}, or simply rewrite with @{text Image_def}. -*} +\ abbreviation Range where "Range(f) \ Image(f, DOMAIN f)" @@ -390,7 +403,7 @@ lemma imageI_exEq [intro]: shows "y \ Image(f,A)" using assms by (auto intro: imageI_eq) -lemma rangeI: -- {* useful special case *} +lemma rangeI: \ \ useful special case \ assumes "\x \ DOMAIN f: y = f[x]" shows "y \ Range(f)" using assms by auto @@ -416,20 +429,20 @@ lemma imageEmpty [simp]: "({} = Image(f,A)) = (A \ DOMAIN f = {})" by auto -subsubsection {* Injective functions *} +subsubsection \ Injective functions \ definition InjectiveOn where "InjectiveOn(f,A) \ \x,y \ A \ DOMAIN f : f[x] = f[y] \ x = y" -abbreviation Injective -- {* special case: injective function *} +abbreviation Injective \ \ special case: injective function \ where "Injective(f) \ InjectiveOn(f, DOMAIN f)" definition Injections where "Injections(S,T) \ { f \ [S \ T] : Injective(f) }" lemmas - setEqualI [where A = "Injections(S,T)", standard, intro!] - setEqualI [where B = "Injections(S,T)", standard, intro!] + setEqualI [where A = "Injections(S,T)" for S T, intro!] + setEqualI [where B = "Injections(S,T)" for S T, intro!] lemma injectiveOnIsBool [intro!,simp]: "isBool(InjectiveOn(f,A))" @@ -439,9 +452,7 @@ lemma boolifyInjectiveOn [simp]: "boolify(InjectiveOn(f,A)) = InjectiveOn(f,A)" by auto -text {* - For the moment, no support by default for automatic reasoning. -*} +text \ For the moment, no support by default for automatic reasoning. \ lemma injectiveOnI: assumes "\x y. \ x \ A; x \ DOMAIN f; y \ A; y \ DOMAIN f; f[x] = f[y] \ \ x = y" @@ -460,7 +471,7 @@ lemma injectiveOnE: shows "P" using assms by (auto simp: InjectiveOn_def) -lemma injectiveOnIff: -- {* useful for simplification *} +lemma injectiveOnIff: \ \ useful for simplification \ assumes "InjectiveOn(f,A)" and "x \ A \ DOMAIN f" and "y \ A \ DOMAIN f" shows "(f[x] = f[y]) = (x = y)" using assms injectiveOnD by auto @@ -470,14 +481,12 @@ lemma injectiveOnSubset: shows "InjectiveOn(f,B)" using assms by (auto simp: InjectiveOn_def) -lemma injectiveOnDifference: +lemma injectiveOnSetminus: assumes "InjectiveOn(f,A)" - shows "InjectiveOn(f, A \\ B)" + shows "InjectiveOn(f, A \ B)" using assms by (auto simp: InjectiveOn_def) -text {* - The existence of an inverse function implies injectivity. -*} +text \The existence of an inverse function implies injectivity.\ lemma inverseThenInjective: assumes inv: "\x. \ x \ A; x \ DOMAIN f \ \ g[f[x]] = x" @@ -492,7 +501,7 @@ proof (rule injectiveOnI) from x1 y1 eq show "x = y" by simp qed -text {* Trivial cases. *} +text \ Trivial cases. \ lemma injectiveOnEmpty [intro!,simp]: "InjectiveOn(f, {})" @@ -507,7 +516,7 @@ lemma "Injective(d :> e)" by simp **) -text {* Injectivity for function extensions. *} +text \ Injectivity and EXCEPT. \ lemma injectiveOnExcept: assumes 1: "InjectiveOn(f, A \\ {v})" and 2: "isAFcn(f)" @@ -541,7 +550,7 @@ proof (rule injectiveOnI) with fx eq show "FALSE" by simp qed with x1 x2 have x3: "x \ (A \\ {v})" "x \ DOMAIN f" by auto - have "y \ v" -- {* symmetrical reasoning *} + have "y \ v" \ \symmetrical reasoning\ proof assume contr: "y = v" with True have fy: "?exc[y] = e" by auto @@ -557,7 +566,7 @@ proof (rule injectiveOnI) qed lemma injectiveOnExtend: - assumes f: "InjectiveOn(f,A)" and g: "InjectiveOn(g,A \\ DOMAIN f)" + assumes f: "InjectiveOn(f,A)" and g: "InjectiveOn(g, A \\ DOMAIN f)" and disj: "Image(f, A) \ Image(g, A \\ DOMAIN f) = {}" shows "InjectiveOn(f @@ g, A)" proof (rule injectiveOnI) @@ -596,25 +605,25 @@ proof (rule injectiveOnI) qed qed -lemma injectiveExtend: -- {* special case *} - assumes 1: "Injective(f)" and 2: "InjectiveOn(g, DOMAIN g \\ DOMAIN f)" - and 3: "Range(f) \ Image(g, DOMAIN g \\ DOMAIN f) = {}" +lemma injectiveExtend: \ \ special case \ + assumes 1: "Injective(f)" and 2: "InjectiveOn(g, DOMAIN g \ DOMAIN f)" + and 3: "Range(f) \ Image(g, DOMAIN g \ DOMAIN f) = {}" shows "Injective(f @@ g)" proof (rule injectiveOnExtend) from 1 show "InjectiveOn(f, DOMAIN (f @@ g))" by (auto simp: InjectiveOn_def) next - from 2 show "InjectiveOn(g, DOMAIN (f @@ g) \\ DOMAIN f)" + from 2 show "InjectiveOn(g, DOMAIN (f @@ g) \ DOMAIN f)" by (auto simp: InjectiveOn_def) next - show "Image(f, DOMAIN (f @@ g)) \ Image(g, DOMAIN (f @@ g) \\ DOMAIN f) = {}" + show "Image(f, DOMAIN (f @@ g)) \ Image(g, DOMAIN (f @@ g) \ DOMAIN f) = {}" proof (clarify) fix x assume xf: "x \ Image(f, DOMAIN (f @@ g))" - and xg: "x \ Image(g, DOMAIN (f @@ g) \\ DOMAIN f)" + and xg: "x \ Image(g, DOMAIN (f @@ g) \ DOMAIN f)" from xf have "x \ Range(f)" by auto moreover - from xg have "x \ Image(g, DOMAIN g \\ DOMAIN f)" by auto + from xg have "x \ Image(g, DOMAIN g \ DOMAIN f)" by auto moreover note 3 ultimately show FALSE by blast @@ -626,9 +635,9 @@ lemma injectiveOnImageInter: shows "Image(f, B \ C) = Image(f,B) \ Image(f,C)" using assms by (auto simp: InjectiveOn_def Image_def) -lemma injectiveOnImageDifference: +lemma injectiveOnImageSetminus: assumes "InjectiveOn(f,A)" and "B \ A" and "C \ A" - shows "Image(f, B \\ C) = Image(f,B) \\ Image(f,C)" + shows "Image(f, B \ C) = Image(f,B) \ Image(f,C)" using assms by (auto simp: InjectiveOn_def Image_def) lemma injectiveImageMember: @@ -639,14 +648,17 @@ using assms by (auto simp: InjectiveOn_def Image_def) lemma injectiveImageSubset: assumes f: "Injective(f)" shows "(Image(f,A) \ Image(f,B)) = (A \ DOMAIN f \ B \ DOMAIN f)" -proof (auto) -- {* the inclusion ``$\supseteq$'' is solved automatically *} - fix x - assume ab: "Image(f,A) \ Image(f,B)" and x: "x \ A" "x \ DOMAIN f" - from x have "f[x] \ Image(f,A)" by (rule imageI) - with ab have "f[x] \ Image(f,B)" .. - then obtain z where z: "z \ B \ DOMAIN f" "f[x] = f[z]" by auto - from f z x have "x = z" by (auto elim: injectiveOnD) - with z show "x \ B" by simp +proof - + { + fix x + assume ab: "Image(f,A) \ Image(f,B)" and x: "x \ A" "x \ DOMAIN f" + from x have "f[x] \ Image(f,A)" by (rule imageI) + with ab have "f[x] \ Image(f,B)" .. + then obtain z where z: "z \ B \ DOMAIN f" "f[x] = f[z]" by auto + from f z x have "x = z" by (auto elim: injectiveOnD) + with z have "x \ B" by simp + } + thus ?thesis by blast \ \the reverse inclusion is proved automatically\ qed lemma injectiveImageEqual: @@ -674,7 +686,7 @@ lemma injectionsE: using assms unfolding Injections_def InjectiveOn_def by blast -subsubsection {* Surjective functions *} +subsubsection \ Surjective functions \ definition Surjective where "Surjective(f,A) \ A \ Range(f)" @@ -683,8 +695,8 @@ definition Surjections where "Surjections(S,T) \ { f \ [S \ T] : Surjective(f,T) }" lemmas - setEqualI [where A = "Surjections(S,T)", standard, intro!] - setEqualI [where B = "Surjections(S,T)", standard, intro!] + setEqualI [where A = "Surjections(S,T)" for S T, intro!] + setEqualI [where B = "Surjections(S,T)" for S T, intro!] lemma surjectiveIsBool [intro!,simp]: "isBool(Surjective(f,A))" @@ -704,10 +716,10 @@ lemma surjectiveD: shows "\x \ DOMAIN f : y = f[x]" using assms by (auto simp: Surjective_def Image_def) -text {* +text \ @{text "\ Surjective(f,A); y \ A; \x. \ x \ DOMAIN f; y = f[x] \ \ P \ \ P"} -*} -lemmas surjectiveE = surjectiveD[THEN bExE, standard] +\ +lemmas surjectiveE = surjectiveD[THEN bExE] lemma surjectiveRange: shows "Surjective(f, Range(f))" @@ -749,20 +761,20 @@ lemma surjectionsRange: using assms by (rule surjectionsE, auto) -subsubsection {* Bijective functions *} +subsubsection \ Bijective functions \ -text {* +text \ Here we do not define a predicate @{text Bijective} because it would require a set parameter for the codomain and would therefore be curiously asymmetrical. -*} +\ definition Bijections where "Bijections(S,T) \ Injections(S,T) \ Surjections(S,T)" lemmas - setEqualI [where A = "Bijections(S,T)", standard, intro!] - setEqualI [where B = "Bijections(S,T)", standard, intro!] + setEqualI [where A = "Bijections(S,T)" for S T, intro!] + setEqualI [where B = "Bijections(S,T)" for S T, intro!] lemma bijectionsI [intro!]: assumes "f \ [S \ T]" and "Injective(f)" and "Surjective(f,T)" @@ -786,7 +798,7 @@ lemma bijectionsE: using 1 by (intro 2, auto simp: Bijections_def Injections_def Surjections_def) -subsubsection {* Inverse of a function *} +subsubsection \ Inverse of a function \ definition Inverse where "Inverse(f) \ [ y \ Range(f) \ CHOOSE x \ DOMAIN f : f[x] = y ]" @@ -842,9 +854,9 @@ proof (intro injectiveOnI, auto) ultimately show "f[x] = f[x']" by simp qed -text {* +text \ For injective functions, @{text Inverse} really inverts the function. -*} +\ lemma injectiveInverse: assumes f: "Injective(f)" and x: "x \ DOMAIN f" @@ -879,9 +891,7 @@ proof (rule surjectiveI) with x show "\y \ DOMAIN Inverse(f) : x = Inverse(f)[y]" by auto qed -text {* - The inverse of a bijection is a bijection. -*} +text \ The inverse of a bijection is a bijection. \ lemma inverseBijections: assumes f: "f \ Bijections(S,T)" diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f8ddcd09..f558c286 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -1,46 +1,41 @@ (* Title: TLA+/NatOrderings.thy Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:28 merz> + Version: Isabelle2020 *) -header {* Orders on natural numbers *} +section \ Orders on natural numbers \ theory NatOrderings imports Peano begin -text {* +text \ Using the sets @{text upto} we can now define the standard ordering on natural numbers. The constant @{text "\"} is defined over the naturals by the axiom (conditional definition) @{text "nat_leq_def"} below; it should be defined over other domains as appropriate later on. We generally define the constant @{text "<"} such that @{text "ab \ a\b"}, over any domain. -*} + iff @{text "a\b \ a\b"}, for any values. +\ -definition leq :: "[c,c] \ c" (infixl "<=" 50) -(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m <= n) \ (m \ upto(n))"*) -where nat_leq_def: "(m <= n) \ (m \ upto(n))" +definition leq :: "[c,c] \ c" (infixl "\" 50) +(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m \ n) \ (m \ upto(n))"*) +where nat_leq_def: "(m \ n) \ (m \ upto(n))" abbreviation (input) - geq :: "[c,c] \ c" (infixl ">=" 50) -where "x >= y \ y <= x" + geq :: "[c,c] \ c" (infixl "\" 50) +where "x \ y \ y \ x" -notation (xsymbols) - leq (infixl "\" 50) and - geq (infixl "\" 50) +notation (ASCII) + leq (infixl "<=" 50) and + geq (infixl ">=" 50) -notation (HTML output) - leq (infixl "\" 50) and - geq (infixl "\" 50) - -subsection {* Operator definitions and generic facts about @{text "<"} *} +subsection \ Operator definitions and generic facts about @{text "<"} \ definition less :: "[c,c] \ c" (infixl "<" 50) where "a < b \ a \ b \ a \ b" @@ -81,7 +76,7 @@ declare leq_neq_trans[simplified,trans] lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" by auto -subsection {* Facts about @{text "\"} over @{text Nat} *} +subsection \ Facts about @{text "\"} over @{text Nat} \ lemma nat_boolify_leq [simp]: "boolify(m \ n) = (m \ n)" by (simp add: nat_leq_def) @@ -92,7 +87,7 @@ by (simp add: nat_leq_def) lemma nat_leq_refl [intro,simp]: "n \ Nat \ n \ n" unfolding nat_leq_def by (rule uptoRefl) -lemma eq_leq_bothE: -- {* reduce equality over integers to double inequality *} +lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ assumes "m \ Nat" and "n \ Nat" and "m = n" and "\ m \ n; n \ m \ \ P" shows "P" using assms by simp @@ -103,55 +98,64 @@ unfolding nat_leq_def by (rule zeroInUpto) lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" by (simp add: nat_leq_def uptoZero) -lemma nat_leq_SuccI [elim!,simp]: (** FIXME: why simp ? **) - assumes "m \ n" and "m \ Nat" and "n \ Nat" +lemma nat_leq_SuccI [elim!(*,simp*)]: + assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" shows "m \ Succ[n]" using assms by (auto simp: nat_leq_def uptoSucc) -lemma nat_leq_Succ: (** FIXME: make this simp ? **) - assumes "m \ Nat" and "n \ Nat" +lemma nat_leq_Succ: + assumes (*"m \ Nat" and*) "n \ Nat" shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" using assms by (auto simp: nat_leq_def uptoSucc) lemma nat_leq_SuccE [elim]: - assumes "m \ Succ[n]" and "m \ Nat" and "n \ Nat" + assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" and "m \ n \ P" and "m = Succ[n] \ P" shows "P" using assms by (auto simp: nat_leq_Succ) lemma nat_leq_limit: - assumes "m \ n" and "\(Succ[m] \ n)" and "m \ Nat" and "n \ Nat" + assumes "m \ n" and "\(Succ[m] \ n)" and (*"m \ Nat" and*) "n \ Nat" shows "m=n" using assms by (auto simp: nat_leq_def intro: uptoLimit) lemma nat_leq_trans [trans]: - assumes "k \ m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" + assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" shows "k \ n" using assms by (auto simp: nat_leq_def elim: uptoTrans) lemma nat_leq_antisym: - assumes "m \ n" and "n \ m" and "m \ Nat" and "n \ Nat" + assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) shows "m = n" using assms by (auto simp add: nat_leq_def elim: uptoAntisym) lemma nat_Succ_not_leq_self [simp]: assumes n: "n \ Nat" shows "(Succ[n] \ n) = FALSE" -using n by (auto dest: nat_leq_antisym) +using n by (auto simp: nat_leq_Succ dest: nat_leq_antisym) -lemma nat_Succ_leqD: +lemma nat_Succ_leqI: + assumes "m \ n" and "m \ Nat" and "n \ Nat" and "m \ n" + shows "Succ[m] \ n" +using assms by (auto dest: nat_leq_limit) + +lemma nat_Succ_leqD [elim]: assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" shows "m \ n" proof - - from m have "m \ Succ[m]" by simp + from m have "m \ Succ[m]" by (simp add: nat_leq_Succ) with leq m n show ?thesis by (elim nat_leq_trans, auto) qed -lemma nat_Succ_leq_Succ: - (* needn't be added to simp: consequence of nat_Succ_leq_iff_less and nat_less_Succ_iff_leq *) +lemma nat_Succ_leq_Succ [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(Succ[m] \ Succ[n]) = (m \ n)" -using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit elim: nat_Succ_leqD) +using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit) + +lemma nat_Succ_leq: + assumes "m \ Nat" and "n \ Nat" + shows "(Succ[m] \ n) = (m \ n \ m \ n)" + using assms by (auto intro: nat_Succ_leqI) lemma nat_leq_linear: "\m \ Nat; n \ Nat\ \ m \ n \ n \ m" unfolding nat_leq_def using uptoLinear . @@ -176,13 +180,13 @@ next qed qed -lemma nat_leq_induct: -- {* sometimes called ``complete induction'' *} +lemma nat_leq_induct: \ \sometimes called ``complete induction''\ assumes "P(0)" and "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(Succ[n])" shows "\n\Nat : P(n)" proof - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp: nat_leq_Succ) + by (intro natInduct, auto simp add: nat_leq_Succ) thus ?thesis by (blast dest: nat_leq_refl) qed @@ -193,39 +197,51 @@ lemma nat_leq_inductE: using assms by (blast dest: nat_leq_induct) -subsection {* Facts about @{text "<"} over @{text Nat} *} +subsection \ Facts about @{text "<"} over @{text Nat} \ -lemma nat_Succ_leq_iff_less [simp]: +lemma nat_less_iff_Succ_leq [simp]: assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] \ n) = (m < n)" + shows "(m < n) = (Succ[m] \ n)" using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) --- {* alternative definition of @{text "<"} over @{text Nat} *} -lemmas nat_less_iff_Succ_leq = sym[OF nat_Succ_leq_iff_less, standard] - -text {* Reduce @{text "\"} to @{text "<"}. *} +text \ Reduce @{text "\"} to @{text "<"}. \ -lemma nat_leq_less: -- {* premises needed for @{text "isBool(m\n)"} and reflexivity *} - assumes "m \ Nat" and "n \ Nat" +lemma nat_leq_less: + assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" -using assms by (auto simp: less_def) + using assms by (auto simp: less_def) -lemma nat_less_Succ_iff_leq [simp]: +lemma nat_leqE: + assumes "m \ n" and "n \ Nat" and 1: "m P" and 2: "m=n \ P" + shows "P" +proof - + from \m \ n\ \n \ Nat\ have "m < n \ m=n" + by (auto simp add: less_def) + with 1 2 show ?thesis by blast +qed + +lemma nat_less_Succ_iff_leq (*[simp]*): assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m \ n)" -using assms -by (simp del: nat_Succ_leq_iff_less add: nat_less_iff_Succ_leq nat_Succ_leq_Succ) + using assms by simp + +lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq] -lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq, standard] +lemma not_leq_Succ_leq [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "((m \ n) = FALSE) = (Succ[n] \ m)" +proof - + from assms nat_leq_linear[of m n] have "((m \ n) = FALSE) = (n < m)" + by (auto simp: less_def dest: nat_leq_antisym) + with assms show ?thesis by simp +qed lemma nat_not_leq_one: assumes "n \ Nat" shows "(\(1 \ n)) = (n = 0)" -using assms by (cases, auto) - -declare nat_not_leq_one[simplified,simp] +using assms by simp -text {* @{text "<"} and @{text "Succ"}. *} +text \ @{text "<"} and @{text "Succ"}. \ lemma nat_Succ_less_mono: assumes "m \ Nat" and "n \ Nat" @@ -240,7 +256,7 @@ using assms by simp lemma nat_not_less0 [simp]: assumes "n \ Nat" shows "(n < 0) = FALSE" -using assms by (auto simp: less_def) +using assms by auto lemma nat_less0E (*[elim!]*): assumes "n < 0" and "n \ Nat" @@ -255,14 +271,14 @@ using assms by auto lemma nat_Succ_lessD: assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" shows "m < n" -using 1[unfolded less_def] 2 3 by simp +using assms by auto lemma nat_less_leq_not_leq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m < n) = (m \ n \ \ n \ m)" using assms by (auto simp: less_def dest: nat_leq_antisym) -text {* Transitivity. *} +text \ Transitivity. \ lemma nat_less_trans (*[trans]*): assumes "k < m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" @@ -289,21 +305,22 @@ lemma nat_less_leq_trans [trans]: shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) -text {* Asymmetry. *} +text \ Asymmetry. \ lemma nat_less_not_sym: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "(n < m) = FALSE" -using assms by (simp add: nat_less_leq_not_leq) - +using assms by auto + lemma nat_less_asym: assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" shows "P" proof (rule contradiction) - assume "\P" with assms show "FALSE" by (auto dest: nat_less_not_sym) + assume "\P" with nat_less_not_sym[OF \m \m \ Nat\ \n \ Nat\] \\P \ n < m\ + show "FALSE" by simp qed -text {* Linearity (totality). *} +text \ Linearity (totality). \ lemma nat_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" @@ -338,9 +355,9 @@ unfolding nat_not_less[OF m n] using assms by simp lemma nat_not_leq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m \ n) = (n < m)" -using assms by (simp add: sym[OF nat_not_less]) +using assms by simp --- {* often useful, but not active by default *} +text \often useful, but not active by default\ lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq lemma nat_not_leq_eq: @@ -366,7 +383,7 @@ using assms by (auto simp: nat_leq_less) lemma nat_antisym_conv2: assumes "m \ n" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_antisym_conv1) +using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) lemma nat_antisym_conv3: assumes "\ n < m" and "m \ Nat" and "n \ Nat" @@ -388,9 +405,9 @@ lemma nat_gt0_not0 (*[simp]*): shows "(0 < n) = (n \ 0)" using assms by (auto simp: nat_neq_iff[simplified]) -lemmas nat_neq0_conv = sym[OF nat_gt0_not0, standard] +lemmas nat_neq0_conv = sym[OF nat_gt0_not0] -text {* Introduction properties *} +text \ Introduction properties \ lemma nat_less_Succ_self (*[iff]*): assumes "n \ Nat" @@ -402,37 +419,37 @@ lemma nat_zero_less_Succ (*[iff]*): shows "0 < Succ[n]" using assms by simp -text {* Elimination properties. *} +text \ Elimination properties \ lemma nat_less_Succ: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m < n \ m = n)" -using assms by (simp add: nat_leq_less) +using assms by (simp add: nat_leq_less[of n m]) lemma nat_less_SuccE: assumes "m < Succ[n]" and "m \ Nat" and "n \ Nat" and "m < n \ P" and "m = n \ P" shows P -using assms by (auto simp: nat_leq_less) +using assms by (auto simp: nat_leq_less[of n m]) lemma nat_less_one (*[simp]*): assumes "n \ Nat" shows "(n < 1) = (n = 0)" using assms by simp -text {* "Less than" is antisymmetric, sort of. *} +text \ "Less than" is antisymmetric, sort of \ lemma nat_less_antisym: assumes m: "m \ Nat" and n: "n \ Nat" shows "\ \(n < m); n < Succ[m] \ \ m = n" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -text {* Lifting @{text "<"} monotonicity to @{text "\"} monotonicity. *} +text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ lemma less_mono_imp_leq_mono: assumes i: "i \ Nat" and j: "j \ Nat" and f: "\n \ Nat : f(n) \ Nat" and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" shows "f(i) \ f(j)" -using ij proof (auto simp: nat_leq_less[OF i j]) +using ij proof (auto simp: nat_leq_less[OF j]) assume "i < j" with i j have "f(i) < f(j)" by (rule mono) thus "f(i) \ f(j)" by (simp add: less_imp_leq) @@ -440,7 +457,7 @@ next from j f show "f(j) \ f(j)" by auto qed -text {* Inductive (?) properties. *} +text \ Inductive (?) properties. \ lemma nat_Succ_lessI: assumes "m \ Nat" and "n \ Nat" and "m < n" and "Succ[m] \ n" @@ -458,9 +475,8 @@ proof - fix n assume n: "n \ Nat" and 1: "i < Succ[n]" and ih: "i < n \ \j\Nat : i \ j \ n = Succ[j]" - from i n 1 have "i < n \ i = n" by (simp add: nat_leq_less) - thus "\j\Nat : i \ j \ Succ[n] = Succ[j]" - proof + from 1 i n show "\j\Nat : i \ j \ Succ[n] = Succ[j]" + proof (rule nat_less_SuccE) assume "i < n" then obtain j where "j \ Nat" and "i \ j" and "n = Succ[j]" by (blast dest: ih) @@ -478,31 +494,48 @@ qed lemma nat_Succ_lessE: assumes major: "Succ[i] < k" and i: "i \ Nat" and k: "k \ Nat" obtains j where "j \ Nat" and "i < j" and "k = Succ[j]" -using assms by (auto elim: nat_lessE) +proof - + from i have "Succ[i] \ Nat" .. + from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" + by (rule nat_lessE) + with i that show ?thesis by simp +qed +(* lemma nat_gt0_implies_Succ: assumes 1: "0 < n" and 2: "n \ Nat" shows "\m : m \ Nat \ n = Succ[m]" using 2 1 by (cases, auto) +*) +lemma nat_ge1_implies_Succ: + assumes 1: "1 \ n" and 2: "n \ Nat" + shows "\m : m \ Nat \ n = Succ[m]" + using 2 1 by (cases, auto) lemma nat_gt0_iff_Succ: assumes n: "n \ Nat" shows "(0 < n) = (\m \ Nat: n = Succ[m])" -using n by (auto dest: nat_gt0_implies_Succ) +using n by (auto dest: nat_ge1_implies_Succ) +(* lemma nat_less_Succ_eq_0_disj: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" using assms by (induct m, auto) +*) +lemma nat_le_eq_0_disj: + assumes "m \ Nat" and "n \ Nat" + shows "(m \ n) = (m = 0 \ (\j\Nat : m = Succ[j] \ j < n))" + using assms by (induct m, auto) lemma nat_less_antisym_false: "\m < n; m \ Nat; n \ Nat\ \ n < m = FALSE" - unfolding less_def using nat_leq_antisym by auto + by auto lemma nat_less_antisym_leq_false: "\m < n; m \ Nat; n \ Nat\ \ n \ m = FALSE" - unfolding less_def using nat_leq_antisym[of m n] by auto + by auto -subsection {* Intervals of natural numbers *} +subsection \Intervals of natural numbers\ definition natInterval :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) where "m .. n \ { k \ Nat : m \ k \ k \ n }" @@ -518,11 +551,11 @@ lemma inNatIntervalE [elim]: using 1 by (intro 2, auto simp add: natInterval_def) lemma inNatInterval_iff: "(k \ m .. n) = (k \ Nat \ m \ k \ k \ n)" -using assms by auto +by auto lemmas - setEqualI [where A = "m .. n", standard, intro] - setEqualI [where B = "m .. n", standard, intro] + setEqualI [where A = "m .. n" for m n, intro] + setEqualI [where B = "m .. n" for m n, intro] lemma lowerInNatInterval [iff]: assumes "m \ n" and "m \ Nat" @@ -538,8 +571,8 @@ lemma gtNotinNatInterval: assumes gt: "m > n" and k: "k \ m .. n" and m: "m \ Nat" and n: "n \ Nat" shows "P" proof - - from k have 1: "m \ k" and 2: "k \ n" and 3: "k \ Nat" by auto - from 1 2 m 3 n have "m \ n" by (rule nat_leq_trans) + from k have 1: "m \ k" and 2: "k \ n" by auto + from 1 2 n have "m \ n" by (rule nat_leq_trans) with m n have "\(n < m)" by (simp add: nat_not_order_simps) from this gt show ?thesis .. qed @@ -576,7 +609,7 @@ using assms by (auto simp: natInterval_def) lemma succNatInterval: assumes "m \ Nat" and "n \ Nat" shows "Succ[m] .. n = (m .. n \ {m})" -using assms by (auto simp: natInterval_def) +using assms by (auto simp: natInterval_def nat_Succ_leq) lemma natIntervalEqual_iff: assumes k: "k \ Nat" and l: "l \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -599,10 +632,10 @@ proof - from False eq m n have 12: "m \ n" by (simp only: natIntervalEmpty_iff nat_not_less) from 11 k eq have 13: "m \ k" by auto from 12 m eq have 14: "k \ m" by auto - from 14 13 k m have 15: "k = m" by (rule nat_leq_antisym) + from 14 13 have 15: "k = m" by (rule nat_leq_antisym) from 11 l eq have 16: "l \ n" by auto from 12 n eq have 17: "n \ l" by auto - from 16 17 l n have "l = n" by (rule nat_leq_antisym) + from 16 17 have "l = n" by (rule nat_leq_antisym) with 15 show ?rhs by blast qed qed diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index ba06d431..6b848505 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -1,30 +1,29 @@ (* Title: TLA+/Peano.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:03 merz> + Version: Isabelle2020 *) -header {* Peano's axioms and natural numbers *} +section \ Peano's axioms and natural numbers \ theory Peano imports FixedPoints Functions begin -text {* +text \ As a preparation for the definition of numbers and arithmetic in \tlaplus{}, we state Peano's axioms for natural numbers and prove the existence of a structure satisfying them. The presentation of the axioms is somewhat simplified compared to the \tlaplus{} book. (Moreover, the existence of such a structure is assumed, but not proven in the book.) -*} +\ -subsection {* The Peano Axioms *} +subsection \ The Peano Axioms \ definition PeanoAxioms :: "[c,c,c] \ c" where - -- {* parameters: the set of natural numbers, zero, and succ function *} + \ \parameters: the set of natural numbers, zero, and successor function\ "PeanoAxioms(N,Z,Sc) \ Z \ N \ Sc \ [N \ N] @@ -32,7 +31,7 @@ definition PeanoAxioms :: "[c,c,c] \ c" where \ (\m,n \ N: Sc[m] = Sc[n] \ m = n) \ (\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S)" -text {* +text \ The existence of a structure satisfying Peano's axioms is proven following the standard ZF construction where @{text "{}"} is zero, @{text "i \ {i}"} is taken as the successor of any natural @@ -42,15 +41,15 @@ text {* set theory). In \tlaplus{}, natural numbers are defined by a sequence of @{text CHOOSE}'s below, so there is no commitment to that particular structure. -*} +\ theorem peanoExists: "\N,Z,Sc : PeanoAxioms(N,Z,Sc)" proof - - let ?sc = "\n. addElt(n,n)" -- {* successor \emph{operator} *} - def expand \ "\S. {{}} \ { ?sc(n) : n \ S}" - def N \ "lfp(infinity, expand)" - def Z \ "{}" - def Sc \ "[n \ N \ ?sc(n)]" -- {* successor \emph{function} *} + let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ + define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" + define N where "N \ lfp(infinity, expand)" + define Z where "Z \ {}" + define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ have mono: "Monotonic(infinity, expand)" using infinity by (auto simp: Monotonic_def expand_def) hence expandN: "expand(N) \ N" @@ -139,12 +138,12 @@ proof - qed -subsection {* Natural numbers: definition and elementary theorems *} +subsection \ Natural numbers: definition and elementary theorems \ -text {* +text \ The structure of natural numbers is now defined to be some set, zero, and successor satisfying Peano's axioms. -*} +\ definition Succ :: "c" where "Succ \ CHOOSE Sc : \N,Z : PeanoAxioms(N,Z,Sc)" @@ -203,8 +202,8 @@ proof - qed lemmas - setEqualI [where A = "Nat", standard, intro!] - setEqualI [where B = "Nat", standard, intro!] + setEqualI [where A = "Nat", intro!] + setEqualI [where B = "Nat", intro!] lemma zeroIsNat [intro!,simp]: "0 \ Nat" using peanoNatZeroSucc by (simp add: PeanoAxioms_def) @@ -215,9 +214,9 @@ using peanoNatZeroSucc by (simp add: PeanoAxioms_def) lemma succIsAFcn [intro!,simp]: "isAFcn(Succ)" using succInNatNat by blast --- {* @{text "DOMAIN Succ = Nat"} *} +\ \@{text "DOMAIN Succ = Nat"}\ lemmas domainSucc [intro!,simp] = funcSetDomain[OF succInNatNat] --- {* @{text "n \ Nat \ Succ[n] \ Nat"} *} +\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ lemmas succIsNat [intro!,simp] = funcSetValue[OF succInNatNat] lemma oneIsNat [intro!,simp]: "1 \ Nat" @@ -236,7 +235,7 @@ lemma [simp]: shows "(0 = Succ[n]) = FALSE" using assms by (auto dest: sym) -lemma succNotZero (*[elim] -- yields "ignoring weak elimination rule" *): +lemma succNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): "\Succ[n] = 0; n \ Nat\ \ P" "\0 = Succ[n]; n \ Nat\ \ P" by (simp+) @@ -262,7 +261,7 @@ proof - thus ?thesis by auto qed --- {* version of above suitable for the inductive reasoning package *} +\ \version of above suitable for the inductive reasoning package\ lemma natInductE [case_names 0 Succ, induct set: Nat]: assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(Succ[n])" shows "P(n)" @@ -313,7 +312,7 @@ lemma succIrrefl_iff [simp]: by (auto dest: succIrrefl) --- {* Induction over two parameters along the ``diagonal''. *} +text \Induction over two parameters along the ``diagonal''.\ lemma diffInduction: assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, Succ[n])" and step: "\m,n\Nat : P(m,n) \ P(Succ[m], Succ[n])" @@ -351,23 +350,23 @@ lemma not0_implies_Suc: "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" by(rule natCases, auto) -subsection {* Initial intervals of natural numbers and ``less than'' *} +subsection \ Initial intervals of natural numbers and ``less than'' \ -text {* +text \ The set of natural numbers up to (and including) a given $n$ is inductively defined as the smallest set of natural numbers that contains $n$ and that is closed under predecessor. NB: ``less than'' is not first-order definable from the Peano axioms, a set-theoretic definition such as the following seems to be unavoidable. -*} +\ definition upto :: "c \ c" where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : Succ[k] \ S })" lemmas - setEqualI [where A = "upto(n)", standard, intro!] - setEqualI [where B = "upto(n)", standard, intro!] + setEqualI [where A = "upto(n)" for n, intro!] + setEqualI [where B = "upto(n)" for n, intro!] lemma uptoNat: "upto(n) \ Nat" unfolding upto_def by (rule lfpSubsetDomain) @@ -399,20 +398,20 @@ lemma uptoSucc: assumes n: "n \ Nat" shows "upto(Succ[n]) = upto(n) \ {Succ[n]}" (is "?lhs = ?rhs") proof - - let "?preds(S)" = "{k \ Nat : Succ[k] \ S}" - let "?f(S,k)" = "{k} \ ?preds(S)" + let ?preds = "\S. {k \ Nat : Succ[k] \ S}" + let ?f = "\S k. {k} \ ?preds(S)" have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" by (auto simp: Monotonic_def) - -- {* ``$\subseteq$'' *} + \ \``$\subseteq$''\ from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto also have "\ \ upto(n)" by (unfold upto_def, rule lfpPreFP, rule mono, rule n) finally have "?f(?rhs, Succ[n]) \ ?rhs" by auto moreover from n have "?rhs \ Nat" - by (intro cupLUB, auto elim: uptoNat[THEN subsetD]) + by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) ultimately have 1: "?lhs \ ?rhs" by (unfold upto_def[where n="Succ[n]"], rule lfpLB) - -- {* ``$\supseteq$'' *} + \ \``$\supseteq$''\ from n mono have 2: "?f(?lhs, Succ[n]) \ ?lhs" unfolding upto_def by (intro lfpPreFP, blast) with n have "?f(?lhs, n) \ ?lhs" by auto @@ -496,8 +495,8 @@ proof - fix n assume "Succ[m] \ upto(n)" and "n \ upto(m)" from this m have "Succ[m] \ upto(m)" by (rule uptoTrans) - with m show "Succ[m] = n" -- {* contradiction *} - by (blast dest: succNotinUpto) + with m show "Succ[m] = n" \ \contradiction\ + by (blast dest: succNotinUpto) qed qed with m n mn nm show ?thesis by blast @@ -540,13 +539,13 @@ next qed -subsection {* Primitive Recursive Functions *} +subsection \ Primitive Recursive Functions \ -text {* +text \ We axiomatize a primitive recursive scheme for functions with one argument and domain on natural numbers. Later, we use it to define addition, multiplication and difference. -*} +\ (** FIXME: replace axiom with proper fixpoint definition **) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index a8e713dd..80d69b78 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -2,79 +2,82 @@ Author: Stephan Merz, LORIA Copyright (C) 2008-2011 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:37 merz> + Version: Isabelle2017 *) -header {* First-order logic for TLA+ *} +section \First-Order Logic for TLA+\ theory PredicateLogic -imports Pure -uses - "~~/src/Tools/misc_legacy.ML" - "~~/src/Tools/intuitionistic.ML" - "~~/src/Provers/splitter.ML" - "~~/src/Provers/hypsubst.ML" - "~~/src/Tools/atomize_elim.ML" - "~~/src/Provers/classical.ML" - "~~/src/Provers/blast.ML" - "~~/src/Provers/quantifier1.ML" - "~~/src/Provers/clasimp.ML" - "~~/src/Tools/IsaPlanner/isand.ML" "~~/src/Tools/IsaPlanner/rw_tools.ML" "~~/src/Tools/IsaPlanner/rw_inst.ML" "~~/src/Tools/IsaPlanner/zipper.ML" "~~/src/Tools/eqsubst.ML" - "~~/src/Tools/induct.ML" - ("simplifier_setup.ML") + imports Pure + keywords + "print_claset" "print_induct_rules" :: diag begin +ML_file "~~/src/Tools/misc_legacy.ML" +ML_file "~~/src/Tools/intuitionistic.ML" +ML_file "~~/src/Tools/IsaPlanner/isand.ML" +ML_file "~~/src/Tools/IsaPlanner/rw_inst.ML" +ML_file "~~/src/Tools/IsaPlanner/zipper.ML" +ML_file "~~/src/Provers/splitter.ML" +ML_file "~~/src/Provers/hypsubst.ML" +ML_file "~~/src/Provers/classical.ML" +ML_file "~~/src/Provers/blast.ML" +ML_file "~~/src/Tools/atomize_elim.ML" +ML_file "~~/src/Provers/quantifier1.ML" +ML_file "~~/src/Provers/clasimp.ML" +ML_file "~~/src/Tools/eqsubst.ML" +ML_file "~~/src/Tools/induct.ML" declare [[ eta_contract = false ]] (* eta contraction doesn't make much sense for first-order logics *) -text {* +text \ We define classical first-order logic as a basis for an encoding of \tlaplus{}. \tlaplus{} is untyped, to the extent that it does not even distinguish between terms and formulas. We therefore - declare a single type $c$ that represents the universe of - ``constants'' rather than introducing the traditional types $i$ and - $o$ of first-order logic that, for example, underly Isabelle/ZF. -*} + declare a single type @{text \c\} that represents the universe of + ``constants'' rather than introducing the types @{text \i\} + and @{text \o\} representing non-Boolean and Boolean values that + are traditional in first-order logic. +\ (* First-order syntax for applications: f(x,y) instead of (f x y) *) setup Pure_Thy.old_appl_syntax_setup -setup {* Intuitionistic.method_setup @{binding iprover} *} +setup \ Intuitionistic.method_setup @{binding iprover} \ typedecl c -text {* +text \ The following (implicit) lifting from the object to the Isabelle meta level is always needed when formalizing a logic. It corresponds - to judgments $\vdash F$ or $\models F$ in standard presentations, + to judgments @{text \\ F\} or @{text \\ F\} in textbook presentations, asserting that a formula is considered true (either because it is an assumption or because it is a theorem). -*} +\ judgment Trueprop :: "c \ prop" ("(_)" 5) -subsection {* Equality *} +subsection \ Equality \ -text {* +text \ The axioms for equality are reflexivity and a rule that asserts that equal terms are interchangeable at the meta level (this is essential for setting up Isabelle's rewriting machinery). In particular, we can derive a substitution rule. -*} +\ axiomatization "eq" :: "c \ c \ c" (infixl "=" 50) where - refl [intro!]: "a = a" + refl [Pure.intro!]: "a = a" and eq_reflection: "t = u \ t \ u" -text {* +text \ Left and right hand sides of definitions are equal. This is the converse of axiom @{thm eq_reflection}. -*} +\ theorem meta_to_obj_eq: assumes df: "t \ u" @@ -97,19 +100,20 @@ proof (elim subst) show "a=a" .. qed +text \@{term "b=a \ P(a) \ P(b)"}\ +lemmas ssubst = sym[THEN subst] + theorem trans [trans]: "a=b \ b=c \ a=c" by (rule subst) -theorems ssubst = sym[THEN subst] - -text {* +text \ \textsc{let}-expressions in \tlaplus{} expressions. \emph{Limitation:} bindings cannot be unfolded selectively. Rewrite with @{text Let_def} in order to expand \emph{all} bindings within an expression or a context. -*} +\ definition Let :: "'b \ ('b \ 'a) \ 'a" @@ -135,95 +139,90 @@ translations "LET x \ a IN e" \ "CONST Let(a, (%x. e))" -subsection {* Propositional logic *} +subsection \ Propositional logic \ -text {* +text \ Propositional logic is introduced in a rather non-standard way by declaring constants @{text TRUE} and @{text FALSE} as well as - conditional expressions. The Boolean connectives - are defined such that they always return either @{text TRUE} or - @{text FALSE}, irrespectively of their arguments. This allows us - to prove many equational laws of propositional logic, which is - useful for automatic reasoning based on rewriting. Note that we - have equivalence as well as equality. The two relations agree over + conditional expressions. The ordinary Boolean connectives + are defined in terms of these basic operators in such a way + that they always return either @{text TRUE} or @{text FALSE}, + irrespectively of their arguments. We also require that equality + returns a Boolean value. In this way, we can prove standard + equational laws of propositional logic, which is useful for + automatic reasoning based on rewriting. Note that we have + equivalence as well as equality. The two relations agree over Boolean values, but equivalence may be stricter weaker than equality over non-Booleans. -*} - -consts - TRUE :: "c" - FALSE :: "c" - cond :: "c \ c \ c \ c" ("(IF (_)/ THEN (_)/ ELSE (_))" 10) -consts - Not :: "c \ c" ("~ _" [40] 40) - conj :: "c \ c \ c" (infixr "&" 35) - disj :: "c \ c \ c" (infixr "|" 30) - imp :: "c \ c \ c" (infixr "=>" 25) - iff :: "c \ c \ c" (infixr "<=>" 25) - -abbreviation not_equal :: "c \ c \ c" (infixl "~=" 50) -where "x ~= y \ ~ (x = y)" - -notation (xsymbols) - Not ("\ _" [40] 40) -and conj (infixr "\" 35) -and disj (infixr "\" 30) -and imp (infixr "\" 25) -and iff (infixr "\" 25) -and not_equal (infix "\" 50) - -notation (HTML output) - Not ("\ _" [40] 40) -and conj (infixr "\" 35) -and disj (infixr "\" 30) -and imp (infixr "\" 25) -and iff (infixr "\" 25) -and not_equal (infix "\" 50) - -defs - not_def: "\ A \ IF A THEN FALSE ELSE TRUE" - conj_def: "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE FALSE" - disj_def: "A \ B \ IF A THEN TRUE ELSE IF B THEN TRUE ELSE FALSE" - imp_def: "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE TRUE" - iff_def: "A \ B \ (A \ B) \ (B \ A)" - -text {* - We adopt the following axioms of propositional logic: + The following axioms govern propositional logic: \begin{enumerate} - \item $A$ is a theorem if and only if it equals @{term TRUE}. - \item @{term FALSE} (more precisely, @{text "\ TRUE"}) implies anything. - \item Conditionals are reasoned about by case distinction. + \item @{term TRUE} is valid, and @{term FALSE} implies anything}, + \item @{text A} is a theorem if and only if it equals @{term TRUE}, + \item conditionals are reasoned about by case distinction, + \item equations evaluate to @{term TRUE} or @{term FALSE}. \end{enumerate} - We also assert that the equality predicate returns either @{term TRUE} - or @{term FALSE}. -*} +\ -axiomatization where - trueI [intro!]: "TRUE" +axiomatization + TRUE :: "c" and - eqTrueI: "A \ A = TRUE" + FALSE :: "c" and - notTrueE [elim!]: "\ TRUE \ A" + cond :: "c \ c \ c \ c" ("(IF (_)/ THEN (_)/ ELSE (_))" 10) +where + trueI [Pure.intro!]: "TRUE" +and + eqTrueI: "A \ A = TRUE" +and + falseE [Pure.elim!]: "FALSE \ A" and - condI: "\ A \ P(t); \ A \ P(e) \ \ P(IF A THEN t ELSE e)" + condI: "\ A \ P(t); (A \ FALSE) \ P(e) \ \ P(IF A THEN t ELSE e)" and - eqBoolean : "x \ y \ (x=y) = FALSE" + eqBoolean : "(x = y \ FALSE) \ (x=y) = FALSE" + + +definition not :: "c \ c" ("\ _" [40] 40) + where "\ A \ IF A THEN FALSE ELSE TRUE" + +definition conj :: "c \ c \ c" (infixr "\" 35) + where "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE FALSE" -text {* - We now derive the standard proof rules of propositional logic. +definition disj :: "c \ c \ c" (infixr "\" 30) + where "A \ B \ IF A THEN TRUE ELSE IF B THEN TRUE ELSE FALSE" + +definition imp :: "c \ c \ c" (infixr "\" 25) + where "A \ B \ IF A THEN (IF B THEN TRUE ELSE FALSE) ELSE TRUE" + +definition iff :: "c \ c \ c" (infixr "\" 25) + where "A \ B \ (A \ B) \ (B \ A)" + +abbreviation not_equal :: "c \ c \ c" (infixl "\" 50) + where "x \ y \ \(x = y)" + +notation (ASCII) + not ("~ _" [40] 40) +and conj (infixr "&" 35) +and disj (infixr "|" 30) +and imp (infixr "=>" 25) +and iff (infixr "<=>" 25) +and not_equal (infix "~=" 50) + +text \ + We now derive the standard laws of propositional logic. The first lemmas are about @{term TRUE}, @{term FALSE}, and conditional expressions. -*} +\ -lemma eqTrueD: -- {* converse of eqTrueI *} +text \ Converse of eqTrueI. \ +lemma eqTrueD: assumes a: "A = TRUE" shows "A" by (unfold a[THEN eq_reflection], rule trueI) -text {* +text \ Assumption @{term TRUE} is useless and can be deleted. -*} +\ lemma TrueAssumption: "(TRUE \ PROP P) == PROP P" proof @@ -237,23 +236,16 @@ lemma condT: "(IF TRUE THEN t ELSE e) = t" proof (rule condI) show "t = t" .. next - assume "\ TRUE" thus "e = t" .. + assume "TRUE \ FALSE" + from this trueI have "FALSE" . + thus "e = t" .. qed lemma notTrue: "(\ TRUE) = FALSE" by (unfold not_def, rule condT) -theorem falseE [elim!]: - assumes f: "FALSE" - shows "A" -proof (rule notTrueE) - have "FALSE = (\ TRUE)" - by (rule sym[OF notTrue]) - hence r: "FALSE \ \ TRUE" - by (rule eq_reflection) - from f show "\ TRUE" - by (unfold r) -qed +lemma notTrueE: "\ TRUE \ A" + unfolding notTrue[THEN eq_reflection] by (rule falseE) lemma condF: "(IF FALSE THEN t ELSE e) = e" proof (rule condI) @@ -277,7 +269,7 @@ proof - by (unfold r, rule condT) qed -lemma condD1 [elim 2]: +lemma condD1 [Pure.elim 2]: assumes c: "IF A THEN P ELSE Q" (is ?if) and a: "A" shows "P" proof (rule eqTrueD) @@ -287,52 +279,71 @@ proof (rule eqTrueD) finally show "P = TRUE" by (rule sym) qed +theorem notI [Pure.intro!]: + assumes hyp: "A \ FALSE" + shows "\ A" +proof (unfold not_def, rule condI) + assume "A" thus "FALSE" by (rule hyp) +next + show "TRUE" .. +qed + +text \ + We do not have @{text "\ A \ A = FALSE"}: this + is true only for Booleans, not for arbitrary values. However, + we have the following theorem, which is just the ordinary + elimination rule for negation. +\ + +lemma notE: + assumes notA: "\ A" and a: "A" + shows "B" +proof - + from notA[unfolded not_def] a have "FALSE" by (rule condD1) + thus ?thesis .. +qed + lemma condElse: assumes na: "\ A" shows "(IF A THEN t ELSE e) = e" proof (rule condI) - assume "A" hence "A = TRUE" - by (rule eqTrueI) - hence r: "A \ TRUE" - by (rule eq_reflection) - from na have "\ TRUE" - by (unfold r) - thus "t = e" .. + assume "A" with na show "t = e" by (rule notE) next show "e = e" .. qed -lemma condD2 [elim 2]: +lemma condD2 [Pure.elim 2]: assumes c: "IF A THEN P ELSE Q" (is ?if) and a: "\ A" shows "Q" -proof (rule eqTrueD) - from c have "?if = TRUE" by (rule eqTrueI) - hence "TRUE = ?if" by (rule sym) - also from a have "?if = Q" by (rule condElse) - finally show "Q = TRUE" by (rule sym) +proof - + from a have "?if = Q" by (rule condElse) + hence "Q = ?if" by (rule sym) + also from c have "?if = TRUE" by (rule eqTrueI) + finally show ?thesis by (rule eqTrueD) qed -text {* - The following theorem shows that we have a classical logic. -*} +text \ + The following theorem confirms that we have a classical logic. +\ lemma cond_id: "(IF A THEN t ELSE t) = t" -proof (rule condI) - show "t=t" .. - show "t=t" .. -qed + by (rule condI) (rule refl)+ -theorem case_split: - assumes p: "P \ Q" - and np: "\ P \ Q" +theorem case_split [case_names True False]: + assumes p: "P \ Q" and np: "\ P \ Q" shows "Q" proof - - from p np have "IF P THEN Q ELSE Q" by (rule condI) + from p np have "IF P THEN Q ELSE Q" + proof (rule condI) + assume "P" thus "P" . + next + assume "P \ FALSE" thus "\P" .. + qed thus "Q" by (unfold eq_reflection[OF cond_id]) qed +text \ Use conditionals in hypotheses. \ theorem condE: - -- {* use conditionals in hypotheses *} assumes p: "P(IF A THEN t ELSE e)" and pos: "\ A; P(t) \ \ B" and neg: "\ \ A; P(e) \ \ B" @@ -341,87 +352,73 @@ proof (rule case_split[of "A"]) assume a: "A" hence r: "IF A THEN t ELSE e \ t" by (unfold eq_reflection[OF condThen]) - with p have "P(t)" - by (unfold r) - with a show "B" - by (rule pos) + with p have "P(t)" by (unfold r) + with a show "B" by (rule pos) next assume a: "\ A" hence r: "IF A THEN t ELSE e \ e" by (unfold eq_reflection[OF condElse]) - with p have "P(e)" - by (unfold r) - with a show "B" - by (rule neg) + with p have "P(e)" by (unfold r) + with a show "B" by (rule neg) qed -text {* +text \ Theorems @{text "condI"} and @{text "condE"} require higher-order unification and are therefore unsuitable for automatic tactics (in particular the \texttt{blast} method). We now derive some special cases that can be given to these methods. -*} +\ --- {* @{text "\A \ t; \A \ e\ \ IF A THEN t ELSE e"} *} -lemmas cond_boolI [intro!] = condI[where P="\ Q. Q"] +lemma cond_boolI [Pure.intro!]: + assumes a: "A \ P" and na: "\A \ Q" + shows "IF A THEN P ELSE Q" +proof (rule condI) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "Q" by (rule na) +qed (rule a) -lemma cond_eqLI [intro!]: +lemma cond_eqLI [Pure.intro!]: assumes 1: "A \ t = v" and 2: "\A \ u = v" shows "(IF A THEN t ELSE u) = v" proof (rule condI) - show "A \ t=v" by (rule 1) + show "A \ t = v" by (rule 1) next - show "\A \ u=v" by (rule 2) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "u = v" by (rule 2) qed -lemma cond_eqRI [intro!]: +lemma cond_eqRI [Pure.intro!]: assumes 1: "A \ v = t" and 2: "\A \ v = u" shows "v = (IF A THEN t ELSE u)" proof (rule condI) show "A \ v = t" by (rule 1) next - show "\A \ v = u" by (rule 2) + assume "A \ FALSE" + hence "\A" by (rule notI) + thus "v = u" by (rule 2) qed --- {* @{text "\IF A THEN t ELSE e; \A;t\ \ B; \\A;e\ \ B\ \ B"} *} +text \ @{text "\IF A THEN P ELSE Q; \A;P\ \ B; \\A;Q\ \ B\ \ B"} \ lemmas cond_boolE [elim!] = condE[where P="\ Q. Q"] -lemma cond_eqLE [elim!]: - assumes maj: "(IF A THEN t ELSE e) = u" - and 1: "\ A; t=u \ \ B" and 2: "\ \A; e=u \ \ B" +lemma cond_eqLE [Pure.elim!]: + assumes "(IF A THEN t ELSE e) = u" + and "\ A; t=u \ \ B" and "\ \A; e=u \ \ B" shows "B" -using maj -proof (rule condE) - show "\ A; t=u \ \ B" by (rule 1) -next - show "\ \A; e=u \ \ B" by (rule 2) -qed + using assms by (rule condE) -lemma cond_eqRE [elim!]: - assumes maj: "u = (IF A THEN t ELSE e)" - and 1: "\ A; u=t \ \ B" and 2: "\ \A; u=e \ \ B" +lemma cond_eqRE [Pure.elim!]: + assumes "u = (IF A THEN t ELSE e)" + and "\ A; u=t \ \ B" and "\ \A; u=e \ \ B" shows "B" -using maj -proof (rule condE) - show "\ A; u=t \ \ B" by (rule 1) -next - show "\ \A; u=e \ \ B" by (rule 2) -qed + using assms by (rule condE) -text {* +text \ Derive standard propositional proof rules, based on the operator definitions in terms of @{text "IF _ THEN _ ELSE _"}. -*} - -theorem notI [intro!]: - assumes hyp: "A \ FALSE" - shows "\ A" -proof (unfold not_def, rule condI) - assume "A" thus "FALSE" - by (rule hyp) -next - show "TRUE" .. -qed +\ lemma false_neq_true: "FALSE \ TRUE" proof @@ -432,10 +429,7 @@ qed lemma false_eq_trueE: assumes ft: "FALSE = TRUE" shows "B" -proof (rule falseE) - from ft show "FALSE" - by (rule eqTrueD) -qed + using false_neq_true ft by (rule notE) lemmas true_eq_falseE = sym[THEN false_eq_trueE] @@ -445,42 +439,22 @@ by iprover lemma A_then_notA_false: assumes a: "A" shows "(\ A) = FALSE" -using a -by (unfold not_def, rule condThen) + using a by (unfold not_def, rule condThen) -text {* +text \ The following is an alternative introduction rule for negation that is useful when we know that @{text A} is Boolean. -*} +\ lemma eq_false_not: assumes a: "A = FALSE" shows "\ A" -proof (rule eqTrueD) - show "(\ A) = TRUE" by (unfold eq_reflection[OF a], rule notFalse) -qed - -text {* - Note that we do not have @{text "\ A \ A = FALSE"}: this - is true only for Booleans, not for arbitrary values. However, - we have the following theorem, which is just the ordinary - elimination rule for negation. -*} - -theorem notE: - assumes notA: "\ A" and a: "A" - shows "B" -proof (rule false_eq_trueE) - from a have "(\ A) = FALSE" - by (rule A_then_notA_false) - hence "FALSE = (\ A)" - by (rule sym) - also from notA have "(\ A) = TRUE" - by (rule eqTrueI) - finally show "FALSE = TRUE" . +proof + assume "A" + thus "FALSE" by (unfold a[THEN eq_reflection]) qed -theorem notE' [elim 2]: +theorem notE' [Pure.elim 2]: assumes notA: "\ A" and r: "\ A \ A" shows "B" using notA @@ -505,18 +479,17 @@ proof with hyp show "FALSE" .. qed -text {* +text \ Some derived proof rules of classical logic. -*} +\ theorem contradiction: assumes hyp: "\ A \ FALSE" shows "A" proof (rule case_split[of "A"]) - assume "\ A" hence "FALSE" - by (rule hyp) + assume "\ A" hence "FALSE" by (rule hyp) thus "A" .. -qed -- {* the other case is trivial *} +qed text \ the other case is trivial \ theorem classical: assumes c: "\ A \ A" @@ -535,17 +508,13 @@ proof (rule contradiction) with a show "FALSE" .. qed -theorem notnotD [dest]: - assumes nn: "\\ A" shows "A" -proof (rule contradiction) - assume "\ A" - with nn show "FALSE" .. -qed +theorem notnotD [Pure.dest]: "\\A \ A" + by (rule swap) -text {* +text \ Note again: @{text A} and @{text "\\ A"} are inter-derivable - (and hence equivalent), but not equal! -*} + (and hence provably equivalent), but not equal! +\ lemma contrapos: assumes b: "\ B" and ab: "A \ B" @@ -569,14 +538,14 @@ proof - thus "A" .. qed -theorem conjI [intro!]: +theorem conjI [Pure.intro!]: assumes a: "A" and b: "B" shows "A \ B" -proof (rule eqTrueD) +proof - from a have "(A \ B) = (IF B THEN TRUE ELSE FALSE)" by (unfold conj_def, rule condThen) also from b have "\ = TRUE" by (rule condThen) - finally show "(A \ B) = TRUE" . + finally show ?thesis by (rule eqTrueD) qed theorem conjD1: @@ -584,8 +553,7 @@ theorem conjD1: shows "A" proof (rule contradiction) assume "\ A" - with ab show "FALSE" - by (unfold conj_def, elim condD2) + with ab show "FALSE" by (unfold conj_def, elim condD2) qed theorem conjD2: @@ -594,12 +562,11 @@ theorem conjD2: proof (rule contradiction) assume b: "\ B" from ab have "A" by (rule conjD1) - with ab have "IF B THEN TRUE ELSE FALSE" - by (unfold conj_def, elim condD1) + with ab have "IF B THEN TRUE ELSE FALSE" by (unfold conj_def, elim condD1) with b show "FALSE" by (elim condD2) qed -theorem conjE [elim!]: +theorem conjE [Pure.elim!]: assumes ab: "A \ B" and c: "A \ B \ C" shows "C" proof (rule c) @@ -607,9 +574,9 @@ proof (rule c) from ab show "B" by (rule conjD2) qed -text {* Disjunction *} +text \ Disjunction \ -theorem disjI1 [elim]: +theorem disjI1 [Pure.elim]: assumes a: "A" shows "A \ B" proof (unfold disj_def, rule) @@ -619,39 +586,28 @@ next from this a show "IF B THEN TRUE ELSE FALSE" .. qed -theorem disjI2 [elim]: +theorem disjI2 [Pure.elim]: assumes b: "B" shows "A \ B" proof (unfold disj_def, rule) show "TRUE" .. next - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE" .. - qed + from b show "IF B THEN TRUE ELSE FALSE" by iprover qed -theorem disjI [intro!]: -- {* classical introduction rule *} +text \ Classical introduction rule for disjunction. \ +theorem disjI [Pure.intro!]: assumes ab: "\ A \ B" shows "A \ B" proof (unfold disj_def, rule) show "TRUE" .. next assume "\ A" - hence b: "B" by (rule ab) - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE".. - qed + hence "B" by (rule ab) + thus "IF B THEN TRUE ELSE FALSE" by iprover qed -theorem disjE [elim!]: +theorem disjE [Pure.elim!]: assumes ab: "A \ B" and ac: "A \ C" and bc: "B \ C" shows "C" proof (rule case_split[where P=A]) @@ -669,30 +625,22 @@ next qed theorem excluded_middle: "A \ \ A" -proof - assume "\ A" thus "\ A" . -qed + by (rule disjI) -text {* Implication *} +text \ Implication \ -theorem impI [intro!]: +theorem impI [Pure.intro!]: assumes ab: "A \ B" shows "A \ B" proof (unfold imp_def, rule) assume "A" - hence b: "B" by (rule ab) - show "IF B THEN TRUE ELSE FALSE" - proof - show "TRUE" .. - next - assume "\ B" - from this b show "FALSE" .. - qed + hence "B" by (rule ab) + thus "IF B THEN TRUE ELSE FALSE" by iprover next show "TRUE" .. qed -theorem mp (*[elim 2]*): +theorem mp (*[Pure.elim 2]*): assumes ab: "A \ B" and a: "A" shows "B" proof (rule contradiction) @@ -702,12 +650,12 @@ proof (rule contradiction) from this notb show "FALSE" by (rule condD2) qed -theorem rev_mp (*[elim 2]*): +theorem rev_mp (*[Pure.elim 2]*): assumes a: "A" and ab: "A \ B" shows "B" using ab a by (rule mp) -theorem impE: +theorem impE [Pure.elim?]: assumes ab: "A \ B" and a: "A" and bc: "B \ C" shows "C" proof - @@ -715,33 +663,31 @@ proof - thus "C" by (rule bc) qed -theorem impCE [elim]: +theorem impCE [Pure.elim]: assumes ab: "A \ B" and b: "B \ P" and a: "\ A \ P" shows "P" -proof (rule classical) - assume contra: "\ P" - have "A" - proof (rule contradiction) - assume "\ A" hence "P" by (rule a) - with contra show "FALSE" .. - qed +proof (rule case_split[of "A"]) + assume "A" with ab have "B" by (rule mp) thus "P" by (rule b) +next + assume "\A" + thus "P" by (rule a) qed theorem impCE': (* used for classical prover *) - assumes ab: "A \ B" and a: "\C \ A" and b: "B \ C" - shows "C" + assumes ab: "A \ B" and a: "\P \ A" and b: "B \ P" + shows "P" proof (rule classical) - assume "\C" + assume "\P" hence "A" by (rule a) with ab have "B" by (rule mp) - thus "C" by (rule b) + thus "P" by (rule b) qed -text {* Equivalence *} +text \ Equivalence \ -theorem iffI [intro!]: +theorem iffI [Pure.intro!]: assumes ab: "A \ B" and ba: "B \ A" shows "A \ B" proof (unfold iff_def, rule) @@ -763,7 +709,7 @@ proof - thus ?thesis by (rule meta_eq_to_iff) qed -theorem iffD1 [elim 2]: +theorem iffD1 [Pure.elim?]: assumes ab: "A \ B" and a: "A" shows "B" using ab @@ -772,7 +718,7 @@ proof (unfold iff_def, elim conjE) from this a show "B" by (rule mp) qed -theorem iffD2 [elim 2]: +theorem iffD2 [Pure.elim?]: assumes ab: "A \ B" and b: "B" shows "A" using ab @@ -789,7 +735,7 @@ proof (rule r) from ab show "B \ A" by (unfold iff_def, elim conjE) qed -theorem iffCE [elim!]: +theorem iffCE [Pure.elim!]: assumes ab: "A \ B" and pos: "\A; B\ \ C" and neg: "\\ A; \ B\ \ C" shows "C" @@ -822,10 +768,10 @@ next qed -subsection {* Predicate Logic *} +subsection \ Predicate Logic \ -text {* - We take Hilbert's $\varepsilon$ as the basic binder and define the +text \ + We take Hilbert's @{text \} as the basic binder and define the other quantifiers as derived connectives. Again, we make sure that quantified formulas evaluate to @{term TRUE} or @{term FALSE}. @@ -836,43 +782,18 @@ text {* and automatic provers such as \texttt{blast} rely on reflection to the object level for reasoning about meta-connectives, which would not be possible with purely first-order quantification. -*} +\ consts Choice :: "('a \ c) \ 'a" - Ex :: "('a \ c) \ c" - All :: "('a \ c) \ c" - -text {* Concrete syntax: several variables as in @{text "\x,y : P(x,y)"}. *} - -nonterminal cidts (* comma-separated idt's *) - -syntax - "" :: "idt \ cidts" ("_" [100] 100) - "@cidts" :: "[idt, cidts] \ cidts" ("_,/ _" [100,100] 100) syntax (* BEWARE: choosing several values in sequence is not equivalent to choosing a tuple of values simultaneously. To avoid confusion, we do not allow several variables in a CHOOSE. *) "@Choice" :: "[idt, c] \ c" ("(3CHOOSE _ :/ _)" [100,10] 10) - "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100,10] 10) - "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100,10] 10) -syntax (xsymbols) - "@Ex" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100, 10] 10) - "@All" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100, 10] 10) translations "CHOOSE x: P" \ "CONST Choice(\x. P)" - (* separate parse and print translations for nested quantifiers because - they operate in opposite directions *) - "\x,xs : P" \ "CONST Ex(\x. \xs : P)" -(** FIXME: print translation doesn't work correctly: - will print \x : \y: P as \x, y : P -- why ??? **) -(* "\x,xs : P" \ "\x : Ex(\xs. P)" *) - "\x : P" \ "CONST Ex(\x. P)" - "\x,xs : P" \ "CONST All(\x. \xs : P)" -(* "\x,xs : P" \ "\x : All(\xs. P)" *) - "\x : P" \ "CONST All(\x. P)" (* Declare the two following axioms separately, otherwise the type checker infers the same type for P, and the type of P in @@ -880,20 +801,46 @@ translations in the first one. *) axiomatization where - chooseI: "P(t) \ P(CHOOSE x: P(x))" + chooseI: "P(t) \ P(CHOOSE x: P(x))" for t::'a axiomatization where choose_det : "(\x. P(x) \ Q(x)) \ (CHOOSE x: P(x)) = (CHOOSE x: Q(x))" -defs - Ex_def: "Ex(P) \ P(CHOOSE x : P(x) = TRUE) = TRUE" - All_def: "All(P) \ \(\x : \ P(x))" +definition Ex :: "('a \ c) \ c" + where "Ex(P) \ P(CHOOSE x : P(x) = TRUE) = TRUE" + +definition All :: "('a \ c) \ c" + where "All(P) \ \ Ex(\x. \ P(x))" + +text \ Concrete syntax: several variables as in @{text "\x,y : P(x,y)"}. \ -text {* +nonterminal cidts (* comma-separated idt's *) + +syntax + "" :: "idt \ cidts" ("_" [100] 100) + "@cidts" :: "[idt, cidts] \ cidts" ("_,/ _" [100,100] 100) + +syntax + "@Ex" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100,10] 10) + "@All" :: "[cidts, c] \ c" ("(3\_ :/ _)" [100,10] 10) +syntax (ASCII) + "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) + "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) +translations + (* separate parse and print translations for nested quantifiers because + they operate in opposite directions *) + "\x,xs : P" \ "CONST Ex(\x. \xs : P)" + "\x,xs : P" \ "\x : CONST Ex(\xs. P)" + "\x : P" \ "CONST Ex(\x. P)" + "\x,xs : P" \ "CONST All(\x. \xs : P)" + "\x,xs : P" \ "\x : CONST All(\xs. P)" + "\x : P" \ "CONST All(\x. P)" + +text \ We introduce two constants @{text arbitrary} and @{text default} - that correspond to unconstrained and overconstrained choice, + that correspond to unconstrained and over-constrained choice, respectively. -*} +\ definition arbitrary::c where "arbitrary \ CHOOSE x : TRUE" @@ -901,7 +848,7 @@ definition arbitrary::c where definition default::c where "default \ CHOOSE x : FALSE" -theorem exI [intro]: +theorem exI [Pure.intro]: assumes hyp: "P(t)" shows "\x : P(x)" proof - @@ -909,7 +856,7 @@ proof - thus ?thesis by (unfold Ex_def, rule chooseI) qed -theorem exE [elim!]: +theorem exE [Pure.elim!]: assumes hyp: "\x : P(x)" and r: "\x. P(x) \ Q" shows "Q" proof - @@ -918,7 +865,7 @@ proof - thus "Q" by (rule r) qed -theorem allI [intro!]: +theorem allI [Pure.intro!]: assumes hyp: "\x. P(x)" shows "\x : P(x)" proof (unfold All_def, rule) @@ -936,7 +883,7 @@ proof (rule contradiction) with hyp show "FALSE" by (unfold All_def, elim notE) qed -theorem allE [elim]: +theorem allE [Pure.elim 2]: assumes hyp: "\x : P(x)" and r: "P(x) \ Q" shows "Q" proof (rule r) @@ -967,18 +914,18 @@ proof (rule q) from p show "P(CHOOSE x : P(x))" by (rule chooseI_ex) qed -lemma choose_equality [intro]: +lemma choose_equality [Pure.intro]: assumes "P(t)" and "\x. P(x) \ x=a" shows "(CHOOSE x : P(x)) = a" -using assms by (rule chooseI2[where Q="\x. x=a"]) + using assms by (rule chooseI2) -lemmas choose_equality' = sym[OF choose_equality, standard, intro] +lemmas choose_equality' = sym[OF choose_equality, Pure.intro] -text {* +text \ Skolemization rule: note that the existential quantifier in the conclusion introduces an operator (of type @{text "c \ c"}), not a value; second-order quantification is necessary here. -*} +\ lemma skolemI: assumes h: "\x : \y: P(x,y)" shows "\f : \x : P(x, f(x))" @@ -1000,16 +947,16 @@ next qed -subsection {* Setting up the automatic proof methods *} +subsection \ Setting up the automatic proof methods \ -subsubsection {* Reflection of meta-level to object-level *} +subsubsection \ Reflection of meta-level to object-level \ -text {* +text \ Our next goal is to getting Isabelle's automated tactics to work for \tlaplus{}. We follow the setup chosen for Isabelle/HOL as far as it applies to \tlaplus{}. The following lemmas, when used as rewrite rules, replace meta- by object-level connectives. -*} +\ lemma atomize_all [atomize]: "(\x. P(x)) \ Trueprop (\x : P(x))" proof @@ -1058,7 +1005,6 @@ next qed qed - lemmas [symmetric,rulify] = atomize_all atomize_imp and [symmetric,defn] = atomize_all atomize_imp atomize_eq @@ -1070,19 +1016,8 @@ lemma tmp2: "(\P x. P(x)) \ (\a::c. Q(a))" . lemma tmp3: "\a. (\x. P(x)) \ P(a)" . -ML {* - let - val thm = @{thm tmp2} - val rew = Object_Logic.atomize (cprop_of thm) - in - Raw_Simplifier.rewrite_rule [rew] thm - end -*} - ** end test cases **) -setup AtomizeElim.setup - lemma atomize_exL[atomize_elim]: "(\x. P(x) \ Q) \ ((\x : P(x)) \ Q)" by rule iprover+ @@ -1095,49 +1030,28 @@ by rule iprover+ lemma atomize_elimL[atomize_elim]: "(\B. (A \ B) \ B) \ Trueprop(A)" .. -subsubsection {* Setting up the classical reasoner *} +subsubsection \ Setting up the classical reasoner \ -text {* +text \ We now instantiate Isabelle's classical reasoner for \tlaplus{}. This includes methods such as \texttt{fast} and \texttt{blast}. -*} +\ lemma thin_refl: "\x=x; PROP W\ \ PROP W" . -ML {* +ML \ (* functions to take apart judgments and formulas, see Isabelle reference manual, section 9.3 *) - fun dest_Trueprop (Const(@{const_name Trueprop}, _) $ P) = P + fun dest_Trueprop (Const(@{const_name \Trueprop\}, _) $ P) = P | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); - fun dest_eq (Const(@{const_name eq}, _) $ t $ u) = (t,u) + fun dest_eq (Const(@{const_name \eq\}, _) $ t $ u) = (t,u) | dest_eq t = raise TERM ("dest_eq", [t]); - fun dest_imp (Const(@{const_name imp}, _) $ A $ B) = (A, B) + fun dest_imp (Const(@{const_name \imp\}, _) $ A $ B) = (A, B) | dest_imp t = raise TERM ("dest_imp", [t]); -(** - structure Hypsubst_Data = - struct - structure Simplifier = Simplifier - val dest_Trueprop = dest_Trueprop - val dest_eq = dest_eq - val dest_imp = dest_imp - val eq_reflection = @{thm eq_reflection} - val rev_eq_reflection = @{thm meta_to_obj_eq} - val imp_intr = @{thm impI} - val rev_mp = @{thm rev_mp} - val subst = @{thm subst} - val sym = @{thm sym} - val thin_refl = @{thm thin_refl} - val prop_subst = @{lemma "PROP P(t) ==> PROP prop (x = t ==> PROP P(x))" - by (unfold prop_def) (drule eq_reflection, unfold)} - end; - structure Hypsubst = HypsubstFun(Hypsubst_Data); - open Hypsubst; -**) - structure Hypsubst = Hypsubst( val dest_Trueprop = dest_Trueprop val dest_eq = dest_eq @@ -1152,19 +1066,6 @@ ML {* ); open Hypsubst; -(** - structure Classical_Data = - struct - val imp_elim = @{thm impCE'} - val not_elim = @{thm notE} - val swap = @{thm swap} - val classical = @{thm classical} - val sizef = Drule.size_of_thm - val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] - end; - structure Classical = ClassicalFun(Classical_Data); -**) - structure Classical = Classical( val imp_elim = @{thm impCE'} val not_elim = @{thm notE} @@ -1176,10 +1077,12 @@ ML {* structure BasicClassical : BASIC_CLASSICAL = Classical; open BasicClassical; -*} +\ +(* setup hypsubst_setup setup Classical.setup +*) declare refl [intro!] and trueI [intro!] @@ -1208,38 +1111,20 @@ declare refl [intro!] and choose_equality [intro] and sym[OF choose_equality, intro] -ML {* -(** - structure Blast = Blast - ( struct - val thy = @{theory} - type claset = Classical.claset - val equality_name = @{const_name PredicateLogic.eq} - val not_name = @{const_name PredicateLogic.Not} - val notE = @{thm notE} - val ccontr = @{thm contradiction} - val contr_tac = Classical.contr_tac - val dup_intr = Classical.dup_intr - val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac - val rep_cs = Classical.rep_cs - val cla_modifiers = Classical.cla_modifiers; - val cla_meth' = Classical.cla_meth' - end ); -**) - +ML \ structure Blast = Blast ( structure Classical = Classical - val Trueprop_const = dest_Const @{const Trueprop} + val Trueprop_const = dest_Const @{const \Trueprop\} val equality_name = @{const_name PredicateLogic.eq} - val not_name = @{const_name PredicateLogic.Not} + val not_name = @{const_name PredicateLogic.not} val notE = @{thm notE} val ccontr = @{thm contradiction} val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac ); -*} +\ -setup Blast.setup +(* setup Blast.setup *) (*** TEST DATA *** @@ -1248,7 +1133,7 @@ apply hypsubst apply (rule refl) done -lemma "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)" +lemma "\x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)" by hypsubst lemma "x = x" @@ -1303,35 +1188,44 @@ by blast *** END TEST DATA ***) -subsubsection {* Setting up the simplifier *} +subsubsection \ Setting up the simplifier \ -text {* +text \ We instantiate the simplifier, Isabelle's generic rewriting machinery. Equational laws for predicate logic will be proven below; they automate much of the purely logical reasoning. -*} +\ + +text \first use of axiom @{text eqBoolean}\ +lemma neq_conv_eq_False: + assumes "x \ y" shows "(x = y) = FALSE" + by (rule eqBoolean) (rule notE[OF assms]) -lemma if_bool_eq_conj: +lemma not_refl_False: + assumes "(x \ x)" shows "A" + using assms by (rule notE) (rule refl) + +lemma if_bool_equiv_conj: "(IF A THEN B ELSE C) \ ((A \ B) \ (\A \ C))" by fast -text {* +text \ A copy of Isabelle's meta-level implication is introduced, which is used internally by the simplifier for fine-tuning congruence rules by simplifying their premises. -*} +\ definition simp_implies :: "[prop, prop] \ prop" (infixr "=simp=>" 1) where - "simp_implies \ op \" + "simp_implies \ (\)" lemma simp_impliesI: - assumes PQ: "(PROP P ==> PROP Q)" + assumes PQ: "(PROP P \ PROP Q)" shows "PROP P =simp=> PROP Q" unfolding simp_implies_def by (rule PQ) lemma simp_impliesE: assumes PQ: "PROP P =simp=> PROP Q" - and P: "PROP P" and QR: "PROP Q ==> PROP R" + and P: "PROP P" and QR: "PROP Q \ PROP R" shows "PROP R" proof - from P have "PROP Q" by (rule PQ [unfolded simp_implies_def]) @@ -1339,32 +1233,29 @@ proof - qed lemma simp_implies_cong: - assumes PP' :"PROP P == PROP P'" - and P'QQ': "PROP P' ==> (PROP Q == PROP Q')" - shows "(PROP P =simp=> PROP Q) == (PROP P' =simp=> PROP Q')" + assumes PP' :"PROP P \ PROP P'" + and P'QQ': "PROP P' \ (PROP Q \ PROP Q')" + shows "(PROP P =simp=> PROP Q) \ (PROP P' =simp=> PROP Q')" unfolding simp_implies_def proof (rule equal_intr_rule) - assume PQ: "PROP P ==> PROP Q" and P': "PROP P'" + assume PQ: "PROP P \ PROP Q" and P': "PROP P'" from PP' [symmetric] and P' have "PROP P" by (rule equal_elim_rule1) hence "PROP Q" by (rule PQ) with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) next - assume P'Q': "PROP P' ==> PROP Q'" and P: "PROP P" + assume P'Q': "PROP P' \ PROP Q'" and P: "PROP P" from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) hence "PROP Q'" by (rule P'Q') with P'QQ' [OF P', symmetric] show "PROP Q" by (rule equal_elim_rule1) qed -use "simplifier_setup.ML" - -setup {* - Simplifier.map_simpset_global (K Simpdata.PL_basic_ss) - #> Simplifier.method_setup Splitter.split_modifiers - #> Splitter.setup - #> Simpdata.clasimp_setup - #> EqSubst.setup -*} +ML_file "simplifier_setup.ML" +ML \open Simpdata\ +setup \ + map_theory_simpset (put_simpset PL_basic_ss) #> + Simplifier.method_setup Splitter.split_modifiers +\ (*** TEST DATA -- note: only basic simplification, no theorems set up *** @@ -1393,10 +1284,13 @@ lemma "x=y \ y=z \ x=z \ f(y) = f(x)" by auto lemma "x=y \ y=z \ x=z" -by force + by force lemma "(\x : P(x)) \ P(a)" -by auto + by auto + +lemma "x \ y \ y = x \ A" + by simp lemma assumes ind: "P(z) \ (\n. P(n) \ P(s(n))) \ (\n. P(n))" @@ -1410,13 +1304,11 @@ lemma shows "Q(n)" using assms by auto -(* without converting all connectives to meta-level (via attribute - [rule_format]), the following loops *) lemma assumes "\Q. (\n: (\i: L(i,n) \ Q(i)) \ Q(n)) \ (\n: Q(n))" and "\n: (\i: L(i,n) \ Q(i)) \ Q(n)" shows "Q(n)" -using assms[rule_format] by auto +using assms(*[rule_format]*) by auto *** END TEST DATA ***) @@ -1438,38 +1330,49 @@ next thus "TRUE = A" by (rule sym) qed +lemma implies_true_equals: "(PROP P \ TRUE) \ Trueprop(TRUE)" + by standard (intro trueI) + +lemma false_implies_equals: "(FALSE \ P) \ Trueprop(TRUE)" + by standard simp_all + +lemma cond_cancel: "(IF x=y THEN y ELSE x) = x" + by fast + lemmas [simp] = triv_forall_equality - TrueAssumption trueprop_eq_true trueprop_true_eq - refl[THEN eqTrueI] -- {* @{text "(x = x) \ TRUE"} *} + TrueAssumption implies_true_equals false_implies_equals + refl[THEN eqTrueI] (* (x = x) = TRUE *) condT (*condThen*) notTrue condF (*condElse*) notFalse - cond_id - false_neq_true[THEN eqBoolean] - not_sym[OF false_neq_true, THEN eqBoolean] + cond_id cond_cancel + false_neq_true[THEN neq_conv_eq_False] (* (FALSE = TRUE) = FALSE *) + not_sym[OF false_neq_true, THEN neq_conv_eq_False] (* (TRUE = FALSE) = FALSE *) iff_refl lemmas [cong] = simp_implies_cong -subsubsection {* Reasoning by cases *} +subsubsection \ Reasoning by cases \ -text {* +text \ The next bit of code sets up reasoning by cases as a proper Isar method, so we can write ``proof cases'' etc. Following the development of FOL, we introduce a set of ``shadow connectives'' that will only be used for this purpose. -*} +\ -theorems cases = case_split [case_names True False] +(* lemmas cases = case_split [case_names True False] *) +context +begin -definition cases_equal where "cases_equal \ eq" -definition cases_implies where "cases_implies \ imp" -definition cases_conj where "cases_conj \ conj" -definition cases_forall where "cases_forall(P) \ \x: P(x)" -definition cases_true where "cases_true \ TRUE" -definition cases_false where "cases_false \ FALSE" +qualified definition cases_equal where "cases_equal \ eq" +qualified definition cases_implies where "cases_implies \ imp" +qualified definition cases_conj where "cases_conj \ conj" +qualified definition cases_forall where "cases_forall(P) \ \x: P(x)" +qualified definition cases_true where "cases_true \ TRUE" +qualified definition cases_false where "cases_false \ FALSE" lemma cases_equal_eq: "(x \ y) \ Trueprop(cases_equal(x, y))" unfolding atomize_eq cases_equal_def . @@ -1488,8 +1391,8 @@ unfolding cases_true_def .. lemmas cases_atomize' = cases_implies_eq cases_conj_eq cases_forall_eq lemmas cases_atomize = cases_atomize' cases_equal_eq -lemmas cases_rulify' [symmetric, standard] = cases_atomize' -lemmas cases_rulify [symmetric, standard] = cases_atomize +lemmas cases_rulify' [symmetric] = cases_atomize' +lemmas cases_rulify [symmetric] = cases_atomize lemmas cases_rulify_fallback = cases_equal_def cases_implies_def cases_conj_def cases_forall_def cases_true_def cases_false_def @@ -1515,48 +1418,50 @@ qed lemmas cases_conj = cases_forall_conj cases_implies_conj cases_conj_curry -ML {* +ML \ structure Induct = Induct ( - val cases_default = @{thm cases} + val cases_default = @{thm case_split} val atomize = @{thms cases_atomize} val rulify = @{thms cases_rulify'} val rulify_fallback = @{thms cases_rulify_fallback} val equal_def = @{thm cases_equal_def} fun dest_def (Const (@{const_name cases_equal}, _) $ t $ u) = SOME (t, u) | dest_def _ = NONE - val trivial_tac = match_tac @{thms cases_trueI} + fun trivial_tac ctxt = match_tac ctxt @{thms cases_trueI} ) -*} +\ -setup {* - Induct.setup #> - Context.theory_map (Induct.map_simpset (fn ss => ss - setmksimps (fn ss => Simpdata.mksimps Simpdata.mksimps_pairs ss #> - map (Simplifier.rewrite_rule (map Thm.symmetric @{thms cases_rulify_fallback}))) +declaration \ + fn _ => Induct.map_simpset (fn ss => ss addsimprocs - [Simplifier.simproc_global @{theory} "swap_cases_false" - ["cases_false ==> PROP P ==> PROP Q"] - (fn _ => fn _ => - (fn _ $ (P as _ $ @{const cases_false}) $ (_ $ Q $ _) => - if P <> Q then SOME Drule.swap_prems_eq else NONE - | _ => NONE)), - Simplifier.simproc_global @{theory} "cases_equal_conj_curry" - ["cases_conj(P, Q) ==> PROP R"] - (fn _ => fn _ => - (fn _ $ (_ $ P) $ _ => - let - fun is_conj (@{const cases_conj} $ P $ Q) = - is_conj P andalso is_conj Q - | is_conj (Const (@{const_name cases_equal}, _) $ _ $ _) = true - | is_conj @{const cases_true} = true - | is_conj @{const cases_false} = true - | is_conj _ = false - in if is_conj P then SOME @{thm cases_conj_curry} else NONE end - | _ => NONE))])) -*} - -text {* Pre-simplification of induction and cases rules *} + [Simplifier.make_simproc @{context} "swap_cases_false" + {lhss = [@{term "cases_false \ PROP P \ PROP Q"}], + proc = fn _ => fn _ => fn ct => + (case Thm.term_of ct of + _ $ (P as _ $ @{const cases_false}) $ (_ $ Q $ _) => + if P <> Q then SOME Drule.swap_prems_eq else NONE + | _ => NONE)}, + Simplifier.make_simproc @{context} "cases_equal_conj_curry" + {lhss = [@{term "cases_conj(P, Q) \ PROP R"}], + proc = fn _ => fn _ => fn ct => + (case Thm.term_of ct of + _ $ (_ $ P) $ _ => + let + fun is_conj (@{const cases_conj} $ P $ Q) = + is_conj P andalso is_conj Q + | is_conj (Const (@{const_name cases_equal}, _) $ _ $ _) = true + | is_conj @{const cases_true} = true + | is_conj @{const cases_false} = true + | is_conj _ = false + in if is_conj P then SOME @{thm cases_conj_curry} else NONE end + | _ => NONE)}] + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> + map (rewrite_rule ctxt (map Thm.symmetric @{thms cases_rulify_fallback})))) +\ + +text \ Pre-simplification of induction and cases rules \ lemma [induct_simp]: "(\x. cases_equal(x, t) \ PROP P(x)) \ PROP P(t)" unfolding cases_equal_def @@ -1593,7 +1498,7 @@ qed lemma [induct_simp]: "(PROP P \ cases_true) \ Trueprop(cases_true)" unfolding cases_true_def by (iprover intro: equal_intr_rule) -lemma [induct_simp]: "(\x. cases_true) \ Trueprop(cases_true)" +lemma [induct_simp]: "(\x::'a::{}. cases_true) \ Trueprop(cases_true)" unfolding cases_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "Trueprop(cases_implies(cases_true, P)) \ Trueprop(P)" @@ -1602,7 +1507,7 @@ lemma [induct_simp]: "Trueprop(cases_implies(cases_true, P)) \ Trueprop(P lemma [induct_simp]: "(x = x) = TRUE" by simp -hide_const cases_forall cases_implies cases_equal cases_conj cases_true cases_false +end (*** BASIC TEST DATA *** @@ -1616,18 +1521,17 @@ qed *** END TEST DATA ***) +subsection \ Propositional simplification \ -subsection {* Propositional simplification *} +subsubsection \ Conversion to Boolean values \ -subsubsection {* Conversion to Boolean values *} - -text {* +text \ Because \tlaplus{} is untyped, equivalence is different from equality, and one has to be careful about stating the laws of propositional logic. For example, although the equivalence @{text "(TRUE \ A) \ A"} is valid, we cannot state the law @{text "(TRUE \ A) = A"} because we have no way of knowing the value of, e.g., @{text "TRUE \ 3"}. - These equalities are valid only if + Such equations are valid only if the connectives are applied to Boolean operands. For automatic reasoning, we are interested in equations that can be used by Isabelle's simplifier. We therefore introduce an auxiliary @@ -1637,7 +1541,7 @@ text {* We will prove below that propositional formulas return a Boolean value when applied to arbitrary arguments. -*} +\ definition boolify :: "c \ c" where "boolify(x) \ IF x THEN TRUE ELSE FALSE" @@ -1645,10 +1549,10 @@ definition boolify :: "c \ c" where definition isBool :: "c \ c" where "isBool(x) \ boolify(x) = x" -text {* +text \ The formulas @{text "P"} and @{text "boolify(P)"} are inter-derivable (but need of course not be equal, unless @{text P} is a Boolean). -*} +\ lemma boolifyI [intro!]: "P \ boolify(P)" unfolding boolify_def by fast @@ -1665,11 +1569,11 @@ lemma boolify_cases: (** too general to be used as a default rule **) shows "P(boolify(x))" unfolding boolify_def using assms by (fast intro: condI) -text {* +text \ @{text boolify} can be defined as @{text "x = TRUE"}. For automatic reasoning, we rewrite the latter to the former, and derive calculational rules for @{text boolify}. -*} +\ lemma [simp]: "(x = TRUE) = boolify(x)" proof (cases "x") @@ -1704,14 +1608,14 @@ by (auto simp add: boolify_def) lemma falseIsBool [intro!,simp]: "isBool(FALSE)" by (unfold isBool_def, rule boolifyFalse) -text {* +text \ The following lemma is used to turn hypotheses @{text "\A"} into - rewrite rules @{text "A = FALSE"}. -*} + rewrite rules @{text "A = FALSE"} when we know that @{text A} is Boolean. +\ lemma boolifyFalseI [intro]: "\ A \ boolify(A) = FALSE" by (auto simp add: boolify_def) -text {* idempotence of @{text boolify} *} +text \ idempotence of @{text boolify} \ lemma boolifyBoolify [simp]: "boolify(boolify(x)) = boolify(x)" by (auto simp add: boolify_def) @@ -1734,18 +1638,16 @@ proof - show ?thesis by (unfold bx) qed -lemmas isBoolE [elim!] = isBoolTrueFalse[THEN disjE, standard] +lemmas isBoolE [elim!] = isBoolTrueFalse[THEN disjE] -lemma boolifyEq [simp]: "boolify(t=u) = (t=u)" (* first use of axiom eqBoolean *) +lemma boolifyEq [simp]: "boolify(t=u) = (t=u)" proof (cases "t=u") case True - hence "(t=u) = TRUE" by (rule eqTrueI) - hence tu: "(t=u) \ TRUE" by (rule eq_reflection) + hence tu: "(t=u) \ TRUE" by (rule eqTrueI[THEN eq_reflection]) show ?thesis by (unfold tu, rule boolifyTrue) next case False - hence "(t=u) = FALSE" by (rule eqBoolean) - hence tu: "(t=u) \ FALSE" by (rule eq_reflection) + hence tu: "(t=u) \ FALSE" by (rule neq_conv_eq_False[THEN eq_reflection]) show "boolify(t=u) = (t=u)" by (unfold tu, rule boolifyFalse) qed @@ -1795,18 +1697,18 @@ by (simp add: iff_def) lemma iffIsBool [intro!,simp]: "isBool(A \ B)" unfolding isBool_def by (rule boolifyIff) -text {* +text \ We can now rewrite equivalences to equations between ``boolified'' arguments, and this gives rise to a technique for proving equations between formulas. -*} +\ lemma boolEqual: assumes "P \ Q" and "isBool(P)" and "isBool(Q)" shows "P = Q" using assms by auto -text {* +text \ The following variant converts equivalences to equations. It might be useful as a (non-conditional) simplification rule, but I suspect that for goals it is more useful to use the standard introduction @@ -1814,12 +1716,10 @@ text {* For assumptions we can use lemma @{text boolEqual} for turning equivalences into conditional rewrites. -*} +\ lemma iffIsBoolifyEqual: "(A \ B) = (boolify(A) = boolify(B))" -proof (rule boolEqual) - show "(A \ B) \ (boolify(A) = boolify(B))" by (auto simp: boolifyFalseI) -qed (simp_all) + by (rule boolEqual) (auto simp: boolifyFalseI) lemma iffThenBoolifyEqual: assumes "A \ B" shows "boolify(A) = boolify(B)" @@ -1834,21 +1734,23 @@ lemma boolEqualIff: shows "(P = Q) = (P \ Q)" using assms by (auto intro: boolEqual) -ML {* +ML \ structure Simpdata = struct open Simpdata; - val mksimps_pairs = [(@{const_name Not}, (@{thms boolifyFalseI}, true)), + val mksimps_pairs = [(@{const_name not}, (@{thms boolifyFalseI}, true)), (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] @ mksimps_pairs; end; open Simpdata; -*} +\ -declaration {* fn _ => - Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) -*} +declaration \ + fn _ => Simplifier.map_ss (fn ss => ss + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +\ (** TEST DATA ** @@ -1864,39 +1766,39 @@ using assms by simp ** END TEST DATA **) -text {* +text \ The following code rewrites @{text "x = y"} to @{text "FALSE"} in the presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. The simplifier is set up so that @{text "y = x"} is already simplified to @{text "FALSE"}, this code adds symmetry of disequality to simplification. -*} +\ lemma symEqLeft: "(x = y) = b \ (y = x) = b" by (auto simp: boolEqualIff) -simproc_setup neq ("x = y") = {* fn _ => -let - val neq_to_EQ_False = @{thm not_sym} RS @{thm eqBoolean} RS @{thm eq_reflection}; - val symEqLeft_to_symEQLeft = @{thm symEqLeft} RS @{thm eq_reflection}; - fun is_neq lhs rhs thm = - (case Thm.prop_of thm of - _ $ (Not' $ (eq' $ l' $ r')) => - Not' = @{const "Not"} andalso eq' = @{const "eq"} andalso - r' aconv lhs andalso l' aconv rhs +simproc_setup neq ("x = y") = \fn _ => + let + val neq_to_EQ_False = @{thm not_sym} RS @{thm neq_conv_eq_False} RS @{thm eq_reflection}; + val symEqLeft_to_symEQLeft = @{thm symEqLeft} RS @{thm eq_reflection}; + fun is_neq eq lhs rhs thm = + (case Thm.prop_of thm of + _ $ (Not $ (eq' $ l' $ r')) => + Not = @{const \not\} andalso eq' = eq andalso + r' aconv lhs andalso l' aconv rhs | _ $ (eq' $ (eq'' $ l' $ r') $ f') => - eq' = @{const "eq"} andalso eq'' = @{const "eq"} andalso - f' = @{const "FALSE"} andalso r' aconv lhs andalso l' aconv rhs - | _ => false); - fun proc ss ct = - (case Thm.term_of ct of - eq $ lhs $ rhs => - (case find_first (is_neq lhs rhs) (Simplifier.prems_of ss) of - SOME thm => SOME ((thm RS symEqLeft_to_symEQLeft) - handle _ => thm RS neq_to_EQ_False) - | NONE => NONE) - | _ => NONE); -in proc end; -*} + eq' = @{const \eq\} andalso eq'' = @{const \eq\} andalso + f' = @{const \FALSE\} andalso r' aconv lhs andalso l' aconv rhs + | _ => false); + fun proc ss ct = + (case Thm.term_of ct of + eq $ lhs $ rhs => + (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of + SOME thm => SOME ((thm RS symEqLeft_to_symEQLeft) + handle _ => thm RS neq_to_EQ_False) + | NONE => NONE) + | _ => NONE); + in proc end +\ (*** TEST *** lemma "a \ b \ (IF b = a THEN t ELSE u) = u" @@ -1951,80 +1853,76 @@ by auto Check what needs to be done to derive the necessary chains of equivalences. **** - -text {* +text \ The following lemmas are used by the simplifier to rearrange quantifiers. -*} +\ lemma uncurry: assumes "P \ Q \ R" shows "P \ Q \ R" using assms by blast +lemma bool_iffI: + assumes "P \ Q" and "Q \ P" + shows "boolify(P) = boolify(Q)" + sorry + lemma iff_allI: assumes "\x. P(x) = Q(x)" shows "(\x : P(x)) = (\x : Q(x))" - using assms by blast + unfolding assms[THEN eq_reflection] by (rule refl) lemma iff_exI: - assumes "!!x. P x = Q x" - shows "(\x. P x) = (\x. Q x)" - using assms by blast + assumes "\x. P(x) = Q(x)" + shows "(\x : P(x)) = (\x : Q(x))" + unfolding assms[THEN eq_reflection] by (rule refl) -lemma all_comm: - "(\x y. P x y) = (\y x. P x y)" +lemma all_comm: "(\x,y : P(x,y)) = (\y,x : P(x,y))" by blast -lemma ex_comm: - "(\x y. P x y) = (\y x. P x y)" +lemma ex_comm: "(\x,y : P(x,y)) = (\y,x : P(x,y))" by blast -ML {* - use "~~/src/Provers/quantifier1.ML"; - local - val uncurry = - prove_goal (the_context()) "P => Q => R ==> P & Q => R" - (fn prems => [cut_facts_tac prems 1, Blast_tac 1]); - val iff_allI = - prove_goal (the_context()) "(!!x. P(x) <=> Q(x)) ==> (ALL x. P(x)) = (ALL x. Q(x))" - (fn prems => [cut_facts_tac prems 1, - auto_tac (claset() addIs [boolEqual], simpset())]); - val iff_exI = - prove_goal (the_context()) "(!!x. P(x) <=> Q(x)) ==> (EX x. P(x)) = (EX x. Q(x))" - (fn prems => [cut_facts_tac prems 1, - auto_tac (claset() addIs [boolEqual], simpset())]); - val all_comm = - prove_goal (the_context()) "(ALL x y. P(x,y)) = (ALL y x. P(x,y))" - (fn _ => [auto_tac (claset() addIs [boolEqual], simpset())]); - val ex_comm = - prove_goal (the_context()) "(EX x y. P(x,y)) = (EX y x. P(x,y))" - (fn _ => [auto_tac (claset() addIs [boolEqual], simpset())]); - val c_type = Type("c", []); - structure Quantifier1 = Quantifier1Fun( - struct - (* syntax *) - fun dest_eq((c as Const("eq",_)) $ s $ t) = SOME(c,s,t) - | dest_eq _ = NONE; - fun dest_conj((c as Const("conj",_)) $ s $ t) = SOME(c,s,t) - | dest_conj _ = NONE; - fun dest_imp((c as Const("imp",_)) $ s $ t) = SOME(c,s,t) - | dest_imp _ = NONE; - val conj = @{term conj}; - val imp = @{term imp}; - (* rules *) - val iff_reflection = thm "eq_reflection"; - val iffI = @{thm iffI}; - val iff_trans = @{thm trans}; - end); - in - -*} +ML \ + structure Quantifier1 = Quantifier1( + (* syntax *) + fun dest_eq(Const(@{const_name \eq\},_) $ s $ t) = SOME(s,t) + | dest_eq _ = NONE; + fun dest_conj(Const(@{const_name \conj\},_) $ s $ t) = SOME(s,t) + | dest_conj _ = NONE; + fun dest_imp(Const(@{const_name \imp\},_) $ s $ t) = SOME(s,t) + | dest_imp _ = NONE; + val conj = @{term conj}; + val imp = @{term imp}; + (* rules *) + val iff_reflection = @{thm eq_reflection}; + val iffI = @{thm bool_iffI}; + val iff_trans = @{thm trans}; + val conjI= @{thm conjI} + val conjE= @{thm conjE} + val impI = @{thm impI} + val mp = @{thm mp} + val uncurry = @{thm uncurry} + val exI = @{thm exI} + val exE = @{thm exE} + val iff_allI = @{thm iff_allI} + val iff_exI = @{thm iff_exI} + val all_comm = @{thm all_comm} + val ex_comm = @{thm ex_comm} + ); +\ + +simproc_setup defined_Ex ("\x : P(x)") = \fn _ => Quantifier1.rearrange_ex\ +simproc_setup defined_All ("\x : P(x)") = \fn _ => Quantifier1.rearrange_all\ + +lemma "\x,y : y \x \ x=y \ x = FALSE" + apply simp ***) -text {* +text \ Orient equations with Boolean constants such that the constant appears on the right-hand side. -*} +\ lemma boolConstEqual [simp]: "(TRUE = P) = (P = TRUE)" @@ -2032,7 +1930,43 @@ lemma boolConstEqual [simp]: by blast+ -subsubsection {* Simplification laws for conditionals *} +subsubsection \Reorienting equations\ + +lemma eq_commute: "(a = b) = (b = a)" + by (auto intro: boolEqual) + +ML \ +signature REORIENT_PROC = +sig + val add : (term -> bool) -> theory -> theory + val proc : morphism -> Proof.context -> cterm -> thm option +end; + +structure Reorient_Proc : REORIENT_PROC = +struct + structure Data = Theory_Data + ( + type T = ((term -> bool) * stamp) list; + val empty = []; + val extend = I; + fun merge data : T = Library.merge (eq_snd (op =)) data; + ); + fun add m = Data.map (cons (m, stamp ())); + fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); + + val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; + fun proc _ ctxt ct = + let + val thy = Proof_Context.theory_of ctxt; + in + case Thm.term_of ct of + (_ $ _ $ u) => if matches thy u then NONE else SOME meta_reorient + | _ => NONE + end; +end; +\ + +subsubsection \ Simplification laws for conditionals \ lemma splitCond [split]: assumes q: "\x. isBool(Q(x))" @@ -2045,7 +1979,8 @@ next thus ?thesis by (auto intro: q) qed -lemma splitCondAsm: -- {* useful with conditionals in hypotheses *} +text \useful with conditionals in hypotheses\ +lemma splitCondAsm: assumes "\x. isBool(Q(x))" shows "Q(IF P THEN t ELSE e) = (\((P \ \Q(t)) \ (\P \ \Q(e))))" using assms by (simp only: splitCond, blast) @@ -2054,7 +1989,8 @@ lemma condCong (*[cong]*): (* strangely, seems to simplify less when active ?! "P = Q \ (IF P THEN t ELSE e) = (IF Q THEN t ELSE e)" by simp -lemma condFullCong: -- {* not active by default, because too expensive *} +text \not active by default, because too expensive\ +lemma condFullCong: "\P = Q; Q \ t=t'; \Q \ e=e'\ \ (IF P THEN t ELSE e) = (IF Q THEN t' ELSE e')" by auto @@ -2065,7 +2001,6 @@ lemma substCond [intro]: using assms by auto lemma cond_simps [simp]: - "(IF x = y THEN y ELSE x) = x" "(IF (IF A THEN B ELSE C) THEN t ELSE e) = (IF (A \ B) \ (\A \ C) THEN t ELSE e)" "(IF A THEN (IF B THEN t ELSE u) ELSE v) = @@ -2073,7 +2008,7 @@ lemma cond_simps [simp]: by auto -subsubsection {* Simplification laws for conjunction *} +subsubsection \ Simplification laws for conjunction \ lemma conj_simps [simp]: "(P \ TRUE) = boolify(P)" @@ -2089,17 +2024,17 @@ lemma conj_simps [simp]: "((P \ Q) \ R) = (P \ Q \ R)" by auto -text {* +text \ The congruence rule for conjunction is occasionally useful, but not active by default. -*} +\ lemma conj_cong: assumes "P = P'" and "P' \ Q = Q'" shows "(P \ Q) = (P' \ Q')" using assms by auto -text {* Commutativity laws are not active by default *} +text \Commutativity laws are not active by default\ lemma conj_comms: "(P \ Q) = (Q \ P)" @@ -2107,7 +2042,7 @@ lemma conj_comms: by auto -subsubsection {* Simplification laws for disjunction *} +subsubsection \Simplification laws for disjunction\ lemma disj_simps [simp]: "(P \ TRUE) = TRUE" @@ -2123,14 +2058,14 @@ lemma disj_simps [simp]: "((P \ Q) \ R) = (P \ Q \ R)" by auto -text {* Congruence rule, not active by default *} +text \Congruence rule, not active by default\ lemma disj_cong: assumes "P = P'" and "\P' \ Q = Q'" shows "(P \ Q) = (P' \ Q')" using assms by auto -text {* Commutativity laws are not active by default *} +text \Commutativity laws are not active by default\ lemma disj_comms: "(P \ Q) = (Q \ P)" @@ -2138,13 +2073,13 @@ lemma disj_comms: by auto -subsubsection {* Simplification laws for negation *} +subsubsection \Simplification laws for negation\ -text {* +text \ Negated formulas create simplifications of the form @{text "A = FALSE"}, we therefore prove two versions of the - following lemmas to complete critical pairs. -*} + following lemmas in order to complete critical pairs. +\ lemma not_simps [simp]: "(\(P \ Q)) = (\P \ \Q)" @@ -2156,14 +2091,14 @@ lemma not_simps [simp]: "(x \ x) = FALSE" "\P. (\(\x : P(x))) = (\x : \P(x))" "\P. (\(\x : P(x))) = (\x : \P(x))" -by (auto simp del: notBoolifyFalse) + by (auto simp del: notBoolifyFalse) declare not_simps [simplified,simp] lemma eqFalse_eqFalse [simp]: "isBool(P) \ ((P = FALSE) = FALSE) = P" by auto -subsubsection {* Simplification laws for implication *} +subsubsection \Simplification laws for implication\ lemma imp_simps [simp]: "(P \ FALSE) = (\P)" @@ -2179,7 +2114,7 @@ lemma imp_cong [cong]: by auto -subsubsection {* Simplification laws for equivalence *} +subsubsection \Simplification laws for equivalence\ lemma iff_simps [simp]: "(TRUE \ P) = boolify(P)" @@ -2194,7 +2129,7 @@ lemma iff_cong [cong]: by auto -subsubsection {* Simplification laws for quantifiers *} +subsubsection \Simplification laws for quantifiers\ (*** subsumed by following simplifications lemma [simp]: "\x : TRUE" @@ -2212,8 +2147,7 @@ lemma quant_simps [simp]: "\ P. (\x : t=x \ P(x)) = boolify(P(t))" by auto - -text {* Miniscoping of quantifiers. *} +text \Miniscoping of quantifiers\ lemma miniscope_ex [simp]: "\P Q. (\x : P(x) \ Q) = ((\x : P(x)) \ Q)" @@ -2238,15 +2172,15 @@ by (rule chooseI, rule refl) declare choose_det [cong] -text {* +text \ A @{text CHOOSE} expression evaluates to @{text default} if the only possible value satisfying the predicate equals @{text default}. Note that the reverse implication is not necessarily true: there could be several values satisfying @{text "P(x)"}, including @{text default}, - and @{text CHOOSE} may return @{text default}. This rule can be useful - for reasoning about \textsc{case} expressions where none of the guards - is true. -*} + and @{text CHOOSE} may return @{text default}. The following rule + can be useful for reasoning about \textsc{case} expressions where + none of the guards is true. +\ lemma equal_default [intro!]: assumes p: "\x : P(x) \ x = default" @@ -2264,11 +2198,9 @@ next unfolding default_def by (blast intro: choose_det) qed -lemmas [intro!] = sym[OF equal_default, standard] +lemmas [intro!] = sym[OF equal_default] -text {* - Similar lemma for @{text arbitrary}. -*} +text \Similar lemma for @{text arbitrary}.\ lemma equal_arbitrary: assumes p: "\x : P(x)" @@ -2279,9 +2211,9 @@ unfolding arbitrary_def proof (rule choose_det) qed -subsubsection {* Distributive laws *} +subsubsection \Distributive laws\ -text {* Not active by default. *} +text \Not active by default\ lemma prop_distrib: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" @@ -2299,7 +2231,7 @@ lemma quant_distrib: by auto -subsubsection {* Further calculational laws *} +subsubsection \Further calculational laws\ lemma cases_simp (*[simp]*): "((P \ Q) \ (\P \ Q)) = boolify(Q)" by auto diff --git a/isabelle/ROOT b/isabelle/ROOT new file mode 100644 index 00000000..53c1d9d9 --- /dev/null +++ b/isabelle/ROOT @@ -0,0 +1,7 @@ +(* build the session using "isabelle build -D ." *) +session "TLA+" = "Pure" + + options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] + theories + Peano + document_files + "root.tex" diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index fa227b21..e37ff62f 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -1,107 +1,92 @@ (* Title: TLA+/SetTheory.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:48 merz> + Version: Isabelle2020 *) -header {* \tlaplus{} Set Theory *} +section \\tlaplus{} Set Theory\ theory SetTheory imports PredicateLogic begin -text {* - This theory defines the version of Zermelo-Fr\"ankel set theory +text \ + This theory defines the version of Zermelo-Fraenkel set theory that underlies \tlaplus{}. -*} +\ -subsection {* Basic syntax and axiomatic basis of set theory. *} +subsection \Basic syntax and axiomatic basis of set theory\ -text {* - We take the set-theoretic constructs of \tlaplus{}, but add generalized - intersection for symmetry and convenience. (Note that @{text "INTER {} = {}"}.) -*} +text \ + The following constants are at the basis of our set theory. +\ consts - emptyset :: "c" ("{}" 100) -- {* empty set *} - upair :: "[c, c] \ c" -- {* unordered pairs *} - addElt :: "[c, c] \ c" -- {* add element to set *} - infinity :: "c" -- {* infinity set *} - SUBSET :: "c \ c" ("SUBSET _" [100]90) -- {* power set *} - UNION :: "c \ c" ("UNION _" [100]90) -- {* generalized union *} - INTER :: "c \ c" ("INTER _" [100]90) -- {* generalized intersection *} - "cup" :: "[c, c] \ c" (infixl "\\cup" 65) -- {* binary union *} - "cap" :: "[c, c] \ c" (infixl "\\cap" 70) -- {* binary intersection *} - "setminus":: "[c, c] \ c" (infixl "\\" 65) -- {* binary set difference *} - "in" :: "[c, c] \ c" (infixl "\\in" 50) -- {* membership relation *} - "subseteq":: "[c, c] \ c" (infixl "\\subseteq" 50) -- {* subset relation *} - subsetOf :: "[c, c \ c] \ c" -- {*@{text "subsetOf(S,p)"} = $\{x \in S : p\}$*} - setOfAll :: "[c, c \ c] \ c" -- {*@{text "setOfAll(S,e)"} = $\{e : x \in S\}$*} - bChoice :: "[c, c \ c] \ c" -- {* bounded choice *} - bAll :: "[c, c \ c] \ c" -- {* bounded universal quantifier *} - bEx :: "[c, c \ c] \ c" -- {* bounded existential quantifier *} - -notation (xsymbols) - "cup" (infixl "\" 65) and - "cap" (infixl "\" 70) and - "setminus" (infixl "\" 65) and - "in" (infixl "\" 50) and - "subseteq" (infixl "\" 50) - -notation (HTML output) - "cup" (infixl "\" 65) and - "cap" (infixl "\" 70) and - "setminus" (infixl "\" 65) and - "in" (infixl "\" 50) and - "subseteq" (infixl "\" 50) - -abbreviation "notin(a, S) \ \(a \ S)" -- {* negated membership *} + "in" :: "[c, c] \ c" (infixl "\" 50) (* set membership *) + emptyset :: "c" ("{}" 100) (* empty set *) + SUBSET :: "c \ c" ("SUBSET _" [100]90) (* power set *) + UNION :: "c \ c" ("UNION _" [100]90) (* generalized union *) + subsetOf :: "[c, c \ c] \ c" (* subsetOf(S,p) = {x \in S : p} *) + setOfAll :: "[c, c \ c] \ c" (* setOfAll(S,e) = {e : x \in S} *) + infinity :: "c" (* infinity set *) + +notation (ASCII) + "in" (infixl "\\in" 50) + +text \Negated set membership\ + +abbreviation "notin(a, S) \ \(a \ S)" notation - "notin" (infixl "\\notin" 50) -notation (xsymbols) - "notin" (infixl "\" 50) -notation (HTML output) "notin" (infixl "\" 50) +notation (ASCII) + "notin" (infixl "\\notin" 50) -abbreviation (input) "supseteq(S, T) \ T \ S" -notation - "supseteq" (infixl "\\supseteq" 50) -notation (xsymbols) - "supseteq" (infixl "\" 50) -notation (HTML output) - "supseteq" (infixl "\" 50) +text \ + We introduce some more notation and operators before stating the + axioms of set theory, in order to make them more readable. +\ -text {* Concrete syntax: proper sub and superset *} +text \ Concrete syntax: set comprehension \ -definition "psubset" :: "[c, c] \ c" (infixl "\\subset" 50) -where "S \\subset T \ S \ T \ S \ T" +(*** + We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} + because logical constants (such as setOfAll) must have fixed arity. + Multiple bindings will be introduced as shorthands involving tuples + when tuples are available -- fortunately we don't need them earlier! -notation (xsymbols) - "psubset" (infixl "\" 50) -notation (HTML output) - "psubset" (infixl "\" 50) + Note that the expression "{x \ S : y \ T}" is ambiguous + (as mentioned in the TLA+ book). In such cases, use logical syntax, e.g. + subsetOf(S, \x. y \ T). See the definition cap_def below. +***) +syntax + "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \ _})") + "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \ _ : _})") +syntax (ASCII) + "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \\in _})") + "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \\in _ : _})") +translations + "{e : x \ S}" \ "CONST setOfAll(S, \x. e)" + "{x \ S : P}" \ "CONST subsetOf(S, \x. P)" -abbreviation (input) "psupset(S,T) \ T \ S" -notation - "psupset" (infix "\\supset" 50) -notation (xsymbols) - "psupset" (infix "\" 50) -notation (HTML output) - "psupset" (infix "\" 50) +text \Unordered pairs\ -lemma psubset_intro [intro!]: - "\ S \ T ; S \ T \ \ S \ T" -unfolding psubset_def by safe +definition upair :: "[c, c] \ c" + where "upair(a,b) \ { IF x={} THEN a ELSE b : x \ SUBSET (SUBSET {}) }" -lemma psubset_elim [elim!]: - "\ S \ T ; \ S \ T ; S \ T \ \ C \ \ C" -unfolding psubset_def by safe +text \Binary set union\ + +definition union :: "[c, c] \ c" (infixl "\" 65) + where "A \ B \ UNION upair(A,B)" +notation (ASCII) + union (infixl "\\union" 65) -text {* Concrete syntax: set enumerations *} +definition addElt :: "[c, c] \ c" + where "addElt(a, A) \ upair(a,a) \ A" + +text \ Concrete syntax: set enumerations \ nonterminal cs @@ -117,7 +102,18 @@ translations abbreviation BOOLEAN :: "c" where "BOOLEAN \ {TRUE, FALSE}" -text {* Concrete syntax: bounded quantification *} +text \Bounded quantification\ + +definition bChoose :: "[c, c \ c] \ c" + where "bChoose(A,P) \ CHOOSE x : x \ A \ P(x)" + +definition bEx :: "[c, c \ c] \ c" + where "bEx(A, P) \ \x : x \ A \ P(x)" + +definition bAll :: "[c, c \ c] \ c" + where "bAll(A, P) \ \x : x \ A \ P(x)" + +text \ Concrete syntax: bounded quantification \ (*** FIXME: allow multiple bindings as in "ALL x:S, y:T. P(x,y)". The following was a valiant attempt to define appropriate syntax @@ -127,23 +123,23 @@ text {* Concrete syntax: bounded quantification *} nonterminal sbinds syntax - "@sbind" :: "[idt,c] \ sbinds" ("_ \\in _") - "@sbinds" :: "[idt,c,sbinds] \ sbinds" ("_ \\in _,/ _") + "@sbind" :: "[idt,c] \ sbinds" ("_ \\in _" [100, 10] 100) + "@sbinds" :: "[idt,c,sbinds] \ sbinds" ("_ \\in _,/ _" [100, 10, 100] 100) syntax (xsymbols) - "@sbind" :: "[idt,c] \ sbinds" ("_ \ _") - "@sbinds" :: "[idt,c, sbinds] \ sbinds" ("_ \ _,/ _") + "@sbind" :: "[idt,c] \ sbinds" ("_ \ _" [100, 10] 100) + "@sbinds" :: "[idt,c, sbinds] \ sbinds" ("_ \ _,/ _" [100, 10, 100] 100) syntax "@bChoice" :: "[sbinds, c] \ c" ("(3CHOOSE _ :/ _)" 10) - "@bEx" :: "[sbinds, c] \ c" ("(3\\E _ :/ _)" 10) - "@bAll" :: "[sbinds, c] \ c" ("(3\\A _ :/ _)" 10) + "@bEx" :: "[sbinds, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) + "@bAll" :: "[sbinds, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) syntax (xsymbols) - "@bEx" :: "[sbinds, c] \ c" ("(3\_ :/ _)" 10) - "@bAll" :: "[sbinds, c] \ c" ("(3\_ :/ _)" 10) + "@bEx" :: "[sbinds, c] \ c" ("(3\_ :/ _)" [100,10] 10) + "@bAll" :: "[sbinds, c] \ c" ("(3\_ :/ _)" [100,10] 10) translations - "ALL x:S, sbs. P" \ "bAll(S, (ALL sbs. (\x. P)))" - "ALL x:S. P" \ "bAll(S, \x. P)" + "\x \ S, bds : P" \ "CONST bAll(S, \x. \bds : P)" + "\x \ S : P" \ "CONST bAll(S, \x. P)" ***) (*** TEMPORARY FIX: @@ -152,16 +148,16 @@ translations ***) syntax - (* Again, only single variable for bounded choice. *) - "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \\in _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) -syntax (xsymbols) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) +syntax (ASCII) + (* Again, only single variable for bounded choice. *) + "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) translations - "CHOOSE x \ S : P" \ "CONST bChoice(S, \x. P)" + "CHOOSE x \ S : P" \ "CONST bChoose(S, \x. P)" (* the following cannot be a print macro because the RHS is non-linear -- need a print translation for display *) "\x, xs \ S : P" \ "CONST bEx(S, \x. \xs \ S : P)" @@ -169,7 +165,7 @@ translations "\x, xs \ S : P" \ "CONST bAll(S, \x. \xs \ S : P)" "\x \ S : P" \ "CONST bAll(S, \x. P)" -print_translation {* +print_translation \ let fun bEx_tr' [S, Abs(x, T, P as (Const (@{const_syntax "bEx"},_) $ S' $ Q))] = (* bEx(S, bEx(S', Q)) => \\E x,y \\in S : Q if S = S' *) @@ -201,9 +197,10 @@ print_translation {* in (Syntax.const "@bAll") $ x' $ S $ P' end | bAll_tr' _ = raise Match; - in [(@{const_syntax "bEx"}, bEx_tr'), (@{const_syntax "bAll"}, bAll_tr')] + in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), + (@{const_syntax "bAll"}, (fn _ => bAll_tr'))] end -*} +\ (*** TEST DATA for parse and print translations *** lemma "\x,y,z \ S : x = y \ y = z" @@ -222,90 +219,86 @@ lemma "\x,y \ S : \z \ T : x=y \ y \ z" oops *** END TEST DATA ***) -text {* Concrete syntax: set comprehension *} - -(*** - We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} - because logical constants (such as setOfAll) must have fixed arity. - Multiple bindings will be introduced as shorthands involving tuples - when tuples are available -- fortunately we don't need them earlier! - - Note that the expression "{x \ S : y \ T}" is ambiguous - (as mentioned in the TLA+ book). In such cases, use logical syntax, e.g. - subsetOf(S, \x. y \ T). See the definition cap_def below. -***) -syntax - "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \\in _})") - "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \\in _ : _})") -syntax (xsymbols) - "@setOfAll" :: "[c, idt, c] \ c" ("(1{_ : _ \ _})") - "@subsetOf" :: "[idt, c, c] \ c" ("(1{_ \ _ : _})") -translations - "{e : x \ S}" \ "CONST setOfAll(S, \x. e)" - "{x \ S : P}" \ "CONST subsetOf(S, \x. P)" - -text {* - The following definitions make the axioms of set theory more readable. - Observe that @{text \} is treated as an uninterpreted predicate symbol. -*} - -defs - bChoose_def: "bChoice(A,P) \ CHOOSE x : x \ A \ P(x)" - bEx_def: "bEx(A, P) \ \x : x \ A \ P(x)" - bAll_def: "bAll(A, P) \ \x : x \ A \ P(x)" - subset_def: "A \ B \ \x \ A : x \ B" - -text {* - We now state a first batch of axioms of set theory: extensionality +text \ + We now state the axioms of set theory: extensionality and the definitions of @{text UNION}, @{text SUBSET}, and @{text setOfAll}. Membership is also asserted to be produce Boolean values---in traditional presentations of ZF set theory this is ensured by distinguishing sorts of - terms and formulas. -*} - + terms and formulas. @{term infinity} is some infinite set, but it is not + uniquely defined. The foundation axiom rules out sets that are ``too big''. +\ axiomatization where inIsBool [intro!,simp]: "isBool(x \ A)" and - extension: "(A = B) \ (A \ B) \ (B \ A)" + extension: "(A = B) \ (\x : x \ A \ x \ B)" +and + UNION: "(A \ UNION S) \ (\B \ S : A \ B)" and - UNION: "(A \ UNION S) \ (\B\S : A \ B)" + SUBSET: "(A \ SUBSET S) \ (\x \ A : x \ S)" and - SUBSET: "(A \ SUBSET S) \ (A \ S)" + setOfAll: "(y \ setOfAll(S, e)) \ (\x : x \ S \ y = e(x))" and - setOfAll: "(y \ { e(x) : x \ S }) \ (\x\S : y = e(x))" + subsetOf: "(y \ subsetOf(S, P)) \ (y \ S \ P(y))" and - subsetOf: "(y \ { x \ S : P(x) }) \ (y \ S \ P(y))" + infinity: "({} \ infinity) \ (\x \ infinity : {x} \ x \ infinity)" +and + foundation: "(A = {}) \ (\x \ A : \y \ x : y \ A)" -text {* - Armed with this understanding, we can now define the remaining operators - of set theory. -*} -defs - upair_def: "upair(a,b) \ { IF x={} THEN a ELSE b : x \ SUBSET (SUBSET {}) }" - cup_def: "A \ B \ UNION upair(A,B)" - addElt_def: "addElt(a, A) \ upair(a,a) \ A" - cap_def: "A \ B \ subsetOf(A, \x. x \ B)" - diff_def: "A \ B \ {x \ A : x \ B}" - INTER_def: "INTER A \ { x \ UNION A : \B \ A : x \ B }" +text \More set operations\ -text {* - The following two axioms complete our presentation of set theory. -*} +definition subset :: "[c, c] \ c" (infixl "\" 50) + where "A \ B \ \x \ A : x \ B" -axiomatization where - -- {* @{term infinity} is some infinite set, but it is not uniquely defined. *} - infinity: "({} \ infinity) \ (\x \ infinity : {x} \ x \ infinity)" -and - -- {* The foundation axiom rules out sets that are ``too big''. *} - foundation: "(A = {}) \ (\x \ A : \y \ x : y \ A)" +definition inter :: "[c, c] \ c" (infixl "\" 70) (* binary intersection *) + where "A \ B \ subsetOf(A, \x. x \ B)" + +definition setminus :: "[c, c] \ c" (infixl "\" 65) (* set difference *) + where "A \ B \ {x \ A : x \ B}" + +definition INTER :: "c \ c" ("INTER _" [100]90) (* generalized intersection *) + where "INTER A \ { x \ UNION A : \B \ A : x \ B }" + +text \Note that @{text "INTER {}"} is @{text "{}"}.}\ + +notation (ASCII) + "inter" (infixl "\\inter" 70) and + "setminus" (infixl "\\" 65) and + "subset" (infixl "\\subseteq" 50) + +abbreviation (input) "supset(S, T) \ T \ S" +notation + "supset" (infixl "\" 50) +notation (ASCII) + "supset" (infixl "\\supseteq" 50) + +definition "psubset" :: "[c, c] \ c" (infixl "\" 50) +where "S \ T \ S \ T \ S \ T" + +notation (ASCII) + "psubset" (infixl "\\subset" 50) -subsection {* Boolean operators *} +abbreviation (input) "psupset(S,T) \ T \ S" +notation + "psupset" (infix "\" 50) +notation (ASCII) + "psupset" (infix "\\supset" 50) + +lemma psubset_intro [intro!]: + "\ S \ T ; S \ T \ \ S \ T" +unfolding psubset_def by safe -text {* +lemma psubset_elim [elim!]: + "\ S \ T ; \ S \ T ; S \ T \ \ C \ \ C" +unfolding psubset_def by safe + + +subsection \ Boolean operators \ + +text \ The following lemmas assert that certain operators always return Boolean values; these are helpful for the automated reasoning methods. -*} +\ lemma boolifyIn [simp]: "boolify(x \ A) = (x \ A)" by (rule inIsBool[unfolded isBool_def]) @@ -331,6 +324,12 @@ by (simp add: subset_def) lemma subsetIsBool [intro!,simp]: "isBool(A \ B)" by (unfold isBool_def, rule boolifySubset) +lemma boolifypSubset [simp]: "boolify(A \ B) = (A \ B)" +by (simp add: psubset_def) + +lemma psubsetIsBool [intro!,simp]: "isBool(A \ B)" +by (unfold isBool_def, rule boolifypSubset) + lemma [intro!]: "\isBool(P); x\S \ P\ \ (x\S) = P" "\isBool(P); P \ x\S\ \ P = (x\S)" @@ -343,7 +342,7 @@ lemma [intro!]: by auto -subsection {* Substitution rules *} +subsection \ Substitution rules \ lemma subst_elem [trans]: assumes "b \ A" and "a=b" @@ -366,7 +365,7 @@ lemma subst_set_rev [trans]: using assms by simp -subsection {* Bounded quantification *} +subsection \ Bounded quantification \ lemma bAllI [intro!]: assumes "\x. x \ A \ P(x)" @@ -379,8 +378,8 @@ lemma bspec [dest?]: using assms unfolding bAll_def by blast (*** currently inactive -- causes several obscure warnings about renamed - variables and seems to slow down large examples such as AtomicBakery -ML {* + variables and seems to slow down large examples such as AtomicBakery ***) +ML \ structure Simpdata = struct @@ -391,12 +390,15 @@ ML {* end; open Simpdata; -*} +\ -declaration {* fn _ => - Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) -*} -***) +declaration \ + fn _ => Simplifier.map_ss (fn ss => ss + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +\ + +(***) lemma bAllE [elim]: assumes "\x\A : P(x)" and "x \ A \ Q" and "P(x) \ Q" @@ -416,7 +418,7 @@ lemma bExI [intro]: shows "\x\A : P(x)" using assms unfolding bEx_def by blast -lemma bExCI: -- {* implicit proof by contradiction *} +lemma bExCI: (* proof by contradiction *) assumes "(\x\A : \P(x)) \ P(a)" and "a \ A" shows "\x\A : P(x)" using assms by blast @@ -499,9 +501,9 @@ qed TODO: investigate the issue and find out what is really needed about type checking. -subsubsection {* Setting up type checking for \tlaplus{}. *} +subsubsection \ Setting up type checking for \tlaplus{}. \ -text {* +text \ Type checking is intended to solve simple goals, for example showing that conjunctions are Boolean, and that the result of a conditional expression is a natural number if both branches yield naturals. Unlike @@ -512,7 +514,7 @@ text {* unknowns. The code is essentially taken from Isabelle/ZF. -*} +\ use "typecheck.ML" setup TypeCheck.setup @@ -538,7 +540,7 @@ by typecheck ***) -subsection {* Simplification of conditional expressions *} +subsection \ Simplification of conditional expressions \ lemma inCond [simp]: "(a \ (IF P THEN S ELSE T)) = ((P \ a \ S) \ (\P \ a \ T))" by (force intro: condI elim: condE) @@ -583,7 +585,7 @@ by (force intro: condI elim: condE) **) -subsection {* Rules for subsets and set equality *} +subsection \ Rules for subsets and set equality \ lemma subsetI [intro!]: assumes "\x. x \ A \ x \ B" @@ -599,7 +601,7 @@ lemma rev_subsetD [trans]: "\ c \ A; A \ B \ \ c \ B" by (rule subsetD) -lemma subsetCE [elim]: -- {* elimination rule for classical logic *} +lemma subsetCE [elim]: assumes "A \ B" and "c \ A \ P" and "c \ B \ P" shows "P" using assms unfolding subset_def by blast @@ -630,7 +632,7 @@ lemma notSubsetE (*[elim!]*): shows "P" using assms by blast -text {* The subset relation is a partial order. *} +text \ The subset relation is a partial order. \ lemma subsetRefl [simp,intro!]: "A \ A" by blast @@ -650,45 +652,48 @@ lemma setEqualI: shows "A = B" by (rule setEqual, (blast intro: assms)+) -text {* +text \ The rule @{text setEqualI} is too general for use as a default introduction rule: we don't want to apply it for Booleans, for example. However, instances where at least one term results from a set constructor are useful. -*} +\ lemmas (** the instances for the empty set are superseded by rules below - setEqualI [where A = "{}", standard, intro!] - setEqualI [where B = "{}", standard, intro!] + setEqualI [where A = "{}", intro!] + setEqualI [where B = "{}", intro!] **) - setEqualI [where A = "addElt(a,C)", standard, intro!] - setEqualI [where B = "addElt(a,C)", standard, intro!] - setEqualI [where A = "SUBSET C", standard, intro!] - setEqualI [where B = "SUBSET C", standard, intro!] - setEqualI [where A = "UNION C", standard, intro!] - setEqualI [where B = "UNION C", standard, intro!] - setEqualI [where A = "INTER C", standard, intro!] - setEqualI [where B = "INTER C", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "C \ D", standard, intro!] - setEqualI [where B = "C \ D", standard, intro!] - setEqualI [where A = "subsetOf(S,P)", standard, intro!] - setEqualI [where B = "subsetOf(S,P)", standard, intro!] - setEqualI [where A = "setOfAll(S,e)", standard, intro!] - setEqualI [where B = "setOfAll(S,e)", standard, intro!] - -lemmas setEqualD1 = extension[THEN iffD1, THEN conjD1, standard] -- {* @{text "A = B \ A \ B"} *} -lemmas setEqualD2 = extension[THEN iffD1, THEN conjD2, standard] -- {* @{text "A = B \ B \ A"} *} - -text {* - We declare the elimination rule for set equalities as an unsafe + setEqualI [where A = "addElt(a,C)" for a C, intro!] + setEqualI [where B = "addElt(a,C)" for a C, intro!] + setEqualI [where A = "SUBSET C" for C, intro!] + setEqualI [where B = "SUBSET C" for C, intro!] + setEqualI [where A = "UNION C" for C, intro!] + setEqualI [where B = "UNION C" for C, intro!] + setEqualI [where A = "INTER C" for C, intro!] + setEqualI [where B = "INTER C" for C, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "C \ D" for C D, intro!] + setEqualI [where B = "C \ D" for C D, intro!] + setEqualI [where A = "subsetOf(S,P)" for S P, intro!] + setEqualI [where B = "subsetOf(S,P)" for S P, intro!] + setEqualI [where A = "setOfAll(S,e)" for S e, intro!] + setEqualI [where B = "setOfAll(S,e)" for S e, intro!] + +lemma setEqualD1: "A = B \ A \ B" + by blast + +lemma setEqualD2: "A = B \ B \ A" + by blast + +text \ + We declare the elimination rule for set equations as an unsafe rule to use with the classical reasoner, so it will be tried if the more obvious uses of equality fail. -*} +\ lemma setEqualE [elim]: assumes "A = B" and "\ c \ A; c \ B \ \ P" and "\ c \ A; c \ B \ \ P" @@ -696,10 +701,10 @@ lemma setEqualE [elim]: using assms by (blast dest: setEqualD1 setEqualD2) lemma setEqual_iff: "(A = B) = (\x : x \ A \ x \ B)" -by (blast intro: setEqualI (*elim: setEqualE*)) +by (blast intro: setEqualI) -subsection {* Set comprehension: @{text setOfAll} and @{text subsetOf} *} +subsection \ Set comprehension: @{text setOfAll} and @{text subsetOf} \ (*** lemma setOfAllI: @@ -735,10 +740,10 @@ lemma setOfAll_cong (*[cong]*): shows "{ e(x) : x \ S } = { f(y) : y \ T }" using assms by auto -text {* +text \ The following rule for showing equality of sets defined by comprehension is probably too general to use by default with the automatic proof methods. -*} +\ lemma setOfAllEqual: assumes "\x. x \ S \ \y\T : e(x) = f(y)" @@ -784,7 +789,7 @@ lemma subsetOfEqual: by (safe elim!: assms) -subsection {* @{text UNION} -- basic rules for generalized union *} +subsection \ @{text UNION} -- basic rules for generalized union \ lemma UNIONI [intro]: assumes "B \ C" and "a \ B" @@ -801,14 +806,14 @@ lemma UNION_iff [simp]: by blast -subsection {* The empty set *} +subsection \ The empty set \ -text {* +text \ Proving that the empty set has no elements is a bit tricky. We first show that the set @{term "{x \ {} : FALSE}"} is empty and then use the foundation axiom to show that it equals the empty set. -*} +\ lemma emptysetEmpty: "a \ {}" proof @@ -820,8 +825,8 @@ proof thus "FALSE" by blast qed -text {* @{text "a \ {} \ P"} *} -lemmas emptyE [elim!] = emptysetEmpty[THEN notE, standard] +text \ @{text "a \ {} \ P"} \ +lemmas emptyE [elim!] = emptysetEmpty[THEN notE] lemma [simp]: "(a \ {}) = FALSE" by blast @@ -844,10 +849,10 @@ lemma nonEmpty [simp]: "({} \ A) = (\x : x \ A)" by (blast+) -text {* Complete critical pairs *} +text \ Complete critical pairs \ lemmas nonEmpty' [simp] = nonEmpty[simplified] --- {* @{text "((A = {}) = FALSE) = (\x : x \ A)"}, - @{text "(({} = A) = FALSE) = (\x : x \ A)"} *} +\ \ @{text "((A = {}) = FALSE) = (\x : x \ A)"}, + @{text "(({} = A) = FALSE) = (\x : x \ A)"} \ lemma emptysetD (*[dest]*): assumes "A = {}" @@ -878,7 +883,7 @@ lemma [simp]: by (blast+) -subsection {* @{text SUBSET} -- the powerset operator *} +subsection \ @{text SUBSET} -- the powerset operator \ lemma SUBSETI [intro!]: assumes "A \ B" @@ -894,18 +899,18 @@ lemma SUBSET_iff [simp]: "(A \ SUBSET B) = (A \ B)" by blast -lemmas emptySUBSET = emptySubset[THEN SUBSETI, standard] -- {* @{term "{} \ SUBSET A"} *} -lemmas selfSUBSET = subsetRefl[THEN SUBSETI, standard] -- {* @{term "A \ SUBSET A"} *} +lemmas emptySUBSET = emptySubset[THEN SUBSETI] \ \ @{term "{} \ SUBSET A"} \ +lemmas selfSUBSET = subsetRefl[THEN SUBSETI] \ \ @{term "A \ SUBSET A"} \ -subsection {* @{text INTER} -- basic rules for generalized intersection *} +subsection \ @{text INTER} -- basic rules for generalized intersection \ -text {* +text \ Generalized intersection is not officially part of \tlaplus{} but can easily be defined as above. Observe that the rules are not exactly dual to those for @{text UNION} because the intersection of the empty set is defined to be the empty set. -*} +\ lemma INTERI [intro]: assumes "\B. B \ C \ a \ B" and "\B : B \ C" @@ -922,62 +927,49 @@ lemma INTER_iff [simp]: by blast -subsection {* Binary union, intersection, and difference: basic rules *} +subsection \ Binary union, intersection, and difference: basic rules \ -text {* +text \ We begin by proving some lemmas about the auxiliary pairing operator @{text upair}. None of these theorems is active by default, as the operator is not part of \tlaplus{} and should not occur in actual reasoning. The dependencies between these operators are quite tricky, therefore the order of the first few lemmas in this section is tightly constrained. -*} +\ lemma upairE: assumes "x \ upair(a,b)" and "x=a \ P" and "x=b \ P" shows "P" using assms by (auto simp: upair_def) -lemma cupE [elim!]: +lemma unionE [elim!]: assumes "x \ A \ B" and "x \ A \ P" and "x \ B \ P" shows "P" -using assms by (auto simp: cup_def elim: upairE) +using assms by (auto simp: union_def elim: upairE) lemma upairI1: "a \ upair(a,b)" by (auto simp: upair_def) -lemma singleton_iff [simp]: "(a \ {b}) = (a = b)" -proof auto - assume a: "a \ {b}" - thus "a = b" - by (auto simp: addElt_def elim: upairE) -next - show "b \ {b}" - by (auto simp: addElt_def cup_def intro: upairI1) -qed - -lemma singletonI (*[intro!]*): "a \ {a}" -by simp - lemma upairI2: "b \ upair(a,b)" by (auto simp: upair_def) lemma upair_iff: "(c \ upair(a,b)) = (c=a \ c=b)" by (blast intro: upairI1 upairI2 elim: upairE) -lemma cupI1: +lemma unionI1: assumes "a \ A" shows "a \ A \ B" -using assms by (auto simp: cup_def upair_iff) +using assms by (auto simp: union_def upair_iff) -lemma cupI2: +lemma unionI2: assumes "a \ B" shows "a \ A \ B" -using assms by (auto simp: cup_def upair_iff) +using assms by (auto simp: union_def upair_iff) -lemma cup_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (auto simp: cup_def upair_iff) +lemma union_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (auto simp: union_def upair_iff) -lemma cupCI [intro!]: +lemma unionCI [intro!]: assumes "c \ B \ c \ A" shows "c \ A \ B" using assms by auto @@ -996,6 +988,12 @@ lemma addEltE [elim!]: shows "P" using assms by auto +lemma singleton_iff (*[simp]*): "(a \ {b}) = (a = b)" + by auto + +lemma singletonI (*[intro!]*): "a \ {a}" +by simp + lemma addEltSubsetI: assumes "a \ B" and "A \ B" shows "addElt(a,A) \ B" @@ -1009,7 +1007,7 @@ using assms by blast lemma addEltSubset_iff: "(addElt(a,A) \ B) = (a \ B \ A \ B)" by blast --- {* Adding the two following lemmas to the simpset breaks proofs. *} +text \ Adding the two following lemmas to the simpset breaks proofs. \ lemma addEltEqual_iff: "(addElt(a,A) = S) = (a \ S \ A \ S \ S \ addElt(a,A))" by blast @@ -1021,48 +1019,48 @@ lemma addEltEqualAddElt: (a \ addElt(b,B) \ A \ addElt(b,B) \ b \ addElt(a,A) \ B \ addElt(a,A))" by (auto simp: addEltEqual_iff) -lemma cap_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (simp add: cap_def) +lemma inter_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (simp add: inter_def) -lemma capI [intro!]: +lemma interI [intro!]: assumes "c \ A" and "c \ B" shows "c \ A \ B" using assms by simp -lemma capD1: +lemma interD1: assumes "c \ A \ B" shows "c \ A" using assms by simp -lemma capD2: +lemma interD2: assumes "c \ A \ B" shows "c \ B" using assms by simp -lemma capE [elim!]: +lemma interE [elim!]: assumes "c \ A \ B" and "\ c \ A; c \ B \ \ P" shows "P" using assms by simp -lemma diff_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" -by (simp add: diff_def) +lemma setminus_iff [simp]: "(c \ A \ B) = (c \ A \ c \ B)" +by (simp add: setminus_def) -lemma diffI [intro!]: +lemma setminusI [intro!]: assumes "c \ A" and "c \ B" shows "c \ A \ B" using assms by simp -lemma diffD1: +lemma setminusD1: assumes "c \ A \ B" shows "c \ A" using assms by simp -lemma diffD2: +lemma setminusD2: assumes "c \ A \ B" shows "c \ B" using assms by simp -lemma diffE [elim!]: +lemma setminusE [elim!]: assumes "c \ A \ B" and "\ c \ A; c \ B \ \ P" shows "P" using assms by simp @@ -1097,7 +1095,7 @@ lemma subsetAddEltE [elim]: using assms by (auto simp: subsetAddElt_iff) -subsection {* Consequences of the foundation axiom *} +subsection \ Consequences of the foundation axiom \ lemma inAsym: assumes hyps: "a \ b" "\P \ b \ a" @@ -1123,7 +1121,7 @@ lemma equalNotIn: using assms by (blast elim: inIrrefl) -subsection {* Miniscoping of bounded quantifiers *} +subsection \ Miniscoping of bounded quantifiers \ lemma miniscope_bAll [simp]: "\P Q. (\x\A : P(x) \ Q) = ((\x\A : P(x)) \ (A = {} \ Q))" @@ -1139,11 +1137,11 @@ lemma miniscope_bAll [simp]: "\P Q. (\x \ {y \ A : P(y)} : Q(x)) = (\y\A: P(y) \ Q(y))" by (blast+) -lemma bAllCup [simp]: +lemma bAllUnion [simp]: "(\x \ A \ B : P(x)) = ((\x\A : P(x)) \ (\x\B : P(x)))" by blast -lemma bAllCap [simp]: +lemma bAllInter [simp]: "(\x \ A \ B : P(x)) = (\x\A : x \ B \ P(x))" by blast @@ -1161,21 +1159,21 @@ lemma miniscope_bEx [simp]: "\P Q. (\x \ {y \ S : P(y)} : Q(x)) = (\y\S : P(y) \ Q(y))" by (blast+) --- {* completing critical pairs for negated assumption *} +text \ completing critical pairs for negated assumption \ lemma notbQuant' [simp]: "\P. ((\x\A : P(x)) = FALSE) = (\x\A : \P(x))" "\P. ((\x\A : P(x)) = FALSE) = (\x\A : \P(x))" by (auto simp: miniscope_bAll[simplified] miniscope_bEx[simplified]) -lemma bExistsCup [simp]: +lemma bExistsUnion [simp]: "(\x \ A \ B : P(x)) = ((\x\A : P(x)) \ (\x\B : P(x)))" by blast -lemma bExistsCap [simp]: +lemma bExistsInter [simp]: "(\x \ A \ B : P(x)) = (\x\A : x \ B \ P(x))" by blast -lemma bQuant_distribs: -- {* not active by default *} +lemma bQuant_distribs: \ \ not active by default \ "(\x\A : P(x) \ Q(x)) = ((\x\A : P(x)) \ (\x\A : Q(x)))" "(\x\A : P(x) \ Q(x)) = ((\x\A : P(x)) \ (\x\A : Q(x)))" by (blast+) @@ -1190,7 +1188,7 @@ lemma bQuantOnePoint [simp]: by (blast+) -subsection {* Simplification of set comprehensions *} +subsection \ Simplification of set comprehensions \ lemma comprehensionSimps [simp]: "\e. setOfAll({}, e) = {}" @@ -1206,24 +1204,24 @@ lemma comprehensionSimps [simp]: "\A e. setOfAll(A, \x. e) = (IF A = {} THEN {} ELSE {e})" by auto -text {* The following are not active by default. *} +text \ The following are not active by default. \ lemma comprehensionDistribs: "\e. { e(x) : x \ A \ B } = {e(x) : x\A} \ {e(x) : x\B}" "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" - -- {* @{text setOfAll} and intersection or difference do not distribute *} + \ \ @{text setOfAll} and intersection or difference do not distribute \ "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" "\P. { x \ A \ B : P(x) } = {x\A : P(x)} \ {x\B : P(x)}" by (blast+) -subsection {* Binary union, intersection, and difference: inclusions and equalities *} +subsection \ Binary union, intersection, and difference: inclusions and equations \ -text {* +text \ The following list contains many simple facts about set theory. Only the most trivial of these are included in the default set of rewriting rules. -*} +\ lemma addEltCommute: "addElt(a, addElt(b, C)) = addElt(b, addElt(a, C))" by blast @@ -1234,59 +1232,59 @@ by blast lemma addEltTwice: "addElt(a, addElt(a, A)) = addElt(a, A)" by blast -lemma capAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN addElt(a, B \ C) ELSE B \ C)" +lemma interAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN addElt(a, B \ C) ELSE B \ C)" by (blast intro: condI elim: condE) -lemma capAddEltRight: "C \ addElt(a,B) = (IF a \ C THEN addElt(a, C \ B) ELSE C \ B)" +lemma interAddEltRight: "C \ addElt(a,B) = (IF a \ C THEN addElt(a, C \ B) ELSE C \ B)" by (blast intro: condI elim: condE) -lemma addEltCap: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" +lemma addEltInter: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" by blast -lemma diffAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN B \ C ELSE addElt(a, B \ C))" +lemma setminusAddEltLeft: "addElt(a,B) \ C = (IF a \ C THEN B \ C ELSE addElt(a, B \ C))" by (blast intro: condI elim: condE) -lemma capSubset: "(C \ A \ B) = (C \ A \ C \ B)" +lemma interSubset: "(C \ A \ B) = (C \ A \ C \ B)" by blast -lemma capLB1: "A \ B \ A" +lemma interLB1: "A \ B \ A" by blast -lemma capLB2: "A \ B \ B" +lemma interLB2: "A \ B \ B" by blast -lemma capGLB: +lemma interGLB: assumes "C \ A" and "C \ B" shows "C \ A \ B" using assms by blast -lemma capEmpty [simp]: +lemma interEmpty [simp]: "A \ {} = {}" "{} \ A = {}" by blast+ -lemma capAbsorb [simp]: "A \ A = A" +lemma interAbsorb [simp]: "A \ A = A" by blast -lemma capLeftAbsorb: "A \ (A \ B) = A \ B" +lemma interLeftAbsorb: "A \ (A \ B) = A \ B" by blast -lemma capCommute: "A \ B = B \ A" +lemma interCommute: "A \ B = B \ A" by blast -lemma capLeftCommute: "A \ (B \ C) = B \ (A \ C)" +lemma interLeftCommute: "A \ (B \ C) = B \ (A \ C)" by blast -lemma capAssoc: "(A \ B) \ C = A \ (B \ C)" +lemma interAssoc: "(A \ B) \ C = A \ (B \ C)" by blast -text {* Intersection is an AC operator: can be added to simp where appropriate *} -lemmas capAC = capAssoc capCommute capLeftCommute capLeftAbsorb +text \ Intersection is an AC operator: can be added to simp where appropriate \ +lemmas interAC = interAssoc interCommute interLeftCommute interLeftAbsorb -lemma subsetOfCap: "{x\A : P(x)} \ B = {x \ A \ B : P(x)}" +lemma subsetOfInter: "{x\A : P(x)} \ B = {x \ A \ B : P(x)}" by blast -lemma capSubsetOf: +lemma interSubsetOf: "B \ {x\A : P(x)} = {x \ B\A : P(x)}" by blast @@ -1298,113 +1296,115 @@ lemma subsetOfConj: "{x\A : P(x) \ Q(x)} = {x\A : P(x)} \ {x\A : Q(x)}" by blast -lemma subsetCup: "(A \ B \ C) = (A \ C \ B \ C)" +lemma subsetUnion: "(A \ B \ C) = (A \ C \ B \ C)" by blast -lemma cupUB1: "A \ A \ B" +lemma unionUB1: "A \ A \ B" by blast -lemma cupUB2: "B \ A \ B" +lemma unionUB2: "B \ A \ B" by blast -lemma cupLUB: +lemma unionLUB: assumes "A \ C" and "B \ C" shows "A \ B \ C" using assms by blast -lemma cupEmpty [simp]: +lemma unionEmpty [simp]: "A \ {} = A" "{} \ A = A" by blast+ -lemma cupAddEltLeft: "addElt(a,B) \ C = addElt(a, B \ C)" +lemma unionAddEltLeft: "addElt(a,B) \ C = addElt(a, B \ C)" by blast -lemma cupAddEltRight: "C \ addElt(a,B) = addElt(a, C \ B)" +lemma unionAddEltRight: "C \ addElt(a,B) = addElt(a, C \ B)" by blast -lemma addEltCup: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" +lemma addEltUnion: "addElt(a, B \ C) = addElt(a,B) \ addElt(a,C)" by blast -lemma cupAbsorb [simp]: "A \ A = A" +lemma unionAbsorb [simp]: "A \ A = A" by blast -lemma cupLeftAbsorb: "A \ (A \ B) = A \ B" +lemma unionLeftAbsorb: "A \ (A \ B) = A \ B" by blast -lemma cupCommute: "A \ B = B \ A" +lemma unionCommute: "A \ B = B \ A" by blast -lemma cupLeftCommute: "A \ (B \ C) = B \ (A \ C)" +lemma unionLeftCommute: "A \ (B \ C) = B \ (A \ C)" by blast -lemma cupAssoc: "(A \ B) \ C = A \ (B \ C)" +lemma unionAssoc: "(A \ B) \ C = A \ (B \ C)" by blast -text {* Union is an AC operator: can be added to simp where appropriate *} -lemmas cupAC = cupAssoc cupCommute cupLeftCommute cupLeftAbsorb +text \ Union is an AC operator: can be added to simp where appropriate \ +lemmas unionAC = unionAssoc unionCommute unionLeftCommute unionLeftAbsorb -text {* Lemmas useful for simplifying enumerated sets are active by default *} +text \ Lemmas useful for simplifying enumerated sets are active by default \ lemmas enumeratedSetSimps [simp] = addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice - capAddEltLeft capAddEltRight cupAddEltLeft cupAddEltRight diffAddEltLeft + interAddEltLeft interAddEltRight unionAddEltLeft unionAddEltRight setminusAddEltLeft -lemma cupEqualEmpty [simp]: "(A \ B = {}) = (A = {} \ B = {})" +(* included below +lemma unionEqualEmpty [simp]: "(A \ B = {}) = (A = {} \ B = {})" by blast +*) -lemma capCupDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma interUnionDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma cupCapDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma unionInterDistrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma diffSubset: "A \ B \ A" +lemma setminusSubset: "A \ B \ A" by blast -lemma subsetDiff: +lemma subsetSetminus: assumes "C \ A" and "C \ B = {}" shows "C \ A \ B" using assms by blast -lemma diffSelf [simp]: "A \ A = {}" +lemma setminusSelf [simp]: "A \ A = {}" by blast -lemma diffDisjoint: +lemma setminusDisjoint: assumes "A \ B = {}" shows "A \ B = A" using assms by blast -lemma emptyDiff [simp]: "{} \ A = {}" +lemma emptySetminus [simp]: "{} \ A = {}" by blast -lemma diffAddElt: "A \ addElt(a,B) = (A \ B) \ {a}" +lemma setminusAddElt: "A \ addElt(a,B) = (A \ B) \ {a}" by blast -lemma cupDiffSelf: "A \ (B \ A) = A \ B" +lemma unionSetminusSelf: "A \ (B \ A) = A \ B" by blast -lemma diffCupLeft: "(A \ B) \ C = (A \ C) \ (B \ C)" +lemma setminusUnionLeft: "(A \ B) \ C = (A \ C) \ (B \ C)" by blast -lemma diffCupRight: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma setminusUnionRight: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma cupDiffLeft: "(A \ B) \ C = (A \ C) \ (A \ (B \ C))" +lemma unionSetminusLeft: "(A \ B) \ C = (A \ C) \ (A \ (B \ C))" by blast -lemma cupDiffRight: "C \ (A \ B) = (C \ A) \ (A \ (B \ C))" +lemma unionSetminusRight: "C \ (A \ B) = (C \ A) \ (A \ (B \ C))" by blast -lemma diffCapLeft: "(A \ B) \ C = A \ (B \ C)" +lemma setminusInterLeft: "(A \ B) \ C = A \ (B \ C)" by blast -lemma diffCapRight: "A \ (B \ C) = (A \ B) \ (A \ C)" +lemma setminusInterRight: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast -lemma capDiffLeft: "(A \ B) \ C = (A \ C) \ B" +lemma interSetminusLeft: "(A \ B) \ C = (A \ C) \ B" by blast -lemma capDiffRight: "C \ (A \ B) = (C \ A) \ B" +lemma interSetminusRight: "C \ (A \ B) = (C \ A) \ B" by blast lemma isEmptySimps [simp]: @@ -1425,7 +1425,7 @@ lemma isEmptySimps [simp]: by (blast+) -subsection {* Generalized union: inclusions and equalities *} +subsection \ Generalized union: inclusions and equalities \ lemma UNIONSimps [simp]: "UNION {} = {}" @@ -1446,19 +1446,19 @@ lemma UNION_LUB: shows "UNION A \ C" using assms by blast -lemma UNIONCupDistrib: "UNION (A \ B) = UNION A \ UNION B" +lemma UNIONunionDistrib: "UNION (A \ B) = UNION A \ UNION B" by blast -lemma UNIONCap: "UNION (A \ B) \ UNION A \ UNION B" +lemma UNIONinter: "UNION (A \ B) \ UNION A \ UNION B" by blast lemma UNIONDisjoint: "((UNION A) \ C = {}) = (\B\A : B \ C = {})" by blast -lemma capUNION: "(UNION A) \ B = UNION { C \ B : C \ A }" +lemma interUNION: "(UNION A) \ B = UNION { C \ B : C \ A }" by blast -lemma diffUNIONLeft: "(UNION A) \ B = UNION { C \ B : C \ A }" +lemma setminusUNIONLeft: "(UNION A) \ B = UNION { C \ B : C \ A }" by blast lemma UNION_mono: @@ -1467,7 +1467,7 @@ lemma UNION_mono: using assms by blast -subsection {* Generalized intersection: inclusions and equalities *} +subsection \ Generalized intersection: inclusions and equations \ lemma INTERSimps [simp]: @@ -1491,23 +1491,23 @@ lemma INTER_GLB: shows "C \ INTER A" using assms by blast -lemma INTERCupDistrib: +lemma INTERunionDistrib: assumes "A \ {}" and "B \ {}" shows "INTER (A \ B) = INTER A \ INTER B" using assms by auto -lemma capINTER: "(INTER A) \ B \ INTER {C \ B : C \ A}" +lemma interINTER: "(INTER A) \ B \ INTER {C \ B : C \ A}" by blast -lemma cupINTER: +lemma unionINTER: assumes "A \ {}" shows "(INTER A) \ B = INTER {C \ B : C \ A}" using assms by auto -lemma diffINTERLeft: "(INTER A) \ B = INTER {C \ B : C \ A}" +lemma setminusINTERLeft: "(INTER A) \ B = INTER {C \ B : C \ A}" by auto -lemma diffINTERRight: +lemma setminusINTERRight: assumes "A \ {}" shows "B \ (INTER A) = UNION {B \ C : C \ A}" using assms by auto @@ -1523,16 +1523,16 @@ lemma INTER_anti_mono: using assms by (auto simp: INTER_def) -subsection {* Powerset: inclusions and equalities *} +subsection \ Powerset: inclusions and equations \ lemma SUBSETEmpty [simp]: "SUBSET {} = { {} }" by blast -lemma SUBSETAddElt: +lemma SUBSETaddElt: "SUBSET addElt(a,A) = SUBSET A \ {addElt(a,X) : X \ SUBSET A}" by (rule setEqualI, auto) -lemma cupSUBSET: "(SUBSET A) \ (SUBSET B) \ SUBSET (A \ B)" +lemma unionSUBSET: "(SUBSET A) \ (SUBSET B) \ SUBSET (A \ B)" by blast lemma UNION_SUBSET [simp]: "UNION (SUBSET A) = A" @@ -1544,7 +1544,7 @@ by blast lemma UNION_in_SUBSET: "(UNION A \ SUBSET B) = (A \ SUBSET (SUBSET B))" by blast -lemma SUBSETcap: "SUBSET (A \ B) = SUBSET A \ SUBSET B" +lemma SUBSETinter: "SUBSET (A \ B) = SUBSET A \ SUBSET B" by blast end diff --git a/isabelle/document/root.tex b/isabelle/document/root.tex index 25b77b15..4e2c13c4 100644 --- a/isabelle/document/root.tex +++ b/isabelle/document/root.tex @@ -31,8 +31,8 @@ \newcommand{\tlaplus}{\textrm{\upshape TLA\textsuperscript{+}}} \newcommand{\deq}{\stackrel{\Delta}{=}} -\renewcommand{\isacharbackslash}{\isamath{\backslash}} -\newcommand{\isactrlconst}{} +% \renewcommand{\isacharbackslash}{\isamath{\backslash}} +% \newcommand{\isactrlconst}{} \begin{document} diff --git a/isabelle/simplifier_setup.ML b/isabelle/simplifier_setup.ML index de942192..6957e6a5 100644 --- a/isabelle/simplifier_setup.ML +++ b/isabelle/simplifier_setup.ML @@ -1,9 +1,8 @@ (* Title: TLA+/simplifier_setup.ML - Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria + Copyright (C) 2008-2020 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2009 - Time-stamp: <2011-10-11 14:59:15 merz> + Version: Isabelle2020 Code for setting up the simplifier for TLA+. The code is essentially copied from FOL and HOL. @@ -19,55 +18,54 @@ struct Negation and equivalence will be handled later. The implicit operator below is Trueprop. *) - fun mk_eq th = case concl_of th of - Const(@{const_name "=="},_) $ _ $ _ => th - | _ $ (Const(@{const_name "eq"},_) $ _ $ _) => th RS @{thm eq_reflection} - | _ $ (Const(@{const_name "Not"},_) $ (Const(@{const_name "eq"},_) $ _ $ _)) => - th RS (@{thm eqBoolean} RS @{thm eq_reflection}) - | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection}; + fun mk_eq th = case Thm.concl_of th of + Const(@{const_name \Pure.eq\},_) $ _ $ _ => th + | _ $ (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _) => th RS @{thm eq_reflection} + | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ + (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => + th RS (@{thm neq_conv_eq_False} RS @{thm eq_reflection}) + | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection} (* turn P into P == TRUE *) - fun mk_eq_True (_: simpset) th = + fun mk_eq_True (_: Proof.context) th = SOME ((th RS @{thm eqTrueI}) RS @{thm eq_reflection}) handle Thm.THM _ => NONE (* Produce theorems of the form (P1 =simp=> ... =simp=> Pn ==> x == y) ==> (P1 =simp=> ... =simp=> Pn ==> x = y) *) - fun lift_meta_to_obj_eq i st = + fun lift_meta_to_obj_eq ctxt i st = let - fun count_imp (Const (@{const_name PredicateLogic.simp_implies}, _) $ _ $ Q) = + fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = 1 + count_imp Q | count_imp _ = 0; - val j = count_imp (Logic.strip_assums_concl (List.nth (prems_of st, i - 1))) + val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) in if j = 0 then @{thm meta_to_obj_eq} else let val Ps = map (fn k => Free ("P" ^ string_of_int k, propT)) (1 upto j); fun mk_simp_implies Q = fold_rev (fn R => fn S => - Const (@{const_name PredicateLogic.simp_implies}, propT --> propT --> propT) $ R $ S) Ps Q - val cT = Type ("c", []); - val x = Free ("x", cT); - val y = Free ("y", cT); - val x_eq_y = Const (@{const_name Trueprop}, cT --> propT) $ (Const (@{const_name "eq"}, [cT, cT] ---> cT) $ x $ y) - in Goal.prove_global (Thm.theory_of_thm st) [] - [mk_simp_implies (Logic.mk_equals (x, y))] - (mk_simp_implies x_eq_y) - (fn {prems, ...} => EVERY - [rewrite_goals_tac @{thms simp_implies_def}, - REPEAT (ares_tac (@{thm meta_to_obj_eq} :: - map (rewrite_rule @{thms simp_implies_def}) prems) 1)]) + Const (@{const_name \PredicateLogic.simp_implies\}, propT --> propT --> propT) $ R $ S) Ps Q + in Goal.prove_global (Proof_Context.theory_of ctxt) [] + [mk_simp_implies @{prop \(x::c) == y\}] + (mk_simp_implies @{prop \(x::c) = y\}) + (fn {context = ctxt, prems} => EVERY + [rewrite_goals_tac ctxt @{thms simp_implies_def}, + REPEAT (assume_tac ctxt 1 ORELSE + resolve_tac ctxt + (@{thm meta_to_obj_eq} :: + map (rewrite_rule ctxt @{thms simp_implies_def}) prems) 1)]) end - end; + end (* Congruence rules expecting = instead of == *) - fun mk_meta_cong (_: simpset) rl = zero_var_indexes - (let val rl' = Seq.hd (TRYALL (fn i => fn st => - rtac (lift_meta_to_obj_eq i st) i st) rl) - in (rl' RS @{thm eq_reflection}) + fun mk_meta_cong ctxt rl = + let val rl' = Seq.hd (TRYALL (fn i => fn st => + resolve_tac ctxt [lift_meta_to_obj_eq ctxt i st] i st) rl) + in (rl' RS @{thm eq_reflection}) handle THM _ => - if can Logic.dest_equals (concl_of rl') then rl' + if can Logic.dest_equals (Thm.concl_of rl') then rl' else error "Conclusion of congruence rules must use =-equality" - end); + end |> zero_var_indexes (* (string * (thm list * bool)) list The following list associates a list of theorems and a flag with @@ -78,71 +76,84 @@ struct allows us to discard "type-correctness" hypotheses in conditional rewrite rules that are added later on. *) val mksimps_pairs = - [(@{const_name imp}, ([@{thm mp}], false)), - (@{const_name conj}, ([@{thm conjD1}, @{thm conjD2}], false)), - (@{const_name All}, ([@{thm spec}], false)), - (@{const_name TRUE}, ([], false)), - (@{const_name FALSE}, ([], false)), - (@{const_name cond}, ([@{thm if_bool_eq_conj} RS @{thm iffD1}], false)) + [(@{const_name \imp\}, ([@{thm mp}], false)), + (@{const_name \conj\}, ([@{thm conjD1}, @{thm conjD2}], false)), + (@{const_name \All\}, ([@{thm spec}], false)), + (@{const_name \TRUE\}, ([], false)), + (@{const_name \FALSE\}, ([], false)), + (@{const_name \cond\}, ([@{thm if_bool_equiv_conj} RS @{thm iffD1}], false)) ]; (* val mk_atomize: (string * (thm list * bool)) list -> thm -> thm list *) - fun mk_atomize pairs = + fun mk_atomize ctxt pairs = let fun atoms thm = let - val simps = global_simpset_of (theory_of_thm thm); fun res th = map (fn rl => th RS rl); (* may raise exception THM *) + val thm_ctxt = Variable.declare_thm thm ctxt fun res_fixed rls = if Thm.maxidx_of (Thm.adjust_maxidx_thm ~1 thm) = ~1 then res thm rls - else Variable.trade (K (fn [thm'] => res thm' rls)) (Variable.global_thm_context thm) [thm]; + else Variable.trade (K (fn [thm'] => res thm' rls)) thm_ctxt [thm]; in - case concl_of thm of - Const(@{const_name Trueprop},_) $ p => + case Thm.concl_of thm of + Const(@{const_name \Trueprop\},_) $ p => (case head_of p of Const(a,_) => (case AList.lookup (op =) pairs a of SOME (rls, b) => let - val rls = maps atoms (res_fixed rls) handle THM _ => [thm] + val rrls = maps atoms (res_fixed rls) handle THM _ => [thm] in - if b then map (simplify simps) rls - else rls + if b then map (simplify ctxt) rrls + else rrls end | NONE => [thm]) | _ => [thm]) | _ => [thm] end; - in atoms end; + in atoms end + + fun mksimps pairs ctxt = + map_filter (try mk_eq) o mk_atomize ctxt pairs o Variable.gen_all ctxt + + fun unsafe_solver_tac ctxt = + let + val sol_thms = + reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI}:: Simplifier.prems_of ctxt; + fun sol_tac i = + FIRST + [resolve_tac ctxt sol_thms i, + assume_tac ctxt i, + eresolve_tac ctxt [@{thm falseE}, @{thm not_refl_False}] i] ORELSE + (match_tac ctxt [@{thm conjI}] + THEN_ALL_NEW sol_tac) i + in + (fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' sol_tac + end; - fun mksimps pairs (_: simpset) = - map_filter (try mk_eq) o mk_atomize pairs o gen_all - - fun unsafe_solver_tac ss = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [resolve_tac (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ss), - atac, etac @{thm falseE}]; - (*the above may instantiate variables*) val unsafe_solver = mk_solver "PredicateLogic unsafe" unsafe_solver_tac; - fun safe_solver_tac ss = - (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' - FIRST' [match_tac (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ss), - eq_assume_tac, ematch_tac @{thms falseE}]; + (*No premature instantiation of variables during simplification*) + fun safe_solver_tac ctxt = + (fn i => REPEAT_DETERM (match_tac ctxt @{thms simp_impliesI} i)) THEN' + FIRST' [match_tac ctxt (reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI} :: Simplifier.prems_of ctxt), + eq_assume_tac, ematch_tac ctxt [@{thm falseE}, @{thm not_refl_False}]]; val safe_solver = mk_solver "PredicateLogic safe" safe_solver_tac; (* Basic infrastructure for simplification *) - val PL_basic_ss = Simplifier.global_context @{theory} empty_ss - setsubgoaler asm_simp_tac + val PL_basic_ss = + empty_simpset @{context} setSSolver safe_solver setSolver unsafe_solver - setmksimps (mksimps mksimps_pairs) - setmkeqTrue mk_eq_True - setmkcong mk_meta_cong; + |> Simplifier.set_subgoaler asm_simp_tac + |> Simplifier.set_mksimps (mksimps mksimps_pairs) + |> Simplifier.set_mkeqTrue mk_eq_True + |> Simplifier.set_mkcong mk_meta_cong + |> simpset_of (* Combination of classical reasoning and simplifier. *) structure Splitter = Splitter ( - val thy = @{theory} + val context = @{context} val mk_eq = mk_eq val meta_eq_to_iff = @{thm meta_eq_to_iff} val iffD = @{thm iffD2} @@ -152,13 +163,11 @@ struct val contrapos = @{thm contrapos} val contrapos2 = @{thm contrapos2} val notnotD = @{thm notnotD} + val safe_tac = Classical.safe_tac ); val split_tac = Splitter.split_tac; val split_inside_tac = Splitter.split_inside_tac; - (* val split_asm_tac = Splitter.split_asm_tac; *) - val op addsplits = Splitter.addsplits; - val op delsplits = Splitter.delsplits; structure Clasimp = Clasimp ( structure Simplifier = Simplifier @@ -171,6 +180,9 @@ struct ); open Clasimp; + fun unfold_tac ctxt ths = + ALLGOALS (full_simp_tac (clear_simpset (put_simpset PL_basic_ss ctxt) addsimps ths)) + end; (* structure Simpdata *) structure Splitter = Simpdata.Splitter; From 5e4070c2c1845af4cd05cbde89310900dede6c9a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 20 Dec 2020 16:48:15 +0200 Subject: [PATCH 064/167] STY: rm trailing whitespace --- isabelle/CaseExpressions.thy | 20 ++-- isabelle/FixedPoints.thy | 2 +- isabelle/Functions.thy | 38 +++---- isabelle/Integers.thy | 102 +++++++++--------- isabelle/NatArith.thy | 194 +++++++++++++++++------------------ isabelle/NatDivision.thy | 94 ++++++++--------- isabelle/NatOrderings.thy | 48 ++++----- isabelle/Peano.thy | 12 +-- isabelle/PredicateLogic.thy | 50 ++++----- isabelle/README.html | 2 +- isabelle/SMT.thy | 146 +++++++++++++------------- isabelle/SetTheory.thy | 24 ++--- isabelle/Strings.thy | 12 +-- isabelle/Tuples.thy | 74 ++++++------- isabelle/Zenon.thy | 8 +- isabelle/simplifier_setup.ML | 22 ++-- 16 files changed, 424 insertions(+), 424 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 10dd8c6f..c9fc0a31 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -41,7 +41,7 @@ definition Case where "Case(ps, es) \ CHOOSE x : x \ (UNION { CaseArm(ps[i], es[i]) : i \ DOMAIN ps })" definition CaseOther where - "CaseOther(ps, es, oth) \ + "CaseOther(ps, es, oth) \ CHOOSE x : x \ (UNION { CaseArm(ps[i], es[i]) : i \ DOMAIN ps }) \ CaseArm((\i \ DOMAIN ps : \ps[i]), oth)" @@ -73,14 +73,14 @@ parse_ast_translation {* (* get_case_constituents extracts the lists of predicates, terms, and default value from the arms of a case expression. The order of the ASTs is reversed. *) - fun get_case_constituents (Ast.Appl[Ast.Constant "_other", t]) = + fun get_case_constituents (Ast.Appl[Ast.Constant "_other", t]) = (* 1st case: single "OTHER" arm *) ([], [], SOME t) - | get_case_constituents (Ast.Appl[Ast.Constant "_case1", p, t]) = + | get_case_constituents (Ast.Appl[Ast.Constant "_case1", p, t]) = (* 2nd case: a single arm, no "OTHER" branch *) ([p], [t], NONE) - | get_case_constituents (Ast.Appl[Ast.Constant "_case2", - Ast.Appl[Ast.Constant "_case1", p, t], + | get_case_constituents (Ast.Appl[Ast.Constant "_case2", + Ast.Appl[Ast.Constant "_case1", p, t], arms]) = (* 3rd case: one arm, followed by remaining arms *) let val (ps, ts, oth) = get_case_constituents arms @@ -125,8 +125,8 @@ print_ast_translation {* | list_from_tp t = [t] in list_from_tp tp end - (* make_case_arms constructs an AST representing the arms of the - CASE expression. The result is an AST of "type" case_arms. + (* make_case_arms constructs an AST representing the arms of the + CASE expression. The result is an AST of "type" case_arms. The lists of predicates and terms are of equal length, oth is optional. The lists can be empty only if "oth" is present, corresponding to a degenerated expression "CASE OTHER -> e". *) @@ -138,7 +138,7 @@ print_ast_translation {* in (* last arm: check if OTHER defaults *) if oth = NONE then arm - else Ast.Appl[Ast.Constant @{syntax_const "_case2"}, arm, + else Ast.Appl[Ast.Constant @{syntax_const "_case2"}, arm, Ast.Appl[Ast.Constant @{syntax_const "_other"}, the oth]] end | make_case_arms (p::ps) (t::ts) oth = @@ -164,13 +164,13 @@ print_ast_translation {* val trms = list_from_tuple tTuple in (* make sure that tuples are of equal length, otherwise give up *) if length prds = length trms - then Ast.Appl[Ast.Constant @{syntax_const "_case_syntax"}, + then Ast.Appl[Ast.Constant @{syntax_const "_case_syntax"}, make_case_arms prds trms (SOME oth)] else Ast.Appl[Ast.Constant "CaseOther", pTuple, tTuple, oth] end | caseother_tr' _ = raise Match in - [(@{const_syntax "Case"}, case_syntax_tr'), + [(@{const_syntax "Case"}, case_syntax_tr'), (@{const_syntax "CaseOther"}, caseother_tr')] end *} diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index e4a086b3..7ddf3970 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -14,7 +14,7 @@ begin text \ As a test for the encoding of \tlaplus{} set theory, we develop the Knaster-Tarski theorems for least and greatest fixed points in the - subset lattice. Again, the proofs essentially follow Paulson's + subset lattice. Again, the proofs essentially follow Paulson's developments for Isabelle/ZF. \ diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index 922271f0..b68d19a2 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -14,12 +14,12 @@ begin subsection \ Syntax and axioms for functions \ text \ - Functions in \tlaplus{} are not defined (e.g., as sets of pairs), + Functions in \tlaplus{} are not defined (e.g., as sets of pairs), but axiomatized, and in fact, pairs and tuples will be defined as special functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined - as functions over products. We follow the development of functions given + as functions over products. We follow the development of functions given in Section 16.1 of ``Specifying Systems''. \ @@ -50,7 +50,7 @@ text \ Two functions are equal if they have the same domain and agree on all arguments within the domain. \ -axiomatization where fcnEqual[elim!]: +axiomatization where fcnEqual[elim!]: "\isAFcn(f); isAFcn(g); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" @@ -62,7 +62,7 @@ text \ consts FuncSet :: "[c,c] \ c" ("([_ \ _])" 900) \ \ function space \ syntax (ASCII) - FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) + FuncSet :: "[c,c] \ c" ("([_ -> _])" 900) axiomatization where FuncSet: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\x\S : f[x] \ T)" @@ -124,13 +124,13 @@ text \ proofs about equality of functions. \ -lemma fcnEqual2[elim!]: +lemma fcnEqual2[elim!]: "\isAFcn(g); isAFcn(f); DOMAIN f = DOMAIN g; \x \ DOMAIN g : f[x]=g[x]\ \ f = g" by (rule fcnEqual) text \ possibly useful as a simplification rule, but cannot be active by default \ -lemma fcnEqualIff: +lemma fcnEqualIff: assumes "isAFcn(f)" and "isAFcn(g)" shows "(f = g) = (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : f[x] = g[x]))" using assms by auto @@ -141,7 +141,7 @@ lemma [intro!]: by auto lemma [intro!]: - "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; + "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; \y \ DOMAIN g : y \ v \ f[y] = g[y]\ \ [g EXCEPT ![v] = w] = f" "\isAFcn(f); DOMAIN f = DOMAIN g; v \ DOMAIN g \ f[v] = w; @@ -271,7 +271,7 @@ lemma functionInFuncSet: using assms by auto lemma exceptInFuncSet[elim!]: - assumes 1: "f \ [S \ U]" and 2: "U \ T" + assumes 1: "f \ [S \ U]" and 2: "U \ T" and 3: "\v \ S; isAFcn(f); DOMAIN f = S; \x \ S : f[x] \ U\ \ e \ T" shows "[f EXCEPT ![v]=e] \ [S \ T]" (is "?exc \ [S \ T]") proof @@ -288,7 +288,7 @@ text \ \ lemma exceptInFuncSetSame: - assumes "f \ [S \ T]" + assumes "f \ [S \ T]" and "\v \ S; isAFcn(f); DOMAIN f = S; \x \ S : f[x] \ T\ \ e \ T" shows "[f EXCEPT ![v]=e] \ [S \ T]" using assms by auto @@ -340,13 +340,13 @@ lemmas fcnEqual[where g = "g @@ h" for g h, intro!] lemma extendEqualIff [simp]: - "isAFcn(f) \ (f = g @@ h) = - (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ - (\x \ DOMAIN g : f[x] = g[x]) \ + "isAFcn(f) \ (f = g @@ h) = + (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ + (\x \ DOMAIN g : f[x] = g[x]) \ (\x \ DOMAIN h \\ DOMAIN g : f[x] = h[x]))" - "isAFcn(f) \ (g @@ h = f) = - (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ - (\x \ DOMAIN g : g[x] = f[x]) \ + "isAFcn(f) \ (g @@ h = f) = + (DOMAIN f = (DOMAIN g) \ (DOMAIN h) \ + (\x \ DOMAIN g : g[x] = f[x]) \ (\x \ DOMAIN h \\ DOMAIN g : h[x] = f[x]))" by auto @@ -431,7 +431,7 @@ by auto subsubsection \ Injective functions \ -definition InjectiveOn +definition InjectiveOn where "InjectiveOn(f,A) \ \x,y \ A \ DOMAIN f : f[x] = f[y] \ x = y" abbreviation Injective \ \ special case: injective function \ @@ -460,7 +460,7 @@ lemma injectiveOnI: using assms by (auto simp: InjectiveOn_def) lemma injectiveOnD: - assumes "f[x] = f[y]" and "InjectiveOn(f,A)" + assumes "f[x] = f[y]" and "InjectiveOn(f,A)" and "x \ A" and "x \ DOMAIN f" and "y \ A" and "y \ DOMAIN f" shows "x = y" using assms by (auto simp: InjectiveOn_def) @@ -493,7 +493,7 @@ lemma inverseThenInjective: shows "InjectiveOn(f,A)" proof (rule injectiveOnI) fix x y - assume x: "x \ A" "x \ DOMAIN f" + assume x: "x \ A" "x \ DOMAIN f" and y: "y \ A" "y \ DOMAIN f" and eq: "f[x] = f[y]" from x have x1: "x = g[f[x]]" by (rule sym[OF inv]) @@ -673,7 +673,7 @@ proof - by auto finally show ?thesis . qed - + lemma injectionsI: assumes "f \ [S \ T]" and "\x y. \ x \ S; y \ S; f[x] = f[y] \ \ x = y" shows "f \ Injections(S,T)" diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 3fa326e5..3502ce19 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -30,17 +30,17 @@ translations (* lemma eq_imp_negeq: "n = m \ -.n = -.m" by simp *) axiomatization where - neg0 [simp]: "-.0 = 0" + neg0 [simp]: "-.0 = 0" and neg_neg [simp]: "-.-.n = n" and negNotInNat [simp]: "-.(Succ[n]) \ Nat" -lemma negNat_noteq_Nat [simp]: +lemma negNat_noteq_Nat [simp]: "\m \ Nat; n \ Nat\ \ (-. Succ[m] = Succ[n]) = FALSE" proof (rule contradiction) assume "(-. Succ[m] = Succ[n]) \ FALSE" - and m: "m \ Nat" and n: "n \ Nat" + and m: "m \ Nat" and n: "n \ Nat" hence "-. Succ[m] = Succ[n]" by auto hence "-. Succ[m] \ Nat" using n by auto with negNotInNat[of m] show FALSE by simp @@ -58,7 +58,7 @@ qed lemma nat_not_eq_inv: "n \ Nat \ n = 0 \ -.n \ n" using not0_implies_Suc[of n] by auto -lemma minusInj [dest]: +lemma minusInj [dest]: assumes hyp: "-.n = -.m" shows "n = m" proof - @@ -66,8 +66,8 @@ proof - thus ?thesis by simp qed -lemma minusInj_iff [simp]: - "-.x = -.y = (x = y)" +lemma minusInj_iff [simp]: + "-.x = -.y = (x = y)" by auto lemma neg0_imp_0 [simp]: "-.n = 0 = (n = 0)" @@ -77,7 +77,7 @@ proof auto thus "n = 0" by simp qed -lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" +lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" by simp lemma notneg0_imp_not0 [dest]: "-.n \ 0 \ n \ 0" @@ -113,7 +113,7 @@ by auto lemma negNat_exists: "-.n \ Nat \ \k \ Nat: n = -.k" by force -lemma nat_eq_negnat_is_0 [simp]: +lemma nat_eq_negnat_is_0 [simp]: assumes "n \ Nat" shows "(n = -.n) = (n = 0)" using assms by (cases "n", auto) @@ -147,10 +147,10 @@ using assms unfolding Int_def by auto -- {* Integer cases over two parameters *} lemma intCases2: assumes m: "m \ Int" and n: "n \ Int" - and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" - and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" - and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" - and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" + and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" + and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" + and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" + and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" shows "P(m,n)" using m proof (cases "m") assume "m \ Nat" @@ -163,14 +163,14 @@ qed lemma intCases3: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" - and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" - and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" - and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" - and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" - and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" - and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" - and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" + and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" + and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" + and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" + and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" + and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" + and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" + and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" + and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" shows "P(m,n,p)" proof (rule intCases2[OF m n]) fix m n @@ -235,11 +235,11 @@ by(simp add: isPos_def) lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \ n \ 0)" unfolding isPos_def using not0_implies_Suc[of n] by auto -(*lemma isNeg_eq_inNegNat: +(*lemma isNeg_eq_inNegNat: "isNeg(n) = (n \ {-.n : n \ Nat } \\ {0})" unfolding isNeg_def by force*) -lemma isNeg_eq_inNegNat: +lemma isNeg_eq_inNegNat: "isNeg(n) = (n \ {-.n : n \ Nat } \ n \ 0)" unfolding isNeg_def by force @@ -250,7 +250,7 @@ lemma negNotNat_isNat: assumes n: "n \ Int" shows "(-.n \ Nat) = FALSE \ n \ Nat" using n by (cases, auto) -lemma noNatisNeg [simp]: +lemma noNatisNeg [simp]: "n \ Nat \ isNeg(n) = FALSE" -- {* No natural number is negative *} unfolding isNeg_def using negNotInNat by blast @@ -270,11 +270,11 @@ lemma isPos_isNeg_false [simp]: "n \ Int \ isPos(n) \ Int \\ {0} \ \ isPos(n) \ isPos(-.n)" unfolding isPos_eq_inNat1 using intNotNatIsNegNat by auto*) -lemma isPos_neg_isNeg [simp]: +lemma isPos_neg_isNeg [simp]: assumes n: "n \ Int" shows "isPos(-.n) = isNeg(n)" by (auto simp: minus_sym isPos_def isNeg_def) -lemma notIsNeg0_isPos: +lemma notIsNeg0_isPos: assumes n: "n \ Int" shows "\\ isNeg(n); n \ 0\ \ isPos(n)" using n by (cases, auto simp: isPos_eq_inNat1 dest: negNat_isNeg) @@ -317,7 +317,7 @@ by (auto simp: sgn_def) lemma sgn1_imp_pos (*[simp]*): "sgn(n) = 1 \ n \ Nat \ n \ 0" unfolding sgn_def isPos_eq_inNat1 by auto -(*lemma neg_imp_negSuc (*[simp]*): +(*lemma neg_imp_negSuc (*[simp]*): assumes h:"n \ {-.n : n \ Nat } \\ {0}" shows "isNeg(n)" unfolding isNeg_def proof - @@ -328,7 +328,7 @@ proof - with 3 show "\k \ Nat : n = -.(Succ[k])" by auto qed*) -lemma sgnm1_imp_neg: +lemma sgnm1_imp_neg: assumes n:"n \ Int" shows "sgn(n) = -.1 \ isNeg(n)" unfolding sgn_def using intThenPosZeroNeg[OF n] by auto @@ -347,7 +347,7 @@ lemma sgnNat_not1: "\n \ Nat; sgn(n) \ 1\ \ n = 0" using sgnNat_is_0or1[of n] by auto -lemma sgnNat_not_neg [simp]: +lemma sgnNat_not_neg [simp]: "n \ Nat \ sgn(n) = -.1 = FALSE" unfolding sgn_def isPos_eq_inNat1 by auto @@ -378,7 +378,7 @@ by (auto simp: sgn_eq_neg1_is_not_nat) lemma sgn_neg_eq_1_false: "\sgn(-.m) = 1; m \ Nat\ \ P" unfolding sgn_def by auto -lemma sgn_minus [simp]: +lemma sgn_minus [simp]: assumes n: "n \ Int" shows "sgn(-.n) = -.sgn(n)" unfolding sgn_def using n by (cases, auto) @@ -400,22 +400,22 @@ lemma absn0 [simp]: "abs(-.0) = 0" by simp **) -lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" +lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" unfolding abs_def by (auto dest: sgnNat_not1) -lemma abs_neg [simp]: +lemma abs_neg [simp]: assumes n: "n \ Int" shows "abs(-.n) = abs(n)" unfolding abs_def using n by (auto dest: sgnNat_not1) -subsection {* Orders on integers *} +subsection {* Orders on integers *} -text {* +text {* We distinguish four cases, depending on the arguments being in Nat or negative. *} -lemmas int_leq_pp_def = nat_leq_def +lemmas int_leq_pp_def = nat_leq_def -- {* 'positive-positive' case, ie: both arguments are naturals *} axiomatization where @@ -427,11 +427,11 @@ and (* theorems int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) -lemma int_boolify_leq [simp]: +lemma int_boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" by(rule intCases2[of a b], simp_all) -lemma int_leq_isBool [intro!,simp]: +lemma int_leq_isBool [intro!,simp]: "\a \ Int; b \ Int\ \ isBool(a \ b)" unfolding isBool_def by auto @@ -450,7 +450,7 @@ by(rule intCases2[of m n], simp_all) subsection {* Addition of integers *} -text {* +text {* Again, we distinguish four cases in the definition of @{text "a + b"}, according to each argument being positive or negative. *} @@ -483,14 +483,14 @@ by (auto simp: int_add_pn_def dest: nat_leq_antisym) text {* Closure *} -lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" +lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" by (rule intCases2[of m n], auto simp: int_add_def) text {* Neutral element *} -lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" +lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" by(rule intCases, auto simp add: int_add_np_def) -lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" +lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" by(rule intCases, auto simp add: int_add_pn_def) text {* Additive inverse element *} @@ -509,7 +509,7 @@ by (rule intCases, auto simp: int_add_def) text {* Commutativity *} -lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" +lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" by(simp add: int_add_def) lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" @@ -548,7 +548,7 @@ by(rule adiff_adiff_left_nat[OF m n p]) declare leq_neq_iff_less [simplified,simp] -lemma int_add_assoc1: +lemma int_add_assoc1: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + -.p) = (m + n) + -.p" apply(rule nat_leq_cases[OF p n]) @@ -601,7 +601,7 @@ apply(rule nat_leq_cases[of "n + p" m]) using adiff_add_assoc2[OF _ p n m, symmetric] apply (simp add: adiff_is_0_eq') apply(rule nat_leq_cases[OF n m], simp_all) apply(rule nat_leq_cases[of p "m -- n"], simp_all) - using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp + using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp using adiff_add_assoc3[symmetric] add_commute_nat[OF n p] apply simp using adiff_add_assoc2[symmetric] add_commute_nat[OF n p] apply simp done @@ -667,7 +667,7 @@ lemma int_add_assoc6: shows "-.m + (-.n + p) = -.(m + n) + p" using assms add_commute_int[of "-.n" p] - add_commute_int[of "-.m" "p + -.n"] + add_commute_int[of "-.m" "p + -.n"] add_commute_int[of "-.(m + n)" p] apply simp apply(rule nat_leq_cases[OF n p], simp_all) apply(rule nat_leq_cases[of m "p -- n"], simp+) @@ -687,11 +687,11 @@ apply(rule nat_leq_cases[OF n p], simp_all) apply(rule adiff_add_assoc[symmetric], simp+) done -lemma add_assoc_int: +lemma add_assoc_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m + (n + p) = (m + n) + p" -using m n p -by (rule intCases3, +using m n p +by (rule intCases3, auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 int_add_assoc4 int_add_assoc5 int_add_assoc6) @@ -785,10 +785,10 @@ apply(rule nat_leq_cases[OF p n]) apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) done -lemma distrib_left_int: +lemma distrib_left_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m * (n + p) = (m * n + m * p)" -apply(rule intCases3[OF m n p], +apply(rule intCases3[OF m n p], simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) apply(rule add_mult_distrib_left_nat, assumption+) apply(rule ppn_distrib_left_nat, assumption+) @@ -828,7 +828,7 @@ done lemma distrib_right_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "(m + n) * p = (m * p + n * p)" -apply(rule intCases3[OF m n p], +apply(rule intCases3[OF m n p], simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) apply(rule add_mult_distrib_right_nat, assumption+) apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) @@ -866,7 +866,7 @@ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" lemma diffIsInt [simp]: -- {* Closure *} - "\m \ Int; n \ Int\ \ m - n \ Int" + "\m \ Int; n \ Int\ \ m - n \ Int" by (simp add: int_diff_def) lemma diff_neg_is_add [simp]: "\m \ Int; n \ Int\ \ m - -.n = m + n" diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index fc2618f9..6b17dc37 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -20,7 +20,7 @@ structure AlgebraSimps = setup AlgebraSimps.setup -text{* +text{* The rewrites accumulated in @{text algebra_simps} deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and @@ -46,17 +46,17 @@ lemma addnatType: assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" using assms unfolding addnat_def by (rule bprimrecType_nat, auto) -lemma addIsNat [intro!,simp]: +lemma addIsNat [intro!,simp]: assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" unfolding nat_add_def[OF assms] using assms addnatType by blast text {* Base case and Inductive case *} -lemma addnat_0: +lemma addnat_0: assumes "m \ Nat" shows "addnat(m)[0] = m" using assms unfolding addnat_def by (rule primrec_natE, auto) -lemma add_0_nat [simp]: +lemma add_0_nat [simp]: assumes "m \ Nat" shows "m + 0 = m" by (simp add: nat_add_def[OF assms] addnat_0[OF assms]) @@ -71,24 +71,24 @@ next with n show "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" by blast qed (auto) -lemma add_Succ_nat [simp]: +lemma add_Succ_nat [simp]: assumes "m \ Nat" and "n \ Nat" shows "m + Succ[n] = Succ[m + n]" using assms by (simp add: nat_add_def addnat_Succ[OF assms]) -lemma add_0_left_nat [simp]: - assumes n: "n \ Nat" +lemma add_0_left_nat [simp]: + assumes n: "n \ Nat" shows "0 + n = n" using n by (induct, auto) -lemma add_Succ_left_nat [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" +lemma add_Succ_left_nat [simp]: + assumes n: "n \ Nat" and m: "m \ Nat" shows "Succ[m] + n = Succ[m + n]" using n apply induct using m by auto -lemma add_Succ_shift_nat: - assumes n: "n \ Nat" and m: "m \ Nat" +lemma add_Succ_shift_nat: + assumes n: "n \ Nat" and m: "m \ Nat" shows "Succ[m] + n = m + Succ[n]" using assms by simp @@ -102,7 +102,7 @@ using assms by auto text {* Associativity *} -lemma add_assoc_nat [algebra_simps]: +lemma add_assoc_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + p) = (m + n) + p" using assms by (induct, simp_all) @@ -120,8 +120,8 @@ lemma add_right_cancel_nat [simp]: using assms by (induct, simp_all) -lemma add_left_commute_nat [algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" +lemma add_left_commute_nat [algebra_simps]: + assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a + (b + c) = b + (a + c)" using assms by(simp only: add_assoc_nat add_commute_nat) @@ -172,7 +172,7 @@ lemma multnatType: unfolding multnat_def by (rule bprimrecType_nat, auto simp: assms) lemma multIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" + assumes m: "m \ Nat" and n: "n \ Nat" shows "m * n \ Nat" unfolding nat_mult_def[OF assms] using assms multnatType by blast @@ -198,17 +198,17 @@ next qed (auto simp: m) lemma mult_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" + assumes "m \ Nat" and "n \ Nat" shows "m * Succ[n] = m * n + m" using assms by (simp add: nat_mult_def multnat_Succ[OF assms]) lemma mult_0_left_nat [simp]: - assumes n: "n \ Nat" + assumes n: "n \ Nat" shows "0 * n = 0" using n by (induct, simp_all) -lemma mult_Succ_left_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" +lemma mult_Succ_left_nat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] * n = m * n + n" using n apply induct using m by (simp_all add: add_ac_nat) @@ -265,7 +265,7 @@ lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp text {* Associativity *} -lemma mult_assoc_nat[algebra_simps]: +lemma mult_assoc_nat[algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m * (n * p) = (m * n) * p" using m apply induct @@ -273,19 +273,19 @@ using assms by (auto simp add: add_mult_distrib_right_nat) text {* Reasoning about @{text "m * n = 0"}, etc. *} -lemma mult_is_0_nat [simp]: +lemma mult_is_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = 0) = (m = 0 \ n = 0)" using m apply induct using n by auto -lemma mult_eq_1_iff_nat [simp]: +lemma mult_eq_1_iff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = 1) = (m = 1 \ n = 1)" using m apply induct using n by (induct, auto)+ -lemma one_eq_mult_iff_nat [simp]: +lemma one_eq_mult_iff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 = m * n) = (m = 1 \ n = 1)" proof - @@ -296,7 +296,7 @@ qed text {* Cancellation rules *} -lemma mult_cancel1_nat [simp]: +lemma mult_cancel1_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k * m = k * n) = (m = n \ (k = 0))" proof - @@ -329,18 +329,18 @@ proof - with k m n show ?thesis by auto qed -lemma mult_cancel2_nat [simp]: +lemma mult_cancel2_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k = n * k) = (m = n \ k = 0)" using assms by (simp add: mult_commute_nat) -lemma Suc_mult_cancel1_nat: +lemma Suc_mult_cancel1_nat: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(Succ[k] * m = Succ[k] * n) = (m = n)" using k m n mult_cancel1_nat[of "Succ[k]" m n] by simp -lemma mult_left_commute_nat[algebra_simps]: +lemma mult_left_commute_nat[algebra_simps]: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * (b * c) = b * (a * c)" using assms by(simp only: mult_commute_nat mult_assoc_nat) @@ -360,7 +360,7 @@ by (simp add: Pred_def) lemma Pred_Succ_nat [simp]: "n \ Nat \ Pred[Succ[n]] = n" unfolding Pred_def by (auto intro: bChooseI2) -lemma Succ_Pred_nat [simp]: +lemma Succ_Pred_nat [simp]: assumes "n \ Nat" and "n \ 0" shows "Succ[Pred[n]] = n" using assms unfolding Pred_def by (cases "n", auto intro: bChooseI2) @@ -372,7 +372,7 @@ using assms by (cases "n", auto) subsection {* Difference of natural numbers *} text {* - We define a form of difference @{text "--"} of natural numbers that cuts off + We define a form of difference @{text "--"} of natural numbers that cuts off at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes called ``arithmetic difference''. *} @@ -395,11 +395,11 @@ unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast text {* Neutral element and Inductive step *} -lemma adiffnat_0: +lemma adiffnat_0: assumes "m \ Nat" shows "adiffnat(m)[0] = m" using assms unfolding adiffnat_def by (rule primrec_natE, auto) -lemma adiff_0_nat [simp]: +lemma adiff_0_nat [simp]: assumes "m \ Nat" shows "m -- 0 = m" by (simp add: nat_adiff_def[OF assms] adiffnat_0[OF assms]) @@ -426,40 +426,40 @@ using n apply induct by (simp_all add: adiff_Succ_nat) text {* Reasoning about @{text "m -- m = 0"}, etc. *} -lemma adiff_Succ_Succ_nat [simp]: +lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] -- Succ[n] = m -- n" using n apply induct using assms by (auto simp add: adiff_Succ_nat) -lemma adiff_self_eq_0_nat [simp]: +lemma adiff_self_eq_0_nat [simp]: assumes m: "m \ Nat" shows "m -- m = 0" using m apply induct by auto text {* Associativity *} -lemma adiff_adiff_left_nat: +lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "(i -- j) -- k = i -- (j + k)" using i j apply (rule diffInduct) using k by auto -lemma Succ_adiff_adiff_nat [simp]: +lemma Succ_adiff_adiff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" using assms by (simp add: adiff_adiff_left_nat) text {* Commutativity *} -lemma adiff_commute_nat: +lemma adiff_commute_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i -- j -- k = i -- k -- j" using assms by (simp add: adiff_adiff_left_nat add_commute_nat) text {* Cancellation rules *} -lemma adiff_add_inverse_nat [simp]: +lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(n + m) -- n = m" using n apply induct @@ -489,20 +489,20 @@ using assms by simp_all text {* Difference distributes over multiplication *} -lemma adiff_mult_distrib_nat [algebra_simps]: +lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m -- n) * k = (m * k) -- (n * k)" using m n apply(rule diffInduct) using k by simp_all -lemma adiff_mult_distrib2_nat [algebra_simps]: +lemma adiff_mult_distrib2_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) -- {* NOT added as rewrites, since sometimes they are used from right-to-left *} lemmas nat_distrib = - add_mult_distrib_right_nat add_mult_distrib_left_nat + add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat @@ -520,25 +520,25 @@ lemma nat_add_left_cancel_leq [simp]: shows "(k + m \ k + n) = (m \ n)" using assms by (induct k) simp_all -lemma nat_add_left_cancel_less [simp]: +lemma nat_add_left_cancel_less [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k + m < k + n) = (m < n)" using k apply induct using assms by simp_all -lemma nat_add_right_cancel_less [simp]: +lemma nat_add_right_cancel_less [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m + k < n + k) = (m < n)" using k apply induct using assms by simp_all -lemma nat_add_right_cancel_leq [simp]: +lemma nat_add_right_cancel_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m + k \ n + k) = (m \ n)" using k apply induct using assms by simp_all -lemma add_gr_0 [simp]: +lemma add_gr_0 [simp]: assumes "n \ Nat" and "m \ Nat" shows "(m + n > 0) = (m > 0 \ n > 0)" using assms by (auto dest: nat_gt0_implies_Succ nat_not_lessD) @@ -576,7 +576,7 @@ subsubsection {* (Partially) Ordered Groups *} -- {* The two following lemmas are just ``one half'' of @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} proved above. *} -lemma add_leq_left_mono: +lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" using assms by simp @@ -591,7 +591,7 @@ lemma add_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c \ d \ a + c \ b + d" using assms - add_leq_right_mono[OF a b c] + add_leq_right_mono[OF a b c] add_leq_left_mono[OF c d b] nat_leq_trans[of "a + c" "b + c" "b + d"] by simp @@ -602,7 +602,7 @@ lemma add_less_left_mono: shows "a < b \ c + a < c + b" using assms by simp -lemma add_less_right_mono: +lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" using assms by simp @@ -612,7 +612,7 @@ lemma add_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c < d \ a + c < b + d" using assms - add_less_right_mono[OF a b c] + add_less_right_mono[OF a b c] add_less_left_mono[OF c d b] nat_less_trans[of "a + c" "b + c" "b + d"] by simp @@ -621,7 +621,7 @@ lemma add_less_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c \ d \ a + c < b + d" using assms - add_less_right_mono[OF a b c] + add_less_right_mono[OF a b c] add_leq_left_mono[OF c d b] nat_less_leq_trans[of "a + c" "b + c" "b + d"] by blast @@ -630,58 +630,58 @@ lemma add_leq_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c < d \ a + c < b + d" using assms - add_leq_right_mono[OF a b c] + add_leq_right_mono[OF a b c] add_less_left_mono[OF c d b] nat_leq_less_trans[of "a + c" "b + c" "b + d"] by blast lemma leq_add1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" + assumes m: "m \ Nat" and n: "n \ Nat" shows "n \ n + m" using assms add_leq_left_mono[of 0 m n] by simp lemma leq_add2 [simp]: - assumes "m \ Nat" and "n \ Nat" + assumes "m \ Nat" and "n \ Nat" shows "n \ m + n" using assms add_leq_right_mono [of 0 m n] by simp lemma less_add_Succ1: - assumes "i \ Nat" and "m \ Nat" + assumes "i \ Nat" and "m \ Nat" shows "i < Succ[i + m]" using assms by simp -lemma less_add_Succ2: - assumes "i \ Nat" and "m \ Nat" +lemma less_add_Succ2: + assumes "i \ Nat" and "m \ Nat" shows "i < Succ[m + i]" using assms by simp -lemma less_iff_Succ_add: - assumes "m \ Nat" and "n \ Nat" +lemma less_iff_Succ_add: + assumes "m \ Nat" and "n \ Nat" shows "(m < n) = (\k \ Nat: n = Succ[m + k])" using assms by (auto intro!: less_imp_Succ_add) -lemma trans_leq_add1: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_leq_add1: + assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i \ j + m" using assms by (auto elim: nat_leq_trans) -lemma trans_leq_add2: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_leq_add2: + assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i \ m + j" using assms by (auto elim: nat_leq_trans) -lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_less_add1: + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < j + m" using assms by (auto elim: nat_less_leq_trans) -lemma trans_less_add2: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" +lemma trans_less_add2: + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" using assms by (auto elim: nat_less_leq_trans) -lemma add_lessD1: - assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" +lemma add_lessD1: + assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "i < k" using assms by (intro nat_leq_less_trans[of "i" "i+j" "k"], simp+) @@ -695,17 +695,17 @@ lemma not_add_less2 [simp]: shows "(j + i < i) = FALSE" using assms by (simp add: add_commute_nat) -lemma add_leqD1: +lemma add_leqD1: assumes "m + k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "m \ n" using assms by (intro nat_leq_trans[of "m" "m+k" "n"], simp+) -lemma add_leqD2: +lemma add_leqD2: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k \ n" using assms by (intro nat_leq_trans[of "k" "m+k" "n"], simp+) -lemma add_leqE: +lemma add_leqE: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "(m \ n \ k \ n \ R) \ R" using assms by (blast dest: add_leqD1 add_leqD2) @@ -720,13 +720,13 @@ subsubsection {* More results about arithmetic difference *} text {* Addition is the inverse of subtraction: if @{term "n \ m"} then @{term "n + (m -- n) = m"}. *} -lemma add_adiff_inverse: +lemma add_adiff_inverse: assumes "m \ Nat" and "n \ Nat" shows "\(m < n) \ n + (m -- n) = m" apply (induct m n rule: diffInduct) using assms by simp_all -lemma le_add_adiff_inverse [simp]: +lemma le_add_adiff_inverse [simp]: assumes "m \ Nat" and "n \ Nat" shows "n \ m \ n + (m -- n) = m" using assms by (simp add: add_adiff_inverse nat_not_order_simps) @@ -736,7 +736,7 @@ lemma le_add_adiff_inverse2 [simp]: shows "n \ m \ (m -- n) + n = m" using assms by (simp add: add_commute_nat) -lemma Succ_adiff_leq: +lemma Succ_adiff_leq: assumes "m \ Nat" and "n \ Nat" shows "n \ m \ Succ[m] -- n = Succ[m -- n]" apply (induct m n rule: diffInduct) @@ -748,13 +748,13 @@ lemma adiff_less_Succ: apply (induct m n rule: diffInduct) using assms by (auto simp: nat_less_Succ) -lemma adiff_leq_self [simp]: +lemma adiff_leq_self [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ m" apply (induct m n rule: diffInduct) using assms by simp_all -lemma leq_iff_add: +lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") proof - @@ -768,7 +768,7 @@ proof - from 1 2 assms show ?thesis by blast qed -lemma less_imp_adiff_less: +lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" proof - @@ -776,18 +776,18 @@ proof - with j k jk show ?thesis by (auto elim: nat_leq_less_trans) qed -lemma adiff_Succ_less (*[simp]*): +lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" shows "0 < n \ n -- Succ[i] < n" using n apply cases using i by auto -lemma adiff_add_assoc: +lemma adiff_add_assoc: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(i + j) -- k = i + (j -- k)" using assms by (induct j k rule: diffInduct, simp+) -lemma adiff_add_assoc2: +lemma adiff_add_assoc2: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(j + i) -- k = (j -- k) + i" using assms by (simp add: add_commute_nat adiff_add_assoc) @@ -802,11 +802,11 @@ lemma adiff_add_assoc4: and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "p -- (m -- n) = (p -- m) + n" using assms - adiff_add_assoc2[OF _ n p m, symmetric] + adiff_add_assoc2[OF _ n p m, symmetric] adiff_add_assoc3[OF _ _ p n m] apply simp using trans_leq_add1[OF _ m p n] by simp -lemma le_imp_adiff_is_add: +lemma le_imp_adiff_is_add: assumes "i \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "(j -- i = k) = (j = k + i)" using assms by auto @@ -821,7 +821,7 @@ lemma adiff_is_0_eq' (*[simp]*): shows "m -- n = 0" using assms by simp -lemma zero_less_adiff [simp]: +lemma zero_less_adiff [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(0 < n -- m) = (m < n)" by (induct m n rule: diffInduct, simp_all add: assms) @@ -859,7 +859,7 @@ lemma leq_adiff_right_false [simp]: shows "m \ m -- n = FALSE" using assms by (simp add: leq_adiff_right_add_left[OF _ m m n]) -lemma leq_adiff_right_imp_0: +lemma leq_adiff_right_imp_0: assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" shows "p = 0" using p h apply (induct) @@ -870,13 +870,13 @@ subsubsection {* Monotonicity of Multiplication *} (* from HOL/Ring_and_Field.thy *) -lemma mult_leq_left_mono: +lemma mult_leq_left_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "c * a \ c * b" using c apply induct using 1 a b by (simp_all add: add_leq_mono) -lemma mult_leq_right_mono: +lemma mult_leq_right_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * c \ b * c" using c apply induct @@ -884,8 +884,8 @@ using 1 a b by (simp_all add: add_leq_mono add_commute_nat) text {* @{text "\"} monotonicity, BOTH arguments *} -lemma mult_leq_mono: - assumes 1: "i \ j" "k \ l" +lemma mult_leq_mono: + assumes 1: "i \ j" "k \ l" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" and l: "l \ Nat" shows "i * k \ j * l" using assms @@ -896,7 +896,7 @@ by simp text {* strict, in 1st argument *} -lemma mult_less_left_mono: +lemma mult_less_left_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "k * i < k * j" using 1 @@ -907,7 +907,7 @@ proof (auto simp: nat_gt0_iff_Succ[OF k]) by (induct m, simp_all add: add_less_mono) qed -lemma mult_less_right_mono: +lemma mult_less_right_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i * k < j * k" using 1 @@ -918,7 +918,7 @@ proof (auto simp: nat_gt0_iff_Succ[OF k]) by (induct m, simp_all add: add_less_mono) qed -lemma nat_0_less_mult_iff [simp]: +lemma nat_0_less_mult_iff [simp]: assumes i: "i \ Nat" and j: "j \ Nat" shows "(0 < i * j) = (0 < i \ 0 < j)" using i apply induct @@ -931,7 +931,7 @@ lemma one_leq_mult_iff (*[simp]*): shows "(1 \ m * n) = (1 \ m \ 1 \ n)" using assms by simp -lemma mult_less_cancel_left [simp]: +lemma mult_less_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m < k * n) = (0 < k \ m < n)" proof (auto intro!: mult_less_left_mono[OF _ _ m n k]) @@ -948,7 +948,7 @@ next qed qed -lemma mult_less_cancel_right [simp]: +lemma mult_less_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k < n * k) = (0 < k \ m < n)" proof (auto intro!: mult_less_right_mono[OF _ _ m n k]) @@ -975,7 +975,7 @@ lemma mult_less_self_right [dest]: shows "k=0" using k assms by (cases, auto) -lemma mult_leq_cancel_left [simp]: +lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) @@ -989,7 +989,7 @@ using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) qed qed -lemma mult_leq_cancel_right [simp]: +lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) @@ -1003,17 +1003,17 @@ using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) qed qed -lemma Suc_mult_less_cancel1: +lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" using assms by (simp del: mult_Succ_left_nat) -lemma Suc_mult_leq_cancel1: +lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m \ Succ[k] * n) = (m \ n)" using assms by (simp del: mult_Succ_left_nat) -lemma nat_leq_square: +lemma nat_leq_square: assumes m: "m \ Nat" shows "m \ m * m" using m by (cases, auto) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 81591110..abef1aed 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -32,18 +32,18 @@ lemma [intro!]: "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" by auto -lemma dvdI [intro]: +lemma dvdI [intro]: assumes a: "a \ Nat" and b: "b \ Nat" and k: "k \ Nat" shows "a = b * k \ b dvd a" unfolding dvd_def[OF a b] using k by blast -lemma dvdE [elim]: +lemma dvdE [elim]: assumes "b dvd a" and "a \ Nat" and "b \ Nat" shows "(\k. \k \ Nat; a = b * k\ \ P) \ P" using assms by (auto simp add: dvd_def) -lemma dvd_refl [iff]: - assumes a: "a \ Nat" +lemma dvd_refl [iff]: + assumes a: "a \ Nat" shows "a dvd a" proof - from a have "a = a*1" by simp @@ -64,21 +64,21 @@ proof - qed lemma dvd_0_left_iff [simp]: - assumes "a \ Nat" + assumes "a \ Nat" shows "(0 dvd a) = (a = 0)" using assms by force -lemma dvd_0_right [iff]: +lemma dvd_0_right [iff]: assumes a: "a \ Nat" shows "a dvd 0" using assms by force -lemma one_dvd [iff]: - assumes a: "a \ Nat" +lemma one_dvd [iff]: + assumes a: "a \ Nat" shows "1 dvd a" using assms by force lemma dvd_mult (*[simp]*): - assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a dvd (b * c)" proof - from dvd a c obtain k where k: "k \ Nat" and "c = a*k" by blast @@ -87,16 +87,16 @@ proof - qed lemma dvd_mult2 (*[simp]*): - assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a dvd (b * c)" using mult_commute_nat[OF b c] dvd_mult[OF dvd a c b] by simp -lemma dvd_triv_right [iff]: +lemma dvd_triv_right [iff]: assumes a: "a \ Nat" and b: "b \ Nat" shows "a dvd (b * a)" using assms by (intro dvd_mult, simp+) -lemma dvd_triv_left [iff]: +lemma dvd_triv_left [iff]: assumes a: "a \ Nat" and b: "b \ Nat" shows "a dvd a * b" using assms by (intro dvd_mult2, simp+) @@ -116,17 +116,17 @@ proof - with a c b' d' show ?thesis by blast qed -lemma dvd_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" +lemma dvd_mult_left: + assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a * b dvd c" shows "a dvd c" proof - - from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" + from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" by (auto simp add: mult_assoc_nat) with a b c show ?thesis by blast qed -lemma dvd_mult_right: +lemma dvd_mult_right: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a*b dvd c" shows "b dvd c" proof - @@ -134,14 +134,14 @@ proof - with b a c show ?thesis by (rule dvd_mult_left) qed -lemma dvd_0_left: +lemma dvd_0_left: assumes "a \ Nat" shows "0 dvd a \ a = 0" using assms by simp lemma dvd_add [iff]: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "a dvd c" + and 1: "a dvd b" and 2: "a dvd c" shows "a dvd (b + c)" proof - from a b 1 obtain b' where b': "b' \ Nat" "b = a * b'" by blast @@ -165,7 +165,7 @@ text {* *} definition divmod_rel where - "divmod_rel(m,n,q,r) \ m = q * n + r + "divmod_rel(m,n,q,r) \ m = q * n + r \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" text {* @{const divmod_rel} is total if $n$ is non-zero. *} @@ -183,7 +183,7 @@ proof - fix m' assume m': "m' \ Nat" and ih: "\q, r \ Nat : m' = q*n + r \ r < n" from ih obtain q' r' - where h1: "m' = q' * n + r'" and h2: "r' < n" + where h1: "m' = q' * n + r'" and h2: "r' < n" and q': "q' \ Nat" and r': "r' \ Nat" by blast show "\q, r \ Nat : Succ[m'] = q * n + r \ r < n" proof (cases "Succ[r'] < n") @@ -204,7 +204,7 @@ qed text {* @{const divmod_rel} has unique solutions in the natural numbers. *} lemma divmod_rel_unique_div: assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" - and m: "m \ Nat" and n: "n \ Nat" + and m: "m \ Nat" and n: "n \ Nat" and q: "q \ Nat" and r: "r \ Nat" and q': "q' \ Nat" and r': "r' \ Nat" shows "q = q'" proof - @@ -270,7 +270,7 @@ lemma divmodNat_divmod_rel [rule_format]: unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]], auto) lemma divmodNat_unique: - assumes h: "divmod_rel(m,n,q,r)" + assumes h: "divmod_rel(m,n,q,r)" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" shows "divmodNat(m,n) = \q,r\" @@ -297,7 +297,7 @@ lemma divIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m div n \ Nat" using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) - + lemma modIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n \ Nat" @@ -431,23 +431,23 @@ lemma mod_div_nat_equality2 [simp]: shows "n * (m div n) + m mod n = m" using assms mult_commute_nat[of "n" "m div n"] by simp -lemma mod_div_nat_equality3 [simp]: +lemma mod_div_nat_equality3 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" shows "m mod n + (m div n) * n = m" using assms add_commute_nat[of "m mod n"] by simp -lemma mod_div_nat_equality4 [simp]: +lemma mod_div_nat_equality4 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" shows "m mod n + n * (m div n) = m" using assms mult_commute_nat[of "n" "m div n"] by simp (** immediate consequence of above -lemma div_mod_nat_equality: +lemma div_mod_nat_equality: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "((m div n) * n + m mod n) + k = m + k" using assms by simp -lemma div_mod_nat_equality2: +lemma div_mod_nat_equality2: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * (m div n) + m mod n) + k = m + k" using assms by simp @@ -487,12 +487,12 @@ lemma div_nat_0: shows "0 div n = 0" using assms by simp -lemma mod_0: +lemma mod_0: assumes "n \ Nat" and "0 < n" shows "0 mod n = 0" using assms by simp -lemma mod_nat_mult_self1 [simp]: +lemma mod_nat_mult_self1 [simp]: assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "(q + m * n) mod n = q mod n" proof - @@ -511,57 +511,57 @@ proof - with assms show ?thesis by (simp del: mod_div_nat_equality) qed -lemma mod_nat_mult_self2 [simp]: +lemma mod_nat_mult_self2 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(q + n * m) mod n = q mod n" using assms by (simp add: mult_commute_nat) -lemma mod_nat_mult_self3 [simp]: +lemma mod_nat_mult_self3 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n + q) mod n = q mod n" using assms by (simp add: add_commute_nat) -lemma mod_nat_mult_self4 [simp]: +lemma mod_nat_mult_self4 [simp]: assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m + q) mod n = q mod n" using assms by (simp add: add_commute_nat) -lemma div_nat_mult_self1_is_id [simp]: +lemma div_nat_mult_self1_is_id [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n) div n = m" using assms div_nat_mult_self1 [of 0 m n] by simp -lemma div_nat_mult_self2_is_id [simp]: +lemma div_nat_mult_self2_is_id [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m) div n = m" using assms div_nat_mult_self2[of 0 n m] by simp -lemma mod_nat_mult_self1_is_0 [simp]: +lemma mod_nat_mult_self1_is_0 [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m * n) mod n = 0" using assms mod_nat_mult_self1 [of 0 m n] by simp -lemma mod_nat_mult_self2_is_0 [simp]: +lemma mod_nat_mult_self2_is_0 [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(n * m) mod n = 0" using assms mod_nat_mult_self2 [of 0 m n] by simp -lemma div_nat_by_1 [simp]: +lemma div_nat_by_1 [simp]: assumes "m \ Nat" shows "m div 1 = m" using assms div_nat_mult_self1_is_id [of m 1] by simp -lemma mod_nat_by_1 [simp]: +lemma mod_nat_by_1 [simp]: assumes "m \ Nat" shows "m mod 1 = 0" using assms mod_nat_mult_self1_is_0[of m 1] by simp -lemma mod_nat_self [simp]: +lemma mod_nat_self [simp]: assumes "n \ Nat" and "0 < n" shows "n mod n = 0" using assms mod_nat_mult_self1_is_0[of 1] by simp -lemma div_nat_self [simp]: +lemma div_nat_self [simp]: assumes "n \ Nat" and "0 < n" shows "n div n = 1" using assms div_nat_mult_self1_is_id [of 1 n] by simp @@ -610,7 +610,7 @@ proof - from 1 2 assms show ?thesis by blast qed -lemma mod_div_nat_trivial [simp]: +lemma mod_div_nat_trivial [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m mod n) div n = 0" proof - @@ -622,7 +622,7 @@ proof - using assms by simp qed -lemma mod_mod_nat_trivial [simp]: +lemma mod_mod_nat_trivial [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "(m mod n) mod n = m mod n" proof - @@ -632,12 +632,12 @@ proof - finally show ?thesis . qed -lemma dvd_nat_imp_mod_0: +lemma dvd_nat_imp_mod_0: assumes "n dvd m" and "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n = 0" using assms by (simp add: dvd_nat_eq_mod_eq_0) -lemma dvd_nat_div_mult_self: +lemma dvd_nat_div_mult_self: assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "(m div n) * n = m" using assms @@ -645,7 +645,7 @@ using assms mod_div_nat_equality[OF m n pos] by simp -lemma dvd_nat_div_mult: +lemma dvd_nat_div_mult: assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and k: "k \ Nat" shows "(m div n) * k = (m * k) div n" @@ -680,7 +680,7 @@ next with 3 4 show "(b div a) dvd (c div a)" by simp qed (auto simp: assms) -lemma dvd_mod_nat_imp_dvd: +lemma dvd_mod_nat_imp_dvd: assumes 1: "k dvd (m mod n)" and 2: "k dvd n" and k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "k dvd m" @@ -690,7 +690,7 @@ proof - with m n pos show ?thesis by simp qed -(*** TODO +(*** TODO text {* Addition respects modular equivalence. *} text {* Multiplication respects modular equivalence. *} diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f558c286..5009aa4e 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -13,7 +13,7 @@ imports Peano begin text \ - Using the sets @{text upto} we can now define the standard ordering on + Using the sets @{text upto} we can now define the standard ordering on natural numbers. The constant @{text "\"} is defined over the naturals by the axiom (conditional definition) @{text "nat_leq_def"} below; it should be defined over other domains as appropriate later on. @@ -119,12 +119,12 @@ lemma nat_leq_limit: shows "m=n" using assms by (auto simp: nat_leq_def intro: uptoLimit) -lemma nat_leq_trans [trans]: +lemma nat_leq_trans [trans]: assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" shows "k \ n" using assms by (auto simp: nat_leq_def elim: uptoTrans) -lemma nat_leq_antisym: +lemma nat_leq_antisym: assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) shows "m = n" using assms by (auto simp add: nat_leq_def elim: uptoAntisym) @@ -139,7 +139,7 @@ lemma nat_Succ_leqI: shows "Succ[m] \ n" using assms by (auto dest: nat_leq_limit) -lemma nat_Succ_leqD [elim]: +lemma nat_Succ_leqD [elim]: assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" shows "m \ n" proof - @@ -253,7 +253,7 @@ lemma nat_Succ_less_SuccE: shows "P" using assms by simp -lemma nat_not_less0 [simp]: +lemma nat_not_less0 [simp]: assumes "n \ Nat" shows "(n < 0) = FALSE" using assms by auto @@ -263,12 +263,12 @@ lemma nat_less0E (*[elim!]*): shows "P" using assms by simp -lemma nat_less_SuccI: +lemma nat_less_SuccI: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "m < Succ[n]" using assms by auto -lemma nat_Succ_lessD: +lemma nat_Succ_lessD: assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" shows "m < n" using assms by auto @@ -300,19 +300,19 @@ lemma nat_leq_less_trans [trans]: shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) -lemma nat_less_leq_trans [trans]: +lemma nat_less_leq_trans [trans]: assumes "k < m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) text \ Asymmetry. \ -lemma nat_less_not_sym: +lemma nat_less_not_sym: assumes "m < n" and "m \ Nat" and "n \ Nat" shows "(n < m) = FALSE" using assms by auto - -lemma nat_less_asym: + +lemma nat_less_asym: assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" shows "P" proof (rule contradiction) @@ -322,12 +322,12 @@ qed text \ Linearity (totality). \ -lemma nat_less_linear: +lemma nat_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" shows "m < n \ m = n \ n < m" unfolding less_def using nat_leq_linear[OF m n] by blast -lemma nat_leq_less_linear: +lemma nat_leq_less_linear: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n \ n < m" using assms nat_less_linear[OF m n] by (auto simp: less_def) @@ -347,7 +347,7 @@ lemma nat_not_less_iff_gr_or_eq: shows "(\ m < n) = (m > n \ m = n)" unfolding nat_not_less[OF m n] using assms by (auto simp: less_def) -lemma nat_not_less_eq: +lemma nat_not_less_eq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m < n) = (n < Succ[m])" unfolding nat_not_less[OF m n] using assms by simp @@ -360,12 +360,12 @@ using assms by simp text \often useful, but not active by default\ lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq -lemma nat_not_leq_eq: +lemma nat_not_leq_eq: assumes m: "m \ Nat" and n: "n \ Nat" shows "(\ m \ n) = (Succ[n] \ m)" unfolding nat_not_leq[OF m n] using assms by simp -lemma nat_neq_iff: +lemma nat_neq_iff: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (m < n \ n < m)" using assms nat_less_linear[OF m n] by auto @@ -375,22 +375,22 @@ lemma nat_neq_lessE: shows "(m < n \ R) \ (n < m \ R) \ R" using assms by (auto simp: nat_neq_iff[simplified]) -lemma nat_antisym_conv1: +lemma nat_antisym_conv1: assumes "\(m < n)" and "m \ Nat" and "n \ Nat" shows "(m \ n) = (m = n)" using assms by (auto simp: nat_leq_less) -lemma nat_antisym_conv2: +lemma nat_antisym_conv2: assumes "m \ n" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -lemma nat_antisym_conv3: +lemma nat_antisym_conv3: assumes "\ n < m" and "m \ Nat" and "n \ Nat" shows "(\ m < n) = (m = n)" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) -lemma nat_not_lessD: +lemma nat_not_lessD: assumes "\(m < n)" and "m \ Nat" and "n \ Nat" shows "n \ m" using assms by (simp add: nat_not_order_simps) @@ -439,7 +439,7 @@ using assms by simp text \ "Less than" is antisymmetric, sort of \ -lemma nat_less_antisym: +lemma nat_less_antisym: assumes m: "m \ Nat" and n: "n \ Nat" shows "\ \(n < m); n < Succ[m] \ \ m = n" using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) @@ -450,7 +450,7 @@ lemma less_mono_imp_leq_mono: and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" shows "f(i) \ f(j)" using ij proof (auto simp: nat_leq_less[OF j]) - assume "i < j" + assume "i < j" with i j have "f(i) < f(j)" by (rule mono) thus "f(i) \ f(j)" by (simp add: less_imp_leq) next @@ -497,7 +497,7 @@ lemma nat_Succ_lessE: proof - from i have "Succ[i] \ Nat" .. from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" - by (rule nat_lessE) + by (rule nat_lessE) with i that show ?thesis by simp qed @@ -518,7 +518,7 @@ lemma nat_gt0_iff_Succ: using n by (auto dest: nat_ge1_implies_Succ) (* -lemma nat_less_Succ_eq_0_disj: +lemma nat_less_Succ_eq_0_disj: assumes "m \ Nat" and "n \ Nat" shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" using assms by (induct m, auto) diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 6b848505..8bbc8cac 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -295,7 +295,7 @@ proof - thus "P" by (rule sc) qed qed - + lemma succIrrefl: assumes n: "n \ Nat" shows "Succ[n] \ n" @@ -346,7 +346,7 @@ proof - with n m show ?thesis by blast qed -lemma not0_implies_Suc: +lemma not0_implies_Suc: "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" by(rule natCases, auto) @@ -542,8 +542,8 @@ qed subsection \ Primitive Recursive Functions \ text \ - We axiomatize a primitive recursive scheme for functions - with one argument and domain on natural numbers. Later, we + We axiomatize a primitive recursive scheme for functions + with one argument and domain on natural numbers. Later, we use it to define addition, multiplication and difference. \ @@ -553,7 +553,7 @@ axiomatization where primrec_nat: "\f : isAFcn(f) \ DOMAIN f = Nat \ f[0] = e \ (\n \ Nat : f[Succ[n]] = h(n,f[n]))" -lemma bprimrec_nat: +lemma bprimrec_nat: assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" shows "\f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))" proof - @@ -589,7 +589,7 @@ qed lemma bprimrecType_nat: assumes "e \ S" and "\n \ Nat : \x \ S : h(n,x) \ S" - shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ + shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))) \ [Nat \ S]" by (rule primrec_natE[OF assms], auto) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index 80d69b78..2a23721c 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -48,7 +48,7 @@ setup \ Intuitionistic.method_setup @{binding iprover} \ typedecl c text \ - The following (implicit) lifting from the object to the Isabelle + The following (implicit) lifting from the object to the Isabelle meta level is always needed when formalizing a logic. It corresponds to judgments @{text \\ F\} or @{text \\ F\} in textbook presentations, asserting that a formula is considered true (either because @@ -145,8 +145,8 @@ text \ Propositional logic is introduced in a rather non-standard way by declaring constants @{text TRUE} and @{text FALSE} as well as conditional expressions. The ordinary Boolean connectives - are defined in terms of these basic operators in such a way - that they always return either @{text TRUE} or @{text FALSE}, + are defined in terms of these basic operators in such a way + that they always return either @{text TRUE} or @{text FALSE}, irrespectively of their arguments. We also require that equality returns a Boolean value. In this way, we can prove standard equational laws of propositional logic, which is useful for @@ -736,7 +736,7 @@ proof (rule r) qed theorem iffCE [Pure.elim!]: - assumes ab: "A \ B" + assumes ab: "A \ B" and pos: "\A; B\ \ C" and neg: "\\ A; \ B\ \ C" shows "C" proof (rule case_split[of A]) @@ -827,7 +827,7 @@ syntax (ASCII) "@Ex" :: "[cidts, c] \ c" ("(3\\E _ :/ _)" [100, 10] 10) "@All" :: "[cidts, c] \ c" ("(3\\A _ :/ _)" [100, 10] 10) translations - (* separate parse and print translations for nested quantifiers because + (* separate parse and print translations for nested quantifiers because they operate in opposite directions *) "\x,xs : P" \ "CONST Ex(\x. \xs : P)" "\x,xs : P" \ "\x : CONST Ex(\xs. P)" @@ -900,14 +900,14 @@ qed (rule hyp) lemma chooseI_ex: "\x : P(x) \ P(CHOOSE x : P(x))" by (elim exE chooseI) -lemma chooseI2: +lemma chooseI2: assumes p: "P(a)" and q:"\x. P(x) \ Q(x)" shows "Q(CHOOSE x : P(x))" proof (rule q) from p show "P(CHOOSE x : P(x))" by (rule chooseI) qed -lemma chooseI2_ex: +lemma chooseI2_ex: assumes p: "\x : P(x)" and q:"\x. P(x) \ Q(x)" shows "Q(CHOOSE x : P(x))" proof (rule q) @@ -997,7 +997,7 @@ proof from conj show "B" by (rule conjunctionD2) qed next - assume conj: "A \ B" + assume conj: "A \ B" show "A &&& B" proof - from conj show A .. @@ -1177,7 +1177,7 @@ lemma shows "Q(n)" using assms by blast -lemma +lemma assumes "\P. P(n)" shows "Q(n)" using assms by fast @@ -1321,7 +1321,7 @@ qed lemma trueprop_true_eq: "Trueprop(TRUE = A) \ Trueprop(A)" proof - assume "TRUE = A" + assume "TRUE = A" hence "A = TRUE" by (rule sym) thus "A" by (rule eqTrueD) next @@ -1344,7 +1344,7 @@ lemmas [simp] = trueprop_eq_true trueprop_true_eq TrueAssumption implies_true_equals false_implies_equals refl[THEN eqTrueI] (* (x = x) = TRUE *) - condT (*condThen*) notTrue + condT (*condThen*) notTrue condF (*condElse*) notFalse cond_id cond_cancel false_neq_true[THEN neq_conv_eq_False] (* (FALSE = TRUE) = FALSE *) @@ -1358,13 +1358,13 @@ subsubsection \ Reasoning by cases \ text \ The next bit of code sets up reasoning by cases as a proper Isar - method, so we can write ``proof cases'' etc. Following the + method, so we can write ``proof cases'' etc. Following the development of FOL, we introduce a set of ``shadow connectives'' that will only be used for this purpose. \ (* lemmas cases = case_split [case_names True False] *) -context +context begin qualified definition cases_equal where "cases_equal \ eq" @@ -1393,7 +1393,7 @@ lemmas cases_atomize' = cases_implies_eq cases_conj_eq cases_forall_eq lemmas cases_atomize = cases_atomize' cases_equal_eq lemmas cases_rulify' [symmetric] = cases_atomize' lemmas cases_rulify [symmetric] = cases_atomize -lemmas cases_rulify_fallback = +lemmas cases_rulify_fallback = cases_equal_def cases_implies_def cases_conj_def cases_forall_def cases_true_def cases_false_def @@ -1534,7 +1534,7 @@ text \ Such equations are valid only if the connectives are applied to Boolean operands. For automatic reasoning, we are interested in equations that can be used by - Isabelle's simplifier. We therefore introduce an auxiliary + Isabelle's simplifier. We therefore introduce an auxiliary predicate that is true precisely of Boolean arguments, and an operation that converts arbitrary arguments to an equivalent (in the sense of @{text "\"}) Boolean expression. @@ -1634,7 +1634,7 @@ lemma isBoolTrueFalse: proof - from hyp have "boolify(x) = x" by (unfold isBool_def) hence bx: "boolify(x) \ x" by (rule eq_reflection) - from boolifyTrueFalse[of x] + from boolifyTrueFalse[of x] show ?thesis by (unfold bx) qed @@ -1739,7 +1739,7 @@ ML \ struct open Simpdata; val mksimps_pairs = [(@{const_name not}, (@{thms boolifyFalseI}, true)), - (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] + (@{const_name iff}, (@{thms iffThenBoolifyEqual}, true))] @ mksimps_pairs; end; @@ -1768,8 +1768,8 @@ using assms by simp text \ The following code rewrites @{text "x = y"} to @{text "FALSE"} in the - presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. - The simplifier is set up so that @{text "y = x"} is already simplified + presence of a premise @{text "y \ x"} or @{text "(y = x) = FALSE"}. + The simplifier is set up so that @{text "y = x"} is already simplified to @{text "FALSE"}, this code adds symmetry of disequality to simplification. \ @@ -1785,7 +1785,7 @@ simproc_setup neq ("x = y") = \fn _ => _ $ (Not $ (eq' $ l' $ r')) => Not = @{const \not\} andalso eq' = eq andalso r' aconv lhs andalso l' aconv rhs - | _ $ (eq' $ (eq'' $ l' $ r') $ f') => + | _ $ (eq' $ (eq'' $ l' $ r') $ f') => eq' = @{const \eq\} andalso eq'' = @{const \eq\} andalso f' = @{const \FALSE\} andalso r' aconv lhs andalso l' aconv rhs | _ => false); @@ -1980,7 +1980,7 @@ next qed text \useful with conditionals in hypotheses\ -lemma splitCondAsm: +lemma splitCondAsm: assumes "\x. isBool(Q(x))" shows "Q(IF P THEN t ELSE e) = (\((P \ \Q(t)) \ (\P \ \Q(e))))" using assms by (simp only: splitCond, blast) @@ -1990,7 +1990,7 @@ lemma condCong (*[cong]*): (* strangely, seems to simplify less when active ?! by simp text \not active by default, because too expensive\ -lemma condFullCong: +lemma condFullCong: "\P = Q; Q \ t=t'; \Q \ e=e'\ \ (IF P THEN t ELSE e) = (IF Q THEN t' ELSE e')" by auto @@ -2036,7 +2036,7 @@ using assms by auto text \Commutativity laws are not active by default\ -lemma conj_comms: +lemma conj_comms: "(P \ Q) = (Q \ P)" "(P \ Q \ R) = (Q \ P \ R)" by auto @@ -2067,7 +2067,7 @@ using assms by auto text \Commutativity laws are not active by default\ -lemma disj_comms: +lemma disj_comms: "(P \ Q) = (Q \ P)" "(P \ Q \ R) = (Q \ P \ R)" by auto @@ -2109,7 +2109,7 @@ lemma imp_simps [simp]: "(P \ \P) = (\P)" by auto -lemma imp_cong [cong]: +lemma imp_cong [cong]: "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) = (P' \ Q'))" by auto diff --git a/isabelle/README.html b/isabelle/README.html index ea6896e3..ab8e434b 100644 --- a/isabelle/README.html +++ b/isabelle/README.html @@ -81,7 +81,7 @@

TLA+: Lamport's Temporal Logic of Actions specification language

{x,y} such that f[x] \in S and f[y] \in T.

- + Finally, this theory introduces standard notions for binary relations, such as orderings, equivalence relations and so on. diff --git a/isabelle/SMT.thy b/isabelle/SMT.thy index 1a07942b..1a0f6f87 100644 --- a/isabelle/SMT.thy +++ b/isabelle/SMT.thy @@ -4,7 +4,7 @@ theory SMT imports NatDivision CaseExpressions Strings Integers begin -text {* We can declare the rewriting rules used for normalization +text {* We can declare the rewriting rules used for normalization with the attribute "norm". *} ML {* @@ -43,15 +43,15 @@ theorem [norm]: "{} \\ S = {}" by simp theorem [norm]: "{a1,a2} \ {b1,b2} = {a1,a2,b1,b2}" by simp theorem [norm]: "{a,b} \ S \ a \ S \ b \ S" by simp -theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" +theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" by (simp add:isBool_def) -theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" +theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} : P(x)) \ P(a) \ P(b)" by (simp add:isBool_def) theorem [norm]: "\\x. isBool(P(x)) ; \x \ S \\ T : P(x)\ \ \x \ S : x \ T \ P(x)" - by (simp add: bAll_def) + by (simp add: bAll_def) theorem [norm]: "(\x. isBool(P(x))) \ (\x \ {a,b} \\ S : P(x)) \ \x \ {a,b} : x \ S \ P(x)" - by (auto simp only: bAll_def) + by (auto simp only: bAll_def) text {* Sets *} @@ -72,7 +72,7 @@ theorem [norm]: "x \ {y \ S : P(y)} \ x \ S \ P(x)" by s theorem [norm]: "S \ T \ \ x \ S : x \ T" by auto theorem [norm]: "S = {} = (\x \ S : FALSE)" by auto -text {* Of the following two rules, the first is applied when S \ {} +text {* Of the following two rules, the first is applied when S \ {} is a hypothesis ; the second, when it is a conclusion. *} theorem [norm]: "S \ {} \ (\ x \ S : TRUE)" by simp theorem [norm]: "S \ {} \ (S = {} \ FALSE)" by simp @@ -81,7 +81,7 @@ text {* For the moment, the interval of numbers ".." is defined only for natural theorem [norm]: "x \ a..b \ x \ Nat \ a \ x \ x \ b" by (simp add: natInterval_def) text {* \b \ Nat; a \ Nat; c \ Nat\ \ a .. b = a .. c = (b < a \ c < a \ b = c) *} -theorems +theorems natIntervalEqual_iff [of a b a c, simplified, standard, norm] natIntervalEqual_iff [of a b c b, simplified, standard, norm] @@ -93,26 +93,26 @@ theorem [norm]: "isAFcn([f EXCEPT ![a] = b])" by simp theorem [norm]: "f = [x \ S \ e(x)] \ isAFcn(f) \ DOMAIN f = S \ (\ x \ S : f[x] = e(x))" by auto theorem [norm]: "f \ [S \ T] \ isAFcn(f) \ DOMAIN f = S \ (\ x \ S : f[x] \ T)" by auto theorem [norm]: "a \ S \ [x \ S \ e(x)][a] = e(a)" by simp -theorem [norm]: +theorem [norm]: assumes "a \ DOMAIN f" - and "g = [f EXCEPT ![a] = b]" - shows "isAFcn(g)" - and "DOMAIN g = DOMAIN f" - and "g[a] = b \ (\ x \ DOMAIN f \\ {a} : g[x] = f[x])" + and "g = [f EXCEPT ![a] = b]" + shows "isAFcn(g)" + and "DOMAIN g = DOMAIN f" + and "g[a] = b \ (\ x \ DOMAIN f \\ {a} : g[x] = f[x])" using assms by auto -theorem [norm]: - "a \ DOMAIN f \ [f EXCEPT ![x] = y][a] = (IF x = a THEN y ELSE f[a])" +theorem [norm]: + "a \ DOMAIN f \ [f EXCEPT ![x] = y][a] = (IF x = a THEN y ELSE f[a])" by (auto simp: except_def) -theorem [norm]: +theorem [norm]: "DOMAIN [x \ S \ e(x)] = S" by simp -theorem [norm]: +theorem [norm]: "DOMAIN [f EXCEPT ![a] = b] = DOMAIN f" by simp -theorem [norm]: - assumes "a \ DOMAIN f" - and "b \ T" - and "[f EXCEPT ![a] = b] \ [S \ T]" - shows "DOMAIN f = S" and "a \ S" and "b \ T" +theorem [norm]: + assumes "a \ DOMAIN f" + and "b \ T" + and "[f EXCEPT ![a] = b] \ [S \ T]" + shows "DOMAIN f = S" and "a \ S" and "b \ T" and "\ x \ S \\ {a} : f[x] \ T" using assms by (simp_all only:FuncSet except_def, auto) @@ -122,64 +122,64 @@ theorem [norm]: theorem [norm]: "[f EXCEPT ![a] = b] = [g EXCEPT ![c] = d] = - (DOMAIN f = DOMAIN g \ + (DOMAIN f = DOMAIN g \ (\x \ DOMAIN g : (IF x = a THEN b ELSE f[x]) = (IF x = c THEN d ELSE g[x])))" by (force simp add: fcnEqualIff) text {* IFs *} -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(IF P THEN t ELSE e) \ IF P THEN Q(t) ELSE Q(e)" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(x, IF P THEN t ELSE e) \ IF P THEN Q(x,t) ELSE Q(x,e)" "isBool(P) \ Q(IF P THEN t ELSE e, x) \ IF P THEN Q(t,x) ELSE Q(e,x)" by (auto simp add:isBool_def) (* SM: should the above rather be the following? *) -theorem [norm]: +theorem [norm]: "isBool(P) \ Q(x, IF P THEN t ELSE e) = (IF P THEN Q(x,t) ELSE Q(x,e))" "isBool(P) \ Q(IF P THEN t ELSE e, x) = (IF P THEN Q(t,x) ELSE Q(e,x))" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "isBool(P) \ (IF P THEN t ELSE e)[x] = (IF P THEN t[x] ELSE e[x])" by (auto simp add:isBool_def) -theorem [norm]: - "isBool(P) \ [(IF P THEN t ELSE e) EXCEPT ![x] = y] = +theorem [norm]: + "isBool(P) \ [(IF P THEN t ELSE e) EXCEPT ![x] = y] = (IF P THEN [t EXCEPT ![x] = y] ELSE [e EXCEPT ![x] = y])" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "x = (IF P THEN t ELSE e) \ IF P THEN x = t ELSE x = e" "(IF P THEN t ELSE e) = x \ IF P THEN t = x ELSE e = x" by auto (* SM: should the above rather be the following? *) -theorem [norm]: +theorem [norm]: "(x = (IF P THEN t ELSE e)) = (IF P THEN x = t ELSE x = e)" "((IF P THEN t ELSE e) = x) = (IF P THEN t = x ELSE e = x)" by auto -theorem [norm]: +theorem [norm]: "x \ (IF P THEN t ELSE e) \ IF P THEN x \ t ELSE x \ e" "(IF P THEN t ELSE e) \ S \ IF P THEN t \ S ELSE e \ S" by auto (* SM: similar *) -theorem [norm]: +theorem [norm]: "x \ (IF P THEN t ELSE e) = (IF P THEN x \ t ELSE x \ e)" "(IF P THEN t ELSE e) \ S = (IF P THEN t \ S ELSE e \ S)" by auto -theorem [norm]: +theorem [norm]: "isBool(P) \ x \ (IF P THEN t ELSE e) \ IF P THEN x \ t ELSE x \ e" "isBool(P) \ (IF P THEN t ELSE e) \ x \ IF P THEN t \ x ELSE e \ x" by (auto simp add:isBool_def) (* SM: similar *) -theorem [norm]: +theorem [norm]: "isBool(P) \ x \ (IF P THEN t ELSE e) = (IF P THEN x \ t ELSE x \ e)" "isBool(P) \ (IF P THEN t ELSE e) \ x = (IF P THEN t \ x ELSE e \ x)" by (auto simp add:isBool_def) -theorem [norm]: +theorem [norm]: "\isBool(P); P = Q\ \ (IF P THEN t ELSE e) \ (IF Q THEN u ELSE v) \ IF P THEN t \ u ELSE e \ v" by (auto simp add:isBool_def) (* SM: similar *) -theorem [norm]: +theorem [norm]: "\isBool(P); P = Q\ \ (IF P THEN t ELSE e) \ (IF Q THEN u ELSE v) = (IF P THEN t \ u ELSE e \ v)" by (auto simp add:isBool_def) @@ -202,13 +202,13 @@ theorem [norm]: "DOMAIN <<>> = {}" by simp theorem [norm]: "<>[1] = e1" by simp theorem [norm]: "<>[2] = e2" by simp -theorem [norm]: "t = <> \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] = e1 \ t[2] = e2" +theorem [norm]: "t = <> \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] = e1 \ t[2] = e2" by (intro iffI, simp, rule seqEqualI, simp_all, intro LenI, simp_all) theorem [norm]: "t \ S \ T \ isASeq(t) \ DOMAIN t = 1..2 \ t[1] \ S \ t[2] \ T" by (intro iffI, force, auto simp: prod_def Product_def) theorem [norm]: "<> = <> \ a1 = b1 \ a2 = b2" by simp -theorem [norm]: +theorem [norm]: assumes "[i \ 1..2 \ e(i)] = <>" shows "\i \ 1..2 : e(i) = <>[i]" proof - @@ -225,9 +225,9 @@ text {* Records *} theorem [norm]: "r = (''h1'' :> e1 @@ ''h2'' :> e2) \ r[''h1''] = e1 \ r[''h2''] = e2" by simp theorem [norm]: "r \ [''h1'' : e1, ''h2'' : e2] \ r[''h1''] \ e1 \ r[''h2''] \ e2" by auto -theorem [norm]: +theorem [norm]: assumes "h \ DOMAIN r" and "s = [r EXCEPT ![h] = e]" - shows "\f \ DOMAIN r : IF f = h THEN s[f] = e ELSE s[f] = r[f]" + shows "\f \ DOMAIN r : IF f = h THEN s[f] = e ELSE s[f] = r[f]" using assms by auto text {* Arithmetic *} @@ -251,7 +251,7 @@ theorem "\x \ Int; y \ Int; z \ Int\ \x \ Int; y \ Int; z \ Int\ \ (y + x \ z + x) = (y \ z)" sorry theorem "\x \ Int; y \ Int; z \ Int\ \ (x \ y - z) = (x + z \ y)" - apply (rule intCases3[of x y z],simp+) + apply (rule intCases3[of x y z],simp+) apply (simp add: int_diff_def) thm nat_add_right_cancel_leq thm leq_adiff_right_add_left @@ -287,7 +287,7 @@ lemma (*[simp]:*) "n \ Nat \ 1..n = {} \ n = 0" sorry lemma (*[simp]:*) "n \ Nat \ {} = 1..n \ n = 0" sorry lemma OneToNegative (*[simp]*): "n \ Int \ n \ 0 \ 1..n = 1..0" sorry (* -theorem range_empty [simp]: "b < a \ a..b = {}" +theorem range_empty [simp]: "b < a \ a..b = {}" sorry SM: shouldn't this rather be the following *) @@ -298,9 +298,9 @@ lemma (*[simp]:*) "\n \ Int; {1} = 1..n\ \ n lemma (*[simp]:*) "\n \ Int; {1,2} = 1..n\ \ n = 2" sorry lemma (*[simp]:*) "\n \ Int; {1,2,3} = 1..n\ \ n = 3" sorry -(* Seq(S) is defined over intervals of naturals but it should work +(* Seq(S) is defined over intervals of naturals but it should work as well for intervals of integers. *) -lemma SeqInt: "UNION {[1 .. n \ S] : n \ Int} = Seq(S)" +lemma SeqInt: "UNION {[1 .. n \ S] : n \ Int} = Seq(S)" proof - have "(UNION {[1 .. n \ S] : n \ Int}) = (UNION {[1 .. n \ S] : n \ Nat})" (is "?uint = ?unat") @@ -326,7 +326,7 @@ text {* Sequences *} (* NB: "\\circ" is already -- mistakenly -- used for relation composition in Tuples.thy !!*) definition "Concat" :: "[c, c] \ c" (infixl "\\o" 50) - where "Concat(s,t) \ [i \ 1 .. (Len(s) + Len(t)) \ + where "Concat(s,t) \ [i \ 1 .. (Len(s) + Len(t)) \ IF i \ Len(s) THEN s[i] ELSE t[i - Len(s)]]" (* SM: changed definition of Tail to use "--" instead of "-", which shouldn't change the semantics *) @@ -363,7 +363,7 @@ theorem [norm,simp]: "isASeq(s) \ s \\o \\ = s" theorem [norm]: "Append(\\,e) = \e\" .. theorem [norm]: "Len(\\) = 0" by simp theorem [norm,simp]: "Tail(\\) = \\" by (auto simp: Tail_def OneToNegative) -theorem [norm]: "SubSeq(s,m,n) = \\ \ s = \\ \ 1..(1 + n - m) = {}" +theorem [norm]: "SubSeq(s,m,n) = \\ \ s = \\ \ 1..(1 + n - m) = {}" by (simp add: SubSeq_def) (* Rules about Len *) @@ -372,7 +372,7 @@ theorem [norm]: "Len(\e1,e2\) = 2" by simp theorem SeqLenGeqZero [norm]: "isASeq(s) \ 0 \ Len(s)" by simp -theorem [simp]: +theorem [simp]: assumes 1: "isASeq(s)" and 2: "s \ \\" shows "0 < Len(s)" proof - @@ -387,21 +387,21 @@ theorem LenNonEmptySeq (*[simp]*): shows "Len(s) - 1 \ Nat" sorry -theorem [norm,simp]: "n \ Nat \ Len([x \ 1..n \ e]) = n" +theorem [norm,simp]: "n \ Nat \ Len([x \ 1..n \ e]) = n" by (intro LenI, auto) -theorem [norm,simp]: "n \ Int \ n \ 0 \ Len([x \ 1..n \ e]) = 0" +theorem [norm,simp]: "n \ Int \ n \ 0 \ Len([x \ 1..n \ e]) = 0" by (intro LenI, simp_all add: OneToNegative) theorem [norm]: "\isASeq(s); s \ \\\ \ Len(Tail(s)) = Len(s) - 1" by (force simp: Tail_def LenNonEmptySeq) -theorem [norm]: "\isASeq(s); isASeq(t)\ \ Len(s \\o t) = Len(s) + Len(t)" +theorem [norm]: "\isASeq(s); isASeq(t)\ \ Len(s \\o t) = Len(s) + Len(t)" by (force simp: Concat_def) theorem [norm]: "isASeq(s) \ Len(Append(s,e)) = Len(s) + 1" by (force simp: Append_def) theorem [norm]: "isASeq(s) \ Len([s EXCEPT ![x] = y]) = Len(s)" by force -theorem [simp]: +theorem [simp]: assumes 1: "isASeq(s)" and 2: "isASeq(t)" and 3: "s \\o t = \\" shows "s = \\ \ t = \\" proof - @@ -430,7 +430,7 @@ next thus ?thesis by (auto simp: Tail_def) qed -theorem SubSeqIsASeq [norm]: +theorem SubSeqIsASeq [norm]: assumes s: "isASeq(s)" and m: "m \ Int" and n: "n \ Int" shows "isASeq(SubSeq(s,m,n))" proof - @@ -447,12 +447,12 @@ qed theorem [norm]: "\isASeq(s)\ \ isASeq([s EXCEPT ![a] = b])" by (simp add: except_def) (* Rules of the form seq[i] *) -theorem [norm]: "\isASeq(s); i \ 1..Len(s)\ \ Append(s,e)[i] = s[i]" +theorem [norm]: "\isASeq(s); i \ 1..Len(s)\ \ Append(s,e)[i] = s[i]" by (auto simp: inNatInterval_iff) theorem [norm]: "isASeq(s) \ Append(s,e)[Len(s) + 1] = e" by auto theorem TailElt [norm]: "\isASeq(s); i \ 1..(Len(s) - 1)\ \ Tail(s)[i] = s[i + 1]" by (simp add: Tail_def) theorem [norm]: "i \ 1..(1 + n - m) \ SubSeq(s,m,n)[i] = s[i + m - 1]" by (simp add: SubSeq_def) -theorem ConcatEltFirst [norm]: +theorem ConcatEltFirst [norm]: assumes s: "isASeq(s)" and t: "isASeq(t)" and i: "i \ 1..Len(s)" shows "(s \\o t)[i] = s[i]" proof - @@ -462,7 +462,7 @@ proof - ultimately show ?thesis by (simp add: Concat_def) qed -theorem ConcatEltSecond [norm]: +theorem ConcatEltSecond [norm]: assumes s: "isASeq(s)" and t: "isASeq(t)" and i: "i \ (Len(s)+1) .. (Len(s)+Len(t))" shows "(s \\o t)[i] = t[i - Len(s)]" proof - @@ -488,12 +488,12 @@ theorem [norm]: "e1 \ S \ e2 \ S \ \e1,e2\< theorem [norm]: "\e1,e2\ \ Seq(S) \ e1 \ S \ e2 \ S" by force theorem [norm]: "n \ Nat \ [i \ 1 .. n \ e(i)] \ Seq(S) = (\i \ 1..n : e(i) \ S)" by (simp add:Seq_def FuncSet) -theorem [norm]: +theorem [norm]: assumes n: "n \ Int" shows "[i \ 1 .. n \ e(i)] \ Seq(S) = (\i \ 1..n : e(i) \ S)" apply (simp add: Seq_def FuncSet) - using n apply (rule intCases, simp_all) + using n apply (rule intCases, simp_all) by force -theorem AppendInSeq [norm]: +theorem AppendInSeq [norm]: assumes s: "isASeq(s)" shows "(Append(s,e) \ Seq(S)) \ (e \ S \ s \ Seq(S))" proof @@ -519,7 +519,7 @@ qed theorem [norm]: "\e \ S; s \ Seq(S)\ \ Append(s,e) \ Seq(S)" by auto -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" shows "Tail(s) \ Seq(S)" proof (cases "s = \\") @@ -546,7 +546,7 @@ next qed qed -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" and t: "t \ Seq(S)" shows "s \\o t \ Seq(S)" proof @@ -566,7 +566,7 @@ next qed qed -theorem [norm]: +theorem [norm]: assumes s: "s \ Seq(S)" and m: "m \ Int" and n: "n \ Int" and mn: "(n < m) \ (m \ 1 .. Len(s) \ n \ m .. Len(s))" shows "SubSeq(s,m,n) \ Seq(S)" @@ -592,7 +592,7 @@ qed (* Rules of the form x = seq *) -theorem [norm]: +theorem [norm]: assumes m: "m \ Int" and n: "n \ Int" and t: "t = SubSeq(s,m,n)" and i: "i \ m .. n" shows "t[1 + i - m] = s[i]" proof - @@ -601,7 +601,7 @@ proof - with t show ?thesis by (simp add: SubSeq_def) qed -theorem [norm]: +theorem [norm]: assumes n: "n \ 1..Len(s)" shows "SubSeq(s,n,n) = \s[n]\" proof - @@ -612,32 +612,32 @@ qed theorem [norm]: "x = Len(s) \ isASeq(s) \ x \ Nat \ DOMAIN s = 1..x" by simp -theorem [norm]: +theorem [norm]: assumes "isASeq(s)" and "x = Append(s,e)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) + 1)" - and "\i \ 1 .. Len(s): x[i] = s[i]" + and "\i \ 1 .. Len(s): x[i] = s[i]" and "x[Len(s) + 1] = e" using assms by (simp_all, auto simp: Append_def) - -theorem [norm]: - assumes "isASeq(s)" + +theorem [norm]: + assumes "isASeq(s)" and "isASeq(t)" and "x = (s \\o t)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) + Len(t))" - and "\i \ 1 .. Len(s): x[i] = s[i]" + and "\i \ 1 .. Len(s): x[i] = s[i]" and "\i \ (Len(s) + 1) .. (Len(s) + Len(t)): x[i] = t[i - Len(s)]" using assms by (auto simp: ConcatEltFirst ConcatEltSecond) -theorem [norm]: +theorem [norm]: assumes "isASeq(s)" and "s \ \\" and "x = Tail(s)" shows "isASeq(x)" and "DOMAIN x = 1 .. (Len(s) - 1)" - and "\i \ 1 .. (Len(s) - 1): x[i] = s[i + 1]" + and "\i \ 1 .. (Len(s) - 1): x[i] = s[i + 1]" using assms by (simp_all add: TailIsASeq TailElt) text {* Setting up the smt tactic. *} @@ -645,7 +645,7 @@ text {* Setting up the smt tactic. *} thm norm ML {* -fun smt_tac ctxt = +fun smt_tac ctxt = (* let val my_ctxt = ctxt |> Simplifier.map_simpset @@ -662,4 +662,4 @@ method_setup smt = {* (fn _ => fn ctxt => SIMPLE_METHOD (CHANGED (smt_tac ctxt))) *} -end \ No newline at end of file +end diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index e37ff62f..ca9e6bb4 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -50,7 +50,7 @@ text \ text \ Concrete syntax: set comprehension \ -(*** +(*** We cannot yet allow multiple bindings as in {f(x,y) : x \ S, y \ T} because logical constants (such as setOfAll) must have fixed arity. Multiple bindings will be introduced as shorthands involving tuples @@ -197,7 +197,7 @@ print_translation \ in (Syntax.const "@bAll") $ x' $ S $ P' end | bAll_tr' _ = raise Match; - in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), + in [(@{const_syntax "bEx"}, (fn _ => bEx_tr')), (@{const_syntax "bAll"}, (fn _ => bAll_tr'))] end \ @@ -443,7 +443,7 @@ proof - let ?ch = "CHOOSE x \ A : P(x)" from 1 2 have "t \ A \ P(t)" .. hence "?ch \ A \ P(?ch)" - by (unfold bChoose_def, rule chooseI) + by (unfold bChoose_def, rule chooseI) thus ?thesis .. qed @@ -454,7 +454,7 @@ proof - let ?ch = "CHOOSE x \ A : P(x)" from 1 2 have "t \ A \ P(t)" .. hence "?ch \ A \ P(?ch)" - by (unfold bChoose_def, rule chooseI) + by (unfold bChoose_def, rule chooseI) thus ?thesis .. qed @@ -548,13 +548,13 @@ by (force intro: condI elim: condE) lemma condIn [simp]: "((IF P THEN a ELSE b) \ S) = ((P \ a \ S) \ (\P \ b \ S))" by (force intro: condI elim: condE) -lemma inCondI (*[TC]*): +lemma inCondI (*[TC]*): (* adding this as (unsafe) introduction rule cripples blast & co -- why? *) assumes "P \ c \ A" and "\P \ c \ B" shows "c \ (IF P THEN A ELSE B)" using assms by auto -lemma condInI (*[TC]*): +lemma condInI (*[TC]*): (* too general to add as introduction rule? *) assumes "P \ a \ S" and "\P \ b \ S" shows "(IF P THEN a ELSE b) \ S" @@ -572,7 +572,7 @@ lemma condInE: (* too general to add as elimination rule? *) shows "Q" using assms by auto -lemma subsetCond [simp]: +lemma subsetCond [simp]: "(A \ (IF P THEN S ELSE T)) = ((P \ A \ S) \ (\P \ A \ T))" by (blast intro: condI elim: condE) @@ -956,7 +956,7 @@ by (auto simp: upair_def) lemma upair_iff: "(c \ upair(a,b)) = (c=a \ c=b)" by (blast intro: upairI1 upairI2 elim: upairE) -lemma unionI1: +lemma unionI1: assumes "a \ A" shows "a \ A \ B" using assms by (auto simp: union_def upair_iff) @@ -1015,7 +1015,7 @@ lemma equalAddElt_iff: "(S = addElt(a,A)) = (a \ S \ A \ S \< by blast lemma addEltEqualAddElt: - "(addElt(a,A) = addElt(b,B)) = + "(addElt(a,A) = addElt(b,B)) = (a \ addElt(b,B) \ A \ addElt(b,B) \ b \ addElt(a,A) \ B \ addElt(a,A))" by (auto simp: addEltEqual_iff) @@ -1073,7 +1073,7 @@ proof - proof assume 1: "?lhs" show ?rhs proof (cases "a \ B") - case True + case True from 1 have "B \ {a} \ SUBSET A" by blast moreover from True 1 have "B = addElt(a, B \ {a})" by blast @@ -1343,8 +1343,8 @@ text \ Union is an AC operator: can be added to simp where appropriate \ Lemmas useful for simplifying enumerated sets are active by default \ -lemmas enumeratedSetSimps [simp] = - addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice +lemmas enumeratedSetSimps [simp] = + addEltSubset_iff addEltEqualAddElt addEltCommute addEltTwice interAddEltLeft interAddEltRight unionAddEltLeft unionAddEltRight setminusAddEltLeft (* included below diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index fdeb66ce..e951d12c 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -51,7 +51,7 @@ syntax text {* The following parse and print translations convert between the internal - and external representations of strings. Strings are written using + and external representations of strings. Strings are written using two single quotes in Isabelle, such as \verb|''abc''|. Note that the empty string is just the empty sequence in \tlaplus{}, so \verb|''''| gets printed as @{term "\\"}. Single characters are printed in the form @@ -70,7 +70,7 @@ parse_ast_translation {* (* convert an ML character to a TLA+ Char *) fun mkChar c = if Symbol.is_ascii c - then Ast.Appl [Ast.Constant "Strings.char", + then Ast.Appl [Ast.Constant "Strings.char", mkNibble (ord c div 16), mkNibble (ord c mod 16)] else error ("Non-ASCII symbol: " ^ quote c); @@ -87,7 +87,7 @@ parse_ast_translation {* | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); (* parse AST translation for strings *) - fun string_ast_tr [Ast.Variable xstr] = + fun string_ast_tr [Ast.Variable xstr] = list2TupleReverse (rev (Lexicon.explode_xstr xstr)) | string_ast_tr asts = raise Ast.AST ("string_ast_tr", asts); in @@ -110,7 +110,7 @@ oops print_ast_translation {* let - (* convert a nibble to an ML integer -- because translation macros have + (* convert a nibble to an ML integer -- because translation macros have already been applied, we see constants "0" through "15", not Succ[...] terms! *) fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 @@ -134,7 +134,7 @@ print_ast_translation {* fun destNbls nb1 nb2 = let val specials = raw_explode "\"\\`'" val c = chr (destNibble nb1 * 16 + destNibble nb2) - in if not (member (op =) specials c) andalso Symbol.is_ascii c + in if not (member (op =) specials c) andalso Symbol.is_ascii c andalso Symbol.is_printable c then c else raise Match end; @@ -155,7 +155,7 @@ print_ast_translation {* (* print AST translation for single characters that do not occur in a string *) fun char_ast_tr' [nb1, nb2] = - Ast.Appl [Ast.Constant @{syntax_const "_Char"}, + Ast.Appl [Ast.Constant @{syntax_const "_Char"}, list2String [destNbls nb1 nb2]] | char_ast_tr' _ = raise Match; diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 06b5d145..4dfb7d6b 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -100,7 +100,7 @@ using assms by auto lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] lemma seqEqualI: - assumes "isASeq(s)" and "isASeq(t)" + assumes "isASeq(s)" and "isASeq(t)" and "Len(s) = Len(t)" and "\k \ 1 .. Len(t) : s[k] = t[k]" shows "s = t" using assms by (intro fcnEqual[of s t], auto) @@ -160,7 +160,7 @@ by (auto elim: seqInSeqRange) subsection {* Sequences via @{text emptySeq} and @{text Append} *} text {* - Sequences can be built from the constructors @{text emptySeq} + Sequences can be built from the constructors @{text emptySeq} (written @{text "\\"}) and Append. *} @@ -264,7 +264,7 @@ lemma appendElt2 (*[simp]*): assumes "isASeq(s)" shows "Append(s,e)[Succ[Len(s)]] = e" using assms by (auto simp: Append_def) - + lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2, standard] lemma isAppend [intro!]: @@ -343,7 +343,7 @@ by (simp add: emptySeq_def) lemma imageAppend [simp]: assumes s: "isASeq(s)" - shows "Image(Append(s,e), A) = + shows "Image(Append(s,e), A) = (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" unfolding appendExtend[OF s] using assms by (auto elim!: inNatIntervalE, force+) @@ -451,7 +451,7 @@ lemma "\0,1\ \ \1,2,3\" by simp lemma "(\a,b\ = \c,d\) = (a=c \ b=d)" by simp ***) -text {* +text {* \tlaplus{} has a form of quantification over tuples written $\exists \langle x,y,z \rangle \in S: P(x,y,z)$. We cannot give a generic definition of this form for arbitrary tuples, but introduce input syntax @@ -496,7 +496,7 @@ text {* @{text x}, @{text y}) such that @{text "f[x] \ S"} and @{text "f[y] \ T"}. Typically, elements of such a function set will be constructed as @{text "(x :> s)@@(y :> t)"}. - This notation for sets of finite functions generalizes similar + This notation for sets of finite functions generalizes similar \tlaplus{} notation for records. Internally, the set is represented as @{text "EnumFuncSet(\x,y\, \S,T\)"}, @@ -528,7 +528,7 @@ qed lemma EnumFuncSetE [elim!]: assumes "f \ EnumFuncSet(doms, rngs)" - and "\f \ [Range(doms) \ UNION Range(rngs)]; + and "\f \ [Range(doms) \ UNION Range(rngs)]; \i \ DOMAIN doms : f[doms[i]] \ rngs[i] \ \ P" shows "P" using assms by (auto simp: EnumFuncSet_def) @@ -550,11 +550,11 @@ parse_ast_translation {* (* get_doms_ranges extracts the lists of arguments and ranges from the arms of a "domrngs" expression. The order of the ASTs is reversed. *) - fun get_doms_ranges (Ast.Appl[Ast.Constant "@domrng", d, r]) = + fun get_doms_ranges (Ast.Appl[Ast.Constant "@domrng", d, r]) = (* base case: one domain, one range *) ([d], [r]) - | get_doms_ranges (Ast.Appl[Ast.Constant "@domrngs", - Ast.Appl[Ast.Constant "@domrng", d, r], + | get_doms_ranges (Ast.Appl[Ast.Constant "@domrngs", + Ast.Appl[Ast.Constant "@domrng", d, r], pairs]) = (* one domrng, followed by remaining doms and ranges *) let val (ds, rs) = get_doms_ranges pairs @@ -645,7 +645,7 @@ text {* *} definition Product -where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : +where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : \i \ 1 .. Len(s) : f[i] \ s[i] }" lemma inProductI [intro!]: @@ -666,7 +666,7 @@ using assms by (auto simp add: Product_def) lemma inProductE [elim!]: assumes "p \ Product(s)" and "isASeq(s)" - and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; + and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; \k \ 1 .. Len(s) : p[k] \ s[k] \ \ P" shows "P" using assms by (auto simp add: Product_def) @@ -836,7 +836,7 @@ lemma inSetOfPairsI [intro]: using assms by (auto simp: setOfPairs_def) lemma inSetOfPairsE [elim!]: -- {* converse true only if $R$ is a relation *} - assumes 1: "z \ setOfPairs(R, e)" + assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" proof - @@ -846,12 +846,12 @@ proof - with pR pz show "P" by (intro 3, auto) qed -lemmas setOfPairsEqualI = +lemmas setOfPairsEqualI = setEqualI [where A = "setOfPairs(R,f)", standard,intro!] setEqualI [where B = "setOfPairs(R,f)", standard,intro!] -lemma setOfPairs_triv [simp]: - assumes s: "R \ A \ B" +lemma setOfPairs_triv [simp]: + assumes s: "R \ A \ B" shows "{ \x,y\ : \x,y\ \ R } = R" using assms by auto @@ -886,7 +886,7 @@ definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" definition rel_comp :: "[c,c] => c" (infixr "\" 75) -- {* binary relation composition *} -where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : +where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) @@ -1018,20 +1018,20 @@ by force subsubsection {* Converse relation *} -lemmas converseEqualI = +lemmas converseEqualI = setEqualI [where A = "r^-1", standard, intro!] setEqualI [where B = "r^-1", standard, intro!] -lemma converse_iff [iff]: +lemma converse_iff [iff]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" using r prodProj by (auto simp: converse_def) -lemma converseI [intro!]: +lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" unfolding converse_def by auto -lemma converseD [sym]: +lemma converseD [sym]: assumes r: "r \ A \ B" shows "\a,b\ \ r^-1 \ \b,a\ \ r" using converse_iff[OF r] by simp @@ -1053,7 +1053,7 @@ proof - show P by simp qed -lemma converse_converse [simp]: +lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" using assms prodProj by (auto elim!: converseE) @@ -1081,18 +1081,18 @@ lemma converse_mono_2: using assms prodProj by auto lemma converse_mono: - assumes r:"r \ A \ B" and s:"s \ A \ B" + assumes r:"r \ A \ B" and s:"s \ A \ B" shows "r^-1 \ s^-1 = (r \ s)" using converse_mono_1[OF r s] converse_mono_2[OF r s] by blast (* from HOL *) -lemma reflexive_converse [simp]: +lemma reflexive_converse [simp]: "r \ A \ B \ reflexive(A, r^-1) = reflexive(A,r)" unfolding reflexive_def by auto -lemma symmetric_converse [simp]: +lemma symmetric_converse [simp]: "r \ A \ B \ symmetric(r^-1) = symmetric(r)" unfolding symmetric_def by auto @@ -1100,11 +1100,11 @@ lemma antisymmetric_converse [simp]: "r \ A \ B \ antisymmetric(r^-1) = antisymmetric(r)" unfolding antisymmetric_def by auto -lemma transitive_converse [simp]: +lemma transitive_converse [simp]: "r \ A \ B \ transitive(r^-1) = transitive(r)" unfolding transitive_def by auto -lemma symmetric_iff_converse_eq: +lemma symmetric_iff_converse_eq: assumes r:"r \ A \ B" shows "symmetric(r) = (r^-1 = r)" proof auto @@ -1125,7 +1125,7 @@ qed subsubsection {* Identity relation over a set *} -lemmas idEqualI = +lemmas idEqualI = setEqualI [where A = "Id(S)", standard, intro!] setEqualI [where B = "Id(S)", standard, intro!] @@ -1175,11 +1175,11 @@ unfolding rel_range_def Id_def by auto subsubsection {* Composition of relations *} -lemmas compEqualI = +lemmas compEqualI = setEqualI [where A = "r \ s", standard, intro!] setEqualI [where B = "r \ s", standard, intro!] -lemma compI [intro]: +lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "\ \a,b\ \ s; \b,c\ \ r \ \ \a,c\ \ r \ s" using assms unfolding rel_comp_def by auto @@ -1189,7 +1189,7 @@ lemma compE [elim!]: shows "(\x y z. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r \ P) \ P" using assms unfolding rel_comp_def by auto -lemma compEpair: +lemma compEpair: assumes "\a,c\ \ r \ s" and "r \ B \ C" and s: "s \ A \ B" shows "\\b. \ \a,b\ \ s; \b,c\ \ r \ \ P \ \ P" using assms by auto @@ -1204,7 +1204,7 @@ lemma rel_comp_in_prodE (*[elim]*): shows "p \ A \ C" using assms by force -lemma converse_comp: +lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "((r \ s)^-1) = (s^-1 \ r^-1)" (is "?lhs = ?rhs") proof @@ -1227,7 +1227,7 @@ next with r s show "x \ ?lhs" by (auto dest: converseSubset) qed -lemma R_comp_Id [simp]: +lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" using r proof auto @@ -1239,7 +1239,7 @@ using r proof auto with 1 show "p \ R \ Id(B)" by simp qed -lemma Id_comp_R [simp]: +lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" using r proof auto @@ -1257,7 +1257,7 @@ unfolding rel_comp_def by auto lemma rel_comp_empty2 [simp]: "R \ {} = {}" unfolding rel_comp_def by auto -lemma comp_assoc: +lemma comp_assoc: assumes t: "T \ A \ B" and s: "S \ B \ C" and r: "R \ C \ D" shows "(R \ S) \ T = R \ (S \ T)" proof @@ -1296,7 +1296,7 @@ next qed qed -lemma rel_comp_mono: +lemma rel_comp_mono: assumes hr': "r' \ r" and hs': "s' \ s" shows "(r' \ s') \ (r \ s)" unfolding rel_comp_def using subrel_dom[OF hs'] subrel_ran[OF hr'] @@ -1358,7 +1358,7 @@ by (blast intro: symmetricI dest: symmetricE) text {* Antisymmetry *} -lemma antisymmetricI [intro]: +lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" unfolding antisymmetric_def by blast diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index d8a797ea..46a9796f 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -347,7 +347,7 @@ by blast lemma zenon_notin_cup_0 : "x \\notin A \\cup B ==> x \\notin A" by blast - + lemma zenon_notin_cup_1 : "x \\notin A \\cup B ==> x \\notin B" by blast @@ -542,7 +542,7 @@ proof - have h10: "x = v | x ~= v" by (rule excluded_middle) show "except (f, v, e)[x] \\in B" using h4 h9 h5 except_def by auto - qed + qed show "except (f, v, e) \\in FuncSet (A, B)" using h6 h7 h8 FuncSet by blast qed @@ -2689,7 +2689,7 @@ proof - using h0 h1 by auto from hf - have hh0: "?gxx" + have hh0: "?gxx" by (rule zenon_disjE1 [OF _ hg1x]) have hi: "cas \\in UNION {CaseArm (<<>>[i], <<>>[i]) : i \\in DOMAIN <<>>}" @@ -2795,7 +2795,7 @@ proof - using h0 h1 by auto from hf - have hh0: "?gxx" + have hh0: "?gxx" by (rule zenon_disjE1 [OF _ hg1x]) have hi: "cas \\in UNION {CaseArm (<<>>[i], <<>>[i]) : i \\in DOMAIN <<>>}" diff --git a/isabelle/simplifier_setup.ML b/isabelle/simplifier_setup.ML index 6957e6a5..8df22c2a 100644 --- a/isabelle/simplifier_setup.ML +++ b/isabelle/simplifier_setup.ML @@ -21,8 +21,8 @@ struct fun mk_eq th = case Thm.concl_of th of Const(@{const_name \Pure.eq\},_) $ _ $ _ => th | _ $ (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _) => th RS @{thm eq_reflection} - | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ - (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => + | _ $ (Const(@{const_name \PredicateLogic.not\},_) $ + (Const(@{const_name \PredicateLogic.eq\},_) $ _ $ _)) => th RS (@{thm neq_conv_eq_False} RS @{thm eq_reflection}) | _ => (th RS @{thm eqTrueI}) RS @{thm eq_reflection} @@ -35,7 +35,7 @@ struct *) fun lift_meta_to_obj_eq ctxt i st = let - fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = + fun count_imp (Const (@{const_name \PredicateLogic.simp_implies\}, _) $ _ $ Q) = 1 + count_imp Q | count_imp _ = 0; val j = count_imp (Logic.strip_assums_concl (Thm.term_of (Thm.cprem_of st i))) @@ -59,7 +59,7 @@ struct (* Congruence rules expecting = instead of == *) fun mk_meta_cong ctxt rl = - let val rl' = Seq.hd (TRYALL (fn i => fn st => + let val rl' = Seq.hd (TRYALL (fn i => fn st => resolve_tac ctxt [lift_meta_to_obj_eq ctxt i st] i st) rl) in (rl' RS @{thm eq_reflection}) handle THM _ => @@ -76,10 +76,10 @@ struct allows us to discard "type-correctness" hypotheses in conditional rewrite rules that are added later on. *) val mksimps_pairs = - [(@{const_name \imp\}, ([@{thm mp}], false)), + [(@{const_name \imp\}, ([@{thm mp}], false)), (@{const_name \conj\}, ([@{thm conjD1}, @{thm conjD2}], false)), - (@{const_name \All\}, ([@{thm spec}], false)), - (@{const_name \TRUE\}, ([], false)), + (@{const_name \All\}, ([@{thm spec}], false)), + (@{const_name \TRUE\}, ([], false)), (@{const_name \FALSE\}, ([], false)), (@{const_name \cond\}, ([@{thm if_bool_equiv_conj} RS @{thm iffD1}], false)) ]; @@ -100,7 +100,7 @@ struct (case head_of p of Const(a,_) => (case AList.lookup (op =) pairs a of - SOME (rls, b) => + SOME (rls, b) => let val rrls = maps atoms (res_fixed rls) handle THM _ => [thm] in @@ -113,10 +113,10 @@ struct end; in atoms end - fun mksimps pairs ctxt = + fun mksimps pairs ctxt = map_filter (try mk_eq) o mk_atomize ctxt pairs o Variable.gen_all ctxt - fun unsafe_solver_tac ctxt = + fun unsafe_solver_tac ctxt = let val sol_thms = reflexive_thm :: @{thm trueI} :: @{thm refl} :: @{thm iff_refl} :: @{thm notFalseI}:: Simplifier.prems_of ctxt; @@ -141,7 +141,7 @@ struct val safe_solver = mk_solver "PredicateLogic safe" safe_solver_tac; (* Basic infrastructure for simplification *) - val PL_basic_ss = + val PL_basic_ss = empty_simpset @{context} setSSolver safe_solver setSolver unsafe_solver From b2acbd7c7818026ea335c3083108954d2d1cb0c5 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 21 Dec 2020 20:57:04 +0200 Subject: [PATCH 065/167] MAI: update `NatArith` to Isabelle2020 The addition of `fixes` statements to theorems in the theory `NatOrderings` is for making explicit the order of theorem parameters, used when instantiating a theorem with explicit substitutions. --- isabelle/NatArith.thy | 572 ++++++++++++++++++++++++++++++-------- isabelle/NatOrderings.thy | 5 + 2 files changed, 462 insertions(+), 115 deletions(-) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 6b17dc37..179a0799 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -2,25 +2,29 @@ Author: Hernan Vanzetto, LORIA Copyright (C) 2009-2011 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 + Version: Isabelle2020 Time-stamp: <2011-10-11 17:39:37 merz> *) -header {* Arithmetic (except division) over natural numbers *} +section \ Arithmetic (except division) over natural numbers \ theory NatArith imports NatOrderings begin -ML{* +named_theorems algebra_simps "algebra simplification rules" + +(* +ML \ structure AlgebraSimps = Named_Thms(val name = "algebra_simps" val description = "algebra simplification rules"); -*} +\ setup AlgebraSimps.setup +*) -text{* +text \ The rewrites accumulated in @{text algebra_simps} deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and @@ -29,10 +33,10 @@ text{* Of course it also works for fields, but it knows nothing about multiplicative inverses or division. This should be catered for by @{text field_simps}. -*} +\ -subsection {* Addition of natural numbers *} +subsection \ Addition of natural numbers \ definition addnat where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]])" @@ -40,7 +44,7 @@ where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ c" (infixl "+" 65) where nat_add_def: "\m \ Nat; n \ Nat\ \ (m + n) \ addnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma addnatType: assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" @@ -50,7 +54,7 @@ lemma addIsNat [intro!,simp]: assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" unfolding nat_add_def[OF assms] using assms addnatType by blast -text {* Base case and Inductive case *} +text \ Base case and Inductive case \ lemma addnat_0: assumes "m \ Nat" shows "addnat(m)[0] = m" @@ -92,7 +96,7 @@ lemma add_Succ_shift_nat: shows "Succ[m] + n = m + Succ[n]" using assms by simp -text {* Commutativity *} +text \ Commutativity \ lemma add_commute_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -100,14 +104,14 @@ lemma add_commute_nat [algebra_simps]: using n apply induct using assms by auto -text {* Associativity *} +text \ Associativity \ lemma add_assoc_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + p) = (m + n) + p" using assms by (induct, simp_all) -text {* Cancellation rules *} +text \ Cancellation rules \ lemma add_left_cancel_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -126,9 +130,9 @@ lemma add_left_commute_nat [algebra_simps]: using assms by(simp only: add_assoc_nat add_commute_nat) (* from HOL/OrderedGroups.thy *) -theorems add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat +lemmas add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat -text {* Reasoning about @{text "m + n = 0"}, etc. *} +text \ Reasoning about @{text "m + n = 0"}, etc. \ lemma add_is_0_nat [iff]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -157,7 +161,7 @@ using n apply (rule natCases) done -subsection {* Multiplication of natural numbers *} +subsection \ Multiplication of natural numbers \ definition multnat where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m)" @@ -165,7 +169,7 @@ where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ c" (infixl "*" 70) where nat_mult_def: "\m \ Nat; n \ Nat\ \ m * n \ multnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma multnatType: assumes "m \ Nat" shows "multnat(m) \ [Nat \ Nat]" @@ -176,16 +180,17 @@ lemma multIsNat [intro!, simp]: shows "m * n \ Nat" unfolding nat_mult_def[OF assms] using assms multnatType by blast -text {* Base case and Inductive step *} +text \ Base case and Inductive step \ lemma multnat_0: assumes "m \ Nat" shows "multnat(m)[0] = 0" unfolding multnat_def by (rule primrec_natE, auto simp: assms) -lemma mult_0_nat [simp]: -- {* neutral element *} +lemma mult_0_nat [simp]: (* -- "neutral element" *) assumes n: "n \ Nat" shows "n * 0 = 0" by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) + lemma multnat_Succ: assumes m: "m \ Nat" and n: "n \ Nat" shows "multnat(m)[Succ[n]] = multnat(m)[n] + m" @@ -213,14 +218,14 @@ lemma mult_Succ_left_nat [simp]: using n apply induct using m by (simp_all add: add_ac_nat) -text {* Commutativity *} +text \ Commutativity \ lemma mult_commute_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m * n = n * m" using assms by (induct, simp_all) -text {* Distributivity *} +text \ Distributivity \ lemma add_mult_distrib_left_nat [algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -257,13 +262,13 @@ proof - by simp qed -text {* Identity element *} +text \ Identity element \ (* used for arithmetic simplification setup *) lemma mult_1_right_nat: "a \ Nat \ a * 1 = a" by simp lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp -text {* Associativity *} +text \ Associativity \ lemma mult_assoc_nat[algebra_simps]: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -271,7 +276,7 @@ lemma mult_assoc_nat[algebra_simps]: using m apply induct using assms by (auto simp add: add_mult_distrib_right_nat) -text {* Reasoning about @{text "m * n = 0"}, etc. *} +text \ Reasoning about @{text "m * n = 0"}, etc. \ lemma mult_is_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -294,7 +299,7 @@ proof - finally show ?thesis . qed -text {* Cancellation rules *} +text \ Cancellation rules \ lemma mult_cancel1_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -346,10 +351,10 @@ lemma mult_left_commute_nat[algebra_simps]: using assms by(simp only: mult_commute_nat mult_assoc_nat) (* from HOL/OrderedGroups.thy *) -theorems mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat +lemmas mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat -subsection {* Predecessor *} +subsection \ Predecessor \ definition Pred where "Pred \ [n \ Nat \ IF n=0 THEN 0 ELSE CHOOSE x \ Nat : n=Succ[x]]" @@ -369,13 +374,13 @@ lemma Pred_in_nat [intro!, simp]: assumes "n \ Nat" shows "Pred[n] \ Nat" using assms by (cases "n", auto) -subsection {* Difference of natural numbers *} +subsection \ Difference of natural numbers \ -text {* +text \ We define a form of difference @{text "--"} of natural numbers that cuts off at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes called ``arithmetic difference''. -*} +\ definition adiffnat where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]])" @@ -383,7 +388,7 @@ where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \< definition adiff (infixl "--" 65) where nat_adiff_def: "\m \ Nat; n \ Nat\ \ (m -- n) \ adiffnat(m)[n]" -text {* Closure *} +text \ Closure \ lemma adiffnatType: assumes "m \ Nat" shows "adiffnat(m) \ [Nat \ Nat]" @@ -393,7 +398,7 @@ lemma adiffIsNat [intro!, simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ Nat" unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast -text {* Neutral element and Inductive step *} +text \ Neutral element and Inductive step \ lemma adiffnat_0: assumes "m \ Nat" shows "adiffnat(m)[0] = m" @@ -424,7 +429,7 @@ lemma adiff_0_eq_0_nat [simp]: shows "0 -- n = 0" using n apply induct by (simp_all add: adiff_Succ_nat) -text {* Reasoning about @{text "m -- m = 0"}, etc. *} +text \ Reasoning about @{text "m -- m = 0"}, etc. \ lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -437,7 +442,7 @@ lemma adiff_self_eq_0_nat [simp]: shows "m -- m = 0" using m apply induct by auto -text {* Associativity *} +text \ Associativity \ lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" @@ -450,14 +455,14 @@ lemma Succ_adiff_adiff_nat [simp]: shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" using assms by (simp add: adiff_adiff_left_nat) -text {* Commutativity *} +text \ Commutativity \ lemma adiff_commute_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i -- j -- k = i -- k -- j" using assms by (simp add: adiff_adiff_left_nat add_commute_nat) -text {* Cancellation rules *} +text \ Cancellation rules \ lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -487,7 +492,7 @@ lemma adiff_add_0_nat [simp]: using n apply induct using assms by simp_all -text {* Difference distributes over multiplication *} +text \ Difference distributes over multiplication \ lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -500,15 +505,15 @@ lemma adiff_mult_distrib2_nat [algebra_simps]: shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) --- {* NOT added as rewrites, since sometimes they are used from right-to-left *} +(* -- "NOT added as rewrites, since sometimes they are used from right-to-left" *) lemmas nat_distrib = add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat -subsection {* Additional arithmetic theorems *} +subsection \ Additional arithmetic theorems \ -subsubsection {* Monotonicity of Addition *} +subsubsection \ Monotonicity of Addition \ lemma Succ_pred [simp]: assumes n: "n \ Nat" @@ -539,9 +544,118 @@ using k apply induct using assms by simp_all lemma add_gr_0 [simp]: - assumes "n \ Nat" and "m \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" -using assms by (auto dest: nat_gt0_implies_Succ nat_not_lessD) + assumes ndom: "n \ Nat" and mdom: "m \ Nat" + shows "(m + n > 0) = (m > 0 \ n > 0)" + proof - + have s1_1: "m + n > 0 ==> m > 0 \ n > 0" + proof - + assume s1_1_asm: "m + n > 0" + have s2_1: "m \ 0 \ n \ 0" + proof - + have s3_1: "m = 0 \ n = 0 ==> FALSE" + proof - + assume s3_1_asm: "m = 0 \ n = 0" + have s4_1: "m + n = 0" + using s3_1_asm ndom mdom add_is_0_nat by auto + have s4_2: "m + n \ 0" + unfolding less_def using s1_1_asm by auto + show "FALSE" + using s4_1 s4_2 by auto + qed + show ?thesis + using s3_1 by auto + qed + show "m > 0 \ n > 0" + using s2_1 ndom mdom nat_gt0_not0 by auto + qed + have s1_2: "m > 0 \ n > 0 ==> m + n > 0" + proof - + assume s1_2_asm: "m > 0 \ n > 0" + have s2_1: "m > 0 ==> m + n > 0" + proof - + assume s2_1_asm: "m > 0" + have s3_1: "m + n > 0 + n" + proof - + have s4_1: "0 < m" + using s2_1_asm by auto + have s4_2: "(0 + n < m + n) = (0 < m)" + using ndom mdom zeroIsNat nat_add_right_cancel_less + by blast + have s4_3: "0 + n < m + n" + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + have s3_2: "0 + n = n" + using ndom add_0_nat add_commute_nat zeroIsNat + by auto + have s3_3: "n >= 0" + using ndom nat_zero_leq by auto + have s3_4: "m + n >= 0 + n" + using s3_1 by (auto simp: less_def) + have s3_5: "m + n >= n" + using s3_4 s3_2 by auto + have s3_6: "m + n >= 0" + using s3_5 s3_3 nat_leq_trans + zeroIsNat ndom mdom addIsNat by auto + have s3_7: "m + n = 0 ==> FALSE" + proof - + assume s3_7_asm: "m + n = 0" + have s4_1: "m = 0 \ n = 0" + using s3_7_asm mdom ndom add_is_0_nat + by auto + show "FALSE" + using s4_1 s2_1_asm by (auto simp: less_def) + qed + show "m + n > 0" + unfolding less_def + using s3_6 s3_7 by auto + qed + have s2_2: "n > 0 ==> m + n > 0" + proof - + assume s2_2_asm: "n > 0" + have s3_1: "m + n > m + 0" + proof - + have s4_1: "0 < n" + using s2_2_asm by auto + have s4_2: "(m + 0 < m + n) = (0 < n)" + using ndom mdom zeroIsNat nat_add_left_cancel_less + by blast + have s4_3: "m + 0 < m + n" + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + have s3_2: "m + 0 = m" + using mdom add_0_nat by auto + have s3_3: "m >= 0" + using mdom nat_zero_leq by auto + have s3_4: "m + n >= m + 0" + using s3_1 by (auto simp: less_def) + have s3_5: "m + n >= m" + using s3_4 s3_2 by auto + have s3_6: "m + n >= 0" + using s3_5 s3_3 nat_leq_trans + zeroIsNat ndom mdom addIsNat by auto + have s3_7: "m + n = 0 ==> FALSE" + proof - + assume s3_7_asm: "m + n = 0" + have s4_1: "m = 0 \ n = 0" + using s3_7_asm mdom ndom add_is_0_nat + by auto + show "FALSE" + using s4_1 s2_2_asm by (auto simp: less_def) + qed + show "m + n > 0" + unfolding less_def + using s3_6 s3_7 by auto + qed + show "m + n > 0" + using s1_2_asm s2_1 s2_2 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed lemma less_imp_Succ_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -571,11 +685,11 @@ apply (induct n p rule: diffInduct) using assms by simp_all -subsubsection {* (Partially) Ordered Groups *} +subsubsection \ (Partially) Ordered Groups \ --- {* The two following lemmas are just ``one half'' of +(* -- "The two following lemmas are just ``one half'' of @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above. *} + proved above." *) lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" @@ -586,7 +700,7 @@ lemma add_leq_right_mono: shows "a \ b \ a + c \ b + c" using assms by simp -text {* non-strict, in both arguments *} +text \ non-strict, in both arguments \ lemma add_leq_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a \ b \ c \ d \ a + c \ b + d" @@ -596,18 +710,43 @@ using assms nat_leq_trans[of "a + c" "b + c" "b + d"] by simp --- {* Similar for strict less than. *} +(* -- "Similar for strict less than." *) lemma add_less_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ c + a < c + b" -using assms by simp + proof - + have s1_1: "a <= b ==> c + a <= c + b" + using a b c add_leq_left_mono by auto + have s1_2: "a \ b ==> (c + a) \ (c + b)" + using a b c add_left_cancel_nat by auto + have s1_3: "a < b ==> c + a < c + b" + proof - + assume s1_3_asm: "a < b" + have s2_1: "a <= b" + using s1_3_asm by (auto simp: less_def) + have s2_2: "a \ b" + using s1_3_asm by (auto simp: less_def) + have s2_3: "c + a <= c + b" + using s1_1 s2_1 by auto + have s2_4: "(c + a) \ (c + b)" + using s1_2 s2_2 by auto + show "c + a < c + b" + using s2_3 s2_4 by (auto simp: less_def) + qed + show "a < b \ c + a < c + b" + using a b c s1_3 by auto + qed +(* +using assms add_leq_left_mono[OF "a" "b" "c"] add_left_cancel_nat[OF "c" "b" "a"] + by (auto simp: nat_less_Succ_iff_leq nat_gt0_not0) +*) lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" -using assms by simp +using assms add_less_left_mono add_commute_nat by auto -text{* Strict monotonicity in both arguments *} +text \ Strict monotonicity in both arguments \ lemma add_less_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" shows "a < b \ c < d \ a + c < b + d" @@ -671,14 +810,36 @@ lemma trans_leq_add2: using assms by (auto elim: nat_leq_trans) lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" + assumes ij: "i < j" and idom: "i \ Nat" and + jdom: "j \ Nat" and mdom: "m \ Nat" shows "i < j + m" -using assms by (auto elim: nat_less_leq_trans) + proof - + have s1_1: "i <= i + m" + proof - + have s2_1: "0 <= m" + using mdom nat_zero_leq by auto + have s2_2: "i + 0 <= i + m" + using idom zeroIsNat mdom s2_1 nat_add_left_cancel_leq by auto + have s2_3: "i + 0 = i" + using idom add_0_nat by auto + show ?thesis + using s2_2 s2_3 by auto + qed + have s1_2: "i + m < j + m" + using ij idom jdom mdom add_less_right_mono by auto + have s1_3: "i + m \ Nat" + using idom mdom addIsNat by auto + have s1_4: "j + m \ Nat" + using jdom mdom addIsNat by auto + show ?thesis + using s1_1 s1_2 idom s1_3 s1_4 nat_leq_less_trans + by auto + qed lemma trans_less_add2: assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" -using assms by (auto elim: nat_less_leq_trans) +using assms trans_less_add1 add_commute_nat by auto lemma add_lessD1: assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" @@ -711,15 +872,32 @@ lemma add_leqE: using assms by (blast dest: add_leqD1 add_leqD2) lemma leq_add_left_false [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "n \ 0" + assumes mdom: "m \ Nat" and ndom: "n \ Nat" and n0: "n \ 0" shows "m + n \ m = FALSE" -using assms nat_leq_less[of "m + n" m] add_eq_self_zero_nat[OF m n] by auto + proof - + have s1_1: "m + n \ m = (m + n < m \ m + n = m)" + proof - + define p where "p \ m + n" + have s2_1: "p \ m = (p < m \ p = m)" + using mdom nat_leq_less by auto + show ?thesis + using s2_1 by (auto simp: p_def) + qed + have s1_2: "(m + n = m) = (n = 0)" + using mdom ndom add_eq_self_zero_nat[of "m" "n"] by auto + have s1_3: "(m + n <= m) = (m + n < m)" + using n0 s1_1 s1_2 by auto + have s1_4: "(m + n < m) = FALSE" + using mdom ndom not_add_less1[of "m" "n"] by auto + show ?thesis + using s1_3 s1_4 by auto + qed -subsubsection {* More results about arithmetic difference *} +subsubsection \ More results about arithmetic difference \ -text {* Addition is the inverse of subtraction: - if @{term "n \ m"} then @{term "n + (m -- n) = m"}. *} +text \ Addition is the inverse of subtraction: + if @{term "n \ m"} then @{term "n + (m -- n) = m"}. \ lemma add_adiff_inverse: assumes "m \ Nat" and "n \ Nat" shows "\(m < n) \ n + (m -- n) = m" @@ -752,7 +930,7 @@ lemma adiff_leq_self [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ m" apply (induct m n rule: diffInduct) -using assms by simp_all +using assms by auto lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -771,10 +949,14 @@ qed lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" -proof - - from j n have "j -- n \ Nat" "j -- n \ j" by simp+ - with j k jk show ?thesis by (auto elim: nat_leq_less_trans) -qed + proof - + have s1_1: "j -- n \ Nat" + using j n by simp+ + have s1_2: "j -- n \ j" + using j n by simp+ + show ?thesis + using j k jk s1_1 s1_2 nat_leq_less_trans by auto + qed lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" @@ -827,15 +1009,17 @@ lemma zero_less_adiff [simp]: by (induct m n rule: diffInduct, simp_all add: assms) lemma less_imp_add_positive: - assumes "i < j" and i: "i \ Nat" and j: "j \ Nat" + assumes ij: "i < j" and i: "i \ Nat" and j: "j \ Nat" shows "\k \ Nat: 0 < k \ i + k = j" proof - - from assms have "i \ j" by auto - with i j have "i + (j -- i) = j" by simp - moreover - from assms have "j -- i \ Nat" "0 < j -- i" by simp+ - ultimately - show ?thesis by blast + have s1_1: "i + (j -- i) = j" + using i j ij le_add_adiff_inverse by (auto simp: less_def) + have s1_2: "j -- i \ Nat" + using assms by simp+ + have s1_3: "0 < j -- i" + using assms zero_less_adiff by auto + show ?thesis + using s1_1 s1_2 s1_3 by blast qed lemma leq_adiff_right_add_left: @@ -866,7 +1050,7 @@ using p h apply (induct) using n by auto -subsubsection {* Monotonicity of Multiplication *} +subsubsection \ Monotonicity of Multiplication \ (* from HOL/Ring_and_Field.thy *) @@ -882,7 +1066,7 @@ lemma mult_leq_right_mono: using c apply induct using 1 a b by (simp_all add: add_leq_mono add_commute_nat) -text {* @{text "\"} monotonicity, BOTH arguments *} +text \ @{text "\"} monotonicity, BOTH arguments \ lemma mult_leq_mono: assumes 1: "i \ j" "k \ l" @@ -894,29 +1078,95 @@ using assms nat_leq_trans[of "i * k" "j * k" "j * l"] by simp -text {* strict, in 1st argument *} +text \ strict, in 1st argument \ lemma mult_less_left_mono: - assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" + assumes ij: "i < j" and k0: "0 < k" and + idom: "i \ Nat" and jdom: "j \ Nat" and kdom: "k \ Nat" shows "k * i < k * j" -using 1 -proof (auto simp: nat_gt0_iff_Succ[OF k]) - fix m - assume m: "m \ Nat" and "i < j" - with m i j show "Succ[m] * i < Succ[m] * j" - by (induct m, simp_all add: add_less_mono) -qed + proof - + have s1_1: "Succ[0] * i < Succ[0] * j" + using mult_1_left_nat ij idom jdom by auto + have s1_2: "\ n. \ + n \ Nat; + Succ[n] * i < Succ[n] * j + \ \ + Succ[Succ[n]] * i < Succ[Succ[n]] * j" + proof - + fix "n" :: "c" + assume s1_2_ndom: "n \ Nat" + assume s1_2_induct: "Succ[n] * i < Succ[n] * j" + have s2_1: "Succ[n] * i + i < Succ[n] * j + j" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "Succ[n] * i \ Nat" + using s3_1 idom multIsNat by auto + have s3_3: "Succ[n] * j \ Nat" + using s3_1 jdom multIsNat by auto + show ?thesis + using s3_2 s3_3 idom jdom s1_2_induct ij + add_less_mono by auto + qed + have s2_2: "i * Succ[n] + i < j * Succ[n] + j" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "Succ[n] * i = i * Succ[n]" + using s3_1 idom mult_commute_nat by auto + have s3_3: "Succ[n] * j = j * Succ[n]" + using s3_1 jdom mult_commute_nat by auto + show ?thesis + using s2_1 s3_2 s3_3 by auto + qed + have s2_3: "i * Succ[Succ[n]] < j * Succ[Succ[n]]" + proof - + have s3_1: "Succ[n] \ Nat" + using s1_2_ndom succIsNat by auto + have s3_2: "i * Succ[Succ[n]] = i * Succ[n] + i" + using idom s3_1 multnat_Succ by (auto simp: nat_mult_def) + have s3_3: "j * Succ[Succ[n]] = j * Succ[n] + j" + using jdom s3_1 multnat_Succ by (auto simp: nat_mult_def) + show ?thesis + using s2_2 s3_2 s3_3 by auto + qed + have s2_4: "i * Succ[Succ[n]] = Succ[Succ[n]] * i" + proof - + have s3_1: "Succ[Succ[n]] \ Nat" + using s1_2_ndom succIsNat by auto + show ?thesis + using s3_1 idom mult_commute_nat by auto + qed + have s2_5: "j * Succ[Succ[n]] = Succ[Succ[n]] * j" + proof - + have s3_1: "Succ[Succ[n]] \ Nat" + using s1_2_ndom succIsNat by auto + show ?thesis + using s3_1 jdom mult_commute_nat by auto + qed + show "Succ[Succ[n]] * i < Succ[Succ[n]] * j" + using s2_3 s2_4 s2_5 by auto + qed + have s1_3: "\ n \ Nat: Succ[n] * i < Succ[n] * j" + using s1_1 s1_2 + natInduct[of "\ q. Succ[q] * i < Succ[q] * j"] by auto + show ?thesis + using s1_3 nat_gt0_iff_Succ[of "k"] k0 kdom by auto + qed lemma mult_less_right_mono: assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "i * k < j * k" -using 1 -proof (auto simp: nat_gt0_iff_Succ[OF k]) - fix m - assume m: "m \ Nat" and "i < j" - with m i j show "i * Succ[m] < j * Succ[m]" - by (induct m, simp_all add: add_less_mono) -qed + proof - + have s1_1: "k * i < k * j" + using 1 i j k mult_less_left_mono by auto + have s1_2: "k * i = i * k" + using i k mult_commute_nat by auto + have s1_3: "k * j = j * k" + using j k mult_commute_nat by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed lemma nat_0_less_mult_iff [simp]: assumes i: "i \ Nat" and j: "j \ Nat" @@ -929,7 +1179,7 @@ done lemma one_leq_mult_iff (*[simp]*): assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 \ m * n) = (1 \ m \ 1 \ n)" -using assms by simp +using assms nat_gt0_not0 by simp lemma mult_less_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" @@ -966,47 +1216,140 @@ next qed lemma mult_less_self_left [dest]: - assumes less: "n*k < n" and n: "n \ Nat" and k: "k \ Nat" + assumes less: "n * k < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" -using k assms by (cases, auto) + proof - + have s1_1: "1 \ k ==> FALSE" + proof - + assume s1_1_asm: "1 \ k" + have s2_1: "n * 1 \ n * k" + using n k oneIsNat s1_1_asm mult_leq_left_mono[of "1" "k" "n"] + by auto + have s2_2: "n \ n * k" + proof - + have s3_1: "n * 1 = n" + using n mult_1_right_nat by auto + show ?thesis + using s2_1 s3_1 by auto + qed + define p where "p \ n" + define q where "q \ n * k" + have s2_3: "q < p" + unfolding p_def q_def using less by auto + have s2_4: "n < p" + proof - + have s3_1: "q \ Nat" + unfolding q_def using n k multIsNat by auto + have s3_2: "p \ Nat" + unfolding p_def using n by auto + have s3_3: "n \ q" + unfolding q_def using s2_2 by auto + show ?thesis + using s3_1 s3_2 n s3_3 s2_3 + nat_leq_less_trans[of "n" "q" "p"] + by auto + qed + have s2_5: "n \ n" + using s2_4 by (auto simp: less_def p_def) + show "FALSE" + using s2_5 by auto + qed + have s1_2: "\ (1 \ k)" + using s1_1 by auto + show ?thesis + using s1_2 k nat_not_leq_one[of "k"] by auto + qed lemma mult_less_self_right [dest]: assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" -using k assms by (cases, auto) + proof - + have s1_1: "n * k < n" + proof - + have s2_1: "k * n = n * k" + using n k mult_commute_nat by auto + show ?thesis + using less s2_1 by auto + qed + show ?thesis + using s1_1 n k mult_less_self_left by auto + qed lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" -using assms proof (auto simp: mult_leq_left_mono nat_neq0_conv[simplified]) - assume 1: "k*m \ k*n" and 2: "0 < k" - show "m \ n" - proof (rule contradiction) - assume "\(m \ n)" - with 2 m n k have "k*n < k*m" by (simp add: nat_not_order_simps mult_less_left_mono) - with m n k have "\(k*m \ k*n)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed + proof - + have s1_1: "(k * m \ k * n) ==> (k = 0 \ m \ n)" + proof - + assume s1_1_asm: "k * m \ k * n" + have s2_1: "\ (k = 0) ==> m \ n" + proof - + assume s2_1_asm: "\ (k = 0)" + have s3_1: "k > 0" + using k nat_gt0_not0 s2_1_asm by auto + have s3_2: "\(m \ n) ==> FALSE" + proof - + assume s3_2_asm: "\(m \ n)" + have s2_1: "n < m" + using s3_2_asm m n nat_not_leq[of "m" "n"] by auto + have s2_2: "k * n < k * m" + using s3_1 s2_1 m n k mult_less_left_mono[of "n" "m" "k"] by auto + with m n k have s2_3: "\ (k * m \ k * n)" + by (simp add: nat_not_order_simps) + show "FALSE" + using s2_2 s2_3 s1_1_asm by auto + qed + show "m \ n" + using s3_2 by auto + qed + show "k = 0 \ m \ n" + using s2_1 by auto + qed + have s1_2: "(k = 0 \ m \ n) ==> (k * m \ k * n)" + proof - + assume s1_2_asm: "k = 0 \ m \ n" + have s2_1: "k = 0 ==> k * m \ k * n" + proof - + assume s2_1_asm: "k = 0" + have s3_1: "k * m = 0" + using s2_1_asm m mult_0_left_nat by auto + have s3_2: "k * n = 0" + using s2_1_asm n mult_0_left_nat by auto + show "k * m \ k * n" + using s3_1 s3_2 by auto + qed + have s2_2: "m \ n ==> k * m \ k * n" + proof - + assume s2_2_asm: "m \ n" + show "k * m \ k * n" + using s2_2_asm m n k mult_leq_left_mono by auto + qed + show "k * m \ k * n" + using s1_2_asm s2_1 s2_2 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" -using assms proof (auto simp: mult_leq_right_mono nat_neq0_conv[simplified]) - assume 1: "m*k \ n*k" and 2: "0 < k" - show "m \ n" - proof (rule contradiction) - assume "\(m \ n)" - with 2 m n k have "n*k < m*k" by (simp add: nat_not_order_simps mult_less_right_mono) - with m n k have "\(m*k \ n*k)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed + proof - + have s1_1: "(k * m \ k * n) = (k = 0 \ m \ n)" + using m n k mult_leq_cancel_left by auto + have s1_2: "k * m = m * k" + using k m mult_commute_nat by auto + have s1_3: "k * n = n * k" + using k n mult_commute_nat by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" -using assms by (simp del: mult_Succ_left_nat) +using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] + nat_zero_less_Succ[of "k"] by auto lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" @@ -1023,7 +1366,7 @@ lemma nat_leq_cube: shows "m \ m * (m * m)" using m by (cases, auto) -text {* Lemma for @{text gcd} *} +text \ Lemma for @{text gcd} \ lemma mult_eq_self_implies_10: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") @@ -1033,5 +1376,4 @@ proof - finally show ?thesis . qed - end diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index 5009aa4e..75b44e09 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -207,6 +207,8 @@ using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) text \ Reduce @{text "\"} to @{text "<"}. \ lemma nat_leq_less: + fixes "m" :: "c" + fixes "n" :: "c" assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" using assms by (auto simp: less_def) @@ -296,6 +298,9 @@ proof - qed lemma nat_leq_less_trans [trans]: + fixes "k" :: "c" + fixes "m" :: "c" + fixes "n" :: "c" assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) From e360a995c9c5defafbd8c14b5ab549442db22f99 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:03:55 +0200 Subject: [PATCH 066/167] MAI: port sectioning and text of theory `NatDivision` to Isabelle2020 --- isabelle/NatDivision.thy | 36 ++++++++++++++++++------------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index abef1aed..fcc8a7f2 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:39:56 merz> *) -header {* The division operators div and mod on Naturals *} +section \ The division operators div and mod on Naturals \ theory NatDivision imports NatArith Tuples begin -subsection {* The divisibility relation *} +subsection \ The divisibility relation \ definition dvd (infixl "dvd" 50) where "a \ Nat \ b \ Nat \ b dvd a \ (\k \ Nat : a = b * k)" @@ -151,9 +151,9 @@ proof - with a b' c' show ?thesis by blast qed -subsection {* Division on @{const Nat} *} +subsection \ Division on @{const Nat} \ -text {* +text \ We define division and modulo over @{const Nat} by means of a characteristic relation with two input arguments @{term "m"}, @{term "n"} and two output arguments @@ -162,13 +162,13 @@ text {* The following definition works for natural numbers, but also for possibly negative integers. Obviously, the second disjunct cannot hold for natural numbers. -*} +\ definition divmod_rel where "divmod_rel(m,n,q,r) \ m = q * n + r \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" -text {* @{const divmod_rel} is total if $n$ is non-zero. *} +text \ @{const divmod_rel} is total if $n$ is non-zero. \ lemma divmod_rel_ex: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" @@ -201,7 +201,7 @@ proof - with pos that show ?thesis by (auto simp: divmod_rel_def) qed -text {* @{const divmod_rel} has unique solutions in the natural numbers. *} +text \ @{const divmod_rel} has unique solutions in the natural numbers. \ lemma divmod_rel_unique_div: assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" and m: "m \ Nat" and n: "n \ Nat" @@ -242,10 +242,10 @@ proof - with assms show ?thesis by (auto simp: divmod_rel_def) qed -text {* +text \ We instantiate divisibility on the natural numbers by means of @{const divmod_rel}: -*} +\ definition divmodNat where "divmodNat(m,n) \ CHOOSE z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" @@ -282,9 +282,9 @@ proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) qed -text {* +text \ We now define division and modulus over natural numbers. -*} +\ definition div (infixr "div" 70) where div_nat_def: "\m \ Nat; n \ Nat\ \ m div n \ divmodNat(m,n)[1]" @@ -331,7 +331,7 @@ lemma mod_nat_less_divisor: shows "m mod n < n" using assms divmod_rel_div_mod_nat[OF assms] by (simp add: divmod_rel_def) -text {* ``Recursive'' computation of @{text div} and @{text mod}. *} +text \ ``Recursive'' computation of @{text div} and @{text mod}. \ lemma divmodNat_base: assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" @@ -386,7 +386,7 @@ proof - qed -text {* The ''recursion'' equations for @{text div} and @{text mod} *} +text \ The ''recursion'' equations for @{text div} and @{text mod} \ lemma div_nat_less [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" @@ -419,7 +419,7 @@ using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] by simp -subsection {* Facts about @{const div} and @{const mod} *} +subsection \ Facts about @{const div} and @{const mod} \ lemma mod_div_nat_equality [simp]: assumes "m \ Nat" and "n \ Nat" and "0 < n" @@ -691,13 +691,13 @@ proof - qed (*** TODO -text {* Addition respects modular equivalence. *} +text \ Addition respects modular equivalence. \ -text {* Multiplication respects modular equivalence. *} +text \ Multiplication respects modular equivalence. \ -text {* Negation respects modular equivalence. *} +text \ Negation respects modular equivalence. \ -text {* Subtraction respects modular equivalence. *} +text \ Subtraction respects modular equivalence. \ *) end From 771a44fa6d69850cfcbbb8c3a0cb4e4319f34d91 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:04:27 +0200 Subject: [PATCH 067/167] API: change numerals to definitions in theory `Peano` because abbreviations do not work well with Isabelle (nested applications of `Succ`). --- isabelle/Peano.thy | 36 ++++++++++++++++++++++++++++++++++-- 1 file changed, 34 insertions(+), 2 deletions(-) diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 8bbc8cac..1f19fe45 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -154,6 +154,38 @@ where "Nat \ DOMAIN Succ" definition zero :: "c" ("0") where "0 \ CHOOSE Z : PeanoAxioms(Nat, Z, Succ)" +abbreviation "one \ Succ[0]" + notation "one" ("1") +definition two :: "c" ("2") + where "2 \ Succ[1]" +definition three :: "c" ("3") + where "3 \ Succ[2]" +definition four :: "c" ("4") + where "4 \ Succ[3]" +definition five :: "c" ("5") + where "five \ Succ[4]" +definition six :: "c" ("6") + where "six \ Succ[5]" +definition seven :: "c" ("7") + where "seven \ Succ[6]" +definition eight :: "c" ("8") + where "eight \ Succ[7]" +definition nine :: "c" ("9") + where "nine \ Succ[8]" +definition ten :: "c" ("10") + where "ten \ Succ[9]" +definition eleven :: "c" ("11") + where "eleven \ Succ[10]" +definition twelve :: "c" ("12") + where "twelve \ Succ[11]" +definition thirteen :: "c" ("13") + where "thirteen \ Succ[12]" +definition fourteen :: "c" ("14") + where "fourteen \ Succ[13]" +definition fifteen :: "c" ("15") + where "fifteen \ Succ[14]" + +(* abbreviation "one \ Succ[0]" notation "one" ("1") abbreviation "two \ Succ[1]" @@ -184,7 +216,7 @@ abbreviation "fourteen \ Succ[13]" notation "fourteen" ("14") abbreviation "fifteen \ Succ[14]" notation "fifteen" ("15") - +*) lemma peanoNatZeroSucc: "PeanoAxioms(Nat, 0, Succ)" proof - @@ -223,7 +255,7 @@ lemma oneIsNat [intro!,simp]: "1 \ Nat" by simp lemma twoIsNat [intro!,simp]: "2 \ Nat" -by simp +unfolding two_def by simp lemma [simp]: assumes "n \ Nat" From b99dbcc7b3ab8442df6275e73bbce8062f1b3e55 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 18:20:09 +0200 Subject: [PATCH 068/167] ENH: add theory `NewIntegers` --- isabelle/NewIntegers.thy | 12578 +++++++++++++++++++++++++++++++++++++ 1 file changed, 12578 insertions(+) create mode 100644 isabelle/NewIntegers.thy diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy new file mode 100644 index 00000000..6c7ae08c --- /dev/null +++ b/isabelle/NewIntegers.thy @@ -0,0 +1,12578 @@ +(* An attempt at defining the integers. + +The steps in Isabelle/Isar are named `si_j` where `s` stands for "step", +`i` is the proof level, and `j` the step number. + +Author: Ioannis Filippidis +All rights reserved. Licensed under 2-clause BSD. +*) +theory NewIntegers +imports + NatArith + (*NatDivision + Tuples*) +begin + +(* +Properties to prove: + typeness, + commutativity, associativity, distributivity + transitivity, symmetry, antisymmetry, reflexivity, + monotonicity +*) + +(*----------------------------------------------------------------------------*) +(* Overriding notation. *) + +no_notation NatOrderings.leq (infixl "<=" 50) +no_notation NatOrderings.geq (infixl ">=" 50) +no_notation NatOrderings.less (infixl "<" 50) +no_notation NatOrderings.greater (infixl ">" 50) +no_notation NatOrderings.natInterval ("(_ .. _)" [90,90] 70) + +no_notation NatArith.arith_add (infixl "+" 65) +(* no_notation NatArith.adiff (infixl "--" 65) *) +no_notation NatArith.mult (infixl "*" 70) + +(* no_notation NatDivision.dvd (infixl "dvd" 50) *) + + +(*----------------------------------------------------------------------------*) +(* The minus operator. *) + + +(* +The minus operator can be defined, instead of axiomatized, as follows. + +minus(x) == + LET singletons == {<>: u \in Nat} + IN + IF x \in Nat THEN + <> + ELSE IF x \in Singletons THEN + CHOOSE u \in Nat: x = <> + ELSE + CHOOSE u \notin (Nat \cup Singletons): TRUE +*) +consts + "minus" :: "[c] \ c" ("-._" [75] 75) + +axiomatization where + neg0 [simp]: "-.0 = 0" +and + neg_neg: "n \ Nat \ -.n \ Nat ==> + -.-.n = n" +and + neg_not_in_nat: + "n \ (Nat \ {0}) ==> + -.n \ Nat" + + +(*----------------------------------------------------------------------------*) +(* The set of integers. *) + +definition Int +where "Int \ + Nat \ {-.n: n \ Nat}" + + +definition negNat :: "c" +where "negNat \ {-.n: n \ Nat}" + + +(*----------------------------------------------------------------------------*) +(* Successor and predecessor on integers. *) + +definition iSucc :: "c" +where "iSucc \ + [i \ Int \ IF i \ Nat THEN Succ[i] + ELSE -.Pred[-.i]]" + + +definition iPred :: "c" +where "iPred \ + [i \ Int \ IF i \ Nat \ {0} + THEN Pred[i] + ELSE -.Succ[-.i]]" + + +(*----------------------------------------------------------------------------*) +(* Recursive definitions on the integers. *) + +(* TODO: replace axiom with proof of existence. *) +axiomatization where + primrec_int: " + \ f: + isAFcn(f) \ + DOMAIN f = Int \ + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n])) + " + +(*----------------------------------------------------------------------------*) +(* Arithmetic operators. *) + + +(* Incrementing and decrementing function. *) +definition addint :: "[c] \ c" +where "addint(m) \ + CHOOSE g \ [Int -> Int]: + g[0] = m \ + (\ n \ Nat: + g[Succ[n]] = iSucc[g[n]] \ + g[-.Succ[n]] = iPred[g[-.n]])" + + +(* Addition on integers. *) +definition add :: "[c, c] \ c" (infixl "+" 65) +where "add(m, n) \ addint(m)[n]" + + +(* Subtraction on integers. *) +definition diff :: "[c, c] \ c" (infixl "-" 65) +where "diff(m, n) \ add(m, -.n)" + + +(* Adding and subtracting function. *) +definition multint :: "[c] \ c" +where "multint(m) \ + CHOOSE g \ [Int -> Int]: + g[0] = 0 \ + (\ n \ Nat: + g[Succ[n]] = add(g[n], m) \ + g[-.Succ[n]] = add(g[-.n], -.m))" + + +(* Multiplication on integers. *) +definition mult :: "[c,c] \ c" (infixl "*" 70) +where "mult(m, n) \ multint(m)[n]" + + +(* Divisibility on the integers. *) +definition dvd :: "[c, c] \ c" (infixl "dvd" 50) +where "b dvd a \ + \ k \ Int: a = mult(b, k)" + + +(* Order on the integers. +This definition is motivated by proofs, where equality is easier to +reason about (adding on both sides, multiplying both sides). +*) +definition leq :: "[c,c] \ c" (infixl "<=" 50) +where "leq(m, n) \ + \ k \ Nat: + add(m, k) = n" + + +abbreviation (input) + geq :: "[c, c] \ c" (infixl ">=" 50) + where "geq(x, y) \ leq(y, x)" + + +definition less :: "[c, c] \ c" (infixl "<" 50) +where "less(a, b) \ leq(a, b) + \ (a \ b)" + + +abbreviation (input) + greater :: "[c, c] \ c" (infixl ">" 50) + where "greater(x, y) \ less(y, x)" + + +definition interval :: "[c, c] \ c" ("(_ .. _)" [90,90] 70) +where "m .. n \ + {k \\in Int: m <= k \ k <= n}" + +(*----------------------------------------------------------------------------*) + + +theorem minus_involutive[simp]: + assumes "n \ Int" + shows "-.-.n = n" + proof - + have s1_1: "n \ Nat \ (\ k \ Nat: n = -.k)" + using assms by (auto simp: Int_def) + have s1_2: "n \ Nat ==> -.-.n = n" + using neg_neg by auto + have s1_3: "n \ Nat ==> -.-.n = n" + proof - + { + assume s2_1: "n \ Nat" + have s2_2: "\ k \ Nat: n = -.k" + using s1_1 s2_1 by auto + have s2_3: "\ k \ Nat: -.n = -.-.k" + using s2_2 by auto + have s2_4: "\ k \ Nat: k = -.n" + using s2_3 neg_neg by auto + have s2_5: "-.n \ Nat" + using s2_4 by auto + have s2_6: "-.-.n = n" + using s2_5 neg_neg by auto + } + from this show "n \ Nat ==> -.-.n = n" by auto + qed + show "-.-.n = n" using s1_2 s1_3 by iprover + qed + + +theorem minus_injective[dest]: + assumes "-.n = -.m" and "n \ Int \ m \ Int" + shows "n = m" + proof - + have "<1>1": "-.-.n = -.-.m" + using assms by simp + have "<1>2": "-.-.n = n" + proof - + have "<2>1": "n \ Int ==> -.-.n = n" + using minus_involutive by auto + have "<2>2": "n \ Int" + using assms by auto + show "-.-.n = n" + using "<2>1" "<2>2" by auto + qed + have "<1>3": "-.-.m = m" + proof - + have "<2>1": "m \ Int ==> -.-.m = m" + using minus_involutive by auto + have "<2>2": "m \ Int" + using assms by auto + show "-.-.m = m" + using "<2>1" "<2>2" by auto + qed + show "n = m" + using "<1>1" "<1>2" "<1>3" by auto + qed + + +theorem minus_eq: + assumes "x = y" + shows "-.x = -.y" + using assms by auto + + +theorem neg_succ_not_in_nat[simp]: + "n \ Nat ==> + -.(Succ[n]) \ Nat" + proof - + { + fix "n" + assume "<2>1": "n \ Nat" + have "<2>2": "Succ[n] \ (Nat \ {0})" + using "<2>1" by auto + have "-.(Succ[n]) \ Nat" + using "<2>2" neg_not_in_nat by auto + } + from this show "n \ Nat ==> + -.(Succ[n]) \ Nat" + by auto + qed + + +theorem neg0_iff_eq0[simp]: + "(-.n = 0) = (n = 0)" + proof auto + assume "<1>1": "-.n = 0" + have "<1>2": "-.n \ Nat" + using "<1>1" by auto + have "<1>3": "-.-.n = n" + using "<1>2" neg_neg by auto + have "<1>4": "-.-.n = 0" + using "<1>1" by auto + show "n = 0" + using "<1>3" "<1>4" by simp + qed + + +theorem neg0_imp_eq0[dest]: + "(-.n = 0) ==> (n = 0)" + by simp + + +theorem not_neg0_imp_not0[dest]: + "(-.n \ 0) ==> (n \ 0)" + by simp + + +(*----------------------------------------------------------------------------*) +(* Properties of the set `Int`. *) + + +theorem nat_in_int[simp]: + "n \ Nat ==> n \ Int" + by (simp add: Int_def) + + +theorem minus_nat_in_int: + "n \ Nat ==> -.n \ Int" + by (auto simp: Int_def) + + +theorem minus_nat_in_neg_nat: + "n \ Nat ==> -.n \ negNat" + by (auto simp: negNat_def) + + +theorem neg_nat_in_int: + "n \ negNat ==> n \ Int" + by (auto simp: negNat_def Int_def) + + +theorem int_disj: + "n \ Int ==> + n \ Nat + \ n \ {-.n: n \ Nat}" + by (auto simp: Int_def) + + +theorem int_disj_nat_negnat: + "n \ Int ==> + n \ Nat + \ n \ negNat" + unfolding negNat_def + using int_disj by auto + + +theorem int_union_nat_negnat: + "Int = Nat \ negNat" + using int_disj_nat_negnat nat_in_int + neg_nat_in_int + by auto + + +theorem int_union_nat_0_negnat: + "Int = (Nat \ {0}) \ negNat" + proof - + have s1_1: "Int = Nat \ negNat" + using int_union_nat_negnat by auto + have s1_2: "0 \ negNat" + unfolding negNat_def + using zeroIsNat neg0 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem int_union_nat_negnat_0: + "Int = Nat \ (negNat \ {0})" + proof - + have s1_1: "Int = Nat \ negNat" + using int_union_nat_negnat by auto + have s1_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem neg_int_eq_int[simp]: + "n \ Int ==> -.n \ Int" + unfolding Int_def by auto + + +theorem minus_negnat_in_int: + "n \ negNat ==> -.n \ Int" + using neg_nat_in_int neg_int_eq_int by auto + + +theorem minus_neg_int_in_nat: + "n \ Int \ n \ Nat ==> + -.n \ Nat \ {0}" + proof - + assume "<1>1": "n \ Int \ n \ Nat" + have "<1>2": "n \ {-.k: k \ Nat}" + using "<1>1" int_disj by iprover + have "<1>3": "\ k \ Nat: n = -.k" + using "<1>2" by auto + have "<1>4": "\ k \ Nat: k = -.n" + using "<1>3" minus_involutive by auto + have "<1>5": "-.n \ Nat" + using "<1>4" by auto + have "<1>6": "-.n \ 0" + proof - + { + assume "<2>1": "-.n = 0" + have "<2>2": "n = 0" + using neg0 "<2>1" by auto + have "<2>3": "n \ Nat" + using "<2>2" by auto + have "<2>4": "n \ Nat" + using "<1>1" by auto + have "<2>5": "FALSE" + using "<2>3" "<2>4" by auto + } + from this show "-.n \ 0" + by auto + qed + show "-.n \ Nat \ {0}" + using "<1>5" "<1>6" by auto + qed + + +theorem neg_negative_in_nat: + assumes hyp: "n \ Int \ n \ Nat" + shows "-. n \ Nat" + proof - + have int: "n \ Int" + using hyp by auto + from int int_disj have disj: + "n \ Nat \ n \ {-. m: m \ Nat}" + by auto + from disj hyp have neg: "n \ {-. m: m \ Nat}" + by auto + from neg have ex1: "\ m \ Nat: n = -. m" + by auto + from ex1 have ex2: "\ m \ Nat: -. n = -. -. m" + by auto + from ex2 neg_neg have ex3: "\ m \ Nat: -. n = m" + by auto + from ex3 show "-. n \ Nat" + by auto + qed + + +theorem minus_nat_0_or_not_in_nat: + "n \ Nat ==> + -.n = 0 \ -.n \ Nat" + using neg0 neg_not_in_nat by auto + + +theorem int_eq_neg_int_is_0[simp]: + "n \ Int ==> (n = -.n) = (n = 0)" + proof - + have "<1>1": "n \ Nat \ {0} ==> + -.n \ Nat" + proof - + assume "<2>1": "n \ Nat \ {0}" + show "-.n \ Nat" + using neg_not_in_nat "<2>1" by auto + qed + have "<1>2": "n \ Nat \ {0} ==> + n \ -.n" + using "<1>1" by force + have "<1>3": "n \ Int \ n \ Nat + ==> -.n \ Nat" + using minus_neg_int_in_nat by auto + have "<1>4": "n \ Int \ n \ Nat + ==> n \ -.n" + using "<1>3" by force + have "<1>5": "n \ Int \ {0} ==> + (n \ Nat \ {0}) \ + (n \ {-.k: k \ Nat})" + using int_disj by auto + have "<1>6": "n \ Int \ {0} ==> + (n \ -.n)" + using "<1>2" "<1>4" "<1>5" by auto + have "<1>7": "n \ Int \ (n \ 0) ==> + (n \ -.n)" + using "<1>6" by auto + have "<1>8": "n \ Int \ (n = -.n) ==> + (n = 0)" + using "<1>7" by auto + have "<1>9": "n \ Int \ (n = 0) ==> + (n = -.n)" + using neg0 by auto + show "n \ Int ==> (n = -.n) = (n = 0)" + using "<1>8" "<1>9" by auto + qed + + +theorem minus_neg_nat_in_nat: + assumes hyp: "n \ negNat" + shows "-.n \ Nat" + proof - + have s1_1: "\ k \ Nat: -.k = n" + using hyp by (auto simp: negNat_def) + have s1_2: "\ k \ Nat: -.-.k = -.n" + using s1_1 by auto + have s1_3: "\ k \ Nat: k = -.n" + using s1_2 neg_neg by auto + show "-.n \ Nat" + using s1_3 by auto + qed + + +theorem minus_neg_nat_0_in_nat_0: + assumes hyp: "n \ negNat \ {0}" + shows "-.n \ Nat \ {0}" + proof - + have s1_1: "-.n \ Nat" + using hyp minus_neg_nat_in_nat by auto + have s1_2: "-.n \ 0" + proof - + { + assume s2_1: "-.n = 0" + have s2_2: "-.-.n = -.0" + using s2_1 by auto + have s2_3: "-.-.n = 0" + using s2_2 neg0 by auto + have s2_4: "n = 0" + proof - + have s3_1: "n \ Int" + using hyp neg_nat_in_int by auto + have s3_2: "-.-.n = n" + using s3_1 minus_involutive + by auto + show ?thesis + using s3_2 s2_3 by auto + qed + have s2_5: "n \ 0" + using hyp by auto + have "FALSE" + using s2_4 s2_5 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem neg_nat_0_not_in_nat: + assumes hyp: "n \ negNat \ {0}" + shows "n \ Nat" + proof auto + assume s1_1: "n \ Nat" + have s1_2: "n \ 0" + using hyp by auto + have s1_3: "n \ Nat \ {0}" + using s1_1 s1_2 by auto + have s1_4: "-.n \ Nat" + using s1_3 neg_not_in_nat by auto + have s1_5: "-.n \ Nat" + using hyp minus_neg_nat_in_nat by auto + show "FALSE" + using s1_4 s1_5 by auto + qed + + +theorem neg_nat_not_in_nat_setminus_0: + assumes hyp: "n \ negNat" + shows "n \ Nat \ {0}" + proof - + have s1_1: "n \ negNat ==> + n \ negNat \ {0} + \ n = 0" + by auto + have s1_2: "n \ negNat \ {0} + ==> n \ Nat \ {0}" + using neg_nat_0_not_in_nat by auto + have s1_3: "n = 0 ==> + n \ Nat \ {0}" + by auto + show "n \ Nat \ {0}" + using hyp s1_1 s1_2 s1_3 by iprover + qed + + +theorem minus_nat_0_in_negnat_0: + assumes hyp: "n \ Nat \ {0}" + shows "-.n \ negNat \ {0}" + proof - + have s1_1: "n \ Nat" + using hyp by auto + have s1_2: "-.n \ negNat" + using s1_1 minus_nat_in_neg_nat by auto + have s1_3: "-.n \ 0" + proof - + { + assume s2_1: "-.n = 0" + have s2_2: "-.-.n = -.0" + using s2_1 by auto + have s2_3: "-.-.n = 0" + using s2_2 by auto + have s2_4: "n = 0" + using s1_1 neg_neg s2_3 by auto + have s2_5: "n \ 0" + using hyp by auto + have "FALSE" + using s2_4 s2_5 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_2 s1_3 by auto + qed + + +theorem minus_one_in_negnat: + "-.1 \\in negNat" + unfolding negNat_def + using oneIsNat by auto + + +theorem minus_one_in_int: + "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + + +theorem minus_one_neq_0: + "-.1 \ 0" + proof - + { + assume s1_1: "-.1 = 0" + have s1_2: "-.-.1 = -.0" + using s1_1 by auto + have s1_3: "-.-.1 = 0" + using s1_2 neg0 by auto + have s1_4: "0 \ Nat" + by auto + have s1_5: "1 = 0" + using s1_3 s1_4 by auto + have "FALSE" + using s1_5 succIrrefl + by auto + } + from this show "-.1 \ 0" + by auto + qed + + +theorem minus_minus_one: + "-.-.1 = 1" + proof - + have s1_1: "1 \ Nat" + using oneIsNat by auto + show "-.-.1 = 1" + using neg_neg s1_1 by auto + qed + + +theorem succ_zero_equality: + assumes hyp: "x = Succ[0]" + shows "x = 1" + using hyp by auto + + +theorem pred_1: + "Pred[1] = 0" + proof - + have s1_1: "1 = Succ[0]" + by auto + have s1_2: "\ x \ Nat: (1 = Succ[x]) => (x = 0)" + using s1_1 succInjIff by auto + show ?thesis + using s1_1 s1_2 choose_equality by auto + qed + + +theorem pred_2: + "Pred[2] = 1" + proof - + have s1_1: "1 \\in Nat \ Succ[1] = 2" + unfolding two_def + using oneIsNat by auto + have s1_2: "\ x. x \\in Nat \ (Succ[x] = 2) ==> (x = 1)" + using oneIsNat succInjIff by (auto simp: two_def) + have s1_4: "Pred[2] = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" + proof - + have s2_1: "Pred[2] = + (IF 2 = 0 THEN 0 ELSE CHOOSE x \ Nat: 2 = Succ[x])" + unfolding Pred_def + using twoIsNat by auto + also have s2_2: "... = (CHOOSE x \ Nat: 2 = Succ[x])" + proof - + have s3_1: "2 \ 0" + using succNotZero by (auto simp: two_def) + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = (CHOOSE x: x \\in Nat \ 2 = Succ[x])" + unfolding bChoose_def + by auto + also have s2_4: "... = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" + proof - + have s3_1: "\ x. (2 = Succ[x]) = (Succ[x] = 2)" + by auto + show ?thesis + using s3_1 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s1_1 s1_2 s1_4 choose_equality[ + of "\ x. x \\in Nat \ Succ[x] = 2" + "1" "1"] + by auto + qed + + +theorem zero_in_int: + "0 \ Int" + using zeroIsNat nat_in_int by auto + + +theorem ipred_0: + shows "iPred[0] = -.1" + proof - + have s1_1: "iPred[0] = -.Succ[-.0]" + proof - + have s2_1: "0 \ Int" + using zero_in_int by auto + have s2_2: "0 \ Nat \ {0}" + by auto + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_2: "... = -.Succ[0]" + using neg0 by auto + also have s1_3: "... = -.1" + by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed + + +theorem isucc_minus_1: + "iSucc[-.1] = 0" + unfolding iSucc_def + using oneIsNat neg_neg pred_1 neg0 by auto + + +theorem negnat_succ_in_nat: + assumes ndom: "n \\in negNat" + shows "Succ[-.n] \\in Nat" + using ndom minus_neg_nat_in_nat succIsNat by auto + + +theorem minus_negnat_succ_in_negnat: + assumes ndom: "n \\in negNat" + shows "-.Succ[-.n] \\in negNat" + using ndom negnat_succ_in_nat minus_nat_in_neg_nat by auto + + +theorem minus_succ_minus_negnat_in_int: + assumes ndom: "n \\in negNat" + shows "-.Succ[-.n] \\in Int" + proof - + have s1_1: "-.n \\in Nat" + using ndom minus_neg_nat_in_nat by auto + have s1_2: "Succ[-.n] \\in Nat" + using s1_1 succIsNat by auto + show ?thesis + using s1_2 minus_nat_in_int by auto + qed + +(*----------------------------------------------------------------------------*) +(* `iSucc` and `iPred` properties. *) + +(* Typeness *) +theorem iSucc_type: + assumes idom: "i \ Int" + shows "iSucc[i] \ Int" + proof - + have s1_1: "i \ Nat ==> + iSucc[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s2_1 by auto + have s2_3: "Succ[i] \ Nat" + using s2_1 succIsNat by auto + have s2_4: "Succ[i] \ Int" + using s2_3 nat_in_int by auto + show "iSucc[i] \ Int" + using s2_2 s2_4 by auto + qed + have s1_2: "i \ Nat ==> + iSucc[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using s2_1 idom by auto + have s2_3: "Pred[-.i] \ Nat" + proof - + have s3_1: "-.i \ Nat" + using idom s2_1 minus_neg_int_in_nat + by auto + show "Pred[-.i] \ Nat" + using s3_1 Pred_in_nat by auto + qed + have s2_4: "-.Pred[-.i] \ Int" + using s2_3 minus_nat_in_int by auto + show "iSucc[i] \ Int" + using s2_2 s2_4 by auto + qed + show "iSucc[i] \ Int" + using s1_1 s1_2 by auto + qed + + +theorem iPred_type: + assumes idom: "i \ Int" + shows "iPred[i] \ Int" + proof - + have s1_1: "i \ Nat \ {0} + ==> iPred[i] \ Int" + proof - + assume s2_1: "i \ Nat \ {0}" + have s2_2: "iPred[i] = Pred[i]" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "Pred[i] \ Nat" + using s2_1 Pred_in_nat by auto + have s2_4: "Pred[i] \ Int" + using s2_3 nat_in_int by auto + show "iPred[i] \ Int" + using s2_2 s2_4 by auto + qed + have s1_2: "i = 0 ==> + iPred[i] \ Int" + proof - + assume s2_1: "i = 0" + have s2_2: "iPred[i] = -.1" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "-.1 \ Int" + using oneIsNat by (auto simp: Int_def) + show "iPred[i] \ Int" + using s2_2 s2_3 by auto + qed + have s1_3: "i \ Nat ==> + iPred[i] \ Int" + proof - + assume s2_1: "i \ Nat" + have s2_2: "iPred[i] = -.Succ[-.i]" + unfolding iPred_def + using idom s2_1 by auto + have s2_3: "-.Succ[-.i] \ Int" + proof - + have s3_1: "-.i \ Nat" + using idom s2_1 minus_neg_int_in_nat + by auto + have s3_2: "Succ[-.i] \ Nat" + using s3_1 succIsNat by auto + show "-.Succ[-.i] \ Int" + using s3_2 minus_nat_in_int by auto + qed + show "iPred[i] \ Int" + using s2_2 s2_3 by auto + qed + show "iPred[i] \ Int" + using s1_1 s1_2 s1_3 by auto + qed + + +(* +THEOREM iSucciPredCommute == + \A i \in Int: iSucc[iPred[i]] = iPred[iSucc[i]] + PROOF + <1>1. SUFFICES ASSUME NEW i \in Int + PROVE iSucc[iPred[i]] = iPred[iSucc[i]] + BY <1>1 + <1>2. CASE i \in Nat \ {0} + <2>1. iSucc[iPred[i]] = i + <3>1. iPred[i] = Pred[i] + <4>1. i \in Nat \ {0} + BY <1>2 + <4> QED + BY <4>1 DEF iPred + <3>2. iSucc[iPred[i]] = Succ[Pred[i]] + <4>1. Pred[i] \in Nat + <5>1. i \in Nat + BY <1>2 + <5> QED + BY <5>1, Pred_in_nat + <4> QED + BY <3>1, <4>1 DEF iSucc + <3>3. Succ[Pred[i]] = i + <4>1. i \in Nat \ {0} + BY <1>2 + <4> QED + BY <4>1, Succ_Pred_nat + <3> QED + BY <3>2, <3>3 + <2>2. iPred[iSucc[i]] = i + <3>1. iSucc[i] = Succ[i] + <4>1. i \in Nat + BY <1>2 + <4> QED + BY <4>1 DEF iSucc + <3>2. iPred[iSucc[i]] = Pred[Succ[i]] + <4>1. Succ[i] \in Nat + <5>1. i \in Nat + BY <1>2 + <5> QED + BY <5>1, succIsNat + <4>2. Succ[i] # 0 + BY <1>2, succNotZero + <4>3. Succ[i] \in Nat \ {0} + BY <4>1, <4>2 + <4>4. iPred[Succ[i]] = Pred[Succ[i]] + BY <4>3 DEF iPred + <4> QED + BY <4>4, <3>1 + <3>3. Pred[Succ[i]] = i + <4>1. i \in Nat + BY <1>2 + <4> QED + BY <4>1, Pred_Succ_nat + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2 + <1>3. CASE i = 0 + <2>1. iSucc[iPred[i]] = 0 + <3>1. iSucc[iPred[i]] = iSucc[-.1] + <4>1. iPred[i] = -.1 + BY <1>3 DEF iPred + <4> QED + BY <4>1 + <3>2. iSucc[-.1] = -.Pred[-.-.1] + <4>1. -.1 \in Int + <5>1. -.1 \in negNat + BY oneIsNat DEF negNat + <5> QED + BY <5>1, neg_nat_in_int + <4>2. -.1 \notin Nat + <5>1. -.1 \in negNat \ {0} + <6>1. -.1 # 0 + BY minus_one_neq_0 + <6>2. -.1 \in negNat + BY oneIsNat DEF negNat + <6> QED + BY <6>1, <6>2 + <5> QED + BY <5>1, neg_nat_0_not_in_nat + <4> QED + BY <4>1, <4>2 DEF iSucc + <3>3. -.Pred[-.-.1] = 0 + <4>1. -.-.1 = 1 + BY minus_minus_one + <4>2. Pred[1] = 0 + BY pred_1 + <4>3. -.0 = 0 + BY neg0 + <4> QED + BY <4>1, <4>2, <4>3 + <3> QED + BY <3>1, <3>2, <3>3 + <2>2. iPred[iSucc[i]] = 0 + <3>1. iSucc[i] = 1 + <4>1. i \in Nat + BY <1>3, zeroIsNat + <4>2. iSucc[i] = Succ[i] + BY <4>1 DEF iSucc + <4>3. Succ[i] = Succ[0] + BY <1>3 + <4>4. Succ[0] = 1 + OBVIOUS + <4> QED + BY <4>2, <4>3, <4>4 + <3>2. iPred[1] = 0 + <4>1. 1 \in Nat \ {0} + BY oneIsNat, succIrrefl + <4>2. iPred[1] = Pred[1] + BY <4>1 DEF iPred + <4>3. Pred[1] = 0 + BY pred_1 + <4> QED + BY <4>2, <4>3 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 + <1>4. CASE i \in negNat + <2>1. iSucc[iPred[i]] = i + <3>1. iPred[i] = -.Succ[-.i] + <4>1. i \in Int + BY <1>1 + <4>2. i \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3>2. -.Succ[-.i] \in negNat \ {0} + <4>1. Succ[-.i] \in Nat \ {0} + <5>1. -.i \ Nat + BY <1>4, minus_neg_nat_in_nat + <5>2. Succ[-.i] \in Nat + BY <5>1, succIsNat + <5>3. Succ[-.i] \ 0 + BY <5>1, succNotZero + <5> QED + BY <5>2, <5>3 + <4> QED + BY <4>1, minus_nat_0_in_negnat_0 + <3>3. iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]] + <4>1. iPred[i] \in negNat \ {0} + BY <3>1, <3>2 + <4>2. iPred[i] \in Int + BY <4>1, neg_nat_in_int + <4>3. iPred[i] \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4>4. iSucc[iPred[i]] = -.Pred[-.iPred[i]] + BY <4>2, <4>3 DEF iSucc + <4> QED + BY <4>4, <3>1 + <3>4. @ = -.Pred[Succ[-.i]] + <4>1. -.i \in Nat + BY <1>4, minus_neg_nat_in_nat + <4>2. Succ[-.i] \in Int + BY <4>1, succIsNat, nat_in_int + <4>3. -.-.Succ[-.i] = Succ[-.i] + BY minus_involutive + <4> QED + BY <4>3 + <3>5. @ = -.-.i + <4>1. -.i \in Nat + BY <1>4, minus_neg_nat_in_nat + <4>2. Pred[Succ[-.i]] = -.i + BY <4>1, Pred_Succ_nat + <4> QED + BY <4>2 + <3>6. @ = i + BY <1>4, minus_involutive + <3> QED + BY <3>3, <3>4, <3>5, <3>6 + <2>2. iPred[iSucc[i]] = i + <3>1. CASE i = 0 + <4>1. iSucc[i] = 1 + <5>1. 0 \in Nat + BY zeroIsNat + <5>2. 0 \in Int + BY <5>1, nat_in_int + <5>3. iSucc[i] = Succ[0] + BY <5>1, <5>2, <3>1 DEF iSucc + <5>4. Succ[0] = 1 + OBVIOUS + <5> QED + BY <5>3, <5>4 + <4>2. iPred[1] = 0 + <5>1. 1 \in Nat \ {0} + BY oneIsNat, succNotZero + <5>2. 1 \in Int + BY <5>1, nat_in_int + <5>3. iPred[1] = Pred[1] + BY <5>1, <5>2 DEF iPred + <5>4. Pred[1] = 0 + BY pred_1 + <5> QED + BY <5>3, <5>4 + <4> QED + BY <4>1, <4>2, <3>1 + <3>2. CASE i # 0 + <4>1. i \in negNat \ {0} + BY <1>4, <3>1 + <4>2. i \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4>3. i \in Int + BY <4>1, neg_nat_in_int + <4>4. iSucc[i] = -.Pred[-.i] + BY <4>2, <4>3 DEF iSucc + <4>5. -.Pred[-.i] \in negNat + <5>1. -.i \in Nat + BY <4>1, minus_neg_nat_in_nat + <5>2. Pred[-.i] \in Nat + BY <5>1, Pred_in_nat + <5> QED + BY <5>2, minus_nat_in_neg_nat + <4>6. -.Pred[-.i] \in Int + BY <4>5, neg_nat_in_int + <4>7. -.Pred[-.i] \notin Nat \ {0} + BY <4>5, neg_nat_not_in_nat_setminus_0 + <4>8. iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]] + BY <4>6, <4>7 DEF iPred + <4>9. -.Succ[-.-.Pred[-.i]] = i + <5>1. -.-.Pred[-.i] = Pred[-.i] + <6>1. -.i \in Nat + BY <4>1, minus_neg_nat_in_nat + <6>2. Pred[-.i] \in Nat + BY <6>1, Pred_in_nat + <6>3. Pred[-.i] \in Int + BY <6>2, nat_in_int + <6> QED + BY <6>3, minus_involutive + <5>2. Succ[Pred[-.i]] = -.i + <6>1. -.i \in Nat \ {0} + BY <4>1, minus_neg_nat_0_in_nat_0 + <6> QED + BY <6>1, Succ_Pred_nat + <5>3. -.-.i = i + BY <4>3, minus_involutive + <5> QED + BY <5>1, <5>2, <5>3 + <4> QED + BY <4>4, <4>8, <4>9 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 + <1> QED + BY <1>1, <1>2, <1>3, <1>4 +*) +theorem iSucciPredCommute: + "\ i \ Int: + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + { + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "i \ Nat \ {0} ==> + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + assume idom: "i \ Nat \ {0}" + have s2_1: "iSucc[iPred[i]] = i" + proof - + have s3_1: "iPred[i] = Pred[i]" + unfolding iPred_def + using s1_1 idom by auto + have s3_2: "iSucc[iPred[i]] = Succ[Pred[i]]" + proof - + have s4_1: "Pred[i] \ Nat" + proof - + have s5_1: "i \ Nat" + using idom by auto + show ?thesis + using s5_1 Pred_in_nat by auto + qed + show ?thesis + unfolding iSucc_def + using s4_1 s3_1 by auto + qed + have s3_3: "Succ[Pred[i]] = i" + using idom Succ_Pred_nat + by auto + show ?thesis + using s3_2 s3_3 by auto + qed + have s2_2: "iPred[iSucc[i]] = i" + proof - + have s3_1: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s1_1 idom by auto + have s3_2: "iPred[iSucc[i]] = Pred[Succ[i]]" + proof - + have s4_1: "Succ[i] \ Nat" + proof - + have s5_1: "i \ Nat" + using idom by auto + show ?thesis + using s5_1 succIsNat by auto + qed + have s4_2: "Succ[i] \ 0" + using idom succNotZero by auto + have s4_3: "Succ[i] \ Nat \ {0}" + using s4_1 s4_2 by auto + have s4_4: "Succ[i] \ Int" + using s4_1 nat_in_int by auto + have s4_5: "iPred[Succ[i]] = Pred[Succ[i]]" + unfolding iPred_def + using s4_3 s4_4 by auto + show ?thesis + using s4_5 s3_1 by auto + qed + have s3_3: "Pred[Succ[i]] = i" + proof - + have s4_1: "i \ Nat" + using idom by auto + show ?thesis + using s4_1 Pred_Succ_nat + by auto + qed + show ?thesis + using s3_2 s3_3 by auto + qed + show ?thesis + using s2_1 s2_2 + by auto + qed + have s1_3: "i \ negNat ==> + iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i + " + proof - + assume idom: "i \ negNat" + have s2_1: "iSucc[iPred[i]] = i" + proof - + have s3_1: "iPred[i] = -.Succ[-.i]" + proof - + have s4_1: "i \ Int" + using s1_1 by auto + have s4_2: "i \ Nat \ {0}" + using idom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + have s3_2: "-.Succ[-.i] \ negNat \ {0}" + proof - + have s4_1: "Succ[-.i] \ Nat \ {0}" + proof - + have s5_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat + by auto + have s5_2: "Succ[-.i] \ Nat" + using s5_1 succIsNat by auto + have s5_3: "Succ[-.i] \ 0" + using s5_1 succNotZero by auto + show ?thesis + using s5_2 s5_3 by auto + qed + show ?thesis + using s4_1 minus_nat_0_in_negnat_0 + by auto + qed + have s3_3: "iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]]" + proof - + have s4_1: "iPred[i] \ negNat \ {0}" + using s3_1 s3_2 by auto + have s4_2: "iPred[i] \ Int" + using s4_1 neg_nat_in_int by auto + have s4_3: "iPred[i] \ Nat" + using s4_1 neg_nat_0_not_in_nat by auto + have s4_4: "iSucc[iPred[i]] = -.Pred[-.iPred[i]]" + unfolding iSucc_def + using s4_2 s4_3 by auto + show ?thesis + using s4_4 s3_1 by auto + qed + have s3_4: "-.Pred[-.-.Succ[-.i]] = -.Pred[Succ[-.i]]" + proof - + have s4_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat by auto + have s4_2: "Succ[-.i] \ Int" + using s4_1 succIsNat nat_in_int by auto + have s4_3: "-.-.Succ[-.i] = Succ[-.i]" + using s4_2 minus_involutive by auto + show ?thesis + using s4_3 by auto + qed + have s3_5: "-.Pred[Succ[-.i]] = -.-.i" + proof - + have s4_1: "-.i \ Nat" + using idom minus_neg_nat_in_nat by auto + have s4_2: "Pred[Succ[-.i]] = -.i" + using s4_1 Pred_Succ_nat by auto + show ?thesis + using s4_2 by auto + qed + have s3_6: "-.-.i = i" + using s1_1 minus_involutive by auto + show ?thesis + using s3_3 s3_4 s3_5 s3_6 by auto + qed + have s2_2: "iPred[iSucc[i]] = i" + proof - + have s3_1: "i = 0 ==> + iPred[iSucc[i]] = i" + proof - + assume i0: "i = 0" + have s4_1: "iSucc[i] = 1" + proof - + have s5_1: "0 \ Nat" + using zeroIsNat by auto + have s5_2: "0 \ Int" + using s5_1 nat_in_int by auto + have s5_3: "iSucc[i] = Succ[0]" + unfolding iSucc_def + using s5_1 s5_2 i0 + by auto + have s5_4: "Succ[0] = 1" + by auto + show ?thesis + using s5_3 s5_4 by auto + qed + have s4_2: "iPred[1] = 0" + proof - + have s5_1: "1 \ Nat \ {0}" + using oneIsNat succNotZero + by auto + have s5_2: "1 \ Int" + using s5_1 nat_in_int by auto + have s5_3: "iPred[1] = Pred[1]" + unfolding iPred_def + using s5_1 s5_2 + by auto + have s5_4: "Pred[1] = 0" + using pred_1 by auto + show ?thesis + using s5_3 s5_4 by auto + qed + show "iPred[iSucc[i]] = i" + using s4_1 s4_2 i0 by auto + qed + have s3_2: "i \ 0 ==> + iPred[iSucc[i]] = i" + proof - + assume inot0: "i \ 0" + have s4_1: "i \ negNat \ {0}" + using idom inot0 by auto + have s4_2: "i \ Nat" + using s4_1 neg_nat_0_not_in_nat + by auto + have s4_3: "i \ Int" + using s4_1 neg_nat_in_int + by auto + have s4_4: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using s4_2 s4_3 by auto + have s4_5: "-.Pred[-.i] \ negNat" + proof - + have s5_1: "-.i \ Nat" + using s4_1 minus_neg_nat_in_nat + by auto + have s5_2: "Pred[-.i] \ Nat" + using s5_1 Pred_in_nat + by auto + show ?thesis + using s5_2 minus_nat_in_neg_nat + by auto + qed + have s4_6: "-.Pred[-.i] \ Int" + using s4_5 neg_nat_in_int + by auto + have s4_7: "-.Pred[-.i] \ Nat \ {0}" + using s4_5 neg_nat_not_in_nat_setminus_0 + by blast + have s4_8: "iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]]" + unfolding iPred_def + using s4_6 s4_7 by auto + have s4_9: "-.Succ[-.-.Pred[-.i]] = i" + proof - + have s5_1: "-.-.Pred[-.i] = Pred[-.i]" + proof - + have s6_1: "-.i \ Nat" + using s4_1 minus_neg_nat_in_nat + by auto + have s6_2: "Pred[-.i] \ Nat" + using s6_1 Pred_in_nat + by auto + have s6_3: "Pred[-.i] \ Int" + using s6_2 nat_in_int + by auto + show ?thesis + using s6_3 minus_involutive + by auto + qed + have s5_2: "Succ[Pred[-.i]] = -.i" + proof - + have s6_1: "-.i \ Nat \ {0}" + using s4_1 minus_neg_nat_0_in_nat_0 + by auto + show ?thesis + using s6_1 Succ_Pred_nat + by auto + qed + have s5_3: "-.-.i = i" + using s4_3 minus_involutive by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + show "iPred[iSucc[i]] = i" + using s4_4 s4_8 s4_9 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + show "iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i" + using s2_1 s2_2 by auto + qed + have s1_4: "Int = (Nat \ {0}) \ negNat" + using int_union_nat_0_negnat by auto + have "iSucc[iPred[i]] = iPred[iSucc[i]] \ + iSucc[iPred[i]] = i \ + iPred[iSucc[i]] = i" + using s1_1 s1_2 s1_3 s1_4 by auto + } + from this show ?thesis + using allI by auto + qed + + +(* +THEOREM iSuccMinusFlip == + \A i \in Int: -.iSucc[-.i] = iPred[i] +PROOF +<1>1. SUFFICES + ASSUME NEW i \in Int + PROVE -.iSucc[-.i] = iPred[i] + BY <1>1 +<1>2. CASE i \in Nat \ {0} + <2>1. i \in Int + BY <1>2, nat_in_int + <2>2. -.iSucc[-.i] = Pred[i] + <3>1. -.iSucc[-.i] = -.-.Pred[-.-.i] + <4>1. -.i \in negNat \ {0} + BY <1>2, minus_nat_0_in_negnat_0 + <4>2. -.i \notin Nat + BY <4>1, neg_nat_0_not_in_nat + <4> QED + BY <2>1, <4>2 DEF iSucc + <3>2. @ = -.-.Pred[i] + <4>1. -.-.i = i + BY <2>1, minus_involutive + <4> QED + BY <4>1 + <3>3. @ = Pred[i] + <4>1. Pred[i] \in Nat + BY <1>2, Pred_in_nat + <4>2. Pred[i] \in Int + BY <4>1, nat_in_int + <4> QED + BY <4>2, minus_involutive + <3> QED + BY <3>1, <3>2, <3>3 + <2>3. iPred[i] = Pred[i] + BY <2>1, <1>2 DEF iPred + <2> QED + BY <2>2, <2>3 +<1>3. CASE i \in negNat + <2>1. i \in Int + BY <1>3, neg_nat_in_int + <2>2. i \notin Nat \ {0} + BY <1>3, neg_nat_not_in_nat_setminus_0 + <2>3. Pred[i] = -.Succ[-.i] + BY <2>1, <2>2 DEF iPred + <2>4. -.iSucc[-.i] = -.Succ[-.i] + <3>1. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3>2. -.i \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>1, <3>2 DEF iSucc + <2> QED + BY <2>3, <2>4 +<1> QED + BY <1>1, <1>2, <1>3, int_union_nat_0_negnat +*) +theorem iSuccMinusFlip: + "\ i \ Int: + -.iSucc[-.i] = iPred[i]" + proof auto + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "i \ Nat \ {0} ==> + -.iSucc[-.i] = iPred[i]" + proof - + assume s1_2_asm: "i \ Nat \ {0}" + have s2_1: "i \ Int" + using s1_2_asm nat_in_int + by auto + have s2_2: "-.iSucc[-.i] = Pred[i]" + proof - + have s3_1: "-.iSucc[-.i] = -.-.Pred[-.-.i]" + proof - + have s4_1: "-.i \ negNat \ {0}" + using s1_2_asm minus_nat_0_in_negnat_0 + by auto + have s4_2: "-.i \ Nat" + using s4_1 neg_nat_0_not_in_nat + by auto + show ?thesis + unfolding iSucc_def + using s2_1 s4_2 by auto + qed + have s3_2: "-.-.Pred[-.-.i] = -.-.Pred[i]" + proof - + have s4_1: "-.-.i = i" + using s2_1 minus_involutive + by auto + show ?thesis + using s4_1 by auto + qed + have s3_3: "-.-.Pred[i] = Pred[i]" + proof - + have s4_1: "Pred[i] \ Nat" + using s1_2_asm Pred_in_nat by auto + have s4_2: "Pred[i] \ Int" + using s4_1 nat_in_int by auto + show ?thesis + using s4_2 minus_involutive by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 by auto + qed + have s2_3: "iPred[i] = Pred[i]" + unfolding iPred_def + using s2_1 s1_2_asm by auto + show "-.iSucc[-.i] = iPred[i]" + using s2_2 s2_3 by auto + qed + have s1_3: "i \ negNat ==> + -.iSucc[-.i] = iPred[i]" + proof - + assume s1_3_asm: "i \ negNat" + have s2_1: "i \ Int" + using s1_3_asm neg_nat_in_int + by auto + have s2_2: "i \ Nat \ {0}" + using s1_3_asm neg_nat_not_in_nat_setminus_0 + by auto + have s2_3: "-.Succ[-.i] = iPred[i]" + unfolding iPred_def + using s2_1 s2_2 by auto + have s2_4: "-.iSucc[-.i] = -.Succ[-.i]" + proof - + have s3_1: "-.i \ Nat" + using s1_3_asm minus_neg_nat_in_nat by auto + have s3_2: "-.i \ Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + show "-.iSucc[-.i] = iPred[i]" + using s2_3 s2_4 by auto + qed + show "-.iSucc[-.i] = iPred[i]" + using s1_1 s1_2 s1_3 int_union_nat_0_negnat + by auto + qed + + +(* +THEOREM iPredMinusFlip == + \A i \in Int: -.iPred[-.i] = iSucc[i] +PROOF +<1>1. SUFFICES + ASSUME NEW i \in Int + PROVE -.iPred[-.i] = iSucc[i] + BY <1>1 +<1>2. -.iSucc[-.-.i] = iPred[-.i] + <2>1. -.i \in Int + BY <1>1, neg_int_eq_int + <2> QED + BY <2>1, iSuccMinusFlip +<1>3. -.iSucc[i] = iPred[-.i] + <2>1. -.-.i = i + BY <1>1, minus_involutive + <2> QED + BY <1>2, <2>1 +<1>4. -.-.iSucc[i] = -.iPred[-.i] + BY <1>3 +<1>5. -.-.iSucc[i] = iSucc[i] + <2>1. iSucc[i] \in Int + BY <1>1, iSucc_type + <2> QED + BY <2>1, minus_involutive +<1> QED + BY <1>4, <1>5 +*) +theorem iPredMinusFlip: + "\ i \ Int: -.iPred[-.i] = iSucc[i]" + proof auto + fix "i" :: "c" + assume s1_1: "i \ Int" + have s1_2: "-.iSucc[-.-.i] = iPred[-.i]" + proof - + have s2_1: "-.i \ Int" + using s1_1 neg_int_eq_int by auto + show ?thesis + using s2_1 iSuccMinusFlip by auto + qed + have s1_3: "-.iSucc[i] = iPred[-.i]" + proof - + have s2_1: "-.-.i = i" + using s1_1 minus_involutive by auto + show ?thesis + using s1_2 s2_1 by auto + qed + have s1_4: "-.-.iSucc[i] = -.iPred[-.i]" + using s1_3 by auto + have s1_5: "-.-.iSucc[i] = iSucc[i]" + proof - + have s2_1: "iSucc[i] \ Int" + using s1_1 iSucc_type by auto + show ?thesis + using s2_1 minus_involutive by auto + qed + show "-.iPred[-.i] = iSucc[i]" + using s1_4 s1_5 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* `neg_nat_induction` *) + + +(* +negNat == {-.n: n \in Nat} + + +THEOREM negNatInduction == + ASSUME + NEW P(n), P(0), + \A i \in negNat: P(i) => P(-.Succ[-.i]) + PROVE + \A i \in negNat: P(i) + PROOF + <1> DEFINE Q(n) == P(-.n) + <1>1. \A n \in Nat: Q(n) + <2>1. Q(0) + BY P(0), -.0 = 0 DEF Q + <2>2. ASSUME NEW n \in Nat, Q(n) + PROVE Q(Succ[n]) + <3>1. P(-.n) + BY <2>2 DEF Q + <3>2. -.n \in negNat + BY <2>2, n \in Nat DEF negNat + <3>3. \A i \in negNat: P(i) => P(-.Succ[-.i]) + OBVIOUS + <3>4. P(-.n) => P(-.Succ[-.-.n]) + BY <3>2, <3>3 + <3>5. P(-.Succ[-.-.n]) + BY <3>1, <3>4 + <3>6. P(-.Succ[n]) + <4>1. -.-.n = n + BY <2>2, minus_involutive, nat_in_int + <4> QED + BY <3>5, <4>1 + <3> QED + BY <3>6 DEF Q + <2> QED + BY <2>1, <2>2, NatInduction + <1>2. ASSUME NEW i \in negNat + PROVE P(i) + <2>1. \E n \in Nat: n = -.i + <3>1. i \in negNat + BY <1>3 + <3>2. \E n \in Nat: i = -.n + BY <3>1 DEF negNat + <3>3. \E n \in Nat: -.i = -.-.n + BY <3>2 + <3>4. \A n \in Nat: -.-.n = n + BY minus_involutive, nat_in_int + <3> QED + BY <3>3, <3>4 + <2>2. \E n \in Nat: Q(-.i) + <3>1. \E n \in Nat: n = -.i /\ Q(n) + BY <1>1, <2>1 + <3>2. \E n \in Nat: n = -.i /\ Q(-.i) + BY <3>1 + <3> QED + BY <3>2 + <2>3. Q(-.i) + BY <2>2 + <2>4. P(-.-.i) + BY <2>3 DEF Q + <2>5. -.-.i = i + <3>1. i \in Int + BY <1>2 DEF negNat, Int + <3> QED + BY <3>1, minus_involutive + <2> QED + BY <2>4, <2>5 + <1> QED + BY <1>2 +*) +theorem neg_nat_induction: + assumes z: "P(0)" and + sc: "\ i. + \ + i \\in negNat; + P(i) + \ \ + P(-.Succ[-.i])" + shows "\ i \ negNat: P(i)" + proof - + { + fix "Q" :: "[c] \ c" + assume q_def: "\ n: Q(n) = P(-.n)" + have s1_1: "\ n \ Nat: Q(n)" + proof - + have s2_1: "Q(0)" + proof - + have s3_1: "P(-.0)" + using z neg0 by auto + show "Q(0)" + using s3_1 q_def allE by auto + qed + have s2_2: "\ n. + \ + n \ Nat; Q(n) + \ \ + Q(Succ[n])" + proof - + fix "n" :: "c" + assume ndom: "n \ Nat" + assume Qn: "Q(n)" + have s3_1: "P(-.n)" + using Qn q_def by auto + have s3_2: "-.n \ negNat" + unfolding negNat_def + using ndom by auto + have s3_3: "\ i. \ i \ negNat; + P(i)\ \ + P(-.Succ[-.i])" + using sc by auto + have s3_4: "P(-.n) ==> + P(-.Succ[-.-.n])" + using s3_2 s3_3 by auto + have s3_5: "P(-.Succ[-.-.n])" + using s3_1 s3_4 by auto + have s3_6: "P(-.Succ[n])" + proof - + have s4_1: "-.-.n = n" + using ndom minus_involutive nat_in_int + by auto + show "P(-.Succ[n])" + using s3_5 s4_1 by auto + qed + show "Q(Succ[n])" + using s3_6 q_def by auto + qed + show "\ n \ Nat: Q(n)" + using s2_1 s2_2 natInduct + by auto + qed + have s1_2: "\ i. i \\in negNat ==> P(i)" + proof - + fix "i" :: "c" + assume idom: "i \ negNat" + have s2_1: "\ n \ Nat: n = -.i" + proof - + have s3_1: "i \ negNat" + using idom by auto + have s3_2: "\ n \ Nat: i = -.n" + using s3_1 by (auto simp: negNat_def) + have s3_3: "\ n \ Nat: -.i = -.-.n" + using s3_2 by auto + have s3_4: "\ n \ Nat: -.-.n = n" + using minus_involutive nat_in_int + by auto + show "\ n \ Nat: n = -.i" + using s3_3 s3_4 by auto + qed + have s2_2: "\ n \ Nat: Q(-.i)" + proof - + have s3_1: "\ n \ Nat: + n = -.i \ Q(n)" + using s1_1 s2_1 by auto + have s3_2: "\ n \ Nat: + n = -.i \ Q(-.i)" + using s3_1 by auto + show "\ n \ Nat: Q(-.i)" + using s3_2 by auto + qed + have s2_3: "Q(-.i)" + using s2_2 by auto + have s2_4: "P(-.-.i)" + using s2_3 q_def by auto + have s2_5: "-.-.i = i" + proof - + have s3_1: "i \ Int" + using idom + by (auto simp: negNat_def Int_def) + show "-.-.i = i" + using s3_1 minus_involutive + by auto + qed + show "P(i)" + using s2_4 s2_5 by auto + qed + have "\ i \ negNat: P(i)" + proof - + have s2_1: + "\ i. i \ negNat => P(i)" + using s1_2 impI by auto + have s2_2: + "\ i: i \ negNat => P(i)" + using s2_1 allI by auto + show "\ i \ negNat: P(i)" + using s2_2 by auto + (* by (rule allI, rule impI, rule s1_2) *) + qed + } + from this have "\ Q: ( + (\ n: Q(n) = P(-.n)) + => + (\ i \ negNat: P(i)) + )" + by auto + from this have " + (\ n: P(-.n) = P(-.n)) + => + (\ i \ negNat: P(i))" + by auto + thus " + (\ i \ negNat: P(i))" + by auto + qed + +(*----------------------------------------------------------------------------*) +(* Primitive recursion in two directions: plus and minus infinity. *) + + +theorem bprimrec_int: + assumes + e: "e \ S" and + succ_h: "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + succ_g: "\ n \ Nat: + \ x \ S: + g(n, x) \ S" + shows "\ f \ [Int -> S]: + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + )" + proof - + from primrec_int[of e h g] obtain f where + 1: "isAFcn(f)" and + 2: "DOMAIN f = Int" and + 3: "f[0] = e" and + 4: "\ n \ Nat: + f[Succ[n]] = h(n, f[n]) + \ f[-.Succ[n]] = g(n, f[-.n])" + by blast + have s1_2: "\ n \ Nat: f[n] \ S" + proof (rule natInduct) + from 3 e show "f[0] \ S" by simp + next + fix n + assume "n \ Nat" and "f[n] \ S" + with succ_h 4 show "f[Succ[n]] \ S" + by force + qed + have s1_3: "\ i \ negNat: f[i] \ S" + proof (rule neg_nat_induction) + from 3 e show "f[0] \ S" by simp + next + fix i + assume ndom: "i \ negNat" and + fidom: "f[i] \ S" + have s2_1: "-.i \ Nat" + using ndom minus_neg_nat_in_nat by auto + have s2_2: "f[-.Succ[-.i]] = g(-.i, f[-.-.i])" + using 4 s2_1 by auto + have s2_3: "g(-.i, f[-.-.i]) = g(-.i, f[i])" + using ndom neg_nat_in_int minus_involutive + by auto + have s2_4: "g(-.i, f[i]) \ S" + using s2_1 fidom succ_g by auto + show "f[-.Succ[-.i]] \ S" + using s2_2 s2_3 s2_4 by auto + qed + have 5: "\ i \ Int: f[i] \ S" + using s1_2 s1_3 int_union_nat_negnat + by auto + with 1 2 3 4 5 show ?thesis + by blast + qed + + +theorem primrec_intE: + assumes e: "e \ S" and + succ_h: "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + succ_g: "\ n \ Nat: + \ x \ S: + g(n, x) \ S" and + f: "f = (CHOOSE r \ [Int -> S]: + r[0] = e \ + (\ n \ Nat: + r[Succ[n]] = h(n, r[n]) \ + r[-.Succ[n]] = g(n, r[-.n]) + ))" (is "f = ?r") and + maj: "\ + f \ [Int -> S]; + f[0] = e; + \ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + \ + \ P" + shows "P" + proof - + from e succ_h succ_g have " + \ r \ [Int -> S]: + r[0] = e \ + (\ n \ Nat: + r[Succ[n]] = h(n, r[n]) \ + r[-.Succ[n]] = g(n, r[-.n]) + )" + by (rule bprimrec_int) + hence "?r \ [Int -> S] \ + ?r[0] = e \ + (\ n \ Nat: + ?r[Succ[n]] = h(n, ?r[n]) \ + ?r[-.Succ[n]] = g(n, ?r[-.n]) + )" + by (rule bChooseI2, auto) + with f maj show ?thesis by auto + qed + + +theorem bprimrecType_int: + assumes "e \ S" and + "\ n \ Nat: + \ x \ S: + h(n, x) \ S" and + "\ n \ Nat: + \ x \ S: + g(n, x) \ S" + shows "(CHOOSE f \ [Int -> S]: + f[0] = e \ + (\ n \ Nat: + f[Succ[n]] = h(n, f[n]) \ + f[-.Succ[n]] = g(n, f[-.n]) + )) \ [Int -> S]" + by (rule primrec_intE[OF assms], auto) + + +(*----------------------------------------------------------------------------*) +(* Typeness of addition and recursive properties. *) + + +theorem addintType: + assumes mdom: "m \ Int" + shows "addint(m) \ [Int -> Int]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: "(CHOOSE f \ [Int -> Int]: + f[0] = m \ + (\ n \ Nat: + f[Succ[n]] = iSucc[f[n]] \ + f[-.Succ[n]] = iPred[f[-.n]] + )) \ [Int -> Int]" + using mdom s1_1 s1_2 + by (rule bprimrecType_int) + show "addint(m) \ [Int -> Int]" + unfolding addint_def + using s1_3 by auto + qed + + +theorem addIsInt: + assumes "m \ Int" and "n \ Int" + shows "add(m, n) \ Int" + unfolding add_def + using assms addintType by blast + + +theorem diffIsInt: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "diff(m, n) \\in Int" + proof - + have s1_1: "diff(m, n) = add(m, -.n)" + unfolding diff_def by auto + have s1_2: "-.n \\in Int" + using ndom neg_int_eq_int by auto + have s1_3: "add(m, -.n) \\in Int" + using mdom s1_2 addIsInt by auto + show ?thesis + using s1_1 s1_3 by auto + qed + + +theorem addint_0: + assumes mdom: "m \ Int" + shows "addint(m)[0] = m" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: " + addint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = m \ + (\ n \ Nat: + r[Succ[n]] = iSucc[r[n]] \ + r[-.Succ[n]] = iPred[r[-.n]]) + )" + unfolding addint_def + by auto + have s1_4: " + \ + addint(m) \ [Int -> Int]; + addint(m)[0] = m; + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + \ + \ addint(m)[0] = m" + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "m" "Int" + "\ n x. iSucc[x]" + "\ n x. iPred[x]" + "addint(m)"] + by auto + qed + + +theorem add_0: + assumes "m \ Int" + shows "m + 0 = m" + unfolding add_def + using assms addint_0 by auto + + +theorem addint_succ: + assumes mdom: "m \ Int" + shows "\ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + iSucc[x] \ Int" + using iSucc_type by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + iPred[x] \ Int" + using iPred_type by auto + have s1_3: " + addint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = m \ + (\ n \ Nat: + r[Succ[n]] = iSucc[r[n]] \ + r[-.Succ[n]] = iPred[r[-.n]]) + )" + unfolding addint_def + by auto + have s1_4: " + \ + addint(m) \ [Int -> Int]; + addint(m)[0] = m; + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + \ + \ + \ n \ Nat: + addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ + addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] + " + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "m" "Int" + "\ n x. iSucc[x]" + "\ n x. iPred[x]" + "addint(m)"] + by auto + qed + + +theorem addint_succ_nat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m + Succ[n] = iSucc[m + n]" + unfolding add_def + using mdom ndom addint_succ[of "m"] spec by auto + + +theorem addint_succ_negnat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m + -.Succ[n] = iPred[m + -.n]" + unfolding add_def + using mdom ndom addint_succ[of "m"] spec by auto + + +theorem nat_add_1: + assumes ndom: "n \\in Nat" + shows "Succ[n] = n + 1" + proof - + have s1_1: "Succ[n] = iSucc[n]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "n \\in Nat" + using ndom by auto + show ?thesis + unfolding iSucc_def + using s2_1 s2_2 by auto + qed + also have s1_2: "... = iSucc[n + 0]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "n + 0 = n" + using s2_1 add_0 by auto + have s2_3: "n = n + 0" + using s2_2[symmetric] by auto + show ?thesis + using s2_3 by auto + qed + also have s1_3: "... = n + Succ[0]" + proof - + have s2_1: "n \\in Int" + using ndom nat_in_int by auto + have s2_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + using s2_1 s2_2 addint_succ_nat by auto + qed + also have s1_4: "... = n + 1" + by auto (* 1 is an abbreviation *) + from calculation show ?thesis . + qed + + +theorem nat_add_in_nat: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" + shows "m + n \\in Nat" + proof - + have s1_1: "m + 0 \\in Nat" + proof - + have s2_1: "m + 0 = m" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + show ?thesis + using s3_1 add_0 by auto + qed + have s2_2: "m \\in Nat" + using mdom by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_2: "\ k. \ + k \\in Nat; + m + k \\in Nat + \ \ + m + Succ[k] \\in Nat" + proof - + fix "k" :: "c" + assume s1_2_kdom: "k \\in Nat" + assume s1_2_induct: "m + k \\in Nat" + have s2_1: "m + Succ[k] = iSucc[m + k]" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + have s3_2: "k \\in Nat" + using s1_2_kdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + have s2_2: "iSucc[m + k] = Succ[m + k]" + proof - + have s3_1: "m + k \\in Nat" + using s1_2_induct by auto + have s3_2: "m + k \\in Int" + using s1_2_induct nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_3: "Succ[m + k] \\in Nat" + using s1_2_induct succIsNat by auto + show "m + Succ[k] \\in Nat" + using s2_1 s2_2 s2_3 by auto + qed + have s1_3: "\ k \ Nat: m + k \\in Nat" + using s1_1 s1_2 natInduct by auto + show ?thesis + using s1_3 ndom spec by auto + qed + + +theorem nat_0_succ: + "\ n \ Nat: + n=0 \ + (\ m \ Nat: n = Succ[m])" + by (rule natInduct, auto) + + +theorem nat_add_0: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" and + mn: "m + n = 0" + shows "m = 0 \ n = 0" + proof - + have s1_1: "n \ 0 ==> FALSE" + proof - + assume s2_1: "n \ 0" + have s2_2: "\ k: k \\in Nat \ n = Succ[k]" + using ndom s2_1 nat_0_succ by auto + def P \ "\ x. x \\in Nat \ n = Succ[x]" + def r \ "CHOOSE x: P(x)" + have s2_3: "r \\in Nat \ n = Succ[r]" + proof - + have s3_1: "\ x: P(x)" + using s2_2 by (auto simp: P_def) + have s3_2: "P(r)" + using s3_1 chooseI_ex by (auto simp: r_def) + show ?thesis + using s3_2 by (auto simp: P_def) + qed + have s2_4: "m + n = m + Succ[r]" + using s2_3 by auto + have s2_5: "m + Succ[r] = iSucc[m + r]" + proof - + have s3_1: "m \\in Int" + using mdom nat_in_int by auto + have s3_2: "r \\in Nat" + using s2_3 by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + have s2_6: "iSucc[m + r] = Succ[m + r]" + proof - + have s3_1: "m + r \\in Nat" + using mdom s2_3 nat_add_in_nat by auto + have s3_2: "m + r \\in Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_7: "Succ[m + r] \ 0" + proof - + have s3_1: "m + r \\in Nat" + using mdom s2_3 nat_add_in_nat by auto + show ?thesis + using s3_1 succNotZero by auto + qed + have s2_7: "m + n \ 0" + using s2_4 s2_5 s2_6 s2_7 by auto + show "FALSE" + using s2_7 mn by auto + qed + have s1_2: "n = 0" + using s1_1 by auto + have s1_3: "m + 0 = 0" + using mn s1_2 by auto + have s1_4: "m = 0" + using s1_3 mdom nat_in_int add_0 by auto + show ?thesis + using s1_2 s1_4 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Typeness of multiplication and recursive properties. *) + + +theorem multType: + assumes mdom: "m \ Int" + shows "multint(m) \ [Int -> Int]" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: "(CHOOSE f \ [Int -> Int]: + f[0] = 0 \ + (\ n \ Nat: + f[Succ[n]] = f[n] + m \ + f[-.Succ[n]] = f[-.n] + -.m + )) \ [Int -> Int]" + using zero_in_int s1_1 s1_2 + by (rule bprimrecType_int[of "0"]) + show "multint(m) \ [Int -> Int]" + unfolding multint_def + using s1_3 by auto + qed + + +theorem multIsInt: + assumes "m \ Int" and "n \ Int" + shows "mult(m, n) \ Int" + unfolding mult_def + using assms multType by blast + + +theorem multint_0: + assumes mdom: "m \ Int" + shows "multint(m)[0] = 0" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: " + multint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = 0 \ + (\ n \ Nat: + r[Succ[n]] = r[n] + m \ + r[-.Succ[n]] = r[-.n] + -.m) + )" + unfolding multint_def + by auto + have s1_4: " + \ + multint(m) \ [Int -> Int]; + multint(m)[0] = 0; + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + \ + \ + multint(m)[0] = 0" + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "0" "Int" + "\ n x. x + m" + "\ n x. x + -.m" + "multint(m)"] + by auto + qed + + +theorem mult_0: + assumes "m \ Int" + shows "m * 0 = 0" + unfolding mult_def + using assms multint_0 by auto + + +theorem multint_succ: + assumes mdom: "m \ Int" + shows "\ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m" + proof - + have s1_1: " + \ n \ Nat: + \ x \ Int: + x + m \ Int" + using mdom addIsInt by auto + have s1_2: " + \ n \ Nat: + \ x \ Int: + x + -.m \ Int" + proof - + have s2_1: "-.m \ Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s2_1 addIsInt by auto + qed + have s1_3: " + multint(m) = (CHOOSE r \ [Int -> Int]: + r[0] = 0 \ + (\ n \ Nat: + r[Succ[n]] = r[n] + m \ + r[-.Succ[n]] = r[-.n] + -.m) + )" + unfolding multint_def + by auto + have s1_4: " + \ + multint(m) \ [Int -> Int]; + multint(m)[0] = 0; + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + \ + \ + \ n \ Nat: + multint(m)[Succ[n]] = multint(m)[n] + m \ + multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m + " + by auto + show ?thesis + using mdom s1_1 s1_2 s1_3 s1_4 + primrec_intE[of + "0" "Int" + "\ n x. x + m" + "\ n x. x + -.m" + "multint(m)"] + by auto + qed + + +theorem multint_succ_nat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m * Succ[n] = m * n + m" + unfolding mult_def + using mdom ndom multint_succ[of "m"] spec by auto + + +theorem multint_succ_negnat: + assumes mdom: "m \ Int" and ndom: "n \ Nat" + shows "m * -.Succ[n] = m * -.n + -.m" + unfolding mult_def + using mdom ndom multint_succ[of "m"] spec by auto + + +theorem nat_mult_in_nat: + assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" + shows "m * n \\in Nat" + proof - + have s1_1: "m * 0 \\in Nat" + proof - + have s2_1: "m * 0 = 0" + using mdom nat_in_int mult_0 by auto + have s2_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_2: "\ k. \ + k \\in Nat; + m * k \\in Nat + \ \ + m * Succ[k] \\in Nat" + proof - + fix "k" :: "c" + assume s1_2_kdom: "k \\in Nat" + assume s1_2_induct: "m * k \\in Nat" + have s2_1: "m * Succ[k] = m * k + m" + using mdom s1_2_kdom multint_succ_nat by auto + have s2_2: "m * k + m \\in Nat" + using s1_2_induct mdom nat_add_in_nat by auto + show "m * Succ[k] \\in Nat" + using s2_1 s2_2 by auto + qed + have s1_3: "\ k \ Nat: m * k \\in Nat" + using s1_1 s1_2 natInduct by auto + show ?thesis + using s1_3 ndom spec by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Commutativity of addition. *) + + +(* +THEOREM pred_irrefl == + ASSUME NEW n \in Nat \ {0} + PROVE Pred[n] # n +PROOF +<1>1. Pred[n] \in Nat + BY n \in Nat, Pred_in_nat +<1>2. Succ[Pred[n]] # Pred[n] + BY <1>1, succIrrefl +<1>3. Succ[Pred[n]] = n + BY n \in Nat \ {0}, Succ_Pred_nat +<1> QED + BY <1>2, <1>3 +*) +theorem pred_irrefl: + assumes ndom: "n \\in Nat \ {0}" + shows "Pred[n] \ n" + proof - + have s1_1: "Pred[n] \\in Nat" + using ndom Pred_in_nat by auto + have s1_2: "Succ[Pred[n]] \ Pred[n]" + using s1_1 succIrrefl by auto + have s1_3: "Succ[Pred[n]] = n" + using ndom Succ_Pred_nat by auto + show ?thesis + using s1_2 s1_3 by auto + qed + + +theorem ipred_plus_1: + assumes hyp: "a \ Int" + shows "iPred[a + 1] = a" + proof - + have s1_1: "iPred[a + 1] = iPred[a + Succ[0]]" + by auto + have s1_2: "... = iPred[iSucc[a + 0]]" + proof - + have s2_1: "a \ Int" + using hyp by auto + have s2_2: "0 \ Nat" + using zeroIsNat by auto + have s2_3: "a + Succ[0] = iSucc[a + 0]" + using s2_1 s2_2 addint_succ_nat by auto + show ?thesis + using s2_3 by auto + qed + have s1_3: "... = iPred[iSucc[a]]" + using hyp add_0 by auto + have s1_4: "... = a" + using hyp iSucciPredCommute spec by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +(* +THEOREM AddCommutativeNatNat == + \A j \in Nat: \A i \in Nat: + j + i = i + j +PROOF +<1>1. \A i \in Nat: 0 + i = i + 0 + <2>1. 0 + 0 = 0 + 0 + <3>1. 0 + 0 = addint(0)[0] + BY DEF + + <3>2. @ = 0 + BY DEF addint + <3> QED + BY <3>1, <3>2 + + <2>2. ASSUME NEW i \in Nat, 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + + <3>1. 0 + Succ[i] = iSucc[0 + i] + BY DEF + + <3>2. @ = iSucc[i + 0] + BY <2>2 + <3>3. @ = iSucc[i] + BY DEF + + <3>4. @ = Succ[i] + BY <2>2, i \in Nat + <3>5. @ = Succ[i] + 0 + BY DEF + + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5 + <2> QED + BY <2>1, <2>2, NatInduction + +<1>2. ASSUME NEW j \in Nat, + \A i \in Nat: j + i = i + j + PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] + + <2>1. Succ[j] + 0 = 0 + Succ[j] + <3>1. Succ[j] + 0 = Succ[j] + BY DEF + + <3>2. @ = Succ[j + 0] + BY DEF + + <3>3. @ = Succ[0 + j] + BY <1>2 + <3>4. @ = 0 + Succ[j] + BY DEF + + <3> QED + BY <3>1, <3>2, <3>3, <3>4 + + <2>2. ASSUME NEW i \in Nat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] + + <3>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] + <4>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. Succ[i] \in Nat + BY <2>2, succIsNat + <5> QED + BY <5>1, <5>2, addint_succ + <4>2. @ = iSucc[i + Succ[j]] + BY <2>2 + <4>3. @ = iSucc[iSucc[i + j]] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, addint_succ + <4>4. @ = iSucc[iSucc[j + i]] + BY <1>2, <2>2 + <4> QED + BY <4>1, <4>2, <4>3, <4>4 + + <3>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] + <4>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. Succ[j] \in Nat + BY <1>2, succIsNat + <5> QED + BY <5>1, <5>2, addint_succ + <4>2. @ = iSucc[j + Succ[i]] + <5>1. Succ[i] \in Nat + BY <2>2, succIsNat + <5> QED + BY <1>2, <5>1 + <4>3. @ = iSucc[iSucc[j + i]] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, addint_succ + <4> QED + BY <4>1, <4>2, <4>3 + + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, NatInduction +<1> QED + BY <1>1, <1>2, NatInduction + + +THEOREM AddCommutative == + \A j \in Int: \A i \in Int: + j + i = i + j +PROOF +<1>1. \A j \in Nat: \A i \in Int: + j + i = i + j + <2>1. \A j \in Nat: \A i \in Nat: + j + i = i + j + \* The proof of step <2>1 is the same with the + \* proof of the theorem `AddCommutativeNatNat`. + <3>1. \A i \in Nat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in Nat, + 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + <5>1. 0 + Succ[i] = iSucc[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Nat + BY <4>2 + <6> QED + BY <6>1, <6>2, addint_succ + DEF addd + <5>2. @ = iSucc[i + 0] + BY <4>2 + <5>3. @ = iSucc[i] + <6>1. i \in Int + BY <4>2, nat_in_int + <6> QED + BY <6>1 addint_0 + <5>4. @ = Succ[i] + BY <4>2 DEF iSucc + <5>5. @ = Succ[i] + 0 + <6>1. Succ[i] \in Nat + BY <4>2, succIsNat + <6>2. Succ[i] \in Int + BY <6>1, nat_in_int + <6> QED + BY <6>2, addint_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, NatInduction + <3>2. ASSUME NEW j \in Nat, + \A i \in Nat: j + i = i + j + PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] + <4>1. Succ[j] + 0 = 0 + Succ[j] + <5>1. Succ[j] + 0 = Succ[j] + <6>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <6> QED + BY <6>1, addint_0 + <5>2. @ = iSucc[j] + BY <3>2, nat_in_int DEF iSucc + <5>3. @ = iSucc[j + 0] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iSucc[0 + j] + <6>1. j + 0 = 0 + j + BY <3>2, zeroIsNat + <6> QED + By <6>1 + <5>5. @ = 0 + Succ[j] + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2, addint_succ + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4>2. ASSUME NEW i \in Nat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] + <5>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] + <6>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] + <7>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6>2. @ = iSucc[i + Succ[j]] + BY <4>2 + <6>3. @ = iSucc[iSucc[i + j]] + <7>1. i \in Int + BY <4>2, nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6>4. @ = iSucc[iSucc[j + i]] + BY <3>2, <4>2 + <6> QED + By <6>1, <6>2, <6>3, <6>4 + <5>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] + <6>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] + <7>1. Succ[i] \in Int + BY <4>2, succIsNat, nat_in_int + <7>2. j \in Nat + BY <3>2, succIsNat + <7> QED + BY <7>1, <7>2, addint_succ + <6>2. @ = iSucc[j + Succ[i]] + <7>1. Succ[i] \in Nat + BY <4>2, succIsNat + <7> QED + BY <3>2, <7>1 + <6>3. @ = iSucc[iSucc[j + i]] + <7>1. j \in Int + BY <3>2, nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ + <6> QED + BY <6>1, <6>2, <6>3 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, NatInduction + <3> QED + BY <3>1, <3>2, NatInduction + <2>2. \A j \in Nat: \A i \in negNat: + j + i = i + j + <3>1. \A i \in negNat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in negNat, + 0 + i = i + 0 + PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 + <5>1. 0 + -.Succ[-.i] = iPred[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <6>3. 0 + -.Succ[-.i] = iPred[0 + -.-.i] + BY <6>1, <6>2, addint_succ_negnat + <6>4. -.-.i = i + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7> QED + BY <7>1, minus_involutive + <6> QED + BY <6>3, <6>4 + <5>2. @ = iPred[i + 0] + BY <4>2 + <5>3. @ = iPred[i] + <6>1. i \in Int + BY <4>2, neg_nat_in_int + <6>2. i + 0 = i + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = -.Succ[-.i] + <6>1. i \in Int + BY <4>2, neg_nat_in_int + <6>2. i \notin Nat \ {0} + BY <4>2, neg_nat_not_in_nat_setminus_0 + <6> QED + BY <6>1, <6>2 DEF iPred + <5>5. @ = -.Succ[-.i] + 0 + <6>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <6>2. Succ[-.i] \in Nat + BY <6>1, succIsNat + <6>3. -.Succ[-.i] \in Int + BY <6>2, minus_nat_in_int + <6> QED + BY <6>3, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3>2. ASSUME NEW j \in Nat, + \A i \in negNat: j + i = i + j + PROVE \A i \in negNat: Succ[j] + i = i + Succ[j] + <4>1. Succ[j] + 0 = 0 + Succ[j] + <5>1. Succ[j] + 0 = Succ[j] + <6>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <6> QED + BY <6>1, add_0 + <5>2. @ = iSucc[j] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2 DEF iSucc + <5>3. @ = iSucc[j + 0] + <6>1. j \in Int + BY <3>2, nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iSucc[0 + j] + <6>1. 0 \in negNat + BY neg0 DEF negNat + <6> QED + BY <3>2, <6>1 + <5>5. @ = 0 + Succ[j] + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Nat + BY <3>2 + <6> QED + BY <6>1, <6>2, addint_succ_nat + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4>2. ASSUME NEW i \in negNat, + Succ[j] + i = i + Succ[j] + PROVE Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j] + <5>1. Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]] + <6>1. Succ[j] + -.Succ[-.i] = iPred[Succ[j] + i] + <7>1. Succ[j] \in Int + BY <3>2, succIsNat, nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7>3. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7>4. Succ[j] + -.Succ[-.i] + = iPred[Succ[j] + -.-.i] + BY <7>1, <7>2, addint_succ_negnat + <7> QED + BY <7>3, <7>4 + <6>2. @ = iPred[i + Succ[j]] + BY <4>2 + <6>3. @ = iPred[iSucc[i + j]] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>4. @ = iPred[iSucc[j + i]] + <7>1. i \in negNat + BY <4>2 + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + i = i + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. -.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]] + <6>1. -.Succ[-.i] + Succ[j] = iSucc[-.Succ[-.i] + j] + <7>1. -.Succ[-.i] \in Int + BY <4>2, minus_neg_nat_in_nat, succIsNat, + minus_nat_in_int + <7>2. j \in Nat + BY <3>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>2. @ = iSucc[j + -.Succ[-.i]] + <7>1. -.Succ[-.i] \in negNat + BY <4>2, minus_neg_nat_in_nat, succIsNat, + minus_nat_in_neg_nat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6>3. @ = iSucc[iPred[j + i]] + <7>1. j \in Int + BY <3>2, nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7>3. j + -.Succ[-.i] = iPred[j + -.-.i] + BY <7>1, <7>2, addint_succ_negnat + <7>4. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>3, <7>4 + <6>4. @ = iPred[iSucc[j + i]] + <7>1. j + i \in Int + <8>1. j \in Int + BY <3>2, nat_in_int + <8>2. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, <8>2, addIsInt + <7> QED + BY <7>1, iSucciPredCommute + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3> QED + BY <3>1, <3>2, NatInduction + <2> QED + BY <2>1, <2>2, int_union_nat_negnat +<1>2. \A j \in negNat: \A i \in Int: + j + i = i + j + <2>1. \A j \in negNat: \A i \in Nat: + j + i = i + j + <3>1. \A i \in Nat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME + NEW i \in Nat, + 0 + i = i + 0 + PROVE 0 + Succ[i] = Succ[i] + 0 + \* The level 5 proof is the same as + \* for step <1>1!<2>1!<3>1!<4>2 . + <5>1. 0 + Succ[i] = iSucc[0 + i] + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Nat + BY <4>2 + <6> QED + BY <6>1, <6>2 DEF addint_succ_nat + <5>2. @ = iSucc[i + 0] + BY <4>2 + <5>3. @ = iSucc[i] + <6>1. i \in Int + BY <4>2, nat_in_int + <6>2. i + 0 = i + BY <6>1, addint_0 + <6> QED + BY <6>2 + <5>4. @ = Succ[i] + <6>1. i \in Nat + BY <4>2 + <6>2. i \in Int + BY <4>2, nat_in_int + <6> QED + BY <6>1, <6>2 DEF iSucc + <5>5. @ = Succ[i] + 0 + <6>1. Succ[i] \in Nat + BY <4>2, succIsNat + <6>2. Succ[i] \in Int + BY <6>1, nat_in_int + <6> QED + BY <6>2, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5 + <4> QED + BY <4>1, <4>2, NatInduction + <3>2. ASSUME + NEW j \in negNat, + \A i \in Nat: j + i = i + j + PROVE + \A i \in Nat: + -.Succ[-.j] + i = i + -.Succ[-.j] + <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] + <5>1. -.Succ[-.j] + 0 = -.Succ[-.j] + <6>1. -.Succ[-.j] \in Int + <7>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>2. Succ[-.j] \in Nat + BY <7>1, succIsNat + <7> QED + BY <7>2, minus_nat_in_int + <6> QED + BY <6>1, add_0 + <5>2. @ = iPred[j] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. j \notin Nat \ {0} + BY <3>2, neg_nat_not_in_nat_setminus_0 + <6> QED + BY <6>1, <6>2 DEF iPred + <5>3. @ = iPred[j + 0] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. j + 0 = j + BY <6>1, add_0 + <6> QED + BY <6>2 + <5>4. @ = iPred[0 + j] + <6>1. 0 \in Nat + BY zeroIsNat + <6>2. \A k \in Nat: j + k = k + j + BY <3>2 + <6>3. j + 0 = 0 + j + BY <6>1, <6>2 + <6> QED + BY <6>3 + <5>5. @ = iPred[0 + -.-.j] + <6>1. j \in Int + BY <3>2, neg_nat_in_int + <6>2. -.-.j = j + BY <6>1, minus_involutive + <6> QED + BY <6>2 + <5>6. @ = 0 + -.Succ[-.j] + <6>1. 0 \in Int + BY zeroIsNat, nat_in_int + <6>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <6> QED + BY <6>1, <6>2, addint_succ_negnat + <5> QED + BY <5>1, <5>2, <5>3, <5>4, <5>5, <5>6 + <4>2. ASSUME + NEW i \in Nat, + -.Succ[-.j] + i = i + -.Succ[-.j] + PROVE + -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j] + <5>1. -.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]] + <6>1. -.Succ[-.j] + Succ[i] = + iSucc[-.Succ[-.j] + i] + <7>1. -.Succ[-.j] \in Int + <8>1. j \in negNat + BY <3>2 + <8>2. -.j \in Nat + BY <8>1, minus_neg_nat_in_nat + <8>3. Succ[-.j] \in Nat + BY <8>2, succIsNat + <8> QED + BY <8>3, minus_nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>2. @ = iSucc[i + -.Succ[-.j]] + BY <4>2 + <6>3. @ = iSucc[iPred[i + j]] + <7>1. i \in Int + BY <4>2, nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>3. -.-.j = j + BY <3>2, neg_nat_in_int, minus_involutive + <7> QED + BY <7>1, <7>2, <7>3, addint_succ_negnat + <6>4. @ = iSucc[iPred[j + i]] + <7>1. i \in Nat + BY <4>2 + <7>2. j + i = i + j + BY <3>2, <7>1 + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]] + <6>1. Succ[i] + -.Succ[-.j] = + iPred[Succ[i] + -.-.j] + <7>1. Succ[i] \in Int + BY <4>2, succIsNat, nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[Succ[i] + j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>2 + <6>3. @ = iPred[j + Succ[i]] + <7>1. Succ[i] \in Nat + BY <4>2, succIsNat + <7>2. \A k \in Nat: j + k = k + j + BY <3>2 + <7>3. j + Succ[i] = Succ[i] + j + BY <7>1, <7>2 + <7> QED + BY <7>3 + <6>4. @ = iPred[iSucc[j + i]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. i \in Nat + BY <4>2 + <7> QED + BY <7>1, <7>2, addint_succ_nat + <6>5. @ = iSucc[iPred[j + i]] + <7>1. j + i \in Int + <8>1. j \in Int + BY <3>2, neg_nat_in_int + <8>2. i \in Int + BY <4>2, nat_in_int + <8> QED + BY <8>1, <8>2, addIsInt + <7> QED + BY <7>1, iSucciPredCommute + <6> QED + BY <6>1, <6>2, <6>3, <6>4, <6>5 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, NatInduction + <3> QED + BY <3>1, <3>2, neg_nat_induction + <2>2. \A j \in negNat: \A i \in negNat: + j + i = i + j + <3>1. \A i \in negNat: 0 + i = i + 0 + <4>1. 0 + 0 = 0 + 0 + OBVIOUS + <4>2. ASSUME NEW i \in negNat, 0 + i = i + 0 + PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 + <5>1. 0 + -.Succ[-.i] = iPred[i] + <6>1. 0 + -.Succ[-.i] = iPred[0 + -.-.i] + <7>1. 0 \in Int + BY zero_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[0 + i] + <7>1. -.-.i = i + BY <4>2, neg_nat_in_int, minus_involutive + <7> QED + BY <7>1 + <6>3. @ = iPred[i + 0] + BY <4>2 + <6>4. @ = iPred[i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. i + 0 = i + BY <7>1, add_0 + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5>2. -.Succ[-.i] + 0 = iPred[i] + <6>1. -.Succ[-.i] + 0 = -.Succ[-.i] + <7>1. -.Succ[-.i] \in Int + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7> QED + BY <7>1, add_0 + <6>2. @ = iPred[i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. i \notin Nat \ {0} + BY <4>2, neg_nat_not_in_nat_setminus_0 + <7> QED + BY <7>1, <7>2 DEF iPred + <6> QED + BY <6>1, <6>2 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3>2. ASSUME NEW j \in negNat, + \A i \in negNat: j + i = i + j + PROVE \A i \in negNat: + -.Succ[-.j] + i = i + -.Succ[-.j] + <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] + <5>1. -.Succ[-.j] + 0 = iPred[j] + <6>1. -.Succ[-.j] + 0 = -.Succ[-.j] + <7>1. -.Succ[-.j] \in Int + <8>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <8>2. Succ[-.j] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7> QED + BY <7>1, add_0 + <6>2. @ = iPred[j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. j \notin Nat \ {0} + BY <3>2, neg_nat_not_in_nat_setminus_0 + <7> QED + BY <7>1, <7>2 DEF iPred + <6> QED + BY <6>1, <6>2 + <5>2. 0 + -.Succ[-.j] = iPred[j] + <6>1. 0 + -.Succ[-.j] = iPred[0 + -.-.j] + <7>1. 0 \in Int + BY zero_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[0 + j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>1, <7>2 + <6>3. @ = iPred[j + 0] + <7>1. 0 \in negNat + BY neg0 DEF negNat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + 0 = 0 + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6>4. @ = iPred[j] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7> QED + BY <7>1, add_0 + <6> QED + BY <6>1, <6>2, <6>3, <6>4 + <5> QED + BY <5>1, <5>2 + <4>2. ASSUME NEW i \in negNat, + -.Succ[-.j] + i = i + -.Succ[-.j] + PROVE -.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j] + <5>1. -.Succ[-.j] + -.Succ[-.i] = iPred[iPred[i + j]] + <6>1. -.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i] + <7>1. -.Succ[-.j] \in Int + <8>1. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <8>2. Succ[-.j] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[-.Succ[-.j] + i] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. -.-.i = i + BY <7>1, minus_involutive + <7> QED + BY <7>1, <7>2 + <6>3. @ = iPred[i + -.Succ[-.j]] + BY <4>2 + <6>4. @ = iPred[iPred[i + -.-.j]] + <7>1. i \in Int + BY <4>2, neg_nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7>3. i + -.Succ[-.j] = iPred[i + -.-.j] + BY <7>1, <7>2, addint_succ_negnat + <7> QED + BY <7>3 + <6>5. @ = iPred[iPred[i + j]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.-.j = j + BY <7>1, minus_involutive + <7> QED + BY <7>2 + <6> QED + BY <6>1, <6>2, <6>3, <6>4, <6>5 + <5>2. -.Succ[-.i] + -.Succ[-.j] = iPred[iPred[i + j]] + <6>1. -.Succ[-.i] + -.Succ[-.j] = + iPred[-.Succ[-.i] + -.-.j] + <7>1. -.Succ[-.i] \in Int + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_int + <7>2. -.j \in Nat + BY <3>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>2. @ = iPred[-.Succ[-.i] + j] + <7>1. -.-.j = j + <8>1. j \in Int + BY <3>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>1 + <6>3. @ = iPred[j + -.Succ[-.i]] + <7>1. -.Succ[-.i] \in negNat + <8>1. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <8>2. Succ[-.i] \in Nat + BY <8>1, succIsNat + <8> QED + BY <8>2, minus_nat_in_neg_nat + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j + BY <7>1, <7>2 + <7> QED + BY <7>3 + <6>4. @ = iPred[iPred[j + -.-.i]] + <7>1. j \in Int + BY <3>2, neg_nat_in_int + <7>2. -.i \in Nat + BY <4>2, minus_neg_nat_in_nat + <7> QED + BY <7>1, <7>2, addint_succ_negnat + <6>5. @ = iPred[iPred[j + i]] + <7>1. -.-.i = i + <8>1. i \in Int + BY <4>2, neg_nat_in_int + <8> QED + BY <8>1, minus_involutive + <7> QED + BY <7>1 + <6>6. @ = iPred[iPred[i + j]] + <7>1. i \in negNat + BY <4>2 + <7>2. \A k \in negNat: j + k = k + j + BY <3>2 + <7>3. j + i = i + j + BY <7>1, <7>2, spec + <7> QED + BY <7>3 + <6> QED + BY <6>1, <6>2, <6>3, <6>4, + <6>5, <6>6 + <5> QED + BY <5>1, <5>2 + <4> QED + BY <4>1, <4>2, neg_nat_induction + <3> QED + BY <3>1, <3>2, neg_nat_induction + <2> QED + BY <2>1, <2>2, int_union_nat_negnat +<1> QED + BY <1>1, <1>2 + +*) +theorem AddCommutative: + "\ j \ Int: \ i \ Int: + j + i = i + j" + proof - + have s1_1: "\ j \ Nat: \ i \ Int: + j + i = i + j" + proof - + have s2_1: "\ j \ Nat: + \ i \ Nat: + j + i = i + j" + proof - + have s3_1: "\ i \ Nat: + 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ Nat; + 0 + i = i + 0 + \ + \ + 0 + Succ[i] = Succ[i] + 0" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + Succ[i] = iSucc[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + unfolding add_def + using s6_1 s6_2 + addint_succ[of "0"] + by auto + qed + have s5_2: "iSucc[0 + i] = iSucc[i + 0]" + using s4_2_asm2 by auto + have s5_3: "iSucc[i + 0] = iSucc[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 nat_in_int + by auto + show ?thesis + unfolding add_def + using s6_1 addint_0[of "i"] + by auto + qed + have s5_4: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using s4_2_asm1 by auto + have s5_5: "Succ[i] = Succ[i] + 0" + proof - + have s6_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + have s6_2: "Succ[i] \ Int" + using s6_1 nat_in_int by auto + show ?thesis + unfolding add_def + using s6_2 addint_0 by auto + qed + show "0 + Succ[i] = Succ[i] + 0" + using s5_1 s5_2 s5_3 s5_4 s5_5 + by auto + qed + show ?thesis + using s4_1 s4_2 + natInduct[of "\ i. 0 + i = i + 0"] + by auto + qed + have s3_2: "\ j. \ + j \ Nat; + \ i \ Nat: j + i = i + j + \ + \ + \ i \ Nat: + Succ[j] + i = i + Succ[j]" + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ Nat" + assume s3_2_asm2: "\ i \ Nat: + j + i = i + j" + have s4_1: "Succ[j] + 0 = 0 + Succ[j]" + proof - + have s5_1: "Succ[j] + 0 = Succ[j]" + proof - + have s6_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat + nat_in_int by auto + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iSucc[j]" + unfolding iSucc_def + using s3_2_asm1 nat_in_int + by auto + also have s5_3: "... = iSucc[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iSucc[0 + j]" + proof - + have s6_1: "j + 0 = 0 + j" + using s3_2_asm2 zeroIsNat by auto + show ?thesis + using s6_1 by auto + qed + also have s5_5: "... = 0 + Succ[j]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + have s6_3: "0 \ Int ==> + j \ Nat ==> + 0 + Succ[j] = iSucc[0 + j]" + using addint_succ_nat + by auto + have s6_4: " + 0 + Succ[j] = iSucc[0 + j]" + using s6_1 s6_2 s6_3 by auto + have s6_5: "iSucc[0 + j] = 0 + Succ[j]" + using s6_4 by auto + show ?thesis + using s6_5 by auto + qed + finally show ?thesis . + qed + have s4_2: "\ i. \ + i \ Nat; + Succ[j] + i = i + Succ[j] + \ \ + Succ[j] + Succ[i] = Succ[i] + Succ[j]" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" + have s5_1: "Succ[j] + Succ[i] = iSucc[iSucc[j + i]]" + proof - + have s6_1: "Succ[j] + Succ[i] = iSucc[Succ[j] + i]" + proof - + have s7_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "i \ Nat" + using s4_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + also have s6_2: "... = iSucc[i + Succ[j]]" + using s4_2_asm2 by auto + also have s6_3: "... = iSucc[iSucc[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + also have s6_4: "... = iSucc[iSucc[j + i]]" + using s3_2_asm2 s4_2_asm1 spec by auto + finally show ?thesis + by auto + qed + have s5_2: "Succ[i] + Succ[j] = iSucc[iSucc[j + i]]" + proof - + have s6_1: "Succ[i] + Succ[j] = iSucc[Succ[i] + j]" + proof - + have s7_1: "Succ[i] \ Int" + using s4_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "j \ Nat" + using s3_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[j + Succ[i]]" + proof - + have s7_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + show ?thesis + using s7_1 s3_2_asm2 by auto + qed + also have s6_3: "... = iSucc[iSucc[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 nat_in_int + by auto + have s7_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat + by auto + qed + finally show ?thesis + by auto + qed + show "Succ[j] + Succ[i] = Succ[i] + Succ[j]" + using s5_1 s5_2 by auto + qed + show "\ i \ Nat: Succ[j] + i = i + Succ[j]" + using s4_1 s4_2 natInduct + by auto + qed + show ?thesis + using s3_1 s3_2 natInduct by auto + qed + have s2_2: "\ j \ Nat: + \ i \ negNat: j + i = i + j" + proof - + have s3_1: "\ i \ negNat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ negNat; + 0 + i = i + 0 + \ \ + 0 + -.Succ[-.i] = -.Succ[-.i] + 0" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ negNat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + -.Succ[-.i] = iPred[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat + by auto + have s6_3: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" + using s6_1 s6_2 addint_succ_negnat + by auto + have s6_4: "-.-.i = i" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s7_1 minus_involutive by auto + qed + show ?thesis + using s6_3 s6_4 by auto + qed + also have s5_2: "... = iPred[i + 0]" + using s4_2_asm2 by auto + also have s5_3: "... = iPred[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s6_2: "i + 0 = i" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = -.Succ[-.i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s6_2: "i \ Nat \ {0}" + using s4_2_asm1 neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s6_1 s6_2 by auto + qed + also have s5_5: "... = -.Succ[-.i] + 0" + proof - + have s6_1: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat by auto + have s6_2: "Succ[-.i] \ Nat" + using s6_1 succIsNat by auto + have s6_3: "-.Succ[-.i] \ Int" + using s6_2 minus_nat_in_int by auto + show ?thesis + using s6_3 add_0 by auto + qed + finally show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" + by auto + qed + show ?thesis + using s4_1 s4_2 neg_nat_induction by auto + qed + have s3_2: "\ j. \ + j \ Nat; + \ i \ negNat: j + i = i + j + \ \ + \ i \ negNat: Succ[j] + i = i + Succ[j]" + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ Nat" + assume s3_2_asm2: "\ i \ negNat: + j + i = i + j" + have s4_1: "Succ[j] + 0 = 0 + Succ[j]" + proof - + have s5_1: "Succ[j] + 0 = Succ[j]" + proof - + have s6_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int by auto + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iSucc[j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + unfolding iSucc_def + using s6_1 s6_2 by auto + qed + also have s5_3: "... = iSucc[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iSucc[0 + j]" + proof - + have s6_1: "0 \ negNat" + unfolding negNat_def + using neg0 by auto + show ?thesis + using s3_2_asm2 s6_1 by auto + qed + also have s5_5: "... = 0 + Succ[j]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s6_1 s6_2 addint_succ_nat by auto + qed + finally show ?thesis + by auto + qed + have s4_2: "\ i. \ + i \ negNat; + Succ[j] + i = i + Succ[j] + \ \ + Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ negNat" + assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" + have s5_1: "Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]]" + proof - + have s6_1: "Succ[j] + -.Succ[-.i] = + iPred[Succ[j] + i]" + proof - + have s7_1: "Succ[j] \ Int" + using s3_2_asm1 succIsNat nat_in_int + by auto + have s7_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat + by auto + have s7_3: "-.-.i = i" + using s4_2_asm1 neg_nat_in_int + minus_involutive by auto + have s7_4: "Succ[j] + -.Succ[-.i] = + iPred[Succ[j] + -.-.i]" + using s7_1 s7_2 addint_succ_negnat + by auto + show ?thesis + using s7_3 s7_4 by auto + qed + also have s6_2: "... = iPred[i + Succ[j]]" + using s4_2_asm2 by auto + also have s6_3: "... = iPred[iSucc[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "i \ negNat" + using s4_2_asm1 by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + i = i + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + finally show ?thesis + by auto + qed + have s5_2: "-.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]]" + proof - + have s6_1: "-.Succ[-.i] + Succ[j] = + iSucc[-.Succ[-.i] + j]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + using s4_2_asm1 minus_neg_nat_in_nat succIsNat + minus_nat_in_int by auto + have s7_2: "j \ Nat" + using s3_2_asm1 by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[j + -.Succ[-.i]]" + proof - + have s7_1: "-.Succ[-.i] \ negNat" + using s4_2_asm1 minus_neg_nat_in_nat succIsNat + minus_nat_in_neg_nat by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + also have s6_3: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_asm1 minus_neg_nat_in_nat by auto + have s7_3: "j + -.Succ[-.i] = iPred[j + -.-.i]" + using s7_1 s7_2 addint_succ_negnat by auto + have s7_4: "-.-.i = i" + proof - + have s8_1: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_3 s7_4 by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "j + i \ Int" + proof - + have s8_1: "j \ Int" + using s3_2_asm1 nat_in_int by auto + have s8_2: "i \ Int" + using s4_2_asm1 neg_nat_in_int by auto + show ?thesis + using s8_1 s8_2 addIsInt by auto + qed + show ?thesis + using s7_1 iSucciPredCommute by auto + qed + finally show ?thesis + by auto + qed + show "Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" + using s5_1 s5_2 by auto + qed + show "\ i \ negNat: Succ[j] + i = i + Succ[j]" + using s4_1 s4_2 neg_nat_induction + by auto + qed + show ?thesis + using s3_1 s3_2 natInduct by auto + qed + show ?thesis + proof - + have s3_1: "\ j. \ i. + j \ Nat => ( + i \ Int => ( + j + i = i + j))" + proof auto + fix "j" :: "c" + fix "i" :: "c" + assume s4_1: "j \ Nat" + assume s4_2: "i \ Int" + have s4_3: "i \ Nat ==> j + i = i + j" + proof - + assume s5_1: "i \ Nat" + show "j + i = i + j" + using s4_1 s5_1 s2_1 by auto + qed + have s4_4: "i \ negNat ==> j + i = i + j" + proof - + assume s5_1: "i \ negNat" + show "j + i = i + j" + using s4_1 s5_1 s2_2 by auto + qed + show "j + i = i + j" + using s4_2 s4_3 s4_4 int_union_nat_negnat by auto + qed + show ?thesis + using s3_1 by auto + qed + qed + have s1_2: "\ j \ negNat: \ i \ Int: + j + i = i + j" + proof - + have s2_1: "\ j \ negNat: \ i \ Nat: + j + i = i + j" + proof - + have s3_1: "\ i \ Nat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ Nat; + 0 + i = i + 0 + \ \ + 0 + Succ[i] = Succ[i] + 0 + " + proof - + fix "i" :: "c" + assume s4_2_asm1: "i \ Nat" + assume s4_2_asm2: "0 + i = i + 0" + have s5_1: "0 + Succ[i] = iSucc[0 + i]" + proof - + have s6_1: "0 \ Int" + using zero_in_int by auto + have s6_2: "i \ Nat" + using s4_2_asm1 by auto + show ?thesis + using s6_1 s6_2 addint_succ_nat by auto + qed + also have s5_2: "... = iSucc[i + 0]" + using s4_2_asm2 by force + also have s5_3: "... = iSucc[i]" + proof - + have s6_1: "i \ Int" + using s4_2_asm1 nat_in_int by auto + have s6_2: "i + 0 = i" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by force + qed + also have s5_4: "... = Succ[i]" + proof - + have s6_1: "i \ Nat" + using s4_2_asm1 by auto + have s6_2: "i \ Int" + using s4_2_asm1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s6_1 s6_2 by auto + qed + also have s5_5: "... = Succ[i] + 0" + proof - + have s6_1: "Succ[i] \ Nat" + using s4_2_asm1 succIsNat by auto + have s6_2: "Succ[i] \ Int" + using s6_1 nat_in_int by auto + show ?thesis + using s6_2 add_0 by auto + qed + finally show "0 + Succ[i] = Succ[i] + 0" + (* . did not work *) + using s5_1 s5_2 s5_3 s5_4 s5_5 by auto + qed + show ?thesis + using s4_1 s4_2 natInduct by auto + qed + have s3_2: "\ j. \ + j \ negNat; + \ i \ Nat: j + i = i + j + \ \ + \ i \ Nat: + -.Succ[-.j] + i = i + -.Succ[-.j] + " + proof - + fix "j" :: "c" + assume s3_2_asm1: "j \ negNat" + assume s3_2_asm2: "\ i \ Nat: + j + i = i + j" + have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" + proof - + have s6_1: "-.Succ[-.j] \ Int" + proof - + have s7_1: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat + by auto + have s7_2: "Succ[-.j] \ Nat" + using s7_1 succIsNat by auto + show ?thesis + using s7_2 minus_nat_in_int by auto + qed + show ?thesis + using s6_1 add_0 by auto + qed + also have s5_2: "... = iPred[j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "j \ Nat \ {0}" + using s3_2_asm1 neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s6_1 s6_2 by auto + qed + also have s5_3: "... = iPred[j + 0]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "j + 0 = j" + using s6_1 add_0 by auto + show ?thesis + using s6_2 by auto + qed + also have s5_4: "... = iPred[0 + j]" + proof - + have s6_1: "0 \ Nat" + using zeroIsNat by auto + have s6_2: "\ k \ Nat: + j + k = k + j" + using s3_2_asm2 by auto + have s6_3: "j + 0 = 0 + j" + using s6_1 s6_2 by auto + show ?thesis + using s6_3 by auto + qed + also have s5_5: "... = iPred[0 + -.-.j]" + proof - + have s6_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s6_2: "-.-.j = j" + using s6_1 minus_involutive by auto + show ?thesis + using s6_2 by auto + qed + also have s5_6: "... = 0 + -.Succ[-.j]" + proof - + have s6_1: "0 \ Int" + using zeroIsNat nat_in_int by auto + have s6_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + show ?thesis + using s6_1 s6_2 addint_succ_negnat by auto + qed + finally show ?thesis . + qed + have s4_2: "\ i. \ + i \ Nat; + -.Succ[-.j] + i = i + -.Succ[-.j] + \ \ + -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ Nat" + assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" + have s5_1: "-.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]]" + proof - + have s6_1: "-.Succ[-.j] + Succ[i] = + iSucc[-.Succ[-.j] + i]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "j \ negNat" + using s3_2_asm1 by auto + have s8_2: "-.j \ Nat" + using s8_1 minus_neg_nat_in_nat by auto + have s8_3: "Succ[-.j] \ Nat" + using s8_2 succIsNat by auto + show ?thesis + using s8_3 minus_nat_in_int by auto + qed + have s7_2: "i \ Nat" + using s4_2_idom by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_2: "... = iSucc[i + -.Succ[-.j]]" + using s4_2_induct by auto + also have s6_3: "... = iSucc[iPred[i + j]]" + proof - + have s7_1: "i \ Int" + using s4_2_idom nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + have s7_3: "-.-.j = j" + using s3_2_asm1 neg_nat_in_int + minus_involutive by auto + show ?thesis + using s7_1 s7_2 s7_3 addint_succ_negnat + by auto + qed + also have s6_4: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "i \ Nat" + using s4_2_idom by auto + have s7_2: "j + i = i + j" + using s3_2_asm2 s7_1 spec by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]]" + proof - + have s6_1: "Succ[i] + -.Succ[-.j] = + iPred[Succ[i] + -.-.j]" + proof - + have s7_1: "Succ[i] \ Int" + using s4_2_idom succIsNat nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_asm1 minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[Succ[i] + j]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_2 by auto + qed + also have s6_3: "... = iPred[j + Succ[i]]" + proof - + have s7_1: "Succ[i] \ Nat" + using s4_2_idom succIsNat by auto + have s7_2: "\ k \ Nat: + j + k = k + j" + using s3_2_asm2 by auto + have s7_3: "j + Succ[i] = Succ[i] + j" + using s7_1 s7_2 by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[iSucc[j + i]]" + proof - + have s7_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s7_2: "i \ Nat" + using s4_2_idom by auto + show ?thesis + using s7_1 s7_2 addint_succ_nat by auto + qed + also have s6_5: "... = iSucc[iPred[j + i]]" + proof - + have s7_1: "j + i \ Int" + proof - + have s8_1: "j \ Int" + using s3_2_asm1 neg_nat_in_int by auto + have s8_2: "i \ Int" + using s4_2_idom nat_in_int by auto + show ?thesis + using s8_1 s8_2 addIsInt by auto + qed + show ?thesis + using s7_1 iSucciPredCommute by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" + using s5_1 s5_2 by auto + qed + show "\ i \ Nat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + using s4_1 s4_2 natInduct by auto + qed + show ?thesis + using s3_1 s3_2 neg_nat_induction by auto + qed + have s2_2: "\ j \ negNat: \ i \ negNat: + j + i = i + j" + proof - + have s3_1: "\ i \ negNat: 0 + i = i + 0" + proof - + have s4_1: "0 + 0 = 0 + 0" + by auto + have s4_2: "\ i. \ + i \ negNat; + 0 + i = i + 0 + \ \ + 0 + -.Succ[-.i] = -.Succ[-.i] + 0" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ negNat" + assume s4_2_induct: "0 + i = i + 0" + have s5_1: "0 + -.Succ[-.i] = iPred[i]" + proof - + have s6_1: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" + proof - + have s7_1: "0 \ Int" + using zero_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[0 + i]" + proof - + have s7_1: "-.-.i = i" + using s4_2_idom neg_nat_in_int + minus_involutive by auto + show ?thesis + using s7_1 by auto + qed + also have s6_3: "... = iPred[i + 0]" + using s4_2_induct by auto + also have s6_4: "... = iPred[i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "i + 0 = i" + using s7_1 add_0 by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "-.Succ[-.i] + 0 = iPred[i]" + proof - + have s6_1: "-.Succ[-.i] + 0 = -.Succ[-.i]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + show ?thesis + using s7_1 add_0 by auto + qed + also have s6_2: "... = iPred[i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "i \ Nat \ {0}" + using s4_2_idom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s7_1 s7_2 by auto + qed + finally show ?thesis . + qed + show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" + using s5_1 s5_2 by auto + qed + show ?thesis + using s4_1 s4_2 neg_nat_induction by auto + qed + have s3_2: "\ j. \ + j \ negNat; + \ i \ negNat: j + i = i + j + \ \ + \ i \ negNat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + proof - + fix "j" :: "c" + assume s3_2_jdom: "j \ negNat" + assume s3_2_induct: "\ i \ negNat: + j + i = i + j" + have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] + 0 = iPred[j]" + proof - + have s6_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.j] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + show ?thesis + using s7_1 add_0 by auto + qed + also have s6_2: "... = iPred[j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "j \ Nat \ {0}" + using s3_2_jdom neg_nat_not_in_nat_setminus_0 + by auto + show ?thesis + unfolding iPred_def + using s7_1 s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "0 + -.Succ[-.j] = iPred[j]" + proof - + have s6_1: "0 + -.Succ[-.j] = iPred[0 + -.-.j]" + proof - + have s7_1: "0 \ Int" + using zero_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[0 + j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_1 s7_2 by auto + qed + also have s6_3: "... = iPred[j + 0]" + proof - + have s7_1: "0 \ negNat" + unfolding negNat_def + using neg0 by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + 0 = 0 + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[j]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + show ?thesis + using s7_1 add_0 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s5_1 s5_2 by auto + qed + have s4_2: "\ i. \ + i \ negNat; + -.Succ[-.j] + i = i + -.Succ[-.j] + \ \ + -.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j]" + proof - + fix "i" :: "c" + assume s4_2_idom: "i \ negNat" + assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" + have s5_1: "-.Succ[-.j] + -.Succ[-.i] = + iPred[iPred[i + j]]" + proof - + have s6_1: "-.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i]" + proof - + have s7_1: "-.Succ[-.j] \ Int" + proof - + have s8_1: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.j] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s7_3: "-.Succ[-.j] \ Int + ==> + -.i \ Nat + ==> + -.Succ[-.j] + -.Succ[-.i] = + iPred[-.Succ[-.j] + -.-.i]" + using addint_succ_negnat[ + of "-.Succ[-.j]" "-.i"] + by auto + show ?thesis + by (rule s7_3, rule s7_1, rule s7_2) + qed + also have s6_2: "... = iPred[-.Succ[-.j] + i]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "-.-.i = i" + using s7_1 minus_involutive by auto + show ?thesis + using s7_1 s7_2 by auto + qed + also have s6_3: "... = iPred[i + -.Succ[-.j]]" + using s4_2_induct by auto + also have s6_4: "... = iPred[iPred[i + -.-.j]]" + proof - + have s7_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat by auto + have s7_3: "i + -.Succ[-.j] = iPred[i + -.-.j]" + using s7_1 s7_2 addint_succ_negnat by auto + show ?thesis + using s7_3 by auto + qed + also have s6_5: "... = iPred[iPred[i + j]]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.-.j = j" + using s7_1 minus_involutive by auto + show ?thesis + using s7_2 by auto + qed + finally show ?thesis . + qed + have s5_2: "-.Succ[-.i] + -.Succ[-.j] = + iPred[iPred[i + j]]" + proof - + have s6_1: "-.Succ[-.i] + -.Succ[-.j] = + iPred[-.Succ[-.i] + -.-.j]" + proof - + have s7_1: "-.Succ[-.i] \ Int" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_int by auto + qed + have s7_2: "-.j \ Nat" + using s3_2_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_2: "... = iPred[-.Succ[-.i] + j]" + proof - + have s7_1: "-.-.j = j" + proof - + have s8_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_1 by auto + qed + also have s6_3: "... = iPred[j + -.Succ[-.i]]" + proof - + have s7_1: "-.Succ[-.i] \ negNat" + proof - + have s8_1: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat + by auto + have s8_2: "Succ[-.i] \ Nat" + using s8_1 succIsNat by auto + show ?thesis + using s8_2 minus_nat_in_neg_nat by auto + qed + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" + using s7_1 s7_2 by auto + show ?thesis + using s7_3 by auto + qed + also have s6_4: "... = iPred[iPred[j + -.-.i]]" + proof - + have s7_1: "j \ Int" + using s3_2_jdom neg_nat_in_int by auto + have s7_2: "-.i \ Nat" + using s4_2_idom minus_neg_nat_in_nat by auto + show ?thesis + using s7_1 s7_2 addint_succ_negnat by auto + qed + also have s6_5: "... = iPred[iPred[j + i]]" + proof - + have s7_1: "-.-.i = i" + proof - + have s8_1: "i \ Int" + using s4_2_idom neg_nat_in_int by auto + show ?thesis + using s8_1 minus_involutive by auto + qed + show ?thesis + using s7_1 by auto + qed + also have s6_6: "... = iPred[iPred[i + j]]" + proof - + have s7_1: "i \ negNat" + using s4_2_idom by auto + have s7_2: "\ k \ negNat: + j + k = k + j" + using s3_2_induct by auto + have s7_3: "j + i = i + j" + using s7_1 s7_2 spec by auto + show ?thesis + using s7_3 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.j] + -.Succ[-.i] = + -.Succ[-.i] + -.Succ[-.j]" + using s5_1 s5_2 by auto + qed + show "\ i \ negNat: + -.Succ[-.j] + i = i + -.Succ[-.j]" + using s4_1 s4_2 neg_nat_induction by auto + qed + show ?thesis + using s3_1 s3_2 neg_nat_induction by auto + qed + show ?thesis + proof - + have s3_1: "\ j. \ i. + j \ negNat => ( + i \ Int => ( + j + i = i + j))" + proof auto + fix "j" :: "c" + fix "i" :: "c" + assume s4_1: "j \ negNat" + assume s4_2: "i \ Int" + have s4_3: "i \ Nat ==> j + i = i + j" + proof - + assume s5_1: "i \ Nat" + show "j + i = i + j" + using s4_1 s5_1 s2_1 by auto + qed + have s4_4: "i \ negNat ==> j + i = i + j" + proof - + assume s5_1: "i \ negNat" + show "j + i = i + j" + using s4_1 s5_1 s2_2 by auto + qed + show "j + i = i + j" + using s4_2 s4_3 s4_4 int_union_nat_negnat by auto + qed + show ?thesis + using s3_1 by auto + qed + qed + show ?thesis + using s1_1 s1_2 int_union_nat_negnat by auto + qed + + +theorem AddCommutative_sequent: + assumes mdom: "m \ Int" and ndom: "n \ Int" + shows "m + n = n + m" + using mdom ndom AddCommutative by auto + + +theorem add_0_left: + assumes mdom: "m \ Int" + shows "0 + m = m" + proof - + have s1_1: "m + 0 = m" + using mdom add_0 by auto + have s1_2: "m + 0 = 0 + m" + using mdom zero_in_int AddCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + + +(*----------------------------------------------------------------------------*) +(* Associativity of addition. *) + + +(* +THEOREM iSuccRightDistributesAdd == + ASSUME NEW a \in Int + PROVE \A b \in Int: iSucc[a + b] = a + iSucc[b] +<1>1. ASSUME NEW b \in Nat + PROVE iSucc[a + b] = a + iSucc[b] + <2>1. iSucc[a + b] = a + Succ[b] + <3>1. a \in Int + OBVIOUS + <3>2. b \in Nat + BY <1>1 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>2. @ = a + iSucc[b] + <3>1. b \in Nat + BY <1>1 + <3>2. b \in Int + BY <1>1, nat_in_int + <3>3. iSucc[b] = Succ[b] + BY <3>1, <3>2 DEF iSucc + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2 +<1>2. \A b \in negNat: iSucc[a + b] = a + iSucc[b] + <2>1. iSucc[a + 0] = a + iSucc[0] + <3>1. iSucc[a + 0] = a + Succ[0] + <4>1. a \in Int + OBVIOUS + <4>2. 0 \in Nat + BY zeroIsNat + <4> QED + BY <4>1, <4>2, addint_succ_nat + <3>2. @ = a + iSucc[0] + <4>1. 0 \in Nat + BY zeroIsNat + <4>2. 0 \in Int + BY zeroIsNat, nat_in_int + <4>3. iSucc[0] = Succ[0] + BY <4>1, <4>2 DEF iSucc + <4> QED + BY <4>3 + <3> QED + BY <3>1, <3>2 + <2>2. ASSUME NEW b \in negNat, + iSucc[a + b] = a + iSucc[b] + PROVE iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]] + <3>1. iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]] + <4>1. a \in Int + OBVIOUS + <4>2. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <4>3. a + -.Succ[-.b] = iPred[a + -.-.b] + BY <4>1, <4>2, addint_succ_negnat + <4>4. -.-.b = b + <5>1. b \in Int + BY <2>2, neg_nat_in_int + <5> QED + BY <5>1, minus_involutive + <4>5. a + -.Succ[-.b] = iPred[a + b] + BY <4>3, <4>4 + <4> QED + BY <4>5 + <3>2. @ = iPred[iSucc[a + b]] + <4>1. a + b \in Int + <5>1. a \in Int + OBVIOUS + <5>2. b \in Int + BY <2>2, neg_nat_in_int + <5> QED + BY <5>1, <5>2, addIsInt + <4> QED + BY <4>1, iSucciPredCommute + <3>3. CASE b = 0 + <4>1. iPred[iSucc[a + b]] = a + <5>1. iPred[iSucc[a + b]] = iPred[iSucc[a + 0]] + BY <3>3 + <5>2. @ = iPred[iSucc[a]] + <6>1. a + 0 = a + BY a \in Int, add_0 + <6> QED + BY <6>1 + <5>3. @ = a + <6>1. a \in Int + OBVIOUS + <6> QED + BY <6>1, iSucciPredCommute, spec + <5> QED + BY <5>1, <5>2, <5>3 + <4>2. a + iSucc[-.Succ[-.b]] = a + <5>1. a + iSucc[-.Succ[-.b]] = a + iSucc[-.Succ[0]] + BY <3>3 + <5>2. @ = a + iSucc[-.1] + OBVIOUS \* 1 is an abbreviation + <5>3. @ = a + -.Pred[-.-.1] + <6>1. -.1 \in Int + BY oneIsNat, minus_nat_in_int + <6>2. -.1 \notin Nat + <7>1. 1 \in Nat \ {0} + BY oneIsNat succNotZero + <7> QED + BY <7>1, minus_nat_0_in_negnat_0, + neg_nat_0_not_in_nat + <6>3. iSucc[-.1] = -.Pred[-.-.1] + BY <6>1, <6>2 DEF iSucc + <6> QED + BY <6>3 + <5>4. @ = a + -.Pred[1] + <6>1. -.-.1 = 1 + <7>1. 1 \in Int + BY oneIsNat nat_in_int + <7> QED + BY <7>1, minus_involutive + <6> QED + BY <6>1 + <5>5. @ = a + -.0 + BY pred_1 + <5>6. @ = a + 0 + BY neg0 + <5>7. @ = a + BY a \in Int, add_0 + <5> QED + BY <5>1, <5>2, <5>3, <5>4, + <5>5, <5>6, <5>7 + <4> QED + BY <3>1, <3>2, <4>1, <4>2 + <3>4. CASE b # 0 + <4>1. iPred[iSucc[a + b]] = + iPred[a + iSucc[b]] + BY <2>2 + <4>2. @ = a + -.Succ[-.iSucc[b]] + <5>1. a \in Int + OBVIOUS + <5>2. /\ -.iSucc[b] \in Nat + /\ iSucc[b] \in negNat + <6>1. b \in negNat \ {0} + BY <2>2, <3>4 + <6>2. iSucc[b] = -.Pred[-.b] + <7>1. b \in Int + BY <6>1, neg_nat_in_int + <7>2. b \notin Nat + BY <6>1, neg_nat_0_not_in_nat + <7> QED + BY <7>1, <7>2 DEF iSucc + <6>3. Pred[-.b] \in Nat + <7>1. -.b \in Nat \ {0} + BY <6>1, minus_neg_nat_0_in_nat_0 + <7> QED + BY <7>1, Pred_in_nat + <6>4. -.iSucc[b] \in Nat + <7>1. -.iSucc[b] = Pred[-.b] + <8>1. -.iSucc[b] = -.-.Pred[-.b] + BY <6>2 + <8>2. Pred[-.b] \in Int + BY <6>3, nat_in_int + <8>3. -.-.Pred[-.b] = Pred[-.b] + BY <8>2, minus_involutive + <8> QED + BY <8>1, <8>3 + <7> QED + BY <6>3, <7>1 + <6>5. iSucc[b] \in negNat + BY <6>2, <6>3, minus_nat_in_neg_nat + <6> QED + BY <6>4, <6>5 + <5>3. a + -.Succ[-.iSucc[b]] = + iPred[a + -.-.iSucc[b]] + BY <5>1, <5>2, addint_succ_negnat + <5>4. -.-.iSucc[b] = iSucc[b] + <6>1. iSucc[b] \in Int + BY <5>2, neg_nat_in_int + <6> QED + BY <6>1, minus_involutive + <5> QED + BY <5>3, <5>4 + <4>3. @ = a + -.Succ[Pred[-.b]] + <5>1. b \in Int + BY <2>2, neg_nat_in_int + <5>2. b \notin Nat + <6>1. b \in negNat \ {0} + BY <2>2, <3>4 + <6> QED + BY <6>1, neg_nat_0_not_in_nat + <5>3. iSucc[b] = -.Pred[-.b] + BY <5>1, <5>2 DEF iSucc + <5>4. -.iSucc[b] = -.-.Pred[-.b] + BY <5>3 + <5>5. -.-.Pred[-.b] = Pred[-.b] + <6>1. -.b \in Nat + BY <2>2, <3>4, minus_neg_nat_0_in_nat_0 + <6>2. Pred[-.b] \in Nat + BY <6>1, Pred_in_nat + <6>3. Pred[-.b] \in Int + BY <6>2, nat_in_int + <6> QED + BY <6>3, minus_involutive + <5>6. -.iSucc[b] = Pred[-.b] + BY <5>4, <5>5 + <5> QED + BY <5>6 + <4>4. @ = a + -.Pred[Succ[-.b]] + <5>1. b \in negNat \ {0} + BY <2>2, <3>4 + <5>2. -.b \in Nat \ {0} + BY <5>1, minus_neg_nat_0_in_nat_0 + <5>3. Succ[Pred[-.b]] = Pred[Succ[-.b]] + BY <5>2, Succ_Pred_nat, Pred_Succ_nat + <5> QED + BY <5>3 + <4>5. @ = a + -.Pred[-.-.Succ[-.b]] + <5>1. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <5>2. Succ[-.b] \in Nat + BY <5>1, succIsNat + <5>3. Succ[-.b] \in Int + BY <5>2, nat_in_int + <5>4. -.-.Succ[-.b] = Succ[-.b] + BY <5>3, minus_involutive + <5> QED + BY <5>4 + <4>6. @ = a + iSucc[-.Succ[-.b]] + <5>1. -.b \in Nat + BY <2>2, minus_neg_nat_in_nat + <5>2. Succ[-.b] \in Nat \ {0} + BY <5>1, succIsNat, succNotZero + <5>3. -.Succ[-.b] \in negNat \ {0} + BY <5>2, minus_nat_0_in_negnat_0 + <5>4. -.Succ[-.b] \in Int + BY <5>3, neg_nat_in_int + <5>5. -.Succ[-.b] \notin Nat + BY <5>3, neg_nat_0_not_in_nat + <5> QED + BY <5>4, <5>5 DEF iSucc + <4> QED + BY <3>1, <3>2, <4>1, <4>2, <4>3, + <4>4, <4>5, <4>6 + <3> QED + BY <3>3, <3>4 + <2> QED + BY <2>1, <2>2, neg_nat_induction +<1> QED + BY <1>1, <1>2 +*) +theorem iSuccRightDistributesAdd: + "\ a \ Int: \ b \ Int: + iSucc[a + b] = a + iSucc[b]" + proof - + { + fix "a" :: "c" + assume adom: "a \ Int" + have s1_1: "\ b \ Nat: iSucc[a + b] = a + iSucc[b]" + proof auto + fix "b" :: "c" + assume bdom: "b \ Nat" + have s2_1: "iSucc[a + b] = a + Succ[b]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b \ Nat" + using bdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + also have s2_2: "... = a + iSucc[b]" + proof - + have s3_1: "b \ Nat" + using bdom by auto + have s3_2: "b \ Int" + using bdom nat_in_int by auto + have s3_3: "iSucc[b] = Succ[b]" + unfolding iSucc_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "iSucc[a + b] = a + iSucc[b]" . + qed + have s1_2: "\ b \ negNat: iSucc[a + b] = a + iSucc[b]" + proof - + have s2_1: "iSucc[a + 0] = a + iSucc[0]" + proof - + have s3_1: "iSucc[a + 0] = a + Succ[0]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s4_1 s4_2 addint_succ_nat by auto + qed + have s3_2: "... = a + iSucc[0]" + proof - + have s4_1: "0 \ Nat" + using zeroIsNat by auto + have s4_2: "0 \ Int" + using zeroIsNat nat_in_int by auto + have s4_3: "iSucc[0] = Succ[0]" + unfolding iSucc_def + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + have s2_2: "\ b. \ + b \ negNat; + iSucc[a + b] = a + iSucc[b] + \ \ + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + fix "b" :: "c" + assume s2_2_bdom: "b \ negNat" + assume s2_2_induct: "iSucc[a + b] = a + iSucc[b]" + have s3_1: "iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s4_3: "a + -.Succ[-.b] = iPred[a + -.-.b]" + using s4_1 s4_2 addint_succ_negnat by auto + have s4_4: "-.-.b = b" + proof - + have s5_1: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + show ?thesis + using s5_1 minus_involutive by auto + qed + have s4_5: "a + -.Succ[-.b] = iPred[a + b]" + using s4_3 s4_4 by auto + show ?thesis + using s4_5 by auto + qed + have s3_2: "... = iPred[iSucc[a + b]]" + proof - + have s4_1: "a + b \ Int" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + show ?thesis + using s5_1 s5_2 addIsInt by auto + qed + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + have s3_3: "b = 0 ==> + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + assume s3_3_bdom: "b = 0" + have s4_1: "iPred[iSucc[a + b]] = a" + proof - + have s5_1: "iPred[iSucc[a + b]] = iPred[iSucc[a + 0]]" + using s3_3_bdom by auto + also have s5_2: "... = iPred[iSucc[a]]" + proof - + have s6_1: "a + 0 = a" + using adom add_0 by auto + show ?thesis + using s6_1 by auto + qed + also have s5_3: "... = a" + using adom iSucciPredCommute spec by auto + finally show ?thesis . + qed + have s4_2: "a + iSucc[-.Succ[-.b]] = a" + proof - + have s5_1: "a + iSucc[-.Succ[-.b]] = + a + iSucc[-.Succ[0]]" + using s3_3_bdom by auto + have s5_2: "... = a + iSucc[-.1]" + by auto + have s5_3: "... = a + -.Pred[-.-.1]" + (* `also` raises error: + empty result sequence + *) + proof - + have s6_1: "-.1 \ Int" + using oneIsNat minus_nat_in_int by auto + have s6_2: "-.1 \ Nat" + proof - + have s7_1: "1 \ Nat \ {0}" + using oneIsNat succNotZero by auto + show ?thesis + using s7_1 minus_nat_0_in_negnat_0 + neg_nat_0_not_in_nat by auto + qed + have s6_3: "iSucc[-.1] = -.Pred[-.-.1]" + unfolding iSucc_def + using s6_1 s6_2 by auto + show ?thesis + using s6_3 by auto + qed + have s5_4: "... = a + -.Pred[1]" + proof - + have s6_1: "-.-.1 = 1" + proof - + have s7_1: "1 \ Int" + using oneIsNat nat_in_int by auto + show ?thesis + using s7_1 minus_involutive by auto + qed + show ?thesis + using s6_1 by auto + qed + have s5_5: "... = a + -.0" + using pred_1 by auto + have s5_6: "... = a + 0" + using neg0 by auto + have s5_7: "... = a" + using adom add_0 by auto + show ?thesis + using s5_1 s5_2 s5_3 s5_4 s5_5 s5_6 s5_7 + by auto + qed + show ?thesis + using s3_1 s3_2 s4_1 s4_2 by auto + qed + have s3_4: "b \ 0 ==> + iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + proof - + assume s3_4_bdom: "b \ 0" + have s4_1: "iPred[iSucc[a + b]] = + iPred[a + iSucc[b]]" + using s2_2_induct by auto + also have s4_2: "... = a + -.Succ[-.iSucc[b]]" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "-.iSucc[b] \ Nat \ + iSucc[b] \ negNat" + proof - + have s6_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + have s6_2: "iSucc[b] = -.Pred[-.b]" + proof - + have s7_1: "b \ Int" + using s6_1 neg_nat_in_int by auto + have s7_2: "b \ Nat" + using s6_1 neg_nat_0_not_in_nat by auto + show ?thesis + unfolding iSucc_def + using s7_1 s7_2 by auto + qed + have s6_3: "Pred[-.b] \ Nat" + proof - + have s7_1: "-.b \ Nat \ {0}" + using s6_1 minus_neg_nat_0_in_nat_0 by auto + show ?thesis + using s7_1 Pred_in_nat by auto + qed + have s6_4: "-.iSucc[b] \ Nat" + proof - + have s7_1: "-.iSucc[b] = Pred[-.b]" + proof - + have s8_1: "-.iSucc[b] = -.-.Pred[-.b]" + using s6_2 by auto + have s8_2: "Pred[-.b] \ Int" + using s6_3 nat_in_int by auto + have s8_3: "-.-.Pred[-.b] = Pred[-.b]" + using s8_2 minus_involutive by auto + show ?thesis + using s8_1 s8_3 by auto + qed + show ?thesis + using s6_3 s7_1 by auto + qed + have s6_5: "iSucc[b] \ negNat" + using s6_2 s6_3 minus_nat_in_neg_nat by auto + show ?thesis + using s6_4 s6_5 by auto + qed + have s5_3: "a + -.Succ[-.iSucc[b]] = + iPred[a + -.-.iSucc[b]]" + using s5_1 s5_2 addint_succ_negnat by auto + have s5_4: "-.-.iSucc[b] = iSucc[b]" + proof - + have s6_1: "iSucc[b] \ Int" + using s5_2 neg_nat_in_int by auto + show ?thesis + using s6_1 minus_involutive by auto + qed + show ?thesis + using s5_3 s5_4 by auto + qed + also have s4_3: "... = a + -.Succ[Pred[-.b]]" + proof - + have s5_1: "b \ Int" + using s2_2_bdom neg_nat_in_int by auto + have s5_2: "b \ Nat" + proof - + have s6_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + show ?thesis + using s6_1 neg_nat_0_not_in_nat by auto + qed + have s5_3: "iSucc[b] = -.Pred[-.b]" + unfolding iSucc_def + using s5_1 s5_2 by auto + have s5_4: "-.iSucc[b] = -.-.Pred[-.b]" + using s5_3 by auto + have s5_5: "-.-.Pred[-.b] = Pred[-.b]" + proof - + have s6_1: "-.b \ Nat" + using s2_2_bdom s3_4_bdom + minus_neg_nat_0_in_nat_0 + by auto + have s6_2: "Pred[-.b] \ Nat" + using s6_1 Pred_in_nat by auto + have s6_3: "Pred[-.b] \ Int" + using s6_2 nat_in_int by auto + show ?thesis + using s6_3 minus_involutive by auto + qed + have s5_6: "-.iSucc[b] = Pred[-.b]" + using s5_4 s5_5 by auto + show ?thesis + using s5_6 by auto + qed + also have s4_4: "... = a + -.Pred[Succ[-.b]]" + proof - + have s5_1: "b \ negNat \ {0}" + using s2_2_bdom s3_4_bdom by auto + have s5_2: "-.b \ Nat \ {0}" + using s5_1 minus_neg_nat_0_in_nat_0 by auto + have s5_3: "Succ[Pred[-.b]] = Pred[Succ[-.b]]" + using s5_2 Succ_Pred_nat Pred_Succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = a + -.Pred[-.-.Succ[-.b]]" + proof - + have s5_1: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s5_2: "Succ[-.b] \ Nat" + using s5_1 succIsNat by auto + have s5_3: "Succ[-.b] \ Int" + using s5_2 nat_in_int by auto + have s5_4: "-.-.Succ[-.b] = Succ[-.b]" + using s5_3 minus_involutive by auto + show ?thesis + using s5_4 by auto + qed + also have s4_6: "... = a + iSucc[-.Succ[-.b]]" + proof - + have s5_1: "-.b \ Nat" + using s2_2_bdom minus_neg_nat_in_nat by auto + have s5_2: "Succ[-.b] \ Nat \ {0}" + using s5_1 succIsNat succNotZero by auto + have s5_3: "-.Succ[-.b] \ negNat \ {0}" + using s5_2 minus_nat_0_in_negnat_0 by auto + have s5_4: "-.Succ[-.b] \ Int" + using s5_3 neg_nat_in_int by auto + have s5_5: "-.Succ[-.b] \ Nat" + using s5_3 neg_nat_0_not_in_nat by auto + show ?thesis + unfolding iSucc_def + using s5_4 s5_5 by auto + qed + finally show "iSucc[a + -.Succ[-.b]] = + a + iSucc[-.Succ[-.b]]" + using s3_1 s3_2 by auto + qed + show "iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" + using s3_3 s3_4 by blast + qed + show ?thesis + using s2_1 s2_2 neg_nat_induction by auto + qed + have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" + using s1_1 s1_2 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem iSuccRightDistributesAdd_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "iSucc[a + b] = a + iSucc[b]" + using adom bdom iSuccRightDistributesAdd by auto + + +(* +THEOREM iPredRightDistributesAdd == + ASSUME NEW a \in Int + PROVE \A b \in Int: iPred[a + b] = a + iPred[b] +PROOF +<1>1. iPred[a + 0] = a + iPred[0] + <2>1. iPred[a + 0] = iPred[a + -.0] + BY neg0 + <2>2. @ = a + -.Succ[0] + <3>1. a \in Int + OBVIOUS + <3>2. 0 \in Nat + BY zeroIsNat + <3> QED + BY <3>1, <3>2, addint_succ_negnat + <2>3. @ = a + -.1 + OBVIOUS \* 1 is an abbreviation + <2>4. @ = a + iPred[0] + BY ipred_0 + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>2. \A b \in Nat: iPred[a + b] = a + iPred[b] + <2>1. ASSUME NEW b \in Nat, + iPred[a + b] = a + iPred[b] + PROVE iPred[a + Succ[b]] = a + iPred[Succ[b]] + <3>1. iPred[a + Succ[b]] = iPred[iSucc[a + b]] + <4>1. a \in Int + OBVIOUS + <4>2. b \in Nat + BY <2>1 + <4>3. a + Succ[b] = iSucc[a + b] + BY <4>1, <4>2, addint_succ_nat + <4> QED + BY <4>3 + <3>2. @ = iSucc[iPred[a + b]] + <4>1. a + b \in Int + <5>1. a \in Int + OBVIOUS + <5>2. b \in Int + BY <2>2, nat_in_int + <5> QED + BY <5>1, <5>2, addIsInt + <4> QED + BY <4>1, iSucciPredCommute + <3>3. @ = iSucc[a + iPred[b]] + BY <2>1 + <3>4. @ = a + iSucc[iPred[b]] + <4>1. b \in Int + BY <2>1, nat_in_int + <4>2. iPred[b] \in Int + BY <4>1, iPred_type + <4>3. a \in Int + OBVIOUS + <4> QED + BY <4>2, <4>3, iSuccRightDistributesAdd + <3>5. @ = a + iPred[iSucc[b]] + <4>1. b \in Int + BY <2>1, nat_in_int + <4> QED + BY <4>1, iSucciPredCommute + <3>6. @ = a + iPred[Succ[b]] + <4>1. b \in Nat + BY <2>1 + <4>2. b \in Int + BY <2>1, nat_in_int + <4>3. iSucc[b] = Succ[b] + BY <4>1, <4>2, nat_in_int DEF iSucc + <4> QED + BY <4>3 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 + <2> QED + BY <2>1, <1>1, NatInduction +<1>3. \A b \in negNat \ {0}: iPred[a + b] = a + iPred[b] + <2>1. ASSUME NEW b \in negNat \ {0} + PROVE iPred[a + b] = a + iPred[b] + <3>1. iPred[a + b] = iPred[a + -.-.b] + <4>1. b \in Int + BY <2>1, neg_nat_in_int + <4>2. -.-.b = b + BY <4>1, minus_involutive + <4> QED + BY <4>2 + <3>2. @ = a + -.Succ[-.b] + <4>1. a \in Int + OBVIOUS + <4>2. -.b \in Nat + BY <2>1, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, addint_succ_negnat + <3>3. @ = a + iPred[b] + <4>1. b \in Int + BY <2>1, neg_nat_in_int + <4>2. b \notin Nat \ {0} + BY <2>1, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3> QED + BY <3>1, <3>2, <3>3 + <2> QED + BY <2>1 +<1> QED + BY <1>2, <1>3 +*) +theorem iPredRightDistributesAdd: + "\ a \ Int: \ b \ Int: + iPred[a + b] = a + iPred[b]" + proof - + { + fix "a" :: "c" + assume adom: "a \ Int" + have s1_1: "iPred[a + 0] = a + iPred[0]" + proof - + have s2_1: "iPred[a + 0] = iPred[a + -.0]" + using neg0 by auto + also have s2_2: "... = a + -.Succ[0]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "0 \ Nat" + using zeroIsNat by auto + show ?thesis + using s3_1 s3_2 addint_succ_negnat by auto + qed + also have s2_3: "... = a + -.1" + by auto + have s2_4: "... = a + iPred[0]" + using ipred_0 by auto + (* `also` raises an error here *) + finally show ?thesis + using s2_4 by auto + qed + have s1_2: "\ b \ Nat: + iPred[a + b] = a + iPred[b]" + proof - + have s2_1: "\ b. \ + b \ Nat; + iPred[a + b] = a + iPred[b] + \ \ + iPred[a + Succ[b]] = a + iPred[Succ[b]]" + proof - + fix "b" :: "c" + assume s2_1_bdom: "b \ Nat" + assume s2_1_induct: "iPred[a + b] = a + iPred[b]" + have s3_1: "iPred[a + Succ[b]] = iPred[iSucc[a + b]]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "b \ Nat" + using s2_1_bdom by auto + have s4_3: "a + Succ[b] = iSucc[a + b]" + using s4_1 s4_2 addint_succ_nat by auto + show ?thesis + using s4_3 by auto + qed + also have s3_2: "... = iSucc[iPred[a + b]]" + proof - + have s4_1: "a + b \ Int" + proof - + have s5_1: "a \ Int" + using adom by auto + have s5_2: "b \ Int" + using s2_1_bdom nat_in_int by auto + show ?thesis + using s5_1 s5_2 addIsInt by auto + qed + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + also have s3_3: "... = iSucc[a + iPred[b]]" + using s2_1_induct by auto + also have s3_4: "... = a + iSucc[iPred[b]]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom nat_in_int by auto + have s4_2: "iPred[b] \ Int" + using s4_1 iPred_type by auto + have s4_3: "a \ Int" + using adom by auto + show ?thesis + using s4_2 s4_3 iSuccRightDistributesAdd + by auto + qed + also have s3_5: "... = a + iPred[iSucc[b]]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom nat_in_int by auto + show ?thesis + using s4_1 iSucciPredCommute by auto + qed + also have s3_6: "... = a + iPred[Succ[b]]" + proof - + have s4_1: "b \ Nat" + using s2_1_bdom by auto + have s4_2: "b \ Int" + using s2_1_bdom nat_in_int by auto + have s4_3: "iSucc[b] = Succ[b]" + unfolding iSucc_def + using s4_1 s4_2 by auto + show ?thesis + using s4_3 by auto + qed + finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . + qed + show ?thesis + using s1_1 s2_1 natInduct by auto + qed + have s1_3: "\ b \ negNat \ {0}: + iPred[a + b] = a + iPred[b]" + proof - + { + fix "b" :: "c" + assume s2_1_bdom: "b \ negNat \ {0}" + { + have s3_1: "iPred[a + b] = iPred[a + -.-.b]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom neg_nat_in_int by auto + have s4_2: "-.-.b = b" + using s4_1 minus_involutive by auto + show ?thesis + using s4_2 by auto + qed + also have s3_2: "... = a + -.Succ[-.b]" + proof - + have s4_1: "a \ Int" + using adom by auto + have s4_2: "-.b \ Nat" + using s2_1_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 addint_succ_negnat by auto + qed + also have s3_3: "... = a + iPred[b]" + proof - + have s4_1: "b \ Int" + using s2_1_bdom neg_nat_in_int by auto + have s4_2: "b \ Nat \ {0}" + proof - + have s5_1: "b \ negNat" + using s2_1_bdom by auto + show ?thesis + using s5_1 neg_nat_not_in_nat_setminus_0 + by auto + qed + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + finally have "iPred[a + b] = a + iPred[b]" . + } + from this have s2_2: "iPred[a + b] = a + iPred[b]" + by auto + } + from this show ?thesis + by auto + qed + have s1_4: "\ b \ Int: + iPred[a + b] = a + iPred[b]" + using s1_2 s1_3 int_union_nat_negnat_0 spec allI by auto + } + from this show ?thesis + by auto + qed + + +theorem iPredRightDistributesAdd_sequent: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "iPred[m + n] = m + iPred[n]" + using mdom ndom iPredRightDistributesAdd by auto + + +theorem isucc_eq_add_1: + assumes idom: "i \\in Int" + shows "iSucc[i] = i + 1" + proof - + have s1_1: "iSucc[i] = iSucc[i + 0]" + using idom add_0 by auto + also have s1_2: "... = i + iSucc[0]" + proof - + have s2_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + show ?thesis + using idom s2_1 iSuccRightDistributesAdd_sequent + by auto + qed + also have s1_3: "... = i + 1" + proof - + have s2_1: "iSucc[0] = Succ[0]" + proof - + have s3_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s3_2: "0 \\in Nat" + using zeroIsNat by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + have s2_2: "Succ[0] = 1" + by auto + have s2_3: "iSucc[0] = 1" + using s2_1 s2_2 by auto + show ?thesis + using s2_3 s1_2 by auto + qed + finally show ?thesis . + qed + + +(* +THEOREM AddAssociative == + \A a \in Int: \A b \in Int: \A c \in Int: + (a + b) + c = a + (b + c) +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE \A c \in Int: + (a + b) + c = a + (b + c) + BY <1>1 +<1>2. (a + b) + 0 = a + (b + 0) + <2>1. (a + b) + 0 = a + b + <3>1. a + b \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, add_0 + <2>2. @ = a + (b + 0) + <3>1. b \in Int + BY <1>1 + <3>2. b + 0 = b + BY <3>1, add_0 + <3> QED + BY <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW c \in Nat, + (a + b) + c = a + (b + c) + PROVE (a + b) + Succ[c] = a + (b + Succ[c]) + <2>1. (a + b) + Succ[c] = iSucc[(a + b) + c] + <3>1. a + b \in Int + BY <1>1, addIsInt + <3>2. c \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>2. @ = iSucc[a + (b + c)] + BY <1>3 \* induction hypothesis + <2>3. @ = a + iSucc[b + c] + <3>1. a \in Int + BY <1>1 + <3>2. b + c \in Int + BY <1>1, <1>3, nat_in_int, addIsInt + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>4. @ = a + (b + iSucc[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>3, nat_in_int + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>5. @ = a + (b + Succ[c]) + <3>1. c \in Nat + BY <1>3 + <3>2. c \in Int + BY <3>1, nat_in_int + <3>3. iSucc[c] = Succ[c] + BY <3>1, <3>2 DEF iSucc + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5 +<1>4. ASSUME NEW c \in negNat, + (a + b) + c = a + (b + c) + PROVE (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c]) + <2>1. (a + b) + -.Succ[-.c] = (a + b) + iPred[c] + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. c \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <3>3. iPred[c] = -.Succ[-.c] + BY <3>1, <3>2 DEF iPred + <3> QED + BY <3>3 + <2>2. @ = iPred[(a + b) + c] + <3>1. a + b \in Int + BY <1>1, addIsInt + <3>2. c \in Int + BY <1>4, neg_nat_in_int + <3>3 QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>3. @ = iPred[a + (b + c)] + BY <1>4 \* induction hypothesis + <2>4. @ = a + iPred[b + c] + <3>1. a \in Int + BY <1>1 + <3>2. b + c \in Int + BY <1>1, <1>4, neg_nat_in_int, addIsInt + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>5. @ = a + (b + iPred[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>4, neg_nat_in_int + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>6. @ = a + (b + -.Succ[-.c]) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. c \notin Nat \ {0} + BY <1>4, neg_nat_not_in_nat_setminus_0 + <3>3. iPred[c] = -.Succ[-.c] + BY <3>1, <3>2 DEF iPred + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 +<1>5. \A c \in Nat: (a + b) + c = a + (b + c) + BY <1>2, <1>3, NatInduction +<1>6. \A c \in negNat: (a + b) + c = a + (b + c) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6 +*) +theorem AddAssociative: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + (a + b) + c = a + (b + c)" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \ Int" + assume bdom: "b \ Int" + have s1_2: "(a + b) + 0 = a + (b + 0)" + proof - + have s2_1: "(a + b) + 0 = a + b" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + show ?thesis + using s3_1 add_0 by auto + qed + have s2_2: "... = a + (b + 0)" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "b + 0 = b" + using s3_1 add_0 by auto + show ?thesis + using s3_2 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ c. \ + c \ Nat; + (a + b) + c = a + (b + c) + \ \ + (a + b) + Succ[c] = a + (b + Succ[c])" + proof - + fix "c" :: "c" + assume s1_3_cdom: "c \ Nat" + assume s1_3_induct: "(a + b) + c = a + (b + c)" + have s2_1: "(a + b) + Succ[c] = iSucc[(a + b) + c]" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + have s3_2: "c \ Nat" + using s1_3_cdom by auto + show ?thesis + using s3_1 s3_2 addint_succ_nat by auto + qed + also have s2_2: "... = iSucc[a + (b + c)]" + using s1_3_induct by auto + also have s2_3: "... = a + iSucc[b + c]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b + c \ Int" + using bdom s1_3_cdom nat_in_int addIsInt by auto + show ?thesis + using s3_1 s3_2 iSuccRightDistributesAdd by auto + qed + also have s2_4: "... = a + (b + iSucc[c])" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "c \ Int" + using s1_3_cdom nat_in_int by auto + have s3_3: "iSucc[b + c] = b + iSucc[c]" + using s3_1 s3_2 iSuccRightDistributesAdd by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = a + (b + Succ[c])" + proof - + have s3_1: "c \ Nat" + using s1_3_cdom by auto + have s3_2: "c \ Int" + using s1_3_cdom nat_in_int by auto + have s3_3: "iSucc[c] = Succ[c]" + unfolding iSucc_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a + b) + Succ[c] = a + (b + Succ[c])" . + qed + have s1_4: "\ c. \ + c \ negNat; + (a + b) + c = a + (b + c) + \ \ + (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" + proof - + fix "c" :: "c" + assume s1_4_cdom: "c \ negNat" + assume s1_4_induct: "(a + b) + c = a + (b + c)" + have s2_1: "(a + b) + -.Succ[-.c] = (a + b) + iPred[c]" + proof - + have s3_1: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "c \ Nat \ {0}" + using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto + have s3_3: "iPred[c] = -.Succ[-.c]" + unfolding iPred_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + also have s2_2: "... = iPred[(a + b) + c]" + proof - + have s3_1: "a + b \ Int" + using adom bdom addIsInt by auto + have s3_2: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by force + qed + also have s2_3: "... = iPred[a + (b + c)]" + using s1_4_induct by auto + also have s2_4: "... = a + iPred[b + c]" + proof - + have s3_1: "a \ Int" + using adom by auto + have s3_2: "b + c \ Int" + using bdom s1_4_cdom neg_nat_in_int addIsInt by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by auto + qed + also have s2_5: "... = a + (b + iPred[c])" + proof - + have s3_1: "b \ Int" + using bdom by auto + have s3_2: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_3: "iPred[b + c] = b + iPred[c]" + using s3_1 s3_2 iPredRightDistributesAdd by force + show ?thesis + using s3_3 by auto + qed + also have s2_6: "... = a + (b + -.Succ[-.c])" + proof - + have s3_1: "c \ Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "c \ Nat \ {0}" + using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto + have s3_3: "iPred[c] = -.Succ[-.c]" + unfolding iPred_def + using s3_1 s3_2 by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" . + qed + have s1_5: "\ c \ Nat: + (a + b) + c = a + (b + c)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ c \ negNat: + (a + b) + c = a + (b + c)" + using s1_2 s1_4 neg_nat_induction by auto + have "\ c \ Int: + (a + b) + c = a + (b + c)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem AddAssociative_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + and cdom: "c \\in Int" + shows "(a + b) + c = a + (b + c)" + using adom bdom cdom AddAssociative by auto + + +theorem AddLeftCommutativity: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + and cdom: "c \\in Int" + shows "b + (a + c) = a + (b + c)" + proof - + have s1_1: "(b + a) + c = (a + b) + c" + using adom bdom by (simp only: AddCommutative_sequent) + show ?thesis + using adom bdom cdom s1_1 by (simp only: AddAssociative_sequent) + qed + + +theorems add_ac_int[algebra_simps] = + AddCommutative_sequent AddAssociative_sequent + AddLeftCommutativity add_0 add_0_left + + +(* A test that the simplification rules work. *) +theorem + assumes "a \\in Int" and "b \\in Int" and "c \\in Int" + shows "a + (b + c) = (a + b) + c" + using assms by (simp add: algebra_simps) + + +(* +THEOREM MinusSuccIsMinusOne == + ASSUME NEW n \in Nat + PROVE -.Succ[n] = -.n + -.1 +<1>1. -.Succ[n] = -.Succ[-.-.n] + <2>1. -.-.n = n + <3>1. n \in Nat + OBVIOUS + <3>2. n \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>2, minus_involutive + <2> QED + BY <2>1 +<1>2. @ = iPred[-.n] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. -.n \notin Nat \ {0} + <3>1. n \in Nat + OBVIOUS + <3>2. -.n \in negNat + BY <3>1, minus_nat_in_neg_nat + <3> QED + BY <3>2, neg_nat_not_in_nat_setminus_0 + <2> QED + BY <2>1, <2>2 DEF iPred +<1>3. @ = iPred[-.n + 0] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. -.n + 0 = -.n + BY <2>1, add_0 + <2> QED + BY <2>2 +<1>4. @ = -.n + iPred[0] + <2>1. -.n \in Int + <3>1. n \in Nat + OBVIOUS + <3> QED + BY <3>1, minus_nat_in_int + <2>2. 0 \in Int + BY zero_in_int + <2> QED + BY <2>1, <2>2, iPredRightDistributesAdd +<1>5. @ = -.n + -.Succ[-.0] + <2>1. 0 \in Int + BY zero_in_int + <2>2. 0 \notin Nat \ {0} + OBVIOUS + <2> QED + BY <2>1, <2>2 DEF iPred +<1>6. @ = -.n + -.Succ[0] + BY neg0 +<1>7. @ = -.n + -.1 + OBVIOUS \* 1 is an abbreviation +<1> QED + BY <1>1, <1>2, <1>3, <1>4, + <1>5, <1>6, <1>7 +*) +theorem MinusSuccIsMinusOne: + assumes ndom: "n \ Nat" + shows "-.Succ[n] = -.n + -.1" + proof - + have s1_1: "-.Succ[n] = -.Succ[-.-.n]" + proof - + have s2_1: "-.-.n = n" + proof - + have s3_1: "n \ Int" + using ndom nat_in_int by auto + show ?thesis + using s3_1 minus_involutive by auto + qed + show ?thesis + using s2_1 by auto + qed + also have s1_2: "... = iPred[-.n]" + proof - + have s2_1: "-.n \ Int" + using ndom minus_nat_in_int by auto + have s2_2: "-.n \ Nat \ {0}" + proof - + have s3_1: "-.n \ negNat" + using ndom minus_nat_in_neg_nat by auto + show ?thesis + using s3_1 neg_nat_not_in_nat_setminus_0 by fast + qed + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_3: "... = iPred[-.n + 0]" + proof - + have s2_1: "-.n \ Int" + proof - + have s3_1: "n \ Nat" + using ndom by auto + show ?thesis + using s3_1 minus_nat_in_int by auto + qed + have s2_2: "-.n + 0 = -.n" + using s2_1 add_0 by auto + show ?thesis + using s2_2 by auto + qed + also have s1_4: "... = -.n + iPred[0]" + proof - + have s2_1: "-.n \ Int" + proof - + have s3_1: "n \ Nat" + using ndom by auto + show ?thesis + using s3_1 minus_nat_in_int by auto + qed + have s2_2: "0 \ Int" + using zero_in_int by auto + show ?thesis + using s2_1 s2_2 iPredRightDistributesAdd by fast + qed + also have s1_5: "... = -.n + -.Succ[-.0]" + proof - + have s2_1: "0 \ Int" + using zero_in_int by auto + have s2_2: "0 \ Nat \ {0}" + by auto + show ?thesis + unfolding iPred_def + using s2_1 s2_2 by auto + qed + also have s1_6: "... = -.n + -.Succ[0]" + using neg0 by auto + also have s1_7: "... = -.n + -.1" + by auto + from calculation show ?thesis + by auto + qed + + +theorem SuccMinusIsPlusOne: + assumes idom: "i \\in negNat" + shows "Succ[-.i] = -.i + 1" + proof - + have s1_1: "-.i \\in Nat" + using idom minus_neg_nat_in_nat by auto + have s1_2: "-.i \\in Int" + using s1_1 nat_in_int by auto + have s1_3: "Succ[-.i] = iSucc[-.i]" + unfolding iSucc_def + using s1_1 s1_2 by auto + also have s1_4: "... = iSucc[-.i + 0]" + using s1_2 add_0 by auto + also have s1_5: "... = -.i + Succ[0]" + using s1_2 zeroIsNat addint_succ_nat by auto + also have s1_6: "... = -.i + 1" + by auto (* 1 is an abbreviation *) + from calculation show ?thesis . + qed + + +(*----------------------------------------------------------------------------*) +(* Properties of addition and difference. *) + + +(* +THEOREM AddNegCancels == + \A i \in Int: i + -.i = 0 +PROOF +<1>1. 0 + -.0 = 0 + <2>1. 0 + -.0 = 0 + 0 + BY neg0 + <2>2. @ = 0 + BY zero_in_int, add_0 + <2> QED + BY <2>1, <2>2 +<1>2. ASSUME NEW i \in Nat, i + -.i = 0 + PROVE Succ[i] + -.Succ[i] = 0 + <2>1. Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i] + <3>1. Succ[i] \in Int + BY <1>2, succIsNat, nat_in_int + <3>2. -.Succ[i] \in Int + BY <3>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, AddCommutative + <2>2. @ = iSucc[-.Succ[i] + i] + <3>1. -.Succ[i] \in Int + BY <1>2, succIsNat, minus_nat_in_int + <3>2. i \in Nat + BY <1>2 + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>3. @ = iSucc[i + -.Succ[i]] + <3>1. i \in Int + BY <1>2, nat_in_int + <3>2. -.Succ[i] \in Int + BY <1>2, succIsNat, minus_nat_in_int + <3>3. -.Succ[i] + i = i + -.Succ[i] + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>4. @ = iSucc[iPred[i + -.i]] + <3>1. i \in Int + BY <1>2, nat_in_int + <3>2. i \in Nat + BY <1>2 + <3> QED + BY <3>1, <3>2, addint_succ_negnat + <2>5. @ = iSucc[iPred[0]] + BY <1>2 + <2>6. @ = iSucc[-.1] + BY ipred_0 + <2>7. @ = 0 + BY isucc_minus_1 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6, <2>7 +<1>3. ASSUME NEW i \in negNat, i + -.i = 0 + PROVE -.Succ[-.i] + -.(-.Succ[-.i]) = 0 + <2>1. -.Succ[-.i] + -.(-.Succ[-.i]) + = -.Succ[-.i] + Succ[-.i] + <3>1. Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, nat_in_int + <3> QED + BY <3>1, minus_involutive + <2>2. @ = iSucc[-.Succ[-.i] + -.i] + <3>1. -.Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, minus_nat_in_int + <3>2. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, addint_succ_nat + <2>3. @ = iSucc[-.i + -.Succ[-.i]] + <3>1. -.Succ[-.i] \in Int + BY <1>3, minus_neg_nat_in_nat, + succIsNat, minus_nat_in_int + <3>2. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>3. -.Succ[-.i] + -.i = -.i + -.Succ[-.i] + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>4. @ = iSucc[iPred[-.i + -.-.i]] + <3>1. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>2. -.i \in Nat + BY <1>3, minus_neg_nat_in_nat + <3>3. -.i + -.Succ[-.i] = iPred[-.i + -.-.i] + BY <3>1, <3>2, addint_succ_negnat + <3> QED + BY <3>3 + <2>5. @ = iSucc[iPred[-.i + i]] + <3>1. -.-.i = i + BY <1>3, neg_nat_in_int, minus_involutive + <3> QED + BY <3>1 + <2>6. @ = iSucc[iPred[i + -.i]] + <3>1. -.i \in Int + BY <1>3, minus_negnat_in_int + <3>2. i \in Int + BY <1>3, neg_nat_in_int + <3>3. -.i + i = i + -.i + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>7. @ = iSucc[iPred[0]] + BY <1>3 \* induction hypothesis + <2>8. @ = iSucc[-.1] + BY ipred_0 + <2>9. @ = 0 + BY isucc_minus_1 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, + <2>6, <2>7, <2>8, <2>9 +<1>4. \A i \in Nat: i + -.i = 0 + BY <1>1, <1>2, NatInduction +<1>5. \A i \in negNat: i + -.i = 0 + BY <1>1, <1>3, neg_nat_induction +<1> QED + BY <1>4, <1>5, int_union_nat_negnat +*) +theorem AddNegCancels: + "\ i \ Int: i + -.i = 0" + proof - + have s1_1: "0 + -.0 = 0" + proof - + have s2_1: "0 + -.0 = 0 + 0" + using neg0 by auto + also have s2_2: "... = 0" + using zero_in_int add_0 by auto + finally show ?thesis . + qed + have s1_2: "\ i. \ + i \\in Nat; + i + -.i = 0 + \ \ + Succ[i] + -.Succ[i] = 0" + proof - + fix "i" :: "c" + assume s1_2_idom: "i \\in Nat" + assume s1_2_induct: "i + -.i = 0" + have s2_1: "Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i]" + proof - + have s3_1: "Succ[i] \\in Int" + using s1_2_idom succIsNat nat_in_int by auto + have s3_2: "-.Succ[i] \\in Int" + using s3_1 neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 AddCommutative by fast + qed + also have s2_2: "... = iSucc[-.Succ[i] + i]" + proof - + have s3_1: "-.Succ[i] \\in Int" + using s1_2_idom succIsNat minus_nat_in_int by auto + have s3_2: "i \\in Nat" + using s1_2_idom by auto + show ?thesis + using s3_1 s3_2 by (rule addint_succ_nat) + qed + also have s2_3: "... = iSucc[i + -.Succ[i]]" + proof - + have s3_1: "i \\in Int" + using s1_2_idom nat_in_int by auto + have s3_2: "-.Succ[i] \\in Int" + using s1_2_idom succIsNat minus_nat_in_int + by auto + have s3_3: "-.Succ[i] + i = i + -.Succ[i]" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = iSucc[iPred[i + -.i]]" + proof - + have s3_1: "i \\in Int" + using s1_2_idom nat_in_int by auto + have s3_2: "i \\in Nat" + using s1_2_idom by auto + show ?thesis + using s3_1 s3_2 addint_succ_negnat by auto + qed + also have s2_5: "... = iSucc[iPred[0]]" + using s1_2_induct by auto + also have s2_6: "... = iSucc[-.1]" + using ipred_0 by auto + also have s2_7: "... = 0" + using isucc_minus_1 by auto + finally show "Succ[i] + -.Succ[i] = 0" . + qed + have s1_3: "\ i. \ + i \\in negNat; + i + -.i = 0 + \ \ + -.Succ[-.i] + -.-.Succ[-.i] = 0" + proof - + fix "i" :: "c" + assume s1_3_idom: "i \\in negNat" + assume s1_3_induct: "i + -.i = 0" + have s2_1: "-.Succ[-.i] + -.(-.Succ[-.i]) + = -.Succ[-.i] + Succ[-.i]" + proof - + have s3_1: "Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat nat_in_int by auto + show ?thesis + using s3_1 minus_involutive by auto + qed + also have s2_2: "... = iSucc[-.Succ[-.i] + -.i]" + proof - + have s3_1: "-.Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat minus_nat_in_int by auto + have s3_2: "-.i \\in Nat" + using s1_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 by (rule addint_succ_nat) + qed + also have s2_3: "... = iSucc[-.i + -.Succ[-.i]]" + proof - + have s3_1: "-.Succ[-.i] \\in Int" + using s1_3_idom minus_neg_nat_in_nat + succIsNat minus_nat_in_int by auto + have s3_2: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_3: "-.Succ[-.i] + -.i = -.i + -.Succ[-.i]" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = iSucc[iPred[-.i + -.-.i]]" + proof - + have s3_1: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_2: "-.i \\in Nat" + using s1_3_idom minus_neg_nat_in_nat by auto + have s3_3: "-.i + -.Succ[-.i] = iPred[-.i + -.-.i]" + using s3_1 s3_2 addint_succ_negnat by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = iSucc[iPred[-.i + i]]" + proof - + have s3_1: "-.-.i = i" + using s1_3_idom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s3_1 by auto + qed + also have s2_6: "... = iSucc[iPred[i + -.i]]" + proof - + have s3_1: "-.i \\in Int" + using s1_3_idom minus_negnat_in_int by auto + have s3_2: "i \\in Int" + using s1_3_idom neg_nat_in_int by auto + have s3_3: "-.i + i = i + -.i" + using s3_1 s3_2 AddCommutative by fast + show ?thesis + using s3_3 by auto + qed + also have s2_7: "... = iSucc[iPred[0]]" + using s1_3_induct by auto + also have s2_8: "... = iSucc[-.1]" + using ipred_0 by auto + also have s2_9: "... = 0" + using isucc_minus_1 by auto + finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . + qed + have s1_4: "\ i \ Nat: i + -.i = 0" + using s1_1 s1_2 natInduct by auto + have s1_5: "\ i \ negNat: i + -.i = 0" + using s1_1 s1_3 neg_nat_induction by auto + show ?thesis + using s1_4 s1_5 int_union_nat_negnat by auto + qed + + +theorem AddNegCancels_sequent: + assumes idom: "i \\in Int" + shows "i + -.i = 0" + using idom AddNegCancels spec by auto + + +theorem AddNegCancels_left: + assumes idom: "i \\in Int" + shows "-.i + i = 0" + proof - + have s1_1: "i + -.i = 0" + using idom AddNegCancels_sequent by auto + have s1_2: "i + -.i = -.i + i" + proof - + have s2_1: "-.i \\in Int" + using idom neg_int_eq_int by auto + show ?thesis + using idom s2_1 AddCommutative_sequent by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM a_plus_b_minus_a_eq_b == + ASSUME NEW a \in Int, NEW b \in Int + PROVE a + (b - a) = b +PROOF +<1>1. a + (b - a) = a + (b + -.a) + BY DEF diff +<1>2. @ = a + (-.a + b) + <2>1. b \in Int + OBVIOUS + <2>2. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>3. b + -.a = -.a + b + BY <2>1, <2>2, AddCommutative + <2> QED + BY <2>3 +<1>3. @ = (a + -.a) + b + <2>1. a \in Int + OBVIOUS + <2>2. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>3. b \in Int + OBVIOUS + <2> QED + BY <2>1, <2>2, <2>3, AddAssociative +<1>4. @ = 0 + b + <2>1. a \in Int + OBVIOUS + <2>2. a + -.a = 0 + BY <2>1, AddNegCancels + <2> QED + BY <2>2 +<1>5. @ = b + BY add_0_left +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5 +*) +theorem a_plus_b_minus_a_eq_b: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "a + (b - a) = b" + proof - + have s1_1: "a + (b - a) = a + (b + -.a)" + unfolding diff_def by auto + also have s1_2: "... = a + (-.a + b)" + proof - + have s2_1: "b \\in Int" + using bdom by auto + have s2_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_3: "b + -.a = -.a + b" + using s2_1 s2_2 AddCommutative by auto + show ?thesis + using s2_3 by auto + qed + also have s1_3: "... = (a + -.a) + b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_3: "b \\in Int" + using bdom by auto + have s2_4: "a \\in Int ==> -.a \\in Int ==> b \\in Int + ==> (a + -.a) + b = a + (-.a + b)" + using AddAssociative_sequent by blast + have s2_5: "(a + -.a) + b = a + (-.a + b)" + using s2_1 s2_2 s2_3 s2_4 by auto + show ?thesis + by (rule s2_5[symmetric]) + qed + also have s1_4: "... = 0 + b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "a + -.a = 0" + using s2_1 AddNegCancels by auto + show ?thesis + using s2_2 by auto + qed + also have s1_5: "... = b" + using bdom add_0_left by auto + finally show ?thesis . + qed + + +(* +THEOREM MinusDistributesAdd == + \A a \in Int: \A b \in Int: + -.(a + b) = (-.a) + (-.b) +PROOF +<1>1. SUFFICES ASSUME NEW a \in Int + PROVE \A b \in Int: + -.(a + b) = (-.a) + (-.b) + BY <1>1 +<1>2. -.(a + 0) = (-.a) + (-.0) + <2>1. -.(a + 0) = -.a + <3>1. a + 0 = a + BY <1>1, add_0 + <3> QED + BY <3>1 + <2>2. -.a + -.0 = -.a + <3>1. -.a + -.0 = -.a + 0 + BY neg0 + <3>2. @ = -.a + <4>1. -.a \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, add_0 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW b \in Nat, + -.(a + b) = (-.a) + (-.b) + PROVE + -.(a + Succ[b]) = (-.a) + (-.Succ[b]) + <2>1. -.(a + Succ[b]) = -.iSucc[a + b] + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3>3. a + Succ[b] = iSucc[a + b] + BY <3>1, <3>2, addint_succ_nat + <3> QED + BY <3>3 + <2>2. @ = -. iSucc[-.((-.a) + (-.b))] + <3>1. a + b = -.(-.a + -.b) + <4>1. -.-.(a + b) = -.(-.a + -.b) + BY <1>3 + <4>2. -.-.(a + b) = a + b + <5>1. a + b \in Int + <6>1. a \in Int + BY <1>1 + <6>2. b \in Int + BY <1>3, nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, minus_involutive + <4> QED + BY <4>1, <4>2 + <3> QED + BY <3>1 + <2>3. @ = iPred[-.a + -.b] + <3>1. -.a + -.b \in Int + <4>1. -.a \in Int + BY <1>1, neg_int_eq_int + <4>2. -.b \in Int + BY <1>3, minus_nat_in_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, iSuccMinusFlip + <2>4. @ = -.a + iPred[-.b] + <3>1. -.a \in Int + BY <1>1, neg_int_eq_int + <3>2. -.b \in Int + BY <1>3, minus_nat_in_int + <3> QED + BY <3>1, <3>2, iPredRightDistributesAdd + <2>5. @ = -.a + -.Succ[b] + <3>1. iPred[-.b] = -.Succ[-.-.b] + <4>1. -.b \in Int + BY <1>3, minus_nat_in_int + <4>2. -.b \notin Nat \ {0} + <5>1. -.b \in negNat + BY <1>3, minus_nat_in_neg_nat + <5> QED + BY <5>1, neg_nat_not_in_nat_setminus_0 + <4> QED + BY <4>1, <4>2 DEF iPred + <3>2. @ = -.Succ[b] + <4>1. -.-.b = b + BY <1>3, nat_in_int, minus_involutive + <4> QED + BY <4>1 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5 +<1>4. ASSUME NEW b \in negNat, + -.(a + b) = (-.a) + (-.b) + PROVE + -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b]) + <2>1. -.(a + -.Succ[-.b]) = -. iPred[a + b] + <3>1. a + -.Succ[-.b] = iPred[a + -.-.b] + <4>1. a \in Int + BY <1>1 + <4>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, addint_succ_negnat + <3>2. a + -.Succ[-.b] = iPred[a + b] + <4>1. -.-.b = b + BY <1>4, neg_nat_in_int, minus_involutive + <4> QED + BY <3>1, <4>1 + <3> QED + BY <3>2 + <2>2. @ = -. iPred[-.((-.a) + (-.b))] + <3>1. a + b = -.(-.a + -.b) + <4>1. -.(a + b) = (-.a) + (-.b) + BY <1>4 + <4>2. -.-.(a + b) = -.(-.a + -.b) + BY <4>1 + <4>3. -.-.(a + b) = a + b + <5>1. a + b \in Int + <6>1. a \in Int + BY <1>1 + <6>2. b \in Int + BY <1>4, neg_nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, minus_involutive + <4> QED + BY <4>2, <4>3 + <3> QED + BY <3>1 + <2>3. @ = iSucc[-.a + -.b] + <3>1. -.a + -.b \in Int + <4>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <4>2. -.b \in Int + BY <1>4, minus_negnat_in_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, iPredMinusFlip + <2>4. @ = -.a + iSucc[-.b] + <3>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <3>2. -.b \in Int + BY <1>4, minus_negnat_in_int + <3> QED + BY <3>1, <3>2, iSuccRightDistributesAdd + <2>5. @ = -.a + Succ[-.b] + <3>1. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>2. -.b \in Int + BY <3>1, nat_in_int + <3> QED + BY <3>1, <3>2 DEF iSucc + <2>6. @ = -.a + (-.-.Succ[-.b]) + <3>1. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>2. Succ[-.b] \in Nat + BY <3>1, succIsNat + <3>3. Succ[-.b] \in Int + BY <3>2, nat_in_int + <3>4. -.-.Succ[-.b] = Succ[-.b] + BY <3>3, minus_involutive + <3> QED + BY <3>4 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 +<1>5. \A b \in Nat: -.(a + b) = (-.a) + (-.b) + BY <1>2, <1>3, NatInduction +<1>6. \A b \in negNat: -.(a + b) = (-.a) + (-.b) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MinusDistributesAdd: + "\ a \ Int: \ b \ Int: + -.(a + b) = (-.a) + (-.b)" + proof - + { + fix "a" :: "c" + assume adom: "a \\in Int" + have s1_2: "-.(a + 0) = (-.a) + (-.0)" + proof - + have s2_1: "-.(a + 0) = -.a" + proof - + have s3_1: "a + 0 = a" + using adom add_0 by auto + show ?thesis + using s3_1 by auto + qed + have s2_2: "-.a + -.0 = -.a" + proof - + have s3_1: "-.a + -.0 = -.a + 0" + using neg0 by auto + also have s3_2: "... = -.a" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + show ?thesis + using s4_1 add_0 by auto + qed + finally show ?thesis . + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ b. \ + b \\in Nat; + -.(a + b) = -.a + -.b + \ \ + -.(a + Succ[b]) = (-.a) + (-.Succ[b])" + proof - + fix "b" :: "c" + assume s1_3_bdom: "b \\in Nat" + assume s1_3_induct: "-.(a + b) = -.a + -.b" + have s2_1: "-.(a + Succ[b]) = -.iSucc[a + b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + have s3_3: "a + Succ[b] = iSucc[a + b]" + using s3_1 s3_2 addint_succ_nat by auto + show ?thesis + using s3_3 by auto + qed + also have s2_2: "... = -. iSucc[-.((-.a) + (-.b))]" + proof - + have s3_1: "a + b = -.(-.a + -.b)" + proof - + have s4_1: "-.-.(a + b) = -.(-.a + -.b)" + using s1_3_induct by auto + have s4_2: "-.-.(a + b) = a + b" + proof - + have s5_1: "a + b \\in Int" + proof - + have s6_1: "a \\in Int" + using adom by auto + have s6_2: "b \\in Int" + using s1_3_bdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 minus_involutive by auto + qed + show ?thesis + using s4_1 s4_2 by auto + qed + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = iPred[-.a + -.b]" + proof - + have s3_1: "-.a + -.b \\in Int" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s4_2: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 iSuccMinusFlip by auto + qed + also have s2_4: "... = -.a + iPred[-.b]" + proof - + have s3_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s3_2: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + show ?thesis + using s3_1 s3_2 iPredRightDistributesAdd by auto + qed + also have s2_5: "... = -.a + -.Succ[b]" + proof - + have s3_1: "iPred[-.b] = -.Succ[-.-.b]" + proof - + have s4_1: "-.b \\in Int" + using s1_3_bdom minus_nat_in_int by auto + have s4_2: "-.b \ Nat \ {0}" + proof - + have s5_1: "-.b \\in negNat" + using s1_3_bdom minus_nat_in_neg_nat by auto + show ?thesis + using s5_1 neg_nat_not_in_nat_setminus_0 by fast + qed + show ?thesis + unfolding iPred_def + using s4_1 s4_2 by auto + qed + also have s3_2: "... = -.Succ[b]" + proof - + have s4_1: "-.-.b = b" + using s1_3_bdom nat_in_int minus_involutive by auto + show ?thesis + using s4_1 by auto + qed + finally have s3_3: "iPred[-.b] = -.Succ[b]" . + show ?thesis + using s3_3 by auto + qed + finally show "-.(a + Succ[b]) = (-.a) + (-.Succ[b])" . + qed + have s1_4: "\ b. \ + b \\in negNat; + -.(a + b) = -.a + -.b + \ \ + -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" + proof - + fix "b" :: "c" + assume s1_4_bdom: "b \\in negNat" + assume s1_4_induct: "-.(a + b) = -.a + -.b" + have s2_1: "-.(a + -.Succ[-.b]) = -. iPred[a + b]" + proof - + have s3_1: "a + -.Succ[-.b] = iPred[a + -.-.b]" + proof - + have s4_1: "a \\in Int" + using adom by auto + have s4_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 addint_succ_negnat by auto + qed + have s3_2: "a + -.Succ[-.b] = iPred[a + b]" + proof - + have s4_1: "-.-.b = b" + using s1_4_bdom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s3_1 s4_1 by auto + qed + show ?thesis + using s3_2 by auto + qed + also have s2_2: "... = -. iPred[-.((-.a) + (-.b))]" + proof - + have s3_1: "a + b = -.(-.a + -.b)" + proof - + have s4_1: "-.(a + b) = (-.a) + (-.b)" + using s1_4_induct by auto + have s4_2: "-.-.(a + b) = -.(-.a + -.b)" + using s4_1 by auto + have s4_3: "-.-.(a + b) = a + b" + proof - + have s5_1: "a + b \\in Int" + proof - + have s6_1: "a \\in Int" + using adom by auto + have s6_2: "b \\in Int" + using s1_4_bdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 minus_involutive by auto + qed + show ?thesis + using s4_2 s4_3 by auto + qed + show ?thesis + using s3_1 by auto + qed + also have s2_3: "... = iSucc[-.a + -.b]" + proof - + have s3_1: "-.a + -.b \\in Int" + proof - + have s4_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s4_2: "-.b \\in Int" + using s1_4_bdom minus_negnat_in_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 iPredMinusFlip by auto + qed + also have s2_4: "... = -.a + iSucc[-.b]" + proof - + have s3_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s3_2: "-.b \\in Int" + using s1_4_bdom minus_negnat_in_int by auto + show ?thesis + using s3_1 s3_2 iSuccRightDistributesAdd by auto + qed + also have s2_5: "... = -.a + Succ[-.b]" + proof - + have s3_1: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + have s3_2: "-.b \\in Int" + using s3_1 nat_in_int by auto + show ?thesis + unfolding iSucc_def + using s3_1 s3_2 by auto + qed + also have s2_6: "... = -.a + (-.-.Succ[-.b])" + proof - + have s3_1: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + have s3_2: "Succ[-.b] \\in Nat" + using s3_1 succIsNat by auto + have s3_3: "Succ[-.b] \\in Int" + using s3_2 nat_in_int by auto + have s3_4: "-.-.Succ[-.b] = Succ[-.b]" + using s3_3 minus_involutive by auto + show ?thesis + using s3_4 by auto + qed + finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . + qed + have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem MinusDistributesAdd_sequent: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "-.(a + b) = (-.a) + (-.b)" + using adom bdom MinusDistributesAdd by auto + + +(* +THEOREM MinusDiff == + ASSUME NEW a \in Int, NEW b \in Int + PROVE -.(a - b) = b - a +PROOF +<1>1. -.(a - b) = -.(a + -.b) + BY DEF diff +<1>2. @ = (-.a) + (-.-.b) + <2>1. a \in Int + OBVIOUS + <2>2. -.b \in Int + BY b \in Int, neg_int_eq_int + <2> QED + BY <2>1, <2>2, MinusDistributesAdd +<1>3. @ = (-.a) + b + BY b \in Int, minus_involutive +<1>4. @ = b + (-.a) + <2>1. -.a \in Int + BY a \in Int, neg_int_eq_int + <2>2. b \in Int + OBVIOUS + <2> QED + BY <2>1, <2>2, AddCommutative +<1>5. @ = b - a + BY DEF diff +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5 +*) +theorem MinusDiff: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "-.(a - b) = b - a" + proof - + have s1_1: "-.(a - b) = -.(a + -.b)" + unfolding diff_def by auto + also have s1_2: "... = -.a + -.-.b" + proof - + have s2_1: "a \\in Int" + using adom by auto + have s2_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + show ?thesis + using s2_1 s2_2 MinusDistributesAdd by auto + qed + also have s1_3: "... = -.a + b" + using bdom minus_involutive by auto + also have s1_4: "... = b + -.a" + proof - + have s2_1: "-.a \\in Int" + using adom neg_int_eq_int by auto + have s2_2: "b \\in Int" + using bdom by auto + show ?thesis + using s2_1 s2_2 AddCommutative by auto + qed + also have s1_5: "... = b - a" + unfolding diff_def by auto + finally show ?thesis . + qed + + +(*----------------------------------------------------------------------------*) +(* Commutativity of multiplication. *) + +(* +THEOREM MultCommutative == + \A i, j \in Int: i * j = j * i +PROOF +<1>1. \A i \in Int: i * 0 = 0 * i + <2>1. 0 * 0 = 0 * 0 + OBVIOUS + <2>2. ASSUME NEW i \in Nat, i * 0 = 0 * i + PROVE Succ[i] * 0 = 0 * Succ[i] + <3>1. 0 * Succ[i] = 0 * i + 0 + <4>1. 0 \in Int + BY zero_in_int + <4>2. i \in Nat + BY <2>2 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3>2. @ = 0 * i + <4>1. 0 * i \in Int + <5>1 0 \in Int + BY zero_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5> QED + BY <5>1, <5>2, multIsInt + <4> QED + BY <4>1, add_0 + <3>3. @ = i * 0 + BY <2>2 + <3>4. @ = 0 + <4>1. i \in Int + BY <2>2, nat_in_int + <4> QED + BY <4>1, mult_0 + <3>5. @ = Succ[i] * 0 + <4>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5 + <2>3. ASSUME NEW i \in negNat, i * 0 = 0 * i + PROVE -.Succ[-.i] * 0 = 0 * -.Succ[-.i] + <3>1. -.Succ[-.i] * 0 = 0 + <4>1. -.Succ[-.i] \in Int + BY <2>3, minus_succ_minus_negnat_in_int + <4> QED + BY <4>1, mult_0 + <3>2. 0 * -.Succ[-.i] = 0 + <4>1. 0 * -.Succ[-.i] = 0 * -.-.i + -.0 + <5>1. 0 \in Int + BY zero_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = 0 * i + -.0 + <5>1. -.-.i = i + BY <2>3, neg_nat_in_int, minus_involutive + <5> QED + BY <5>1 + <4>3. @ = 0 * i + 0 + BY neg0 + <4>4. @ = 0 * i + <5>1. 0 * i \in Int + <6>1. 0 \in Int + BY zero_in_int + <6>2. i \in Int + BY <2>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5> QED + BY <5>1, add_0 + <4>5. @ = i * 0 + BY <2>3 \* induction hypothesis + <4>6. @ = 0 + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5> QED + BY <5>1, mult_0 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, + <4>5, <4>6 + <3> QED + BY <3>1, <3>2 + <2>4. \A i \in Nat: i * 0 = 0 * i + BY <2>1, <2>2, NatInduction + <2>5. \A i \in negNat: i * 0 = 0 * i + BY <2>1, <2>3, neg_nat_induction + <2> QED + BY <2>4, <2>5, int_union_nat_negnat +<1>2. ASSUME NEW j \in Nat, + \A i \in Int: i * j = j * i + PROVE \A i \in Int: i * Succ[j] = Succ[j] * i + <2>1. 0 * Succ[j] = Succ[j] * 0 + <3>1. 0 * Succ[j] = 0 + <4>1. 0 * Succ[j] = 0 * j + 0 + <5>1. 0 \in Int + BY zero_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = 0 * j + <5>1. 0 * j \in Int + <6>1. 0 \in Int + BY zero_in_int + <6>2. j \in Int + BY <1>2, nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5> QED + BY <5>1, add_0 + <4>3. @ = j * 0 + BY <1>2, zero_in_int, spec + <4>4. @ = 0 + <5>1. j \in Int + BY <1>2, nat_in_int + <5> QED + BY <5>1, mult_0 + <4> QED + BY <4>1, <4>2, <4>3, <4>4 + <3>2. Succ[j] * 0 = 0 + <4>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2 + <2>2. ASSUME NEW i \in Nat, + i * Succ[j] = Succ[j] * i + PROVE Succ[i] * Succ[j] = Succ[j] * Succ[i] + <3>1. Succ[i] * Succ[j] = i * j + ((i + j) + 1) + <4>1. Succ[i] * Succ[j] = Succ[i] * j + Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = j * Succ[i] + Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. Succ[i] * j = j * Succ[i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>3. @ = (j * i + j) + Succ[i] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5>3. j * Succ[i] = j * i + j + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>4. @ = (i * j + j) + Succ[i] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>5. @ = (i * j + j) + (i + 1) + <5>1. i \in Nat + BY <2>2 + <5>2. Succ[i] = i + 1 + BY <5>1, nat_add_1 + <5> QED + BY <5>2 + <4>6. @ = i * j + (j + (i + 1)) + <5>1. i * j \in Int + BY <1>2, <2>2, nat_in_int, multIsInt + <5>2. j \in Int + BY <1>2, nat_in_int + <5>3. i + 1 \in Int + BY <2>2, oneIsNat, nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + ((j + i) + 1) + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. j + (i + 1) = (j + i) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>8. @ = i * j + ((i + j) + 1) + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. j + i = i + j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8 + <3>2. Succ[j] * Succ[i] = i * j + ((i + j) + 1) + <4>1. Succ[j] * Succ[i] = Succ[j] * i + Succ[j] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = i * Succ[j] + Succ[j] + BY <2>2 \* induction hypothesis + <4>3. @ = (i * j + i) + Succ[j] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Nat + BY <1>2 + <5>3. i * Succ[j] = i * j + i + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>4. @ = (i * j + i) + (j + 1) + <5>1. j \in Nat + BY <1>2 + <5>2. Succ[j] = j + 1 + BY <5>1, nat_add_1 + <5> QED + BY <5>2 + <4>5. @ = i * j + (i + (j + 1)) + <5>1. i * j \in Int + BY <1>2, <2>2, nat_in_int, multIsInt + <5>2. i \in Int + BY <2>2, nat_in_int + <5>3. j + 1 \in Int + BY <1>2, oneIsNat, nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>6. @ = i * j + ((i + j) + 1) + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. j \in Int + BY <1>2, nat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. i + (j + 1) = (i + j) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, <4>6 + <3> QED + BY <3>1, <3>2 + <2>3. ASSUME NEW i \in negNat, + i * Succ[j] = Succ[j] * i + PROVE -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i] + <3>1. -.Succ[-.i] * Succ[j] = + i * j + ((i + -.j) + -.1) + <4>1. -.Succ[-.i] * Succ[j] = + -.Succ[-.i] * j + -.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. j \in Nat + BY <1>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = j * -.Succ[-.i] + -.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>3. @ = (j * -.-.i + -.j) + -.Succ[-.i] + <5>1. j \in Int + BY <1>2, nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>3. j * -.Succ[-.i] = j * -.-.i + -.j + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>4. @ = (j * i + -.j) + -.Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>5. @ = (i * j + -.j) + -.Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>2 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>6. @ = (i * j + -.j) + (i + -.1) + <5>1. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>2. -.Succ[-.i] = -.-.i + -.1 + BY <5>1, MinusSuccIsMinusOne + <5>3. -.-.i = i + BY <2>3, neg_nat_in_int, minus_involutive + <5>4. -.Succ[-.i] = i + -.1 + BY <5>2, <5>3 + <5> QED + BY <5>4 + <4>7. @ = i * j + (-.j + (i + -.1)) + <5>1. i * j \in Int + BY <1>2, nat_in_int, <2>3, neg_nat_in_int, + multIsInt + <5>2. -.j \in Int + BY <1>2, minus_nat_in_int + <5>3. i + -.1 \in Int + BY <2>3, minus_one_in_negnat, + neg_nat_in_int, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + ((-.j + i) + -.1) + <5>1. -.j \in Int + BY <1>2, minus_nat_in_int + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <5>4. -.j + (i + -.1) = (-.j + i) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>9. @ = i * j + ((i + -.j) + -.1) + <5>1. -.j \in Int + BY <1>2, minus_nat_in_int + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.j + i = i + -.j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, + <4>5, <4>6, <4>7, <4>8, <4>9 + <3>2. Succ[j] * -.Succ[-.i] = + i * j + ((i + -.j) + -.1) + <4>1. Succ[j] * -.Succ[-.i] = + Succ[j] * -.-.i + -.Succ[j] + <5>1. Succ[j] \in Int + BY <1>2, succIsNat, nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = Succ[j] * i + -.Succ[j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = i * Succ[j] + -.Succ[j] + BY <2>3 \* induction hypothesis + <4>4. @ = (i * j + i) + -.Succ[j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. j \in Nat + BY <1>2 + <5>3. i * Succ[j] = i * j + i + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>5. @ = (i * j + i) + (-.j + -.1) + <5>1. j \in Nat + BY <1>2 + <5>2. -.Succ[j] = -.j + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>6. @ = i * j + (i + (-.j + -.1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>2, nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. i \in Int + BY <2>3, neg_nat_in_int + <5>3. -.j + -.1 \in Int + <6>1. -.j \in Int + BY <1>2, minus_nat_in_int + <6>2. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + ((i + -.j) + -.1) + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.j \in Int + BY <1>2, minus_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_negnat, neg_nat_in_int + <5>4. i + (-.j + -.1) = (I + -.j) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, NatInduction, neg_nat_induction, + int_union_nat_negnat +<1>3. ASSUME NEW j \in negNat, + \A i \in Int: i * j = j * i + PROVE \A i \in Int: i * -.Succ[-.j] = -.Succ[-.j] * i + <2>1. 0 * -.Succ[-.j] = -.Succ[-.j] * 0 + <3>1. 0 * -.Succ[-.j] = 0 * -.-.j + -.0 + <4>1. 0 \in Int + BY zeroIsNat nat_in_int + <4>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3>2. @ = 0 * -.-.j + <4>1. 0 * -.-.j \in Int + <5>1. 0 \in Int + BY zeroIsNat nat_in_int + <5>2. -.-.j \in Int + BY <1>3, neg_nat_in_int, minus_involutive + <5> QED + BY <5>1, <5>2, multIsInt + <4>2. 0 * -.-.j + 0 = 0 * -.-.j + BY <4>1, add_0 + <4> QED + BY <4>2, neg0 + <3>3. @ = 0 * j + <4>1. j \in Int + BY <1>3, neg_nat_in_int + <4> QED + BY <4>1, minus_involutive + <3>4. @ = j * 0 + <4>1. 0 \in Int + BY zeroIsNat, nat_in_int + <4>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <4> QED + BY <4>1, <4>2, spec + <3>5. @ = 0 + <4>1. j \in Int + BY <1>3, neg_nat_in_int + <4> QED + BY <4>1, mult_0 + <3>6. @ = -.Succ[-.j] * 0 + <4>1. -.Succ[-.j] \in Int + BY <1>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <4> QED + BY <4>1, mult_0 + <3> QED + BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 + <2>2. ASSUME NEW i \in Nat, + i * -.Succ[-.j] = -.Succ[-.j] * i + PROVE Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i] + <3>1. Succ[i] * -.Succ[-.j] = i * j + ((-.i + j) + -.1) + <4>1. Succ[i] * -.Succ[-.j] = + Succ[i] * -.-.j + -.Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = Succ[i] * j + -.Succ[i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = j * Succ[i] + -.Succ[i] + <5>1. Succ[i] \in Int + BY <2>2, succIsNat, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. Succ[i] * j = j * Succ[i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>4. @ = (j * i + j) + -.Succ[i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. i \in Nat + BY <2>2 + <5>3. j * Succ[i] = j * i + j + BY <5>1, <5>2, multint_succ_nat + <5> QED + BY <5>3 + <4>5. @ = (i * j + j) + -.Succ[i] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>6. @ = i * j + (j + -.Succ[i]) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>2, nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. j \in Int + BY <1>3, neg_nat_in_int + <5>3. -.Succ[i] \in Int + BY <2>2, succIsNat, minus_nat_in_int + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>7. @ = i * j + (j + (-.i + -.1)) + <5>1. i \in Nat + BY <2>2 + <5>2. -.Succ[i] = -.i + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>8. @ = i * j + ((j + -.i) + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_int + <5>4. j + (-.i + -.1) = (j + -.i) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>9. @ = i * j + ((-.i + j) + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. j + -.i = -.i + j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9 + <3>2. -.Succ[-.j] * Succ[i] = i * j + ((-.i + j) + -.1) + <4>1. -.Succ[-.j] * Succ[i] = + -.Succ[-.j] * i + -.Succ[-.j] + <5>1. -.Succ[-.j] \in Int + BY <1>3, minus_negnat_succ_in_negnat, + neg_nat_in_int + <5>2. i \in Nat + BY <2>2 + <5> QED + BY <5>1, <5>2, multint_succ_nat + <4>2. @ = i * -.Succ[-.j] + -.Succ[-.j] + BY <2>2 \* induction hypothesis + <4>3. @ = (i * -.-.j + -.i) + -.Succ[-.j] + <5>1. i \in Int + BY <2>2, nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>3. i * -.Succ[-.j] = i * -.-.j + -.i + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>4. @ = (i * j + -.i) + -.Succ[-.j] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>5. @ = (i * j + -.i) + (-.-.j + -.1) + <5>1. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>2. -.Succ[-.j] = -.-.j + -.1 + BY <5>1, MinusSuccIsMinusOne + <5> QED + BY <5>2 + <4>6. @ = (i * j + -.i) + (j + -.1) + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = i * j + (-.i + (j + -.1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>2, nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.i \in Int + BY <2>2, minus_nat_in_int + <5>3. j + -.1 \in Int + <6>1. j \in Int + BY <1>3, neg_nat_in_int + <6>2. -.1 \in Int + BY minus_one_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + ((-.i + j) + -.1) + <5>1. -.i \in Int + BY <2>2, minus_nat_in_int + <5>2. j \in Int + BY <1>3, neg_nat_in_int + <5>3. -.1 \in Int + BY minus_one_in_int + <5>4. -.i + (j + -.1) = (-.i + j) + -.1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8 + <3> QED + BY <3>1, <3>2 + <2>3. ASSUME NEW i \in negNat, + i * -.Succ[-.j] = -.Succ[-.j] * i + PROVE -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i] + <3>1. -.Succ[-.i] * -.Succ[-.j] = + i * j + ((-.i + -.j) + 1) + <4>1. -.Succ[-.i] * -.Succ[-.j] = + -.Succ[-.i] * -.-.j + -.-.Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, + neg_nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = -.Succ[-.i] * j + -.-.Succ[-.i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = -.Succ[-.i] * j + Succ[-.i] + <5>1. Succ[-.i] \in Int + BY <2>3, negnat_succ_in_nat, nat_in_int + <5>2. -.-.Succ[-.i] = Succ[-.i] + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>4. @ = j * -.Succ[-.i] + Succ[-.i] + <5>1. -.Succ[-.i] \in Int + BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>5. @ = (j * -.-.i + -.j) + Succ[-.i] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5>3. j * -.Succ[-.i] = j * -.-.i + -.j + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>6. @ = (j * i + -.j) + Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = (i * j + -.j) + Succ[-.i] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. \A k \in Int: k * j = j * k + BY <1>3 \* induction hypothesis + <5>3. i * j = j * i + BY <5>1, <5>2, spec + <5> QED + BY <5>3 + <4>8. @ = (i * j + -.j) + (-.i + 1) + <5>1. i \in negNat + BY <2>3 + <5>2. Succ[-.i] = -.i + 1 + BY <5>1, SuccMinusIsPlusOne + <5> QED + BY <5>2 + <4>9. @ = i * j + (-.j + (-.i + 1)) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>3. -.i + 1 \in Int + <6>1. -.i \in Int + BY <2>3, minus_negnat_in_int + <6>2. 1 \in Int + BY oneIsNat, nat_in_int + <6> QED + BY <6>1, <6>2, addIsInt + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>10. @ = i * j + ((-.j + -.i) + 1) + <5>1. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. -.j + (-.i + 1) = (-.j + -.i) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4>11. @ = i * j + ((-.i + -.j) + 1) + <5>1. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. -.j + -.i = -.i + -.j + BY <5>1, <5>2, AddCommutative + <5> QED + BY <5>3 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9, + <4>10, <4>11 + <3>2. -.Succ[-.j] * -.Succ[-.i] = + i * j + ((-.i + -.j) + 1) + <4>1. -.Succ[-.j] * -.Succ[-.i] = + -.Succ[-.j] * -.-.i + -.-.Succ[-.j] + <5>1. -.Succ[-.j] \in Int + BY <1>3, minus_succ_minus_negnat_in_int + <5>2. -.i \in Nat + BY <2>3, minus_neg_nat_in_nat + <5> QED + BY <5>1, <5>2, multint_succ_negnat + <4>2. @ = -.Succ[-.j] * i + -.-.Succ[-.j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.-.i = i + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>3. @ = -.Succ[-.j] * i + Succ[-.j] + <5>1. Succ[-.j] \in Int + BY <1>3, negnat_succ_in_nat, nat_in_int + <5>2. -.-.Succ[-.j] = Succ[-.j] + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>4. @ = i * -.Succ[-.j] + Succ[-.j] + BY <2>3 \* induction hypothesis + <4>5. @ = (i * -.-.j + -.i) + Succ[-.j] + <5>1. i \in Int + BY <2>3, neg_nat_in_int + <5>2. -.j \in Nat + BY <1>3, minus_neg_nat_in_nat + <5>3. i * -.Succ[-.j] = i * -.-.j + -.i + BY <5>1, <5>2, multint_succ_negnat + <5> QED + BY <5>3 + <4>6. @ = (i * j + -.i) + Succ[-.j] + <5>1. j \in Int + BY <1>3, neg_nat_in_int + <5>2. -.-.j = j + BY <5>1, minus_involutive + <5> QED + BY <5>2 + <4>7. @ = i * j + (-.i + Succ[-.j]) + <5>1. i * j \in Int + <6>1. i \in Int + BY <2>3, neg_nat_in_int + <6>2. j \in Int + BY <1>3, neg_nat_in_int + <6> QED + BY <6>1, <6>2, multIsInt + <5>2. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>3. Succ[-.j] \in Int + BY <1>3, negnat_succ_in_nat, nat_in_int + <5> QED + BY <5>1, <5>2, <5>3, AddAssociative + <4>8. @ = i * j + (-.i + (-.j + 1)) + <5>1. Succ[-.j] = -.j + 1 + BY <1>3, SuccMinusIsPlusOne + <5> QED + BY <5>1 + <4>9. @ = i * j + ((-.i + -.j) + 1) + <5>1. -.i \in Int + BY <2>3, minus_negnat_in_int + <5>2. -.j \in Int + BY <1>3, minus_negnat_in_int + <5>3. 1 \in Int + BY oneIsNat, nat_in_int + <5>4. -.i + (-.j + 1) = (-.i + -.j) + 1 + BY <5>1, <5>2, <5>3, AddAssociative + <5> QED + BY <5>4 + <4> QED + BY <4>1, <4>2, <4>3, <4>4, <4>5, + <4>6, <4>7, <4>8, <4>9 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, NatInduction, neg_nat_induction, + int_union_nat_negnat +<1> QED + BY <1>1, <1>2, <1>3, NatInduction, neg_nat_induction, + int_union_nat_negnat +*) +theorem MultCommutative: + "\ i \ Int: \ j \ Int: + i * j = j * i" + proof - + have s1_1: "\ i \ Int: i * 0 = 0 * i" + proof - + have s2_1: "0 * 0 = 0 * 0" + by auto + have s2_2: "\ i. \ + i \\in Nat; + i * 0 = 0 * i + \ \ + Succ[i] * 0 = 0 * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * 0 = 0 * i" + have s3_1: "0 * Succ[i] = 0 * i + 0" + proof - + have s4_1: "0 \\in Int" + using zero_in_int by auto + have s4_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + also have s3_2: "... = 0 * i" + proof - + have s4_1: "0 * i \\in Int" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + show ?thesis + using s4_1 add_0 by auto + qed + also have s3_3: "... = i * 0" + using s2_2_induct by auto + also have s3_4: "... = 0" + proof - + have s4_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + also have s3_5: "... = Succ[i] * 0" + proof - + have s4_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + finally show "Succ[i] * 0 = 0 * Succ[i]" by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * 0 = 0 * i + \ \ + -.Succ[-.i] * 0 = 0 * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * 0 = 0 * i" + have s3_1: "-.Succ[-.i] * 0 = 0" + proof - + have s4_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_succ_minus_negnat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + have s3_2: "0 * -.Succ[-.i] = 0" + proof - + have s4_1: "0 * -.Succ[-.i] = 0 * -.-.i + -.0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = 0 * i + -.0" + proof - + have s5_1: "-.-.i = i" + using s2_3_idom neg_nat_in_int + minus_involutive by auto + show ?thesis + using s5_1 by auto + qed + also have s4_3: "... = 0 * i + 0" + using neg0 by auto + also have s4_4: "... = 0 * i" + proof - + have s5_1: "0 * i \\in Int" + proof - + have s6_1: "0 \\in Int" + using zero_in_int by auto + have s6_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + show ?thesis + using s5_1 add_0 by auto + qed + also have s4_5: "... = i * 0" + using s2_3_induct by auto + also have s4_6: "... = 0" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + show ?thesis + using s5_1 mult_0 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * 0 = 0 * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: i * 0 = 0 * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: i * 0 = 0 * i" + using s2_1 s2_3 neg_nat_induction by auto + show ?thesis + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_2: "\ j. \ + j \\in Nat; + \ i \ Int: i * j = j * i + \ \ + \ i \ Int: i * Succ[j] = Succ[j] * i" + proof - + fix "j" :: "c" + assume s1_2_jdom: "j \\in Nat" + assume s1_2_induct: "\ i \ Int: i * j = j * i" + have s2_1: "0 * Succ[j] = Succ[j] * 0" + proof - + have s3_1: "0 * Succ[j] = 0" + proof - + have s4_1: "0 * Succ[j] = 0 * j + 0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = 0 * j" + proof - + have s5_1: "0 * j \\in Int" + proof - + have s6_1: "0 \\in Int" + using zero_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + show ?thesis + using s5_1 add_0 by auto + qed + also have s4_3: "... = j * 0" + proof - + have s5_1: "0 \\in Int" + using zero_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + show ?thesis + using s5_1 s5_2 spec by fast + qed + also have s4_4: "... = 0" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s5_1 mult_0 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * 0 = 0" + proof - + have s4_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + have s2_2: "\ i. \ + i \\in Nat; + i * Succ[j] = Succ[j] * i + \ \ + Succ[i] * Succ[j] = Succ[j] * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * Succ[j] = Succ[j] * i" + have s3_1: "Succ[i] * Succ[j] = i * j + ((i + j) + 1)" + proof - + have s4_1: "Succ[i] * Succ[j] = Succ[i] * j + Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = j * Succ[i] + Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_2_induct by auto + have s5_3: "Succ[i] * j = j * Succ[i]" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_3: "... = (j * i + j) + Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + have s5_3: "j * Succ[i] = j * i + j" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + j) + Succ[i]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_2_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + j) + (i + 1)" + proof - + have s5_1: "i \\in Nat" + using s2_2_idom by auto + have s5_2: "Succ[i] = i + 1" + using s5_1 nat_add_1 by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = i * j + (j + (i + 1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_3: "i + 1 \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "1 \\in Nat" + using oneIsNat nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + ((j + i) + 1)" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "j + (i + 1) = (j + i) + 1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_8: "... = i * j + ((i + j) + 1)" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "j + i = i + j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * Succ[i] = i * j + ((i + j) + 1)" + proof - + have s4_1: "Succ[j] * Succ[i] = Succ[j] * i + Succ[j]" + proof - + have s5_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = i * Succ[j] + Succ[j]" + using s2_2_induct by auto + also have s4_3: "... = (i * j + i) + Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + have s5_3: "i * Succ[j] = i * j + i" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + i) + (j + 1)" + proof - + have s5_1: "j \\in Nat" + using s1_2_jdom by auto + have s5_2: "Succ[j] = j + 1" + using s5_1 nat_add_1 by fast + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = i * j + (i + (j + 1))" + proof - + have s5_1: "i * j \\in Int" + using s1_2_jdom s2_2_idom nat_in_int multIsInt by auto + have s5_2: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_3: "j + 1 \\in Int" + using s1_2_jdom oneIsNat nat_in_int addIsInt by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_6: "... = i * j + ((i + j) + 1)" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "i + (j + 1) = (i + j) + 1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "Succ[i] * Succ[j] = Succ[j] * Succ[i]" + using s3_1 s3_2 by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * Succ[j] = Succ[j] * i + \ \ + -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * Succ[j] = Succ[j] * i" + have s3_1: "-.Succ[-.i] * Succ[j] = + i * j + ((i + -.j) + -.1)" + proof - + have s4_1: "-.Succ[-.i] * Succ[j] = + -.Succ[-.i] * j + -.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = j * -.Succ[-.i] + -.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_3: "... = (j * -.-.i + -.j) + -.Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_2_jdom nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (j * i + -.j) + -.Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = (i * j + -.j) + -.Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_2_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by fast + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (i * j + -.j) + (i + -.1)" + proof - + have s5_1: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_2: "-.Succ[-.i] = -.-.i + -.1" + using s5_1 MinusSuccIsMinusOne by fast + have s5_3: "-.-.i = i" + using s2_3_idom neg_nat_in_int minus_involutive + by auto + have s5_4: "-.Succ[-.i] = i + -.1" + using s5_2 s5_3 by auto + show ?thesis + using s5_4 by auto + qed + also have s4_7: "... = i * j + (-.j + (i + -.1))" + proof - + have s5_1: "i * j \\in Int" + using s1_2_jdom nat_in_int s2_3_idom + neg_nat_in_int multIsInt by auto + have s5_2: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_3: "i + -.1 \\in Int" + using s2_3_idom minus_one_in_negnat + neg_nat_in_int addIsInt by auto + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + ((-.j + i) + -.1)" + proof - + have s5_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + have s5_4: " -.j + (i + -.1) = (-.j + i) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_9: "... = i * j + ((i + -.j) + -.1)" + proof - + have s5_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.j + i = i + -.j" + using s5_1 s5_2 AddCommutative_sequent + by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "Succ[j] * -.Succ[-.i] = + i * j + ((i + -.j) + -.1)" + proof - + have s4_1: "Succ[j] * -.Succ[-.i] = + Succ[j] * -.-.i + -.Succ[j]" + proof - + have s5_1: "Succ[j] \\in Int" + using s1_2_jdom succIsNat nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = Succ[j] * i + -.Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = i * Succ[j] + -.Succ[j]" + using s2_3_induct by auto + also have s4_4: "... = (i * j + i) + -.Succ[j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "j \\in Nat" + using s1_2_jdom by auto + have s5_3: "i * Succ[j] = i * j + i" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + i) + (-.j + -.1)" + proof - + have s5_1: "j \\in Nat" + using s1_2_jdom by auto + have s5_2: "-.Succ[j] = -.j + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = i * j + (i + (-.j + -.1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_2_jdom nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_3: "-.j + -.1 \\in Int" + proof - + have s6_1: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s6_2: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + ((i + -.j) + -.1)" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.j \\in Int" + using s1_2_jdom minus_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_negnat neg_nat_in_int by auto + have s5_4: "i + (-.j + -.1) = (i + -.j) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" + using s2_1 s2_3 neg_nat_induction by auto + show "\ i \ Int: i * Succ[j] = Succ[j] * i" + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_3: "\ j. \ + j \\in negNat; + \ i \ Int: i * j = j * i + \ \ + \ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" + proof - + fix "j" :: "c" + assume s1_3_jdom: "j \\in negNat" + assume s1_3_induct: "\ i \ Int: i * j = j * i" + have s2_1: "0 * -.Succ[-.j] = -.Succ[-.j] * 0" + proof - + have s3_1: "0 * -.Succ[-.j] = 0 * -.-.j + -.0" + proof - + have s4_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s4_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + also have s3_2: "... = 0 * -.-.j" + proof - + have s4_1: "0 * -.-.j \\in Int" + proof - + have s5_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s5_2: "-.-.j \\in Int" + using s1_3_jdom neg_nat_in_int minus_involutive + by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + have s4_2: "0 * -.-.j + 0 = 0 * -.-.j" + using s4_1 add_0 by auto + show ?thesis + using s4_2 neg0 by auto + qed + also have s3_3: "... = 0 * j" + proof - + have s4_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s4_1 minus_involutive by auto + qed + also have s3_4: "... = j * 0" + proof - + have s4_1: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s4_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + show ?thesis + using s4_1 s4_2 spec by fast + qed + also have s3_5: "... = 0" + proof - + have s4_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + also have s3_6: "... = -.Succ[-.j] * 0" + proof - + have s4_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + show ?thesis + using s4_1 mult_0 by auto + qed + finally show ?thesis . + qed + have s2_2: "\ i. \ + i \\in Nat; + i * -.Succ[-.j] = -.Succ[-.j] * i + \ \ + Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" + proof - + fix "i" :: "c" + assume s2_2_idom: "i \\in Nat" + assume s2_2_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" + have s3_1: "Succ[i] * -.Succ[-.j] = + i * j + ((-.i + j) + -.1)" + proof - + have s4_1: "Succ[i] * -.Succ[-.j] = + Succ[i] * -.-.j + -.Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = Succ[i] * j + -.Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = j * Succ[i] + -.Succ[i]" + proof - + have s5_1: "Succ[i] \\in Int" + using s2_2_idom succIsNat nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + have s5_3: "Succ[i] * j = j * Succ[i]" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (j * i + j) + -.Succ[i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + have s5_3: "j * Succ[i] = j * i + j" + using s5_1 s5_2 multint_succ_nat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (i * j + j) + -.Succ[i]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "\ k \ Int: k * j = j * k" + using s1_3_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = i * j + (j + -.Succ[i])" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_3: "-.Succ[i] \\in Int" + using s2_2_idom succIsNat minus_nat_in_int by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_7: "... = i * j + (j + (-.i + -.1))" + proof - + have s5_1: "i \\in Nat" + using s2_2_idom by auto + have s5_2: "-.Succ[i] = -.i + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_8: "... = i * j + ((j + -.i) + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_int by auto + have s5_4: "j + (-.i + -.1) = (j + -.i) + -.1" + using s5_1 s5_2 s5_3 + AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_9: "... = i * j + ((-.i + j) + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "j + -.i = -.i + j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "-.Succ[-.j] * Succ[i] = + i * j + ((-.i + j) + -.1)" + proof - + have s4_1: "-.Succ[-.j] * Succ[i] = + -.Succ[-.j] * i + -.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "i \\in Nat" + using s2_2_idom by auto + show ?thesis + using s5_1 s5_2 multint_succ_nat by auto + qed + also have s4_2: "... = i * -.Succ[-.j] + -.Succ[-.j]" + using s2_2_induct by auto + also have s4_3: "... = (i * -.-.j + -.i) + -.Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_4: "... = (i * j + -.i) + -.Succ[-.j]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_5: "... = (i * j + -.i) + (-.-.j + -.1)" + proof - + have s5_1: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_2: "-.Succ[-.j] = -.-.j + -.1" + using s5_1 MinusSuccIsMinusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_6: "... = (i * j + -.i) + (j + -.1)" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = i * j + (-.i + (j + -.1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_2_idom nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_3: "j + -.1 \\in Int" + proof - + have s6_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s6_2: "-.1 \\in Int" + using minus_one_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + ((-.i + j) + -.1)" + proof - + have s5_1: "-.i \\in Int" + using s2_2_idom minus_nat_in_int by auto + have s5_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_3: "-.1 \\in Int" + using minus_one_in_int by auto + have s5_4: "-.i + (j + -.1) = (-.i + j) + -.1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" + using s3_1 s3_2 by auto + qed + have s2_3: "\ i. \ + i \\in negNat; + i * -.Succ[-.j] = -.Succ[-.j] * i + \ \ + -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" + proof - + fix "i" :: "c" + assume s2_3_idom: "i \\in negNat" + assume s2_3_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" + have s3_1: "-.Succ[-.i] * -.Succ[-.j] = + i * j + ((-.i + -.j) + 1)" + proof - + have s4_1: "-.Succ[-.i] * -.Succ[-.j] = + -.Succ[-.i] * -.-.j + -.-.Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = -.Succ[-.i] * j + -.-.Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = -.Succ[-.i] * j + Succ[-.i]" + proof - + have s5_1: "Succ[-.i] \\in Int" + using s2_3_idom negnat_succ_in_nat nat_in_int by auto + have s5_2: "-.-.Succ[-.i] = Succ[-.i]" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_4: "... = j * -.Succ[-.i] + Succ[-.i]" + proof - + have s5_1: "-.Succ[-.i] \\in Int" + using s2_3_idom minus_negnat_succ_in_negnat + neg_nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_3_induct by auto + have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_5: "... = (j * -.-.i + -.j) + Succ[-.i]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (j * i + -.j) + Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = (i * j + -.j) + Succ[-.i]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "\ k \ Int: + k * j = j * k" + using s1_3_induct by auto + have s5_3: "i * j = j * i" + using s5_1 s5_2 spec by auto + show ?thesis + using s5_3 by auto + qed + also have s4_8: "... = (i * j + -.j) + (-.i + 1)" + proof - + have s5_1: "i \\in negNat" + using s2_3_idom by auto + have s5_2: "Succ[-.i] = -.i + 1" + using s5_1 SuccMinusIsPlusOne by fast + show ?thesis + using s5_2 by auto + qed + also have s4_9: "... = i * j + (-.j + (-.i + 1))" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_3: "-.i + 1 \\in Int" + proof - + have s6_1: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s6_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + show ?thesis + using s6_1 s6_2 addIsInt by auto + qed + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_10: "... = i * j + ((-.j + -.i) + 1)" + proof - + have s5_1: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "-.j + (-.i + 1) = (-.j + -.i) + 1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + also have s4_11: "... = i * j + ((-.i + -.j) + 1)" + proof - + have s5_1: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "-.j + -.i = -.i + -.j" + using s5_1 s5_2 AddCommutative_sequent by auto + show ?thesis + using s5_3 by auto + qed + finally show ?thesis . + qed + have s3_2: "-.Succ[-.j] * -.Succ[-.i] = + i * j + ((-.i + -.j) + 1)" + proof - + have s4_1: "-.Succ[-.j] * -.Succ[-.i] = + -.Succ[-.j] * -.-.i + -.-.Succ[-.j]" + proof - + have s5_1: "-.Succ[-.j] \\in Int" + using s1_3_jdom minus_succ_minus_negnat_in_int + by auto + have s5_2: "-.i \\in Nat" + using s2_3_idom minus_neg_nat_in_nat by auto + show ?thesis + using s5_1 s5_2 multint_succ_negnat by auto + qed + also have s4_2: "... = -.Succ[-.j] * i + -.-.Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.-.i = i" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_3: "... = -.Succ[-.j] * i + Succ[-.j]" + proof - + have s5_1: "Succ[-.j] \\in Int" + using s1_3_jdom negnat_succ_in_nat + nat_in_int by auto + have s5_2: "-.-.Succ[-.j] = Succ[-.j]" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_4: "... = i * -.Succ[-.j] + Succ[-.j]" + using s2_3_induct by auto + also have s4_5: "... = (i * -.-.j + -.i) + Succ[-.j]" + proof - + have s5_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s5_2: "-.j \\in Nat" + using s1_3_jdom minus_neg_nat_in_nat by auto + have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" + using s5_1 s5_2 multint_succ_negnat by auto + show ?thesis + using s5_3 by auto + qed + also have s4_6: "... = (i * j + -.i) + Succ[-.j]" + proof - + have s5_1: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + have s5_2: "-.-.j = j" + using s5_1 minus_involutive by auto + show ?thesis + using s5_2 by auto + qed + also have s4_7: "... = i * j + (-.i + Succ[-.j])" + proof - + have s5_1: "i * j \\in Int" + proof - + have s6_1: "i \\in Int" + using s2_3_idom neg_nat_in_int by auto + have s6_2: "j \\in Int" + using s1_3_jdom neg_nat_in_int by auto + show ?thesis + using s6_1 s6_2 multIsInt by auto + qed + have s5_2: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_3: "Succ[-.j] \\in Int" + using s1_3_jdom negnat_succ_in_nat nat_in_int by auto + show ?thesis + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + qed + also have s4_8: "... = i * j + (-.i + (-.j + 1))" + proof - + have s5_1: "Succ[-.j] = -.j + 1" + using s1_3_jdom SuccMinusIsPlusOne by fast + show ?thesis + using s5_1 by auto + qed + also have s4_9: "... = i * j + ((-.i + -.j) + 1)" + proof - + have s5_1: "-.i \\in Int" + using s2_3_idom minus_negnat_in_int by auto + have s5_2: "-.j \\in Int" + using s1_3_jdom minus_negnat_in_int by auto + have s5_3: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s5_4: "-.i + (-.j + 1) = (-.i + -.j) + 1" + using s5_1 s5_2 s5_3 AddAssociative_sequent by auto + show ?thesis + using s5_4 by auto + qed + finally show ?thesis . + qed + show "-.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" + using s3_1 s3_2 by auto + qed + have s2_4: "\ i \ Nat: + i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_1 s2_2 natInduct by auto + have s2_5: "\ i \ negNat: + i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_1 s2_3 neg_nat_induction by auto + show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" + using s2_4 s2_5 int_union_nat_negnat by auto + qed + have s1_4: "\ j \ Nat: \ i \ Int: + i * j = j * i" + using s1_1 s1_2 natInduct by auto + have s1_5: "\ j \ negNat: \ i \ Int: + i * j = j * i" + using s1_1 s1_3 neg_nat_induction by auto + show ?thesis + using s1_4 s1_5 int_union_nat_negnat by auto + qed + + +theorem MultCommutative_sequent: + assumes idom: "i \\in Int" and jdom: "j \\in Int" + shows "i * j = j * i" + using idom jdom MultCommutative spec by auto + + +(* +THEOREM mult_0_left == + ASSUME m \in Int + PROVE 0 * m = 0 + BY mult_0, MultCommutative +*) +theorem mult_0_left: + assumes ndom: "n \\in Int" + shows "0 * n = 0" + proof - + have s1_1: "n * 0 = 0" + using ndom mult_0 by auto + have s1_2: "n * 0 = 0 * n" + using ndom zero_in_int MultCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Associativity of multiplication. *) + + +(* +THEOREM MultLeftDistributesAdd == + \A a, b, c \in Int: + a * (b + c) = a * b + a * c +PROOF +<1>1. SUFFICES + ASSUME NEW b \in Int, NEW c \in Int + PROVE \A a \in Int: a * (b + c) = a * b + a * c + BY <1>1 +<1>2. 0 * (b + c) = 0 * b + 0 * c + <2>1. 0 * (b + c) = 0 + <3>1. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, mult_0_left + <2>2. 0 * b + 0 * c = 0 + <3>1. 0 * b = 0 + BY <1>1, mult_0_left + <3>2. 0 * c = 0 + BY <1>1, mult_0_left + <3> QED + BY <3>1, <3>2, zero_in_int, add_0 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW a \in Nat, + a * (b + c) = a * b + a * c + PROVE Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c + <2>1. Succ[a] * (b + c) = (b + c) * Succ[a] + <3>1. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, MultCommutative + <2>2. @ = (b + c) * a + (b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. a \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>3. @ = a * (b + c) + (b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. a \in Int + BY <1>3, nat_in_int + <3>3. (b + c) * a = a * (b + c) + BY <3>1, <3>2, MultCommutative + <3> QED + BY <3>3 + <2>4. @ = (a * b + a * c) + (b + c) + BY <1>3 \* induction hypothesis + <2>5. @ = a * b + (a * c + (b + c)) + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>6. @ = a * b + ((a * c + b) + c) + <3>1. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. c \in Int + BY <1>1 + <3>4. a * c + (b + c) = (a * c + b) + c + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>7. @ = a * b + ((b + a * c) + c) + <3>1. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. a * c + b = b + a * c + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>8. @ = a * b + (b + (a * c + c)) + <3>1. b \in Int + BY <1>1 + <3>2. a * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. c \in Int + BY <1>1 + <3>4. (b + a * c) + c = b + (a * c + c) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>9. @ = (a * b + b) + (a * c + c) + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. b \in Int + BY <1>1 + <3>3. (a * c + c) \in Int + BY <1>1, <1>3, nat_in_int, addIsInt, multIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>10. @ = (b * a + b) + (c * a + c) + <3>1. a * b = b * a + <4>1. a \in Int + BY <1>3, nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. a * c = c * a + <4>1. a \in Int + BY <1>3, nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2>11. @ = b * Succ[a] + c * Succ[a] + <3>1. b * Succ[a] = b * a + b + <4>1. b \in Int + BY <1>1 + <4>2. a \in Nat + BY <1>3 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3>2. c * Succ[a] = c * a + c + <4>1. c \in Int + BY <1>1 + <4>2. a \in Nat + BY <1>3 + <4> QED + BY <4>1, <4>2, multint_succ_nat + <3> QED + BY <3>1, <3>2 + <2>12. @ = Succ[a] * b + Succ[a] * c + <3>1. b * Succ[a] = Succ[a] * b + <4>1. b \in Int + BY <1>1 + <4>2. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. c * Succ[a] = Succ[a] * c + <4>1. c \in Int + BY <1>1 + <4>2. Succ[a] \in Int + BY <1>3, succIsNat, nat_in_int + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, + <2>5, <2>6, <2>7, <2>8, <2>9, + <2>10, <2>11, <2>12 +<1>4. ASSUME NEW a \in negNat, + a * (b + c) = a * b + a * c + PROVE -.Succ[-.a] * (b + c) = -.Succ[-.a] * b + -.Succ[-.a] * c + <2>1. -.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a] + <3>1. -.Succ[-.a] \in Int + BY <1>4, minus_succ_minus_negnat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3> QED + BY <3>1, <3>2, MultCommutative + <2>2. @ = (b + c) * -.-.a + -.(b + c) + <3>1. b + c \in Int + BY <1>1, addIsInt + <3>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>3. @ = (b + c) * a + -.(b + c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>4. @ = a * (b + c) + -.(b + c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. b + c \in Int + BY <1>1, addIsInt + <3>3. (b + c) * a = a * (b + c) + BY <3>1, <3>2, MultCommutative + <3> QED + BY <3>3 + <2>5. @ = (a * b + a * c) + -.(b + c) + BY <1>4 \* induction hypothesis + <2>6. @ = (a * b + a * c) + (-.b + -.c) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Int + BY <1>1 + <3>3. -.(b + c) = -.b + -.c + BY <3>1, <3>2, MinusDistributesAdd, spec + <3> QED + BY <3>3 + <2>7. @ = a * b + (a * c + (-.b + -.c)) + <3>1. a * b \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.b + -.c \in Int + <4>1. -.b \in Int + BY <1>1, neg_int_eq_int + <4>2. -.c \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>8. @ = a * b + ((a * c + -.b) + -.c) + <3>1. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. -.c \in Int + BY <1>1, neg_int_eq_int + <3>4. a * c + (-.b + -.c) = + (a * c + -.b) + -.c + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>9. @ = a * b + ((-.b + a * c) + -.c) + <3>1. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. a * c + -.b = -.b + a * c + BY <3>1, <3>2, AddCommutative + <3> QED + BY <3>3 + <2>10. @ = a * b + (-.b + (a * c + -.c)) + <3>1. -.b \in Int + BY <1>1, neg_int_eq_int + <3>2. a * c \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.c \in Int + BY <1>1, neg_int_eq_int + <3>4. (-.b + a * c) + -.c = + -.b + (a * c + -.c) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <3>4 + <2>11. @ = (a * b + -.b) + (a * c + -.c) + <3>1. a * b \in Int + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.b \in Int + BY <1>1, neg_int_eq_int + <3>3. a * c + -.c \in Int + <4>1. a * c \in Int + <5>1. a \in Int + BY <1>4, neg_nat_in_int + <5>2. c \in Int + BY <1>1 + <5> QED + BY <5>1, <5>2, multIsInt + <4>2. -.c \in Int + BY <1>1, neg_int_eq_int + <4> QED + BY <4>1, <4>2, addIsInt + <3> QED + BY <3>1, <3>2, <3>3, AddAssociative + <2>12. @ = (b * a + -.b) + (c * a + -.c) + <3>1. a * b = b * a + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. b \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3>2. a * c = c * a + <4>1. a \in Int + BY <1>4, neg_nat_in_int + <4>2. c \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, MultCommutative + <3> QED + BY <3>1, <3>2 + <2>13. @ = (b * -.-.a + -.b) + (c * -.-.a + -.c) + <3>1. a \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>14. @ = b * -.Succ[-.a] + c * -.Succ[-.a] + <3>1. b * -.Succ[-.a] = b * -.-.a + -.b + <4>1. b \in Int + BY <1>1 + <4>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3>2. c * -.Succ[-.a] = c * -.-.a + -.c + <4>1. c \in Int + BY <1>1 + <4>2. -.a \in Nat + BY <1>4, minus_neg_nat_in_nat + <4> QED + BY <4>1, <4>2, multint_succ_negnat + <3> QED + BY <3>1, <3>2 + <2>15. @ = -.Succ[-.a] * b + -.Succ[-.a] * c + <3>1. -.Succ[-.a] \in Int + BY <1>4, minus_succ_minus_negnat_in_int + <3>2. b * -.Succ[-.a] = -.Succ[-.a] * b + <4>1. b \in Int + BY <1>1 + <4> QED + BY <3>1, <4>1, MultCommutative + <3>3. c * -.Succ[-.a] = -.Succ[-.a] * c + <4>1. c \in Int + BY <1>1 + <4> QED + BY <3>1, <4>1, MultCommutative + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, <2>5, 2<6>, + <2>7, <2>8, <2>9, <2>10, <2>11, + <2>12, <2>13, <2>14, <2>15 +<1>5. \A a \in Nat: a * (b + c) = a * b + a * c + BY <1>2, <1>3, NatInduction +<1>6. \A a \in negNat: a * (b + c) = a * b + a * c + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MultLeftDistributesAdd_forall: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + a * (b + c) = a * b + a * c" + proof - + { + fix "b" :: "c" + fix "c" :: "c" + assume bdom: "b \\in Int" + assume cdom: "c \\in Int" + have s1_2: "0 * (b + c) = 0 * b + 0 * c" + proof - + have s2_1: "0 * (b + c) = 0" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 mult_0_left by auto + qed + have s2_2: "0 * b + 0 * c = 0" + proof - + have s3_1: "0 * b = 0" + using bdom mult_0_left by auto + have s3_2: "0 * c = 0" + using cdom mult_0_left by auto + show ?thesis + using s3_1 s3_2 zero_in_int add_0 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ a. \ + a \\in Nat; + a * (b + c) = a * b + a * c + \ \ + Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" + proof - + fix "a" :: "c" + assume s1_3_adom: "a \\in Nat" + assume s1_3_induct: "a * (b + c) = a * b + a * c" + have s2_1: "Succ[a] * (b + c) = (b + c) * Succ[a]" + proof - + have s3_1: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 MultCommutative_sequent by auto + qed + also have s2_2: "... = (b + c) * a + (b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_3: "... = a * (b + c) + (b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s3_3: "(b + c) * a = a * (b + c)" + using s3_1 s3_2 MultCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_4: "... = (a * b + a * c) + (b + c)" + using s1_3_induct by auto + also have s2_5: "... = a * b + (a * c + (b + c))" + proof - + have s3_1: "a * b \\in Int" + using bdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_3: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_6: "... = a * b + ((a * c + b) + c)" + proof - + have s3_1: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "c \\in Int" + using cdom by auto + have s3_4: "a * c + (b + c) = (a * c + b) + c" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_7: "... = a * b + ((b + a * c) + c)" + proof - + have s3_1: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "a * c + b = b + a * c" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_8: "... = a * b + (b + (a * c + c))" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "a * c \\in Int" + using cdom s1_3_adom nat_in_int multIsInt by auto + have s3_3: "c \\in Int" + using cdom by auto + have s3_4: "(b + a * c) + c = b + (a * c + c)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_9: "... = (a * b + b) + (a * c + c)" + proof - + have s3_1: "a * b \\in Int" + using bdom s1_3_adom nat_in_int multIsInt by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "a * c + c \\in Int" + using cdom s1_3_adom nat_in_int addIsInt multIsInt by auto + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_10: "... = (b * a + b) + (c * a + c)" + proof - + have s3_1: "a * b = b * a" + proof - + have s4_1: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + have s3_2: "a * c = c * a" + proof - + have s4_1: "a \\in Int" + using s1_3_adom nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_11: "... = b * Succ[a] + c * Succ[a]" + proof - + have s3_1: "b * Succ[a] = b * a + b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + have s3_2: "c * Succ[a] = c * a + c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "a \\in Nat" + using s1_3_adom by auto + show ?thesis + using s4_1 s4_2 multint_succ_nat by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_12: "... = Succ[a] * b + Succ[a] * c" + proof - + have s3_1: "b * Succ[a] = Succ[a] * b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + show ?thesis + using s4_1 s4_2 MultCommutative by auto + qed + have s3_2: "c * Succ[a] = Succ[a] * c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "Succ[a] \\in Int" + using s1_3_adom succIsNat nat_in_int by auto + show ?thesis + using s4_1 s4_2 MultCommutative by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + finally show "Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" . + qed + have s1_4: "\ a. \ + a \\in negNat; + a * (b + c) = a * b + a * c + \ \ + -.Succ[-.a] * (b + c) = + -.Succ[-.a] * b + -.Succ[-.a] * c" + proof - + fix "a" :: "c" + assume s1_4_adom: "a \\in negNat" + assume s1_4_induct: "a * (b + c) = a * b + a * c" + have s2_1: "-.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a]" + proof - + have s3_1: "-.Succ[-.a] \\in Int" + using s1_4_adom minus_succ_minus_negnat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + show ?thesis + using s3_1 s3_2 MultCommutative by auto + qed + also have s2_2: "... = (b + c) * -.-.a + -.(b + c)" + proof - + have s3_1: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_3: "... = (b + c) * a + -.(b + c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_4: "... = a * (b + c) + -.(b + c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "b + c \\in Int" + using bdom cdom addIsInt by auto + have s3_3: "(b + c) * a = a * (b + c)" + using s3_1 s3_2 MultCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = (a * b + a * c) + -.(b + c)" + using s1_4_induct by auto + also have s2_6: "... = (a * b + a * c) + (-.b + -.c)" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "c \\in Int" + using cdom by auto + have s3_3: "-.(b + c) = -.b + -.c" + using s3_1 s3_2 MinusDistributesAdd spec by auto + show ?thesis + using s3_3 by auto + qed + also have s2_7: "... = a * b + (a * c + (-.b + -.c))" + proof - + have s3_1: "a * b \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.b + -.c \\in Int" + proof - + have s4_1: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s4_2: "-.c \\in Int" + using cdom neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_8: "... = a * b + ((a * c + -.b) + -.c)" + proof - + have s3_1: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "-.c \\in Int" + using cdom neg_int_eq_int by auto + have s3_4: "a * c + (-.b + -.c) = + (a * c + -.b) + -.c" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_9: "... = a * b + ((-.b + a * c) + -.c)" + proof - + have s3_1: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "a * c + -.b = -.b + a * c" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s3_3 by auto + qed + also have s2_10: "... = a * b + (-.b + (a * c + -.c))" + proof - + have s3_1: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_2: "a * c \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.c \\in Int" + using cdom neg_int_eq_int by auto + have s3_4: "(-.b + a * c) + -.c = + -.b + (a * c + -.c)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s3_4 by auto + qed + also have s2_11: "... = (a * b + -.b) + (a * c + -.c)" + proof - + have s3_1: "a * b \\in Int" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s3_3: "a * c + -.c \\in Int" + proof - + have s4_1: "a * c \\in Int" + proof - + have s5_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s5_2: "c \\in Int" + using cdom by auto + show ?thesis + using s5_1 s5_2 multIsInt by auto + qed + have s4_2: "-.c \\in Int" + using cdom neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 addIsInt by auto + qed + show ?thesis + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + qed + also have s2_12: "... = (b * a + -.b) + (c * a + -.c)" + proof - + have s3_1: "a * b = b * a" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "b \\in Int" + using bdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + have s3_2: "a * c = c * a" + proof - + have s4_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s4_2: "c \\in Int" + using cdom by auto + show ?thesis + using s4_1 s4_2 MultCommutative_sequent by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_13: "... = (b * -.-.a + -.b) + (c * -.-.a + -.c)" + proof - + have s3_1: "a \\in Int" + using s1_4_adom neg_nat_in_int by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_14: "... = b * -.Succ[-.a] + c * -.Succ[-.a]" + proof - + have s3_1: "b * -.Succ[-.a] = b * -.-.a + -.b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + have s4_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + have s3_2: "c * -.Succ[-.a] = c * -.-.a + -.c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + have s4_2: "-.a \\in Nat" + using s1_4_adom minus_neg_nat_in_nat by auto + show ?thesis + using s4_1 s4_2 multint_succ_negnat by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_15: "... = -.Succ[-.a] * b + -.Succ[-.a] * c" + proof - + have s3_1: "-.Succ[-.a] \\in Int" + using s1_4_adom minus_succ_minus_negnat_in_int by auto + have s3_2: "b * -.Succ[-.a] = -.Succ[-.a] * b" + proof - + have s4_1: "b \\in Int" + using bdom by auto + show ?thesis + using s3_1 s4_1 MultCommutative_sequent by auto + qed + have s3_3: "c * -.Succ[-.a] = -.Succ[-.a] * c" + proof - + have s4_1: "c \\in Int" + using cdom by auto + show ?thesis + using s3_1 s4_1 MultCommutative_sequent by auto + qed + show ?thesis + using s3_2 s3_3 by auto + qed + finally show "-.Succ[-.a] * (b + c) = + -.Succ[-.a] * b + -.Succ[-.a] * c" . + qed + have s1_5: "\ a \ Nat: + a * (b + c) = a * b + a * c" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ a \ negNat: + a * (b + c) = a * b + a * c" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ a \ Int: + a * (b + c) = a * b + a * c" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +theorem MultLeftDistributesAdd: + assumes adom: "a \\in Int" and bdom: "b \\in Int" and + cdom: "c \\in Int" + shows "a * (b + c) = a * b + a * c" + using adom bdom cdom MultLeftDistributesAdd_forall by auto + + +theorem MultRightDistributesAdd: + assumes adom: "a \\in Int" and bdom: "b \\in Int" and + cdom: "c \\in Int" + shows "(b + c) * a = b * a + c * a" + proof - + have s1_1: "a * (b + c) = a * b + a * c" + using adom bdom cdom MultLeftDistributesAdd by auto + have s1_2: "a * (b + c) = (b + c) * a" + using adom bdom cdom addIsInt MultCommutative by auto + have s1_3: "a * b = b * a" + using adom bdom MultCommutative by auto + have s1_4: "a * c = c * a" + using adom cdom MultCommutative by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorems mult_distributes[algebra_simps] = + MultLeftDistributesAdd MultRightDistributesAdd + + +(* +THEOREM MinusCommutesRightMult == + \A a, b \in Int: -.(a * b) = a * -.b +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int + PROVE \A b \in Int: -.(a * b) = a * -.b + BY <1>1 +<1>2. -.(a * 0) = a * -.0 + <2>1. -.(a * 0) = 0 + BY <1>1, mult_0, neg0 + <2>2. a * -.0 = 0 + BY <1>1, neg0, mult_0 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW b \in Nat, + -.(a * b) = a * -.b + PROVE -.(a * Succ[b]) = a * -.Succ[b] + <2>1. -.(a * Succ[b]) = -.(a * b + a) + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>2. @ = -.(a * b) + -.a + <3>1. a * b \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>2. a \in Int + BY <1>1 + <3> QED + BY <3>1, <3>2, MinusDistributesAdd + <2>3. @ = a * -.b + -.a + BY <1>3 \* induction hypothesis + <2>4. @ = a * -.Succ[b] + <3>1. a \in Int + BY <1>1 + <3>2. b \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>4. ASSUME NEW b \in negNat, + -.(a * b) = a * -.b + PROVE -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b] + <2>1. -.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a) + <3>1. a \in Int + BY <1>1 + <3>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>2. @ = -.(a * -.-.b) + -.-.a + <3>1. a * -.-.b \in Int + <4>1. a \in Int + BY <1>1 + <4>2. -.-.b \in Int + BY <1>4, minus_negnat_in_int, + neg_int_eq_int + <4> QED + BY <4>1, <4>2, multIsInt + <3>2. -.a \in Int + BY <1>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, MinusDistributesAdd + <2>3. @ = -.(a * -.-.b) + a + <3>1. a \in Int + BY <1>1 + <3>2. -.-.a = a + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>4. @ = -.(a * b) + a + <3>1. b \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.b = b + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>5. @ = a * -.b + a + BY <1>4 \* induction hypothesis + <2>6. @ = a * Succ[-.b] + <3>1. a \in Int + BY <1>1 + <3>2. -.b \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>7. @ = a * -.-.Succ[-.b] + <3>1. Succ[-.b] \in Int + BY <1>4, negnat_succ_in_nat, nat_in_int + <3>2. -.-.Succ[-.b] = Succ[-.b] + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2> QED + BY <2>1, <2>2, <2>3, ,2>4, + <2>5, <2>6, <2>7 +<1>5. \A b \in Nat: -.(a * b) = a * -.b + BY <1>2, <1>3, NatInduction +<1>6. \A b \in negNat: -.(a * b) = a * -.b + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MinusCommutesRightMult: + "\ a \ Int: \ b \ Int: + -.(a * b) = a * -.b" + proof - + { + fix "a" :: "c" + assume adom: "a \\in Int" + have s1_2: "-.(a * 0) = a * -.0" + proof - + have s2_1: "-.(a * 0) = 0" + using adom mult_0 neg0 by auto + have s2_2: "a * -.0 = 0" + using adom neg0 mult_0 by auto + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ b. \ + b \\in Nat; + -.(a * b) = a * -.b + \ \ + -.(a * Succ[b]) = a * -.Succ[b]" + proof - + fix "b" :: "c" + assume s1_3_bdom: "b \\in Nat" + assume s1_3_induct: "-.(a * b) = a * -.b" + have s2_1: "-.(a * Succ[b]) = -.(a * b + a)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_2: "... = -.(a * b) + -.a" + proof - + have s3_1: "a * b \\in Int" + using adom s1_3_bdom nat_in_int multIsInt by auto + have s3_2: "a \\in Int" + using adom by auto + show ?thesis + using s3_1 s3_2 MinusDistributesAdd by auto + qed + also have s2_3: "... = a * -.b + -.a" + using s1_3_induct by auto + also have s2_4: "... = a * -.Succ[b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Nat" + using s1_3_bdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + finally show "-.(a * Succ[b]) = a * -.Succ[b]" . + qed + have s1_4: "\ b. \ + b \\in negNat; + -.(a * b) = a * -.b + \ \ + -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" + proof - + fix "b" :: "c" + assume s1_4_bdom: "b \\in negNat" + assume s1_4_induct: "-.(a * b) = a * -.b" + have s2_1: "-.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_2: "... = -.(a * -.-.b) + -.-.a" + proof - + have s3_1: "a * -.-.b \\in Int" + proof - + have s4_1: "a \\in Int" + using adom by auto + have s4_2: "-.-.b \\in Int" + using s1_4_bdom minus_negnat_in_int + neg_int_eq_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_2: "-.a \\in Int" + using adom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 MinusDistributesAdd by auto + qed + also have s2_3: "... = -.(a * -.-.b) + a" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.-.a = a" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_4: "... = -.(a * b) + a" + proof - + have s3_1: "b \\in Int" + using s1_4_bdom neg_nat_in_int by auto + have s3_2: "-.-.b = b" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_5: "... = a * -.b + a" + using s1_4_induct by auto + also have s2_6: "... = a * Succ[-.b]" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "-.b \\in Nat" + using s1_4_bdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_7: "... = a * -.-.Succ[-.b]" + proof - + have s3_1: "Succ[-.b] \\in Int" + using s1_4_bdom negnat_succ_in_nat nat_in_int by auto + have s3_2: "-.-.Succ[-.b] = Succ[-.b]" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + finally show "-.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" . + qed + have s1_5: "\ b \ Nat: + -.(a * b) = a * -.b" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ b \ negNat: + -.(a * b) = a * -.b" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ b \ Int: + -.(a * b) = a * -.b" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + +(* +THEOREM MinusCommutesLeftMult == + \A a, b \in Int: -.(a * b) = -.a * b +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE -.(a * b) = -.a * b + BY <1>1 +<1>2. -.(a * b) = -.(b * a) + BY <1>1, MultCommutative +<1>3. @ = b * -.a + BY <1>1, MinusCommutesRightMult +<1>4. @ = -.a * b + BY <1>1, neg_int_eq_int, MultCommutative +<1> QED + BY <1>2, <1>3, <1>4 +*) +theorem MinusCommutesLeftMult: + "\ a \ Int: \ b \ Int: + -.(a * b) = -.a * b" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \\in Int" + assume bdom: "b \\in Int" + have s1_2: "-.(a * b) = -.(b * a)" + using adom bdom MultCommutative_sequent by auto + also have s1_3: "... = b * -.a" + using adom bdom MinusCommutesRightMult by auto + also have s1_4: "... = -.a * b" + using adom bdom neg_int_eq_int + MultCommutative_sequent by auto + finally have "-.(a * b) = -.a * b" + by auto + } + from this show ?thesis + by auto + qed + + +(* +THEOREM MultAssociative == + \A a \in Int: \A b \in Int: \A c \in Int: + (a * b) * c = a * (b * c) +PROOF +<1>1. SUFFICES + ASSUME NEW a \in Int, NEW b \in Int + PROVE \A c \in Int: (a * b) * c = a * (b * c) + BY <1>1, allI +<1>2. (a * b) * 0 = a * (b * 0) + <2>1. (a * b) * 0 = 0 + <3>1. a * b \in Int + BY <1>1, multIsInt + <3> QED + BY <3>1, mult_0 + <2>2. a * (b * 0) = 0 + <3>1. b * 0 = 0 + BY <1>1, mult_0 + <3>2. a * 0 = 0 + BY <1>1, mult_0 + <3> QED + BY <3>1, <3>2 + <2> QED + BY <2>1, <2>2 +<1>3. ASSUME NEW c \in Nat, + (a * b) * c = a * (b * c) + PROVE (a * b) * Succ[c] = a * (b * Succ[c]) + <2>1. (a * b) * Succ[c] = (a * b) * c + (a * b) + <3>1. a * b \in Int + BY <1>1, multIsInt + <3>2. c \in Nat + BY <1>3 + <3> QED + BY <3>1, <3>2, multint_succ_nat + <2>2. @ = a * (b * c) + (a * b) + BY <1>3 \* induction hypothesis + <2>3. @ = a * (b * c + b) + <3>1. a \in Int + BY <1>1 + <3>2. b * c \in Int + BY <1>1, <1>3, nat_in_int, multIsInt + <3>3. b \in Int + BY <1>1 + <3> QED + BY <3>1, <3>2, <3>3, MultLeftDistributesAdd + <2>4. @ = a * (b * Succ[c]) + <3>1. b \in Int + BY <1>1 + <3>2. c \in Nat + BY <1>3 + <3>3. b * Succ[c] = b * c + b + BY <3>1, <3>2, multint_succ_nat + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4 +<1>4. ASSUME NEW c \in negNat, + (a * b) * c = a * (b * c) + PROVE (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c]) + <2>1. (a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b) + <3>1. a * b \in Int + BY <1>1, multIsInt + <3>2. -.c \in Nat + BY <1>4, minus_neg_nat_in_nat + <3> QED + BY <3>1, <3>2, multint_succ_negnat + <2>2. @ = (a * b) * c + -.(a * b) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.c = c + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>3. @ = a * (b * c) + -.(a * b) + BY <1>4 \* induction hypothesis + <2>4. @ = a * (b * c) + a * -.b + <3>1. a \in Int + BY <1>1 + <3>2. b \in Int + BY <1>1 + <3>3. -.(a * b) = a * -.b + BY <3>1, <3>2, MinusCommutesRightMult + <3> QED + BY <3>3 + <2>5. @ = a * (b * c + -.b) + <3>1. a \in Int + BY <1>1 + <3>2. b * c \in Int + BY <1>1 <1>4, neg_nat_in_int, multIsInt + <3>3. -.b \in Int + BY <1>1, neg_int_eq_int + <3> QED + BY <3>1, <3>2, <3>3, MultLeftDistributesAdd + <2>6. @ = a * (b * -.-.c + -.b) + <3>1. c \in Int + BY <1>4, neg_nat_in_int + <3>2. -.-.c = c + BY <3>1, minus_involutive + <3> QED + BY <3>2 + <2>7. @ = a * (b * -.Succ[-.c]) + <3>1. b \in Int + BY <1>1 + <3>2. -.c \in Nat + BY <1>4, minus_neg_nat_in_nat + <3>3. b * -.Succ[-.c] = b * -.-.c + -.b + BY <3>1, <3>2, multint_succ_negnat + <3> QED + BY <3>3 + <2> QED + BY <2>1, <2>2, <2>3, <2>4, + <2>5, <2>6, <2>7 +<1>5. \A c \in Nat: (a * b) * c = a * (b * c) + BY <1>2, <1>3, NatInduction +<1>6. \A c \in negNat: (a * b) * c = a * (b * c) + BY <1>2, <1>4, neg_nat_induction +<1> QED + BY <1>5, <1>6, int_union_nat_negnat +*) +theorem MultAssociative: + "\ a \ Int: + \ b \ Int: + \ c \ Int: + (a * b) * c = a * (b * c)" + proof - + { + fix "a" :: "c" + fix "b" :: "c" + assume adom: "a \\in Int" + assume bdom: "b \\in Int" + have s1_2: "(a * b) * 0 = a * (b * 0)" + proof - + have s2_1: "(a * b) * 0 = 0" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + show ?thesis + using s3_1 mult_0 by auto + qed + have s2_2: "a * (b * 0) = 0" + proof - + have s3_1: "b * 0 = 0" + using bdom mult_0 by auto + have s3_2: "a * 0 = 0" + using adom mult_0 by auto + show ?thesis + using s3_1 s3_2 by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_3: "\ c. \ + c \\in Nat; + (a * b) * c = a * (b * c) + \ \ + (a * b) * Succ[c] = a * (b * Succ[c])" + proof - + fix "c" :: "c" + assume s1_3_cdom: "c \\in Nat" + assume s1_3_induct: "(a * b) * c = a * (b * c)" + have s2_1: "(a * b) * Succ[c] = (a * b) * c + (a * b)" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + have s3_2: "c \\in Nat" + using s1_3_cdom by auto + show ?thesis + using s3_1 s3_2 multint_succ_nat by auto + qed + also have s2_2: "... = a * (b * c) + (a * b)" + using s1_3_induct by auto + also have s2_3: "... = a * (b * c + b)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b * c \\in Int" + using bdom s1_3_cdom nat_in_int multIsInt by auto + have s3_3: "b \\in Int" + using bdom by auto + show ?thesis + using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] by auto + qed + also have s2_4: "... = a * (b * Succ[c])" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "c \\in Nat" + using s1_3_cdom by auto + have s3_3: "b * Succ[c] = b * c + b" + using s3_1 s3_2 multint_succ_nat by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a * b) * Succ[c] = a * (b * Succ[c])" . + qed + have s1_4: "\ c. \ + c \\in negNat; + (a * b) * c = a * (b * c) + \ \ + (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" + proof - + fix "c" :: "c" + assume s1_4_cdom: "c \\in negNat" + assume s1_4_induct: "(a * b) * c = a * (b * c)" + have s2_1: "(a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b)" + proof - + have s3_1: "a * b \\in Int" + using adom bdom multIsInt by auto + have s3_2: "-.c \\in Nat" + using s1_4_cdom minus_neg_nat_in_nat by auto + show ?thesis + using s3_1 s3_2 multint_succ_negnat by auto + qed + also have s2_2: "... = (a * b) * c + -.(a * b)" + proof - + have s3_1: "c \\in Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "-.-.c = c" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_3: "... = a * (b * c) + -.(a * b)" + using s1_4_induct by auto + also have s2_4: "... = a * (b * c) + a * -.b" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b \\in Int" + using bdom by auto + have s3_3: "-.(a * b) = a * -.b" + using s3_1 s3_2 MinusCommutesRightMult by auto + show ?thesis + using s3_3 by auto + qed + also have s2_5: "... = a * (b * c + -.b)" + proof - + have s3_1: "a \\in Int" + using adom by auto + have s3_2: "b * c \\in Int" + using bdom s1_4_cdom neg_nat_in_int multIsInt by auto + have s3_3: "-.b \\in Int" + using bdom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] + by auto + qed + also have s2_6: "... = a * (b * -.-.c + -.b)" + proof - + have s3_1: "c \\in Int" + using s1_4_cdom neg_nat_in_int by auto + have s3_2: "-.-.c = c" + using s3_1 minus_involutive by auto + show ?thesis + using s3_2 by auto + qed + also have s2_7: "... = a * (b * -.Succ[-.c])" + proof - + have s3_1: "b \\in Int" + using bdom by auto + have s3_2: "-.c \\in Nat" + using s1_4_cdom minus_neg_nat_in_nat by auto + have s3_3: "b * -.Succ[-.c] = b * -.-.c + -.b" + using s3_1 s3_2 multint_succ_negnat by auto + show ?thesis + using s3_3 by auto + qed + finally show "(a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" . + qed + have s1_5: "\ c \ Nat: + (a * b) * c = a * (b * c)" + using s1_2 s1_3 natInduct by auto + have s1_6: "\ c \ negNat: + (a * b) * c = a * (b * c)" + using s1_2 s1_4 neg_nat_induction by auto + have s1_7: "\ c \ Int: + (a * b) * c = a * (b * c)" + using s1_5 s1_6 int_union_nat_negnat by auto + } + from this show ?thesis + by auto + qed + + + +(*----------------------------------------------------------------------------*) +(* Properties of `leq` and `add` and `diff`. *) + + +theorem LeqIsBool: + "(m <= n) \\in BOOLEAN" + unfolding leq_def by auto + + +theorem LeqReflexive: + assumes ndom: "n \\in Int" + shows "n <= n" + proof - + have s1_1: "n + 0 = n" + using ndom add_0 by auto + have s1_2: "0 \\in Nat" + using zeroIsNat by auto + have s1_3: "\ r \ Nat: n + r = n" + using s1_1 s1_2 by auto + show ?thesis + unfolding leq_def + using s1_3 by auto + qed + + +theorem LeqTransitive: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m <= n" and nk: "n <= k" + shows "m <= k" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + (* PICK t \in Nat: n + t = k *) + def Q \ "\ x. x \\in Nat \ n + x = k" + def t \ "CHOOSE x: Q(x)" + have s1_2: "t \\in Nat \ n + t = k" + proof - + have s2_1: "\\E x: Q(x)" + using nk by (auto simp: leq_def Q_def) + have s2_2: "Q(t)" + unfolding t_def + using s2_1 chooseI_ex[of Q] by auto + show ?thesis + using s2_2 by (auto simp: Q_def) + qed + have s1_3: "(m + r) + t = k" + using s1_1 s1_2 by auto + have s1_4: "m + (r + t) = k" + proof - + have s2_1: "m \\in Int" + using mdom by auto + have s2_2: "r \\in Int" + using s1_1 nat_in_int by auto + have s2_3: "t \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "(m + r) + t = m + (r + t)" + using s2_1 s2_2 s2_3 AddAssociative_sequent by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "r + t \\in Nat" + using s1_1 s1_2 nat_add_in_nat by auto + have s1_6: "\ q \ Nat: m + q = k" + using s1_4 s1_5 by auto + show ?thesis + unfolding leq_def + using s1_6 by auto + qed + + +theorem leq_0: + assumes ndom: "n \\in Int" + shows "(n \\in Nat) = (0 <= n)" + proof - + have s1_1: "n \\in Nat ==> 0 <= n" + proof - + assume s1_1_ndom: "n \\in Nat" + have s2_1: "0 + n = n" + using ndom add_0_left by auto + have s2_2: "\ r \ Nat: 0 + r = n" + using s2_1 s1_1_ndom by auto + show "0 <= n" + unfolding leq_def + using s2_2 by auto + qed + have s1_2: "0 <= n ==> n \\in Nat" + proof - + assume s1_2_leq: "0 <= n" + def P \ "\ x. x \\in Nat \ 0 + x = n" + def r \ "CHOOSE x: P(x)" + have s2_1: "r \\in Nat \ 0 + r = n" + proof - + have s3_1: "\\E x: P(x)" + using s1_2_leq by (auto simp: leq_def P_def) + have s3_2: "P(r)" + unfolding r_def + using s3_1 chooseI_ex[of P] by auto + show ?thesis + using s3_2 by (auto simp: P_def) + qed + have s2_2: "0 + r = r" + proof - + have s3_1: "r \\in Int" + using s2_1 nat_in_int by auto + show ?thesis + using s3_1 add_0_left by auto + qed + have s2_3: "r = n" + using s2_1 s2_2 by auto + show "n \\in Nat" + using s2_1 s2_3 by auto + qed + have s1_3: "(0 <= n) \\in BOOLEAN" + using LeqIsBool by auto + have s1_4: "(n \\in Nat) \\in BOOLEAN" + using inIsBool isBoolTrueFalse by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorem eq_leq_both: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m = n" + shows "m <= n \ n <= m" + proof - + have s1_1: "m <= n" + proof - + have s2_1: "m + 0 = m" + using mdom add_0 by auto + have s2_2: "m + 0 = n" + using s2_1 mn by auto + have s2_3: "0 \\in Nat" + using zeroIsNat by auto + have s2_4: "\ r \ Nat: m + r = n" + using s2_2 s2_3 by auto + show ?thesis + unfolding leq_def + using s2_4 by auto + qed + have s1_2: "n <= m" + proof - + have s2_1: "n + 0 = n" + using ndom add_0 by auto + have s2_2: "n + 0 = m" + using s2_1 mn by auto + have s2_3: "0 \\in Nat" + using zeroIsNat by auto + have s2_4: "\ r \ Nat: n + r = m" + using s2_2 s2_3 by auto + show ?thesis + unfolding leq_def + using s2_4 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* The inverse of the theorem `eq_leq_both`. *) +theorem leq_both_eq: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m <= n" and nm: "n <= m" + shows "m = n" + proof - + def P \ "\ x. x \\in Nat \ m + x = n" + def p \ "CHOOSE x: P(x)" + def Q \ "\ x. x \\in Nat \ n + x = m" + def q \ "CHOOSE x: Q(x)" + have s1_1: "p \\in Nat \ m + p = n" + proof - + have s2_1: "\ x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(p)" + unfolding p_def using s2_1 chooseI_ex[of P] + by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "q \\in Nat \ n + q = m" + proof - + have s2_1: "\ x: Q(x)" + using nm by (auto simp: leq_def Q_def) + have s2_2: "Q(q)" + unfolding q_def using s2_1 chooseI_ex[of Q] + by auto + show ?thesis + using s2_2 by (auto simp: Q_def) + qed + have s1_3: "(m + p) + (n + q) = n + m" + using s1_1 s1_2 by auto + have s1_4: "p + q = 0" + proof - + have s2_1: "(m + p) + (n + q) = n + m" + using s1_3 by auto + have s2_2: "m + (p + (n + q)) = n + m" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "p \\in Int" + using s1_1 nat_in_int by auto + have s3_3: "n + q \\in Int" + using ndom s1_2 nat_in_int addIsInt by auto + have s3_4: "(m + p) + (n + q) = + m + (p + (n + q))" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + have s2_3: "m + (p + (q + n)) = n + m" + proof - + have s3_1: "n \\in Int" + using ndom by auto + have s3_2: "q \\in Int" + using s1_2 nat_in_int by auto + have s3_3: "n + q = q + n" + using s3_1 s3_2 AddCommutative_sequent + by auto + show ?thesis + using s2_2 s3_3 by auto + qed + have s2_4: "m + ((p + q) + n) = n + m" + proof - + have s3_1: "p \\in Int" + using s1_1 nat_in_int by auto + have s3_2: "q \\in Int" + using s1_2 nat_in_int by auto + have s3_3: "n \\in Int" + using ndom by auto + have s3_4: "p + (q + n) = (p + q) + n" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_3 s3_4 by auto + qed + have s2_5: "m + (n + (p + q)) = n + m" + proof - + have s3_1: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "(p + q) + n = n + (p + q)" + using s3_1 s3_2 + AddCommutative_sequent by auto + show ?thesis + using s2_4 s3_3 by auto + qed + have s2_6: "(m + n) + (p + q) = n + m" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_4: "m + (n + (p + q)) = (m + n) + (p + q)" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_5 s3_4 by auto + qed + have s2_7: "-.(m + n) + ((m + n) + (p + q)) = + -.(m + n) + (n + m)" + using s2_6 by auto + have s2_8: "(-.(m + n) + (m + n)) + (p + q) = + -.(m + n) + (n + m)" + proof - + have s3_1: "m + n \\in Int" + using mdom ndom addIsInt by auto + have s3_2: "-.(m + n) \\in Int" + using s3_1 neg_int_eq_int by auto + have s3_3: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + have s3_4: "-.(m + n) + ((m + n) + (p + q)) = + (-.(m + n) + (m + n)) + (p + q)" + using s3_1 s3_2 s3_3 + AddAssociative_sequent by auto + show ?thesis + using s2_7 s3_4 by auto + qed + have s2_9: "(-.(m + n) + (m + n)) + (p + q) = + -.(m + n) + (m + n)" + proof - + have s3_1: "m \\in Int" + using mdom by auto + have s3_2: "n \\in Int" + using ndom by auto + have s3_3: "n + m = m + n" + using s3_1 s3_2 AddCommutative_sequent by auto + show ?thesis + using s2_8 s3_3 by auto + qed + have s2_10: "0 + (p + q) = 0" + proof - + have s3_1: "-.(m + n) + (m + n) = 0" + proof - + have s4_1: "m + n \\in Int" + using mdom ndom addIsInt by auto + show ?thesis + using s4_1 AddNegCancels_left by auto + qed + show ?thesis + using s2_9 s3_1 by auto + qed + have s2_11: "p + q = 0" + proof - + have s3_1: "p + q \\in Int" + using s1_1 s1_2 nat_in_int addIsInt by auto + show ?thesis + using s2_10 s3_1 add_0_left by auto + qed + show ?thesis + using s2_11 by auto + qed + have s1_5: "p = 0 \ q = 0" + using s1_1 s1_2 s1_4 nat_add_0 by auto + have s1_6: "m + 0 = n" + using s1_1 s1_5 by auto + show ?thesis + using s1_6 mdom add_0 by auto + qed + + +theorem eq_iff_leq_both: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "(m = n) = ((m <= n) \ (n <= m))" + using mdom ndom eq_leq_both leq_both_eq by auto + + +(* +THEOREM dichotomy_leq == + ASSUME NEW a \in Int, NEW b \in Int + PROVE (a <= b) \ (b <= a) +PROOF +<1>1. CASE (a - b) \in Nat + <2>1. b + (a - b) = a + BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b + <2>2. \E n \in Nat: b + n = a + BY <1>1, <2>1 + <2>3. b <= a + BY <2>2 DEF leq + <2> QED + BY <2>3 +<1>2. CASE (a - b) \notin Nat + <2>1. -.(a - b) \in Nat + <3>1. (a - b) \in Int + <4>1. a - b = a + -.b + BY DEF diff + <4>2. a \in Int + OBVIOUS + <4>3. -.b \in Int + BY b \in Int, neg_int_eq_int + <4>4. a + -.b \in Int + BY <4>2, <4>3, addIsInt + <4> QED + BY <4>1, <4>4 + <3> QED + BY <3>1, <1>2, neg_negative_in_nat + <2>2. (b - a) \in Nat + <3>1. -.(a - b) = b - a + BY a \in Int, b \in Int, MinusDiff + <3> QED + BY <2>1, <3>1 + <2>3. a + (b - a) = b + BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b + <2>4. \E n \in Nat: a + n = b + BY <2>2, <2>3 + <2>5. a <= b + BY <2>4 DEF leq + <2> QED + BY <2>5 +<1> QED + BY <1>1, <1>2 +*) +theorem dichotomy_leq: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(a <= b) \ (b <= a)" + proof - + have s1_1: "a - b \\in Nat ==> + (a <= b) \ (b <= a)" + proof - + assume s1_1_asm: "a - b \\in Nat" + have s2_1: "b + (a - b) = a" + using adom bdom a_plus_b_minus_a_eq_b by auto + have s2_2: "\ n \ Nat: b + n = a" + using s1_1_asm s2_1 by auto + have s2_3: "b <= a" + unfolding leq_def + using s2_2 by auto + show "(a <= b) \ (b <= a)" + using s2_3 by auto + qed + have s1_2: "a - b \\notin Nat ==> + (a <= b) \ (b <= a)" + proof - + assume s1_2_asm: "a - b \\notin Nat" + have s2_1: "-.(a - b) \\in Nat" + proof - + have s3_1: "(a - b) \\in Int" + proof - + have s4_1: "a - b = a + -.b" + unfolding diff_def by auto + have s4_2: "a \\in Int" + using adom by auto + have s4_3: "-.b \\in Int" + using bdom neg_int_eq_int by auto + have s4_4: "a + -.b \\in Int" + using s4_2 s4_3 addIsInt by auto + show ?thesis + using s4_1 s4_4 by auto + qed + show ?thesis + using s3_1 s1_2_asm neg_negative_in_nat by auto + qed + have s2_2: "(b - a) \\in Nat" + proof - + have s3_1: "-.(a - b) = b - a" + using adom bdom MinusDiff by auto + show ?thesis + using s2_1 s3_1 by auto + qed + have s2_3: "a + (b - a) = b" + using adom bdom a_plus_b_minus_a_eq_b by auto + have s2_4: "\ n \ Nat: a + n = b" + using s2_2 s2_3 by auto + have s2_5: "a <= b" + unfolding leq_def + using s2_4 by auto + show "(a <= b) \ (b <= a)" + using s2_5 by auto + qed + show ?thesis + using s1_1 s1_2 by auto + qed + + +theorem trichotomy_less: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(a < b) \ (a = b) \ (b < a)" + using adom bdom dichotomy_leq by (auto simp: less_def) + + +theorem trichotomy_less_0: + assumes ndom: "n \\in Int" + shows "(n < 0) \ (n = 0) \ (0 < n)" + using ndom zero_in_int trichotomy_less by auto + + +(* +THEOREM leq_mult_monotonic == + \A m \in Int: \A n \in Int: + (m <= n) => (\A k \in Nat: k * m <= k * n) +PROOF +<1>1. SUFFICES + ASSUME NEW m \in Int, NEW n \in Int, NEW k \in Nat, + m <= n + PROVE k * m <= k * n + BY <1>1 +<1>2. PICK r \in Nat: m + r = n + BY <1>1 DEF leq +<1>3. k * (m + r) = k * n + BY <1>2 +<1>4. k * m + k * r = k * n + <2>1. k \in Int + BY <1>1, nat_in_int + <2>2. m \in Int + BY <1>1 + <2>3. r \in Int + BY <1>2, nat_in_int + <2>4. k * (m + r) = k * m + k * r + BY <2>1, <2>2, <2>3, MultLeftDistributesAdd + <2> QED + BY <1>3, <2>4 +<1>5. k * r \in Nat + BY <1>1, <1>2, nat_mult_in_nat +<1>6. \E x \in Nat: k * m + x = k * n + BY <1>4, <1>5 +<1> QED + BY <1>6 DEF leq +*) +theorem leq_mult_monotonic: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Nat" and mn: "m <= n" + shows "k * m <= k * n" + proof - + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_2: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_3: "k * (m + r) = k * n" + using s1_2 by auto + have s1_4: "k * m + k * r = k * n" + proof - + have s2_1: "k \\in Int" + using kdom nat_in_int by auto + have s2_2: "m \\in Int" + using mdom by auto + have s2_3: "r \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "k * (m + r) = k * m + k * r" + using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "k * r \\in Nat" + using kdom s1_2 nat_mult_in_nat by auto + have s1_6: "\ x \ Nat: k * m + x = k * n" + using s1_4 s1_5 by auto + show ?thesis + unfolding leq_def + using s1_6 by auto + qed + + +(* +THEOREM leq_mult_monotonic_neg: + \A m \in Int: \A n \in Int: + (m <= n) => (\A k \in negNat: k * n <= k * m) +PROOF +<1>1. SUFFICES + ASSUME NEW m \in Int, NEW n \in Int, NEW k \in negNat, + m <= n + PROVE k * n <= k * m + BY <1>1 +<1>2. PICK r \in Nat: m + r = n + BY <1>1 DEF leq +<1>3. k * (m + r) = k * n + BY <1>2 +<1>4. k * m + k * r = k * n + <2>1. k \in Int + BY <1>1, neg_nat_in_int + <2>2. m \in Int + BY <1>1 + <2>3. r \in Int + BY <1>2, nat_in_int + <2>4. k * (m + r) = k * m + k * r + BY <2>1, <2>2, <2>3, MultLeftDistributesAdd + <2> QED + BY <1>3, <2>4 +<1>5. k * m = k * n + -.(k * r) + <2>1. (k * m + k * r) + -.(k * r) = k * n + -.(k * r) + BY <1>4 + <2>2. k * m + (k * r + -.(k * r)) = k * n + -.(k * r) + <3>1. k * m \in Int + BY <1>1, neg_nat_in_int, multIsInt + <3>2. k * r \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. r \in Int + BY <1>2, nat_in_int + <4> QED + BY <4>1, <4>2, multIsInt + <3>3. -.(k * r) \in Int + BY <3>2, neg_int_eq_int + <3>4. (k * m + k * r) + -.(k * r) = + k * m + (k * r + -.(k * r)) + BY <3>1, <3>2, <3>3, AddAssociative + <3> QED + BY <2>1, <3>4 + <2>3. k * r + -.(k * r) = 0 + <3>1. k * r \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. r \in Int + BY <1>2, nat_in_int + <4> QED + BY <4>1, <4>2, multIsInt + <3> QED + BY <3>1, AddNegCancels + <2>4. k * m + 0 = k * n + -.(k * r) + BY <2>2, <2>3 + <2>5. k * m + 0 = k * m + <3>1. k * m \in Int + <4>1. k \in Int + BY <1>1, neg_nat_in_int + <4>2. m \in Int + BY <1>1 + <4> QED + BY <4>1, <4>2, multIsInt + <3> QED + BY <3>1, add_0 + <2> QED + BY <2>4, <2>5 +<1>6. k * n + (-.k * r) = k * m + <2>1. k * n + -.(k * r) = k * m + BY <1>5 \* symmetric + <2>2. -.(k * r) = -.k * r + <3>1. k \in Int + BY <1>1, neg_nat_in_int + <3>2. r \in Int + BY <1>2, nat_in_int + <3> QED + BY <3>1, <3>2, MinusCommutesLeftMult + <2> QED + BY <2>1, <2>2 +<1>7. -.k * r \in Nat + <2>1. -.k \in Nat + BY <1>1, minus_neg_nat_in_nat + <2>2. r \in Nat + BY <1>2 + <2> QED + BY <2>1, <2>2, nat_mult_in_nat +<1>8. \E x \in Nat: k * n + x = k * m + BY <1>6, <1>7 +<1> QED + BY <1>8 DEF leq +*) +theorem leq_mult_monotonic_neg: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in negNat" and mn: "m <= n" + shows "k * n <= k * m" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_2: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_3: "k * (m + r) = k * n" + using s1_2 by auto + have s1_4: "k * m + k * r = k * n" + proof - + have s2_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s2_2: "m \\in Int" + using mdom by auto + have s2_3: "r \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "k * (m + r) = k * m + k * r" + using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto + show ?thesis + using s1_3 s2_4 by auto + qed + have s1_5: "k * m = k * n + -.(k * r)" + proof - + have s2_1: "(k * m + k * r) + -.(k * r) = k * n + -.(k * r)" + using s1_4 by auto + have s2_2: "k * m + (k * r + -.(k * r)) = k * n + -.(k * r)" + proof - + have s3_1: "k * m \\in Int" + using mdom kdom neg_nat_in_int multIsInt by auto + have s3_2: "k * r \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + have s3_3: "-.(k * r) \\in Int" + using s3_2 neg_int_eq_int by auto + have s3_4: "(k * m + k * r) + -.(k * r) = + k * m + (k * r + -.(k * r))" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + have s2_3: "k * r + -.(k * r) = 0" + proof - + have s3_1: "k * r \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + show ?thesis + using s3_1 AddNegCancels by auto + qed + have s2_4: "k * m + 0 = k * n + -.(k * r)" + using s2_2 s2_3 by auto + have s2_5: "k * m + 0 = k * m" + proof - + have s3_1: "k * m \\in Int" + proof - + have s4_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s4_2: "m \\in Int" + using mdom by auto + show ?thesis + using s4_1 s4_2 multIsInt by auto + qed + show ?thesis + using s3_1 add_0 by auto + qed + show ?thesis + using s2_4 s2_5 by auto + qed + have s1_6: "k * n + (-.k * r) = k * m" + proof - + have s2_1: "k * n + -.(k * r) = k * m" + using s1_5 by auto + have s2_2: "-.(k * r) = -.k * r" + proof - + have s3_1: "k \\in Int" + using kdom neg_nat_in_int by auto + have s3_2: "r \\in Int" + using s1_2 nat_in_int by auto + show ?thesis + using s3_1 s3_2 MinusCommutesLeftMult by auto + qed + show ?thesis + using s2_1 s2_2 by auto + qed + have s1_7: "-.k * r \\in Nat" + proof - + have s2_1: "-.k \\in Nat" + using kdom minus_neg_nat_in_nat by auto + have s2_2: "r \\in Nat" + using s1_2 by auto + show ?thesis + using s2_1 s2_2 nat_mult_in_nat by auto + qed + have s1_8: "\ x \ Nat: k * n + x = k * m" + using s1_6 s1_7 by auto + show ?thesis + unfolding leq_def + using s1_8 by auto + qed + + +(*----------------------------------------------------------------------------*) + + +(* Monotonicity of addition with respect to order. *) +theorem leq_add_monotonic_right: + assumes mdom: "m \ Int" and ndom: "n \ Int" and + kdom: "k \ Int" and mn: "m <= n" + shows "m + k <= n + k" + proof - + have s1_1: "\ r: r \ Nat \ m + r = n" + using mn by (auto simp: leq_def) + def r \ "CHOOSE r: r \ Nat \ m + r = n" + have s1_2: "r \ Nat \ m + r = n" + unfolding r_def + by (rule chooseI_ex, rule s1_1) + have s1_3: "(m + r) + k = n + k" + using s1_2 by auto + have s1_4: "(m + k) + r = n + k" + proof - + have s2_1: "(r + m) + k = n + k" + proof - + have s3_1: "m \ Int" + using mdom by auto + have s3_2: "r \ Int" + using s1_2 nat_in_int by auto + have s3_3: "m + r = r + m" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s1_3 s3_3 by auto + qed + have s2_2: "r + (m + k) = n + k" + proof - + have s3_1: "r \ Int" + using s1_2 nat_in_int by auto + have s3_2: "m \ Int" + using mdom by auto + have s3_3: "k \ Int" + using kdom by auto + have s3_4: "(r + m) + k = r + (m + k)" + using s3_1 s3_2 s3_3 AddAssociative_sequent by auto + show ?thesis + using s2_1 s3_4 by auto + qed + show ?thesis + proof - + have s3_1: "r \ Int" + using s1_2 nat_in_int by auto + have s3_2: "m + k \ Int" + using mdom kdom addIsInt by auto + have s3_3: "r + (m + k) = (m + k) + r" + using s3_1 s3_2 AddCommutative by auto + show ?thesis + using s2_2 s3_3 by auto + qed + qed + have s1_5: "\ t \ Nat: + (m + k) + t = n + k" + using s1_2 s1_4 by auto + show "m + k <= n + k" + using s1_5 by (auto simp: leq_def) + qed + + +theorem eq_add_inj_right: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m + k = n + k" + shows "m = n" + proof - + have s1_1: "(m + k) + -.k = (n + k) + -.k" + using mn by auto + have s1_2: "m + (k + -.k) = n + (k + -.k)" + proof - + have minus_kdom: "-.k \\in Int" + using kdom neg_int_eq_int by auto + have s2_1: "(m + k) + -.k = m + (k + -.k)" + using mdom kdom minus_kdom AddAssociative_sequent + by auto + have s2_2: "(n + k) + -.k = n + (k + -.k)" + using ndom kdom minus_kdom AddAssociative_sequent + by auto + show ?thesis + using s1_1 s2_1 s2_2 by auto + qed + have s1_3: "m + 0 = n + 0" + proof - + have s2_1: "k + -.k = 0" + using kdom AddNegCancels_sequent by auto + show ?thesis + using s1_2 s2_1 by auto + qed + have s1_4: "m + 0 = m" + using mdom add_0 by auto + have s1_5: "n + 0 = n" + using ndom add_0 by auto + show ?thesis + using s1_3 s1_4 s1_5 by auto + qed + + +theorem less_add_monotonic_right: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m < n" + shows "m + k < n + k" + proof - + have s1_1: "m + k <= n + k" + proof - + have s2_1: "m <= n" + using mn by (auto simp: less_def) + show ?thesis + using mdom ndom kdom s2_1 leq_add_monotonic_right by auto + qed + have s1_2: "m + k \ n + k" + proof - + have s2_1: "m + k = n + k ==> FALSE" + proof - + assume s2_1_asm: "m + k = n + k" + have s3_1: "m = n" + using s2_1_asm mdom ndom kdom eq_add_inj_right + by auto + have s3_2: "m \ n" + using mn by (auto simp: less_def) + show "FALSE" + using s3_1 s3_2 by auto + qed + show ?thesis + using s2_1 by auto + qed + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_add_monotonic_left: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + kdom: "k \\in Int" and mn: "m < n" + shows "k + m < k + n" + proof - + have s1_1: "m + k < n + k" + using mdom ndom kdom mn less_add_monotonic_right by auto + have s1_2: "m + k = k + m" + using mdom kdom AddCommutative_sequent by auto + have s1_3: "n + k = k + n" + using ndom kdom AddCommutative_sequent by auto + show ?thesis + using s1_1 s1_2 s1_3 by auto + qed + + +theorem minus_less: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m < n" + shows "-.n < -.m" + proof - + have s1_1: "m + -.m < n + -.m" + proof - + have s2_1: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using mdom s2_1 ndom mn less_add_monotonic_right by auto + qed + have s1_2: "0 < n + -.m" + proof - + have s2_1: "m + -.m = 0" + using mdom AddNegCancels_sequent by auto + show ?thesis + using s1_1 s2_1 by auto + qed + have s1_3: "-.n + 0 < -.n + (n + -.m)" + proof - + have s2_1: "-.n \\in Int" + using ndom neg_int_eq_int by auto + have s2_2: "0 \\in Int" + using zeroIsNat nat_in_int by auto + have s2_3: "n + -.m \\in Int" + proof - + have s3_1: "n \\in Int" + using ndom by auto + have s3_2: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s3_1 s3_2 addIsInt by auto + qed + show ?thesis + using s1_2 s2_1 s2_2 s2_3 less_add_monotonic_left + by auto + qed + have s1_4: "-.n < -.n + (n + -.m)" + proof - + have s2_1: "-.n + 0 = -.n" + using ndom neg_int_eq_int add_0 by auto + show ?thesis + using s1_3 s2_1 by auto + qed + have s1_5: "-.n + (n + -.m) = -.m" + proof - + have s2_1: "-.n + (n + -.m) = (-.n + n) + -.m" + using ndom mdom neg_int_eq_int AddAssociative_sequent + by auto + have s2_2: "-.n + n = 0" + using ndom AddNegCancels_left by auto + have s2_3: "(-.n + n) + -.m = 0 + -.m" + using s2_2 by auto + have s2_4: "0 + -.m = -.m" + proof - + have s3_1: "-.m \\in Int" + using mdom neg_int_eq_int by auto + show ?thesis + using s3_1 add_0_left by auto + qed + show ?thesis + using s2_1 s2_3 s2_4 by auto + qed + show ?thesis + using s1_4 s1_5 by auto + qed + + +theorem leq_diff_add: + assumes ndom: "n \ Int" and mdom: "m \ Int" and + kdom: "k \ Int" and nmk: "n - m <= k" + shows "n <= k + m" + proof - + def P \ "\ x. x \\in Nat \ (n - m) + x = k" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ (n - m) + r = k" + proof - + have s2_1: "\\E x: P(x)" + using nmk by (auto simp: leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "m + ((n - m) + r) = m + k" + using s1_1 by auto + have s1_3: "(m + -.m) + (n + r) = k + m" + proof - + have s2_1: "m + ((n - m) + r) = (m + -.m) + (n + r)" + proof - + have s3_1: "m + ((n - m) + r) = + m + ((n + -.m) + r)" + unfolding diff_def by auto + also have s3_2: "... = m + ((-.m + n) + r)" + using mdom neg_int_eq_int ndom + AddCommutative_sequent by auto + also have s3_3: "... = (m + (-.m + n)) + r" + using mdom neg_int_eq_int ndom s1_1 nat_in_int + addIsInt AddAssociative_sequent by auto + also have s3_4: "... = ((m + -.m) + n) + r" + using mdom neg_int_eq_int ndom + AddAssociative_sequent by auto + also have s3_5: "... = (m + -.m) + (n + r)" + using mdom neg_int_eq_int addIsInt ndom + s1_1 nat_in_int AddAssociative_sequent by auto + finally show ?thesis . + qed + have s2_2: "m + k = k + m" + using mdom kdom AddCommutative_sequent by auto + show ?thesis + using s1_2 s2_1 s2_2 by auto + qed + have s1_4: "n + r = k + m" + proof - + have s2_1: "m + -.m = 0" + using mdom AddNegCancels_sequent by auto + have s2_2: "(m + -.m) + (n + r) = n + r" + proof - + have s3_1: "(m + -.m) + (n + r) = 0 + (n + r)" + using s2_1 by auto + have s3_2: "0 + (n + r) = n + r" + proof - + have s4_1: "n + r \\in Int" + using ndom s1_1 nat_in_int addIsInt by auto + show ?thesis + using s4_1 add_0_left by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + show ?thesis + using s1_3 s2_2 by auto + qed + have s1_5: "\ x \ Nat: n + x = k + m" + using s1_1 s1_4 by auto + show ?thesis + unfolding leq_def + using s1_5 by auto + qed + + +theorem leq_diff_nat: + assumes mdom: "m \\in Int" and ndom: "n \\in Nat" + shows "m - n <= m" + proof - + have s1_1: "(m - n) + n = m" + proof - + have s2_1: "-.n + n = 0" + proof - + have s3_1: "n \\in Int" + using ndom nat_in_int by auto + show ?thesis + using s3_1 AddNegCancels_left by auto + qed + have s2_2: "m + (-.n + n) = m + 0" + using s2_1 by auto + have s2_3: "(m + -.n) + n = m" + proof - + have s3_1: "m + (-.n + n) = (m + -.n) + n" + proof - + have s4_1: "m \\in Int" + using mdom by auto + have s4_2: "-.n \\in Int" + using ndom minus_nat_in_int by auto + have s4_3: "n \\in Int" + using ndom nat_in_int by auto + show ?thesis + using s4_1 s4_2 s4_3 AddAssociative_sequent by auto + qed + have s3_2: "m + 0 = m" + using mdom add_0 by auto + show ?thesis + using s2_2 s3_1 s3_2 by auto + qed + show ?thesis + unfolding diff_def + using s2_3 by auto + qed + have s1_2: "\ r \ Nat: (m - n) + r = m" + using ndom s1_1 by auto + show ?thesis + unfolding leq_def + using s1_2 by auto + qed + + +(* See also the theorem `less_leq_trans`. *) +theorem leq_less_trans: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + and kdom: "k \\in Int" and + mn: "m <= n" and nk: "n < k" + shows "m < k" + proof - + have s1_1: "(n <= k) \ (n \ k)" + using nk by (auto simp: less_def) + have s1_2: "m <= k" + proof - + have s2_2: "n <= k" + using s1_1 by auto + have s2_3: "m <= n" + using mn by auto + show ?thesis + using s2_2 s2_3 mdom ndom kdom LeqTransitive + by iprover + qed + have s1_3: "m \ k" + proof - + { + assume s3_1: "m = k" + have s3_2: "k <= n" + using mn s3_1 by auto + have s3_3: "n <= k" + using nk by (auto simp: less_def) + have s3_4: "n = k" + using s3_2 s3_3 ndom kdom leq_both_eq + by auto + have "FALSE" + using s3_4 s1_1 by auto + } + from this show "m \ k" + by auto + qed + show ?thesis + unfolding less_def + using s1_2 s1_3 by auto + qed + + +(* See also the theorem `leq_less_trans`. *) +theorem less_leq_trans: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + and kdom: "k \\in Int" and + mn: "m < n" and nk: "n <= k" + shows "m < k" + proof - + have s1_1: "(m <= n) \ (m \ n)" + using mn by (auto simp: less_def) + have s1_2: "m <= k" + proof - + have s2_1: "m <= n" + using s1_1 by auto + have s2_2: "n <= k" + using nk by auto + show ?thesis + using s2_1 s2_2 mdom ndom kdom LeqTransitive + by iprover + qed + have s1_3: "m \ k" + proof - + have s2_1: "m = k ==> FALSE" + proof - + assume s2_1_asm: "m = k" + have s3_1: "n <= m" + using nk s2_1_asm by auto + have s3_2: "m <= n" + using mn by (auto simp: less_def) + have s3_3: "m = n" + using s3_1 s3_2 mdom ndom leq_both_eq + by auto + show "FALSE" + using s3_3 s1_1 by auto + qed + show ?thesis + using s2_1 by auto + qed + show ?thesis + unfolding less_def + using s1_2 s1_3 by auto + qed + + +theorem less_not_leq: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" + shows "(m < n) = (~ (n <= m))" + proof - + have s1_1: "m < n ==> ~ (n <= m)" + proof - + assume mn: "m < n" + have s2_1: "(m <= n) \ (m \ n)" + using mn by (auto simp: less_def) + have s2_2: "n <= m ==> FALSE" + proof - + assume nm: "n <= m" + have s3_1: "m = n" + using mdom ndom s2_1 nm leq_both_eq by blast + show "FALSE" + using s2_1 s3_1 by auto + qed + show "~ (n <= m)" + using s2_2 by auto + qed + have s1_2: "(~ (n <= m)) ==> m < n" + proof - + assume nm: "~ (n <= m)" + have s2_1: "(m <= n) \ (n <= m)" + using mdom ndom dichotomy_leq by auto + have s2_2: "m <= n" + using nm s2_1 by auto + have s2_3: "m = n ==> FALSE" + proof - + assume eq: "m = n" + have s3_1: "n <= n" + using ndom LeqReflexive by auto + have s3_2: "n <= m" + using eq s3_1 by auto + show "FALSE" + using nm s3_2 by auto + qed + have s2_4: "m \ n" + using s2_3 by auto + show "m < n" + unfolding less_def + using s2_2 s2_4 by auto + qed + have s1_3: "(~ (n <= m)) \\in BOOLEAN" + by auto + have s1_4: "(m < n) \\in BOOLEAN" + unfolding less_def by auto + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed + + +theorem leq_linear: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(~ (a <= b)) \ (a = b) \ (~ (b <= a))" + using adom bdom trichotomy_less[of "a" "b"] + less_not_leq[of "a" "b"] + less_not_leq[of "b" "a"] by auto + + +theorem leq_geq_linear: + assumes adom: "a \\in Int" and bdom: "b \\in Int" + shows "(~ (b >= a)) \ (a = b) \ (~ (b <= a))" + using adom bdom leq_linear by auto + + +(* +THEOREM less_is_leq_plus_one == + ASSUME NEW m \in Int, NEW n \in Int, m < n + PROVE m + 1 <= n +PROOF +<1>1. PICK r \in Nat: m + r = n + BY m < n DEF less, leq +<1>2. PICK k \in Nat: r = Succ[k] + <2>1. r = 0 \/ \E k \in Nat: r = Succ[k] + BY <1>1, nat_0_succ + <2>2. ASSUME r = 0 + PROVE FALSE + <3>1. m + 0 = n + BY <1>1, <2>2 + <3>2. m = n + BY <3>1, m \in Int, add_0 + <3>3. m # n + BY m < n DEF less + <3> QED + BY <3>2, <3>3 + <2> QED + BY <2>1, <2>2 +<1>3. m + Succ[k] = n + BY <1>1, <1>2 +<1>4. m + (k + 1) = n + <2>1. Succ[k] = k + 1 + BY <1>2, nat_add_1 + <2> QED + BY <1>3, <2>1 +<1>5. m + (1 + k) = n + <2>1. k \in Int + BY <1>2, nat_in_int + <2>2. 1 \in Int + BY oneIsNat, nat_in_int + <2>3. k + 1 = 1 + k + BY <2>1, <2>2, AddCommutative + <2> QED + BY <1>4, <2>3 +<1>6. (m + 1) + k = n + <2>1. m \in Int + OBVIOUS + <2>2. 1 \in Int + BY oneIsNat nat_in_int + <2>3. k \in Int + BY <1>2, nat_in_int + <2>4. m + (1 + k) = (m + 1) + k + BY <2>1, <2>2, <2>3, AddAssociative + <2> QED + BY <1>5, <2>4 +<1>7. \E q \in Nat: (m + 1) + q = n + BY <1>2, <1>6 +<1> QED + BY <1>7 DEF leq +*) +theorem less_is_leq_plus_one: + assumes mdom: "m \\in Int" and ndom: "n \\in Int" and + mn: "m < n" + shows "m + 1 <= n" + proof - + (* PICK r \in Nat: m + r = n *) + def P \ "\ x. x \\in Nat \ m + x = n" + def r \ "CHOOSE x: P(x)" + have s1_1: "r \\in Nat \ m + r = n" + proof - + have s2_1: "\\E x: P(x)" + using mn by (auto simp: less_def leq_def P_def) + have s2_2: "P(r)" + unfolding r_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + (* PICK k \in Nat: r = Succ[k] *) + def Q \ "\ x. x \\in Nat \ r = Succ[x]" + def k \ "CHOOSE x: Q(x)" + have s1_2: "k \\in Nat \ r = Succ[k]" + proof - + have s2_1: "r = 0 \ + (\ x \ Nat: r = Succ[x])" + using s1_1 nat_0_succ by auto + have s2_2: "r = 0 ==> FALSE" + proof - + assume s2_2_asm: "r = 0" + have s3_1: "m + 0 = n" + using s1_1 s2_2_asm by auto + have s3_2: "m = n" + using s3_1 mdom add_0 by auto + have s3_3: "m \ n" + using mn by (auto simp: less_def) + show ?thesis + using s3_2 s3_3 by auto + qed + have s2_3: "\ x \ Nat: r = Succ[x]" + using s2_1 s2_2 by auto + have s2_4: "\\E x: Q(x)" + unfolding Q_def using s2_3 by auto + have s2_5: "Q(k)" + unfolding k_def + using s2_4 chooseI_ex[of Q] by auto + show ?thesis + using s2_5 by (auto simp: Q_def) + qed + have s1_3: "m + Succ[k] = n" + using s1_1 s1_2 by auto + have s1_4: "m + (k + 1) = n" + proof - + have s2_1: "Succ[k] = k + 1" + using s1_2 nat_add_1 by fast + show ?thesis + using s1_3 s2_1 by auto + qed + have s1_5: "m + (1 + k) = n" + proof - + have s2_1: "k \\in Int" + using s1_2 nat_in_int by auto + have s2_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s2_3: "k + 1 = 1 + k" + using s2_1 s2_2 AddCommutative_sequent by auto + show ?thesis + using s1_4 s2_3 by auto + qed + have s1_6: "(m + 1) + k = n" + proof - + have s2_1: "m \\in Int" + using mdom by auto + have s2_2: "1 \\in Int" + using oneIsNat nat_in_int by auto + have s2_3: "k \\in Int" + using s1_2 nat_in_int by auto + have s2_4: "m + (1 + k) = (m + 1) + k" + using s2_1 s2_2 s2_3 AddAssociative_sequent by auto + show ?thesis + using s1_5 s2_4 by auto + qed + have s1_7: "\ q \ Nat: (m + 1) + q = n" + using s1_2 s1_6 by auto + show ?thesis + unfolding leq_def using s1_7 by auto + qed + + +(* +THEOREM less_succ_nat == + ASSUME NEW n \in Nat + PROVE n < Succ[n] +PROOF +<1>1. n <= Succ[n] + <2>1. Succ[n] = n + 1 + BY n \in Nat, nat_add_1 + <2>2. 1 \in Nat + BY oneIsNat + <2>3. \E r \in Nat: n + r = Succ[n] + BY <2>1, <2>2 + <2> QED + BY <2>3 DEF leq +<1>2. n # Succ[n] + BY n \in Nat, succIrrefl +<1> QED + BY <1>1, <1>2 DEF less +*) +theorem less_succ_nat: + assumes ndom: "n \\in Nat" + shows "n < Succ[n]" + proof - + have s1_1: "n <= Succ[n]" + proof - + have s2_1: "Succ[n] = n + 1" + using ndom nat_add_1 by fast + have s2_2: "1 \\in Nat" + using oneIsNat by auto + have s2_3: "\\E r \\in Nat: n + r = Succ[n]" + using s2_1 s2_2 by auto + show ?thesis + unfolding leq_def + using s2_3 by auto + qed + have s1_2: "n \ Succ[n]" + using ndom succIrrefl by auto + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_pred_nat: + assumes ndom: "n \\in Nat \ {0}" + shows "Pred[n] < n" + proof - + have s1_1: "Pred[n] <= n" + proof - + have s2_1: "Pred[n] \\in Nat" + using ndom Pred_in_nat by auto + have s2_2: "Succ[Pred[n]] = n" + using ndom Succ_Pred_nat by auto + have s2_3: "Pred[n] + 1 = n" + proof - + have s3_1: "Succ[Pred[n]] = Pred[n] + 1" + using s2_1 nat_add_1 by fast + show ?thesis + using s2_2 s3_1 by auto + qed + have s2_4: "1 \\in Nat" + using oneIsNat by auto + have s2_5: "\\E x \\in Nat: Pred[n] + x = n" + using s2_3 s2_4 by auto + show ?thesis + unfolding leq_def + using s2_5 by auto + qed + have s1_2: "Pred[n] \ n" + using ndom pred_irrefl by auto + show ?thesis + unfolding less_def + using s1_1 s1_2 by auto + qed + + +theorem less_isucc: + assumes idom: "i \\in Int" + shows "i < iSucc[i]" + proof - + have s1_1: "i \\in Nat ==> i < iSucc[i]" + proof - + assume s1_1_asm: "i \\in Nat" + have s2_1: "i < Succ[i]" + using s1_1_asm less_succ_nat by auto + have s2_2: "iSucc[i] = Succ[i]" + unfolding iSucc_def + using idom s1_1_asm by auto + show "i < iSucc[i]" + using s2_1 s2_2 by auto + qed + have s1_2: "i \\notin Nat ==> i < iSucc[i]" + proof - + assume s1_2_asm: "i \\notin Nat" + have s2_1: "iSucc[i] = -.Pred[-.i]" + unfolding iSucc_def + using idom s1_2_asm by auto + have s2_2: "-.i \\in Nat \ {0}" + using idom s1_2_asm minus_neg_int_in_nat by auto + have s2_3: "Pred[-.i] < -.i" + using s2_2 less_pred_nat by auto + have s2_4: "-.-.i < -.Pred[-.i]" + proof - + have s3_1: "Pred[-.i] \\in Int" + using s2_2 Pred_in_nat nat_in_int by auto + have s3_2: "-.i \\in Int" + using s2_2 nat_in_int by auto + show ?thesis + using s2_3 s3_1 s3_2 minus_less by auto + qed + have s2_5: "-.-.i = i" + using idom minus_involutive by auto + show "i < iSucc[i]" + using s2_1 s2_4 s2_5 by auto + qed + show ?thesis + using s1_1 s1_2 by iprover + qed + + +theorem less_i_add_1: + assumes idom: "i \\in Int" + shows "i < i + 1" + proof - + have s1_1: "i < iSucc[i]" + using idom less_isucc by auto + have s1_2: "iSucc[i] = i + 1" + using idom isucc_eq_add_1 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM i_less_j_nat == + ASSUME NEW i \in Int, NEW j \in Nat, i <= j + PROVE i < Succ[j] +PROOF +<1>1. j < Succ[j] + BY j \in Nat, less_succ_nat +<1>2. j \in Int + BY j \in Nat, nat_in_int +<1>3. Succ[j] \in Int + BY j \in Nat, succIsNat, nat_in_int +<1> QED + BY i \in Int, <1>2, <1>3, i <= j, <1>1, + leq_less_trans +*) +theorem i_less_succ_j_nat: + assumes idom: "i \\in Int" and jdom: "j \\in Nat" and + ij: "i <= j" + shows "i < Succ[j]" + proof - + have s1_1: "j < Succ[j]" + using jdom less_succ_nat by auto + have s1_2: "j \\in Int" + using jdom nat_in_int by auto + have s1_3: "Succ[j] \\in Int" + using jdom succIsNat nat_in_int by auto + show ?thesis + using idom s1_2 s1_3 ij s1_1 leq_less_trans by auto + qed + + +theorem less_add_1: + assumes idom: "i \\in Int" and jdom: "j \\in Int" and + ij: "i <= j" + shows "i < j + 1" + proof - + have s1_1: "j < j + 1" + using jdom less_i_add_1 by auto + have s1_2: "j + 1 \\in Int" + proof - + have s2_1: "1 \\in Int" + using oneIsNat nat_in_int by auto + show ?thesis + using jdom s2_1 addIsInt by auto + qed + show ?thesis + using idom jdom s1_2 ij s1_1 leq_less_trans by auto + qed + + +(*----------------------------------------------------------------------------*) +(* Intervals of integers. *) + + +(* +THEOREM SplitNat0 == + Nat = {0} \cup {n + 1: n \in Nat} +PROOF +<1>1. ASSUME NEW n \in Nat + PROVE n \in {0} \cup {k + 1: k \in Nat} + <2>1. n = 0 \/ \E m \in Nat: n = Succ[m] + BY <1>1, nat_0_succ + <2>2. (\E m \in Nat: n = Succ[m]) <=> + n \in {Succ[m]: m \in Nat} + OBVIOUS + <2>3. ASSUME NEW m \in Nat + PROVE Succ[m] = m + 1 + BY <2>3, nat_add_1 + <2>4. {Succ[m]: m \in Nat} = + {m + 1: m \in Nat} + BY <2>3 + <2>5. n \in {0} \/ n \in {m + 1: m \in Nat} + BY <2>1, <2>2, <2>4 + <2> QED + BY <2>5 +<1>2. ASSUME NEW n \in {0} \cup {k + 1: k \in Nat} + PROVE n \in Nat + <2>1. CASE n = 0 + BY <2>1, zeroIsNat + <2>2. CASE n \in {k + 1: k \in Nat} + <3>1. PICK k \in Nat: n = k + 1 + BY <2>2 + <3>2. k + 1 \in Nat + BY <3>1, oneIsNat, nat_add_in_nat + <3> QED + BY <3>1, <3>2 + <2> QED + BY <1>2, <2>1, <2>2 +<1> QED + BY <1>1, <1>2 +*) +theorem SplitNat0: + "Nat = {0} \ {n + 1: n \\in Nat}" + proof - + have s1_1: "\\A n \\in Nat: n \\in {0} \ {k + 1: k \\in Nat}" + proof - + { + fix "n" :: "c" + assume s1_1_asm: "n \\in Nat" + have s2_1: "n = 0 \ (\\E m \\in Nat: n = Succ[m])" + using s1_1_asm nat_0_succ by auto + have s2_2: "\\E m \\in Nat: n = Succ[m] ==> + n \\in {Succ[m]: m \\in Nat}" + by auto + have s2_3: "{Succ[m]: m \\in Nat} = {m + 1: m \\in Nat}" + proof - + have s3_1: "\\A y \\in {Succ[m]: m \\in Nat}: + y \\in {m + 1: m \\in Nat}" + proof - + { + fix "y" :: "c" + assume s4_1: "y \\in {Succ[m]: m \\in Nat}" + have s4_2: "\\E m \\in Nat: y = Succ[m]" + using s4_1 setOfAll by auto + have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" + using s4_2 by (auto simp: bEx_def) + (* PICK r \in Nat: y = Succ[r] *) + def P \ "\ x. x \\in Nat \ y = Succ[x]" + def r \ "CHOOSE x: P(x)" + have s4_4: "r \\in Nat \ y = Succ[r]" + proof - + have s5_1: "\\E x: P(x)" + using s4_3 by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_5: "r \\in Nat \ y = r + 1" + proof - + have s5_1: "r \\in Nat" + using s4_4 by auto + have s5_2: "Succ[r] = r + 1" + using s5_1 nat_add_1 by fast + have s5_3: "y = Succ[r]" + using s4_4 by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + have "y \\in {m + 1: m \\in Nat}" + using s4_5 setOfAll_eqI by auto + } + from this show ?thesis + by auto + qed + have s3_2: "\\A y \\in {m + 1: m \\in Nat}: + y \\in {Succ[m]: m \\in Nat}" + proof - + { + fix "y" :: "c" + assume s4_1: "y \\in {m + 1: m \\in Nat}" + have s4_2: "\\E m \\in Nat: y = m + 1" + using s4_1 setOfAll by auto + have s4_3: "\\E m: m \\in Nat \ y = m + 1" + using s4_2 by (auto simp: bEx_def) + (* PICK r \in Nat: y = r + 1 *) + def P \ "\ x. x \\in Nat \ y = x + 1" + def r \ "CHOOSE x: P(x)" + have s4_4: "r \\in Nat \ y = r + 1" + proof - + have s5_1: "\\E x: P(x)" + using s4_3 by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_5: "r \\in Nat \ y = Succ[r]" + proof - + have s5_1: "r \\in Nat" + using s4_4 by auto + have s5_2: "Succ[r] = r + 1" + using s5_1 nat_add_1 by fast + have s5_3: "y = r + 1" + using s4_4 by auto + show ?thesis + using s5_1 s5_2 s5_3 by auto + qed + have "y \\in {Succ[m]: m \\in Nat}" + using s4_5 setOfAll_eqI by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s3_1 s3_2 extension + by (auto simp: subset_def) + qed + have s2_4: "n \\in {0} \ n \\in {m + 1: m \\in Nat}" + using s2_1 s2_2 s2_3 by auto + have s2_5: "n \\in {0} \ {k + 1: k \\in Nat}" + using s2_4 by auto + } + from this show ?thesis + by auto + qed + have s1_2: "\\A n \\in {0} \ {k + 1: k \\in Nat}: n \\in Nat" + proof - + { + fix "n" :: "c" + assume s1_2_asm: "n \\in {0} \ {k + 1: k \\in Nat}" + have s2_1: "n = 0 ==> n \\in Nat" + proof - + assume s2_1_asm: "n = 0" + show "n \\in Nat" + using s2_1_asm zeroIsNat by auto + qed + have s2_2: "n \\in {k + 1: k \\in Nat} ==> n \\in Nat" + proof - + assume s2_2_asm: "n \\in {k + 1: k \\in Nat}" + have s3_1: "\\E k \\in Nat: n = k + 1" + using s2_2_asm by auto + have s3_2: "\\E k: k \\in Nat \ n = k + 1" + using s3_1 by auto + (* PICK k \in Nat: n = k + 1 *) + def P \ "\ x. x \\in Nat \ n = x + 1" + def k \ "CHOOSE x: P(x)" + have s3_3: "k \\in Nat \ n = k + 1" + proof - + have s4_1: "\\E x: P(x)" + using s3_2 by (auto simp: P_def) + have s4_2: "P(k)" + unfolding k_def + using s4_1 chooseI_ex[of P] by auto + show ?thesis + using s4_2 by (auto simp: P_def) + qed + have s3_4: "k + 1 \\in Nat" + using s3_3 oneIsNat nat_add_in_nat by auto + show "n \\in Nat" + using s3_3 s3_4 by auto + qed + have s2_3: "n \\in Nat" + using s1_2_asm s2_1 s2_2 by auto + } + from this show ?thesis + by auto + qed + show ?thesis + using s1_1 s1_2 extension + by (auto simp: subset_def) + qed + + +(* +THEOREM NatIsAdd0 == + Nat = {k + 0: k \in Nat} +PROOF +<1>1. Nat = {k: k \in Nat} + OBVIOUS +<1>2. ASSUME NEW k \in Nat + PROVE k + 0 = k + BY <1>2, nat_in_int, add_0 +<1>3. {k: k \in Nat} = + {k + 0: k \in Nat} + BY <1>2 +<1> QED + BY <1>1, <1>3 +*) +theorem NatIsAdd0: + "Nat = {k + 0: k \\in Nat}" + proof - + have s1_1: "Nat = {k: k \\in Nat}" + by auto + have s1_2: "{k: k \\in Nat} = {k + 0: k \\in Nat}" + using nat_in_int add_0 by auto + show ?thesis + using s1_1 s1_2 by auto + qed + + +(* +THEOREM SplitAddi == + ASSUME NEW i \in Nat + PROVE {k + i: k \in Nat} = + {i} \cup {k + Succ[i]: k \in Nat} +PROOF +<1>1. n \in {k + i: k \in Nat} + = \E k \in Nat: n = k + i + OBVIOUS +<1>2. @ = \E k \in {0} \cup {q + 1: q \in Nat}: + n = k + i + BY SplitNat0 +<1>3. @ = \/ \E k \in {0}: n = k + i + \/ \E k \in {q + 1: q \in Nat}: n = k + i + OBVIOUS +<1>4. @ = \/ n = 0 + i + \/ \E k: /\ k \in {q + 1: q \in Nat} + /\ n = k + i + OBVIOUS +<1>5. @ = \/ n = i + \/ \E k: /\ \E q \in Nat: k = q + 1 + /\ n = k + i + BY i \in Nat, nat_in_int, add_0_left +<1>6. @ = \/ n = i + \/ \E k: \E q \in Nat: /\ k = q + 1 + /\ n = k + i + OBVIOUS +<1>7. @ = \/ n = i + \/ \E q \in Nat: n = (q + 1) + i + OBVIOUS +<1>8. @ = \/ n \in {i} + \/ n \in {q + (i + 1): q \in Nat} + BY AddCommutative, AddAssociative, + oneIsNat, nat_in_int +<1>9. @ = \/ n \in {i} + \/ n \in {q + Succ[i]: q \in Nat} + BY nat_add_1 +<1>10. @ = n \in {i} \cup {q + Succ[i]: q \in Nat} + OBVIOUS +<1> QED + BY <1>1, <1>2, <1>3, <1>4, <1>5, + <1>6, <1>7, <1>8, <1>9, <1>10 +*) +theorem SplitAddi: + assumes idom: "i \\in Nat" + shows "{k + i: k \\in Nat} = + ({i} \ {k + Succ[i]: k \\in Nat})" + proof - + { + fix "n" :: "c" + have s1_1: "n \\in {k + i: k \\in Nat} + = (\\E k \\in Nat: n = k + i)" + by auto + also have s1_2: "... = ( + \\E k \\in ({0} \ {q + 1: q \\in Nat}): + n = k + i)" + using SplitNat0 by auto + also have s1_3: "... = ( + (\\E k \\in {0}: n = k + i) \ + (\\E k \\in {q + 1: q \\in Nat}: n = k + i))" + by auto + also have s1_4: "... = ( + (n = 0 + i) \ + (\\E k: (k \\in {q + 1: q \\in Nat}) \ n = k + i))" + by auto + also have s1_5: "... = ( + (n = i) \ + (\\E k: (\\E q \\in Nat: k = q + 1) \ n = k + i))" + using idom nat_in_int add_0_left by auto + also have s1_6: "... = ( + (n = i) \ + (\\E k: \\E q \\in Nat: k = q + 1 \ n = k + i))" + by auto + also have s1_7: "... = ( + (n = i) \ + (\\E q \\in Nat: n = (q + 1) + i))" + by auto + also have s1_8: "... = ( + (n = i) \ + (\\E q \\in Nat: n = q + (i + 1)))" + using AddCommutative_sequent AddAssociative_sequent + oneIsNat idom nat_in_int by auto + also have s1_9: "... = ( + (n = i) \ + (\\E q \\in Nat: n = q + Succ[i]))" + using idom nat_add_1[symmetric] by force + also have s1_10: "... = ( + (n \\in {i}) \ + (n \\in {q + Succ[i]: q \\in Nat}))" + by auto + also have s1_11: "... = ( + n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + by auto + finally have s1_12: " + (n \\in {k + i: k \\in Nat}) + = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" . + } + from this have s1_13: "\ n. (n \\in {k + i: k \\in Nat}) + = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + by auto + have s1_14: "\ n. (n \\in {k + i: k \\in Nat}) + ==> (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" + using s1_13 by auto + have s1_15: "\ n. (n \\in ({i} \ {k + Succ[i]: k \\in Nat})) + ==> (n \\in {k + i: k \\in Nat})" + using s1_13 by auto + show ?thesis + using s1_14 s1_15 + by (rule setEqualI, blast) + qed + + +(* +THEOREM LeqNatOpenInterval == + ASSUME NEW i \in Nat, + NEW j \in {k + Succ[i]: k \in Nat} + PROVE ~ (j <= i) +PROOF +<1>1. PICK k \in Nat: j = k + Succ[i] + OBVIOUS +<1>2. Succ[i] <= j + BY <1>1 DEF leq +<1>3. i < Succ[i] + BY i \in Nat, less_succ_nat +<1>4. i < j + <2>1. i \in Int + BY i \in Nat, nat_in_int + <2>2. j \in Int + BY <1>1, i \in Nat, succIsNat, nat_in_int, + addIsInt + <2>3. Succ[i] \in Int + BY i \in Nat, succIsNat, nat_in_int + <2> QED + BY <2>1, <2>2, <2>3, <1>3, <1>2, + leq_less_trans +<1>5. i \in Int + BY i \in Nat, nat_in_int +<1>6. j \in Int + BY <1>1, i \in Nat, succIsNat, nat_in_int, + addIsInt +<1> QED + BY <1>4, <1>5, <1>6, less_not_leq +*) +theorem LeqNatOpenInterval: + assumes idom: "i \\in Nat" and + jdom: "j \\in {k + Succ[i]: k \\in Nat}" + shows "~ (j <= i)" + proof - + def P \ "\ x. x \\in Nat \ j = x + Succ[i]" + def k \ "CHOOSE x: P(x)" + have s1_1: "k \\in Nat \ j = k + Succ[i]" + proof - + have s2_1: "\\E x: P(x)" + using jdom by (auto simp: P_def) + have s2_2: "P(k)" + unfolding k_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_2: "Succ[i] <= j" + proof - + have s2_1: "k + Succ[i] = j" + using s1_1 by auto + have s2_2: "Succ[i] + k = j" + using s2_1 s1_1 idom succIsNat nat_in_int + AddCommutative_sequent by auto + have s2_3: "k \\in Nat" + using s1_1 by auto + show ?thesis + unfolding leq_def + using s2_2 s2_3 by auto + qed + have s1_3: "i < Succ[i]" + using idom less_succ_nat by auto + have s1_4: "i < j" + proof - + have s2_1: "i \\in Int" + using idom nat_in_int by auto + have s2_2: "j \\in Int" + proof - + have s3_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s3_2: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s3_3: "j = k + Succ[i]" + using s1_1 by auto + show ?thesis + using s3_1 s3_2 s3_3 addIsInt by auto + qed + have s2_3: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + show ?thesis + using s2_1 s2_2 s2_3 s1_3 s1_2 + less_leq_trans by iprover + qed + have s1_5: "i \\in Int" + using idom nat_in_int by auto + have s1_6: "j \\in Int" + proof - + have s2_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s2_2: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s2_3: "j = k + Succ[i]" + using s1_1 by auto + show ?thesis + using s2_1 s2_2 s2_3 addIsInt by auto + qed + show ?thesis + using s1_4 s1_5 s1_6 less_not_leq by auto + qed + + +(* +negNatIsAdd0 == + negNat = {-.k + -.0: k \in Nat} +PROOF +<1>1. negNat = {-.k: k \in Nat} + BY DEF negNat +<1>2. ASSUME NEW k \in Nat + PROVE -.k + -.0 = -.k + <2>1. -.k + -.0 = -.k + 0 + BY neg0 + <2>2. @ = -.k + <3>1. -.k \in Int + BY <1>2, minus_nat_in_int + <3> QED + BY <3>1, add_0 + <2> QED + BY <2>1, <2>2 +<1>3. {-.k: k \in Nat} = + (-.k + -.0: k \in Nat} + BY <1>2 +<1> QED + BY <1>1, <1>3 +*) +theorem negNatIsAdd0: + "negNat = {-.k + -.0: k \\in Nat}" + proof - + have s1_1: "negNat = {-.k: k \\in Nat}" + unfolding negNat_def by auto + have s1_2: "\ k. k \\in Nat ==> -.k + -.0 = -.k" + proof - + fix "k" :: "c" + assume s1_2_asm: "k \\in Nat" + have s2_1: "-.k + -.0 = -.k + 0" + using neg0 by auto + also have s2_2: "... = -.k" + proof - + have s3_1: "-.k \\in Int" + using s1_2_asm minus_nat_in_int by auto + show ?thesis + using s3_1 add_0 by auto + qed + finally show "-.k + -.0 = -.k" . + qed + have s1_3: "{-.k: k \\in Nat} = + {-.k + -.0: k \\in Nat}" + using s1_2 by auto + show ?thesis + using s1_1 s1_3 by auto + qed + + +theorem negNatSplitAddi: + assumes idom: "i \\in Nat" + shows "{-.k + -.i: k \\in Nat} = + ({-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + proof - + have s1_1: "\ n. (n \\in {-.k + -.i: k \\in Nat}) + = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + proof - + fix "n" :: "c" + have s2_1: "n \\in {-.k + -.i: k \\in Nat} + = (\\E k \\in Nat: n = -.k + -.i)" + by auto + also have s2_2: "... = ( + \\E k \\in ({0} \ {q + 1: q \\in Nat}): + n = -.k + -.i)" + using SplitNat0 by auto + also have s2_3: "... = ( + (\\E k \\in {0}: n = -.k + -.i) \ + (\\E k \\in {q + 1: q \\in Nat}: n = -.k + -.i) + )" + by auto + also have s2_4: "... = ( + (n = -.0 + -.i) \ + (\\E k: (k \\in {q + 1: q \\in Nat}) \ (n = -.k + -.i)))" + by auto + also have s2_5: "... = ( + (n = -.i) \ + (\\E k: (\\E q \\in Nat: k = q + 1) \ n = -.k + -.i))" + using idom minus_nat_in_int add_0_left by auto + also have s2_6: "... = ( + (n = -.i) \ + (\\E k: \\E q \\in Nat: k = q + 1 \ (n = -.k + -.i)))" + by auto + also have s2_7: "... = ( + (n = -.i) \ + (\\E q \\in Nat: n = -.(q + 1) + -.i))" + by auto + also have s2_8: "... = ( + (n = -.i) \ + (\\E q \\in Nat: n = -.q + -.Succ[i]))" + proof - + have s3_1: "\\E q \\in Nat: n = -.(q + 1) + -.i ==> + \\E q \\in Nat: n = -.q + -.Succ[i]" + proof - + assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" + def P \ "\ x. x \\in Nat \ + n = -.(x + 1) + -.i" + def r \ "CHOOSE x: P(x)" + have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" + proof - + have s5_1: "\\E x: P(x)" + using s3_1_asm by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_2: "n = (-.r + -.1) + -.i" + using s4_1 oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_3: "n = -.r + (-.i + -.1)" + using s4_2 s4_1 idom oneIsNat nat_in_int minus_nat_in_int + AddAssociative_sequent AddCommutative_sequent by auto + have s4_4: "n = -.r + -.(i + 1)" + using s4_3 idom oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_5: "n = -.r + -.Succ[i]" + proof - + have s5_1: "Succ[i] = i + 1" + using idom nat_add_1 by fast + show ?thesis + using s5_1 s4_4 by auto + qed + show "\\E q \\in Nat: n = -.q + -.Succ[i]" + using s4_1 s4_5 by auto + qed + have s3_2: "\\E q \\in Nat: n = -.q + -.Succ[i] ==> + \\E q \\in Nat: n = -.(q + 1) + -.i" + proof - + assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" + def P \ "\ x. x \\in Nat \ + n = -.x + -.Succ[i]" + def r \ "CHOOSE x: P(x)" + have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" + proof - + have s5_1: "\\E x: P(x)" + using s3_2_asm by (auto simp: P_def) + have s5_2: "P(r)" + unfolding r_def + using s5_1 chooseI_ex[of P] by auto + show ?thesis + using s5_2 by (auto simp: P_def) + qed + have s4_2: "n = -.r + -.(i + 1)" + proof - + have s5_1: "Succ[i] = i + 1" + using idom nat_add_1 by fast + show ?thesis + using s4_1 s5_1 by auto + qed + have s4_3: "n = -.r + (-.i + -.1)" + using s4_2 idom oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + have s4_4: "n = (-.r + -.1) + -.i" + using s4_3 s4_1 idom oneIsNat minus_nat_in_int + AddAssociative_sequent AddCommutative_sequent + by auto + have s4_5: "n = -.(r + 1) + -.i" + using s4_4 s4_1 oneIsNat nat_in_int + MinusDistributesAdd_sequent by auto + show "\\E q \\in Nat: n = -.(q + 1) + -.i" + using s4_1 s4_5 by auto + qed + show ?thesis + using s3_1 s3_2 by auto + qed + also have s2_9: "... = ( + (n \\in {-.i}) \ + (n \\in {-.k + -.Succ[i]: k \\in Nat}))" + by auto + also have s2_10: "... = ( + n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + by auto + finally show "(n \\in {-.k + -.i: k \\in Nat}) + = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" . + qed + have s1_2: "\ n. (n \\in {-.k + -.i: k \\in Nat}) + ==> (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" + using s1_1 by auto + have s1_3: "\ n. (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat}) + ==> (n \\in {-.k + -.i: k \\in Nat})" + using s1_1 by auto + show ?thesis + using s1_2 s1_3 by (rule setEqualI, blast) + qed + + +theorem LeqnegNatOpenInterval: + assumes idom: "i \\in Nat" and + jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" + shows "~ (-.i <= j)" + proof - + def P \ "\ x. x \\in Nat \ j = -.x + -.Succ[i]" + def k \ "CHOOSE x: P(x)" + have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" + proof - + have s2_1: "\\E x: P(x)" + using jdom by (auto simp: P_def) + have s2_2: "P(k)" + unfolding k_def + using s2_1 chooseI_ex[of P] by auto + show ?thesis + using s2_2 by (auto simp: P_def) + qed + have s1_types: "j \\in Int \ -.i \\in Int \ + -.j \\in Int \ Succ[i] \\in Int \ + -.k \\in Int \ -.Succ[i] \\in Int \ + k \\in Int \ i \\in Int" + proof - + have s2_1: "k \\in Int" + using s1_1 nat_in_int by auto + have s2_2: "-.k \\in Int" + using s2_1 neg_int_eq_int by auto + have s2_3: "i \\in Int" + using idom nat_in_int by auto + have s2_4: "-.i \\in Int" + using s2_3 neg_int_eq_int by auto + have s2_5: "Succ[i] \\in Int" + using idom succIsNat nat_in_int by auto + have s2_6: "-.Succ[i] \\in Int" + using s2_5 neg_int_eq_int by auto + have s2_7: "j \\in Int" + proof - + have s2_1: "j = -.k + -.Succ[i]" + using s1_1 by auto + show ?thesis + using s2_1 s2_2 s2_6 addIsInt by auto + qed + have s2_8: "-.j \\in Int" + using s2_7 neg_int_eq_int by auto + show ?thesis + using s2_1 s2_2 s2_3 s2_4 s2_5 + s2_6 s2_7 s2_8 by auto + qed + have s1_2: "Succ[i] <= -.j" + proof - + have s2_1: "Succ[i] + k = -.j" + proof - + have s3_1: "j = -.k + -.Succ[i]" + using s1_1 by auto + have s3_2: "-.j + j = -.j + (-.k + -.Succ[i])" + using s3_1 by auto + have s3_3: "0 = -.j + (-.k + -.Succ[i])" + using s3_2 s1_types AddNegCancels_left by auto + have s3_4: "0 + Succ[i] = (-.j + (-.k + -.Succ[i])) + Succ[i]" + using s3_3 by auto + have s3_5: "0 + Succ[i] = ((-.j + -.k) + -.Succ[i]) + Succ[i]" + using s3_4 s1_types AddAssociative_sequent by auto + have s3_6: "0 + Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" + using s3_5 s1_types addIsInt AddAssociative_sequent by auto + have s3_7: "Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" + proof - + have s4_1: "Succ[i] \\in Int" + using s1_types by auto + have s4_2: "0 + Succ[i] = Succ[i]" + using s4_1 add_0_left by fast + show ?thesis + using s3_6 s4_2 by auto + qed + have s3_8: "Succ[i] = (-.j + -.k) + 0" + using s3_7 s1_types AddNegCancels_left by auto + have s3_9: "Succ[i] = -.j + -.k" + proof - + have s4_1: "-.j + -.k \\in Int" + using s1_types addIsInt by auto + show ?thesis + using s3_8 s4_1 add_0 by auto + qed + have s3_10: "Succ[i] + k = (-.j + -.k) + k" + using s3_9 by auto + have s3_11: "Succ[i] + k = -.j + (-.k + k)" + using s3_10 s1_types AddAssociative_sequent by auto + have s3_12: "Succ[i] + k = -.j + 0" + using s3_11 s1_types AddNegCancels_left by auto + show ?thesis + using s3_12 s1_types add_0 by auto + qed + have s2_2: "\\E r \\in Nat: Succ[i] + r = -.j" + using s1_1 s2_1 by auto + show ?thesis + unfolding leq_def + using s2_2 by auto + qed + have s1_3: "i < Succ[i]" + using idom less_succ_nat by fast + have s1_4: "i < -.j" + using s1_types s1_2 s1_3 less_leq_trans by auto + have s1_5: "j < -.i" + proof - + have s2_1: "-.-.j < -.i" + proof - + have s3_1: "i \\in Int" + using s1_types by auto + have s3_2: "-.j \\in Int" + using s1_types by auto + show ?thesis + using s3_1 s3_2 s1_4 minus_less by auto + qed + have s2_2: "-.-.j = j" + using s1_types minus_involutive by auto + show ?thesis + using s2_1 s2_2 by auto + qed + show ?thesis + using s1_5 s1_types less_not_leq by auto + qed + + +theorem split_interval_nat_negnat: + shows "{k \\in Int: a <= k \ k <= b} = + {k \\in negNat: a <= k \ k <= b} \ + {k \\in Nat: a <= k \ k <= b}" + using int_union_nat_negnat by auto + + +end From 797ba38b95b6c7bc33e0ac95fda7665b4a91ed05 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:14:12 +0200 Subject: [PATCH 069/167] API: add TLA+ syntax for bounded quantifiers --- isabelle/SetTheory.thy | 2 ++ 1 file changed, 2 insertions(+) diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index ca9e6bb4..8729d7e5 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -151,6 +151,8 @@ syntax "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) syntax (ASCII) (* Again, only single variable for bounded choice. *) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) From 3e6b8e9fec0e2e770b6924ce3d7326eb7a002f62 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:26:25 +0200 Subject: [PATCH 070/167] MAI: update theory `NewIntegers` to Isabelle2020 --- isabelle/NewIntegers.thy | 234 ++++++++++++++++++++++++++------------- 1 file changed, 156 insertions(+), 78 deletions(-) diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy index 6c7ae08c..103b7221 100644 --- a/isabelle/NewIntegers.thy +++ b/isabelle/NewIntegers.thy @@ -2190,7 +2190,7 @@ theorem nat_add_in_nat: using s2_1 s2_2 s2_3 by auto qed have s1_3: "\ k \ Nat: m + k \\in Nat" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ k. m + k \\in Nat"] by auto show ?thesis using s1_3 ndom spec by auto qed @@ -2213,8 +2213,9 @@ theorem nat_add_0: assume s2_1: "n \ 0" have s2_2: "\ k: k \\in Nat \ n = Succ[k]" using ndom s2_1 nat_0_succ by auto - def P \ "\ x. x \\in Nat \ n = Succ[x]" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ n = Succ[x]" + define r where "r \ CHOOSE x: P(x)" have s2_3: "r \\in Nat \ n = Succ[r]" proof - have s3_1: "\ x: P(x)" @@ -2468,7 +2469,7 @@ theorem nat_mult_in_nat: using s2_1 s2_2 by auto qed have s1_3: "\ k \ Nat: m * k \\in Nat" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ k. m * k \\in Nat"] by auto show ?thesis using s1_3 ndom spec by auto qed @@ -3443,7 +3444,7 @@ theorem AddCommutative: show ?thesis using s4_1 s4_2 natInduct[of "\ i. 0 + i = i + 0"] - by auto + by blast qed have s3_2: "\ j. \ j \ Nat; @@ -3584,11 +3585,15 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show "\ i \ Nat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 natInduct - by auto + using s4_1 s4_2 + natInduct[of "\ i. Succ[j] + i = i + Succ[j]"] + by auto qed show ?thesis - using s3_1 s3_2 natInduct by auto + using s3_1 s3_2 + natInduct[of "\ j. + \ i \ Nat: j + i = i + j"] + by auto qed have s2_2: "\ j \ Nat: \ i \ negNat: j + i = i + j" @@ -3663,7 +3668,9 @@ theorem AddCommutative: by auto qed show ?thesis - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ Nat; @@ -3844,11 +3851,16 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show "\ i \ negNat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 neg_nat_induction - by auto + using s4_1 s4_2 + neg_nat_induction[of + "\ i. Succ[j] + i = i + Succ[j]"] + by auto qed show ?thesis - using s3_1 s3_2 natInduct by auto + using s3_1 s3_2 + natInduct[of "\ j. + \ i \ negNat: j + i = i + j"] + by auto qed show ?thesis proof - @@ -3944,7 +3956,9 @@ theorem AddCommutative: using s5_1 s5_2 s5_3 s5_4 s5_5 by auto qed show ?thesis - using s4_1 s4_2 natInduct by auto + using s4_1 s4_2 + natInduct[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ negNat; @@ -4146,10 +4160,16 @@ theorem AddCommutative: qed show "\ i \ Nat: -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 natInduct by auto + using s4_1 s4_2 + natInduct[of "\ i. + -.Succ[-.j] + i = i + -.Succ[-.j]"] + by auto qed show ?thesis - using s3_1 s3_2 neg_nat_induction by auto + using s3_1 s3_2 + neg_nat_induction[of "\ j. + \ i \ Nat: j + i = i + j"] + by auto qed have s2_2: "\ j \ negNat: \ i \ negNat: j + i = i + j" @@ -4233,7 +4253,9 @@ theorem AddCommutative: using s5_1 s5_2 by auto qed show ?thesis - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. 0 + i = i + 0"] + by blast qed have s3_2: "\ j. \ j \ negNat; @@ -4492,10 +4514,16 @@ theorem AddCommutative: qed show "\ i \ negNat: -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 neg_nat_induction by auto + using s4_1 s4_2 + neg_nat_induction[of "\ i. + -.Succ[-.j] + i = i + -.Succ[-.j]"] + by auto qed show ?thesis - using s3_1 s3_2 neg_nat_induction by auto + using s3_1 s3_2 + neg_nat_induction[of "\ j. + \ i \ negNat: j + i = i + j"] + by auto qed show ?thesis proof - @@ -5125,7 +5153,10 @@ theorem iSuccRightDistributesAdd: using s3_3 s3_4 by blast qed show ?thesis - using s2_1 s2_2 neg_nat_induction by auto + using s2_1 s2_2 + neg_nat_induction[of "\ b. + iSucc[a + b] = a + iSucc[b]"] + by auto qed have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" using s1_1 s1_2 int_union_nat_negnat by auto @@ -5346,7 +5377,10 @@ theorem iPredRightDistributesAdd: finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . qed show ?thesis - using s1_1 s2_1 natInduct by auto + using s1_1 s2_1 + natInduct[of "\ b. + iPred[a + b] = a + iPred[b]"] + by auto qed have s1_3: "\ b \ negNat \ {0}: iPred[a + b] = a + iPred[b]" @@ -5718,10 +5752,14 @@ theorem AddAssociative: qed have s1_5: "\ c \ Nat: (a + b) + c = a + (b + c)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ c. + (a + b) + c = a + (b + c)"] + by auto have s1_6: "\ c \ negNat: (a + b) + c = a + (b + c)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ c. + (a + b) + c = a + (b + c)"] + by auto have "\ c \ Int: (a + b) + c = a + (b + c)" using s1_5 s1_6 int_union_nat_negnat by auto @@ -5750,7 +5788,7 @@ theorem AddLeftCommutativity: qed -theorems add_ac_int[algebra_simps] = +lemmas add_ac_int[algebra_simps] = AddCommutative_sequent AddAssociative_sequent AddLeftCommutativity add_0 add_0_left @@ -6192,9 +6230,13 @@ theorem AddNegCancels: finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . qed have s1_4: "\ i \ Nat: i + -.i = 0" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ i. + i + -.i = 0"] + by auto have s1_5: "\ i \ negNat: i + -.i = 0" - using s1_1 s1_3 neg_nat_induction by auto + using s1_1 s1_3 neg_nat_induction[of "\ i. + i + -.i = 0"] + by auto show ?thesis using s1_4 s1_5 int_union_nat_negnat by auto qed @@ -6726,9 +6768,13 @@ theorem MinusDistributesAdd: finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . qed have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ b. + -.(a + b) = (-.a) + (-.b)"] + by auto have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ b. + -.(a + b) = (-.a) + (-.b)"] + by auto have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" using s1_5 s1_6 int_union_nat_negnat by auto } @@ -7760,9 +7806,13 @@ theorem MultCommutative: using s3_1 s3_2 by auto qed have s2_4: "\ i \ Nat: i * 0 = 0 * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * 0 = 0 * i"] + by blast have s2_5: "\ i \ negNat: i * 0 = 0 * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * 0 = 0 * i"] + by blast show ?thesis using s2_4 s2_5 int_union_nat_negnat by auto qed @@ -8222,9 +8272,13 @@ theorem MultCommutative: using s3_1 s3_2 by auto qed have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * Succ[j] = Succ[j] * i"] + by auto have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * Succ[j] = Succ[j] * i"] + by auto show "\ i \ Int: i * Succ[j] = Succ[j] * i" using s2_4 s2_5 int_union_nat_negnat by auto qed @@ -8759,19 +8813,27 @@ theorem MultCommutative: qed have s2_4: "\ i \ Nat: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_2 natInduct by auto + using s2_1 s2_2 natInduct[of "\ i. + i * -.Succ[-.j] = -.Succ[-.j] * i"] + by auto have s2_5: "\ i \ negNat: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_3 neg_nat_induction by auto + using s2_1 s2_3 neg_nat_induction[of "\ i. + i * -.Succ[-.j] = -.Succ[-.j] * i"] + by auto show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" using s2_4 s2_5 int_union_nat_negnat by auto qed have s1_4: "\ j \ Nat: \ i \ Int: i * j = j * i" - using s1_1 s1_2 natInduct by auto + using s1_1 s1_2 natInduct[of "\ j. + \ i \ Int: i * j = j * i"] + by blast have s1_5: "\ j \ negNat: \ i \ Int: i * j = j * i" - using s1_1 s1_3 neg_nat_induction by auto + using s1_1 s1_3 neg_nat_induction[of "\ j. + \ i \ Int: i * j = j * i"] + by blast show ?thesis using s1_4 s1_5 int_union_nat_negnat by auto qed @@ -9637,10 +9699,14 @@ theorem MultLeftDistributesAdd_forall: qed have s1_5: "\ a \ Nat: a * (b + c) = a * b + a * c" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ a. + a * (b + c) = a * b + a * c"] + by auto have s1_6: "\ a \ negNat: a * (b + c) = a * b + a * c" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ a. + a * (b + c) = a * b + a * c"] + by auto have s1_7: "\ a \ Int: a * (b + c) = a * b + a * c" using s1_5 s1_6 int_union_nat_negnat by auto @@ -9675,7 +9741,7 @@ theorem MultRightDistributesAdd: qed -theorems mult_distributes[algebra_simps] = +lemmas mult_distributes[algebra_simps] = MultLeftDistributesAdd MultRightDistributesAdd @@ -9918,10 +9984,14 @@ theorem MinusCommutesRightMult: qed have s1_5: "\ b \ Nat: -.(a * b) = a * -.b" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ b. + -.(a * b) = a * -.b"] + by auto have s1_6: "\ b \ negNat: -.(a * b) = a * -.b" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ b. + -.(a * b) = a * -.b"] + by auto have s1_7: "\ b \ Int: -.(a * b) = a * -.b" using s1_5 s1_6 int_union_nat_negnat by auto @@ -10243,10 +10313,14 @@ theorem MultAssociative: qed have s1_5: "\ c \ Nat: (a * b) * c = a * (b * c)" - using s1_2 s1_3 natInduct by auto + using s1_2 s1_3 natInduct[of "\ c. + (a * b) * c = a * (b * c)"] + by auto have s1_6: "\ c \ negNat: (a * b) * c = a * (b * c)" - using s1_2 s1_4 neg_nat_induction by auto + using s1_2 s1_4 neg_nat_induction[of "\ c. + (a * b) * c = a * (b * c)"] + by auto have s1_7: "\ c \ Int: (a * b) * c = a * (b * c)" using s1_5 s1_6 int_union_nat_negnat by auto @@ -10288,8 +10362,8 @@ theorem LeqTransitive: shows "m <= k" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -10301,8 +10375,8 @@ theorem LeqTransitive: using s2_2 by (auto simp: P_def) qed (* PICK t \in Nat: n + t = k *) - def Q \ "\ x. x \\in Nat \ n + x = k" - def t \ "CHOOSE x: Q(x)" + define Q where "Q \ \ x. x \\in Nat \ n + x = k" + define t where "t \ CHOOSE x: Q(x)" have s1_2: "t \\in Nat \ n + t = k" proof - have s2_1: "\\E x: Q(x)" @@ -10356,8 +10430,8 @@ theorem leq_0: have s1_2: "0 <= n ==> n \\in Nat" proof - assume s1_2_leq: "0 <= n" - def P \ "\ x. x \\in Nat \ 0 + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ 0 + x = n" + define r where "r \ CHOOSE x: P(x)" have s2_1: "r \\in Nat \ 0 + r = n" proof - have s3_1: "\\E x: P(x)" @@ -10433,10 +10507,10 @@ theorem leq_both_eq: mn: "m <= n" and nm: "n <= m" shows "m = n" proof - - def P \ "\ x. x \\in Nat \ m + x = n" - def p \ "CHOOSE x: P(x)" - def Q \ "\ x. x \\in Nat \ n + x = m" - def q \ "CHOOSE x: Q(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define p where "p \ CHOOSE x: P(x)" + define Q where "Q \ \ x. x \\in Nat \ n + x = m" + define q where "q \ CHOOSE x: Q(x)" have s1_1: "p \\in Nat \ m + p = n" proof - have s2_1: "\ x: P(x)" @@ -10754,8 +10828,8 @@ theorem leq_mult_monotonic: kdom: "k \\in Nat" and mn: "m <= n" shows "k * m <= k * n" proof - - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_2: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -10890,8 +10964,8 @@ theorem leq_mult_monotonic_neg: shows "k * n <= k * m" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_2: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -11019,7 +11093,7 @@ theorem leq_add_monotonic_right: proof - have s1_1: "\ r: r \ Nat \ m + r = n" using mn by (auto simp: leq_def) - def r \ "CHOOSE r: r \ Nat \ m + r = n" + define r where "r \ CHOOSE r: r \ Nat \ m + r = n" have s1_2: "r \ Nat \ m + r = n" unfolding r_def by (rule chooseI_ex, rule s1_1) @@ -11231,8 +11305,8 @@ theorem leq_diff_add: kdom: "k \ Int" and nmk: "n - m <= k" shows "n <= k + m" proof - - def P \ "\ x. x \\in Nat \ (n - m) + x = k" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ (n - m) + x = k" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ (n - m) + r = k" proof - have s2_1: "\\E x: P(x)" @@ -11556,8 +11630,8 @@ theorem less_is_leq_plus_one: shows "m + 1 <= n" proof - (* PICK r \in Nat: m + r = n *) - def P \ "\ x. x \\in Nat \ m + x = n" - def r \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ m + x = n" + define r where "r \ CHOOSE x: P(x)" have s1_1: "r \\in Nat \ m + r = n" proof - have s2_1: "\\E x: P(x)" @@ -11569,8 +11643,8 @@ theorem less_is_leq_plus_one: using s2_2 by (auto simp: P_def) qed (* PICK k \in Nat: r = Succ[k] *) - def Q \ "\ x. x \\in Nat \ r = Succ[x]" - def k \ "CHOOSE x: Q(x)" + define Q where "Q \ \ x. x \\in Nat \ r = Succ[x]" + define k where "k \ CHOOSE x: Q(x)" have s1_2: "k \\in Nat \ r = Succ[k]" proof - have s2_1: "r = 0 \ @@ -11888,8 +11962,9 @@ theorem SplitNat0: have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" using s4_2 by (auto simp: bEx_def) (* PICK r \in Nat: y = Succ[r] *) - def P \ "\ x. x \\in Nat \ y = Succ[x]" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ y = Succ[x]" + define r where "r \ CHOOSE x: P(x)" have s4_4: "r \\in Nat \ y = Succ[r]" proof - have s5_1: "\\E x: P(x)" @@ -11928,8 +12003,9 @@ theorem SplitNat0: have s4_3: "\\E m: m \\in Nat \ y = m + 1" using s4_2 by (auto simp: bEx_def) (* PICK r \in Nat: y = r + 1 *) - def P \ "\ x. x \\in Nat \ y = x + 1" - def r \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ y = x + 1" + define r where "r \ CHOOSE x: P(x)" have s4_4: "r \\in Nat \ y = r + 1" proof - have s5_1: "\\E x: P(x)" @@ -11988,8 +12064,9 @@ theorem SplitNat0: have s3_2: "\\E k: k \\in Nat \ n = k + 1" using s3_1 by auto (* PICK k \in Nat: n = k + 1 *) - def P \ "\ x. x \\in Nat \ n = x + 1" - def k \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ n = x + 1" + define k where "k \ CHOOSE x: P(x)" have s3_3: "k \\in Nat \ n = k + 1" proof - have s4_1: "\\E x: P(x)" @@ -12013,7 +12090,7 @@ theorem SplitNat0: qed show ?thesis using s1_1 s1_2 extension - by (auto simp: subset_def) + by blast qed @@ -12192,8 +12269,8 @@ theorem LeqNatOpenInterval: jdom: "j \\in {k + Succ[i]: k \\in Nat}" shows "~ (j <= i)" proof - - def P \ "\ x. x \\in Nat \ j = x + Succ[i]" - def k \ "CHOOSE x: P(x)" + define P where "P \ \ x. x \\in Nat \ j = x + Succ[i]" + define k where "k \ CHOOSE x: P(x)" have s1_1: "k \\in Nat \ j = k + Succ[i]" proof - have s2_1: "\\E x: P(x)" @@ -12354,9 +12431,9 @@ theorem negNatSplitAddi: \\E q \\in Nat: n = -.q + -.Succ[i]" proof - assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" - def P \ "\ x. x \\in Nat \ + define P where "P \ \ x. x \\in Nat \ n = -.(x + 1) + -.i" - def r \ "CHOOSE x: P(x)" + define r where "r \ CHOOSE x: P(x)" have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" proof - have s5_1: "\\E x: P(x)" @@ -12390,9 +12467,9 @@ theorem negNatSplitAddi: \\E q \\in Nat: n = -.(q + 1) + -.i" proof - assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" - def P \ "\ x. x \\in Nat \ + define P where "P \ \ x. x \\in Nat \ n = -.x + -.Succ[i]" - def r \ "CHOOSE x: P(x)" + define r where "r \ CHOOSE x: P(x)" have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" proof - have s5_1: "\\E x: P(x)" @@ -12452,8 +12529,9 @@ theorem LeqnegNatOpenInterval: jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" shows "~ (-.i <= j)" proof - - def P \ "\ x. x \\in Nat \ j = -.x + -.Succ[i]" - def k \ "CHOOSE x: P(x)" + define P where "P \ + \ x. x \\in Nat \ j = -.x + -.Succ[i]" + define k where "k \ CHOOSE x: P(x)" have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" proof - have s2_1: "\\E x: P(x)" From 93a14b96ca24351b5a8f8245e5fc2fee1c801082 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:35:26 +0200 Subject: [PATCH 071/167] MAI: update sectioning and text of theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 84 ++++++++++++++++++++++----------------------- 1 file changed, 42 insertions(+), 42 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 4dfb7d6b..d23d908a 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:39:46 merz> *) -header {* Tuples and Relations in \tlaplus{} *} +section \ Tuples and Relations in \tlaplus{} \ theory Tuples imports NatOrderings begin -text {* +text \ We develop a theory of tuples and relations in \tlaplus{}. Tuples are functions whose domains are intervals of the form $1 .. n$, for some natural number $n$, and relations are sets of tuples. In particular, @@ -21,15 +21,15 @@ text {* when defining functions recursively, as we have a fixed point theorem on sets but not on functions.) We also introduce standard notions for binary relations, such as orderings, equivalence relations and so on. -*} +\ -subsection {* Sequences and Tuples *} +subsection \ Sequences and Tuples \ -text {* +text \ Tuples and sequences are the same mathematical objects in \tlaplus{}, so we give elementary definitions for sequences here. Further operations on sequences require arithmetic and will be introduced in a separate theory. -*} +\ definition Seq -- {* set of finite sequences with elements from $S$ *} where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" @@ -157,12 +157,12 @@ lemma isASeqInSeq: "isASeq(s) = (\S: s \ Seq(S))" by (auto elim: seqInSeqRange) -subsection {* Sequences via @{text emptySeq} and @{text Append} *} +subsection \ Sequences via @{text emptySeq} and @{text Append} \ -text {* +text \ Sequences can be built from the constructors @{text emptySeq} (written @{text "\\"}) and Append. -*} +\ definition emptySeq ("(<< >>)") where "<< >> \ [x \ 1 .. 0 \ {}]" @@ -328,10 +328,10 @@ lemma appendEqualIff [simp]: shows "(Append(s,e) = Append(t,f)) = (s = t \ e = f)" using appendD1[OF s t] appendD2[OF s t] by auto -text {* +text \ The following lemma gives a possible alternative definition of @{text Append}. -*} +\ lemma appendExtend: assumes "isASeq(s)" @@ -348,10 +348,10 @@ lemma imageAppend [simp]: unfolding appendExtend[OF s] using assms by (auto elim!: inNatIntervalE, force+) -text {* +text \ Inductive reasoning about sequences, based on @{term "\\"} and @{text Append}. -*} +\ lemma seqInduct [case_names empty append, induct set: Seq]: assumes s: "s \ Seq(S)" @@ -411,12 +411,12 @@ lemma seqCases [case_names Empty Append, cases set: Seq]: using assms by (auto dest: seqEmptyOrAppend) -subsection {* Enumerated sequences *} +subsection \ Enumerated sequences \ -text {* +text \ We introduce the conventional syntax @{text "\a,b,c\"} for tuples and enumerated sequences, based on the above constructors. -*} +\ nonterminal tpl @@ -451,12 +451,12 @@ lemma "\0,1\ \ \1,2,3\" by simp lemma "(\a,b\ = \c,d\) = (a=c \ b=d)" by simp ***) -text {* +text \ \tlaplus{} has a form of quantification over tuples written $\exists \langle x,y,z \rangle \in S: P(x,y,z)$. We cannot give a generic definition of this form for arbitrary tuples, but introduce input syntax for tuples of length up to $5$. -*} +\ syntax "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) @@ -488,9 +488,9 @@ translations "\\x,y,z,u\ \ S : P" \ "\x,y,z,u : \x,y,z,u\ \ S \ P" "\\x,y,z,u,v\ \ S : P" \ "\x,y,z,u,v : \x,y,z,u,v\ \ S \ P" -subsection {* Sets of finite functions *} +subsection \ Sets of finite functions \ -text {* +text \ We introduce notation such as @{text "[x: S, y: T]"} to designate the set of finite functions @{text f} with domain @{text "{x,y}"} (for constants @{text x}, @{text y}) such that @{text "f[x] \ S"} and @{text "f[y] \ T"}. @@ -502,7 +502,7 @@ text {* Internally, the set is represented as @{text "EnumFuncSet(\x,y\, \S,T\)"}, using appropriate translation functions between the internal and external representations. -*} +\ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : @@ -635,14 +635,14 @@ by auto ***) -subsection {* Set product *} +subsection \ Set product \ -text {* +text \ The cartesian product of two sets $A$ and $B$ is the set of pairs whose first component is in $A$ and whose second component is in $B$. We generalize the definition of products to an arbitrary number of sets: $Product(\langle A_1,\ldots,A_n \rangle) = A_1 \times\cdots\times A_n$. -*} +\ definition Product where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : @@ -707,7 +707,7 @@ lemma "\a\ \ Product(\A,B\)" by auto ***) -text {* Special case: binary product *} +text \ Special case: binary product \ definition prod :: "c \ c \ c" (infixr "\\X" 100) where @@ -813,7 +813,7 @@ unfolding prod_def Product_def using assms by auto -subsection {* Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} *} +subsection \ Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} \ definition setOfPairs :: "[c, [c,c]\c] \ c" where "setOfPairs(R, f) \ { f(p[1], p[2]) : p \ R }" @@ -874,7 +874,7 @@ lemma setOfPairsEqual: using assms by (auto, blast+) -subsection {* Basic notions about binary relations *} +subsection \ Basic notions about binary relations \ definition rel_domain :: "c => c" where "rel_domain(r) \ { p[1] : p \ r }" @@ -896,7 +896,7 @@ definition Id :: "c \ c" -- {* diagonal: identity over a set *} where "Id(A) \ { \x,x\ : x \ A }" -text {* Properties of relations *} +text \ Properties of relations \ definition reflexive -- {* reflexivity over a set *} where "reflexive(A,r) \ \x \ A: \x,x\ \ r" @@ -953,7 +953,7 @@ lemma equivalenceIsBool[intro!,simp]: "isBool(equivalence(A,r))" unfolding isBool_def by (rule boolifyEquivalence) -subsubsection {* Domain and Range *} +subsubsection \ Domain and Range \ lemma prod_in_dom_x_ran: assumes "r \ A \ B" and "p \ r" @@ -1016,7 +1016,7 @@ lemma in_ran_imp_in_B: "r \ A \ B \ p \ rel by force -subsubsection {* Converse relation *} +subsubsection \ Converse relation \ lemmas converseEqualI = setEqualI [where A = "r^-1", standard, intro!] @@ -1123,7 +1123,7 @@ next qed -subsubsection {* Identity relation over a set *} +subsubsection \ Identity relation over a set \ lemmas idEqualI = setEqualI [where A = "Id(S)", standard, intro!] @@ -1173,7 +1173,7 @@ lemma ran_Id [simp]: "rel_range(Id(A)) = A" unfolding rel_range_def Id_def by auto -subsubsection {* Composition of relations *} +subsubsection \ Composition of relations \ lemmas compEqualI = setEqualI [where A = "r \ s", standard, intro!] @@ -1330,9 +1330,9 @@ unfolding rel_comp_def proof auto qed -subsubsection {* Properties of relations *} +subsubsection \ Properties of relations \ -text {* Reflexivity *} +text \ Reflexivity \ lemma reflI [intro]: "(\x. x \ A \ \x,x\ \ r) \ reflexive(A,r)" unfolding reflexive_def by blast @@ -1344,7 +1344,7 @@ lemma reflexive_empty (*[simp]*): "reflexive({}, {})" by auto -text {* Symmetry *} +text \ Symmetry \ lemma symmetricI: "\ \x y. \x,y\ \ r \ \y,x\ \ r \ \ symmetric(r)" unfolding symmetric_def by blast @@ -1356,7 +1356,7 @@ lemma symmetric_Int: "\ symmetric(r); symmetric(s) \ \ Antisymmetry \ lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" @@ -1372,7 +1372,7 @@ lemma antisym_empty (*[simp]*): "antisymmetric({})" by blast -text {* Transitivity *} +text \ Transitivity \ lemma transitiveI [intro]: "(\x y z. \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r) \ transitive(r)" @@ -1388,7 +1388,7 @@ lemma transitive_iff_comp_subset: "transitive(r) = (r \ r \ r)" unfolding transitive_def rel_comp_def by (auto elim!: subsetD) -text {* Irreflexivity *} +text \ Irreflexivity \ lemma irreflexiveI [intro]: "\ \x. x \ A \ \x,x\ \ r \ \ irreflexive(A,r)" unfolding irreflexive_def by blast @@ -1397,7 +1397,7 @@ lemma irreflexiveE [dest]: "\ irreflexive(A,r); x \ A \ \ Equivalence Relations \ (**************** NOT USED ANYWHERE ************************************ definition @@ -1422,11 +1422,11 @@ abbreviation --{* Abbreviation for the common case where the relations are identical *} ***************************************************************************) -text{* @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} *} +text \ @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} \ (* from Suppes, Theorem 70 *) -text {* First half: ``only if'' part *} +text \ First half: ``only if'' part \ lemma sym_trans_comp_subset: assumes "r \ A \ A" and "symmetric(r)" and "transitive(r)" @@ -1460,7 +1460,7 @@ using eq sym_trans_comp_subset[OF r] refl_comp_subset[OF r] unfolding equivalence_def by (intro setEqual, simp+) -text {* Second half: ``if'' part, needs totality of relation $r$ *} +text \ Second half: ``if'' part, needs totality of relation $r$ \ lemma comp_equivI: assumes dom: "rel_domain(r) = A" and r: "r \ A \ A" and comp: "r^-1 \ r = r" From d208966d2d96e002eae20bb13170134b54299b6b Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:38:38 +0200 Subject: [PATCH 072/167] MAI: port comments `--` in theory `Tuples` to Isabelle2020 The `-- {* ... *}` have been changed to comments, not knowing another equivalent syntax in Isabelle2020. --- isabelle/Tuples.thy | 50 ++++++++++++++++++++++----------------------- 1 file changed, 25 insertions(+), 25 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index d23d908a..64cb6464 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -31,13 +31,13 @@ text \ on sequences require arithmetic and will be introduced in a separate theory. \ -definition Seq -- {* set of finite sequences with elements from $S$ *} +definition Seq (* -- \ set of finite sequences with elements from $S$ \ *) where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" -definition isASeq -- {* characteristic predicate for sequences or tuples *} +definition isASeq (* -- \ characteristic predicate for sequences or tuples \close> *) where "isASeq(s) \ isAFcn(s) \ (\n \ Nat : DOMAIN s = 1 .. n)" -definition Len -- {* length of a sequence *} +definition Len (* -- \ length of a sequence \ *) where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: @@ -80,7 +80,7 @@ lemma SeqIsAFcn (*[elim!]*): shows "isAFcn(s)" using assms by auto --- {* @{text "s \ Seq(S) \ isAFcn(s)"} *} +(* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn, standard] lemma LenInNat [simp]: @@ -88,7 +88,7 @@ lemma LenInNat [simp]: shows "Len(s) \ Nat" using assms by auto --- {* @{text "s \ Seq(S) \ Len(s) \ Nat"} *} +(* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat, standard] lemma DomainSeqLen [simp]: @@ -96,7 +96,7 @@ lemma DomainSeqLen [simp]: shows "DOMAIN s = 1 .. Len(s)" using assms by auto --- {* @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} *} +(* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] lemma seqEqualI: @@ -121,7 +121,7 @@ lemma SeqI [intro!]: shows "s \ Seq(S)" using assms by (auto simp: Seq_def) -lemma SeqI': -- {* closer to the definition but probably less useful *} +lemma SeqI': (* -- \ closer to the definition but probably less useful \ *) assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" using assms by (auto simp: Seq_def) @@ -179,7 +179,7 @@ where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" by (auto simp: emptySeq_def isASeq_def) --- {* @{text "isAFcn(\\)"} *} +(* -- \ @{text "isAFcn(\\)"} \ *) lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] lemma lenEmptySeq [simp]: "Len(\\) = 0" @@ -199,7 +199,7 @@ lemma appendIsASeq [simp,intro!]: using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) --- {* @{text "isASeq(s) \ isAFcn(Append(s,e))"} *} +(* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn, standard] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" @@ -213,11 +213,11 @@ lemma isEmptySeq [intro!]: "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" by auto --- {* immediate consequence of @{text isEmptySeq} *} +(* -- \ immediate consequence of @{text isEmptySeq} \ *) lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto --- {* Symmetric equation could be a useful rewrite rule (it is applied by TLC) *} +(* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn, standard] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" @@ -243,7 +243,7 @@ lemma lenAppend [simp]: shows "Len(Append(s,e)) = Succ[Len(s)]" using assms by (intro LenI, auto simp: Append_def) --- {* @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} *} +(* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend, standard] lemma appendElt [simp]: @@ -398,7 +398,7 @@ proof - with s show ?thesis unfolding Seq_def by auto qed --- {* example of an inductive proof about sequences *} +(* -- \ example of an inductive proof about sequences \ *) lemma seqEmptyOrAppend: assumes "s \ Seq(S)" shows "s = \\ \ (\s' \ Seq(S): \e \ S : s = Append(s',e))" @@ -508,7 +508,7 @@ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" -lemmas -- {* establish set equality for sets of enumerated functions *} +lemmas (* -- \ establish set equality for sets of enumerated functions \ *) setEqualI [where A = "EnumFuncSet(doms, rngs)", standard, intro!] setEqualI [where B = "EnumFuncSet(doms, rngs)", standard, intro!] @@ -835,7 +835,7 @@ lemma inSetOfPairsI [intro]: shows "a \ setOfPairs(R, e)" using assms by (auto simp: setOfPairs_def) -lemma inSetOfPairsE [elim!]: -- {* converse true only if $R$ is a relation *} +lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" @@ -885,20 +885,20 @@ where "rel_range(r) \ { p[2] : p \ r }" definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" -definition rel_comp :: "[c,c] => c" (infixr "\" 75) -- {* binary relation composition *} +definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" -definition Id :: "c \ c" -- {* diagonal: identity over a set *} +definition Id :: "c \ c" (* -- \ diagonal: identity over a set \ *) where "Id(A) \ { \x,x\ : x \ A }" text \ Properties of relations \ -definition reflexive -- {* reflexivity over a set *} +definition reflexive (* -- \ reflexivity over a set \ *) where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" @@ -907,7 +907,7 @@ unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" unfolding isBool_def by (rule boolifyReflexive) -definition symmetric -- {* symmetric relation *} +definition symmetric (* -- \ symmetric relation \ *) where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" @@ -916,7 +916,7 @@ unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" unfolding isBool_def by (rule boolifySymmetric) -definition antisymmetric -- {* antisymmetric relation *} +definition antisymmetric (* -- \ antisymmetric relation \ *) where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" @@ -925,7 +925,7 @@ unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" unfolding isBool_def by (rule boolifyAntisymmetric) -definition transitive -- {* transitivity predicate *} +definition transitive (* -- \ transitivity predicate \ *) where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" @@ -934,7 +934,7 @@ unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" unfolding isBool_def by (rule boolifyTransitive) -definition irreflexive -- {* irreflexivity predicate *} +definition irreflexive (* -- \ irreflexivity predicate \ *) where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" @@ -943,7 +943,7 @@ unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" unfolding isBool_def by (rule boolifyIrreflexive) -definition equivalence :: "[c,c] \ c" -- {* (partial) equivalence relation *} +definition equivalence :: "[c,c] \ c" (* -- \ (partial) equivalence relation \ *) where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" @@ -1043,7 +1043,7 @@ lemma converseE [elim]: (** consider [elim!] ?? **) assumes yx: "yx \ r^-1" and r: "r \ A \ B" and p: "\x y. yx = \y,x\ \ \x,y\ \ r \ P" shows "P" - -- {* More general than @{text converseD}, as it ``splits'' the member of the relation. *} + (* -- \ More general than @{text converseD}, as it ``splits'' the member of the relation. \ *) proof - from prodProj[OF subsetD[OF converseSubset[OF r] yx]] have 2: "yx = \yx[1], yx[2]\" . with yx have 3: "\yx[2], yx[1]\ \ r" @@ -1419,7 +1419,7 @@ abbreviation abbreviation RESPECTS2 ::"[c \ c \ c, c] \ c" (infixr "respects2 " 80) where "f respects2 r \ congruent2(r,r,f)" - --{* Abbreviation for the common case where the relations are identical *} + (* -- \ Abbreviation for the common case where the relations are identical \ *) ***************************************************************************) text \ @{term r} is an equivalence relation iff @{term "converse(r) \ r = r"} \ From 519db6e5153827a72bb9ca730b4b4eec758ea0d4 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:44:10 +0200 Subject: [PATCH 073/167] MAI: rm attribute `standard` from `lemmas` because in Isabelle2020 this attribute is not needed, and unsupported. This conclusion is based on the file `NEWS` in the Isabelle `hg` repository. --- isabelle/Tuples.thy | 40 ++++++++++++++++++++-------------------- 1 file changed, 20 insertions(+), 20 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 64cb6464..e939c1bb 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -81,7 +81,7 @@ lemma SeqIsAFcn (*[elim!]*): using assms by auto (* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) -lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn, standard] +lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" @@ -89,7 +89,7 @@ lemma LenInNat [simp]: using assms by auto (* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) -lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat, standard] +lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] lemma DomainSeqLen [simp]: assumes "isASeq(s)" @@ -97,7 +97,7 @@ lemma DomainSeqLen [simp]: using assms by auto (* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) -lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen, standard] +lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: assumes "isASeq(s)" and "isASeq(t)" @@ -200,7 +200,7 @@ using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) (* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) -lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn, standard] +lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" by (simp add: emptySeq_def) @@ -218,7 +218,7 @@ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto (* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) -lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn, standard] +lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" by auto @@ -244,28 +244,28 @@ lemma lenAppend [simp]: using assms by (intro LenI, auto simp: Append_def) (* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) -lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend, standard] +lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] lemma appendElt [simp]: assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Succ[Len(s)]" shows "Append(s,e)[k] = (IF k = Succ[Len(s)] THEN e ELSE s[k])" using assms by (auto simp: Append_def) -lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt, standard] +lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt] lemma appendElt1 (*[simp]*): assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Len(s)" shows "Append(s,e)[k] = s[k]" using assms by (auto simp: Append_def) -lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1, standard] +lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1] lemma appendElt2 (*[simp]*): assumes "isASeq(s)" shows "Append(s,e)[Succ[Len(s)]] = e" using assms by (auto simp: Append_def) -lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2, standard] +lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2] lemma isAppend [intro!]: assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. Succ[Len(s)]" and s: "isASeq(s)" @@ -280,7 +280,7 @@ next by (auto simp: Append_def) qed -lemmas isAppend' [intro!] = isAppend[symmetric, standard] +lemmas isAppend' [intro!] = isAppend[symmetric] lemma appendInSeq [simp]: assumes s: "s \ Seq(S)" and e: "e \ S" @@ -509,8 +509,8 @@ where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNI \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" lemmas (* -- \ establish set equality for sets of enumerated functions \ *) - setEqualI [where A = "EnumFuncSet(doms, rngs)", standard, intro!] - setEqualI [where B = "EnumFuncSet(doms, rngs)", standard, intro!] + setEqualI [where A = "EnumFuncSet(doms, rngs)", intro!] + setEqualI [where B = "EnumFuncSet(doms, rngs)", intro!] lemma EnumFuncSetI [intro!,simp]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" @@ -847,8 +847,8 @@ proof - qed lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)", standard,intro!] - setEqualI [where B = "setOfPairs(R,f)", standard,intro!] + setEqualI [where A = "setOfPairs(R,f)",intro!] + setEqualI [where B = "setOfPairs(R,f)",intro!] lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" @@ -1019,8 +1019,8 @@ by force subsubsection \ Converse relation \ lemmas converseEqualI = - setEqualI [where A = "r^-1", standard, intro!] - setEqualI [where B = "r^-1", standard, intro!] + setEqualI [where A = "r^-1", intro!] + setEqualI [where B = "r^-1", intro!] lemma converse_iff [iff]: assumes r: "r \ A \ B" @@ -1126,8 +1126,8 @@ qed subsubsection \ Identity relation over a set \ lemmas idEqualI = - setEqualI [where A = "Id(S)", standard, intro!] - setEqualI [where B = "Id(S)", standard, intro!] + setEqualI [where A = "Id(S)", intro!] + setEqualI [where B = "Id(S)", intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" unfolding Id_def by auto @@ -1176,8 +1176,8 @@ unfolding rel_range_def Id_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s", standard, intro!] - setEqualI [where B = "r \ s", standard, intro!] + setEqualI [where A = "r \ s", intro!] + setEqualI [where B = "r \ s", intro!] lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From c65f420058f20fae7f656aa4199c41e0ac29a4f8 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:44:47 +0200 Subject: [PATCH 074/167] MAI: update `def` to `define ... where` for Isabelle2020 in the theory `Tuples`. --- isabelle/Tuples.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index e939c1bb..e0884db3 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -369,8 +369,8 @@ proof - proof fix sn assume sn: "sn \ [1 .. Succ[n] \ S]" - def so \ "[k \ 1 .. n \ sn[k]]" - def lst \ "sn[Succ[n]]" + define so where "so \ [k \ 1 .. n \ sn[k]]" + define lst where "lst \ sn[Succ[n]]" have 1: "sn = Append(so, lst)" proof from sn show "isAFcn(sn)" by simp From 6691c323b4b3f5a12ca1367c2d60a940861e554b Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:46:10 +0200 Subject: [PATCH 075/167] MAI: update ML section delimiters in theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index e0884db3..5bd08f33 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -541,7 +541,7 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") -parse_ast_translation {* +parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. The order of elements is reversed. *) @@ -571,9 +571,9 @@ parse_ast_translation {* in [("@EnumFuncSet", enum_funcset_tr)] end -*} +\ -print_ast_translation {* +print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] | list_from_tuple (Ast.Appl[Ast.Constant "@tuple", tp]) = @@ -602,8 +602,7 @@ print_ast_translation {* in [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] end -*} - +\ (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" From 403a5fdde8cde983c7d94b6cc7a2e7eb51b846a5 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:47:13 +0200 Subject: [PATCH 076/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. --- isabelle/Tuples.thy | 2 ++ 1 file changed, 2 insertions(+) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5bd08f33..96c35178 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -541,6 +541,7 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") +(* parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -603,6 +604,7 @@ print_ast_translation \ [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] end \ +*) (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" From 84c010e614dd20cfd5ae2c15c1863710f262ecd0 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Wed, 23 Dec 2020 21:48:28 +0200 Subject: [PATCH 077/167] MAI: update `lemmas` of theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 96c35178..0ca96fe4 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -848,8 +848,8 @@ proof - qed lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)",intro!] - setEqualI [where B = "setOfPairs(R,f)",intro!] + setEqualI [where A = "setOfPairs(R,f)" for R f, intro!] + setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" @@ -1020,8 +1020,8 @@ by force subsubsection \ Converse relation \ lemmas converseEqualI = - setEqualI [where A = "r^-1", intro!] - setEqualI [where B = "r^-1", intro!] + setEqualI [where A = "r^-1" for r, intro!] + setEqualI [where B = "r^-1" for r, intro!] lemma converse_iff [iff]: assumes r: "r \ A \ B" @@ -1127,8 +1127,8 @@ qed subsubsection \ Identity relation over a set \ lemmas idEqualI = - setEqualI [where A = "Id(S)", intro!] - setEqualI [where B = "Id(S)", intro!] + setEqualI [where A = "Id(S)" for S, intro!] + setEqualI [where B = "Id(S)" for S, intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" unfolding Id_def by auto @@ -1177,8 +1177,8 @@ unfolding rel_range_def Id_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s", intro!] - setEqualI [where B = "r \ s", intro!] + setEqualI [where A = "r \ s" for r s, intro!] + setEqualI [where B = "r \ s" for r s, intro!] lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From a2881a3c76bd2148670336f7713986e3340ca1aa Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 25 Dec 2020 14:03:52 +0200 Subject: [PATCH 078/167] MAI: update theory `Tuples` to Isabelle2020 --- isabelle/Tuples.thy | 398 +++++++++++++++++++++++++++++++++++++------- 1 file changed, 336 insertions(+), 62 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 0ca96fe4..47d556ce 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -336,7 +336,28 @@ text \ lemma appendExtend: assumes "isASeq(s)" shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" -using assms by force + proof - + have s1_1: "Append(s, e) = + [k \ 1 .. Succ[Len(s)] \ + IF k = Succ[Len(s)] THEN e ELSE s[k]]" + unfolding Append_def by auto + have s1_2: "(Succ[Len(s)] :> e) = + [x \ {Succ[Len(s)]} \ e]" + unfolding oneArg_def by auto + define p where "p \ s @@ (Succ[Len(s)] :> e)" + have s1_3: "p = + [x \ (DOMAIN s) \ {Succ[Len(s)]} + \ IF x \ DOMAIN s THEN s[x] ELSE e]" + unfolding p_def extend_def using s1_2 by auto + have s1_4: "(DOMAIN s) \ {Succ[Len(s)]} + = 1 .. Succ[Len(s)]" + using assms by auto + have s1_5: "p = [x \ 1..Succ[Len(s)] \ + IF x \ DOMAIN s THEN s[x] ELSE e]" + using s1_3 s1_4 by auto + show ?thesis + using s1_1 s1_5 assms by (auto simp: p_def) + qed lemma imageEmptySeq [simp]: "Image(\\, A) = {}" by (simp add: emptySeq_def) @@ -725,6 +746,7 @@ by (auto simp add: prod_def) lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" +unfolding two_def using assms by (auto simp add: prod_def) lemma inProd': @@ -750,7 +772,7 @@ proof next show "\a,b\[1] = a" by simp next - show "\a,b\[2] = b" by simp + show "\a,b\[2] = b" unfolding two_def by simp qed lemma inProdE [elim]: @@ -829,12 +851,13 @@ translations lemma inSetOfPairsI_ex: assumes "\\x,y\ \ R : a = e(x,y)" shows "a \ { e(x,y) : \x,y\ \ R }" -using assms by (auto simp: setOfPairs_def) +using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsI [intro]: assumes "a = e(x,y)" and "\x,y\ \ R" shows "a \ setOfPairs(R, e)" -using assms by (auto simp: setOfPairs_def) +unfolding two_def +using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) assumes 1: "z \ setOfPairs(R, e)" @@ -853,27 +876,187 @@ lemmas setOfPairsEqualI = lemma setOfPairs_triv [simp]: assumes s: "R \ A \ B" - shows "{ \x,y\ : \x,y\ \ R } = R" -using assms by auto + shows "{ \x, y\ : \x, y\ \ R } = R" + proof - + define Q where "Q \ + { \x, y\ : \x, y\ \ R }" + have s1_1: "Q = {<>: p \ R}" + unfolding Q_def setOfPairs_def by auto + have s1_2: "\ p. p \ R ==> p \ A \ B" + using s by auto + have s1_3: "\ p. p \ R ==> p = <>" + proof - + fix "p" :: "c" + assume s1_3_asm: "p \\in R" + have s2_1: "p \ A \ B" + using s1_3_asm s by auto + show "p = <>" + using s2_1 prodProj by auto + qed + define e :: "[c] => c" where "e \ \ x. x" + define f :: "[c] => c" where "f \ \ x. <>" + have s1_4: "\ x. x \ R ==> e(x) = f(x)" + unfolding e_def f_def using s1_3 by auto + have s1_5: "{e(x): x \ R} = {f(y): y \ R}" + apply (rule setOfAll_cong[of "R" "R" "e" "f"]) + apply (simp) + by (rule s1_4) + show ?thesis + using s1_1 s1_5 by (auto simp: Q_def e_def f_def) + qed lemma setOfPairs_cong (*[cong]*): - assumes 1: "R = S" and 2: "S \ A \ B" and 3: "\x y. \x,y\ \ S \ e(x,y) = f(x,y)" - shows "{ e(x,y) : \x,y\ \ R } = { f(u,v) : \u,v\ \ S }" -using assms proof (auto) - fix u v - let ?p = "\u,v\" - assume uv: "?p \ S" - with 3 have "f(u,v) = e(?p[1], ?p[2])" by simp - with uv show "f(u,v) \ setOfPairs(S, e)" by auto -qed + assumes 1: "R = S" and + 2: "S \ A \ B" and + 3: "\x y. \x, y\ \ S + \ + e(x, y) = f(x, y)" + shows "{ e(x, y) : \x, y\ \ R } = + { f(u, v) : \u, v\ \ S }" + proof - + have s1_1: "{ e(x, y) : \x, y\ \ R } = + { e(x, y) : \x, y\ \ S }" + using 1 by auto + have s1_2: "{ e(x, y) : \x, y\ \ S } = + { e(p[1], p[2]) : p \ S }" + by (auto simp: setOfPairs_def) + have s1_3: "{ f(u, v) : \u, v\ \ S } = + { f(p[1], p[2]) : p \ S }" + by (auto simp: setOfPairs_def) + define P where "P \ { e(p[1], p[2]): p \ S }" + define Q where "Q \ { f(p[1], p[2]): p \ S }" + have s1_4: "P = Q" + proof - + have s2_1: "\ x. x \ P <=> x \ Q" + proof - + fix "x" :: "c" + have s3_1: "x \\in P ==> x \\in Q" + proof - + assume s3_1_asm: "x \\in P" + have s4_1: "x \ { e(p[1], p[2]): p \ S }" + using s3_1_asm by (auto simp: P_def) + have s4_2: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_1 setOfAll by auto + have s4_3: "\\E p: p \\in S \ x = f(p[1], p[2])" + using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) + have s4_4: "x \ { f(p[1], p[2]): p \ S }" + using s4_3 setOfAll by auto + show "x \\in Q" + unfolding Q_def using s4_4 by auto + qed + have s3_2: "x \\in Q ==> x \\in P" + proof - + assume s3_2_asm: "x \\in Q" + have s4_1: "x \ { f(p[1], p[2]): p \ S}" + using s3_2_asm by (auto simp: Q_def) + have s4_2: "\\E p: p \\in S \ x = f(p[1], p[2])" + using s4_1 setOfAll by auto + have s4_3: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) + have s4_4: "x \ { e(p[1], p[2]): p \ S }" + using s4_3 setOfAll by auto + show "x \\in P" + unfolding P_def using s4_4 by auto + qed + show "x \ P <=> x \ Q" + using s3_1 s3_2 by auto + qed + have s2_2: "\ x: x \ P <=> x \ Q" + using s2_1 allI by auto + show ?thesis + using s2_2 extension by auto + qed + show ?thesis + using s1_1 s1_2 s1_3 s1_4 + by (auto simp: P_def Q_def) + qed lemma setOfPairsEqual: - assumes 1: "\x y. \x,y\ \ S \ \\x',y'\ \ T : e(x,y) = f(x',y')" - and 2: "\x' y'. \x',y'\ \ T \ \\x,y\\S : f(x',y') = e(x,y)" - and 3: "S \ A \ B" and 4: "T \ C \ D" - shows "{ e(x,y) : \x,y\ \ S } = { f(x,y) : \x,y\ \ T }" -using assms by (auto, blast+) - + assumes 1: "\x y. + \x, y\ \ S + \ \\x', y'\ \ T : + e(x, y) = f(x', y')" + and 2: "\x' y'. + \x', y'\ \ T + \ \\x, y\ \ S : + f(x', y') = e(x, y)" + and 3: "S \ A \ B" and + 4: "T \ C \ D" + shows "{ e(x, y) : \x, y\ \ S } = + { f(x, y) : \x, y\ \ T }" + proof - + have s1_1: "{ e(x, y) : \x, y\ \ S } = + { e(p[1], p[2]): p \ S }" + by (auto simp: setOfPairs_def) + have s1_2: "{ f(x, y) : \x, y\ \ T } = + { f(p[1], p[2]): p \ T }" + by (auto simp: setOfPairs_def) + define P where "P \ { e(p[1], p[2]): p \ S }" + define Q where "Q \ { f(p[1], p[2]): p \ T }" + have s1_3: "P = Q" + proof - + have s2_1: "\ x. x \ P <=> x \ Q" + proof - + fix "x" :: "c" + have s3_1: "x \\in P ==> x \\in Q" + proof - + assume s3_1_asm: "x \\in P" + have s4_1: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s3_1_asm setOfAll by (auto simp: P_def) + have s4_2: "\\E p: p \\in S \ + \p[1], p[2]\ \ S \ + x = e(p[1], p[2])" + using s4_1 pairProj_in_rel 3 by auto + have s4_3: "\\E p: p \\in S \ + x = e(p[1], p[2]) \ + (\\u, v\ \ T: + e(p[1], p[2]) = f(u, v))" + using s4_2 1 by auto + have s4_4: "\\u, v\ \ T: + x = f(u, v)" + using s4_3 by auto + have s4_5: "\\E u, v: <> \\in T \ x = f(u, v)" + using s4_4 by auto + have s4_6: "\\E p: p \\in T \ x = f(p[1], p[2])" + using s4_5 4 pairProj_in_rel by (auto simp: two_def) + show "x \\in Q" + using s4_6 setOfAll by (auto simp: Q_def) + qed + have s3_2: "x \\in Q ==> x \\in P" + proof - + assume s3_2_asm: "x \\in Q" + have s4_1: "\\E p: p \\in T \ x = f(p[1], p[2])" + using s3_2_asm setOfAll by (auto simp: Q_def) + have s4_2: "\\E p: p \\in T \ + <> \\in T \ + x = f(p[1], p[2])" + using s4_1 pairProj_in_rel 4 by auto + have s4_3: "\\E p: p \\in T \ + x = f(p[1], p[2]) \ + (\\u, v\ \ S: + f(p[1], p[2]) = e(u, v))" + using s4_2 2 by auto + have s4_4: "\\u, v\ \ S: + x = e(u, v)" + using s4_3 by auto + have s4_5: "\\E u, v: <> \\in S \ x = e(u, v)" + using s4_4 by auto + have s4_6: "\\E p: p \\in S \ x = e(p[1], p[2])" + using s4_5 3 pairProj_in_rel by (auto simp: two_def) + show "x \\in P" + using s4_6 setOfAll by (auto simp: P_def) + qed + show "x \\in P <=> x \\in Q" + using s3_1 s3_2 by auto + qed + have s2_2: "\ x: x \ P <=> x \ Q" + using s2_1 allI by auto + show ?thesis + using s2_2 extension by auto + qed + show ?thesis + using s1_1 s1_2 s1_3 by (auto simp: P_def Q_def) + qed subsection \ Basic notions about binary relations \ @@ -986,7 +1169,7 @@ unfolding rel_range_def using pairProj_in_prod by auto lemma in_rel_rangeI [iff]: assumes "\x,y\ \ r" shows "y \ rel_range(r)" -unfolding rel_range_def using assms by auto +unfolding rel_range_def two_def using assms by auto lemma in_rel_rangeE [elim]: assumes y: "y \ rel_range(r)" and r: "r \ A \ B" and p: "\x. \x,y\ \ r \ P" @@ -1026,11 +1209,11 @@ lemmas converseEqualI = lemma converse_iff [iff]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" -using r prodProj by (auto simp: converse_def) +using r prodProj by (auto simp: converse_def two_def) lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" -unfolding converse_def by auto +unfolding converse_def two_def by auto lemma converseD [sym]: assumes r: "r \ A \ B" @@ -1057,10 +1240,52 @@ qed lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" -using assms prodProj by (auto elim!: converseE) + proof - + have s1_1: "\ x. x \\in A \ B ==> x = <>" + using prodProj by (auto simp: two_def) + have s1_2: "\ x. x \\in r ==> x \\in A \ B" + using s1_1 r by auto + have s1_3: "\ x. x \\in r ==> x = <>" + using s1_1 s1_2 by auto + show ?thesis + using s1_3 assms prodProj pairProj_in_rel + by (auto simp: converse_def two_def) + qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" -using prodProj by auto + proof - + have s1_1: "(A \ B)^-1 = + {<>: p \ A \ B}" + unfolding converse_def by auto + have s1_2: "B \ A = + {<>: p \ B \ A}" + using setOfPairs_triv prodProj by auto + have s1_3: "\ x. x \\in B \\X A ==> + \\E p \\in A \\X B: x = <>" + proof - + fix "x" :: "c" + assume s1_5_asm: "x \\in B \\X A" + have s2_1: "<> \\in A \\X B" + using s1_5_asm by (auto simp: two_def) + define q where "q \ <>" + have s2_2: "q \\in A \\X B \ q = <>" + using s2_1 by (auto simp: q_def) + have s2_3: "x = <>" + using s1_5_asm prodProj by auto + have s2_4: "q[1] = x[2] \ q[2] = x[1]" + using s2_2 by (auto simp: two_def) + have s2_5: "x = <>" + using s2_3 s2_4 by auto + show "\\E p \\in A \\X B: x = <>" + using s2_2 s2_5 by auto + qed + have s1_4: "\ p. p \\in A \\X B ==> + p[1] \\in A \ p[2] \\in B" + using pairProj_in_rel inProd prodProj + by (auto simp: two_def) + show ?thesis + using s1_1 s1_2 s1_3 s1_4 by auto + qed lemma converse_empty [simp]: "converse({}) = {}" by auto @@ -1106,23 +1331,44 @@ lemma transitive_converse [simp]: unfolding transitive_def by auto lemma symmetric_iff_converse_eq: - assumes r:"r \ A \ B" + assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" -proof auto - fix p - assume "symmetric(r)" and "p \ r^-1" - with r show "p \ r" by (auto elim!: converseE simp add: symmetric_def) -next - fix p - assume 1: "symmetric(r)" and 2: "p \ r" - from r 2 have 3: "p = \p[1],p[2]\" by (intro prodProj, auto) - with 1 2 have "\p[2],p[1]\ \ r" by (force simp add: symmetric_def) - with 3 show "p \ r^-1" by (auto dest: converseI) -next - assume "r^-1 = r" thus "symmetric(r)" - by (auto simp add: symmetric_def) -qed - + proof - + have s1_1: "symmetric(r) ==> r^-1 = r" + proof - + assume s1_1_asm: "symmetric(r)" + have s2_1: "\ p. p \\in r^-1 ==> p \\in r" + proof - + fix "p" :: "c" + assume s2_1_asm: "p \\in r^-1" + show "p \\in r" + using r s1_1_asm s2_1_asm + by (auto elim!: converseE simp: symmetric_def) + qed + have s2_2: "\ p. p \\in r ==> p \\in r^-1" + proof - + fix "p" :: "c" + assume s2_2_asm: "p \\in r" + have s3_1: "p = <>" + using s2_2_asm r by (intro prodProj, auto) + have s3_2: "<> \\in r" + using s1_1_asm s2_2_asm s3_1 by (auto simp: symmetric_def) + show "p \\in r^-1" + using s3_1 s3_2 by (auto dest: converseI) + qed + show "r^-1 = r" + using s2_1 s2_2 setEqualI[of "r^-1" "r"] + by blast + qed + have s1_2: "r^-1 = r ==> symmetric(r)" + proof - + assume s1_2_asm: "r^-1 = r" + show "symmetric(r)" + using s1_2_asm by (auto simp: symmetric_def) + qed + show ?thesis + using s1_1 s1_2 by auto + qed subsubsection \ Identity relation over a set \ @@ -1165,13 +1411,14 @@ lemma Id_eqI: "a = b \ a \ A \ \a,b\ by simp lemma converse_Id [simp]: "Id(A)^-1 = Id(A)" -by auto +by (auto simp: Id_def two_def) + lemma dom_Id [simp]: "rel_domain(Id(A)) = A" unfolding rel_domain_def Id_def by auto lemma ran_Id [simp]: "rel_range(Id(A)) = A" -unfolding rel_range_def Id_def by auto +unfolding rel_range_def Id_def two_def by auto subsubsection \ Composition of relations \ @@ -1208,7 +1455,7 @@ using assms by force lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "((r \ s)^-1) = (s^-1 \ r^-1)" (is "?lhs = ?rhs") -proof +proof (rule setEqualI) fix x assume x: "x \ ?lhs" from s r have "r \ s \ A \ C" by (rule rel_comp_in_prod) @@ -1231,26 +1478,53 @@ qed lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" -using r proof auto - fix p - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) - with 1 show "p \ R \ Id(B)" by simp -qed + using r + proof - + have s1_1: "\ p. p \\in R ==> p \\in R \ Id(B)" + proof - + fix "p" :: "c" + assume p: "p \ R" + with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) + from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) + with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) + with 1 show "p \ R \ Id(B)" by simp + qed + have s1_2: "\ p. p \\in R \ Id(B) ==> p \\in R" + proof - + fix "p" :: "c" + assume s1_2_asm: "p \\in R \ Id(B)" + show "p \\in R" + using r s1_2_asm by auto + qed + show ?thesis + using s1_1 s1_2 setEqualI[of "R \ Id(B)" "R"] by blast + qed + lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" -using r proof auto - fix p - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) - with 1 show "p \ Id(B) \ R" by simp -qed + using r + proof - + have s1_1: "\ p. p \\in R ==> p \\in Id(B) \ R" + proof - + fix "p" :: "c" + assume p: "p \ R" + with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) + from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) + with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) + with 1 show "p \ Id(B) \ R" by simp + qed + have s1_2: "\ p. p \\in Id(B) \ R ==> p \\in R" + proof - + fix "p" :: "c" + assume s1_2_asm: "p \\in Id(B) \ R" + show "p \\in R" + using r s1_2_asm by auto + qed + show ?thesis + using s1_1 s1_2 setEqualI[of "Id(B) \ R" "R"] by blast + qed lemma rel_comp_empty1 [simp]: "{} \ R = {}" unfolding rel_comp_def by auto @@ -1261,7 +1535,7 @@ unfolding rel_comp_def by auto lemma comp_assoc: assumes t: "T \ A \ B" and s: "S \ B \ C" and r: "R \ C \ D" shows "(R \ S) \ T = R \ (S \ T)" -proof +proof (rule setEqualI) fix p assume p: "p \ (R \ S) \ T" from r s have "R \ S \ B \ D" by simp From 9dc91a83ddf9e6ff3692411a7cfa796bbbef8ed0 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 25 Dec 2020 14:04:28 +0200 Subject: [PATCH 079/167] STY: shorter lines --- isabelle/Tuples.thy | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 47d556ce..7c7bfc97 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1070,8 +1070,13 @@ definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) -where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : - \x,z : p = \x,z\ \ (\y: \x,y\ \ s \ \y,z\ \ r) }" +where "r \ s \ + { p \ rel_domain(s) \ rel_range(r) : + \ x, z : + p = \x, z\ \ + (\ y: + \x, y\ \ s \ + \y, z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" From 8bbc5b8cb87ea8014c2500193d94591a2997e2bf Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 13:32:13 +0200 Subject: [PATCH 080/167] DRAFT: comment attribute `intro!` in `lemmas compEqualI` in theory `Tuples`, because it gives rise to non-terminating behavior in Isabelle2020. --- isabelle/Tuples.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 7c7bfc97..d644e8b5 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1429,8 +1429,8 @@ unfolding rel_range_def Id_def two_def by auto subsubsection \ Composition of relations \ lemmas compEqualI = - setEqualI [where A = "r \ s" for r s, intro!] - setEqualI [where B = "r \ s" for r s, intro!] + setEqualI [where A = "r \ s" for r s] (* , intro!] *) + setEqualI [where B = "r \ s" for r s] (* , intro!] *) lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" From 7b66160d1f0d808c451975a41a8ffe8d27b0acdb Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 13:51:04 +0200 Subject: [PATCH 081/167] MAI: update theory `NatDivision` after change of numerals to definitions --- isabelle/NatDivision.thy | 19 ++++++++++++------- 1 file changed, 12 insertions(+), 7 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index fcc8a7f2..1d2869aa 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -256,7 +256,7 @@ lemma divmodNatPairEx: proof - from assms obtain q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" by (rule divmod_rel_ex) - thus ?thesis by force + thus ?thesis unfolding two_def by force qed lemma divmodNatInNatNat: @@ -279,6 +279,7 @@ proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) fix z assume "z \ Nat \ Nat" and "divmod_rel(m,n,z[1],z[2])" with m n q r h show "z = \q,r\" + unfolding two_def by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) qed @@ -301,18 +302,20 @@ using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) lemma modIsNat [iff]: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "m mod n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def) +using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def two_def) lemma divmodNat_div_mod: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" shows "divmodNat(m,n) = \m div n, m mod n\" -unfolding div_nat_def[OF m n] mod_nat_def[OF m n] using divmodNatInNatNat[OF assms] +unfolding div_nat_def[OF m n] mod_nat_def[OF m n] two_def +using divmodNatInNatNat[OF assms] by force lemma divmod_rel_div_mod_nat: assumes "m \ Nat" and "n \ Nat" and "0 < n" shows "divmod_rel(m, n, m div n, m mod n)" -using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] by simp +using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] +by (auto simp: two_def) lemma div_nat_unique: assumes h: "divmod_rel(m,n,q,r)" @@ -324,7 +327,7 @@ lemma mod_nat_unique: assumes h: "divmod_rel(m,n,q,r)" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" shows "m mod n = r" -unfolding mod_nat_def[OF m n] using divmodNat_unique[OF assms] by simp +unfolding mod_nat_def[OF m n] two_def using divmodNat_unique[OF assms] by simp lemma mod_nat_less_divisor: assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" @@ -342,10 +345,12 @@ proof - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" by (simp add: divmodNat_divmod_rel) from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - with 1 2 less n have "?dm[1] * n < n" by (auto simp: divmod_rel_def elim!: add_lessD1) + with 1 2 less n have "?dm[1] * n < n" + unfolding two_def + by (auto simp: divmod_rel_def elim!: add_lessD1) with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) - with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def) + with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) with 3 prodProj[OF 2] show ?thesis by simp qed From 3b44804fd2e5ed74e4eaba04a93b9917386f7d33 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 16:56:43 +0200 Subject: [PATCH 082/167] MAI: update theory `NatDivision` to Isabelle2020 --- isabelle/NatDivision.thy | 38 +++++++++++++++++++++++++++++--------- 1 file changed, 29 insertions(+), 9 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 1d2869aa..94eefffd 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -221,12 +221,18 @@ proof - assume "\(x \ x')" with nat have "x' < x" by (simp add: nat_not_leq[simplified]) with nat obtain k where k: "k \ Nat" "x = Succ[x'+k]" - by (auto simp: less_iff_Succ_add) + using less_iff_Succ_add by auto with eq nat n have "x'*n + (k*n + n + y) = x'*n + y'" by (simp add: add_mult_distrib_right_nat add_assoc_nat) with nat k n have "k*n + n + y = y'" by simp - with less k n nat have "(k*n + y) + n < n" by (simp add: add_ac_nat) - with k n nat show "FALSE" by simp + with less k n nat have s3_1: "(k*n + y) + n < n" by (simp add: add_ac_nat) + define j where "j \ k * n + y" + have s3_2: "j + n < n" + unfolding j_def using s3_1 by auto + have s3_3: "j \\in Nat" + unfolding j_def using multIsNat addIsNat k n nat by auto + show "FALSE" + using s3_2 s3_3 n not_add_less2[of "n" "j"] by auto qed } from this[OF q r q' r'] this[OF q' r' q r] q q' mqr mqr' rn rn' @@ -345,9 +351,19 @@ proof - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" by (simp add: divmodNat_divmod_rel) from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - with 1 2 less n have "?dm[1] * n < n" - unfolding two_def - by (auto simp: divmod_rel_def elim!: add_lessD1) + have "?dm[1] * n < n" + proof - + have s2_1: "m = ?dm[1] * n + ?dm[2]" + using 1 by (auto simp: divmod_rel_def) + have s2_2: "?dm[1] * n + ?dm[2] < n" + using less s2_1 by auto + have s2_3: "?dm[1] \\in Nat" + using 2 by auto + have s2_4: "?dm[2] \\in Nat" + using 2 by (auto simp: two_def) + show ?thesis + using s2_2 s2_3 s2_4 multIsNat n add_lessD1 by auto + qed with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) @@ -364,8 +380,11 @@ proof - have 2: "m div n \ 0" proof assume "m div n = 0" - with 1 m n pos have "m < n" by (auto simp: divmod_rel_def) - with geq m n show "FALSE" by (auto simp: nat_less_leq_not_leq) + with 1 m n pos have s2_1: "m < n" by (auto simp: divmod_rel_def) + have s2_2: "\ (n \ m)" + using s2_1 n m nat_not_leq by auto + show "FALSE" + using geq s2_2 by auto qed with m n pos obtain k where k1: "k \ Nat" and k2: "m div n = Succ[k]" using not0_implies_Suc[of "m div n"] by auto @@ -621,7 +640,8 @@ lemma mod_div_nat_trivial [simp]: proof - from assms have "m div n + (m mod n) div n = (m mod n + (m div n) * n) div n" - by (simp add: mod_nat_less_divisor) + using div_nat_mult_self1[of "m mod n" "m div n" "n"] + by (auto simp: mod_nat_less_divisor) also from assms have "... = m div n + 0" by simp finally show ?thesis using assms by simp From e71c6d17b71c9fd2ec2646a5a936f21a5eb0a012 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 16:56:58 +0200 Subject: [PATCH 083/167] STY: shorter lines --- isabelle/NatDivision.thy | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy index 94eefffd..cdfa1eb4 100644 --- a/isabelle/NatDivision.thy +++ b/isabelle/NatDivision.thy @@ -165,8 +165,11 @@ text \ \ definition divmod_rel where - "divmod_rel(m,n,q,r) \ m = q * n + r - \ ((0 < n \ 0 \ r \ r < n) \ (n < 0 \ r \ 0 \ n < r))" + "divmod_rel(m,n,q,r) \ + m = q * n + r + \ + ((0 < n \ 0 \ r \ r < n) \ + (n < 0 \ r \ 0 \ n < r))" text \ @{const divmod_rel} is total if $n$ is non-zero. \ From 6e3be13c760c80e59170400a38224549a62f27c8 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:03:24 +0200 Subject: [PATCH 084/167] MAI: update delimiters in theory `CaseExpressions` to Isabelle2020 --- isabelle/CaseExpressions.thy | 22 +++++++++++----------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index c9fc0a31..d7fcf341 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -6,13 +6,13 @@ Time-stamp: <2011-10-11 17:40:22 merz> *) -header {* Case expressions *} +section \ Case expressions \ theory CaseExpressions imports Tuples begin -text {* +text \ A \textsc{case} expression in \tlaplus{} has the form \[ \textsc{case } p_1 \rightarrow e_1 @@ -32,7 +32,7 @@ text {* guard of the \textsc{other}-branch, when present, is the conjunction of the negated guards of all other branches, so every guard appears twice (and will be simplified twice) in a @{text CaseOther} expression. -*} +\ definition CaseArm -- {* preliminary construct to convert case arm into set *} where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" @@ -64,7 +64,7 @@ syntax (HTML output) "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") -parse_ast_translation {* +parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. The order of elements is reversed. *) @@ -99,10 +99,10 @@ parse_ast_translation {* in [("_case_syntax", case_syntax_tr)] end -*} +\ (** debugging ** -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" oops @@ -113,10 +113,10 @@ oops lemma "CASE OTHER \ g" (* degenerated case *) oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) **) -print_ast_translation {* +print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] | list_from_tuple (Ast.Appl[Ast.Constant "@tuple", tp]) = @@ -173,10 +173,10 @@ print_ast_translation {* [(@{const_syntax "Case"}, case_syntax_tr'), (@{const_syntax "CaseOther"}, caseother_tr')] end -*} +\ (** debugging ** -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" oops @@ -187,7 +187,7 @@ oops lemma "CASE OTHER \ g" (* degenerated case *) oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) **) lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def From a781b250d55fde2d2f6cc270ed497b3b2195528f Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:04:05 +0200 Subject: [PATCH 085/167] MAI: update `--` comments in theory `CaseExpressions` to Isabelle2020 --- isabelle/CaseExpressions.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index d7fcf341..bcb75e57 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -34,7 +34,7 @@ text \ twice (and will be simplified twice) in a @{text CaseOther} expression. \ -definition CaseArm -- {* preliminary construct to convert case arm into set *} +definition CaseArm (* -- preliminary construct to convert case arm into set *) where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" definition Case where From e98574f5b676f07ea14b7a28b55d0082312a71d1 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:06:12 +0200 Subject: [PATCH 086/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. Mark them as "TODO". --- isabelle/CaseExpressions.thy | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index bcb75e57..09c2ca77 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -64,6 +64,7 @@ syntax (HTML output) "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") +(* TODO parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -100,6 +101,7 @@ parse_ast_translation \ [("_case_syntax", case_syntax_tr)] end \ +*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) @@ -116,6 +118,7 @@ oops (*ML \ reset Syntax.trace_ast; \*) **) +(* TODO print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] @@ -174,6 +177,7 @@ print_ast_translation \ (@{const_syntax "CaseOther"}, caseother_tr')] end \ +*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) From 9934f539c5f55b66298abd783d4bdb218f444076 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:15:13 +0200 Subject: [PATCH 087/167] MAI: update delimiters in theory `Strings` to Isabelle2020 --- isabelle/Strings.thy | 36 ++++++++++++++++++------------------ 1 file changed, 18 insertions(+), 18 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index e951d12c..b274d2cf 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -6,18 +6,18 @@ Time-stamp: <2011-10-11 17:40:15 merz> *) -header {* Characters and strings *} +section \ Characters and strings \ theory Strings imports Tuples begin -subsection {* Characters *} +subsection \ Characters \ -text {* +text \ Characters are represented as pairs of hexadecimal digits (also called \emph{nibbles}). -*} +\ definition Nibble where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" @@ -40,7 +40,7 @@ lemma isChar [simp]: "(c \ Char) = (\a,b \ Nibble : c = char(a,b unfolding Char_def by auto -subsection {* Strings *} +subsection \ Strings \ definition String where "String \ Seq(Char)" @@ -49,7 +49,7 @@ syntax "_Char" :: "xstr \ c" ("CHAR _") "_String" :: "xstr \ c" ("_") -text {* +text \ The following parse and print translations convert between the internal and external representations of strings. Strings are written using two single quotes in Isabelle, such as \verb|''abc''|. Note that the @@ -57,9 +57,9 @@ text {* printed as @{term "\\"}. Single characters are printed in the form \verb|CHAR ''a''|: Isabelle doesn't provide single characters in its lexicon. -*} +\ -parse_ast_translation {* +parse_ast_translation \ let (* convert an ML integer to a nibble *) fun mkNibble n = @@ -93,11 +93,11 @@ parse_ast_translation {* in [("_Char", char_ast_tr), ("_String", string_ast_tr)] end -*} +\ (** debug **) -(*ML {* set Syntax.trace_ast; *}*) +(*ML \ set Syntax.trace_ast; \*) lemma "''a''" oops @@ -105,10 +105,10 @@ oops lemma "CHAR ''a''" oops -(*ML {* reset Syntax.trace_ast; *}*) +(*ML \ reset Syntax.trace_ast; \*) (**) -print_ast_translation {* +print_ast_translation \ let (* convert a nibble to an ML integer -- because translation macros have already been applied, we see constants "0" through "15", not Succ[...] terms! *) @@ -167,11 +167,11 @@ print_ast_translation {* in [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] end -*} +\ (*** examples ** (* not sure if the following is the way to set tracing? *) -(* ML {* val _ = Config.put Ast.trace true @{context} *} *) +(* ML \ val _ = Config.put Ast.trace true @{context} \ *) lemma "CHAR ''a'' \ CHAR ''b''" by simp @@ -191,13 +191,13 @@ by simp lemma "''abc''[2] \ CHAR ''a''" by simp -(* ML {* val _ = Config.put Ast.trace false @{context} *} *) +(* ML \ val _ = Config.put Ast.trace false @{context} \ *) **) -subsection {* Records and sets of records *} +subsection \ Records and sets of records \ -text {* +text \ Records are simply represented as enumerated functions with string arguments, such as @{text "(''foo'' :> 1) @ (''bar'' :> TRUE)"}. Similarly, there is no specific @{text "EXCEPT"} construct for records; use the standard one for @@ -207,7 +207,7 @@ text {* record syntax in Isabelle seems difficult, because the Isabelle lexer distinguishes between identifiers and strings: the latter must be surrounded by two single quotes. -*} +\ (** Examples ** From 4500d004e01948607f7bfe88d68d668b866d602a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:15:59 +0200 Subject: [PATCH 088/167] MAI: update `--` comments in theory `Strings` to Isabelle2020 --- isabelle/Strings.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index b274d2cf..41b761f5 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -22,7 +22,7 @@ text \ definition Nibble where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" -definition char -- {* @{text char} is intended to be applied to nibbles *} +definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" lemma charInj [simp]: "(char(a,b) = char(c,d)) = (a=c \ b=d)" From 0b7258b84a74458b85abd60630663ace56a64f0c Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:22:05 +0200 Subject: [PATCH 089/167] MAI: update `xstr` to `str_token` for Isabelle2020 based on the file `NEWS` of the Isabelle `hg` repository. --- isabelle/Strings.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 41b761f5..63c91b52 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -46,8 +46,8 @@ definition String where "String \ Seq(Char)" syntax - "_Char" :: "xstr \ c" ("CHAR _") - "_String" :: "xstr \ c" ("_") + "_Char" :: "str_token \ c" ("CHAR _") + "_String" :: "str_token \ c" ("_") text \ The following parse and print translations convert between the internal From 76faf9f1d986a8b04c116433e411efd7309462c2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:22:31 +0200 Subject: [PATCH 090/167] DRAFT: comment `parse_ast_translation` and `print_ast_translation` because these ML sections need to be updated to Isabelle2020. Mark them as "TODO". --- isabelle/Strings.thy | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 63c91b52..162b00f2 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -59,6 +59,7 @@ text \ lexicon. \ +(* TODO parse_ast_translation \ let (* convert an ML integer to a nibble *) @@ -94,6 +95,7 @@ parse_ast_translation \ [("_Char", char_ast_tr), ("_String", string_ast_tr)] end \ +*) (** debug **) @@ -108,6 +110,7 @@ oops (*ML \ reset Syntax.trace_ast; \*) (**) +(* TODO print_ast_translation \ let (* convert a nibble to an ML integer -- because translation macros have @@ -168,6 +171,7 @@ print_ast_translation \ [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] end \ +*) (*** examples ** (* not sure if the following is the way to set tracing? *) From 905d8840ad3bfff20366131e1f4db14f79cc9a24 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:24:06 +0200 Subject: [PATCH 091/167] DRAFT: comment two `lemma`s that are for debugging because the first lemma raises an error about undefined constant `"_String"`. --- isabelle/Strings.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 162b00f2..cabb62c8 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -98,7 +98,7 @@ parse_ast_translation \ *) -(** debug **) +(** debug ** (*ML \ set Syntax.trace_ast; \*) lemma "''a''" @@ -108,7 +108,7 @@ lemma "CHAR ''a''" oops (*ML \ reset Syntax.trace_ast; \*) -(**) +**) (* TODO print_ast_translation \ From e75014c899fc00c2217444c658b87124095d1b44 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:27:27 +0200 Subject: [PATCH 092/167] MAI: update delimiters in theory `Constant` to Isabelle2020 --- isabelle/Constant.thy | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index 23c186be..f7aff81d 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -6,15 +6,15 @@ Time-stamp: <2011-10-11 17:40:28 merz> *) -header {* Main theory for constant-level Isabelle/\tlaplus{} *} +section \ Main theory for constant-level Isabelle/\tlaplus{} \ theory Constant imports NatDivision CaseExpressions Strings Integers begin -text {* +text \ This is just an umbrella for the component theories. -*} +\ end From f01de6b50c8fba32f650df27aaccb61aa514ef4a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:37:48 +0200 Subject: [PATCH 093/167] MAI: update delimiters in theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 61 ++++++++++++++++++++++--------------------- 1 file changed, 31 insertions(+), 30 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 3502ce19..166ea6c3 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -7,13 +7,13 @@ *) -header {* The Integers as a superset of natural numbers *} +section \ The Integers as a superset of natural numbers \ theory Integers imports Tuples NatArith begin -subsection {* The minus sign *} +subsection \ The minus sign \ consts "minus" :: "c \ c" ("-._" [75] 75) @@ -123,7 +123,7 @@ lemma (*[simp]*) "\x \ Nat : -.1 = -.x" by auto lemma (*[simp]*) "x \ Nat \ (1 = -.x) = FALSE" by (auto simp: sym[OF minus_sym]) -subsection {* The set of Integers *} +subsection \ The set of Integers \ definition Int where "Int \ Nat \ {-.n : n \ Nat}" @@ -200,7 +200,7 @@ lemma intNotNatIsNegNat: "\n \ Nat; n \ Int\ \ Predicates ''is positive'' and 'is negative' \ definition isPos -- {* Predicate ''is positive'' *} where "isPos(n) \ \k \ Nat: n = Succ[k]" @@ -288,7 +288,7 @@ lemma intThenPosZeroNeg: by (auto elim: notIsNeg0_isPos[OF n]) -subsection {* Signum function and absolute value *} +subsection \ Signum function and absolute value \ definition sgn -- {* signum function *} where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" @@ -383,7 +383,7 @@ lemma sgn_minus [simp]: shows "sgn(-.n) = -.sgn(n)" unfolding sgn_def using n by (cases, auto) -text {* Absolute value *} +text \ Absolute value \ lemma absIsNat [simp]: assumes n: "n \ Int" shows "abs(n) \ Nat" @@ -408,12 +408,12 @@ lemma abs_neg [simp]: unfolding abs_def using n by (auto dest: sgnNat_not1) -subsection {* Orders on integers *} +subsection \ Orders on integers \ -text {* +text \ We distinguish four cases, depending on the arguments being in Nat or negative. -*} +\ lemmas int_leq_pp_def = nat_leq_def -- {* 'positive-positive' case, ie: both arguments are naturals *} @@ -448,18 +448,19 @@ lemma neg_le_iff_le [simp]: by(rule intCases2[of m n], simp_all) -subsection {* Addition of integers *} +subsection \ Addition of integers \ -text {* +text \ Again, we distinguish four cases in the definition of @{text "a + b"}, according to each argument being positive or negative. -*} +\ +(* cf. NatArith *) (** The following is rejected by Isabelle because the two definitions are not distinguishable by argument types. defs (unchecked overloaded) int_add_def: "\a \ Int; b \ Int \ \ a + b \ - IF a \ Nat \ b \ Nat THEN addnat(a)[b] (* cf. NatArith *) + IF a \ Nat \ b \ Nat THEN addnat(a)[b] ELSE IF isNeg(a) \ isNeg(b) THEN -.(addnat(-.a)[-.b]) ELSE IF isNeg(a) THEN IF -.a \ b THEN b -- a ELSE -.(a -- b) ELSE IF a \ -.b THEN -.(b -- a) ELSE a -- b" @@ -481,19 +482,19 @@ theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" by (auto simp: int_add_pn_def dest: nat_leq_antisym) -text {* Closure *} +text \ Closure \ lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" by (rule intCases2[of m n], auto simp: int_add_def) -text {* Neutral element *} +text \ Neutral element \ lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" by(rule intCases, auto simp add: int_add_np_def) lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" by(rule intCases, auto simp add: int_add_pn_def) -text {* Additive inverse element *} +text \ Additive inverse element \ lemma add_inverse_nat [simp]: "n \ Nat \ n + -.n = 0" by(simp add: int_add_pn_def) @@ -507,7 +508,7 @@ by (rule intCases, auto simp: int_add_def) lemma add_inverse2_int [simp]: "n \ Int \ -.n + n = 0" by (rule intCases, auto simp: int_add_def) -text {* Commutativity *} +text \ Commutativity \ lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" by(simp add: int_add_def) @@ -515,7 +516,7 @@ by(simp add: int_add_def) lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" by(rule intCases2[of m n], auto simp add: int_add_def add_commute_nat) -text {* Associativity *} +text \ Associativity \ lemma add_pn_eq_adiff [simp]: "\m \ n; m \ Nat; n \ Nat\ \ m + -.n = -.(n -- m)" @@ -695,7 +696,7 @@ by (rule intCases3, auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 int_add_assoc4 int_add_assoc5 int_add_assoc6) -text {* Minus sign distributes over addition *} +text \ Minus sign distributes over addition \ lemma minus_distrib_pn_int [simp]: "m \ Nat \ n \ Nat \ -.(m + -.n) = -.m + n" @@ -713,7 +714,7 @@ lemma int_add_minus_distrib [simp]: by (rule intCases2[OF m n], simp_all) -subsection {* Multiplication of integers *} +subsection \ Multiplication of integers \ axiomatization where int_mult_pn_def: "\a \ Nat; b \ Nat\ \ a * -.b = -.(a * b)" @@ -724,12 +725,12 @@ and theorems int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) -text {* Closure *} +text \ Closure \ lemma multIsInt [simp]: "\a \ Int; b \ Int\ \ a * b \ Int" by (rule intCases2[of a b], simp_all add: int_mult_def) -text {* Neutral element *} +text \ Neutral element \ lemma mult_0_right_int [simp]: "a \ Int \ a * 0 = 0" by (rule intCases[of a], simp_all add: int_mult_np_def) @@ -737,12 +738,12 @@ by (rule intCases[of a], simp_all add: int_mult_np_def) lemma mult_0_left_int [simp]: "a \ Int \ 0 * a = 0" by (rule intCases[of a], simp_all add: int_mult_pn_def) -text {* Commutativity *} +text \ Commutativity \ lemma mult_commute_int: "\a \ Int; b \ Int\ \ a * b = b * a" by (rule intCases2[of a b], simp_all add: int_mult_def mult_commute_nat) -text {* Identity element *} +text \ Identity element \ lemma mult_1_right_int [simp]: "a \ Int \ a * 1 = a" by (rule intCases[of a], simp_all add: int_mult_def) @@ -750,14 +751,14 @@ by (rule intCases[of a], simp_all add: int_mult_def) lemma mult_1_left_int [simp]: "a \ Int \ 1 * a = a" by (rule intCases[of a], simp_all add: int_mult_def) -text {* Associativity *} +text \ Associativity \ lemma mult_assoc_int: assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" shows "m * (n * p) = (m * n) * p" by(rule intCases3[OF m n p], simp_all add: mult_assoc_nat int_mult_def) -text {* Distributivity *} +text \ Distributivity \ lemma ppn_distrib_left_nat: (* ppn stands for m=positive, n=positive, p=negative *) assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" @@ -840,7 +841,7 @@ apply(rule intCases3[OF m n p], apply(simp only: add_mult_distrib_right_nat) done -text {* Minus sign distributes over multiplication *} +text \ Minus sign distributes over multiplication \ lemma minus_mult_left_int: assumes m: "m \ Int" and n: "n \ Int" @@ -853,14 +854,14 @@ lemma minus_mult_right_int: by (rule intCases2[OF m n], simp_all add: int_mult_def) -subsection {* Difference of integers *} +subsection \ Difference of integers \ -text {* +text \ Difference over integers is simply defined as addition of the complement. Note that this difference, noted @{text "-"}, is different from the difference over natural numbers, noted @{text "--"}, even for two natural numbers, because the latter cuts off at $0$. -*} +\ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" From 4df449ac6c9c7756188f70fac93b5a37299953fc Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 17:40:15 +0200 Subject: [PATCH 094/167] MAI: update `--` comments in theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 166ea6c3..f899a391 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -18,7 +18,7 @@ subsection \ The minus sign \ consts "minus" :: "c \ c" ("-._" [75] 75) -syntax -- {* syntax for negative naturals *} +syntax (* -- syntax for negative naturals *) "-.0" :: "c" ("-.0") "-.1" :: "c" ("-.1") "-.2" :: "c" ("-.2") @@ -144,7 +144,7 @@ lemma intCases [case_names Positive Negative, cases set: Int]: shows "P" using assms unfolding Int_def by auto --- {* Integer cases over two parameters *} +(* -- Integer cases over two parameters *) lemma intCases2: assumes m: "m \ Int" and n: "n \ Int" and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" @@ -202,10 +202,10 @@ unfolding Int_def by auto subsection \ Predicates ''is positive'' and 'is negative' \ -definition isPos -- {* Predicate ''is positive'' *} +definition isPos (* -- Predicate ''is positive'' *) where "isPos(n) \ \k \ Nat: n = Succ[k]" -definition isNeg -- {* Predicate ''is negative'' *} +definition isNeg (* -- Predicate ''is negative'' *) where "isNeg(n) \ \k \ Nat: n = -.Succ[k]" lemma boolify_isPos [simp]: "boolify(isPos(n)) = (isPos(n))" @@ -251,7 +251,7 @@ lemma negNotNat_isNat: using n by (cases, auto) lemma noNatisNeg [simp]: - "n \ Nat \ isNeg(n) = FALSE" -- {* No natural number is negative *} + "n \ Nat \ isNeg(n) = FALSE" (* -- No natural number is negative *) unfolding isNeg_def using negNotInNat by blast lemma negNat_isNeg [intro]: "\m \ Nat; m \ 0\ \ isNeg(-.m)" @@ -290,10 +290,10 @@ by (auto elim: notIsNeg0_isPos[OF n]) subsection \ Signum function and absolute value \ -definition sgn -- {* signum function *} +definition sgn (* -- signum function *) where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" -definition abs -- {* absolute value *} +definition abs (* -- absolute value *) where "abs(n) \ IF sgn(n) = -.1 THEN -.n ELSE n" lemma sgnInInt [simp]: "n \ Int \ sgn(n) \ Int" @@ -416,7 +416,7 @@ text \ \ lemmas int_leq_pp_def = nat_leq_def - -- {* 'positive-positive' case, ie: both arguments are naturals *} + (* -- 'positive-positive' case, ie: both arguments are naturals *) axiomatization where int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = FALSE" @@ -438,7 +438,7 @@ unfolding isBool_def by auto lemma int_leq_refl [iff]: "n \ Int \ n \ n" by(rule intCases, auto) -lemma eq_leq_bothE: -- {* reduce equality over integers to double inequality *} +lemma eq_leq_bothE: (* -- reduce equality over integers to double inequality *) assumes "m \ Int" and "n \ Int" and "m = n" and "\m \ n; n \ m\ \ P" shows "P" using assms by simp @@ -467,7 +467,7 @@ defs (unchecked overloaded) **) lemmas - int_add_pp_def = nat_add_def -- {* both numbers are positive, ie. naturals *} + int_add_pp_def = nat_add_def (* -- both numbers are positive, ie. naturals *) axiomatization where int_add_pn_def: "\a \ Nat; b \ Nat\ \ a + (-.b) \ IF a \ b THEN -.(b -- a) ELSE a -- b" @@ -477,7 +477,7 @@ and int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) - --{* When we use these definitions, we don't want to unfold the 'pp' case *} + (* -- When we use these definitions, we don't want to unfold the 'pp' case *) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" by (auto simp: int_add_pn_def dest: nat_leq_antisym) @@ -866,7 +866,7 @@ text \ definition diff (infixl "-" 65) where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" -lemma diffIsInt [simp]: -- {* Closure *} +lemma diffIsInt [simp]: (* -- Closure *) "\m \ Int; n \ Int\ \ m - n \ Int" by (simp add: int_diff_def) From dea11d653c5fd8d49fcf2cb125ccdd10ec8a3955 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 18:52:02 +0200 Subject: [PATCH 095/167] MAI: update to `lemmas` for Isabelle2020 in theory `Integers` --- isabelle/Integers.thy | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index f899a391..650c63ea 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -425,7 +425,7 @@ and and int_leq_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ -.b = (b \ a)" -(* theorems int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) +(* lemmas int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) lemma int_boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" @@ -476,7 +476,7 @@ int_add_np_def: "\a \ Nat; b \ Nat\ \ (- and int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" -theorems int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) +lemmas int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) (* -- When we use these definitions, we don't want to unfold the 'pp' case *) lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" @@ -723,7 +723,7 @@ and and int_mult_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a * -.b = a * b" -theorems int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) +lemmas int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) text \ Closure \ From 445c7113f3290c1a02f3caa05475db38f46b0b29 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 26 Dec 2020 18:52:42 +0200 Subject: [PATCH 096/167] MAI: rm nonexistent `assms` (raises error in Isabelle2020) --- isabelle/Integers.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 650c63ea..e9159290 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -883,7 +883,7 @@ lemma diff_self_eq_0_int [simp]: "m \ Int \ m - m = 0" by (simp add: int_diff_def) lemma neg_diff_is_diff [simp]: "\m \ Int; n \ Int\ \ -.(m - n) = n - m" -using assms by (simp add: int_diff_def add_commute_int) +by (simp add: int_diff_def add_commute_int) lemma diff_nat_is_add_neg: "\m \ Nat; n \ Nat\ \ m - n = m + -.n" by (simp add: int_diff_def) From 9822c72c6a15ac17c6cf3d581697c3e190b62781 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:17:37 +0200 Subject: [PATCH 097/167] ENH: add theorem `leq_adiff_left_add_right_equiv` to the theory `NatArith`, to be used for proving the theorem `int_add_assoc1` in the theory `Integers`. --- isabelle/NatArith.thy | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 179a0799..912320e2 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -1027,6 +1027,11 @@ lemma leq_adiff_right_add_left: shows "m \ n -- k = (m + k \ n)" using assms by (induct n k rule: diffInduct, simp+) +lemma leq_adiff_left_add_right_equiv: + assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" + shows "(n -- k \ m) = (n \ m + k)" +using assms by (induct n k rule: diffInduct, simp+) + lemma leq_adiff_left_add_right: assumes 1: "n -- p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "n \ m + p" From bedb5d5a0c9a4e2f5981009a03a664b35ba99114 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:20:07 +0200 Subject: [PATCH 098/167] MAI: partially update theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 182 +++++++++++++++++++++++++++++++++++++++--- 1 file changed, 169 insertions(+), 13 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index e9159290..d314b323 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -552,19 +552,175 @@ declare leq_neq_iff_less [simplified,simp] lemma int_add_assoc1: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (n + -.p) = (m + n) + -.p" -apply(rule nat_leq_cases[OF p n]) - using assms apply simp_all - apply(rule nat_leq_cases[of p "m + n"], simp_all) - apply(simp add: adiff_add_assoc[OF _ m n p]) - apply(rule nat_leq_cases[of p "m + n"], simp+) - apply(rule nat_leq_cases[of "p -- n" m], simp+) - apply(rule adiff_add_assoc3, simp+) - apply(rule adiff_add_assoc5, simp+) - apply(rule nat_leq_cases[of "p -- n" m], simp_all) - apply(rule adiff_add_assoc6, simp_all) - apply(simp only: add_commute_nat[of m n]) - apply(rule adiff_adiff_left_nat, simp+) -done +proof - +have s1_1: "n \ p ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s1_1_asm: "n \ p" + have s2_1: "n + -.p = -.(p -- n)" + using s1_1_asm n p int_add_pn_def[of "n" "p"] by auto + have s2_2: "p \ m + n ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s2_2_asm: "p \ m + n" + have s3_1: "m + (n + -.p) = m -- (p -- n)" + proof - + have s4_1: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_2: "m + -.(p -- n) = m -- (p -- n)" + proof - + have s5_1: "p -- n \ m" + using s1_1_asm s2_2_asm + leq_adiff_left_add_right_equiv[of "n" "p" "m"] + p n m by auto + show ?thesis + using s5_1 p n m adiffIsNat + int_add_pn_def by auto + qed + show ?thesis + using s4_1 s4_2 by auto + qed + have s3_2: "(m + n) + -.p = (m + n) -- p" + using s2_2_asm p n m addIsNat int_add_pn_def by auto + have s3_3: "m -- (p -- n) = (m + n) -- p" + using s1_1_asm s2_2_asm + adiff_add_assoc3[of "n" "p" "m"] m n p by auto + show "m + (n + -.p) = (m + n) + -.p" + using s3_1 s3_2 s3_3 by auto + qed + have s2_3: "\ (p \ m + n) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s2_3_asm: "\ (p \ m + n)" + have s3_1: "m + n < p" + using s2_3_asm p m n addIsNat nat_not_leq by auto + have s3_2: "m + n \ p" + using s3_1 by (auto simp: less_def) + have s3_3: "(m + n) + -.p = -.(p -- (m + n))" + using s3_2 m n p addIsNat int_add_pn_def by auto + have s3_4: "p -- n \ m ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s3_4_asm: "p -- n \ m" + have s4_1: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_2: "m + -.(p -- n) = m -- (p -- n)" + using s3_4_asm p n m adiffIsNat int_add_pn_def + by auto + have s4_3: "m + (n + -.p) = m -- (p -- n)" + using s4_1 s4_2 by auto + have s4_4: "m -- (p -- n) = -.(p -- (m + n))" + using m n p s1_1_asm s3_2 s3_4_asm + adiff_add_assoc6 by auto + show "m + (n + -.p) = (m + n) + -.p" + using s3_3 s4_3 s4_4[symmetric] by auto + qed + have s3_5: "\ (p -- n \ m) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s3_5_asm: "\ (p -- n \ m)" + have s4_1: "m \ p -- n" + proof - + have s5_1: "m < p -- n" + using p n m adiffIsNat nat_not_leq s3_5_asm by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "m + (n + -.p) = m + -.(p -- n)" + using s2_1 by auto + have s4_3: "m + -.(p -- n) = -.((p -- n) -- m)" + using s4_1 s3_5_asm p n m adiffIsNat + int_add_pn_def[of "m" "p -- n"] + by auto + have s4_4: "m + (n + -.p) = -.((p -- n) -- m)" + using s4_2 s4_3 by auto + have s4_5: "(p -- n) -- m = p -- (n + m)" + using p n m adiff_adiff_left_nat by auto + have s4_6: "(p -- n) -- m = p -- (m + n)" + using n m add_commute_nat s4_5 by auto + have s4_7: "-.((p -- n) -- m) = -.(p -- (m + n))" + using s4_6 by auto + show "m + (n + -.p) = (m + n) + -.p" + using s4_7 s4_4 s3_3 by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s3_4 s3_5 nat_leq_isBool[of "p -- n" "m"] + isBoolTrueFalse[of "p -- n \ m"] + by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s2_2 s2_3 + nat_leq_isBool[of "p" "m + n"] + isBoolTrueFalse[of "p \ m + n"] + by auto + qed +have s1_2: "\ (n \ p) ==> + m + (n + -.p) = (m + n) + -.p" + proof - + assume s1_2_asm: "\ (n \ p)" + have s2_1: "n + -.p = n -- p" + proof - + have s3_1: "\n \ Nat; p \ Nat\ + \ + n + (-.p) \ + IF n \ p THEN -.(p -- n) ELSE n -- p" + using int_add_pn_def[of "n" "p"] by auto + have s3_2: "n + (-.p) \ n -- p" + using s3_1 s1_2_asm n p by auto + show ?thesis + using s3_2 by auto + qed + have s2_2: "\ ((m + n) \ p)" + proof - + have s3_1: "p < n" + using p n s1_2_asm nat_not_leq by auto + have s3_2: "p < m + n" + using m p n s3_1 trans_less_add2 by auto + show ?thesis + using s3_2 p m n addIsNat nat_not_leq by auto + qed + have s2_3: "(m + n) + -.p = (m + n) -- p" + proof - + have s3_1: "m + n \\in Nat" + using m n addIsNat by auto + have s3_2: "\(m + n) \\in Nat; p \\in Nat\ + \ + (m + n) + (-.p) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE ((m + n) -- p)" + using int_add_pn_def[of "m + n" "p"] by auto + have s3_3: "(m + n) + (-.p) \ (m + n) -- p" + using s3_1 p s3_2 s2_2 by auto + show ?thesis + using s3_3 by auto + qed + have s2_4: "(m + (n + -.p) = (m + n) + -.p) = + (m + (n -- p) = (m + n) -- p)" + using s2_1 s2_3 by auto + have s2_5: "(m + n) -- p = m + (n -- p)" + proof - + have s3_1: "p \ n" + proof - + have s4_1: "p < n" + using s1_2_asm p n nat_not_leq by auto + show ?thesis + using s4_1 by (auto simp: less_def) + qed + show ?thesis + using m n p s3_1 adiff_add_assoc[of "p" "n" "m"] by auto + qed + show "m + (n + -.p) = (m + n) + -.p" + using s2_4 s2_5[symmetric] by auto + qed +show ?thesis + proof - + have s2_1: "((n \ p) = TRUE) \ ((n \ p) = FALSE)" + using nat_leq_isBool[of "n" "p"] isBoolTrueFalse[of "n \ p"] + by auto + show ?thesis + using s1_1 s1_2 s2_1 by auto + qed +qed lemma int_add_assoc2: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" From 4f2a1fa2f2ac64377b9a813abdca0fb066682bf2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:20:27 +0200 Subject: [PATCH 099/167] STY: shorter lines --- isabelle/NatOrderings.thy | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index 75b44e09..f0407ee6 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -162,7 +162,10 @@ unfolding nat_leq_def using uptoLinear . lemma nat_leq_cases: assumes m: "m \ Nat" and n: "n \ Nat" - and leq: "m \ n \ P" and geq: "\n \ m; n \ m\ \ P" + and leq: "m \ n \ P" and + geq: "\n \ m; + n \ m + \ \ P" shows "P" proof (cases "m \ n") case True thus "P" by (rule leq) From 2e154e059922d9c9856a6f18ca339132bc4f7835 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sun, 27 Dec 2020 22:23:48 +0200 Subject: [PATCH 100/167] DEV: ignore directory `isabelle/output/` which contains the document generated by Isabelle for the TLA+ theories. rebasing the isabelle2020 branch on master, fixing merge conflict by adding /isabelle/output to .gitignore, as on master --- .gitignore | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/.gitignore b/.gitignore index cb918ca5..fdbae785 100644 --- a/.gitignore +++ b/.gitignore @@ -28,7 +28,4 @@ /tools/installer/tlaps-source-release.sh # `tlapm` binary tlapm -# installation for testing, -# created by the installer -/test/bin/ -/test/lib/ +/isabelle/output/ From 3de7a35b33270c225ab5e7b17dd95541f3ecf10b Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 20:20:48 +0200 Subject: [PATCH 101/167] MAI: update theory `Integers` to Isabelle2020 --- isabelle/Integers.thy | 292 +++++++++++++++++++++++++++++++++++++++--- 1 file changed, 274 insertions(+), 18 deletions(-) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index d314b323..c4398167 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -725,24 +725,280 @@ qed lemma int_add_assoc2: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "m + (-.p + n) = (m + -.p) + n" -using assms apply (simp add: add_commute_int[of "-.p" n]) - using int_add_assoc1[OF m n p] apply simp - apply(rule nat_leq_cases[of p "m + n"], simp_all) - apply(rule nat_leq_cases[OF p m], simp_all) - apply(rule adiff_add_assoc2, simp_all) - apply (simp add: add_commute_int[of "-.(p -- m)" n]) - apply(simp only: add_commute_nat[OF m n]) - apply(rule nat_leq_cases[of "p -- m" n], simp_all) - apply(rule adiff_add_assoc3[symmetric], simp+) - apply(rule adiff_add_assoc5[symmetric], simp+) - apply(rule nat_leq_cases[OF p m], simp_all) - apply (simp add: add_commute_nat[OF m n]) - apply (simp add: add_commute_int[of "-.(p -- m)" n]) - apply(rule nat_leq_cases[of "p -- m" n], simp_all) - apply(simp add: add_commute_nat[OF m n]) - apply(rule adiff_add_assoc6[symmetric], simp+) - apply(rule adiff_adiff_left_nat[symmetric], simp+) -done +proof - +have s1_1: "(m + n) + -.p = (m + -.p) + n" + proof - + have s2_1: "m + n \ p ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s2_1_asm: "m + n \ p" + have s3_1: "(m + n) + -.p = -.(p -- (m + n))" + proof - + have s4_1: "\m + n \ Nat; p \ Nat\ + \ ((m + n) + (-.p)) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE + (m + n) -- p" + using int_add_pn_def[of "m + n" "p"] by auto + have s4_2: "(m + n) + -.p \ -.(p -- (m + n))" + using s4_1 s2_1_asm m n p addIsNat by auto + show ?thesis + using s4_2 by auto + qed + have s3_2: "m \ p" + using s2_1_asm add_leqD1 m n p by auto + have s3_3: "(m + -.p) + n = -.(p -- m) + n" + proof - + have s4_1: "m + -.p = -.(p -- m)" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ -.(p -- m)" + using s5_1 m p s3_2 by auto + show ?thesis + using s5_2 by auto + qed + show ?thesis + using s4_1 by auto + qed + have s3_4: "n \ p -- m ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_4_asm: "n \ p -- m" + have s4_1: "(m + -.p) + n = -.((p -- m) -- n)" + proof - + have s5_1: "-.(p -- m) + n = n + -.(p -- m)" + using p n m adiffIsNat add_commute_pn_nat by auto + have s5_2: "n + -.(p -- m) = -.((p -- m) -- n)" + proof - + have s6_1: "\n \ Nat; p -- m \ Nat\ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s6_2: "n + -.(p -- m) \ -.((p -- m) -- n)" + using s6_1 n m p adiffIsNat s3_4_asm by auto + show ?thesis + using s6_2 by auto + qed + show ?thesis + using s3_3 s5_1 s5_2 by auto + qed + have s4_2: "(p -- m) -- n = p -- (m + n)" + using p n m adiff_adiff_left_nat by auto + show "(m + n) + -.p = (m + -.p) + n" + using s3_1 s4_1 s4_2 by auto + qed + have s3_5: "\ (n \ p -- m) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_5_asm: "\ (n \ p -- m)" + have s4_1: "p -- m \ n" + proof - + have s5_1: "p -- m < n" + using s3_5_asm nat_not_leq p n m adiffIsNat by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "-.(p -- m) + n = n -- (p -- m)" + proof - + have s5_1: "-.(p -- m) + n = n + -.(p -- m)" + using p n m add_commute_pn_nat adiffIsNat by auto + have s5_2: "n + -.(p -- m) = n -- (p -- m)" + proof - + have s6_1: "\n \ Nat; p -- m \ Nat\ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s6_2: "n + -.(p -- m) \ n -- (p -- m)" + using s6_1 s3_5_asm p n m adiffIsNat by auto + show ?thesis + using s6_2 by auto + qed + show ?thesis + using s5_1 s5_2 by auto + qed + have s4_3: "n -- (p -- m) = -.(p -- (m + n))" + using adiff_add_assoc6 m n p s3_2 s2_1_asm + s4_1 add_commute_nat by auto + show "(m + n) + -.p = (m + -.p) + n" + using s3_1 s3_3 s4_2 s4_3[symmetric] by auto + qed + show ?thesis + using s3_4 s3_5 nat_leq_isBool[of "n" "p -- m"] + isBoolTrueFalse[of "n \ p -- m"] by auto + qed + have s2_2: "\ (m + n \ p) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s2_2_asm: "\ (m + n \ p)" + have s3_1: "(m + n) + -.p = (m + n) -- p" + proof - + have s4_1: "\(m + n) \ Nat; p \ Nat\ + \ ((m + n) + (-.p)) \ + IF m + n \ p THEN + -.(p -- (m + n)) ELSE + (m + n) -- p" + using int_add_pn_def[of "m + n" "p"] by auto + have s4_2: "((m + n) + (-.p)) \ ((m + n) -- p)" + proof - + have s5_1: "m + n \\in Nat" + using m n addIsNat by auto + show "((m + n) + (-.p)) \ ((m + n) -- p)" + using s2_2_asm s4_1[OF s5_1 p] by auto + qed + show ?thesis + using s4_2 by auto + qed + have s3_2: "m \ p ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_2_asm: "m \ p" + have s4_1: "m + -.p = -.(p -- m)" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ -.(p -- m)" + using s5_1 m p s3_2_asm by auto + show ?thesis + using s5_2 by auto + qed + have s4_4: "p \ m + n" + proof - + have s5_1: "p < m + n" + using s2_2_asm p m n addIsNat nat_not_leq by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "n \ p -- m ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s4_2_asm: "n \ p -- m" + have s5_1: "(m + -.p) + n = -.((p -- m) -- n)" + proof - + have s6_1: "(m + -.p) + n = -.(p -- m) + n" + using s4_1 by auto + have s6_2: "-.(p -- m) + n = n + -.(p -- m)" + using p m n adiffIsNat add_commute_pn_nat by auto + have s6_3: "n + -.(p -- m) = -.((p -- m) -- n)" + proof - + have s7_1: "\n \ Nat; p -- m \ Nat + \ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s7_2: "n + -.(p -- m) \ -.((p -- m) -- n)" + using s7_1 p n m adiffIsNat s4_2_asm by auto + show ?thesis + using s7_2 by auto + qed + show ?thesis + using s6_1 s6_2 s6_3 by auto + qed + have s5_2: "-.((p -- m) -- n) = (m + n) -- p" + using adiff_add_assoc5 s4_4 s4_2_asm s3_2_asm + add_commute_nat + m n p by auto + show "(m + n) + -.p = (m + -.p) + n" + using s5_1 s5_2 s3_1 by auto + qed + have s4_3: "\ (n \ p -- m) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s4_3_asm: "\ (n \ p -- m)" + have s5_1: "(m + n) -- p = n -- (p -- m)" + using m n p adiff_add_assoc3[of "m" "p" "n"] + s3_2_asm s4_4 add_commute_nat + by auto + have s5_2: "(m + -.p) + n = n -- (p -- m)" + proof - + have s6_1: "(m + -.p) + n = -.(p -- m) + n" + using s4_1 by auto + have s6_2: "-.(p -- m) + n = n + -.(p -- m)" + using s6_1 n p m adiffIsNat + add_commute_pn_nat by auto + have s6_3: "n + -.(p -- m) = n -- (p -- m)" + proof - + have s7_1: "\n \ Nat; p -- m \ Nat + \ + \ (n + (-.(p -- m))) \ + IF n \ p -- m THEN + -.((p -- m) -- n) ELSE + n -- (p -- m)" + using int_add_pn_def[of "n" "p -- m"] by auto + have s7_2: "n + -.(p -- m) \ n -- (p -- m)" + using s7_1 s4_3_asm n p m adiffIsNat + by auto + show ?thesis + using s7_2 by auto + qed + show ?thesis + using s6_1 s6_2 s6_3 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s5_1 s3_1 s5_2 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s4_2 s4_3 nat_leq_isBool[of "n" "p -- m"] + isBoolTrueFalse[of "n \ p -- m"] by auto + qed + have s3_3: "\ (m \ p) ==> + (m + n) + -.p = (m + -.p) + n" + proof - + assume s3_3_asm: "\ (m \ p)" + have s4_1: "p \ m" + proof - + have s5_1: "p < m" + using s3_3_asm p m nat_not_leq by auto + show ?thesis + using s5_1 by (auto simp: less_def) + qed + have s4_2: "(m + n) -- p = (m -- p) + n" + using adiff_add_assoc2 s4_1 p n m by auto + have s4_3: "m + -.p = m -- p" + proof - + have s5_1: "\m \ Nat; p \ Nat\ + \ (m + (-.p)) \ + IF m \ p THEN + -.(p -- m) ELSE + m -- p" + using int_add_pn_def[of "m" "p"] by auto + have s5_2: "m + -.p \ m -- p" + using s5_1 m p s3_3_asm by auto + show ?thesis + using s5_2 by auto + qed + show "(m + n) + -.p = (m + -.p) + n" + using s4_2 s3_1 s4_3 by auto + qed + show ?thesis + using s3_2 s3_3 nat_leq_isBool[of "m" "p"] + isBoolTrueFalse[of "m \ p"] by auto + qed + show ?thesis + using s2_1 s2_2 nat_leq_isBool[of "m + n" "p"] + isBoolTrueFalse[of "m + n \ p"] by auto + qed +have s1_2: "m + (n + -.p) = (m + -.p) + n" + using s1_1 int_add_assoc1[OF m n p] by auto +have s1_3: "n + -.p = -.p + n" + using n p add_commute_pn_nat[of "n" "p"] by auto +show ?thesis + using s1_2 s1_3 by auto +qed declare leq_neq_iff_less [simplified,simp del] From 313af1caf0087810c2f8e5e0f2d5c4eabe2aa9c2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:35:12 +0200 Subject: [PATCH 102/167] MAI: update delimiters in theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 46a9796f..133039c9 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -11,7 +11,7 @@ theory Zenon imports Constant begin -text {* The following lemmas make a cleaner meta-object reification *} +text \ The following lemmas make a cleaner meta-object reification \ lemma atomize_meta_bAll [atomize]: "(\x. (x \ S \ P(x))) \ Trueprop (\ x \ S : P(x))" From 384fc98133b8b17ff9d0cf854c3be201e679034a Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:36:51 +0200 Subject: [PATCH 103/167] DRAFT: define abbreviations `\\cup` and `\\cap` in theory `Zenon` This syntax was present in the theory `SetTheory` in its version for Isabelle2011-1. The abbreviation in the theory `Zenon` is for use within the theory `Zenon`, and possibly by proofs generated by `zenon`. --- isabelle/Zenon.thy | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 133039c9..f90e1d6f 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -336,6 +336,14 @@ by blast (* INTER -- needed ? *) +abbreviation + cup :: "[c, c] \ c" (infixl "\\cup" 65) + where "a \\cup b \ a \ b" + +abbreviation + cap :: "[c, c] \ c" (infixl "\\cap" 70) + where "a \\cap b \ a \ b" + lemma zenon_in_cup : "x \\in A \\cup B ==> (x \\in A ==> FALSE) ==> (x \\in B ==> FALSE) ==> FALSE" by blast From 62d9d8c0ce95b11a836d0129010b3930338c600d Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:37:39 +0200 Subject: [PATCH 104/167] MAI: partially update theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index f90e1d6f..702cbe57 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -944,7 +944,8 @@ proof - assume h2: "~c ==> P(b) ==> FALSE" have h2x: "~c ==> ~P(b)" using h2 by auto have h3: "~P (IF c THEN a ELSE b)" - by (rule condI [of c "\ x . ~P (x)", OF h1x h2x]) + using atomize_not condI[of c "\ x . ~P (x)", OF h1x h2x] + by (auto) show FALSE by (rule notE [OF h3 main]) qed @@ -1211,7 +1212,7 @@ lemma zenon_tuple_acc_1 : lemma zenon_tuple_acc_2 : "isASeq (r) ==> k \\in Nat ==> 0 < k ==> k \ Len (r) ==> x = r[k] ==> x = Append (r, e) [k]" -by auto +using appendElt1 by auto definition zenon_ss :: "c \ c" where "zenon_ss (n) \ IF n \\in Nat THEN Succ[n] ELSE 1" From 6aba81abfa1925103bc8fad7d7a881e6a9c3c622 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:40:29 +0200 Subject: [PATCH 105/167] DRAFT: use `EnumFuncSet` instead of `[... : ...]` syntax because the required `parse_ast_translation` and `print_ast_translation` in the theory `Tuples` are currently commented, until they are updated to Isabelle2020. --- isabelle/Zenon.thy | 36 +++++++++++++++++++++++++++++------- 1 file changed, 29 insertions(+), 7 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 702cbe57..460cc898 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1263,7 +1263,9 @@ lemma zenon_inrecordsetI1 : assumes hfcn: "isAFcn (r)" assumes hdom: "DOMAIN r = {l1x}" assumes h1: "r[l1x] \\in s1x" - shows "r \\in [l1x : s1x]" + (* TODO + shows "r \\in [l1x : s1x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x}" @@ -1324,7 +1326,9 @@ lemma zenon_inrecordsetI2 : assumes hdom: "DOMAIN r = {l1x, l2x}" assumes h1: "r[l1x] \\in s1x" assumes h2: "r[l2x] \\in s2x" - shows "r \\in [l1x : s1x, l2x : s2x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x}" @@ -1404,7 +1408,9 @@ lemma zenon_inrecordsetI3 : assumes h1: "r[l1x] \\in s1x" assumes h2: "r[l2x] \\in s2x" assumes h3: "r[l3x] \\in s3x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x]" *) + shows "r \\in EnumFuncSet(<>, <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x}" @@ -1503,7 +1509,11 @@ lemma zenon_inrecordsetI4 : assumes h2: "r[l2x] \\in s2x" assumes h3: "r[l3x] \\in s3x" assumes h4: "r[l4x] \\in s4x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x}" @@ -1621,7 +1631,11 @@ lemma zenon_inrecordsetI5 : assumes h3: "r[l3x] \\in s3x" assumes h4: "r[l4x] \\in s4x" assumes h5: "r[l5x] \\in s5x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x}" @@ -1758,7 +1772,11 @@ lemma zenon_inrecordsetI6 : assumes h4: "r[l4x] \\in s4x" assumes h5: "r[l5x] \\in s5x" assumes h6: "r[l6x] \\in s6x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x, l6x}" @@ -1914,7 +1932,11 @@ lemma zenon_inrecordsetI7 : assumes h5: "r[l5x] \\in s5x" assumes h6: "r[l6x] \\in s6x" assumes h7: "r[l7x] \\in s7x" - shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x, l7x : s7x]" + (* TODO + shows "r \\in [l1x : s1x, l2x : s2x, l3x : s3x, l4x : s4x, l5x : s5x, l6x : s6x, l7x : s7x]" *) + shows "r \\in EnumFuncSet( + <>, + <>)" proof - let ?doms = "<>" let ?domset = "{l1x, l2x, l3x, l4x, l5x, l6x, l7x}" From 56c65bda5329b4aa19cb2efe659807e37726ac5d Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Mon, 28 Dec 2020 21:57:07 +0200 Subject: [PATCH 106/167] MAI: update definition syntax in theory `Zenon` to Isabelle2020 --- isabelle/Zenon.thy | 179 ++++++++++++++++++++++++++------------------- 1 file changed, 103 insertions(+), 76 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 460cc898..49cbfd67 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -2108,7 +2108,7 @@ lemma zenon_caseother0 : proof - fix P e0 assume h: "P (CASE OTHER -> e0)" - def cas == "CASE OTHER -> e0" + define cas where "cas \ CASE OTHER -> e0" have hh: "P (cas)" using h by (fold cas_def) assume hoth: "P (e0) ==> FALSE" have hb: "(\ i \ DOMAIN <<>> : ~<<>>[i]) = TRUE" by auto @@ -2666,13 +2666,14 @@ lemma zenon_case1 : assumes hoth: "~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c1x" (is "?dcs") + define dcs where "dcs \ c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2682,10 +2683,10 @@ proof - using hoth hh1 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq ( + define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -2740,26 +2741,28 @@ lemma zenon_caseother1 : assumes hoth: "~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq ( + define dcs where "dcs \ c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -2845,13 +2848,14 @@ lemma zenon_case2 : assumes hoth: "~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c2x | c1x" (is "?dcs") + define dcs where "dcs \ c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2861,10 +2865,10 @@ proof - using hoth hh1 hh2 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq ( + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -2923,26 +2927,28 @@ lemma zenon_caseother2 : assumes hoth: "~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq ( + define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3033,13 +3039,14 @@ lemma zenon_case3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3050,10 +3057,12 @@ proof - using hoth hh1 hh2 hh3 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3118,26 +3127,29 @@ lemma zenon_caseother3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3235,13 +3247,14 @@ lemma zenon_case4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3253,10 +3266,12 @@ proof - using hoth hh1 hh2 hh3 hh4 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3327,26 +3342,30 @@ lemma zenon_caseother4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c4x | c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) @@ -3451,13 +3470,14 @@ lemma zenon_case5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c5x | c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3470,10 +3490,12 @@ proof - using hoth hh1 hh2 hh3 hh4 hh5 by blast next assume ha: "?dcs" - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" (is "?scs") - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" (is "?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" @@ -3550,26 +3572,31 @@ lemma zenon_caseother5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - def cs == "<>" (is "?cs") - def es == "<>" (is "?es") - def arms == "UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" + define cs where "cs \ <>" (is "?cs") + define es where "es \ <>" (is "?es") + define arms where "arms \ + UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" (is "?arms") - def cas == "?cas" + define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - def dcs == "c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") - def scs == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define dcs where "dcs \ + c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define scs where "scs \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" (is "?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) - def ses == "zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( + define ses where "ses \ + zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" (is "?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) - def armoth == "CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" + define armoth where "armoth \ + CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" (is "?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) From 7aeb6815362f960a128fa6a8b9fce6b803720781 Mon Sep 17 00:00:00 2001 From: merz Date: Sun, 3 Jan 2021 09:38:00 +0100 Subject: [PATCH 107/167] pass (mainly) over NatArith --- isabelle/FixedPoints.thy | 2 +- isabelle/Functions.thy | 2 +- isabelle/NatArith.thy | 826 ++++++++++++++---------------------- isabelle/NatOrderings.thy | 30 +- isabelle/Peano.thy | 2 +- isabelle/PredicateLogic.thy | 13 +- isabelle/SetTheory.thy | 8 +- 7 files changed, 338 insertions(+), 545 deletions(-) diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index 7ddf3970..28162aea 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -1,6 +1,6 @@ (* Title: TLA+/FixedPoints.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index b68d19a2..ca831784 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -1,6 +1,6 @@ (* Title: TLA+/Functions.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy index 912320e2..1e58c382 100644 --- a/isabelle/NatArith.thy +++ b/isabelle/NatArith.thy @@ -1,9 +1,8 @@ (* Title: TLA+/NatArith.thy Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 - Time-stamp: <2011-10-11 17:39:37 merz> *) section \ Arithmetic (except division) over natural numbers \ @@ -152,7 +151,7 @@ lemma one_is_add_nat: using m apply (rule natCases) using n by (induct, auto)+ -lemma add_eq_self_zero_nat [simp]: +lemma add_eq_self_zero_nat1 [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(m + n = m) = (n = 0)" using n apply (rule natCases) @@ -160,6 +159,11 @@ using n apply (rule natCases) using m apply induct apply auto done +lemma add_eq_self_zero_nat2 [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(n + m = m) = (n = 0)" +using assms by (simp add: add_commute_nat) + subsection \ Multiplication of natural numbers \ @@ -186,7 +190,7 @@ lemma multnat_0: assumes "m \ Nat" shows "multnat(m)[0] = 0" unfolding multnat_def by (rule primrec_natE, auto simp: assms) -lemma mult_0_nat [simp]: (* -- "neutral element" *) +lemma mult_0_nat [simp]: \ \ neutral element \ assumes n: "n \ Nat" shows "n * 0 = 0" by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) @@ -427,28 +431,26 @@ using assms by (simp add: nat_adiff_def adiffnat_Succ[OF assms]) lemma adiff_0_eq_0_nat [simp]: assumes n: "n \ Nat" shows "0 -- n = 0" -using n apply induct by (simp_all add: adiff_Succ_nat) +using assms by (induct) (simp_all add: adiff_Succ_nat) text \ Reasoning about @{text "m -- m = 0"}, etc. \ lemma adiff_Succ_Succ_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "Succ[m] -- Succ[n] = m -- n" -using n apply induct -using assms by (auto simp add: adiff_Succ_nat) + using n by induct (auto simp add: m adiff_Succ_nat) lemma adiff_self_eq_0_nat [simp]: assumes m: "m \ Nat" shows "m -- m = 0" -using m apply induct by auto +using m by induct auto text \ Associativity \ lemma adiff_adiff_left_nat: assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" shows "(i -- j) -- k = i -- (j + k)" -using i j apply (rule diffInduct) -using k by auto +using i j by (rule diffInduct) (auto simp: k) lemma Succ_adiff_adiff_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" @@ -467,8 +469,7 @@ text \ Cancellation rules \ lemma adiff_add_inverse_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(n + m) -- n = m" -using n apply induct -using assms by auto +using n by induct (simp_all add: m) lemma adiff_add_inverse2_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" @@ -478,8 +479,7 @@ using assms by (simp add: add_commute_nat [of m n]) lemma adiff_cancel_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(k + m) -- (k + n) = m -- n" -using k apply induct -using assms by simp_all +using k by induct (simp_all add: m n) lemma adiff_cancel2_nat [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" @@ -489,23 +489,21 @@ using assms by (simp add: add_commute_nat) lemma adiff_add_0_nat [simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "n -- (n + m) = 0" -using n apply induct -using assms by simp_all +using n by induct (simp_all add: m) text \ Difference distributes over multiplication \ lemma adiff_mult_distrib_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "(m -- n) * k = (m * k) -- (n * k)" -using m n apply(rule diffInduct) -using k by simp_all +using m n by (rule diffInduct) (simp_all add: k) lemma adiff_mult_distrib2_nat [algebra_simps]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" shows "k * (m -- n) = (k * m) -- (k * n)" using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) -(* -- "NOT added as rewrites, since sometimes they are used from right-to-left" *) +\ \ NOT added as rewrites, since sometimes they are used from right to left \ lemmas nat_distrib = add_mult_distrib_right_nat add_mult_distrib_left_nat adiff_mult_distrib_nat adiff_mult_distrib2_nat @@ -525,137 +523,31 @@ lemma nat_add_left_cancel_leq [simp]: shows "(k + m \ k + n) = (m \ n)" using assms by (induct k) simp_all -lemma nat_add_left_cancel_less [simp]: +lemma nat_add_right_cancel_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m < k + n) = (m < n)" -using k apply induct -using assms by simp_all + shows "(m + k \ n + k) = (m \ n)" +using assms by (induct k) simp_all -lemma nat_add_right_cancel_less [simp]: +lemma nat_add_left_cancel_Succ_leq [simp]: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k < n + k) = (m < n)" -using k apply induct -using assms by simp_all + shows "(Succ[k + m] \ k + n) = (Succ[m] \ n)" +using assms by (induct k) simp_all -lemma nat_add_right_cancel_leq [simp]: +(* immediate corollary *) +lemma nat_add_left_cancel_less: assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k \ n + k) = (m \ n)" -using k apply induct -using assms by simp_all + shows "(k + m < k + n) = (m < n)" +using assms by simp -lemma add_gr_0 [simp]: - assumes ndom: "n \ Nat" and mdom: "m \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" - proof - - have s1_1: "m + n > 0 ==> m > 0 \ n > 0" - proof - - assume s1_1_asm: "m + n > 0" - have s2_1: "m \ 0 \ n \ 0" - proof - - have s3_1: "m = 0 \ n = 0 ==> FALSE" - proof - - assume s3_1_asm: "m = 0 \ n = 0" - have s4_1: "m + n = 0" - using s3_1_asm ndom mdom add_is_0_nat by auto - have s4_2: "m + n \ 0" - unfolding less_def using s1_1_asm by auto - show "FALSE" - using s4_1 s4_2 by auto - qed - show ?thesis - using s3_1 by auto - qed - show "m > 0 \ n > 0" - using s2_1 ndom mdom nat_gt0_not0 by auto - qed - have s1_2: "m > 0 \ n > 0 ==> m + n > 0" - proof - - assume s1_2_asm: "m > 0 \ n > 0" - have s2_1: "m > 0 ==> m + n > 0" - proof - - assume s2_1_asm: "m > 0" - have s3_1: "m + n > 0 + n" - proof - - have s4_1: "0 < m" - using s2_1_asm by auto - have s4_2: "(0 + n < m + n) = (0 < m)" - using ndom mdom zeroIsNat nat_add_right_cancel_less - by blast - have s4_3: "0 + n < m + n" - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - have s3_2: "0 + n = n" - using ndom add_0_nat add_commute_nat zeroIsNat - by auto - have s3_3: "n >= 0" - using ndom nat_zero_leq by auto - have s3_4: "m + n >= 0 + n" - using s3_1 by (auto simp: less_def) - have s3_5: "m + n >= n" - using s3_4 s3_2 by auto - have s3_6: "m + n >= 0" - using s3_5 s3_3 nat_leq_trans - zeroIsNat ndom mdom addIsNat by auto - have s3_7: "m + n = 0 ==> FALSE" - proof - - assume s3_7_asm: "m + n = 0" - have s4_1: "m = 0 \ n = 0" - using s3_7_asm mdom ndom add_is_0_nat - by auto - show "FALSE" - using s4_1 s2_1_asm by (auto simp: less_def) - qed - show "m + n > 0" - unfolding less_def - using s3_6 s3_7 by auto - qed - have s2_2: "n > 0 ==> m + n > 0" - proof - - assume s2_2_asm: "n > 0" - have s3_1: "m + n > m + 0" - proof - - have s4_1: "0 < n" - using s2_2_asm by auto - have s4_2: "(m + 0 < m + n) = (0 < n)" - using ndom mdom zeroIsNat nat_add_left_cancel_less - by blast - have s4_3: "m + 0 < m + n" - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - have s3_2: "m + 0 = m" - using mdom add_0_nat by auto - have s3_3: "m >= 0" - using mdom nat_zero_leq by auto - have s3_4: "m + n >= m + 0" - using s3_1 by (auto simp: less_def) - have s3_5: "m + n >= m" - using s3_4 s3_2 by auto - have s3_6: "m + n >= 0" - using s3_5 s3_3 nat_leq_trans - zeroIsNat ndom mdom addIsNat by auto - have s3_7: "m + n = 0 ==> FALSE" - proof - - assume s3_7_asm: "m + n = 0" - have s4_1: "m = 0 \ n = 0" - using s3_7_asm mdom ndom add_is_0_nat - by auto - show "FALSE" - using s4_1 s2_2_asm by (auto simp: less_def) - qed - show "m + n > 0" - unfolding less_def - using s3_6 s3_7 by auto - qed - show "m + n > 0" - using s1_2_asm s2_1 s2_2 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed +lemma nat_add_right_cancel_Succ_less [simp]: + assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" + shows "(Succ[m + k] \ n + k) = (Succ[m] \ n)" +using assms by (induct k) simp_all + +lemma nat_add_right_cancel_less: + assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" + shows "(m + k < n + k) = (m < n)" +using assms by simp lemma less_imp_Succ_add: assumes m: "m \ Nat" and n: "n \ Nat" @@ -687,9 +579,9 @@ using assms by simp_all subsubsection \ (Partially) Ordered Groups \ -(* -- "The two following lemmas are just ``one half'' of - @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above." *) +\ \ The two following lemmas are just ``one half'' of + @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} + proved above." \ lemma add_leq_left_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a \ b \ c + a \ c + b" @@ -710,41 +602,16 @@ using assms nat_leq_trans[of "a + c" "b + c" "b + d"] by simp -(* -- "Similar for strict less than." *) +\ \ Similar for strict less than. \ lemma add_less_left_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + assumes "a \ Nat" and "b \ Nat" and "c \ Nat" shows "a < b \ c + a < c + b" - proof - - have s1_1: "a <= b ==> c + a <= c + b" - using a b c add_leq_left_mono by auto - have s1_2: "a \ b ==> (c + a) \ (c + b)" - using a b c add_left_cancel_nat by auto - have s1_3: "a < b ==> c + a < c + b" - proof - - assume s1_3_asm: "a < b" - have s2_1: "a <= b" - using s1_3_asm by (auto simp: less_def) - have s2_2: "a \ b" - using s1_3_asm by (auto simp: less_def) - have s2_3: "c + a <= c + b" - using s1_1 s2_1 by auto - have s2_4: "(c + a) \ (c + b)" - using s1_2 s2_2 by auto - show "c + a < c + b" - using s2_3 s2_4 by (auto simp: less_def) - qed - show "a < b \ c + a < c + b" - using a b c s1_3 by auto - qed -(* -using assms add_leq_left_mono[OF "a" "b" "c"] add_left_cancel_nat[OF "c" "b" "a"] - by (auto simp: nat_less_Succ_iff_leq nat_gt0_not0) -*) + using assms by (simp add: add_leq_left_mono less_def) lemma add_less_right_mono: assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a < b \ a + c < b + c" -using assms add_less_left_mono add_commute_nat by auto +using assms by (auto simp: add_less_left_mono add_commute_nat) text \ Strict monotonicity in both arguments \ lemma add_less_mono: @@ -784,6 +651,36 @@ lemma leq_add2 [simp]: shows "n \ m + n" using assms add_leq_right_mono [of 0 m n] by simp +lemma trans_leq_add1 [elim]: + assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m+k" +using assms by (simp add: nat_leq_trans) + +lemma trans_leq_add2 [elim]: + assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ k+m" +using assms by (simp add: nat_leq_trans) + +lemma add_leqD1 [elim]: + assumes "n+k \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m" +using assms by (simp add: nat_leq_trans[of "n" "n+k" "m"]) + +lemma add_leqD2 [elim]: + assumes "k+n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "n \ m" +using assms by (simp add: nat_leq_trans[of "n" "k+n" "m"]) + +lemma add_Succ_leqD1 [elim]: + assumes "Succ[n+k] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "Succ[n] \ m" +using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[n+k]" "m"]) + +lemma add_Succ_leqD2 [elim]: + assumes "Succ[k+n] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" + shows "Succ[n] \ m" +using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[k+n]" "m"]) + lemma less_add_Succ1: assumes "i \ Nat" and "m \ Nat" shows "i < Succ[i + m]" @@ -799,99 +696,88 @@ lemma less_iff_Succ_add: shows "(m < n) = (\k \ Nat: n = Succ[m + k])" using assms by (auto intro!: less_imp_Succ_add) -lemma trans_leq_add1: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i \ j + m" -using assms by (auto elim: nat_leq_trans) - -lemma trans_leq_add2: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i \ m + j" -using assms by (auto elim: nat_leq_trans) - lemma trans_less_add1: - assumes ij: "i < j" and idom: "i \ Nat" and - jdom: "j \ Nat" and mdom: "m \ Nat" + assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < j + m" - proof - - have s1_1: "i <= i + m" - proof - - have s2_1: "0 <= m" - using mdom nat_zero_leq by auto - have s2_2: "i + 0 <= i + m" - using idom zeroIsNat mdom s2_1 nat_add_left_cancel_leq by auto - have s2_3: "i + 0 = i" - using idom add_0_nat by auto - show ?thesis - using s2_2 s2_3 by auto - qed - have s1_2: "i + m < j + m" - using ij idom jdom mdom add_less_right_mono by auto - have s1_3: "i + m \ Nat" - using idom mdom addIsNat by auto - have s1_4: "j + m \ Nat" - using jdom mdom addIsNat by auto - show ?thesis - using s1_1 s1_2 idom s1_3 s1_4 nat_leq_less_trans - by auto - qed +using assms by auto lemma trans_less_add2: assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" shows "i < m + j" -using assms trans_less_add1 add_commute_nat by auto +using assms by auto lemma add_lessD1: assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" shows "i < k" -using assms by (intro nat_leq_less_trans[of "i" "i+j" "k"], simp+) +using assms by auto -lemma not_add_less1 [simp]: +lemma add_lessD2: + assumes "j+i < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" + shows "i < k" +using assms by auto + +lemma not_Succ_add_le1: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(Succ[i + j] \ i) = FALSE" +using assms by simp + +lemma not_Succ_add_le2: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(Succ[j + i] \ i) = FALSE" +using assms by simp + +lemma not_add_less1: assumes i: "i \ Nat" and j: "j \ Nat" shows "(i + j < i) = FALSE" -by (auto dest: add_lessD1[OF _ i j i]) +using assms by simp -lemma not_add_less2 [simp]: +lemma not_add_less2: assumes "i \ Nat" and "j \ Nat" shows "(j + i < i) = FALSE" -using assms by (simp add: add_commute_nat) - -lemma add_leqD1: - assumes "m + k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "m \ n" -using assms by (intro nat_leq_trans[of "m" "m+k" "n"], simp+) - -lemma add_leqD2: - assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k \ n" -using assms by (intro nat_leq_trans[of "k" "m+k" "n"], simp+) +using assms by simp lemma add_leqE: assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "(m \ n \ k \ n \ R) \ R" using assms by (blast dest: add_leqD1 add_leqD2) -lemma leq_add_left_false [simp]: - assumes mdom: "m \ Nat" and ndom: "n \ Nat" and n0: "n \ 0" - shows "m + n \ m = FALSE" - proof - - have s1_1: "m + n \ m = (m + n < m \ m + n = m)" - proof - - define p where "p \ m + n" - have s2_1: "p \ m = (p < m \ p = m)" - using mdom nat_leq_less by auto - show ?thesis - using s2_1 by (auto simp: p_def) - qed - have s1_2: "(m + n = m) = (n = 0)" - using mdom ndom add_eq_self_zero_nat[of "m" "n"] by auto - have s1_3: "(m + n <= m) = (m + n < m)" - using n0 s1_1 s1_2 by auto - have s1_4: "(m + n < m) = FALSE" - using mdom ndom not_add_less1[of "m" "n"] by auto - show ?thesis - using s1_3 s1_4 by auto - qed +lemma leq_add_self1 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(m+n \ m) = (n=0)" +proof - + { + assume "m+n \ m" + with assms have "m+n = m" by (intro nat_leq_antisym) simp_all + } + with assms show ?thesis by auto +qed + +lemma leq_add_self2 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(n+m \ m) = (n=0)" +proof - + { + assume "n+m \ m" + with assms have "n+m = m" by (intro nat_leq_antisym) simp_all + } + with assms show ?thesis by auto +qed + +lemma add_gt_0: + assumes "m \ Nat" and "n \ Nat" + shows "(m + n > 0) = (m > 0 \ n > 0)" +proof - + from assms have "m + n \ Nat" by simp + hence "(m + n > 0) = (m + n \ 0)" by (rule nat_gt0_not0) + also from assms have "\ = (m \ 0 \ n \ 0)" by simp + finally show ?thesis using assms by (simp only: nat_gt0_not0) +qed + +(* The above lemma follows from the following simplification rule. *) +lemma add_ge_1 [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(m + n \ 1) = (m \ 1 \ n \ 1)" + using assms by (auto simp add: add_gt_0) subsubsection \ More results about arithmetic difference \ @@ -936,33 +822,31 @@ lemma leq_iff_add: assumes m: "m \ Nat" and n: "n \ Nat" shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") proof - - have 1: "?lhs \ ?rhs" + have "?lhs \ ?rhs" proof assume mn: "m \ n" with m n have "n = m + (n -- m)" by simp with m n show "?rhs" by blast qed - from assms have 2: "?rhs \ ?lhs" by auto - from 1 2 assms show ?thesis by blast + moreover + from assms have "?rhs \ ?lhs" by auto + ultimately show ?thesis by blast qed lemma less_imp_adiff_less: assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" shows "j -- n < k" - proof - - have s1_1: "j -- n \ Nat" - using j n by simp+ - have s1_2: "j -- n \ j" - using j n by simp+ - show ?thesis - using j k jk s1_1 s1_2 nat_leq_less_trans by auto - qed +proof - + from j n have "j -- n \ Nat" by simp + moreover + from j n have s1_2: "j -- n \ j" by simp + ultimately show ?thesis using j k jk nat_leq_less_trans by auto +qed lemma adiff_Succ_less (*[simp]*): assumes i: "i \ Nat" and n: "n \ Nat" shows "0 < n \ n -- Succ[i] < n" -using n apply cases -using i by auto +using n by cases (auto simp: i) lemma adiff_add_assoc: assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" @@ -1003,23 +887,24 @@ lemma adiff_is_0_eq' (*[simp]*): shows "m -- n = 0" using assms by simp -lemma zero_less_adiff [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(0 < n -- m) = (m < n)" +lemma one_leq_adiff [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(1 \ n -- m) = (Succ[m] \ n)" by (induct m n rule: diffInduct, simp_all add: assms) +lemma zero_less_adiff: + assumes "m \ Nat" and "n \ Nat" + shows "(0 < n -- m) = (m < n)" +using assms by simp + lemma less_imp_add_positive: - assumes ij: "i < j" and i: "i \ Nat" and j: "j \ Nat" + assumes "i < j" and "i \ Nat" and "j \ Nat" shows "\k \ Nat: 0 < k \ i + k = j" proof - - have s1_1: "i + (j -- i) = j" - using i j ij le_add_adiff_inverse by (auto simp: less_def) - have s1_2: "j -- i \ Nat" - using assms by simp+ - have s1_3: "0 < j -- i" - using assms zero_less_adiff by auto - show ?thesis - using s1_1 s1_2 s1_3 by blast + thm leq_iff_add + from assms obtain k where k: "k \ Nat" "j = Succ[i]+k" + by (auto simp: leq_iff_add) + with assms show ?thesis by force qed lemma leq_adiff_right_add_left: @@ -1043,16 +928,10 @@ lemma leq_adiff_trans: apply(rule nat_leq_trans[of "p -- n" p m]) using assms adiff_leq_self[OF p n] by simp_all -lemma leq_adiff_right_false [simp]: - assumes "n \ 0" "n \ m" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ m -- n = FALSE" -using assms by (simp add: leq_adiff_right_add_left[OF _ m m n]) - -lemma leq_adiff_right_imp_0: - assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" - shows "p = 0" -using p h apply (induct) -using n by auto +lemma leq_adiff_right [simp]: + assumes "n \ m" and "m \ Nat" and "n \ Nat" + shows "(m \ m -- n) = (n = 0)" + using assms by (simp add: leq_adiff_right_add_left) subsubsection \ Monotonicity of Multiplication \ @@ -1062,14 +941,12 @@ subsubsection \ Monotonicity of Multiplication \ lemma mult_leq_left_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "c * a \ c * b" -using c apply induct -using 1 a b by (simp_all add: add_leq_mono) +using c by induct (simp_all add: add_leq_mono 1 a b) lemma mult_leq_right_mono: assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" shows "a * c \ b * c" -using c apply induct -using 1 a b by (simp_all add: add_leq_mono add_commute_nat) +using c by induct (simp_all add: add_leq_mono add_commute_nat 1 a b) text \ @{text "\"} monotonicity, BOTH arguments \ @@ -1081,280 +958,193 @@ using assms mult_leq_right_mono[OF _ i j k] mult_leq_left_mono[OF _ k l j] nat_leq_trans[of "i * k" "j * k" "j * l"] -by simp + by simp + +lemma self_leq_mult_left: + assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" + shows "a \ a * b" +proof - + have "1 \ Nat" by simp + from 1 this b a have "a * 1 \ a * b" by (rule mult_leq_left_mono) + with a b show ?thesis by simp +qed -text \ strict, in 1st argument \ +lemma self_leq_mult_right: + assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" + shows "a \ b * a" +using assms by (simp add: self_leq_mult_left mult_commute_nat) -lemma mult_less_left_mono: - assumes ij: "i < j" and k0: "0 < k" and - idom: "i \ Nat" and jdom: "j \ Nat" and kdom: "k \ Nat" - shows "k * i < k * j" - proof - - have s1_1: "Succ[0] * i < Succ[0] * j" - using mult_1_left_nat ij idom jdom by auto - have s1_2: "\ n. \ - n \ Nat; - Succ[n] * i < Succ[n] * j - \ \ - Succ[Succ[n]] * i < Succ[Succ[n]] * j" - proof - - fix "n" :: "c" - assume s1_2_ndom: "n \ Nat" - assume s1_2_induct: "Succ[n] * i < Succ[n] * j" - have s2_1: "Succ[n] * i + i < Succ[n] * j + j" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "Succ[n] * i \ Nat" - using s3_1 idom multIsNat by auto - have s3_3: "Succ[n] * j \ Nat" - using s3_1 jdom multIsNat by auto - show ?thesis - using s3_2 s3_3 idom jdom s1_2_induct ij - add_less_mono by auto - qed - have s2_2: "i * Succ[n] + i < j * Succ[n] + j" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "Succ[n] * i = i * Succ[n]" - using s3_1 idom mult_commute_nat by auto - have s3_3: "Succ[n] * j = j * Succ[n]" - using s3_1 jdom mult_commute_nat by auto - show ?thesis - using s2_1 s3_2 s3_3 by auto - qed - have s2_3: "i * Succ[Succ[n]] < j * Succ[Succ[n]]" - proof - - have s3_1: "Succ[n] \ Nat" - using s1_2_ndom succIsNat by auto - have s3_2: "i * Succ[Succ[n]] = i * Succ[n] + i" - using idom s3_1 multnat_Succ by (auto simp: nat_mult_def) - have s3_3: "j * Succ[Succ[n]] = j * Succ[n] + j" - using jdom s3_1 multnat_Succ by (auto simp: nat_mult_def) - show ?thesis - using s2_2 s3_2 s3_3 by auto - qed - have s2_4: "i * Succ[Succ[n]] = Succ[Succ[n]] * i" - proof - - have s3_1: "Succ[Succ[n]] \ Nat" - using s1_2_ndom succIsNat by auto - show ?thesis - using s3_1 idom mult_commute_nat by auto - qed - have s2_5: "j * Succ[Succ[n]] = Succ[Succ[n]] * j" - proof - - have s3_1: "Succ[Succ[n]] \ Nat" - using s1_2_ndom succIsNat by auto - show ?thesis - using s3_1 jdom mult_commute_nat by auto - qed - show "Succ[Succ[n]] * i < Succ[Succ[n]] * j" - using s2_3 s2_4 s2_5 by auto - qed - have s1_3: "\ n \ Nat: Succ[n] * i < Succ[n] * j" - using s1_1 s1_2 - natInduct[of "\ q. Succ[q] * i < Succ[q] * j"] by auto - show ?thesis - using s1_3 nat_gt0_iff_Succ[of "k"] k0 kdom by auto - qed +text \ Similar lemmas for @{text "<"} \ + +lemma mult_Succ_leq_right_mono1: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + shows "Succ[a * Succ[c]] \ b * Succ[c]" +using c proof induct + case 0 + with assms show ?case by simp +next + fix n + assume n: "n \ Nat" and ih: "Succ[a * Succ[n]] \ b * Succ[n]" + from ab a b have 1: "a \ b" + by (simp add: nat_leq_trans[of "a" "Succ[a]" "b"]) + from a n have "Succ[a * Succ[Succ[n]]] = Succ[a * Succ[n]] + a" + by simp + also from a b n ih have "\ \ b* Succ[n] + a" + by simp + also from a b n 1 have "\ \ b * Succ[n] + b" + by simp + finally show "Succ[a * Succ[Succ[n]]] \ b * Succ[Succ[n]]" + using b n by simp +qed + +lemma mult_Succ_leq_right_mono2: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" + shows "Succ[a * c] \ b * c" +proof - + from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" + by (blast dest: nat_ge1_implies_Succ) + with ab a b show ?thesis by (auto dest: mult_Succ_leq_right_mono1[of a b k]) +qed lemma mult_less_right_mono: - assumes 1: "i < j" "0 < k" and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "i * k < j * k" - proof - - have s1_1: "k * i < k * j" - using 1 i j k mult_less_left_mono by auto - have s1_2: "k * i = i * k" - using i k mult_commute_nat by auto - have s1_3: "k * j = j * k" - using j k mult_commute_nat by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed + assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" + shows "a * c < b * c" +using assms by (simp add: mult_Succ_leq_right_mono2) -lemma nat_0_less_mult_iff [simp]: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(0 < i * j) = (0 < i \ 0 < j)" -using i apply induct - using j apply simp - using j apply(induct, simp_all) -done +lemma mult_Succ_leq_left_mono1: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" + shows "Succ[Succ[c] * a] \ Succ[c] * b" +proof - + have "Succ[Succ[c] * a] = Succ[a * Succ[c]]" + using a c by (simp add: mult_commute_nat) + also have "\ \ b * Succ[c]" + using assms by (rule mult_Succ_leq_right_mono1) + also have "\ = Succ[c] * b" + using b c by (simp add: mult_commute_nat) + finally show ?thesis . +qed -lemma one_leq_mult_iff (*[simp]*): +lemma mult_Succ_leq_left_mono2: + assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" + shows "Succ[c * a] \ c * b" +proof - + from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" + by (blast dest: nat_ge1_implies_Succ) + with ab a b show ?thesis by (auto dest: mult_Succ_leq_left_mono1[of a b k]) +qed + +lemma mult_less_left_mono: + assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" + shows "c * a < c * b" +using assms by (simp add: mult_Succ_leq_left_mono2) + +lemma one_leq_mult_iff[simp]: assumes m: "m \ Nat" and n: "n \ Nat" shows "(1 \ m * n) = (1 \ m \ 1 \ n)" using assms nat_gt0_not0 by simp -lemma mult_less_cancel_left [simp]: +lemma nat_0_less_mult_iff: + assumes i: "i \ Nat" and j: "j \ Nat" + shows "(0 < i * j) = (0 < i \ 0 < j)" +using assms by simp + +lemma mult_Succ_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m < k * n) = (0 < k \ m < n)" -proof (auto intro!: mult_less_left_mono[OF _ _ m n k]) - assume "k*m < k*n" - from k m n this show "0 < k" by (cases k, simp_all) + shows "(Succ[k * m] \ k * n) = (1 \ k \ Succ[m] \ n)" +proof (auto intro: mult_Succ_leq_left_mono2[OF _ m n k _]) + assume "Succ[k * m] \ k * n" + from k m n this show "1 \ k" by (cases k) simp_all next - assume 1: "k*m < k*n" - show "m < n" + assume 1: "Succ[k * m] \ k * n" + show "Succ[m] \ n" proof (rule contradiction) - assume "\(m < n)" + assume "\(Succ[m] \ n)" with m n k have "k*n \ k*m" by (simp add: nat_not_order_simps mult_leq_left_mono) - with m n k have "\(k*m < k*n)" by (simp add: nat_not_order_simps) + with m n k have "\(Succ[k*m] \ k*n)" by (simp add: nat_not_order_simps) with 1 show "FALSE" by simp qed qed -lemma mult_less_cancel_right [simp]: +lemma mult_less_cancel_left: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k < n * k) = (0 < k \ m < n)" -proof (auto intro!: mult_less_right_mono[OF _ _ m n k]) - assume "m*k < n*k" - from k m n this show "0 < k" by (cases k, simp_all) + shows "(k * m < k * n) = (0 < k \ m < n)" +using assms by simp + +lemma mult_Succ_leq_cancel_right [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" + shows "(Succ[m * k] \ n * k) = (1 \ k \ Succ[m] \ n)" +proof (auto intro: mult_Succ_leq_right_mono2[OF _ m n k _]) + assume "Succ[m * k] \ n * k" + from k m n this show "1 \ k" by (cases k) simp_all next - assume 1: "m*k < n*k" - show "m < n" + assume 1: "Succ[m * k] \ n * k" + show "Succ[m] \ n" proof (rule contradiction) - assume "\(m < n)" + assume "\(Succ[m] \ n)" with m n k have "n*k \ m*k" by (simp add: nat_not_order_simps mult_leq_right_mono) - with m n k have "\(m*k < n*k)" by (simp add: nat_not_order_simps) + with m n k have "\(Succ[m*k] \ n*k)" by (simp add: nat_not_order_simps) with 1 show "FALSE" by simp qed qed -lemma mult_less_self_left [dest]: - assumes less: "n * k < n" and n: "n \ Nat" and k: "k \ Nat" +lemma mult_less_cancel_right: + assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" + shows "(m * k < n * k) = (0 < k \ m < n)" +using assms by simp + +lemma mult_Succ_leq_self_left [dest]: + assumes nk: "Succ[n*k] \ n" and n: "n \ Nat" and k: "k \ Nat" + shows "k = 0" +proof (rule contradiction) + assume "k \ 0" + with k nat_neq0_conv have "1 \ k" by simp + with n k have "n \ n * k" by (rule self_leq_mult_left) + with n k have "\(Succ[n*k] \ n)" by (simp add: nat_not_order_simps) + from this nk show "FALSE" .. +qed + +lemma mult_less_self_left: + assumes "n * k < n" and "n \ Nat" and "k \ Nat" shows "k=0" - proof - - have s1_1: "1 \ k ==> FALSE" - proof - - assume s1_1_asm: "1 \ k" - have s2_1: "n * 1 \ n * k" - using n k oneIsNat s1_1_asm mult_leq_left_mono[of "1" "k" "n"] - by auto - have s2_2: "n \ n * k" - proof - - have s3_1: "n * 1 = n" - using n mult_1_right_nat by auto - show ?thesis - using s2_1 s3_1 by auto - qed - define p where "p \ n" - define q where "q \ n * k" - have s2_3: "q < p" - unfolding p_def q_def using less by auto - have s2_4: "n < p" - proof - - have s3_1: "q \ Nat" - unfolding q_def using n k multIsNat by auto - have s3_2: "p \ Nat" - unfolding p_def using n by auto - have s3_3: "n \ q" - unfolding q_def using s2_2 by auto - show ?thesis - using s3_1 s3_2 n s3_3 s2_3 - nat_leq_less_trans[of "n" "q" "p"] - by auto - qed - have s2_5: "n \ n" - using s2_4 by (auto simp: less_def p_def) - show "FALSE" - using s2_5 by auto - qed - have s1_2: "\ (1 \ k)" - using s1_1 by auto - show ?thesis - using s1_2 k nat_not_leq_one[of "k"] by auto - qed +using assms by auto -lemma mult_less_self_right [dest]: +lemma mult_Succ_leq_self_right [dest]: + assumes less: "Succ[k*n] \ n" and n: "n \ Nat" and k: "k \ Nat" + shows "k=0" +using assms by (auto simp: mult_commute_nat[OF k n]) + +lemma mult_less_self_right: assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" shows "k=0" - proof - - have s1_1: "n * k < n" - proof - - have s2_1: "k * n = n * k" - using n k mult_commute_nat by auto - show ?thesis - using less s2_1 by auto - qed - show ?thesis - using s1_1 n k mult_less_self_left by auto - qed +using assms by auto lemma mult_leq_cancel_left [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(k * m \ k * n) = (k = 0 \ m \ n)" - proof - - have s1_1: "(k * m \ k * n) ==> (k = 0 \ m \ n)" - proof - - assume s1_1_asm: "k * m \ k * n" - have s2_1: "\ (k = 0) ==> m \ n" - proof - - assume s2_1_asm: "\ (k = 0)" - have s3_1: "k > 0" - using k nat_gt0_not0 s2_1_asm by auto - have s3_2: "\(m \ n) ==> FALSE" - proof - - assume s3_2_asm: "\(m \ n)" - have s2_1: "n < m" - using s3_2_asm m n nat_not_leq[of "m" "n"] by auto - have s2_2: "k * n < k * m" - using s3_1 s2_1 m n k mult_less_left_mono[of "n" "m" "k"] by auto - with m n k have s2_3: "\ (k * m \ k * n)" - by (simp add: nat_not_order_simps) - show "FALSE" - using s2_2 s2_3 s1_1_asm by auto - qed - show "m \ n" - using s3_2 by auto - qed - show "k = 0 \ m \ n" - using s2_1 by auto - qed - have s1_2: "(k = 0 \ m \ n) ==> (k * m \ k * n)" - proof - - assume s1_2_asm: "k = 0 \ m \ n" - have s2_1: "k = 0 ==> k * m \ k * n" - proof - - assume s2_1_asm: "k = 0" - have s3_1: "k * m = 0" - using s2_1_asm m mult_0_left_nat by auto - have s3_2: "k * n = 0" - using s2_1_asm n mult_0_left_nat by auto - show "k * m \ k * n" - using s3_1 s3_2 by auto - qed - have s2_2: "m \ n ==> k * m \ k * n" - proof - - assume s2_2_asm: "m \ n" - show "k * m \ k * n" - using s2_2_asm m n k mult_leq_left_mono by auto - qed - show "k * m \ k * n" - using s1_2_asm s2_1 s2_2 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed +proof - + { + assume 1: "k*m \ k*n" and 2: "k \ 0" and 3: "\(m \ n)" + from k 2 nat_gt0_not0 have "k > 0" by simp + with 3 k m n have "\(k * m \ k * n)" by (simp add: nat_not_order_simps) + from this 1 have "FALSE" .. + } + moreover have "k=0 \ k*m \ k*n" + using assms by simp + moreover have "m \ n \ k*m \ k*n" + using assms by (simp add: mult_leq_left_mono) + ultimately show ?thesis by blast +qed lemma mult_leq_cancel_right [simp]: assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" shows "(m * k \ n * k) = (k = 0 \ m \ n)" - proof - - have s1_1: "(k * m \ k * n) = (k = 0 \ m \ n)" - using m n k mult_leq_cancel_left by auto - have s1_2: "k * m = m * k" - using k m mult_commute_nat by auto - have s1_3: "k * n = n * k" - using k n mult_commute_nat by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed +using assms by (simp add: mult_commute_nat) lemma Suc_mult_less_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" shows "(Succ[k] * m < Succ[k] * n) = (m < n)" using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] - nat_zero_less_Succ[of "k"] by auto + nat_zero_less_Succ[of "k"] by auto lemma Suc_mult_leq_cancel1: assumes "m \ Nat" and "n \ Nat" and "k \ Nat" @@ -1377,7 +1167,7 @@ lemma mult_eq_self_implies_10: shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") proof - from assms have "(m*n = m) = (m*n = m*1)" by simp - also have "... = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) + also have "\ = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) finally show ?thesis . qed diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy index f0407ee6..97e45a2f 100644 --- a/isabelle/NatOrderings.thy +++ b/isabelle/NatOrderings.thy @@ -1,6 +1,6 @@ (* Title: TLA+/NatOrderings.thy Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) @@ -34,6 +34,11 @@ notation (ASCII) leq (infixl "<=" 50) and geq (infixl ">=" 50) +lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" +by (rule ssubst) + +lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" +by (rule subst) subsection \ Operator definitions and generic facts about @{text "<"} \ @@ -50,6 +55,12 @@ by (simp add: less_def) lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" +by (rule ssubst) + +lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" +by (rule subst) + lemma less_imp_leq [elim!]: "a < b \ a \ b" unfolding less_def by simp @@ -98,16 +109,22 @@ unfolding nat_leq_def by (rule zeroInUpto) lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" by (simp add: nat_leq_def uptoZero) -lemma nat_leq_SuccI [elim!(*,simp*)]: +lemma nat_leq_SuccI [(*elim!,*)intro]: assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" shows "m \ Succ[n]" using assms by (auto simp: nat_leq_def uptoSucc) +(* Do not make this "simp": interferes with other rules, e.g. simplifying "Succ[m] \ Succ[n]" *) lemma nat_leq_Succ: assumes (*"m \ Nat" and*) "n \ Nat" shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" using assms by (auto simp: nat_leq_def uptoSucc) +lemma nat_is_leq_Succ [simp]: + assumes "n \ Nat" + shows "n \ Succ[n]" +using assms by (simp add: nat_leq_Succ) + lemma nat_leq_SuccE [elim]: assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" and "m \ n \ P" and "m = Succ[n] \ P" @@ -189,7 +206,7 @@ lemma nat_leq_induct: \ \sometimes called ``complete induction'' shows "\n\Nat : P(n)" proof - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp add: nat_leq_Succ) + by (intro natInduct, auto simp: nat_leq_Succ) thus ?thesis by (blast dest: nat_leq_refl) qed @@ -210,8 +227,6 @@ using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) text \ Reduce @{text "\"} to @{text "<"}. \ lemma nat_leq_less: - fixes "m" :: "c" - fixes "n" :: "c" assumes (*"m \ Nat" and*) "n \ Nat" shows "m \ n = (m < n \ m = n)" using assms by (auto simp: less_def) @@ -301,9 +316,6 @@ proof - qed lemma nat_leq_less_trans [trans]: - fixes "k" :: "c" - fixes "m" :: "c" - fixes "n" :: "c" assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" shows "k < n" using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) @@ -522,7 +534,7 @@ lemma nat_ge1_implies_Succ: lemma nat_gt0_iff_Succ: assumes n: "n \ Nat" - shows "(0 < n) = (\m \ Nat: n = Succ[m])" + shows "(1 \ n) = (\m \ Nat: n = Succ[m])" using n by (auto dest: nat_ge1_implies_Succ) (* diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy index 1f19fe45..90fb821e 100644 --- a/isabelle/Peano.thy +++ b/isabelle/Peano.thy @@ -1,6 +1,6 @@ (* Title: TLA+/Peano.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index 2a23721c..afade712 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -1,8 +1,8 @@ (* Title: TLA+/PredicateLogic.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2017 + Version: Isabelle2020 *) section \First-Order Logic for TLA+\ @@ -1079,11 +1079,6 @@ ML \ \ -(* -setup hypsubst_setup -setup Classical.setup -*) - declare refl [intro!] and trueI [intro!] and conjI [intro!] @@ -1124,8 +1119,6 @@ ML \ ); \ -(* setup Blast.setup *) - (*** TEST DATA *** lemma "x = y \ y = x" @@ -1615,7 +1608,7 @@ text \ lemma boolifyFalseI [intro]: "\ A \ boolify(A) = FALSE" by (auto simp add: boolify_def) -text \ idempotence of @{text boolify} \ +\ \ idempotency of @{text boolify} \ lemma boolifyBoolify [simp]: "boolify(boolify(x)) = boolify(x)" by (auto simp add: boolify_def) diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index 8729d7e5..707897f4 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -1,6 +1,6 @@ (* Title: TLA+/SetTheory.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2020 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD Version: Isabelle2020 *) @@ -151,13 +151,11 @@ syntax "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ \ _ :/ _)" [100,0,0]10) "@bEx" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) "@bAll" :: "[cidts, c, c] \ c" ("(3\_ \ _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) syntax (ASCII) (* Again, only single variable for bounded choice. *) "@bChoice" :: "[idt, c, c] \ c" ("(3CHOOSE _ in _ :/ _)" [100,0,0] 10) - "@bEx" :: "[cidts, c, c] \ c" ("(3EX _ in _ :/ _)" [100,0,0] 10) - "@bAll" :: "[cidts, c, c] \ c" ("(3ALL _ in _ :/ _)" [100,0,0] 10) + "@bEx" :: "[cidts, c, c] \ c" ("(3\\E _ \\in _ :/ _)" [100,0,0] 10) + "@bAll" :: "[cidts, c, c] \ c" ("(3\\A _ \\in _ :/ _)" [100,0,0] 10) translations "CHOOSE x \ S : P" \ "CONST bChoose(S, \x. P)" (* the following cannot be a print macro because the RHS is non-linear From f0ac8cc6dfc7b65b725770c4ca30303af5a71036 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Fri, 22 Jan 2021 12:15:32 +0200 Subject: [PATCH 108/167] API: update to keywords of Isabelle 2021 The script `isabelle_keywords_get.sh` is outdated, because the file it gets the keywords from is not part of recent Isabelle distributions. Instead, a Python script `isabelle_keywords_update.py` has been added, which relies on the Scala script `outer_keywords.scala`, which gets the Isabelle 2021 keywords for the logic `Pure`. rebasing branch isabelle2020 on master: removing src/isabelle_keywords_get.sh, modified src/isabelle_keywords.ml as in master --- src/isabelle_keywords.ml | 547 +++++++++++++------------------- src/isabelle_keywords_get.sh | 20 -- src/isabelle_keywords_update.py | 22 ++ 3 files changed, 238 insertions(+), 351 deletions(-) delete mode 100755 src/isabelle_keywords_get.sh create mode 100644 src/isabelle_keywords_update.py diff --git a/src/isabelle_keywords.ml b/src/isabelle_keywords.ml index c773a054..4a3f13fc 100644 --- a/src/isabelle_keywords.ml +++ b/src/isabelle_keywords.ml @@ -1,333 +1,218 @@ -(* AUTOMATICALLY GENERATED by isabelle_keywords_get.sh - DO NOT EDIT *) +(* AUTOMATICALLY GENERATED by `update_isabelle_keywords.py` - DO NOT EDIT *) +Revision.f "$Rev$";; let v = [ - "."; - ".."; - "Isabelle.command"; - "ML"; - "ML_command"; - "ML_prf"; - "ML_val"; - "ProofGeneral.inform_file_processed"; - "ProofGeneral.inform_file_retracted"; - "ProofGeneral.kill_proof"; - "ProofGeneral.pr"; - "ProofGeneral.process_pgip"; - "ProofGeneral.restart"; - "ProofGeneral.undo"; - "abbreviation"; - "advanced"; - "also"; - "and"; - "apply"; - "apply_end"; - "arities"; - "assume"; - "assumes"; - "atom_decl"; - "attach"; - "attribute_setup"; - "avoids"; - "ax_specification"; - "axiomatization"; - "axioms"; - "back"; - "begin"; - "binder"; - "boogie_end"; - "boogie_open"; - "boogie_status"; - "boogie_vc"; - "by"; - "cannot_undo"; - "case"; - "case_eqns"; - "cd"; - "chapter"; - "checking"; - "class"; - "class_deps"; - "classes"; - "classrel"; - "codatatype"; - "code_abort"; - "code_class"; - "code_const"; - "code_datatype"; - "code_deps"; - "code_include"; - "code_instance"; - "code_library"; - "code_module"; - "code_modulename"; - "code_monad"; - "code_pred"; - "code_reflect"; - "code_reserved"; - "code_thms"; - "code_type"; - "coinductive"; - "coinductive_set"; - "commit"; - "con_defs"; - "congs"; - "constrains"; - "consts"; - "consts_code"; - "contains"; - "context"; - "corollary"; - "cpodef"; - "datatype"; - "datatypes"; - "declaration"; - "declare"; - "def"; - "default_sort"; - "defer"; - "defer_recdef"; - "defines"; - "definition"; - "defs"; - "disable_pr"; - "display_drafts"; - "domain"; - "domain_isomorphism"; - "domaindef"; - "domains"; - "done"; - "elimination"; - "enable_pr"; - "end"; - "enriched_type"; - "equivariance"; - "exit"; - "export_code"; - "extract"; - "extract_type"; - "file"; - "finalconsts"; - "finally"; - "find_consts"; - "find_theorems"; - "fix"; - "fixes"; - "fixrec"; - "for"; - "from"; - "full_prf"; - "fun"; - "function"; - "functions"; - "guess"; - "have"; - "header"; - "help"; - "hence"; - "hide_class"; - "hide_const"; - "hide_fact"; - "hide_type"; - "hints"; - "identifier"; - "if"; - "imports"; - "in"; - "induction"; - "inductive"; - "inductive_cases"; - "inductive_set"; - "inductive_simps"; - "infix"; - "infixl"; - "infixr"; - "init_toplevel"; - "instance"; - "instantiation"; - "interpret"; - "interpretation"; - "intros"; - "is"; - "judgment"; - "kill"; - "kill_thy"; - "lazy"; - "lemma"; - "lemmas"; - "let"; - "linear_undo"; - "local_setup"; - "locale"; - "method_setup"; - "module_name"; - "monos"; - "moreover"; - "morphisms"; - "next"; - "nitpick"; - "nitpick_params"; - "no_notation"; - "no_syntax"; - "no_translations"; - "no_type_notation"; - "nominal_datatype"; - "nominal_inductive"; - "nominal_inductive2"; - "nominal_primrec"; - "nonterminal"; - "notation"; - "note"; - "notepad"; - "notes"; - "obtain"; - "obtains"; - "oops"; - "open"; - "oracle"; - "output"; - "overloaded"; - "overloading"; - "parse_ast_translation"; - "parse_translation"; - "partial_function"; - "pcpodef"; - "permissive"; - "pervasive"; - "pr"; - "prefer"; - "presume"; - "pretty_setmargin"; - "prf"; - "primrec"; - "print_abbrevs"; - "print_antiquotations"; - "print_ast_translation"; - "print_attributes"; - "print_binds"; - "print_cases"; - "print_claset"; - "print_classes"; - "print_codeproc"; - "print_codesetup"; - "print_coercion_maps"; - "print_coercions"; - "print_commands"; - "print_configs"; - "print_context"; - "print_dependencies"; - "print_drafts"; - "print_facts"; - "print_induct_rules"; - "print_interps"; - "print_locale"; - "print_locales"; - "print_methods"; - "print_orders"; - "print_quotconsts"; - "print_quotients"; - "print_quotmaps"; - "print_rules"; - "print_simpset"; - "print_statement"; - "print_syntax"; - "print_tcset"; - "print_theorems"; - "print_theory"; - "print_trans_rules"; - "print_translation"; - "proof"; - "prop"; - "pwd"; - "qed"; - "quickcheck"; - "quickcheck_params"; - "quit"; - "quotient_definition"; - "quotient_type"; - "realizability"; - "realizers"; - "recdef"; - "recdef_tc"; - "record"; - "recursor_eqns"; - "refute"; - "refute_params"; - "remove_thy"; - "rep_datatype"; - "schematic_corollary"; - "schematic_lemma"; - "schematic_theorem"; - "sect"; - "section"; - "setup"; - "show"; - "shows"; - "simproc_setup"; - "sledgehammer"; - "sledgehammer_params"; - "smt_status"; - "solve_direct"; - "sorry"; - "spark_end"; - "spark_open"; - "spark_proof_functions"; - "spark_status"; - "spark_types"; - "spark_vc"; - "specification"; - "statespace"; - "structure"; - "subclass"; - "sublocale"; - "subsect"; - "subsection"; - "subsubsect"; - "subsubsection"; - "syntax"; - "syntax_declaration"; - "term"; - "termination"; - "text"; - "text_raw"; - "then"; - "theorem"; - "theorems"; - "theory"; - "thm"; - "thm_deps"; - "thus"; - "thy_deps"; - "translations"; - "try"; - "try_methods"; - "txt"; - "txt_raw"; - "typ"; - "type_elims"; - "type_intros"; - "type_notation"; - "type_synonym"; - "typed_print_translation"; - "typedecl"; - "typedef"; - "types"; - "types_code"; - "ultimately"; - "unchecked"; - "undo"; - "undos_proof"; - "unfolding"; - "unsafe"; - "unused_thms"; - "use"; - "use_thy"; - "uses"; - "using"; - "value"; - "values"; - "welcome"; - "where"; - "with"; - "write"; - "{"; - "}"; + "ML"; + "ML_command"; + "ML_export"; + "ML_file"; + "ML_file_debug"; + "ML_file_no_debug"; + "ML_prf"; + "ML_val"; + "ROOTS_file"; + "SML_export"; + "SML_file"; + "SML_file_debug"; + "SML_file_no_debug"; + "SML_import"; + "abbreviation"; + "abbrevs"; + "alias"; + "also"; + "and"; + "apply"; + "apply_end"; + "assume"; + "assumes"; + "attribute_setup"; + "axiomatization"; + "back"; + "begin"; + "bibtex_file"; + "binder"; + "bundle"; + "by"; + "case"; + "chapter"; + "class"; + "class_deps"; + "code_datatype"; + "compile_generated_files"; + "consider"; + "constrains"; + "consts"; + "context"; + "corollary"; + "declaration"; + "declare"; + "default_sort"; + "defer"; + "define"; + "defines"; + "definition"; + "done"; + "end"; + "experiment"; + "export_files"; + "export_generated_files"; + "export_prefix"; + "external_file"; + "external_files"; + "extract"; + "extract_type"; + "finally"; + "find_consts"; + "find_theorems"; + "fix"; + "fixes"; + "for"; + "from"; + "full_prf"; + "generate_file"; + "global_interpretation"; + "guess"; + "have"; + "help"; + "hence"; + "hide_class"; + "hide_const"; + "hide_fact"; + "hide_type"; + "if"; + "imports"; + "in"; + "include"; + "includes"; + "including"; + "infix"; + "infixl"; + "infixr"; + "instance"; + "instantiation"; + "interpret"; + "interpretation"; + "is"; + "judgment"; + "keywords"; + "lemma"; + "lemmas"; + "let"; + "local_setup"; + "locale"; + "locale_deps"; + "method_setup"; + "moreover"; + "named_theorems"; + "next"; + "no_notation"; + "no_syntax"; + "no_translations"; + "no_type_notation"; + "nonterminal"; + "notation"; + "note"; + "notepad"; + "notes"; + "obtain"; + "obtains"; + "oops"; + "open"; + "opening"; + "oracle"; + "output"; + "overloaded"; + "overloading"; + "paragraph"; + "parse_ast_translation"; + "parse_translation"; + "pervasive"; + "prefer"; + "premises"; + "presume"; + "prf"; + "print_ML_antiquotations"; + "print_abbrevs"; + "print_antiquotations"; + "print_ast_translation"; + "print_attributes"; + "print_bundles"; + "print_cases"; + "print_classes"; + "print_codesetup"; + "print_commands"; + "print_context"; + "print_definitions"; + "print_defn_rules"; + "print_facts"; + "print_interps"; + "print_locale"; + "print_locales"; + "print_methods"; + "print_options"; + "print_rules"; + "print_simpset"; + "print_state"; + "print_statement"; + "print_syntax"; + "print_term_bindings"; + "print_theorems"; + "print_theory"; + "print_trans_rules"; + "print_translation"; + "private"; + "proof"; + "prop"; + "proposition"; + "qed"; + "qualified"; + "realizability"; + "realizers"; + "rewrites"; + "schematic_goal"; + "section"; + "setup"; + "show"; + "shows"; + "simproc_setup"; + "sorry"; + "structure"; + "subclass"; + "subgoal"; + "sublocale"; + "subparagraph"; + "subsection"; + "subsubsection"; + "supply"; + "syntax"; + "syntax_declaration"; + "term"; + "text"; + "text_raw"; + "then"; + "theorem"; + "theory"; + "thm"; + "thm_deps"; + "thm_oracles"; + "thus"; + "thy_deps"; + "translations"; + "txt"; + "typ"; + "type_alias"; + "type_notation"; + "type_synonym"; + "typed_print_translation"; + "typedecl"; + "ultimately"; + "unbundle"; + "unchecked"; + "unfolding"; + "unused_thms"; + "using"; + "welcome"; + "when"; + "where"; + "with"; + "write"; ];; diff --git a/src/isabelle_keywords_get.sh b/src/isabelle_keywords_get.sh deleted file mode 100755 index 8ddc6c59..00000000 --- a/src/isabelle_keywords_get.sh +++ /dev/null @@ -1,20 +0,0 @@ -#!/bin/sh - -# Copyright (C) 2012 INRIA and Microsoft Corporation - -# This script generates the "isabelle_keywords.ml" file. -# It assumes a working installation of Isabelle is in the PATH. - -eval `isabelle getenv ISABELLE_HOME` - -{ - echo '(* AUTOMATICALLY GENERATED by isabelle_keywords_get.sh - DO NOT EDIT *)' - echo 'let v = [' - - sed -n -e '/^.*"\([^"]*\)".*$/s//\1/p' $ISABELLE_HOME/etc/isar-keywords*.el \ - | sed -e 's/\\\\\././g' \ - | sort -u \ - | sed -e 's/.*/ "&";/' - - echo '];;' -} > isabelle_keywords.ml diff --git a/src/isabelle_keywords_update.py b/src/isabelle_keywords_update.py new file mode 100644 index 00000000..ee7ad720 --- /dev/null +++ b/src/isabelle_keywords_update.py @@ -0,0 +1,22 @@ +"""Generate OCaml list of Isabelle keywords. + +This script reads the output of the invocation: +./Isabelle2021-RC2.app/Isabelle/bin/isabelle outer_keywords Pure +which is contained in the file with name equal to `infile`. +""" +infile = 'isabelle_keywords.txt' + + +def main(): + with open(infile, 'r') as f: + text = f.read() + for line in text.split('\n'): + s = '"{line}";'.format(line=line) + if not s: + continue + print(s) + + +if __name__ == '__main__': + main() + From 80d213f37790b200021344b17c8fd940ee9ee81c Mon Sep 17 00:00:00 2001 From: Isabelle Group <> Date: Sat, 23 Jan 2021 12:40:46 +0200 Subject: [PATCH 109/167] ENH: add script for Isabelle 2021 logic keywords Kindly provided by Makarius Wenzel. --- src/outer_keywords.scala | 51 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 51 insertions(+) create mode 100644 src/outer_keywords.scala diff --git a/src/outer_keywords.scala b/src/outer_keywords.scala new file mode 100644 index 00000000..b51073ca --- /dev/null +++ b/src/outer_keywords.scala @@ -0,0 +1,51 @@ +// DESCRIPTION: print outer syntax keywords of specified theory + +object Tool extends isabelle.Isabelle_Tool.Body { + import isabelle._ + + def outer_keywords( + options: Options, + theory_name: String, + dirs: List[Path] = Nil, + progress: Progress = new Progress) + { + val sessions_structure = Sessions.load_structure(options, dirs = dirs) + val session_name = sessions_structure.bootstrap.theory_qualifier(theory_name) + val deps = Sessions.deps(sessions_structure.selection(session_name)) + + deps(session_name).loaded_theory_syntax(theory_name) match { + case None => error("Bad theory " + quote(theory_name)) + case Some(syntax) => + val keywords = syntax.keywords.kinds.keySet.toList.sorted + keywords.foreach(progress.echo) + } + } + + def apply(args: List[String]) + { + var dirs: List[Path] = Nil + + val getopts = Getopts(""" +Usage: isabelle outer_keywords [OPTIONS] THEORY + + Options are: + -d DIR include session directory + + Print outer syntax keywords of specific theory (long name), e.g. + "Pure", "Main", "HOL-Library.Multiset". +""", + "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg)))) + + val more_args = getopts(args) + val theory_name = + more_args match { + case List(name) => name + case _ => getopts.usage() + } + + val options = Options.init() + val progress = new Console_Progress() + + outer_keywords(options, theory_name, dirs = dirs, progress = progress) + } +} From 71942c5c7cafb49f65a2d1a321a412b06d4e5580 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 12:53:02 +0200 Subject: [PATCH 110/167] DOC: how to install and use `outer_keywords.scala` --- src/isabelle_keywords_update.py | 6 +++--- src/outer_keywords.scala | 13 +++++++++++++ 2 files changed, 16 insertions(+), 3 deletions(-) diff --git a/src/isabelle_keywords_update.py b/src/isabelle_keywords_update.py index ee7ad720..6315e8ec 100644 --- a/src/isabelle_keywords_update.py +++ b/src/isabelle_keywords_update.py @@ -1,8 +1,9 @@ """Generate OCaml list of Isabelle keywords. This script reads the output of the invocation: -./Isabelle2021-RC2.app/Isabelle/bin/isabelle outer_keywords Pure -which is contained in the file with name equal to `infile`. + isabelle outer_keywords Pure > isabelle_keywords.txt +The script `outer_keywords.scala` is present in this directory, +and contains its installation and usage documentation. """ infile = 'isabelle_keywords.txt' @@ -19,4 +20,3 @@ def main(): if __name__ == '__main__': main() - diff --git a/src/outer_keywords.scala b/src/outer_keywords.scala index b51073ca..72c61e3a 100644 --- a/src/outer_keywords.scala +++ b/src/outer_keywords.scala @@ -1,4 +1,17 @@ // DESCRIPTION: print outer syntax keywords of specified theory +// +// INSTALLATION: +// Place this Scala script in the directory $ISABELLE_HOME_USER/Tools +// and to the file $ISABELLE_HOME_USER/etc/settings the line +// ISABELLE_TOOLS="$ISABELLE_TOOLS:$ISABELLE_HOME_USER/Tools" +// To find the path ISABELLE_HOME_USER the tool `isabelle getenv` can be used +// by invoking: +// isabelle getenv ISABELLE_HOME_USER +// For example, when this script was first used for `tlapm`, this path was +// $HOME/.isabelle/Isabelle2021-RC2 +// +// Invoke the script with: +// isabelle outer_keywords Pure object Tool extends isabelle.Isabelle_Tool.Body { import isabelle._ From c8c6d782b70ac2660d9ffcfe38b20830c1dca620 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:04:55 +0200 Subject: [PATCH 111/167] MAI: use theories `Constant` and `Zenon` in `ROOT` The attribute `global` is included to enable directly importing these theories within user theories. --- isabelle/ROOT | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/isabelle/ROOT b/isabelle/ROOT index 53c1d9d9..b84420fd 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -2,6 +2,7 @@ session "TLA+" = "Pure" + options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] theories - Peano + Constant (global) + Zenon (global) document_files "root.tex" From f30c5cbfad6fb405391df0deff38841213671bd5 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:05:32 +0200 Subject: [PATCH 112/167] DRAFT: comment document generation for Isabelle theories because it currently raises errors. --- isabelle/ROOT | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/isabelle/ROOT b/isabelle/ROOT index b84420fd..84a6f8b6 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,8 +1,12 @@ (* build the session using "isabelle build -D ." *) session "TLA+" = "Pure" + + (* options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] + *) theories Constant (global) Zenon (global) + (* document_files "root.tex" + *) From df2034bbb2d0337eba59ba1bf2a87d43dc6e549f Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:06:29 +0200 Subject: [PATCH 113/167] MAI: add theorem `leq_adiff_right_imp_0` to theory `Integers` because this theorem is used within the theory `Integers`. --- isabelle/Integers.thy | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index c4398167..776ddd84 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -1044,6 +1044,12 @@ apply(rule nat_leq_cases[of m "n + p" ], simp_all) apply(rule adiff_adiff_left_nat[symmetric], simp+) done +lemma leq_adiff_right_imp_0: + assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" + shows "p = 0" +using p h apply (induct) +using n by auto + lemma int_add_assoc5: assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" shows "-.m + (n + -.p) = -.m + n + -.p" From 0c2833cc66d8fb224a91a161e3205108b2674a14 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:07:01 +0200 Subject: [PATCH 114/167] REL: update `isabelle/Makefile` for Isabelle2020 --- isabelle/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/isabelle/Makefile b/isabelle/Makefile index 93e3091d..cd752015 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -10,4 +10,4 @@ rebuild: isabelle usedir -b -i true -d pdf Pure TLA+ heap-only: - isabelle usedir -b Pure TLA+ + isabelle build -v -d . -b TLA+ From 7fa2eba22c51c629d848a872fe008874f5d36df2 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:08:10 +0200 Subject: [PATCH 115/167] MAI: update theory `Zenon` to use `\\A i \\in` syntax following changes to bounded quantification syntax in the theory `SetTheory`. --- isabelle/Zenon.thy | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 49cbfd67..c00ac8c5 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1047,7 +1047,7 @@ proof (rule inProductE [OF a b]) assume h1: "isASeq(p)" assume h2: "Len(p) = Len(s)" assume h3: "p \\in [1 .. Len (s) -> UNION Range (s)]" - assume h4: "ALL k in 1 .. Len (s) : p[k] \\in s[k]" + assume h4: "\\A k \\in 1 .. Len (s) : p[k] \\in s[k]" show "p[i] \\in s[i]" (is "?c") proof (rule bAllE [where x = i, OF h4]) assume h5: "i \\notin 1 .. Len(s)" @@ -1137,7 +1137,7 @@ lemma zenon_in_recordset_field : shows "r[doms[i]] \\in rngs[i]" proof (rule EnumFuncSetE [OF a]) assume h1: "r \\in [Range(doms) -> UNION Range(rngs)]" - assume h2: "ALL i in DOMAIN doms : r[doms[i]] \\in rngs[i]" + assume h2: "\\A i \\in DOMAIN doms : r[doms[i]] \\in rngs[i]" show "r[doms[i]] \\in rngs[i]" (is "?c") proof (rule bAllE [where x = i, OF h2]) assume h3: "i \\notin DOMAIN doms" @@ -1198,12 +1198,12 @@ lemma zenon_tuple_dom_3 : "isASeq (s) & isASeq (t) & DOMAIN s = DOMAIN t ==> DOMAIN s = DOMAIN t" by auto -lemma zenon_all_rec_1 : "ALL i in {} : r[doms[i]] \\in rngs[i]" by auto +lemma zenon_all_rec_1 : "\\A i \\in {} : r[doms[i]] \\in rngs[i]" by auto lemma zenon_all_rec_2 : - "ALL i in s : r[doms[i]] \\in rngs[i] ==> + "\\A i \\in s : r[doms[i]] \\in rngs[i] ==> r[doms[j]] \\in rngs[j] ==> - ALL i in addElt (j, s) : r[doms[i]] \\in rngs[i]" + \\A i \\in addElt (j, s) : r[doms[i]] \\in rngs[i]" by auto lemma zenon_tuple_acc_1 : @@ -1296,7 +1296,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1360,7 +1360,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1443,7 +1443,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1547,7 +1547,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1670,7 +1670,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1812,7 +1812,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" @@ -1973,7 +1973,7 @@ proof - by (simp only: DomainSeqLen [OF hdomseq], (rule zenon_dom_app_2)+, rule zenon_dom_app_1) have hind: "?indices = DOMAIN ?doms" by (rule conjE [OF hindx], elim conjE) - show "ALL i in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" + show "\\A i \\in DOMAIN ?doms : r[?doms[i]] \\in ?rngs[i]" proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" From d7dd0c1633083c0b1364fe8b1350a066c4ee3f98 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:09:13 +0200 Subject: [PATCH 116/167] DRAFT: comment auto-generated proof rules for `CASE` expressions in the theory `Zenon`, because currently these raise errors, due to an update that is needed in the theory `CaseExpressions`. --- isabelle/Zenon.thy | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index c00ac8c5..4c5d5c2d 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -2100,7 +2100,7 @@ proof - qed (*** END AUTO_GENERATED STUFF ***) - +(* (* ------------- Proof rules for CASE expressions ------ *) lemma zenon_caseother0 : @@ -3694,7 +3694,7 @@ proof - by (rule zenon_case_empty_union [OF hi]) qed qed - +*) (*** END AUTO_GENERATED STUFF ***) end From ea5575c8d332d93e19a0b0252c754f2d99aaf168 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:15:02 +0200 Subject: [PATCH 117/167] MAI: update `tlapm` to call `isabelle process` of Isabelle2020 As part of this change, the Isabelle/TLA+ theories are installed under the directory `${prefix}/lib/isabelle_tla/`. where `${prefix}` the value of the corresponding variable in the file `Makefile`. --- src/params.ml | 14 +++++++++++--- tools/Makefile.in | 6 ++++++ 2 files changed, 17 insertions(+), 3 deletions(-) diff --git a/src/params.ml b/src/params.ml index e674ebdc..75943500 100644 --- a/src/params.ml +++ b/src/params.ml @@ -59,6 +59,13 @@ let library_path = let d = Filename.concat d "tlaps" in d +let isabelle_tla_path = + let d = Sys.executable_name in + let d = Filename.dirname (Filename.dirname d) in + let d = Filename.concat d "lib" in + let d = Filename.concat d "isabelle_tla" in + d + type executable = | Unchecked of string * string * string (* exec, command, version_command *) | User of string (* command *) @@ -127,11 +134,12 @@ let isabelle_success_string = "((TLAPS SUCCESS))" let isabelle = let cmd = - Printf.sprintf "isabelle-process -r -q -e \"(use_thy \\\"$file\\\"; \ - writeln \\\"%s\\\");\" TLA+" + Printf.sprintf "isabelle process -e \"(use_thy \\\"$file\\\"; \ + writeln \\\"%s\\\");\" -d %s -l TLA+" isabelle_success_string + isabelle_tla_path in - make_exec "isabelle-process" cmd "isabelle version" + make_exec "isabelle process" cmd "isabelle version" ;; let set_fast_isabelle () = diff --git a/tools/Makefile.in b/tools/Makefile.in index 613a1438..d4ccacf5 100644 --- a/tools/Makefile.in +++ b/tools/Makefile.in @@ -10,6 +10,7 @@ VPATH = @srcdir@ bindir = @bindir@ LIBPATH = @libdir@/tlaps +ISATLAPATH = @libdir@/isabelle_tla all: check-configure tlapm@EXE@ @@ -101,6 +102,8 @@ install: mkdir -p -m 0755 ${LIBPATH} ${INSTALL} -m 0644 library/*.tla ${LIBPATH} @TAR@ -cf - examples | ( cd ${LIBPATH} && @TAR@ -xf - ) + mkdir -p -m 0755 ${ISATLAPATH} + ${INSTALL} -m 0644 isabelle/*.thy isabelle/*.ML isabelle/ROOT ${ISATLAPATH} .PHONY: uninstall uninstall: @@ -108,6 +111,9 @@ uninstall: if test -n "`ls -A ${LIBPATH}`"; \ then rm -rf ${LIBPATH} ; \ fi + if test -n "`ls -A ${ISATLAPATH}`"; \ + then rm -rf ${ISATLAPATH} ; \ + fi .PHONY: newversion newversion: From 1e11d3dcbcab4273f500ed01ebdfc16d70bfd0d7 Mon Sep 17 00:00:00 2001 From: Ioannis Filippidis Date: Sat, 23 Jan 2021 14:17:06 +0200 Subject: [PATCH 118/167] MAI: include `using assms` only if there exist assumptions in the generated Isabelle theories. This change is needed because Isabelle2020 raises an error if `using assms` appears without any `assumes` keywords preceding it. --- src/backend/isabelle.ml | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/src/backend/isabelle.ml b/src/backend/isabelle.ml index 9e8dabb9..328782f6 100644 --- a/src/backend/isabelle.ml +++ b/src/backend/isabelle.ml @@ -27,6 +27,7 @@ let bump : ctx -> ctx = Ctx.bump let length : ctx -> int = Ctx.length +let emitted_assumptions = ref false let cook x = "is" ^ pickle x @@ -503,6 +504,7 @@ and pp_print_sequent_outer cx ff sq = match Deque.front sq.context with pp_print_sequent_outer ncx ff { sq with context = hs } | Fresh (nm, _, _, Bounded (b, Visible)) -> let (ncx, nm) = adj cx nm in + emitted_assumptions := true; fprintf ff "fixes %s@,assumes %s_in : @[\"(%s \\ %a)\"@]@," nm nm nm (pp_print_expr false cx) b ; pp_print_sequent_outer ncx ff { sq with context = hs } @@ -515,6 +517,7 @@ and pp_print_sequent_outer cx ff sq = match Deque.front sq.context with begin try let null = make_formatter (fun _ _ _ -> ()) (fun () -> ()) in pp_print_expr false cx null e; (* may raise Unsupported *) + emitted_assumptions := true; fprintf ff "assumes v'%d: \"%a\"@," (Ctx.length cx) (pp_print_expr false cx) e with Unsupported msg -> @@ -663,6 +666,7 @@ let thy_temp ob tac tempname thyout = Printf.fprintf thyout "lemma ob'%d: (* %s *)\n" obid obfp; let ff = Format.formatter_of_out_channel thyout in begin try + emitted_assumptions := false; pp_print_expr true Ctx.dot ff (Sequent ob.obl.core @@ nowhere); with Unsupported msg -> failwith ("isabelle : unsupported operator " ^ msg) @@ -671,7 +675,10 @@ let thy_temp ob tac tempname thyout = Printf.fprintf thyout "(is \"PROP ?ob'%d\")\n" obid; Printf.fprintf thyout "proof -\n"; Printf.fprintf thyout "show \"PROP ?ob'%d\"\n" obid; - Printf.fprintf thyout "using assms by %s\n" tac; + if !emitted_assumptions then + Printf.fprintf thyout "using assms by %s\n" tac + else + Printf.fprintf thyout "by %s\n" tac; Printf.fprintf thyout "qed\n"; Printf.fprintf thyout "end\n"; ;; From 6ddeccbdfe8613f571a418cd579bebacf226a8a5 Mon Sep 17 00:00:00 2001 From: merz Date: Sat, 23 Jan 2021 14:37:28 +0100 Subject: [PATCH 119/167] fixed syntax translations for Isa2020 --- isabelle/CaseExpressions.thy | 45 ++- isabelle/Tuples.thy | 525 ++++++++--------------------------- 2 files changed, 136 insertions(+), 434 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 09c2ca77..6507100a 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -1,9 +1,8 @@ (* Title: TLA+/CaseExpressions.thy Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:22 merz> + Version: Isabelle2020 *) section \ Case expressions \ @@ -34,7 +33,7 @@ text \ twice (and will be simplified twice) in a @{text CaseOther} expression. \ -definition CaseArm (* -- preliminary construct to convert case arm into set *) +definition CaseArm \ \ preliminary construct to convert case arm into set \ where "CaseArm(p,e) \ IF p THEN {e} ELSE {}" definition Case where @@ -49,22 +48,16 @@ nonterminal case_arm and case_arms syntax "_case_syntax":: "case_arms \ c" ("(CASE _ )" 10) - "_case1" :: "[c, c] \ case_arm" ("(2_ ->/ _)" 10) - "" :: "case_arm \ case_arms" ("_") - "_other" :: "c \ case_arms" ("OTHER -> _") - "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ [] _") - -syntax (xsymbols) "_case1" :: "[c, c] \ case_arm" ("(2_ \/ _)" 10) + "" :: "case_arm \ case_arms" ("_") "_other" :: "c \ case_arms" ("OTHER \ _") "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") -syntax (HTML output) - "_case1" :: "[c, c] \ case_arm" ("(2_ \/ _)" 10) - "_other" :: "c \ case_arms" ("OTHER \ _") - "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ \ _") +syntax (ASCII) + "_case1" :: "[c, c] \ case_arm" ("(2_ ->/ _)" 10) + "_other" :: "c \ case_arms" ("OTHER -> _") + "_case2" :: "[case_arm, case_arms] \ case_arms" ("_/ [] _") -(* TODO parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -98,15 +91,15 @@ parse_ast_translation \ end | case_syntax_tr _ = raise Match; in - [("_case_syntax", case_syntax_tr)] + [("_case_syntax", K case_syntax_tr)] end \ -*) + (** debugging ** (*ML \ set Syntax.trace_ast; \*) -lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" +lemma "CASE a \ b \ c \ d \ e \ f \ OTHER \ g" oops lemma "CASE a \ b" @@ -118,7 +111,6 @@ oops (*ML \ reset Syntax.trace_ast; \*) **) -(* TODO print_ast_translation \ let fun list_from_tuple (Ast.Constant @{const_syntax "emptySeq"}) = [] @@ -173,16 +165,15 @@ print_ast_translation \ end | caseother_tr' _ = raise Match in - [(@{const_syntax "Case"}, case_syntax_tr'), - (@{const_syntax "CaseOther"}, caseother_tr')] + [(@{const_syntax "Case"}, K case_syntax_tr'), + (@{const_syntax "CaseOther"}, K caseother_tr')] end \ -*) (** debugging ** (*ML \ set Syntax.trace_ast; \*) -lemma "CASE a \ b [] c \ d [] e \ f [] OTHER \ g" +lemma "CASE a \ b \ c \ d \ e \ f \ OTHER \ g" oops lemma "CASE a \ b" @@ -199,15 +190,15 @@ lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def (** test cases ** lemma - "i=1 \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1) = 0" -by simp + "i=1 \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = 0" +by (simp add: two_def) lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1) = default" + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" by auto lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 [] i=1 \ 0 [] i=2 \ 1 [] OTHER \ a) = a" + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" by auto **) diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index d644e8b5..5ecc293f 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Tuples.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:46 merz> + Version: Isabelle2020 *) section \ Tuples and Relations in \tlaplus{} \ @@ -31,13 +30,13 @@ text \ on sequences require arithmetic and will be introduced in a separate theory. \ -definition Seq (* -- \ set of finite sequences with elements from $S$ \ *) +definition Seq \ \ set of finite sequences with elements from $S$ \ where "Seq(S) \ UNION { [ 1 .. n \ S] : n \ Nat }" -definition isASeq (* -- \ characteristic predicate for sequences or tuples \close> *) +definition isASeq \ \ characteristic predicate for sequences or tuples \ where "isASeq(s) \ isAFcn(s) \ (\n \ Nat : DOMAIN s = 1 .. n)" -definition Len (* -- \ length of a sequence \ *) +definition Len \ \ length of a sequence \ where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: @@ -80,15 +79,15 @@ lemma SeqIsAFcn (*[elim!]*): shows "isAFcn(s)" using assms by auto -(* -- \ @{text "s \ Seq(S) \ isAFcn(s)"} \ *) +\ \ @{text "s \ Seq(S) \ isAFcn(s)"} \ lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" shows "Len(s) \ Nat" -using assms by auto +using assms by (rule isASeqE) -(* -- \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ *) +\ \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] lemma DomainSeqLen [simp]: @@ -96,7 +95,7 @@ lemma DomainSeqLen [simp]: shows "DOMAIN s = 1 .. Len(s)" using assms by auto -(* -- \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ *) +\ \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: @@ -121,7 +120,7 @@ lemma SeqI [intro!]: shows "s \ Seq(S)" using assms by (auto simp: Seq_def) -lemma SeqI': (* -- \ closer to the definition but probably less useful \ *) +lemma SeqI': \ \ closer to the definition but probably less useful \ assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" using assms by (auto simp: Seq_def) @@ -141,7 +140,7 @@ qed lemma seqFuncSet: assumes "s \ Seq(S)" shows "s \ [1 .. Len(s) \ S]" -using assms by auto +using assms by (rule SeqE) lemma seqElt [elim!]: assumes "s \ Seq(S)" and "n \ Nat" and "1 \ n" and "n \ Len(s)" @@ -164,14 +163,11 @@ text \ (written @{text "\\"}) and Append. \ -definition emptySeq ("(<< >>)") -where "<< >> \ [x \ 1 .. 0 \ {}]" +definition emptySeq ("(\\)") +where "\\ \ [x \ 1 .. 0 \ {}]" -notation (xsymbols) - emptySeq ("(\\)") - -notation (HTML output) - emptySeq ("(\\)") +notation (ASCII) + emptySeq ("(<< >>)") definition Append :: "[c,c] \ c" where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len(s)] THEN e ELSE s[k]]" @@ -179,7 +175,7 @@ where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" by (auto simp: emptySeq_def isASeq_def) -(* -- \ @{text "isAFcn(\\)"} \ *) +\ \ @{text "isAFcn(\\)"} \ lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] lemma lenEmptySeq [simp]: "Len(\\) = 0" @@ -199,7 +195,7 @@ lemma appendIsASeq [simp,intro!]: using s unfolding Append_def by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) -(* -- \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ *) +\ \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" @@ -213,11 +209,11 @@ lemma isEmptySeq [intro!]: "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" by auto -(* -- \ immediate consequence of @{text isEmptySeq} \ *) +\ \ immediate consequence of @{text isEmptySeq} \ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" by auto -(* -- \ Symmetric equation could be a useful rewrite rule (it is applied by TLC) \ *) +\ \ The reverse equation could be a useful rewrite rule (it is applied by TLC) \ lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" @@ -243,7 +239,7 @@ lemma lenAppend [simp]: shows "Len(Append(s,e)) = Succ[Len(s)]" using assms by (intro LenI, auto simp: Append_def) -(* -- \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ *) +\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] lemma appendElt [simp]: @@ -336,28 +332,7 @@ text \ lemma appendExtend: assumes "isASeq(s)" shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" - proof - - have s1_1: "Append(s, e) = - [k \ 1 .. Succ[Len(s)] \ - IF k = Succ[Len(s)] THEN e ELSE s[k]]" - unfolding Append_def by auto - have s1_2: "(Succ[Len(s)] :> e) = - [x \ {Succ[Len(s)]} \ e]" - unfolding oneArg_def by auto - define p where "p \ s @@ (Succ[Len(s)] :> e)" - have s1_3: "p = - [x \ (DOMAIN s) \ {Succ[Len(s)]} - \ IF x \ DOMAIN s THEN s[x] ELSE e]" - unfolding p_def extend_def using s1_2 by auto - have s1_4: "(DOMAIN s) \ {Succ[Len(s)]} - = 1 .. Succ[Len(s)]" - using assms by auto - have s1_5: "p = [x \ 1..Succ[Len(s)] \ - IF x \ DOMAIN s THEN s[x] ELSE e]" - using s1_3 s1_4 by auto - show ?thesis - using s1_1 s1_5 assms by (auto simp: p_def) - qed +using assms by (intro isAppend') auto lemma imageEmptySeq [simp]: "Image(\\, A) = {}" by (simp add: emptySeq_def) @@ -367,7 +342,7 @@ lemma imageAppend [simp]: shows "Image(Append(s,e), A) = (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" unfolding appendExtend[OF s] -using assms by (auto elim!: inNatIntervalE, force+) +using assms by auto (force+) text \ Inductive reasoning about sequences, based on @{term "\\"} and @@ -419,7 +394,7 @@ proof - with s show ?thesis unfolding Seq_def by auto qed -(* -- \ example of an inductive proof about sequences \ *) +\ \ example of an inductive proof about sequences \ lemma seqEmptyOrAppend: assumes "s \ Seq(S)" shows "s = \\ \ (\s' \ Seq(S): \e \ S : s = Append(s',e))" @@ -444,13 +419,10 @@ nonterminal tpl syntax "" :: "c \ tpl" ("_") "@app" :: "[tpl, c] \ tpl" ("_,/ _") - "@tuple" :: "tpl \ c" ("<<(_)>>") - -syntax (xsymbols) "@tuple" :: "tpl \ c" ("\(_)\") -syntax (HTML output) - "@tuple" :: "tpl \ c" ("\(_)\") +syntax (ASCII) + "@tuple" :: "tpl \ c" ("<<(_)>>") translations "\ tp, x \" \ "CONST Append(\tp\, x)" @@ -458,13 +430,13 @@ translations (*** examples *** -lemma "Len(\a,b,c\) = 3" by simp +lemma "Len(\a,b,c\) = 3" by (simp add: three_def two_def) lemma "Append(\a,b\, c) = \a,b,c\" .. lemma "\a,b,c\[1] = a" by simp -lemma "\a,b,c\[2] = b" by simp -lemma "\a,b,c\[3] = c" by simp +lemma "\a,b,c\[2] = b" by (simp add: two_def) +lemma "\a,b,c\[3] = c" by (simp add: two_def three_def) lemma "\0,1\ \ \\" by simp lemma "\0,1\ \ \1\" by simp lemma "\0,1\ \ \1,2\" by simp @@ -480,16 +452,6 @@ text \ \ syntax - "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) - "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) - "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) - "@bEx5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) - "@bAll2" :: "[idt,idt,c,c] \ c" ("(3ALL <<_,_>> in _ :/ _)" [100,100,0,0] 10) - "@bAll3" :: "[idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) - "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) - "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) - -syntax (xsymbols) "@bEx2" :: "[idt,idt,c,c] \ c" ("(3\\_,_\ \ _ :/ _)" [100,100,0,0] 10) "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3\\_,_,_\ \ _ :/ _)" [100,100,100,0,0] 10) "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_\ \ _ :/ _)" [100,100,100,100,0,0] 10) @@ -499,6 +461,16 @@ syntax (xsymbols) "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_\ \ _ :/ _)" [100,100,100,100,0,0] 10) "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3\\_,_,_,_,_\ \ _ :/ _)" [100,100,100,100,100,0,0] 10) +syntax (ASCII) + "@bEx2" :: "[idt,idt,c,c] \ c" ("(3EX <<_,_>> in _ :/ _)" [100,100,0,0] 10) + "@bEx3" :: "[idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) + "@bEx4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) + "@bEx5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3EX <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) + "@bAll2" :: "[idt,idt,c,c] \ c" ("(3ALL <<_,_>> in _ :/ _)" [100,100,0,0] 10) + "@bAll3" :: "[idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_>> in _ :/ _)" [100,100,100,0,0] 10) + "@bAll4" :: "[idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_>> in _ :/ _)" [100,100,100,100,0,0] 10) + "@bAll5" :: "[idt,idt,idt,idt,idt,c,c] \ c" ("(3ALL <<_,_,_,_,_>> in _ :/ _)" [100,100,100,100,100,0,0] 10) + translations "\\x,y\ \ S : P" \ "\x,y : \x,y\ \ S \ P" "\\x,y,z\ \ S : P" \ "\x,y,z : \x,y,z\ \ S \ P" @@ -529,9 +501,9 @@ definition EnumFuncSet :: "c \ c \ c" where "EnumFuncSet(doms, rngs) \ { f \ [Range(doms) \ UNION Range(rngs)] : \i \ DOMAIN doms : f[doms[i]] \ rngs[i] }" -lemmas (* -- \ establish set equality for sets of enumerated functions \ *) - setEqualI [where A = "EnumFuncSet(doms, rngs)", intro!] - setEqualI [where B = "EnumFuncSet(doms, rngs)", intro!] +lemmas \ \ establish set equality for sets of enumerated functions \ + setEqualI [where A = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] + setEqualI [where B = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] lemma EnumFuncSetI [intro!,simp]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" @@ -562,7 +534,6 @@ syntax "@domrngs" :: "[domrng, domrngs] \ domrngs" ("_,/ _") "@EnumFuncSet":: "domrngs \ c" ("[_]") -(* parse_ast_translation \ let (* make_tuple converts a list of ASTs to a tuple formed from these ASTs. @@ -591,7 +562,7 @@ parse_ast_translation \ end | enum_funcset_tr _ = raise Match; in - [("@EnumFuncSet", enum_funcset_tr)] + [("@EnumFuncSet", K enum_funcset_tr)] end \ @@ -622,35 +593,35 @@ print_ast_translation \ end | enum_funcset_tr' _ = raise Match in - [(@{const_syntax "EnumFuncSet"}, enum_funcset_tr')] + [(@{const_syntax "EnumFuncSet"}, K enum_funcset_tr')] end \ -*) + (*** Examples *** lemma "(1 :> TRUE) \ [1 : BOOLEAN]" by auto lemma "(1 :> TRUE) @@ (5 :> 2) \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE @@ 5 :> 2) = (5 :> 2 @@ 1 :> TRUE)" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (2 :> {}) \ [1 : BOOLEAN, 2: SUBSET Nat, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![1] = FALSE] \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![5] = {}] \ [1 : BOOLEAN, 5 : SUBSET Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (1 :> {}) \ [1 : BOOLEAN, 5 : Nat]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "[1 : BOOLEAN, 5 : Nat] = [5 : Nat, 1 : BOOLEAN]" -by auto +by (auto simp: two_def three_def four_def five_def) lemma "\f \ [1 : BOOLEAN, 2: Nat]; g \ [2: Nat, 1: BOOLEAN]; f[1] = g[1]; f[2] = g[2]\ \ f = g" by auto @@ -732,12 +703,10 @@ lemma "\a\ \ Product(\A,B\)" by auto text \ Special case: binary product \ definition - prod :: "c \ c \ c" (infixr "\\X" 100) where - "A \\X B \ Product(\A,B\)" -notation (xsymbols) - prod (infixr "\" 100) -notation (HTML output) - prod (infixr "\" 100) + prod :: "c \ c \ c" (infixr "\" 100) where + "A \ B \ Product(\A,B\)" +notation (ASCII) + prod (infixr "\\X" 100) lemma inProd [simp]: "(\a,b\ \ A \ B) = (a \ A \ b \ B)" @@ -746,8 +715,7 @@ by (auto simp add: prod_def) lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" -unfolding two_def -using assms by (auto simp add: prod_def) +using assms by (auto simp: two_def prod_def) lemma inProd': "(p \ A \ B) = (\a \ A : \ b \ B : p = \a,b\)" @@ -821,13 +789,12 @@ using assms by (auto dest: pairProj_in_prod) lemma relProj2 [elim]: (** consider [elim!] ?? **) assumes "\a,b\ \ r" and "r \ A \ B" shows "b \ B" -using assms by (auto dest: pairProj_in_prod) + using assms by (auto dest: pairProj_in_prod) lemma setOfAllPairs_eq_r (*[simp]*): assumes r: "r \ A \ B" shows "{\p[1], p[2]\ : p \ r} = r" -apply auto -using subsetD[OF r, THEN prodProj[of _ A B]] by simp_all +using subsetD[OF r, THEN prodProj] by force lemma subsetsInProd: assumes "A' \ A" and "B' \ B" @@ -842,9 +809,9 @@ definition setOfPairs :: "[c, [c,c]\c] \ c" where "setOfPairs(R, f) \ { f(p[1], p[2]) : p \ R }" syntax - "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : <<_,_>> \\in _})") -syntax (xsymbols) "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : \_,_\ \ _})") +syntax (ASCII) + "@setOfPairs" :: "[c, idt, idt, c] \ c" ("(1{_ : <<_,_>> \\in _})") translations "{e : \x,y\ \ R}" \ "CONST setOfPairs(R, \x y. e)" @@ -859,7 +826,7 @@ lemma inSetOfPairsI [intro]: unfolding two_def using assms by (auto simp: setOfPairs_def two_def) -lemma inSetOfPairsE [elim!]: (* -- \ converse true only if $R$ is a relation \ *) +lemma inSetOfPairsE [elim!]: \ \ the converse is true only if $R$ is a relation \ assumes 1: "z \ setOfPairs(R, e)" and 2: "R \ A \ B" and 3: "\x y. \ \x,y\ \ R; z = e(x,y) \ \ P" shows "P" @@ -875,188 +842,27 @@ lemmas setOfPairsEqualI = setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] lemma setOfPairs_triv [simp]: - assumes s: "R \ A \ B" + assumes "R \ A \ B" shows "{ \x, y\ : \x, y\ \ R } = R" - proof - - define Q where "Q \ - { \x, y\ : \x, y\ \ R }" - have s1_1: "Q = {<>: p \ R}" - unfolding Q_def setOfPairs_def by auto - have s1_2: "\ p. p \ R ==> p \ A \ B" - using s by auto - have s1_3: "\ p. p \ R ==> p = <>" - proof - - fix "p" :: "c" - assume s1_3_asm: "p \\in R" - have s2_1: "p \ A \ B" - using s1_3_asm s by auto - show "p = <>" - using s2_1 prodProj by auto - qed - define e :: "[c] => c" where "e \ \ x. x" - define f :: "[c] => c" where "f \ \ x. <>" - have s1_4: "\ x. x \ R ==> e(x) = f(x)" - unfolding e_def f_def using s1_3 by auto - have s1_5: "{e(x): x \ R} = {f(y): y \ R}" - apply (rule setOfAll_cong[of "R" "R" "e" "f"]) - apply (simp) - by (rule s1_4) - show ?thesis - using s1_1 s1_5 by (auto simp: Q_def e_def f_def) - qed +using assms by auto lemma setOfPairs_cong (*[cong]*): assumes 1: "R = S" and 2: "S \ A \ B" and - 3: "\x y. \x, y\ \ S - \ - e(x, y) = f(x, y)" + 3: "\x y. \x, y\ \ S \ e(x, y) = f(x, y)" shows "{ e(x, y) : \x, y\ \ R } = { f(u, v) : \u, v\ \ S }" - proof - - have s1_1: "{ e(x, y) : \x, y\ \ R } = - { e(x, y) : \x, y\ \ S }" - using 1 by auto - have s1_2: "{ e(x, y) : \x, y\ \ S } = - { e(p[1], p[2]) : p \ S }" - by (auto simp: setOfPairs_def) - have s1_3: "{ f(u, v) : \u, v\ \ S } = - { f(p[1], p[2]) : p \ S }" - by (auto simp: setOfPairs_def) - define P where "P \ { e(p[1], p[2]): p \ S }" - define Q where "Q \ { f(p[1], p[2]): p \ S }" - have s1_4: "P = Q" - proof - - have s2_1: "\ x. x \ P <=> x \ Q" - proof - - fix "x" :: "c" - have s3_1: "x \\in P ==> x \\in Q" - proof - - assume s3_1_asm: "x \\in P" - have s4_1: "x \ { e(p[1], p[2]): p \ S }" - using s3_1_asm by (auto simp: P_def) - have s4_2: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_1 setOfAll by auto - have s4_3: "\\E p: p \\in S \ x = f(p[1], p[2])" - using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) - have s4_4: "x \ { f(p[1], p[2]): p \ S }" - using s4_3 setOfAll by auto - show "x \\in Q" - unfolding Q_def using s4_4 by auto - qed - have s3_2: "x \\in Q ==> x \\in P" - proof - - assume s3_2_asm: "x \\in Q" - have s4_1: "x \ { f(p[1], p[2]): p \ S}" - using s3_2_asm by (auto simp: Q_def) - have s4_2: "\\E p: p \\in S \ x = f(p[1], p[2])" - using s4_1 setOfAll by auto - have s4_3: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_2 2 3 pairProj_in_rel by (auto simp: two_def) - have s4_4: "x \ { e(p[1], p[2]): p \ S }" - using s4_3 setOfAll by auto - show "x \\in P" - unfolding P_def using s4_4 by auto - qed - show "x \ P <=> x \ Q" - using s3_1 s3_2 by auto - qed - have s2_2: "\ x: x \ P <=> x \ Q" - using s2_1 allI by auto - show ?thesis - using s2_2 extension by auto - qed - show ?thesis - using s1_1 s1_2 s1_3 s1_4 - by (auto simp: P_def Q_def) - qed +using assms by (force intro!: inSetOfPairsI_ex) lemma setOfPairsEqual: - assumes 1: "\x y. - \x, y\ \ S - \ \\x', y'\ \ T : - e(x, y) = f(x', y')" - and 2: "\x' y'. - \x', y'\ \ T - \ \\x, y\ \ S : - f(x', y') = e(x, y)" - and 3: "S \ A \ B" and - 4: "T \ C \ D" + assumes 1: "\x y. \x, y\ \ S \ \\x', y'\ \ T : e(x, y) = f(x', y')" + and 2: "\x' y'. \x', y'\ \ T \ \\x, y\ \ S : f(x', y') = e(x, y)" + and 3: "S \ A \ B" + and 4: "T \ C \ D" shows "{ e(x, y) : \x, y\ \ S } = { f(x, y) : \x, y\ \ T }" - proof - - have s1_1: "{ e(x, y) : \x, y\ \ S } = - { e(p[1], p[2]): p \ S }" - by (auto simp: setOfPairs_def) - have s1_2: "{ f(x, y) : \x, y\ \ T } = - { f(p[1], p[2]): p \ T }" - by (auto simp: setOfPairs_def) - define P where "P \ { e(p[1], p[2]): p \ S }" - define Q where "Q \ { f(p[1], p[2]): p \ T }" - have s1_3: "P = Q" - proof - - have s2_1: "\ x. x \ P <=> x \ Q" - proof - - fix "x" :: "c" - have s3_1: "x \\in P ==> x \\in Q" - proof - - assume s3_1_asm: "x \\in P" - have s4_1: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s3_1_asm setOfAll by (auto simp: P_def) - have s4_2: "\\E p: p \\in S \ - \p[1], p[2]\ \ S \ - x = e(p[1], p[2])" - using s4_1 pairProj_in_rel 3 by auto - have s4_3: "\\E p: p \\in S \ - x = e(p[1], p[2]) \ - (\\u, v\ \ T: - e(p[1], p[2]) = f(u, v))" - using s4_2 1 by auto - have s4_4: "\\u, v\ \ T: - x = f(u, v)" - using s4_3 by auto - have s4_5: "\\E u, v: <> \\in T \ x = f(u, v)" - using s4_4 by auto - have s4_6: "\\E p: p \\in T \ x = f(p[1], p[2])" - using s4_5 4 pairProj_in_rel by (auto simp: two_def) - show "x \\in Q" - using s4_6 setOfAll by (auto simp: Q_def) - qed - have s3_2: "x \\in Q ==> x \\in P" - proof - - assume s3_2_asm: "x \\in Q" - have s4_1: "\\E p: p \\in T \ x = f(p[1], p[2])" - using s3_2_asm setOfAll by (auto simp: Q_def) - have s4_2: "\\E p: p \\in T \ - <> \\in T \ - x = f(p[1], p[2])" - using s4_1 pairProj_in_rel 4 by auto - have s4_3: "\\E p: p \\in T \ - x = f(p[1], p[2]) \ - (\\u, v\ \ S: - f(p[1], p[2]) = e(u, v))" - using s4_2 2 by auto - have s4_4: "\\u, v\ \ S: - x = e(u, v)" - using s4_3 by auto - have s4_5: "\\E u, v: <> \\in S \ x = e(u, v)" - using s4_4 by auto - have s4_6: "\\E p: p \\in S \ x = e(p[1], p[2])" - using s4_5 3 pairProj_in_rel by (auto simp: two_def) - show "x \\in P" - using s4_6 setOfAll by (auto simp: P_def) - qed - show "x \\in P <=> x \\in Q" - using s3_1 s3_2 by auto - qed - have s2_2: "\ x: x \ P <=> x \ Q" - using s2_1 allI by auto - show ?thesis - using s2_2 extension by auto - qed - show ?thesis - using s1_1 s1_2 s1_3 by (auto simp: P_def Q_def) - qed +using assms by (force intro!: inSetOfPairsI_ex) + subsection \ Basic notions about binary relations \ @@ -1069,25 +875,22 @@ where "rel_range(r) \ { p[2] : p \ r }" definition converse :: "c => c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" -definition rel_comp :: "[c,c] => c" (infixr "\" 75) (* -- \ binary relation composition \ *) +\ \ composition of binary relations \ +definition rel_comp :: "[c,c] => c" (infixr "\" 75) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : - \ x, z : - p = \x, z\ \ - (\ y: - \x, y\ \ s \ - \y, z\ \ r) }" + \ x, z : p = \x, z\ \ (\ y: \x, y\ \ s \ \y, z\ \ r) }" definition rel_image :: "[c,c] => c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" -definition Id :: "c \ c" (* -- \ diagonal: identity over a set \ *) +definition Id :: "c \ c" \ \ diagonal: identity over a set \ where "Id(A) \ { \x,x\ : x \ A }" text \ Properties of relations \ -definition reflexive (* -- \ reflexivity over a set \ *) +definition reflexive \ \ reflexivity over a set \ where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" @@ -1096,7 +899,7 @@ unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" unfolding isBool_def by (rule boolifyReflexive) -definition symmetric (* -- \ symmetric relation \ *) +definition symmetric \ \ symmetric relation \ where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" @@ -1105,7 +908,7 @@ unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" unfolding isBool_def by (rule boolifySymmetric) -definition antisymmetric (* -- \ antisymmetric relation \ *) +definition antisymmetric \ \ antisymmetric relation \ where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" @@ -1114,7 +917,7 @@ unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" unfolding isBool_def by (rule boolifyAntisymmetric) -definition transitive (* -- \ transitivity predicate \ *) +definition transitive \ \ transitivity predicate \ where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" @@ -1123,7 +926,7 @@ unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" unfolding isBool_def by (rule boolifyTransitive) -definition irreflexive (* -- \ irreflexivity predicate \ *) +definition irreflexive \ \ irreflexivity predicate \ where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" @@ -1132,7 +935,7 @@ unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" unfolding isBool_def by (rule boolifyIrreflexive) -definition equivalence :: "[c,c] \ c" (* -- \ (partial) equivalence relation \ *) +definition equivalence :: "[c,c] \ c" \ \ (partial) equivalence relation \ where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" @@ -1211,7 +1014,7 @@ lemmas converseEqualI = setEqualI [where A = "r^-1" for r, intro!] setEqualI [where B = "r^-1" for r, intro!] -lemma converse_iff [iff]: +lemma converse_iff [simp]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" using r prodProj by (auto simp: converse_def two_def) @@ -1236,8 +1039,7 @@ lemma converseE [elim]: (** consider [elim!] ?? **) proof - from prodProj[OF subsetD[OF converseSubset[OF r] yx]] have 2: "yx = \yx[1], yx[2]\" . with yx have 3: "\yx[2], yx[1]\ \ r" - unfolding converse_def apply auto - using r prodProj by auto + unfolding converse_def using r prodProj by auto from p[of "yx[1]" "yx[2]"] 2 3 show P by simp qed @@ -1245,52 +1047,18 @@ qed lemma converse_converse [simp]: assumes r: "r \ A \ B" shows "(r^-1)^-1 = r" - proof - - have s1_1: "\ x. x \\in A \ B ==> x = <>" - using prodProj by (auto simp: two_def) - have s1_2: "\ x. x \\in r ==> x \\in A \ B" - using s1_1 r by auto - have s1_3: "\ x. x \\in r ==> x = <>" - using s1_1 s1_2 by auto - show ?thesis - using s1_3 assms prodProj pairProj_in_rel - by (auto simp: converse_def two_def) - qed +proof - + { + fix x + assume x: "x \ r" + have "x = \x[1], x[2]\" + using x r by (intro prodProj) auto + } + then show ?thesis by (force simp: converse_def two_def) +qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" - proof - - have s1_1: "(A \ B)^-1 = - {<>: p \ A \ B}" - unfolding converse_def by auto - have s1_2: "B \ A = - {<>: p \ B \ A}" - using setOfPairs_triv prodProj by auto - have s1_3: "\ x. x \\in B \\X A ==> - \\E p \\in A \\X B: x = <>" - proof - - fix "x" :: "c" - assume s1_5_asm: "x \\in B \\X A" - have s2_1: "<> \\in A \\X B" - using s1_5_asm by (auto simp: two_def) - define q where "q \ <>" - have s2_2: "q \\in A \\X B \ q = <>" - using s2_1 by (auto simp: q_def) - have s2_3: "x = <>" - using s1_5_asm prodProj by auto - have s2_4: "q[1] = x[2] \ q[2] = x[1]" - using s2_2 by (auto simp: two_def) - have s2_5: "x = <>" - using s2_3 s2_4 by auto - show "\\E p \\in A \\X B: x = <>" - using s2_2 s2_5 by auto - qed - have s1_4: "\ p. p \\in A \\X B ==> - p[1] \\in A \ p[2] \\in B" - using pairProj_in_rel inProd prodProj - by (auto simp: two_def) - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed +by auto lemma converse_empty [simp]: "converse({}) = {}" by auto @@ -1338,42 +1106,8 @@ lemma transitive_converse [simp]: lemma symmetric_iff_converse_eq: assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" - proof - - have s1_1: "symmetric(r) ==> r^-1 = r" - proof - - assume s1_1_asm: "symmetric(r)" - have s2_1: "\ p. p \\in r^-1 ==> p \\in r" - proof - - fix "p" :: "c" - assume s2_1_asm: "p \\in r^-1" - show "p \\in r" - using r s1_1_asm s2_1_asm - by (auto elim!: converseE simp: symmetric_def) - qed - have s2_2: "\ p. p \\in r ==> p \\in r^-1" - proof - - fix "p" :: "c" - assume s2_2_asm: "p \\in r" - have s3_1: "p = <>" - using s2_2_asm r by (intro prodProj, auto) - have s3_2: "<> \\in r" - using s1_1_asm s2_2_asm s3_1 by (auto simp: symmetric_def) - show "p \\in r^-1" - using s3_1 s3_2 by (auto dest: converseI) - qed - show "r^-1 = r" - using s2_1 s2_2 setEqualI[of "r^-1" "r"] - by blast - qed - have s1_2: "r^-1 = r ==> symmetric(r)" - proof - - assume s1_2_asm: "r^-1 = r" - show "symmetric(r)" - using s1_2_asm by (auto simp: symmetric_def) - qed - show ?thesis - using s1_1 s1_2 by auto - qed +using assms by (auto simp: symmetric_def) + subsubsection \ Identity relation over a set \ @@ -1483,53 +1217,30 @@ qed lemma R_comp_Id [simp]: assumes r: "R \ B \ C" shows "R \ Id(B) = R" - using r - proof - - have s1_1: "\ p. p \\in R ==> p \\in R \ Id(B)" - proof - - fix "p" :: "c" - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[1] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ R \ Id(B)" by (intro compI, auto) - with 1 show "p \ R \ Id(B)" by simp - qed - have s1_2: "\ p. p \\in R \ Id(B) ==> p \\in R" - proof - - fix "p" :: "c" - assume s1_2_asm: "p \\in R \ Id(B)" - show "p \\in R" - using r s1_2_asm by auto - qed - show ?thesis - using s1_1 s1_2 setEqualI[of "R \ Id(B)" "R"] by blast - qed - +proof (rule setEqualI) + fix x + assume "x \ R \ Id(B)" + with r show "x \ R" by auto +next + fix x + assume x: "x \ R" + with r obtain b c where "b \ B" "c \ C" "x = \b,c\" by force + with x r show "x \ R \ Id(B)" by auto +qed lemma Id_comp_R [simp]: assumes r: "R \ A \ B" shows "Id(B) \ R = R" - using r - proof - - have s1_1: "\ p. p \\in R ==> p \\in Id(B) \ R" - proof - - fix "p" :: "c" - assume p: "p \ R" - with r have 1: "p = \p[1], p[2]\" (is "p = ?pp") by (intro prodProj, auto) - from p r have "p[2] \ B" by (auto dest: pairProj_in_prod) - with 1 p r have "?pp \ Id(B) \ R" by (intro compI, auto) - with 1 show "p \ Id(B) \ R" by simp - qed - have s1_2: "\ p. p \\in Id(B) \ R ==> p \\in R" - proof - - fix "p" :: "c" - assume s1_2_asm: "p \\in Id(B) \ R" - show "p \\in R" - using r s1_2_asm by auto - qed - show ?thesis - using s1_1 s1_2 setEqualI[of "Id(B) \ R" "R"] by blast - qed +proof (rule setEqualI) + fix x + assume "x \ Id(B) \ R" + with r show "x \ R" by auto +next + fix x + assume x: "x \ R" + with r obtain a b where "a \ A" "b \ B" "x = \a,b\" by force + with x r show "x \ Id(B) \ R" by auto +qed lemma rel_comp_empty1 [simp]: "{} \ R = {}" unfolding rel_comp_def by auto @@ -1553,7 +1264,7 @@ proof (rule setEqualI) fix u assume "\y,u\ \ S" and "\u,z\ \ R" with r s t 2 have "\x,z\ \ R \ (S \ T)" - by (intro compI, auto elim!: relProj1 relProj2) + by (intro compI) (auto elim!: relProj1 relProj2) with 1 show ?thesis by simp qed qed @@ -1570,7 +1281,7 @@ next fix u assume "\x,u\ \ T" and "\u,y\ \ S" with r s t 3 have "\x,z\ \ (R \ S) \ T" - by (intro compI, auto elim!: relProj1 relProj2) + by (intro compI) (auto elim!: relProj1 relProj2) with 1 show ?thesis by simp qed qed @@ -1759,7 +1470,7 @@ proof - have sym: "symmetric(r)" proof - from comp have "r^-1 = (r^-1 \ r)^-1" by simp - also from r r1 have "... = r^-1 \ r" by (simp add: converse_comp) + also from r r1 have "\ = r^-1 \ r" by (simp add: converse_comp) finally have "r^-1 = r" by (simp add: comp) with r show ?thesis by (simp add: symmetric_iff_converse_eq) qed From e6605aed9eeb88328604cdad3ce3841a9e90b16a Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 29 Jan 2021 18:54:56 +0100 Subject: [PATCH 120/167] remaining syntax translations for Isabelle-2020 --- isabelle/Constant.thy | 5 +-- isabelle/Strings.thy | 101 +++++++++++++++++++++++++++--------------- 2 files changed, 68 insertions(+), 38 deletions(-) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index f7aff81d..dba29b43 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Constant.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:28 merz> + Version: Isabelle2020 *) section \ Main theory for constant-level Isabelle/\tlaplus{} \ diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index cabb62c8..42152c84 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -1,9 +1,8 @@ (* Title: TLA+/Strings.thy Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:15 merz> + Version: Isabelle2020 *) section \ Characters and strings \ @@ -20,7 +19,23 @@ text \ \ definition Nibble -where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" +(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) + where "Nibble \ {0, + 1, + Succ[1], + Succ[Succ[1]], + Succ[Succ[Succ[1]]], + Succ[Succ[Succ[Succ[1]]]], + Succ[Succ[Succ[Succ[Succ[1]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]], + Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]]]}" definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" @@ -30,7 +45,8 @@ by (simp add: char_def) (*** example ** lemma "char(3,13) \ char(14,2)" -by simp + by (simp add: two_def three_def four_def five_def six_def seven_def eight_def + nine_def ten_def eleven_def twelve_def thirteen_def fourteen_def) ***) definition Char @@ -59,7 +75,6 @@ text \ lexicon. \ -(* TODO parse_ast_translation \ let (* convert an ML integer to a nibble *) @@ -75,46 +90,64 @@ parse_ast_translation \ mkNibble (ord c div 16), mkNibble (ord c mod 16)] else error ("Non-ASCII symbol: " ^ quote c); + (* remove two leading quotes from a list of characters *) + fun trim_quotes ("'" :: ( "'" :: cs)) = cs + | trim_quotes cs = cs + (* convert a list of ML characters into a TLA+ string, in reverse order *) fun list2TupleReverse [] = Ast.Constant "Tuples.emptySeq" | list2TupleReverse (c :: cs) = Ast.Appl [Ast.Constant "Tuples.Append", list2TupleReverse cs, mkChar c]; (* parse AST translation for characters *) - fun char_ast_tr [Ast.Variable xstr] = - (case Lexicon.explode_xstr xstr of + fun char_ast_tr [Ast.Variable str] = + let val trimmed_str = (trim_quotes (rev (trim_quotes (Symbol.explode str)))) + in (case trimmed_str of [c] => mkChar c - | _ => error ("Expected single character, not " ^ xstr)) - | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); + | _ => error ("Expected single character, not " ^ str)) + end + | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts); (* parse AST translation for strings *) - fun string_ast_tr [Ast.Variable xstr] = - list2TupleReverse (rev (Lexicon.explode_xstr xstr)) + fun string_ast_tr [Ast.Variable str] = + list2TupleReverse (trim_quotes (rev ((trim_quotes (Symbol.explode str))))) | string_ast_tr asts = raise Ast.AST ("string_ast_tr", asts); in - [("_Char", char_ast_tr), ("_String", string_ast_tr)] + [("_Char", K char_ast_tr), ("_String", K string_ast_tr)] end \ -*) (** debug ** -(*ML \ set Syntax.trace_ast; \*) +(*declare [[syntax_ast_trace = true]]*) lemma "''a''" oops +lemma "''''" +oops + +lemma "''abc''" +oops + lemma "CHAR ''a''" oops -(*ML \ reset Syntax.trace_ast; \*) +(* declare [[syntax_ast_trace = false]] *) **) -(* TODO print_ast_translation \ let - (* convert a nibble to an ML integer -- because translation macros have - already been applied, we see constants "0" through "15", not Succ[...] terms! *) + (* convert a nibble to an ML integer *) + fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 + | destNibble (Ast.Constant @{const_syntax "one"}) = 1 + | destNibble (Ast.Appl [Ast.Constant @{const_syntax "Functions.fapply"}, + Ast.Constant @{const_syntax "Peano.Succ"}, nb]) + = (destNibble nb) + 1 + | destNibble _ = raise Match + + (* the following version should be used when 2 .. 15 are abbreviations, not definitions *) +(* fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 | destNibble (Ast.Constant @{const_syntax "two"}) = 2 @@ -132,7 +165,7 @@ print_ast_translation \ | destNibble (Ast.Constant @{const_syntax "fourteen"}) = 14 | destNibble (Ast.Constant @{const_syntax "fifteen"}) = 15 | destNibble _ = raise Match; - +*) (* convert a pair of nibbles to an ML character *) fun destNbls nb1 nb2 = let val specials = raw_explode "\"\\`'" @@ -154,13 +187,12 @@ print_ast_translation \ (* convert a list of TLA+ characters to the output representation of a TLA+ string *) fun list2String cs = Ast.Appl [Ast.Constant "_inner_string", - Ast.Variable (Lexicon.implode_xstr cs)]; + Ast.Variable (Lexicon.implode_str cs)]; (* print AST translation for single characters that do not occur in a string *) fun char_ast_tr' [nb1, nb2] = - Ast.Appl [Ast.Constant @{syntax_const "_Char"}, - list2String [destNbls nb1 nb2]] - | char_ast_tr' _ = raise Match; + Ast.Appl [Ast.Constant @{syntax_const "_Char"}, list2String [destNbls nb1 nb2]] + | char_ast_tr' _ = raise Match (* print AST translation for non-empty literal strings, fails (by raising exception Match) @@ -168,14 +200,12 @@ print_ast_translation \ fun string_ast_tr' [args] = list2String (map destChar (tuple2List args)) | string_ast_tr' _ = raise Match; in - [(@{const_syntax "char"}, char_ast_tr'), ("@tuple", string_ast_tr')] + [(@{const_syntax "char"}, K char_ast_tr'), ("@tuple", K string_ast_tr')] end \ -*) (*** examples ** -(* not sure if the following is the way to set tracing? *) -(* ML \ val _ = Config.put Ast.trace true @{context} \ *) +(*declare [[syntax_ast_trace = true]]*) lemma "CHAR ''a'' \ CHAR ''b''" by simp @@ -193,9 +223,10 @@ lemma "''ab''[1] = CHAR ''a''" by simp lemma "''abc''[2] \ CHAR ''a''" -by simp +by (simp add: two_def) + +(* declare [[syntax_ast_trace = false]] *) -(* ML \ val _ = Config.put Ast.trace false @{context} \ *) **) @@ -220,7 +251,7 @@ by auto lemma "r \ [''bar'' : BOOLEAN, ''foo'' : Nat] \ [r EXCEPT ![''foo''] = 3] \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by force (* "by auto" also works, but is slow *) +by (force simp: two_def three_def) (* "by auto" also works, but is slow *) lemma "(''a'' :> 1) \ (''b'' :> 1)" by simp @@ -229,19 +260,19 @@ lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1)" by simp lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1 @@ ''b'' :> 3)" -by simp +by (simp add: three_def) lemma "(''a'' :> 1 @@ ''b'' :> 2) = (''b'' :> 2 @@ ''a'' :> 1)" by simp lemma "''ab'' = [i \ {1,2} \ ''ab''[i]]" -by auto +by (auto simp: two_def) lemma "''ab'' = (1 :> CHAR ''a'') @@ (2 :> CHAR ''b'')" -by auto +by (auto simp: two_def) lemma "Len(''ab'') = 2" -by simp +by (simp add: two_def) **) From fdcd7a55c21ed0bf1b321d369ffb69946d7eaa98 Mon Sep 17 00:00:00 2001 From: merz Date: Sun, 7 Feb 2021 09:42:39 +0100 Subject: [PATCH 121/167] restored document preparation --- isabelle/PredicateLogic.thy | 6 +++--- isabelle/ROOT | 11 +++++------ isabelle/ROOT.ML | 15 --------------- isabelle/SetTheory.thy | 2 +- 4 files changed, 9 insertions(+), 25 deletions(-) delete mode 100644 isabelle/ROOT.ML diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index afade712..cd9411a9 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -157,10 +157,10 @@ text \ The following axioms govern propositional logic: \begin{enumerate} - \item @{term TRUE} is valid, and @{term FALSE} implies anything}, - \item @{text A} is a theorem if and only if it equals @{term TRUE}, + \item @{text TRUE} is valid, and @{text FALSE} implies anything, + \item @{text A} is a theorem if and only if it equals @{text TRUE}, \item conditionals are reasoned about by case distinction, - \item equations evaluate to @{term TRUE} or @{term FALSE}. + \item equations evaluate to @{text TRUE} or @{text FALSE}. \end{enumerate} \ diff --git a/isabelle/ROOT b/isabelle/ROOT index 84a6f8b6..2bd32182 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,12 +1,11 @@ -(* build the session using "isabelle build -D ." *) +(* build the session using "isabelle build -b -D ." *) +chapter "TLA+" session "TLA+" = "Pure" + - (* - options [document = pdf, document_output = "output", document_variants="document:outline=/proof"] - *) + options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + theories Constant (global) Zenon (global) - (* + document_files "root.tex" - *) diff --git a/isabelle/ROOT.ML b/isabelle/ROOT.ML deleted file mode 100644 index f174a865..00000000 --- a/isabelle/ROOT.ML +++ /dev/null @@ -1,15 +0,0 @@ -(* Title: TLA+/ROOT.ML - Author: Stephan Merz, LORIA - Copyright (C) 2008-2011 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2011 - Time-stamp: <2011-07-26 16:53:42 merz> - -Encoding of TLA+ constant level as an object logic in Isabelle. -*) - -val banner = "TLA+"; -writeln banner; - -use_thy "Constant"; -use_thy "Zenon"; diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index 707897f4..bbb6fbe0 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -259,7 +259,7 @@ definition setminus :: "[c, c] \ c" (infixl "\" 65) ( definition INTER :: "c \ c" ("INTER _" [100]90) (* generalized intersection *) where "INTER A \ { x \ UNION A : \B \ A : x \ B }" -text \Note that @{text "INTER {}"} is @{text "{}"}.}\ +text \Note that @{text "INTER {}"} is @{text "{}"}.\ notation (ASCII) "inter" (infixl "\\inter" 70) and From c2676982377fb99f0e63ee6b70685ec8cdc69f42 Mon Sep 17 00:00:00 2001 From: merz Date: Thu, 23 Dec 2021 18:49:50 +0100 Subject: [PATCH 122/167] lemmas for SMT axioms --- isabelle/Integers.thy | 4 +- isabelle/NewSMT.thy | 281 ++++++++++++++++++++++++++++++++++++++++++ isabelle/ROOT | 1 + 3 files changed, 284 insertions(+), 2 deletions(-) create mode 100644 isabelle/NewSMT.thy diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index 776ddd84..fdcc141f 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -419,7 +419,7 @@ lemmas int_leq_pp_def = nat_leq_def (* -- 'positive-positive' case, ie: both arguments are naturals *) axiomatization where - int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = FALSE" + int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = (a = 0 \ b = 0)" and int_leq_np_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ b = TRUE" and @@ -445,7 +445,7 @@ using assms by simp lemma neg_le_iff_le [simp]: "\m \ Int; n \ Int \ \ -.n \ -.m = (m \ n)" -by(rule intCases2[of m n], simp_all) + by (rule intCases2[of m n]) auto subsection \ Addition of integers \ diff --git a/isabelle/NewSMT.thy b/isabelle/NewSMT.thy new file mode 100644 index 00000000..766ae093 --- /dev/null +++ b/isabelle/NewSMT.thy @@ -0,0 +1,281 @@ +section \Axioms used by the new SMT backend\ + +theory NewSMT + imports Constant +begin + +subsection \Set Theory\ + +lemma ChooseDef: "P(x) \ P(Choice(P))" + by (rule chooseI) + +lemma SetExtensionality: "(\z : z \ a \ z \ b) \ a = b" + by (simp add: extension) + +lemma SubseteqIntro: "(\z : z \ a \ z \ b) \ a \ b" + by blast + +lemma SubseteqElim: "a \ b \ x \ a \ x \ b" + by blast + +text \ + The two following axioms are stated for $n$-ary enumerations, + we prove their instances for $n=2$. +\ +lemma EnumIntro: "a \ {a,b} \ b \ {a,b}" + by blast + +lemma EnumElim: "x \ {a,b} \ x=a \ x=b" + by blast + +lemma SubsetDef: "x \ SUBSET a \ x \ a" + by blast + +lemma UnionIntro: "x \ y \ y \ a \ x \ UNION a" + by blast + +lemma UnionElim: "x \ UNION a \ (\y : x \ y \ y \ a)" + by blast + +lemma SetstDef: "x \ subsetOf(a,P) \ x \ a \ P(x)" + by (rule subsetOf) + +text \ + The two following axioms are stated for an $n$-ary + operator $f$ but Isabelle only provides the construct + for unary operators. +\ +lemma SetofIntro: "y \ a \ f(y) \ setOfAll(a,f)" + by blast + +lemma SetofElim: "x \ setOfAll(a,f) \ (\y : y \ a \ x = f(y))" + by blast + +lemma CupDef: "x \ a \ b \ x \ a \ x \ b" + by blast + +lemma CapDef: "x \ a \ b \ x \ a \ x \ b" + by blast + +lemma SetminusDef: "x \ a \ b \ x \ a \ x \ b" + by blast + + +subsection \Functions\ + +lemma FcnExtensionality: + assumes "isAFcn(f)" "isAFcn(g)" "DOMAIN f = DOMAIN g" + "\x \ DOMAIN f : f[x] = g[x]" + shows "f = g" + using assms by auto + +lemma FcnIsafcn: "isAFcn(Fcn(a,F))" + by (rule fcnIsAFcn) + +lemma FcnDom: "DOMAIN Fcn(a,F) = a" + by (rule DOMAIN) + +lemma FcnApp: "x \ a \ Fcn(a,F)[x] = F(x)" + by auto + +lemma ArrowIntro: + assumes "isAFcn(f)" "DOMAIN f = a" "\x \ a : f[x] \ b" + shows "f \ [a \ b]" + using assms by (rule inFuncSet) + +lemma ArrowElim1: "f \ [a \ b] \ isAFcn(f) \ DOMAIN f = a" + by blast + +lemma ArrowElim2: "f \ [a \ b] \ x \ a \ f[x] \ b" + by blast + +lemma ExceptIsafcn: "isAFcn(except(f,x,y))" + by blast + +lemma ExceptDom: "DOMAIN except(f,x,y) = DOMAIN f" + by (rule domainExcept) + +lemma ExceptApp1: "x \ DOMAIN f \ except(f,x,y)[x] = y" + by auto + +lemma ExceptApp2: "z \ DOMAIN f \ z \ x \ except(f,x,y)[z] = f[z]" + by auto + + +subsection \Integers\ + +lemma PlusTyping: "x \ Int \ y \ Int \ x+y \ Int" + by simp + +lemma UminusTyping: "x \ Int \ -.x \ Int" + by simp + +lemma MinusTyping: "x \ Int \ y \ Int \ x-y \ Int" + by simp + +lemma TimesTyping: "x \ Int \ y \ Int \ x*y \ Int" + by simp + +(* Isabelle currently provides integer division only for natural + numbers, so we cannot prove the following. + +lemma DivTyping: "x \ Int \ y \ Int \ y \ 0 \ x \\div y \ Int" +*) + +lemma NatDef: "x \ Nat \ x \ Int \ 0 \ x" +proof + assume "x \ Nat" + then show "x \ Int \ 0 \ x" by simp +next + assume "x \ Int \ 0 \ x" + hence "x \ Int" "0 \ x" by blast+ + from \x \ Int\ show "x \ Nat" + proof (rule intCases) + fix m + assume "m \ Nat" "x = -.m" + with \0 \ x\ show ?thesis + by (cases m) auto + qed +qed + +(* Isabelle currently does not support integer intervals, + so we cannot prove the following. + +lemma RangeDef: "x \ a .. b \ x \ Int \ a \ x \ x \ b" +*) + + +subsection \Booleans\ + +text \ + The following lemmas do not exactly correspond to the SMT + axioms because Isabelle/TLA+ is untyped and does not have + cast operators. +\ + +lemma BooleanIntro: "TRUE \ BOOLEAN \ FALSE \ BOOLEAN" + by blast + +lemma BooleanElim: "x \ BOOLEAN \ x = TRUE \ x = FALSE" + by blast + + +subsection \Tuples and Records\ + +text \ + The following SMT axioms are for tuples or records + of arbitrary length, we prove instances for pairs. +\ + +lemma TupIsafcn: "isAFcn(\x,y\)" + by blast + +lemma TupDom: "DOMAIN \x,y\ = {1,2}" + by (auto simp: two_def) + +lemma TupApp: "\x,y\[1] = x \ \x,y\[2] = y" + by (auto simp: two_def) + +lemma ProductIntro: "x \ a \ y \ b \ \x,y\ \ a \ b" + by auto + +lemma ProductElim: "x \ a \ b \ x = \x[1], x[2]\" + by (rule prodProj) + +lemma RecordIsafcn: "isAFcn(''foo'' :> x @@ ''bar'' :> y)" + by auto + +lemma RecordDom: "DOMAIN (''foo'' :> x @@ ''bar'' :> y) = {''foo'', ''bar''}" + by auto + +lemma RecordApp: + "(''foo'' :> x @@ ''bar'' :> y)[''foo''] = x" + "(''foo'' :> x @@ ''bar'' :> y)[''bar''] = y" + by auto + +lemma RectIntro: + assumes "x \ a" "y \ b" + shows "(''foo'' :> x @@ ''bar'' :> y) \ [''foo'' : a, ''bar'' : b]" + using assms by auto + +lemma RectElim: + assumes "x \ [''foo'' : a, ''bar'' : b]" + shows "x = (''foo'' :> x[''foo''] @@ ''bar'' :> x[''bar''])" + using assms by force + + +subsection \Sequences\ + +text \ + Somewhat rewritten, in particular the original hypothesis + @{text "\i: i \ DOMAIN s \ 1 \ i \ i \ Len(s)"} + doesn't imply that @{text "DOMAIN s \ Nat"}. + Could instead assume + @{text "DOMAIN s = 1 .. Len(s)"} and + @{text "\i \ DOMAIN s : s[i] \ a"}. +\ +lemma SeqIntro: + assumes "isAFcn(s)" "Len(s) \ Nat" + "\i : i \ DOMAIN s \ i \ Nat \ 1 \ i \ i \ Len(s)" + "\i \ Nat : 1 \ i \ i \ Len(s) \ s[i] \ a" + shows "s \ Seq(a)" + using assms by auto + +lemma SeqElim1: + "s \ Seq(a) \ isAFcn(s) \ Len(s) \ Nat \ DOMAIN s = 1 .. Len(s)" + by auto + +text \ + Similarly rewritten as theorem @{text SeqIntro} above. +\ +lemma SeqElim2: + "s \ Seq(a) \ i \ Nat \ 1 \ i \ i \ Len(s) \ s[i] \ a" + by auto + +lemma LenDef: + "x \ Nat \ DOMAIN s = 1 .. x \ Len(s) = x" + by auto + +text \ + Isabelle/TLA+ currently doesn't know about concatenation of sequences, + so we cannot prove lemmas corresponding to the SMT axioms about concatenation. +\ + +lemma AppendTyping: "s \ Seq(a) \ x \ a \ Append(s,x) \ Seq(a)" + by auto + +lemma AppendLen: + assumes "Len(s) \ Nat" + shows "Len(Append(s,x)) = Len(s)+1" +proof + from assms show "DOMAIN Append(s,x) = 1 .. (Len(s)+1)" by auto +next + from assms show "Len(s)+1 \ Nat" by auto +qed + +text \ + Some changes w.r.t. the original formulations of the axioms about @{text Append}. +\ +lemma AppendApp1: + assumes "Len(s) \ Nat" (* missing from original formulation *) + "i \ Nat" (* changed from Int because Isabelle doesn't support Int intervals *) + "1 \ i" "i \ Len(s)" + shows "Append(s,x)[i] = s[i]" + using assms unfolding Append_def by auto + +lemma AppendApp2: + "Len(s) \ Nat \ Append(s,x)[Len(s)+1] = x" + unfolding Append_def by auto + +text \ + Isabelle/TLA+ doesn't know about the remaining operations on sequences. +\ + +lemma TupSeqTyping: "x \ a \ y \ a \ \x,y\ \ Seq(a)" + by auto + +lemma TupSeqLen: "Len(\x,y\) = 2" + by (auto simp: two_def) + + +end diff --git a/isabelle/ROOT b/isabelle/ROOT index 2bd32182..30960a90 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -6,6 +6,7 @@ session "TLA+" = "Pure" + theories Constant (global) Zenon (global) +(* NewSMT (global) *) document_files "root.tex" From bbd123b17cf72c7685f52a78ff7c3fc6cb2a41b8 Mon Sep 17 00:00:00 2001 From: merz Date: Wed, 16 Mar 2022 16:19:52 +0100 Subject: [PATCH 123/167] refactoring of integer theories, update to Isabelle 2021-1 --- isabelle/CaseExpressions.thy | 16 +- isabelle/Constant.thy | 6 +- isabelle/IntegerArithmetic.thy | 2140 ++++++++++++++++++++++++++ isabelle/IntegerDivision.thy | 713 +++++++++ isabelle/Integers.thy | 2579 +++++++++++++++++--------------- isabelle/README.html | 61 +- isabelle/Strings.thy | 48 +- isabelle/Tuples.thy | 511 ++++--- isabelle/Zenon.thy | 498 ++---- 9 files changed, 4716 insertions(+), 1856 deletions(-) create mode 100644 isabelle/IntegerArithmetic.thy create mode 100644 isabelle/IntegerDivision.thy diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index 6507100a..f9c6cbeb 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -1,8 +1,8 @@ (* Title: TLA+/CaseExpressions.thy - Author: Stephan Merz, LORIA - Copyright (C) 2009-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Case expressions \ @@ -191,15 +191,15 @@ lemmas Case_simps [simp] = CaseArm_def Case_def CaseOther_def lemma "i=1 \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = 0" -by (simp add: two_def) + by simp lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" -by auto + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1) = default" + by auto lemma - "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" -by auto + "i \ {0,1,2} \ (CASE i=0 \ 2 \ i=1 \ 0 \ i=2 \ 1 \ OTHER \ a) = a" + by auto **) diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index dba29b43..d085827c 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -1,14 +1,14 @@ (* Title: TLA+/Constant.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Copyright (C) 2008-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Main theory for constant-level Isabelle/\tlaplus{} \ theory Constant -imports NatDivision CaseExpressions Strings Integers +imports IntegerDivision CaseExpressions Strings begin text \ diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy new file mode 100644 index 00000000..07aaf7a6 --- /dev/null +++ b/isabelle/IntegerArithmetic.thy @@ -0,0 +1,2140 @@ +(* Title: TLA+/IntegerArithmetic.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation + License: BSD + Version: Isabelle2021-1 +*) + +section \ Arithmetic (except division) for the integers \ + +theory IntegerArithmetic +imports Integers +begin + +named_theorems algebra_simps "algebra simplification rules" + +text \ + The rewrites accumulated in @{text algebra_simps} deal with the + classical algebraic structures of groups, rings and family. They simplify + terms by multiplying everything out (in case of a ring) and bringing sums and + products into a canonical form (by ordered rewriting). As a result these + rewrites decide the equational theory of groups and rings but also help + with inequalities. +\ + + +subsection \ Addition and difference of integers \ + +text \ + We define integer addition as a recursive function. + Subtraction is just addition of the complement, and + is introduced as an abbreviation. +\ + +definition pr_addint where +"pr_addint(a) \ int_primrec(Int, a, \n x y. succ[x], \n x y. pred[y])" + +definition addint :: "[c,c] \ c" (infixl "+" 65) +where "\a \ Int; b \ Int\ \ (a + b) \ pr_addint(a)[b]" + +abbreviation subint (infixl "-" 65) +where "a - b \ a + (-b)" + +text \ Closure \ + +lemma pr_addintType: + assumes "a \ Int" shows "pr_addint(a) \ [Int \ Int]" + using assms unfolding pr_addint_def by (rule int_primrecType) auto + +lemma addintType [iff]: "\a \ Int; b \ Int\ \ a+b \ Int" + unfolding addint_def by (blast dest: pr_addintType) + +text \ Base case and successor cases \ + +lemma pr_addint_0: "a \ Int \ pr_addint(a)[0] = a" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma add_0_int [simp]: "a \ Int \ a + 0 = a" + by (simp add: addint_def pr_addint_0) + +lemma pr_addint_SuccNat: + "\a \ Int; n \ Nat\ \ pr_addint(a)[succ[n]] = succ[pr_addint(a)[n]]" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma pr_addint_uminusSuccNat: + "\a \ Int; n \ Nat\ \ pr_addint(a)[-succ[n]] = pred[pr_addint(a)[-n]]" + by (rule int_primrecE[OF pr_addint_def[THEN meta_to_obj_eq]]) auto + +lemma addint_succ [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + succ[b] = succ[a+b]" + using b a + by (cases b) + (auto dest: pr_addintType + simp: addint_def pr_addint_SuccNat pr_addint_uminusSuccNat) + +lemma addint_pred [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + pred[b] = pred[a+b]" + using b a + by (cases b) + (auto dest: pr_addintType + simp: addint_def pr_addint_SuccNat pr_addint_uminusSuccNat pred0 sym[OF uminusSucc]) + +lemma subint_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "a - succ[b] = pred[a-b]" + using assms by (simp add: uminusSucc) + +lemma subint_pred [simp]: + assumes "a \ Int" and "b \ Int" + shows "a - pred[b] = succ[a-b]" + using assms by (simp add: uminusPred) + +lemma addintNat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m + n \ Nat" + using n m by induct auto + +lemma addint_0_left [simp]: + assumes "a \ Int" + shows "0 + a = a" + using assms + by induct (auto simp: uminusSucc) + +lemma add_succ_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[a] + b = succ[a + b]" + using b a by (induct b) auto + +lemma add_pred_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[a] + b = pred[a + b]" + using b a by (induct b) auto + +lemma add_uminus_succ_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "-succ[a] + b = pred[-a + b]" + using assms by (simp add: uminusSucc) + +lemma add_uminus_pred_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "-pred[a] + b = succ[-a + b]" + using assms by (simp add: uminusPred) + +lemma succ_pred_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[pred[a] + b] = a + b" + using b a by induct auto + +lemma pred_succ_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[succ[a] + b] = a + b" + using b a by induct auto + +lemma subint_self [simp]: + assumes "a \ Int" + shows "a - a = 0" + using assms by induct auto + +lemma addint_uminus_self [simp]: + assumes "a \ Int" + shows "-a + a = 0" + using assms by induct auto + +lemma uminus_add [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "-(a + b) = -a - b" + using b a by induct (auto simp: uminusSucc) + +text \ Commutativity \ + +lemma add_commute [algebra_simps]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a + b = b + a" + using b a by (induct b) (auto simp: uminusSucc) + +text \ Associativity \ + +lemma add_assoc [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + (b + c) = (a + b) + c" + using assms by induct simp+ + +lemma add_left_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + (b + c) = b + (a + c)" + using assms by (simp only: add_assoc add_commute) + +lemma add_right_commute: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a + b + c = a + c + b" +proof - + from assms have "a + b + c = a + (b + c)" + by (rule sym[OF add_assoc]) + also from assms have "\ = a + (c + b)" + by (simp add: add_commute) + finally show ?thesis + using assms by (simp add: add_assoc) +qed + +lemmas add_ac = add_assoc add_commute add_left_commute + + +text \ Cancellation rules \ + +lemma add_left_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b = a + c) = (b = c)" + using assms by induct simp+ + +lemma add_right_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(b + a = c + a) = (b = c)" + using assms by induct simp+ + +lemma add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b = a) = (b = 0)" + using assms by induct (simp add: uminusSucc)+ + +lemma add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (a = a + b) = (b = 0)" + "\a \ Int; b \ Int\ \ (b + a = a) = (b = 0)" + "\a \ Int; b \ Int\ \ (a = b + a) = (b = 0)" + by (auto simp: add_commute dest: sym) + +lemma succ_add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a+b] = a) = (b = -1)" +proof - + { + assume "succ[a+b] = a" + with assms have "a + succ[b] = a" by simp + with assms add_eq_self1[of "a" "succ[b]"] + have "succ[b] = 0" by simp + hence "pred[succ[b]] = pred[0]" by simp + with assms have "b = -1" by (simp add: pred0) + } + with assms show ?thesis by auto +qed + +lemma succ_add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (succ[b+a] = a) = (b = -1)" + "\a \ Int; b \ Int\ \ (a = succ[a+b]) = (b = -1)" + "\a \ Int; b \ Int\ \ (a = succ[b+a]) = (b = -1)" + by (auto simp: add_commute dest: sym) + +lemma pred_add_eq_self1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a+b] = a) = (b = 1)" +proof - + { + assume "pred[a+b] = a" + with assms have "a + pred[b] = a" by simp + with assms add_eq_self1[of "a" "pred[b]"] + have "pred[b] = 0" by simp + hence "succ[pred[b]] = succ[0]" by simp + with assms have "b = 1" by simp + } + with assms show ?thesis by auto +qed + +lemma pred_add_eq_self2 [simp]: + "\a \ Int; b \ Int\ \ (pred[b+a] = a) = (b = 1)" + "\a \ Int; b \ Int\ \ (a = pred[a+b]) = (b = 1)" + "\a \ Int; b \ Int\ \ (a = pred[b+a]) = (b = 1)" + by (auto simp: add_commute dest: sym) + + +text \ Rules about integer difference \ + +lemma diff_diff_left: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(a - b) - c = a - (b + c)" + using assms by (simp add: add_assoc) + +lemma add_diff_same_1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b) - a = b" + using assms by (simp add: add_ac) + +lemma add_diff_same_2 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a + b) - b = a" +proof - + from assms have "(a + b) - b = a + (b - b)" + by (intro sym[OF add_assoc]) auto + with assms show ?thesis by simp +qed + +lemma diff_add_same_1 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a - b) + b = a" + using assms add_diff_same_2[of "a" "-b"] by simp + +lemma diff_add_same_2 [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a + b) + a = b" + using assms by (simp add: add_ac) + +lemma add_diff_cancel_left [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b) - (a + c) = b - c" + using assms by (simp add: add_ac) + +lemma add_diff_cancel_right [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b) - (c + b) = a - c" (is "?lhs = ?rhs") +proof - + from assms have "?lhs = a + (b - (c + b))" + by (intro sym[OF add_assoc]) simp+ + also have "\ = ?rhs" + using assms by (simp add: add_ac) + finally show ?thesis . +qed + +lemma add_left_diff_right: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b = c) = (a = c - b)" +proof - + { + assume "a + b = c" + hence "(a + b) - b = c - b" by simp + with assms have "a = c - b" by simp + } + moreover + { + assume "a = c - b" + hence "a + b = (c - b) + b" by simp + with assms have "a + b = c" by simp + } + ultimately show ?thesis by auto +qed + +lemma diff_0_eq [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a - b = 0) = (a = b)" +proof - + { + assume "a - b = 0" + hence "(a - b) + b = 0 + b" by simp + with assms have "a = b" by simp + } + with assms show ?thesis by auto +qed + +lemma add_right_diff_left: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a = b + c) = (c = a - b)" + using assms by (auto simp: add_left_diff_right add_assoc) + +lemma diff_0_eqs [simp]: + "\a \ Int; b \ Int\ \ (-a + b = 0) = (a = b)" + "\a \ Int; b \ Int\ \ (0 = a - b) = (a = b)" + "\a \ Int; b \ Int\ \ (0 = -a + b) = (a = b)" + by (auto simp: add_commute dest: sym) + +text \ Reasoning about @{text "m + n = 0"}, etc. \ + +lemma add_is_0_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(m + n = 0) = (m = 0 \ n = 0)" + using m apply cases + using n by (induct, auto) + +lemma zero_is_add_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(0 = m + n) = (m = 0 \ n = 0)" + using m apply cases + using n by (induct, auto) + +lemma add_is_1_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(m + n = 1) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" + using m apply cases + using n by (induct, auto simp: add_is_0_nat) + +lemma one_is_add_nat: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "(1 = m + n) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" + using m apply (rule natCases) + using n by (induct, auto simp: zero_is_add_nat) + + +subsection \ Multiplication of natural numbers \ + +definition pr_multint where +"pr_multint(a) \ int_primrec(Int, 0, \n x y. x+a, \n x y. y-a)" + +definition multint :: "[c,c] \ c" (infixl "*" 70) +where "\a \ Int; b \ Int\ \ (a * b) \ pr_multint(a)[b]" + +text \ Closure \ + +lemma pr_multintType: + assumes "a \ Int" shows "pr_multint(a) \ [Int \ Int]" + using assms unfolding pr_multint_def + by (intro int_primrecType) auto + +lemma multintType [iff]: "\a \ Int; b \ Int\ \ a*b \ Int" + unfolding multint_def by (blast dest: pr_multintType) + +text \ Base case and successor cases \ + +lemma pr_multint_0: "a \ Int \ pr_multint(a)[0] = 0" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma times_0 [simp]: "a \ Int \ a * 0 = 0" + by (simp add: multint_def pr_multint_0) + +lemma pr_multint_SuccNat: + "\a \ Int; n \ Nat\ \ pr_multint(a)[succ[n]] = pr_multint(a)[n] + a" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma pr_multint_uminusSuccNat: + "\a \ Int; n \ Nat\ \ pr_multint(a)[-succ[n]] = pr_multint(a)[-n] - a" + by (rule int_primrecE[OF pr_multint_def[THEN meta_to_obj_eq]]) auto + +lemma times_succ [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * succ[b] = a*b + a" + using b a + by (cases b) + (simp add: multint_def pr_multint_SuccNat pr_multint_uminusSuccNat + pr_multintType[THEN funcSetValue])+ + +lemma times_pred [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * pred[b] = a*b - a" + using b a + by (cases b) + (simp add: multint_def pr_multint_SuccNat pr_multint_uminusSuccNat + pr_multintType[THEN funcSetValue] pred0 sym[OF uminusSucc])+ + +lemma times_uminus_right [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a * (-b) = -(a * b)" +using b proof induct + case 0 + with a show ?case by simp +next + case (pos n) + with a show ?case by (simp add: uminusSucc) +next + case (neg n) + with a show ?case + by simp (simp add: uminusSucc) +qed + +lemma times_nat [simp]: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m * n \ Nat" + using n m by induct simp+ + +lemma times_0_left [simp]: + assumes "a \ Int" + shows "0 * a = 0" + using assms by induct simp+ + +lemma times_succ_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "succ[a] * b = a * b + b" +using b proof induct + case 0 + with a show ?case by simp +next + case (pos n) + with a show ?case by (simp add: add_ac) +next + case (neg n) + with a show ?case + (* the two applications of simp cannot be combined ... *) + by (simp add: add_right_commute) (simp add: uminusSucc) +qed + +lemma times_pred_left [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "pred[a] * b = a*b - b" + using b a by induct (simp add: add_right_commute)+ + +text \ Commutativity \ + +lemma times_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" + shows "a * b = b * a" +using assms by induct (simp add: uminusSucc)+ + +lemma times_uminus_left [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a) * b = -(a * b)" + using assms by (simp add: times_commute) + +text \ Distributivity \ + +lemma add_times_distrib_left [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b + c) = a * b + a * c" + using assms by induct (simp add: uminusSucc add_right_commute add_assoc)+ + +lemma add_times_distrib_right [algebra_simps]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(b + c) * a = b * a + c * a" + using assms by (simp add: times_commute add_times_distrib_left) + +\ \ NOT added as rewrites, since sometimes they are used from right to left \ +lemmas int_distrib = add_times_distrib_left add_times_distrib_right + +(* +text \ Identity element \ + +(no longer?) used for arithmetic simplification setup +lemma times_1_right: "a \ Int \ a * 1 = a" by simp +lemma times_1_left: "a \ Int \ 1 * a = a" by simp +lemma times_uminus_1_right: "a \ Int \ a * (-1) = -a" by simp +lemma times_uminus_1_left: "a \ Int \ (-1) * a = -a" by simp +*) + +text \ Associativity \ + +lemma times_assoc [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b * c) = (a * b) * c" + using assms by induct (auto simp add: add_times_distrib_right) + +lemma times_left_commute [algebra_simps]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * (b * c) = b * (a * c)" +using assms by (simp only: times_commute times_assoc) + +(* from HOL/OrderedGroups.thy *) +lemmas times_ac = times_assoc times_commute times_left_commute + +lemma times_right_commute: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a * b * c = a * c * b" +proof - + from assms have "a * b * c = a * (b * c)" + by (rule sym[OF times_assoc]) + also from assms have "\ = a * (c * b)" + by (simp add: times_commute) + finally show ?thesis + using assms by (simp add: times_assoc) +qed + +text \ Reasoning about @{text "m * n = 0"}, etc. \ + +lemma times_is_0_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = 0) = (a = 0 \ b = 0)" +using a proof cases + case 0 + with b show ?thesis by simp +next + case (pos n) + with b show ?thesis + by cases auto +next + case na: (neg n) + with b show ?thesis + proof cases + case 0 + with na show ?thesis by simp + next + case (pos m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 0" + by (simp add: sym[OF uminusSucc]) + with na pos show ?thesis by simp + next + case (neg m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 0" + by (simp add: sym[OF uminusSucc]) + with na neg show ?thesis by simp + qed +qed + +lemma zero_istimes_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(0 = a * b) = (a = 0 \ b = 0)" + using assms by (auto dest: sym) + +lemma times_is_1_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = 1) = ((a = 1 \ b = 1) \ (a = -1 \ b = -1))" +using a proof induct + case 0 + with b show ?case by simp +next + case (pos n) + with b show ?case by cases (auto simp: add_is_0_nat) +next + case na: (neg n) + with b show ?case + proof cases + case 0 + with na show ?thesis by simp + next + case (pos m) + with na have "n * m + m + n \ Nat" + by simp + hence "pred[-(n * m + m + n)] \ 1" + by (simp add: sym[OF uminusSucc]) + with na pos show ?thesis by simp + next + case (neg m) + with na show ?thesis by (auto simp: zero_is_add_nat) + qed +qed + +lemma one_is_times_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(1 = a * b) = ((a = 1 \ b = 1) \ (a = -1 \ b = -1))" + using assms by (auto dest: sym) + +text \ Cancellation rules \ + +lemma times_cancel1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "(a * b = a * c) = (b = c \ (a = 0))" +proof - + { + assume "a \ 0" "a*b = a*c" + from b c \a*b = a*c\ have "b = c" + proof (induct arbitrary: c) + case 0 + with a \a \ 0\ show ?case by simp + next + fix n c + assume n: "n \ Nat" and c: "c \ Int" and eq: "a * succ[n] = a * c" + and ih: "\d. \d \ Int; a*n = a *d\ \ n = d" + from a n have "a * n = (a * succ[n]) - a" by simp + also have "\ = a*c - a" using eq by simp + finally have "a * n = a * pred[c]" using a c by simp + with c ih have "n = pred[c]" by blast + with n c show "succ[n] = c" by simp + next + fix n c + assume n: "n \ Nat" and c: "c \ Int" and eq: "a * -succ[n] = a * c" + and ih: "\d. \d \ Int; a*(-n) = a *d\ \ -n = d" + from a n have "a * (-n) = (a * -succ[n]) + a" by simp + also have "\ = a * c + a" using eq by simp + finally have "a * (-n) = a * succ[c]" using a c by simp + with c ih have "-n = succ[c]" by blast + with n c show "-succ[n] = c" by (simp add: uminusSucc) + qed + } + with assms show ?thesis by auto +qed + +lemma times_cancel2 [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a * c = b * c) = (a = b \ c = 0)" +proof - + from assms have "a*c = c*a" "b*c = c*b" + by (simp add: times_commute)+ + with assms show ?thesis + by simp +qed + +lemma mult_eq_self1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b = a) = (b = 1 \ a = 0)" (is "?lhs = ?rhs") +proof - + from assms have "?lhs = (a*b = a*1)" by simp + also from assms have "\ = ?rhs" by (intro times_cancel1) simp+ + finally show ?thesis . +qed + +lemma mult_eq_self2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(b * a = a) = (b = 1 \ a = 0)" + using assms by (simp add: times_commute) + + +subsection \Order on integer numbers\ + +text \ + We define the standard order @{text "\"} on integers using a conditional + definition, and we define the constant @{text "<"} such that @{text "ab \ a\b"}, for any values. +\ + +definition leq :: "[c,c] \ c" (infixl "\" 50) +where int_leq_def: "\a \ Int; b \ Int\ \ a \ b \ b - a \ Nat" + +abbreviation (input) + geq :: "[c,c] \ c" (infixl "\" 50) +where "x \ y \ y \ x" + +notation (ASCII) + leq (infixl "<=" 50) and + geq (infixl ">=" 50) + +lemma boolify_leq [simp]: "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" +by (simp add: int_leq_def) + +lemma leq_isBool [intro!,simp]: "\a \ Int; b \ Int\ \ isBool(a \ b)" +by (simp add: int_leq_def) + +lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" +by (rule ssubst) + +lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" +by (rule subst) + +subsection \ Operator definitions and generic facts about @{text "<"} \ + +definition less :: "[c,c] \ c" (infixl "<" 50) +where "a < b \ a \ b \ a \ b" + +abbreviation (input) + greater :: "[c,c] \ c" (infixl ">" 50) +where "x > y \ y < x" + +lemma boolify_less [simp]: "boolify(a < b) = (a < b)" +by (simp add: less_def) + +lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" +by (rule ssubst) + +lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" +by (rule subst) + +lemma less_imp_leq [elim!]: "a < b \ a \ b" +unfolding less_def by simp + +lemma less_irrefl [simp]: "(a < a) = FALSE" +unfolding less_def by simp + +lemma less_irreflE [elim!]: "a < a \ R" +by simp + +lemma less_not_refl: "a < b \ a \ b" +by auto + +lemma neq_leq_trans [trans]: "a \ b \ a \ b \ a < b" +by (simp add: less_def) + +declare neq_leq_trans[simplified,trans] + +lemma leq_neq_trans [trans,elim!]: "a \ b \ a \ b \ a < b" +by (simp add: less_def) + +declare leq_neq_trans[simplified,trans] + +(* Don't add to [simp]: will be tried on all disequalities! *) +lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" +by auto + +subsection \ Facts about @{text "\"} over @{text Int} \ + +lemma int_leq_refl [intro,simp]: "a \ Int \ a \ a" + by (simp add: int_leq_def) + +lemma int_leq_antisym [elim]: + assumes ab: "a \ b" and ba: "b \ a" and a: "a \ Int" and b: "b \ Int" + shows "a = b" +proof - + from a b ab have "b - a \ Nat" + by (simp add: int_leq_def) + moreover + from a b ba have "-(b - a) \ Nat" + by (simp add: int_leq_def add_commute) + ultimately have "b - a = 0" + by simp + with a b show ?thesis + by simp +qed + +declare int_leq_antisym[THEN sym, elim] + +lemma int_leq_trans [trans]: + assumes "a \ b" and "b \ c" + and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a \ c" +proof - + from a b c have "c - a = (c - b) + (b - a)" + by (simp add: add_assoc) + with assms show ?thesis + by (simp add: int_leq_def) +qed + +lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ + assumes "a \ Int" and "b \ Int" and "a = b" and "\ a \ b; b \ a \ \ P" + shows "P" + using assms by simp + +lemma int_leq_linear: + assumes "a \ Int" and "b \ Int" + shows "a \ b \ b \ a" +proof - + from assms have "b - a \ Int" by simp + hence "b - a \ Nat \ -(b - a) \ Nat" by (auto elim: intElim) + with assms show ?thesis by (simp add: int_leq_def add_commute) +qed + +lemma int_leq_uminus [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a \ -b) = (b \ a)" + using assms by (simp add: int_leq_def add_commute) + +lemma uminus_leq_zero [simp]: + "a \ Int \ (-a \ 0) = (0 \ a)" + "a \ Int \ (0 \ -a) = (a \ 0)" + by (simp add: int_leq_def)+ + +lemma int_leq_succ [iff]: + "a \ Int \ a \ succ[a]" + by (simp add: int_leq_def) + +lemma int_pred_leq [iff]: + "a \ Int \ pred[a] \ a" + by (simp add: int_leq_def) + +lemma int_leq_iff_zero_leq_diff: + (* loops, not suitable as a rewrite rule *) + assumes "a \ Int" and "b \ Int" + shows "(a \ b) = (0 \ b-a)" + using assms by (simp add: int_leq_def) + +lemma int_leqI: + "\a \ Int; b \ Int; 0 \ b-a\ \ a \ b" + by (simp add: int_leq_def) + +lemma int_leqE: + "\a \ b; a \ Int; b \ Int; 0 \ b-a \ P\ \ P" + by (simp add: int_leq_def) + +text \ + From now on, we reduce the set @{text "Nat"} to the set of + non-negative integers. +\ +lemma nat_iff_int_geq0 [simp]: "n \ Nat = (n \ Int \ 0 \ n)" + by (auto simp: int_leq_def) + +declare natIsInt [simp del] +declare succInNat [simp del] +declare predInNat [simp del] + +lemma succ_geq_0 [simp]: + assumes "a \ Int" + shows "(0 \ succ[a]) = (0 \ a \ a = -1)" + using assms succInNat[OF assms] by auto + +lemma pred_geq_0 [simp]: + assumes "a \ Int" + shows "(0 \ pred[a]) = (0 \ a \ a \ 0)" + using assms predInNat[OF assms] by auto + +lemma succ_leq_0 [simp]: + assumes "a \ Int" + shows "(succ[a] \ 0) = (a \ 0 \ a \ 0)" +proof - + from assms have "(succ[a] \ 0) = (-0 \ -succ[a])" + by (intro sym[OF int_leq_uminus]) simp+ + with assms show ?thesis by (simp add: uminusSucc) +qed + +lemma pred_leq_0 [simp]: + assumes "a \ Int" + shows "(pred[a] \ 0) = (a \ 0 \ a = 1)" +proof - + from assms have "(pred[a] \ 0) = (-0 \ -pred[a])" + by (intro sym[OF int_leq_uminus]) simp+ + with assms show ?thesis by (simp add: uminusPred) +qed + +lemma nat_zero_leq: "n \ Nat \ 0 \ n" + by simp + +lemma nat_leq_zero [simp]: + "\n \ Int; 0 \ n\ \ (n \ 0) = (n = 0)" + by auto + +lemma uminus_leq_nat: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "-m \ n" + using assms by (simp add: int_leq_trans[of "-m" "0" "n"]) + +lemma int_succ_leq_succ [simp]: + "\a \ Int; b \ Int\ \ (succ[a] \ succ[b]) = (a \ b)" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "succ[a]" "succ[b]"] + by simp + +lemma int_succ_leq_succI [intro!]: + "\a \ Int; b \ Int; a \ b\ \ succ[a] \ succ[b]" + by simp + +lemma int_succ_leq_succE [elim!]: + "\succ[a] \ succ[b]; a \ Int; b \ Int; a \ b \ P\ \ P" + by simp + +lemma int_pred_leq_pred [simp]: + "\a \ Int; b \ Int\ \ (pred[a] \ pred[b]) = (a \ b)" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "pred[a]" "pred[b]"] + by simp + +lemma int_pred_leq_predI [intro!]: + "\a \ Int; b \ Int; a \ b\ \ pred[a] \ pred[b]" + by simp + +lemma int_succ_leq_predE [elim!]: + "\pred[a] \ pred[b]; a \ Int; b \ Int; a \ b \ P\ \ P" + by simp + +lemma int_leq_succI [intro]: + "\a \ b; a \ Int; b \ Int\ \ a \ succ[b]" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "a" "succ[b]"] + by simp + +lemma int_pred_leqI [intro]: + "\a \ b; a \ Int; b \ Int\ \ pred[a] \ b" + using int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "pred[a]" "b"] + by simp + +lemma int_pred_leq_is_leq_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] \ b) = (a \ succ[b])" +proof - + from assms have "(pred[a] \ b) = (succ[pred[a]] \ succ[b])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with assms show ?thesis by simp +qed + +lemma int_leq_pred_is_succ_leq [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ pred[b]) = (succ[a] \ b)" +proof - + from assms have "(a \ pred[b]) = (succ[a] \ succ[pred[b]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with assms show ?thesis by simp +qed + +lemma int_leq_succ_iff [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a \ succ[b]) = (a \ b \ a = succ[b])" +proof - + { + assume "b - a = -1" + with assms have "b = -1 + a" + by (simp add: add_left_diff_right) + with assms have "a = succ[b]" + by simp + } + with assms int_leq_iff_zero_leq_diff[of "a" "b"] + int_leq_iff_zero_leq_diff[of "a" "succ[b]"] + show ?thesis by auto +qed + +lemma int_pred_leq_iff: + assumes a: "a \ Int" and b: "b \ Int" + shows "(pred[a] \ b) = (a \ b \ a = succ[b])" + using assms by simp + +lemma int_leq_succE [elim!]: + assumes "a \ succ[b]" and "a \ Int" and "b \ Int" + and "a \ b \ P" and "a = succ[b] \ P" + shows "P" + using assms by auto + +lemma int_pred_leqE [elim!]: + assumes "pred[a] \ b" and "a \ Int" and "b \ Int" + and "a \ b \ P" and "a = succ[b] \ P" + shows "P" + using assms by auto + +lemma int_succ_leq_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a] \ b) = (a < b)" +proof - + { + assume "succ[a] \ b" + with assms have "a \ b" + by (simp add: int_leq_trans[of "a" "succ[a]" "b"]) + moreover + have "a \ b" + proof + assume "a = b" + with assms int_leq_succ[of "a"] \succ[a] \ b\ have "succ[a] = a" + by (intro int_leq_antisym) simp+ + with assms show "FALSE" + by simp + qed + ultimately have "a < b" + by (simp add: less_def) + } + moreover + { + assume "a < b" + have "succ[a] \ b" + proof (rule int_leqI) + from assms have "b - a \ Int" by simp + thus "0 \ b - succ[a]" + proof (cases) + case 0 + with assms \a < b\ show ?thesis + by (simp add: less_def) + next + case (pos n) + with assms show ?thesis + by simp + next + case (neg n) + with assms \a < b\ show ?thesis + by (simp add: less_def int_leq_iff_zero_leq_diff[of "a" "b"]) + qed + qed (simp add: assms)+ + } + ultimately show ?thesis + using assms by (intro boolEqual) auto +qed + +lemma int_leq_pred_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ pred[b]) = (a < b)" + using assms by simp + +lemma int_leq_uminus_succ_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a \ -succ[b]) = (a < -b)" + using assms by (simp add: uminusSucc) + +lemma int_uminus_pred_leq_iff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-pred[a] \ b) = (-a < b)" + using assms by (simp add: uminusPred) + +lemma int_succ_leqI [intro]: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "succ[a] \ b" + using assms by simp + +lemma int_succ_leqE [elim!]: + assumes "succ[a] \ b" and "a \ Int" and "b \ Int" + and "\a < b\ \ P" + shows "P" + using assms by simp + +lemma int_leq_predI [intro]: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "a \ pred[b]" + using assms by simp + +lemma int_leq_predE [elim!]: + assumes "a \ pred[b]" and "a \ Int" and "b \ Int" + and "a < b \ P" + shows "P" + using assms by simp + +lemma int_not_succ_leq [simp]: + "a \ Int \ (succ[a] \ a) = FALSE" + by auto + +lemma int_not_leq_pred [simp]: + "a \ Int \ (a \ pred[a]) = FALSE" + by auto + +lemma int_leq_limit_1: + assumes "a \ b" and "\(succ[a] \ b)" and "a \ Int" and "b \ Int" + shows "a=b" + using assms by (auto simp: less_def) + +lemma int_leq_limit_2: + assumes "a \ b" and "\(a \ pred[b])" and "a \ Int" and "b \ Int" + shows "a=b" + using assms by (auto elim: int_leq_limit_1) + +text \ + Case distinction and induction rules involving @{text "\"}. +\ + +lemma int_leq_cases: + assumes a: "a \ Int" and b: "b \ Int" + and "a \ b \ P" and "\b \ a; b \ a\ \ P" + shows "P" + using assms int_leq_linear[OF a b] by auto + +lemma nat_leq_induct: \ \sometimes called ``complete induction''\ + assumes base: "P(0)" + and step: "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(succ[n])" + shows "\n\Nat : P(n)" +proof - + have "\n\Nat : \m\Nat : m \ n \ P(m)" + proof (rule natInduct) + from base show "\m \ Nat : m \ 0 \ P(m)" by auto + next + fix n + assume "n \ Nat" "\m \ Nat : m \ n \ P(m)" + with step show "\m \ Nat : m \ succ[n] \ P(m)" by auto + qed + thus ?thesis by auto +qed + +lemma nat_leq_inductE: + assumes "n \ Nat" + and "P(0)" and "\n. \n \ Nat; \m\Nat : m \ n \ P(m)\ \ P(succ[n])" + shows "P(n)" + using assms by (blast dest: nat_leq_induct) + +lemma int_leq_induct: \ \analogous induction principle for the integers\ + assumes base: "P(0)" + and step: "\n\Nat : (\i\Int : -n \ i \ i \ n \ P(i)) \ P(succ[n]) \ P(-succ[n])" + shows "\a\Int : P(a)" +proof - + have "\n\Nat : \a\Int : -n \ a \ a \ n \ P(a)" + proof (rule natInduct) + from base show "\a \ Int : -0 \ a \ a \ 0 \ P(a)" + by (auto dest: int_leq_antisym) + next + fix n + assume n: "n \ Nat" and ih: "\a \ Int : -n \ a \ a \ n \ P(a)" + show "\a \ Int : -succ[n] \ a \ a \ succ[n] \ P(a)" + proof (clarify) + fix a + assume m: "a \ Int" "-succ[n] \ a" "a \ succ[n]" + from n ih step have "P(succ[n]) \ P(-succ[n])" by blast + moreover { + assume "-n \ a" "a \ n" + with ih \a \ Int\ have "P(a)" by blast + } + ultimately show "P(a)" + using n m by (auto simp: uminusSucc) + qed + qed + moreover + { + fix n + assume "n \ Nat" + hence "-n \ Int \ n \ Int \ -n \ -n \ -n \ n \ n \ n" + by (simp add: int_leq_trans[of "-n" "0" "n"]) + } + ultimately have "\n \ Nat : P(n) \ P(-n)" by blast + thus ?thesis by (blast elim: intElim) +qed + +lemma int_leq_inductE: + assumes "a \ Int" + and "P(0)" + and "\n. \n \ Nat; \i\Int : -n \ i \ i \ n \ P(i)\ \ P(succ[n])" + and "\n. \n \ Nat; \i\Int : -n \ i \ i \ n \ P(i)\ \ P(-succ[n])" + shows "P(a)" + using assms by (blast dest: int_leq_induct) + + +subsection \ Facts about @{text "<"} over @{text Int} \ + +text \ Reduce @{text "\"} to @{text "<"}. \ + +lemma int_leq_less: + assumes "a \ Int" and "b \ Int" + shows "a \ b = (a < b \ a = b)" + using assms by (auto simp: less_def) + +lemma int_leq_lessE: + assumes "a \ b" and "a \ Int" and "b \ Int" + and "a < b \ P" and "a=b \ P" + shows "P" + using assms by (auto simp: int_leq_less) + +lemma int_less_iff_zero_less_diff: + assumes "a \ Int" and "b \ Int" + shows "(a < b) = (0 < b-a)" + using assms unfolding less_def + by (auto simp: int_leq_iff_zero_leq_diff[of "a" "b"]) + +lemma int_lessI: + "\a \ Int; b \ Int; 0 < b-a\ \ a < b" + by (simp add: int_less_iff_zero_less_diff[of "a" "b"]) + +lemma int_lessE: + "\a < b; a \ Int; b \ Int; 0 < b-a \ P\ \ P" + by (simp add: int_less_iff_zero_less_diff[of "a" "b"]) + +lemma int_leq_not_less: + assumes "a \ Int" and "b \ Int" and "a \ b" + shows "(b < a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_not_leq: + assumes "a \ Int" and "b \ Int" and "a < b" + shows "(b \ a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +text \ + Declaring the above lemmas as simplification rules makes the prover + very slow, but the instances where @{text "a"} is @{text "0"} are useful. +\ +lemmas zero_leq_not_less_zero [simp] = int_leq_not_less[of 0, simplified] +lemmas zero_less_not_leq_zero [simp] = int_less_not_leq[of 0, simplified] + +lemma int_less_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a < succ[b]) = (a \ b)" + using assms by (auto simp: less_def int_leq_succI) + +lemmas int_leq_is_less_succ = sym[OF int_less_succ] + +lemma int_less_succE: + assumes "a < succ[b]" and "a \ Int" and "b \ Int" + and "a < b \ P" and "a = b \ P" + shows P + using assms by (auto elim: int_leq_lessE) + +lemma int_pred_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] < b) = (a \ b)" + using assms by (auto simp: less_def) + +lemmas int_leq_is_pred_less = sym[OF int_pred_less] + +lemma not_leq_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "((a \ b) = FALSE) = (b < a)" + using assms int_leq_linear[of a b] + by (auto simp: less_def dest: int_leq_antisym) + +lemma not_less_leq [simp]: + assumes "a \ Int" and "b \ Int" + shows "((a < b) = FALSE) = (b \ a)" + using assms int_leq_linear[of a b] + by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_uminus [simp]: + assumes "a \ Int" and "b \ Int" + shows "(-a < -b) = (b < a)" + using assms by (auto simp: less_def) + +lemma int_uminus_zero [simp]: + "a \ Int \ (-a < 0) = (0 < a)" + "a \ Int \ (0 < -a) = (a < 0)" + by (auto simp: less_def) + +text \ @{text "<"} and @{text "succ / pred"}. \ + +lemma int_succ_less_succ [simp]: + "\a \ Int; b \ Int\ \ (succ[a] < succ[b]) = (a < b)" + by (auto simp: less_def int_leq_succI) + +lemma int_succ_less_succI [intro]: + "\a \ Int; b \ Int; a < b\ \ succ[a] < succ[b]" + by simp + +lemma int_succ_less_succE [elim]: + assumes "succ[a] < succ[b]" and "a \ Int" and "b \ Int" and "a < b \ P" + shows "P" + using assms by simp + +lemma int_pred_less_pred [simp]: + "\a \ Int; b \ Int\ \ (pred[a] < pred[b]) = (a < b)" + by (simp add: less_def) + +lemma int_pred_less_predI [intro]: + "\a \ Int; b \ Int; a < b\ \ pred[a] < pred[b]" + by simp + +lemma int_pred_less_predE [elim]: + assumes "pred[a] < pred[b]" and "a \ Int" and "b \ Int" and "a < b \ P" + shows "P" + using assms by simp + +lemma int_less_succ_same [iff]: + "a \ Int \ a < succ[a]" + by (simp add: less_def) + +lemma int_pred_less_same [iff]: + "a \ Int \ pred[a] < a" + by (simp add: less_def) + +lemma int_pred_less_is_less_succ [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] < b) = (a < succ[b])" + using assms by (auto simp: less_def) + +lemma int_less_pred_is_succ_less [simp]: + assumes "a \ Int" and "b \ Int" + shows "(a < pred[b]) = (succ[a] < b)" + using assms by (auto simp: less_def) + +lemma succ_nat_not_neg [simp]: + "\n \ Int; 0 \ n\ \ (succ[n] < 0) = FALSE" + by (auto simp: less_def) + +lemma int_less_leq_not_leq: + assumes "a \ Int" and "b \ Int" + shows "(a < b) = (a \ b \ \ b \ a)" + using assms by (auto simp: less_def dest: int_leq_antisym) + +text \ Transitivity. \ + +lemma int_less_trans [trans]: + assumes "a < b" and "b < c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma nat_less_trans_Succ: \ \stronger than the above\ + assumes lt1: "a < b" and lt2: "b < c" + and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "succ[a] < c" + using assms +proof - + from a b lt1 have "succ[succ[a]] \ succ[b]" + by (auto simp: less_def int_leq_succI) + also from b c lt2 have "\ \ c" + by (auto simp: less_def) + finally show ?thesis + using a b c by (auto simp: less_def) +qed + +lemma int_leq_less_trans [trans]: + assumes "a \ b" and "b < c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma int_less_leq_trans [trans]: + assumes "a < b" and "b \ c" and "a \ Int" and "b \ Int" and "c \ Int" + shows "a < c" + using assms unfolding less_def + by (auto elim!: int_leq_trans dest: int_leq_antisym) + +lemma nat_zero_less_succ [iff]: + "\n \ Int; 0 \ n\ \ 0 < succ[n]" + by (simp add: int_leq_less_trans) + +text \ Asymmetry. \ + +lemma int_less_not_sym: + assumes "a < b" and "a \ Int" and "b \ Int" + shows "(b < a) = FALSE" + using assms by (auto simp: less_def dest: int_leq_antisym) + +lemma int_less_asym: + assumes 1: "a < b" and 2: "a \ Int" and 3: "b \ Int" and 4: "\P \ b < a" + shows "P" + using int_less_not_sym[OF 1 2 3] 4 by auto + +text \ Linearity (totality). \ + +lemma int_less_linear: + assumes a: "a \ Int" and b: "b \ Int" + shows "a < b \ a = b \ b < a" + unfolding less_def using int_leq_linear[OF a b] by blast + +lemma int_leq_less_linear: + assumes a: "a \ Int" and b: "b \ Int" + shows "a \ b \ b < a" + using assms int_less_linear[OF a b] by (auto simp: less_def) + +lemma int_less_cases [case_names less equal greater]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a < b \ P) \ (a = b \ P) \ (b < a \ P) \ P" + using int_less_linear[OF a b] by blast + +lemma int_neq_iff: + assumes a: "a \ Int" and b: "b \ Int" + shows "((a = b) = FALSE) = (a < b \ b < a)" + using assms int_less_linear[OF a b] by auto + +lemma int_neq_lessE: + assumes "a \ b" and "a \ Int" and "b \ Int" + shows "(a < b \ P) \ (b < a \ P) \ P" + using assms by (auto simp: int_neq_iff) + +lemma int_antisym_conv1: + assumes "\ a < b" and "a \ Int" and "b \ Int" + shows "(a \ b) = (a = b)" + using assms by (auto dest: int_leq_antisym) + +lemma int_antisym_conv2: + assumes "a \ b" and "a \ Int" and "b \ Int" + shows "(\ a < b) = (b = a)" + using assms by (auto dest: int_leq_antisym) + +lemma int_antisym_conv3: + assumes "\ a < b" and "a \ Int" and "b \ Int" + shows "(\ b < a) = (a = b)" + using assms by (auto dest: int_leq_antisym) + +lemma nat_gt0_not0: + assumes "n \ Int" and "0 \ n" + shows "(0 < n) = (n \ 0)" + using assms by auto + +lemmas nat_neq0_conv = sym[OF nat_gt0_not0, simplified, simp] + +lemma nat_less_one (*[simp]*): + assumes "n \ Int" and "0 \ n" + shows "(n < 1) = (n = 0)" + using assms by auto + +text \ "Less than" is antisymmetric, sort of \ + +lemma int_less_antisym: + assumes "a \ Int" and "b \ Int" + shows "\ \(b < a); b < succ[a] \ \ a = b" + using assms by (auto dest: int_leq_antisym) + +text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ +lemma int_less_mono_imp_leq_mono: + assumes ij: "i \ j" and i: "i \ Int" and j: "j \ Int" + and f: "\a \ Int : f(a) \ Int" + and mono: "\i j. \i \ Int; j \ Int; i \ f(i) < f(j)" + shows "f(i) \ f(j)" +proof (rule int_leq_lessE[OF ij i j]) + assume "i < j" + with i j have "f(i) < f(j)" by (rule mono) + thus ?thesis by auto +next + assume "i = j" + with i f show ?thesis by simp +qed + +lemma nat_less_mono_imp_leq_mono: + assumes ij: "i \ j" and i: "i \ Nat" and j: "j \ Nat" + and f: "\a \ Nat : f(a) \ Int" + and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" + shows "f(i) \ f(j)" +proof (rule int_leq_lessE[OF ij]) + assume "i < j" + with i j have "f(i) < f(j)" by (rule mono) + thus ?thesis by auto +next + assume "i = j" + with i f show ?thesis by simp +qed (auto simp: i[simplified] j[simplified]) + +lemma int_succ_lessI: + assumes "a \ Int" and "b \ Int" and "a < b" and "succ[a] \ b" + shows "succ[a] < b" + using assms by (auto simp: less_def) + +lemma int_less_antisym_false: "\a < b; a \ Int; b \ Int\ \ b < a = FALSE" + by auto + +lemma nat_less_antisym_leq_false: "\a < b; a \ Int; b \ Int\ \ b \ a = FALSE" + by auto + + +subsection \ Additional arithmetic theorems \ + +subsubsection \ Monotonicity of Addition \ + +lemma int_add_left_cancel_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(i + j \ i + k) = (j \ k)" + using assms by (induct i) (simp del: int_leq_succ_iff)+ + +lemma int_add_right_cancel_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(j + i \ k + i) = (j \ k)" + using assms by (induct i) (simp del: int_leq_succ_iff)+ + +lemma int_add_left_cancel_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(i + j < i + k) = (j < k)" + using assms by (simp add: less_def) + +lemma int_add_right_cancel_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(j + i < k + i) = (j < k)" + using assms by (simp add: less_def) + +lemma int_add_left_cancel_succ_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[i + j] \ i + k) = (succ[j] \ k)" +using assms proof (induct i) + case (neg n) + then have "(-n + j \ pred[-n + k]) = (succ[-n+j] \ succ[pred[-n + k]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with neg show ?case by simp +qed (simp del: int_leq_succ_iff)+ + +lemma int_add_right_cancel_succ_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[j + i] \ k + i) = (succ[j] \ k)" +using assms proof (induct i) + case (neg n) + then have "(j - n \ pred[k - n]) = (succ[j-n] \ succ[pred[k-n]])" + by (intro sym[OF int_succ_leq_succ]) simp+ + with neg show ?case by simp +qed (simp del: int_leq_succ_iff)+ + +lemma int_add_left_cancel_pred_leq [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(pred[i + j] \ i + k) = (pred[j] \ k)" +proof - + from assms have "pred[i+j] = i + pred[j]" by simp + hence "(pred[i + j] \ i + k) = (i + pred[j] \ i + k)" by simp + also have "\ = (pred[j] \ k)" + using assms by (intro int_add_left_cancel_leq) simp+ + finally show ?thesis . +qed + +lemma int_add_right_cancel_pred_less [simp]: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(pred[j + i] \ k + i) = (pred[j] \ k)" +proof - + from assms have "pred[j+i] = pred[j] + i" by simp + hence "(pred[j+i] \ k + i) = (pred[j] + i \ k + i)" by simp + also have "\ = (pred[j] \ k)" + using assms by (intro int_add_right_cancel_leq) simp+ + finally show ?thesis . +qed + +lemma int_leq_imp_add: + assumes i: "i \ Int" and j: "j \ Int" and ij: "i \ j" + shows "\k \ Nat : j = i+k" +proof - + from assms have "j-i \ Nat" + by (simp add: int_leq_iff_zero_leq_diff[of "i" "j"]) + moreover + from i j have "j = i + (j-i)" + by (simp add: add_ac) + ultimately show ?thesis .. +qed + +lemma int_less_imp_succ_add: + assumes ij: "i < j" and i: "i \ Int" and j: "j \ Int" + shows "\k \ Nat: j = succ[i + k]" +proof - + from assms have "succ[i] \ j" + by simp + with i j have "j - succ[i] \ Nat" + by (simp add: int_less_iff_zero_less_diff[of "i" "j"]) + moreover + from i j have "j = succ[i + (j - succ[i])]" + by (simp add: add_ac) + ultimately show ?thesis .. +qed + + +subsubsection \ (Partially) Ordered Groups \ + +\ \ The two following lemmas are just ``one half'' of + @{text int_add_left_cancel_leq} and @{text int_add_right_cancel_leq} + proved above." \ +lemma add_leq_left_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a \ b \ c + a \ c + b" + using assms by simp + +lemma add_leq_right_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a \ b \ a + c \ b + c" + using assms by simp + +text \ non-strict, in both arguments \ +lemma add_leq_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a \ b \ c \ d \ a + c \ b + d" + using assms + add_leq_right_mono[of a b c] + add_leq_left_mono[of c d b] + int_leq_trans[of "a + c" "b + c" "b + d"] + by simp + +\ \ Similar for strict less than. \ +lemma add_less_left_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a < b \ c + a < c + b" + using assms by (simp add: add_leq_left_mono less_def) + +lemma add_less_right_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "a < b \ a + c < b + c" + using assms by (simp add: add_leq_right_mono less_def) + +text \ Strict monotonicity in both arguments \ +lemma add_less_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a < b \ c < d \ a + c < b + d" + using assms + add_less_right_mono[of a b c] + add_less_left_mono[of c d b] + int_less_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma add_less_leq_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a < b \ c \ d \ a + c < b + d" + using assms + add_less_right_mono[of a b c] + add_leq_left_mono[of c d b] + int_less_leq_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma add_leq_diff1: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b \ c) = (a \ c - b)" +proof - + from assms have "(a + b \ c) = ((a + b) - b \ c - b)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis by simp +qed + +lemma add_leq_diff2: + assumes "a \ Int" and "b \ Int" and "c \ Int" + shows "(a + b \ c) = (b \ c - a)" +proof - + from assms have "(a + b \ c) = ((a + b) - a \ c - a)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis by simp +qed + +lemma add_leq_less_mono: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "d \ Int" + shows "a \ b \ c < d \ a + c < b + d" + using assms + add_leq_right_mono[of a b c] + add_less_left_mono[of c d b] + int_leq_less_trans[of "a + c" "b + c" "b + d"] + by simp + +lemma leq_add_nat1 [iff]: + assumes "a \ Int" and "0 \ a" and "b \ Int" + shows "b \ b + a" + using assms add_leq_left_mono[of 0 a b] by simp + +lemma leq_add_nat2 [iff]: + assumes "a \ Int" and "0 \ a" and "b \ Int" + shows "b \ a + b" + using assms add_leq_right_mono [of 0 a b] by simp + +lemma trans_leq_add_nat1 [simp] (* was [elim]*): + assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" + shows "a \ b+m" + using assms by (simp add: int_leq_trans[of "a" "b" "b+m"]) + +lemma trans_leq_add_nat2 [simp] (* was [elim] *): + assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" + shows "a \ m+b" + using assms by (simp add: int_leq_trans[of "a" "b" "m+b"]) + +lemma add_nat_leqD1 [elim]: + assumes "a+k \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "a \ b" + using assms by (simp add: int_leq_trans[of "a" "a+k" "b"]) + +lemma add_nat_leqD2 [elim]: + assumes "k+a \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "a \ b" + using assms by (simp add: int_leq_trans[of "a" "k+a" "b"]) + +lemma add_nat_pred_leqD1 [elim]: + assumes "pred[a+k] \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "pred[a] \ b" +proof - + { + assume 1: "a + k = succ[b]" + have "a \ succ[b]" + proof (rule int_leqI) + from assms 1 have "succ[b] - a = k" + by (simp add: add_left_diff_right add_ac) + with \0 \ k\ show "0 \ succ[b] - a" + by simp + qed (simp add: assms)+ + } + with assms show ?thesis + by (auto simp del: int_leq_succ_iff) +qed + +lemma add_nat_pred_leqD2 [elim]: + assumes "pred[k+a] \ b" and "a \ Int" and "b \ Int" and "k \ Int" and "0 \ k" + shows "pred[a] \ b" +proof (rule add_nat_pred_leqD1) + from assms show "pred[a+k] \ b" + by (simp add: add_commute) +qed (simp add: assms)+ + +lemma less_succ_add_nat1 (*[simp]*): + assumes "i \ Int" and "m \ Int" and "0 \ m" + shows "i < succ[i + m]" + using assms by simp + +lemma less_add_Succ2 (*[simp]*): + assumes "i \ Int" and "m \ Int" and "0 \ m" + shows "i < succ[m + i]" + using assms by simp + +lemma diff_nat_leq_self1 [simp]: + assumes "i \ Int" and "n \ Int" and "0 \ n" + shows "i - n \ i" +proof - + from assms have "i-n \ (i-n)+n" + by (intro leq_add_nat1) simp+ + with assms show ?thesis by simp +qed + +lemma diff_nat_leq_self2 [simp]: + assumes "i \ Int" and "n \ Int" and "0 \ n" + shows "-n + i \ i" + using assms by (simp add: add_commute) + +lemma trans_leq_diff_nat1 [simp]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i-m \ j" + using assms by (simp add: int_leq_trans[of "i-m" "i" "j"]) + +lemma trans_leq_diff_nat2 [simp]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "-m + i \ j" + using assms by (simp add: int_leq_trans[of "-m + i" "i" "j"]) + +lemma int_leq_iff_add: + assumes "i \ Int" and "j \ Int" + shows "(i \ j) = (\k \ Nat: j = i + k)" (is "?lhs = ?rhs") + using assms by (auto intro: int_leq_imp_add) + +lemma int_less_iff_succ_add: + assumes "i \ Int" and "j \ Int" + shows "(i < j) = (\k \ Nat: j = succ[i + k])" (is "?lhs = ?rhs") + using assms by (auto intro: int_less_imp_succ_add) + +lemma leq_add_self1 [simp]: + assumes "i \ Int" and "j \ Int" + shows "(i + j \ i) = (j \ 0)" +proof - + from assms have "(i+j \ i) = (i+j-i \ i-i)" + by (intro sym[OF int_add_right_cancel_leq]) simp+ + with assms show ?thesis + by simp +qed + +lemma leq_add_self2 [simp]: + assumes "i \ Int" and "j \ Int" + shows "(j + i \ i) = (j \ 0)" + using assms by (simp add: add_commute) + +lemma trans_less_add_nat1 [simp]: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i < j + m" + using assms by (simp add: int_less_leq_trans[of "i" "j" "j+m"]) + +lemma trans_less_add_nat2 [simp]: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i < m + j" + using assms by (simp add: int_less_leq_trans[of "i" "j" "m+j"]) + +lemma trans_less_diff_nat1: + assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" + shows "i - m < j" + using assms by (simp add: int_leq_less_trans[of "i-m" "i" "j"]) + +lemma add_nat_lessD1: + assumes "i+k < j" and "i \ Int" and "j \ Int" and "k \ Int" and "0 \ k" + shows "i < j" + using assms by (simp add: int_leq_less_trans[of "i" "i+k" "j"]) + +lemma add_lessD2: + assumes "k+i < j" and "i \ Int" and "j \ Int" and "k \ Int" and "0 \ k" + shows "i < j" + using assms by (simp add: int_leq_less_trans[of "i" "k+i" "j"]) + +lemma nat_add_gt_0: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "(m + n > 0) = (m > 0 \ n > 0)" + using assms by auto + +lemma add_ge_1: + assumes "m \ Int" and "0 \ m" and "n \ Int" and "0 \ n" + shows "(m + n \ 1) = (m \ 1 \ n \ 1)" + using assms by auto + +lemma less_imp_add_positive: + assumes "i < j" and "i \ Int" and "j \ Int" + shows "\k \ Nat: 0 < k \ i + k = j" +proof - + from assms obtain k where k: "k \ Nat" "i + succ[k] = j" + by (auto dest: int_less_imp_succ_add) + from k have "0 < succ[k]" "succ[k] \ Nat" by simp+ + with assms k show ?thesis by blast +qed + + +subsubsection \ Monotonicity of Multiplication \ + +lemma times_leq_left_pos_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "c * a \ c * b" + using c by induct (simp add: add_leq_mono 1 a b)+ + +lemma times_leq_left_neg_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "(-c) * b \ (-c) * a" +using c proof (induct) + case 0 + with a b show ?case by simp +next + case (Succ n) + from a b 1 have "-b \ -a" by simp + with a b Succ show ?case by (simp add: add_leq_mono) +qed + +lemma times_leq_right_pos_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "a * c \ b * c" + using c by induct (simp add: add_leq_mono 1 a b)+ + +lemma times_leq_right_neg_mono: + assumes 1: "a \ b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Nat" + shows "b * (-c) \ a * (-c)" +using c proof (induct) + case 0 + with a b show ?case by simp +next + case (Succ n) + from a b 1 have "-b \ -a" by simp + with a b Succ show ?case by (simp add: add_leq_mono) +qed + +lemma nat_leq_times_left: + assumes a: "a \ Nat" and b: "b \ Int" and 1: "1 \ b" + shows "a \ a * b" +proof - + have "1 \ Int" by simp + from 1 this b a have "a * 1 \ a * b" by (rule times_leq_left_pos_mono) + with a b show ?thesis by simp +qed + +lemma nat_leq_times_right: + assumes "a \ Nat" and "b \ Int" and "1 \ b" + shows "a \ b * a" + using assms by (simp add: nat_leq_times_left times_commute) + +text \ Similar lemmas for @{text "<"} \ + +lemma times_less_left_pos_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "c * a < c * b" +proof - + from c 0 obtain n where "n \ Nat" "c = succ[n]" + by cases auto + from \n \ Nat\ have "succ[n] * a < succ[n] * b" + by induct (simp add: add_less_mono 1 a b)+ + with \c = succ[n]\ show ?thesis by simp +qed + +lemma times_less_left_neg_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "c * b < c * a" +proof - + from c 0 obtain n where "n \ Nat" "c = -succ[n]" + by cases auto + from a b 1 have "-b < -a" by simp + with \n \ Nat\ have "(-succ[n]) * b < (-succ[n]) * a" + by (induct) (simp add: add_less_mono a b)+ + with \c = -succ[n]\ show ?thesis by simp +qed + +lemma times_less_right_pos_mono: + assumes "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and "0 < c" + shows "a * c < b * c" + using times_less_left_pos_mono[OF assms] a b c by (simp add: times_commute) + +lemma times_less_right_neg_mono: + assumes 1: "a < b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and "c < 0" + shows "b * c < a * c" + using times_less_left_neg_mono[OF assms] a b c by (simp add: times_commute) + +lemma nat_less_times_left: + assumes a: "a \ Int" and 0: "0 < a" and b: "b \ Int" and 1: "1 < b" + shows "a < a * b" +proof - + have "1 \ Int" by simp + from 1 this b a 0 have "a * 1 < a * b" by (rule times_less_left_pos_mono) + with a b show ?thesis by simp +qed + +lemma nat_less_times_right: + assumes "a \ Int" and 0: "0 < a" and "b \ Int" and "1 < b" + shows "a < b * a" + using assms by (simp add: nat_less_times_left times_commute) + +lemma times_leq_left_pos_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "(c * a \ c * b) = (a \ b)" (is "?lhs = ?rhs") +proof (rule boolEqual) + { + assume lhs: "?lhs" and contr: "\?rhs" + from a b contr have "b < a" by simp + from this b a c 0 have "c * b < c * a" by (rule times_less_left_pos_mono) + with a b c have "\?lhs" by simp + from this lhs have "FALSE" .. + } + moreover + from c 0 have "c \ Nat" + by (simp add: less_def) + ultimately show "?lhs \ ?rhs" + using times_leq_left_pos_mono[of a b c] assms by blast +qed (simp add: a b c)+ + +lemma times_leq_right_pos_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "0 < c" + shows "(a * c \ b * c) = (a \ b)" (is "?lhs = ?rhs") +proof - + from a b c have "(a * c \ b * c) = (c * a \ c * b)" + by (simp add: times_commute) + with assms show ?thesis + by simp +qed + +lemma times_leq_left_neg_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "(c * a \ c * b) = (b \ a)" (is "?lhs = ?rhs") +proof (rule boolEqual) + { + assume lhs: "?lhs" and contr: "\?rhs" + from a b contr have "a < b" by simp + from this a b c 0 have "c * b < c * a" by (rule times_less_left_neg_mono) + with a b c have "\?lhs" by simp + from this lhs have "FALSE" .. + } + moreover + from c 0 obtain d where "d \ Nat" "c = -d" + by (cases c) force+ + ultimately show "?lhs \ ?rhs" + using times_leq_left_neg_mono[of b a d] assms by blast +qed (simp add: a b c)+ + +lemma times_leq_right_neg_cancel [simp]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and 0: "c < 0" + shows "(a * c \ b * c) = (b \ a)" (is "?lhs = ?rhs") +proof - + from a b c have "(a * c \ b * c) = (c * a \ c * b)" + by (simp add: times_commute) + with assms show ?thesis + by simp +qed + +lemma times_less_left_pos_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(c * a < c * b) = (a < b)" + unfolding less_def using assms by auto + +lemma times_less_left_neg_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "c < 0" + shows "(c * a < c * b) = (b < a)" + unfolding less_def using assms by auto + +lemma times_less_right_pos_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * c < b * c) = (a < b)" + unfolding less_def using assms by auto + +lemma times_less_right_neg_cancel [simp]: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "c < 0" + shows "(a * c < b * c) = (b < a)" + unfolding less_def using assms by auto + +lemma times_pos [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(0 < a * b) = ((0 < a \ 0 < b) \ (a < 0 \ b < 0))" +using a proof (cases) + case 0 + with b show ?thesis by simp +next + case (pos n) + hence a': "a \ Nat" "0 < a" + by auto + from a have "(0 < a * b) = (a * 0 < a * b)" + by simp + also from a \0 < a\ b have "\ = (0 < b)" + by (intro times_less_left_pos_cancel) simp+ + finally show ?thesis + using a' by simp +next + case (neg n) + hence a': "a \ Int" "a < 0" + by auto + from a have "(0 < a * b) = (a * 0 < a * b)" + by simp + also from a' b have "\ = (b < 0)" + by (intro times_less_left_neg_cancel) simp+ + finally show ?thesis + using a' by (simp add: int_less_not_sym[of a 0]) +qed + +lemma times_neg: + assumes a: "a \ Int" and b: "b \ Int" + shows "(a * b < 0) = ((0 < a \ b < 0) \ (a < 0 \ 0 < b))" +proof - + from assms have "(a * b < 0) = (0 < (-a) * b)" + by simp + also from assms have "\ = ((0 < -a \ 0 < b) \ (-a < 0 \ b < 0))" + by (intro times_pos) simp+ + finally show ?thesis + using assms by auto +qed + + +subsection \Intervals of integers\ + +definition intvl :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) +where "i .. j \ { k \ Int : i \ k \ k \ j }" + +lemma intvlI [iff]: + assumes "k \ Int" and "i \ k" and "k \ j" + shows "k \ i .. j" + using assms by (simp add: intvl_def) + +lemma intvlE [elim]: + assumes "k \ i .. j" and "\k \ Int; i \ k; k \ j\ \ P" + shows P + using assms by (simp add: intvl_def) + +lemma intvlIff [simp]: "(k \ i .. j) = (k \ Int \ i \ k \ k \ j)" + by auto + +lemmas + setEqualI [where A = "i .. j" for i j, intro] + setEqualI [where B = "i .. j" for i j, intro] + +lemma gtNotinIntvl: + assumes gt: "i > j" and k: "k \ i .. j" and i: "i \ Int" and j: "j \ Int" + shows "P" +proof - + from k have "k \ Int" "i \ k" "k \ j" by auto + with i j have "i \ j" by (auto elim: int_leq_trans) + with i j have "\(j < i)" by simp + from this gt show ?thesis .. +qed + +lemma intvlIsEmpty: + assumes "i \ Int" and "j \ Int" and "i > j" + shows "i .. j = {}" + using assms by (blast dest: gtNotinIntvl) + +lemma intvlEmpty_iff: + assumes i: "i \ Int" and j: "j \ Int" + shows "(i .. j = {}) = (i > j)" +proof - + { + assume mt: "i .. j = {}" + have "i > j" + proof (rule contradiction) + assume "\(i > j)" + with assms have "i \ j" by simp + with i have "i \ i ..j" by simp + with mt show "FALSE" by simp + qed + } + moreover + from i j have "i > j \ i..j = {}" + by (blast dest: intvlIsEmpty) + ultimately show ?thesis + by blast +qed + +lemma intvlSingleton [simp]: + "i \ Int \ i .. i = {i}" + by (auto dest: int_leq_antisym) + +lemma intvlSucc [simp]: + assumes "i \ Int" and "j \ Int" and "i \ succ[j]" + shows "i .. succ[j] = (i..j \ {succ[j]})" + using assms by auto + +lemma succIntvl: + assumes "i \ Int" and "j \ Int" + shows "succ[i] .. j = (i .. j \ {i})" + using assms unfolding intvl_def by auto + +lemma intvlEqual_iff: + assumes i: "i \ Int" and j: "j \ Int" and k: "k \ Int" and l: "l \ Int" + shows "(i .. j = k .. l) = ((i > j \ k > l) \ (i = k \ j = l))" (is "?lhs = ?rhs") +proof (cases "i > j") + case True + with i j have "i .. j = {}" + by (simp add: intvlEmpty_iff) + moreover + from k l have "(k .. l = {}) = (k > l)" + by (simp add: intvlEmpty_iff) + ultimately show ?thesis + using assms True by force +next + case False + with i j have ij: "i \ j" by simp + with i j have "i \ i .. j" "j \ i .. j" by simp+ + { + assume eq: "i..j = k..l" and idx: "i \ k \ j \ l" + from eq \i \ i..j\ intvlIsEmpty[OF k l] have "\(k > l)" + by blast + with k l have "k \ k..l" "l \ k..l" by simp+ + from idx have "FALSE" + proof + assume "i \ k" + from this i k show "FALSE" + proof (rule int_neq_lessE) + assume "i < k" + with i k have "i \ k .. l" by auto + with eq \i \ i .. j\ show "FALSE" by simp + next + assume "k < i" + with i k have "k \ i .. j" by auto + with eq \k \ k..l\ show "FALSE" by simp + qed + next + assume "j \ l" + from this j l show "FALSE" + proof (rule int_neq_lessE) + assume "j < l" + with j l have "l \ i .. j" by auto + with eq \l \ k ..l\ show "FALSE" by simp + next + assume "l < j" + with j l have "j \ k..l" by auto + with eq \j \ i..j\ show "FALSE" by simp + qed + qed + } + thus ?thesis using False by auto +qed + +lemma zerotoInj [simp]: + assumes "l \ Nat" and "m \ Int" and "n \ Int" + shows "(0 .. l = m .. n) = (m=0 \ l=n)" +using assms by (auto simp: intvlEqual_iff) + +lemma zerotoInj' [simp]: + assumes "k \ Int" and "l \ Int" and "n \ Nat" + shows "(k .. l = 0 .. n) = (k=0 \ l=n)" +using assms by (auto simp: intvlEqual_iff) + +lemma zerotoEmpty [simp]: + assumes "m \ Nat" + shows "succ[m] .. 0 = {}" + using assms by (intro intvlIsEmpty) auto + +lemma onetoInj [simp]: + assumes "l \ Nat" and "m \ Int" and "n \ Int" and "l \ 0 \ 1 \ l" + shows "(1 .. l = m .. n) = (m=1 \ l=n)" + using assms by (auto simp: intvlEqual_iff) + +lemma onetoInj' [simp]: + assumes "k \ Int" and "l \ Int" and "n \ Nat" and "n\0 \ 1 \ n" + shows "(k .. l = 1 .. n) = (k=1 \ l=n)" + using assms by (auto simp: intvlEqual_iff) + +lemma SuccInIntvlSucc: + assumes "m \ k" and "k \ Int" and "m \ Int" and "n \ Int" + shows "(succ[k] \ m .. succ[n]) = (k \ m .. n)" + using assms by auto + +lemma SuccInIntvlSuccSucc: + assumes "i \ Int" and "j \ Int" and "k \ Int" + shows "(succ[i] \ succ[j] .. succ[k]) = (i \ j .. k)" + using assms by auto + +end diff --git a/isabelle/IntegerDivision.thy b/isabelle/IntegerDivision.thy new file mode 100644 index 00000000..2da1b4eb --- /dev/null +++ b/isabelle/IntegerDivision.thy @@ -0,0 +1,713 @@ +(* Title: TLA+/IntegerDivision.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation + License: BSD + Version: Isabelle2021-1 +*) + +section \ The division operators div and mod on the integers \ + +theory IntegerDivision +imports IntegerArithmetic +begin + +subsection \ The divisibility relation \ + +definition dvd (infixl "dvd" 50) +where "a \ Int \ b \ Int \ b dvd a \ (\k \ Int : a = b * k)" + +lemma boolify_dvd [simp]: + assumes "a \ Int" and "b \ Int" + shows "boolify(b dvd a) = (b dvd a)" +using assms by (simp add: dvd_def) + +lemma dvdIsBool [intro!,simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "isBool(b dvd a)" +using assms by (simp add: dvd_def) + +lemma [intro!]: + "\isBool(P); isBool(a dvd b); (a dvd b) \ P\ \ (a dvd b) = P" + "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" +by auto + +lemma dvdI [intro]: + assumes a: "a \ Int" and b: "b \ Int" and k: "k \ Int" + shows "a = b * k \ b dvd a" +unfolding dvd_def[OF a b] using k by blast + +lemma dvdE [elim]: + assumes "b dvd a" and "a \ Int" and "b \ Int" + shows "(\k. \k \ Int; a = b * k\ \ P) \ P" +using assms by (auto simp: dvd_def) + +lemma dvd_refl [iff]: + assumes a: "a \ Int" + shows "a dvd a" +proof - + from a have "a = a*1" by simp + with a show ?thesis by blast +qed + +lemma dvd_trans [trans]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and 1: "a dvd b" and 2: "b dvd c" + shows "a dvd c" +proof - + from a b 1 obtain k where k: "k \ Int" and "b = a * k" by blast + moreover + from b c 2 obtain l where l: "l \ Int" and "c = b * l" by blast + ultimately have h: "c = a * (k * l)" + using a b c by (simp add: times_assoc) + thus ?thesis using a c k l by blast +qed + +lemma dvd_0_left_iff [simp]: + assumes "a \ Int" + shows "(0 dvd a) = (a = 0)" +using assms by auto + +lemma dvd_0_right [iff]: + assumes a: "a \ Int" shows "a dvd 0" +using assms by force + +lemma one_dvd [iff]: + assumes a: "a \ Int" + shows "1 dvd a" +using assms by force + +lemma dvd_mult (*[simp]*): + assumes dvd: "a dvd c" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a dvd (b * c)" +proof - + from dvd a c obtain k where k: "k \ Int" and "c = a*k" by blast + with a b c have "b*c = a * (b*k)" by (simp add: times_left_commute) + with a b c k show ?thesis by blast +qed + +lemma dvd_mult2 (*[simp]*): + assumes dvd: "a dvd b" and a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + shows "a dvd (b * c)" +using times_commute[OF b c] dvd_mult[OF dvd a c b] by simp + +lemma dvd_triv_right [iff]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a dvd (b * a)" +using assms by (intro dvd_mult) simp+ + +lemma dvd_triv_left [iff]: + assumes a: "a \ Int" and b: "b \ Int" + shows "a dvd a * b" +using assms by (intro dvd_mult2) simp+ + +lemma mult_dvd_mono: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and d: "d \ Int" + and 1: "a dvd b" and 2: "c dvd d" + shows "(a * c) dvd (b * d)" +proof - + from a b 1 obtain b' where b': "b = a * b'" "b' \ Int" by blast + from c d 2 obtain d' where d': "d = c * d'" "d' \ Int" by blast + with b' a b c d + times_assoc[of a b' "c * d'"] + times_left_commute[of b' c d'] + times_assoc[of a c "b' * d'"] + have "b * d = (a * c) * (b' * d')" by simp + with a c b' d' show ?thesis by blast +qed + +lemma dvd_mult_left: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and h: "a * b dvd c" + shows "a dvd c" +proof - + from h a b c obtain k where k: "k \ Int" "c = a*(b*k)" + by (auto simp add: times_assoc) + with a b c show ?thesis by blast +qed + +lemma dvd_mult_right: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" and h: "a*b dvd c" + shows "b dvd c" +proof - + from h a b c have "b*a dvd c" by (simp add: times_ac) + with b a c show ?thesis by (rule dvd_mult_left) +qed + +lemma dvd_0_left: + assumes "a \ Int" + shows "0 dvd a \ a = 0" + using assms by simp + +lemma dvd_add [iff]: + assumes a: "a \ Int" and b: "b \ Int" and c: "c \ Int" + and 1: "a dvd b" and 2: "a dvd c" + shows "a dvd (b + c)" +proof - + from a b 1 obtain b' where b': "b' \ Int" "b = a * b'" by blast + from a c 2 obtain c' where c': "c' \ Int" "c = a * c'" by blast + from a b c b' c' + have "b + c = a * (b' + c')" by (simp add: add_times_distrib_left) + with a b' c' show ?thesis by blast +qed + +subsection \ Division and modulus on @{const Int} \ + +text \ + \tlaplus{} defines integer division only for the case + where the divisor is positive. Two different extensions + are commonly introduced for negative divisors, and we + leave this case undefined in order to avoid confusion. +\ + +definition div (infixr "div" 70) +where int_div_def: "\a \ Int; b \ Int\ \ + a div b \ CHOOSE q \ Int : \r \ 0 .. (b-1) : a = b * q + r" + +definition mod (infixr "%" 70) +where "a % b \ a - b * (a div b)" + +lemma div_exists: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" + shows "\q \ Int : \r \ 0 .. (b-1) : a = b * q + r" +using a proof induct + case 0 + from b have "0 = b * 0 + 0" "0 \ 0 .. (b-1)" by simp+ + thus ?case by blast +next + case (pos n) + then obtain q r where qr: "q \ Int" "r \ 0 .. (b-1)" "n = b * q + r" + by blast + show ?case + proof (cases "r = b-1") + case True + with \n \ Nat\ b qr + have "succ[n] = b * succ[q] + 0" "0 \ 0 .. (b-1)" + by simp+ + with \q \ Int\ show ?thesis by blast + next + case False + with b qr + have "succ[r] \ 0 .. (b-1)" "succ[n] = b * q + succ[r]" + by (auto simp: less_def) + with \q \ Int\ show ?thesis by blast + qed +next + case (neg n) + then obtain q r where qr: "q \ Int" "r \ 0 .. (b-1)" "-n = b * q + r" + by blast + show ?case + proof (cases "r = 0") + case True + with \n \ Nat\ b qr + have "-succ[n] = b * pred[q] + (b - 1)" "b-1 \ 0 .. (b-1)" + by (simp add: uminusSucc[of n])+ + with \q \ Int\ show ?thesis by blast + next + case False + with b qr \n \ Nat\ + have "pred[r] \ 0 .. (b-1)" "-succ[n] = b * q + pred[r]" + by (auto simp: uminusSucc[of n]) + with \q \ Int\ show ?thesis by blast + qed +qed + +lemma div_mod_unique: + assumes qr: "a = b * q + r" "q \ Int" "r \ 0 .. (b-1)" + and qr': "a = b * q' + r'" "q' \ Int" "r' \ 0 .. (b-1)" + and ab: "a \ Int" "b \ Int" "0 < b" + shows "q' = q" "r' = r" +proof - + { + fix x x' y y' + assume xy: "x \ Int" "x' \ Int" "y \ Nat" "y' \ Nat" + and y': "y' < b" + and eq: "b * x + y = b * x' + y'" + have "x \ x'" + proof (rule contradiction) + assume "\(x \ x')" + with xy obtain k where k: "k \ Nat" "x = succ[x' + k]" + by (auto dest: int_less_imp_succ_add) + with xy ab eq have "b*x' + (b*k + b + y) = b*x' + y'" + by (simp add: add_times_distrib_left add_assoc) + with xy ab k have "b*k + b + y = y'" by simp + with xy ab k y' have "(b*k + y) + b < b" by (simp add: add_ac) + with xy ab k add_less_right_mono[of "(b*k + y) + b" "b" "-b"] + have "b*k + y < 0" by simp + moreover + from xy ab k times_leq_right_pos_mono[of "0" "b" "k"] + have "0 \ b*k + y" by (simp add: less_def) + ultimately show "FALSE" + using xy ab k by simp + qed + } + from this[of q q' r r'] this[of q' q r' r] assms + show 1: "q' = q" by (auto dest: int_leq_antisym) + with add_right_diff_left[of a "b*q" r] add_right_diff_left[of a "b*q'" r'] assms + show "r' = r" by simp +qed + +lemma div_type [iff]: + assumes ab: "a \ Int" "b \ Int" and "0 < b" + shows "a div b \ Int" + using div_exists[OF assms] unfolding int_div_def[OF ab] + by (rule bChooseI2) + +lemma mod_type [iff]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "a % b \ Int" + using assms unfolding mod_def by simp + +lemma mod_range: + assumes ab: "a \ Int" "b \ Int" and "0 < b" + shows "a % b \ 0 .. (b-1)" +proof - + from assms have "a div b \ Int" by simp + from div_exists[OF assms] + have "\r \ 0 .. (b-1) : a = b * (a div b) + r" + unfolding int_div_def[OF ab] by (rule bChooseI2) + then obtain r where "r \ 0 .. (b-1)" "a = b * (a div b) + r" + by blast + with assms show ?thesis + by (simp add: mod_def add_right_diff_left[of "a" "b * (a div b)" r]) +qed + +lemma mod_div_equality: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) + b * (a div b) = a" + using assms unfolding mod_def by simp + +lemma mod_div_equality2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "b * (a div b) + (a % b) = a" + using assms by (simp add: mod_div_equality add_commute) + +lemma mod_div_equality3: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a div b) * b + (a % b) = a" + using assms by (simp add: mod_div_equality2 times_commute) + +lemma mod_div_equality4: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) + (a div b) * b = a" + using assms by (simp add: mod_div_equality times_commute) + +lemma div_equal: + assumes "a = b*q + r" and "a \ Int" and "b \ Int" and "0 < b" + and "q \ Int" and "r \ 0 .. (b-1)" + shows "a div b = q" + using assms mod_div_equality2[of a b] mod_range[of a b] + div_mod_unique[of a b q r "a div b" "a % b"] + by auto + +lemma mod_equal: + assumes "a = b*q + r" and "a \ Int" and "b \ Int" and "0 < b" + and "q \ Int" and "r \ 0 .. (b-1)" + shows "a % b = r" + using assms mod_div_equality2[of a b] mod_range[of a b] + div_mod_unique[of a b q r "a div b" "a % b"] + by auto + + +text \ ``Recursive'' laws about @{text div} and @{text "%"}. \ + +lemma div_mod_nat_less_divisor: + assumes "a \ 0 .. (b-1)" "b \ Int" "0 < b" + shows "a div b = 0" "a % b = a" +proof - + from assms have 1: "a = b*0 + a" by simp + from 1[THEN div_equal] assms show "a div b = 0" + by blast + from 1[THEN mod_equal] assms show "a % b = a" + by blast +qed + +lemma div_is_zero_iff [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a div b = 0) = (a \ 0 .. (b-1))" + using assms mod_div_equality2[of a b] mod_range[of a b] + by (auto simp: div_mod_nat_less_divisor) + +lemma mod_is_self_iff [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b = a) = (a \ 0 .. (b-1))" + using assms mod_div_equality2[of a b] mod_range[of a b] + by (auto simp: div_mod_nat_less_divisor) + +lemma div_mod_minus_less_divisor: + assumes "a \ (-b) .. (-1)" "b \ Int" "0 < b" + shows "a div b = -1" "a % b = b+a" +proof - + from assms have "0 \ b+a" + using add_leq_left_mono[of "-b" "a" "b"] by simp + moreover + from assms have "b+a < b" + using int_leq_pred_iff[of a 0] add_less_left_mono[of a 0 b] + by (simp add: pred0) + ultimately have 1: "b+a \ 0 .. (b-1)" + using assms by simp + from assms have 2: "a = b*(-1) + (b+a)" by (simp add: add_assoc) + from 1 2[THEN div_equal] assms show "a div b = -1" + by blast + from 1 2[THEN mod_equal] assms show "a % b = b+a" + by blast +qed + +lemma div_is_minusone_iff [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a div b = -1) = (a \ (-b) .. (-1))" +proof - + { + assume "a div b = -1" + with assms mod_div_equality2[of a b] + have 1: "a = -b + (a % b)" by simp + from assms mod_range[of a b] have "-b \ -b + (a % b)" + by simp + moreover + from assms mod_range[of a b] have "-b + (a % b) < -b + b" + by (intro add_leq_less_mono) simp+ + ultimately have "a \ (-b) .. (-1)" + using assms 1 by simp + } + moreover + { + assume "a \ (-b) .. (-1)" + with assms div_mod_minus_less_divisor[of a b] have "a div b = -1" + by blast + } + ultimately show ?thesis + using assms by blast +qed + +lemma div_mod_plus_divisor1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a+b) div b = succ[a div b]" "(a+b) % b = a % b" +proof - + from assms have "a = b * (a div b) + (a % b)" + by (simp add: mod_div_equality2) + with assms have "a+b = b * (a div b) + (a % b) + b" + by simp + with assms have 1: "a+b = b * succ[a div b] + (a % b)" + by (simp add: add_ac) + from 1[THEN div_equal] assms mod_range[of a b] + show "(a+b) div b = succ[a div b]" + by blast + from 1[THEN mod_equal] assms mod_range[of a b] + show "(a+b) % b = a % b" + by blast +qed + +lemma div_mod_plus_divisor2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b+a) div b = succ[a div b]" "(b+a) % b = a % b" + using assms by (simp add: add_commute)+ + +lemma div_mod_minus_divisor1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a-b) div b = pred[a div b]" "(a-b) % b = a % b" +proof - + from assms have "a = b * (a div b) + (a % b)" + by (simp add: mod_div_equality2) + with assms have "a-b = b * (a div b) + (a % b) - b" + by simp + with assms have 1: "a-b = b * pred[a div b] + (a % b)" + by (simp add: add_right_commute) + from 1[THEN div_equal] assms mod_range[of a b] + show "(a-b) div b = pred[a div b]" + by blast + from 1[THEN mod_equal] assms mod_range[of a b] + show "(a-b) % b = a % b" + by blast +qed + +lemma div_mod_minus_divisor2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(-b + a) div b = pred[a div b]" "(-b + a) % b = a % b" + using assms by (simp add: add_commute)+ + + +subsection \ Facts about @{text div} and @{text "%"} \ + +lemma div_times_divisor1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + c*b) div b = c + (a div b)" + using c by (induct) (simp add: a b add_assoc)+ + +lemma div_times_divisor2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + b*c) div b = c + (a div b)" + using c by (induct) (simp add: a b add_assoc)+ + +lemma div_times_divisor3 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(b*c + a) div b = c + (a div b)" + using assms by (simp add: add_commute) + +lemma div_times_divisor4 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(c*b + a) div b = c + (a div b)" + using assms by (simp add: add_commute) + +lemma mod_times_divisor1 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + c*b) % b = a % b" + using c by (induct) (simp add: a b add_assoc)+ + +lemma mod_times_divisor2 [simp]: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "(a + b*c) % b = a % b" + using c by (induct) (simp add: a b add_assoc)+ + +lemma mod_times_divisor3 [simp]: + assumes "a \ Int" and "b \ Int" "0 < b" and "c \ Int" + shows "(b*c + a) % b = a % b" + using assms by (simp add: add_commute) + +lemma mod_times_divisor4 [simp]: + assumes "a \ Int" and "b \ Int" "0 < b" and "c \ Int" + shows "(c*b + a) % b = a % b" + using assms by (simp add: add_commute) + +lemma times_div_self1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a*b) div b = a" + using assms div_times_divisor1[of 0 b a] by simp + +lemma times_div_self2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b*a) div b = a" + using assms by (simp add: times_commute) + +lemma times_mod_self1 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(a*b) % b = 0" + using assms mod_times_divisor1[of 0 b a] by simp + +lemma times_mod_self2 [simp]: + assumes "a \ Int" "b \ Int" "0 < b" + shows "(b*a) % b = 0" + using assms by (simp add: times_commute) + +lemma div_mod_by_one [simp]: + "a \ Int \ a div 1 = a" + "a \ Int \ a % 1 = 0" + using times_div_self1[of a 1] times_mod_self1[of a 1] by simp+ + +lemma div_mod_by_self [simp]: + "\a \ Int; 0 < a\ \ a div a = 1" + "\a \ Int; 0 < a\ \ a % a = 0" + using times_div_self1[of 1 a] times_mod_self1[of 1 a] by simp+ + +lemma mod_div_same [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) div b = 0" + using assms mod_range[of a b] by simp + +lemma mod_mod_same [simp]: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(a % b) % b = a % b" + using assms mod_times_divisor1[of "a % b" "b" "a div b"] mod_range[of a b] + by simp + +lemma div_mod_decomposition: + assumes a: "a \ Int" and b: "b \ Int" "0 < b" + obtains q r where "q \ Int" and "r \ 0 .. (b-1)" + and "q = a div b" and "r = a % b" and "a = b*q + r" + using assms mod_div_equality2[of a b, THEN sym] mod_range[of a b] + by blast + + +subsection \Further arithmetic laws for division and modulus\ + +lemma add_div: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a + b) div c = a div c + b div c + (a%c + b%c) div c" +proof - + from assms have "a + b = (c * (a div c) + a%c) + (c * (b div c) + b%c)" + by (simp add: mod_div_equality2) + with assms have "a + b = c * (a div c + b div c) + (a%c + b%c)" + by (simp add: add_times_distrib_left add_commute add_left_commute) + with assms show ?thesis by simp +qed + +lemma add_mod: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a + b) % c = ((a % c) + (b % c)) % c" +proof - + from assms have "(a%c + b%c) % c = (c * (a div c + b div c) + (a%c + b%c)) % c" + by simp + moreover + from assms have "c * (a div c + b div c) + (a%c + b%c) + = (c * (a div c) + a%c) + (c * (b div c) + b%c)" + by (simp add: add_times_distrib_left add_commute add_left_commute) + with assms have "c * (a div c + b div c) + (a%c + b%c) = a+b" + by (simp add: mod_div_equality2) + ultimately show ?thesis by simp +qed + +lemma times_div: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * b) div c = a * (b div c) + (a div c) * (b%c) + ((a%c) * (b%c)) div c" +proof - + from assms have "a * b = (c * (a div c) + a%c) * (c * (b div c) + b%c)" + by (simp add: mod_div_equality2) + also from assms + have "\ = (c * (a div c) + a%c) * (c * (b div c)) + (c * (a div c) + a%c) * (b%c)" + by (simp add: add_times_distrib_left) + also from assms + have "\ = (c * (a div c) * (c * (b div c)) + + (a%c) * (c * (b div c))) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: add_times_distrib_right) + also from assms + have "\ = (c * c * (a div c) * (b div c) + + (a%c) * c * (b div c)) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: times_assoc times_right_commute) + also from assms + have "\ = (c * c * (a div c) + (a%c) * c) * (b div c) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: add_times_distrib_right) + also from assms + have "\ = (c * c * (a div c) + c * (a%c)) * (b div c) + + (c * (a div c) * (b%c) + (a%c) * (b%c))" + by (simp add: times_commute) + also from assms + have "\ = (c * (c * (a div c)) + c * (a%c)) * (b div c) + + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + by (simp add: times_assoc) + also from assms + have "\ = (c * (c * (a div c) + (a%c))) * (b div c) + + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + by (simp add: add_times_distrib_left) + finally have "a * b = c * (a * (b div c)) + (c * ((a div c) * (b%c)) + (a%c) * (b%c))" + using assms mod_div_equality2[of a c] + by (simp add: times_assoc) + with assms have "(a * b) div c = a * (b div c) + (c * ((a div c) * (b%c)) + (a%c) * (b%c)) div c" + by simp + moreover + from assms + have "(c * ((a div c) * (b%c)) + (a%c) * (b%c)) div c = (a div c) * (b%c) + ((a%c) * (b%c)) div c" + by simp + ultimately + have "(a * b) div c = a * (b div c) + ((a div c) * (b%c) + ((a%c) * (b%c)) div c)" + by simp + with assms show ?thesis by (simp add: add_assoc) +qed + +lemma times_mod: + assumes "a \ Int" and "b \ Int" and "c \ Int" and "0 < c" + shows "(a * b) % c = ((a % c) * (b % c)) % c" +proof - + from assms + have "((a % c) * (b % c)) % c + = (c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c)) % c" + by simp + moreover + from assms + have "c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c) + = (c * (a div c) + a % c) * (c * (b div c) + b % c)" + by (simp add: add_times_distrib_left add_times_distrib_right + times_commute times_left_commute + add_commute add_left_commute) + with assms + have "c * ((a div c) * c * (b div c) + (a div c) * (b % c) + (a % c) * (b div c)) + + (a % c) * (b % c) + = a * b" + by (simp add: mod_div_equality2) + ultimately show ?thesis by simp +qed + +lemma uminus_div1: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) div b = -(a div b) + (-(a%b)) div b" +proof - + from assms have "-a = -(b * (a div b) + a%b)" + by (simp add: mod_div_equality2) + with assms have "-a = b * (-(a div b)) - (a % b)" + by simp + with assms show ?thesis + by (simp del: times_uminus_right) +qed + +lemma uminus_div2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) div b = -(a div b) + (IF a%b = 0 THEN 0 ELSE -1)" +proof (cases "a % b = 0") + case True + with assms show ?thesis by (simp add: uminus_div1) +next + case False + with assms mod_range[of a b] have "-(a % b) \ (-b) .. (-1)" + by auto + with assms have "(-(a % b)) div b = -1" + by (simp add: div_mod_minus_less_divisor) + with assms False uminus_div1[of a b] + show ?thesis by simp +qed + +lemma uminus_mod1: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) % b = (-(a % b)) % b" +proof - + from assms have "(-a) % b = (-(b * (a div b) + a % b)) % b" + by (simp add: mod_div_equality2) + moreover + from assms have "-(b * (a div b) + a % b) = b * (-(a div b)) - a % b" + by simp + ultimately show ?thesis + using assms mod_times_divisor3[of "-(a % b)" "b" "-(a div b)"] + by simp +qed + +lemma uminus_mod2: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(-a) % b = (IF a % b = 0 THEN 0 ELSE b - (a%b))" +proof (cases "a % b = 0") + case True + with assms show ?thesis by (simp add: uminus_mod1) +next + case False + with assms mod_range[of a b] have "-(a % b) \ (-b) .. (-1)" + by auto + with assms have "(-(a % b)) % b = b - (a % b)" + by (simp add: div_mod_minus_less_divisor) + with assms uminus_mod1[of a b] False + show ?thesis by simp +qed + + +subsection \Divisibility and @{text div} / @{text mod}\ + +lemma dvd_is_mod0: + assumes "a \ Int" and "b \ Int" and "0 < b" + shows "(b dvd a) = (a % b = 0)" + using assms mod_div_equality2[of a b] by auto + +lemma dvd_times_div_self: + assumes "b dvd a" and "a \ Int" and "b \ Int" and "0 < b" + shows "b * (a div b) = a" + using assms dvd_is_mod0[of a b] mod_div_equality2[of a b] by simp + +lemma dvd_times_div_distrib: + assumes dvd: "b dvd a" and a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "c * (a div b) = (c*a) div b" +proof - + from dvd a b obtain d where d: "d \ Int" "a = b*d" by auto + with a b c have "c*a = b * (c*d)" by (simp add: times_left_commute) + with a b c d show ?thesis by simp +qed + +lemma dvd_mod_imp_dvd: + assumes "a dvd (c % b)" and "a dvd b" + and a: "a \ Int" and b: "b \ Int" "0 < b" and c: "c \ Int" + shows "a dvd c" +proof - + from assms have "a dvd ((c div b) * b + c % b)" + by (simp add: dvd_mult) + with a b c show ?thesis + by (simp add: mod_div_equality3) +qed + +end diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index fdcc141f..bfbefc7f 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -1,1314 +1,1537 @@ -(* Title: Integers.thy - Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation +(* Title: TLA+/Integers.thy + Author: Stephan Merz, LORIA + Copyright (C) 2008-2021 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:40:02 merz> - + Version: Isabelle2021-1 *) -section \ The Integers as a superset of natural numbers \ +section \ The set of integer numbers \ theory Integers -imports Tuples NatArith +imports FixedPoints Functions begin -subsection \ The minus sign \ - -consts - "minus" :: "c \ c" ("-._" [75] 75) - -syntax (* -- syntax for negative naturals *) - "-.0" :: "c" ("-.0") - "-.1" :: "c" ("-.1") - "-.2" :: "c" ("-.2") -translations - "-.0" \ "-.(0)" - "-.1" \ "-.(1)" - "-.2" \ "-.(2)" - -(* lemma eq_imp_negeq: "n = m \ -.n = -.m" by simp *) - -axiomatization where - neg0 [simp]: "-.0 = 0" -and - neg_neg [simp]: "-.-.n = n" -and - negNotInNat [simp]: "-.(Succ[n]) \ Nat" - -lemma negNat_noteq_Nat [simp]: - "\m \ Nat; n \ Nat\ \ (-. Succ[m] = Succ[n]) = FALSE" -proof (rule contradiction) - assume "(-. Succ[m] = Succ[n]) \ FALSE" - and m: "m \ Nat" and n: "n \ Nat" - hence "-. Succ[m] = Succ[n]" by auto - hence "-. Succ[m] \ Nat" using n by auto - with negNotInNat[of m] show FALSE by simp -qed +subsection \Extended Peano axioms\ -lemma negNat_noteq_Nat2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m] = -. Succ[n]) = FALSE" -proof auto - assume "Succ[m] = -. Succ[n]" - hence "-. Succ[n] = Succ[m]" by simp - with m n show "FALSE" by simp -qed +text \ + We extend the standard Peano axioms by a complement operation + for constructing negative integers, and we prove the existence + of a structure satisfying these axioms. +\ -lemma nat_not_eq_inv: "n \ Nat \ n = 0 \ -.n \ n" - using not0_implies_Suc[of n] by auto +definition ExtendedPeano :: "[c,c,c,c,c] \ c" where + \ \Parameters: the set of integer and natural numbers, zero, + and successor and complement functions.\ + "ExtendedPeano(I,N,Z,sc,cp) \ + Z \ N + \ sc \ [N \ N] + \ (\n \ N : sc[n] \ Z) + \ (\m,n \ N : sc[m] = sc[n] \ m = n) + \ (\S \ SUBSET N : Z \ S \ (\n\S : sc[n] \ S) \ N \ S) + \ cp \ [I \ I] + \ I = N \ {cp[n] : n \ N} + \ cp[Z] = Z + \ (\k \ N : cp[k] \ N \ k = Z) + \ (\k,l \ I : cp[k] = cp[l] \ k = l) + \ (\k \ I : cp[cp[k]] = k)" -lemma minusInj [dest]: - assumes hyp: "-.n = -.m" - shows "n = m" +text \ + We now prove the existence of a structure satisfying the + extended Peano axioms. For @{text N}, we take the standard + ZF construction where @{text "{}"} is zero, where + @{text "i \ {i}"} is the successor of any natural number + @{text i}, and where the set @{text N} is defined as + the least set that contains zero and is closed under successor + (this is a subset of the infinity set asserted to exist in ZF + set theory). We then pick a value @{text e} that does not + occur in the construction of @{text N} and define the + complement function by adding or removing that value to / from + the set representing the argument. Later, integers are defined + using a sequence of @{text CHOOSE}'s, so there is no commitment + to that particular structure. +\ + +theorem extendedPeanoExists: "\I,N,Z,sc,cp : ExtendedPeano(I,N,Z,sc,cp)" proof - - from hyp have "-.-.n = -.-.m" by simp - thus ?thesis by simp -qed + let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ + define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" + define N where "N \ lfp(infinity, expand)" + define Z where "Z \ {}" + define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ + have mono: "Monotonic(infinity, expand)" + using infinity by (auto simp: Monotonic_def expand_def) + hence expandN: "expand(N) \ N" + by (unfold N_def, rule lfpPreFP) + from expandN have 1: "Z \ N" + by (auto simp: expand_def Z_def) + have 2: "Sc \ [N \ N]" + proof (unfold Sc_def, rule functionInFuncSet) + show "\n \ N : ?sc(n) \ N" using expandN by (auto simp: expand_def) + qed + have 3: "\m\N : Sc[m] \ Z" + unfolding Z_def Sc_def by auto + have 4: "\m,n \ N : Sc[m] = Sc[n] \ m = n" + proof (clarify) + fix m n + assume "m \ N" and "n \ N" and "Sc[m] = Sc[n]" + hence eq: "?sc(m) = ?sc(n)" by (simp add: Sc_def) + show "m = n" + proof (rule setEqual) + show "m \ n" + proof (rule subsetI) + fix x + assume x: "x \ m" show "x \ n" + proof (rule contradiction) + assume "x \ n" + with x eq have "n \ m" by auto + moreover + from eq have "m \ ?sc(n)" by auto + ultimately + show "FALSE" by (blast elim: inAsym) + qed + qed + next + show "n \ m" + proof (rule subsetI) + fix x + assume x: "x \ n" show "x \ m" + proof (rule contradiction) + assume "x \ m" + with x eq have "m \ n" by auto + moreover + from eq have "n \ ?sc(m)" by auto + ultimately + show "FALSE" by (blast elim: inAsym) + qed + qed + qed + qed + have 5: "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" + proof (clarify del: subsetI) + fix S + assume sub: "S \ N" and Z: "Z \ S" and Sc: "\n\S : Sc[n] \ S" + show "N \ S" + proof (unfold N_def, rule lfpLB) + show "expand(S) \ S" + proof (auto simp: expand_def) + from Z show "{} \ S" by (simp add: Z_def) + next + fix n + assume n: "n \ S" + with Sc have "Sc[n] \ S" .. + moreover + from n sub have "n \ N" by auto + hence "Sc[n] = ?sc(n)" by (simp add: Sc_def) + ultimately show "?sc(n) \ S" by simp + qed + next + have "N \ infinity" + by (unfold N_def, rule lfpSubsetDomain) + with sub show "S \ infinity" by auto + qed + qed + + define e where "e \ CHOOSE e : e \ UNION {n : n \ N}" + define cp where "cp \ \n. IF n = Z THEN n ELSE IF e \ n THEN n \ {e} ELSE n \ {e}" + define I where "I \ N \ {cp(n) : n \ N}" + let ?cp = "[k \ I \ cp(k)]" + + have "\e : e \ UNION {n : n \ N}" by (blast intro: inIrrefl) + hence "e \ UNION {n : n \ N}" unfolding e_def by (rule chooseI_ex) + hence e: "\n \ N : e \ n" by blast + + have cpZ: "cp(Z) = Z" by (simp add: cp_def) + + have diff: "\k \ I \ {Z} : cp(k) \ k" + proof (clarify) + fix k + assume "k \ I" "k \ Z" "cp(k) = k" + show "FALSE" + proof (cases "k \ N") + case True + with e \k \ Z\ have "e \ k \ cp(k) = k \ {e}" by (simp add: cp_def) + with \cp(k) = k\ show ?thesis by blast + next + case False + with \k \ I\ obtain n where "n \ N" "k = cp(n)" by (auto simp: I_def) + with \k \ Z\ have "n \ Z" by (auto simp: cp_def) + with e \n \ N\ \k = cp(n)\ \k \ Z\ + have "k = n \ {e}" "cp(k) = k \ {e}" by (auto simp: cp_def) + with \cp(k) = k\ show ?thesis by auto + qed + qed + + from e have cpN: "\n \ N : cp(cp(n)) = n" + by (auto simp: Z_def cp_def) + + from e have cpNN: "\k \ N : cp(k) \ N \ k = Z" + by (auto simp: cp_def) + + { + fix k + assume k: "k \ I" "cp(k) = Z" + have "k = Z" + proof (cases "k \ N") + case True + with k 1 cpNN show ?thesis by blast + next + case False + with \k \ I\ obtain n where "n \ N" "k = cp(n)" + by (auto simp: I_def) + with \cp(k) = Z\ cpN have "n = Z" by simp + with \k = cp(n)\ show ?thesis by (simp add: cp_def) + qed + } + with cpZ have cpisZ: "\k \ I: cp(k) = Z \ k = Z" by blast + + from cpN have cpI: "\k \ I : cp(k) \ I" by (auto simp: I_def) + hence 6: "?cp \ [I \ I]" by (rule functionInFuncSet) + + have 7: "I = N \ {?cp[n] : n \ N}" by (force simp: I_def) + + from cpZ 1 have 8: "?cp[Z] = Z" by (simp add: I_def) + + from cpNN have 9: "\n \ N : ?cp[n] \ N \ n = Z" by (simp add: I_def) + + { + fix k l + assume dom: "k \ I" "l \ I" and eq: "cp(k) = cp(l)" and kl: "k \ l" + have "FALSE" + proof (cases "k \ N") + case True + have "k \ Z" + proof + assume "k = Z" + with eq cpZ dom kl cpisZ show "FALSE" by auto + qed + with \k \ N\ e have cpk: "e \ k" "cp(k) = k \ {e}" by (auto simp: cp_def) + have "l \ N" + proof (rule contradiction) + assume "l \ N" + with \l \ I\ obtain n where "n \ N" "l = cp(n)" + by (auto simp: I_def) + with cpk cpN eq have "n = k \ {e}" by simp + with \n \ N\ e show "FALSE" by blast + qed + have "l \ Z" + proof + assume "l = Z" + with eq cpZ dom kl cpisZ show "FALSE" by auto + qed + with \l \ N\ e have "e \ l" "cp(l) = l \ {e}" by (auto simp: cp_def) + with cpk eq have "k \ l" "l \ k" by auto + with kl show ?thesis by (auto dest: setEqual) + next + case False + with dom obtain nk where nk: "nk \ N" "k = cp(nk)" + by (auto simp: I_def) + with cpN eq have cpl: "cp(l) \ N" by simp + have "l \ N" + proof + assume "l \ N" + with cpl cpNN have "l = Z" by simp + with eq dom cpZ cpisZ False 1 show "FALSE" by simp + qed + with dom obtain nl where "nl \ N" "l = cp(nl)" + by (auto simp: I_def) + with nk eq kl cpN show ?thesis by simp + qed + } + hence 10: "\k,l \ I : ?cp[k] = ?cp[l] \ k = l" by auto -lemma minusInj_iff [simp]: - "-.x = -.y = (x = y)" -by auto + from cpI cpN have 11: "\k \ I : ?cp[?cp[k]] = k" + by (auto simp: I_def) -lemma neg0_imp_0 [simp]: "-.n = 0 = (n = 0)" -proof auto - assume "-.n = 0" - hence "-.-.n = 0" by simp - thus "n = 0" by simp + from 1 2 3 4 5 6 7 8 9 10 11 have "ExtendedPeano(I,N,Z,Sc,?cp)" + unfolding ExtendedPeano_def by blast + thus ?thesis by blast qed -lemma neg0_eq_0 [dest]: "-.n = 0 \ (n = 0)" -by simp +subsection \The structure of integer numbers\ -lemma notneg0_imp_not0 [dest]: "-.n \ 0 \ n \ 0" -by auto +text \ + The integer numbers are now defined as some structure + satisfying the extended Peano axioms. +\ -lemma not0_imp_notNat [simp]: "n \ Nat \ n \ 0 \ -.n \ Nat" - using not0_implies_Suc[of n] by auto +definition Succ :: "c" where + "Succ \ CHOOSE sc : \cp,I,N,Z : ExtendedPeano(I,N,Z,sc,cp)" -lemma negSuccNotZero [simp]: "n \ Nat \ (-. Succ[n] = 0) = FALSE" -by auto +definition Nat :: "c" where + "Nat \ DOMAIN Succ" -lemma negSuccNotZero2 [simp]: "n \ Nat \ (0 = -. Succ[n]) = FALSE" -proof auto - assume n: "n \ Nat" and 1: "0 = -. Succ[n]" - from 1 have "-. Succ[n] = 0" by simp - with n show FALSE by simp -qed +definition zero :: "c" ("0") where + "0 \ CHOOSE Z : \cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" -lemma negInNat_imp_false [dest]: "-.Succ[n] \ Nat \ FALSE" - using negNotInNat[of n] by simp +definition intCplt :: "c" where + "intCplt \ CHOOSE cp : \I : ExtendedPeano(I,Nat,0,Succ,cp)" -lemma negInNatFalse [simp]: "-.Succ[n] \ Nat = FALSE" - using negNotInNat[of n] by auto +definition Int :: "c" where + "Int \ DOMAIN intCplt" -lemma n_negn_inNat_is0 [simp]: - assumes "n \ Nat" - shows "-.n \ Nat = (n = 0)" -using assms by (cases "n", auto) +definition uminus :: "c \ c" ("-_" [80] 80) where + "-z \ intCplt[z]" -lemma minus_sym: "-.a = b = (a = -.b)" -by auto +lemmas + setEqualI [where A = "Nat", intro!] + setEqualI [where B = "Nat", intro!] + setEqualI [where A = "Int", intro!] + setEqualI [where B = "Int", intro!] -lemma negNat_exists: "-.n \ Nat \ \k \ Nat: n = -.k" -by force +lemma intExtendedPeano: "ExtendedPeano(Int,Nat,0,Succ,intCplt)" +proof - + have "\cp,I,N,Z : ExtendedPeano(I,N,Z,Succ,cp)" + proof (unfold Succ_def, rule chooseI_ex) + from extendedPeanoExists + show "\Sc,cp,I,N,Z : ExtendedPeano(I,N,Z,Sc,cp)" by blast + qed + then obtain N Z where PNZ: "\cp,I : ExtendedPeano(I,N,Z,Succ,cp)" by blast + hence "Succ \ [N \ N]" + by (simp add: ExtendedPeano_def) + hence "N = Nat" + by (simp add: Nat_def) + with PNZ have "\cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" by simp + hence "\cp,I : ExtendedPeano(I,Nat,0,Succ,cp)" + unfolding zero_def by (rule chooseI) + hence "\I : ExtendedPeano(I,Nat,0,Succ,intCplt)" + unfolding intCplt_def by (rule chooseI_ex) + then obtain I where IEP: "ExtendedPeano(I,Nat,0,Succ,intCplt)" + by blast + hence "intCplt \ [I \ I]" + by (auto simp: ExtendedPeano_def) + hence "I = Int" + by (simp add: Int_def) + with IEP show ?thesis by simp +qed -lemma nat_eq_negnat_is_0 [simp]: - assumes "n \ Nat" - shows "(n = -.n) = (n = 0)" -using assms by (cases "n", auto) +lemma natIsInt [simp,intro]: "n \ Nat \ n \ Int" + using intExtendedPeano by (auto simp: ExtendedPeano_def) -(* used for simplification in additions *) -lemma (*[simp]*) "\x \ Nat : -.1 = -.x" by auto -lemma (*[simp]*) "x \ Nat \ (1 = -.x) = FALSE" by (auto simp: sym[OF minus_sym]) +lemma uminusIsInt [simp,intro!]: "n \ Int \ -n \ Int" + using intExtendedPeano by (auto simp: ExtendedPeano_def uminus_def) +text \ + Fundamental lemmas about zero, successor, and complement. +\ -subsection \ The set of Integers \ +lemma zeroIsNat [intro!,simp]: "0 \ Nat" + using intExtendedPeano by (simp add: ExtendedPeano_def) -definition Int -where "Int \ Nat \ {-.n : n \ Nat}" +(* redundant, but will become relevant later *) +lemma zeroIsInt [intro!,simp]: "0 \ Int" + by simp -lemma natInInt [simp]: "n \ Nat \ n \ Int" -by (simp add: Int_def) +lemma SuccInNatNat [intro!,simp]: "Succ \ [Nat \ Nat]" + using intExtendedPeano by (simp add: ExtendedPeano_def) -lemma intDisj: "n \ Int \ n \ Nat \ n \ {-.n : n \ Nat}" -by (auto simp: Int_def) +lemma SuccIsAFcn [intro!,simp]: "isAFcn(Succ)" + using SuccInNatNat by blast -lemma negint_eq_int [simp]: "-.n \ Int = (n \ Int)" -unfolding Int_def by force +\ \@{text "DOMAIN Succ = Nat"}\ +lemmas domainSucc [intro!,simp] = funcSetDomain[OF SuccInNatNat] +\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ +lemmas SuccIsNat [intro!,simp] = funcSetValue[OF SuccInNatNat] -lemma intCases [case_names Positive Negative, cases set: Int]: - assumes n: "n \ Int" - and sc: "n \ Nat \ P" - and nsc: "\m. \m \ Nat; n = -.m\ \ P" - shows "P" -using assms unfolding Int_def by auto - -(* -- Integer cases over two parameters *) -lemma intCases2: - assumes m: "m \ Int" and n: "n \ Int" - and pp: "\m n. \m \ Nat; n \ Nat\ \ P(m,n)" - and pn: "\m n. \m \ Nat; n \ Nat\ \ P(m,-.n)" - and np: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,n)" - and nn: "\m n. \m \ Nat; n \ Nat\ \ P(-.m,-.n)" - shows "P(m,n)" -using m proof (cases "m") - assume "m \ Nat" - from n this pp pn show "P(m,n)" by (cases "n", auto) -next - fix m' - assume "m' \ Nat" "m = -. m'" - from n this np nn show "P(m,n)" by (cases "n", auto) +lemma [simp]: + assumes "n \ Nat" + shows "(Succ[n] = 0) = FALSE" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def) + +lemma [simp]: + assumes n: "n \ Nat" + shows "(0 = Succ[n]) = FALSE" + using assms by (auto dest: sym) + +lemma SuccNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): + "\Succ[n] = 0; n \ Nat\ \ P" + "\0 = Succ[n]; n \ Nat\ \ P" +by (simp+) + +lemma SuccInj [dest]: + assumes "Succ[m] = Succ[n]" and "m \ Nat" and "n \ Nat" + shows "m=n" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def) + +lemma SuccInjIff [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(Succ[m] = Succ[n]) = (m = n)" +using assms by auto + +\ \The primitive induction rule for natural numbers, will be superseded below.\ +lemma pr_natInduct: + assumes z: "P(0)" + and sc: "\n. \n \ Nat; P(n)\ \ P(Succ[n])" + shows "\n\Nat : P(n)" +proof - + let ?P = "{n \ Nat : P(n)}" + have "?P \ SUBSET Nat" by auto + moreover + from z have "0 \ ?P" by simp + moreover + from sc have "\n \ ?P : Succ[n] \ ?P" by simp + ultimately have "Nat \ ?P" + using intExtendedPeano by (simp add: ExtendedPeano_def) + thus ?thesis by auto qed -lemma intCases3: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - and ppp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,p)" - and ppn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,n,-.p)" - and pnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,p)" - and pnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(m,-.n,-.p)" - and npp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,p)" - and npn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,n,-.p)" - and nnp: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,p)" - and nnn: "\m n p. \m \ Nat; n \ Nat; p \ Nat\ \ P(-.m,-.n,-.p)" - shows "P(m,n,p)" -proof (rule intCases2[OF m n]) - fix m n - assume "m \ Nat" and "n \ Nat" - from p this ppp ppn show "P(m,n,p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this pnp pnn show "P(m, -.n, p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this npp npn show "P(-.m, n, p)" by (cases "p", auto) -next - fix m n - assume "m \ Nat" and "n \ Nat" - from p this nnp nnn show "P(-.m, -.n, p)" by (cases "p", auto) +lemma pr_natCases: + assumes n: "n \ Nat" + and 0: "n = 0 \ P" + and suc: "\m. \m \ Nat; n = Succ[m]\ \ P" + shows "P" +proof - + have "\n \ Nat : n = 0 \ (\m \ Nat : n = Succ[m])" + by (rule pr_natInduct) auto + with assms show ?thesis by blast qed -lemma int_eq_negint_is_0 [simp]: "n \ Int \ n = -.n = (n = 0)" -by(rule intCases, auto) - -lemma intNotNatIsNeg: "\n \ Nat; n \ Int\ \ \k \ Nat: n = -.k" -unfolding Int_def by auto - -lemma intNotNatIsNegNat: "\n \ Nat; n \ Int\ \ -.n \ Nat" -unfolding Int_def by auto - - -subsection \ Predicates ''is positive'' and 'is negative' \ - -definition isPos (* -- Predicate ''is positive'' *) -where "isPos(n) \ \k \ Nat: n = Succ[k]" - -definition isNeg (* -- Predicate ''is negative'' *) -where "isNeg(n) \ \k \ Nat: n = -.Succ[k]" - -lemma boolify_isPos [simp]: "boolify(isPos(n)) = (isPos(n))" -by (simp add: isPos_def) - -lemma isPos_isBool [intro!,simp]: "isBool(isPos(n))" -by (simp add: isPos_def) +lemma [simp]: "-0 = 0" + using intExtendedPeano by (simp add: ExtendedPeano_def uminus_def) -lemma boolify_isNeg [simp]: "boolify(isNeg(n)) = (isNeg(n))" -by (simp add: isNeg_def) - -lemma isNeg_isBool [intro!,simp]: "isBool(isNeg(n))" -by (simp add: isNeg_def) - -lemma zeroNotPos [dest]: "isPos(0) \ FALSE" by (auto simp: isPos_def) -lemma zeroNotNeg [dest]: "isNeg(0) \ FALSE" by (auto simp: isNeg_def) - -lemma natIsPos [simp]: "n \ Nat \ isPos(Succ[n])" by(simp add: isPos_def) -lemma negIsNeg [simp]: "n \ Nat \ isNeg(-.Succ[n])" by(simp add: isNeg_def) - -lemma negIsNotPos [simp]: "n \ Nat \ isPos(-.Succ[n]) = FALSE" -by(simp add: isPos_def) - -(*lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \\ {0})" - unfolding isPos_def using not0_implies_Suc[of n] by auto*) - -lemma isPos_eq_inNat1: "isPos(n) = (n \ Nat \ n \ 0)" -unfolding isPos_def using not0_implies_Suc[of n] by auto - -(*lemma isNeg_eq_inNegNat: - "isNeg(n) = (n \ {-.n : n \ Nat } \\ {0})" - unfolding isNeg_def by force*) - -lemma isNeg_eq_inNegNat: - "isNeg(n) = (n \ {-.n : n \ Nat } \ n \ 0)" -unfolding isNeg_def by force - -lemma intIsPos_isNat: "n \ Int \ isPos(n) \ n \ Nat" -by(auto simp: isPos_def) - -lemma negNotNat_isNat: - assumes n: "n \ Int" shows "(-.n \ Nat) = FALSE \ n \ Nat" -using n by (cases, auto) - -lemma noNatisNeg [simp]: - "n \ Nat \ isNeg(n) = FALSE" (* -- No natural number is negative *) -unfolding isNeg_def using negNotInNat by blast - -lemma negNat_isNeg [intro]: "\m \ Nat; m \ 0\ \ isNeg(-.m)" -unfolding isNeg_eq_inNegNat by auto - -lemma nat_is_0_or_pos: "(n = 0 \ isPos(n)) = (n \ Nat)" -unfolding isPos_def by force - -lemma isNeg_dichotomy (*[simp]*): "n \ Int \ isNeg(-.n) \ isNeg(n) = FALSE" -unfolding isNeg_def by auto +lemma uminusNat [simp]: + assumes "n \ Nat" + shows "(-n \ Nat) = (n = 0)" +proof - + { + assume "-n \ Nat" + with assms intExtendedPeano have "n = 0" + by (auto simp: ExtendedPeano_def uminus_def) + } + thus ?thesis by auto +qed -lemma isPos_isNeg_false [simp]: "n \ Int \ isPos(n) \ isNeg(n) = FALSE" - unfolding isPos_def by force +lemma uminusInj [dest]: + assumes "-a = -b" and "a \ Int" and "b \ Int" + shows "a = b" + using assms intExtendedPeano + unfolding ExtendedPeano_def uminus_def by blast -(*lemma notPosIsNegPos: - "n \ Int \\ {0} \ \ isPos(n) \ isPos(-.n)" -unfolding isPos_eq_inNat1 using intNotNatIsNegNat by auto*) +lemma uminusInjIff [simp]: + "\a \ Int; b \ Int\ \ (-a = -b) = (a = b)" + by auto -lemma isPos_neg_isNeg [simp]: - assumes n: "n \ Int" shows "isPos(-.n) = isNeg(n)" -by (auto simp: minus_sym isPos_def isNeg_def) +lemma uminusUminus [simp]: + "a \ Int \ --a = a" + using intExtendedPeano by (simp add: ExtendedPeano_def uminus_def) -lemma notIsNeg0_isPos: - assumes n: "n \ Int" - shows "\\ isNeg(n); n \ 0\ \ isPos(n)" -using n by (cases, auto simp: isPos_eq_inNat1 dest: negNat_isNeg) - -lemma notIsPos_notNat [simp]: "\\ isPos(n); n \ 0\ \ n \ Nat = FALSE" -by (auto simp: isPos_eq_inNat1) +lemma uminusSwap: + assumes "a \ Int" and "b \ Int" + shows "(-a = b) = (a = -b)" +proof - + from assms have "(-a = b) = (--a = -b)" by auto + with assms show ?thesis by simp +qed -lemma intThenPosZeroNeg: +lemma intElim: assumes n: "n \ Int" - shows "isNeg(n) \ n = 0 \ isPos(n)" -by (auto elim: notIsNeg0_isPos[OF n]) - - -subsection \ Signum function and absolute value \ + and pos: "n \ Nat \ P" and neg: "\m. \m \ Nat; n = -m\ \ P" + shows "P" + using assms intExtendedPeano by (auto simp: ExtendedPeano_def uminus_def) -definition sgn (* -- signum function *) -where "sgn(n) \ IF n = 0 THEN 0 ELSE (IF isPos(n) THEN 1 ELSE -.1)" +lemma uminusZero [simp]: + assumes "a \ Int" + shows "(-a = 0) = (a = 0)" + using assms by (simp add: uminusSwap) -definition abs (* -- absolute value *) -where "abs(n) \ IF sgn(n) = -.1 THEN -.n ELSE n" +lemma uminusZero' [simp]: + assumes "a \ Int" + shows "(0 = -a) = (a = 0)" +using assms by (auto dest: sym) -lemma sgnInInt [simp]: "n \ Int \ sgn(n) \ Int" -by (auto simp: sgn_def) +lemma uminusReflZero: + assumes "a \ Int" "-a = a" + shows "a = 0" +proof - + { + fix k + assume "k \ Nat" "-k = k" + hence "-k \ Nat" by simp + with \k \ Nat\ have "k = 0" by simp + } + from assms this show ?thesis + by (elim intElim) (force+) +qed -lemma sgn0 [simp]: "sgn(0) = 0" -by (simp add: sgn_def) +lemma [simp]: + "a \ Int \ (-a = a) = (a = 0)" + "a \ Int \ (a = -a) = (a = 0)" + by (auto dest: uminusReflZero) -lemma sgnPos [simp]: "n \ Nat \ sgn(Succ[n]) = 1" -by (simp add: sgn_def) +lemma pr_intCases: + assumes a: "a \ Int" + and 0: "a = 0 \ P" + and pos: "\n. \ n \ Nat; a = Succ[n] \ \ P" + and neg: "\n. \ n \ Nat; a = -Succ[n] \ \ P" + shows "P" +using a proof (rule intElim) + assume "a \ Nat" + with 0 pos show ?thesis by (auto elim: pr_natCases) +next + fix m + assume "m \ Nat" "a = -m" + with 0 neg show ?thesis by (auto elim: pr_natCases) +qed -lemma sgnNeg [simp]: "n \ Nat \ sgn(-.Succ[n]) = -.1" -by (simp add: sgn_def) -lemma sgn0_imp_0: "sgn(n) = 0 \ n = 0" -by (auto simp: sgn_def) +subsection \Successor and predecessor on integer numbers.\ -lemma sgn0_iff_0 [simp]: "(sgn(n) = 0) = (n = 0)" -by (auto simp: sgn_def) +text \ + We extend the successor function for the set of integer numbers + and define the predecessor function as its inverse. These + functions are denoted @{text "succ"} and @{text "pred"} and will + replace the primitive function @{text "Succ"} for reasoning + about integers. +\ -lemma sgn1_imp_pos (*[simp]*): "sgn(n) = 1 \ n \ Nat \ n \ 0" -unfolding sgn_def isPos_eq_inNat1 by auto +definition succ where + "succ \ [a \ Int \ IF a \ Nat THEN Succ[a] + ELSE -(CHOOSE m \ Nat : a = -Succ[m])]" + +abbreviation "one \ succ[0]" +notation "one" ("1") +(* +abbreviation "two \ succ[1]" +notation "two" ("2") +abbreviation "three \ succ[2]" +notation "three" ("3") +abbreviation "four \ succ[3]" +notation "four" ("4") +abbreviation "five \ succ[4]" +notation "five" ("5") +abbreviation "six \ succ[5]" +notation "six" ("6") +abbreviation "seven \ succ[6]" +notation "seven" ("7") +abbreviation "eight \ succ[7]" +notation "eight" ("8") +abbreviation "nine \ succ[8]" +notation "nine" ("9") +abbreviation "ten \ succ[9]" +notation "ten" ("10") +abbreviation "eleven \ succ[10]" +notation "eleven" ("11") +abbreviation "twelve \ succ[11]" +notation "twelve" ("12") +abbreviation "thirteen \ succ[12]" +notation "thirteen" ("13") +abbreviation "fourteen \ succ[13]" +notation "fourteen" ("14") +abbreviation "fifteen \ succ[14]" +notation "fifteen" ("15") +*) -(*lemma neg_imp_negSuc (*[simp]*): - assumes h:"n \ {-.n : n \ Nat } \\ {0}" shows "isNeg(n)" -unfolding isNeg_def +definition two :: "c" ("2") + where "2 \ succ[1]" +definition three :: "c" ("3") + where "3 \ succ[2]" +definition four :: "c" ("4") + where "4 \ succ[3]" +definition five :: "c" ("5") + where "five \ succ[4]" +definition six :: "c" ("6") + where "six \ succ[5]" +definition seven :: "c" ("7") + where "seven \ succ[6]" +definition eight :: "c" ("8") + where "eight \ succ[7]" +definition nine :: "c" ("9") + where "nine \ succ[8]" +definition ten :: "c" ("10") + where "ten \ succ[9]" +definition eleven :: "c" ("11") + where "eleven \ succ[10]" +definition twelve :: "c" ("12") + where "twelve \ succ[11]" +definition thirteen :: "c" ("13") + where "thirteen \ succ[12]" +definition fourteen :: "c" ("14") + where "fourteen \ succ[13]" +definition fifteen :: "c" ("15") + where "fifteen \ succ[14]" + +lemma succType: "succ \ [Int \ Int]" proof - - from h have 1:"\k \ Nat : n = -.k" and 2:"n \ 0" by auto - from 1 obtain k where k:"k \ Nat" and 3:"n = -.k" by auto - with 2 have "k \ 0" by auto - with k not0_implies_Suc[of k] have "\m \ Nat : k = Succ[m]" by auto - with 3 show "\k \ Nat : n = -.(Succ[k])" by auto -qed*) - -lemma sgnm1_imp_neg: - assumes n:"n \ Int" shows "sgn(n) = -.1 \ isNeg(n)" -unfolding sgn_def using intThenPosZeroNeg[OF n] by auto + { + fix a + assume "a \ Int" + hence "succ[a] \ Int" + by (rule pr_intCases) (auto simp: succ_def intro: bChooseI2) + } + thus ?thesis by (auto simp: succ_def) +qed -lemma isPos_sgn [simp]: "isPos(sgn(n)) = isPos(n)" -unfolding isPos_def sgn_def by force +lemma succIsAFcn [intro!,simp]: "isAFcn(succ)" + using succType by blast + +\ \@{text "DOMAIN succ = Int"}\ +lemmas [intro!,simp] = funcSetDomain[OF succType] +\ \@{text "a \ Int \ succ[a] \ Int"}\ +lemmas [intro!,simp] = funcSetValue[OF succType] + +lemma succIsSucc: "n \ Nat \ succ[n] = Succ[n]" + by (simp add: succ_def) + +lemma succUminusSuccNat: "n \ Nat \ succ[-succ[n]] = -n" + by (auto simp: succ_def intro: bChooseI2) + +lemma succUminusSucc [simp]: + assumes "a \ Int" + shows "succ[-succ[a]] = -a" +using assms proof (rule pr_intCases) + fix n + assume n: "n \ Nat" "a = -Succ[n]" + hence "succ[a] = -n" + by (simp add: sym[OF succIsSucc] succUminusSuccNat) + with n show ?thesis + by (simp add: succIsSucc) +qed (simp add: succUminusSuccNat)+ + +lemma succInNat [simp]: + assumes "a \ Int" + shows "(succ[a] \ Nat) = (a \ Nat \ {-1})" +proof - + { + assume "a \ Nat" hence "succ[a] \ Nat" + by (simp add: succIsSucc) + } + moreover + { + assume "a = -1" hence "succ[a] \ Nat" + by simp + } + moreover + { + assume s: "succ[a] \ Nat" + from assms have "a \ Nat \ {-1}" + proof (rule pr_intCases) + fix n + assume "n \ Nat" "a = -Succ[n]" + with s show ?thesis by (simp add: sym[OF succIsSucc]) + qed (simp+) + } + ultimately show ?thesis by blast +qed -lemma sgnNat_is_0or1 (*[simp]*): - "n \ Nat \ sgn(n) = 0 \ sgn(n) = 1" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma succIsUminus [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(succ[a] = -b) = (a = -succ[b])" +proof - + { + assume "succ[a] = -b" + with b have "succ[b] = succ[-succ[a]]" by simp + also from a have "\ = -a" by simp + finally have "a = -succ[b]" using a by simp + } + with assms show ?thesis by auto +qed -lemma sgnNat_not0: - "\n \ Nat; sgn(n) \ 0\ \ sgn(n) = 1" -using sgnNat_is_0or1[of n] by auto +lemma uminusIsSucc [simp]: + assumes a: "a \ Int" and b: "b \ Int" + shows "(-a = succ[b]) = (b = -succ[a])" + using assms by (auto simp: uminusSwap) -lemma sgnNat_not1: - "\n \ Nat; sgn(n) \ 1\ \ n = 0" -using sgnNat_is_0or1[of n] by auto +lemma succIs0: (* better not add to simp *) + assumes "a \ Int" + shows "(succ[a] = 0) = (a = -1)" +proof - + from assms have "(succ[a] = -0) = (a = -1)" + by (rule succIsUminus) simp + thus ?thesis by simp +qed -lemma sgnNat_not_neg [simp]: - "n \ Nat \ sgn(n) = -.1 = FALSE" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma zeroIsSucc: + assumes "a \ Int" + shows "(0 = succ[a]) = (a = -1)" + by (auto simp: sym[OF succIs0[OF assms]]) + +lemma succNatNotZero (*[elim] -- "ignoring weak elimination rule" *): + "\succ[n] = 0; n \ Nat\ \ P" + "\0 = succ[n]; n \ Nat\ \ P" + by (auto simp: succIsSucc) + +lemma succNatZeroIff [simp]: + "n \ Nat \ (succ[n] = 0) = FALSE" + "n \ Nat \ (0 = succ[n]) = FALSE" + by (auto simp: succIsSucc) + +lemma succInj [dest]: + assumes eq: "succ[a] = succ[b]" and a: "a \ Int" and b: "b \ Int" + shows "a = b" +using a proof (rule pr_intCases) + assume a0: "a = 0" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" with a0 show ?thesis by simp + next + fix n + assume "n \ Nat" "b = Succ[n]" + with a0 eq show ?thesis by (simp add: succIsSucc) + next + fix n + assume "n \ Nat" "b = -Succ[n]" + with a0 eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +next + fix n + assume n: "n \ Nat" "a = Succ[n]" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" + with n eq show ?thesis by (simp add: succIsSucc) + next + fix m + assume "m \ Nat" "b = Succ[m]" + with n eq show ?thesis by (simp add: succIsSucc) + next + fix m + assume "m \ Nat" "b = -Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +next + fix n + assume n: "n \ Nat" "a = -Succ[n]" + from b show ?thesis + proof (rule pr_intCases) + assume "b = 0" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + next + fix m + assume "m \ Nat" "b = Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + next + fix m + assume "m \ Nat" "b = -Succ[m]" + with n eq show ?thesis by (simp add: sym[OF succIsSucc]) + qed +qed -lemma notNat_imp_sgn_neg1 [intro]: "n \ Nat \ sgn(n) = -.1" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma succInjIff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(succ[a] = succ[b]) = (a = b)" + using assms by auto + +lemma intIsSucc: + assumes "a \ Int" + shows "\b \ Int : a = succ[b]" +using assms proof (rule pr_intCases) + assume "a = 0" + hence "a = succ[-1]" by simp + moreover have "-1 \ Int" by simp + ultimately show ?thesis by blast +next + fix n + assume n: "n \ Nat" "a = Succ[n]" + hence "a = succ[n]" by (simp add: succIsSucc) + with n show ?thesis by blast +next + fix n + assume n: "n \ Nat" "a = -Succ[n]" + hence "a = succ[-succ[Succ[n]]]" by simp + with n show ?thesis by blast +qed -lemma eqSgnNat_imp_nat: "sgn(m) = sgn(n) \ m \ Nat \ n \ Nat" -unfolding sgn_def isPos_eq_inNat1 by auto +definition pred where + "pred = [a \ Int \ CHOOSE b \ Int : a = succ[b]]" -lemma eqSgn_imp_0_nat [simp]: "n \ Nat \ sgn(n) = sgn(-.n) = (n = 0)" -unfolding sgn_def isPos_def by force +lemma predType: "pred \ [Int \ Int]" + unfolding pred_def by (auto intro: intIsSucc[THEN bChooseI2]) -lemma eqSgn_imp_0_nat2 [simp]: "n \ Nat \ sgn(-.n) = sgn(n) = (n = 0)" -unfolding sgn_def isPos_def by force +lemma predIsAFcn [intro!,simp]: "isAFcn(pred)" + using predType by blast -lemma eqSgn_imp_0 [simp]: "n \ Int \ sgn(n) = sgn(-.n) = (n = 0)" -by(rule intCases, auto) +\ \@{text "DOMAIN pred = Int"}\ +lemmas [intro!,simp] = funcSetDomain[OF predType] +\ \@{text "a \ Int \ pred[a] \ Int"}\ +lemmas [intro!,simp] = funcSetValue[OF predType] -(*lemma notSgnNegs [simp]: "n \ Int \\ {0} \ sgn(n) = sgn(-.n) = FALSE" - by(rule intCases[of n], simp_all) *) +lemma predValue: + assumes "b \ Int" and "a = succ[b]" + shows "b = pred[a]" + using assms unfolding pred_def by (auto intro: bChooseI) -lemma sgn_eq_neg1_is_not_nat (*[simp]*): "(sgn(n) = -.1) = (n \ Nat \ n \ 0)" -unfolding sgn_def isPos_eq_inNat1 by auto +lemma pred0: "pred[0] = -1" + by (rule sym[OF predValue]) auto -lemma sgn_not_neg1_is_nat [simp]: "((sgn(n) = -.1) = FALSE) = (n \ Nat)" -by (auto simp: sgn_eq_neg1_is_not_nat) +lemma predSucc [simp]: + assumes "a \ Int" + shows "pred[succ[a]] = a" + by (rule predValue[OF assms, THEN sym, OF refl]) -lemma sgn_neg_eq_1_false: "\sgn(-.m) = 1; m \ Nat\ \ P" -unfolding sgn_def by auto +lemma succPred [simp]: + assumes "a \ Int" + shows "succ[pred[a]] = a" + using assms unfolding pred_def by (auto intro: intIsSucc[THEN bChooseI2]) -lemma sgn_minus [simp]: - assumes n: "n \ Int" - shows "sgn(-.n) = -.sgn(n)" -unfolding sgn_def using n by (cases, auto) +lemma predInj [dest]: + assumes eq: "pred[a] = pred[b]" and ab: "a \ Int" "b \ Int" + shows "a = b" +proof - + from eq have "succ[pred[a]] = succ[pred[b]]" + by simp + with ab show ?thesis by simp +qed -text \ Absolute value \ +lemma predInjIff [simp]: + assumes "a \ Int" and "b \ Int" + shows "(pred[a] = pred[b]) = (a = b)" + using assms by auto -lemma absIsNat [simp]: - assumes n: "n \ Int" shows "abs(n) \ Nat" -unfolding abs_def using intNotNatIsNegNat[OF _ n] by auto +lemma uminusSucc (*[simp]*): + assumes "a \ Int" + shows "-succ[a] = pred[-a]" +proof - + from assms have "succ[-succ[a]] = succ[pred[-a]]" + by simp + with assms show ?thesis + by blast +qed -lemma absNat [simp]: "n \ Nat \ abs(n) = n" -unfolding abs_def by auto +lemma uminusPred (*[simp]*): + assumes "a \ Int" + shows "-pred[a] = succ[-a]" + using assms by auto -lemma abs0 [simp]: "abs(0) = 0" -unfolding abs_def by simp +text \ + We restate the induction and case splitting theorems for + natural numbers and integers in terms of @{text "succ"} and + @{text "pred"}. +\ +lemma natInduct: + assumes "P(0)" and "\n. \ n \ Nat; P(n) \ \ P(succ[n])" + shows "\n \ Nat : P(n)" + using assms by (intro pr_natInduct) (auto simp: succIsSucc) -(** -lemma absn0 [simp]: "abs(-.0) = 0" -by simp -**) +\ \version of above suitable for the inductive reasoning package\ +lemma natInductE [case_names 0 Succ, induct set: Nat]: + assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(succ[n])" + shows "P(n)" +using bspec[OF natInduct, where P=P] assms by blast -lemma abs_negNat [simp]: "n \ Nat \ abs(-.n) = n" -unfolding abs_def by (auto dest: sgnNat_not1) +(*** EXAMPLE INDUCTION PROOFS *** -lemma abs_neg [simp]: - assumes n: "n \ Int" shows "abs(-.n) = abs(n)" -unfolding abs_def using n by (auto dest: sgnNat_not1) +lemma "\n\Nat : n=0 \ (\m \ Nat : n = succ[m])" +by (rule natInduct, auto) +lemma + assumes 1: "n \ Nat" + shows "n=0 \ (\m \ Nat : n = succ[m])" +using 1 by (induct, auto) -subsection \ Orders on integers \ +*** END EXAMPLE ***) -text \ - We distinguish four cases, depending on the arguments being in - Nat or negative. -\ +lemma natCases [case_names 0 succ, cases set: Nat]: + assumes n: "n \ Nat" + and z: "n=0 \ P" and sc: "\m. \m \ Nat; n = succ[m]\ \ P" + shows "P" +proof - + from n have "n=0 \ (\m \ Nat : n = succ[m])" + by (induct, auto) + thus ?thesis + proof + assume "n=0" thus "P" by (rule z) + next + assume "\m\Nat : n = succ[m]" + then obtain m where "m \ Nat" and "n = succ[m]" .. + thus "P" by (rule sc) + qed +qed -lemmas int_leq_pp_def = nat_leq_def - (* -- 'positive-positive' case, ie: both arguments are naturals *) +lemma not0_implies_Suc: + "\n \ Nat; n \ 0\ \ \m \ Nat: n = succ[m]" + by (rule natCases) auto + +text \Induction over two parameters along the ``diagonal''.\ +lemma diffInduction: + assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, succ[n])" + and step: "\m,n\Nat : P(m,n) \ P(succ[m], succ[n])" + shows "\m,n\Nat : P(m,n)" +proof (rule natInduct) + show "\n\Nat : P(0,n)" + using b1 b2 by (intro natInduct) auto +next + fix m + assume m: "m \ Nat" and ih: "\n\Nat : P(m,n)" + show "\n\Nat : P(succ[m],n)" + proof (rule bAllI) + fix n + assume "n \ Nat" thus "P(succ[m],n)" + proof (cases) + case 0 with b1 m show ?thesis by auto + next + case succ with step ih m show ?thesis by auto + qed + qed +qed -axiomatization where - int_leq_pn_def [simp]: "\a \ Nat; b \ Nat\ \ a \ -.b = (a = 0 \ b = 0)" -and - int_leq_np_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ b = TRUE" -and - int_leq_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a \ -.b = (b \ a)" +lemma diffInduct: + assumes n: "n \ Nat" and m: "m \ Nat" + and b1: "\m. m\Nat \ P(m,0)" and b2: "\n. n\Nat \ P(0, succ[n])" + and step: "\m n. \m \ Nat; n\Nat; P(m,n) \ \ P(succ[m], succ[n])" + shows "P(m,n)" +proof - + have "\m,n\Nat : P(m,n)" + by (rule diffInduction, auto intro: b1 b2 step) + with n m show ?thesis by blast +qed -(* lemmas int_leq_def = int_leq_pn_def int_leq_np_def int_leq_nn_def *) +lemma intInduct: + assumes z: "P(0)" + and pos: "\n. \n \ Nat; P(n); P(-n)\ \ P(succ[n])" + and neg: "\n. \n \ Nat; P(n); P(-n)\ \ P(-succ[n])" + shows "\a\Int : P(a)" +proof - + from assms have 1: "\n \ Nat : P(n) \ P(-n)" + by (intro natInduct) auto + show ?thesis + proof + fix a + assume "a \ Int" + from this 1 show "P(a)" + by (elim intElim) auto + qed +qed -lemma int_boolify_leq [simp]: - "\a \ Int; b \ Int\ \ boolify(a \ b) = (a \ b)" -by(rule intCases2[of a b], simp_all) +\ \turn the above lemma into an induction method\ +lemma intInductE [case_names 0 pos neg, induct set: Int]: + assumes "a \ Int" and "P(0)" + and "\n. \n \ Nat; P(n); P(-n)\ \ P(succ[n])" + and "\n. \n \ Nat; P(n); P(-n)\ \ P(-succ[n])" + shows "P(a)" + using bspec[OF intInduct, where P=P] assms by blast + +lemma intZeroPosNeg: + assumes "a \ Int" + shows "a = 0 \ (\n \ Nat : a = succ[n]) \ (\n \ Nat : a = -succ[n])" +using assms by (induct a) auto + +lemma intCases [case_names 0 pos neg, cases set: Int]: + assumes "a \ Int" + and "a = 0 \ P" + and "\n. \n \ Nat; a = succ[n]\ \ P" + and "\n. \n \ Nat; a = -succ[n]\ \ P" + shows "P" + using assms by (blast dest: intZeroPosNeg) + +lemma succIrrefl: + assumes "a \ Int" + shows "succ[a] \ a" + using assms by induct auto + +lemma succIrreflE (*[elim] -- don't: "ignoring weak elimination rule" *): + "\succ[a] = a; a \ Int\ \ P" + "\a = succ[a]; a \ Int\ \ P" +by (auto dest: succIrrefl) + +lemma succIrrefl_iff [simp]: + "a \ Int \ (succ[a] = a) = FALSE" + "a \ Int \ (a = succ[a]) = FALSE" +by (auto dest: succIrrefl) + +lemma predIrrefl: + assumes "a \ Int" + shows "pred[a] \ a" +proof + assume "pred[a] = a" + hence "succ[pred[a]] = succ[a]" by simp + with assms show "FALSE" by simp +qed -lemma int_leq_isBool [intro!,simp]: - "\a \ Int; b \ Int\ \ isBool(a \ b)" -unfolding isBool_def by auto +lemma predIrreflE (*[elim]*): + "\pred[a] = a; a \ Int\ \ P" + "\a = pred[a]; a \ Int\ \ P" +by (auto dest: predIrrefl) -lemma int_leq_refl [iff]: "n \ Int \ n \ n" -by(rule intCases, auto) +lemma predIrrefl_iff [simp]: + "a \ Int \ (pred[a] = a) = FALSE" + "a \ Int \ (a = pred[a]) = FALSE" +by (auto dest: predIrrefl) -lemma eq_leq_bothE: (* -- reduce equality over integers to double inequality *) - assumes "m \ Int" and "n \ Int" and "m = n" and "\m \ n; n \ m\ \ P" - shows "P" -using assms by simp +lemma predInNat [simp]: + assumes "a \ Int" + shows "(pred[a] \ Nat) = (a \ Nat \ {0})" + using assms by cases (auto simp: pred0 sym[OF uminusSucc]) -lemma neg_le_iff_le [simp]: - "\m \ Int; n \ Int \ \ -.n \ -.m = (m \ n)" - by (rule intCases2[of m n]) auto +lemma succNatNotNeg [simp]: + assumes "m \ Nat" and "n \ Nat" + shows "(succ[m] = -n) = FALSE" "(n = -succ[m]) = FALSE" + "(-n = succ[m]) = FALSE" "(-succ[m] = n) = FALSE" + using assms by auto -subsection \ Addition of integers \ +subsection \ Initial intervals of natural numbers \ text \ - Again, we distinguish four cases in the definition of @{text "a + b"}, - according to each argument being positive or negative. + The set of natural numbers up to (and including) a given $n$ is + inductively defined as the smallest set of natural numbers that + contains $n$ and all numbers whose successor is in the set. + + NB: ``less than'' is not first-order definable from the Peano axioms, + a set-theoretic definition such as the following seems to be unavoidable. \ -(* cf. NatArith *) -(** The following is rejected by Isabelle because the two definitions - are not distinguishable by argument types. -defs (unchecked overloaded) - int_add_def: "\a \ Int; b \ Int \ \ a + b \ - IF a \ Nat \ b \ Nat THEN addnat(a)[b] - ELSE IF isNeg(a) \ isNeg(b) THEN -.(addnat(-.a)[-.b]) - ELSE IF isNeg(a) THEN IF -.a \ b THEN b -- a ELSE -.(a -- b) - ELSE IF a \ -.b THEN -.(b -- a) ELSE a -- b" -**) +definition upto :: "c \ c" +where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : succ[k] \ S })" lemmas - int_add_pp_def = nat_add_def (* -- both numbers are positive, ie. naturals *) + setEqualI [where A = "upto(n)" for n, intro!] + setEqualI [where B = "upto(n)" for n, intro!] -axiomatization where -int_add_pn_def: "\a \ Nat; b \ Nat\ \ a + (-.b) \ IF a \ b THEN -.(b -- a) ELSE a -- b" -and -int_add_np_def: "\a \ Nat; b \ Nat\ \ (-.a) + b \ IF b \ a THEN -.(a -- b) ELSE b -- a" -and -int_add_nn_def [simp]: "\a \ Nat; b \ Nat\ \ (-.a) + (-.b) = -.(a + b)" +lemma uptoNat: "upto(n) \ Nat" + unfolding upto_def by (rule lfpSubsetDomain) -lemmas int_add_def = int_add_pn_def int_add_np_def (*int_add_nn_def*) - (* -- When we use these definitions, we don't want to unfold the 'pp' case *) +lemma uptoPred: + assumes suc: "succ[m] \ upto(n)" and m: "m \ Nat" and n: "n \ Nat" + shows "m \ upto(n)" +proof - + let ?f = "\S. {n} \ {k\Nat : succ[k] \ S}" + from n have mono: "Monotonic(Nat, ?f)" + unfolding Monotonic_def by blast + from m suc have 1: "m \ ?f(upto(n))" by auto + from mono have 2: "?f(upto(n)) \ upto(n)" + unfolding upto_def by (rule lfpPreFP) + from 1 2 show ?thesis by blast +qed -lemma int_add_neg_eq_natDiff [simp]: "\n \ m; m \ Nat; n \ Nat\ \ m + (-.n) = m -- n" -by (auto simp: int_add_pn_def dest: nat_leq_antisym) +lemma uptoZero: "upto(0) = {0}" +proof (rule setEqual) + have "{0} \ { k \ Nat : succ[k] \ {0} } \ {0}" by auto + thus "upto(0) \ {0}" + unfolding upto_def by (rule lfpLB, auto) +next + show "{0} \ upto(0)" + unfolding upto_def by (rule lfpGLB, auto) +qed -text \ Closure \ +lemma uptoSucc: + assumes n: "n \ Nat" + shows "upto(succ[n]) = upto(n) \ {succ[n]}" (is "?lhs = ?rhs") +proof - + let ?preds = "\S. {k \ Nat : succ[k] \ S}" + let ?f = "\S k. {k} \ ?preds(S)" + have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" + by (auto simp: Monotonic_def) + \ \``$\subseteq$''\ + from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto + also have "\ \ upto(n)" + by (unfold upto_def, rule lfpPreFP, rule mono, rule n) + finally have "?f(?rhs, succ[n]) \ ?rhs" by auto + moreover from n have "?rhs \ Nat" + by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) + ultimately have 1: "?lhs \ ?rhs" + by (unfold upto_def[where n="succ[n]"], rule lfpLB) + \ \``$\supseteq$''\ + from n mono have 2: "?f(?lhs, succ[n]) \ ?lhs" + unfolding upto_def by (intro lfpPreFP) auto + with n have "?f(?lhs, n) \ ?lhs" by auto + moreover have "?lhs \ Nat" by (rule uptoNat) + ultimately have 3: "upto(n) \ ?lhs" + unfolding upto_def[where n=n] by (rule lfpLB) + from 2 have 4: "succ[n] \ ?lhs" by auto + from 3 4 have "?rhs \ ?lhs" by auto + with 1 show ?thesis by (rule setEqual) +qed -lemma addIsInt [simp]: "\m \ Int; n \ Int\ \ m + n \ Int" -by (rule intCases2[of m n], auto simp: int_add_def) +lemma uptoRefl: + assumes n: "n \ Nat" + shows "n \ upto(n)" +using n proof (cases) + case 0 thus ?thesis by (simp add: uptoZero) +next + case (succ m) thus ?thesis by (auto simp: uptoSucc) +qed -text \ Neutral element \ +lemma zeroInUpto: + assumes n: "n \ Nat" + shows "0 \ upto(n)" +using n by (induct, auto simp: uptoZero uptoSucc) -lemma add_0_right_int [simp]: "n \ Int \ n + 0 = n" -by(rule intCases, auto simp add: int_add_np_def) -lemma add_0_left_int [simp]: "n \ Int \ 0 + n = n" -by(rule intCases, auto simp add: int_add_pn_def) +lemma SuccNotUptoZero: + assumes "n \ Nat" and "succ[n] \ upto(0)" + shows "P" +using assms by (auto simp: uptoZero) -text \ Additive inverse element \ +lemma uptoTrans: + assumes "k \ upto(m)" and "m \ upto(n)" and "n \ Nat" + shows "k \ upto(n)" +proof - + have "\n\Nat : m \ upto(n) \ upto(m) \ upto(n)" + by (rule natInduct, auto simp: uptoZero uptoSucc) + with assms show ?thesis by blast +qed -lemma add_inverse_nat [simp]: "n \ Nat \ n + -.n = 0" -by(simp add: int_add_pn_def) +lemma succNotinUpto: + assumes n: "n \ Nat" + shows "succ[n] \ upto(n)" +using n proof (induct) + show "1 \ upto(0)" by (auto simp: uptoZero) +next + fix n + assume n: "n \ Nat" and ih: "succ[n] \ upto(n)" + show "succ[succ[n]] \ upto(succ[n])" + proof (auto simp: uptoSucc n) + assume "succ[succ[n]] \ upto(n)" + with n have "succ[n] \ upto(n)" + by (auto elim: uptoPred) + with ih show "FALSE" .. + qed +qed -lemma add_inverse2_nat [simp]: "n \ Nat \ -.n + n = 0" -by(simp add: int_add_np_def) +lemma uptoLimit: + assumes m: "m \ upto(n)" and suc: "succ[m] \ upto(n)" and n: "n \ Nat" + shows "m=n" +proof - + from m uptoNat have mNat: "m \ Nat" by blast + from n have "\m\Nat: m \ upto(n) \ succ[m] \ upto(n) \ m=n" (is "?P(n)") + by (induct, auto simp: uptoZero uptoSucc) + with mNat m suc show ?thesis by blast +qed -lemma add_inverse_int [simp]: "n \ Int \ n + -.n = 0" -by (rule intCases, auto simp: int_add_def) +lemma uptoAntisym: + assumes mn: "m \ upto(n)" and nm: "n \ upto(m)" + shows "m=n" +proof - + from mn uptoNat have m: "m \ Nat" by blast + from nm uptoNat have n: "n \ Nat" by blast + have "\m,n\Nat : m \ upto(n) \ n \ upto(m) \ m=n" (is "\m,n\Nat : ?P(m,n)") + proof (rule natInduct) + show "\n\Nat : ?P(0,n)" by (auto simp: uptoZero) + next + fix m + assume m: "m \ Nat" and ih: "\n\Nat : ?P(m,n)" + show "\n\Nat : ?P(succ[m],n)" + proof (auto simp: uptoSucc m) + fix n + assume "succ[m] \ upto(n)" and "n \ upto(m)" + from this m have "succ[m] \ upto(m)" by (rule uptoTrans) + with m show "succ[m] = n" \ \contradiction\ + by (blast dest: succNotinUpto) + qed + qed + with m n mn nm show ?thesis by blast +qed -lemma add_inverse2_int [simp]: "n \ Int \ -.n + n = 0" -by (rule intCases, auto simp: int_add_def) +lemma uptoInj: + assumes n: "n \ Nat" and m: "m \ Nat" + shows "(upto(n) = upto(m)) = (n = m)" +proof (auto) + assume 1: "upto(n) = upto(m)" + from n have "n \ upto(n)" by (rule uptoRefl) + with 1 have "n \ upto(m)" by auto + moreover + from m have "m \ upto(m)" by (rule uptoRefl) + with 1 have "m \ upto(n)" by auto + ultimately + show "n = m" by (rule uptoAntisym) +qed -text \ Commutativity \ +lemma uptoLinear: + assumes m: "m \ Nat" and n: "n \ Nat" + shows "m \ upto(n) \ n \ upto(m)" (is "?P(m,n)") +using m proof induct + from n show "?P(0,n)" by (auto simp: zeroInUpto) +next + fix k + assume k: "k \ Nat" and ih: "?P(k,n)" + from k show "?P(succ[k],n)" + proof (auto simp: uptoSucc) + assume kn: "(succ[k] \ upto(n)) = FALSE" + show "n \ upto(k)" + proof (rule contradiction) + assume c: "n \ upto(k)" + with ih have "k \ upto(n)" by simp + from this kn n have "k = n" by (rule uptoLimit[simplified]) + with n have "n \ upto(k)" by (simp add: uptoRefl) + with c show "FALSE" .. + qed + qed +qed -lemma add_commute_pn_nat: "\m \ Nat; n \ Nat\ \ m + -.n = -.n + m" -by(simp add: int_add_def) +lemma uptoInduct: + assumes n: "n \ Nat" + and base: "P(0)" + and step: "\m. \m \ Nat; m \ upto(n) \ P(m); succ[m] \ upto(n)\ \ P(succ[m])" + shows "\m \ upto(n) : P(m)" +proof - + have "\m \ Nat: m \ upto(n) \ P(m)" + by (rule natInduct) (auto simp: base dest: step) + with uptoNat show ?thesis by auto +qed -lemma add_commute_int: "\m \ Int; n \ Int\ \ m + n = n + m" - by(rule intCases2[of m n], auto simp add: int_add_def add_commute_nat) -text \ Associativity \ +subsection \ Primitive recursive functions over natural numbers \ -lemma add_pn_eq_adiff [simp]: - "\m \ n; m \ Nat; n \ Nat\ \ m + -.n = -.(n -- m)" -by (simp add: int_add_def) +text \ + We justify the definition of function over natural numbers by + primitive recursion. The idea is to construct a sequence of + approximations, and then obtain the function by diagonalization. +\ -lemma adiff_add_assoc5: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ p; p \ m + n; m \ p -- n\ \ -.(p -- n -- m) = m + n -- p" -apply (induct p n rule: diffInduct) -using assms by (auto dest: nat_leq_antisym) +definition primrec_nat_upto where + "primrec_nat_upto(base, step, f, n) \ + isAFcn(f) \ DOMAIN f = Nat + \ f[0] = base + \ (\m \ Nat : succ[m] \ upto(n) \ f[succ[m]] = step(m, f[m]))" + +lemma primrec_nat_upto_deterministic: + assumes n: "n \ Nat" and m: "m \ upto(n)" + and f: "primrec_nat_upto(base, step, f, n)" + and g: "primrec_nat_upto(base, step, g, m)" + and k: "k \ upto(m)" + shows "f[k] = g[k]" +proof - + from m uptoNat have mNat: "m \ Nat" by auto + hence "\k \ upto(m): f[k] = g[k]" + proof (rule uptoInduct) + from f g show "f[0] = g[0]" unfolding primrec_nat_upto_def by simp + next + fix k + assume kNat: "k \ Nat" and ih: "k \ upto(m) \ f[k] = g[k]" + and sk: "succ[k] \ upto(m)" + from sk kNat g have "g[succ[k]] = step(k, g[k])" + unfolding primrec_nat_upto_def by simp + moreover + from sk m n have "succ[k] \ upto(n)" + by (rule uptoTrans) + with kNat f have "f[succ[k]] = step(k, f[k])" + unfolding primrec_nat_upto_def by simp + moreover + from sk kNat mNat have "k \ upto(m)" + by (rule uptoPred) + ultimately show "f[succ[k]] = g[succ[k]]" + using ih by simp + qed + with k show ?thesis by blast +qed -lemma adiff_add_assoc6: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ p; m + n \ p; p -- n \ m\ \ m -- (p -- n) = -.(p -- (m + n))" -apply (induct p n rule: diffInduct) -using assms by (auto dest: nat_leq_antisym) +lemma primrec_nat_upto_exists: + "\n \ Nat : \f : primrec_nat_upto(base, step, f, n)" +proof (rule natInduct) + have "primrec_nat_upto(base, step, [n \ Nat \ base], 0)" + unfolding primrec_nat_upto_def by (auto simp: uptoZero) + thus "\f : primrec_nat_upto(base, step, f, 0)" .. +next + fix n + assume n: "n \ Nat" + and ih: "\f : primrec_nat_upto(base, step, f, n)" + from n have ssn: "succ[succ[n]] \ n" + by induct (simp+) + from ih obtain f where f: "primrec_nat_upto(base, step, f, n)" .. + define g where "g \ [f EXCEPT ![succ[n]] = step(n, f[n])]" + from f n have "isAFcn(g) \ DOMAIN g = Nat \ g[0] = base" + by (simp add: primrec_nat_upto_def g_def) + moreover + { + fix m + assume 1: "m \ Nat" and 2: "succ[m] \ upto(n)" + have "m \ succ[n]" + proof + assume 3: "m = succ[n]" + from n have "n \ upto(succ[succ[n]])" + by (simp add: uptoSucc uptoRefl) + with 2 3 have "succ[succ[n]] = n" + by (auto dest: uptoAntisym) + with ssn show "FALSE" .. + qed + with f n 1 2 have "g[succ[m]] = step(m, g[m])" + by (auto simp: primrec_nat_upto_def g_def) + } + moreover + from f n have "g[succ[n]] = step(n, g[n])" + by (simp add: primrec_nat_upto_def g_def) + ultimately have "primrec_nat_upto(base, step, g, succ[n])" + using n by (auto simp: primrec_nat_upto_def uptoSucc) + thus "\f : primrec_nat_upto(base, step, f, succ[n])" .. +qed -lemma adiff_add_assoc7: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\p + n \ m; m \ n\ \ -.(m -- (p + n)) = n -- m + p" -apply (induct n m rule: diffInduct) -using assms by simp_all +lemma primrec_nat: + "\f : isAFcn(f) \ DOMAIN f = Nat + \ f[0] = base \ (\n \ Nat : f[succ[n]] = step(n,f[n]))" +proof - + from primrec_nat_upto_exists[THEN fcnConstruct] obtain F where + F: "isAFcn(F) \ DOMAIN F = Nat + \ (\n\Nat : primrec_nat_upto(base, step, F[n], n))" + by blast + define g where "g \ [n \ Nat \ F[n][n]]" + have "isAFcn(g)" "DOMAIN g = Nat" + unfolding g_def by (simp+) + moreover + from F have "g[0] = base" + by (simp add: primrec_nat_upto_def g_def) + moreover + have "\n \ Nat : g[succ[n]] = step(n, g[n])" + proof + fix n + assume n: "n \ Nat" + hence 1: "n \ upto(n)" "n \ upto(succ[n])" "succ[n] \ upto(succ[n])" + by (simp add: uptoSucc uptoRefl)+ + moreover + from F n have 2: "primrec_nat_upto(base, step, F[n], n)" + and 3: "primrec_nat_upto(base, step, F[succ[n]], succ[n])" + by auto + ultimately have 4: "F[succ[n]][n] = F[n][n]" + using n by (auto intro: primrec_nat_upto_deterministic[where n="succ[n]"]) + with n 1 3 show "g[succ[n]] = step(n, g[n])" + by (auto simp: g_def primrec_nat_upto_def) + qed + ultimately + show ?thesis by blast +qed -lemma adiff_add_assoc8: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\n \ m; p \ m -- n; p \ m; m -- p \ n\ \ m -- n -- p = -.(n -- (m -- p))" -using adiff_add_assoc6[OF n p m] apply simp -using leq_adiff_right_add_left[OF _ p m n] add_commute_nat[OF p n] apply simp -by(rule adiff_adiff_left_nat[OF m n p]) +lemma bprimrec_nat: + assumes base: "base \ S" and step: "\n \ Nat : \x \ S : step(n,x) \ S" + shows "\f \ [Nat \ S] : f[0] = base \ (\n \ Nat: f[succ[n]] = step(n,f[n]))" +proof - + from primrec_nat[of base step] obtain f where + 1: "isAFcn(f)" and 2: "DOMAIN f = Nat" + and 3: "f[0] = base" and 4: "\n\Nat : f[succ[n]] = step(n,f[n])" + by blast + have "\n\Nat : f[n] \ S" + proof (rule natInduct) + from 3 base show "f[0] \ S" by simp + next + fix n + assume "n \ Nat" and "f[n] \ S" + with step 4 show "f[succ[n]] \ S" by force + qed + with 1 2 3 4 show ?thesis + by blast +qed -declare leq_neq_iff_less [simplified,simp] +definition nat_primrec where + "nat_primrec(S, base, step) \ + CHOOSE f \ [Nat \ S] : f[0] = base + \ (\n \ Nat : f[succ[n]] = step(n, f[n]))" -lemma int_add_assoc1: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (n + -.p) = (m + n) + -.p" +lemma nat_primrecE: + assumes f: "f = nat_primrec(S, base, step)" (is "f = ?g") + and base: "base \ S" + and step: "\n \ Nat : \x \ S : step(n,x) \ S" + and maj: "\ f \ [Nat \ S]; f[0] = base; \n \ Nat : f[succ[n]] = step(n, f[n]) \ \ P" + shows "P" proof - -have s1_1: "n \ p ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s1_1_asm: "n \ p" - have s2_1: "n + -.p = -.(p -- n)" - using s1_1_asm n p int_add_pn_def[of "n" "p"] by auto - have s2_2: "p \ m + n ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s2_2_asm: "p \ m + n" - have s3_1: "m + (n + -.p) = m -- (p -- n)" - proof - - have s4_1: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_2: "m + -.(p -- n) = m -- (p -- n)" - proof - - have s5_1: "p -- n \ m" - using s1_1_asm s2_2_asm - leq_adiff_left_add_right_equiv[of "n" "p" "m"] - p n m by auto - show ?thesis - using s5_1 p n m adiffIsNat - int_add_pn_def by auto - qed - show ?thesis - using s4_1 s4_2 by auto - qed - have s3_2: "(m + n) + -.p = (m + n) -- p" - using s2_2_asm p n m addIsNat int_add_pn_def by auto - have s3_3: "m -- (p -- n) = (m + n) -- p" - using s1_1_asm s2_2_asm - adiff_add_assoc3[of "n" "p" "m"] m n p by auto - show "m + (n + -.p) = (m + n) + -.p" - using s3_1 s3_2 s3_3 by auto - qed - have s2_3: "\ (p \ m + n) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s2_3_asm: "\ (p \ m + n)" - have s3_1: "m + n < p" - using s2_3_asm p m n addIsNat nat_not_leq by auto - have s3_2: "m + n \ p" - using s3_1 by (auto simp: less_def) - have s3_3: "(m + n) + -.p = -.(p -- (m + n))" - using s3_2 m n p addIsNat int_add_pn_def by auto - have s3_4: "p -- n \ m ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s3_4_asm: "p -- n \ m" - have s4_1: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_2: "m + -.(p -- n) = m -- (p -- n)" - using s3_4_asm p n m adiffIsNat int_add_pn_def - by auto - have s4_3: "m + (n + -.p) = m -- (p -- n)" - using s4_1 s4_2 by auto - have s4_4: "m -- (p -- n) = -.(p -- (m + n))" - using m n p s1_1_asm s3_2 s3_4_asm - adiff_add_assoc6 by auto - show "m + (n + -.p) = (m + n) + -.p" - using s3_3 s4_3 s4_4[symmetric] by auto - qed - have s3_5: "\ (p -- n \ m) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s3_5_asm: "\ (p -- n \ m)" - have s4_1: "m \ p -- n" - proof - - have s5_1: "m < p -- n" - using p n m adiffIsNat nat_not_leq s3_5_asm by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "m + (n + -.p) = m + -.(p -- n)" - using s2_1 by auto - have s4_3: "m + -.(p -- n) = -.((p -- n) -- m)" - using s4_1 s3_5_asm p n m adiffIsNat - int_add_pn_def[of "m" "p -- n"] - by auto - have s4_4: "m + (n + -.p) = -.((p -- n) -- m)" - using s4_2 s4_3 by auto - have s4_5: "(p -- n) -- m = p -- (n + m)" - using p n m adiff_adiff_left_nat by auto - have s4_6: "(p -- n) -- m = p -- (m + n)" - using n m add_commute_nat s4_5 by auto - have s4_7: "-.((p -- n) -- m) = -.(p -- (m + n))" - using s4_6 by auto - show "m + (n + -.p) = (m + n) + -.p" - using s4_7 s4_4 s3_3 by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s3_4 s3_5 nat_leq_isBool[of "p -- n" "m"] - isBoolTrueFalse[of "p -- n \ m"] - by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s2_2 s2_3 - nat_leq_isBool[of "p" "m + n"] - isBoolTrueFalse[of "p \ m + n"] - by auto - qed -have s1_2: "\ (n \ p) ==> - m + (n + -.p) = (m + n) + -.p" - proof - - assume s1_2_asm: "\ (n \ p)" - have s2_1: "n + -.p = n -- p" - proof - - have s3_1: "\n \ Nat; p \ Nat\ - \ - n + (-.p) \ - IF n \ p THEN -.(p -- n) ELSE n -- p" - using int_add_pn_def[of "n" "p"] by auto - have s3_2: "n + (-.p) \ n -- p" - using s3_1 s1_2_asm n p by auto - show ?thesis - using s3_2 by auto - qed - have s2_2: "\ ((m + n) \ p)" - proof - - have s3_1: "p < n" - using p n s1_2_asm nat_not_leq by auto - have s3_2: "p < m + n" - using m p n s3_1 trans_less_add2 by auto - show ?thesis - using s3_2 p m n addIsNat nat_not_leq by auto - qed - have s2_3: "(m + n) + -.p = (m + n) -- p" - proof - - have s3_1: "m + n \\in Nat" - using m n addIsNat by auto - have s3_2: "\(m + n) \\in Nat; p \\in Nat\ - \ - (m + n) + (-.p) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE ((m + n) -- p)" - using int_add_pn_def[of "m + n" "p"] by auto - have s3_3: "(m + n) + (-.p) \ (m + n) -- p" - using s3_1 p s3_2 s2_2 by auto - show ?thesis - using s3_3 by auto - qed - have s2_4: "(m + (n + -.p) = (m + n) + -.p) = - (m + (n -- p) = (m + n) -- p)" - using s2_1 s2_3 by auto - have s2_5: "(m + n) -- p = m + (n -- p)" - proof - - have s3_1: "p \ n" - proof - - have s4_1: "p < n" - using s1_2_asm p n nat_not_leq by auto - show ?thesis - using s4_1 by (auto simp: less_def) - qed - show ?thesis - using m n p s3_1 adiff_add_assoc[of "p" "n" "m"] by auto - qed - show "m + (n + -.p) = (m + n) + -.p" - using s2_4 s2_5[symmetric] by auto - qed -show ?thesis - proof - - have s2_1: "((n \ p) = TRUE) \ ((n \ p) = FALSE)" - using nat_leq_isBool[of "n" "p"] isBoolTrueFalse[of "n \ p"] - by auto - show ?thesis - using s1_1 s1_2 s2_1 by auto - qed + from base step have "\ g \ [Nat \ S] : g[0] = base \ (\n \ Nat: g[succ[n]] = step(n,g[n]))" + by (rule bprimrec_nat) + hence "?g \ [Nat \ S] \ ?g[0] = base \ (\n \ Nat: ?g[succ[n]] = step(n,?g[n]))" + unfolding nat_primrec_def by (rule bChooseI2, auto) + with f maj show ?thesis by blast qed -lemma int_add_assoc2: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (-.p + n) = (m + -.p) + n" +lemma nat_primrecType: + assumes "base \ S" and "\n \ Nat : \x \ S : step(n,x) \ S" + shows "nat_primrec(S, base, step) \ [Nat \ S]" proof - -have s1_1: "(m + n) + -.p = (m + -.p) + n" - proof - - have s2_1: "m + n \ p ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s2_1_asm: "m + n \ p" - have s3_1: "(m + n) + -.p = -.(p -- (m + n))" - proof - - have s4_1: "\m + n \ Nat; p \ Nat\ - \ ((m + n) + (-.p)) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE - (m + n) -- p" - using int_add_pn_def[of "m + n" "p"] by auto - have s4_2: "(m + n) + -.p \ -.(p -- (m + n))" - using s4_1 s2_1_asm m n p addIsNat by auto - show ?thesis - using s4_2 by auto - qed - have s3_2: "m \ p" - using s2_1_asm add_leqD1 m n p by auto - have s3_3: "(m + -.p) + n = -.(p -- m) + n" - proof - - have s4_1: "m + -.p = -.(p -- m)" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ -.(p -- m)" - using s5_1 m p s3_2 by auto - show ?thesis - using s5_2 by auto - qed - show ?thesis - using s4_1 by auto - qed - have s3_4: "n \ p -- m ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_4_asm: "n \ p -- m" - have s4_1: "(m + -.p) + n = -.((p -- m) -- n)" - proof - - have s5_1: "-.(p -- m) + n = n + -.(p -- m)" - using p n m adiffIsNat add_commute_pn_nat by auto - have s5_2: "n + -.(p -- m) = -.((p -- m) -- n)" - proof - - have s6_1: "\n \ Nat; p -- m \ Nat\ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s6_2: "n + -.(p -- m) \ -.((p -- m) -- n)" - using s6_1 n m p adiffIsNat s3_4_asm by auto - show ?thesis - using s6_2 by auto - qed - show ?thesis - using s3_3 s5_1 s5_2 by auto - qed - have s4_2: "(p -- m) -- n = p -- (m + n)" - using p n m adiff_adiff_left_nat by auto - show "(m + n) + -.p = (m + -.p) + n" - using s3_1 s4_1 s4_2 by auto - qed - have s3_5: "\ (n \ p -- m) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_5_asm: "\ (n \ p -- m)" - have s4_1: "p -- m \ n" - proof - - have s5_1: "p -- m < n" - using s3_5_asm nat_not_leq p n m adiffIsNat by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "-.(p -- m) + n = n -- (p -- m)" - proof - - have s5_1: "-.(p -- m) + n = n + -.(p -- m)" - using p n m add_commute_pn_nat adiffIsNat by auto - have s5_2: "n + -.(p -- m) = n -- (p -- m)" - proof - - have s6_1: "\n \ Nat; p -- m \ Nat\ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s6_2: "n + -.(p -- m) \ n -- (p -- m)" - using s6_1 s3_5_asm p n m adiffIsNat by auto - show ?thesis - using s6_2 by auto - qed - show ?thesis - using s5_1 s5_2 by auto - qed - have s4_3: "n -- (p -- m) = -.(p -- (m + n))" - using adiff_add_assoc6 m n p s3_2 s2_1_asm - s4_1 add_commute_nat by auto - show "(m + n) + -.p = (m + -.p) + n" - using s3_1 s3_3 s4_2 s4_3[symmetric] by auto - qed - show ?thesis - using s3_4 s3_5 nat_leq_isBool[of "n" "p -- m"] - isBoolTrueFalse[of "n \ p -- m"] by auto - qed - have s2_2: "\ (m + n \ p) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s2_2_asm: "\ (m + n \ p)" - have s3_1: "(m + n) + -.p = (m + n) -- p" - proof - - have s4_1: "\(m + n) \ Nat; p \ Nat\ - \ ((m + n) + (-.p)) \ - IF m + n \ p THEN - -.(p -- (m + n)) ELSE - (m + n) -- p" - using int_add_pn_def[of "m + n" "p"] by auto - have s4_2: "((m + n) + (-.p)) \ ((m + n) -- p)" - proof - - have s5_1: "m + n \\in Nat" - using m n addIsNat by auto - show "((m + n) + (-.p)) \ ((m + n) -- p)" - using s2_2_asm s4_1[OF s5_1 p] by auto - qed - show ?thesis - using s4_2 by auto - qed - have s3_2: "m \ p ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_2_asm: "m \ p" - have s4_1: "m + -.p = -.(p -- m)" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ -.(p -- m)" - using s5_1 m p s3_2_asm by auto - show ?thesis - using s5_2 by auto - qed - have s4_4: "p \ m + n" - proof - - have s5_1: "p < m + n" - using s2_2_asm p m n addIsNat nat_not_leq by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "n \ p -- m ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s4_2_asm: "n \ p -- m" - have s5_1: "(m + -.p) + n = -.((p -- m) -- n)" - proof - - have s6_1: "(m + -.p) + n = -.(p -- m) + n" - using s4_1 by auto - have s6_2: "-.(p -- m) + n = n + -.(p -- m)" - using p m n adiffIsNat add_commute_pn_nat by auto - have s6_3: "n + -.(p -- m) = -.((p -- m) -- n)" - proof - - have s7_1: "\n \ Nat; p -- m \ Nat - \ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s7_2: "n + -.(p -- m) \ -.((p -- m) -- n)" - using s7_1 p n m adiffIsNat s4_2_asm by auto - show ?thesis - using s7_2 by auto - qed - show ?thesis - using s6_1 s6_2 s6_3 by auto - qed - have s5_2: "-.((p -- m) -- n) = (m + n) -- p" - using adiff_add_assoc5 s4_4 s4_2_asm s3_2_asm - add_commute_nat - m n p by auto - show "(m + n) + -.p = (m + -.p) + n" - using s5_1 s5_2 s3_1 by auto - qed - have s4_3: "\ (n \ p -- m) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s4_3_asm: "\ (n \ p -- m)" - have s5_1: "(m + n) -- p = n -- (p -- m)" - using m n p adiff_add_assoc3[of "m" "p" "n"] - s3_2_asm s4_4 add_commute_nat - by auto - have s5_2: "(m + -.p) + n = n -- (p -- m)" - proof - - have s6_1: "(m + -.p) + n = -.(p -- m) + n" - using s4_1 by auto - have s6_2: "-.(p -- m) + n = n + -.(p -- m)" - using s6_1 n p m adiffIsNat - add_commute_pn_nat by auto - have s6_3: "n + -.(p -- m) = n -- (p -- m)" - proof - - have s7_1: "\n \ Nat; p -- m \ Nat - \ - \ (n + (-.(p -- m))) \ - IF n \ p -- m THEN - -.((p -- m) -- n) ELSE - n -- (p -- m)" - using int_add_pn_def[of "n" "p -- m"] by auto - have s7_2: "n + -.(p -- m) \ n -- (p -- m)" - using s7_1 s4_3_asm n p m adiffIsNat - by auto - show ?thesis - using s7_2 by auto - qed - show ?thesis - using s6_1 s6_2 s6_3 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s5_1 s3_1 s5_2 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s4_2 s4_3 nat_leq_isBool[of "n" "p -- m"] - isBoolTrueFalse[of "n \ p -- m"] by auto - qed - have s3_3: "\ (m \ p) ==> - (m + n) + -.p = (m + -.p) + n" - proof - - assume s3_3_asm: "\ (m \ p)" - have s4_1: "p \ m" - proof - - have s5_1: "p < m" - using s3_3_asm p m nat_not_leq by auto - show ?thesis - using s5_1 by (auto simp: less_def) - qed - have s4_2: "(m + n) -- p = (m -- p) + n" - using adiff_add_assoc2 s4_1 p n m by auto - have s4_3: "m + -.p = m -- p" - proof - - have s5_1: "\m \ Nat; p \ Nat\ - \ (m + (-.p)) \ - IF m \ p THEN - -.(p -- m) ELSE - m -- p" - using int_add_pn_def[of "m" "p"] by auto - have s5_2: "m + -.p \ m -- p" - using s5_1 m p s3_3_asm by auto - show ?thesis - using s5_2 by auto - qed - show "(m + n) + -.p = (m + -.p) + n" - using s4_2 s3_1 s4_3 by auto - qed - show ?thesis - using s3_2 s3_3 nat_leq_isBool[of "m" "p"] - isBoolTrueFalse[of "m \ p"] by auto - qed - show ?thesis - using s2_1 s2_2 nat_leq_isBool[of "m + n" "p"] - isBoolTrueFalse[of "m + n \ p"] by auto - qed -have s1_2: "m + (n + -.p) = (m + -.p) + n" - using s1_1 int_add_assoc1[OF m n p] by auto -have s1_3: "n + -.p = -.p + n" - using n p add_commute_pn_nat[of "n" "p"] by auto -show ?thesis - using s1_2 s1_3 by auto + define f where "f \ nat_primrec(S, base, step)" + from f_def[THEN meta_to_obj_eq] assms show ?thesis + by (rule nat_primrecE) (simp add: f_def) qed -declare leq_neq_iff_less [simplified,simp del] - -lemma int_add_assoc3: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + -.(n + p) = m + -.n + -.p" -apply(rule nat_leq_cases[of "n + p" m]) - using assms apply simp_all - apply(rule nat_leq_cases[OF n m], simp_all) - apply(rule nat_leq_cases[of p "m -- n"], simp_all) - apply(rule adiff_adiff_left_nat[symmetric], simp+) - using adiff_add_assoc6 add_commute_nat[OF n p] apply simp - using adiff_add_assoc2[OF _ p n m, symmetric] apply (simp add: adiff_is_0_eq') - apply(rule nat_leq_cases[OF n m], simp_all) - apply(rule nat_leq_cases[of p "m -- n"], simp_all) - using adiff_add_assoc5[symmetric] add_commute_nat[OF n p] apply simp - using adiff_add_assoc3[symmetric] add_commute_nat[OF n p] apply simp - using adiff_add_assoc2[symmetric] add_commute_nat[OF n p] apply simp -done - -lemma int_add_assoc4: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (n + p) = (-.m + n) + p" -using assms add_commute_int[of "-.m" "n + p"] add_commute_int[of "-.m" n] apply simp -apply(rule nat_leq_cases[of m "n + p" ], simp_all) - apply(rule nat_leq_cases[OF m n], simp_all) - apply(rule adiff_add_assoc2, simp+) - apply(simp add: add_commute_int[of "-.(m -- n)" p]) - apply(rule nat_leq_cases[of "m -- n" p], simp_all) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc3[symmetric]) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc5[symmetric]) - apply(rule nat_leq_cases[OF m n], simp_all) - apply(simp only: add_commute_nat[of n p]) - apply(rule adiff_add_assoc7, simp_all) - apply(simp add: add_commute_int[of "-.(m -- n)" p]) - apply(rule nat_leq_cases[of "m -- n" p], simp+) - apply(simp only: add_commute_nat[of n p]) - apply(simp only: adiff_add_assoc6[symmetric]) - apply(simp only: add_commute_nat[of n p]) - apply(simp add: add_commute_nat[of p n]) - apply(rule adiff_adiff_left_nat[symmetric], simp+) -done - -lemma leq_adiff_right_imp_0: - assumes h:"n \ n -- p" "p \ n" and n: "n \ Nat" and p: "p \ Nat" - shows "p = 0" -using p h apply (induct) -using n by auto - -lemma int_add_assoc5: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (n + -.p) = -.m + n + -.p" -using assms -apply(simp add: add_commute_int[of "-.m" "n + -.p"] add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[OF p n], simp_all) - apply(rule nat_leq_cases[of m "n -- p"], simp+) - apply(rule nat_leq_cases[of m n], simp+) - apply(rule nat_leq_cases[of p "n -- m"], simp_all) - apply(rule adiff_commute_nat[OF n p m]) - apply(rule adiff_add_assoc8, simp+) - using nat_leq_trans[of n m "n -- p"] apply simp - using leq_adiff_right_imp_0[OF _ _ n p] nat_leq_antisym[of m n] apply simp - apply(rule nat_leq_cases[OF m n], simp_all) - apply(rule nat_leq_cases[of p "n -- m"], simp_all) - apply(rule adiff_add_assoc8[symmetric], simp_all) - using leq_adiff_left_add_right[OF _ p n m] - add_commute_nat[OF p m] - apply(simp add: adiff_add_assoc3) - apply(simp add: adiff_add_assoc4) - apply(rule nat_leq_cases[of m n], simp_all) - apply(rule nat_leq_cases[of p "n -- m"], simp+) - using nat_leq_trans[of n p "n -- m"] apply simp - using leq_adiff_right_imp_0[OF _ _ n m] apply simp - using nat_leq_antisym[of n p] apply simp - apply(rule minusInj, simp) - apply(rule adiff_add_assoc4[symmetric], simp+) - apply(simp add: adiff_add_assoc2[symmetric]) - apply(simp add: add_commute_nat) -done - -lemma int_add_assoc6: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m + (-.n + p) = -.(m + n) + p" -using assms - add_commute_int[of "-.n" p] - add_commute_int[of "-.m" "p + -.n"] - add_commute_int[of "-.(m + n)" p] apply simp -apply(rule nat_leq_cases[OF n p], simp_all) - apply(rule nat_leq_cases[of m "p -- n"], simp+) - apply(rule nat_leq_cases[of "m + n" p], simp+) - apply(simp only: add_commute_nat[of m n]) - apply(rule adiff_adiff_left_nat, simp_all) - apply(simp only: minus_sym[symmetric]) - apply(rule adiff_add_assoc5, simp_all) - apply(rule nat_leq_cases[of "m + n" p], simp_all) - apply(simp only: minus_sym) - apply(rule adiff_add_assoc6, simp_all) - apply(rule adiff_add_assoc3, simp_all) - apply(rule nat_leq_cases[of "m + n" p], simp_all) - apply(simp only: minus_sym) - apply(rule adiff_add_assoc7[symmetric], simp_all) - apply(simp add: add_commute_nat[of "n -- p" m]) - apply(rule adiff_add_assoc[symmetric], simp+) -done - -lemma add_assoc_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m + (n + p) = (m + n) + p" -using m n p -by (rule intCases3, - auto simp: add_assoc_nat int_add_assoc1 int_add_assoc2 int_add_assoc3 - int_add_assoc4 int_add_assoc5 int_add_assoc6) - -text \ Minus sign distributes over addition \ - -lemma minus_distrib_pn_int [simp]: - "m \ Nat \ n \ Nat \ -.(m + -.n) = -.m + n" -apply(simp add: add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[of n m], simp_all) -done - -lemma minus_distrib_np_int [simp]: - "m \ Nat \ n \ Nat \ -.(-.m + n) = m + -.n" -by(simp add: add_commute_int) - -lemma int_add_minus_distrib [simp]: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m + n) = -.m + -.n" -by (rule intCases2[OF m n], simp_all) - - -subsection \ Multiplication of integers \ - -axiomatization where - int_mult_pn_def: "\a \ Nat; b \ Nat\ \ a * -.b = -.(a * b)" -and - int_mult_np_def: "\a \ Nat; b \ Nat\ \ -.a * b = -.(a * b)" -and - int_mult_nn_def [simp]: "\a \ Nat; b \ Nat\ \ -.a * -.b = a * b" - -lemmas int_mult_def = int_mult_pn_def int_mult_np_def (*int_mult_nn_def*) - -text \ Closure \ - -lemma multIsInt [simp]: "\a \ Int; b \ Int\ \ a * b \ Int" -by (rule intCases2[of a b], simp_all add: int_mult_def) - -text \ Neutral element \ - -lemma mult_0_right_int [simp]: "a \ Int \ a * 0 = 0" -by (rule intCases[of a], simp_all add: int_mult_np_def) - -lemma mult_0_left_int [simp]: "a \ Int \ 0 * a = 0" -by (rule intCases[of a], simp_all add: int_mult_pn_def) - -text \ Commutativity \ - -lemma mult_commute_int: "\a \ Int; b \ Int\ \ a * b = b * a" -by (rule intCases2[of a b], simp_all add: int_mult_def mult_commute_nat) - -text \ Identity element \ - -lemma mult_1_right_int [simp]: "a \ Int \ a * 1 = a" -by (rule intCases[of a], simp_all add: int_mult_def) - -lemma mult_1_left_int [simp]: "a \ Int \ 1 * a = a" -by (rule intCases[of a], simp_all add: int_mult_def) - -text \ Associativity \ - -lemma mult_assoc_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m * (n * p) = (m * n) * p" -by(rule intCases3[OF m n p], simp_all add: mult_assoc_nat int_mult_def) - -text \ Distributivity \ - -lemma ppn_distrib_left_nat: (* ppn stands for m=positive, n=positive, p=negative *) - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n + -.p) = m * n + -.(m * p)" -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"]) - using assms apply(simp_all add: adiff_mult_distrib2_nat int_mult_def) -done - -lemma npn_distrib_left_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m * (n + -.p) = -.(m * n) + m * p" -using assms apply (simp add: add_commute_int[of "-.(m * n)" "m * p"]) -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"], simp_all) - apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) -done - -lemma nnp_distrib_left_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "-.m * (-.n + p) = m * n + -.(m * p)" -using assms apply (simp add: add_commute_int[of "-.n" p]) -apply(rule nat_leq_cases[OF p n]) - apply(rule nat_leq_cases[of "m * p" "m * n"], simp_all) - apply (auto simp: adiff_mult_distrib2_nat int_mult_def dest: nat_leq_antisym) -done - -lemma distrib_left_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "m * (n + p) = (m * n + m * p)" -apply(rule intCases3[OF m n p], - simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) - apply(rule add_mult_distrib_left_nat, assumption+) - apply(rule ppn_distrib_left_nat, assumption+) - apply(simp add: add_commute_int, rule ppn_distrib_left_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_left_nat)+ - apply(rule npn_distrib_left_nat, assumption+) - apply(rule nnp_distrib_left_nat, assumption+) - apply(simp only: add_mult_distrib_left_nat) -done - -lemma pnp_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + -.n) * p = m * p + -.(n * p)" -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - using assms apply(simp_all add: adiff_mult_distrib_nat int_mult_def) -done - -lemma pnn_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + -.n) * -.p = -.(m * p) + n * p" -using assms apply (simp add: add_commute_int[of "-.(m * p)" "n * p"]) -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - apply (auto simp: adiff_mult_distrib_nat int_mult_def dest: nat_leq_antisym) -done - -lemma npn_distrib_right_nat: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(-.m + n) * -.p = m * p + -.(n * p)" -using assms apply (simp add: add_commute_int[of "-.m" n]) -apply(rule nat_leq_cases[OF n m]) - apply(rule nat_leq_cases[of "n * p" "m * p"]) - apply (auto simp: adiff_mult_distrib_nat int_mult_def dest: nat_leq_antisym) -done - -lemma distrib_right_int: - assumes m: "m \ Int" and n: "n \ Int" and p: "p \ Int" - shows "(m + n) * p = (m * p + n * p)" -apply(rule intCases3[OF m n p], - simp_all only: int_mult_def int_add_nn_def int_mult_nn_def addIsNat) - apply(rule add_mult_distrib_right_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) - apply(rule pnp_distrib_right_nat, assumption+) - apply(rule pnn_distrib_right_nat, assumption+) - apply(simp add: add_commute_int, rule pnp_distrib_right_nat, assumption+) - apply(rule npn_distrib_right_nat, assumption+) - apply(simp only: int_add_nn_def multIsNat add_mult_distrib_right_nat) - apply(simp only: add_mult_distrib_right_nat) -done - -text \ Minus sign distributes over multiplication \ - -lemma minus_mult_left_int: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m * n) = -.m * n" -by (rule intCases2[OF m n], simp_all add: int_mult_def) - -lemma minus_mult_right_int: - assumes m: "m \ Int" and n: "n \ Int" - shows "-.(m * n) = m * -.n" -by (rule intCases2[OF m n], simp_all add: int_mult_def) - - -subsection \ Difference of integers \ + +subsection \ Primitive recursive functions over the integers \ text \ - Difference over integers is simply defined as addition of the complement. - Note that this difference, noted @{text "-"}, is different from the - difference over natural numbers, noted @{text "--"}, even for two natural - numbers, because the latter cuts off at $0$. + We introduce a similar construction for functions defined + over the integers. We have two step functions for positive and + negative integers and allow both of them to refer to the value + of the function at the predecessor and its conjugate. \ -definition diff (infixl "-" 65) -where int_diff_def: "\m \ Int; n \ Int\ \ m - n = m + -.n" - -lemma diffIsInt [simp]: (* -- Closure *) - "\m \ Int; n \ Int\ \ m - n \ Int" -by (simp add: int_diff_def) - -lemma diff_neg_is_add [simp]: "\m \ Int; n \ Int\ \ m - -.n = m + n" -by (simp add: int_diff_def) - -lemma diff_0_right_int [simp]: "m \ Int \ m - 0 = m" -by (simp add: int_diff_def) +definition primrec_int_upto where + "primrec_int_upto(base, pstep, nstep, f, n) \ + isAFcn(f) \ DOMAIN f = Int + \ f[0] = base + \ (\m \ Nat : succ[m] \ upto(n) \ + f[succ[m]] = pstep(m, f[m], f[-m]) + \ f[-succ[m]] = nstep(m, f[m], f[-m]))" + +lemma primrec_int_upto_deterministic: + assumes n: "n \ Nat" and m: "m \ upto(n)" + and f: "primrec_int_upto(base, pstep, nstep, f, n)" + and g: "primrec_int_upto(base, pstep, nstep, g, m)" + and k: "k \ upto(m)" + shows "f[k] = g[k]" "f[-k] = g[-k]" +proof - + from m uptoNat have mNat: "m \ Nat" by auto + hence "\k \ upto(m): f[k] = g[k] \ f[-k] = g[-k]" + proof (rule uptoInduct) + from f g show "f[0] = g[0] \ f[-0] = g[-0]" + unfolding primrec_int_upto_def by simp + next + fix k + assume kNat: "k \ Nat" + and ih: "k \ upto(m) \ f[k] = g[k] \ f[-k] = g[-k]" + and sk: "succ[k] \ upto(m)" + from sk kNat g + have "g[succ[k]] = pstep(k, g[k], g[-k]) + \ g[-succ[k]] = nstep(k, g[k], g[-k])" + unfolding primrec_int_upto_def by simp + moreover + from sk m n have "succ[k] \ upto(n)" + by (rule uptoTrans) + with kNat f + have "f[succ[k]] = pstep(k, f[k], f[-k]) + \ f[-succ[k]] = nstep(k, f[k], f[-k])" + unfolding primrec_int_upto_def by simp + moreover + from sk kNat mNat have "k \ upto(m)" + by (rule uptoPred) + ultimately + show "f[succ[k]] = g[succ[k]] \ f[-succ[k]] = g[-succ[k]]" + using ih by simp + qed + with k show "f[k] = g[k]" "f[-k] = g[-k]" by (blast+) +qed -lemma diff_0_left_int [simp]: "n \ Int \ 0 - n = -.n" -by (simp add: int_diff_def) +lemma primrec_int_upto_exists: + "\n \ Nat : \f : primrec_int_upto(base, pstep, nstep, f, n)" +proof (rule natInduct) + have "primrec_int_upto(base, pstep, nstep, [n \ Int \ base], 0)" + unfolding primrec_int_upto_def by (auto simp: uptoZero) + thus "\f : primrec_int_upto(base, pstep, nstep, f, 0)" .. +next + fix n + assume n: "n \ Nat" + and ih: "\f : primrec_int_upto(base, pstep, nstep, f, n)" + from n have ssn: "succ[succ[n]] \ n" + by induct (simp+) + from ih obtain f where f: "primrec_int_upto(base, pstep, nstep, f, n)" .. + define g where "g \ [m \ Int \ + IF m = succ[n] THEN pstep(n, f[n], f[-n]) + ELSE IF m = -succ[n] THEN nstep(n, f[n], f[-n]) + ELSE f[m]]" + have "isAFcn(g) \ DOMAIN g = Int" + by (simp add: g_def) + moreover + from f n have "g[0] = base" + by (simp add: primrec_int_upto_def g_def) + moreover + { + fix m + assume 1: "m \ Nat" and 2: "succ[m] \ upto(n)" + have m1: "m \ succ[n]" + proof + assume 3: "m = succ[n]" + from n have "n \ upto(succ[succ[n]])" + by (simp add: uptoSucc uptoRefl) + with 2 3 have "succ[succ[n]] = n" + by (auto dest: uptoAntisym) + with ssn show "FALSE" .. + qed + from 1 n have m2: "m \ -succ[n]" by auto + from m1 m2 f n 1 2 + have "g[succ[m]] = pstep(m, g[m], g[-m]) + \ g[-succ[m]] = nstep(m, g[m], g[-m])" + by (auto simp: primrec_int_upto_def g_def) + } + moreover + from f n have "g[succ[n]] = pstep(n, g[n], g[-n])" + by (simp add: g_def) + moreover + from f n have "g[-succ[n]] = nstep(n, g[n], g[-n])" + by (simp add: g_def) + ultimately have "primrec_int_upto(base, pstep, nstep, g, succ[n])" + using n by (auto simp: primrec_int_upto_def uptoSucc) + thus "\f : primrec_int_upto(base, pstep, nstep, f, succ[n])" .. +qed -lemma diff_self_eq_0_int [simp]: "m \ Int \ m - m = 0" -by (simp add: int_diff_def) +lemma primrec_int: + "\f : isAFcn(f) \ DOMAIN f = Int + \ f[0] = base + \ (\n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]))" +proof - + from primrec_int_upto_exists[THEN fcnConstruct] obtain F where + F: "isAFcn(F) \ DOMAIN F = Nat + \ (\n\Nat : primrec_int_upto(base, pstep, nstep, F[n], n))" + by blast + define g where "g \ [i \ Int \ IF i \ Nat THEN F[i][i] ELSE F[-i][i]]" + have "isAFcn(g)" "DOMAIN g = Int" + unfolding g_def by (simp+) + moreover + from F have "g[0] = base" + by (simp add: primrec_int_upto_def g_def) + moreover + have "\n \ Nat : g[succ[n]] = pstep(n, g[n], g[-n]) + \ g[-succ[n]] = nstep(n, g[n], g[-n])" + (is "\n \ Nat : ?P(n)") + proof + fix n + assume n: "n \ Nat" + hence 1: "n \ upto(n)" "n \ upto(succ[n])" "succ[n] \ upto(succ[n])" + by (simp add: uptoSucc uptoRefl)+ + moreover + from F n have 2: "primrec_int_upto(base, pstep, nstep, F[n], n)" + and 3: "primrec_int_upto(base, pstep, nstep, F[succ[n]], succ[n])" + by auto + with 1 have 4: "F[succ[n]][n] = F[n][n]" "F[succ[n]][-n] = F[n][-n]" + using n by (auto intro: primrec_int_upto_deterministic[where n = "succ[n]"]) + moreover + from n have "g[-n] = F[n][-n]" + by (auto simp add: g_def) + ultimately show "?P(n)" + using 3 n by (auto simp: g_def primrec_int_upto_def) + qed + ultimately + show ?thesis by blast +qed -lemma neg_diff_is_diff [simp]: "\m \ Int; n \ Int\ \ -.(m - n) = n - m" -by (simp add: int_diff_def add_commute_int) +lemma bprimrec_int: + assumes base: "base \ S" + and pstep: "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and nstep: "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + shows "\f \ [Int \ S] : + f[0] = base + \ (\n \ Nat: f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat: f[-succ[n]] = nstep(n, f[n], f[-n]))" +proof - + from primrec_int[of base pstep nstep] obtain f where + 1: "isAFcn(f)" and 2: "DOMAIN f = Int" + and 3: "f[0] = base" + and 4: "\n\Nat : f[succ[n]] = pstep(n, f[n], f[-n])" + and 5: "\n\Nat : f[-succ[n]] = nstep(n, f[n], f[-n])" + by blast + from 3 4 5 base pstep nstep have "\n\Int : f[n] \ S" + by (intro intInduct) auto + with 1 2 3 4 5 show ?thesis + by blast +qed -lemma diff_nat_is_add_neg: "\m \ Nat; n \ Nat\ \ m - n = m + -.n" -by (simp add: int_diff_def) +definition int_primrec where + "int_primrec(S, base, pstep, nstep) \ + CHOOSE f \ [Int \ S] : f[0] = base + \ (\n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n])) + \ (\n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]))" + +lemma int_primrecE: + assumes f: "f = int_primrec(S, base, pstep, nstep)" (is "f = ?g") + and base: "base \ S" + and pstep: "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and nstep: "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + and maj: "\ f \ [Int \ S]; f[0] = base; + \n \ Nat : f[succ[n]] = pstep(n, f[n], f[-n]); + \n \ Nat : f[-succ[n]] = nstep(n, f[n], f[-n]) \ \ P" + shows "P" +proof - + from base pstep nstep + have "\g \ [Int \ S] : g[0] = base + \ (\n \ Nat: g[succ[n]] = pstep(n, g[n], g[-n])) + \ (\n \ Nat: g[-succ[n]] = nstep(n, g[n], g[-n]))" + by (rule bprimrec_int) + hence "?g \ [Int \ S] \ ?g[0] = base + \ (\n \ Nat: ?g[succ[n]] = pstep(n, ?g[n], ?g[-n])) + \ (\n \ Nat: ?g[-succ[n]] = nstep(n, ?g[n], ?g[-n]))" + unfolding int_primrec_def by (rule bChooseI2) auto + with f maj show ?thesis by blast +qed -(** -lemma diff_neg_nat_is_add [simp]: "\m \ Nat; n \ Nat\ \ m - -.n = m + n" -by simp -**) +lemma int_primrecType: + assumes "base \ S" + and "\n \ Nat : \x,y \ S : pstep(n,x,y) \ S" + and "\n \ Nat : \x,y \ S : nstep(n,x,y) \ S" + shows "int_primrec(S, base, pstep, nstep) \ [Int \ S]" +proof - + define f where "f \ int_primrec(S, base, pstep, nstep)" + from f_def[THEN meta_to_obj_eq] assms show ?thesis + by (rule int_primrecE) (simp add: f_def) +qed end diff --git a/isabelle/README.html b/isabelle/README.html index ab8e434b..77e5d740 100644 --- a/isabelle/README.html +++ b/isabelle/README.html @@ -9,9 +9,9 @@ -

TLA+: Lamport's Temporal Logic of Actions specification language

+

TLA+: Lamport's specification language based on the Temporal Logic of Actions

-

TLA +

TLA is a linear-time temporal logic introduced by Leslie Lamport in The Temporal Logic of Actions (ACM TOPLAS 16(3), 1994, 872-923). Unlike other temporal logics, both systems and properties @@ -20,10 +20,14 @@

TLA+: Lamport's Temporal Logic of Actions specification language

structural relations such as refinement, parallel composition, and hiding.

-

TLA+ is a language for specifying and verifying concurrent and distributed systems. It is based on a variant of Zermelo-Fraenkel set theory for describing the data structures manipulated by the algorithms to be verified, and on TLA for describing their dynamic behavior. TLA+ has been applied to numerous case studies. +

TLA+ is a language for specifying and verifying concurrent and distributed +systems. It is based on a variant of Zermelo-Fraenkel set theory for describing +the data structures manipulated by the algorithms to be verified, and on TLA for +describing their dynamic behavior. TLA+ has been applied to numerous case +studies.

-

This directory formalizes TLA+ in Isabelle (version 2009-1), as follows: +

This directory formalizes TLA+ in Isabelle (version 2021-1), as follows:

  • PredicateLogic theory @@ -44,14 +48,14 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

    The FixedPoints theory develops the Knaster-Tarski theorems for least and greatest fixed points in the subset lattice. This development is used for the construction of - natural numbers; it also serves as a test for the encoding of TLA+ + integers; it also serves as a test for the encoding of TLA+ set theory.
  • Functions theory: functions in TLA+ are not defined (e.g., as sets of pairs), - but axiomatized, and in fact, pairs and tuples are later defined as special + but axiomatized, and in fact, pairs and tuples are later defined as functions. Incidentally, this approach helps us to identify functional values, and to automate the reasoning about them. This theory considers only unary functions; functions with multiple arguments are defined @@ -59,19 +63,32 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

  • - Peano theory: Peano's axioms for natural numbers. We prove - the existence of a structure satisfying these axioms and derive the - set of natural numbers. + Integers theory: the structure of integer numbers is introduced + using an extended version of Peano's axioms. We prove the existence of a + structure satisfying these axioms and define the integers as some such + structure (using CHOOSE expressions), together with a subet of + natural numbers. We define the successor and predecessor functions over + integers and prove induction rules for natural numbers and integers. + We also justify the definition of functions over integers by primitive + recursion.
  • - NatOrderings theory: introduces the order <= over - natural numbers. It also defines a < b (for arbitrary - values) as a <= b & a # b. Proves many lemmas over - these orders. Also defines intervals m .. n for natural - numbers m and n. + IntegerArithmetic theory: introduces addition (including + subtraction) and multiplication over integers, as well as the order + <= over integers and a < b (for arbitrary + values) as a <= b & a # b. Proves many lemmas for + integer arithmetic. Also defines intervals m .. n for + integers m and n.
  • +
  • + IntegerDivision theory: introduces integer division and + modulus, as well as the divisibility relation. In accordance with + the definition in the standard library of TLA+, div and + % (modulus) are defined only if the second argument is a + positive integer. +
  • The Tuples theory defines tuples and relations in TLA+. Tuples are functions whose domains are intervals of the form @@ -104,18 +121,6 @@

    TLA+: Lamport's Temporal Logic of Actions specification language

    in theory Tuples.
  • -
  • - NatArith theory: addition, subtraction, and multiplication - over natural numbers. The subtraction (noted --) is defined - such that the result is cut off at 0: this is different from - the standard difference over integers. -
  • - -
  • - NatDivision theory: divisibility relation, division and - modulus for natural numbers. -
  • -

@@ -142,7 +147,5 @@

TLA+: Lamport's Temporal Logic of Actions specification language

Stephan Merz
- -Last modified: Mon Jun 8 11:15:49 CET 2009 - +Last modified: Tue Mar 15 15:05:26 CET 2022 diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index 42152c84..da52b7a4 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -19,23 +19,23 @@ text \ \ definition Nibble -(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) +(* where "Nibble \ {0,1,2,3,4,5,6,7,8,9,10,11,12,13,14,15}" *) where "Nibble \ {0, 1, - Succ[1], - Succ[Succ[1]], - Succ[Succ[Succ[1]]], - Succ[Succ[Succ[Succ[1]]]], - Succ[Succ[Succ[Succ[Succ[1]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]], - Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[Succ[1]]]]]]]]]]]]]]}" + succ[1], + succ[succ[1]], + succ[succ[succ[1]]], + succ[succ[succ[succ[1]]]], + succ[succ[succ[succ[succ[1]]]]], + succ[succ[succ[succ[succ[succ[1]]]]]], + succ[succ[succ[succ[succ[succ[succ[1]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]]], + succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[succ[1]]]]]]]]]]]]]]}" definition char (* -- @{text char} is intended to be applied to nibbles *) where "char(a,b) \ \a,b\" @@ -80,8 +80,8 @@ parse_ast_translation \ (* convert an ML integer to a nibble *) fun mkNibble n = if n = 0 - then Ast.Constant "Peano.zero" - else Ast.Appl [Ast.Constant "Functions.fapply", Ast.Constant "Peano.Succ", mkNibble (n-1)]; + then Ast.Constant "Integers.zero" + else Ast.Appl [Ast.Constant "Functions.fapply", Ast.Constant "Integers.succ", mkNibble (n-1)]; (* convert an ML character to a TLA+ Char *) fun mkChar c = @@ -139,14 +139,14 @@ oops print_ast_translation \ let (* convert a nibble to an ML integer *) + (* version to be used if 2 .. 15 are definitions, not abbreviations *) fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 | destNibble (Ast.Appl [Ast.Constant @{const_syntax "Functions.fapply"}, - Ast.Constant @{const_syntax "Peano.Succ"}, nb]) + Ast.Constant @{const_syntax "Integers.succ"}, nb]) = (destNibble nb) + 1 | destNibble _ = raise Match - - (* the following version should be used when 2 .. 15 are abbreviations, not definitions *) + (* version to be used when 2 .. 15 are abbreviations, not definitions *) (* fun destNibble (Ast.Constant @{const_syntax "zero"}) = 0 | destNibble (Ast.Constant @{const_syntax "one"}) = 1 @@ -247,11 +247,11 @@ text \ (** Examples ** lemma "(''foo'' :> 1 @@ ''bar'' :> TRUE) \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by auto +by auto (* slow *) lemma "r \ [''bar'' : BOOLEAN, ''foo'' : Nat] \ [r EXCEPT ![''foo''] = 3] \ [''bar'' : BOOLEAN, ''foo'' : Nat]" -by (force simp: two_def three_def) (* "by auto" also works, but is slow *) +by (force simp: two_def three_def) (* even slower, "auto" also works *) lemma "(''a'' :> 1) \ (''b'' :> 1)" by simp @@ -260,7 +260,7 @@ lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1)" by simp lemma "(''a'' :> 1 @@ ''b'' :> 2) \ (''a'' :> 1 @@ ''b'' :> 3)" -by (simp add: three_def) +by (simp add: two_def three_def) lemma "(''a'' :> 1 @@ ''b'' :> 2) = (''b'' :> 2 @@ ''a'' :> 1)" by simp @@ -278,7 +278,7 @@ by (simp add: two_def) end -(* NB: Make sure that the following are proved automatically: +(* NB: Make sure that the following is proved automatically once concatenation is defined: THEOREM Thm1 == "ab" = "a" \o "b" OBVIOUS diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5ecc293f..5aa1bd04 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1,14 +1,14 @@ (* Title: TLA+/Tuples.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2022 Inria and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2021-1 *) section \ Tuples and Relations in \tlaplus{} \ theory Tuples -imports NatOrderings +imports IntegerArithmetic begin text \ @@ -16,10 +16,8 @@ text \ functions whose domains are intervals of the form $1 .. n$, for some natural number $n$, and relations are sets of tuples. In particular, \tlaplus{} distinguishes between a function and its graph, and we have - functions to convert between the two. (This is useful, for example, - when defining functions recursively, as we have a fixed point theorem - on sets but not on functions.) We also introduce standard notions for - binary relations, such as orderings, equivalence relations and so on. + functions to convert between the two. We also introduce standard notions + for binary relations, such as orderings, equivalence relations and so on. \ subsection \ Sequences and Tuples \ @@ -41,96 +39,104 @@ where "Len(s) \ CHOOSE n \ Nat : DOMAIN s = 1 .. n" lemma isASeqIsBool [intro!,simp]: "isBool(isASeq(s))" -by (simp add: isASeq_def) + by (simp add: isASeq_def) lemma boolifyIsASeq [simp]: "boolify(isASeq(s)) = isASeq(s)" -by auto + by auto lemma isASeqI [intro (*,simp*)]: assumes "isAFcn(s)" and "n \ Nat" and "DOMAIN s = 1 .. n" shows "isASeq(s)" -using assms by (auto simp: isASeq_def) + using assms by (auto simp: isASeq_def) lemma SeqIsASeq [elim!]: assumes "s \ Seq(S)" shows "isASeq(s)" -using assms by (auto simp: Seq_def) + using assms by (auto simp: Seq_def) lemma LenI [intro]: - assumes "DOMAIN s = 1 .. n" and "n \ Nat" + assumes dom: "DOMAIN s = 1 .. n" and n: "n \ Nat" shows "Len(s) = n" -proof (unfold Len_def, rule bChooseI2) - from assms show "\x \ Nat : DOMAIN s = 1 .. x" by blast +unfolding Len_def proof (rule bChooseI2) + from assms show "\m \ Nat : DOMAIN s = 1 .. m" + by blast next fix m assume "m \ Nat" and "DOMAIN s = 1 .. m" - with assms show "m = n" by auto + with dom n show "m = n" + by (auto simp: intvlEqual_iff) qed lemma isASeqE [elim]: - assumes "isASeq(s)" - and "\isAFcn(s); DOMAIN s = 1 .. Len(s); Len(s) \ Nat\ \ P" - shows "P" -using assms by (auto simp: isASeq_def dest: LenI) + assumes s: "isASeq(s)" + and p: "\isAFcn(s); DOMAIN s = 1 .. Len(s); Len(s) \ Nat\ \ P" + shows "P" + using assms unfolding isASeq_def by (blast dest: LenI) +\ \The corresponding lemma for @{text "Seq(S)"} will be proved later.\ lemma SeqIsAFcn (*[elim!]*): assumes "isASeq(s)" shows "isAFcn(s)" -using assms by auto + using assms by (rule isASeqE) \ \ @{text "s \ Seq(S) \ isAFcn(s)"} \ lemmas SeqIsAFcn' (*[elim!]*) = SeqIsASeq[THEN SeqIsAFcn] lemma LenInNat [simp]: assumes "isASeq(s)" - shows "Len(s) \ Nat" -using assms by (rule isASeqE) + shows "Len(s) \ Int" "0 \ Len(s)" + using assms by (rule isASeqE, simp)+ -\ \ @{text "s \ Seq(S) \ Len(s) \ Nat"} \ -lemmas LenInNat' [simp] = SeqIsASeq[THEN LenInNat] +lemma LenInNat' [simp]: + assumes "s \ Seq(S)" + shows "Len(s) \ Int" "0 \ Len(s)" + using assms[THEN SeqIsASeq] by simp+ lemma DomainSeqLen [simp]: assumes "isASeq(s)" shows "DOMAIN s = 1 .. Len(s)" -using assms by auto + using assms by auto \ \ @{text "s \ Seq(S) \ DOMAIN s = 1 .. Len(s)"} \ -lemmas DomainSeqLen' (*[simp,elim!]*) = SeqIsASeq[THEN DomainSeqLen] +lemmas DomainSeqLen' [simp] = SeqIsASeq[THEN DomainSeqLen] lemma seqEqualI: assumes "isASeq(s)" and "isASeq(t)" and "Len(s) = Len(t)" and "\k \ 1 .. Len(t) : s[k] = t[k]" shows "s = t" -using assms by (intro fcnEqual[of s t], auto) + using assms by (intro fcnEqual[of s t]) (simp add: SeqIsAFcn)+ lemma seqEqualE: assumes "isASeq(s)" and "isASeq(t)" and "s=t" and "\ Len(s) = Len(t); \k \ 1 .. Len(t) : s[k] = t[k] \ \ P" shows "P" -using assms by auto + using assms by auto lemma seqEqualIff: assumes "isASeq(s)" and "isASeq(t)" shows "(s = t) = (Len(s) = Len(t) \ (\k \ 1 .. Len(t) : s[k] = t[k]))" -by (auto elim: seqEqualI[OF assms] seqEqualE[OF assms]) + by (auto elim: seqEqualI[OF assms] seqEqualE[OF assms]) -lemma SeqI [intro!]: - assumes "isASeq(s)" and "\k. k \ 1 .. Len(s) \ s[k] \ S" +lemma SeqI [intro]: + assumes s: "isASeq(s)" and "\k. k \ 1 .. Len(s) \ s[k] \ S" shows "s \ Seq(S)" -using assms by (auto simp: Seq_def) +proof - + from assms have "s \ [1 .. Len(s) \ S]" by auto + with s show ?thesis unfolding Seq_def by auto +qed lemma SeqI': \ \ closer to the definition but probably less useful \ assumes "s \ [1 .. n \ S]" and "n \ Nat" shows "s \ Seq(S)" -using assms by (auto simp: Seq_def) + using assms by (auto simp: Seq_def) lemma SeqE [elim]: assumes s: "s \ Seq(S)" and p: "\s \ [1 .. Len(s) \ S]; Len(s) \ Nat\ \ P" shows "P" proof (rule p) - from s show "Len(s) \ Nat" by (rule LenInNat') + from s show "Len(s) \ Nat" by simp next from s obtain n where "n \ Nat" and "s \ [1 .. n \ S]" by (auto simp: Seq_def) @@ -140,20 +146,20 @@ qed lemma seqFuncSet: assumes "s \ Seq(S)" shows "s \ [1 .. Len(s) \ S]" -using assms by (rule SeqE) + using assms by (rule SeqE) -lemma seqElt [elim!]: - assumes "s \ Seq(S)" and "n \ Nat" and "1 \ n" and "n \ Len(s)" +lemma seqElt [elim]: + assumes "s \ Seq(S)" and "n \ Int" and "0 < n" and "n \ Len(s)" shows "s[n] \ S" -using assms by auto + using assms by auto lemma seqInSeqRange: assumes "isASeq(s)" shows "s \ Seq(Range(s))" -using assms by auto + using assms by auto lemma isASeqInSeq: "isASeq(s) = (\S: s \ Seq(S))" -by (auto elim: seqInSeqRange) + by (auto elim: seqInSeqRange) subsection \ Sequences via @{text emptySeq} and @{text Append} \ @@ -170,15 +176,15 @@ notation (ASCII) emptySeq ("(<< >>)") definition Append :: "[c,c] \ c" -where "Append(s,e) \ [k \ 1 .. Succ[Len(s)] \ IF k = Succ[Len(s)] THEN e ELSE s[k]]" +where "Append(s,e) \ [k \ 1 .. succ[Len(s)] \ IF k = succ[Len(s)] THEN e ELSE s[k]]" -lemma emptySeqIsASeq [simp,intro!]: "isASeq(\\)" -by (auto simp: emptySeq_def isASeq_def) +lemma emptySeqIsASeq [iff]: "isASeq(\\)" + unfolding emptySeq_def by force \ \ @{text "isAFcn(\\)"} \ -lemmas emptySeqIsAFcn [simp,intro!] = emptySeqIsASeq[THEN SeqIsAFcn] +lemmas emptySeqIsAFcn [iff] = emptySeqIsASeq[THEN SeqIsAFcn] -lemma lenEmptySeq [simp]: "Len(\\) = 0" +lemma lenEmptySeq [iff]: "Len(\\) = 0" by (auto simp: emptySeq_def) lemma emptySeqInSeq (*[simp,intro!]*): "\\ \ Seq(S)" @@ -187,101 +193,105 @@ by auto lemma SeqNotEmpty [simp]: "(Seq(S) = {}) = FALSE" "({} = Seq(S)) = FALSE" -by auto + by (force intro: emptySeqInSeq)+ -lemma appendIsASeq [simp,intro!]: +lemma appendIsASeq [iff]: assumes s: "isASeq(s)" shows "isASeq(Append(s,e))" -using s unfolding Append_def -by (rule isASeqE, intro isASeqI, auto simp del: natIntervalSucc) + using s unfolding Append_def + by (rule isASeqE, intro isASeqI, auto) + +\ \@{text "isASeq(s) \ isAFcn(Append(s,e))"}\ +lemmas appendIsAFcn [iff] = appendIsASeq[THEN SeqIsAFcn] -\ \ @{text "isASeq(s) \ isAFcn(Append(s,e))"} \ -lemmas appendIsAFcn [simp,intro!] = appendIsASeq[THEN SeqIsAFcn] +\ \@{text "s \ Seq(S) \ isAFcn(Append(s,e))"}\ +lemmas appendIsAFcn' [iff] = SeqIsASeq[THEN appendIsAFcn] lemma domainEmptySeq [simp]: "DOMAIN \\ = {}" -by (simp add: emptySeq_def) + by (simp add: emptySeq_def) -lemma domainAppend [simp]: "DOMAIN Append(s,e) = 1 .. Succ[Len(s)]" -by (simp add: Append_def) +lemma domainAppend [simp]: "DOMAIN Append(s,e) = 1 .. succ[Len(s)]" + by (simp add: Append_def) -lemma isEmptySeq [intro!]: +lemma isEmptySeq [intro]: "\isAFcn(f); DOMAIN f = {}\ \ f = \\" "\isAFcn(f); DOMAIN f = {}\ \ \\ = f" -by auto + by auto \ \ immediate consequence of @{text isEmptySeq} \ lemma emptySeqEmptyFcn: "\\ = [x \ {} \ y]" -by auto + by auto \ \ The reverse equation could be a useful rewrite rule (it is applied by TLC) \ lemmas emptyFcnEmptySeq = sym[OF emptySeqEmptyFcn] lemma emptyDomainIsEmptySeq [simp]: "(f \ [{} \ S]) = (f = \\)" -by auto + by auto lemma seqLenZeroIsEmpty (*used to be [simp], but emptySeqIff seems more useful*): assumes "isASeq(s)" shows "(Len(s) = 0) = (s = \\)" -using assms by auto + using assms by auto lemma emptySeqIff [simp]: assumes "isAFcn(s)" shows "(s = \\) = (DOMAIN s = {} \ Len(s) = 0)" -using assms by auto + using assms by auto lemma emptySeqIff' [simp]: assumes "isAFcn(s)" shows "(\\ = s) = (DOMAIN s = {} \ Len(s) = 0)" -using assms by auto + using assms by auto lemma lenAppend [simp]: assumes "isASeq(s)" - shows "Len(Append(s,e)) = Succ[Len(s)]" -using assms by (intro LenI, auto simp: Append_def) + shows "Len(Append(s,e)) = succ[Len(s)]" + using assms unfolding Append_def by (intro LenI) force+ -\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = Succ[Len(s)]"} \ +\ \ @{text "s \ Seq(S) \ Len(Append(s,e)) = succ[Len(s)]"} \ lemmas lenAppend' [simp] = SeqIsASeq[THEN lenAppend] +lemma appendNotEmpty [simp]: + assumes "isASeq(s)" + shows "(Append(s,e) = \\) = FALSE" + using assms by auto + lemma appendElt [simp]: - assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Succ[Len(s)]" - shows "Append(s,e)[k] = (IF k = Succ[Len(s)] THEN e ELSE s[k])" -using assms by (auto simp: Append_def) + assumes "isASeq(s)" and "k \ Int" and "0 < k" and "k \ succ[Len(s)]" + shows "Append(s,e)[k] = (IF k = succ[Len(s)] THEN e ELSE s[k])" + using assms by (auto simp: Append_def) lemmas appendElt' [simp] = SeqIsASeq[THEN appendElt] -lemma appendElt1 (*[simp]*): - assumes "isASeq(s)" and "k \ Nat" and "0 < k" and "k \ Len(s)" +lemma appendElt1 [simp]: + assumes "isASeq(s)" and "k \ Int" and "0 < k" and "k \ Len(s)" shows "Append(s,e)[k] = s[k]" -using assms by (auto simp: Append_def) + using assms unfolding Append_def by (force simp: int_leq_succI) -lemmas appendElt1' (*[simp]*) = SeqIsASeq[THEN appendElt1] +lemmas appendElt1' [simp] = SeqIsASeq[THEN appendElt1] -lemma appendElt2 (*[simp]*): +lemma appendElt2 [simp]: assumes "isASeq(s)" - shows "Append(s,e)[Succ[Len(s)]] = e" -using assms by (auto simp: Append_def) + shows "Append(s,e)[succ[Len(s)]] = e" + using assms unfolding Append_def by force -lemmas appendElt2' (*[simp]*) = SeqIsASeq[THEN appendElt2] +lemmas appendElt2' [simp] = SeqIsASeq[THEN appendElt2] -lemma isAppend [intro!]: - assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. Succ[Len(s)]" and s: "isASeq(s)" - and elt1: "\n \ 1 .. Len(s) : f[n] = s[n]" and elt2: "f[Succ[Len(s)]] = e" +lemma isAppend [intro]: + assumes f: "isAFcn(f)" and dom: "DOMAIN f = 1 .. succ[Len(s)]" and s: "isASeq(s)" + and elt1: "\n \ 1 .. Len(s) : f[n] = s[n]" and elt2: "f[succ[Len(s)]] = e" shows "f = Append(s,e)" proof (rule fcnEqual[OF f]) - from s show "isAFcn(Append(s,e))" by simp -next - from dom show "DOMAIN f = DOMAIN Append(s,e)" by simp -next from s elt1 elt2 show "\x \ DOMAIN Append(s, e) : f[x] = Append(s, e)[x]" - by (auto simp: Append_def) -qed + unfolding Append_def by force +qed (simp add: s dom)+ -lemmas isAppend' [intro!] = isAppend[symmetric] +lemmas isAppend' [intro] = isAppend[symmetric] lemma appendInSeq [simp]: assumes s: "s \ Seq(S)" and e: "e \ S" shows "Append(s,e) \ Seq(S)" -using assms by (force simp: nat_leq_Succ) + using assms by (intro SeqI) auto lemma appendD1: assumes s: "isASeq(s)" and t: "isASeq(t)" and app: "Append(s,e) = Append(t,f)" @@ -292,15 +302,15 @@ proof - from s have 1: "isASeq(?s1)" by simp from t have 2: "isASeq(?t1)" by simp from 1 2 app have len: "Len(?s1) = Len(?t1)" and elt: "\k \ 1 .. Len(?t1) : ?s1[k] = ?t1[k]" - by (blast elim: seqEqualE)+ - from s t len have ls: "Len(s) = Len(t)" by simp + by (auto elim: seqEqualE) + from s t len have ls: "Len(s) = Len(t)" by auto thus ?thesis - proof (rule seqEqualI[OF s t], auto) + proof (rule seqEqualI[OF s t], clarify) fix k assume k: "k \ 1 .. Len(t)" - with s ls have "s[k] = ?s1[k]" by (intro sym[OF appendElt1], auto) - also from k elt t have "... = ?t1[k]" by auto - also from t k have "... = t[k]" by (intro appendElt1, auto) + with s ls have "s[k] = ?s1[k]" by (intro sym[OF appendElt1]) auto + also from k elt t have "\ = ?t1[k]" by auto + also from t k have "... = t[k]" by (intro appendElt1) auto finally show "s[k] = t[k]" . qed qed @@ -314,15 +324,17 @@ proof - from s have 1: "isASeq(?s1)" by simp from t have 2: "isASeq(?t1)" by simp from 1 2 app have "Len(?s1) = Len(?t1)" and "\k \ 1 .. Len(?t1) : ?s1[k] = ?t1[k]" - by (blast elim: seqEqualE)+ - with s t have "?s1[Len(?s1)] = ?t1[Len(?t1)]" by auto + by (auto elim: seqEqualE)+ + moreover + from s t app have "Len(?s1) \ 1 .. Len(?t1)" by auto + ultimately have "?s1[Len(?s1)] = ?t1[Len(?t1)]" by auto with s t show ?thesis by simp qed lemma appendEqualIff [simp]: assumes s: "isASeq(s)" and t: "isASeq(t)" shows "(Append(s,e) = Append(t,f)) = (s = t \ e = f)" -using appendD1[OF s t] appendD2[OF s t] by auto + using appendD1[OF s t] appendD2[OF s t] by auto text \ The following lemma gives a possible alternative definition of @@ -331,18 +343,35 @@ text \ lemma appendExtend: assumes "isASeq(s)" - shows "Append(s,e) = s @@ (Succ[Len(s)] :> e)" -using assms by (intro isAppend') auto + shows "Append(s,e) = s @@ (succ[Len(s)] :> e)" (is "_ = ?rhs") + using assms by (intro isAppend') force+ lemma imageEmptySeq [simp]: "Image(\\, A) = {}" -by (simp add: emptySeq_def) + by (simp add: emptySeq_def) lemma imageAppend [simp]: assumes s: "isASeq(s)" shows "Image(Append(s,e), A) = - (IF Succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" -unfolding appendExtend[OF s] -using assms by auto (force+) + (IF succ[Len(s)] \ A THEN addElt(e, Image(s,A)) ELSE Image(s,A))" +proof - + { + fix i + assume "i \ A" "i \ Int" "0 < i" "i \ Len(s)" + with s have "\j \ A : j \ Int \ 1 \ j \ j \ Len(s) \ + s[i] = (s @@ succ[Len(s)] :> e)[j]" + by auto + } + moreover + { + fix i + assume "i \ A" "i \ Int" "0 < i" "i \ Len(s)" + with s have "\j \ A : j \ Int \ 1 \ j \ j \ Len(s) \ + s[i] = s[j]" + by auto + } + ultimately show ?thesis + unfolding appendExtend[OF s] using s by auto +qed text \ Inductive reasoning about sequences, based on @{term "\\"} and @@ -357,21 +386,21 @@ lemma seqInduct [case_names empty append, induct set: Seq]: proof - have "\n \ Nat : \s \ [1 .. n \ S] : P(s)" (is "\n \ Nat : ?A(n)") proof (rule natInduct) - from base show "?A(0)" by (auto del: funcSetE') + from base show "?A(0)" by auto next fix n assume n: "n \ Nat" and ih: "?A(n)" - show "?A(Succ[n])" + show "?A(succ[n])" proof fix sn - assume sn: "sn \ [1 .. Succ[n] \ S]" + assume sn: "sn \ [1 .. succ[n] \ S]" define so where "so \ [k \ 1 .. n \ sn[k]]" - define lst where "lst \ sn[Succ[n]]" + define lst where "lst \ sn[succ[n]]" have 1: "sn = Append(so, lst)" proof from sn show "isAFcn(sn)" by simp next - from sn n show "DOMAIN sn = 1 .. Succ[Len(so)]" + from sn n show "DOMAIN sn = 1 .. succ[Len(so)]" by (simp add: so_def LenI) next from n show "isASeq(so)" by (force simp: so_def) @@ -379,11 +408,11 @@ proof - from n show "\k \ 1 .. Len(so) : sn[k] = so[k]" by (auto simp: so_def LenI) next - from n show "sn[Succ[Len(so)]] = lst" + from n show "sn[succ[Len(so)]] = lst" by (simp add: so_def lst_def LenI) qed from sn n have 2: "so \ [1 .. n \ S]" - by (force simp: so_def) + unfolding so_def by force with ih have 3: "P(so)" .. from 2 n have 4: "so \ Seq(S)" unfolding Seq_def by auto @@ -430,7 +459,8 @@ translations (*** examples *** -lemma "Len(\a,b,c\) = 3" by (simp add: three_def two_def) +lemma "Len(\a,b,c\) = 3" + by (simp add: two_def three_def) lemma "Append(\a,b\, c) = \a,b,c\" .. @@ -505,7 +535,7 @@ lemmas \ \ establish set equality for sets of enumerated function setEqualI [where A = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] setEqualI [where B = "EnumFuncSet(doms, rngs)" for doms rngs, intro!] -lemma EnumFuncSetI [intro!,simp]: +lemma EnumFuncSetI [intro!]: assumes 1: "isAFcn(f)" and 2: "DOMAIN f = Range(doms)" and 3: "DOMAIN rngs = DOMAIN doms" and 4: "\i \ DOMAIN doms: f[doms[i]] \ rngs[i]" @@ -598,76 +628,78 @@ print_ast_translation \ \ (*** Examples *** - lemma "(1 :> TRUE) \ [1 : BOOLEAN]" -by auto + by auto lemma "(1 :> TRUE) @@ (5 :> 2) \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE @@ 5 :> 2) = (5 :> 2 @@ 1 :> TRUE)" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (2 :> {}) \ [1 : BOOLEAN, 2: SUBSET Nat, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![1] = FALSE] \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[((1 :> TRUE) @@ (5 :> 2)) EXCEPT ![5] = {}] \ [1 : BOOLEAN, 5 : SUBSET Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "(1 :> TRUE) @@ (5 :> 2) @@ (1 :> {}) \ [1 : BOOLEAN, 5 : Nat]" -by (auto simp: two_def three_def four_def five_def) + by (auto simp: two_def three_def four_def five_def) lemma "[1 : BOOLEAN, 5 : Nat] = [5 : Nat, 1 : BOOLEAN]" -by (auto simp: two_def three_def four_def five_def) + by auto lemma "\f \ [1 : BOOLEAN, 2: Nat]; g \ [2: Nat, 1: BOOLEAN]; f[1] = g[1]; f[2] = g[2]\ \ f = g" -by auto - + by auto ***) subsection \ Set product \ text \ - The cartesian product of two sets $A$ and $B$ is the set of pairs + The Cartesian product of two sets $A$ and $B$ is the set of pairs whose first component is in $A$ and whose second component is in $B$. We generalize the definition of products to an arbitrary number of sets: $Product(\langle A_1,\ldots,A_n \rangle) = A_1 \times\cdots\times A_n$. \ definition Product -where "Product(s) \ { f \ [1 .. Len(s) \ UNION Range(s)] : +where "Product(s) \ { f \ [1 .. Len(s) \ UNION {s[i] : i \ 1 .. Len(s)}] : \i \ 1 .. Len(s) : f[i] \ s[i] }" +lemmas \ \ establish set equality for products \ + setEqualI [where A = "Product(s)" for s, intro!] + setEqualI [where B = "Product(s)" for s, intro!] + lemma inProductI [intro!]: assumes "isASeq(p)" and "isASeq(s)" and "Len(p) = Len(s)" and "\k \ 1 .. Len(s) : p[k] \ s[k]" shows "p \ Product(s)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductIsASeq: assumes "p \ Product(s)" and "isASeq(s)" shows "isASeq(p)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductLen: assumes "p \ Product(s)" and "isASeq(s)" shows "Len(p) = Len(s)" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto lemma inProductE [elim!]: assumes "p \ Product(s)" and "isASeq(s)" and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; \k \ 1 .. Len(s) : p[k] \ s[k] \ \ P" shows "P" -using assms by (auto simp add: Product_def) + using assms unfolding Product_def by auto (*** examples *** lemma "Product(\\) = { \\ }" by auto -lemma "\2,1\ \ Product(\Nat, Nat\)" by auto +lemma "\2,1\ \ Product(\Nat, Nat\)" by (auto simp: two_def) lemma assumes "x \ X" and "y \ Y" and "z \ Z" @@ -708,14 +740,18 @@ definition notation (ASCII) prod (infixr "\\X" 100) +lemmas \ \ establish set equality for binary products \ + setEqualI [where A = "prod(X, Y)" for X Y, intro!] + setEqualI [where B = "prod(X, Y)" for X Y, intro!] + lemma inProd [simp]: "(\a,b\ \ A \ B) = (a \ A \ b \ B)" -by (auto simp add: prod_def) + unfolding prod_def by auto lemma prodProj: assumes p: "p \ A \ B" shows "p = \p[1], p[2]\" -using assms by (auto simp: two_def prod_def) + using assms unfolding prod_def by (auto simp: two_def) lemma inProd': "(p \ A \ B) = (\a \ A : \ b \ B : p = \a,b\)" @@ -730,7 +766,7 @@ qed lemma inProdI [intro]: assumes a: "a \ A" and b: "b \ B" and p: "P(\a,b\)" shows "\p \ A \ B : P(p)" -using assms by (intro bExI[of "\a,b\"], simp+) + using assms by (intro bExI[of "\a,b\"]) simp+ lemma inProdI': assumes a: "a \ A" and b: "b \ B" @@ -740,7 +776,7 @@ proof next show "\a,b\[1] = a" by simp next - show "\a,b\[2] = b" unfolding two_def by simp + show "\a,b\[2] = b" by (simp add: two_def) qed lemma inProdE [elim]: @@ -750,43 +786,41 @@ lemma inProdE [elim]: using assms by (auto simp add: inProd') (*** examples *** -lemma "\2,1\ \ Nat \ Nat" by simp +lemma "\2,1\ \ Nat \ Nat" by (simp add: two_def) ***) lemma prodEmptyIff [simp]: "(A \ B = {}) = ((A = {}) \ (B = {}))" -proof auto - fix a b - assume a: "a \ A" and b: "b \ B" and prod: "A \ B = {}" - from a b have "\a,b\ \ A \ B" by simp - with prod show "FALSE" by blast +proof - + { + fix a b + assume a: "a \ A" and b: "b \ B" and prod: "A \ B = {}" + from a b have "\a,b\ \ A \ B" by simp + with prod have "FALSE" by blast + } + thus ?thesis by auto qed lemma prodEmptyIff' [simp]: "({} = A \ B) = ((A = {}) \ (B = {}))" -proof auto - fix a b - assume a: "a \ A" and b: "b \ B" and prod: "{} = A \ B" - from a b have "\a,b\ \ A \ B" by simp - with prod show "FALSE" by blast -qed + by (auto dest: sym) lemma pairProj_in_rel (*[simp]*): assumes r: "r \ A \ B" and p: "p \ r" shows "\p[1],p[2]\ \ r" -using p prodProj[OF rev_subsetD[OF p r], symmetric] by simp + using p prodProj[OF rev_subsetD[OF p r], symmetric] by simp lemma pairProj_in_prod (*[simp]*): assumes r: "r \ A \ B" and p: "p \ r" shows "\p[1],p[2]\ \ A \ B" -using subsetD[OF r p] prodProj[OF rev_subsetD[OF p r], symmetric] by simp + using subsetD[OF r p] prodProj[OF rev_subsetD[OF p r], symmetric] by simp -lemma relProj1 [elim]: (** consider [elim!] ?? **) +lemma relProj1 [elim]: assumes "\a,b\ \ r" and "r \ A \ B" shows "a \ A" -using assms by (auto dest: pairProj_in_prod) + using assms by (auto dest: pairProj_in_prod) -lemma relProj2 [elim]: (** consider [elim!] ?? **) +lemma relProj2 [elim]: assumes "\a,b\ \ r" and "r \ A \ B" shows "b \ B" using assms by (auto dest: pairProj_in_prod) @@ -794,13 +828,12 @@ lemma relProj2 [elim]: (** consider [elim!] ?? **) lemma setOfAllPairs_eq_r (*[simp]*): assumes r: "r \ A \ B" shows "{\p[1], p[2]\ : p \ r} = r" -using subsetD[OF r, THEN prodProj] by force + using subsetD[OF r, THEN prodProj] by force lemma subsetsInProd: assumes "A' \ A" and "B' \ B" shows "A' \ B' \ A \ B" -unfolding prod_def Product_def -using assms by auto + unfolding prod_def Product_def using assms by auto subsection \ Syntax for setOfPairs: @{text "{e : \x,y\ \ R}"} \ @@ -815,16 +848,19 @@ syntax (ASCII) translations "{e : \x,y\ \ R}" \ "CONST setOfPairs(R, \x y. e)" +lemmas \ \ establish set equality for sets of pairs \ + setEqualI [where A = "setOfPairs(R, f)" for R f, intro!] + setEqualI [where B = "setOfPairs(R, f)" for R f, intro!] + lemma inSetOfPairsI_ex: assumes "\\x,y\ \ R : a = e(x,y)" shows "a \ { e(x,y) : \x,y\ \ R }" -using assms by (auto simp: setOfPairs_def two_def) + using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsI [intro]: assumes "a = e(x,y)" and "\x,y\ \ R" shows "a \ setOfPairs(R, e)" -unfolding two_def -using assms by (auto simp: setOfPairs_def two_def) + using assms by (auto simp: setOfPairs_def two_def) lemma inSetOfPairsE [elim!]: \ \ the converse is true only if $R$ is a relation \ assumes 1: "z \ setOfPairs(R, e)" @@ -837,14 +873,10 @@ proof - with pR pz show "P" by (intro 3, auto) qed -lemmas setOfPairsEqualI = - setEqualI [where A = "setOfPairs(R,f)" for R f, intro!] - setEqualI [where B = "setOfPairs(R,f)" for R f, intro!] - lemma setOfPairs_triv [simp]: assumes "R \ A \ B" shows "{ \x, y\ : \x, y\ \ R } = R" -using assms by auto + using assms by auto lemma setOfPairs_cong (*[cong]*): assumes 1: "R = S" and @@ -852,7 +884,7 @@ lemma setOfPairs_cong (*[cong]*): 3: "\x y. \x, y\ \ S \ e(x, y) = f(x, y)" shows "{ e(x, y) : \x, y\ \ R } = { f(u, v) : \u, v\ \ S }" -using assms by (force intro!: inSetOfPairsI_ex) + using assms by (force intro!: inSetOfPairsI_ex) lemma setOfPairsEqual: assumes 1: "\x y. \x, y\ \ S \ \\x', y'\ \ T : e(x, y) = f(x', y')" @@ -861,27 +893,27 @@ lemma setOfPairsEqual: and 4: "T \ C \ D" shows "{ e(x, y) : \x, y\ \ S } = { f(x, y) : \x, y\ \ T }" -using assms by (force intro!: inSetOfPairsI_ex) + using assms by (force intro!: inSetOfPairsI_ex) subsection \ Basic notions about binary relations \ -definition rel_domain :: "c => c" +definition rel_domain :: "c \ c" where "rel_domain(r) \ { p[1] : p \ r }" -definition rel_range :: "c => c" +definition rel_range :: "c \ c" where "rel_range(r) \ { p[2] : p \ r }" -definition converse :: "c => c" ("(_^-1)" [1000] 999) +definition converse :: "c \ c" ("(_^-1)" [1000] 999) where "r^-1 \ { \p[2],p[1]\ : p \ r}" \ \ composition of binary relations \ -definition rel_comp :: "[c,c] => c" (infixr "\" 75) +definition rel_comp :: "[c,c] \ c" (infixr "\" 75) where "r \ s \ { p \ rel_domain(s) \ rel_range(r) : \ x, z : p = \x, z\ \ (\ y: \x, y\ \ s \ \y, z\ \ r) }" -definition rel_image :: "[c,c] => c" (infixl "``" 90) +definition rel_image :: "[c,c] \ c" (infixl "``" 90) where "r `` A \ {y \ rel_range(r) : \x \ A: \x,y\ \ r}" definition Id :: "c \ c" \ \ diagonal: identity over a set \ @@ -894,55 +926,55 @@ definition reflexive \ \ reflexivity over a set \ where "reflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyReflexive [simp]: "boolify(reflexive(A,r)) = reflexive(A,r)" -unfolding reflexive_def by simp + unfolding reflexive_def by simp lemma reflexiveIsBool[intro!,simp]: "isBool(reflexive(A,r))" -unfolding isBool_def by (rule boolifyReflexive) + unfolding isBool_def by (rule boolifyReflexive) definition symmetric \ \ symmetric relation \ where "symmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r" lemma boolifySymmetric [simp]: "boolify(symmetric(r)) = symmetric(r)" -unfolding symmetric_def by simp + unfolding symmetric_def by simp lemma symmetricIsBool[intro!,simp]: "isBool(symmetric(r))" -unfolding isBool_def by (rule boolifySymmetric) + unfolding isBool_def by (rule boolifySymmetric) definition antisymmetric \ \ antisymmetric relation \ -where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" + where "antisymmetric(r) \ \x,y: \x,y\ \ r \ \y,x\ \ r \ x = y" lemma boolifyAntisymmetric [simp]: "boolify(antisymmetric(r)) = antisymmetric(r)" -unfolding antisymmetric_def by simp + unfolding antisymmetric_def by simp lemma antisymmetricIsBool[intro!,simp]: "isBool(antisymmetric(r))" -unfolding isBool_def by (rule boolifyAntisymmetric) + unfolding isBool_def by (rule boolifyAntisymmetric) definition transitive \ \ transitivity predicate \ where "transitive(r) \ \x,y,z: \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r" lemma boolifyTransitive [simp]: "boolify(transitive(r)) = transitive(r)" -unfolding transitive_def by simp + unfolding transitive_def by simp lemma transitiveIsBool[intro!,simp]: "isBool(transitive(r))" -unfolding isBool_def by (rule boolifyTransitive) + unfolding isBool_def by (rule boolifyTransitive) definition irreflexive \ \ irreflexivity predicate \ where "irreflexive(A,r) \ \x \ A: \x,x\ \ r" lemma boolifyIrreflexive [simp]: "boolify(irreflexive(A,r)) = irreflexive(A,r)" -unfolding irreflexive_def by simp + unfolding irreflexive_def by simp lemma irreflexiveIsBool[intro!,simp]: "isBool(irreflexive(A,r))" -unfolding isBool_def by (rule boolifyIrreflexive) + unfolding isBool_def by (rule boolifyIrreflexive) definition equivalence :: "[c,c] \ c" \ \ (partial) equivalence relation \ where "equivalence(A,r) \ reflexive(A,r) \ symmetric(r) \ transitive(r)" lemma boolifyEquivalence [simp]: "boolify(equivalence(A,r)) = equivalence(A,r)" -unfolding equivalence_def by simp + unfolding equivalence_def by simp lemma equivalenceIsBool[intro!,simp]: "isBool(equivalence(A,r))" -unfolding isBool_def by (rule boolifyEquivalence) + unfolding isBool_def by (rule boolifyEquivalence) subsubsection \ Domain and Range \ @@ -950,13 +982,12 @@ subsubsection \ Domain and Range \ lemma prod_in_dom_x_ran: assumes "r \ A \ B" and "p \ r" shows "\p[1],p[2]\ \ rel_domain(r) \ rel_range(r)" -unfolding inProd rel_domain_def rel_range_def -using assms by auto + unfolding inProd rel_domain_def rel_range_def using assms by auto lemma in_rel_domainI [iff]: assumes "\x,y\ \ r" shows "x \ rel_domain(r)" -unfolding rel_domain_def using assms by auto + unfolding rel_domain_def using assms by auto lemma in_rel_domainE [elim]: assumes x: "x \ rel_domain(r)" and r: "r \ A \ B" and p: "\y. \x,y\ \ r \ P" @@ -969,15 +1000,15 @@ proof - qed lemma rel_domain (*[simp]*): "r \ A \ B \ rel_domain(r) \ A" -unfolding rel_domain_def using pairProj_in_prod by auto + unfolding rel_domain_def using pairProj_in_prod by auto lemma rel_range (*[simp]*): "r \ A \ B \ rel_range(r) \ B" -unfolding rel_range_def using pairProj_in_prod by auto + unfolding rel_range_def using pairProj_in_prod by auto lemma in_rel_rangeI [iff]: assumes "\x,y\ \ r" shows "y \ rel_range(r)" -unfolding rel_range_def two_def using assms by auto + unfolding rel_range_def using assms by (auto simp: two_def) lemma in_rel_rangeE [elim]: assumes y: "y \ rel_range(r)" and r: "r \ A \ B" and p: "\x. \x,y\ \ r \ P" @@ -990,22 +1021,22 @@ proof - qed lemma dom_in_A (*[simp]*): "rel_domain ({ p \ A \ B : P(p) }) \ A" -by auto + by auto lemma ran_in_B (*[simp]*): "rel_range ({ p \ A \ B : P(p) }) \ B" -by auto + by auto lemma subrel_dom: "r' \ r \ x \ rel_domain(r') \ x \ rel_domain(r)" -unfolding rel_domain_def by auto + unfolding rel_domain_def by auto lemma subrel_ran: "r' \ r \ x \ rel_range(r') \ x \ rel_range(r)" -unfolding rel_range_def by auto + unfolding rel_range_def by auto lemma in_dom_imp_in_A: "r \ A \ B \ x \ rel_domain(r) \ x \ A" -by force + by force lemma in_ran_imp_in_B: "r \ A \ B \ p \ rel_range(r) \ p \ B" -by force + by force subsubsection \ Converse relation \ @@ -1017,21 +1048,21 @@ lemmas converseEqualI = lemma converse_iff [simp]: assumes r: "r \ A \ B" shows "(\a,b\ \ r^-1) = (\b,a\ \ r)" -using r prodProj by (auto simp: converse_def two_def) + using r prodProj by (auto simp: converse_def two_def) lemma converseI [intro!]: shows "\a,b\ \ r \ \b,a\ \ r^-1" -unfolding converse_def two_def by auto + unfolding converse_def by (auto simp: two_def) lemma converseD [sym]: assumes r: "r \ A \ B" shows "\a,b\ \ r^-1 \ \b,a\ \ r" -using converse_iff[OF r] by simp + using converse_iff[OF r] by simp lemma converseSubset: "r \ A \ B \ r^-1 \ B \ A" -unfolding converse_def using pairProj_in_prod by auto + unfolding converse_def using pairProj_in_prod by auto -lemma converseE [elim]: (** consider [elim!] ?? **) +lemma converseE [elim]: assumes yx: "yx \ r^-1" and r: "r \ A \ B" and p: "\x y. yx = \y,x\ \ \x,y\ \ r \ P" shows "P" @@ -1058,10 +1089,10 @@ proof - qed lemma converse_prod [simp]: "(A \ B)^-1 = B \ A" -by auto + by auto lemma converse_empty [simp]: "converse({}) = {}" -by auto + by auto lemma converse_mono_1: assumes r: "r \ A \ B" and s: "s \ A \ B" and sub: "r^-1 \ s^-1" @@ -1077,13 +1108,12 @@ qed lemma converse_mono_2: assumes "r \ A \ B" and "s \ A \ B" and "r \ s" shows "r^-1 \ s^-1" -using assms prodProj by auto + using assms prodProj by auto lemma converse_mono: assumes r:"r \ A \ B" and s:"s \ A \ B" shows "r^-1 \ s^-1 = (r \ s)" -using converse_mono_1[OF r s] converse_mono_2[OF r s] -by blast + using converse_mono_1[OF r s] converse_mono_2[OF r s] by blast (* from HOL *) @@ -1106,7 +1136,7 @@ lemma transitive_converse [simp]: lemma symmetric_iff_converse_eq: assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" -using assms by (auto simp: symmetric_def) + using assms by (auto simp: symmetric_def) subsubsection \ Identity relation over a set \ @@ -1116,48 +1146,47 @@ lemmas idEqualI = setEqualI [where B = "Id(S)" for S, intro!] lemma IdI [iff]: "x \ S \ \x,x\ \ Id(S)" -unfolding Id_def by auto + unfolding Id_def by auto lemma IdI' [intro]: "x \ S \ p = \x,x\ \ p \ Id(S)" -unfolding Id_def by auto + unfolding Id_def by auto lemma IdE [elim!]: "p \ Id(S) \ (\x. x \ S \ p = \x,x\ \ P) \ P" -unfolding Id_def by auto + unfolding Id_def by auto lemma Id_iff: "(\a,b\ \ Id(S)) = (a = b \ a \ S)" -by auto + by auto lemma Id_subset_Prod [simp]: "Id(S) \ S \ S" -unfolding Id_def by auto + unfolding Id_def by auto lemma reflexive_Id: "reflexive(S,Id(S))" -unfolding reflexive_def by auto + unfolding reflexive_def by auto lemma antisymmetric_Id [simp]: "antisymmetric(Id(S))" -unfolding antisymmetric_def by auto + unfolding antisymmetric_def by auto lemma symmetric_Id [simp]: "symmetric(Id(S))" -unfolding symmetric_def by auto + unfolding symmetric_def by auto lemma transitive_Id [simp]: "transitive(Id(S))" -unfolding transitive_def by auto + unfolding transitive_def by auto lemma Id_empty [simp]: "Id({}) = {}" -unfolding Id_def by simp + unfolding Id_def by simp lemma Id_eqI: "a = b \ a \ A \ \a,b\ \ Id(A)" -by simp + by simp lemma converse_Id [simp]: "Id(A)^-1 = Id(A)" -by (auto simp: Id_def two_def) - + by (auto simp: Id_def) lemma dom_Id [simp]: "rel_domain(Id(A)) = A" -unfolding rel_domain_def Id_def by auto + unfolding rel_domain_def Id_def by auto lemma ran_Id [simp]: "rel_range(Id(A)) = A" -unfolding rel_range_def Id_def two_def by auto + unfolding rel_range_def Id_def by (auto simp: two_def) subsubsection \ Composition of relations \ @@ -1169,27 +1198,27 @@ lemmas compEqualI = lemma compI [intro]: assumes r: "r \ B \ C" and s: "s \ A \ B" shows "\ \a,b\ \ s; \b,c\ \ r \ \ \a,c\ \ r \ s" -using assms unfolding rel_comp_def by auto + using assms unfolding rel_comp_def by auto lemma compE [elim!]: assumes "xz \ r \ s" and "r \ B \ C" and "s \ A \ B" shows "(\x y z. xz = \x,z\ \ \x,y\ \ s \ \y,z\ \ r \ P) \ P" -using assms unfolding rel_comp_def by auto + using assms unfolding rel_comp_def by auto lemma compEpair: assumes "\a,c\ \ r \ s" and "r \ B \ C" and s: "s \ A \ B" shows "\\b. \ \a,b\ \ s; \b,c\ \ r \ \ P \ \ P" -using assms by auto + using assms by auto lemma rel_comp_in_prod [iff]: assumes s: "s \ A \ B" and r: "r \ B \ C" shows "r \ s \ A \ C" -using assms by force + using assms by force lemma rel_comp_in_prodE (*[elim]*): assumes "p \ r \ s" and "s \ A \ B" and r: "r \ B \ C" shows "p \ A \ C" -using assms by force + using assms by force lemma converse_comp: assumes r: "r \ B \ C" and s: "s \ A \ B" @@ -1326,66 +1355,66 @@ subsubsection \ Properties of relations \ text \ Reflexivity \ lemma reflI [intro]: "(\x. x \ A \ \x,x\ \ r) \ reflexive(A,r)" -unfolding reflexive_def by blast + unfolding reflexive_def by blast lemma reflexiveD [elim!]: "reflexive(A,r) \ a \ A \ \a,a\ \ r" -unfolding reflexive_def by blast + unfolding reflexive_def by blast lemma reflexive_empty (*[simp]*): "reflexive({}, {})" -by auto + by auto text \ Symmetry \ lemma symmetricI: "\ \x y. \x,y\ \ r \ \y,x\ \ r \ \ symmetric(r)" -unfolding symmetric_def by blast + unfolding symmetric_def by blast lemma symmetricE: "\ symmetric(r); \x,y\ \ r \ \ \y,x\ \ r" -unfolding symmetric_def by blast + unfolding symmetric_def by blast lemma symmetric_Int: "\ symmetric(r); symmetric(s) \ \ symmetric(r \ s)" -by (blast intro: symmetricI dest: symmetricE) + by (blast intro: symmetricI dest: symmetricE) text \ Antisymmetry \ lemma antisymmetricI [intro]: "\ \x y. \ \x,y\ \ r; \y,x\ \ r \ \ x = y \ \ antisymmetric(r)" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisymmetricE [elim]: "\ antisymmetric(r); \x,y\ \ r; \y,x\ \ r \ \ x = y" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisymmetricSubset: "r \ s \ antisymmetric(s) \ antisymmetric(r)" -unfolding antisymmetric_def by blast + unfolding antisymmetric_def by blast lemma antisym_empty (*[simp]*): "antisymmetric({})" -by blast + by blast text \ Transitivity \ lemma transitiveI [intro]: "(\x y z. \x,y\ \ r \ \y,z\ \ r \ \x,z\ \ r) \ transitive(r)" -unfolding transitive_def by blast + unfolding transitive_def by blast lemma transD [elim]: "\ transitive(r); \x,y\ \ r; \y,z\ \ r \ \ \x,z\ \ r" -unfolding transitive_def by blast + unfolding transitive_def by blast lemma trans_Int: "transitive(r) \ transitive(s) \ transitive(r \ s)" -by fast + by fast lemma transitive_iff_comp_subset: "transitive(r) = (r \ r \ r)" -unfolding transitive_def rel_comp_def by (auto elim!: subsetD) + unfolding transitive_def rel_comp_def by (auto elim!: subsetD) text \ Irreflexivity \ lemma irreflexiveI [intro]: "\ \x. x \ A \ \x,x\ \ r \ \ irreflexive(A,r)" -unfolding irreflexive_def by blast + unfolding irreflexive_def by blast lemma irreflexiveE [dest]: "\ irreflexive(A,r); x \ A \ \ \x,x\ \ r" -unfolding irreflexive_def by blast + unfolding irreflexive_def by blast subsubsection \ Equivalence Relations \ @@ -1422,7 +1451,7 @@ text \ First half: ``only if'' part \ lemma sym_trans_comp_subset: assumes "r \ A \ A" and "symmetric(r)" and "transitive(r)" shows "r^-1 \ r \ r" -using assms by (simp add: symmetric_iff_converse_eq transitive_iff_comp_subset) + using assms by (simp add: symmetric_iff_converse_eq transitive_iff_comp_subset) lemma refl_comp_subset: assumes r: "r \ A \ A" and refl: "reflexive(A,r)" diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 4c5d5c2d..dae0051c 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1,8 +1,7 @@ (* Zenon.thy --- Support lemmas for Zenon proofs * Author: Damien Doligez - * Copyright (C) 2008-2011 INRIA and Microsoft Corporation - * Version: Isabelle2011-1 - * Time-stamp: <2011-10-11 17:41:20 merz> + * Copyright (C) 2008-2022 INRIA and Microsoft Corporation + * Version: Isabelle2021-1 *) (* isabelle usedir -b Pure TLA+ *) @@ -1090,7 +1089,7 @@ using zenon_sa_seq by (auto simp add: zenon_sa_def) lemma zenon_sa_diff_0a : "zenon_sa (zenon_sa (s1, e1), e2) ~= zenon_sa (<<>>, f2)" -using zenon_sa_def zenon_sa_seq by auto +using zenon_sa_seq by (auto simp: zenon_sa_def) lemma zenon_sa_diff_0b : "zenon_sa (<<>>, f2) ~= zenon_sa (zenon_sa (s1, e1), e2)" @@ -1117,10 +1116,10 @@ lemma zenon_in_nat_0 : "~(0 \\in Nat) ==> FALSE" by blast lemma zenon_in_nat_1 : "~(1 \\in Nat) ==> FALSE" -by blast +by simp lemma zenon_in_nat_2 : "~(2 \\in Nat) ==> FALSE" -by blast +by (simp add: two_def) lemma zenon_in_nat_succ : "~(Succ[x] \\in Nat) ==> (~(x \\in Nat) ==> FALSE) ==> FALSE" @@ -1175,7 +1174,10 @@ lemma zenon_range_1 : "isASeq (<<>>) & {} = Range (<<>>)" by auto lemma zenon_range_2 : assumes h: "(isASeq (s) & a = Range (s))" shows "(isASeq (Append (s, x)) & addElt (x, a) = Range (Append (s, x)))" -using h by auto +proof - + from assms have "succ[Len(s)] \ DOMAIN Append(s,x)" by auto + with assms show ?thesis by auto +qed lemma zenon_set_rev_1 : "a = {} \\cup c ==> c = a" by auto @@ -1207,7 +1209,8 @@ lemma zenon_all_rec_2 : by auto lemma zenon_tuple_acc_1 : - "isASeq (r) ==> Len (r) = n ==> x = Append (r, x) [Succ[n]]" by auto + "isASeq (r) ==> Len (r) = n ==> x = Append (r, x) [succ[n]]" + by auto lemma zenon_tuple_acc_2 : "isASeq (r) ==> k \\in Nat ==> 0 < k ==> k \ Len (r) ==> @@ -1215,32 +1218,52 @@ lemma zenon_tuple_acc_2 : using appendElt1 by auto definition zenon_ss :: "c \ c" -where "zenon_ss (n) \ IF n \\in Nat THEN Succ[n] ELSE 1" +where "zenon_ss (n) \ IF n \\in Nat THEN succ[n] ELSE 1" -lemma zenon_ss_nat : "zenon_ss(x) \\in Nat" by (auto simp add: zenon_ss_def) +lemma zenon_ss_nat : "zenon_ss(x) \\in Nat" +unfolding zenon_ss_def proof (rule condI) + assume "x \ Nat" thus "succ[x] \ Nat" by simp +next + show "1 \ Nat" by simp +qed -lemma zenon_ss_1 : "Succ[0] = zenon_ss(0)" by (auto simp add: zenon_ss_def) +lemma zenon_ss_1 : "succ[0] = zenon_ss(0)" by (simp add: zenon_ss_def) -lemma zenon_ss_2 : "Succ[zenon_ss(x)] = zenon_ss(zenon_ss(x))" by (auto simp add: zenon_ss_def) +lemma zenon_ss_2 : "succ[zenon_ss(x)] = zenon_ss(zenon_ss(x))" +proof (cases "x \ Nat") + case True + then show ?thesis by (simp add: zenon_ss_def) +next + case False + hence "zenon_ss(x) = 1" by (auto simp: zenon_ss_def) + thus ?thesis by (simp add: zenon_ss_def) +qed lemma zenon_zero_lt : "0 < zenon_ss(x)" - by (simp add: zenon_ss_def, rule disjE [of "x \\in Nat" "x \\notin Nat"], rule excluded_middle, simp+) +proof (cases "x \ Nat") + case True + then show ?thesis by (simp add: zenon_ss_def) +next + case False + hence "zenon_ss(x) = 1" by (auto simp: zenon_ss_def) + then show ?thesis by simp +qed lemma zenon_ss_le_sa_1 : "zenon_ss(0) <= Len (zenon_sa (s, x))" - by (auto simp add: zenon_ss_def zenon_sa_def, rule disjE [of "isASeq (s)" "~isASeq (s)"], rule excluded_middle, simp+) + by (cases "isASeq(s)") (auto simp: zenon_ss_def zenon_sa_def) lemma zenon_ss_le_sa_2 : fixes x y z assumes h0: "zenon_ss (x) <= Len (zenon_sa (s, y))" shows "zenon_ss (zenon_ss (x)) <= Len (zenon_sa (zenon_sa (s, y), z))" proof - - have h1: "Succ [Len (zenon_sa (s, y))] = Len (zenon_sa (zenon_sa (s, y), z))" + have h1: "succ [Len (zenon_sa (s, y))] = Len (zenon_sa (zenon_sa (s, y), z))" using zenon_sa_seq by (auto simp add: zenon_sa_def) have h2: "Len (zenon_sa (s, y)) \\in Nat" - using zenon_sa_seq by (rule LenInNat) - have h3: "Succ [zenon_ss (x)] \ Succ [Len (zenon_sa (s, y))]" - using zenon_ss_nat h2 h0 by (simp only: nat_Succ_leq_Succ) + using zenon_sa_seq by simp + have h3: "succ [zenon_ss (x)] \ succ [Len (zenon_sa (s, y))]" + using zenon_ss_nat h2 h0 by (intro int_succ_leq_succI) simp+ show ?thesis using h3 by (auto simp add: zenon_ss_2 h1) qed @@ -1250,8 +1273,8 @@ by auto lemma zenon_dom_app_2 : assumes h: "isASeq (s) & n = Len (s) & x = 1 .. Len (s)" - shows "isASeq (Append (s, y)) & Succ[n] = Len (Append (s, y)) - & addElt (Succ[n], x) = 1 .. Len (Append (s, y))" (is "?a & ?b & ?c") + shows "isASeq (Append (s, y)) & succ[n] = Len (Append (s, y)) + & addElt (succ[n], x) = 1 .. Len (Append (s, y))" using h by auto (* generic proof rules instantiated for small n *) @@ -1272,7 +1295,7 @@ proof - let ?domsetrev = "{l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" + let ?n1n = "succ[?n0n]" let ?indices = "{?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1335,8 +1358,8 @@ proof - let ?domsetrev = "{l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" let ?indices = "{?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1364,39 +1387,18 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) - qed qed qed @@ -1417,9 +1419,9 @@ proof - let ?domsetrev = "{l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" let ?indices = "{?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1447,53 +1449,23 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) @@ -1520,10 +1492,10 @@ proof - let ?domsetrev = "{l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" let ?indices = "{?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1551,70 +1523,30 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) @@ -1642,11 +1574,11 @@ proof - let ?domsetrev = "{l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" let ?indices = "{?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1674,87 +1606,37 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) @@ -1783,12 +1665,12 @@ proof - let ?domsetrev = "{l6x, l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" - let ?n6n = "Succ[?n5n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" + let ?n6n = "succ[?n5n]" let ?indices = "{?n6n, ?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1816,104 +1698,44 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) next have hn: "l6x = ?doms[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s6x = ?rngs[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n6n]] \\in ?rngs[?n6n]" by (rule subst[OF hs], rule subst[OF hn], rule h6) @@ -1943,13 +1765,13 @@ proof - let ?domsetrev = "{l7x, l6x, l5x, l4x, l3x, l2x, l1x}" let ?rngs = "<>" let ?n0n = "0" - let ?n1n = "Succ[?n0n]" - let ?n2n = "Succ[?n1n]" - let ?n3n = "Succ[?n2n]" - let ?n4n = "Succ[?n3n]" - let ?n5n = "Succ[?n4n]" - let ?n6n = "Succ[?n5n]" - let ?n7n = "Succ[?n6n]" + let ?n1n = "succ[?n0n]" + let ?n2n = "succ[?n1n]" + let ?n3n = "succ[?n2n]" + let ?n4n = "succ[?n3n]" + let ?n5n = "succ[?n4n]" + let ?n6n = "succ[?n5n]" + let ?n7n = "succ[?n6n]" let ?indices = "{?n7n, ?n6n, ?n5n, ?n4n, ?n3n, ?n2n, ?n1n}" have hdomx : "?domsetrev = DOMAIN r" by (rule zenon_set_rev_1, (rule zenon_set_rev_2)+, rule zenon_set_rev_3, @@ -1977,121 +1799,51 @@ proof - proof (rule subst [OF hind], (rule zenon_all_rec_2)+, rule zenon_all_rec_1) have hn: "l1x = ?doms[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s1x = ?rngs[?n1n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n1n]] \\in ?rngs[?n1n]" by (rule subst[OF hs], rule subst[OF hn], rule h1) next have hn: "l2x = ?doms[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s2x = ?rngs[?n2n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n2n]] \\in ?rngs[?n2n]" by (rule subst[OF hs], rule subst[OF hn], rule h2) next have hn: "l3x = ?doms[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s3x = ?rngs[?n3n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n3n]] \\in ?rngs[?n3n]" by (rule subst[OF hs], rule subst[OF hn], rule h3) next have hn: "l4x = ?doms[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s4x = ?rngs[?n4n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n4n]] \\in ?rngs[?n4n]" by (rule subst[OF hs], rule subst[OF hn], rule h4) next have hn: "l5x = ?doms[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s5x = ?rngs[?n5n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n5n]] \\in ?rngs[?n5n]" by (rule subst[OF hs], rule subst[OF hn], rule h5) next have hn: "l6x = ?doms[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s6x = ?rngs[?n6n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n6n]] \\in ?rngs[?n6n]" by (rule subst[OF hs], rule subst[OF hn], rule h6) next have hn: "l7x = ?doms[?n7n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp have hs: "s7x = ?rngs[?n7n]" - by (((rule zenon_tuple_acc_2, safe, simp only: zenon_ss_1 zenon_ss_2, - rule zenon_zero_lt, - simp only: zenon_ss_1 zenon_ss_2 zenon_sa_1 zenon_sa_2, - ((rule zenon_ss_le_sa_2)+)?, rule zenon_ss_le_sa_1 - )+)?, - rule zenon_tuple_acc_1, auto) + by simp show "r[?doms[?n7n]] \\in ?rngs[?n7n]" by (rule subst[OF hs], rule subst[OF hn], rule h7) From 7841cd8dd8324762f13a125af3c56f19398db4a5 Mon Sep 17 00:00:00 2001 From: merz Date: Fri, 18 Mar 2022 19:20:01 +0100 Subject: [PATCH 124/167] fixed Zenon.thy and examples, cleanup of old theory files --- isabelle/NatArith.thy | 1174 --- isabelle/NatDivision.thy | 731 -- isabelle/NatOrderings.thy | 709 -- isabelle/NewIntegers.thy | 12656 -------------------------- isabelle/Peano.thy | 629 -- isabelle/Zenon.thy | 184 +- isabelle/examples/Allocator.thy | 69 +- isabelle/examples/AtomicBakeryG.thy | 910 +- isabelle/examples/ROOT | 18 + isabelle/examples/ROOT.ML | 13 - 10 files changed, 521 insertions(+), 16572 deletions(-) delete mode 100644 isabelle/NatArith.thy delete mode 100644 isabelle/NatDivision.thy delete mode 100644 isabelle/NatOrderings.thy delete mode 100644 isabelle/NewIntegers.thy delete mode 100644 isabelle/Peano.thy create mode 100644 isabelle/examples/ROOT delete mode 100644 isabelle/examples/ROOT.ML diff --git a/isabelle/NatArith.thy b/isabelle/NatArith.thy deleted file mode 100644 index 1e58c382..00000000 --- a/isabelle/NatArith.thy +++ /dev/null @@ -1,1174 +0,0 @@ -(* Title: TLA+/NatArith.thy - Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - -section \ Arithmetic (except division) over natural numbers \ - -theory NatArith -imports NatOrderings -begin - -named_theorems algebra_simps "algebra simplification rules" - -(* -ML \ -structure AlgebraSimps = - Named_Thms(val name = "algebra_simps" - val description = "algebra simplification rules"); -\ - -setup AlgebraSimps.setup -*) - -text \ - The rewrites accumulated in @{text algebra_simps} deal with the - classical algebraic structures of groups, rings and family. They simplify - terms by multiplying everything out (in case of a ring) and bringing sums and - products into a canonical form (by ordered rewriting). As a result these - rewrites decide group and ring equalities but also help with inequalities. - - Of course it also works for fields, but it knows nothing about multiplicative - inverses or division. This should be catered for by @{text field_simps}. -\ - - -subsection \ Addition of natural numbers \ - -definition addnat -where "addnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]])" - -definition arith_add :: "[c,c] \ c" (infixl "+" 65) -where nat_add_def: "\m \ Nat; n \ Nat\ \ (m + n) \ addnat(m)[n]" - -text \ Closure \ - -lemma addnatType: - assumes "m \ Nat" shows "addnat(m) \ [Nat \ Nat]" -using assms unfolding addnat_def by (rule bprimrecType_nat, auto) - -lemma addIsNat [intro!,simp]: - assumes "m \ Nat" and "n \ Nat" shows "m + n \ Nat" -unfolding nat_add_def[OF assms] using assms addnatType by blast - -text \ Base case and Inductive case \ - -lemma addnat_0: - assumes "m \ Nat" shows "addnat(m)[0] = m" -using assms unfolding addnat_def by (rule primrec_natE, auto) - -lemma add_0_nat [simp]: - assumes "m \ Nat" shows "m + 0 = m" -by (simp add: nat_add_def[OF assms] addnat_0[OF assms]) - -lemma addnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" -proof (rule primrec_natE[OF m]) - show "addnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Succ[g[x]]))" - unfolding addnat_def .. -next - assume "\n \ Nat : addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" - with n show "addnat(m)[Succ[n]] = Succ[addnat(m)[n]]" by blast -qed (auto) - -lemma add_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "m + Succ[n] = Succ[m + n]" -using assms by (simp add: nat_add_def addnat_Succ[OF assms]) - -lemma add_0_left_nat [simp]: - assumes n: "n \ Nat" - shows "0 + n = n" -using n by (induct, auto) - -lemma add_Succ_left_nat [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "Succ[m] + n = Succ[m + n]" -using n apply induct -using m by auto - -lemma add_Succ_shift_nat: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "Succ[m] + n = m + Succ[n]" -using assms by simp - -text \ Commutativity \ - -lemma add_commute_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m + n = n + m" -using n apply induct -using assms by auto - -text \ Associativity \ - -lemma add_assoc_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m + (n + p) = (m + n) + p" -using assms by (induct, simp_all) - -text \ Cancellation rules \ - -lemma add_left_cancel_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(m + n = m + p) = (n = p)" -using assms by (induct, simp_all) - -lemma add_right_cancel_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(n + m = p + m) = (n = p)" -using assms by (induct, simp_all) - - -lemma add_left_commute_nat [algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a + (b + c) = b + (a + c)" -using assms by(simp only: add_assoc_nat add_commute_nat) - -(* from HOL/OrderedGroups.thy *) -lemmas add_ac_nat = add_assoc_nat add_commute_nat add_left_commute_nat - -text \ Reasoning about @{text "m + n = 0"}, etc. \ - -lemma add_is_0_nat [iff]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = 0) = (m = 0 \ n = 0)" -using m apply (rule natCases) -using n by (induct, auto) - -lemma add_is_1_nat: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = 1) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" -using m apply (rule natCases) -using n by (induct, auto) - -lemma one_is_add_nat: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 = m + n) = (m = 1 \ n = 0 \ m = 0 \ n = 1)" -using m apply (rule natCases) -using n by (induct, auto)+ - -lemma add_eq_self_zero_nat1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n = m) = (n = 0)" -using n apply (rule natCases) - using m apply simp - using m apply induct apply auto -done - -lemma add_eq_self_zero_nat2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(n + m = m) = (n = 0)" -using assms by (simp add: add_commute_nat) - - -subsection \ Multiplication of natural numbers \ - -definition multnat -where "multnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m)" - -definition mult :: "[c,c] \ c" (infixl "*" 70) -where nat_mult_def: "\m \ Nat; n \ Nat\ \ m * n \ multnat(m)[n]" - -text \ Closure \ - -lemma multnatType: - assumes "m \ Nat" shows "multnat(m) \ [Nat \ Nat]" -unfolding multnat_def by (rule bprimrecType_nat, auto simp: assms) - -lemma multIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m * n \ Nat" -unfolding nat_mult_def[OF assms] using assms multnatType by blast - -text \ Base case and Inductive step \ - -lemma multnat_0: - assumes "m \ Nat" shows "multnat(m)[0] = 0" -unfolding multnat_def by (rule primrec_natE, auto simp: assms) - -lemma mult_0_nat [simp]: \ \ neutral element \ - assumes n: "n \ Nat" shows "n * 0 = 0" -by (simp add: nat_mult_def[OF assms] multnat_0[OF assms]) - - -lemma multnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "multnat(m)[Succ[n]] = multnat(m)[n] + m" -proof (rule primrec_natE) - show "multnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = 0 \ (\x \ Nat : g[Succ[x]] = g[x] + m))" - unfolding multnat_def .. -next - assume "\n \ Nat : multnat(m)[Succ[n]] = multnat(m)[n] + m" - with n show "multnat(m)[Succ[n]] = multnat(m)[n] + m" by blast -qed (auto simp: m) - -lemma mult_Succ_nat [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "m * Succ[n] = m * n + m" -using assms by (simp add: nat_mult_def multnat_Succ[OF assms]) - -lemma mult_0_left_nat [simp]: - assumes n: "n \ Nat" - shows "0 * n = 0" -using n by (induct, simp_all) - -lemma mult_Succ_left_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "Succ[m] * n = m * n + n" -using n apply induct -using m by (simp_all add: add_ac_nat) - -text \ Commutativity \ - -lemma mult_commute_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m * n = n * m" -using assms by (induct, simp_all) - -text \ Distributivity \ - -lemma add_mult_distrib_left_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n + p) = m * n + m * p" -using assms apply induct -proof auto - fix m - assume "m \ Nat" "m * (n + p) = m * n + m * p" - with n p - add_assoc_nat[of "m * n + m * p" n p] - add_assoc_nat[of "m * n" "m * p" n] - add_commute_nat[of "m * p" n] - add_assoc_nat[of "m * n" n "m * p"] - add_assoc_nat[of "m * n + n" "m * p" p] - show "m * n + m * p + (n + p) = m * n + n + (m * p + p)" - by simp -qed - -lemma add_mult_distrib_right_nat [algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "(n + p) * m = n * m + p * m" -using m apply induct -using n p apply auto -proof - - fix m - assume "m \ Nat" "(n + p) * m = n * m + p * m" - with n p - add_assoc_nat[of "n * m + p * m" n p] - add_assoc_nat[of "n * m" "p * m" n] - add_commute_nat[of "p * m" n] - add_assoc_nat[of "n * m" n "p * m"] - add_assoc_nat[of "n * m + n" "p * m" p] - show "n * m + p * m + (n + p) = n * m + n + (p * m + p)" - by simp -qed - -text \ Identity element \ - -(* used for arithmetic simplification setup *) -lemma mult_1_right_nat: "a \ Nat \ a * 1 = a" by simp -lemma mult_1_left_nat: "a \ Nat \ 1 * a = a" by simp - -text \ Associativity \ - -lemma mult_assoc_nat[algebra_simps]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "m * (n * p) = (m * n) * p" -using m apply induct -using assms by (auto simp add: add_mult_distrib_right_nat) - -text \ Reasoning about @{text "m * n = 0"}, etc. \ - -lemma mult_is_0_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = 0) = (m = 0 \ n = 0)" -using m apply induct -using n by auto - -lemma mult_eq_1_iff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = 1) = (m = 1 \ n = 1)" -using m apply induct -using n by (induct, auto)+ - -lemma one_eq_mult_iff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 = m * n) = (m = 1 \ n = 1)" -proof - - have "(1 = m * n) = (m * n = 1)" by auto - also from assms have "... = (m = 1 \ n = 1)" by simp - finally show ?thesis . -qed - -text \ Cancellation rules \ - -lemma mult_cancel1_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k * m = k * n) = (m = n \ (k = 0))" -proof - - have "k \ 0 \ k * m = k * n \ m = n" - using n m proof (induct arbitrary: m) - fix m - assume "k \ 0" "k * m = k * 0" "m \ Nat" - with k show "m = 0" by simp - next - fix n m - assume - n': "n \ Nat" and h1:"\m. \k \ 0; k * m = k * n; m \ Nat\ \ m = n" - and k0: "k \ 0" and h2: "k * m = k * Succ[n]" and m':"m \ Nat" - from m' show "m = Succ[n]" - proof (rule natCases) - assume "m = 0" - with k have "k * m = 0" by simp - moreover - from k k0 n' have "k * Succ[n] \ 0" by simp - moreover - note h2 - ultimately show ?thesis by simp - next - fix w - assume w:"w \ Nat" and h3:"m = Succ[w]" - with k n' h2 have "k * w = k * n" by simp - with k k0 w h1[of w] h3 show ?thesis by simp - qed - qed - with k m n show ?thesis by auto -qed - -lemma mult_cancel2_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k = n * k) = (m = n \ k = 0)" -using assms by (simp add: mult_commute_nat) - -lemma Suc_mult_cancel1_nat: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[k] * m = Succ[k] * n) = (m = n)" -using k m n mult_cancel1_nat[of "Succ[k]" m n] by simp - - -lemma mult_left_commute_nat[algebra_simps]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a * (b * c) = b * (a * c)" -using assms by(simp only: mult_commute_nat mult_assoc_nat) - -(* from HOL/OrderedGroups.thy *) -lemmas mult_ac_nat = mult_assoc_nat mult_commute_nat mult_left_commute_nat - - -subsection \ Predecessor \ - -definition Pred -where "Pred \ [n \ Nat \ IF n=0 THEN 0 ELSE CHOOSE x \ Nat : n=Succ[x]]" - -lemma Pred_0_nat [simp]: "Pred[0] = 0" -by (simp add: Pred_def) - -lemma Pred_Succ_nat [simp]: "n \ Nat \ Pred[Succ[n]] = n" -unfolding Pred_def by (auto intro: bChooseI2) - -lemma Succ_Pred_nat [simp]: - assumes "n \ Nat" and "n \ 0" - shows "Succ[Pred[n]] = n" -using assms unfolding Pred_def by (cases "n", auto intro: bChooseI2) - -lemma Pred_in_nat [intro!, simp]: - assumes "n \ Nat" shows "Pred[n] \ Nat" -using assms by (cases "n", auto) - -subsection \ Difference of natural numbers \ - -text \ - We define a form of difference @{text "--"} of natural numbers that cuts off - at $0$, that is @{text "m -- n = 0"} if @{text "m < n"}. This is sometimes - called ``arithmetic difference''. -\ - -definition adiffnat -where "adiffnat(m) \ CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]])" - -definition adiff (infixl "--" 65) -where nat_adiff_def: "\m \ Nat; n \ Nat\ \ (m -- n) \ adiffnat(m)[n]" - -text \ Closure \ - -lemma adiffnatType: - assumes "m \ Nat" shows "adiffnat(m) \ [Nat \ Nat]" -using assms unfolding adiffnat_def by (rule bprimrecType_nat, auto) - -lemma adiffIsNat [intro!, simp]: - assumes m: "m \ Nat" and n: "n \ Nat" shows "m -- n \ Nat" -unfolding nat_adiff_def[OF assms] using assms adiffnatType by blast - -text \ Neutral element and Inductive step \ - -lemma adiffnat_0: - assumes "m \ Nat" shows "adiffnat(m)[0] = m" -using assms unfolding adiffnat_def by (rule primrec_natE, auto) - -lemma adiff_0_nat [simp]: - assumes "m \ Nat" shows "m -- 0 = m" -by (simp add: nat_adiff_def[OF assms] adiffnat_0[OF assms]) - -lemma adiffnat_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" -proof (rule primrec_natE[OF m]) - show "adiffnat(m) = (CHOOSE g \ [Nat \ Nat] : g[0] = m \ (\x \ Nat : g[Succ[x]] = Pred[g[x]]))" - unfolding adiffnat_def .. -next - assume "\n \ Nat : adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" - with n show "adiffnat(m)[Succ[n]] = Pred[adiffnat(m)[n]]" by blast -qed (auto) - -lemma adiff_Succ_nat: - assumes "m \ Nat" and "n \ Nat" - shows "m -- Succ[n] = Pred[m -- n]" -using assms by (simp add: nat_adiff_def adiffnat_Succ[OF assms]) - -lemma adiff_0_eq_0_nat [simp]: - assumes n: "n \ Nat" - shows "0 -- n = 0" -using assms by (induct) (simp_all add: adiff_Succ_nat) - -text \ Reasoning about @{text "m -- m = 0"}, etc. \ - -lemma adiff_Succ_Succ_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "Succ[m] -- Succ[n] = m -- n" - using n by induct (auto simp add: m adiff_Succ_nat) - -lemma adiff_self_eq_0_nat [simp]: - assumes m: "m \ Nat" - shows "m -- m = 0" -using m by induct auto - -text \ Associativity \ - -lemma adiff_adiff_left_nat: - assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "(i -- j) -- k = i -- (j + k)" -using i j by (rule diffInduct) (auto simp: k) - -lemma Succ_adiff_adiff_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[m] -- n) -- Succ[k] = (m -- n) -- k" -using assms by (simp add: adiff_adiff_left_nat) - -text \ Commutativity \ - -lemma adiff_commute_nat: - assumes i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "i -- j -- k = i -- k -- j" -using assms by (simp add: adiff_adiff_left_nat add_commute_nat) - -text \ Cancellation rules \ - -lemma adiff_add_inverse_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(n + m) -- n = m" -using n by induct (simp_all add: m) - -lemma adiff_add_inverse2_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m + n) -- n = m" -using assms by (simp add: add_commute_nat [of m n]) - -lemma adiff_cancel_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m) -- (k + n) = m -- n" -using k by induct (simp_all add: m n) - -lemma adiff_cancel2_nat [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k) -- (n + k) = m -- n" -using assms by (simp add: add_commute_nat) - -lemma adiff_add_0_nat [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "n -- (n + m) = 0" -using n by induct (simp_all add: m) - -text \ Difference distributes over multiplication \ - -lemma adiff_mult_distrib_nat [algebra_simps]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m -- n) * k = (m * k) -- (n * k)" -using m n by (rule diffInduct) (simp_all add: k) - -lemma adiff_mult_distrib2_nat [algebra_simps]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "k * (m -- n) = (k * m) -- (k * n)" -using assms by (simp add: adiff_mult_distrib_nat mult_commute_nat [of k]) - -\ \ NOT added as rewrites, since sometimes they are used from right to left \ -lemmas nat_distrib = - add_mult_distrib_right_nat add_mult_distrib_left_nat - adiff_mult_distrib_nat adiff_mult_distrib2_nat - - -subsection \ Additional arithmetic theorems \ - -subsubsection \ Monotonicity of Addition \ - -lemma Succ_pred [simp]: - assumes n: "n \ Nat" - shows "n > 0 \ Succ[n -- 1] = n" -using assms by (simp add: adiff_Succ_nat[OF n zeroIsNat] nat_gt0_not0[OF n]) - -lemma nat_add_left_cancel_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m \ k + n) = (m \ n)" -using assms by (induct k) simp_all - -lemma nat_add_right_cancel_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k \ n + k) = (m \ n)" -using assms by (induct k) simp_all - -lemma nat_add_left_cancel_Succ_leq [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[k + m] \ k + n) = (Succ[m] \ n)" -using assms by (induct k) simp_all - -(* immediate corollary *) -lemma nat_add_left_cancel_less: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k + m < k + n) = (m < n)" -using assms by simp - -lemma nat_add_right_cancel_Succ_less [simp]: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m + k] \ n + k) = (Succ[m] \ n)" -using assms by (induct k) simp_all - -lemma nat_add_right_cancel_less: - assumes k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(m + k < n + k) = (m < n)" -using assms by simp - -lemma less_imp_Succ_add: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m < n \ (\k \ Nat: n = Succ[m + k])" (is "_ \ ?P(n)") -using n proof (induct) - case 0 with m show ?case by simp -next - fix n - assume n: "n \ Nat" and ih: "m ?P(n)" and suc: "m < Succ[n]" - from suc m n show "?P(Succ[n])" - proof (rule nat_less_SuccE) - assume "m Nat" and "n = Succ[m + k]" by (blast dest: ih) - with m n have "Succ[k] \ Nat" and "Succ[n] = Succ[m + Succ[k]]" by auto - thus ?thesis .. - next - assume "m=n" - with n have "Succ[n] = Succ[m + 0]" by simp - thus ?thesis by blast - qed -qed - -lemma nat_leq_trans_add_left_false [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "\m + n \ p; p \ n\ \ (m + n < p) = FALSE" -apply (induct n p rule: diffInduct) -using assms by simp_all - - -subsubsection \ (Partially) Ordered Groups \ - -\ \ The two following lemmas are just ``one half'' of - @{text nat_add_left_cancel_leq} and @{text nat_add_right_cancel_leq} - proved above." \ -lemma add_leq_left_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a \ b \ c + a \ c + b" -using assms by simp - -lemma add_leq_right_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a \ b \ a + c \ b + c" -using assms by simp - -text \ non-strict, in both arguments \ -lemma add_leq_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a \ b \ c \ d \ a + c \ b + d" -using assms - add_leq_right_mono[OF a b c] - add_leq_left_mono[OF c d b] - nat_leq_trans[of "a + c" "b + c" "b + d"] -by simp - -\ \ Similar for strict less than. \ -lemma add_less_left_mono: - assumes "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "a < b \ c + a < c + b" - using assms by (simp add: add_leq_left_mono less_def) - -lemma add_less_right_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a < b \ a + c < b + c" -using assms by (auto simp: add_less_left_mono add_commute_nat) - -text \ Strict monotonicity in both arguments \ -lemma add_less_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a < b \ c < d \ a + c < b + d" -using assms - add_less_right_mono[OF a b c] - add_less_left_mono[OF c d b] - nat_less_trans[of "a + c" "b + c" "b + d"] -by simp - -lemma add_less_leq_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a < b \ c \ d \ a + c < b + d" -using assms - add_less_right_mono[OF a b c] - add_leq_left_mono[OF c d b] - nat_less_leq_trans[of "a + c" "b + c" "b + d"] -by blast - -lemma add_leq_less_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - shows "a \ b \ c < d \ a + c < b + d" -using assms - add_leq_right_mono[OF a b c] - add_less_left_mono[OF c d b] - nat_leq_less_trans[of "a + c" "b + c" "b + d"] -by blast - -lemma leq_add1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "n \ n + m" -using assms add_leq_left_mono[of 0 m n] by simp - -lemma leq_add2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m + n" -using assms add_leq_right_mono [of 0 m n] by simp - -lemma trans_leq_add1 [elim]: - assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m+k" -using assms by (simp add: nat_leq_trans) - -lemma trans_leq_add2 [elim]: - assumes "n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ k+m" -using assms by (simp add: nat_leq_trans) - -lemma add_leqD1 [elim]: - assumes "n+k \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m" -using assms by (simp add: nat_leq_trans[of "n" "n+k" "m"]) - -lemma add_leqD2 [elim]: - assumes "k+n \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "n \ m" -using assms by (simp add: nat_leq_trans[of "n" "k+n" "m"]) - -lemma add_Succ_leqD1 [elim]: - assumes "Succ[n+k] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "Succ[n] \ m" -using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[n+k]" "m"]) - -lemma add_Succ_leqD2 [elim]: - assumes "Succ[k+n] \ m" and "n \ Nat" and "m \ Nat" and "k \ Nat" - shows "Succ[n] \ m" -using assms by (simp add: nat_leq_trans[of "Succ[n]" "Succ[k+n]" "m"]) - -lemma less_add_Succ1: - assumes "i \ Nat" and "m \ Nat" - shows "i < Succ[i + m]" -using assms by simp - -lemma less_add_Succ2: - assumes "i \ Nat" and "m \ Nat" - shows "i < Succ[m + i]" -using assms by simp - -lemma less_iff_Succ_add: - assumes "m \ Nat" and "n \ Nat" - shows "(m < n) = (\k \ Nat: n = Succ[m + k])" -using assms by (auto intro!: less_imp_Succ_add) - -lemma trans_less_add1: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i < j + m" -using assms by auto - -lemma trans_less_add2: - assumes "i < j" and "i \ Nat" and "j \ Nat" and "m \ Nat" - shows "i < m + j" -using assms by auto - -lemma add_lessD1: - assumes "i+j < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "i < k" -using assms by auto - -lemma add_lessD2: - assumes "j+i < k" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "i < k" -using assms by auto - -lemma not_Succ_add_le1: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(Succ[i + j] \ i) = FALSE" -using assms by simp - -lemma not_Succ_add_le2: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(Succ[j + i] \ i) = FALSE" -using assms by simp - -lemma not_add_less1: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(i + j < i) = FALSE" -using assms by simp - -lemma not_add_less2: - assumes "i \ Nat" and "j \ Nat" - shows "(j + i < i) = FALSE" -using assms by simp - -lemma add_leqE: - assumes "m+k \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(m \ n \ k \ n \ R) \ R" -using assms by (blast dest: add_leqD1 add_leqD2) - -lemma leq_add_self1 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m+n \ m) = (n=0)" -proof - - { - assume "m+n \ m" - with assms have "m+n = m" by (intro nat_leq_antisym) simp_all - } - with assms show ?thesis by auto -qed - -lemma leq_add_self2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(n+m \ m) = (n=0)" -proof - - { - assume "n+m \ m" - with assms have "n+m = m" by (intro nat_leq_antisym) simp_all - } - with assms show ?thesis by auto -qed - -lemma add_gt_0: - assumes "m \ Nat" and "n \ Nat" - shows "(m + n > 0) = (m > 0 \ n > 0)" -proof - - from assms have "m + n \ Nat" by simp - hence "(m + n > 0) = (m + n \ 0)" by (rule nat_gt0_not0) - also from assms have "\ = (m \ 0 \ n \ 0)" by simp - finally show ?thesis using assms by (simp only: nat_gt0_not0) -qed - -(* The above lemma follows from the following simplification rule. *) -lemma add_ge_1 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m + n \ 1) = (m \ 1 \ n \ 1)" - using assms by (auto simp add: add_gt_0) - - -subsubsection \ More results about arithmetic difference \ - -text \ Addition is the inverse of subtraction: - if @{term "n \ m"} then @{term "n + (m -- n) = m"}. \ -lemma add_adiff_inverse: - assumes "m \ Nat" and "n \ Nat" - shows "\(m < n) \ n + (m -- n) = m" -apply (induct m n rule: diffInduct) -using assms by simp_all - -lemma le_add_adiff_inverse [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ n + (m -- n) = m" -using assms by (simp add: add_adiff_inverse nat_not_order_simps) - -lemma le_add_adiff_inverse2 [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ (m -- n) + n = m" -using assms by (simp add: add_commute_nat) - -lemma Succ_adiff_leq: - assumes "m \ Nat" and "n \ Nat" - shows "n \ m \ Succ[m] -- n = Succ[m -- n]" -apply (induct m n rule: diffInduct) -using assms by simp_all - -lemma adiff_less_Succ: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m -- n < Succ[m]" -apply (induct m n rule: diffInduct) -using assms by (auto simp: nat_less_Succ) - -lemma adiff_leq_self [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m -- n \ m" -apply (induct m n rule: diffInduct) -using assms by auto - -lemma leq_iff_add: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n = (\k \ Nat: n = m + k)" (is "?lhs = ?rhs") -proof - - have "?lhs \ ?rhs" - proof - assume mn: "m \ n" - with m n have "n = m + (n -- m)" by simp - with m n show "?rhs" by blast - qed - moreover - from assms have "?rhs \ ?lhs" by auto - ultimately show ?thesis by blast -qed - -lemma less_imp_adiff_less: - assumes jk: "j < k" and j: "j \ Nat" and k: "k \ Nat" and n: "n \ Nat" - shows "j -- n < k" -proof - - from j n have "j -- n \ Nat" by simp - moreover - from j n have s1_2: "j -- n \ j" by simp - ultimately show ?thesis using j k jk nat_leq_less_trans by auto -qed - -lemma adiff_Succ_less (*[simp]*): - assumes i: "i \ Nat" and n: "n \ Nat" - shows "0 < n \ n -- Succ[i] < n" -using n by cases (auto simp: i) - -lemma adiff_add_assoc: - assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(i + j) -- k = i + (j -- k)" -using assms by (induct j k rule: diffInduct, simp+) - -lemma adiff_add_assoc2: - assumes "k \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(j + i) -- k = (j -- k) + i" -using assms by (simp add: add_commute_nat adiff_add_assoc) - -lemma adiff_add_assoc3: - assumes "n \ p" and "p \ m+n" and "m \ Nat" and "n \ Nat" and "p \ Nat" - shows "m -- (p -- n) = m + n -- p" -using assms by (induct p n rule: diffInduct, simp+) - -lemma adiff_add_assoc4: - assumes 1: "n \ m" and 2: "m -- n \ p" and 3: "m \ p" - and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "p -- (m -- n) = (p -- m) + n" -using assms - adiff_add_assoc2[OF _ n p m, symmetric] - adiff_add_assoc3[OF _ _ p n m] apply simp -using trans_leq_add1[OF _ m p n] by simp - -lemma le_imp_adiff_is_add: - assumes "i \ j" and "i \ Nat" and "j \ Nat" and "k \ Nat" - shows "(j -- i = k) = (j = k + i)" -using assms by auto - -lemma adiff_is_0_eq [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m -- n = 0) = (m \ n)" -by (induct m n rule: diffInduct, simp_all add: assms) - -lemma adiff_is_0_eq' (*[simp]*): - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "m -- n = 0" -using assms by simp - -lemma one_leq_adiff [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(1 \ n -- m) = (Succ[m] \ n)" -by (induct m n rule: diffInduct, simp_all add: assms) - -lemma zero_less_adiff: - assumes "m \ Nat" and "n \ Nat" - shows "(0 < n -- m) = (m < n)" -using assms by simp - -lemma less_imp_add_positive: - assumes "i < j" and "i \ Nat" and "j \ Nat" - shows "\k \ Nat: 0 < k \ i + k = j" -proof - - thm leq_iff_add - from assms obtain k where k: "k \ Nat" "j = Succ[i]+k" - by (auto simp: leq_iff_add) - with assms show ?thesis by force -qed - -lemma leq_adiff_right_add_left: - assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "m \ n -- k = (m + k \ n)" -using assms by (induct n k rule: diffInduct, simp+) - -lemma leq_adiff_left_add_right_equiv: - assumes "k \ n" and "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(n -- k \ m) = (n \ m + k)" -using assms by (induct n k rule: diffInduct, simp+) - -lemma leq_adiff_left_add_right: - assumes 1: "n -- p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "n \ m + p" -using assms by (induct n p rule: diffInduct, simp+) - -lemma leq_adiff_trans: - assumes "p \ m" and m: "m \ Nat" and n: "n \ Nat" and p: "p \ Nat" - shows "p -- n \ m" -apply(rule nat_leq_trans[of "p -- n" p m]) -using assms adiff_leq_self[OF p n] by simp_all - -lemma leq_adiff_right [simp]: - assumes "n \ m" and "m \ Nat" and "n \ Nat" - shows "(m \ m -- n) = (n = 0)" - using assms by (simp add: leq_adiff_right_add_left) - - -subsubsection \ Monotonicity of Multiplication \ - -(* from HOL/Ring_and_Field.thy *) - -lemma mult_leq_left_mono: - assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "c * a \ c * b" -using c by induct (simp_all add: add_leq_mono 1 a b) - -lemma mult_leq_right_mono: - assumes 1:"a \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a * c \ b * c" -using c by induct (simp_all add: add_leq_mono add_commute_nat 1 a b) - -text \ @{text "\"} monotonicity, BOTH arguments \ - -lemma mult_leq_mono: - assumes 1: "i \ j" "k \ l" - and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" and l: "l \ Nat" - shows "i * k \ j * l" -using assms - mult_leq_right_mono[OF _ i j k] - mult_leq_left_mono[OF _ k l j] - nat_leq_trans[of "i * k" "j * k" "j * l"] - by simp - -lemma self_leq_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" - shows "a \ a * b" -proof - - have "1 \ Nat" by simp - from 1 this b a have "a * 1 \ a * b" by (rule mult_leq_left_mono) - with a b show ?thesis by simp -qed - -lemma self_leq_mult_right: - assumes a: "a \ Nat" and b: "b \ Nat" and 1: "1 \ b" - shows "a \ b * a" -using assms by (simp add: self_leq_mult_left mult_commute_nat) - -text \ Similar lemmas for @{text "<"} \ - -lemma mult_Succ_leq_right_mono1: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "Succ[a * Succ[c]] \ b * Succ[c]" -using c proof induct - case 0 - with assms show ?case by simp -next - fix n - assume n: "n \ Nat" and ih: "Succ[a * Succ[n]] \ b * Succ[n]" - from ab a b have 1: "a \ b" - by (simp add: nat_leq_trans[of "a" "Succ[a]" "b"]) - from a n have "Succ[a * Succ[Succ[n]]] = Succ[a * Succ[n]] + a" - by simp - also from a b n ih have "\ \ b* Succ[n] + a" - by simp - also from a b n 1 have "\ \ b * Succ[n] + b" - by simp - finally show "Succ[a * Succ[Succ[n]]] \ b * Succ[Succ[n]]" - using b n by simp -qed - -lemma mult_Succ_leq_right_mono2: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" - shows "Succ[a * c] \ b * c" -proof - - from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" - by (blast dest: nat_ge1_implies_Succ) - with ab a b show ?thesis by (auto dest: mult_Succ_leq_right_mono1[of a b k]) -qed - -lemma mult_less_right_mono: - assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "a * c < b * c" -using assms by (simp add: mult_Succ_leq_right_mono2) - -lemma mult_Succ_leq_left_mono1: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "Succ[Succ[c] * a] \ Succ[c] * b" -proof - - have "Succ[Succ[c] * a] = Succ[a * Succ[c]]" - using a c by (simp add: mult_commute_nat) - also have "\ \ b * Succ[c]" - using assms by (rule mult_Succ_leq_right_mono1) - also have "\ = Succ[c] * b" - using b c by (simp add: mult_commute_nat) - finally show ?thesis . -qed - -lemma mult_Succ_leq_left_mono2: - assumes ab: "Succ[a] \ b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and 1: "1 \ c" - shows "Succ[c * a] \ c * b" -proof - - from c 1 obtain k where k: "k \ Nat" "c = Succ[k]" - by (blast dest: nat_ge1_implies_Succ) - with ab a b show ?thesis by (auto dest: mult_Succ_leq_left_mono1[of a b k]) -qed - -lemma mult_less_left_mono: - assumes "a < b" and "0 < c" and "a \ Nat" and "b \ Nat" and "c \ Nat" - shows "c * a < c * b" -using assms by (simp add: mult_Succ_leq_left_mono2) - -lemma one_leq_mult_iff[simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(1 \ m * n) = (1 \ m \ 1 \ n)" -using assms nat_gt0_not0 by simp - -lemma nat_0_less_mult_iff: - assumes i: "i \ Nat" and j: "j \ Nat" - shows "(0 < i * j) = (0 < i \ 0 < j)" -using assms by simp - -lemma mult_Succ_leq_cancel_left [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[k * m] \ k * n) = (1 \ k \ Succ[m] \ n)" -proof (auto intro: mult_Succ_leq_left_mono2[OF _ m n k _]) - assume "Succ[k * m] \ k * n" - from k m n this show "1 \ k" by (cases k) simp_all -next - assume 1: "Succ[k * m] \ k * n" - show "Succ[m] \ n" - proof (rule contradiction) - assume "\(Succ[m] \ n)" - with m n k have "k*n \ k*m" by (simp add: nat_not_order_simps mult_leq_left_mono) - with m n k have "\(Succ[k*m] \ k*n)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed - -lemma mult_less_cancel_left: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m < k * n) = (0 < k \ m < n)" -using assms by simp - -lemma mult_Succ_leq_cancel_right [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(Succ[m * k] \ n * k) = (1 \ k \ Succ[m] \ n)" -proof (auto intro: mult_Succ_leq_right_mono2[OF _ m n k _]) - assume "Succ[m * k] \ n * k" - from k m n this show "1 \ k" by (cases k) simp_all -next - assume 1: "Succ[m * k] \ n * k" - show "Succ[m] \ n" - proof (rule contradiction) - assume "\(Succ[m] \ n)" - with m n k have "n*k \ m*k" by (simp add: nat_not_order_simps mult_leq_right_mono) - with m n k have "\(Succ[m*k] \ n*k)" by (simp add: nat_not_order_simps) - with 1 show "FALSE" by simp - qed -qed - -lemma mult_less_cancel_right: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k < n * k) = (0 < k \ m < n)" -using assms by simp - -lemma mult_Succ_leq_self_left [dest]: - assumes nk: "Succ[n*k] \ n" and n: "n \ Nat" and k: "k \ Nat" - shows "k = 0" -proof (rule contradiction) - assume "k \ 0" - with k nat_neq0_conv have "1 \ k" by simp - with n k have "n \ n * k" by (rule self_leq_mult_left) - with n k have "\(Succ[n*k] \ n)" by (simp add: nat_not_order_simps) - from this nk show "FALSE" .. -qed - -lemma mult_less_self_left: - assumes "n * k < n" and "n \ Nat" and "k \ Nat" - shows "k=0" -using assms by auto - -lemma mult_Succ_leq_self_right [dest]: - assumes less: "Succ[k*n] \ n" and n: "n \ Nat" and k: "k \ Nat" - shows "k=0" -using assms by (auto simp: mult_commute_nat[OF k n]) - -lemma mult_less_self_right: - assumes less: "k*n < n" and n: "n \ Nat" and k: "k \ Nat" - shows "k=0" -using assms by auto - -lemma mult_leq_cancel_left [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(k * m \ k * n) = (k = 0 \ m \ n)" -proof - - { - assume 1: "k*m \ k*n" and 2: "k \ 0" and 3: "\(m \ n)" - from k 2 nat_gt0_not0 have "k > 0" by simp - with 3 k m n have "\(k * m \ k * n)" by (simp add: nat_not_order_simps) - from this 1 have "FALSE" .. - } - moreover have "k=0 \ k*m \ k*n" - using assms by simp - moreover have "m \ n \ k*m \ k*n" - using assms by (simp add: mult_leq_left_mono) - ultimately show ?thesis by blast -qed - -lemma mult_leq_cancel_right [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and k: "k \ Nat" - shows "(m * k \ n * k) = (k = 0 \ m \ n)" -using assms by (simp add: mult_commute_nat) - -lemma Suc_mult_less_cancel1: - assumes "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(Succ[k] * m < Succ[k] * n) = (m < n)" -using assms mult_less_cancel_left[of "m" "n" "Succ[k]"] - nat_zero_less_Succ[of "k"] by auto - -lemma Suc_mult_leq_cancel1: - assumes "m \ Nat" and "n \ Nat" and "k \ Nat" - shows "(Succ[k] * m \ Succ[k] * n) = (m \ n)" -using assms by (simp del: mult_Succ_left_nat) - -lemma nat_leq_square: - assumes m: "m \ Nat" - shows "m \ m * m" -using m by (cases, auto) - -lemma nat_leq_cube: - assumes m: "m \ Nat" - shows "m \ m * (m * m)" -using m by (cases, auto) - -text \ Lemma for @{text gcd} \ -lemma mult_eq_self_implies_10: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m * n = m) = (n = 1 \ m = 0)" (is "?lhs = ?rhs") -proof - - from assms have "(m*n = m) = (m*n = m*1)" by simp - also have "\ = ?rhs" by (rule mult_cancel1_nat[OF m n oneIsNat]) - finally show ?thesis . -qed - -end diff --git a/isabelle/NatDivision.thy b/isabelle/NatDivision.thy deleted file mode 100644 index cdfa1eb4..00000000 --- a/isabelle/NatDivision.thy +++ /dev/null @@ -1,731 +0,0 @@ -(* Title: TLA+/NatDivision.thy - Author: Hernan Vanzetto - Copyright (C) 2009-2011 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:39:56 merz> -*) - -section \ The division operators div and mod on Naturals \ - -theory NatDivision -imports NatArith Tuples -begin - -subsection \ The divisibility relation \ - -definition dvd (infixl "dvd" 50) -where "a \ Nat \ b \ Nat \ b dvd a \ (\k \ Nat : a = b * k)" - -lemma boolify_dvd [simp]: - assumes "a \ Nat" and "b \ Nat" - shows "boolify(b dvd a) = (b dvd a)" -using assms by (simp add: dvd_def) - -lemma dvdIsBool [intro!,simp]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "isBool(b dvd a)" -using assms by (simp add: dvd_def) - -lemma [intro!]: - "\isBool(P); isBool(a dvd b); (a dvd b) \ P\ \ (a dvd b) = P" - "\isBool(P); isBool(a dvd b); P \ (a dvd b)\ \ P = (a dvd b)" -by auto - -lemma dvdI [intro]: - assumes a: "a \ Nat" and b: "b \ Nat" and k: "k \ Nat" - shows "a = b * k \ b dvd a" -unfolding dvd_def[OF a b] using k by blast - -lemma dvdE [elim]: - assumes "b dvd a" and "a \ Nat" and "b \ Nat" - shows "(\k. \k \ Nat; a = b * k\ \ P) \ P" -using assms by (auto simp add: dvd_def) - -lemma dvd_refl [iff]: - assumes a: "a \ Nat" - shows "a dvd a" -proof - - from a have "a = a*1" by simp - with a show ?thesis by blast -qed - -lemma dvd_trans [trans]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "b dvd c" - shows "a dvd c" -proof - - from a b 1 obtain k where k: "k \ Nat" and "b = a * k" by blast - moreover - from b c 2 obtain l where l: "l \ Nat" and "c = b * l" by blast - ultimately have h:"c = a * (k * l)" - using a b c by (simp add: mult_assoc_nat) - thus ?thesis using a c k l by blast -qed - -lemma dvd_0_left_iff [simp]: - assumes "a \ Nat" - shows "(0 dvd a) = (a = 0)" -using assms by force - -lemma dvd_0_right [iff]: - assumes a: "a \ Nat" shows "a dvd 0" -using assms by force - -lemma one_dvd [iff]: - assumes a: "a \ Nat" - shows "1 dvd a" -using assms by force - -lemma dvd_mult (*[simp]*): - assumes dvd: "a dvd c" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a dvd (b * c)" -proof - - from dvd a c obtain k where k: "k \ Nat" and "c = a*k" by blast - with a b c have "b*c = a * (b*k)" by (simp add: mult_left_commute_nat) - with a b c k show ?thesis by blast -qed - -lemma dvd_mult2 (*[simp]*): - assumes dvd: "a dvd b" and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - shows "a dvd (b * c)" -using mult_commute_nat[OF b c] dvd_mult[OF dvd a c b] by simp - -lemma dvd_triv_right [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "a dvd (b * a)" -using assms by (intro dvd_mult, simp+) - -lemma dvd_triv_left [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" - shows "a dvd a * b" -using assms by (intro dvd_mult2, simp+) - -lemma mult_dvd_mono: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and d: "d \ Nat" - and 1: "a dvd b" and 2: "c dvd d" - shows "(a * c) dvd (b * d)" -proof - - from a b 1 obtain b' where b': "b = a * b'" "b' \ Nat" by blast - from c d 2 obtain d' where d': "d = c * d'" "d' \ Nat" by blast - with b' a b c d - mult_assoc_nat[of a b' "c * d'"] - mult_left_commute_nat[of b' c d'] - mult_assoc_nat[of a c "b' * d'"] - have "b * d = (a * c) * (b' * d')" by simp - with a c b' d' show ?thesis by blast -qed - -lemma dvd_mult_left: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and h: "a * b dvd c" - shows "a dvd c" -proof - - from h a b c obtain k where k: "k \ Nat" "c = a*(b*k)" - by (auto simp add: mult_assoc_nat) - with a b c show ?thesis by blast -qed - -lemma dvd_mult_right: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and h: "a*b dvd c" - shows "b dvd c" -proof - - from h a b c have "b*a dvd c" by (simp add: mult_ac_nat) - with b a c show ?thesis by (rule dvd_mult_left) -qed - -lemma dvd_0_left: - assumes "a \ Nat" - shows "0 dvd a \ a = 0" -using assms by simp - -lemma dvd_add [iff]: - assumes a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" - and 1: "a dvd b" and 2: "a dvd c" - shows "a dvd (b + c)" -proof - - from a b 1 obtain b' where b': "b' \ Nat" "b = a * b'" by blast - from a c 2 obtain c' where c': "c' \ Nat" "c = a * c'" by blast - from a b c b' c' - have "b + c = a * (b' + c')" by (simp add: add_mult_distrib_left_nat) - with a b' c' show ?thesis by blast -qed - -subsection \ Division on @{const Nat} \ - -text \ - We define division and modulo over @{const Nat} by means - of a characteristic relation with two input arguments - @{term "m"}, @{term "n"} and two output arguments - @{term "q"}(uotient) and @{term "r"}(emainder). - - The following definition works for natural numbers, but also for - possibly negative integers. Obviously, the second disjunct cannot - hold for natural numbers. -\ - -definition divmod_rel where - "divmod_rel(m,n,q,r) \ - m = q * n + r - \ - ((0 < n \ 0 \ r \ r < n) \ - (n < 0 \ r \ 0 \ n < r))" - -text \ @{const divmod_rel} is total if $n$ is non-zero. \ - -lemma divmod_rel_ex: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - obtains q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" -proof - - have "\q,r \ Nat : m = q*n + r \ r < n" - using m proof (induct) - case 0 - from n pos have "0 = 0 * n + 0 \ 0 < n" by simp - then show ?case by blast - next - fix m' - assume m': "m' \ Nat" and ih: "\q, r \ Nat : m' = q*n + r \ r < n" - from ih obtain q' r' - where h1: "m' = q' * n + r'" and h2: "r' < n" - and q': "q' \ Nat" and r': "r' \ Nat" by blast - show "\q, r \ Nat : Succ[m'] = q * n + r \ r < n" - proof (cases "Succ[r'] < n") - case True - from h1 h2 m' q' n r' have "Succ[m'] = q' * n + Succ[r']" by simp - with True q' r' show ?thesis by blast - next - case False - with n r' have "n \ Succ[r']" by (simp add: nat_not_less[simplified]) - with r' n h2 have "n = Succ[r']" by (intro nat_leq_antisym, simp+) - with h1 m' q' r' have "Succ[m'] = Succ[q'] * n + 0" by (simp add: add_ac_nat) - with pos q' show ?thesis by blast - qed - qed - with pos that show ?thesis by (auto simp: divmod_rel_def) -qed - -text \ @{const divmod_rel} has unique solutions in the natural numbers. \ -lemma divmod_rel_unique_div: - assumes 1: "divmod_rel(m,n,q,r)" and 2: "divmod_rel(m,n,q',r')" - and m: "m \ Nat" and n: "n \ Nat" - and q: "q \ Nat" and r: "r \ Nat" and q': "q' \ Nat" and r': "r' \ Nat" - shows "q = q'" -proof - - from n 1 have pos: "0 < n" and mqr: "m = q*n+r" and rn: "r < n" - by (auto simp: divmod_rel_def) - from n 2 have mqr': "m = q'*n+r'" and rn': "r' < n" - by (auto simp: divmod_rel_def) - { - fix x y x' y' - assume nat: "x \ Nat" "y \ Nat" "x' \ Nat" "y' \ Nat" - and eq: "x*n + y = x'*n + y'" and less: "y' < n" - have "x \ x'" - proof (rule contradiction) - assume "\(x \ x')" - with nat have "x' < x" by (simp add: nat_not_leq[simplified]) - with nat obtain k where k: "k \ Nat" "x = Succ[x'+k]" - using less_iff_Succ_add by auto - with eq nat n have "x'*n + (k*n + n + y) = x'*n + y'" - by (simp add: add_mult_distrib_right_nat add_assoc_nat) - with nat k n have "k*n + n + y = y'" by simp - with less k n nat have s3_1: "(k*n + y) + n < n" by (simp add: add_ac_nat) - define j where "j \ k * n + y" - have s3_2: "j + n < n" - unfolding j_def using s3_1 by auto - have s3_3: "j \\in Nat" - unfolding j_def using multIsNat addIsNat k n nat by auto - show "FALSE" - using s3_2 s3_3 n not_add_less2[of "n" "j"] by auto - qed - } - from this[OF q r q' r'] this[OF q' r' q r] q q' mqr mqr' rn rn' - show ?thesis by (intro nat_leq_antisym, simp+) -qed - -lemma divmod_rel_unique_mod: - assumes "divmod_rel(m,n,q,r)" and "divmod_rel(m,n,q',r')" - and "m \ Nat" and "n \ Nat" and "q \ Nat" and "r \ Nat" and "q' \ Nat" and "r' \ Nat" - shows "r = r'" -proof - - from assms have "q = q'" by (rule divmod_rel_unique_div) - with assms show ?thesis by (auto simp: divmod_rel_def) -qed - -text \ - We instantiate divisibility on the natural numbers by - means of @{const divmod_rel}: -\ - -definition divmodNat -where "divmodNat(m,n) \ CHOOSE z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" - -lemma divmodNatPairEx: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "\z \ Nat \ Nat : divmod_rel(m,n,z[1],z[2])" -proof - - from assms obtain q r where "q \ Nat" "r \ Nat" "divmod_rel(m,n,q,r)" - by (rule divmod_rel_ex) - thus ?thesis unfolding two_def by force -qed - -lemma divmodNatInNatNat: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "divmodNat(m,n) \ Nat \ Nat" -unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]]) - -lemma divmodNat_divmod_rel [rule_format]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "z = divmodNat(m,n) \ divmod_rel(m,n,z[1],z[2])" -unfolding divmodNat_def by (rule bChooseI2[OF divmodNatPairEx[OF assms]], auto) - -lemma divmodNat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - and q: "q \ Nat" and r: "r \ Nat" - shows "divmodNat(m,n) = \q,r\" -unfolding divmodNat_def -proof (rule bChooseI2[OF divmodNatPairEx[OF m n pos]]) - fix z - assume "z \ Nat \ Nat" and "divmod_rel(m,n,z[1],z[2])" - with m n q r h show "z = \q,r\" - unfolding two_def - by (auto elim!: inProdE elim: divmod_rel_unique_div divmod_rel_unique_mod) -qed - -text \ - We now define division and modulus over natural numbers. -\ - -definition div (infixr "div" 70) -where div_nat_def: "\m \ Nat; n \ Nat\ \ m div n \ divmodNat(m,n)[1]" - -definition mod (infixr "mod" 70) -where mod_nat_def: "\m \ Nat; n \ Nat\ \ m mod n \ divmodNat(m,n)[2]" - - -lemma divIsNat [iff]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "m div n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: div_nat_def) - -lemma modIsNat [iff]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "m mod n \ Nat" -using divmodNatInNatNat[OF assms] assms by (auto simp: mod_nat_def two_def) - -lemma divmodNat_div_mod: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "divmodNat(m,n) = \m div n, m mod n\" -unfolding div_nat_def[OF m n] mod_nat_def[OF m n] two_def -using divmodNatInNatNat[OF assms] -by force - -lemma divmod_rel_div_mod_nat: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "divmod_rel(m, n, m div n, m mod n)" -using divmodNat_divmod_rel[OF assms sym[OF divmodNat_div_mod[OF assms]]] -by (auto simp: two_def) - -lemma div_nat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" - shows "m div n = q" -unfolding div_nat_def[OF m n] using divmodNat_unique[OF assms] by simp - -lemma mod_nat_unique: - assumes h: "divmod_rel(m,n,q,r)" - and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and q: "q \ Nat" and r: "r \ Nat" - shows "m mod n = r" -unfolding mod_nat_def[OF m n] two_def using divmodNat_unique[OF assms] by simp - -lemma mod_nat_less_divisor: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "m mod n < n" -using assms divmod_rel_div_mod_nat[OF assms] by (simp add: divmod_rel_def) - -text \ ``Recursive'' computation of @{text div} and @{text mod}. \ - -lemma divmodNat_base: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "divmodNat(m,n) = \0, m\" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - let ?dm = "divmodNat(m,n)" - from m n pos have 1: "divmod_rel(m,n,?dm[1],?dm[2])" - by (simp add: divmodNat_divmod_rel) - from m n pos have 2: "?dm \ Nat \ Nat" by (rule divmodNatInNatNat) - have "?dm[1] * n < n" - proof - - have s2_1: "m = ?dm[1] * n + ?dm[2]" - using 1 by (auto simp: divmod_rel_def) - have s2_2: "?dm[1] * n + ?dm[2] < n" - using less s2_1 by auto - have s2_3: "?dm[1] \\in Nat" - using 2 by auto - have s2_4: "?dm[2] \\in Nat" - using 2 by (auto simp: two_def) - show ?thesis - using s2_2 s2_3 s2_4 multIsNat n add_lessD1 by auto - qed - with 2 n have 3: "?dm[1] = 0" by (intro mult_less_self_right, auto) - - with 1 2 m n have "?dm[2] = m" by (auto simp: divmod_rel_def two_def) - - with 3 prodProj[OF 2] show ?thesis by simp -qed - -lemma divmodNat_step: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "divmodNat(m,n) = \Succ[(m -- n) div n], (m -- n) mod n\" -proof - - from m n pos have 1: "divmod_rel(m, n, m div n, m mod n)" - by (rule divmod_rel_div_mod_nat) - have 2: "m div n \ 0" - proof - assume "m div n = 0" - with 1 m n pos have s2_1: "m < n" by (auto simp: divmod_rel_def) - have s2_2: "\ (n \ m)" - using s2_1 n m nat_not_leq by auto - show "FALSE" - using geq s2_2 by auto - qed - with m n pos obtain k where k1: "k \ Nat" and k2: "m div n = Succ[k]" - using not0_implies_Suc[of "m div n"] by auto - with 1 m n pos have "m = n + k*n + m mod n" - by (auto simp: divmod_rel_def add_commute_nat) - moreover - from m n k1 pos geq have "... -- n = k*n + m mod n" - by (simp add: adiff_add_assoc2) - ultimately - have "m -- n = k*n + m mod n" by simp - with pos m n 1 have "divmod_rel(m -- n, n, k, m mod n)" - by (auto simp: divmod_rel_def) - - with k1 m n pos have "divmodNat(m -- n, n) = \k, m mod n\" - by (intro divmodNat_unique, simp+) - moreover - from m n pos have "divmodNat(m -- n, n) = \(m--n) div n, (m--n) mod n\" - by (intro divmodNat_div_mod, simp+) - ultimately - have "m div n = Succ[(m--n) div n]" and "m mod n = (m--n) mod n" - using m n k2 by auto - thus ?thesis by (simp add: divmodNat_div_mod[OF m n pos]) -qed - - -text \ The ''recursion'' equations for @{text div} and @{text mod} \ - -lemma div_nat_less [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "m div n = 0" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - from divmodNat_base[OF m n less] divmodNat_div_mod[OF m n pos] show ?thesis - by simp -qed - -lemma div_nat_geq: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "m div n = Succ[(m -- n) div n]" -using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] -by simp - -lemma mod_nat_less [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and less: "m < n" - shows "m mod n = m" -proof - - from assms have pos: "0 < n" by (intro nat_leq_less_trans[of "0" "m" "n"], simp+) - from divmodNat_base[OF m n less] divmodNat_div_mod[OF m n pos] show ?thesis - by simp -qed - -lemma mod_nat_geq: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" and geq: "n \ m" - shows "m mod n = (m -- n) mod n" -using divmodNat_step[OF assms] divmodNat_div_mod[OF m n pos] -by simp - - -subsection \ Facts about @{const div} and @{const mod} \ - -lemma mod_div_nat_equality [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m div n) * n + m mod n = m" -using divmod_rel_div_mod_nat [OF assms] by (simp add: divmod_rel_def) - -lemma mod_div_nat_equality2 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "n * (m div n) + m mod n = m" -using assms mult_commute_nat[of "n" "m div n"] by simp - -lemma mod_div_nat_equality3 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" - shows "m mod n + (m div n) * n = m" -using assms add_commute_nat[of "m mod n"] by simp - -lemma mod_div_nat_equality4 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and "0 < n" - shows "m mod n + n * (m div n) = m" -using assms mult_commute_nat[of "n" "m div n"] by simp - -(** immediate consequence of above -lemma div_mod_nat_equality: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "((m div n) * n + m mod n) + k = m + k" -using assms by simp - -lemma div_mod_nat_equality2: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * (m div n) + m mod n) + k = m + k" -using assms by simp -**) - -lemma div_nat_mult_self1 [simp]: - assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(q + m * n) div n = m + (q div n)" (is "?P(m)") -using m proof (induct m) - from assms show "?P(0)" by simp -next - fix k - assume k: "k \ Nat" and ih: "?P(k)" - from n q k have "n \ q + (k*n + n)" by (simp add: add_assoc_nat) - with q k n pos have "(q + (k*n + n)) div n = Succ[(q + k*n) div n]" - by (simp add: div_nat_geq add_assoc_nat) - with ih q m n k pos show "?P(Succ[k])" by simp -qed - -lemma div_nat_mult_self2 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(q + n * m) div n = m + q div n" -using assms by (simp add: mult_commute_nat) - -lemma div_nat_mult_self3 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(m * n + q) div n = m + q div n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_mult_self4 [simp]: - assumes "q \ Nat" and "n \ Nat" and "m \ Nat" and "0 < n" - shows "(n * m + q) div n = m + q div n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_0: - assumes "n \ Nat" and "0 < n" - shows "0 div n = 0" -using assms by simp - -lemma mod_0: - assumes "n \ Nat" and "0 < n" - shows "0 mod n = 0" -using assms by simp - -lemma mod_nat_mult_self1 [simp]: - assumes q: "q \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(q + m * n) mod n = q mod n" -proof - - from assms have "m*n + q = q + m*n" - by (simp add: add_commute_nat) - also from assms have "... = ((q + m*n) div n) * n + (q + m*n) mod n" - by (intro sym[OF mod_div_nat_equality], simp+) - also from assms have "... = (m + q div n) * n + (q + m*n) mod n" - by simp - also from assms have "... = m*n + ((q div n) * n + (q + m*n) mod n)" - by (simp add: add_mult_distrib_right_nat add_assoc_nat) - finally have "q = (q div n) * n + (q + m*n) mod n" - using assms by simp - with q n pos have "(q div n) * n + (q + m*n) mod n = (q div n) * n + q mod n" - by simp - with assms show ?thesis by (simp del: mod_div_nat_equality) -qed - -lemma mod_nat_mult_self2 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(q + n * m) mod n = q mod n" -using assms by (simp add: mult_commute_nat) - -lemma mod_nat_mult_self3 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n + q) mod n = q mod n" -using assms by (simp add: add_commute_nat) - -lemma mod_nat_mult_self4 [simp]: - assumes "q \ Nat" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m + q) mod n = q mod n" -using assms by (simp add: add_commute_nat) - -lemma div_nat_mult_self1_is_id [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n) div n = m" -using assms div_nat_mult_self1 [of 0 m n] by simp - -lemma div_nat_mult_self2_is_id [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m) div n = m" -using assms div_nat_mult_self2[of 0 n m] by simp - -lemma mod_nat_mult_self1_is_0 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m * n) mod n = 0" -using assms mod_nat_mult_self1 [of 0 m n] by simp - -lemma mod_nat_mult_self2_is_0 [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(n * m) mod n = 0" -using assms mod_nat_mult_self2 [of 0 m n] by simp - -lemma div_nat_by_1 [simp]: - assumes "m \ Nat" - shows "m div 1 = m" -using assms div_nat_mult_self1_is_id [of m 1] by simp - -lemma mod_nat_by_1 [simp]: - assumes "m \ Nat" - shows "m mod 1 = 0" -using assms mod_nat_mult_self1_is_0[of m 1] by simp - -lemma mod_nat_self [simp]: - assumes "n \ Nat" and "0 < n" - shows "n mod n = 0" -using assms mod_nat_mult_self1_is_0[of 1] by simp - -lemma div_nat_self [simp]: - assumes "n \ Nat" and "0 < n" - shows "n div n = 1" -using assms div_nat_mult_self1_is_id [of 1 n] by simp - -lemma div_nat_add_self1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m + n) div n = m div n + 1" -using assms div_nat_mult_self1[OF m oneIsNat n pos] by simp - -lemma div_nat_add_self2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(n + m) div n = m div n + 1" -using assms div_nat_mult_self3[OF m n oneIsNat pos] by simp - -lemma mod_nat_add_self1 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m + n) mod n = m mod n" -using assms mod_nat_mult_self1[OF m oneIsNat n pos] by simp - -lemma mod_nat_add_self2 [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(n + m) mod n = m mod n" -using assms mod_nat_mult_self3[OF m oneIsNat n pos] by simp - -lemma div_mod_nat_decomp: - assumes m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - obtains q r where "q \ Nat" and "r \ Nat" - and "q = m div n" and "r = m mod n" and "m = q * n + r" -proof - - from m n pos have "m = (m div n) * n + (m mod n)" by simp - with assms that show ?thesis by blast -qed - -lemma dvd_nat_eq_mod_eq_0: - assumes "m \ Nat" and "n \ Nat" and "0 < m" - shows "(m dvd n) = (n mod m = 0)" (is "?lhs = ?rhs") -proof - - from assms have 1: "?lhs \ ?rhs" by auto - have 2: "?rhs \ ?lhs" - proof - assume mod: "n mod m = 0" - with assms mod_div_nat_equality[of n m] have "(n div m) * m = n" by simp - with assms have "n = m * (n div m)" by (simp add: mult_commute_nat) - with assms show "m dvd n" by blast - qed - from 1 2 assms show ?thesis by blast -qed - -lemma mod_div_nat_trivial [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m mod n) div n = 0" -proof - - from assms - have "m div n + (m mod n) div n = (m mod n + (m div n) * n) div n" - using div_nat_mult_self1[of "m mod n" "m div n" "n"] - by (auto simp: mod_nat_less_divisor) - also from assms have "... = m div n + 0" by simp - finally show ?thesis - using assms by simp -qed - -lemma mod_mod_nat_trivial [simp]: - assumes "m \ Nat" and "n \ Nat" and "0 < n" - shows "(m mod n) mod n = m mod n" -proof - - from assms mod_nat_mult_self1[of "m mod n" "m div n" "n"] - have "(m mod n) mod n = (m mod n + (m div n) * n) mod n" by simp - also from assms have "... = m mod n" by simp - finally show ?thesis . -qed - -lemma dvd_nat_imp_mod_0: - assumes "n dvd m" and "m \ Nat" and "n \ Nat" and "0 < n" - shows "m mod n = 0" -using assms by (simp add: dvd_nat_eq_mod_eq_0) - -lemma dvd_nat_div_mult_self: - assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "(m div n) * n = m" -using assms - dvd_nat_imp_mod_0[OF assms] - mod_div_nat_equality[OF m n pos] -by simp - -lemma dvd_nat_div_mult: - assumes dvd: "n dvd m" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - and k: "k \ Nat" - shows "(m div n) * k = (m * k) div n" -proof - - from dvd m n obtain l where l: "l \ Nat" "m = n*l" by auto - with m n k have "m * k = n * (l*k)" by (simp add: mult_assoc_nat) - with m n k l pos show ?thesis by simp -qed - -lemma div_nat_dvd_div [simp]: - assumes 1: "a dvd b" and 2: "a dvd c" - and a: "a \ Nat" and b: "b \ Nat" and c: "c \ Nat" and pos: "0 < a" - shows "(b div a) dvd (c div a) = (b dvd c)" -proof (auto) - assume lhs: "(b div a) dvd (c div a)" - with a b c pos have "((b div a) * a) dvd ((c div a) * a)" - by (intro mult_dvd_mono, simp+) - moreover - from 1 a b pos have "(b div a) * a = b" by (simp add: dvd_nat_div_mult_self) - moreover - from 2 a c pos have "(c div a) * a = c" by (simp add: dvd_nat_div_mult_self) - ultimately show "b dvd c" by simp -next - assume rhs: "b dvd c" - with b c obtain k where k: "k \ Nat" "c = b*k" by auto - from 1 a b obtain l where l: "l \ Nat" "b = a*l" by auto - with a pos have 3: "b div a = l" by simp - from 2 a c obtain m where m: "m \ Nat" "c = a*m" by auto - with a pos have 4: "c div a = m" by simp - from k l m a pos mult_assoc_nat[of a l k, symmetric] have "m = l*k" by auto - with k l m have "l dvd m" by auto - with 3 4 show "(b div a) dvd (c div a)" by simp -qed (auto simp: assms) - -lemma dvd_mod_nat_imp_dvd: - assumes 1: "k dvd (m mod n)" and 2: "k dvd n" - and k: "k \ Nat" and m: "m \ Nat" and n: "n \ Nat" and pos: "0 < n" - shows "k dvd m" -proof - - from assms have "k dvd ((m div n) * n + m mod n)" - by (simp add: dvd_mult del: mod_div_nat_equality) - with m n pos show ?thesis by simp -qed - -(*** TODO -text \ Addition respects modular equivalence. \ - -text \ Multiplication respects modular equivalence. \ - -text \ Negation respects modular equivalence. \ - -text \ Subtraction respects modular equivalence. \ -*) - -end diff --git a/isabelle/NatOrderings.thy b/isabelle/NatOrderings.thy deleted file mode 100644 index 97e45a2f..00000000 --- a/isabelle/NatOrderings.thy +++ /dev/null @@ -1,709 +0,0 @@ -(* Title: TLA+/NatOrderings.thy - Author: Hernan Vanzetto and Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - - -section \ Orders on natural numbers \ - -theory NatOrderings -imports Peano -begin - -text \ - Using the sets @{text upto} we can now define the standard ordering on - natural numbers. The constant @{text "\"} is defined over the naturals - by the axiom (conditional definition) @{text "nat_leq_def"} below; it - should be defined over other domains as appropriate later on. - - We generally define the constant @{text "<"} such that @{text "ab \ a\b"}, for any values. -\ - -definition leq :: "[c,c] \ c" (infixl "\" 50) -(*where nat_leq_def: "\m \ Nat; n \ Nat\ \ (m \ n) \ (m \ upto(n))"*) -where nat_leq_def: "(m \ n) \ (m \ upto(n))" - -abbreviation (input) - geq :: "[c,c] \ c" (infixl "\" 50) -where "x \ y \ y \ x" - -notation (ASCII) - leq (infixl "<=" 50) and - geq (infixl ">=" 50) - -lemma eq_leq_trans [trans]: "\m = n; n \ p\ \ m \ p" -by (rule ssubst) - -lemma leq_eq_trans [trans]: "\m \ n; n = p\ \ m \ p" -by (rule subst) - -subsection \ Operator definitions and generic facts about @{text "<"} \ - -definition less :: "[c,c] \ c" (infixl "<" 50) -where "a < b \ a \ b \ a \ b" - -abbreviation (input) - greater :: "[c,c] \ c" (infixl ">" 50) -where "x > y \ y < x" - -lemma boolify_less [simp]: "boolify(a < b) = (a < b)" -by (simp add: less_def) - -lemma less_isBool [intro!,simp]: "isBool(am = n; n < p\ \ m < p" -by (rule ssubst) - -lemma less_eq_trans [trans]: "\m < n; n = p\ \ m < p" -by (rule subst) - -lemma less_imp_leq [elim!]: "a < b \ a \ b" -unfolding less_def by simp - -lemma less_irrefl [simp]: "(a < a) = FALSE" -unfolding less_def by simp - -lemma less_irreflE [elim!]: "a < a \ R" -by simp - -lemma less_not_refl: "a < b \ a \ b" -by auto - -lemma neq_leq_trans [trans]: "a \ b \ a \ b \ a < b" -by (simp add: less_def) - -declare neq_leq_trans[simplified,trans] - -lemma leq_neq_trans [trans,elim!]: "a \ b \ a \ b \ a < b" -by (simp add: less_def) - -declare leq_neq_trans[simplified,trans] - -(* Don't add to [simp]: will be tried on all disequalities! *) -lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" -by auto - -subsection \ Facts about @{text "\"} over @{text Nat} \ - -lemma nat_boolify_leq [simp]: "boolify(m \ n) = (m \ n)" -by (simp add: nat_leq_def) - -lemma nat_leq_isBool [intro,simp]: "isBool(m \ n)" -by (simp add: nat_leq_def) - -lemma nat_leq_refl [intro,simp]: "n \ Nat \ n \ n" -unfolding nat_leq_def by (rule uptoRefl) - -lemma eq_leq_bothE: \ \reduce equality over integers to double inequality\ - assumes "m \ Nat" and "n \ Nat" and "m = n" and "\ m \ n; n \ m \ \ P" - shows "P" -using assms by simp - -lemma nat_zero_leq [simp]: "n \ Nat \ 0 \ n" -unfolding nat_leq_def by (rule zeroInUpto) - -lemma nat_leq_zero [simp]: "n \ Nat \ (n \ 0) = (n = 0)" -by (simp add: nat_leq_def uptoZero) - -lemma nat_leq_SuccI [(*elim!,*)intro]: - assumes "m \ n" (*and "m \ Nat"*) and "n \ Nat" - shows "m \ Succ[n]" -using assms by (auto simp: nat_leq_def uptoSucc) - -(* Do not make this "simp": interferes with other rules, e.g. simplifying "Succ[m] \ Succ[n]" *) -lemma nat_leq_Succ: - assumes (*"m \ Nat" and*) "n \ Nat" - shows "(m \ Succ[n]) = (m \ n \ m = Succ[n])" -using assms by (auto simp: nat_leq_def uptoSucc) - -lemma nat_is_leq_Succ [simp]: - assumes "n \ Nat" - shows "n \ Succ[n]" -using assms by (simp add: nat_leq_Succ) - -lemma nat_leq_SuccE [elim]: - assumes "m \ Succ[n]" and (*"m \ Nat" and*) "n \ Nat" - and "m \ n \ P" and "m = Succ[n] \ P" - shows "P" -using assms by (auto simp: nat_leq_Succ) - -lemma nat_leq_limit: - assumes "m \ n" and "\(Succ[m] \ n)" and (*"m \ Nat" and*) "n \ Nat" - shows "m=n" -using assms by (auto simp: nat_leq_def intro: uptoLimit) - -lemma nat_leq_trans [trans]: - assumes "k \ m" and "m \ n" (*and "k \ Nat" and "m \ Nat"*) and "n \ Nat" - shows "k \ n" -using assms by (auto simp: nat_leq_def elim: uptoTrans) - -lemma nat_leq_antisym: - assumes "m \ n" and "n \ m" (*and "m \ Nat" and "n \ Nat"*) - shows "m = n" -using assms by (auto simp add: nat_leq_def elim: uptoAntisym) - -lemma nat_Succ_not_leq_self [simp]: - assumes n: "n \ Nat" - shows "(Succ[n] \ n) = FALSE" -using n by (auto simp: nat_leq_Succ dest: nat_leq_antisym) - -lemma nat_Succ_leqI: - assumes "m \ n" and "m \ Nat" and "n \ Nat" and "m \ n" - shows "Succ[m] \ n" -using assms by (auto dest: nat_leq_limit) - -lemma nat_Succ_leqD [elim]: - assumes leq: "Succ[m] \ n" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ n" -proof - - from m have "m \ Succ[m]" by (simp add: nat_leq_Succ) - with leq m n show ?thesis by (elim nat_leq_trans, auto) -qed - -lemma nat_Succ_leq_Succ [simp]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(Succ[m] \ Succ[n]) = (m \ n)" -using m n by (auto simp: nat_leq_Succ intro: nat_leq_limit) - -lemma nat_Succ_leq: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] \ n) = (m \ n \ m \ n)" - using assms by (auto intro: nat_Succ_leqI) - -lemma nat_leq_linear: "\m \ Nat; n \ Nat\ \ m \ n \ n \ m" -unfolding nat_leq_def using uptoLinear . - -lemma nat_leq_cases: - assumes m: "m \ Nat" and n: "n \ Nat" - and leq: "m \ n \ P" and - geq: "\n \ m; - n \ m - \ \ P" - shows "P" -proof (cases "m \ n") - case True thus "P" by (rule leq) -next - case False - with m n have nm: "n \ m" by (blast dest: nat_leq_linear) - thus "P" - proof (cases "n=m") - case True - with m have "m \ n" by simp - thus "P" by (rule leq) - next - case False - with nm show "P" by (rule geq) - qed -qed - -lemma nat_leq_induct: \ \sometimes called ``complete induction''\ - assumes "P(0)" - and "\n\Nat : (\m\Nat : m \ n \ P(m)) \ P(Succ[n])" - shows "\n\Nat : P(n)" -proof - - from assms have "\n\Nat : \m\Nat : m \ n \ P(m)" - by (intro natInduct, auto simp: nat_leq_Succ) - thus ?thesis by (blast dest: nat_leq_refl) -qed - -lemma nat_leq_inductE: - assumes "n \ Nat" - and "P(0)" and "\n. \n \ Nat; \m\Nat : m \ n \ P(m)\ \ P(Succ[n])" - shows "P(n)" -using assms by (blast dest: nat_leq_induct) - - -subsection \ Facts about @{text "<"} over @{text Nat} \ - -lemma nat_less_iff_Succ_leq [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(m < n) = (Succ[m] \ n)" -using assms by (auto simp: less_def dest: nat_Succ_leqD nat_leq_limit) - -text \ Reduce @{text "\"} to @{text "<"}. \ - -lemma nat_leq_less: - assumes (*"m \ Nat" and*) "n \ Nat" - shows "m \ n = (m < n \ m = n)" - using assms by (auto simp: less_def) - -lemma nat_leqE: - assumes "m \ n" and "n \ Nat" and 1: "m P" and 2: "m=n \ P" - shows "P" -proof - - from \m \ n\ \n \ Nat\ have "m < n \ m=n" - by (auto simp add: less_def) - with 1 2 show ?thesis by blast -qed - -lemma nat_less_Succ_iff_leq (*[simp]*): - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m \ n)" - using assms by simp - -lemmas nat_leq_iff_less_Succ = sym[OF nat_less_Succ_iff_leq] - -lemma not_leq_Succ_leq [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "((m \ n) = FALSE) = (Succ[n] \ m)" -proof - - from assms nat_leq_linear[of m n] have "((m \ n) = FALSE) = (n < m)" - by (auto simp: less_def dest: nat_leq_antisym) - with assms show ?thesis by simp -qed - -lemma nat_not_leq_one: - assumes "n \ Nat" - shows "(\(1 \ n)) = (n = 0)" -using assms by simp - -text \ @{text "<"} and @{text "Succ"}. \ - -lemma nat_Succ_less_mono: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] < Succ[n]) = (m < n)" -using assms by simp - -lemma nat_Succ_less_SuccE: - assumes "Succ[m] < Succ[n]" and "m \ Nat" and "n \ Nat" and "m < n \ P" - shows "P" -using assms by simp - -lemma nat_not_less0 [simp]: - assumes "n \ Nat" - shows "(n < 0) = FALSE" -using assms by auto - -lemma nat_less0E (*[elim!]*): - assumes "n < 0" and "n \ Nat" - shows "P" -using assms by simp - -lemma nat_less_SuccI: - assumes "m < n" and "m \ Nat" and "n \ Nat" - shows "m < Succ[n]" -using assms by auto - -lemma nat_Succ_lessD: - assumes 1: "Succ[m] < n" and 2: "m \ Nat" and 3: "n \ Nat" - shows "m < n" -using assms by auto - -lemma nat_less_leq_not_leq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m < n) = (m \ n \ \ n \ m)" -using assms by (auto simp: less_def dest: nat_leq_antisym) - -text \ Transitivity. \ - -lemma nat_less_trans (*[trans]*): - assumes "k < m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -lemma nat_less_trans_Succ [trans]: - assumes lt1: "i < j" and lt2: "j < k" - and i: "i \ Nat" and j: "j \ Nat" and k: "k \ Nat" - shows "Succ[i] < k" -proof - - from i j lt1 have "Succ[Succ[i]] \ Succ[j]" by simp - also from j k lt2 have "Succ[j] \ k" by simp - finally show ?thesis using i j k by simp -qed - -lemma nat_leq_less_trans [trans]: - assumes "k \ m" and "m < n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -lemma nat_less_leq_trans [trans]: - assumes "k < m" and "m \ n" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "k < n" -using assms by (auto simp: less_def dest: nat_leq_trans nat_leq_antisym) - -text \ Asymmetry. \ - -lemma nat_less_not_sym: - assumes "m < n" and "m \ Nat" and "n \ Nat" - shows "(n < m) = FALSE" -using assms by auto - -lemma nat_less_asym: - assumes "m < n" and "m \ Nat" and "n \ Nat" and "\P \ n < m" - shows "P" -proof (rule contradiction) - assume "\P" with nat_less_not_sym[OF \m \m \ Nat\ \n \ Nat\] \\P \ n < m\ - show "FALSE" by simp -qed - -text \ Linearity (totality). \ - -lemma nat_less_linear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m < n \ m = n \ n < m" -unfolding less_def using nat_leq_linear[OF m n] by blast - -lemma nat_leq_less_linear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n \ n < m" -using assms nat_less_linear[OF m n] by (auto simp: less_def) - -lemma nat_less_cases [case_names less equal greater]: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m < n \ P) \ (m = n \ P) \ (n < m \ P) \ P" -using nat_less_linear[OF m n] by blast - -lemma nat_not_less: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (n \ m)" -using assms nat_leq_linear[OF m n] by (auto simp: less_def dest: nat_leq_antisym) - -lemma nat_not_less_iff_gr_or_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (m > n \ m = n)" -unfolding nat_not_less[OF m n] using assms by (auto simp: less_def) - -lemma nat_not_less_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m < n) = (n < Succ[m])" -unfolding nat_not_less[OF m n] using assms by simp - -lemma nat_not_leq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m \ n) = (n < m)" -using assms by simp - -text \often useful, but not active by default\ -lemmas nat_not_order_simps[simplified] = nat_not_less nat_not_leq - -lemma nat_not_leq_eq: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(\ m \ n) = (Succ[n] \ m)" -unfolding nat_not_leq[OF m n] using assms by simp - -lemma nat_neq_iff: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ n = (m < n \ n < m)" -using assms nat_less_linear[OF m n] by auto - -lemma nat_neq_lessE: - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "(m < n \ R) \ (n < m \ R) \ R" -using assms by (auto simp: nat_neq_iff[simplified]) - -lemma nat_antisym_conv1: - assumes "\(m < n)" and "m \ Nat" and "n \ Nat" - shows "(m \ n) = (m = n)" -using assms by (auto simp: nat_leq_less) - -lemma nat_antisym_conv2: - assumes "m \ n" and "m \ Nat" and "n \ Nat" - shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -lemma nat_antisym_conv3: - assumes "\ n < m" and "m \ Nat" and "n \ Nat" - shows "(\ m < n) = (m = n)" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -lemma nat_not_lessD: - assumes "\(m < n)" and "m \ Nat" and "n \ Nat" - shows "n \ m" -using assms by (simp add: nat_not_order_simps) - -lemma nat_not_lessI: - assumes "n \ m" and "m \ Nat" and "n \ Nat" - shows "\(m < n)" -using assms by (simp add: nat_not_order_simps) - -lemma nat_gt0_not0 (*[simp]*): - assumes "n \ Nat" - shows "(0 < n) = (n \ 0)" -using assms by (auto simp: nat_neq_iff[simplified]) - -lemmas nat_neq0_conv = sym[OF nat_gt0_not0] - -text \ Introduction properties \ - -lemma nat_less_Succ_self (*[iff]*): - assumes "n \ Nat" - shows "n < Succ[n]" -using assms by simp - -lemma nat_zero_less_Succ (*[iff]*): - assumes "n \ Nat" - shows "0 < Succ[n]" -using assms by simp - -text \ Elimination properties \ - -lemma nat_less_Succ: - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m < n \ m = n)" -using assms by (simp add: nat_leq_less[of n m]) - -lemma nat_less_SuccE: - assumes "m < Succ[n]" and "m \ Nat" and "n \ Nat" - and "m < n \ P" and "m = n \ P" - shows P -using assms by (auto simp: nat_leq_less[of n m]) - -lemma nat_less_one (*[simp]*): - assumes "n \ Nat" - shows "(n < 1) = (n = 0)" -using assms by simp - -text \ "Less than" is antisymmetric, sort of \ - -lemma nat_less_antisym: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "\ \(n < m); n < Succ[m] \ \ m = n" -using assms by (auto simp: nat_not_order_simps elim: nat_leq_antisym) - -text \ Lifting @{text "<"} monotonicity to @{text "\"} monotonicity \ -lemma less_mono_imp_leq_mono: - assumes i: "i \ Nat" and j: "j \ Nat" and f: "\n \ Nat : f(n) \ Nat" - and ij: "i \ j" and mono: "\i j. \i \ Nat; j \ Nat; i \ f(i) < f(j)" - shows "f(i) \ f(j)" -using ij proof (auto simp: nat_leq_less[OF j]) - assume "i < j" - with i j have "f(i) < f(j)" by (rule mono) - thus "f(i) \ f(j)" by (simp add: less_imp_leq) -next - from j f show "f(j) \ f(j)" by auto -qed - -text \ Inductive (?) properties. \ - -lemma nat_Succ_lessI: - assumes "m \ Nat" and "n \ Nat" and "m < n" and "Succ[m] \ n" - shows "Succ[m] < n" -using assms by (simp add: leq_neq_iff_less[simplified]) - -lemma nat_lessE: - assumes major: "i < k" and i: "i \ Nat" and k: "k \ Nat" - obtains j where "j \ Nat" and "i \ j" and "k = Succ[j]" -proof - - from k major have "\j\Nat : i \ j \ k = Succ[j]" - proof (induct k) - case 0 with i show ?case by simp - next - fix n - assume n: "n \ Nat" and 1: "i < Succ[n]" - and ih: "i < n \ \j\Nat : i \ j \ n = Succ[j]" - from 1 i n show "\j\Nat : i \ j \ Succ[n] = Succ[j]" - proof (rule nat_less_SuccE) - assume "i < n" - then obtain j where "j \ Nat" and "i \ j" and "n = Succ[j]" - by (blast dest: ih) - with i have "Succ[j] \ Nat" and "i \ Succ[j]" and "Succ[n] = Succ[Succ[j]]" - by auto - thus ?thesis by blast - next - assume "i = n" - with i show ?thesis by blast - qed - qed - with that show ?thesis by blast -qed - -lemma nat_Succ_lessE: - assumes major: "Succ[i] < k" and i: "i \ Nat" and k: "k \ Nat" - obtains j where "j \ Nat" and "i < j" and "k = Succ[j]" -proof - - from i have "Succ[i] \ Nat" .. - from major this k obtain j where "j \ Nat" "Succ[i] \ j" "k = Succ[j]" - by (rule nat_lessE) - with i that show ?thesis by simp -qed - -(* -lemma nat_gt0_implies_Succ: - assumes 1: "0 < n" and 2: "n \ Nat" - shows "\m : m \ Nat \ n = Succ[m]" -using 2 1 by (cases, auto) -*) -lemma nat_ge1_implies_Succ: - assumes 1: "1 \ n" and 2: "n \ Nat" - shows "\m : m \ Nat \ n = Succ[m]" - using 2 1 by (cases, auto) - -lemma nat_gt0_iff_Succ: - assumes n: "n \ Nat" - shows "(1 \ n) = (\m \ Nat: n = Succ[m])" -using n by (auto dest: nat_ge1_implies_Succ) - -(* -lemma nat_less_Succ_eq_0_disj: - assumes "m \ Nat" and "n \ Nat" - shows "(m < Succ[n]) = (m = 0 \ (\j\Nat: m = Succ[j] \ j < n))" -using assms by (induct m, auto) -*) -lemma nat_le_eq_0_disj: - assumes "m \ Nat" and "n \ Nat" - shows "(m \ n) = (m = 0 \ (\j\Nat : m = Succ[j] \ j < n))" - using assms by (induct m, auto) - -lemma nat_less_antisym_false: "\m < n; m \ Nat; n \ Nat\ \ n < m = FALSE" - by auto - -lemma nat_less_antisym_leq_false: "\m < n; m \ Nat; n \ Nat\ \ n \ m = FALSE" - by auto - - -subsection \Intervals of natural numbers\ - -definition natInterval :: "[c,c] \ c" ("(_ .. _)" [90,90] 70) -where "m .. n \ { k \ Nat : m \ k \ k \ n }" - -lemma inNatIntervalI [intro!,simp]: - assumes "k \ Nat" and "m \ k" and "k \ n" - shows "k \ m .. n" -using assms by (simp add: natInterval_def) - -lemma inNatIntervalE [elim]: - assumes 1: "k \ m .. n" and 2: "\k \ Nat; m \ k; k \ n\ \ P" - shows P -using 1 by (intro 2, auto simp add: natInterval_def) - -lemma inNatInterval_iff: "(k \ m .. n) = (k \ Nat \ m \ k \ k \ n)" -by auto - -lemmas - setEqualI [where A = "m .. n" for m n, intro] - setEqualI [where B = "m .. n" for m n, intro] - -lemma lowerInNatInterval [iff]: - assumes "m \ n" and "m \ Nat" - shows "m \ m .. n" -using assms by (simp add: natInterval_def) - -lemma upperInNatInterval [iff]: - assumes "m \ n" and "n \ Nat" - shows "n \ m .. n" -using assms by (simp add: natInterval_def) - -lemma gtNotinNatInterval: - assumes gt: "m > n" and k: "k \ m .. n" and m: "m \ Nat" and n: "n \ Nat" - shows "P" -proof - - from k have 1: "m \ k" and 2: "k \ n" by auto - from 1 2 n have "m \ n" by (rule nat_leq_trans) - with m n have "\(n < m)" by (simp add: nat_not_order_simps) - from this gt show ?thesis .. -qed - -lemma natIntervalIsEmpty: - assumes "m \ Nat" and "n \ Nat" and "m > n" - shows "m .. n = {}" -using assms by (blast dest: gtNotinNatInterval) - -lemma natIntervalEmpty_iff: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "(m .. n = {}) = (m > n)" -proof (auto dest: natIntervalIsEmpty[OF m n]) - assume mt: "m .. n = {}" - show "n < m" - proof (rule contradiction) - assume "\(n < m)" - with m n have "m \ n" by (simp add: nat_not_order_simps) - from this m have "m \ m .. n" by (rule lowerInNatInterval) - with mt show "FALSE" by blast - qed -qed - -lemma natIntervalSingleton [simp]: - assumes "n \ Nat" - shows "n .. n = {n}" -using assms by (auto dest: nat_leq_antisym) - -lemma natIntervalSucc [simp]: - assumes "m \ Nat" and "n \ Nat" and "m \ Succ[n]" - shows "m .. Succ[n] = addElt(Succ[n], m .. n)" -using assms by (auto simp: natInterval_def) - -lemma succNatInterval: - assumes "m \ Nat" and "n \ Nat" - shows "Succ[m] .. n = (m .. n \ {m})" -using assms by (auto simp: natInterval_def nat_Succ_leq) - -lemma natIntervalEqual_iff: - assumes k: "k \ Nat" and l: "l \ Nat" and m: "m \ Nat" and n: "n \ Nat" - shows "(k .. l = m .. n) = ((k > l \ m > n) \ (k = m \ l = n))" (is "?lhs = ?rhs") -proof - - have 1: "?lhs \ ?rhs" - proof - assume eq: "?lhs" - show ?rhs - proof (cases "k .. l = {}") - case True - with k l have "k > l" by (simp only: natIntervalEmpty_iff) - moreover - from True eq m n have "m > n" by (simp only: natIntervalEmpty_iff) - ultimately - show ?rhs by blast - next - case False - with k l have 11: "k \ l" by (simp only: natIntervalEmpty_iff nat_not_less) - from False eq m n have 12: "m \ n" by (simp only: natIntervalEmpty_iff nat_not_less) - from 11 k eq have 13: "m \ k" by auto - from 12 m eq have 14: "k \ m" by auto - from 14 13 have 15: "k = m" by (rule nat_leq_antisym) - from 11 l eq have 16: "l \ n" by auto - from 12 n eq have 17: "n \ l" by auto - from 16 17 have "l = n" by (rule nat_leq_antisym) - with 15 show ?rhs by blast - qed - qed - have 2: "?rhs \ ?lhs" - proof auto - assume lk: "l < k" and nm: "n < m" - from k l lk have "k .. l = {}" by (rule natIntervalIsEmpty) - moreover - from m n nm have "m .. n = {}" by (rule natIntervalIsEmpty) - ultimately - show ?lhs by auto - qed - from 1 2 show ?thesis by blast -qed - -lemma zerotoInj [simp]: - assumes "l \ Nat" and "m \ Nat" and "n \ Nat" - shows "(0 .. l = m .. n) = (m=0 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma zerotoInj' [simp]: - assumes "k \ Nat" and "l \ Nat" and "n \ Nat" - shows "(k .. l = 0 .. n) = (k=0 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma zerotoEmpty [simp]: - assumes "m \ Nat" - shows "Succ[m] .. 0 = {}" -using assms by auto - -lemma onetoInj [simp]: - assumes "l \ Nat" and "m \ Nat" and "n \ Nat" and "l\0 \ m=1" - shows "(1 .. l = m .. n) = (m=1 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma onetoInj' [simp]: - assumes "k \ Nat" and "l \ Nat" and "n \ Nat" and "n\0 \ k=1" - shows "(k .. l = 1 .. n) = (k=1 \ l=n)" -using assms by (auto simp: natIntervalEqual_iff) - -lemma SuccInNatIntervalSucc: - assumes "m \ k" and "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(Succ[k] \ m .. Succ[n]) = (k \ m .. n)" -using assms by auto - -lemma SuccInNatIntervalSuccSucc: - assumes "k \ Nat" and "m \ Nat" and "n \ Nat" - shows "(Succ[k] \ Succ[m] .. Succ[n]) = (k \ m .. n)" -using assms by auto - -end diff --git a/isabelle/NewIntegers.thy b/isabelle/NewIntegers.thy deleted file mode 100644 index 103b7221..00000000 --- a/isabelle/NewIntegers.thy +++ /dev/null @@ -1,12656 +0,0 @@ -(* An attempt at defining the integers. - -The steps in Isabelle/Isar are named `si_j` where `s` stands for "step", -`i` is the proof level, and `j` the step number. - -Author: Ioannis Filippidis -All rights reserved. Licensed under 2-clause BSD. -*) -theory NewIntegers -imports - NatArith - (*NatDivision - Tuples*) -begin - -(* -Properties to prove: - typeness, - commutativity, associativity, distributivity - transitivity, symmetry, antisymmetry, reflexivity, - monotonicity -*) - -(*----------------------------------------------------------------------------*) -(* Overriding notation. *) - -no_notation NatOrderings.leq (infixl "<=" 50) -no_notation NatOrderings.geq (infixl ">=" 50) -no_notation NatOrderings.less (infixl "<" 50) -no_notation NatOrderings.greater (infixl ">" 50) -no_notation NatOrderings.natInterval ("(_ .. _)" [90,90] 70) - -no_notation NatArith.arith_add (infixl "+" 65) -(* no_notation NatArith.adiff (infixl "--" 65) *) -no_notation NatArith.mult (infixl "*" 70) - -(* no_notation NatDivision.dvd (infixl "dvd" 50) *) - - -(*----------------------------------------------------------------------------*) -(* The minus operator. *) - - -(* -The minus operator can be defined, instead of axiomatized, as follows. - -minus(x) == - LET singletons == {<>: u \in Nat} - IN - IF x \in Nat THEN - <> - ELSE IF x \in Singletons THEN - CHOOSE u \in Nat: x = <> - ELSE - CHOOSE u \notin (Nat \cup Singletons): TRUE -*) -consts - "minus" :: "[c] \ c" ("-._" [75] 75) - -axiomatization where - neg0 [simp]: "-.0 = 0" -and - neg_neg: "n \ Nat \ -.n \ Nat ==> - -.-.n = n" -and - neg_not_in_nat: - "n \ (Nat \ {0}) ==> - -.n \ Nat" - - -(*----------------------------------------------------------------------------*) -(* The set of integers. *) - -definition Int -where "Int \ - Nat \ {-.n: n \ Nat}" - - -definition negNat :: "c" -where "negNat \ {-.n: n \ Nat}" - - -(*----------------------------------------------------------------------------*) -(* Successor and predecessor on integers. *) - -definition iSucc :: "c" -where "iSucc \ - [i \ Int \ IF i \ Nat THEN Succ[i] - ELSE -.Pred[-.i]]" - - -definition iPred :: "c" -where "iPred \ - [i \ Int \ IF i \ Nat \ {0} - THEN Pred[i] - ELSE -.Succ[-.i]]" - - -(*----------------------------------------------------------------------------*) -(* Recursive definitions on the integers. *) - -(* TODO: replace axiom with proof of existence. *) -axiomatization where - primrec_int: " - \ f: - isAFcn(f) \ - DOMAIN f = Int \ - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n])) - " - -(*----------------------------------------------------------------------------*) -(* Arithmetic operators. *) - - -(* Incrementing and decrementing function. *) -definition addint :: "[c] \ c" -where "addint(m) \ - CHOOSE g \ [Int -> Int]: - g[0] = m \ - (\ n \ Nat: - g[Succ[n]] = iSucc[g[n]] \ - g[-.Succ[n]] = iPred[g[-.n]])" - - -(* Addition on integers. *) -definition add :: "[c, c] \ c" (infixl "+" 65) -where "add(m, n) \ addint(m)[n]" - - -(* Subtraction on integers. *) -definition diff :: "[c, c] \ c" (infixl "-" 65) -where "diff(m, n) \ add(m, -.n)" - - -(* Adding and subtracting function. *) -definition multint :: "[c] \ c" -where "multint(m) \ - CHOOSE g \ [Int -> Int]: - g[0] = 0 \ - (\ n \ Nat: - g[Succ[n]] = add(g[n], m) \ - g[-.Succ[n]] = add(g[-.n], -.m))" - - -(* Multiplication on integers. *) -definition mult :: "[c,c] \ c" (infixl "*" 70) -where "mult(m, n) \ multint(m)[n]" - - -(* Divisibility on the integers. *) -definition dvd :: "[c, c] \ c" (infixl "dvd" 50) -where "b dvd a \ - \ k \ Int: a = mult(b, k)" - - -(* Order on the integers. -This definition is motivated by proofs, where equality is easier to -reason about (adding on both sides, multiplying both sides). -*) -definition leq :: "[c,c] \ c" (infixl "<=" 50) -where "leq(m, n) \ - \ k \ Nat: - add(m, k) = n" - - -abbreviation (input) - geq :: "[c, c] \ c" (infixl ">=" 50) - where "geq(x, y) \ leq(y, x)" - - -definition less :: "[c, c] \ c" (infixl "<" 50) -where "less(a, b) \ leq(a, b) - \ (a \ b)" - - -abbreviation (input) - greater :: "[c, c] \ c" (infixl ">" 50) - where "greater(x, y) \ less(y, x)" - - -definition interval :: "[c, c] \ c" ("(_ .. _)" [90,90] 70) -where "m .. n \ - {k \\in Int: m <= k \ k <= n}" - -(*----------------------------------------------------------------------------*) - - -theorem minus_involutive[simp]: - assumes "n \ Int" - shows "-.-.n = n" - proof - - have s1_1: "n \ Nat \ (\ k \ Nat: n = -.k)" - using assms by (auto simp: Int_def) - have s1_2: "n \ Nat ==> -.-.n = n" - using neg_neg by auto - have s1_3: "n \ Nat ==> -.-.n = n" - proof - - { - assume s2_1: "n \ Nat" - have s2_2: "\ k \ Nat: n = -.k" - using s1_1 s2_1 by auto - have s2_3: "\ k \ Nat: -.n = -.-.k" - using s2_2 by auto - have s2_4: "\ k \ Nat: k = -.n" - using s2_3 neg_neg by auto - have s2_5: "-.n \ Nat" - using s2_4 by auto - have s2_6: "-.-.n = n" - using s2_5 neg_neg by auto - } - from this show "n \ Nat ==> -.-.n = n" by auto - qed - show "-.-.n = n" using s1_2 s1_3 by iprover - qed - - -theorem minus_injective[dest]: - assumes "-.n = -.m" and "n \ Int \ m \ Int" - shows "n = m" - proof - - have "<1>1": "-.-.n = -.-.m" - using assms by simp - have "<1>2": "-.-.n = n" - proof - - have "<2>1": "n \ Int ==> -.-.n = n" - using minus_involutive by auto - have "<2>2": "n \ Int" - using assms by auto - show "-.-.n = n" - using "<2>1" "<2>2" by auto - qed - have "<1>3": "-.-.m = m" - proof - - have "<2>1": "m \ Int ==> -.-.m = m" - using minus_involutive by auto - have "<2>2": "m \ Int" - using assms by auto - show "-.-.m = m" - using "<2>1" "<2>2" by auto - qed - show "n = m" - using "<1>1" "<1>2" "<1>3" by auto - qed - - -theorem minus_eq: - assumes "x = y" - shows "-.x = -.y" - using assms by auto - - -theorem neg_succ_not_in_nat[simp]: - "n \ Nat ==> - -.(Succ[n]) \ Nat" - proof - - { - fix "n" - assume "<2>1": "n \ Nat" - have "<2>2": "Succ[n] \ (Nat \ {0})" - using "<2>1" by auto - have "-.(Succ[n]) \ Nat" - using "<2>2" neg_not_in_nat by auto - } - from this show "n \ Nat ==> - -.(Succ[n]) \ Nat" - by auto - qed - - -theorem neg0_iff_eq0[simp]: - "(-.n = 0) = (n = 0)" - proof auto - assume "<1>1": "-.n = 0" - have "<1>2": "-.n \ Nat" - using "<1>1" by auto - have "<1>3": "-.-.n = n" - using "<1>2" neg_neg by auto - have "<1>4": "-.-.n = 0" - using "<1>1" by auto - show "n = 0" - using "<1>3" "<1>4" by simp - qed - - -theorem neg0_imp_eq0[dest]: - "(-.n = 0) ==> (n = 0)" - by simp - - -theorem not_neg0_imp_not0[dest]: - "(-.n \ 0) ==> (n \ 0)" - by simp - - -(*----------------------------------------------------------------------------*) -(* Properties of the set `Int`. *) - - -theorem nat_in_int[simp]: - "n \ Nat ==> n \ Int" - by (simp add: Int_def) - - -theorem minus_nat_in_int: - "n \ Nat ==> -.n \ Int" - by (auto simp: Int_def) - - -theorem minus_nat_in_neg_nat: - "n \ Nat ==> -.n \ negNat" - by (auto simp: negNat_def) - - -theorem neg_nat_in_int: - "n \ negNat ==> n \ Int" - by (auto simp: negNat_def Int_def) - - -theorem int_disj: - "n \ Int ==> - n \ Nat - \ n \ {-.n: n \ Nat}" - by (auto simp: Int_def) - - -theorem int_disj_nat_negnat: - "n \ Int ==> - n \ Nat - \ n \ negNat" - unfolding negNat_def - using int_disj by auto - - -theorem int_union_nat_negnat: - "Int = Nat \ negNat" - using int_disj_nat_negnat nat_in_int - neg_nat_in_int - by auto - - -theorem int_union_nat_0_negnat: - "Int = (Nat \ {0}) \ negNat" - proof - - have s1_1: "Int = Nat \ negNat" - using int_union_nat_negnat by auto - have s1_2: "0 \ negNat" - unfolding negNat_def - using zeroIsNat neg0 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem int_union_nat_negnat_0: - "Int = Nat \ (negNat \ {0})" - proof - - have s1_1: "Int = Nat \ negNat" - using int_union_nat_negnat by auto - have s1_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem neg_int_eq_int[simp]: - "n \ Int ==> -.n \ Int" - unfolding Int_def by auto - - -theorem minus_negnat_in_int: - "n \ negNat ==> -.n \ Int" - using neg_nat_in_int neg_int_eq_int by auto - - -theorem minus_neg_int_in_nat: - "n \ Int \ n \ Nat ==> - -.n \ Nat \ {0}" - proof - - assume "<1>1": "n \ Int \ n \ Nat" - have "<1>2": "n \ {-.k: k \ Nat}" - using "<1>1" int_disj by iprover - have "<1>3": "\ k \ Nat: n = -.k" - using "<1>2" by auto - have "<1>4": "\ k \ Nat: k = -.n" - using "<1>3" minus_involutive by auto - have "<1>5": "-.n \ Nat" - using "<1>4" by auto - have "<1>6": "-.n \ 0" - proof - - { - assume "<2>1": "-.n = 0" - have "<2>2": "n = 0" - using neg0 "<2>1" by auto - have "<2>3": "n \ Nat" - using "<2>2" by auto - have "<2>4": "n \ Nat" - using "<1>1" by auto - have "<2>5": "FALSE" - using "<2>3" "<2>4" by auto - } - from this show "-.n \ 0" - by auto - qed - show "-.n \ Nat \ {0}" - using "<1>5" "<1>6" by auto - qed - - -theorem neg_negative_in_nat: - assumes hyp: "n \ Int \ n \ Nat" - shows "-. n \ Nat" - proof - - have int: "n \ Int" - using hyp by auto - from int int_disj have disj: - "n \ Nat \ n \ {-. m: m \ Nat}" - by auto - from disj hyp have neg: "n \ {-. m: m \ Nat}" - by auto - from neg have ex1: "\ m \ Nat: n = -. m" - by auto - from ex1 have ex2: "\ m \ Nat: -. n = -. -. m" - by auto - from ex2 neg_neg have ex3: "\ m \ Nat: -. n = m" - by auto - from ex3 show "-. n \ Nat" - by auto - qed - - -theorem minus_nat_0_or_not_in_nat: - "n \ Nat ==> - -.n = 0 \ -.n \ Nat" - using neg0 neg_not_in_nat by auto - - -theorem int_eq_neg_int_is_0[simp]: - "n \ Int ==> (n = -.n) = (n = 0)" - proof - - have "<1>1": "n \ Nat \ {0} ==> - -.n \ Nat" - proof - - assume "<2>1": "n \ Nat \ {0}" - show "-.n \ Nat" - using neg_not_in_nat "<2>1" by auto - qed - have "<1>2": "n \ Nat \ {0} ==> - n \ -.n" - using "<1>1" by force - have "<1>3": "n \ Int \ n \ Nat - ==> -.n \ Nat" - using minus_neg_int_in_nat by auto - have "<1>4": "n \ Int \ n \ Nat - ==> n \ -.n" - using "<1>3" by force - have "<1>5": "n \ Int \ {0} ==> - (n \ Nat \ {0}) \ - (n \ {-.k: k \ Nat})" - using int_disj by auto - have "<1>6": "n \ Int \ {0} ==> - (n \ -.n)" - using "<1>2" "<1>4" "<1>5" by auto - have "<1>7": "n \ Int \ (n \ 0) ==> - (n \ -.n)" - using "<1>6" by auto - have "<1>8": "n \ Int \ (n = -.n) ==> - (n = 0)" - using "<1>7" by auto - have "<1>9": "n \ Int \ (n = 0) ==> - (n = -.n)" - using neg0 by auto - show "n \ Int ==> (n = -.n) = (n = 0)" - using "<1>8" "<1>9" by auto - qed - - -theorem minus_neg_nat_in_nat: - assumes hyp: "n \ negNat" - shows "-.n \ Nat" - proof - - have s1_1: "\ k \ Nat: -.k = n" - using hyp by (auto simp: negNat_def) - have s1_2: "\ k \ Nat: -.-.k = -.n" - using s1_1 by auto - have s1_3: "\ k \ Nat: k = -.n" - using s1_2 neg_neg by auto - show "-.n \ Nat" - using s1_3 by auto - qed - - -theorem minus_neg_nat_0_in_nat_0: - assumes hyp: "n \ negNat \ {0}" - shows "-.n \ Nat \ {0}" - proof - - have s1_1: "-.n \ Nat" - using hyp minus_neg_nat_in_nat by auto - have s1_2: "-.n \ 0" - proof - - { - assume s2_1: "-.n = 0" - have s2_2: "-.-.n = -.0" - using s2_1 by auto - have s2_3: "-.-.n = 0" - using s2_2 neg0 by auto - have s2_4: "n = 0" - proof - - have s3_1: "n \ Int" - using hyp neg_nat_in_int by auto - have s3_2: "-.-.n = n" - using s3_1 minus_involutive - by auto - show ?thesis - using s3_2 s2_3 by auto - qed - have s2_5: "n \ 0" - using hyp by auto - have "FALSE" - using s2_4 s2_5 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem neg_nat_0_not_in_nat: - assumes hyp: "n \ negNat \ {0}" - shows "n \ Nat" - proof auto - assume s1_1: "n \ Nat" - have s1_2: "n \ 0" - using hyp by auto - have s1_3: "n \ Nat \ {0}" - using s1_1 s1_2 by auto - have s1_4: "-.n \ Nat" - using s1_3 neg_not_in_nat by auto - have s1_5: "-.n \ Nat" - using hyp minus_neg_nat_in_nat by auto - show "FALSE" - using s1_4 s1_5 by auto - qed - - -theorem neg_nat_not_in_nat_setminus_0: - assumes hyp: "n \ negNat" - shows "n \ Nat \ {0}" - proof - - have s1_1: "n \ negNat ==> - n \ negNat \ {0} - \ n = 0" - by auto - have s1_2: "n \ negNat \ {0} - ==> n \ Nat \ {0}" - using neg_nat_0_not_in_nat by auto - have s1_3: "n = 0 ==> - n \ Nat \ {0}" - by auto - show "n \ Nat \ {0}" - using hyp s1_1 s1_2 s1_3 by iprover - qed - - -theorem minus_nat_0_in_negnat_0: - assumes hyp: "n \ Nat \ {0}" - shows "-.n \ negNat \ {0}" - proof - - have s1_1: "n \ Nat" - using hyp by auto - have s1_2: "-.n \ negNat" - using s1_1 minus_nat_in_neg_nat by auto - have s1_3: "-.n \ 0" - proof - - { - assume s2_1: "-.n = 0" - have s2_2: "-.-.n = -.0" - using s2_1 by auto - have s2_3: "-.-.n = 0" - using s2_2 by auto - have s2_4: "n = 0" - using s1_1 neg_neg s2_3 by auto - have s2_5: "n \ 0" - using hyp by auto - have "FALSE" - using s2_4 s2_5 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_2 s1_3 by auto - qed - - -theorem minus_one_in_negnat: - "-.1 \\in negNat" - unfolding negNat_def - using oneIsNat by auto - - -theorem minus_one_in_int: - "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - - -theorem minus_one_neq_0: - "-.1 \ 0" - proof - - { - assume s1_1: "-.1 = 0" - have s1_2: "-.-.1 = -.0" - using s1_1 by auto - have s1_3: "-.-.1 = 0" - using s1_2 neg0 by auto - have s1_4: "0 \ Nat" - by auto - have s1_5: "1 = 0" - using s1_3 s1_4 by auto - have "FALSE" - using s1_5 succIrrefl - by auto - } - from this show "-.1 \ 0" - by auto - qed - - -theorem minus_minus_one: - "-.-.1 = 1" - proof - - have s1_1: "1 \ Nat" - using oneIsNat by auto - show "-.-.1 = 1" - using neg_neg s1_1 by auto - qed - - -theorem succ_zero_equality: - assumes hyp: "x = Succ[0]" - shows "x = 1" - using hyp by auto - - -theorem pred_1: - "Pred[1] = 0" - proof - - have s1_1: "1 = Succ[0]" - by auto - have s1_2: "\ x \ Nat: (1 = Succ[x]) => (x = 0)" - using s1_1 succInjIff by auto - show ?thesis - using s1_1 s1_2 choose_equality by auto - qed - - -theorem pred_2: - "Pred[2] = 1" - proof - - have s1_1: "1 \\in Nat \ Succ[1] = 2" - unfolding two_def - using oneIsNat by auto - have s1_2: "\ x. x \\in Nat \ (Succ[x] = 2) ==> (x = 1)" - using oneIsNat succInjIff by (auto simp: two_def) - have s1_4: "Pred[2] = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" - proof - - have s2_1: "Pred[2] = - (IF 2 = 0 THEN 0 ELSE CHOOSE x \ Nat: 2 = Succ[x])" - unfolding Pred_def - using twoIsNat by auto - also have s2_2: "... = (CHOOSE x \ Nat: 2 = Succ[x])" - proof - - have s3_1: "2 \ 0" - using succNotZero by (auto simp: two_def) - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = (CHOOSE x: x \\in Nat \ 2 = Succ[x])" - unfolding bChoose_def - by auto - also have s2_4: "... = (CHOOSE x: x \\in Nat \ Succ[x] = 2)" - proof - - have s3_1: "\ x. (2 = Succ[x]) = (Succ[x] = 2)" - by auto - show ?thesis - using s3_1 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s1_1 s1_2 s1_4 choose_equality[ - of "\ x. x \\in Nat \ Succ[x] = 2" - "1" "1"] - by auto - qed - - -theorem zero_in_int: - "0 \ Int" - using zeroIsNat nat_in_int by auto - - -theorem ipred_0: - shows "iPred[0] = -.1" - proof - - have s1_1: "iPred[0] = -.Succ[-.0]" - proof - - have s2_1: "0 \ Int" - using zero_in_int by auto - have s2_2: "0 \ Nat \ {0}" - by auto - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_2: "... = -.Succ[0]" - using neg0 by auto - also have s1_3: "... = -.1" - by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed - - -theorem isucc_minus_1: - "iSucc[-.1] = 0" - unfolding iSucc_def - using oneIsNat neg_neg pred_1 neg0 by auto - - -theorem negnat_succ_in_nat: - assumes ndom: "n \\in negNat" - shows "Succ[-.n] \\in Nat" - using ndom minus_neg_nat_in_nat succIsNat by auto - - -theorem minus_negnat_succ_in_negnat: - assumes ndom: "n \\in negNat" - shows "-.Succ[-.n] \\in negNat" - using ndom negnat_succ_in_nat minus_nat_in_neg_nat by auto - - -theorem minus_succ_minus_negnat_in_int: - assumes ndom: "n \\in negNat" - shows "-.Succ[-.n] \\in Int" - proof - - have s1_1: "-.n \\in Nat" - using ndom minus_neg_nat_in_nat by auto - have s1_2: "Succ[-.n] \\in Nat" - using s1_1 succIsNat by auto - show ?thesis - using s1_2 minus_nat_in_int by auto - qed - -(*----------------------------------------------------------------------------*) -(* `iSucc` and `iPred` properties. *) - -(* Typeness *) -theorem iSucc_type: - assumes idom: "i \ Int" - shows "iSucc[i] \ Int" - proof - - have s1_1: "i \ Nat ==> - iSucc[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s2_1 by auto - have s2_3: "Succ[i] \ Nat" - using s2_1 succIsNat by auto - have s2_4: "Succ[i] \ Int" - using s2_3 nat_in_int by auto - show "iSucc[i] \ Int" - using s2_2 s2_4 by auto - qed - have s1_2: "i \ Nat ==> - iSucc[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using s2_1 idom by auto - have s2_3: "Pred[-.i] \ Nat" - proof - - have s3_1: "-.i \ Nat" - using idom s2_1 minus_neg_int_in_nat - by auto - show "Pred[-.i] \ Nat" - using s3_1 Pred_in_nat by auto - qed - have s2_4: "-.Pred[-.i] \ Int" - using s2_3 minus_nat_in_int by auto - show "iSucc[i] \ Int" - using s2_2 s2_4 by auto - qed - show "iSucc[i] \ Int" - using s1_1 s1_2 by auto - qed - - -theorem iPred_type: - assumes idom: "i \ Int" - shows "iPred[i] \ Int" - proof - - have s1_1: "i \ Nat \ {0} - ==> iPred[i] \ Int" - proof - - assume s2_1: "i \ Nat \ {0}" - have s2_2: "iPred[i] = Pred[i]" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "Pred[i] \ Nat" - using s2_1 Pred_in_nat by auto - have s2_4: "Pred[i] \ Int" - using s2_3 nat_in_int by auto - show "iPred[i] \ Int" - using s2_2 s2_4 by auto - qed - have s1_2: "i = 0 ==> - iPred[i] \ Int" - proof - - assume s2_1: "i = 0" - have s2_2: "iPred[i] = -.1" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "-.1 \ Int" - using oneIsNat by (auto simp: Int_def) - show "iPred[i] \ Int" - using s2_2 s2_3 by auto - qed - have s1_3: "i \ Nat ==> - iPred[i] \ Int" - proof - - assume s2_1: "i \ Nat" - have s2_2: "iPred[i] = -.Succ[-.i]" - unfolding iPred_def - using idom s2_1 by auto - have s2_3: "-.Succ[-.i] \ Int" - proof - - have s3_1: "-.i \ Nat" - using idom s2_1 minus_neg_int_in_nat - by auto - have s3_2: "Succ[-.i] \ Nat" - using s3_1 succIsNat by auto - show "-.Succ[-.i] \ Int" - using s3_2 minus_nat_in_int by auto - qed - show "iPred[i] \ Int" - using s2_2 s2_3 by auto - qed - show "iPred[i] \ Int" - using s1_1 s1_2 s1_3 by auto - qed - - -(* -THEOREM iSucciPredCommute == - \A i \in Int: iSucc[iPred[i]] = iPred[iSucc[i]] - PROOF - <1>1. SUFFICES ASSUME NEW i \in Int - PROVE iSucc[iPred[i]] = iPred[iSucc[i]] - BY <1>1 - <1>2. CASE i \in Nat \ {0} - <2>1. iSucc[iPred[i]] = i - <3>1. iPred[i] = Pred[i] - <4>1. i \in Nat \ {0} - BY <1>2 - <4> QED - BY <4>1 DEF iPred - <3>2. iSucc[iPred[i]] = Succ[Pred[i]] - <4>1. Pred[i] \in Nat - <5>1. i \in Nat - BY <1>2 - <5> QED - BY <5>1, Pred_in_nat - <4> QED - BY <3>1, <4>1 DEF iSucc - <3>3. Succ[Pred[i]] = i - <4>1. i \in Nat \ {0} - BY <1>2 - <4> QED - BY <4>1, Succ_Pred_nat - <3> QED - BY <3>2, <3>3 - <2>2. iPred[iSucc[i]] = i - <3>1. iSucc[i] = Succ[i] - <4>1. i \in Nat - BY <1>2 - <4> QED - BY <4>1 DEF iSucc - <3>2. iPred[iSucc[i]] = Pred[Succ[i]] - <4>1. Succ[i] \in Nat - <5>1. i \in Nat - BY <1>2 - <5> QED - BY <5>1, succIsNat - <4>2. Succ[i] # 0 - BY <1>2, succNotZero - <4>3. Succ[i] \in Nat \ {0} - BY <4>1, <4>2 - <4>4. iPred[Succ[i]] = Pred[Succ[i]] - BY <4>3 DEF iPred - <4> QED - BY <4>4, <3>1 - <3>3. Pred[Succ[i]] = i - <4>1. i \in Nat - BY <1>2 - <4> QED - BY <4>1, Pred_Succ_nat - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2 - <1>3. CASE i = 0 - <2>1. iSucc[iPred[i]] = 0 - <3>1. iSucc[iPred[i]] = iSucc[-.1] - <4>1. iPred[i] = -.1 - BY <1>3 DEF iPred - <4> QED - BY <4>1 - <3>2. iSucc[-.1] = -.Pred[-.-.1] - <4>1. -.1 \in Int - <5>1. -.1 \in negNat - BY oneIsNat DEF negNat - <5> QED - BY <5>1, neg_nat_in_int - <4>2. -.1 \notin Nat - <5>1. -.1 \in negNat \ {0} - <6>1. -.1 # 0 - BY minus_one_neq_0 - <6>2. -.1 \in negNat - BY oneIsNat DEF negNat - <6> QED - BY <6>1, <6>2 - <5> QED - BY <5>1, neg_nat_0_not_in_nat - <4> QED - BY <4>1, <4>2 DEF iSucc - <3>3. -.Pred[-.-.1] = 0 - <4>1. -.-.1 = 1 - BY minus_minus_one - <4>2. Pred[1] = 0 - BY pred_1 - <4>3. -.0 = 0 - BY neg0 - <4> QED - BY <4>1, <4>2, <4>3 - <3> QED - BY <3>1, <3>2, <3>3 - <2>2. iPred[iSucc[i]] = 0 - <3>1. iSucc[i] = 1 - <4>1. i \in Nat - BY <1>3, zeroIsNat - <4>2. iSucc[i] = Succ[i] - BY <4>1 DEF iSucc - <4>3. Succ[i] = Succ[0] - BY <1>3 - <4>4. Succ[0] = 1 - OBVIOUS - <4> QED - BY <4>2, <4>3, <4>4 - <3>2. iPred[1] = 0 - <4>1. 1 \in Nat \ {0} - BY oneIsNat, succIrrefl - <4>2. iPred[1] = Pred[1] - BY <4>1 DEF iPred - <4>3. Pred[1] = 0 - BY pred_1 - <4> QED - BY <4>2, <4>3 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 - <1>4. CASE i \in negNat - <2>1. iSucc[iPred[i]] = i - <3>1. iPred[i] = -.Succ[-.i] - <4>1. i \in Int - BY <1>1 - <4>2. i \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3>2. -.Succ[-.i] \in negNat \ {0} - <4>1. Succ[-.i] \in Nat \ {0} - <5>1. -.i \ Nat - BY <1>4, minus_neg_nat_in_nat - <5>2. Succ[-.i] \in Nat - BY <5>1, succIsNat - <5>3. Succ[-.i] \ 0 - BY <5>1, succNotZero - <5> QED - BY <5>2, <5>3 - <4> QED - BY <4>1, minus_nat_0_in_negnat_0 - <3>3. iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]] - <4>1. iPred[i] \in negNat \ {0} - BY <3>1, <3>2 - <4>2. iPred[i] \in Int - BY <4>1, neg_nat_in_int - <4>3. iPred[i] \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4>4. iSucc[iPred[i]] = -.Pred[-.iPred[i]] - BY <4>2, <4>3 DEF iSucc - <4> QED - BY <4>4, <3>1 - <3>4. @ = -.Pred[Succ[-.i]] - <4>1. -.i \in Nat - BY <1>4, minus_neg_nat_in_nat - <4>2. Succ[-.i] \in Int - BY <4>1, succIsNat, nat_in_int - <4>3. -.-.Succ[-.i] = Succ[-.i] - BY minus_involutive - <4> QED - BY <4>3 - <3>5. @ = -.-.i - <4>1. -.i \in Nat - BY <1>4, minus_neg_nat_in_nat - <4>2. Pred[Succ[-.i]] = -.i - BY <4>1, Pred_Succ_nat - <4> QED - BY <4>2 - <3>6. @ = i - BY <1>4, minus_involutive - <3> QED - BY <3>3, <3>4, <3>5, <3>6 - <2>2. iPred[iSucc[i]] = i - <3>1. CASE i = 0 - <4>1. iSucc[i] = 1 - <5>1. 0 \in Nat - BY zeroIsNat - <5>2. 0 \in Int - BY <5>1, nat_in_int - <5>3. iSucc[i] = Succ[0] - BY <5>1, <5>2, <3>1 DEF iSucc - <5>4. Succ[0] = 1 - OBVIOUS - <5> QED - BY <5>3, <5>4 - <4>2. iPred[1] = 0 - <5>1. 1 \in Nat \ {0} - BY oneIsNat, succNotZero - <5>2. 1 \in Int - BY <5>1, nat_in_int - <5>3. iPred[1] = Pred[1] - BY <5>1, <5>2 DEF iPred - <5>4. Pred[1] = 0 - BY pred_1 - <5> QED - BY <5>3, <5>4 - <4> QED - BY <4>1, <4>2, <3>1 - <3>2. CASE i # 0 - <4>1. i \in negNat \ {0} - BY <1>4, <3>1 - <4>2. i \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4>3. i \in Int - BY <4>1, neg_nat_in_int - <4>4. iSucc[i] = -.Pred[-.i] - BY <4>2, <4>3 DEF iSucc - <4>5. -.Pred[-.i] \in negNat - <5>1. -.i \in Nat - BY <4>1, minus_neg_nat_in_nat - <5>2. Pred[-.i] \in Nat - BY <5>1, Pred_in_nat - <5> QED - BY <5>2, minus_nat_in_neg_nat - <4>6. -.Pred[-.i] \in Int - BY <4>5, neg_nat_in_int - <4>7. -.Pred[-.i] \notin Nat \ {0} - BY <4>5, neg_nat_not_in_nat_setminus_0 - <4>8. iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]] - BY <4>6, <4>7 DEF iPred - <4>9. -.Succ[-.-.Pred[-.i]] = i - <5>1. -.-.Pred[-.i] = Pred[-.i] - <6>1. -.i \in Nat - BY <4>1, minus_neg_nat_in_nat - <6>2. Pred[-.i] \in Nat - BY <6>1, Pred_in_nat - <6>3. Pred[-.i] \in Int - BY <6>2, nat_in_int - <6> QED - BY <6>3, minus_involutive - <5>2. Succ[Pred[-.i]] = -.i - <6>1. -.i \in Nat \ {0} - BY <4>1, minus_neg_nat_0_in_nat_0 - <6> QED - BY <6>1, Succ_Pred_nat - <5>3. -.-.i = i - BY <4>3, minus_involutive - <5> QED - BY <5>1, <5>2, <5>3 - <4> QED - BY <4>4, <4>8, <4>9 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 - <1> QED - BY <1>1, <1>2, <1>3, <1>4 -*) -theorem iSucciPredCommute: - "\ i \ Int: - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - { - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "i \ Nat \ {0} ==> - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - assume idom: "i \ Nat \ {0}" - have s2_1: "iSucc[iPred[i]] = i" - proof - - have s3_1: "iPred[i] = Pred[i]" - unfolding iPred_def - using s1_1 idom by auto - have s3_2: "iSucc[iPred[i]] = Succ[Pred[i]]" - proof - - have s4_1: "Pred[i] \ Nat" - proof - - have s5_1: "i \ Nat" - using idom by auto - show ?thesis - using s5_1 Pred_in_nat by auto - qed - show ?thesis - unfolding iSucc_def - using s4_1 s3_1 by auto - qed - have s3_3: "Succ[Pred[i]] = i" - using idom Succ_Pred_nat - by auto - show ?thesis - using s3_2 s3_3 by auto - qed - have s2_2: "iPred[iSucc[i]] = i" - proof - - have s3_1: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s1_1 idom by auto - have s3_2: "iPred[iSucc[i]] = Pred[Succ[i]]" - proof - - have s4_1: "Succ[i] \ Nat" - proof - - have s5_1: "i \ Nat" - using idom by auto - show ?thesis - using s5_1 succIsNat by auto - qed - have s4_2: "Succ[i] \ 0" - using idom succNotZero by auto - have s4_3: "Succ[i] \ Nat \ {0}" - using s4_1 s4_2 by auto - have s4_4: "Succ[i] \ Int" - using s4_1 nat_in_int by auto - have s4_5: "iPred[Succ[i]] = Pred[Succ[i]]" - unfolding iPred_def - using s4_3 s4_4 by auto - show ?thesis - using s4_5 s3_1 by auto - qed - have s3_3: "Pred[Succ[i]] = i" - proof - - have s4_1: "i \ Nat" - using idom by auto - show ?thesis - using s4_1 Pred_Succ_nat - by auto - qed - show ?thesis - using s3_2 s3_3 by auto - qed - show ?thesis - using s2_1 s2_2 - by auto - qed - have s1_3: "i \ negNat ==> - iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i - " - proof - - assume idom: "i \ negNat" - have s2_1: "iSucc[iPred[i]] = i" - proof - - have s3_1: "iPred[i] = -.Succ[-.i]" - proof - - have s4_1: "i \ Int" - using s1_1 by auto - have s4_2: "i \ Nat \ {0}" - using idom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - have s3_2: "-.Succ[-.i] \ negNat \ {0}" - proof - - have s4_1: "Succ[-.i] \ Nat \ {0}" - proof - - have s5_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat - by auto - have s5_2: "Succ[-.i] \ Nat" - using s5_1 succIsNat by auto - have s5_3: "Succ[-.i] \ 0" - using s5_1 succNotZero by auto - show ?thesis - using s5_2 s5_3 by auto - qed - show ?thesis - using s4_1 minus_nat_0_in_negnat_0 - by auto - qed - have s3_3: "iSucc[iPred[i]] = -.Pred[-.-.Succ[-.i]]" - proof - - have s4_1: "iPred[i] \ negNat \ {0}" - using s3_1 s3_2 by auto - have s4_2: "iPred[i] \ Int" - using s4_1 neg_nat_in_int by auto - have s4_3: "iPred[i] \ Nat" - using s4_1 neg_nat_0_not_in_nat by auto - have s4_4: "iSucc[iPred[i]] = -.Pred[-.iPred[i]]" - unfolding iSucc_def - using s4_2 s4_3 by auto - show ?thesis - using s4_4 s3_1 by auto - qed - have s3_4: "-.Pred[-.-.Succ[-.i]] = -.Pred[Succ[-.i]]" - proof - - have s4_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat by auto - have s4_2: "Succ[-.i] \ Int" - using s4_1 succIsNat nat_in_int by auto - have s4_3: "-.-.Succ[-.i] = Succ[-.i]" - using s4_2 minus_involutive by auto - show ?thesis - using s4_3 by auto - qed - have s3_5: "-.Pred[Succ[-.i]] = -.-.i" - proof - - have s4_1: "-.i \ Nat" - using idom minus_neg_nat_in_nat by auto - have s4_2: "Pred[Succ[-.i]] = -.i" - using s4_1 Pred_Succ_nat by auto - show ?thesis - using s4_2 by auto - qed - have s3_6: "-.-.i = i" - using s1_1 minus_involutive by auto - show ?thesis - using s3_3 s3_4 s3_5 s3_6 by auto - qed - have s2_2: "iPred[iSucc[i]] = i" - proof - - have s3_1: "i = 0 ==> - iPred[iSucc[i]] = i" - proof - - assume i0: "i = 0" - have s4_1: "iSucc[i] = 1" - proof - - have s5_1: "0 \ Nat" - using zeroIsNat by auto - have s5_2: "0 \ Int" - using s5_1 nat_in_int by auto - have s5_3: "iSucc[i] = Succ[0]" - unfolding iSucc_def - using s5_1 s5_2 i0 - by auto - have s5_4: "Succ[0] = 1" - by auto - show ?thesis - using s5_3 s5_4 by auto - qed - have s4_2: "iPred[1] = 0" - proof - - have s5_1: "1 \ Nat \ {0}" - using oneIsNat succNotZero - by auto - have s5_2: "1 \ Int" - using s5_1 nat_in_int by auto - have s5_3: "iPred[1] = Pred[1]" - unfolding iPred_def - using s5_1 s5_2 - by auto - have s5_4: "Pred[1] = 0" - using pred_1 by auto - show ?thesis - using s5_3 s5_4 by auto - qed - show "iPred[iSucc[i]] = i" - using s4_1 s4_2 i0 by auto - qed - have s3_2: "i \ 0 ==> - iPred[iSucc[i]] = i" - proof - - assume inot0: "i \ 0" - have s4_1: "i \ negNat \ {0}" - using idom inot0 by auto - have s4_2: "i \ Nat" - using s4_1 neg_nat_0_not_in_nat - by auto - have s4_3: "i \ Int" - using s4_1 neg_nat_in_int - by auto - have s4_4: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using s4_2 s4_3 by auto - have s4_5: "-.Pred[-.i] \ negNat" - proof - - have s5_1: "-.i \ Nat" - using s4_1 minus_neg_nat_in_nat - by auto - have s5_2: "Pred[-.i] \ Nat" - using s5_1 Pred_in_nat - by auto - show ?thesis - using s5_2 minus_nat_in_neg_nat - by auto - qed - have s4_6: "-.Pred[-.i] \ Int" - using s4_5 neg_nat_in_int - by auto - have s4_7: "-.Pred[-.i] \ Nat \ {0}" - using s4_5 neg_nat_not_in_nat_setminus_0 - by blast - have s4_8: "iPred[-.Pred[-.i]] = -.Succ[-.-.Pred[-.i]]" - unfolding iPred_def - using s4_6 s4_7 by auto - have s4_9: "-.Succ[-.-.Pred[-.i]] = i" - proof - - have s5_1: "-.-.Pred[-.i] = Pred[-.i]" - proof - - have s6_1: "-.i \ Nat" - using s4_1 minus_neg_nat_in_nat - by auto - have s6_2: "Pred[-.i] \ Nat" - using s6_1 Pred_in_nat - by auto - have s6_3: "Pred[-.i] \ Int" - using s6_2 nat_in_int - by auto - show ?thesis - using s6_3 minus_involutive - by auto - qed - have s5_2: "Succ[Pred[-.i]] = -.i" - proof - - have s6_1: "-.i \ Nat \ {0}" - using s4_1 minus_neg_nat_0_in_nat_0 - by auto - show ?thesis - using s6_1 Succ_Pred_nat - by auto - qed - have s5_3: "-.-.i = i" - using s4_3 minus_involutive by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - show "iPred[iSucc[i]] = i" - using s4_4 s4_8 s4_9 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - show "iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i" - using s2_1 s2_2 by auto - qed - have s1_4: "Int = (Nat \ {0}) \ negNat" - using int_union_nat_0_negnat by auto - have "iSucc[iPred[i]] = iPred[iSucc[i]] \ - iSucc[iPred[i]] = i \ - iPred[iSucc[i]] = i" - using s1_1 s1_2 s1_3 s1_4 by auto - } - from this show ?thesis - using allI by auto - qed - - -(* -THEOREM iSuccMinusFlip == - \A i \in Int: -.iSucc[-.i] = iPred[i] -PROOF -<1>1. SUFFICES - ASSUME NEW i \in Int - PROVE -.iSucc[-.i] = iPred[i] - BY <1>1 -<1>2. CASE i \in Nat \ {0} - <2>1. i \in Int - BY <1>2, nat_in_int - <2>2. -.iSucc[-.i] = Pred[i] - <3>1. -.iSucc[-.i] = -.-.Pred[-.-.i] - <4>1. -.i \in negNat \ {0} - BY <1>2, minus_nat_0_in_negnat_0 - <4>2. -.i \notin Nat - BY <4>1, neg_nat_0_not_in_nat - <4> QED - BY <2>1, <4>2 DEF iSucc - <3>2. @ = -.-.Pred[i] - <4>1. -.-.i = i - BY <2>1, minus_involutive - <4> QED - BY <4>1 - <3>3. @ = Pred[i] - <4>1. Pred[i] \in Nat - BY <1>2, Pred_in_nat - <4>2. Pred[i] \in Int - BY <4>1, nat_in_int - <4> QED - BY <4>2, minus_involutive - <3> QED - BY <3>1, <3>2, <3>3 - <2>3. iPred[i] = Pred[i] - BY <2>1, <1>2 DEF iPred - <2> QED - BY <2>2, <2>3 -<1>3. CASE i \in negNat - <2>1. i \in Int - BY <1>3, neg_nat_in_int - <2>2. i \notin Nat \ {0} - BY <1>3, neg_nat_not_in_nat_setminus_0 - <2>3. Pred[i] = -.Succ[-.i] - BY <2>1, <2>2 DEF iPred - <2>4. -.iSucc[-.i] = -.Succ[-.i] - <3>1. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3>2. -.i \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>1, <3>2 DEF iSucc - <2> QED - BY <2>3, <2>4 -<1> QED - BY <1>1, <1>2, <1>3, int_union_nat_0_negnat -*) -theorem iSuccMinusFlip: - "\ i \ Int: - -.iSucc[-.i] = iPred[i]" - proof auto - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "i \ Nat \ {0} ==> - -.iSucc[-.i] = iPred[i]" - proof - - assume s1_2_asm: "i \ Nat \ {0}" - have s2_1: "i \ Int" - using s1_2_asm nat_in_int - by auto - have s2_2: "-.iSucc[-.i] = Pred[i]" - proof - - have s3_1: "-.iSucc[-.i] = -.-.Pred[-.-.i]" - proof - - have s4_1: "-.i \ negNat \ {0}" - using s1_2_asm minus_nat_0_in_negnat_0 - by auto - have s4_2: "-.i \ Nat" - using s4_1 neg_nat_0_not_in_nat - by auto - show ?thesis - unfolding iSucc_def - using s2_1 s4_2 by auto - qed - have s3_2: "-.-.Pred[-.-.i] = -.-.Pred[i]" - proof - - have s4_1: "-.-.i = i" - using s2_1 minus_involutive - by auto - show ?thesis - using s4_1 by auto - qed - have s3_3: "-.-.Pred[i] = Pred[i]" - proof - - have s4_1: "Pred[i] \ Nat" - using s1_2_asm Pred_in_nat by auto - have s4_2: "Pred[i] \ Int" - using s4_1 nat_in_int by auto - show ?thesis - using s4_2 minus_involutive by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 by auto - qed - have s2_3: "iPred[i] = Pred[i]" - unfolding iPred_def - using s2_1 s1_2_asm by auto - show "-.iSucc[-.i] = iPred[i]" - using s2_2 s2_3 by auto - qed - have s1_3: "i \ negNat ==> - -.iSucc[-.i] = iPred[i]" - proof - - assume s1_3_asm: "i \ negNat" - have s2_1: "i \ Int" - using s1_3_asm neg_nat_in_int - by auto - have s2_2: "i \ Nat \ {0}" - using s1_3_asm neg_nat_not_in_nat_setminus_0 - by auto - have s2_3: "-.Succ[-.i] = iPred[i]" - unfolding iPred_def - using s2_1 s2_2 by auto - have s2_4: "-.iSucc[-.i] = -.Succ[-.i]" - proof - - have s3_1: "-.i \ Nat" - using s1_3_asm minus_neg_nat_in_nat by auto - have s3_2: "-.i \ Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - show "-.iSucc[-.i] = iPred[i]" - using s2_3 s2_4 by auto - qed - show "-.iSucc[-.i] = iPred[i]" - using s1_1 s1_2 s1_3 int_union_nat_0_negnat - by auto - qed - - -(* -THEOREM iPredMinusFlip == - \A i \in Int: -.iPred[-.i] = iSucc[i] -PROOF -<1>1. SUFFICES - ASSUME NEW i \in Int - PROVE -.iPred[-.i] = iSucc[i] - BY <1>1 -<1>2. -.iSucc[-.-.i] = iPred[-.i] - <2>1. -.i \in Int - BY <1>1, neg_int_eq_int - <2> QED - BY <2>1, iSuccMinusFlip -<1>3. -.iSucc[i] = iPred[-.i] - <2>1. -.-.i = i - BY <1>1, minus_involutive - <2> QED - BY <1>2, <2>1 -<1>4. -.-.iSucc[i] = -.iPred[-.i] - BY <1>3 -<1>5. -.-.iSucc[i] = iSucc[i] - <2>1. iSucc[i] \in Int - BY <1>1, iSucc_type - <2> QED - BY <2>1, minus_involutive -<1> QED - BY <1>4, <1>5 -*) -theorem iPredMinusFlip: - "\ i \ Int: -.iPred[-.i] = iSucc[i]" - proof auto - fix "i" :: "c" - assume s1_1: "i \ Int" - have s1_2: "-.iSucc[-.-.i] = iPred[-.i]" - proof - - have s2_1: "-.i \ Int" - using s1_1 neg_int_eq_int by auto - show ?thesis - using s2_1 iSuccMinusFlip by auto - qed - have s1_3: "-.iSucc[i] = iPred[-.i]" - proof - - have s2_1: "-.-.i = i" - using s1_1 minus_involutive by auto - show ?thesis - using s1_2 s2_1 by auto - qed - have s1_4: "-.-.iSucc[i] = -.iPred[-.i]" - using s1_3 by auto - have s1_5: "-.-.iSucc[i] = iSucc[i]" - proof - - have s2_1: "iSucc[i] \ Int" - using s1_1 iSucc_type by auto - show ?thesis - using s2_1 minus_involutive by auto - qed - show "-.iPred[-.i] = iSucc[i]" - using s1_4 s1_5 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* `neg_nat_induction` *) - - -(* -negNat == {-.n: n \in Nat} - - -THEOREM negNatInduction == - ASSUME - NEW P(n), P(0), - \A i \in negNat: P(i) => P(-.Succ[-.i]) - PROVE - \A i \in negNat: P(i) - PROOF - <1> DEFINE Q(n) == P(-.n) - <1>1. \A n \in Nat: Q(n) - <2>1. Q(0) - BY P(0), -.0 = 0 DEF Q - <2>2. ASSUME NEW n \in Nat, Q(n) - PROVE Q(Succ[n]) - <3>1. P(-.n) - BY <2>2 DEF Q - <3>2. -.n \in negNat - BY <2>2, n \in Nat DEF negNat - <3>3. \A i \in negNat: P(i) => P(-.Succ[-.i]) - OBVIOUS - <3>4. P(-.n) => P(-.Succ[-.-.n]) - BY <3>2, <3>3 - <3>5. P(-.Succ[-.-.n]) - BY <3>1, <3>4 - <3>6. P(-.Succ[n]) - <4>1. -.-.n = n - BY <2>2, minus_involutive, nat_in_int - <4> QED - BY <3>5, <4>1 - <3> QED - BY <3>6 DEF Q - <2> QED - BY <2>1, <2>2, NatInduction - <1>2. ASSUME NEW i \in negNat - PROVE P(i) - <2>1. \E n \in Nat: n = -.i - <3>1. i \in negNat - BY <1>3 - <3>2. \E n \in Nat: i = -.n - BY <3>1 DEF negNat - <3>3. \E n \in Nat: -.i = -.-.n - BY <3>2 - <3>4. \A n \in Nat: -.-.n = n - BY minus_involutive, nat_in_int - <3> QED - BY <3>3, <3>4 - <2>2. \E n \in Nat: Q(-.i) - <3>1. \E n \in Nat: n = -.i /\ Q(n) - BY <1>1, <2>1 - <3>2. \E n \in Nat: n = -.i /\ Q(-.i) - BY <3>1 - <3> QED - BY <3>2 - <2>3. Q(-.i) - BY <2>2 - <2>4. P(-.-.i) - BY <2>3 DEF Q - <2>5. -.-.i = i - <3>1. i \in Int - BY <1>2 DEF negNat, Int - <3> QED - BY <3>1, minus_involutive - <2> QED - BY <2>4, <2>5 - <1> QED - BY <1>2 -*) -theorem neg_nat_induction: - assumes z: "P(0)" and - sc: "\ i. - \ - i \\in negNat; - P(i) - \ \ - P(-.Succ[-.i])" - shows "\ i \ negNat: P(i)" - proof - - { - fix "Q" :: "[c] \ c" - assume q_def: "\ n: Q(n) = P(-.n)" - have s1_1: "\ n \ Nat: Q(n)" - proof - - have s2_1: "Q(0)" - proof - - have s3_1: "P(-.0)" - using z neg0 by auto - show "Q(0)" - using s3_1 q_def allE by auto - qed - have s2_2: "\ n. - \ - n \ Nat; Q(n) - \ \ - Q(Succ[n])" - proof - - fix "n" :: "c" - assume ndom: "n \ Nat" - assume Qn: "Q(n)" - have s3_1: "P(-.n)" - using Qn q_def by auto - have s3_2: "-.n \ negNat" - unfolding negNat_def - using ndom by auto - have s3_3: "\ i. \ i \ negNat; - P(i)\ \ - P(-.Succ[-.i])" - using sc by auto - have s3_4: "P(-.n) ==> - P(-.Succ[-.-.n])" - using s3_2 s3_3 by auto - have s3_5: "P(-.Succ[-.-.n])" - using s3_1 s3_4 by auto - have s3_6: "P(-.Succ[n])" - proof - - have s4_1: "-.-.n = n" - using ndom minus_involutive nat_in_int - by auto - show "P(-.Succ[n])" - using s3_5 s4_1 by auto - qed - show "Q(Succ[n])" - using s3_6 q_def by auto - qed - show "\ n \ Nat: Q(n)" - using s2_1 s2_2 natInduct - by auto - qed - have s1_2: "\ i. i \\in negNat ==> P(i)" - proof - - fix "i" :: "c" - assume idom: "i \ negNat" - have s2_1: "\ n \ Nat: n = -.i" - proof - - have s3_1: "i \ negNat" - using idom by auto - have s3_2: "\ n \ Nat: i = -.n" - using s3_1 by (auto simp: negNat_def) - have s3_3: "\ n \ Nat: -.i = -.-.n" - using s3_2 by auto - have s3_4: "\ n \ Nat: -.-.n = n" - using minus_involutive nat_in_int - by auto - show "\ n \ Nat: n = -.i" - using s3_3 s3_4 by auto - qed - have s2_2: "\ n \ Nat: Q(-.i)" - proof - - have s3_1: "\ n \ Nat: - n = -.i \ Q(n)" - using s1_1 s2_1 by auto - have s3_2: "\ n \ Nat: - n = -.i \ Q(-.i)" - using s3_1 by auto - show "\ n \ Nat: Q(-.i)" - using s3_2 by auto - qed - have s2_3: "Q(-.i)" - using s2_2 by auto - have s2_4: "P(-.-.i)" - using s2_3 q_def by auto - have s2_5: "-.-.i = i" - proof - - have s3_1: "i \ Int" - using idom - by (auto simp: negNat_def Int_def) - show "-.-.i = i" - using s3_1 minus_involutive - by auto - qed - show "P(i)" - using s2_4 s2_5 by auto - qed - have "\ i \ negNat: P(i)" - proof - - have s2_1: - "\ i. i \ negNat => P(i)" - using s1_2 impI by auto - have s2_2: - "\ i: i \ negNat => P(i)" - using s2_1 allI by auto - show "\ i \ negNat: P(i)" - using s2_2 by auto - (* by (rule allI, rule impI, rule s1_2) *) - qed - } - from this have "\ Q: ( - (\ n: Q(n) = P(-.n)) - => - (\ i \ negNat: P(i)) - )" - by auto - from this have " - (\ n: P(-.n) = P(-.n)) - => - (\ i \ negNat: P(i))" - by auto - thus " - (\ i \ negNat: P(i))" - by auto - qed - -(*----------------------------------------------------------------------------*) -(* Primitive recursion in two directions: plus and minus infinity. *) - - -theorem bprimrec_int: - assumes - e: "e \ S" and - succ_h: "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - succ_g: "\ n \ Nat: - \ x \ S: - g(n, x) \ S" - shows "\ f \ [Int -> S]: - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - )" - proof - - from primrec_int[of e h g] obtain f where - 1: "isAFcn(f)" and - 2: "DOMAIN f = Int" and - 3: "f[0] = e" and - 4: "\ n \ Nat: - f[Succ[n]] = h(n, f[n]) - \ f[-.Succ[n]] = g(n, f[-.n])" - by blast - have s1_2: "\ n \ Nat: f[n] \ S" - proof (rule natInduct) - from 3 e show "f[0] \ S" by simp - next - fix n - assume "n \ Nat" and "f[n] \ S" - with succ_h 4 show "f[Succ[n]] \ S" - by force - qed - have s1_3: "\ i \ negNat: f[i] \ S" - proof (rule neg_nat_induction) - from 3 e show "f[0] \ S" by simp - next - fix i - assume ndom: "i \ negNat" and - fidom: "f[i] \ S" - have s2_1: "-.i \ Nat" - using ndom minus_neg_nat_in_nat by auto - have s2_2: "f[-.Succ[-.i]] = g(-.i, f[-.-.i])" - using 4 s2_1 by auto - have s2_3: "g(-.i, f[-.-.i]) = g(-.i, f[i])" - using ndom neg_nat_in_int minus_involutive - by auto - have s2_4: "g(-.i, f[i]) \ S" - using s2_1 fidom succ_g by auto - show "f[-.Succ[-.i]] \ S" - using s2_2 s2_3 s2_4 by auto - qed - have 5: "\ i \ Int: f[i] \ S" - using s1_2 s1_3 int_union_nat_negnat - by auto - with 1 2 3 4 5 show ?thesis - by blast - qed - - -theorem primrec_intE: - assumes e: "e \ S" and - succ_h: "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - succ_g: "\ n \ Nat: - \ x \ S: - g(n, x) \ S" and - f: "f = (CHOOSE r \ [Int -> S]: - r[0] = e \ - (\ n \ Nat: - r[Succ[n]] = h(n, r[n]) \ - r[-.Succ[n]] = g(n, r[-.n]) - ))" (is "f = ?r") and - maj: "\ - f \ [Int -> S]; - f[0] = e; - \ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - \ - \ P" - shows "P" - proof - - from e succ_h succ_g have " - \ r \ [Int -> S]: - r[0] = e \ - (\ n \ Nat: - r[Succ[n]] = h(n, r[n]) \ - r[-.Succ[n]] = g(n, r[-.n]) - )" - by (rule bprimrec_int) - hence "?r \ [Int -> S] \ - ?r[0] = e \ - (\ n \ Nat: - ?r[Succ[n]] = h(n, ?r[n]) \ - ?r[-.Succ[n]] = g(n, ?r[-.n]) - )" - by (rule bChooseI2, auto) - with f maj show ?thesis by auto - qed - - -theorem bprimrecType_int: - assumes "e \ S" and - "\ n \ Nat: - \ x \ S: - h(n, x) \ S" and - "\ n \ Nat: - \ x \ S: - g(n, x) \ S" - shows "(CHOOSE f \ [Int -> S]: - f[0] = e \ - (\ n \ Nat: - f[Succ[n]] = h(n, f[n]) \ - f[-.Succ[n]] = g(n, f[-.n]) - )) \ [Int -> S]" - by (rule primrec_intE[OF assms], auto) - - -(*----------------------------------------------------------------------------*) -(* Typeness of addition and recursive properties. *) - - -theorem addintType: - assumes mdom: "m \ Int" - shows "addint(m) \ [Int -> Int]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: "(CHOOSE f \ [Int -> Int]: - f[0] = m \ - (\ n \ Nat: - f[Succ[n]] = iSucc[f[n]] \ - f[-.Succ[n]] = iPred[f[-.n]] - )) \ [Int -> Int]" - using mdom s1_1 s1_2 - by (rule bprimrecType_int) - show "addint(m) \ [Int -> Int]" - unfolding addint_def - using s1_3 by auto - qed - - -theorem addIsInt: - assumes "m \ Int" and "n \ Int" - shows "add(m, n) \ Int" - unfolding add_def - using assms addintType by blast - - -theorem diffIsInt: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "diff(m, n) \\in Int" - proof - - have s1_1: "diff(m, n) = add(m, -.n)" - unfolding diff_def by auto - have s1_2: "-.n \\in Int" - using ndom neg_int_eq_int by auto - have s1_3: "add(m, -.n) \\in Int" - using mdom s1_2 addIsInt by auto - show ?thesis - using s1_1 s1_3 by auto - qed - - -theorem addint_0: - assumes mdom: "m \ Int" - shows "addint(m)[0] = m" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: " - addint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = m \ - (\ n \ Nat: - r[Succ[n]] = iSucc[r[n]] \ - r[-.Succ[n]] = iPred[r[-.n]]) - )" - unfolding addint_def - by auto - have s1_4: " - \ - addint(m) \ [Int -> Int]; - addint(m)[0] = m; - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - \ - \ addint(m)[0] = m" - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "m" "Int" - "\ n x. iSucc[x]" - "\ n x. iPred[x]" - "addint(m)"] - by auto - qed - - -theorem add_0: - assumes "m \ Int" - shows "m + 0 = m" - unfolding add_def - using assms addint_0 by auto - - -theorem addint_succ: - assumes mdom: "m \ Int" - shows "\ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - iSucc[x] \ Int" - using iSucc_type by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - iPred[x] \ Int" - using iPred_type by auto - have s1_3: " - addint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = m \ - (\ n \ Nat: - r[Succ[n]] = iSucc[r[n]] \ - r[-.Succ[n]] = iPred[r[-.n]]) - )" - unfolding addint_def - by auto - have s1_4: " - \ - addint(m) \ [Int -> Int]; - addint(m)[0] = m; - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - \ - \ - \ n \ Nat: - addint(m)[Succ[n]] = iSucc[addint(m)[n]] \ - addint(m)[-.Succ[n]] = iPred[addint(m)[-.n]] - " - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "m" "Int" - "\ n x. iSucc[x]" - "\ n x. iPred[x]" - "addint(m)"] - by auto - qed - - -theorem addint_succ_nat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m + Succ[n] = iSucc[m + n]" - unfolding add_def - using mdom ndom addint_succ[of "m"] spec by auto - - -theorem addint_succ_negnat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m + -.Succ[n] = iPred[m + -.n]" - unfolding add_def - using mdom ndom addint_succ[of "m"] spec by auto - - -theorem nat_add_1: - assumes ndom: "n \\in Nat" - shows "Succ[n] = n + 1" - proof - - have s1_1: "Succ[n] = iSucc[n]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "n \\in Nat" - using ndom by auto - show ?thesis - unfolding iSucc_def - using s2_1 s2_2 by auto - qed - also have s1_2: "... = iSucc[n + 0]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "n + 0 = n" - using s2_1 add_0 by auto - have s2_3: "n = n + 0" - using s2_2[symmetric] by auto - show ?thesis - using s2_3 by auto - qed - also have s1_3: "... = n + Succ[0]" - proof - - have s2_1: "n \\in Int" - using ndom nat_in_int by auto - have s2_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - using s2_1 s2_2 addint_succ_nat by auto - qed - also have s1_4: "... = n + 1" - by auto (* 1 is an abbreviation *) - from calculation show ?thesis . - qed - - -theorem nat_add_in_nat: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" - shows "m + n \\in Nat" - proof - - have s1_1: "m + 0 \\in Nat" - proof - - have s2_1: "m + 0 = m" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - show ?thesis - using s3_1 add_0 by auto - qed - have s2_2: "m \\in Nat" - using mdom by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_2: "\ k. \ - k \\in Nat; - m + k \\in Nat - \ \ - m + Succ[k] \\in Nat" - proof - - fix "k" :: "c" - assume s1_2_kdom: "k \\in Nat" - assume s1_2_induct: "m + k \\in Nat" - have s2_1: "m + Succ[k] = iSucc[m + k]" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - have s3_2: "k \\in Nat" - using s1_2_kdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - have s2_2: "iSucc[m + k] = Succ[m + k]" - proof - - have s3_1: "m + k \\in Nat" - using s1_2_induct by auto - have s3_2: "m + k \\in Int" - using s1_2_induct nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_3: "Succ[m + k] \\in Nat" - using s1_2_induct succIsNat by auto - show "m + Succ[k] \\in Nat" - using s2_1 s2_2 s2_3 by auto - qed - have s1_3: "\ k \ Nat: m + k \\in Nat" - using s1_1 s1_2 natInduct[of "\ k. m + k \\in Nat"] by auto - show ?thesis - using s1_3 ndom spec by auto - qed - - -theorem nat_0_succ: - "\ n \ Nat: - n=0 \ - (\ m \ Nat: n = Succ[m])" - by (rule natInduct, auto) - - -theorem nat_add_0: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" and - mn: "m + n = 0" - shows "m = 0 \ n = 0" - proof - - have s1_1: "n \ 0 ==> FALSE" - proof - - assume s2_1: "n \ 0" - have s2_2: "\ k: k \\in Nat \ n = Succ[k]" - using ndom s2_1 nat_0_succ by auto - define P where "P \ - \ x. x \\in Nat \ n = Succ[x]" - define r where "r \ CHOOSE x: P(x)" - have s2_3: "r \\in Nat \ n = Succ[r]" - proof - - have s3_1: "\ x: P(x)" - using s2_2 by (auto simp: P_def) - have s3_2: "P(r)" - using s3_1 chooseI_ex by (auto simp: r_def) - show ?thesis - using s3_2 by (auto simp: P_def) - qed - have s2_4: "m + n = m + Succ[r]" - using s2_3 by auto - have s2_5: "m + Succ[r] = iSucc[m + r]" - proof - - have s3_1: "m \\in Int" - using mdom nat_in_int by auto - have s3_2: "r \\in Nat" - using s2_3 by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - have s2_6: "iSucc[m + r] = Succ[m + r]" - proof - - have s3_1: "m + r \\in Nat" - using mdom s2_3 nat_add_in_nat by auto - have s3_2: "m + r \\in Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_7: "Succ[m + r] \ 0" - proof - - have s3_1: "m + r \\in Nat" - using mdom s2_3 nat_add_in_nat by auto - show ?thesis - using s3_1 succNotZero by auto - qed - have s2_7: "m + n \ 0" - using s2_4 s2_5 s2_6 s2_7 by auto - show "FALSE" - using s2_7 mn by auto - qed - have s1_2: "n = 0" - using s1_1 by auto - have s1_3: "m + 0 = 0" - using mn s1_2 by auto - have s1_4: "m = 0" - using s1_3 mdom nat_in_int add_0 by auto - show ?thesis - using s1_2 s1_4 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Typeness of multiplication and recursive properties. *) - - -theorem multType: - assumes mdom: "m \ Int" - shows "multint(m) \ [Int -> Int]" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: "(CHOOSE f \ [Int -> Int]: - f[0] = 0 \ - (\ n \ Nat: - f[Succ[n]] = f[n] + m \ - f[-.Succ[n]] = f[-.n] + -.m - )) \ [Int -> Int]" - using zero_in_int s1_1 s1_2 - by (rule bprimrecType_int[of "0"]) - show "multint(m) \ [Int -> Int]" - unfolding multint_def - using s1_3 by auto - qed - - -theorem multIsInt: - assumes "m \ Int" and "n \ Int" - shows "mult(m, n) \ Int" - unfolding mult_def - using assms multType by blast - - -theorem multint_0: - assumes mdom: "m \ Int" - shows "multint(m)[0] = 0" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: " - multint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = 0 \ - (\ n \ Nat: - r[Succ[n]] = r[n] + m \ - r[-.Succ[n]] = r[-.n] + -.m) - )" - unfolding multint_def - by auto - have s1_4: " - \ - multint(m) \ [Int -> Int]; - multint(m)[0] = 0; - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - \ - \ - multint(m)[0] = 0" - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "0" "Int" - "\ n x. x + m" - "\ n x. x + -.m" - "multint(m)"] - by auto - qed - - -theorem mult_0: - assumes "m \ Int" - shows "m * 0 = 0" - unfolding mult_def - using assms multint_0 by auto - - -theorem multint_succ: - assumes mdom: "m \ Int" - shows "\ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m" - proof - - have s1_1: " - \ n \ Nat: - \ x \ Int: - x + m \ Int" - using mdom addIsInt by auto - have s1_2: " - \ n \ Nat: - \ x \ Int: - x + -.m \ Int" - proof - - have s2_1: "-.m \ Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s2_1 addIsInt by auto - qed - have s1_3: " - multint(m) = (CHOOSE r \ [Int -> Int]: - r[0] = 0 \ - (\ n \ Nat: - r[Succ[n]] = r[n] + m \ - r[-.Succ[n]] = r[-.n] + -.m) - )" - unfolding multint_def - by auto - have s1_4: " - \ - multint(m) \ [Int -> Int]; - multint(m)[0] = 0; - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - \ - \ - \ n \ Nat: - multint(m)[Succ[n]] = multint(m)[n] + m \ - multint(m)[-.Succ[n]] = multint(m)[-.n] + -.m - " - by auto - show ?thesis - using mdom s1_1 s1_2 s1_3 s1_4 - primrec_intE[of - "0" "Int" - "\ n x. x + m" - "\ n x. x + -.m" - "multint(m)"] - by auto - qed - - -theorem multint_succ_nat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m * Succ[n] = m * n + m" - unfolding mult_def - using mdom ndom multint_succ[of "m"] spec by auto - - -theorem multint_succ_negnat: - assumes mdom: "m \ Int" and ndom: "n \ Nat" - shows "m * -.Succ[n] = m * -.n + -.m" - unfolding mult_def - using mdom ndom multint_succ[of "m"] spec by auto - - -theorem nat_mult_in_nat: - assumes mdom: "m \\in Nat" and ndom: "n \\in Nat" - shows "m * n \\in Nat" - proof - - have s1_1: "m * 0 \\in Nat" - proof - - have s2_1: "m * 0 = 0" - using mdom nat_in_int mult_0 by auto - have s2_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_2: "\ k. \ - k \\in Nat; - m * k \\in Nat - \ \ - m * Succ[k] \\in Nat" - proof - - fix "k" :: "c" - assume s1_2_kdom: "k \\in Nat" - assume s1_2_induct: "m * k \\in Nat" - have s2_1: "m * Succ[k] = m * k + m" - using mdom s1_2_kdom multint_succ_nat by auto - have s2_2: "m * k + m \\in Nat" - using s1_2_induct mdom nat_add_in_nat by auto - show "m * Succ[k] \\in Nat" - using s2_1 s2_2 by auto - qed - have s1_3: "\ k \ Nat: m * k \\in Nat" - using s1_1 s1_2 natInduct[of "\ k. m * k \\in Nat"] by auto - show ?thesis - using s1_3 ndom spec by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Commutativity of addition. *) - - -(* -THEOREM pred_irrefl == - ASSUME NEW n \in Nat \ {0} - PROVE Pred[n] # n -PROOF -<1>1. Pred[n] \in Nat - BY n \in Nat, Pred_in_nat -<1>2. Succ[Pred[n]] # Pred[n] - BY <1>1, succIrrefl -<1>3. Succ[Pred[n]] = n - BY n \in Nat \ {0}, Succ_Pred_nat -<1> QED - BY <1>2, <1>3 -*) -theorem pred_irrefl: - assumes ndom: "n \\in Nat \ {0}" - shows "Pred[n] \ n" - proof - - have s1_1: "Pred[n] \\in Nat" - using ndom Pred_in_nat by auto - have s1_2: "Succ[Pred[n]] \ Pred[n]" - using s1_1 succIrrefl by auto - have s1_3: "Succ[Pred[n]] = n" - using ndom Succ_Pred_nat by auto - show ?thesis - using s1_2 s1_3 by auto - qed - - -theorem ipred_plus_1: - assumes hyp: "a \ Int" - shows "iPred[a + 1] = a" - proof - - have s1_1: "iPred[a + 1] = iPred[a + Succ[0]]" - by auto - have s1_2: "... = iPred[iSucc[a + 0]]" - proof - - have s2_1: "a \ Int" - using hyp by auto - have s2_2: "0 \ Nat" - using zeroIsNat by auto - have s2_3: "a + Succ[0] = iSucc[a + 0]" - using s2_1 s2_2 addint_succ_nat by auto - show ?thesis - using s2_3 by auto - qed - have s1_3: "... = iPred[iSucc[a]]" - using hyp add_0 by auto - have s1_4: "... = a" - using hyp iSucciPredCommute spec by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -(* -THEOREM AddCommutativeNatNat == - \A j \in Nat: \A i \in Nat: - j + i = i + j -PROOF -<1>1. \A i \in Nat: 0 + i = i + 0 - <2>1. 0 + 0 = 0 + 0 - <3>1. 0 + 0 = addint(0)[0] - BY DEF + - <3>2. @ = 0 - BY DEF addint - <3> QED - BY <3>1, <3>2 - - <2>2. ASSUME NEW i \in Nat, 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - - <3>1. 0 + Succ[i] = iSucc[0 + i] - BY DEF + - <3>2. @ = iSucc[i + 0] - BY <2>2 - <3>3. @ = iSucc[i] - BY DEF + - <3>4. @ = Succ[i] - BY <2>2, i \in Nat - <3>5. @ = Succ[i] + 0 - BY DEF + - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5 - <2> QED - BY <2>1, <2>2, NatInduction - -<1>2. ASSUME NEW j \in Nat, - \A i \in Nat: j + i = i + j - PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] - - <2>1. Succ[j] + 0 = 0 + Succ[j] - <3>1. Succ[j] + 0 = Succ[j] - BY DEF + - <3>2. @ = Succ[j + 0] - BY DEF + - <3>3. @ = Succ[0 + j] - BY <1>2 - <3>4. @ = 0 + Succ[j] - BY DEF + - <3> QED - BY <3>1, <3>2, <3>3, <3>4 - - <2>2. ASSUME NEW i \in Nat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] - - <3>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] - <4>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. Succ[i] \in Nat - BY <2>2, succIsNat - <5> QED - BY <5>1, <5>2, addint_succ - <4>2. @ = iSucc[i + Succ[j]] - BY <2>2 - <4>3. @ = iSucc[iSucc[i + j]] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, addint_succ - <4>4. @ = iSucc[iSucc[j + i]] - BY <1>2, <2>2 - <4> QED - BY <4>1, <4>2, <4>3, <4>4 - - <3>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] - <4>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. Succ[j] \in Nat - BY <1>2, succIsNat - <5> QED - BY <5>1, <5>2, addint_succ - <4>2. @ = iSucc[j + Succ[i]] - <5>1. Succ[i] \in Nat - BY <2>2, succIsNat - <5> QED - BY <1>2, <5>1 - <4>3. @ = iSucc[iSucc[j + i]] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, addint_succ - <4> QED - BY <4>1, <4>2, <4>3 - - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, NatInduction -<1> QED - BY <1>1, <1>2, NatInduction - - -THEOREM AddCommutative == - \A j \in Int: \A i \in Int: - j + i = i + j -PROOF -<1>1. \A j \in Nat: \A i \in Int: - j + i = i + j - <2>1. \A j \in Nat: \A i \in Nat: - j + i = i + j - \* The proof of step <2>1 is the same with the - \* proof of the theorem `AddCommutativeNatNat`. - <3>1. \A i \in Nat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in Nat, - 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - <5>1. 0 + Succ[i] = iSucc[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Nat - BY <4>2 - <6> QED - BY <6>1, <6>2, addint_succ - DEF addd - <5>2. @ = iSucc[i + 0] - BY <4>2 - <5>3. @ = iSucc[i] - <6>1. i \in Int - BY <4>2, nat_in_int - <6> QED - BY <6>1 addint_0 - <5>4. @ = Succ[i] - BY <4>2 DEF iSucc - <5>5. @ = Succ[i] + 0 - <6>1. Succ[i] \in Nat - BY <4>2, succIsNat - <6>2. Succ[i] \in Int - BY <6>1, nat_in_int - <6> QED - BY <6>2, addint_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, NatInduction - <3>2. ASSUME NEW j \in Nat, - \A i \in Nat: j + i = i + j - PROVE \A i \in Nat: Succ[j] + i = i + Succ[j] - <4>1. Succ[j] + 0 = 0 + Succ[j] - <5>1. Succ[j] + 0 = Succ[j] - <6>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <6> QED - BY <6>1, addint_0 - <5>2. @ = iSucc[j] - BY <3>2, nat_in_int DEF iSucc - <5>3. @ = iSucc[j + 0] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iSucc[0 + j] - <6>1. j + 0 = 0 + j - BY <3>2, zeroIsNat - <6> QED - By <6>1 - <5>5. @ = 0 + Succ[j] - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2, addint_succ - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4>2. ASSUME NEW i \in Nat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + Succ[i] = Succ[i] + Succ[j] - <5>1. Succ[j] + Succ[i] = iSucc[iSucc[j + i]] - <6>1. Succ[j] + Succ[i] = iSucc[Succ[j] + i] - <7>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6>2. @ = iSucc[i + Succ[j]] - BY <4>2 - <6>3. @ = iSucc[iSucc[i + j]] - <7>1. i \in Int - BY <4>2, nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6>4. @ = iSucc[iSucc[j + i]] - BY <3>2, <4>2 - <6> QED - By <6>1, <6>2, <6>3, <6>4 - <5>2. Succ[i] + Succ[j] = iSucc[iSucc[j + i]] - <6>1. Succ[i] + Succ[j] = iSucc[Succ[i] + j] - <7>1. Succ[i] \in Int - BY <4>2, succIsNat, nat_in_int - <7>2. j \in Nat - BY <3>2, succIsNat - <7> QED - BY <7>1, <7>2, addint_succ - <6>2. @ = iSucc[j + Succ[i]] - <7>1. Succ[i] \in Nat - BY <4>2, succIsNat - <7> QED - BY <3>2, <7>1 - <6>3. @ = iSucc[iSucc[j + i]] - <7>1. j \in Int - BY <3>2, nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ - <6> QED - BY <6>1, <6>2, <6>3 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, NatInduction - <3> QED - BY <3>1, <3>2, NatInduction - <2>2. \A j \in Nat: \A i \in negNat: - j + i = i + j - <3>1. \A i \in negNat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in negNat, - 0 + i = i + 0 - PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 - <5>1. 0 + -.Succ[-.i] = iPred[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <6>3. 0 + -.Succ[-.i] = iPred[0 + -.-.i] - BY <6>1, <6>2, addint_succ_negnat - <6>4. -.-.i = i - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7> QED - BY <7>1, minus_involutive - <6> QED - BY <6>3, <6>4 - <5>2. @ = iPred[i + 0] - BY <4>2 - <5>3. @ = iPred[i] - <6>1. i \in Int - BY <4>2, neg_nat_in_int - <6>2. i + 0 = i - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = -.Succ[-.i] - <6>1. i \in Int - BY <4>2, neg_nat_in_int - <6>2. i \notin Nat \ {0} - BY <4>2, neg_nat_not_in_nat_setminus_0 - <6> QED - BY <6>1, <6>2 DEF iPred - <5>5. @ = -.Succ[-.i] + 0 - <6>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <6>2. Succ[-.i] \in Nat - BY <6>1, succIsNat - <6>3. -.Succ[-.i] \in Int - BY <6>2, minus_nat_in_int - <6> QED - BY <6>3, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3>2. ASSUME NEW j \in Nat, - \A i \in negNat: j + i = i + j - PROVE \A i \in negNat: Succ[j] + i = i + Succ[j] - <4>1. Succ[j] + 0 = 0 + Succ[j] - <5>1. Succ[j] + 0 = Succ[j] - <6>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <6> QED - BY <6>1, add_0 - <5>2. @ = iSucc[j] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2 DEF iSucc - <5>3. @ = iSucc[j + 0] - <6>1. j \in Int - BY <3>2, nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iSucc[0 + j] - <6>1. 0 \in negNat - BY neg0 DEF negNat - <6> QED - BY <3>2, <6>1 - <5>5. @ = 0 + Succ[j] - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Nat - BY <3>2 - <6> QED - BY <6>1, <6>2, addint_succ_nat - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4>2. ASSUME NEW i \in negNat, - Succ[j] + i = i + Succ[j] - PROVE Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j] - <5>1. Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]] - <6>1. Succ[j] + -.Succ[-.i] = iPred[Succ[j] + i] - <7>1. Succ[j] \in Int - BY <3>2, succIsNat, nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7>3. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7>4. Succ[j] + -.Succ[-.i] - = iPred[Succ[j] + -.-.i] - BY <7>1, <7>2, addint_succ_negnat - <7> QED - BY <7>3, <7>4 - <6>2. @ = iPred[i + Succ[j]] - BY <4>2 - <6>3. @ = iPred[iSucc[i + j]] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>4. @ = iPred[iSucc[j + i]] - <7>1. i \in negNat - BY <4>2 - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + i = i + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. -.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]] - <6>1. -.Succ[-.i] + Succ[j] = iSucc[-.Succ[-.i] + j] - <7>1. -.Succ[-.i] \in Int - BY <4>2, minus_neg_nat_in_nat, succIsNat, - minus_nat_in_int - <7>2. j \in Nat - BY <3>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>2. @ = iSucc[j + -.Succ[-.i]] - <7>1. -.Succ[-.i] \in negNat - BY <4>2, minus_neg_nat_in_nat, succIsNat, - minus_nat_in_neg_nat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6>3. @ = iSucc[iPred[j + i]] - <7>1. j \in Int - BY <3>2, nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7>3. j + -.Succ[-.i] = iPred[j + -.-.i] - BY <7>1, <7>2, addint_succ_negnat - <7>4. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>3, <7>4 - <6>4. @ = iPred[iSucc[j + i]] - <7>1. j + i \in Int - <8>1. j \in Int - BY <3>2, nat_in_int - <8>2. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, <8>2, addIsInt - <7> QED - BY <7>1, iSucciPredCommute - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3> QED - BY <3>1, <3>2, NatInduction - <2> QED - BY <2>1, <2>2, int_union_nat_negnat -<1>2. \A j \in negNat: \A i \in Int: - j + i = i + j - <2>1. \A j \in negNat: \A i \in Nat: - j + i = i + j - <3>1. \A i \in Nat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME - NEW i \in Nat, - 0 + i = i + 0 - PROVE 0 + Succ[i] = Succ[i] + 0 - \* The level 5 proof is the same as - \* for step <1>1!<2>1!<3>1!<4>2 . - <5>1. 0 + Succ[i] = iSucc[0 + i] - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Nat - BY <4>2 - <6> QED - BY <6>1, <6>2 DEF addint_succ_nat - <5>2. @ = iSucc[i + 0] - BY <4>2 - <5>3. @ = iSucc[i] - <6>1. i \in Int - BY <4>2, nat_in_int - <6>2. i + 0 = i - BY <6>1, addint_0 - <6> QED - BY <6>2 - <5>4. @ = Succ[i] - <6>1. i \in Nat - BY <4>2 - <6>2. i \in Int - BY <4>2, nat_in_int - <6> QED - BY <6>1, <6>2 DEF iSucc - <5>5. @ = Succ[i] + 0 - <6>1. Succ[i] \in Nat - BY <4>2, succIsNat - <6>2. Succ[i] \in Int - BY <6>1, nat_in_int - <6> QED - BY <6>2, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5 - <4> QED - BY <4>1, <4>2, NatInduction - <3>2. ASSUME - NEW j \in negNat, - \A i \in Nat: j + i = i + j - PROVE - \A i \in Nat: - -.Succ[-.j] + i = i + -.Succ[-.j] - <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] - <5>1. -.Succ[-.j] + 0 = -.Succ[-.j] - <6>1. -.Succ[-.j] \in Int - <7>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>2. Succ[-.j] \in Nat - BY <7>1, succIsNat - <7> QED - BY <7>2, minus_nat_in_int - <6> QED - BY <6>1, add_0 - <5>2. @ = iPred[j] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. j \notin Nat \ {0} - BY <3>2, neg_nat_not_in_nat_setminus_0 - <6> QED - BY <6>1, <6>2 DEF iPred - <5>3. @ = iPred[j + 0] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. j + 0 = j - BY <6>1, add_0 - <6> QED - BY <6>2 - <5>4. @ = iPred[0 + j] - <6>1. 0 \in Nat - BY zeroIsNat - <6>2. \A k \in Nat: j + k = k + j - BY <3>2 - <6>3. j + 0 = 0 + j - BY <6>1, <6>2 - <6> QED - BY <6>3 - <5>5. @ = iPred[0 + -.-.j] - <6>1. j \in Int - BY <3>2, neg_nat_in_int - <6>2. -.-.j = j - BY <6>1, minus_involutive - <6> QED - BY <6>2 - <5>6. @ = 0 + -.Succ[-.j] - <6>1. 0 \in Int - BY zeroIsNat, nat_in_int - <6>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <6> QED - BY <6>1, <6>2, addint_succ_negnat - <5> QED - BY <5>1, <5>2, <5>3, <5>4, <5>5, <5>6 - <4>2. ASSUME - NEW i \in Nat, - -.Succ[-.j] + i = i + -.Succ[-.j] - PROVE - -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j] - <5>1. -.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]] - <6>1. -.Succ[-.j] + Succ[i] = - iSucc[-.Succ[-.j] + i] - <7>1. -.Succ[-.j] \in Int - <8>1. j \in negNat - BY <3>2 - <8>2. -.j \in Nat - BY <8>1, minus_neg_nat_in_nat - <8>3. Succ[-.j] \in Nat - BY <8>2, succIsNat - <8> QED - BY <8>3, minus_nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>2. @ = iSucc[i + -.Succ[-.j]] - BY <4>2 - <6>3. @ = iSucc[iPred[i + j]] - <7>1. i \in Int - BY <4>2, nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>3. -.-.j = j - BY <3>2, neg_nat_in_int, minus_involutive - <7> QED - BY <7>1, <7>2, <7>3, addint_succ_negnat - <6>4. @ = iSucc[iPred[j + i]] - <7>1. i \in Nat - BY <4>2 - <7>2. j + i = i + j - BY <3>2, <7>1 - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]] - <6>1. Succ[i] + -.Succ[-.j] = - iPred[Succ[i] + -.-.j] - <7>1. Succ[i] \in Int - BY <4>2, succIsNat, nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[Succ[i] + j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>2 - <6>3. @ = iPred[j + Succ[i]] - <7>1. Succ[i] \in Nat - BY <4>2, succIsNat - <7>2. \A k \in Nat: j + k = k + j - BY <3>2 - <7>3. j + Succ[i] = Succ[i] + j - BY <7>1, <7>2 - <7> QED - BY <7>3 - <6>4. @ = iPred[iSucc[j + i]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. i \in Nat - BY <4>2 - <7> QED - BY <7>1, <7>2, addint_succ_nat - <6>5. @ = iSucc[iPred[j + i]] - <7>1. j + i \in Int - <8>1. j \in Int - BY <3>2, neg_nat_in_int - <8>2. i \in Int - BY <4>2, nat_in_int - <8> QED - BY <8>1, <8>2, addIsInt - <7> QED - BY <7>1, iSucciPredCommute - <6> QED - BY <6>1, <6>2, <6>3, <6>4, <6>5 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, NatInduction - <3> QED - BY <3>1, <3>2, neg_nat_induction - <2>2. \A j \in negNat: \A i \in negNat: - j + i = i + j - <3>1. \A i \in negNat: 0 + i = i + 0 - <4>1. 0 + 0 = 0 + 0 - OBVIOUS - <4>2. ASSUME NEW i \in negNat, 0 + i = i + 0 - PROVE 0 + -.Succ[-.i] = -.Succ[-.i] + 0 - <5>1. 0 + -.Succ[-.i] = iPred[i] - <6>1. 0 + -.Succ[-.i] = iPred[0 + -.-.i] - <7>1. 0 \in Int - BY zero_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[0 + i] - <7>1. -.-.i = i - BY <4>2, neg_nat_in_int, minus_involutive - <7> QED - BY <7>1 - <6>3. @ = iPred[i + 0] - BY <4>2 - <6>4. @ = iPred[i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. i + 0 = i - BY <7>1, add_0 - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5>2. -.Succ[-.i] + 0 = iPred[i] - <6>1. -.Succ[-.i] + 0 = -.Succ[-.i] - <7>1. -.Succ[-.i] \in Int - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7> QED - BY <7>1, add_0 - <6>2. @ = iPred[i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. i \notin Nat \ {0} - BY <4>2, neg_nat_not_in_nat_setminus_0 - <7> QED - BY <7>1, <7>2 DEF iPred - <6> QED - BY <6>1, <6>2 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3>2. ASSUME NEW j \in negNat, - \A i \in negNat: j + i = i + j - PROVE \A i \in negNat: - -.Succ[-.j] + i = i + -.Succ[-.j] - <4>1. -.Succ[-.j] + 0 = 0 + -.Succ[-.j] - <5>1. -.Succ[-.j] + 0 = iPred[j] - <6>1. -.Succ[-.j] + 0 = -.Succ[-.j] - <7>1. -.Succ[-.j] \in Int - <8>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <8>2. Succ[-.j] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7> QED - BY <7>1, add_0 - <6>2. @ = iPred[j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. j \notin Nat \ {0} - BY <3>2, neg_nat_not_in_nat_setminus_0 - <7> QED - BY <7>1, <7>2 DEF iPred - <6> QED - BY <6>1, <6>2 - <5>2. 0 + -.Succ[-.j] = iPred[j] - <6>1. 0 + -.Succ[-.j] = iPred[0 + -.-.j] - <7>1. 0 \in Int - BY zero_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[0 + j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>1, <7>2 - <6>3. @ = iPred[j + 0] - <7>1. 0 \in negNat - BY neg0 DEF negNat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + 0 = 0 + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6>4. @ = iPred[j] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7> QED - BY <7>1, add_0 - <6> QED - BY <6>1, <6>2, <6>3, <6>4 - <5> QED - BY <5>1, <5>2 - <4>2. ASSUME NEW i \in negNat, - -.Succ[-.j] + i = i + -.Succ[-.j] - PROVE -.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j] - <5>1. -.Succ[-.j] + -.Succ[-.i] = iPred[iPred[i + j]] - <6>1. -.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i] - <7>1. -.Succ[-.j] \in Int - <8>1. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <8>2. Succ[-.j] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[-.Succ[-.j] + i] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. -.-.i = i - BY <7>1, minus_involutive - <7> QED - BY <7>1, <7>2 - <6>3. @ = iPred[i + -.Succ[-.j]] - BY <4>2 - <6>4. @ = iPred[iPred[i + -.-.j]] - <7>1. i \in Int - BY <4>2, neg_nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7>3. i + -.Succ[-.j] = iPred[i + -.-.j] - BY <7>1, <7>2, addint_succ_negnat - <7> QED - BY <7>3 - <6>5. @ = iPred[iPred[i + j]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.-.j = j - BY <7>1, minus_involutive - <7> QED - BY <7>2 - <6> QED - BY <6>1, <6>2, <6>3, <6>4, <6>5 - <5>2. -.Succ[-.i] + -.Succ[-.j] = iPred[iPred[i + j]] - <6>1. -.Succ[-.i] + -.Succ[-.j] = - iPred[-.Succ[-.i] + -.-.j] - <7>1. -.Succ[-.i] \in Int - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_int - <7>2. -.j \in Nat - BY <3>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>2. @ = iPred[-.Succ[-.i] + j] - <7>1. -.-.j = j - <8>1. j \in Int - BY <3>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>1 - <6>3. @ = iPred[j + -.Succ[-.i]] - <7>1. -.Succ[-.i] \in negNat - <8>1. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <8>2. Succ[-.i] \in Nat - BY <8>1, succIsNat - <8> QED - BY <8>2, minus_nat_in_neg_nat - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + -.Succ[-.i] = -.Succ[-.i] + j - BY <7>1, <7>2 - <7> QED - BY <7>3 - <6>4. @ = iPred[iPred[j + -.-.i]] - <7>1. j \in Int - BY <3>2, neg_nat_in_int - <7>2. -.i \in Nat - BY <4>2, minus_neg_nat_in_nat - <7> QED - BY <7>1, <7>2, addint_succ_negnat - <6>5. @ = iPred[iPred[j + i]] - <7>1. -.-.i = i - <8>1. i \in Int - BY <4>2, neg_nat_in_int - <8> QED - BY <8>1, minus_involutive - <7> QED - BY <7>1 - <6>6. @ = iPred[iPred[i + j]] - <7>1. i \in negNat - BY <4>2 - <7>2. \A k \in negNat: j + k = k + j - BY <3>2 - <7>3. j + i = i + j - BY <7>1, <7>2, spec - <7> QED - BY <7>3 - <6> QED - BY <6>1, <6>2, <6>3, <6>4, - <6>5, <6>6 - <5> QED - BY <5>1, <5>2 - <4> QED - BY <4>1, <4>2, neg_nat_induction - <3> QED - BY <3>1, <3>2, neg_nat_induction - <2> QED - BY <2>1, <2>2, int_union_nat_negnat -<1> QED - BY <1>1, <1>2 - -*) -theorem AddCommutative: - "\ j \ Int: \ i \ Int: - j + i = i + j" - proof - - have s1_1: "\ j \ Nat: \ i \ Int: - j + i = i + j" - proof - - have s2_1: "\ j \ Nat: - \ i \ Nat: - j + i = i + j" - proof - - have s3_1: "\ i \ Nat: - 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ Nat; - 0 + i = i + 0 - \ - \ - 0 + Succ[i] = Succ[i] + 0" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + Succ[i] = iSucc[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - unfolding add_def - using s6_1 s6_2 - addint_succ[of "0"] - by auto - qed - have s5_2: "iSucc[0 + i] = iSucc[i + 0]" - using s4_2_asm2 by auto - have s5_3: "iSucc[i + 0] = iSucc[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 nat_in_int - by auto - show ?thesis - unfolding add_def - using s6_1 addint_0[of "i"] - by auto - qed - have s5_4: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using s4_2_asm1 by auto - have s5_5: "Succ[i] = Succ[i] + 0" - proof - - have s6_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - have s6_2: "Succ[i] \ Int" - using s6_1 nat_in_int by auto - show ?thesis - unfolding add_def - using s6_2 addint_0 by auto - qed - show "0 + Succ[i] = Succ[i] + 0" - using s5_1 s5_2 s5_3 s5_4 s5_5 - by auto - qed - show ?thesis - using s4_1 s4_2 - natInduct[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ Nat; - \ i \ Nat: j + i = i + j - \ - \ - \ i \ Nat: - Succ[j] + i = i + Succ[j]" - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ Nat" - assume s3_2_asm2: "\ i \ Nat: - j + i = i + j" - have s4_1: "Succ[j] + 0 = 0 + Succ[j]" - proof - - have s5_1: "Succ[j] + 0 = Succ[j]" - proof - - have s6_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat - nat_in_int by auto - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iSucc[j]" - unfolding iSucc_def - using s3_2_asm1 nat_in_int - by auto - also have s5_3: "... = iSucc[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iSucc[0 + j]" - proof - - have s6_1: "j + 0 = 0 + j" - using s3_2_asm2 zeroIsNat by auto - show ?thesis - using s6_1 by auto - qed - also have s5_5: "... = 0 + Succ[j]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - have s6_3: "0 \ Int ==> - j \ Nat ==> - 0 + Succ[j] = iSucc[0 + j]" - using addint_succ_nat - by auto - have s6_4: " - 0 + Succ[j] = iSucc[0 + j]" - using s6_1 s6_2 s6_3 by auto - have s6_5: "iSucc[0 + j] = 0 + Succ[j]" - using s6_4 by auto - show ?thesis - using s6_5 by auto - qed - finally show ?thesis . - qed - have s4_2: "\ i. \ - i \ Nat; - Succ[j] + i = i + Succ[j] - \ \ - Succ[j] + Succ[i] = Succ[i] + Succ[j]" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" - have s5_1: "Succ[j] + Succ[i] = iSucc[iSucc[j + i]]" - proof - - have s6_1: "Succ[j] + Succ[i] = iSucc[Succ[j] + i]" - proof - - have s7_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "i \ Nat" - using s4_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - also have s6_2: "... = iSucc[i + Succ[j]]" - using s4_2_asm2 by auto - also have s6_3: "... = iSucc[iSucc[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - also have s6_4: "... = iSucc[iSucc[j + i]]" - using s3_2_asm2 s4_2_asm1 spec by auto - finally show ?thesis - by auto - qed - have s5_2: "Succ[i] + Succ[j] = iSucc[iSucc[j + i]]" - proof - - have s6_1: "Succ[i] + Succ[j] = iSucc[Succ[i] + j]" - proof - - have s7_1: "Succ[i] \ Int" - using s4_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "j \ Nat" - using s3_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[j + Succ[i]]" - proof - - have s7_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - show ?thesis - using s7_1 s3_2_asm2 by auto - qed - also have s6_3: "... = iSucc[iSucc[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 nat_in_int - by auto - have s7_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat - by auto - qed - finally show ?thesis - by auto - qed - show "Succ[j] + Succ[i] = Succ[i] + Succ[j]" - using s5_1 s5_2 by auto - qed - show "\ i \ Nat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 - natInduct[of "\ i. Succ[j] + i = i + Succ[j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - natInduct[of "\ j. - \ i \ Nat: j + i = i + j"] - by auto - qed - have s2_2: "\ j \ Nat: - \ i \ negNat: j + i = i + j" - proof - - have s3_1: "\ i \ negNat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ negNat; - 0 + i = i + 0 - \ \ - 0 + -.Succ[-.i] = -.Succ[-.i] + 0" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ negNat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + -.Succ[-.i] = iPred[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat - by auto - have s6_3: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" - using s6_1 s6_2 addint_succ_negnat - by auto - have s6_4: "-.-.i = i" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s7_1 minus_involutive by auto - qed - show ?thesis - using s6_3 s6_4 by auto - qed - also have s5_2: "... = iPred[i + 0]" - using s4_2_asm2 by auto - also have s5_3: "... = iPred[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s6_2: "i + 0 = i" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = -.Succ[-.i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s6_2: "i \ Nat \ {0}" - using s4_2_asm1 neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s6_1 s6_2 by auto - qed - also have s5_5: "... = -.Succ[-.i] + 0" - proof - - have s6_1: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat by auto - have s6_2: "Succ[-.i] \ Nat" - using s6_1 succIsNat by auto - have s6_3: "-.Succ[-.i] \ Int" - using s6_2 minus_nat_in_int by auto - show ?thesis - using s6_3 add_0 by auto - qed - finally show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" - by auto - qed - show ?thesis - using s4_1 s4_2 - neg_nat_induction[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ Nat; - \ i \ negNat: j + i = i + j - \ \ - \ i \ negNat: Succ[j] + i = i + Succ[j]" - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ Nat" - assume s3_2_asm2: "\ i \ negNat: - j + i = i + j" - have s4_1: "Succ[j] + 0 = 0 + Succ[j]" - proof - - have s5_1: "Succ[j] + 0 = Succ[j]" - proof - - have s6_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int by auto - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iSucc[j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - unfolding iSucc_def - using s6_1 s6_2 by auto - qed - also have s5_3: "... = iSucc[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iSucc[0 + j]" - proof - - have s6_1: "0 \ negNat" - unfolding negNat_def - using neg0 by auto - show ?thesis - using s3_2_asm2 s6_1 by auto - qed - also have s5_5: "... = 0 + Succ[j]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s6_1 s6_2 addint_succ_nat by auto - qed - finally show ?thesis - by auto - qed - have s4_2: "\ i. \ - i \ negNat; - Succ[j] + i = i + Succ[j] - \ \ - Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ negNat" - assume s4_2_asm2: "Succ[j] + i = i + Succ[j]" - have s5_1: "Succ[j] + -.Succ[-.i] = iPred[iSucc[j + i]]" - proof - - have s6_1: "Succ[j] + -.Succ[-.i] = - iPred[Succ[j] + i]" - proof - - have s7_1: "Succ[j] \ Int" - using s3_2_asm1 succIsNat nat_in_int - by auto - have s7_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat - by auto - have s7_3: "-.-.i = i" - using s4_2_asm1 neg_nat_in_int - minus_involutive by auto - have s7_4: "Succ[j] + -.Succ[-.i] = - iPred[Succ[j] + -.-.i]" - using s7_1 s7_2 addint_succ_negnat - by auto - show ?thesis - using s7_3 s7_4 by auto - qed - also have s6_2: "... = iPred[i + Succ[j]]" - using s4_2_asm2 by auto - also have s6_3: "... = iPred[iSucc[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "i \ negNat" - using s4_2_asm1 by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + i = i + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - finally show ?thesis - by auto - qed - have s5_2: "-.Succ[-.i] + Succ[j] = iPred[iSucc[j + i]]" - proof - - have s6_1: "-.Succ[-.i] + Succ[j] = - iSucc[-.Succ[-.i] + j]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - using s4_2_asm1 minus_neg_nat_in_nat succIsNat - minus_nat_in_int by auto - have s7_2: "j \ Nat" - using s3_2_asm1 by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[j + -.Succ[-.i]]" - proof - - have s7_1: "-.Succ[-.i] \ negNat" - using s4_2_asm1 minus_neg_nat_in_nat succIsNat - minus_nat_in_neg_nat by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - also have s6_3: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_asm1 minus_neg_nat_in_nat by auto - have s7_3: "j + -.Succ[-.i] = iPred[j + -.-.i]" - using s7_1 s7_2 addint_succ_negnat by auto - have s7_4: "-.-.i = i" - proof - - have s8_1: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_3 s7_4 by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "j + i \ Int" - proof - - have s8_1: "j \ Int" - using s3_2_asm1 nat_in_int by auto - have s8_2: "i \ Int" - using s4_2_asm1 neg_nat_in_int by auto - show ?thesis - using s8_1 s8_2 addIsInt by auto - qed - show ?thesis - using s7_1 iSucciPredCommute by auto - qed - finally show ?thesis - by auto - qed - show "Succ[j] + -.Succ[-.i] = -.Succ[-.i] + Succ[j]" - using s5_1 s5_2 by auto - qed - show "\ i \ negNat: Succ[j] + i = i + Succ[j]" - using s4_1 s4_2 - neg_nat_induction[of - "\ i. Succ[j] + i = i + Succ[j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - natInduct[of "\ j. - \ i \ negNat: j + i = i + j"] - by auto - qed - show ?thesis - proof - - have s3_1: "\ j. \ i. - j \ Nat => ( - i \ Int => ( - j + i = i + j))" - proof auto - fix "j" :: "c" - fix "i" :: "c" - assume s4_1: "j \ Nat" - assume s4_2: "i \ Int" - have s4_3: "i \ Nat ==> j + i = i + j" - proof - - assume s5_1: "i \ Nat" - show "j + i = i + j" - using s4_1 s5_1 s2_1 by auto - qed - have s4_4: "i \ negNat ==> j + i = i + j" - proof - - assume s5_1: "i \ negNat" - show "j + i = i + j" - using s4_1 s5_1 s2_2 by auto - qed - show "j + i = i + j" - using s4_2 s4_3 s4_4 int_union_nat_negnat by auto - qed - show ?thesis - using s3_1 by auto - qed - qed - have s1_2: "\ j \ negNat: \ i \ Int: - j + i = i + j" - proof - - have s2_1: "\ j \ negNat: \ i \ Nat: - j + i = i + j" - proof - - have s3_1: "\ i \ Nat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ Nat; - 0 + i = i + 0 - \ \ - 0 + Succ[i] = Succ[i] + 0 - " - proof - - fix "i" :: "c" - assume s4_2_asm1: "i \ Nat" - assume s4_2_asm2: "0 + i = i + 0" - have s5_1: "0 + Succ[i] = iSucc[0 + i]" - proof - - have s6_1: "0 \ Int" - using zero_in_int by auto - have s6_2: "i \ Nat" - using s4_2_asm1 by auto - show ?thesis - using s6_1 s6_2 addint_succ_nat by auto - qed - also have s5_2: "... = iSucc[i + 0]" - using s4_2_asm2 by force - also have s5_3: "... = iSucc[i]" - proof - - have s6_1: "i \ Int" - using s4_2_asm1 nat_in_int by auto - have s6_2: "i + 0 = i" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by force - qed - also have s5_4: "... = Succ[i]" - proof - - have s6_1: "i \ Nat" - using s4_2_asm1 by auto - have s6_2: "i \ Int" - using s4_2_asm1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s6_1 s6_2 by auto - qed - also have s5_5: "... = Succ[i] + 0" - proof - - have s6_1: "Succ[i] \ Nat" - using s4_2_asm1 succIsNat by auto - have s6_2: "Succ[i] \ Int" - using s6_1 nat_in_int by auto - show ?thesis - using s6_2 add_0 by auto - qed - finally show "0 + Succ[i] = Succ[i] + 0" - (* . did not work *) - using s5_1 s5_2 s5_3 s5_4 s5_5 by auto - qed - show ?thesis - using s4_1 s4_2 - natInduct[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ negNat; - \ i \ Nat: j + i = i + j - \ \ - \ i \ Nat: - -.Succ[-.j] + i = i + -.Succ[-.j] - " - proof - - fix "j" :: "c" - assume s3_2_asm1: "j \ negNat" - assume s3_2_asm2: "\ i \ Nat: - j + i = i + j" - have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" - proof - - have s6_1: "-.Succ[-.j] \ Int" - proof - - have s7_1: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat - by auto - have s7_2: "Succ[-.j] \ Nat" - using s7_1 succIsNat by auto - show ?thesis - using s7_2 minus_nat_in_int by auto - qed - show ?thesis - using s6_1 add_0 by auto - qed - also have s5_2: "... = iPred[j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "j \ Nat \ {0}" - using s3_2_asm1 neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s6_1 s6_2 by auto - qed - also have s5_3: "... = iPred[j + 0]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "j + 0 = j" - using s6_1 add_0 by auto - show ?thesis - using s6_2 by auto - qed - also have s5_4: "... = iPred[0 + j]" - proof - - have s6_1: "0 \ Nat" - using zeroIsNat by auto - have s6_2: "\ k \ Nat: - j + k = k + j" - using s3_2_asm2 by auto - have s6_3: "j + 0 = 0 + j" - using s6_1 s6_2 by auto - show ?thesis - using s6_3 by auto - qed - also have s5_5: "... = iPred[0 + -.-.j]" - proof - - have s6_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s6_2: "-.-.j = j" - using s6_1 minus_involutive by auto - show ?thesis - using s6_2 by auto - qed - also have s5_6: "... = 0 + -.Succ[-.j]" - proof - - have s6_1: "0 \ Int" - using zeroIsNat nat_in_int by auto - have s6_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - show ?thesis - using s6_1 s6_2 addint_succ_negnat by auto - qed - finally show ?thesis . - qed - have s4_2: "\ i. \ - i \ Nat; - -.Succ[-.j] + i = i + -.Succ[-.j] - \ \ - -.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ Nat" - assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" - have s5_1: "-.Succ[-.j] + Succ[i] = iSucc[iPred[j + i]]" - proof - - have s6_1: "-.Succ[-.j] + Succ[i] = - iSucc[-.Succ[-.j] + i]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "j \ negNat" - using s3_2_asm1 by auto - have s8_2: "-.j \ Nat" - using s8_1 minus_neg_nat_in_nat by auto - have s8_3: "Succ[-.j] \ Nat" - using s8_2 succIsNat by auto - show ?thesis - using s8_3 minus_nat_in_int by auto - qed - have s7_2: "i \ Nat" - using s4_2_idom by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_2: "... = iSucc[i + -.Succ[-.j]]" - using s4_2_induct by auto - also have s6_3: "... = iSucc[iPred[i + j]]" - proof - - have s7_1: "i \ Int" - using s4_2_idom nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - have s7_3: "-.-.j = j" - using s3_2_asm1 neg_nat_in_int - minus_involutive by auto - show ?thesis - using s7_1 s7_2 s7_3 addint_succ_negnat - by auto - qed - also have s6_4: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "i \ Nat" - using s4_2_idom by auto - have s7_2: "j + i = i + j" - using s3_2_asm2 s7_1 spec by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "Succ[i] + -.Succ[-.j] = iSucc[iPred[j + i]]" - proof - - have s6_1: "Succ[i] + -.Succ[-.j] = - iPred[Succ[i] + -.-.j]" - proof - - have s7_1: "Succ[i] \ Int" - using s4_2_idom succIsNat nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_asm1 minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[Succ[i] + j]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_2 by auto - qed - also have s6_3: "... = iPred[j + Succ[i]]" - proof - - have s7_1: "Succ[i] \ Nat" - using s4_2_idom succIsNat by auto - have s7_2: "\ k \ Nat: - j + k = k + j" - using s3_2_asm2 by auto - have s7_3: "j + Succ[i] = Succ[i] + j" - using s7_1 s7_2 by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[iSucc[j + i]]" - proof - - have s7_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s7_2: "i \ Nat" - using s4_2_idom by auto - show ?thesis - using s7_1 s7_2 addint_succ_nat by auto - qed - also have s6_5: "... = iSucc[iPred[j + i]]" - proof - - have s7_1: "j + i \ Int" - proof - - have s8_1: "j \ Int" - using s3_2_asm1 neg_nat_in_int by auto - have s8_2: "i \ Int" - using s4_2_idom nat_in_int by auto - show ?thesis - using s8_1 s8_2 addIsInt by auto - qed - show ?thesis - using s7_1 iSucciPredCommute by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.j] + Succ[i] = Succ[i] + -.Succ[-.j]" - using s5_1 s5_2 by auto - qed - show "\ i \ Nat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 - natInduct[of "\ i. - -.Succ[-.j] + i = i + -.Succ[-.j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - neg_nat_induction[of "\ j. - \ i \ Nat: j + i = i + j"] - by auto - qed - have s2_2: "\ j \ negNat: \ i \ negNat: - j + i = i + j" - proof - - have s3_1: "\ i \ negNat: 0 + i = i + 0" - proof - - have s4_1: "0 + 0 = 0 + 0" - by auto - have s4_2: "\ i. \ - i \ negNat; - 0 + i = i + 0 - \ \ - 0 + -.Succ[-.i] = -.Succ[-.i] + 0" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ negNat" - assume s4_2_induct: "0 + i = i + 0" - have s5_1: "0 + -.Succ[-.i] = iPred[i]" - proof - - have s6_1: "0 + -.Succ[-.i] = iPred[0 + -.-.i]" - proof - - have s7_1: "0 \ Int" - using zero_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[0 + i]" - proof - - have s7_1: "-.-.i = i" - using s4_2_idom neg_nat_in_int - minus_involutive by auto - show ?thesis - using s7_1 by auto - qed - also have s6_3: "... = iPred[i + 0]" - using s4_2_induct by auto - also have s6_4: "... = iPred[i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "i + 0 = i" - using s7_1 add_0 by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "-.Succ[-.i] + 0 = iPred[i]" - proof - - have s6_1: "-.Succ[-.i] + 0 = -.Succ[-.i]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - show ?thesis - using s7_1 add_0 by auto - qed - also have s6_2: "... = iPred[i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "i \ Nat \ {0}" - using s4_2_idom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s7_1 s7_2 by auto - qed - finally show ?thesis . - qed - show "0 + -.Succ[-.i] = -.Succ[-.i] + 0" - using s5_1 s5_2 by auto - qed - show ?thesis - using s4_1 s4_2 - neg_nat_induction[of "\ i. 0 + i = i + 0"] - by blast - qed - have s3_2: "\ j. \ - j \ negNat; - \ i \ negNat: j + i = i + j - \ \ - \ i \ negNat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - proof - - fix "j" :: "c" - assume s3_2_jdom: "j \ negNat" - assume s3_2_induct: "\ i \ negNat: - j + i = i + j" - have s4_1: "-.Succ[-.j] + 0 = 0 + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] + 0 = iPred[j]" - proof - - have s6_1: "-.Succ[-.j] + 0 = -.Succ[-.j]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.j] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - show ?thesis - using s7_1 add_0 by auto - qed - also have s6_2: "... = iPred[j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "j \ Nat \ {0}" - using s3_2_jdom neg_nat_not_in_nat_setminus_0 - by auto - show ?thesis - unfolding iPred_def - using s7_1 s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "0 + -.Succ[-.j] = iPred[j]" - proof - - have s6_1: "0 + -.Succ[-.j] = iPred[0 + -.-.j]" - proof - - have s7_1: "0 \ Int" - using zero_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[0 + j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_1 s7_2 by auto - qed - also have s6_3: "... = iPred[j + 0]" - proof - - have s7_1: "0 \ negNat" - unfolding negNat_def - using neg0 by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + 0 = 0 + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[j]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - show ?thesis - using s7_1 add_0 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s5_1 s5_2 by auto - qed - have s4_2: "\ i. \ - i \ negNat; - -.Succ[-.j] + i = i + -.Succ[-.j] - \ \ - -.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j]" - proof - - fix "i" :: "c" - assume s4_2_idom: "i \ negNat" - assume s4_2_induct: "-.Succ[-.j] + i = i + -.Succ[-.j]" - have s5_1: "-.Succ[-.j] + -.Succ[-.i] = - iPred[iPred[i + j]]" - proof - - have s6_1: "-.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i]" - proof - - have s7_1: "-.Succ[-.j] \ Int" - proof - - have s8_1: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.j] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s7_3: "-.Succ[-.j] \ Int - ==> - -.i \ Nat - ==> - -.Succ[-.j] + -.Succ[-.i] = - iPred[-.Succ[-.j] + -.-.i]" - using addint_succ_negnat[ - of "-.Succ[-.j]" "-.i"] - by auto - show ?thesis - by (rule s7_3, rule s7_1, rule s7_2) - qed - also have s6_2: "... = iPred[-.Succ[-.j] + i]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "-.-.i = i" - using s7_1 minus_involutive by auto - show ?thesis - using s7_1 s7_2 by auto - qed - also have s6_3: "... = iPred[i + -.Succ[-.j]]" - using s4_2_induct by auto - also have s6_4: "... = iPred[iPred[i + -.-.j]]" - proof - - have s7_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat by auto - have s7_3: "i + -.Succ[-.j] = iPred[i + -.-.j]" - using s7_1 s7_2 addint_succ_negnat by auto - show ?thesis - using s7_3 by auto - qed - also have s6_5: "... = iPred[iPred[i + j]]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.-.j = j" - using s7_1 minus_involutive by auto - show ?thesis - using s7_2 by auto - qed - finally show ?thesis . - qed - have s5_2: "-.Succ[-.i] + -.Succ[-.j] = - iPred[iPred[i + j]]" - proof - - have s6_1: "-.Succ[-.i] + -.Succ[-.j] = - iPred[-.Succ[-.i] + -.-.j]" - proof - - have s7_1: "-.Succ[-.i] \ Int" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_int by auto - qed - have s7_2: "-.j \ Nat" - using s3_2_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_2: "... = iPred[-.Succ[-.i] + j]" - proof - - have s7_1: "-.-.j = j" - proof - - have s8_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_1 by auto - qed - also have s6_3: "... = iPred[j + -.Succ[-.i]]" - proof - - have s7_1: "-.Succ[-.i] \ negNat" - proof - - have s8_1: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat - by auto - have s8_2: "Succ[-.i] \ Nat" - using s8_1 succIsNat by auto - show ?thesis - using s8_2 minus_nat_in_neg_nat by auto - qed - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + -.Succ[-.i] = -.Succ[-.i] + j" - using s7_1 s7_2 by auto - show ?thesis - using s7_3 by auto - qed - also have s6_4: "... = iPred[iPred[j + -.-.i]]" - proof - - have s7_1: "j \ Int" - using s3_2_jdom neg_nat_in_int by auto - have s7_2: "-.i \ Nat" - using s4_2_idom minus_neg_nat_in_nat by auto - show ?thesis - using s7_1 s7_2 addint_succ_negnat by auto - qed - also have s6_5: "... = iPred[iPred[j + i]]" - proof - - have s7_1: "-.-.i = i" - proof - - have s8_1: "i \ Int" - using s4_2_idom neg_nat_in_int by auto - show ?thesis - using s8_1 minus_involutive by auto - qed - show ?thesis - using s7_1 by auto - qed - also have s6_6: "... = iPred[iPred[i + j]]" - proof - - have s7_1: "i \ negNat" - using s4_2_idom by auto - have s7_2: "\ k \ negNat: - j + k = k + j" - using s3_2_induct by auto - have s7_3: "j + i = i + j" - using s7_1 s7_2 spec by auto - show ?thesis - using s7_3 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.j] + -.Succ[-.i] = - -.Succ[-.i] + -.Succ[-.j]" - using s5_1 s5_2 by auto - qed - show "\ i \ negNat: - -.Succ[-.j] + i = i + -.Succ[-.j]" - using s4_1 s4_2 - neg_nat_induction[of "\ i. - -.Succ[-.j] + i = i + -.Succ[-.j]"] - by auto - qed - show ?thesis - using s3_1 s3_2 - neg_nat_induction[of "\ j. - \ i \ negNat: j + i = i + j"] - by auto - qed - show ?thesis - proof - - have s3_1: "\ j. \ i. - j \ negNat => ( - i \ Int => ( - j + i = i + j))" - proof auto - fix "j" :: "c" - fix "i" :: "c" - assume s4_1: "j \ negNat" - assume s4_2: "i \ Int" - have s4_3: "i \ Nat ==> j + i = i + j" - proof - - assume s5_1: "i \ Nat" - show "j + i = i + j" - using s4_1 s5_1 s2_1 by auto - qed - have s4_4: "i \ negNat ==> j + i = i + j" - proof - - assume s5_1: "i \ negNat" - show "j + i = i + j" - using s4_1 s5_1 s2_2 by auto - qed - show "j + i = i + j" - using s4_2 s4_3 s4_4 int_union_nat_negnat by auto - qed - show ?thesis - using s3_1 by auto - qed - qed - show ?thesis - using s1_1 s1_2 int_union_nat_negnat by auto - qed - - -theorem AddCommutative_sequent: - assumes mdom: "m \ Int" and ndom: "n \ Int" - shows "m + n = n + m" - using mdom ndom AddCommutative by auto - - -theorem add_0_left: - assumes mdom: "m \ Int" - shows "0 + m = m" - proof - - have s1_1: "m + 0 = m" - using mdom add_0 by auto - have s1_2: "m + 0 = 0 + m" - using mdom zero_in_int AddCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - - -(*----------------------------------------------------------------------------*) -(* Associativity of addition. *) - - -(* -THEOREM iSuccRightDistributesAdd == - ASSUME NEW a \in Int - PROVE \A b \in Int: iSucc[a + b] = a + iSucc[b] -<1>1. ASSUME NEW b \in Nat - PROVE iSucc[a + b] = a + iSucc[b] - <2>1. iSucc[a + b] = a + Succ[b] - <3>1. a \in Int - OBVIOUS - <3>2. b \in Nat - BY <1>1 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>2. @ = a + iSucc[b] - <3>1. b \in Nat - BY <1>1 - <3>2. b \in Int - BY <1>1, nat_in_int - <3>3. iSucc[b] = Succ[b] - BY <3>1, <3>2 DEF iSucc - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2 -<1>2. \A b \in negNat: iSucc[a + b] = a + iSucc[b] - <2>1. iSucc[a + 0] = a + iSucc[0] - <3>1. iSucc[a + 0] = a + Succ[0] - <4>1. a \in Int - OBVIOUS - <4>2. 0 \in Nat - BY zeroIsNat - <4> QED - BY <4>1, <4>2, addint_succ_nat - <3>2. @ = a + iSucc[0] - <4>1. 0 \in Nat - BY zeroIsNat - <4>2. 0 \in Int - BY zeroIsNat, nat_in_int - <4>3. iSucc[0] = Succ[0] - BY <4>1, <4>2 DEF iSucc - <4> QED - BY <4>3 - <3> QED - BY <3>1, <3>2 - <2>2. ASSUME NEW b \in negNat, - iSucc[a + b] = a + iSucc[b] - PROVE iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]] - <3>1. iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]] - <4>1. a \in Int - OBVIOUS - <4>2. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <4>3. a + -.Succ[-.b] = iPred[a + -.-.b] - BY <4>1, <4>2, addint_succ_negnat - <4>4. -.-.b = b - <5>1. b \in Int - BY <2>2, neg_nat_in_int - <5> QED - BY <5>1, minus_involutive - <4>5. a + -.Succ[-.b] = iPred[a + b] - BY <4>3, <4>4 - <4> QED - BY <4>5 - <3>2. @ = iPred[iSucc[a + b]] - <4>1. a + b \in Int - <5>1. a \in Int - OBVIOUS - <5>2. b \in Int - BY <2>2, neg_nat_in_int - <5> QED - BY <5>1, <5>2, addIsInt - <4> QED - BY <4>1, iSucciPredCommute - <3>3. CASE b = 0 - <4>1. iPred[iSucc[a + b]] = a - <5>1. iPred[iSucc[a + b]] = iPred[iSucc[a + 0]] - BY <3>3 - <5>2. @ = iPred[iSucc[a]] - <6>1. a + 0 = a - BY a \in Int, add_0 - <6> QED - BY <6>1 - <5>3. @ = a - <6>1. a \in Int - OBVIOUS - <6> QED - BY <6>1, iSucciPredCommute, spec - <5> QED - BY <5>1, <5>2, <5>3 - <4>2. a + iSucc[-.Succ[-.b]] = a - <5>1. a + iSucc[-.Succ[-.b]] = a + iSucc[-.Succ[0]] - BY <3>3 - <5>2. @ = a + iSucc[-.1] - OBVIOUS \* 1 is an abbreviation - <5>3. @ = a + -.Pred[-.-.1] - <6>1. -.1 \in Int - BY oneIsNat, minus_nat_in_int - <6>2. -.1 \notin Nat - <7>1. 1 \in Nat \ {0} - BY oneIsNat succNotZero - <7> QED - BY <7>1, minus_nat_0_in_negnat_0, - neg_nat_0_not_in_nat - <6>3. iSucc[-.1] = -.Pred[-.-.1] - BY <6>1, <6>2 DEF iSucc - <6> QED - BY <6>3 - <5>4. @ = a + -.Pred[1] - <6>1. -.-.1 = 1 - <7>1. 1 \in Int - BY oneIsNat nat_in_int - <7> QED - BY <7>1, minus_involutive - <6> QED - BY <6>1 - <5>5. @ = a + -.0 - BY pred_1 - <5>6. @ = a + 0 - BY neg0 - <5>7. @ = a - BY a \in Int, add_0 - <5> QED - BY <5>1, <5>2, <5>3, <5>4, - <5>5, <5>6, <5>7 - <4> QED - BY <3>1, <3>2, <4>1, <4>2 - <3>4. CASE b # 0 - <4>1. iPred[iSucc[a + b]] = - iPred[a + iSucc[b]] - BY <2>2 - <4>2. @ = a + -.Succ[-.iSucc[b]] - <5>1. a \in Int - OBVIOUS - <5>2. /\ -.iSucc[b] \in Nat - /\ iSucc[b] \in negNat - <6>1. b \in negNat \ {0} - BY <2>2, <3>4 - <6>2. iSucc[b] = -.Pred[-.b] - <7>1. b \in Int - BY <6>1, neg_nat_in_int - <7>2. b \notin Nat - BY <6>1, neg_nat_0_not_in_nat - <7> QED - BY <7>1, <7>2 DEF iSucc - <6>3. Pred[-.b] \in Nat - <7>1. -.b \in Nat \ {0} - BY <6>1, minus_neg_nat_0_in_nat_0 - <7> QED - BY <7>1, Pred_in_nat - <6>4. -.iSucc[b] \in Nat - <7>1. -.iSucc[b] = Pred[-.b] - <8>1. -.iSucc[b] = -.-.Pred[-.b] - BY <6>2 - <8>2. Pred[-.b] \in Int - BY <6>3, nat_in_int - <8>3. -.-.Pred[-.b] = Pred[-.b] - BY <8>2, minus_involutive - <8> QED - BY <8>1, <8>3 - <7> QED - BY <6>3, <7>1 - <6>5. iSucc[b] \in negNat - BY <6>2, <6>3, minus_nat_in_neg_nat - <6> QED - BY <6>4, <6>5 - <5>3. a + -.Succ[-.iSucc[b]] = - iPred[a + -.-.iSucc[b]] - BY <5>1, <5>2, addint_succ_negnat - <5>4. -.-.iSucc[b] = iSucc[b] - <6>1. iSucc[b] \in Int - BY <5>2, neg_nat_in_int - <6> QED - BY <6>1, minus_involutive - <5> QED - BY <5>3, <5>4 - <4>3. @ = a + -.Succ[Pred[-.b]] - <5>1. b \in Int - BY <2>2, neg_nat_in_int - <5>2. b \notin Nat - <6>1. b \in negNat \ {0} - BY <2>2, <3>4 - <6> QED - BY <6>1, neg_nat_0_not_in_nat - <5>3. iSucc[b] = -.Pred[-.b] - BY <5>1, <5>2 DEF iSucc - <5>4. -.iSucc[b] = -.-.Pred[-.b] - BY <5>3 - <5>5. -.-.Pred[-.b] = Pred[-.b] - <6>1. -.b \in Nat - BY <2>2, <3>4, minus_neg_nat_0_in_nat_0 - <6>2. Pred[-.b] \in Nat - BY <6>1, Pred_in_nat - <6>3. Pred[-.b] \in Int - BY <6>2, nat_in_int - <6> QED - BY <6>3, minus_involutive - <5>6. -.iSucc[b] = Pred[-.b] - BY <5>4, <5>5 - <5> QED - BY <5>6 - <4>4. @ = a + -.Pred[Succ[-.b]] - <5>1. b \in negNat \ {0} - BY <2>2, <3>4 - <5>2. -.b \in Nat \ {0} - BY <5>1, minus_neg_nat_0_in_nat_0 - <5>3. Succ[Pred[-.b]] = Pred[Succ[-.b]] - BY <5>2, Succ_Pred_nat, Pred_Succ_nat - <5> QED - BY <5>3 - <4>5. @ = a + -.Pred[-.-.Succ[-.b]] - <5>1. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <5>2. Succ[-.b] \in Nat - BY <5>1, succIsNat - <5>3. Succ[-.b] \in Int - BY <5>2, nat_in_int - <5>4. -.-.Succ[-.b] = Succ[-.b] - BY <5>3, minus_involutive - <5> QED - BY <5>4 - <4>6. @ = a + iSucc[-.Succ[-.b]] - <5>1. -.b \in Nat - BY <2>2, minus_neg_nat_in_nat - <5>2. Succ[-.b] \in Nat \ {0} - BY <5>1, succIsNat, succNotZero - <5>3. -.Succ[-.b] \in negNat \ {0} - BY <5>2, minus_nat_0_in_negnat_0 - <5>4. -.Succ[-.b] \in Int - BY <5>3, neg_nat_in_int - <5>5. -.Succ[-.b] \notin Nat - BY <5>3, neg_nat_0_not_in_nat - <5> QED - BY <5>4, <5>5 DEF iSucc - <4> QED - BY <3>1, <3>2, <4>1, <4>2, <4>3, - <4>4, <4>5, <4>6 - <3> QED - BY <3>3, <3>4 - <2> QED - BY <2>1, <2>2, neg_nat_induction -<1> QED - BY <1>1, <1>2 -*) -theorem iSuccRightDistributesAdd: - "\ a \ Int: \ b \ Int: - iSucc[a + b] = a + iSucc[b]" - proof - - { - fix "a" :: "c" - assume adom: "a \ Int" - have s1_1: "\ b \ Nat: iSucc[a + b] = a + iSucc[b]" - proof auto - fix "b" :: "c" - assume bdom: "b \ Nat" - have s2_1: "iSucc[a + b] = a + Succ[b]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b \ Nat" - using bdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - also have s2_2: "... = a + iSucc[b]" - proof - - have s3_1: "b \ Nat" - using bdom by auto - have s3_2: "b \ Int" - using bdom nat_in_int by auto - have s3_3: "iSucc[b] = Succ[b]" - unfolding iSucc_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "iSucc[a + b] = a + iSucc[b]" . - qed - have s1_2: "\ b \ negNat: iSucc[a + b] = a + iSucc[b]" - proof - - have s2_1: "iSucc[a + 0] = a + iSucc[0]" - proof - - have s3_1: "iSucc[a + 0] = a + Succ[0]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s4_1 s4_2 addint_succ_nat by auto - qed - have s3_2: "... = a + iSucc[0]" - proof - - have s4_1: "0 \ Nat" - using zeroIsNat by auto - have s4_2: "0 \ Int" - using zeroIsNat nat_in_int by auto - have s4_3: "iSucc[0] = Succ[0]" - unfolding iSucc_def - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - have s2_2: "\ b. \ - b \ negNat; - iSucc[a + b] = a + iSucc[b] - \ \ - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - fix "b" :: "c" - assume s2_2_bdom: "b \ negNat" - assume s2_2_induct: "iSucc[a + b] = a + iSucc[b]" - have s3_1: "iSucc[a + -.Succ[-.b]] = iSucc[iPred[a + b]]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s4_3: "a + -.Succ[-.b] = iPred[a + -.-.b]" - using s4_1 s4_2 addint_succ_negnat by auto - have s4_4: "-.-.b = b" - proof - - have s5_1: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - show ?thesis - using s5_1 minus_involutive by auto - qed - have s4_5: "a + -.Succ[-.b] = iPred[a + b]" - using s4_3 s4_4 by auto - show ?thesis - using s4_5 by auto - qed - have s3_2: "... = iPred[iSucc[a + b]]" - proof - - have s4_1: "a + b \ Int" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - show ?thesis - using s5_1 s5_2 addIsInt by auto - qed - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - have s3_3: "b = 0 ==> - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - assume s3_3_bdom: "b = 0" - have s4_1: "iPred[iSucc[a + b]] = a" - proof - - have s5_1: "iPred[iSucc[a + b]] = iPred[iSucc[a + 0]]" - using s3_3_bdom by auto - also have s5_2: "... = iPred[iSucc[a]]" - proof - - have s6_1: "a + 0 = a" - using adom add_0 by auto - show ?thesis - using s6_1 by auto - qed - also have s5_3: "... = a" - using adom iSucciPredCommute spec by auto - finally show ?thesis . - qed - have s4_2: "a + iSucc[-.Succ[-.b]] = a" - proof - - have s5_1: "a + iSucc[-.Succ[-.b]] = - a + iSucc[-.Succ[0]]" - using s3_3_bdom by auto - have s5_2: "... = a + iSucc[-.1]" - by auto - have s5_3: "... = a + -.Pred[-.-.1]" - (* `also` raises error: - empty result sequence - *) - proof - - have s6_1: "-.1 \ Int" - using oneIsNat minus_nat_in_int by auto - have s6_2: "-.1 \ Nat" - proof - - have s7_1: "1 \ Nat \ {0}" - using oneIsNat succNotZero by auto - show ?thesis - using s7_1 minus_nat_0_in_negnat_0 - neg_nat_0_not_in_nat by auto - qed - have s6_3: "iSucc[-.1] = -.Pred[-.-.1]" - unfolding iSucc_def - using s6_1 s6_2 by auto - show ?thesis - using s6_3 by auto - qed - have s5_4: "... = a + -.Pred[1]" - proof - - have s6_1: "-.-.1 = 1" - proof - - have s7_1: "1 \ Int" - using oneIsNat nat_in_int by auto - show ?thesis - using s7_1 minus_involutive by auto - qed - show ?thesis - using s6_1 by auto - qed - have s5_5: "... = a + -.0" - using pred_1 by auto - have s5_6: "... = a + 0" - using neg0 by auto - have s5_7: "... = a" - using adom add_0 by auto - show ?thesis - using s5_1 s5_2 s5_3 s5_4 s5_5 s5_6 s5_7 - by auto - qed - show ?thesis - using s3_1 s3_2 s4_1 s4_2 by auto - qed - have s3_4: "b \ 0 ==> - iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - proof - - assume s3_4_bdom: "b \ 0" - have s4_1: "iPred[iSucc[a + b]] = - iPred[a + iSucc[b]]" - using s2_2_induct by auto - also have s4_2: "... = a + -.Succ[-.iSucc[b]]" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "-.iSucc[b] \ Nat \ - iSucc[b] \ negNat" - proof - - have s6_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - have s6_2: "iSucc[b] = -.Pred[-.b]" - proof - - have s7_1: "b \ Int" - using s6_1 neg_nat_in_int by auto - have s7_2: "b \ Nat" - using s6_1 neg_nat_0_not_in_nat by auto - show ?thesis - unfolding iSucc_def - using s7_1 s7_2 by auto - qed - have s6_3: "Pred[-.b] \ Nat" - proof - - have s7_1: "-.b \ Nat \ {0}" - using s6_1 minus_neg_nat_0_in_nat_0 by auto - show ?thesis - using s7_1 Pred_in_nat by auto - qed - have s6_4: "-.iSucc[b] \ Nat" - proof - - have s7_1: "-.iSucc[b] = Pred[-.b]" - proof - - have s8_1: "-.iSucc[b] = -.-.Pred[-.b]" - using s6_2 by auto - have s8_2: "Pred[-.b] \ Int" - using s6_3 nat_in_int by auto - have s8_3: "-.-.Pred[-.b] = Pred[-.b]" - using s8_2 minus_involutive by auto - show ?thesis - using s8_1 s8_3 by auto - qed - show ?thesis - using s6_3 s7_1 by auto - qed - have s6_5: "iSucc[b] \ negNat" - using s6_2 s6_3 minus_nat_in_neg_nat by auto - show ?thesis - using s6_4 s6_5 by auto - qed - have s5_3: "a + -.Succ[-.iSucc[b]] = - iPred[a + -.-.iSucc[b]]" - using s5_1 s5_2 addint_succ_negnat by auto - have s5_4: "-.-.iSucc[b] = iSucc[b]" - proof - - have s6_1: "iSucc[b] \ Int" - using s5_2 neg_nat_in_int by auto - show ?thesis - using s6_1 minus_involutive by auto - qed - show ?thesis - using s5_3 s5_4 by auto - qed - also have s4_3: "... = a + -.Succ[Pred[-.b]]" - proof - - have s5_1: "b \ Int" - using s2_2_bdom neg_nat_in_int by auto - have s5_2: "b \ Nat" - proof - - have s6_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - show ?thesis - using s6_1 neg_nat_0_not_in_nat by auto - qed - have s5_3: "iSucc[b] = -.Pred[-.b]" - unfolding iSucc_def - using s5_1 s5_2 by auto - have s5_4: "-.iSucc[b] = -.-.Pred[-.b]" - using s5_3 by auto - have s5_5: "-.-.Pred[-.b] = Pred[-.b]" - proof - - have s6_1: "-.b \ Nat" - using s2_2_bdom s3_4_bdom - minus_neg_nat_0_in_nat_0 - by auto - have s6_2: "Pred[-.b] \ Nat" - using s6_1 Pred_in_nat by auto - have s6_3: "Pred[-.b] \ Int" - using s6_2 nat_in_int by auto - show ?thesis - using s6_3 minus_involutive by auto - qed - have s5_6: "-.iSucc[b] = Pred[-.b]" - using s5_4 s5_5 by auto - show ?thesis - using s5_6 by auto - qed - also have s4_4: "... = a + -.Pred[Succ[-.b]]" - proof - - have s5_1: "b \ negNat \ {0}" - using s2_2_bdom s3_4_bdom by auto - have s5_2: "-.b \ Nat \ {0}" - using s5_1 minus_neg_nat_0_in_nat_0 by auto - have s5_3: "Succ[Pred[-.b]] = Pred[Succ[-.b]]" - using s5_2 Succ_Pred_nat Pred_Succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = a + -.Pred[-.-.Succ[-.b]]" - proof - - have s5_1: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s5_2: "Succ[-.b] \ Nat" - using s5_1 succIsNat by auto - have s5_3: "Succ[-.b] \ Int" - using s5_2 nat_in_int by auto - have s5_4: "-.-.Succ[-.b] = Succ[-.b]" - using s5_3 minus_involutive by auto - show ?thesis - using s5_4 by auto - qed - also have s4_6: "... = a + iSucc[-.Succ[-.b]]" - proof - - have s5_1: "-.b \ Nat" - using s2_2_bdom minus_neg_nat_in_nat by auto - have s5_2: "Succ[-.b] \ Nat \ {0}" - using s5_1 succIsNat succNotZero by auto - have s5_3: "-.Succ[-.b] \ negNat \ {0}" - using s5_2 minus_nat_0_in_negnat_0 by auto - have s5_4: "-.Succ[-.b] \ Int" - using s5_3 neg_nat_in_int by auto - have s5_5: "-.Succ[-.b] \ Nat" - using s5_3 neg_nat_0_not_in_nat by auto - show ?thesis - unfolding iSucc_def - using s5_4 s5_5 by auto - qed - finally show "iSucc[a + -.Succ[-.b]] = - a + iSucc[-.Succ[-.b]]" - using s3_1 s3_2 by auto - qed - show "iSucc[a + -.Succ[-.b]] = a + iSucc[-.Succ[-.b]]" - using s3_3 s3_4 by blast - qed - show ?thesis - using s2_1 s2_2 - neg_nat_induction[of "\ b. - iSucc[a + b] = a + iSucc[b]"] - by auto - qed - have s1_3: "\ b \ Int: iSucc[a + b] = a + iSucc[b]" - using s1_1 s1_2 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem iSuccRightDistributesAdd_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "iSucc[a + b] = a + iSucc[b]" - using adom bdom iSuccRightDistributesAdd by auto - - -(* -THEOREM iPredRightDistributesAdd == - ASSUME NEW a \in Int - PROVE \A b \in Int: iPred[a + b] = a + iPred[b] -PROOF -<1>1. iPred[a + 0] = a + iPred[0] - <2>1. iPred[a + 0] = iPred[a + -.0] - BY neg0 - <2>2. @ = a + -.Succ[0] - <3>1. a \in Int - OBVIOUS - <3>2. 0 \in Nat - BY zeroIsNat - <3> QED - BY <3>1, <3>2, addint_succ_negnat - <2>3. @ = a + -.1 - OBVIOUS \* 1 is an abbreviation - <2>4. @ = a + iPred[0] - BY ipred_0 - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>2. \A b \in Nat: iPred[a + b] = a + iPred[b] - <2>1. ASSUME NEW b \in Nat, - iPred[a + b] = a + iPred[b] - PROVE iPred[a + Succ[b]] = a + iPred[Succ[b]] - <3>1. iPred[a + Succ[b]] = iPred[iSucc[a + b]] - <4>1. a \in Int - OBVIOUS - <4>2. b \in Nat - BY <2>1 - <4>3. a + Succ[b] = iSucc[a + b] - BY <4>1, <4>2, addint_succ_nat - <4> QED - BY <4>3 - <3>2. @ = iSucc[iPred[a + b]] - <4>1. a + b \in Int - <5>1. a \in Int - OBVIOUS - <5>2. b \in Int - BY <2>2, nat_in_int - <5> QED - BY <5>1, <5>2, addIsInt - <4> QED - BY <4>1, iSucciPredCommute - <3>3. @ = iSucc[a + iPred[b]] - BY <2>1 - <3>4. @ = a + iSucc[iPred[b]] - <4>1. b \in Int - BY <2>1, nat_in_int - <4>2. iPred[b] \in Int - BY <4>1, iPred_type - <4>3. a \in Int - OBVIOUS - <4> QED - BY <4>2, <4>3, iSuccRightDistributesAdd - <3>5. @ = a + iPred[iSucc[b]] - <4>1. b \in Int - BY <2>1, nat_in_int - <4> QED - BY <4>1, iSucciPredCommute - <3>6. @ = a + iPred[Succ[b]] - <4>1. b \in Nat - BY <2>1 - <4>2. b \in Int - BY <2>1, nat_in_int - <4>3. iSucc[b] = Succ[b] - BY <4>1, <4>2, nat_in_int DEF iSucc - <4> QED - BY <4>3 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 - <2> QED - BY <2>1, <1>1, NatInduction -<1>3. \A b \in negNat \ {0}: iPred[a + b] = a + iPred[b] - <2>1. ASSUME NEW b \in negNat \ {0} - PROVE iPred[a + b] = a + iPred[b] - <3>1. iPred[a + b] = iPred[a + -.-.b] - <4>1. b \in Int - BY <2>1, neg_nat_in_int - <4>2. -.-.b = b - BY <4>1, minus_involutive - <4> QED - BY <4>2 - <3>2. @ = a + -.Succ[-.b] - <4>1. a \in Int - OBVIOUS - <4>2. -.b \in Nat - BY <2>1, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, addint_succ_negnat - <3>3. @ = a + iPred[b] - <4>1. b \in Int - BY <2>1, neg_nat_in_int - <4>2. b \notin Nat \ {0} - BY <2>1, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3> QED - BY <3>1, <3>2, <3>3 - <2> QED - BY <2>1 -<1> QED - BY <1>2, <1>3 -*) -theorem iPredRightDistributesAdd: - "\ a \ Int: \ b \ Int: - iPred[a + b] = a + iPred[b]" - proof - - { - fix "a" :: "c" - assume adom: "a \ Int" - have s1_1: "iPred[a + 0] = a + iPred[0]" - proof - - have s2_1: "iPred[a + 0] = iPred[a + -.0]" - using neg0 by auto - also have s2_2: "... = a + -.Succ[0]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "0 \ Nat" - using zeroIsNat by auto - show ?thesis - using s3_1 s3_2 addint_succ_negnat by auto - qed - also have s2_3: "... = a + -.1" - by auto - have s2_4: "... = a + iPred[0]" - using ipred_0 by auto - (* `also` raises an error here *) - finally show ?thesis - using s2_4 by auto - qed - have s1_2: "\ b \ Nat: - iPred[a + b] = a + iPred[b]" - proof - - have s2_1: "\ b. \ - b \ Nat; - iPred[a + b] = a + iPred[b] - \ \ - iPred[a + Succ[b]] = a + iPred[Succ[b]]" - proof - - fix "b" :: "c" - assume s2_1_bdom: "b \ Nat" - assume s2_1_induct: "iPred[a + b] = a + iPred[b]" - have s3_1: "iPred[a + Succ[b]] = iPred[iSucc[a + b]]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "b \ Nat" - using s2_1_bdom by auto - have s4_3: "a + Succ[b] = iSucc[a + b]" - using s4_1 s4_2 addint_succ_nat by auto - show ?thesis - using s4_3 by auto - qed - also have s3_2: "... = iSucc[iPred[a + b]]" - proof - - have s4_1: "a + b \ Int" - proof - - have s5_1: "a \ Int" - using adom by auto - have s5_2: "b \ Int" - using s2_1_bdom nat_in_int by auto - show ?thesis - using s5_1 s5_2 addIsInt by auto - qed - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - also have s3_3: "... = iSucc[a + iPred[b]]" - using s2_1_induct by auto - also have s3_4: "... = a + iSucc[iPred[b]]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom nat_in_int by auto - have s4_2: "iPred[b] \ Int" - using s4_1 iPred_type by auto - have s4_3: "a \ Int" - using adom by auto - show ?thesis - using s4_2 s4_3 iSuccRightDistributesAdd - by auto - qed - also have s3_5: "... = a + iPred[iSucc[b]]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom nat_in_int by auto - show ?thesis - using s4_1 iSucciPredCommute by auto - qed - also have s3_6: "... = a + iPred[Succ[b]]" - proof - - have s4_1: "b \ Nat" - using s2_1_bdom by auto - have s4_2: "b \ Int" - using s2_1_bdom nat_in_int by auto - have s4_3: "iSucc[b] = Succ[b]" - unfolding iSucc_def - using s4_1 s4_2 by auto - show ?thesis - using s4_3 by auto - qed - finally show "iPred[a + Succ[b]] = a + iPred[Succ[b]]" . - qed - show ?thesis - using s1_1 s2_1 - natInduct[of "\ b. - iPred[a + b] = a + iPred[b]"] - by auto - qed - have s1_3: "\ b \ negNat \ {0}: - iPred[a + b] = a + iPred[b]" - proof - - { - fix "b" :: "c" - assume s2_1_bdom: "b \ negNat \ {0}" - { - have s3_1: "iPred[a + b] = iPred[a + -.-.b]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom neg_nat_in_int by auto - have s4_2: "-.-.b = b" - using s4_1 minus_involutive by auto - show ?thesis - using s4_2 by auto - qed - also have s3_2: "... = a + -.Succ[-.b]" - proof - - have s4_1: "a \ Int" - using adom by auto - have s4_2: "-.b \ Nat" - using s2_1_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 addint_succ_negnat by auto - qed - also have s3_3: "... = a + iPred[b]" - proof - - have s4_1: "b \ Int" - using s2_1_bdom neg_nat_in_int by auto - have s4_2: "b \ Nat \ {0}" - proof - - have s5_1: "b \ negNat" - using s2_1_bdom by auto - show ?thesis - using s5_1 neg_nat_not_in_nat_setminus_0 - by auto - qed - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - finally have "iPred[a + b] = a + iPred[b]" . - } - from this have s2_2: "iPred[a + b] = a + iPred[b]" - by auto - } - from this show ?thesis - by auto - qed - have s1_4: "\ b \ Int: - iPred[a + b] = a + iPred[b]" - using s1_2 s1_3 int_union_nat_negnat_0 spec allI by auto - } - from this show ?thesis - by auto - qed - - -theorem iPredRightDistributesAdd_sequent: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "iPred[m + n] = m + iPred[n]" - using mdom ndom iPredRightDistributesAdd by auto - - -theorem isucc_eq_add_1: - assumes idom: "i \\in Int" - shows "iSucc[i] = i + 1" - proof - - have s1_1: "iSucc[i] = iSucc[i + 0]" - using idom add_0 by auto - also have s1_2: "... = i + iSucc[0]" - proof - - have s2_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - show ?thesis - using idom s2_1 iSuccRightDistributesAdd_sequent - by auto - qed - also have s1_3: "... = i + 1" - proof - - have s2_1: "iSucc[0] = Succ[0]" - proof - - have s3_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s3_2: "0 \\in Nat" - using zeroIsNat by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - have s2_2: "Succ[0] = 1" - by auto - have s2_3: "iSucc[0] = 1" - using s2_1 s2_2 by auto - show ?thesis - using s2_3 s1_2 by auto - qed - finally show ?thesis . - qed - - -(* -THEOREM AddAssociative == - \A a \in Int: \A b \in Int: \A c \in Int: - (a + b) + c = a + (b + c) -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE \A c \in Int: - (a + b) + c = a + (b + c) - BY <1>1 -<1>2. (a + b) + 0 = a + (b + 0) - <2>1. (a + b) + 0 = a + b - <3>1. a + b \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, add_0 - <2>2. @ = a + (b + 0) - <3>1. b \in Int - BY <1>1 - <3>2. b + 0 = b - BY <3>1, add_0 - <3> QED - BY <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW c \in Nat, - (a + b) + c = a + (b + c) - PROVE (a + b) + Succ[c] = a + (b + Succ[c]) - <2>1. (a + b) + Succ[c] = iSucc[(a + b) + c] - <3>1. a + b \in Int - BY <1>1, addIsInt - <3>2. c \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>2. @ = iSucc[a + (b + c)] - BY <1>3 \* induction hypothesis - <2>3. @ = a + iSucc[b + c] - <3>1. a \in Int - BY <1>1 - <3>2. b + c \in Int - BY <1>1, <1>3, nat_in_int, addIsInt - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>4. @ = a + (b + iSucc[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>3, nat_in_int - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>5. @ = a + (b + Succ[c]) - <3>1. c \in Nat - BY <1>3 - <3>2. c \in Int - BY <3>1, nat_in_int - <3>3. iSucc[c] = Succ[c] - BY <3>1, <3>2 DEF iSucc - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5 -<1>4. ASSUME NEW c \in negNat, - (a + b) + c = a + (b + c) - PROVE (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c]) - <2>1. (a + b) + -.Succ[-.c] = (a + b) + iPred[c] - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. c \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <3>3. iPred[c] = -.Succ[-.c] - BY <3>1, <3>2 DEF iPred - <3> QED - BY <3>3 - <2>2. @ = iPred[(a + b) + c] - <3>1. a + b \in Int - BY <1>1, addIsInt - <3>2. c \in Int - BY <1>4, neg_nat_in_int - <3>3 QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>3. @ = iPred[a + (b + c)] - BY <1>4 \* induction hypothesis - <2>4. @ = a + iPred[b + c] - <3>1. a \in Int - BY <1>1 - <3>2. b + c \in Int - BY <1>1, <1>4, neg_nat_in_int, addIsInt - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>5. @ = a + (b + iPred[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>4, neg_nat_in_int - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>6. @ = a + (b + -.Succ[-.c]) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. c \notin Nat \ {0} - BY <1>4, neg_nat_not_in_nat_setminus_0 - <3>3. iPred[c] = -.Succ[-.c] - BY <3>1, <3>2 DEF iPred - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 -<1>5. \A c \in Nat: (a + b) + c = a + (b + c) - BY <1>2, <1>3, NatInduction -<1>6. \A c \in negNat: (a + b) + c = a + (b + c) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6 -*) -theorem AddAssociative: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - (a + b) + c = a + (b + c)" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \ Int" - assume bdom: "b \ Int" - have s1_2: "(a + b) + 0 = a + (b + 0)" - proof - - have s2_1: "(a + b) + 0 = a + b" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - show ?thesis - using s3_1 add_0 by auto - qed - have s2_2: "... = a + (b + 0)" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "b + 0 = b" - using s3_1 add_0 by auto - show ?thesis - using s3_2 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ c. \ - c \ Nat; - (a + b) + c = a + (b + c) - \ \ - (a + b) + Succ[c] = a + (b + Succ[c])" - proof - - fix "c" :: "c" - assume s1_3_cdom: "c \ Nat" - assume s1_3_induct: "(a + b) + c = a + (b + c)" - have s2_1: "(a + b) + Succ[c] = iSucc[(a + b) + c]" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - have s3_2: "c \ Nat" - using s1_3_cdom by auto - show ?thesis - using s3_1 s3_2 addint_succ_nat by auto - qed - also have s2_2: "... = iSucc[a + (b + c)]" - using s1_3_induct by auto - also have s2_3: "... = a + iSucc[b + c]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b + c \ Int" - using bdom s1_3_cdom nat_in_int addIsInt by auto - show ?thesis - using s3_1 s3_2 iSuccRightDistributesAdd by auto - qed - also have s2_4: "... = a + (b + iSucc[c])" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "c \ Int" - using s1_3_cdom nat_in_int by auto - have s3_3: "iSucc[b + c] = b + iSucc[c]" - using s3_1 s3_2 iSuccRightDistributesAdd by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = a + (b + Succ[c])" - proof - - have s3_1: "c \ Nat" - using s1_3_cdom by auto - have s3_2: "c \ Int" - using s1_3_cdom nat_in_int by auto - have s3_3: "iSucc[c] = Succ[c]" - unfolding iSucc_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a + b) + Succ[c] = a + (b + Succ[c])" . - qed - have s1_4: "\ c. \ - c \ negNat; - (a + b) + c = a + (b + c) - \ \ - (a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" - proof - - fix "c" :: "c" - assume s1_4_cdom: "c \ negNat" - assume s1_4_induct: "(a + b) + c = a + (b + c)" - have s2_1: "(a + b) + -.Succ[-.c] = (a + b) + iPred[c]" - proof - - have s3_1: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "c \ Nat \ {0}" - using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto - have s3_3: "iPred[c] = -.Succ[-.c]" - unfolding iPred_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - also have s2_2: "... = iPred[(a + b) + c]" - proof - - have s3_1: "a + b \ Int" - using adom bdom addIsInt by auto - have s3_2: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by force - qed - also have s2_3: "... = iPred[a + (b + c)]" - using s1_4_induct by auto - also have s2_4: "... = a + iPred[b + c]" - proof - - have s3_1: "a \ Int" - using adom by auto - have s3_2: "b + c \ Int" - using bdom s1_4_cdom neg_nat_in_int addIsInt by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by auto - qed - also have s2_5: "... = a + (b + iPred[c])" - proof - - have s3_1: "b \ Int" - using bdom by auto - have s3_2: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_3: "iPred[b + c] = b + iPred[c]" - using s3_1 s3_2 iPredRightDistributesAdd by force - show ?thesis - using s3_3 by auto - qed - also have s2_6: "... = a + (b + -.Succ[-.c])" - proof - - have s3_1: "c \ Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "c \ Nat \ {0}" - using s1_4_cdom neg_nat_not_in_nat_setminus_0 by auto - have s3_3: "iPred[c] = -.Succ[-.c]" - unfolding iPred_def - using s3_1 s3_2 by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a + b) + -.Succ[-.c] = a + (b + -.Succ[-.c])" . - qed - have s1_5: "\ c \ Nat: - (a + b) + c = a + (b + c)" - using s1_2 s1_3 natInduct[of "\ c. - (a + b) + c = a + (b + c)"] - by auto - have s1_6: "\ c \ negNat: - (a + b) + c = a + (b + c)" - using s1_2 s1_4 neg_nat_induction[of "\ c. - (a + b) + c = a + (b + c)"] - by auto - have "\ c \ Int: - (a + b) + c = a + (b + c)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem AddAssociative_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - and cdom: "c \\in Int" - shows "(a + b) + c = a + (b + c)" - using adom bdom cdom AddAssociative by auto - - -theorem AddLeftCommutativity: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - and cdom: "c \\in Int" - shows "b + (a + c) = a + (b + c)" - proof - - have s1_1: "(b + a) + c = (a + b) + c" - using adom bdom by (simp only: AddCommutative_sequent) - show ?thesis - using adom bdom cdom s1_1 by (simp only: AddAssociative_sequent) - qed - - -lemmas add_ac_int[algebra_simps] = - AddCommutative_sequent AddAssociative_sequent - AddLeftCommutativity add_0 add_0_left - - -(* A test that the simplification rules work. *) -theorem - assumes "a \\in Int" and "b \\in Int" and "c \\in Int" - shows "a + (b + c) = (a + b) + c" - using assms by (simp add: algebra_simps) - - -(* -THEOREM MinusSuccIsMinusOne == - ASSUME NEW n \in Nat - PROVE -.Succ[n] = -.n + -.1 -<1>1. -.Succ[n] = -.Succ[-.-.n] - <2>1. -.-.n = n - <3>1. n \in Nat - OBVIOUS - <3>2. n \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>2, minus_involutive - <2> QED - BY <2>1 -<1>2. @ = iPred[-.n] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. -.n \notin Nat \ {0} - <3>1. n \in Nat - OBVIOUS - <3>2. -.n \in negNat - BY <3>1, minus_nat_in_neg_nat - <3> QED - BY <3>2, neg_nat_not_in_nat_setminus_0 - <2> QED - BY <2>1, <2>2 DEF iPred -<1>3. @ = iPred[-.n + 0] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. -.n + 0 = -.n - BY <2>1, add_0 - <2> QED - BY <2>2 -<1>4. @ = -.n + iPred[0] - <2>1. -.n \in Int - <3>1. n \in Nat - OBVIOUS - <3> QED - BY <3>1, minus_nat_in_int - <2>2. 0 \in Int - BY zero_in_int - <2> QED - BY <2>1, <2>2, iPredRightDistributesAdd -<1>5. @ = -.n + -.Succ[-.0] - <2>1. 0 \in Int - BY zero_in_int - <2>2. 0 \notin Nat \ {0} - OBVIOUS - <2> QED - BY <2>1, <2>2 DEF iPred -<1>6. @ = -.n + -.Succ[0] - BY neg0 -<1>7. @ = -.n + -.1 - OBVIOUS \* 1 is an abbreviation -<1> QED - BY <1>1, <1>2, <1>3, <1>4, - <1>5, <1>6, <1>7 -*) -theorem MinusSuccIsMinusOne: - assumes ndom: "n \ Nat" - shows "-.Succ[n] = -.n + -.1" - proof - - have s1_1: "-.Succ[n] = -.Succ[-.-.n]" - proof - - have s2_1: "-.-.n = n" - proof - - have s3_1: "n \ Int" - using ndom nat_in_int by auto - show ?thesis - using s3_1 minus_involutive by auto - qed - show ?thesis - using s2_1 by auto - qed - also have s1_2: "... = iPred[-.n]" - proof - - have s2_1: "-.n \ Int" - using ndom minus_nat_in_int by auto - have s2_2: "-.n \ Nat \ {0}" - proof - - have s3_1: "-.n \ negNat" - using ndom minus_nat_in_neg_nat by auto - show ?thesis - using s3_1 neg_nat_not_in_nat_setminus_0 by fast - qed - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_3: "... = iPred[-.n + 0]" - proof - - have s2_1: "-.n \ Int" - proof - - have s3_1: "n \ Nat" - using ndom by auto - show ?thesis - using s3_1 minus_nat_in_int by auto - qed - have s2_2: "-.n + 0 = -.n" - using s2_1 add_0 by auto - show ?thesis - using s2_2 by auto - qed - also have s1_4: "... = -.n + iPred[0]" - proof - - have s2_1: "-.n \ Int" - proof - - have s3_1: "n \ Nat" - using ndom by auto - show ?thesis - using s3_1 minus_nat_in_int by auto - qed - have s2_2: "0 \ Int" - using zero_in_int by auto - show ?thesis - using s2_1 s2_2 iPredRightDistributesAdd by fast - qed - also have s1_5: "... = -.n + -.Succ[-.0]" - proof - - have s2_1: "0 \ Int" - using zero_in_int by auto - have s2_2: "0 \ Nat \ {0}" - by auto - show ?thesis - unfolding iPred_def - using s2_1 s2_2 by auto - qed - also have s1_6: "... = -.n + -.Succ[0]" - using neg0 by auto - also have s1_7: "... = -.n + -.1" - by auto - from calculation show ?thesis - by auto - qed - - -theorem SuccMinusIsPlusOne: - assumes idom: "i \\in negNat" - shows "Succ[-.i] = -.i + 1" - proof - - have s1_1: "-.i \\in Nat" - using idom minus_neg_nat_in_nat by auto - have s1_2: "-.i \\in Int" - using s1_1 nat_in_int by auto - have s1_3: "Succ[-.i] = iSucc[-.i]" - unfolding iSucc_def - using s1_1 s1_2 by auto - also have s1_4: "... = iSucc[-.i + 0]" - using s1_2 add_0 by auto - also have s1_5: "... = -.i + Succ[0]" - using s1_2 zeroIsNat addint_succ_nat by auto - also have s1_6: "... = -.i + 1" - by auto (* 1 is an abbreviation *) - from calculation show ?thesis . - qed - - -(*----------------------------------------------------------------------------*) -(* Properties of addition and difference. *) - - -(* -THEOREM AddNegCancels == - \A i \in Int: i + -.i = 0 -PROOF -<1>1. 0 + -.0 = 0 - <2>1. 0 + -.0 = 0 + 0 - BY neg0 - <2>2. @ = 0 - BY zero_in_int, add_0 - <2> QED - BY <2>1, <2>2 -<1>2. ASSUME NEW i \in Nat, i + -.i = 0 - PROVE Succ[i] + -.Succ[i] = 0 - <2>1. Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i] - <3>1. Succ[i] \in Int - BY <1>2, succIsNat, nat_in_int - <3>2. -.Succ[i] \in Int - BY <3>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, AddCommutative - <2>2. @ = iSucc[-.Succ[i] + i] - <3>1. -.Succ[i] \in Int - BY <1>2, succIsNat, minus_nat_in_int - <3>2. i \in Nat - BY <1>2 - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>3. @ = iSucc[i + -.Succ[i]] - <3>1. i \in Int - BY <1>2, nat_in_int - <3>2. -.Succ[i] \in Int - BY <1>2, succIsNat, minus_nat_in_int - <3>3. -.Succ[i] + i = i + -.Succ[i] - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>4. @ = iSucc[iPred[i + -.i]] - <3>1. i \in Int - BY <1>2, nat_in_int - <3>2. i \in Nat - BY <1>2 - <3> QED - BY <3>1, <3>2, addint_succ_negnat - <2>5. @ = iSucc[iPred[0]] - BY <1>2 - <2>6. @ = iSucc[-.1] - BY ipred_0 - <2>7. @ = 0 - BY isucc_minus_1 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6, <2>7 -<1>3. ASSUME NEW i \in negNat, i + -.i = 0 - PROVE -.Succ[-.i] + -.(-.Succ[-.i]) = 0 - <2>1. -.Succ[-.i] + -.(-.Succ[-.i]) - = -.Succ[-.i] + Succ[-.i] - <3>1. Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, nat_in_int - <3> QED - BY <3>1, minus_involutive - <2>2. @ = iSucc[-.Succ[-.i] + -.i] - <3>1. -.Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, minus_nat_in_int - <3>2. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, addint_succ_nat - <2>3. @ = iSucc[-.i + -.Succ[-.i]] - <3>1. -.Succ[-.i] \in Int - BY <1>3, minus_neg_nat_in_nat, - succIsNat, minus_nat_in_int - <3>2. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>3. -.Succ[-.i] + -.i = -.i + -.Succ[-.i] - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>4. @ = iSucc[iPred[-.i + -.-.i]] - <3>1. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>2. -.i \in Nat - BY <1>3, minus_neg_nat_in_nat - <3>3. -.i + -.Succ[-.i] = iPred[-.i + -.-.i] - BY <3>1, <3>2, addint_succ_negnat - <3> QED - BY <3>3 - <2>5. @ = iSucc[iPred[-.i + i]] - <3>1. -.-.i = i - BY <1>3, neg_nat_in_int, minus_involutive - <3> QED - BY <3>1 - <2>6. @ = iSucc[iPred[i + -.i]] - <3>1. -.i \in Int - BY <1>3, minus_negnat_in_int - <3>2. i \in Int - BY <1>3, neg_nat_in_int - <3>3. -.i + i = i + -.i - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>7. @ = iSucc[iPred[0]] - BY <1>3 \* induction hypothesis - <2>8. @ = iSucc[-.1] - BY ipred_0 - <2>9. @ = 0 - BY isucc_minus_1 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, - <2>6, <2>7, <2>8, <2>9 -<1>4. \A i \in Nat: i + -.i = 0 - BY <1>1, <1>2, NatInduction -<1>5. \A i \in negNat: i + -.i = 0 - BY <1>1, <1>3, neg_nat_induction -<1> QED - BY <1>4, <1>5, int_union_nat_negnat -*) -theorem AddNegCancels: - "\ i \ Int: i + -.i = 0" - proof - - have s1_1: "0 + -.0 = 0" - proof - - have s2_1: "0 + -.0 = 0 + 0" - using neg0 by auto - also have s2_2: "... = 0" - using zero_in_int add_0 by auto - finally show ?thesis . - qed - have s1_2: "\ i. \ - i \\in Nat; - i + -.i = 0 - \ \ - Succ[i] + -.Succ[i] = 0" - proof - - fix "i" :: "c" - assume s1_2_idom: "i \\in Nat" - assume s1_2_induct: "i + -.i = 0" - have s2_1: "Succ[i] + -.Succ[i] = -.Succ[i] + Succ[i]" - proof - - have s3_1: "Succ[i] \\in Int" - using s1_2_idom succIsNat nat_in_int by auto - have s3_2: "-.Succ[i] \\in Int" - using s3_1 neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 AddCommutative by fast - qed - also have s2_2: "... = iSucc[-.Succ[i] + i]" - proof - - have s3_1: "-.Succ[i] \\in Int" - using s1_2_idom succIsNat minus_nat_in_int by auto - have s3_2: "i \\in Nat" - using s1_2_idom by auto - show ?thesis - using s3_1 s3_2 by (rule addint_succ_nat) - qed - also have s2_3: "... = iSucc[i + -.Succ[i]]" - proof - - have s3_1: "i \\in Int" - using s1_2_idom nat_in_int by auto - have s3_2: "-.Succ[i] \\in Int" - using s1_2_idom succIsNat minus_nat_in_int - by auto - have s3_3: "-.Succ[i] + i = i + -.Succ[i]" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = iSucc[iPred[i + -.i]]" - proof - - have s3_1: "i \\in Int" - using s1_2_idom nat_in_int by auto - have s3_2: "i \\in Nat" - using s1_2_idom by auto - show ?thesis - using s3_1 s3_2 addint_succ_negnat by auto - qed - also have s2_5: "... = iSucc[iPred[0]]" - using s1_2_induct by auto - also have s2_6: "... = iSucc[-.1]" - using ipred_0 by auto - also have s2_7: "... = 0" - using isucc_minus_1 by auto - finally show "Succ[i] + -.Succ[i] = 0" . - qed - have s1_3: "\ i. \ - i \\in negNat; - i + -.i = 0 - \ \ - -.Succ[-.i] + -.-.Succ[-.i] = 0" - proof - - fix "i" :: "c" - assume s1_3_idom: "i \\in negNat" - assume s1_3_induct: "i + -.i = 0" - have s2_1: "-.Succ[-.i] + -.(-.Succ[-.i]) - = -.Succ[-.i] + Succ[-.i]" - proof - - have s3_1: "Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat nat_in_int by auto - show ?thesis - using s3_1 minus_involutive by auto - qed - also have s2_2: "... = iSucc[-.Succ[-.i] + -.i]" - proof - - have s3_1: "-.Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat minus_nat_in_int by auto - have s3_2: "-.i \\in Nat" - using s1_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 by (rule addint_succ_nat) - qed - also have s2_3: "... = iSucc[-.i + -.Succ[-.i]]" - proof - - have s3_1: "-.Succ[-.i] \\in Int" - using s1_3_idom minus_neg_nat_in_nat - succIsNat minus_nat_in_int by auto - have s3_2: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_3: "-.Succ[-.i] + -.i = -.i + -.Succ[-.i]" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = iSucc[iPred[-.i + -.-.i]]" - proof - - have s3_1: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_2: "-.i \\in Nat" - using s1_3_idom minus_neg_nat_in_nat by auto - have s3_3: "-.i + -.Succ[-.i] = iPred[-.i + -.-.i]" - using s3_1 s3_2 addint_succ_negnat by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = iSucc[iPred[-.i + i]]" - proof - - have s3_1: "-.-.i = i" - using s1_3_idom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s3_1 by auto - qed - also have s2_6: "... = iSucc[iPred[i + -.i]]" - proof - - have s3_1: "-.i \\in Int" - using s1_3_idom minus_negnat_in_int by auto - have s3_2: "i \\in Int" - using s1_3_idom neg_nat_in_int by auto - have s3_3: "-.i + i = i + -.i" - using s3_1 s3_2 AddCommutative by fast - show ?thesis - using s3_3 by auto - qed - also have s2_7: "... = iSucc[iPred[0]]" - using s1_3_induct by auto - also have s2_8: "... = iSucc[-.1]" - using ipred_0 by auto - also have s2_9: "... = 0" - using isucc_minus_1 by auto - finally show "-.Succ[-.i] + -.-.Succ[-.i] = 0" . - qed - have s1_4: "\ i \ Nat: i + -.i = 0" - using s1_1 s1_2 natInduct[of "\ i. - i + -.i = 0"] - by auto - have s1_5: "\ i \ negNat: i + -.i = 0" - using s1_1 s1_3 neg_nat_induction[of "\ i. - i + -.i = 0"] - by auto - show ?thesis - using s1_4 s1_5 int_union_nat_negnat by auto - qed - - -theorem AddNegCancels_sequent: - assumes idom: "i \\in Int" - shows "i + -.i = 0" - using idom AddNegCancels spec by auto - - -theorem AddNegCancels_left: - assumes idom: "i \\in Int" - shows "-.i + i = 0" - proof - - have s1_1: "i + -.i = 0" - using idom AddNegCancels_sequent by auto - have s1_2: "i + -.i = -.i + i" - proof - - have s2_1: "-.i \\in Int" - using idom neg_int_eq_int by auto - show ?thesis - using idom s2_1 AddCommutative_sequent by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM a_plus_b_minus_a_eq_b == - ASSUME NEW a \in Int, NEW b \in Int - PROVE a + (b - a) = b -PROOF -<1>1. a + (b - a) = a + (b + -.a) - BY DEF diff -<1>2. @ = a + (-.a + b) - <2>1. b \in Int - OBVIOUS - <2>2. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>3. b + -.a = -.a + b - BY <2>1, <2>2, AddCommutative - <2> QED - BY <2>3 -<1>3. @ = (a + -.a) + b - <2>1. a \in Int - OBVIOUS - <2>2. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>3. b \in Int - OBVIOUS - <2> QED - BY <2>1, <2>2, <2>3, AddAssociative -<1>4. @ = 0 + b - <2>1. a \in Int - OBVIOUS - <2>2. a + -.a = 0 - BY <2>1, AddNegCancels - <2> QED - BY <2>2 -<1>5. @ = b - BY add_0_left -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5 -*) -theorem a_plus_b_minus_a_eq_b: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "a + (b - a) = b" - proof - - have s1_1: "a + (b - a) = a + (b + -.a)" - unfolding diff_def by auto - also have s1_2: "... = a + (-.a + b)" - proof - - have s2_1: "b \\in Int" - using bdom by auto - have s2_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_3: "b + -.a = -.a + b" - using s2_1 s2_2 AddCommutative by auto - show ?thesis - using s2_3 by auto - qed - also have s1_3: "... = (a + -.a) + b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_3: "b \\in Int" - using bdom by auto - have s2_4: "a \\in Int ==> -.a \\in Int ==> b \\in Int - ==> (a + -.a) + b = a + (-.a + b)" - using AddAssociative_sequent by blast - have s2_5: "(a + -.a) + b = a + (-.a + b)" - using s2_1 s2_2 s2_3 s2_4 by auto - show ?thesis - by (rule s2_5[symmetric]) - qed - also have s1_4: "... = 0 + b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "a + -.a = 0" - using s2_1 AddNegCancels by auto - show ?thesis - using s2_2 by auto - qed - also have s1_5: "... = b" - using bdom add_0_left by auto - finally show ?thesis . - qed - - -(* -THEOREM MinusDistributesAdd == - \A a \in Int: \A b \in Int: - -.(a + b) = (-.a) + (-.b) -PROOF -<1>1. SUFFICES ASSUME NEW a \in Int - PROVE \A b \in Int: - -.(a + b) = (-.a) + (-.b) - BY <1>1 -<1>2. -.(a + 0) = (-.a) + (-.0) - <2>1. -.(a + 0) = -.a - <3>1. a + 0 = a - BY <1>1, add_0 - <3> QED - BY <3>1 - <2>2. -.a + -.0 = -.a - <3>1. -.a + -.0 = -.a + 0 - BY neg0 - <3>2. @ = -.a - <4>1. -.a \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, add_0 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW b \in Nat, - -.(a + b) = (-.a) + (-.b) - PROVE - -.(a + Succ[b]) = (-.a) + (-.Succ[b]) - <2>1. -.(a + Succ[b]) = -.iSucc[a + b] - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3>3. a + Succ[b] = iSucc[a + b] - BY <3>1, <3>2, addint_succ_nat - <3> QED - BY <3>3 - <2>2. @ = -. iSucc[-.((-.a) + (-.b))] - <3>1. a + b = -.(-.a + -.b) - <4>1. -.-.(a + b) = -.(-.a + -.b) - BY <1>3 - <4>2. -.-.(a + b) = a + b - <5>1. a + b \in Int - <6>1. a \in Int - BY <1>1 - <6>2. b \in Int - BY <1>3, nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, minus_involutive - <4> QED - BY <4>1, <4>2 - <3> QED - BY <3>1 - <2>3. @ = iPred[-.a + -.b] - <3>1. -.a + -.b \in Int - <4>1. -.a \in Int - BY <1>1, neg_int_eq_int - <4>2. -.b \in Int - BY <1>3, minus_nat_in_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, iSuccMinusFlip - <2>4. @ = -.a + iPred[-.b] - <3>1. -.a \in Int - BY <1>1, neg_int_eq_int - <3>2. -.b \in Int - BY <1>3, minus_nat_in_int - <3> QED - BY <3>1, <3>2, iPredRightDistributesAdd - <2>5. @ = -.a + -.Succ[b] - <3>1. iPred[-.b] = -.Succ[-.-.b] - <4>1. -.b \in Int - BY <1>3, minus_nat_in_int - <4>2. -.b \notin Nat \ {0} - <5>1. -.b \in negNat - BY <1>3, minus_nat_in_neg_nat - <5> QED - BY <5>1, neg_nat_not_in_nat_setminus_0 - <4> QED - BY <4>1, <4>2 DEF iPred - <3>2. @ = -.Succ[b] - <4>1. -.-.b = b - BY <1>3, nat_in_int, minus_involutive - <4> QED - BY <4>1 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5 -<1>4. ASSUME NEW b \in negNat, - -.(a + b) = (-.a) + (-.b) - PROVE - -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b]) - <2>1. -.(a + -.Succ[-.b]) = -. iPred[a + b] - <3>1. a + -.Succ[-.b] = iPred[a + -.-.b] - <4>1. a \in Int - BY <1>1 - <4>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, addint_succ_negnat - <3>2. a + -.Succ[-.b] = iPred[a + b] - <4>1. -.-.b = b - BY <1>4, neg_nat_in_int, minus_involutive - <4> QED - BY <3>1, <4>1 - <3> QED - BY <3>2 - <2>2. @ = -. iPred[-.((-.a) + (-.b))] - <3>1. a + b = -.(-.a + -.b) - <4>1. -.(a + b) = (-.a) + (-.b) - BY <1>4 - <4>2. -.-.(a + b) = -.(-.a + -.b) - BY <4>1 - <4>3. -.-.(a + b) = a + b - <5>1. a + b \in Int - <6>1. a \in Int - BY <1>1 - <6>2. b \in Int - BY <1>4, neg_nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, minus_involutive - <4> QED - BY <4>2, <4>3 - <3> QED - BY <3>1 - <2>3. @ = iSucc[-.a + -.b] - <3>1. -.a + -.b \in Int - <4>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <4>2. -.b \in Int - BY <1>4, minus_negnat_in_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, iPredMinusFlip - <2>4. @ = -.a + iSucc[-.b] - <3>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <3>2. -.b \in Int - BY <1>4, minus_negnat_in_int - <3> QED - BY <3>1, <3>2, iSuccRightDistributesAdd - <2>5. @ = -.a + Succ[-.b] - <3>1. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>2. -.b \in Int - BY <3>1, nat_in_int - <3> QED - BY <3>1, <3>2 DEF iSucc - <2>6. @ = -.a + (-.-.Succ[-.b]) - <3>1. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>2. Succ[-.b] \in Nat - BY <3>1, succIsNat - <3>3. Succ[-.b] \in Int - BY <3>2, nat_in_int - <3>4. -.-.Succ[-.b] = Succ[-.b] - BY <3>3, minus_involutive - <3> QED - BY <3>4 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, <2>6 -<1>5. \A b \in Nat: -.(a + b) = (-.a) + (-.b) - BY <1>2, <1>3, NatInduction -<1>6. \A b \in negNat: -.(a + b) = (-.a) + (-.b) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MinusDistributesAdd: - "\ a \ Int: \ b \ Int: - -.(a + b) = (-.a) + (-.b)" - proof - - { - fix "a" :: "c" - assume adom: "a \\in Int" - have s1_2: "-.(a + 0) = (-.a) + (-.0)" - proof - - have s2_1: "-.(a + 0) = -.a" - proof - - have s3_1: "a + 0 = a" - using adom add_0 by auto - show ?thesis - using s3_1 by auto - qed - have s2_2: "-.a + -.0 = -.a" - proof - - have s3_1: "-.a + -.0 = -.a + 0" - using neg0 by auto - also have s3_2: "... = -.a" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - show ?thesis - using s4_1 add_0 by auto - qed - finally show ?thesis . - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ b. \ - b \\in Nat; - -.(a + b) = -.a + -.b - \ \ - -.(a + Succ[b]) = (-.a) + (-.Succ[b])" - proof - - fix "b" :: "c" - assume s1_3_bdom: "b \\in Nat" - assume s1_3_induct: "-.(a + b) = -.a + -.b" - have s2_1: "-.(a + Succ[b]) = -.iSucc[a + b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - have s3_3: "a + Succ[b] = iSucc[a + b]" - using s3_1 s3_2 addint_succ_nat by auto - show ?thesis - using s3_3 by auto - qed - also have s2_2: "... = -. iSucc[-.((-.a) + (-.b))]" - proof - - have s3_1: "a + b = -.(-.a + -.b)" - proof - - have s4_1: "-.-.(a + b) = -.(-.a + -.b)" - using s1_3_induct by auto - have s4_2: "-.-.(a + b) = a + b" - proof - - have s5_1: "a + b \\in Int" - proof - - have s6_1: "a \\in Int" - using adom by auto - have s6_2: "b \\in Int" - using s1_3_bdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 minus_involutive by auto - qed - show ?thesis - using s4_1 s4_2 by auto - qed - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = iPred[-.a + -.b]" - proof - - have s3_1: "-.a + -.b \\in Int" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s4_2: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 iSuccMinusFlip by auto - qed - also have s2_4: "... = -.a + iPred[-.b]" - proof - - have s3_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s3_2: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - show ?thesis - using s3_1 s3_2 iPredRightDistributesAdd by auto - qed - also have s2_5: "... = -.a + -.Succ[b]" - proof - - have s3_1: "iPred[-.b] = -.Succ[-.-.b]" - proof - - have s4_1: "-.b \\in Int" - using s1_3_bdom minus_nat_in_int by auto - have s4_2: "-.b \ Nat \ {0}" - proof - - have s5_1: "-.b \\in negNat" - using s1_3_bdom minus_nat_in_neg_nat by auto - show ?thesis - using s5_1 neg_nat_not_in_nat_setminus_0 by fast - qed - show ?thesis - unfolding iPred_def - using s4_1 s4_2 by auto - qed - also have s3_2: "... = -.Succ[b]" - proof - - have s4_1: "-.-.b = b" - using s1_3_bdom nat_in_int minus_involutive by auto - show ?thesis - using s4_1 by auto - qed - finally have s3_3: "iPred[-.b] = -.Succ[b]" . - show ?thesis - using s3_3 by auto - qed - finally show "-.(a + Succ[b]) = (-.a) + (-.Succ[b])" . - qed - have s1_4: "\ b. \ - b \\in negNat; - -.(a + b) = -.a + -.b - \ \ - -.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" - proof - - fix "b" :: "c" - assume s1_4_bdom: "b \\in negNat" - assume s1_4_induct: "-.(a + b) = -.a + -.b" - have s2_1: "-.(a + -.Succ[-.b]) = -. iPred[a + b]" - proof - - have s3_1: "a + -.Succ[-.b] = iPred[a + -.-.b]" - proof - - have s4_1: "a \\in Int" - using adom by auto - have s4_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 addint_succ_negnat by auto - qed - have s3_2: "a + -.Succ[-.b] = iPred[a + b]" - proof - - have s4_1: "-.-.b = b" - using s1_4_bdom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s3_1 s4_1 by auto - qed - show ?thesis - using s3_2 by auto - qed - also have s2_2: "... = -. iPred[-.((-.a) + (-.b))]" - proof - - have s3_1: "a + b = -.(-.a + -.b)" - proof - - have s4_1: "-.(a + b) = (-.a) + (-.b)" - using s1_4_induct by auto - have s4_2: "-.-.(a + b) = -.(-.a + -.b)" - using s4_1 by auto - have s4_3: "-.-.(a + b) = a + b" - proof - - have s5_1: "a + b \\in Int" - proof - - have s6_1: "a \\in Int" - using adom by auto - have s6_2: "b \\in Int" - using s1_4_bdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 minus_involutive by auto - qed - show ?thesis - using s4_2 s4_3 by auto - qed - show ?thesis - using s3_1 by auto - qed - also have s2_3: "... = iSucc[-.a + -.b]" - proof - - have s3_1: "-.a + -.b \\in Int" - proof - - have s4_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s4_2: "-.b \\in Int" - using s1_4_bdom minus_negnat_in_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 iPredMinusFlip by auto - qed - also have s2_4: "... = -.a + iSucc[-.b]" - proof - - have s3_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s3_2: "-.b \\in Int" - using s1_4_bdom minus_negnat_in_int by auto - show ?thesis - using s3_1 s3_2 iSuccRightDistributesAdd by auto - qed - also have s2_5: "... = -.a + Succ[-.b]" - proof - - have s3_1: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - have s3_2: "-.b \\in Int" - using s3_1 nat_in_int by auto - show ?thesis - unfolding iSucc_def - using s3_1 s3_2 by auto - qed - also have s2_6: "... = -.a + (-.-.Succ[-.b])" - proof - - have s3_1: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - have s3_2: "Succ[-.b] \\in Nat" - using s3_1 succIsNat by auto - have s3_3: "Succ[-.b] \\in Int" - using s3_2 nat_in_int by auto - have s3_4: "-.-.Succ[-.b] = Succ[-.b]" - using s3_3 minus_involutive by auto - show ?thesis - using s3_4 by auto - qed - finally show "-.(a + -.Succ[-.b]) = (-.a) + (-.-.Succ[-.b])" . - qed - have s1_5: "\ b \ Nat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_3 natInduct[of "\ b. - -.(a + b) = (-.a) + (-.b)"] - by auto - have s1_6: "\ b \ negNat: -.(a + b) = (-.a) + (-.b)" - using s1_2 s1_4 neg_nat_induction[of "\ b. - -.(a + b) = (-.a) + (-.b)"] - by auto - have s1_7: "\ b \ Int: -.(a + b) = (-.a) + (-.b)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem MinusDistributesAdd_sequent: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "-.(a + b) = (-.a) + (-.b)" - using adom bdom MinusDistributesAdd by auto - - -(* -THEOREM MinusDiff == - ASSUME NEW a \in Int, NEW b \in Int - PROVE -.(a - b) = b - a -PROOF -<1>1. -.(a - b) = -.(a + -.b) - BY DEF diff -<1>2. @ = (-.a) + (-.-.b) - <2>1. a \in Int - OBVIOUS - <2>2. -.b \in Int - BY b \in Int, neg_int_eq_int - <2> QED - BY <2>1, <2>2, MinusDistributesAdd -<1>3. @ = (-.a) + b - BY b \in Int, minus_involutive -<1>4. @ = b + (-.a) - <2>1. -.a \in Int - BY a \in Int, neg_int_eq_int - <2>2. b \in Int - OBVIOUS - <2> QED - BY <2>1, <2>2, AddCommutative -<1>5. @ = b - a - BY DEF diff -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5 -*) -theorem MinusDiff: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "-.(a - b) = b - a" - proof - - have s1_1: "-.(a - b) = -.(a + -.b)" - unfolding diff_def by auto - also have s1_2: "... = -.a + -.-.b" - proof - - have s2_1: "a \\in Int" - using adom by auto - have s2_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - show ?thesis - using s2_1 s2_2 MinusDistributesAdd by auto - qed - also have s1_3: "... = -.a + b" - using bdom minus_involutive by auto - also have s1_4: "... = b + -.a" - proof - - have s2_1: "-.a \\in Int" - using adom neg_int_eq_int by auto - have s2_2: "b \\in Int" - using bdom by auto - show ?thesis - using s2_1 s2_2 AddCommutative by auto - qed - also have s1_5: "... = b - a" - unfolding diff_def by auto - finally show ?thesis . - qed - - -(*----------------------------------------------------------------------------*) -(* Commutativity of multiplication. *) - -(* -THEOREM MultCommutative == - \A i, j \in Int: i * j = j * i -PROOF -<1>1. \A i \in Int: i * 0 = 0 * i - <2>1. 0 * 0 = 0 * 0 - OBVIOUS - <2>2. ASSUME NEW i \in Nat, i * 0 = 0 * i - PROVE Succ[i] * 0 = 0 * Succ[i] - <3>1. 0 * Succ[i] = 0 * i + 0 - <4>1. 0 \in Int - BY zero_in_int - <4>2. i \in Nat - BY <2>2 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3>2. @ = 0 * i - <4>1. 0 * i \in Int - <5>1 0 \in Int - BY zero_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5> QED - BY <5>1, <5>2, multIsInt - <4> QED - BY <4>1, add_0 - <3>3. @ = i * 0 - BY <2>2 - <3>4. @ = 0 - <4>1. i \in Int - BY <2>2, nat_in_int - <4> QED - BY <4>1, mult_0 - <3>5. @ = Succ[i] * 0 - <4>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5 - <2>3. ASSUME NEW i \in negNat, i * 0 = 0 * i - PROVE -.Succ[-.i] * 0 = 0 * -.Succ[-.i] - <3>1. -.Succ[-.i] * 0 = 0 - <4>1. -.Succ[-.i] \in Int - BY <2>3, minus_succ_minus_negnat_in_int - <4> QED - BY <4>1, mult_0 - <3>2. 0 * -.Succ[-.i] = 0 - <4>1. 0 * -.Succ[-.i] = 0 * -.-.i + -.0 - <5>1. 0 \in Int - BY zero_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = 0 * i + -.0 - <5>1. -.-.i = i - BY <2>3, neg_nat_in_int, minus_involutive - <5> QED - BY <5>1 - <4>3. @ = 0 * i + 0 - BY neg0 - <4>4. @ = 0 * i - <5>1. 0 * i \in Int - <6>1. 0 \in Int - BY zero_in_int - <6>2. i \in Int - BY <2>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5> QED - BY <5>1, add_0 - <4>5. @ = i * 0 - BY <2>3 \* induction hypothesis - <4>6. @ = 0 - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5> QED - BY <5>1, mult_0 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, - <4>5, <4>6 - <3> QED - BY <3>1, <3>2 - <2>4. \A i \in Nat: i * 0 = 0 * i - BY <2>1, <2>2, NatInduction - <2>5. \A i \in negNat: i * 0 = 0 * i - BY <2>1, <2>3, neg_nat_induction - <2> QED - BY <2>4, <2>5, int_union_nat_negnat -<1>2. ASSUME NEW j \in Nat, - \A i \in Int: i * j = j * i - PROVE \A i \in Int: i * Succ[j] = Succ[j] * i - <2>1. 0 * Succ[j] = Succ[j] * 0 - <3>1. 0 * Succ[j] = 0 - <4>1. 0 * Succ[j] = 0 * j + 0 - <5>1. 0 \in Int - BY zero_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = 0 * j - <5>1. 0 * j \in Int - <6>1. 0 \in Int - BY zero_in_int - <6>2. j \in Int - BY <1>2, nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5> QED - BY <5>1, add_0 - <4>3. @ = j * 0 - BY <1>2, zero_in_int, spec - <4>4. @ = 0 - <5>1. j \in Int - BY <1>2, nat_in_int - <5> QED - BY <5>1, mult_0 - <4> QED - BY <4>1, <4>2, <4>3, <4>4 - <3>2. Succ[j] * 0 = 0 - <4>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2 - <2>2. ASSUME NEW i \in Nat, - i * Succ[j] = Succ[j] * i - PROVE Succ[i] * Succ[j] = Succ[j] * Succ[i] - <3>1. Succ[i] * Succ[j] = i * j + ((i + j) + 1) - <4>1. Succ[i] * Succ[j] = Succ[i] * j + Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = j * Succ[i] + Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. Succ[i] * j = j * Succ[i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>3. @ = (j * i + j) + Succ[i] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5>3. j * Succ[i] = j * i + j - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>4. @ = (i * j + j) + Succ[i] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>5. @ = (i * j + j) + (i + 1) - <5>1. i \in Nat - BY <2>2 - <5>2. Succ[i] = i + 1 - BY <5>1, nat_add_1 - <5> QED - BY <5>2 - <4>6. @ = i * j + (j + (i + 1)) - <5>1. i * j \in Int - BY <1>2, <2>2, nat_in_int, multIsInt - <5>2. j \in Int - BY <1>2, nat_in_int - <5>3. i + 1 \in Int - BY <2>2, oneIsNat, nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + ((j + i) + 1) - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. j + (i + 1) = (j + i) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>8. @ = i * j + ((i + j) + 1) - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. j + i = i + j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8 - <3>2. Succ[j] * Succ[i] = i * j + ((i + j) + 1) - <4>1. Succ[j] * Succ[i] = Succ[j] * i + Succ[j] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = i * Succ[j] + Succ[j] - BY <2>2 \* induction hypothesis - <4>3. @ = (i * j + i) + Succ[j] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Nat - BY <1>2 - <5>3. i * Succ[j] = i * j + i - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>4. @ = (i * j + i) + (j + 1) - <5>1. j \in Nat - BY <1>2 - <5>2. Succ[j] = j + 1 - BY <5>1, nat_add_1 - <5> QED - BY <5>2 - <4>5. @ = i * j + (i + (j + 1)) - <5>1. i * j \in Int - BY <1>2, <2>2, nat_in_int, multIsInt - <5>2. i \in Int - BY <2>2, nat_in_int - <5>3. j + 1 \in Int - BY <1>2, oneIsNat, nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>6. @ = i * j + ((i + j) + 1) - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. j \in Int - BY <1>2, nat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. i + (j + 1) = (i + j) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, <4>6 - <3> QED - BY <3>1, <3>2 - <2>3. ASSUME NEW i \in negNat, - i * Succ[j] = Succ[j] * i - PROVE -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i] - <3>1. -.Succ[-.i] * Succ[j] = - i * j + ((i + -.j) + -.1) - <4>1. -.Succ[-.i] * Succ[j] = - -.Succ[-.i] * j + -.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. j \in Nat - BY <1>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = j * -.Succ[-.i] + -.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>3. @ = (j * -.-.i + -.j) + -.Succ[-.i] - <5>1. j \in Int - BY <1>2, nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>3. j * -.Succ[-.i] = j * -.-.i + -.j - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>4. @ = (j * i + -.j) + -.Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>5. @ = (i * j + -.j) + -.Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>2 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>6. @ = (i * j + -.j) + (i + -.1) - <5>1. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>2. -.Succ[-.i] = -.-.i + -.1 - BY <5>1, MinusSuccIsMinusOne - <5>3. -.-.i = i - BY <2>3, neg_nat_in_int, minus_involutive - <5>4. -.Succ[-.i] = i + -.1 - BY <5>2, <5>3 - <5> QED - BY <5>4 - <4>7. @ = i * j + (-.j + (i + -.1)) - <5>1. i * j \in Int - BY <1>2, nat_in_int, <2>3, neg_nat_in_int, - multIsInt - <5>2. -.j \in Int - BY <1>2, minus_nat_in_int - <5>3. i + -.1 \in Int - BY <2>3, minus_one_in_negnat, - neg_nat_in_int, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + ((-.j + i) + -.1) - <5>1. -.j \in Int - BY <1>2, minus_nat_in_int - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <5>4. -.j + (i + -.1) = (-.j + i) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>9. @ = i * j + ((i + -.j) + -.1) - <5>1. -.j \in Int - BY <1>2, minus_nat_in_int - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.j + i = i + -.j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, - <4>5, <4>6, <4>7, <4>8, <4>9 - <3>2. Succ[j] * -.Succ[-.i] = - i * j + ((i + -.j) + -.1) - <4>1. Succ[j] * -.Succ[-.i] = - Succ[j] * -.-.i + -.Succ[j] - <5>1. Succ[j] \in Int - BY <1>2, succIsNat, nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = Succ[j] * i + -.Succ[j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = i * Succ[j] + -.Succ[j] - BY <2>3 \* induction hypothesis - <4>4. @ = (i * j + i) + -.Succ[j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. j \in Nat - BY <1>2 - <5>3. i * Succ[j] = i * j + i - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>5. @ = (i * j + i) + (-.j + -.1) - <5>1. j \in Nat - BY <1>2 - <5>2. -.Succ[j] = -.j + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>6. @ = i * j + (i + (-.j + -.1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>2, nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. i \in Int - BY <2>3, neg_nat_in_int - <5>3. -.j + -.1 \in Int - <6>1. -.j \in Int - BY <1>2, minus_nat_in_int - <6>2. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + ((i + -.j) + -.1) - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.j \in Int - BY <1>2, minus_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_negnat, neg_nat_in_int - <5>4. i + (-.j + -.1) = (I + -.j) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, NatInduction, neg_nat_induction, - int_union_nat_negnat -<1>3. ASSUME NEW j \in negNat, - \A i \in Int: i * j = j * i - PROVE \A i \in Int: i * -.Succ[-.j] = -.Succ[-.j] * i - <2>1. 0 * -.Succ[-.j] = -.Succ[-.j] * 0 - <3>1. 0 * -.Succ[-.j] = 0 * -.-.j + -.0 - <4>1. 0 \in Int - BY zeroIsNat nat_in_int - <4>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3>2. @ = 0 * -.-.j - <4>1. 0 * -.-.j \in Int - <5>1. 0 \in Int - BY zeroIsNat nat_in_int - <5>2. -.-.j \in Int - BY <1>3, neg_nat_in_int, minus_involutive - <5> QED - BY <5>1, <5>2, multIsInt - <4>2. 0 * -.-.j + 0 = 0 * -.-.j - BY <4>1, add_0 - <4> QED - BY <4>2, neg0 - <3>3. @ = 0 * j - <4>1. j \in Int - BY <1>3, neg_nat_in_int - <4> QED - BY <4>1, minus_involutive - <3>4. @ = j * 0 - <4>1. 0 \in Int - BY zeroIsNat, nat_in_int - <4>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <4> QED - BY <4>1, <4>2, spec - <3>5. @ = 0 - <4>1. j \in Int - BY <1>3, neg_nat_in_int - <4> QED - BY <4>1, mult_0 - <3>6. @ = -.Succ[-.j] * 0 - <4>1. -.Succ[-.j] \in Int - BY <1>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <4> QED - BY <4>1, mult_0 - <3> QED - BY <3>1, <3>2, <3>3, <3>4, <3>5, <3>6 - <2>2. ASSUME NEW i \in Nat, - i * -.Succ[-.j] = -.Succ[-.j] * i - PROVE Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i] - <3>1. Succ[i] * -.Succ[-.j] = i * j + ((-.i + j) + -.1) - <4>1. Succ[i] * -.Succ[-.j] = - Succ[i] * -.-.j + -.Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = Succ[i] * j + -.Succ[i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = j * Succ[i] + -.Succ[i] - <5>1. Succ[i] \in Int - BY <2>2, succIsNat, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. Succ[i] * j = j * Succ[i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>4. @ = (j * i + j) + -.Succ[i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. i \in Nat - BY <2>2 - <5>3. j * Succ[i] = j * i + j - BY <5>1, <5>2, multint_succ_nat - <5> QED - BY <5>3 - <4>5. @ = (i * j + j) + -.Succ[i] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>6. @ = i * j + (j + -.Succ[i]) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>2, nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. j \in Int - BY <1>3, neg_nat_in_int - <5>3. -.Succ[i] \in Int - BY <2>2, succIsNat, minus_nat_in_int - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>7. @ = i * j + (j + (-.i + -.1)) - <5>1. i \in Nat - BY <2>2 - <5>2. -.Succ[i] = -.i + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>8. @ = i * j + ((j + -.i) + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_int - <5>4. j + (-.i + -.1) = (j + -.i) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>9. @ = i * j + ((-.i + j) + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. j + -.i = -.i + j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9 - <3>2. -.Succ[-.j] * Succ[i] = i * j + ((-.i + j) + -.1) - <4>1. -.Succ[-.j] * Succ[i] = - -.Succ[-.j] * i + -.Succ[-.j] - <5>1. -.Succ[-.j] \in Int - BY <1>3, minus_negnat_succ_in_negnat, - neg_nat_in_int - <5>2. i \in Nat - BY <2>2 - <5> QED - BY <5>1, <5>2, multint_succ_nat - <4>2. @ = i * -.Succ[-.j] + -.Succ[-.j] - BY <2>2 \* induction hypothesis - <4>3. @ = (i * -.-.j + -.i) + -.Succ[-.j] - <5>1. i \in Int - BY <2>2, nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>3. i * -.Succ[-.j] = i * -.-.j + -.i - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>4. @ = (i * j + -.i) + -.Succ[-.j] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>5. @ = (i * j + -.i) + (-.-.j + -.1) - <5>1. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>2. -.Succ[-.j] = -.-.j + -.1 - BY <5>1, MinusSuccIsMinusOne - <5> QED - BY <5>2 - <4>6. @ = (i * j + -.i) + (j + -.1) - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = i * j + (-.i + (j + -.1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>2, nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.i \in Int - BY <2>2, minus_nat_in_int - <5>3. j + -.1 \in Int - <6>1. j \in Int - BY <1>3, neg_nat_in_int - <6>2. -.1 \in Int - BY minus_one_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + ((-.i + j) + -.1) - <5>1. -.i \in Int - BY <2>2, minus_nat_in_int - <5>2. j \in Int - BY <1>3, neg_nat_in_int - <5>3. -.1 \in Int - BY minus_one_in_int - <5>4. -.i + (j + -.1) = (-.i + j) + -.1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8 - <3> QED - BY <3>1, <3>2 - <2>3. ASSUME NEW i \in negNat, - i * -.Succ[-.j] = -.Succ[-.j] * i - PROVE -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i] - <3>1. -.Succ[-.i] * -.Succ[-.j] = - i * j + ((-.i + -.j) + 1) - <4>1. -.Succ[-.i] * -.Succ[-.j] = - -.Succ[-.i] * -.-.j + -.-.Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, - neg_nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = -.Succ[-.i] * j + -.-.Succ[-.i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = -.Succ[-.i] * j + Succ[-.i] - <5>1. Succ[-.i] \in Int - BY <2>3, negnat_succ_in_nat, nat_in_int - <5>2. -.-.Succ[-.i] = Succ[-.i] - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>4. @ = j * -.Succ[-.i] + Succ[-.i] - <5>1. -.Succ[-.i] \in Int - BY <2>3, minus_negnat_succ_in_negnat, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. -.Succ[-.i] * j = j * -.Succ[-.i] - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>5. @ = (j * -.-.i + -.j) + Succ[-.i] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5>3. j * -.Succ[-.i] = j * -.-.i + -.j - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>6. @ = (j * i + -.j) + Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = (i * j + -.j) + Succ[-.i] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. \A k \in Int: k * j = j * k - BY <1>3 \* induction hypothesis - <5>3. i * j = j * i - BY <5>1, <5>2, spec - <5> QED - BY <5>3 - <4>8. @ = (i * j + -.j) + (-.i + 1) - <5>1. i \in negNat - BY <2>3 - <5>2. Succ[-.i] = -.i + 1 - BY <5>1, SuccMinusIsPlusOne - <5> QED - BY <5>2 - <4>9. @ = i * j + (-.j + (-.i + 1)) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>3. -.i + 1 \in Int - <6>1. -.i \in Int - BY <2>3, minus_negnat_in_int - <6>2. 1 \in Int - BY oneIsNat, nat_in_int - <6> QED - BY <6>1, <6>2, addIsInt - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>10. @ = i * j + ((-.j + -.i) + 1) - <5>1. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. -.j + (-.i + 1) = (-.j + -.i) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4>11. @ = i * j + ((-.i + -.j) + 1) - <5>1. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. -.j + -.i = -.i + -.j - BY <5>1, <5>2, AddCommutative - <5> QED - BY <5>3 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9, - <4>10, <4>11 - <3>2. -.Succ[-.j] * -.Succ[-.i] = - i * j + ((-.i + -.j) + 1) - <4>1. -.Succ[-.j] * -.Succ[-.i] = - -.Succ[-.j] * -.-.i + -.-.Succ[-.j] - <5>1. -.Succ[-.j] \in Int - BY <1>3, minus_succ_minus_negnat_in_int - <5>2. -.i \in Nat - BY <2>3, minus_neg_nat_in_nat - <5> QED - BY <5>1, <5>2, multint_succ_negnat - <4>2. @ = -.Succ[-.j] * i + -.-.Succ[-.j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.-.i = i - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>3. @ = -.Succ[-.j] * i + Succ[-.j] - <5>1. Succ[-.j] \in Int - BY <1>3, negnat_succ_in_nat, nat_in_int - <5>2. -.-.Succ[-.j] = Succ[-.j] - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>4. @ = i * -.Succ[-.j] + Succ[-.j] - BY <2>3 \* induction hypothesis - <4>5. @ = (i * -.-.j + -.i) + Succ[-.j] - <5>1. i \in Int - BY <2>3, neg_nat_in_int - <5>2. -.j \in Nat - BY <1>3, minus_neg_nat_in_nat - <5>3. i * -.Succ[-.j] = i * -.-.j + -.i - BY <5>1, <5>2, multint_succ_negnat - <5> QED - BY <5>3 - <4>6. @ = (i * j + -.i) + Succ[-.j] - <5>1. j \in Int - BY <1>3, neg_nat_in_int - <5>2. -.-.j = j - BY <5>1, minus_involutive - <5> QED - BY <5>2 - <4>7. @ = i * j + (-.i + Succ[-.j]) - <5>1. i * j \in Int - <6>1. i \in Int - BY <2>3, neg_nat_in_int - <6>2. j \in Int - BY <1>3, neg_nat_in_int - <6> QED - BY <6>1, <6>2, multIsInt - <5>2. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>3. Succ[-.j] \in Int - BY <1>3, negnat_succ_in_nat, nat_in_int - <5> QED - BY <5>1, <5>2, <5>3, AddAssociative - <4>8. @ = i * j + (-.i + (-.j + 1)) - <5>1. Succ[-.j] = -.j + 1 - BY <1>3, SuccMinusIsPlusOne - <5> QED - BY <5>1 - <4>9. @ = i * j + ((-.i + -.j) + 1) - <5>1. -.i \in Int - BY <2>3, minus_negnat_in_int - <5>2. -.j \in Int - BY <1>3, minus_negnat_in_int - <5>3. 1 \in Int - BY oneIsNat, nat_in_int - <5>4. -.i + (-.j + 1) = (-.i + -.j) + 1 - BY <5>1, <5>2, <5>3, AddAssociative - <5> QED - BY <5>4 - <4> QED - BY <4>1, <4>2, <4>3, <4>4, <4>5, - <4>6, <4>7, <4>8, <4>9 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, NatInduction, neg_nat_induction, - int_union_nat_negnat -<1> QED - BY <1>1, <1>2, <1>3, NatInduction, neg_nat_induction, - int_union_nat_negnat -*) -theorem MultCommutative: - "\ i \ Int: \ j \ Int: - i * j = j * i" - proof - - have s1_1: "\ i \ Int: i * 0 = 0 * i" - proof - - have s2_1: "0 * 0 = 0 * 0" - by auto - have s2_2: "\ i. \ - i \\in Nat; - i * 0 = 0 * i - \ \ - Succ[i] * 0 = 0 * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * 0 = 0 * i" - have s3_1: "0 * Succ[i] = 0 * i + 0" - proof - - have s4_1: "0 \\in Int" - using zero_in_int by auto - have s4_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - also have s3_2: "... = 0 * i" - proof - - have s4_1: "0 * i \\in Int" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - show ?thesis - using s4_1 add_0 by auto - qed - also have s3_3: "... = i * 0" - using s2_2_induct by auto - also have s3_4: "... = 0" - proof - - have s4_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - also have s3_5: "... = Succ[i] * 0" - proof - - have s4_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - finally show "Succ[i] * 0 = 0 * Succ[i]" by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * 0 = 0 * i - \ \ - -.Succ[-.i] * 0 = 0 * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * 0 = 0 * i" - have s3_1: "-.Succ[-.i] * 0 = 0" - proof - - have s4_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_succ_minus_negnat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - have s3_2: "0 * -.Succ[-.i] = 0" - proof - - have s4_1: "0 * -.Succ[-.i] = 0 * -.-.i + -.0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = 0 * i + -.0" - proof - - have s5_1: "-.-.i = i" - using s2_3_idom neg_nat_in_int - minus_involutive by auto - show ?thesis - using s5_1 by auto - qed - also have s4_3: "... = 0 * i + 0" - using neg0 by auto - also have s4_4: "... = 0 * i" - proof - - have s5_1: "0 * i \\in Int" - proof - - have s6_1: "0 \\in Int" - using zero_in_int by auto - have s6_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - show ?thesis - using s5_1 add_0 by auto - qed - also have s4_5: "... = i * 0" - using s2_3_induct by auto - also have s4_6: "... = 0" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - show ?thesis - using s5_1 mult_0 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * 0 = 0 * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: i * 0 = 0 * i" - using s2_1 s2_2 natInduct[of "\ i. - i * 0 = 0 * i"] - by blast - have s2_5: "\ i \ negNat: i * 0 = 0 * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * 0 = 0 * i"] - by blast - show ?thesis - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_2: "\ j. \ - j \\in Nat; - \ i \ Int: i * j = j * i - \ \ - \ i \ Int: i * Succ[j] = Succ[j] * i" - proof - - fix "j" :: "c" - assume s1_2_jdom: "j \\in Nat" - assume s1_2_induct: "\ i \ Int: i * j = j * i" - have s2_1: "0 * Succ[j] = Succ[j] * 0" - proof - - have s3_1: "0 * Succ[j] = 0" - proof - - have s4_1: "0 * Succ[j] = 0 * j + 0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = 0 * j" - proof - - have s5_1: "0 * j \\in Int" - proof - - have s6_1: "0 \\in Int" - using zero_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - show ?thesis - using s5_1 add_0 by auto - qed - also have s4_3: "... = j * 0" - proof - - have s5_1: "0 \\in Int" - using zero_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - show ?thesis - using s5_1 s5_2 spec by fast - qed - also have s4_4: "... = 0" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s5_1 mult_0 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * 0 = 0" - proof - - have s4_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - have s2_2: "\ i. \ - i \\in Nat; - i * Succ[j] = Succ[j] * i - \ \ - Succ[i] * Succ[j] = Succ[j] * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * Succ[j] = Succ[j] * i" - have s3_1: "Succ[i] * Succ[j] = i * j + ((i + j) + 1)" - proof - - have s4_1: "Succ[i] * Succ[j] = Succ[i] * j + Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = j * Succ[i] + Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_2_induct by auto - have s5_3: "Succ[i] * j = j * Succ[i]" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_3: "... = (j * i + j) + Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - have s5_3: "j * Succ[i] = j * i + j" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + j) + Succ[i]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_2_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + j) + (i + 1)" - proof - - have s5_1: "i \\in Nat" - using s2_2_idom by auto - have s5_2: "Succ[i] = i + 1" - using s5_1 nat_add_1 by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = i * j + (j + (i + 1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_3: "i + 1 \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "1 \\in Nat" - using oneIsNat nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + ((j + i) + 1)" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "j + (i + 1) = (j + i) + 1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_8: "... = i * j + ((i + j) + 1)" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "j + i = i + j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * Succ[i] = i * j + ((i + j) + 1)" - proof - - have s4_1: "Succ[j] * Succ[i] = Succ[j] * i + Succ[j]" - proof - - have s5_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = i * Succ[j] + Succ[j]" - using s2_2_induct by auto - also have s4_3: "... = (i * j + i) + Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - have s5_3: "i * Succ[j] = i * j + i" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + i) + (j + 1)" - proof - - have s5_1: "j \\in Nat" - using s1_2_jdom by auto - have s5_2: "Succ[j] = j + 1" - using s5_1 nat_add_1 by fast - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = i * j + (i + (j + 1))" - proof - - have s5_1: "i * j \\in Int" - using s1_2_jdom s2_2_idom nat_in_int multIsInt by auto - have s5_2: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_3: "j + 1 \\in Int" - using s1_2_jdom oneIsNat nat_in_int addIsInt by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_6: "... = i * j + ((i + j) + 1)" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "i + (j + 1) = (i + j) + 1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "Succ[i] * Succ[j] = Succ[j] * Succ[i]" - using s3_1 s3_2 by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * Succ[j] = Succ[j] * i - \ \ - -.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * Succ[j] = Succ[j] * i" - have s3_1: "-.Succ[-.i] * Succ[j] = - i * j + ((i + -.j) + -.1)" - proof - - have s4_1: "-.Succ[-.i] * Succ[j] = - -.Succ[-.i] * j + -.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = j * -.Succ[-.i] + -.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_3: "... = (j * -.-.i + -.j) + -.Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_2_jdom nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (j * i + -.j) + -.Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = (i * j + -.j) + -.Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_2_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by fast - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (i * j + -.j) + (i + -.1)" - proof - - have s5_1: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_2: "-.Succ[-.i] = -.-.i + -.1" - using s5_1 MinusSuccIsMinusOne by fast - have s5_3: "-.-.i = i" - using s2_3_idom neg_nat_in_int minus_involutive - by auto - have s5_4: "-.Succ[-.i] = i + -.1" - using s5_2 s5_3 by auto - show ?thesis - using s5_4 by auto - qed - also have s4_7: "... = i * j + (-.j + (i + -.1))" - proof - - have s5_1: "i * j \\in Int" - using s1_2_jdom nat_in_int s2_3_idom - neg_nat_in_int multIsInt by auto - have s5_2: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_3: "i + -.1 \\in Int" - using s2_3_idom minus_one_in_negnat - neg_nat_in_int addIsInt by auto - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + ((-.j + i) + -.1)" - proof - - have s5_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - have s5_4: " -.j + (i + -.1) = (-.j + i) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_9: "... = i * j + ((i + -.j) + -.1)" - proof - - have s5_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.j + i = i + -.j" - using s5_1 s5_2 AddCommutative_sequent - by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "Succ[j] * -.Succ[-.i] = - i * j + ((i + -.j) + -.1)" - proof - - have s4_1: "Succ[j] * -.Succ[-.i] = - Succ[j] * -.-.i + -.Succ[j]" - proof - - have s5_1: "Succ[j] \\in Int" - using s1_2_jdom succIsNat nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = Succ[j] * i + -.Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = i * Succ[j] + -.Succ[j]" - using s2_3_induct by auto - also have s4_4: "... = (i * j + i) + -.Succ[j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "j \\in Nat" - using s1_2_jdom by auto - have s5_3: "i * Succ[j] = i * j + i" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + i) + (-.j + -.1)" - proof - - have s5_1: "j \\in Nat" - using s1_2_jdom by auto - have s5_2: "-.Succ[j] = -.j + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = i * j + (i + (-.j + -.1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_2_jdom nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_3: "-.j + -.1 \\in Int" - proof - - have s6_1: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s6_2: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + ((i + -.j) + -.1)" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.j \\in Int" - using s1_2_jdom minus_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_negnat neg_nat_in_int by auto - have s5_4: "i + (-.j + -.1) = (i + -.j) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * Succ[j] = Succ[j] * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_2 natInduct[of "\ i. - i * Succ[j] = Succ[j] * i"] - by auto - have s2_5: "\ i \ negNat: i * Succ[j] = Succ[j] * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * Succ[j] = Succ[j] * i"] - by auto - show "\ i \ Int: i * Succ[j] = Succ[j] * i" - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_3: "\ j. \ - j \\in negNat; - \ i \ Int: i * j = j * i - \ \ - \ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" - proof - - fix "j" :: "c" - assume s1_3_jdom: "j \\in negNat" - assume s1_3_induct: "\ i \ Int: i * j = j * i" - have s2_1: "0 * -.Succ[-.j] = -.Succ[-.j] * 0" - proof - - have s3_1: "0 * -.Succ[-.j] = 0 * -.-.j + -.0" - proof - - have s4_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s4_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - also have s3_2: "... = 0 * -.-.j" - proof - - have s4_1: "0 * -.-.j \\in Int" - proof - - have s5_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s5_2: "-.-.j \\in Int" - using s1_3_jdom neg_nat_in_int minus_involutive - by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - have s4_2: "0 * -.-.j + 0 = 0 * -.-.j" - using s4_1 add_0 by auto - show ?thesis - using s4_2 neg0 by auto - qed - also have s3_3: "... = 0 * j" - proof - - have s4_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s4_1 minus_involutive by auto - qed - also have s3_4: "... = j * 0" - proof - - have s4_1: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s4_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - show ?thesis - using s4_1 s4_2 spec by fast - qed - also have s3_5: "... = 0" - proof - - have s4_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - also have s3_6: "... = -.Succ[-.j] * 0" - proof - - have s4_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - show ?thesis - using s4_1 mult_0 by auto - qed - finally show ?thesis . - qed - have s2_2: "\ i. \ - i \\in Nat; - i * -.Succ[-.j] = -.Succ[-.j] * i - \ \ - Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" - proof - - fix "i" :: "c" - assume s2_2_idom: "i \\in Nat" - assume s2_2_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" - have s3_1: "Succ[i] * -.Succ[-.j] = - i * j + ((-.i + j) + -.1)" - proof - - have s4_1: "Succ[i] * -.Succ[-.j] = - Succ[i] * -.-.j + -.Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = Succ[i] * j + -.Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = j * Succ[i] + -.Succ[i]" - proof - - have s5_1: "Succ[i] \\in Int" - using s2_2_idom succIsNat nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - have s5_3: "Succ[i] * j = j * Succ[i]" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (j * i + j) + -.Succ[i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - have s5_3: "j * Succ[i] = j * i + j" - using s5_1 s5_2 multint_succ_nat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (i * j + j) + -.Succ[i]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "\ k \ Int: k * j = j * k" - using s1_3_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = i * j + (j + -.Succ[i])" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_3: "-.Succ[i] \\in Int" - using s2_2_idom succIsNat minus_nat_in_int by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_7: "... = i * j + (j + (-.i + -.1))" - proof - - have s5_1: "i \\in Nat" - using s2_2_idom by auto - have s5_2: "-.Succ[i] = -.i + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_8: "... = i * j + ((j + -.i) + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_int by auto - have s5_4: "j + (-.i + -.1) = (j + -.i) + -.1" - using s5_1 s5_2 s5_3 - AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_9: "... = i * j + ((-.i + j) + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "j + -.i = -.i + j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "-.Succ[-.j] * Succ[i] = - i * j + ((-.i + j) + -.1)" - proof - - have s4_1: "-.Succ[-.j] * Succ[i] = - -.Succ[-.j] * i + -.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "i \\in Nat" - using s2_2_idom by auto - show ?thesis - using s5_1 s5_2 multint_succ_nat by auto - qed - also have s4_2: "... = i * -.Succ[-.j] + -.Succ[-.j]" - using s2_2_induct by auto - also have s4_3: "... = (i * -.-.j + -.i) + -.Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_4: "... = (i * j + -.i) + -.Succ[-.j]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_5: "... = (i * j + -.i) + (-.-.j + -.1)" - proof - - have s5_1: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_2: "-.Succ[-.j] = -.-.j + -.1" - using s5_1 MinusSuccIsMinusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_6: "... = (i * j + -.i) + (j + -.1)" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = i * j + (-.i + (j + -.1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_2_idom nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_3: "j + -.1 \\in Int" - proof - - have s6_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s6_2: "-.1 \\in Int" - using minus_one_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + ((-.i + j) + -.1)" - proof - - have s5_1: "-.i \\in Int" - using s2_2_idom minus_nat_in_int by auto - have s5_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_3: "-.1 \\in Int" - using minus_one_in_int by auto - have s5_4: "-.i + (j + -.1) = (-.i + j) + -.1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "Succ[i] * -.Succ[-.j] = -.Succ[-.j] * Succ[i]" - using s3_1 s3_2 by auto - qed - have s2_3: "\ i. \ - i \\in negNat; - i * -.Succ[-.j] = -.Succ[-.j] * i - \ \ - -.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" - proof - - fix "i" :: "c" - assume s2_3_idom: "i \\in negNat" - assume s2_3_induct: "i * -.Succ[-.j] = -.Succ[-.j] * i" - have s3_1: "-.Succ[-.i] * -.Succ[-.j] = - i * j + ((-.i + -.j) + 1)" - proof - - have s4_1: "-.Succ[-.i] * -.Succ[-.j] = - -.Succ[-.i] * -.-.j + -.-.Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = -.Succ[-.i] * j + -.-.Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = -.Succ[-.i] * j + Succ[-.i]" - proof - - have s5_1: "Succ[-.i] \\in Int" - using s2_3_idom negnat_succ_in_nat nat_in_int by auto - have s5_2: "-.-.Succ[-.i] = Succ[-.i]" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_4: "... = j * -.Succ[-.i] + Succ[-.i]" - proof - - have s5_1: "-.Succ[-.i] \\in Int" - using s2_3_idom minus_negnat_succ_in_negnat - neg_nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_3_induct by auto - have s5_3: "-.Succ[-.i] * j = j * -.Succ[-.i]" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_5: "... = (j * -.-.i + -.j) + Succ[-.i]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - have s5_3: "j * -.Succ[-.i] = j * -.-.i + -.j" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (j * i + -.j) + Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = (i * j + -.j) + Succ[-.i]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "\ k \ Int: - k * j = j * k" - using s1_3_induct by auto - have s5_3: "i * j = j * i" - using s5_1 s5_2 spec by auto - show ?thesis - using s5_3 by auto - qed - also have s4_8: "... = (i * j + -.j) + (-.i + 1)" - proof - - have s5_1: "i \\in negNat" - using s2_3_idom by auto - have s5_2: "Succ[-.i] = -.i + 1" - using s5_1 SuccMinusIsPlusOne by fast - show ?thesis - using s5_2 by auto - qed - also have s4_9: "... = i * j + (-.j + (-.i + 1))" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_3: "-.i + 1 \\in Int" - proof - - have s6_1: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s6_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - show ?thesis - using s6_1 s6_2 addIsInt by auto - qed - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_10: "... = i * j + ((-.j + -.i) + 1)" - proof - - have s5_1: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "-.j + (-.i + 1) = (-.j + -.i) + 1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - also have s4_11: "... = i * j + ((-.i + -.j) + 1)" - proof - - have s5_1: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "-.j + -.i = -.i + -.j" - using s5_1 s5_2 AddCommutative_sequent by auto - show ?thesis - using s5_3 by auto - qed - finally show ?thesis . - qed - have s3_2: "-.Succ[-.j] * -.Succ[-.i] = - i * j + ((-.i + -.j) + 1)" - proof - - have s4_1: "-.Succ[-.j] * -.Succ[-.i] = - -.Succ[-.j] * -.-.i + -.-.Succ[-.j]" - proof - - have s5_1: "-.Succ[-.j] \\in Int" - using s1_3_jdom minus_succ_minus_negnat_in_int - by auto - have s5_2: "-.i \\in Nat" - using s2_3_idom minus_neg_nat_in_nat by auto - show ?thesis - using s5_1 s5_2 multint_succ_negnat by auto - qed - also have s4_2: "... = -.Succ[-.j] * i + -.-.Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.-.i = i" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_3: "... = -.Succ[-.j] * i + Succ[-.j]" - proof - - have s5_1: "Succ[-.j] \\in Int" - using s1_3_jdom negnat_succ_in_nat - nat_in_int by auto - have s5_2: "-.-.Succ[-.j] = Succ[-.j]" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_4: "... = i * -.Succ[-.j] + Succ[-.j]" - using s2_3_induct by auto - also have s4_5: "... = (i * -.-.j + -.i) + Succ[-.j]" - proof - - have s5_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s5_2: "-.j \\in Nat" - using s1_3_jdom minus_neg_nat_in_nat by auto - have s5_3: "i * -.Succ[-.j] = i * -.-.j + -.i" - using s5_1 s5_2 multint_succ_negnat by auto - show ?thesis - using s5_3 by auto - qed - also have s4_6: "... = (i * j + -.i) + Succ[-.j]" - proof - - have s5_1: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - have s5_2: "-.-.j = j" - using s5_1 minus_involutive by auto - show ?thesis - using s5_2 by auto - qed - also have s4_7: "... = i * j + (-.i + Succ[-.j])" - proof - - have s5_1: "i * j \\in Int" - proof - - have s6_1: "i \\in Int" - using s2_3_idom neg_nat_in_int by auto - have s6_2: "j \\in Int" - using s1_3_jdom neg_nat_in_int by auto - show ?thesis - using s6_1 s6_2 multIsInt by auto - qed - have s5_2: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_3: "Succ[-.j] \\in Int" - using s1_3_jdom negnat_succ_in_nat nat_in_int by auto - show ?thesis - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - qed - also have s4_8: "... = i * j + (-.i + (-.j + 1))" - proof - - have s5_1: "Succ[-.j] = -.j + 1" - using s1_3_jdom SuccMinusIsPlusOne by fast - show ?thesis - using s5_1 by auto - qed - also have s4_9: "... = i * j + ((-.i + -.j) + 1)" - proof - - have s5_1: "-.i \\in Int" - using s2_3_idom minus_negnat_in_int by auto - have s5_2: "-.j \\in Int" - using s1_3_jdom minus_negnat_in_int by auto - have s5_3: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s5_4: "-.i + (-.j + 1) = (-.i + -.j) + 1" - using s5_1 s5_2 s5_3 AddAssociative_sequent by auto - show ?thesis - using s5_4 by auto - qed - finally show ?thesis . - qed - show "-.Succ[-.i] * -.Succ[-.j] = -.Succ[-.j] * -.Succ[-.i]" - using s3_1 s3_2 by auto - qed - have s2_4: "\ i \ Nat: - i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_2 natInduct[of "\ i. - i * -.Succ[-.j] = -.Succ[-.j] * i"] - by auto - have s2_5: "\ i \ negNat: - i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_1 s2_3 neg_nat_induction[of "\ i. - i * -.Succ[-.j] = -.Succ[-.j] * i"] - by auto - show "\ i \ Int: i * -.Succ[-.j] = -.Succ[-.j] * i" - using s2_4 s2_5 int_union_nat_negnat by auto - qed - have s1_4: "\ j \ Nat: \ i \ Int: - i * j = j * i" - using s1_1 s1_2 natInduct[of "\ j. - \ i \ Int: i * j = j * i"] - by blast - have s1_5: "\ j \ negNat: \ i \ Int: - i * j = j * i" - using s1_1 s1_3 neg_nat_induction[of "\ j. - \ i \ Int: i * j = j * i"] - by blast - show ?thesis - using s1_4 s1_5 int_union_nat_negnat by auto - qed - - -theorem MultCommutative_sequent: - assumes idom: "i \\in Int" and jdom: "j \\in Int" - shows "i * j = j * i" - using idom jdom MultCommutative spec by auto - - -(* -THEOREM mult_0_left == - ASSUME m \in Int - PROVE 0 * m = 0 - BY mult_0, MultCommutative -*) -theorem mult_0_left: - assumes ndom: "n \\in Int" - shows "0 * n = 0" - proof - - have s1_1: "n * 0 = 0" - using ndom mult_0 by auto - have s1_2: "n * 0 = 0 * n" - using ndom zero_in_int MultCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Associativity of multiplication. *) - - -(* -THEOREM MultLeftDistributesAdd == - \A a, b, c \in Int: - a * (b + c) = a * b + a * c -PROOF -<1>1. SUFFICES - ASSUME NEW b \in Int, NEW c \in Int - PROVE \A a \in Int: a * (b + c) = a * b + a * c - BY <1>1 -<1>2. 0 * (b + c) = 0 * b + 0 * c - <2>1. 0 * (b + c) = 0 - <3>1. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, mult_0_left - <2>2. 0 * b + 0 * c = 0 - <3>1. 0 * b = 0 - BY <1>1, mult_0_left - <3>2. 0 * c = 0 - BY <1>1, mult_0_left - <3> QED - BY <3>1, <3>2, zero_in_int, add_0 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW a \in Nat, - a * (b + c) = a * b + a * c - PROVE Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c - <2>1. Succ[a] * (b + c) = (b + c) * Succ[a] - <3>1. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, MultCommutative - <2>2. @ = (b + c) * a + (b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. a \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>3. @ = a * (b + c) + (b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. a \in Int - BY <1>3, nat_in_int - <3>3. (b + c) * a = a * (b + c) - BY <3>1, <3>2, MultCommutative - <3> QED - BY <3>3 - <2>4. @ = (a * b + a * c) + (b + c) - BY <1>3 \* induction hypothesis - <2>5. @ = a * b + (a * c + (b + c)) - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>6. @ = a * b + ((a * c + b) + c) - <3>1. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. c \in Int - BY <1>1 - <3>4. a * c + (b + c) = (a * c + b) + c - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>7. @ = a * b + ((b + a * c) + c) - <3>1. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. a * c + b = b + a * c - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>8. @ = a * b + (b + (a * c + c)) - <3>1. b \in Int - BY <1>1 - <3>2. a * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. c \in Int - BY <1>1 - <3>4. (b + a * c) + c = b + (a * c + c) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>9. @ = (a * b + b) + (a * c + c) - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. b \in Int - BY <1>1 - <3>3. (a * c + c) \in Int - BY <1>1, <1>3, nat_in_int, addIsInt, multIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>10. @ = (b * a + b) + (c * a + c) - <3>1. a * b = b * a - <4>1. a \in Int - BY <1>3, nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. a * c = c * a - <4>1. a \in Int - BY <1>3, nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2>11. @ = b * Succ[a] + c * Succ[a] - <3>1. b * Succ[a] = b * a + b - <4>1. b \in Int - BY <1>1 - <4>2. a \in Nat - BY <1>3 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3>2. c * Succ[a] = c * a + c - <4>1. c \in Int - BY <1>1 - <4>2. a \in Nat - BY <1>3 - <4> QED - BY <4>1, <4>2, multint_succ_nat - <3> QED - BY <3>1, <3>2 - <2>12. @ = Succ[a] * b + Succ[a] * c - <3>1. b * Succ[a] = Succ[a] * b - <4>1. b \in Int - BY <1>1 - <4>2. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. c * Succ[a] = Succ[a] * c - <4>1. c \in Int - BY <1>1 - <4>2. Succ[a] \in Int - BY <1>3, succIsNat, nat_in_int - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, - <2>5, <2>6, <2>7, <2>8, <2>9, - <2>10, <2>11, <2>12 -<1>4. ASSUME NEW a \in negNat, - a * (b + c) = a * b + a * c - PROVE -.Succ[-.a] * (b + c) = -.Succ[-.a] * b + -.Succ[-.a] * c - <2>1. -.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a] - <3>1. -.Succ[-.a] \in Int - BY <1>4, minus_succ_minus_negnat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3> QED - BY <3>1, <3>2, MultCommutative - <2>2. @ = (b + c) * -.-.a + -.(b + c) - <3>1. b + c \in Int - BY <1>1, addIsInt - <3>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>3. @ = (b + c) * a + -.(b + c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>4. @ = a * (b + c) + -.(b + c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. b + c \in Int - BY <1>1, addIsInt - <3>3. (b + c) * a = a * (b + c) - BY <3>1, <3>2, MultCommutative - <3> QED - BY <3>3 - <2>5. @ = (a * b + a * c) + -.(b + c) - BY <1>4 \* induction hypothesis - <2>6. @ = (a * b + a * c) + (-.b + -.c) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Int - BY <1>1 - <3>3. -.(b + c) = -.b + -.c - BY <3>1, <3>2, MinusDistributesAdd, spec - <3> QED - BY <3>3 - <2>7. @ = a * b + (a * c + (-.b + -.c)) - <3>1. a * b \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.b + -.c \in Int - <4>1. -.b \in Int - BY <1>1, neg_int_eq_int - <4>2. -.c \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>8. @ = a * b + ((a * c + -.b) + -.c) - <3>1. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. -.c \in Int - BY <1>1, neg_int_eq_int - <3>4. a * c + (-.b + -.c) = - (a * c + -.b) + -.c - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>9. @ = a * b + ((-.b + a * c) + -.c) - <3>1. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. a * c + -.b = -.b + a * c - BY <3>1, <3>2, AddCommutative - <3> QED - BY <3>3 - <2>10. @ = a * b + (-.b + (a * c + -.c)) - <3>1. -.b \in Int - BY <1>1, neg_int_eq_int - <3>2. a * c \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.c \in Int - BY <1>1, neg_int_eq_int - <3>4. (-.b + a * c) + -.c = - -.b + (a * c + -.c) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <3>4 - <2>11. @ = (a * b + -.b) + (a * c + -.c) - <3>1. a * b \in Int - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.b \in Int - BY <1>1, neg_int_eq_int - <3>3. a * c + -.c \in Int - <4>1. a * c \in Int - <5>1. a \in Int - BY <1>4, neg_nat_in_int - <5>2. c \in Int - BY <1>1 - <5> QED - BY <5>1, <5>2, multIsInt - <4>2. -.c \in Int - BY <1>1, neg_int_eq_int - <4> QED - BY <4>1, <4>2, addIsInt - <3> QED - BY <3>1, <3>2, <3>3, AddAssociative - <2>12. @ = (b * a + -.b) + (c * a + -.c) - <3>1. a * b = b * a - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. b \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3>2. a * c = c * a - <4>1. a \in Int - BY <1>4, neg_nat_in_int - <4>2. c \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, MultCommutative - <3> QED - BY <3>1, <3>2 - <2>13. @ = (b * -.-.a + -.b) + (c * -.-.a + -.c) - <3>1. a \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>14. @ = b * -.Succ[-.a] + c * -.Succ[-.a] - <3>1. b * -.Succ[-.a] = b * -.-.a + -.b - <4>1. b \in Int - BY <1>1 - <4>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3>2. c * -.Succ[-.a] = c * -.-.a + -.c - <4>1. c \in Int - BY <1>1 - <4>2. -.a \in Nat - BY <1>4, minus_neg_nat_in_nat - <4> QED - BY <4>1, <4>2, multint_succ_negnat - <3> QED - BY <3>1, <3>2 - <2>15. @ = -.Succ[-.a] * b + -.Succ[-.a] * c - <3>1. -.Succ[-.a] \in Int - BY <1>4, minus_succ_minus_negnat_in_int - <3>2. b * -.Succ[-.a] = -.Succ[-.a] * b - <4>1. b \in Int - BY <1>1 - <4> QED - BY <3>1, <4>1, MultCommutative - <3>3. c * -.Succ[-.a] = -.Succ[-.a] * c - <4>1. c \in Int - BY <1>1 - <4> QED - BY <3>1, <4>1, MultCommutative - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, <2>5, 2<6>, - <2>7, <2>8, <2>9, <2>10, <2>11, - <2>12, <2>13, <2>14, <2>15 -<1>5. \A a \in Nat: a * (b + c) = a * b + a * c - BY <1>2, <1>3, NatInduction -<1>6. \A a \in negNat: a * (b + c) = a * b + a * c - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MultLeftDistributesAdd_forall: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - a * (b + c) = a * b + a * c" - proof - - { - fix "b" :: "c" - fix "c" :: "c" - assume bdom: "b \\in Int" - assume cdom: "c \\in Int" - have s1_2: "0 * (b + c) = 0 * b + 0 * c" - proof - - have s2_1: "0 * (b + c) = 0" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 mult_0_left by auto - qed - have s2_2: "0 * b + 0 * c = 0" - proof - - have s3_1: "0 * b = 0" - using bdom mult_0_left by auto - have s3_2: "0 * c = 0" - using cdom mult_0_left by auto - show ?thesis - using s3_1 s3_2 zero_in_int add_0 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ a. \ - a \\in Nat; - a * (b + c) = a * b + a * c - \ \ - Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" - proof - - fix "a" :: "c" - assume s1_3_adom: "a \\in Nat" - assume s1_3_induct: "a * (b + c) = a * b + a * c" - have s2_1: "Succ[a] * (b + c) = (b + c) * Succ[a]" - proof - - have s3_1: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 MultCommutative_sequent by auto - qed - also have s2_2: "... = (b + c) * a + (b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_3: "... = a * (b + c) + (b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s3_3: "(b + c) * a = a * (b + c)" - using s3_1 s3_2 MultCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_4: "... = (a * b + a * c) + (b + c)" - using s1_3_induct by auto - also have s2_5: "... = a * b + (a * c + (b + c))" - proof - - have s3_1: "a * b \\in Int" - using bdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_3: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_6: "... = a * b + ((a * c + b) + c)" - proof - - have s3_1: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "c \\in Int" - using cdom by auto - have s3_4: "a * c + (b + c) = (a * c + b) + c" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_7: "... = a * b + ((b + a * c) + c)" - proof - - have s3_1: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "a * c + b = b + a * c" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_8: "... = a * b + (b + (a * c + c))" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "a * c \\in Int" - using cdom s1_3_adom nat_in_int multIsInt by auto - have s3_3: "c \\in Int" - using cdom by auto - have s3_4: "(b + a * c) + c = b + (a * c + c)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_9: "... = (a * b + b) + (a * c + c)" - proof - - have s3_1: "a * b \\in Int" - using bdom s1_3_adom nat_in_int multIsInt by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "a * c + c \\in Int" - using cdom s1_3_adom nat_in_int addIsInt multIsInt by auto - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_10: "... = (b * a + b) + (c * a + c)" - proof - - have s3_1: "a * b = b * a" - proof - - have s4_1: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - have s3_2: "a * c = c * a" - proof - - have s4_1: "a \\in Int" - using s1_3_adom nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_11: "... = b * Succ[a] + c * Succ[a]" - proof - - have s3_1: "b * Succ[a] = b * a + b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - have s3_2: "c * Succ[a] = c * a + c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "a \\in Nat" - using s1_3_adom by auto - show ?thesis - using s4_1 s4_2 multint_succ_nat by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_12: "... = Succ[a] * b + Succ[a] * c" - proof - - have s3_1: "b * Succ[a] = Succ[a] * b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - show ?thesis - using s4_1 s4_2 MultCommutative by auto - qed - have s3_2: "c * Succ[a] = Succ[a] * c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "Succ[a] \\in Int" - using s1_3_adom succIsNat nat_in_int by auto - show ?thesis - using s4_1 s4_2 MultCommutative by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - finally show "Succ[a] * (b + c) = Succ[a] * b + Succ[a] * c" . - qed - have s1_4: "\ a. \ - a \\in negNat; - a * (b + c) = a * b + a * c - \ \ - -.Succ[-.a] * (b + c) = - -.Succ[-.a] * b + -.Succ[-.a] * c" - proof - - fix "a" :: "c" - assume s1_4_adom: "a \\in negNat" - assume s1_4_induct: "a * (b + c) = a * b + a * c" - have s2_1: "-.Succ[-.a] * (b + c) = (b + c) * -.Succ[-.a]" - proof - - have s3_1: "-.Succ[-.a] \\in Int" - using s1_4_adom minus_succ_minus_negnat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - show ?thesis - using s3_1 s3_2 MultCommutative by auto - qed - also have s2_2: "... = (b + c) * -.-.a + -.(b + c)" - proof - - have s3_1: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_3: "... = (b + c) * a + -.(b + c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_4: "... = a * (b + c) + -.(b + c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "b + c \\in Int" - using bdom cdom addIsInt by auto - have s3_3: "(b + c) * a = a * (b + c)" - using s3_1 s3_2 MultCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = (a * b + a * c) + -.(b + c)" - using s1_4_induct by auto - also have s2_6: "... = (a * b + a * c) + (-.b + -.c)" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "c \\in Int" - using cdom by auto - have s3_3: "-.(b + c) = -.b + -.c" - using s3_1 s3_2 MinusDistributesAdd spec by auto - show ?thesis - using s3_3 by auto - qed - also have s2_7: "... = a * b + (a * c + (-.b + -.c))" - proof - - have s3_1: "a * b \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.b + -.c \\in Int" - proof - - have s4_1: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s4_2: "-.c \\in Int" - using cdom neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_8: "... = a * b + ((a * c + -.b) + -.c)" - proof - - have s3_1: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "-.c \\in Int" - using cdom neg_int_eq_int by auto - have s3_4: "a * c + (-.b + -.c) = - (a * c + -.b) + -.c" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_9: "... = a * b + ((-.b + a * c) + -.c)" - proof - - have s3_1: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "a * c + -.b = -.b + a * c" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s3_3 by auto - qed - also have s2_10: "... = a * b + (-.b + (a * c + -.c))" - proof - - have s3_1: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_2: "a * c \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.c \\in Int" - using cdom neg_int_eq_int by auto - have s3_4: "(-.b + a * c) + -.c = - -.b + (a * c + -.c)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s3_4 by auto - qed - also have s2_11: "... = (a * b + -.b) + (a * c + -.c)" - proof - - have s3_1: "a * b \\in Int" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s3_3: "a * c + -.c \\in Int" - proof - - have s4_1: "a * c \\in Int" - proof - - have s5_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s5_2: "c \\in Int" - using cdom by auto - show ?thesis - using s5_1 s5_2 multIsInt by auto - qed - have s4_2: "-.c \\in Int" - using cdom neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 addIsInt by auto - qed - show ?thesis - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - qed - also have s2_12: "... = (b * a + -.b) + (c * a + -.c)" - proof - - have s3_1: "a * b = b * a" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "b \\in Int" - using bdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - have s3_2: "a * c = c * a" - proof - - have s4_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s4_2: "c \\in Int" - using cdom by auto - show ?thesis - using s4_1 s4_2 MultCommutative_sequent by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_13: "... = (b * -.-.a + -.b) + (c * -.-.a + -.c)" - proof - - have s3_1: "a \\in Int" - using s1_4_adom neg_nat_in_int by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_14: "... = b * -.Succ[-.a] + c * -.Succ[-.a]" - proof - - have s3_1: "b * -.Succ[-.a] = b * -.-.a + -.b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - have s4_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - have s3_2: "c * -.Succ[-.a] = c * -.-.a + -.c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - have s4_2: "-.a \\in Nat" - using s1_4_adom minus_neg_nat_in_nat by auto - show ?thesis - using s4_1 s4_2 multint_succ_negnat by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_15: "... = -.Succ[-.a] * b + -.Succ[-.a] * c" - proof - - have s3_1: "-.Succ[-.a] \\in Int" - using s1_4_adom minus_succ_minus_negnat_in_int by auto - have s3_2: "b * -.Succ[-.a] = -.Succ[-.a] * b" - proof - - have s4_1: "b \\in Int" - using bdom by auto - show ?thesis - using s3_1 s4_1 MultCommutative_sequent by auto - qed - have s3_3: "c * -.Succ[-.a] = -.Succ[-.a] * c" - proof - - have s4_1: "c \\in Int" - using cdom by auto - show ?thesis - using s3_1 s4_1 MultCommutative_sequent by auto - qed - show ?thesis - using s3_2 s3_3 by auto - qed - finally show "-.Succ[-.a] * (b + c) = - -.Succ[-.a] * b + -.Succ[-.a] * c" . - qed - have s1_5: "\ a \ Nat: - a * (b + c) = a * b + a * c" - using s1_2 s1_3 natInduct[of "\ a. - a * (b + c) = a * b + a * c"] - by auto - have s1_6: "\ a \ negNat: - a * (b + c) = a * b + a * c" - using s1_2 s1_4 neg_nat_induction[of "\ a. - a * (b + c) = a * b + a * c"] - by auto - have s1_7: "\ a \ Int: - a * (b + c) = a * b + a * c" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -theorem MultLeftDistributesAdd: - assumes adom: "a \\in Int" and bdom: "b \\in Int" and - cdom: "c \\in Int" - shows "a * (b + c) = a * b + a * c" - using adom bdom cdom MultLeftDistributesAdd_forall by auto - - -theorem MultRightDistributesAdd: - assumes adom: "a \\in Int" and bdom: "b \\in Int" and - cdom: "c \\in Int" - shows "(b + c) * a = b * a + c * a" - proof - - have s1_1: "a * (b + c) = a * b + a * c" - using adom bdom cdom MultLeftDistributesAdd by auto - have s1_2: "a * (b + c) = (b + c) * a" - using adom bdom cdom addIsInt MultCommutative by auto - have s1_3: "a * b = b * a" - using adom bdom MultCommutative by auto - have s1_4: "a * c = c * a" - using adom cdom MultCommutative by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -lemmas mult_distributes[algebra_simps] = - MultLeftDistributesAdd MultRightDistributesAdd - - -(* -THEOREM MinusCommutesRightMult == - \A a, b \in Int: -.(a * b) = a * -.b -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int - PROVE \A b \in Int: -.(a * b) = a * -.b - BY <1>1 -<1>2. -.(a * 0) = a * -.0 - <2>1. -.(a * 0) = 0 - BY <1>1, mult_0, neg0 - <2>2. a * -.0 = 0 - BY <1>1, neg0, mult_0 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW b \in Nat, - -.(a * b) = a * -.b - PROVE -.(a * Succ[b]) = a * -.Succ[b] - <2>1. -.(a * Succ[b]) = -.(a * b + a) - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>2. @ = -.(a * b) + -.a - <3>1. a * b \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>2. a \in Int - BY <1>1 - <3> QED - BY <3>1, <3>2, MinusDistributesAdd - <2>3. @ = a * -.b + -.a - BY <1>3 \* induction hypothesis - <2>4. @ = a * -.Succ[b] - <3>1. a \in Int - BY <1>1 - <3>2. b \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>4. ASSUME NEW b \in negNat, - -.(a * b) = a * -.b - PROVE -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b] - <2>1. -.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a) - <3>1. a \in Int - BY <1>1 - <3>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>2. @ = -.(a * -.-.b) + -.-.a - <3>1. a * -.-.b \in Int - <4>1. a \in Int - BY <1>1 - <4>2. -.-.b \in Int - BY <1>4, minus_negnat_in_int, - neg_int_eq_int - <4> QED - BY <4>1, <4>2, multIsInt - <3>2. -.a \in Int - BY <1>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, MinusDistributesAdd - <2>3. @ = -.(a * -.-.b) + a - <3>1. a \in Int - BY <1>1 - <3>2. -.-.a = a - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>4. @ = -.(a * b) + a - <3>1. b \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.b = b - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>5. @ = a * -.b + a - BY <1>4 \* induction hypothesis - <2>6. @ = a * Succ[-.b] - <3>1. a \in Int - BY <1>1 - <3>2. -.b \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>7. @ = a * -.-.Succ[-.b] - <3>1. Succ[-.b] \in Int - BY <1>4, negnat_succ_in_nat, nat_in_int - <3>2. -.-.Succ[-.b] = Succ[-.b] - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2> QED - BY <2>1, <2>2, <2>3, ,2>4, - <2>5, <2>6, <2>7 -<1>5. \A b \in Nat: -.(a * b) = a * -.b - BY <1>2, <1>3, NatInduction -<1>6. \A b \in negNat: -.(a * b) = a * -.b - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MinusCommutesRightMult: - "\ a \ Int: \ b \ Int: - -.(a * b) = a * -.b" - proof - - { - fix "a" :: "c" - assume adom: "a \\in Int" - have s1_2: "-.(a * 0) = a * -.0" - proof - - have s2_1: "-.(a * 0) = 0" - using adom mult_0 neg0 by auto - have s2_2: "a * -.0 = 0" - using adom neg0 mult_0 by auto - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ b. \ - b \\in Nat; - -.(a * b) = a * -.b - \ \ - -.(a * Succ[b]) = a * -.Succ[b]" - proof - - fix "b" :: "c" - assume s1_3_bdom: "b \\in Nat" - assume s1_3_induct: "-.(a * b) = a * -.b" - have s2_1: "-.(a * Succ[b]) = -.(a * b + a)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_2: "... = -.(a * b) + -.a" - proof - - have s3_1: "a * b \\in Int" - using adom s1_3_bdom nat_in_int multIsInt by auto - have s3_2: "a \\in Int" - using adom by auto - show ?thesis - using s3_1 s3_2 MinusDistributesAdd by auto - qed - also have s2_3: "... = a * -.b + -.a" - using s1_3_induct by auto - also have s2_4: "... = a * -.Succ[b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Nat" - using s1_3_bdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - finally show "-.(a * Succ[b]) = a * -.Succ[b]" . - qed - have s1_4: "\ b. \ - b \\in negNat; - -.(a * b) = a * -.b - \ \ - -.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" - proof - - fix "b" :: "c" - assume s1_4_bdom: "b \\in negNat" - assume s1_4_induct: "-.(a * b) = a * -.b" - have s2_1: "-.(a * -.Succ[-.b]) = -.(a * -.-.b + -.a)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_2: "... = -.(a * -.-.b) + -.-.a" - proof - - have s3_1: "a * -.-.b \\in Int" - proof - - have s4_1: "a \\in Int" - using adom by auto - have s4_2: "-.-.b \\in Int" - using s1_4_bdom minus_negnat_in_int - neg_int_eq_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_2: "-.a \\in Int" - using adom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 MinusDistributesAdd by auto - qed - also have s2_3: "... = -.(a * -.-.b) + a" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.-.a = a" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_4: "... = -.(a * b) + a" - proof - - have s3_1: "b \\in Int" - using s1_4_bdom neg_nat_in_int by auto - have s3_2: "-.-.b = b" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_5: "... = a * -.b + a" - using s1_4_induct by auto - also have s2_6: "... = a * Succ[-.b]" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "-.b \\in Nat" - using s1_4_bdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_7: "... = a * -.-.Succ[-.b]" - proof - - have s3_1: "Succ[-.b] \\in Int" - using s1_4_bdom negnat_succ_in_nat nat_in_int by auto - have s3_2: "-.-.Succ[-.b] = Succ[-.b]" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - finally show "-.(a * -.Succ[-.b]) = a * -.-.Succ[-.b]" . - qed - have s1_5: "\ b \ Nat: - -.(a * b) = a * -.b" - using s1_2 s1_3 natInduct[of "\ b. - -.(a * b) = a * -.b"] - by auto - have s1_6: "\ b \ negNat: - -.(a * b) = a * -.b" - using s1_2 s1_4 neg_nat_induction[of "\ b. - -.(a * b) = a * -.b"] - by auto - have s1_7: "\ b \ Int: - -.(a * b) = a * -.b" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - -(* -THEOREM MinusCommutesLeftMult == - \A a, b \in Int: -.(a * b) = -.a * b -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE -.(a * b) = -.a * b - BY <1>1 -<1>2. -.(a * b) = -.(b * a) - BY <1>1, MultCommutative -<1>3. @ = b * -.a - BY <1>1, MinusCommutesRightMult -<1>4. @ = -.a * b - BY <1>1, neg_int_eq_int, MultCommutative -<1> QED - BY <1>2, <1>3, <1>4 -*) -theorem MinusCommutesLeftMult: - "\ a \ Int: \ b \ Int: - -.(a * b) = -.a * b" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \\in Int" - assume bdom: "b \\in Int" - have s1_2: "-.(a * b) = -.(b * a)" - using adom bdom MultCommutative_sequent by auto - also have s1_3: "... = b * -.a" - using adom bdom MinusCommutesRightMult by auto - also have s1_4: "... = -.a * b" - using adom bdom neg_int_eq_int - MultCommutative_sequent by auto - finally have "-.(a * b) = -.a * b" - by auto - } - from this show ?thesis - by auto - qed - - -(* -THEOREM MultAssociative == - \A a \in Int: \A b \in Int: \A c \in Int: - (a * b) * c = a * (b * c) -PROOF -<1>1. SUFFICES - ASSUME NEW a \in Int, NEW b \in Int - PROVE \A c \in Int: (a * b) * c = a * (b * c) - BY <1>1, allI -<1>2. (a * b) * 0 = a * (b * 0) - <2>1. (a * b) * 0 = 0 - <3>1. a * b \in Int - BY <1>1, multIsInt - <3> QED - BY <3>1, mult_0 - <2>2. a * (b * 0) = 0 - <3>1. b * 0 = 0 - BY <1>1, mult_0 - <3>2. a * 0 = 0 - BY <1>1, mult_0 - <3> QED - BY <3>1, <3>2 - <2> QED - BY <2>1, <2>2 -<1>3. ASSUME NEW c \in Nat, - (a * b) * c = a * (b * c) - PROVE (a * b) * Succ[c] = a * (b * Succ[c]) - <2>1. (a * b) * Succ[c] = (a * b) * c + (a * b) - <3>1. a * b \in Int - BY <1>1, multIsInt - <3>2. c \in Nat - BY <1>3 - <3> QED - BY <3>1, <3>2, multint_succ_nat - <2>2. @ = a * (b * c) + (a * b) - BY <1>3 \* induction hypothesis - <2>3. @ = a * (b * c + b) - <3>1. a \in Int - BY <1>1 - <3>2. b * c \in Int - BY <1>1, <1>3, nat_in_int, multIsInt - <3>3. b \in Int - BY <1>1 - <3> QED - BY <3>1, <3>2, <3>3, MultLeftDistributesAdd - <2>4. @ = a * (b * Succ[c]) - <3>1. b \in Int - BY <1>1 - <3>2. c \in Nat - BY <1>3 - <3>3. b * Succ[c] = b * c + b - BY <3>1, <3>2, multint_succ_nat - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4 -<1>4. ASSUME NEW c \in negNat, - (a * b) * c = a * (b * c) - PROVE (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c]) - <2>1. (a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b) - <3>1. a * b \in Int - BY <1>1, multIsInt - <3>2. -.c \in Nat - BY <1>4, minus_neg_nat_in_nat - <3> QED - BY <3>1, <3>2, multint_succ_negnat - <2>2. @ = (a * b) * c + -.(a * b) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.c = c - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>3. @ = a * (b * c) + -.(a * b) - BY <1>4 \* induction hypothesis - <2>4. @ = a * (b * c) + a * -.b - <3>1. a \in Int - BY <1>1 - <3>2. b \in Int - BY <1>1 - <3>3. -.(a * b) = a * -.b - BY <3>1, <3>2, MinusCommutesRightMult - <3> QED - BY <3>3 - <2>5. @ = a * (b * c + -.b) - <3>1. a \in Int - BY <1>1 - <3>2. b * c \in Int - BY <1>1 <1>4, neg_nat_in_int, multIsInt - <3>3. -.b \in Int - BY <1>1, neg_int_eq_int - <3> QED - BY <3>1, <3>2, <3>3, MultLeftDistributesAdd - <2>6. @ = a * (b * -.-.c + -.b) - <3>1. c \in Int - BY <1>4, neg_nat_in_int - <3>2. -.-.c = c - BY <3>1, minus_involutive - <3> QED - BY <3>2 - <2>7. @ = a * (b * -.Succ[-.c]) - <3>1. b \in Int - BY <1>1 - <3>2. -.c \in Nat - BY <1>4, minus_neg_nat_in_nat - <3>3. b * -.Succ[-.c] = b * -.-.c + -.b - BY <3>1, <3>2, multint_succ_negnat - <3> QED - BY <3>3 - <2> QED - BY <2>1, <2>2, <2>3, <2>4, - <2>5, <2>6, <2>7 -<1>5. \A c \in Nat: (a * b) * c = a * (b * c) - BY <1>2, <1>3, NatInduction -<1>6. \A c \in negNat: (a * b) * c = a * (b * c) - BY <1>2, <1>4, neg_nat_induction -<1> QED - BY <1>5, <1>6, int_union_nat_negnat -*) -theorem MultAssociative: - "\ a \ Int: - \ b \ Int: - \ c \ Int: - (a * b) * c = a * (b * c)" - proof - - { - fix "a" :: "c" - fix "b" :: "c" - assume adom: "a \\in Int" - assume bdom: "b \\in Int" - have s1_2: "(a * b) * 0 = a * (b * 0)" - proof - - have s2_1: "(a * b) * 0 = 0" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - show ?thesis - using s3_1 mult_0 by auto - qed - have s2_2: "a * (b * 0) = 0" - proof - - have s3_1: "b * 0 = 0" - using bdom mult_0 by auto - have s3_2: "a * 0 = 0" - using adom mult_0 by auto - show ?thesis - using s3_1 s3_2 by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_3: "\ c. \ - c \\in Nat; - (a * b) * c = a * (b * c) - \ \ - (a * b) * Succ[c] = a * (b * Succ[c])" - proof - - fix "c" :: "c" - assume s1_3_cdom: "c \\in Nat" - assume s1_3_induct: "(a * b) * c = a * (b * c)" - have s2_1: "(a * b) * Succ[c] = (a * b) * c + (a * b)" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - have s3_2: "c \\in Nat" - using s1_3_cdom by auto - show ?thesis - using s3_1 s3_2 multint_succ_nat by auto - qed - also have s2_2: "... = a * (b * c) + (a * b)" - using s1_3_induct by auto - also have s2_3: "... = a * (b * c + b)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b * c \\in Int" - using bdom s1_3_cdom nat_in_int multIsInt by auto - have s3_3: "b \\in Int" - using bdom by auto - show ?thesis - using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] by auto - qed - also have s2_4: "... = a * (b * Succ[c])" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "c \\in Nat" - using s1_3_cdom by auto - have s3_3: "b * Succ[c] = b * c + b" - using s3_1 s3_2 multint_succ_nat by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a * b) * Succ[c] = a * (b * Succ[c])" . - qed - have s1_4: "\ c. \ - c \\in negNat; - (a * b) * c = a * (b * c) - \ \ - (a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" - proof - - fix "c" :: "c" - assume s1_4_cdom: "c \\in negNat" - assume s1_4_induct: "(a * b) * c = a * (b * c)" - have s2_1: "(a * b) * -.Succ[-.c] = (a * b) * -.-.c + -.(a * b)" - proof - - have s3_1: "a * b \\in Int" - using adom bdom multIsInt by auto - have s3_2: "-.c \\in Nat" - using s1_4_cdom minus_neg_nat_in_nat by auto - show ?thesis - using s3_1 s3_2 multint_succ_negnat by auto - qed - also have s2_2: "... = (a * b) * c + -.(a * b)" - proof - - have s3_1: "c \\in Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "-.-.c = c" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_3: "... = a * (b * c) + -.(a * b)" - using s1_4_induct by auto - also have s2_4: "... = a * (b * c) + a * -.b" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b \\in Int" - using bdom by auto - have s3_3: "-.(a * b) = a * -.b" - using s3_1 s3_2 MinusCommutesRightMult by auto - show ?thesis - using s3_3 by auto - qed - also have s2_5: "... = a * (b * c + -.b)" - proof - - have s3_1: "a \\in Int" - using adom by auto - have s3_2: "b * c \\in Int" - using bdom s1_4_cdom neg_nat_in_int multIsInt by auto - have s3_3: "-.b \\in Int" - using bdom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 s3_3 MultLeftDistributesAdd[symmetric] - by auto - qed - also have s2_6: "... = a * (b * -.-.c + -.b)" - proof - - have s3_1: "c \\in Int" - using s1_4_cdom neg_nat_in_int by auto - have s3_2: "-.-.c = c" - using s3_1 minus_involutive by auto - show ?thesis - using s3_2 by auto - qed - also have s2_7: "... = a * (b * -.Succ[-.c])" - proof - - have s3_1: "b \\in Int" - using bdom by auto - have s3_2: "-.c \\in Nat" - using s1_4_cdom minus_neg_nat_in_nat by auto - have s3_3: "b * -.Succ[-.c] = b * -.-.c + -.b" - using s3_1 s3_2 multint_succ_negnat by auto - show ?thesis - using s3_3 by auto - qed - finally show "(a * b) * -.Succ[-.c] = a * (b * -.Succ[-.c])" . - qed - have s1_5: "\ c \ Nat: - (a * b) * c = a * (b * c)" - using s1_2 s1_3 natInduct[of "\ c. - (a * b) * c = a * (b * c)"] - by auto - have s1_6: "\ c \ negNat: - (a * b) * c = a * (b * c)" - using s1_2 s1_4 neg_nat_induction[of "\ c. - (a * b) * c = a * (b * c)"] - by auto - have s1_7: "\ c \ Int: - (a * b) * c = a * (b * c)" - using s1_5 s1_6 int_union_nat_negnat by auto - } - from this show ?thesis - by auto - qed - - - -(*----------------------------------------------------------------------------*) -(* Properties of `leq` and `add` and `diff`. *) - - -theorem LeqIsBool: - "(m <= n) \\in BOOLEAN" - unfolding leq_def by auto - - -theorem LeqReflexive: - assumes ndom: "n \\in Int" - shows "n <= n" - proof - - have s1_1: "n + 0 = n" - using ndom add_0 by auto - have s1_2: "0 \\in Nat" - using zeroIsNat by auto - have s1_3: "\ r \ Nat: n + r = n" - using s1_1 s1_2 by auto - show ?thesis - unfolding leq_def - using s1_3 by auto - qed - - -theorem LeqTransitive: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m <= n" and nk: "n <= k" - shows "m <= k" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - (* PICK t \in Nat: n + t = k *) - define Q where "Q \ \ x. x \\in Nat \ n + x = k" - define t where "t \ CHOOSE x: Q(x)" - have s1_2: "t \\in Nat \ n + t = k" - proof - - have s2_1: "\\E x: Q(x)" - using nk by (auto simp: leq_def Q_def) - have s2_2: "Q(t)" - unfolding t_def - using s2_1 chooseI_ex[of Q] by auto - show ?thesis - using s2_2 by (auto simp: Q_def) - qed - have s1_3: "(m + r) + t = k" - using s1_1 s1_2 by auto - have s1_4: "m + (r + t) = k" - proof - - have s2_1: "m \\in Int" - using mdom by auto - have s2_2: "r \\in Int" - using s1_1 nat_in_int by auto - have s2_3: "t \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "(m + r) + t = m + (r + t)" - using s2_1 s2_2 s2_3 AddAssociative_sequent by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "r + t \\in Nat" - using s1_1 s1_2 nat_add_in_nat by auto - have s1_6: "\ q \ Nat: m + q = k" - using s1_4 s1_5 by auto - show ?thesis - unfolding leq_def - using s1_6 by auto - qed - - -theorem leq_0: - assumes ndom: "n \\in Int" - shows "(n \\in Nat) = (0 <= n)" - proof - - have s1_1: "n \\in Nat ==> 0 <= n" - proof - - assume s1_1_ndom: "n \\in Nat" - have s2_1: "0 + n = n" - using ndom add_0_left by auto - have s2_2: "\ r \ Nat: 0 + r = n" - using s2_1 s1_1_ndom by auto - show "0 <= n" - unfolding leq_def - using s2_2 by auto - qed - have s1_2: "0 <= n ==> n \\in Nat" - proof - - assume s1_2_leq: "0 <= n" - define P where "P \ \ x. x \\in Nat \ 0 + x = n" - define r where "r \ CHOOSE x: P(x)" - have s2_1: "r \\in Nat \ 0 + r = n" - proof - - have s3_1: "\\E x: P(x)" - using s1_2_leq by (auto simp: leq_def P_def) - have s3_2: "P(r)" - unfolding r_def - using s3_1 chooseI_ex[of P] by auto - show ?thesis - using s3_2 by (auto simp: P_def) - qed - have s2_2: "0 + r = r" - proof - - have s3_1: "r \\in Int" - using s2_1 nat_in_int by auto - show ?thesis - using s3_1 add_0_left by auto - qed - have s2_3: "r = n" - using s2_1 s2_2 by auto - show "n \\in Nat" - using s2_1 s2_3 by auto - qed - have s1_3: "(0 <= n) \\in BOOLEAN" - using LeqIsBool by auto - have s1_4: "(n \\in Nat) \\in BOOLEAN" - using inIsBool isBoolTrueFalse by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -theorem eq_leq_both: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m = n" - shows "m <= n \ n <= m" - proof - - have s1_1: "m <= n" - proof - - have s2_1: "m + 0 = m" - using mdom add_0 by auto - have s2_2: "m + 0 = n" - using s2_1 mn by auto - have s2_3: "0 \\in Nat" - using zeroIsNat by auto - have s2_4: "\ r \ Nat: m + r = n" - using s2_2 s2_3 by auto - show ?thesis - unfolding leq_def - using s2_4 by auto - qed - have s1_2: "n <= m" - proof - - have s2_1: "n + 0 = n" - using ndom add_0 by auto - have s2_2: "n + 0 = m" - using s2_1 mn by auto - have s2_3: "0 \\in Nat" - using zeroIsNat by auto - have s2_4: "\ r \ Nat: n + r = m" - using s2_2 s2_3 by auto - show ?thesis - unfolding leq_def - using s2_4 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* The inverse of the theorem `eq_leq_both`. *) -theorem leq_both_eq: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m <= n" and nm: "n <= m" - shows "m = n" - proof - - define P where "P \ \ x. x \\in Nat \ m + x = n" - define p where "p \ CHOOSE x: P(x)" - define Q where "Q \ \ x. x \\in Nat \ n + x = m" - define q where "q \ CHOOSE x: Q(x)" - have s1_1: "p \\in Nat \ m + p = n" - proof - - have s2_1: "\ x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(p)" - unfolding p_def using s2_1 chooseI_ex[of P] - by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "q \\in Nat \ n + q = m" - proof - - have s2_1: "\ x: Q(x)" - using nm by (auto simp: leq_def Q_def) - have s2_2: "Q(q)" - unfolding q_def using s2_1 chooseI_ex[of Q] - by auto - show ?thesis - using s2_2 by (auto simp: Q_def) - qed - have s1_3: "(m + p) + (n + q) = n + m" - using s1_1 s1_2 by auto - have s1_4: "p + q = 0" - proof - - have s2_1: "(m + p) + (n + q) = n + m" - using s1_3 by auto - have s2_2: "m + (p + (n + q)) = n + m" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "p \\in Int" - using s1_1 nat_in_int by auto - have s3_3: "n + q \\in Int" - using ndom s1_2 nat_in_int addIsInt by auto - have s3_4: "(m + p) + (n + q) = - m + (p + (n + q))" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - have s2_3: "m + (p + (q + n)) = n + m" - proof - - have s3_1: "n \\in Int" - using ndom by auto - have s3_2: "q \\in Int" - using s1_2 nat_in_int by auto - have s3_3: "n + q = q + n" - using s3_1 s3_2 AddCommutative_sequent - by auto - show ?thesis - using s2_2 s3_3 by auto - qed - have s2_4: "m + ((p + q) + n) = n + m" - proof - - have s3_1: "p \\in Int" - using s1_1 nat_in_int by auto - have s3_2: "q \\in Int" - using s1_2 nat_in_int by auto - have s3_3: "n \\in Int" - using ndom by auto - have s3_4: "p + (q + n) = (p + q) + n" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_3 s3_4 by auto - qed - have s2_5: "m + (n + (p + q)) = n + m" - proof - - have s3_1: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "(p + q) + n = n + (p + q)" - using s3_1 s3_2 - AddCommutative_sequent by auto - show ?thesis - using s2_4 s3_3 by auto - qed - have s2_6: "(m + n) + (p + q) = n + m" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_4: "m + (n + (p + q)) = (m + n) + (p + q)" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_5 s3_4 by auto - qed - have s2_7: "-.(m + n) + ((m + n) + (p + q)) = - -.(m + n) + (n + m)" - using s2_6 by auto - have s2_8: "(-.(m + n) + (m + n)) + (p + q) = - -.(m + n) + (n + m)" - proof - - have s3_1: "m + n \\in Int" - using mdom ndom addIsInt by auto - have s3_2: "-.(m + n) \\in Int" - using s3_1 neg_int_eq_int by auto - have s3_3: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - have s3_4: "-.(m + n) + ((m + n) + (p + q)) = - (-.(m + n) + (m + n)) + (p + q)" - using s3_1 s3_2 s3_3 - AddAssociative_sequent by auto - show ?thesis - using s2_7 s3_4 by auto - qed - have s2_9: "(-.(m + n) + (m + n)) + (p + q) = - -.(m + n) + (m + n)" - proof - - have s3_1: "m \\in Int" - using mdom by auto - have s3_2: "n \\in Int" - using ndom by auto - have s3_3: "n + m = m + n" - using s3_1 s3_2 AddCommutative_sequent by auto - show ?thesis - using s2_8 s3_3 by auto - qed - have s2_10: "0 + (p + q) = 0" - proof - - have s3_1: "-.(m + n) + (m + n) = 0" - proof - - have s4_1: "m + n \\in Int" - using mdom ndom addIsInt by auto - show ?thesis - using s4_1 AddNegCancels_left by auto - qed - show ?thesis - using s2_9 s3_1 by auto - qed - have s2_11: "p + q = 0" - proof - - have s3_1: "p + q \\in Int" - using s1_1 s1_2 nat_in_int addIsInt by auto - show ?thesis - using s2_10 s3_1 add_0_left by auto - qed - show ?thesis - using s2_11 by auto - qed - have s1_5: "p = 0 \ q = 0" - using s1_1 s1_2 s1_4 nat_add_0 by auto - have s1_6: "m + 0 = n" - using s1_1 s1_5 by auto - show ?thesis - using s1_6 mdom add_0 by auto - qed - - -theorem eq_iff_leq_both: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "(m = n) = ((m <= n) \ (n <= m))" - using mdom ndom eq_leq_both leq_both_eq by auto - - -(* -THEOREM dichotomy_leq == - ASSUME NEW a \in Int, NEW b \in Int - PROVE (a <= b) \ (b <= a) -PROOF -<1>1. CASE (a - b) \in Nat - <2>1. b + (a - b) = a - BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b - <2>2. \E n \in Nat: b + n = a - BY <1>1, <2>1 - <2>3. b <= a - BY <2>2 DEF leq - <2> QED - BY <2>3 -<1>2. CASE (a - b) \notin Nat - <2>1. -.(a - b) \in Nat - <3>1. (a - b) \in Int - <4>1. a - b = a + -.b - BY DEF diff - <4>2. a \in Int - OBVIOUS - <4>3. -.b \in Int - BY b \in Int, neg_int_eq_int - <4>4. a + -.b \in Int - BY <4>2, <4>3, addIsInt - <4> QED - BY <4>1, <4>4 - <3> QED - BY <3>1, <1>2, neg_negative_in_nat - <2>2. (b - a) \in Nat - <3>1. -.(a - b) = b - a - BY a \in Int, b \in Int, MinusDiff - <3> QED - BY <2>1, <3>1 - <2>3. a + (b - a) = b - BY a \in Int, b \in Int, a_plus_b_minus_a_eq_b - <2>4. \E n \in Nat: a + n = b - BY <2>2, <2>3 - <2>5. a <= b - BY <2>4 DEF leq - <2> QED - BY <2>5 -<1> QED - BY <1>1, <1>2 -*) -theorem dichotomy_leq: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(a <= b) \ (b <= a)" - proof - - have s1_1: "a - b \\in Nat ==> - (a <= b) \ (b <= a)" - proof - - assume s1_1_asm: "a - b \\in Nat" - have s2_1: "b + (a - b) = a" - using adom bdom a_plus_b_minus_a_eq_b by auto - have s2_2: "\ n \ Nat: b + n = a" - using s1_1_asm s2_1 by auto - have s2_3: "b <= a" - unfolding leq_def - using s2_2 by auto - show "(a <= b) \ (b <= a)" - using s2_3 by auto - qed - have s1_2: "a - b \\notin Nat ==> - (a <= b) \ (b <= a)" - proof - - assume s1_2_asm: "a - b \\notin Nat" - have s2_1: "-.(a - b) \\in Nat" - proof - - have s3_1: "(a - b) \\in Int" - proof - - have s4_1: "a - b = a + -.b" - unfolding diff_def by auto - have s4_2: "a \\in Int" - using adom by auto - have s4_3: "-.b \\in Int" - using bdom neg_int_eq_int by auto - have s4_4: "a + -.b \\in Int" - using s4_2 s4_3 addIsInt by auto - show ?thesis - using s4_1 s4_4 by auto - qed - show ?thesis - using s3_1 s1_2_asm neg_negative_in_nat by auto - qed - have s2_2: "(b - a) \\in Nat" - proof - - have s3_1: "-.(a - b) = b - a" - using adom bdom MinusDiff by auto - show ?thesis - using s2_1 s3_1 by auto - qed - have s2_3: "a + (b - a) = b" - using adom bdom a_plus_b_minus_a_eq_b by auto - have s2_4: "\ n \ Nat: a + n = b" - using s2_2 s2_3 by auto - have s2_5: "a <= b" - unfolding leq_def - using s2_4 by auto - show "(a <= b) \ (b <= a)" - using s2_5 by auto - qed - show ?thesis - using s1_1 s1_2 by auto - qed - - -theorem trichotomy_less: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(a < b) \ (a = b) \ (b < a)" - using adom bdom dichotomy_leq by (auto simp: less_def) - - -theorem trichotomy_less_0: - assumes ndom: "n \\in Int" - shows "(n < 0) \ (n = 0) \ (0 < n)" - using ndom zero_in_int trichotomy_less by auto - - -(* -THEOREM leq_mult_monotonic == - \A m \in Int: \A n \in Int: - (m <= n) => (\A k \in Nat: k * m <= k * n) -PROOF -<1>1. SUFFICES - ASSUME NEW m \in Int, NEW n \in Int, NEW k \in Nat, - m <= n - PROVE k * m <= k * n - BY <1>1 -<1>2. PICK r \in Nat: m + r = n - BY <1>1 DEF leq -<1>3. k * (m + r) = k * n - BY <1>2 -<1>4. k * m + k * r = k * n - <2>1. k \in Int - BY <1>1, nat_in_int - <2>2. m \in Int - BY <1>1 - <2>3. r \in Int - BY <1>2, nat_in_int - <2>4. k * (m + r) = k * m + k * r - BY <2>1, <2>2, <2>3, MultLeftDistributesAdd - <2> QED - BY <1>3, <2>4 -<1>5. k * r \in Nat - BY <1>1, <1>2, nat_mult_in_nat -<1>6. \E x \in Nat: k * m + x = k * n - BY <1>4, <1>5 -<1> QED - BY <1>6 DEF leq -*) -theorem leq_mult_monotonic: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Nat" and mn: "m <= n" - shows "k * m <= k * n" - proof - - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_2: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_3: "k * (m + r) = k * n" - using s1_2 by auto - have s1_4: "k * m + k * r = k * n" - proof - - have s2_1: "k \\in Int" - using kdom nat_in_int by auto - have s2_2: "m \\in Int" - using mdom by auto - have s2_3: "r \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "k * (m + r) = k * m + k * r" - using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "k * r \\in Nat" - using kdom s1_2 nat_mult_in_nat by auto - have s1_6: "\ x \ Nat: k * m + x = k * n" - using s1_4 s1_5 by auto - show ?thesis - unfolding leq_def - using s1_6 by auto - qed - - -(* -THEOREM leq_mult_monotonic_neg: - \A m \in Int: \A n \in Int: - (m <= n) => (\A k \in negNat: k * n <= k * m) -PROOF -<1>1. SUFFICES - ASSUME NEW m \in Int, NEW n \in Int, NEW k \in negNat, - m <= n - PROVE k * n <= k * m - BY <1>1 -<1>2. PICK r \in Nat: m + r = n - BY <1>1 DEF leq -<1>3. k * (m + r) = k * n - BY <1>2 -<1>4. k * m + k * r = k * n - <2>1. k \in Int - BY <1>1, neg_nat_in_int - <2>2. m \in Int - BY <1>1 - <2>3. r \in Int - BY <1>2, nat_in_int - <2>4. k * (m + r) = k * m + k * r - BY <2>1, <2>2, <2>3, MultLeftDistributesAdd - <2> QED - BY <1>3, <2>4 -<1>5. k * m = k * n + -.(k * r) - <2>1. (k * m + k * r) + -.(k * r) = k * n + -.(k * r) - BY <1>4 - <2>2. k * m + (k * r + -.(k * r)) = k * n + -.(k * r) - <3>1. k * m \in Int - BY <1>1, neg_nat_in_int, multIsInt - <3>2. k * r \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. r \in Int - BY <1>2, nat_in_int - <4> QED - BY <4>1, <4>2, multIsInt - <3>3. -.(k * r) \in Int - BY <3>2, neg_int_eq_int - <3>4. (k * m + k * r) + -.(k * r) = - k * m + (k * r + -.(k * r)) - BY <3>1, <3>2, <3>3, AddAssociative - <3> QED - BY <2>1, <3>4 - <2>3. k * r + -.(k * r) = 0 - <3>1. k * r \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. r \in Int - BY <1>2, nat_in_int - <4> QED - BY <4>1, <4>2, multIsInt - <3> QED - BY <3>1, AddNegCancels - <2>4. k * m + 0 = k * n + -.(k * r) - BY <2>2, <2>3 - <2>5. k * m + 0 = k * m - <3>1. k * m \in Int - <4>1. k \in Int - BY <1>1, neg_nat_in_int - <4>2. m \in Int - BY <1>1 - <4> QED - BY <4>1, <4>2, multIsInt - <3> QED - BY <3>1, add_0 - <2> QED - BY <2>4, <2>5 -<1>6. k * n + (-.k * r) = k * m - <2>1. k * n + -.(k * r) = k * m - BY <1>5 \* symmetric - <2>2. -.(k * r) = -.k * r - <3>1. k \in Int - BY <1>1, neg_nat_in_int - <3>2. r \in Int - BY <1>2, nat_in_int - <3> QED - BY <3>1, <3>2, MinusCommutesLeftMult - <2> QED - BY <2>1, <2>2 -<1>7. -.k * r \in Nat - <2>1. -.k \in Nat - BY <1>1, minus_neg_nat_in_nat - <2>2. r \in Nat - BY <1>2 - <2> QED - BY <2>1, <2>2, nat_mult_in_nat -<1>8. \E x \in Nat: k * n + x = k * m - BY <1>6, <1>7 -<1> QED - BY <1>8 DEF leq -*) -theorem leq_mult_monotonic_neg: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in negNat" and mn: "m <= n" - shows "k * n <= k * m" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_2: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_3: "k * (m + r) = k * n" - using s1_2 by auto - have s1_4: "k * m + k * r = k * n" - proof - - have s2_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s2_2: "m \\in Int" - using mdom by auto - have s2_3: "r \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "k * (m + r) = k * m + k * r" - using s2_1 s2_2 s2_3 MultLeftDistributesAdd by auto - show ?thesis - using s1_3 s2_4 by auto - qed - have s1_5: "k * m = k * n + -.(k * r)" - proof - - have s2_1: "(k * m + k * r) + -.(k * r) = k * n + -.(k * r)" - using s1_4 by auto - have s2_2: "k * m + (k * r + -.(k * r)) = k * n + -.(k * r)" - proof - - have s3_1: "k * m \\in Int" - using mdom kdom neg_nat_in_int multIsInt by auto - have s3_2: "k * r \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - have s3_3: "-.(k * r) \\in Int" - using s3_2 neg_int_eq_int by auto - have s3_4: "(k * m + k * r) + -.(k * r) = - k * m + (k * r + -.(k * r))" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - have s2_3: "k * r + -.(k * r) = 0" - proof - - have s3_1: "k * r \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - show ?thesis - using s3_1 AddNegCancels by auto - qed - have s2_4: "k * m + 0 = k * n + -.(k * r)" - using s2_2 s2_3 by auto - have s2_5: "k * m + 0 = k * m" - proof - - have s3_1: "k * m \\in Int" - proof - - have s4_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s4_2: "m \\in Int" - using mdom by auto - show ?thesis - using s4_1 s4_2 multIsInt by auto - qed - show ?thesis - using s3_1 add_0 by auto - qed - show ?thesis - using s2_4 s2_5 by auto - qed - have s1_6: "k * n + (-.k * r) = k * m" - proof - - have s2_1: "k * n + -.(k * r) = k * m" - using s1_5 by auto - have s2_2: "-.(k * r) = -.k * r" - proof - - have s3_1: "k \\in Int" - using kdom neg_nat_in_int by auto - have s3_2: "r \\in Int" - using s1_2 nat_in_int by auto - show ?thesis - using s3_1 s3_2 MinusCommutesLeftMult by auto - qed - show ?thesis - using s2_1 s2_2 by auto - qed - have s1_7: "-.k * r \\in Nat" - proof - - have s2_1: "-.k \\in Nat" - using kdom minus_neg_nat_in_nat by auto - have s2_2: "r \\in Nat" - using s1_2 by auto - show ?thesis - using s2_1 s2_2 nat_mult_in_nat by auto - qed - have s1_8: "\ x \ Nat: k * n + x = k * m" - using s1_6 s1_7 by auto - show ?thesis - unfolding leq_def - using s1_8 by auto - qed - - -(*----------------------------------------------------------------------------*) - - -(* Monotonicity of addition with respect to order. *) -theorem leq_add_monotonic_right: - assumes mdom: "m \ Int" and ndom: "n \ Int" and - kdom: "k \ Int" and mn: "m <= n" - shows "m + k <= n + k" - proof - - have s1_1: "\ r: r \ Nat \ m + r = n" - using mn by (auto simp: leq_def) - define r where "r \ CHOOSE r: r \ Nat \ m + r = n" - have s1_2: "r \ Nat \ m + r = n" - unfolding r_def - by (rule chooseI_ex, rule s1_1) - have s1_3: "(m + r) + k = n + k" - using s1_2 by auto - have s1_4: "(m + k) + r = n + k" - proof - - have s2_1: "(r + m) + k = n + k" - proof - - have s3_1: "m \ Int" - using mdom by auto - have s3_2: "r \ Int" - using s1_2 nat_in_int by auto - have s3_3: "m + r = r + m" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s1_3 s3_3 by auto - qed - have s2_2: "r + (m + k) = n + k" - proof - - have s3_1: "r \ Int" - using s1_2 nat_in_int by auto - have s3_2: "m \ Int" - using mdom by auto - have s3_3: "k \ Int" - using kdom by auto - have s3_4: "(r + m) + k = r + (m + k)" - using s3_1 s3_2 s3_3 AddAssociative_sequent by auto - show ?thesis - using s2_1 s3_4 by auto - qed - show ?thesis - proof - - have s3_1: "r \ Int" - using s1_2 nat_in_int by auto - have s3_2: "m + k \ Int" - using mdom kdom addIsInt by auto - have s3_3: "r + (m + k) = (m + k) + r" - using s3_1 s3_2 AddCommutative by auto - show ?thesis - using s2_2 s3_3 by auto - qed - qed - have s1_5: "\ t \ Nat: - (m + k) + t = n + k" - using s1_2 s1_4 by auto - show "m + k <= n + k" - using s1_5 by (auto simp: leq_def) - qed - - -theorem eq_add_inj_right: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m + k = n + k" - shows "m = n" - proof - - have s1_1: "(m + k) + -.k = (n + k) + -.k" - using mn by auto - have s1_2: "m + (k + -.k) = n + (k + -.k)" - proof - - have minus_kdom: "-.k \\in Int" - using kdom neg_int_eq_int by auto - have s2_1: "(m + k) + -.k = m + (k + -.k)" - using mdom kdom minus_kdom AddAssociative_sequent - by auto - have s2_2: "(n + k) + -.k = n + (k + -.k)" - using ndom kdom minus_kdom AddAssociative_sequent - by auto - show ?thesis - using s1_1 s2_1 s2_2 by auto - qed - have s1_3: "m + 0 = n + 0" - proof - - have s2_1: "k + -.k = 0" - using kdom AddNegCancels_sequent by auto - show ?thesis - using s1_2 s2_1 by auto - qed - have s1_4: "m + 0 = m" - using mdom add_0 by auto - have s1_5: "n + 0 = n" - using ndom add_0 by auto - show ?thesis - using s1_3 s1_4 s1_5 by auto - qed - - -theorem less_add_monotonic_right: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m < n" - shows "m + k < n + k" - proof - - have s1_1: "m + k <= n + k" - proof - - have s2_1: "m <= n" - using mn by (auto simp: less_def) - show ?thesis - using mdom ndom kdom s2_1 leq_add_monotonic_right by auto - qed - have s1_2: "m + k \ n + k" - proof - - have s2_1: "m + k = n + k ==> FALSE" - proof - - assume s2_1_asm: "m + k = n + k" - have s3_1: "m = n" - using s2_1_asm mdom ndom kdom eq_add_inj_right - by auto - have s3_2: "m \ n" - using mn by (auto simp: less_def) - show "FALSE" - using s3_1 s3_2 by auto - qed - show ?thesis - using s2_1 by auto - qed - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_add_monotonic_left: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - kdom: "k \\in Int" and mn: "m < n" - shows "k + m < k + n" - proof - - have s1_1: "m + k < n + k" - using mdom ndom kdom mn less_add_monotonic_right by auto - have s1_2: "m + k = k + m" - using mdom kdom AddCommutative_sequent by auto - have s1_3: "n + k = k + n" - using ndom kdom AddCommutative_sequent by auto - show ?thesis - using s1_1 s1_2 s1_3 by auto - qed - - -theorem minus_less: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m < n" - shows "-.n < -.m" - proof - - have s1_1: "m + -.m < n + -.m" - proof - - have s2_1: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using mdom s2_1 ndom mn less_add_monotonic_right by auto - qed - have s1_2: "0 < n + -.m" - proof - - have s2_1: "m + -.m = 0" - using mdom AddNegCancels_sequent by auto - show ?thesis - using s1_1 s2_1 by auto - qed - have s1_3: "-.n + 0 < -.n + (n + -.m)" - proof - - have s2_1: "-.n \\in Int" - using ndom neg_int_eq_int by auto - have s2_2: "0 \\in Int" - using zeroIsNat nat_in_int by auto - have s2_3: "n + -.m \\in Int" - proof - - have s3_1: "n \\in Int" - using ndom by auto - have s3_2: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s3_1 s3_2 addIsInt by auto - qed - show ?thesis - using s1_2 s2_1 s2_2 s2_3 less_add_monotonic_left - by auto - qed - have s1_4: "-.n < -.n + (n + -.m)" - proof - - have s2_1: "-.n + 0 = -.n" - using ndom neg_int_eq_int add_0 by auto - show ?thesis - using s1_3 s2_1 by auto - qed - have s1_5: "-.n + (n + -.m) = -.m" - proof - - have s2_1: "-.n + (n + -.m) = (-.n + n) + -.m" - using ndom mdom neg_int_eq_int AddAssociative_sequent - by auto - have s2_2: "-.n + n = 0" - using ndom AddNegCancels_left by auto - have s2_3: "(-.n + n) + -.m = 0 + -.m" - using s2_2 by auto - have s2_4: "0 + -.m = -.m" - proof - - have s3_1: "-.m \\in Int" - using mdom neg_int_eq_int by auto - show ?thesis - using s3_1 add_0_left by auto - qed - show ?thesis - using s2_1 s2_3 s2_4 by auto - qed - show ?thesis - using s1_4 s1_5 by auto - qed - - -theorem leq_diff_add: - assumes ndom: "n \ Int" and mdom: "m \ Int" and - kdom: "k \ Int" and nmk: "n - m <= k" - shows "n <= k + m" - proof - - define P where "P \ \ x. x \\in Nat \ (n - m) + x = k" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ (n - m) + r = k" - proof - - have s2_1: "\\E x: P(x)" - using nmk by (auto simp: leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "m + ((n - m) + r) = m + k" - using s1_1 by auto - have s1_3: "(m + -.m) + (n + r) = k + m" - proof - - have s2_1: "m + ((n - m) + r) = (m + -.m) + (n + r)" - proof - - have s3_1: "m + ((n - m) + r) = - m + ((n + -.m) + r)" - unfolding diff_def by auto - also have s3_2: "... = m + ((-.m + n) + r)" - using mdom neg_int_eq_int ndom - AddCommutative_sequent by auto - also have s3_3: "... = (m + (-.m + n)) + r" - using mdom neg_int_eq_int ndom s1_1 nat_in_int - addIsInt AddAssociative_sequent by auto - also have s3_4: "... = ((m + -.m) + n) + r" - using mdom neg_int_eq_int ndom - AddAssociative_sequent by auto - also have s3_5: "... = (m + -.m) + (n + r)" - using mdom neg_int_eq_int addIsInt ndom - s1_1 nat_in_int AddAssociative_sequent by auto - finally show ?thesis . - qed - have s2_2: "m + k = k + m" - using mdom kdom AddCommutative_sequent by auto - show ?thesis - using s1_2 s2_1 s2_2 by auto - qed - have s1_4: "n + r = k + m" - proof - - have s2_1: "m + -.m = 0" - using mdom AddNegCancels_sequent by auto - have s2_2: "(m + -.m) + (n + r) = n + r" - proof - - have s3_1: "(m + -.m) + (n + r) = 0 + (n + r)" - using s2_1 by auto - have s3_2: "0 + (n + r) = n + r" - proof - - have s4_1: "n + r \\in Int" - using ndom s1_1 nat_in_int addIsInt by auto - show ?thesis - using s4_1 add_0_left by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - show ?thesis - using s1_3 s2_2 by auto - qed - have s1_5: "\ x \ Nat: n + x = k + m" - using s1_1 s1_4 by auto - show ?thesis - unfolding leq_def - using s1_5 by auto - qed - - -theorem leq_diff_nat: - assumes mdom: "m \\in Int" and ndom: "n \\in Nat" - shows "m - n <= m" - proof - - have s1_1: "(m - n) + n = m" - proof - - have s2_1: "-.n + n = 0" - proof - - have s3_1: "n \\in Int" - using ndom nat_in_int by auto - show ?thesis - using s3_1 AddNegCancels_left by auto - qed - have s2_2: "m + (-.n + n) = m + 0" - using s2_1 by auto - have s2_3: "(m + -.n) + n = m" - proof - - have s3_1: "m + (-.n + n) = (m + -.n) + n" - proof - - have s4_1: "m \\in Int" - using mdom by auto - have s4_2: "-.n \\in Int" - using ndom minus_nat_in_int by auto - have s4_3: "n \\in Int" - using ndom nat_in_int by auto - show ?thesis - using s4_1 s4_2 s4_3 AddAssociative_sequent by auto - qed - have s3_2: "m + 0 = m" - using mdom add_0 by auto - show ?thesis - using s2_2 s3_1 s3_2 by auto - qed - show ?thesis - unfolding diff_def - using s2_3 by auto - qed - have s1_2: "\ r \ Nat: (m - n) + r = m" - using ndom s1_1 by auto - show ?thesis - unfolding leq_def - using s1_2 by auto - qed - - -(* See also the theorem `less_leq_trans`. *) -theorem leq_less_trans: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - and kdom: "k \\in Int" and - mn: "m <= n" and nk: "n < k" - shows "m < k" - proof - - have s1_1: "(n <= k) \ (n \ k)" - using nk by (auto simp: less_def) - have s1_2: "m <= k" - proof - - have s2_2: "n <= k" - using s1_1 by auto - have s2_3: "m <= n" - using mn by auto - show ?thesis - using s2_2 s2_3 mdom ndom kdom LeqTransitive - by iprover - qed - have s1_3: "m \ k" - proof - - { - assume s3_1: "m = k" - have s3_2: "k <= n" - using mn s3_1 by auto - have s3_3: "n <= k" - using nk by (auto simp: less_def) - have s3_4: "n = k" - using s3_2 s3_3 ndom kdom leq_both_eq - by auto - have "FALSE" - using s3_4 s1_1 by auto - } - from this show "m \ k" - by auto - qed - show ?thesis - unfolding less_def - using s1_2 s1_3 by auto - qed - - -(* See also the theorem `leq_less_trans`. *) -theorem less_leq_trans: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - and kdom: "k \\in Int" and - mn: "m < n" and nk: "n <= k" - shows "m < k" - proof - - have s1_1: "(m <= n) \ (m \ n)" - using mn by (auto simp: less_def) - have s1_2: "m <= k" - proof - - have s2_1: "m <= n" - using s1_1 by auto - have s2_2: "n <= k" - using nk by auto - show ?thesis - using s2_1 s2_2 mdom ndom kdom LeqTransitive - by iprover - qed - have s1_3: "m \ k" - proof - - have s2_1: "m = k ==> FALSE" - proof - - assume s2_1_asm: "m = k" - have s3_1: "n <= m" - using nk s2_1_asm by auto - have s3_2: "m <= n" - using mn by (auto simp: less_def) - have s3_3: "m = n" - using s3_1 s3_2 mdom ndom leq_both_eq - by auto - show "FALSE" - using s3_3 s1_1 by auto - qed - show ?thesis - using s2_1 by auto - qed - show ?thesis - unfolding less_def - using s1_2 s1_3 by auto - qed - - -theorem less_not_leq: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" - shows "(m < n) = (~ (n <= m))" - proof - - have s1_1: "m < n ==> ~ (n <= m)" - proof - - assume mn: "m < n" - have s2_1: "(m <= n) \ (m \ n)" - using mn by (auto simp: less_def) - have s2_2: "n <= m ==> FALSE" - proof - - assume nm: "n <= m" - have s3_1: "m = n" - using mdom ndom s2_1 nm leq_both_eq by blast - show "FALSE" - using s2_1 s3_1 by auto - qed - show "~ (n <= m)" - using s2_2 by auto - qed - have s1_2: "(~ (n <= m)) ==> m < n" - proof - - assume nm: "~ (n <= m)" - have s2_1: "(m <= n) \ (n <= m)" - using mdom ndom dichotomy_leq by auto - have s2_2: "m <= n" - using nm s2_1 by auto - have s2_3: "m = n ==> FALSE" - proof - - assume eq: "m = n" - have s3_1: "n <= n" - using ndom LeqReflexive by auto - have s3_2: "n <= m" - using eq s3_1 by auto - show "FALSE" - using nm s3_2 by auto - qed - have s2_4: "m \ n" - using s2_3 by auto - show "m < n" - unfolding less_def - using s2_2 s2_4 by auto - qed - have s1_3: "(~ (n <= m)) \\in BOOLEAN" - by auto - have s1_4: "(m < n) \\in BOOLEAN" - unfolding less_def by auto - show ?thesis - using s1_1 s1_2 s1_3 s1_4 by auto - qed - - -theorem leq_linear: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(~ (a <= b)) \ (a = b) \ (~ (b <= a))" - using adom bdom trichotomy_less[of "a" "b"] - less_not_leq[of "a" "b"] - less_not_leq[of "b" "a"] by auto - - -theorem leq_geq_linear: - assumes adom: "a \\in Int" and bdom: "b \\in Int" - shows "(~ (b >= a)) \ (a = b) \ (~ (b <= a))" - using adom bdom leq_linear by auto - - -(* -THEOREM less_is_leq_plus_one == - ASSUME NEW m \in Int, NEW n \in Int, m < n - PROVE m + 1 <= n -PROOF -<1>1. PICK r \in Nat: m + r = n - BY m < n DEF less, leq -<1>2. PICK k \in Nat: r = Succ[k] - <2>1. r = 0 \/ \E k \in Nat: r = Succ[k] - BY <1>1, nat_0_succ - <2>2. ASSUME r = 0 - PROVE FALSE - <3>1. m + 0 = n - BY <1>1, <2>2 - <3>2. m = n - BY <3>1, m \in Int, add_0 - <3>3. m # n - BY m < n DEF less - <3> QED - BY <3>2, <3>3 - <2> QED - BY <2>1, <2>2 -<1>3. m + Succ[k] = n - BY <1>1, <1>2 -<1>4. m + (k + 1) = n - <2>1. Succ[k] = k + 1 - BY <1>2, nat_add_1 - <2> QED - BY <1>3, <2>1 -<1>5. m + (1 + k) = n - <2>1. k \in Int - BY <1>2, nat_in_int - <2>2. 1 \in Int - BY oneIsNat, nat_in_int - <2>3. k + 1 = 1 + k - BY <2>1, <2>2, AddCommutative - <2> QED - BY <1>4, <2>3 -<1>6. (m + 1) + k = n - <2>1. m \in Int - OBVIOUS - <2>2. 1 \in Int - BY oneIsNat nat_in_int - <2>3. k \in Int - BY <1>2, nat_in_int - <2>4. m + (1 + k) = (m + 1) + k - BY <2>1, <2>2, <2>3, AddAssociative - <2> QED - BY <1>5, <2>4 -<1>7. \E q \in Nat: (m + 1) + q = n - BY <1>2, <1>6 -<1> QED - BY <1>7 DEF leq -*) -theorem less_is_leq_plus_one: - assumes mdom: "m \\in Int" and ndom: "n \\in Int" and - mn: "m < n" - shows "m + 1 <= n" - proof - - (* PICK r \in Nat: m + r = n *) - define P where "P \ \ x. x \\in Nat \ m + x = n" - define r where "r \ CHOOSE x: P(x)" - have s1_1: "r \\in Nat \ m + r = n" - proof - - have s2_1: "\\E x: P(x)" - using mn by (auto simp: less_def leq_def P_def) - have s2_2: "P(r)" - unfolding r_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - (* PICK k \in Nat: r = Succ[k] *) - define Q where "Q \ \ x. x \\in Nat \ r = Succ[x]" - define k where "k \ CHOOSE x: Q(x)" - have s1_2: "k \\in Nat \ r = Succ[k]" - proof - - have s2_1: "r = 0 \ - (\ x \ Nat: r = Succ[x])" - using s1_1 nat_0_succ by auto - have s2_2: "r = 0 ==> FALSE" - proof - - assume s2_2_asm: "r = 0" - have s3_1: "m + 0 = n" - using s1_1 s2_2_asm by auto - have s3_2: "m = n" - using s3_1 mdom add_0 by auto - have s3_3: "m \ n" - using mn by (auto simp: less_def) - show ?thesis - using s3_2 s3_3 by auto - qed - have s2_3: "\ x \ Nat: r = Succ[x]" - using s2_1 s2_2 by auto - have s2_4: "\\E x: Q(x)" - unfolding Q_def using s2_3 by auto - have s2_5: "Q(k)" - unfolding k_def - using s2_4 chooseI_ex[of Q] by auto - show ?thesis - using s2_5 by (auto simp: Q_def) - qed - have s1_3: "m + Succ[k] = n" - using s1_1 s1_2 by auto - have s1_4: "m + (k + 1) = n" - proof - - have s2_1: "Succ[k] = k + 1" - using s1_2 nat_add_1 by fast - show ?thesis - using s1_3 s2_1 by auto - qed - have s1_5: "m + (1 + k) = n" - proof - - have s2_1: "k \\in Int" - using s1_2 nat_in_int by auto - have s2_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s2_3: "k + 1 = 1 + k" - using s2_1 s2_2 AddCommutative_sequent by auto - show ?thesis - using s1_4 s2_3 by auto - qed - have s1_6: "(m + 1) + k = n" - proof - - have s2_1: "m \\in Int" - using mdom by auto - have s2_2: "1 \\in Int" - using oneIsNat nat_in_int by auto - have s2_3: "k \\in Int" - using s1_2 nat_in_int by auto - have s2_4: "m + (1 + k) = (m + 1) + k" - using s2_1 s2_2 s2_3 AddAssociative_sequent by auto - show ?thesis - using s1_5 s2_4 by auto - qed - have s1_7: "\ q \ Nat: (m + 1) + q = n" - using s1_2 s1_6 by auto - show ?thesis - unfolding leq_def using s1_7 by auto - qed - - -(* -THEOREM less_succ_nat == - ASSUME NEW n \in Nat - PROVE n < Succ[n] -PROOF -<1>1. n <= Succ[n] - <2>1. Succ[n] = n + 1 - BY n \in Nat, nat_add_1 - <2>2. 1 \in Nat - BY oneIsNat - <2>3. \E r \in Nat: n + r = Succ[n] - BY <2>1, <2>2 - <2> QED - BY <2>3 DEF leq -<1>2. n # Succ[n] - BY n \in Nat, succIrrefl -<1> QED - BY <1>1, <1>2 DEF less -*) -theorem less_succ_nat: - assumes ndom: "n \\in Nat" - shows "n < Succ[n]" - proof - - have s1_1: "n <= Succ[n]" - proof - - have s2_1: "Succ[n] = n + 1" - using ndom nat_add_1 by fast - have s2_2: "1 \\in Nat" - using oneIsNat by auto - have s2_3: "\\E r \\in Nat: n + r = Succ[n]" - using s2_1 s2_2 by auto - show ?thesis - unfolding leq_def - using s2_3 by auto - qed - have s1_2: "n \ Succ[n]" - using ndom succIrrefl by auto - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_pred_nat: - assumes ndom: "n \\in Nat \ {0}" - shows "Pred[n] < n" - proof - - have s1_1: "Pred[n] <= n" - proof - - have s2_1: "Pred[n] \\in Nat" - using ndom Pred_in_nat by auto - have s2_2: "Succ[Pred[n]] = n" - using ndom Succ_Pred_nat by auto - have s2_3: "Pred[n] + 1 = n" - proof - - have s3_1: "Succ[Pred[n]] = Pred[n] + 1" - using s2_1 nat_add_1 by fast - show ?thesis - using s2_2 s3_1 by auto - qed - have s2_4: "1 \\in Nat" - using oneIsNat by auto - have s2_5: "\\E x \\in Nat: Pred[n] + x = n" - using s2_3 s2_4 by auto - show ?thesis - unfolding leq_def - using s2_5 by auto - qed - have s1_2: "Pred[n] \ n" - using ndom pred_irrefl by auto - show ?thesis - unfolding less_def - using s1_1 s1_2 by auto - qed - - -theorem less_isucc: - assumes idom: "i \\in Int" - shows "i < iSucc[i]" - proof - - have s1_1: "i \\in Nat ==> i < iSucc[i]" - proof - - assume s1_1_asm: "i \\in Nat" - have s2_1: "i < Succ[i]" - using s1_1_asm less_succ_nat by auto - have s2_2: "iSucc[i] = Succ[i]" - unfolding iSucc_def - using idom s1_1_asm by auto - show "i < iSucc[i]" - using s2_1 s2_2 by auto - qed - have s1_2: "i \\notin Nat ==> i < iSucc[i]" - proof - - assume s1_2_asm: "i \\notin Nat" - have s2_1: "iSucc[i] = -.Pred[-.i]" - unfolding iSucc_def - using idom s1_2_asm by auto - have s2_2: "-.i \\in Nat \ {0}" - using idom s1_2_asm minus_neg_int_in_nat by auto - have s2_3: "Pred[-.i] < -.i" - using s2_2 less_pred_nat by auto - have s2_4: "-.-.i < -.Pred[-.i]" - proof - - have s3_1: "Pred[-.i] \\in Int" - using s2_2 Pred_in_nat nat_in_int by auto - have s3_2: "-.i \\in Int" - using s2_2 nat_in_int by auto - show ?thesis - using s2_3 s3_1 s3_2 minus_less by auto - qed - have s2_5: "-.-.i = i" - using idom minus_involutive by auto - show "i < iSucc[i]" - using s2_1 s2_4 s2_5 by auto - qed - show ?thesis - using s1_1 s1_2 by iprover - qed - - -theorem less_i_add_1: - assumes idom: "i \\in Int" - shows "i < i + 1" - proof - - have s1_1: "i < iSucc[i]" - using idom less_isucc by auto - have s1_2: "iSucc[i] = i + 1" - using idom isucc_eq_add_1 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM i_less_j_nat == - ASSUME NEW i \in Int, NEW j \in Nat, i <= j - PROVE i < Succ[j] -PROOF -<1>1. j < Succ[j] - BY j \in Nat, less_succ_nat -<1>2. j \in Int - BY j \in Nat, nat_in_int -<1>3. Succ[j] \in Int - BY j \in Nat, succIsNat, nat_in_int -<1> QED - BY i \in Int, <1>2, <1>3, i <= j, <1>1, - leq_less_trans -*) -theorem i_less_succ_j_nat: - assumes idom: "i \\in Int" and jdom: "j \\in Nat" and - ij: "i <= j" - shows "i < Succ[j]" - proof - - have s1_1: "j < Succ[j]" - using jdom less_succ_nat by auto - have s1_2: "j \\in Int" - using jdom nat_in_int by auto - have s1_3: "Succ[j] \\in Int" - using jdom succIsNat nat_in_int by auto - show ?thesis - using idom s1_2 s1_3 ij s1_1 leq_less_trans by auto - qed - - -theorem less_add_1: - assumes idom: "i \\in Int" and jdom: "j \\in Int" and - ij: "i <= j" - shows "i < j + 1" - proof - - have s1_1: "j < j + 1" - using jdom less_i_add_1 by auto - have s1_2: "j + 1 \\in Int" - proof - - have s2_1: "1 \\in Int" - using oneIsNat nat_in_int by auto - show ?thesis - using jdom s2_1 addIsInt by auto - qed - show ?thesis - using idom jdom s1_2 ij s1_1 leq_less_trans by auto - qed - - -(*----------------------------------------------------------------------------*) -(* Intervals of integers. *) - - -(* -THEOREM SplitNat0 == - Nat = {0} \cup {n + 1: n \in Nat} -PROOF -<1>1. ASSUME NEW n \in Nat - PROVE n \in {0} \cup {k + 1: k \in Nat} - <2>1. n = 0 \/ \E m \in Nat: n = Succ[m] - BY <1>1, nat_0_succ - <2>2. (\E m \in Nat: n = Succ[m]) <=> - n \in {Succ[m]: m \in Nat} - OBVIOUS - <2>3. ASSUME NEW m \in Nat - PROVE Succ[m] = m + 1 - BY <2>3, nat_add_1 - <2>4. {Succ[m]: m \in Nat} = - {m + 1: m \in Nat} - BY <2>3 - <2>5. n \in {0} \/ n \in {m + 1: m \in Nat} - BY <2>1, <2>2, <2>4 - <2> QED - BY <2>5 -<1>2. ASSUME NEW n \in {0} \cup {k + 1: k \in Nat} - PROVE n \in Nat - <2>1. CASE n = 0 - BY <2>1, zeroIsNat - <2>2. CASE n \in {k + 1: k \in Nat} - <3>1. PICK k \in Nat: n = k + 1 - BY <2>2 - <3>2. k + 1 \in Nat - BY <3>1, oneIsNat, nat_add_in_nat - <3> QED - BY <3>1, <3>2 - <2> QED - BY <1>2, <2>1, <2>2 -<1> QED - BY <1>1, <1>2 -*) -theorem SplitNat0: - "Nat = {0} \ {n + 1: n \\in Nat}" - proof - - have s1_1: "\\A n \\in Nat: n \\in {0} \ {k + 1: k \\in Nat}" - proof - - { - fix "n" :: "c" - assume s1_1_asm: "n \\in Nat" - have s2_1: "n = 0 \ (\\E m \\in Nat: n = Succ[m])" - using s1_1_asm nat_0_succ by auto - have s2_2: "\\E m \\in Nat: n = Succ[m] ==> - n \\in {Succ[m]: m \\in Nat}" - by auto - have s2_3: "{Succ[m]: m \\in Nat} = {m + 1: m \\in Nat}" - proof - - have s3_1: "\\A y \\in {Succ[m]: m \\in Nat}: - y \\in {m + 1: m \\in Nat}" - proof - - { - fix "y" :: "c" - assume s4_1: "y \\in {Succ[m]: m \\in Nat}" - have s4_2: "\\E m \\in Nat: y = Succ[m]" - using s4_1 setOfAll by auto - have s4_3: "\\E m: m \\in Nat \ y = Succ[m]" - using s4_2 by (auto simp: bEx_def) - (* PICK r \in Nat: y = Succ[r] *) - define P where "P \ - \ x. x \\in Nat \ y = Succ[x]" - define r where "r \ CHOOSE x: P(x)" - have s4_4: "r \\in Nat \ y = Succ[r]" - proof - - have s5_1: "\\E x: P(x)" - using s4_3 by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_5: "r \\in Nat \ y = r + 1" - proof - - have s5_1: "r \\in Nat" - using s4_4 by auto - have s5_2: "Succ[r] = r + 1" - using s5_1 nat_add_1 by fast - have s5_3: "y = Succ[r]" - using s4_4 by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - have "y \\in {m + 1: m \\in Nat}" - using s4_5 setOfAll_eqI by auto - } - from this show ?thesis - by auto - qed - have s3_2: "\\A y \\in {m + 1: m \\in Nat}: - y \\in {Succ[m]: m \\in Nat}" - proof - - { - fix "y" :: "c" - assume s4_1: "y \\in {m + 1: m \\in Nat}" - have s4_2: "\\E m \\in Nat: y = m + 1" - using s4_1 setOfAll by auto - have s4_3: "\\E m: m \\in Nat \ y = m + 1" - using s4_2 by (auto simp: bEx_def) - (* PICK r \in Nat: y = r + 1 *) - define P where "P \ - \ x. x \\in Nat \ y = x + 1" - define r where "r \ CHOOSE x: P(x)" - have s4_4: "r \\in Nat \ y = r + 1" - proof - - have s5_1: "\\E x: P(x)" - using s4_3 by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_5: "r \\in Nat \ y = Succ[r]" - proof - - have s5_1: "r \\in Nat" - using s4_4 by auto - have s5_2: "Succ[r] = r + 1" - using s5_1 nat_add_1 by fast - have s5_3: "y = r + 1" - using s4_4 by auto - show ?thesis - using s5_1 s5_2 s5_3 by auto - qed - have "y \\in {Succ[m]: m \\in Nat}" - using s4_5 setOfAll_eqI by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s3_1 s3_2 extension - by (auto simp: subset_def) - qed - have s2_4: "n \\in {0} \ n \\in {m + 1: m \\in Nat}" - using s2_1 s2_2 s2_3 by auto - have s2_5: "n \\in {0} \ {k + 1: k \\in Nat}" - using s2_4 by auto - } - from this show ?thesis - by auto - qed - have s1_2: "\\A n \\in {0} \ {k + 1: k \\in Nat}: n \\in Nat" - proof - - { - fix "n" :: "c" - assume s1_2_asm: "n \\in {0} \ {k + 1: k \\in Nat}" - have s2_1: "n = 0 ==> n \\in Nat" - proof - - assume s2_1_asm: "n = 0" - show "n \\in Nat" - using s2_1_asm zeroIsNat by auto - qed - have s2_2: "n \\in {k + 1: k \\in Nat} ==> n \\in Nat" - proof - - assume s2_2_asm: "n \\in {k + 1: k \\in Nat}" - have s3_1: "\\E k \\in Nat: n = k + 1" - using s2_2_asm by auto - have s3_2: "\\E k: k \\in Nat \ n = k + 1" - using s3_1 by auto - (* PICK k \in Nat: n = k + 1 *) - define P where "P \ - \ x. x \\in Nat \ n = x + 1" - define k where "k \ CHOOSE x: P(x)" - have s3_3: "k \\in Nat \ n = k + 1" - proof - - have s4_1: "\\E x: P(x)" - using s3_2 by (auto simp: P_def) - have s4_2: "P(k)" - unfolding k_def - using s4_1 chooseI_ex[of P] by auto - show ?thesis - using s4_2 by (auto simp: P_def) - qed - have s3_4: "k + 1 \\in Nat" - using s3_3 oneIsNat nat_add_in_nat by auto - show "n \\in Nat" - using s3_3 s3_4 by auto - qed - have s2_3: "n \\in Nat" - using s1_2_asm s2_1 s2_2 by auto - } - from this show ?thesis - by auto - qed - show ?thesis - using s1_1 s1_2 extension - by blast - qed - - -(* -THEOREM NatIsAdd0 == - Nat = {k + 0: k \in Nat} -PROOF -<1>1. Nat = {k: k \in Nat} - OBVIOUS -<1>2. ASSUME NEW k \in Nat - PROVE k + 0 = k - BY <1>2, nat_in_int, add_0 -<1>3. {k: k \in Nat} = - {k + 0: k \in Nat} - BY <1>2 -<1> QED - BY <1>1, <1>3 -*) -theorem NatIsAdd0: - "Nat = {k + 0: k \\in Nat}" - proof - - have s1_1: "Nat = {k: k \\in Nat}" - by auto - have s1_2: "{k: k \\in Nat} = {k + 0: k \\in Nat}" - using nat_in_int add_0 by auto - show ?thesis - using s1_1 s1_2 by auto - qed - - -(* -THEOREM SplitAddi == - ASSUME NEW i \in Nat - PROVE {k + i: k \in Nat} = - {i} \cup {k + Succ[i]: k \in Nat} -PROOF -<1>1. n \in {k + i: k \in Nat} - = \E k \in Nat: n = k + i - OBVIOUS -<1>2. @ = \E k \in {0} \cup {q + 1: q \in Nat}: - n = k + i - BY SplitNat0 -<1>3. @ = \/ \E k \in {0}: n = k + i - \/ \E k \in {q + 1: q \in Nat}: n = k + i - OBVIOUS -<1>4. @ = \/ n = 0 + i - \/ \E k: /\ k \in {q + 1: q \in Nat} - /\ n = k + i - OBVIOUS -<1>5. @ = \/ n = i - \/ \E k: /\ \E q \in Nat: k = q + 1 - /\ n = k + i - BY i \in Nat, nat_in_int, add_0_left -<1>6. @ = \/ n = i - \/ \E k: \E q \in Nat: /\ k = q + 1 - /\ n = k + i - OBVIOUS -<1>7. @ = \/ n = i - \/ \E q \in Nat: n = (q + 1) + i - OBVIOUS -<1>8. @ = \/ n \in {i} - \/ n \in {q + (i + 1): q \in Nat} - BY AddCommutative, AddAssociative, - oneIsNat, nat_in_int -<1>9. @ = \/ n \in {i} - \/ n \in {q + Succ[i]: q \in Nat} - BY nat_add_1 -<1>10. @ = n \in {i} \cup {q + Succ[i]: q \in Nat} - OBVIOUS -<1> QED - BY <1>1, <1>2, <1>3, <1>4, <1>5, - <1>6, <1>7, <1>8, <1>9, <1>10 -*) -theorem SplitAddi: - assumes idom: "i \\in Nat" - shows "{k + i: k \\in Nat} = - ({i} \ {k + Succ[i]: k \\in Nat})" - proof - - { - fix "n" :: "c" - have s1_1: "n \\in {k + i: k \\in Nat} - = (\\E k \\in Nat: n = k + i)" - by auto - also have s1_2: "... = ( - \\E k \\in ({0} \ {q + 1: q \\in Nat}): - n = k + i)" - using SplitNat0 by auto - also have s1_3: "... = ( - (\\E k \\in {0}: n = k + i) \ - (\\E k \\in {q + 1: q \\in Nat}: n = k + i))" - by auto - also have s1_4: "... = ( - (n = 0 + i) \ - (\\E k: (k \\in {q + 1: q \\in Nat}) \ n = k + i))" - by auto - also have s1_5: "... = ( - (n = i) \ - (\\E k: (\\E q \\in Nat: k = q + 1) \ n = k + i))" - using idom nat_in_int add_0_left by auto - also have s1_6: "... = ( - (n = i) \ - (\\E k: \\E q \\in Nat: k = q + 1 \ n = k + i))" - by auto - also have s1_7: "... = ( - (n = i) \ - (\\E q \\in Nat: n = (q + 1) + i))" - by auto - also have s1_8: "... = ( - (n = i) \ - (\\E q \\in Nat: n = q + (i + 1)))" - using AddCommutative_sequent AddAssociative_sequent - oneIsNat idom nat_in_int by auto - also have s1_9: "... = ( - (n = i) \ - (\\E q \\in Nat: n = q + Succ[i]))" - using idom nat_add_1[symmetric] by force - also have s1_10: "... = ( - (n \\in {i}) \ - (n \\in {q + Succ[i]: q \\in Nat}))" - by auto - also have s1_11: "... = ( - n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - by auto - finally have s1_12: " - (n \\in {k + i: k \\in Nat}) - = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" . - } - from this have s1_13: "\ n. (n \\in {k + i: k \\in Nat}) - = (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - by auto - have s1_14: "\ n. (n \\in {k + i: k \\in Nat}) - ==> (n \\in ({i} \ {k + Succ[i]: k \\in Nat}))" - using s1_13 by auto - have s1_15: "\ n. (n \\in ({i} \ {k + Succ[i]: k \\in Nat})) - ==> (n \\in {k + i: k \\in Nat})" - using s1_13 by auto - show ?thesis - using s1_14 s1_15 - by (rule setEqualI, blast) - qed - - -(* -THEOREM LeqNatOpenInterval == - ASSUME NEW i \in Nat, - NEW j \in {k + Succ[i]: k \in Nat} - PROVE ~ (j <= i) -PROOF -<1>1. PICK k \in Nat: j = k + Succ[i] - OBVIOUS -<1>2. Succ[i] <= j - BY <1>1 DEF leq -<1>3. i < Succ[i] - BY i \in Nat, less_succ_nat -<1>4. i < j - <2>1. i \in Int - BY i \in Nat, nat_in_int - <2>2. j \in Int - BY <1>1, i \in Nat, succIsNat, nat_in_int, - addIsInt - <2>3. Succ[i] \in Int - BY i \in Nat, succIsNat, nat_in_int - <2> QED - BY <2>1, <2>2, <2>3, <1>3, <1>2, - leq_less_trans -<1>5. i \in Int - BY i \in Nat, nat_in_int -<1>6. j \in Int - BY <1>1, i \in Nat, succIsNat, nat_in_int, - addIsInt -<1> QED - BY <1>4, <1>5, <1>6, less_not_leq -*) -theorem LeqNatOpenInterval: - assumes idom: "i \\in Nat" and - jdom: "j \\in {k + Succ[i]: k \\in Nat}" - shows "~ (j <= i)" - proof - - define P where "P \ \ x. x \\in Nat \ j = x + Succ[i]" - define k where "k \ CHOOSE x: P(x)" - have s1_1: "k \\in Nat \ j = k + Succ[i]" - proof - - have s2_1: "\\E x: P(x)" - using jdom by (auto simp: P_def) - have s2_2: "P(k)" - unfolding k_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_2: "Succ[i] <= j" - proof - - have s2_1: "k + Succ[i] = j" - using s1_1 by auto - have s2_2: "Succ[i] + k = j" - using s2_1 s1_1 idom succIsNat nat_in_int - AddCommutative_sequent by auto - have s2_3: "k \\in Nat" - using s1_1 by auto - show ?thesis - unfolding leq_def - using s2_2 s2_3 by auto - qed - have s1_3: "i < Succ[i]" - using idom less_succ_nat by auto - have s1_4: "i < j" - proof - - have s2_1: "i \\in Int" - using idom nat_in_int by auto - have s2_2: "j \\in Int" - proof - - have s3_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s3_2: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s3_3: "j = k + Succ[i]" - using s1_1 by auto - show ?thesis - using s3_1 s3_2 s3_3 addIsInt by auto - qed - have s2_3: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - show ?thesis - using s2_1 s2_2 s2_3 s1_3 s1_2 - less_leq_trans by iprover - qed - have s1_5: "i \\in Int" - using idom nat_in_int by auto - have s1_6: "j \\in Int" - proof - - have s2_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s2_2: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s2_3: "j = k + Succ[i]" - using s1_1 by auto - show ?thesis - using s2_1 s2_2 s2_3 addIsInt by auto - qed - show ?thesis - using s1_4 s1_5 s1_6 less_not_leq by auto - qed - - -(* -negNatIsAdd0 == - negNat = {-.k + -.0: k \in Nat} -PROOF -<1>1. negNat = {-.k: k \in Nat} - BY DEF negNat -<1>2. ASSUME NEW k \in Nat - PROVE -.k + -.0 = -.k - <2>1. -.k + -.0 = -.k + 0 - BY neg0 - <2>2. @ = -.k - <3>1. -.k \in Int - BY <1>2, minus_nat_in_int - <3> QED - BY <3>1, add_0 - <2> QED - BY <2>1, <2>2 -<1>3. {-.k: k \in Nat} = - (-.k + -.0: k \in Nat} - BY <1>2 -<1> QED - BY <1>1, <1>3 -*) -theorem negNatIsAdd0: - "negNat = {-.k + -.0: k \\in Nat}" - proof - - have s1_1: "negNat = {-.k: k \\in Nat}" - unfolding negNat_def by auto - have s1_2: "\ k. k \\in Nat ==> -.k + -.0 = -.k" - proof - - fix "k" :: "c" - assume s1_2_asm: "k \\in Nat" - have s2_1: "-.k + -.0 = -.k + 0" - using neg0 by auto - also have s2_2: "... = -.k" - proof - - have s3_1: "-.k \\in Int" - using s1_2_asm minus_nat_in_int by auto - show ?thesis - using s3_1 add_0 by auto - qed - finally show "-.k + -.0 = -.k" . - qed - have s1_3: "{-.k: k \\in Nat} = - {-.k + -.0: k \\in Nat}" - using s1_2 by auto - show ?thesis - using s1_1 s1_3 by auto - qed - - -theorem negNatSplitAddi: - assumes idom: "i \\in Nat" - shows "{-.k + -.i: k \\in Nat} = - ({-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - proof - - have s1_1: "\ n. (n \\in {-.k + -.i: k \\in Nat}) - = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - proof - - fix "n" :: "c" - have s2_1: "n \\in {-.k + -.i: k \\in Nat} - = (\\E k \\in Nat: n = -.k + -.i)" - by auto - also have s2_2: "... = ( - \\E k \\in ({0} \ {q + 1: q \\in Nat}): - n = -.k + -.i)" - using SplitNat0 by auto - also have s2_3: "... = ( - (\\E k \\in {0}: n = -.k + -.i) \ - (\\E k \\in {q + 1: q \\in Nat}: n = -.k + -.i) - )" - by auto - also have s2_4: "... = ( - (n = -.0 + -.i) \ - (\\E k: (k \\in {q + 1: q \\in Nat}) \ (n = -.k + -.i)))" - by auto - also have s2_5: "... = ( - (n = -.i) \ - (\\E k: (\\E q \\in Nat: k = q + 1) \ n = -.k + -.i))" - using idom minus_nat_in_int add_0_left by auto - also have s2_6: "... = ( - (n = -.i) \ - (\\E k: \\E q \\in Nat: k = q + 1 \ (n = -.k + -.i)))" - by auto - also have s2_7: "... = ( - (n = -.i) \ - (\\E q \\in Nat: n = -.(q + 1) + -.i))" - by auto - also have s2_8: "... = ( - (n = -.i) \ - (\\E q \\in Nat: n = -.q + -.Succ[i]))" - proof - - have s3_1: "\\E q \\in Nat: n = -.(q + 1) + -.i ==> - \\E q \\in Nat: n = -.q + -.Succ[i]" - proof - - assume s3_1_asm: "\\E q \\in Nat: n = -.(q + 1) + -.i" - define P where "P \ \ x. x \\in Nat \ - n = -.(x + 1) + -.i" - define r where "r \ CHOOSE x: P(x)" - have s4_1: "r \\in Nat \ n = -.(r + 1) + -.i" - proof - - have s5_1: "\\E x: P(x)" - using s3_1_asm by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_2: "n = (-.r + -.1) + -.i" - using s4_1 oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_3: "n = -.r + (-.i + -.1)" - using s4_2 s4_1 idom oneIsNat nat_in_int minus_nat_in_int - AddAssociative_sequent AddCommutative_sequent by auto - have s4_4: "n = -.r + -.(i + 1)" - using s4_3 idom oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_5: "n = -.r + -.Succ[i]" - proof - - have s5_1: "Succ[i] = i + 1" - using idom nat_add_1 by fast - show ?thesis - using s5_1 s4_4 by auto - qed - show "\\E q \\in Nat: n = -.q + -.Succ[i]" - using s4_1 s4_5 by auto - qed - have s3_2: "\\E q \\in Nat: n = -.q + -.Succ[i] ==> - \\E q \\in Nat: n = -.(q + 1) + -.i" - proof - - assume s3_2_asm: "\\E q \\in Nat: n = -.q + -.Succ[i]" - define P where "P \ \ x. x \\in Nat \ - n = -.x + -.Succ[i]" - define r where "r \ CHOOSE x: P(x)" - have s4_1: "r \\in Nat \ n = -.r + -.Succ[i]" - proof - - have s5_1: "\\E x: P(x)" - using s3_2_asm by (auto simp: P_def) - have s5_2: "P(r)" - unfolding r_def - using s5_1 chooseI_ex[of P] by auto - show ?thesis - using s5_2 by (auto simp: P_def) - qed - have s4_2: "n = -.r + -.(i + 1)" - proof - - have s5_1: "Succ[i] = i + 1" - using idom nat_add_1 by fast - show ?thesis - using s4_1 s5_1 by auto - qed - have s4_3: "n = -.r + (-.i + -.1)" - using s4_2 idom oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - have s4_4: "n = (-.r + -.1) + -.i" - using s4_3 s4_1 idom oneIsNat minus_nat_in_int - AddAssociative_sequent AddCommutative_sequent - by auto - have s4_5: "n = -.(r + 1) + -.i" - using s4_4 s4_1 oneIsNat nat_in_int - MinusDistributesAdd_sequent by auto - show "\\E q \\in Nat: n = -.(q + 1) + -.i" - using s4_1 s4_5 by auto - qed - show ?thesis - using s3_1 s3_2 by auto - qed - also have s2_9: "... = ( - (n \\in {-.i}) \ - (n \\in {-.k + -.Succ[i]: k \\in Nat}))" - by auto - also have s2_10: "... = ( - n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - by auto - finally show "(n \\in {-.k + -.i: k \\in Nat}) - = (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" . - qed - have s1_2: "\ n. (n \\in {-.k + -.i: k \\in Nat}) - ==> (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat})" - using s1_1 by auto - have s1_3: "\ n. (n \\in {-.i} \ {-.k + -.Succ[i]: k \\in Nat}) - ==> (n \\in {-.k + -.i: k \\in Nat})" - using s1_1 by auto - show ?thesis - using s1_2 s1_3 by (rule setEqualI, blast) - qed - - -theorem LeqnegNatOpenInterval: - assumes idom: "i \\in Nat" and - jdom: "j \\in {-.k + -.Succ[i]: k \\in Nat}" - shows "~ (-.i <= j)" - proof - - define P where "P \ - \ x. x \\in Nat \ j = -.x + -.Succ[i]" - define k where "k \ CHOOSE x: P(x)" - have s1_1: "k \\in Nat \ j = -.k + -.Succ[i]" - proof - - have s2_1: "\\E x: P(x)" - using jdom by (auto simp: P_def) - have s2_2: "P(k)" - unfolding k_def - using s2_1 chooseI_ex[of P] by auto - show ?thesis - using s2_2 by (auto simp: P_def) - qed - have s1_types: "j \\in Int \ -.i \\in Int \ - -.j \\in Int \ Succ[i] \\in Int \ - -.k \\in Int \ -.Succ[i] \\in Int \ - k \\in Int \ i \\in Int" - proof - - have s2_1: "k \\in Int" - using s1_1 nat_in_int by auto - have s2_2: "-.k \\in Int" - using s2_1 neg_int_eq_int by auto - have s2_3: "i \\in Int" - using idom nat_in_int by auto - have s2_4: "-.i \\in Int" - using s2_3 neg_int_eq_int by auto - have s2_5: "Succ[i] \\in Int" - using idom succIsNat nat_in_int by auto - have s2_6: "-.Succ[i] \\in Int" - using s2_5 neg_int_eq_int by auto - have s2_7: "j \\in Int" - proof - - have s2_1: "j = -.k + -.Succ[i]" - using s1_1 by auto - show ?thesis - using s2_1 s2_2 s2_6 addIsInt by auto - qed - have s2_8: "-.j \\in Int" - using s2_7 neg_int_eq_int by auto - show ?thesis - using s2_1 s2_2 s2_3 s2_4 s2_5 - s2_6 s2_7 s2_8 by auto - qed - have s1_2: "Succ[i] <= -.j" - proof - - have s2_1: "Succ[i] + k = -.j" - proof - - have s3_1: "j = -.k + -.Succ[i]" - using s1_1 by auto - have s3_2: "-.j + j = -.j + (-.k + -.Succ[i])" - using s3_1 by auto - have s3_3: "0 = -.j + (-.k + -.Succ[i])" - using s3_2 s1_types AddNegCancels_left by auto - have s3_4: "0 + Succ[i] = (-.j + (-.k + -.Succ[i])) + Succ[i]" - using s3_3 by auto - have s3_5: "0 + Succ[i] = ((-.j + -.k) + -.Succ[i]) + Succ[i]" - using s3_4 s1_types AddAssociative_sequent by auto - have s3_6: "0 + Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" - using s3_5 s1_types addIsInt AddAssociative_sequent by auto - have s3_7: "Succ[i] = (-.j + -.k) + (-.Succ[i] + Succ[i])" - proof - - have s4_1: "Succ[i] \\in Int" - using s1_types by auto - have s4_2: "0 + Succ[i] = Succ[i]" - using s4_1 add_0_left by fast - show ?thesis - using s3_6 s4_2 by auto - qed - have s3_8: "Succ[i] = (-.j + -.k) + 0" - using s3_7 s1_types AddNegCancels_left by auto - have s3_9: "Succ[i] = -.j + -.k" - proof - - have s4_1: "-.j + -.k \\in Int" - using s1_types addIsInt by auto - show ?thesis - using s3_8 s4_1 add_0 by auto - qed - have s3_10: "Succ[i] + k = (-.j + -.k) + k" - using s3_9 by auto - have s3_11: "Succ[i] + k = -.j + (-.k + k)" - using s3_10 s1_types AddAssociative_sequent by auto - have s3_12: "Succ[i] + k = -.j + 0" - using s3_11 s1_types AddNegCancels_left by auto - show ?thesis - using s3_12 s1_types add_0 by auto - qed - have s2_2: "\\E r \\in Nat: Succ[i] + r = -.j" - using s1_1 s2_1 by auto - show ?thesis - unfolding leq_def - using s2_2 by auto - qed - have s1_3: "i < Succ[i]" - using idom less_succ_nat by fast - have s1_4: "i < -.j" - using s1_types s1_2 s1_3 less_leq_trans by auto - have s1_5: "j < -.i" - proof - - have s2_1: "-.-.j < -.i" - proof - - have s3_1: "i \\in Int" - using s1_types by auto - have s3_2: "-.j \\in Int" - using s1_types by auto - show ?thesis - using s3_1 s3_2 s1_4 minus_less by auto - qed - have s2_2: "-.-.j = j" - using s1_types minus_involutive by auto - show ?thesis - using s2_1 s2_2 by auto - qed - show ?thesis - using s1_5 s1_types less_not_leq by auto - qed - - -theorem split_interval_nat_negnat: - shows "{k \\in Int: a <= k \ k <= b} = - {k \\in negNat: a <= k \ k <= b} \ - {k \\in Nat: a <= k \ k <= b}" - using int_union_nat_negnat by auto - - -end diff --git a/isabelle/Peano.thy b/isabelle/Peano.thy deleted file mode 100644 index 90fb821e..00000000 --- a/isabelle/Peano.thy +++ /dev/null @@ -1,629 +0,0 @@ -(* Title: TLA+/Peano.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation - License: BSD - Version: Isabelle2020 -*) - -section \ Peano's axioms and natural numbers \ - -theory Peano -imports FixedPoints Functions -begin - -text \ - As a preparation for the definition of numbers and arithmetic - in \tlaplus{}, we state Peano's axioms for natural numbers and - prove the existence of a structure satisfying them. The presentation - of the axioms is somewhat simplified compared to the \tlaplus{} book. - (Moreover, the existence of such a structure is assumed, but not - proven in the book.) -\ - -subsection \ The Peano Axioms \ - -definition PeanoAxioms :: "[c,c,c] \ c" where - \ \parameters: the set of natural numbers, zero, and successor function\ - "PeanoAxioms(N,Z,Sc) \ - Z \ N - \ Sc \ [N \ N] - \ (\n \ N : Sc[n] \ Z) - \ (\m,n \ N: Sc[m] = Sc[n] \ m = n) - \ (\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S)" - -text \ - The existence of a structure satisfying Peano's axioms is proven - following the standard ZF construction where @{text "{}"} is zero, - @{text "i \ {i}"} is taken as the successor of any natural - number @{text i}, and the set of natural numbers is defined as - the least set that contains zero and is closed under successor - (this is a subset of the infinity set asserted to exist in ZF - set theory). In \tlaplus{}, natural numbers are defined by a sequence - of @{text CHOOSE}'s below, so there is no commitment to that - particular structure. -\ - -theorem peanoExists: "\N,Z,Sc : PeanoAxioms(N,Z,Sc)" -proof - - let ?sc = "\n. addElt(n,n)" \ \ successor \emph{operator} \ - define expand where "expand \ \S. {{}} \ { ?sc(n) : n \ S}" - define N where "N \ lfp(infinity, expand)" - define Z where "Z \ {}" - define Sc where "Sc \ [n \ N \ ?sc(n)]" \ \successor \emph{function}\ - have mono: "Monotonic(infinity, expand)" - using infinity by (auto simp: Monotonic_def expand_def) - hence expandN: "expand(N) \ N" - by (unfold N_def, rule lfpPreFP) - from expandN have 1: "Z \ N" - by (auto simp: expand_def Z_def) - have 2: "Sc \ [N \ N]" - proof (unfold Sc_def, rule functionInFuncSet) - show "\n \ N : ?sc(n) \ N" using expandN by (auto simp: expand_def) - qed - have 3: "\m\N : Sc[m] \ Z" - unfolding Z_def Sc_def by auto - have 4: "\m,n \ N : Sc[m] = Sc[n] \ m = n" - proof (clarify) - fix m n - assume "m \ N" and "n \ N" and "Sc[m] = Sc[n]" - hence eq: "?sc(m) = ?sc(n)" by (simp add: Sc_def) - show "m = n" - proof (rule setEqual) - show "m \ n" - proof (rule subsetI) - fix x - assume x: "x \ m" show "x \ n" - proof (rule contradiction) - assume "x \ n" - with x eq have "n \ m" by auto - moreover - from eq have "m \ ?sc(n)" by auto - ultimately - show "FALSE" by (blast elim: inAsym) - qed - qed - next - show "n \ m" - proof (rule subsetI) - fix x - assume x: "x \ n" show "x \ m" - proof (rule contradiction) - assume "x \ m" - with x eq have "m \ n" by auto - moreover - from eq have "n \ ?sc(m)" by auto - ultimately - show "FALSE" by (blast elim: inAsym) - qed - qed - qed - qed - have 5: "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" - proof (clarify del: subsetI) - fix S - assume sub: "S \ N" and Z: "Z \ S" and Sc: "\n\S : Sc[n] \ S" - show "N \ S" - proof (unfold N_def, rule lfpLB) - show "expand(S) \ S" - proof (auto simp: expand_def) - from Z show "{} \ S" by (simp add: Z_def) - next - fix n - assume n: "n \ S" - with Sc have "Sc[n] \ S" .. - moreover - from n sub have "n \ N" by auto - hence "Sc[n] = ?sc(n)" by (simp add: Sc_def) - ultimately show "?sc(n) \ S" by simp - qed - next - have "N \ infinity" - by (unfold N_def, rule lfpSubsetDomain) - with sub show "S \ infinity" by auto - qed - qed - from 1 2 3 4 5 have "PeanoAxioms(N,Z,Sc)" - unfolding PeanoAxioms_def by blast - thus ?thesis by blast -qed - -lemma peanoInduct: - assumes pa: "PeanoAxioms(N,Z,Sc)" - and "S \ N" and "Z \ S" and "\n. n \ S \ Sc[n] \ S" - shows "N \ S" -proof - - from pa have "\S \ SUBSET N : Z \ S \ (\n\S : Sc[n] \ S) \ N \ S" - unfolding PeanoAxioms_def by blast - with assms show ?thesis by blast -qed - - -subsection \ Natural numbers: definition and elementary theorems \ - -text \ - The structure of natural numbers is now defined to be some set, - zero, and successor satisfying Peano's axioms. -\ - -definition Succ :: "c" -where "Succ \ CHOOSE Sc : \N,Z : PeanoAxioms(N,Z,Sc)" - -definition Nat :: "c" -where "Nat \ DOMAIN Succ" - -definition zero :: "c" ("0") -where "0 \ CHOOSE Z : PeanoAxioms(Nat, Z, Succ)" - -abbreviation "one \ Succ[0]" - notation "one" ("1") -definition two :: "c" ("2") - where "2 \ Succ[1]" -definition three :: "c" ("3") - where "3 \ Succ[2]" -definition four :: "c" ("4") - where "4 \ Succ[3]" -definition five :: "c" ("5") - where "five \ Succ[4]" -definition six :: "c" ("6") - where "six \ Succ[5]" -definition seven :: "c" ("7") - where "seven \ Succ[6]" -definition eight :: "c" ("8") - where "eight \ Succ[7]" -definition nine :: "c" ("9") - where "nine \ Succ[8]" -definition ten :: "c" ("10") - where "ten \ Succ[9]" -definition eleven :: "c" ("11") - where "eleven \ Succ[10]" -definition twelve :: "c" ("12") - where "twelve \ Succ[11]" -definition thirteen :: "c" ("13") - where "thirteen \ Succ[12]" -definition fourteen :: "c" ("14") - where "fourteen \ Succ[13]" -definition fifteen :: "c" ("15") - where "fifteen \ Succ[14]" - -(* -abbreviation "one \ Succ[0]" -notation "one" ("1") -abbreviation "two \ Succ[1]" -notation "two" ("2") -abbreviation "three \ Succ[2]" -notation "three" ("3") -abbreviation "four \ Succ[3]" -notation "four" ("4") -abbreviation "five \ Succ[4]" -notation "five" ("5") -abbreviation "six \ Succ[5]" -notation "six" ("6") -abbreviation "seven \ Succ[6]" -notation "seven" ("7") -abbreviation "eight \ Succ[7]" -notation "eight" ("8") -abbreviation "nine \ Succ[8]" -notation "nine" ("9") -abbreviation "ten \ Succ[9]" -notation "ten" ("10") -abbreviation "eleven \ Succ[10]" -notation "eleven" ("11") -abbreviation "twelve \ Succ[11]" -notation "twelve" ("12") -abbreviation "thirteen \ Succ[12]" -notation "thirteen" ("13") -abbreviation "fourteen \ Succ[13]" -notation "fourteen" ("14") -abbreviation "fifteen \ Succ[14]" -notation "fifteen" ("15") -*) - -lemma peanoNatZeroSucc: "PeanoAxioms(Nat, 0, Succ)" -proof - - have "\N,Z : PeanoAxioms(N,Z,Succ)" - proof (unfold Succ_def, rule chooseI_ex) - from peanoExists show "\Sc,N,Z : PeanoAxioms(N,Z,Sc)" by blast - qed - then obtain N Z where PNZ: "PeanoAxioms(N,Z,Succ)" by blast - hence "Succ \ [N \ N]" - by (simp add: PeanoAxioms_def) - hence "N = Nat" - by (simp add: Nat_def) - with PNZ have "PeanoAxioms(Nat, Z, Succ)" by simp - thus ?thesis by (unfold zero_def, rule chooseI) -qed - -lemmas - setEqualI [where A = "Nat", intro!] - setEqualI [where B = "Nat", intro!] - -lemma zeroIsNat [intro!,simp]: "0 \ Nat" -using peanoNatZeroSucc by (simp add: PeanoAxioms_def) - -lemma succInNatNat [intro!,simp]: "Succ \ [Nat \ Nat]" -using peanoNatZeroSucc by (simp add: PeanoAxioms_def) - -lemma succIsAFcn [intro!,simp]: "isAFcn(Succ)" -using succInNatNat by blast - -\ \@{text "DOMAIN Succ = Nat"}\ -lemmas domainSucc [intro!,simp] = funcSetDomain[OF succInNatNat] -\ \@{text "n \ Nat \ Succ[n] \ Nat"}\ -lemmas succIsNat [intro!,simp] = funcSetValue[OF succInNatNat] - -lemma oneIsNat [intro!,simp]: "1 \ Nat" -by simp - -lemma twoIsNat [intro!,simp]: "2 \ Nat" -unfolding two_def by simp - -lemma [simp]: - assumes "n \ Nat" - shows "(Succ[n] = 0) = FALSE" -using assms peanoNatZeroSucc by (auto simp: PeanoAxioms_def) - -lemma [simp]: - assumes n: "n \ Nat" - shows "(0 = Succ[n]) = FALSE" -using assms by (auto dest: sym) - -lemma succNotZero (*[elim] \ don't: produces "ignoring weak elimination rule"*): - "\Succ[n] = 0; n \ Nat\ \ P" - "\0 = Succ[n]; n \ Nat\ \ P" -by (simp+) - -lemma succInj [dest]: - assumes "Succ[m] = Succ[n]" and "m \ Nat" and "n \ Nat" - shows "m=n" -using peanoNatZeroSucc assms by (auto simp: PeanoAxioms_def) - -lemma succInjIff [simp]: - assumes "m \ Nat" and "n \ Nat" - shows "(Succ[m] = Succ[n]) = (m = n)" -using assms by auto - -lemma natInduct: - assumes z: "P(0)" - and sc: "\n. \n \ Nat; P(n)\ \ P(Succ[n])" - shows "\n\Nat : P(n)" -proof - - let ?P = "{n \ Nat : P(n)}" - from peanoNatZeroSucc have "Nat \ ?P" - by (rule peanoInduct, auto simp: z sc) - thus ?thesis by auto -qed - -\ \version of above suitable for the inductive reasoning package\ -lemma natInductE [case_names 0 Succ, induct set: Nat]: - assumes "n \ Nat" and "P(0)" and "\n. \n \ Nat; P(n)\ \ P(Succ[n])" - shows "P(n)" -using bspec[OF natInduct, where P=P] assms by blast - -(*** EXAMPLE INDUCTION PROOFS *** - -lemma "\n\Nat : n=0 \ (\m \ Nat : n = Succ[m])" -by (rule natInduct, auto) - -lemma - assumes 1: "n \ Nat" - shows "n=0 \ (\m \ Nat : n = Succ[m])" -using 1 by (induct, auto) - -*** END EXAMPLE ***) - -lemma natCases [case_names 0 Succ, cases set: Nat]: - assumes n: "n \ Nat" - and z: "n=0 \ P" and sc: "\m. \m \ Nat; n = Succ[m]\ \ P" - shows "P" -proof - - from n have "n=0 \ (\m \ Nat : n = Succ[m])" - by (induct, auto) - thus ?thesis - proof - assume "n=0" thus "P" by (rule z) - next - assume "\m\Nat : n = Succ[m]" - then obtain m where "m \ Nat" and "n = Succ[m]" .. - thus "P" by (rule sc) - qed -qed - -lemma succIrrefl: - assumes n: "n \ Nat" - shows "Succ[n] \ n" -using n by (induct, auto) - -lemma succIrreflE (*[elim] -- don't: "ignoring weak elimination rule" *): - "\Succ[n] = n; n \ Nat\ \ P" - "\n = Succ[n]; n \ Nat\ \ P" -by (auto dest: succIrrefl) - -lemma succIrrefl_iff [simp]: - "n \ Nat \ (Succ[n] = n) = FALSE" - "n \ Nat \ (n = Succ[n]) = FALSE" -by (auto dest: succIrrefl) - - -text \Induction over two parameters along the ``diagonal''.\ -lemma diffInduction: - assumes b1: "\m\Nat : P(m,0)" and b2: "\n\Nat : P(0, Succ[n])" - and step: "\m,n\Nat : P(m,n) \ P(Succ[m], Succ[n])" - shows "\m,n\Nat : P(m,n)" -proof (rule natInduct) - show "\n\Nat : P(0,n)" - using b1 b2 by (intro natInduct, auto) -next - fix m - assume m: "m \ Nat" and ih: "\n\Nat : P(m,n)" - show "\n\Nat : P(Succ[m],n)" - proof (rule bAllI) - fix n - assume "n \ Nat" thus "P(Succ[m],n)" - proof (cases) - case 0 with b1 m show ?thesis by auto - next - case Succ with step ih m show ?thesis by auto - qed - qed -qed - -lemma diffInduct: - assumes n: "n \ Nat" and m: "m \ Nat" - and b1: "\m. m\Nat \ P(m,0)" and b2: "\n. n\Nat \ P(0, Succ[n])" - and step: "\m n. \m \ Nat; n\Nat; P(m,n) \ \ P(Succ[m], Succ[n])" - shows "P(m,n)" -proof - - have "\m,n\Nat : P(m,n)" - by (rule diffInduction, auto intro: b1 b2 step) - with n m show ?thesis by blast -qed - -lemma not0_implies_Suc: - "\n \ Nat; n \ 0\ \ \m \ Nat: n = Succ[m]" -by(rule natCases, auto) - -subsection \ Initial intervals of natural numbers and ``less than'' \ - -text \ - The set of natural numbers up to (and including) a given $n$ is - inductively defined as the smallest set of natural numbers that - contains $n$ and that is closed under predecessor. - - NB: ``less than'' is not first-order definable from the Peano axioms, - a set-theoretic definition such as the following seems to be unavoidable. -\ - -definition upto :: "c \ c" -where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : Succ[k] \ S })" - -lemmas - setEqualI [where A = "upto(n)" for n, intro!] - setEqualI [where B = "upto(n)" for n, intro!] - -lemma uptoNat: "upto(n) \ Nat" - unfolding upto_def by (rule lfpSubsetDomain) - -lemma uptoPred: - assumes Suc: "Succ[m] \ upto(n)" and m: "m \ Nat" and n: "n \ Nat" - shows "m \ upto(n)" -proof - - let ?f = "\S. {n} \ {k\Nat : Succ[k] \ S}" - from n have mono: "Monotonic(Nat, ?f)" - unfolding Monotonic_def by blast - from m Suc have 1: "m \ ?f(upto(n))" by auto - from mono have 2: "?f(upto(n)) \ upto(n)" - unfolding upto_def by (rule lfpPreFP) - from 1 2 show ?thesis by blast -qed - -lemma uptoZero: "upto(0) = {0}" -proof (rule setEqual) - have "{0} \ { k \ Nat : Succ[k] \ {0} } \ {0}" by auto - thus "upto(0) \ {0}" - unfolding upto_def by (rule lfpLB, auto) -next - show "{0} \ upto(0)" - unfolding upto_def by (rule lfpGLB, auto) -qed - -lemma uptoSucc: - assumes n: "n \ Nat" - shows "upto(Succ[n]) = upto(n) \ {Succ[n]}" (is "?lhs = ?rhs") -proof - - let ?preds = "\S. {k \ Nat : Succ[k] \ S}" - let ?f = "\S k. {k} \ ?preds(S)" - have mono: "\k. k \ Nat \ Monotonic(Nat, \S. ?f(S,k))" - by (auto simp: Monotonic_def) - \ \``$\subseteq$''\ - from n have "?preds(?rhs) \ ?f(upto(n), n)" by auto - also have "\ \ upto(n)" - by (unfold upto_def, rule lfpPreFP, rule mono, rule n) - finally have "?f(?rhs, Succ[n]) \ ?rhs" by auto - moreover from n have "?rhs \ Nat" - by (intro unionLUB, auto elim: uptoNat[THEN subsetD]) - ultimately have 1: "?lhs \ ?rhs" - by (unfold upto_def[where n="Succ[n]"], rule lfpLB) - \ \``$\supseteq$''\ - from n mono have 2: "?f(?lhs, Succ[n]) \ ?lhs" - unfolding upto_def by (intro lfpPreFP, blast) - with n have "?f(?lhs, n) \ ?lhs" by auto - moreover have "?lhs \ Nat" by (rule uptoNat) - ultimately have 3: "upto(n) \ ?lhs" - unfolding upto_def[where n=n] by (rule lfpLB) - from 2 have 4: "Succ[n] \ ?lhs" by auto - from 3 4 have "?rhs \ ?lhs" by auto - with 1 show ?thesis by (rule setEqual) -qed - -lemma uptoRefl: - assumes n: "n \ Nat" - shows "n \ upto(n)" -using n proof (cases) - case 0 thus ?thesis by (simp add: uptoZero) -next - case Succ thus ?thesis by (auto simp: uptoSucc) -qed - -lemma zeroInUpto: - assumes n: "n \ Nat" - shows "0 \ upto(n)" -using n by (induct, auto simp: uptoZero uptoSucc) - -lemma SuccNotUptoZero: - assumes "n \ Nat" and "Succ[n] \ upto(0)" - shows "P" -using assms by (auto simp: uptoZero) - -lemma uptoTrans: - assumes "k \ upto(m)" and "m \ upto(n)" and "n \ Nat" - shows "k \ upto(n)" -proof - - have "\n\Nat : m \ upto(n) \ upto(m) \ upto(n)" - by (rule natInduct, auto simp: uptoZero uptoSucc) - with assms show ?thesis by blast -qed - -lemma succNotinUpto: - assumes n: "n \ Nat" - shows "Succ[n] \ upto(n)" -using n proof (induct) - show "1 \ upto(0)" by (auto simp: uptoZero) -next - fix n - assume n: "n \ Nat" and ih: "Succ[n] \ upto(n)" - show "Succ[Succ[n]] \ upto(Succ[n])" - proof (auto simp: uptoSucc n) - assume "Succ[Succ[n]] \ upto(n)" - with n have "Succ[n] \ upto(n)" - by (auto elim: uptoPred) - with ih show "FALSE" .. - qed -qed - -lemma uptoLimit: - assumes m: "m \ upto(n)" and suc: "Succ[m] \ upto(n)" and n: "n \ Nat" - shows "m=n" -proof - - from m uptoNat have mNat: "m \ Nat" by blast - from n have "\m\Nat: m \ upto(n) \ Succ[m] \ upto(n) \ m=n" (is "?P(n)") - by (induct, auto simp: uptoZero uptoSucc) - with mNat m suc show ?thesis by blast -qed - -lemma uptoAntisym: - assumes mn: "m \ upto(n)" and nm: "n \ upto(m)" - shows "m=n" -proof - - from mn uptoNat have m: "m \ Nat" by blast - from nm uptoNat have n: "n \ Nat" by blast - have "\m,n\Nat : m \ upto(n) \ n \ upto(m) \ m=n" (is "\m,n\Nat : ?P(m,n)") - proof (rule natInduct) - show "\n\Nat : ?P(0,n)" by (auto simp: uptoZero) - next - fix m - assume m: "m \ Nat" and ih: "\n\Nat : ?P(m,n)" - show "\n\Nat : ?P(Succ[m],n)" - proof (auto simp: uptoSucc m) - fix n - assume "Succ[m] \ upto(n)" and "n \ upto(m)" - from this m have "Succ[m] \ upto(m)" by (rule uptoTrans) - with m show "Succ[m] = n" \ \contradiction\ - by (blast dest: succNotinUpto) - qed - qed - with m n mn nm show ?thesis by blast -qed - -lemma uptoInj [simp]: - assumes n: "n \ Nat" and m: "m \ Nat" - shows "(upto(n) = upto(m)) = (n = m)" -proof (auto) - assume 1: "upto(n) = upto(m)" - from n have "n \ upto(n)" by (rule uptoRefl) - with 1 have "n \ upto(m)" by auto - moreover - from m have "m \ upto(m)" by (rule uptoRefl) - with 1 have "m \ upto(n)" by auto - ultimately - show "n = m" by (rule uptoAntisym) -qed - -lemma uptoLinear: - assumes m: "m \ Nat" and n: "n \ Nat" - shows "m \ upto(n) \ n \ upto(m)" (is "?P(m,n)") -using m proof induct - from n show "?P(0,n)" by (auto simp: zeroInUpto) -next - fix k - assume k: "k \ Nat" and ih: "?P(k,n)" - from k show "?P(Succ[k],n)" - proof (auto simp: uptoSucc) - assume kn: "(Succ[k] \ upto(n)) = FALSE" - show "n \ upto(k)" - proof (rule contradiction) - assume c: "n \ upto(k)" - with ih have "k \ upto(n)" by simp - from this kn n have "k = n" by (rule uptoLimit[simplified]) - with n have "n \ upto(k)" by (simp add: uptoRefl) - with c show "FALSE" .. - qed - qed -qed - - -subsection \ Primitive Recursive Functions \ - -text \ - We axiomatize a primitive recursive scheme for functions - with one argument and domain on natural numbers. Later, we - use it to define addition, multiplication and difference. -\ - -(** FIXME: replace axiom with proper fixpoint definition **) - -axiomatization where - primrec_nat: "\f : isAFcn(f) \ DOMAIN f = Nat - \ f[0] = e \ (\n \ Nat : f[Succ[n]] = h(n,f[n]))" - -lemma bprimrec_nat: - assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" - shows "\f \ [Nat \ S] : f[0] = e \ (\n \ Nat: f[Succ[n]] = h(n,f[n]))" -proof - - from primrec_nat[of e h] obtain f where - 1: "isAFcn(f)" and 2: "DOMAIN f = Nat" - and 3: "f[0] = e" and 4: "\n\Nat : f[Succ[n]] = h(n,f[n])" - by blast - have "\n\Nat : f[n] \ S" - proof (rule natInduct) - from 3 e show "f[0] \ S" by simp - next - fix n - assume "n \ Nat" and "f[n] \ S" - with suc 4 show "f[Succ[n]] \ S" by force - qed - with 1 2 3 4 show ?thesis - by blast -qed - -lemma primrec_natE: - assumes e: "e \ S" and suc: "\n \ Nat : \x \ S : h(n,x) \ S" - and f: "f = (CHOOSE g \ [Nat \ S] : g[0] = e \ (\n \ Nat: g[Succ[n]] = h(n,g[n])))" - (is "f = ?g") - and maj: "\ f \ [Nat \ S]; f[0] = e; \n \ Nat : f[Succ[n]] = h(n, f[n]) \ \ P" - shows "P" -proof - - from e suc have "\ g \ [Nat \ S] : g[0] = e \ (\n \ Nat: g[Succ[n]] = h(n,g[n]))" - by (rule bprimrec_nat) - hence "?g \ [Nat \ S] \ ?g[0] = e \ (\n \ Nat: ?g[Succ[n]] = h(n,?g[n]))" - by (rule bChooseI2, auto) - with f maj show ?thesis by blast -qed - -lemma bprimrecType_nat: - assumes "e \ S" and "\n \ Nat : \x \ S : h(n,x) \ S" - shows "(CHOOSE f \ [Nat \ S] : f[0] = e \ - (\n \ Nat: f[Succ[n]] = h(n,f[n]))) - \ [Nat \ S]" -by (rule primrec_natE[OF assms], auto) - -end diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index dae0051c..91cf9114 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1852,7 +1852,7 @@ proof - qed (*** END AUTO_GENERATED STUFF ***) -(* + (* ------------- Proof rules for CASE expressions ------ *) lemma zenon_caseother0 : @@ -1920,7 +1920,7 @@ proof (rule boolEqual, simp only: zenon_seqify_appseq, rule iffI) assume h4: "~ c" with h1 h6 obtain i where i: "i \ DOMAIN zenon_seqify(cs)" "Append(zenon_seqify(cs), c)[i]" - by auto + by force with h6 have "zenon_seqify(cs)[i]" by (auto simp: leq_neq_iff_less[simplified]) with i show "\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i]" by blast qed @@ -1929,11 +1929,13 @@ next assume h1: "c | (\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i])" show "\ i \ DOMAIN Append (zenon_seqify (cs), c) - : Append (zenon_seqify (cs), c)[i]" (is "?g") + : Append (zenon_seqify (cs), c)[i]" + (is "?g" is "\i \ DOMAIN ?app : Append(?seq,c)[i]") proof (rule disjE [OF h1]) assume h2: c - show "?g" - using h2 h0 by auto + with h0 have "succ[Len(?seq)] \ DOMAIN ?app" "?app[succ[Len(?seq)]]" + by auto + thus "?g" by blast next assume h2: "\ i \ DOMAIN zenon_seqify (cs) : zenon_seqify (cs)[i]" @@ -1944,15 +1946,15 @@ next have h4: "i \\in DOMAIN Append (zenon_seqify (cs), c)" using h0 h3 by auto assume h5: "zenon_seqify (cs)[i]" - have h6: "i ~= Succ[Len (zenon_seqify (cs))]" + have h6: "i ~= succ[Len (zenon_seqify (cs))]" using h0 h3 by auto have h7: "Append (zenon_seqify (cs), c)[i]" using h6 h5 h3 h0 by force show "?g" - using h4 h7 by auto + using h4 h7 by blast qed qed -qed (simp_all) +qed (simp+) lemma zenon_case_seq_empty : assumes h0: "\ i \ DOMAIN zenon_seqify (<<>>) @@ -2009,14 +2011,14 @@ proof proof (rule bExE [OF h5]) fix i assume h6: "i \\in DOMAIN Append (zenon_seqify(cs), c)" - have h7: "i = Succ [Len (zenon_seqify (cs))] + have h7: "i = succ [Len (zenon_seqify (cs))] | i \\in DOMAIN (zenon_seqify (cs))" using h3c h6 by auto assume h8: "x \\in CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i])" have h9: "i \\in DOMAIN (zenon_seqify (cs))" (is "?g") proof (rule disjE [OF h7]) - assume h10: "i = Succ [Len (zenon_seqify (cs))]" + assume h10: "i = succ [Len (zenon_seqify (cs))]" have h11: FALSE using h8 h10 h4 h3c h3e h2 by auto show "?g" using h11 by auto next @@ -2101,10 +2103,16 @@ next using h4 zenon_seqifyIsASeq by auto have h5: "x \\in CaseArm (c, e)" using h2 by blast + have "succ[Len(zenon_seqify(cs))] \\in DOMAIN Append (zenon_seqify (cs), c)" + "CaseArm (Append (zenon_seqify (cs), c)[succ[Len(zenon_seqify(cs))]], + Append (zenon_seqify (es), e)[succ[Len(zenon_seqify(cs))]]) + = CaseArm (c, e)" + using h3 zenon_seqifyIsASeq by auto + with h5 have h6: "x \\in UNION {CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) : i \\in DOMAIN Append (zenon_seqify (cs), c)}" - using h5 zenon_seqifyIsASeq appendElt2 h3 by auto + by blast show "?g" using h4 h6 by auto next @@ -2151,12 +2159,12 @@ next \\in {CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) : i \\in DOMAIN Append (zenon_seqify (cs), c)}" - using h10 by auto + using h10 by blast qed qed show "?g" using h4 h7 by blast qed -qed (simp_all) +qed (simp+) lemma zenon_case_len_simpl : fixes cs c es e @@ -2173,7 +2181,7 @@ next show "Len (Append (zenon_seqify (cs), c)) = Len (Append (zenon_seqify (es), e))" using h1 zenon_seqifyIsASeq by auto -qed (simp_all) +qed (simp+) lemma zenon_case_empty_union : fixes x @@ -2209,9 +2217,17 @@ next fix i assume h7: "i \\in DOMAIN Append (zenon_seqify (cs), c)" show "~Append (zenon_seqify (cs), c)[i]" - using g0 h7 h6 zenon_seqifyIsASeq by (unfold Append_def, auto) + proof (cases "i \\in DOMAIN zenon_seqify(cs)") + case True + with h6 zenon_seqifyIsASeq show ?thesis by force + next + case False + with h7 zenon_seqifyIsASeq[of cs] + have "i = succ[Len(zenon_seqify(cs))]" by (auto simp: int_less_not_leq) + with g0 zenon_seqifyIsASeq show ?thesis by auto + qed qed -qed (simp_all) +qed (simp+) lemma zenon_case_oth_simpl_l2 : fixes cs c es e @@ -2244,9 +2260,9 @@ proof assume h9: "i \\in DOMAIN Append (zenon_seqify (cs), c)" assume h10: "CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) = B" - have h11: "i = Succ[Len (zenon_seqify (cs))] ==> FALSE" + have h11: "i = succ[Len (zenon_seqify (cs))] ==> FALSE" proof - - assume h12: "i = Succ[Len (zenon_seqify (cs))]" + assume h12: "i = succ[Len (zenon_seqify (cs))]" have h13: "B = CaseArm (c, e)" using h10 h12 g0 appendElt2 zenon_seqifyIsASeq by auto show FALSE @@ -2418,14 +2434,14 @@ lemma zenon_case1 : assumes hoth: "~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c1x" (is "?dcs") + define dcs where "dcs \ c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2437,10 +2453,10 @@ proof - assume ha: "?dcs" define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2493,29 +2509,29 @@ lemma zenon_caseother1 : assumes hoth: "~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq ( <<>>, c1x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq ( <<>>, e1x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2600,14 +2616,14 @@ lemma zenon_case2 : assumes hoth: "~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c2x | c1x" (is "?dcs") + define dcs where "dcs \ c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2619,10 +2635,10 @@ proof - assume ha: "?dcs" define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2679,29 +2695,29 @@ lemma zenon_caseother2 : assumes hoth: "~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2791,14 +2807,14 @@ lemma zenon_case3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2812,11 +2828,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -2879,30 +2895,30 @@ lemma zenon_caseother3 : assumes hoth: "~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -2999,14 +3015,14 @@ lemma zenon_case4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3021,11 +3037,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -3094,31 +3110,31 @@ lemma zenon_caseother4 : assumes hoth: "~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + define dcs where "dcs \ c4x | c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3222,14 +3238,14 @@ lemma zenon_case5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) - define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "?dcs") + define dcs where "dcs \ c5x | c4x | c3x | c2x | c1x" (is "_ \ ?dcs") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3245,11 +3261,11 @@ proof - define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" - (is "?scs") + (is "_ \ ?scs") define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" - (is "?ses") + (is "_ \ ?ses") have ha1: "\ i \ DOMAIN ?scs : ?scs[i]" using ha zenon_case_seq_empty by (simp only: zenon_case_seq_simpl zenon_seqify_empty, blast) @@ -3324,32 +3340,32 @@ lemma zenon_caseother5 : assumes hoth: "~c5x & ~c4x & ~c3x & ~c2x & ~c1x & TRUE ==> P (oth) ==> FALSE" shows FALSE proof - - define cs where "cs \ <>" (is "?cs") - define es where "es \ <>" (is "?es") + define cs where "cs \ <>" (is "_ \ ?cs") + define es where "es \ <>" (is "_ \ ?es") define arms where "arms \ UNION {CaseArm (?cs[i], ?es[i]) : i \\in DOMAIN ?cs}" - (is "?arms") + (is "_ \ ?arms") define cas where "cas \ ?cas" have h0: "P (cas)" using h by (fold cas_def) define dcs where "dcs \ - c5x | c4x | c3x | c2x | c1x | FALSE" (is "?dcs") + c5x | c4x | c3x | c2x | c1x | FALSE" (is "_ \ ?dcs") define scs where "scs \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, c1x), c2x), c3x), c4x), c5x))" - (is "?scs") + (is "_ \ ?scs") have hscs : "?cs = ?scs" by (simp only: zenon_seqify_appseq zenon_seqify_empty) define ses where "ses \ zenon_seqify (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq (zenon_appseq ( <<>>, e1x), e2x), e3x), e4x), e5x))" - (is "?ses") + (is "_ \ ?ses") have hses : "?es = ?ses" by (simp only: zenon_seqify_appseq zenon_seqify_empty) have hlen: "Len (?scs) = Len (?ses)" (is "?hlen") by (simp only: zenon_case_len_simpl) define armoth where "armoth \ CaseArm (\ i \ DOMAIN ?cs : ~?cs[i], oth)" - (is "?armoth") + (is "_ \ ?armoth") show FALSE proof (rule zenon_em [of "?dcs"]) assume ha: "~(?dcs)" @@ -3446,7 +3462,7 @@ proof - by (rule zenon_case_empty_union [OF hi]) qed qed -*) + (*** END AUTO_GENERATED STUFF ***) end diff --git a/isabelle/examples/Allocator.thy b/isabelle/examples/Allocator.thy index 2c2647f9..30542f61 100644 --- a/isabelle/examples/Allocator.thy +++ b/isabelle/examples/Allocator.thy @@ -1,18 +1,29 @@ -(* Title: TLA+/AtomicBakery.thy - Author: Stephan Merz, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation +(* Title: TLA+/Allocator.thy + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:38:22 merz> + Version: Isabelle2021-1 *) +section \A Simple Resource Allocator\ + theory Allocator imports Constant begin +text \ + We define and prove the safety of a simple-minded resource allocator + where clients can request and return resources. The allocator can + hand out resources that are available. Since the allocator does + not keep track of the fact if clients will be able to satisfy all + their requests, it may get into a situation where clients hold + part of their requested resources but will wait forever for the + outstanding resources, held by other clients. +\ + declare funcSetValue [dest] -consts -- {* constant parameters of TLA+ SimpleAllocator module *} +consts \ \ constant parameters of TLA+ SimpleAllocator module \ Client :: c Resource :: c @@ -28,14 +39,14 @@ lemma InitTypeInvariant: "Init(unsat,alloc) \ TypeInvariant(unsat,alloc)" by (auto simp: Init_def TypeInvariant_def) --- {* Alternative formulation as proof rule (meta-level implication) *} +\ \ Alternative formulation as proof rule (meta-level implication) \ lemma InitTypeInvariant': assumes "Init(unsat,alloc)" shows "TypeInvariant(unsat,alloc)" using assms by (auto simp: Init_def TypeInvariant_def) definition available where - "available(alloc) \ Resource \\ (UNION {alloc[c] : c \ Client})" + "available(alloc) \ Resource \ (UNION {alloc[c] : c \ Client})" definition Request where "Request(c,S,unsat,alloc,unsat',alloc') \ @@ -49,7 +60,7 @@ lemma RequestTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Request_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma RequestTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Request(c,S,unsat,alloc,unsat',alloc')" @@ -60,7 +71,7 @@ definition Allocate where "Allocate(c,S,unsat,alloc,unsat',alloc') \ S \ {} \ S \ available(alloc) \ unsat[c] \ alloc' = [alloc EXCEPT ![c] = alloc[c] \ S] - \ unsat' = [unsat EXCEPT ![c] = alloc[c] \\ S]" + \ unsat' = [unsat EXCEPT ![c] = alloc[c] \ S]" lemma AllocateTypeInvariant: "TypeInvariant(unsat,alloc) @@ -68,7 +79,7 @@ lemma AllocateTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Allocate_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma AllocateTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Allocate(c,S,unsat,alloc,unsat',alloc')" @@ -78,7 +89,7 @@ using assms unfolding TypeInvariant_def Allocate_def by auto definition Return where "Return(c,S,unsat,alloc,unsat',alloc') \ S \ {} \ S \ alloc[c] - \ alloc' = [alloc EXCEPT ![c] = alloc[c] \\ S] + \ alloc' = [alloc EXCEPT ![c] = alloc[c] \ S] \ unsat' = unsat" lemma ReturnTypeInvariant: @@ -87,7 +98,7 @@ lemma ReturnTypeInvariant: \ TypeInvariant(unsat',alloc')" unfolding TypeInvariant_def Return_def by auto --- {* Alternative formulation *} +\ \ Alternative formulation \ lemma ReturnTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "c \ Client" and "S \ SUBSET Resource" and "Return(c,S,unsat,alloc,unsat',alloc')" @@ -109,14 +120,14 @@ by (blast intro: RequestTypeInvariant [rule_format] AllocateTypeInvariant [rule_format] ReturnTypeInvariant [rule_format]) --- {* Alternative formulation, using the rule formats of the individual lemmas *} +\ \ Alternative formulation, using the rule formats of the individual lemmas \ lemma NextTypeInvariant': assumes "TypeInvariant(unsat,alloc)" and "Next(unsat,alloc,unsat',alloc')" shows "TypeInvariant(unsat',alloc')" using assms unfolding Next_def by (blast intro: RequestTypeInvariant' AllocateTypeInvariant' ReturnTypeInvariant') --- {* Direct proof, without use of lemmas for subactions *} +\ \ Direct proof, without use of lemmas for subactions \ lemma NextTypeInvariant'': "TypeInvariant(unsat,alloc) \ Next(unsat,alloc,unsat',alloc') \ TypeInvariant(unsat',alloc')" @@ -124,7 +135,7 @@ unfolding TypeInvariant_def Next_def Request_def Allocate_def Return_def by auto -text {* Proof of mutual exclusion: no two processes ever hold a common resource *} +text \ Proof of mutual exclusion: no two processes ever hold a common resource \ definition Mutex where "Mutex(alloc) \ @@ -133,10 +144,10 @@ definition Mutex where lemma InitMutex: "Init(unsat,alloc) \ Mutex(alloc)" unfolding Init_def Mutex_def by auto -text {* +text \ The @{text Request} action trivially preserves mutual exclusion because it leaves the value of @{text alloc} unchanged. -*} +\ lemma RequestMutex: "Mutex(alloc) @@ -144,10 +155,10 @@ lemma RequestMutex: \ Mutex(alloc')" unfolding Request_def by auto -text {* +text \ The @{text Return} action decreases the set of allocated resources, so preservation of mutual exclusion follows easily. Note the use of the type invariant. -*} +\ lemma ReturnMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -174,10 +185,10 @@ proof (clarify) qed qed -text {* +text \ In fact, the proof is easy for Isabelle's automatic proof methods. Note the use of @{text clarsimp} to simplify the goal before the heavy lifting. -*} +\ lemma "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -185,10 +196,10 @@ lemma \ Mutex(alloc')" by (clarsimp simp: Mutex_def TypeInvariant_def Return_def) auto -text {* +text \ The proof for the @{text Allocate} action requires some case analysis. Unfortunately, we need to prove two symmetric assertions. -*} +\ lemma AllocateMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) @@ -206,13 +217,13 @@ proof (clarify) and rc: "r \ alloc'[c]" and rc': "r \ alloc'[c']" show "c' = c" proof (cases "c=clt \ c'=clt") - case False -- {* the simple case first *} + case False \ \ the simple case first \ with all c c' tpg have "alloc'[c] = alloc[c] \ alloc'[c'] = alloc[c']" by (auto simp: Allocate_def TypeInvariant_def) with mux rc rc' c c' r show "c' = c" by (auto simp: Mutex_def) next - case True -- {* one of @{term c} or @{term c'} is @{term clt} *} + case True \ \ one of @{term c} or @{term c'} is @{term clt} \ thus "c' = c" proof assume c1: "c = clt" @@ -233,7 +244,7 @@ proof (clarify) using all c1 c tpg by (auto simp: Allocate_def TypeInvariant_def) with rc show "FALSE" by auto qed - next -- {* symmetric argument *} + next \ \ symmetric argument \ assume c1: "c' = clt" show "c' = c" proof (rule contradiction) @@ -258,10 +269,10 @@ proof (clarify) qed -text {* +text \ Using the above lemmas, it is straightforward to derive that the next-state action preserves mutual exclusion. -*} +\ lemma NextMutex: "Mutex(alloc) \ TypeInvariant(unsat,alloc) \ Next(unsat,alloc,unsat',alloc') diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index f04f5428..8e0f0a91 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -1,15 +1,23 @@ (* Title: TLA+/AtomicBakery.thy Author: Hernan Vanzetto, LORIA - Copyright (C) 2009-2011 INRIA and Microsoft Corporation + Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2011-1 - Time-stamp: <2011-10-11 17:37:50 merz> + Version: Isabelle2021-1 *) +section \Safety Proof of the Atomic Version of the Bakery Algorithm\ + theory AtomicBakeryG imports Constant begin +text \ + This is the version of the atomic Bakery algorithm used for + presenting TLAPS, verified using only Isabelle/TLA+. The + proof is quite tedious, compared to the proof that makes use + of the different back-end provers available in TLAPS. +\ + (************************************************** \algorithm AtomicBakery { @@ -80,7 +88,7 @@ definition Init where definition p1 where "p1(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p1'' - \ unread' = [unread EXCEPT ![self] = P \\ {self}] + \ unread' = [unread EXCEPT ![self] = P \ {self}] \ max' = [max EXCEPT ![self] = 0] \ flag' = [flag EXCEPT ![self] = TRUE] \ pc' = [pc EXCEPT ![self] = ''p2''] @@ -91,7 +99,7 @@ definition p2 where pc[self] = ''p2'' \ (IF unread[self] \ {} THEN \i \ unread[self] : - ( unread' = [unread EXCEPT ![self] = unread[self] \\ {i}] + ( unread' = [unread EXCEPT ![self] = unread[self] \ {i}] \ (IF num[i] > max[self] THEN max' = [max EXCEPT ![self] = num[i]] ELSE (max' = max))) @@ -111,7 +119,7 @@ definition p4 where "p4(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p4'' \ flag' = [flag EXCEPT ![self] = FALSE] - \ unread' = [unread EXCEPT ![self] = P \\ {self}] + \ unread' = [unread EXCEPT ![self] = P \ {self}] \ pc' = [pc EXCEPT ![self] = ''p5''] \ num' = num \ max' = max \ nxt' = nxt" @@ -133,11 +141,11 @@ definition p6 where \ (IF self > nxt[self] THEN num[nxt[self]] > num[self] ELSE num[nxt[self]] \ num[self])) - \ unread' = [unread EXCEPT ![self] = unread[self] \\ {nxt[self]}] + \ unread' = [unread EXCEPT ![self] = unread[self] \ {nxt[self]}] \ pc' = [pc EXCEPT ![self] = ''p5''] \ num' = num \ flag' = flag \ max' = max \ nxt' = nxt" -definition p7 where -- {* Critical section *} +definition p7 where \ \ Critical section \ "p7(self,unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt') \ pc[self] = ''p7'' \ TRUE @@ -193,12 +201,12 @@ definition TypeOK where \ max \ [P \ Nat \ {defaultInitValue}] \ (\j \ P : (pc[j] \ {''p2'', ''p3''}) \ max[j] \ Nat) \ nxt \ [P \ P \ {defaultInitValue}] - \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ P \\ {i}) + \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ P \ {i}) \ pc \ [P \ {''p1'',''p2'',''p3'',''p4'',''p5'',''p6'',''p7'',''p8''}]" ---{* The type invariant in p6 should be - @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \\ {i})"} +\ \ The type invariant in p6 should be + @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \ {i})"} but it works anyway as it is. -*} +\ (** definition TypeOK where (** version for the alternative initial condition **) @@ -212,10 +220,10 @@ definition TypeOK where (** version for the alternative initial condition **) \ nxt \ [P \ P] \ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ i) \ pc \ [P \ {''p1'',''p2'',''p3'',''p4'',''p5'',''p6'',''p7'',''p8''}]" ---{* The type invariant in p6 should be - @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \\ {i})"} +\ \ The type invariant in p6 should be + @{text "\ (\i \ P : (pc[i] = ''p6'') \ nxt[i] \ unread[i] \ {i})"} but it works anyway as it is. -*} +\ **) definition GG where @@ -239,12 +247,12 @@ definition IInv where ((num[i] = 0) = (pc[i] \ {''p1'', ''p2'', ''p3''})) \ (flag[i] = (pc[i] \ {''p2'', ''p3'', ''p4''})) \ (pc[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread[i]) \\ {i} : After(j,i,unread,max,flag,pc,num,nxt))) + (\j \ (P \ unread[i]) \ {i} : After(j,i,unread,max,flag,pc,num,nxt))) \ ( pc[i] = ''p6'' \ ( (pc[nxt[i]] = ''p2'') \ i \ unread[nxt[i]] \ (pc[nxt[i]] = ''p3'')) \ max[nxt[i]] \ num[i]) - \ ((pc[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread,max,flag,pc,num,nxt)))" + \ ((pc[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread,max,flag,pc,num,nxt)))" definition Inv where "Inv(unread,max,flag,pc,num,nxt) \ @@ -261,8 +269,8 @@ theorem GGIrreflexive: assumes i: "i \ P" and j: "j \ P" and 1: "i \ j" and 2: "num[i] \ Nat \\ {0}" and 3: "num[j] \ Nat \\ {0}" shows "\(GG(i,j,num) \ GG(j,i,num))" -unfolding GG_def using assms nat_less_linear[of i j] -by (auto dest: procInNat nat_less_not_sym nat_less_leq_trans) +unfolding GG_def using assms int_less_linear[of i j] +by (auto dest: procInNat int_less_not_sym int_less_leq_trans) theorem InitImpliesTypeOK: "Init(unread,max,flag,pc,num,nxt) \ TypeOK(unread,max,flag,pc,num,nxt)" @@ -313,15 +321,27 @@ next with i j pc type show "unread'[j] \ P \ j \ unread'[j]" by (auto simp: TypeOK_def) qed - from u type i have n: "num[i] \ Nat" - by (auto simp: TypeOK_def) + from u \i \ unread[self]\ have "i \ ProcSet" by blast + with type have n: "num[i] \ Nat" by (auto simp: TypeOK_def) with i type have m': "max' \ [P \ Nat \ {defaultInitValue}]" by (auto simp: TypeOK_def) - from n i type have m'': "\j \ P : pc'[j] \ {''p2'',''p3''} \ max'[j] \ Nat" - (* FIXME: - - why doesn't Isabelle do the case distinction by itself? - - why do we need condElse? *) - by (cases "max[self] < num[i]", auto simp: TypeOK_def (*condThen*) condElse) + have m'': "\j \ P : pc'[j] \ {''p2'',''p3''} \ max'[j] \ Nat" + proof (clarify) + fix j + assume "j \ ProcSet" "pc'[j] \ {''p2'', ''p3''}" + with type i have mj: "max[j] \ Nat" + unfolding TypeOK_def by auto + show "max'[j] \ Nat" + proof (cases "j = self") + case True + with i mj type \j \ ProcSet\ n show ?thesis + by (cases "max[self] < num[i]") (auto simp: TypeOK_def) + next + case False + with i mj type \j \ ProcSet\ show ?thesis + by (auto simp: TypeOK_def) + qed + qed from triv u' u'' m' m'' show ?thesis by (auto simp add: TypeOK_def) qed @@ -390,7 +410,7 @@ qed lemma leq_neq_trans' (*[dest!]*): "a \ b \ b \ a \ a < b" - by (drule not_sym) (rule leq_neq_trans) + by (simp add: less_def) theorem InvInvariant: assumes inv: "Inv(unread,max,flag,pc,num,nxt)" @@ -407,26 +427,29 @@ proof auto assume type: "TypeOK(unread, max, flag, pc, num, nxt)" and iinv: "\j \ ProcSet: IInv(j, unread, max, flag, pc, num, nxt)" and i: "i \ ProcSet" - -- {* auxiliary definition that is used in several places of the proof below *} - def after \ "\k. pc[k] = ''p1'' \ + \ \ auxiliary definition that is used in several places of the proof below \ + define after where + "after \ \k. pc[k] = ''p1'' \ pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ pc[k] = ''p3'' \ num[i] \ max[k] \ (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. - -- {* iinv3 and iinv5: particular parts of the invariant, taken to the meta-level for then being instantiated with the proper variables, to ease the work of the classical reasoner. *} + \ \ iinv3 and iinv5: particular parts of the invariant, + taken to the meta-level for then being instantiated with the proper variables, + to ease the work of the classical reasoner. \ from iinv have iinv3: "\i j. - \pc[i] \ {''p5'', ''p6''}; i \ ProcSet; j \ ProcSet \\ unread[i] \\ {i}\ + \pc[i] \ {''p5'', ''p6''}; i \ ProcSet; j \ ProcSet \ unread[i] \ {i}\ \ After(j, i, unread, max, flag, pc, num, nxt)" proof - fix i j assume pci: "pc[i] \ {''p5'', ''p6''}" - and i: "i \ ProcSet" and j:"j \ ProcSet \\ unread[i] \\ {i}" + and i: "i \ ProcSet" and j:"j \ ProcSet \ unread[i] \ {i}" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. hence "pc[i] \ {''p5'', ''p6''} \ - (\j \ ProcSet \\ unread[i] \\ {i} : + (\j \ ProcSet \ unread[i] \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" unfolding IInv_def by auto with pci i j @@ -436,7 +459,7 @@ proof auto from iinv have iinv5: "\i j. - \pc[i] \ {''p7'', ''p8''}; i \ ProcSet; j \ ProcSet \\ {i}\ + \pc[i] \ {''p7'', ''p8''}; i \ ProcSet; j \ ProcSet \ {i}\ \ After(j, i, unread, max, flag, pc, num, nxt)" proof - fix i j @@ -444,7 +467,7 @@ proof auto and i: "i \ ProcSet" and j:"j \ ProcSet \\ {i}" from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. hence "pc[i] \ {''p7'', ''p8''} \ - (\j \ ProcSet \\ {i} : + (\j \ ProcSet \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" unfolding IInv_def by auto with pci i j @@ -452,7 +475,7 @@ proof auto by auto qed - -- {* This also is an instantiation of a type invariant, since auto can't resolve it in a reasonable time. *} + \ \ This also is an instantiation of a type invariant, since auto can't resolve it in a reasonable time. \ from type i have nxti: "pc[i] = ''p6'' \ nxt[i] \ ProcSet \ nxt[i] \ i" by (auto simp: TypeOK_def) @@ -474,210 +497,89 @@ proof auto have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p1_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True with selfi p1 type iinvi self i pc' show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: "proof(..., auto)" doesn't finish in reasonable time **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - thus "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE" - by auto - qed - next - case False with selfi p1 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: "proof(..., auto/force) doesn't finish in reasonable time - and auto / force / blast cannot handle the following in one fell swoop - so we have to resort to very low-level reasoning. **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" by auto - assume "\k \ ProcSet \\ unread[i] \\ {i}: - pc[k] = ''p1'' \ - pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ - pc[k] = ''p3'' \ num[i] \ max[k] \ - (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ - GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" - hence aft: "\k \ ProcSet \\ unread[i] \\ {i}: after(k)" unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - hence "pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - unfolding after_def . - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by simp + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + next + case False with selfi p1 type iinvi self i pc' j show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto qed qed - qed - have 4: "pc'[i] = ''p6'' - \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] - \ (pc'[nxt'[i]] = ''p3'')) - \ max'[nxt'[i]] \ num'[i]" - proof (cases "self = nxt[i]") - case True - with selfi p1 type iinvi self i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p1_def) - next - case False - with type self i have "pc[i] = ''p6'' \ nxt[i] \ ProcSet" - by (auto simp: TypeOK_def) - with False selfi p1 type iinvi self i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p1_def) - qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" - proof (rule+) - fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" - show "After(j,i,unread',max',flag',pc',num',nxt')" - proof (cases "self = j") - case True with selfi p1 type iinvi self i pc' show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: same problems as in 3 above **) - assume "(j = i) = FALSE" "pc[i] = ''p7'' \ pc[i] = ''p8''" - thus "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE" - by auto - qed + have 4: "pc'[i] = ''p6'' + \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] + \ (pc'[nxt'[i]] = ''p3'')) + \ max'[nxt'[i]] \ num'[i]" + proof (cases "self = nxt[i]") + case True + with selfi p1 type iinvi self i show ?thesis + by (clarsimp simp: TypeOK_def IInv_def p1_def) next - case False with selfi p1 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p1_def nat_gt0_not0) - (** FIXME: again, similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" by auto - assume "\k \ ProcSet \\ {i} : - pc[k] = ''p1'' \ - pc[k] = ''p2'' \ (i \ unread[k] \ num[i] \ max[k]) \ - pc[k] = ''p3'' \ num[i] \ max[k] \ - (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ - GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" - hence aft: "\k \ ProcSet \\ {i} : after(k)" unfolding after_def . - with j have "after(j)" by blast - hence "pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - unfolding after_def . - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by simp + case False + with type self i have "pc[i] = ''p6'' \ nxt[i] \ ProcSet" + by (auto simp: TypeOK_def) + with False selfi p1 type iinvi self i show ?thesis + by (clarsimp simp: TypeOK_def IInv_def p1_def) + qed + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + proof (rule+) + fix j + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + show "After(j,i,unread',max',flag',pc',num',nxt')" + proof (cases "self = j") + case True with selfi p1 type iinvi self i pc' show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + next + case False with selfi p1 type iinvi self i pc' j show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto qed qed + from 1 2 3 4 5 show ?thesis + unfolding IInv_def by blast qed - from 1 2 3 4 5 show ?thesis - unfolding IInv_def by blast - qed - next + next assume p2: "p2(self, unread, max, flag, pc, num, nxt, unread', max', flag', pc', num', nxt')" show ?thesis proof (cases "self = i") assume selfi: "self = i" - show ?thesis - proof (cases "unread[self] = {}") - case True - with selfi p2 type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p2_def) - next - case False - from type have "isAFcn(pc)" - (* FIXME: needed for proving that pc'=pc, but why doesn't "simp ..." suffice *) - by (force simp add: TypeOK_def) - with False p2 obtain k where - k: "pc[self] = ''p2''" "k \ unread[self]" - "unread' = [unread EXCEPT ![self] = unread[self] \\ {k}]" - "max' = (IF num[k] > max[self] THEN [max EXCEPT ![self] = num[k]] ELSE max)" - "pc' = pc" "num' = num" "flag' = flag" "nxt' = nxt" - by (auto simp: p2_def) - with selfi type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def) - qed + with selfi p2 type iinvi i show ?thesis + by (auto simp: p2_def TypeOK_def IInv_def) next assume selfi: "self \ i" show ?thesis proof (cases "unread[self] = {}") assume empty: "unread[self] = {}" - with selfi p2 type iinvi self i + from i selfi self p2 type empty + have unch: "pc'[i] = pc[i]" "unread' = unread" "max' = max" "num' = num" "flag' = flag" + by (auto simp: TypeOK_def p2_def) + from empty selfi p2 type iinvi self i have 1: "(num'[i] = 0) = (pc'[i] \ {''p1'', ''p2'', ''p3''})" by (clarsimp simp: TypeOK_def IInv_def p2_def) from empty selfi p2 type iinvi self i have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p2_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" - from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" - by (auto simp: TypeOK_def) + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" + from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" + unfolding IInv_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") - case True with empty selfi p2 type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: again, some manual help necessary **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" - "(j \ unread[i]) = FALSE" "pc[j] = ''p2''" - "unread[j] = {}" - "\j \ ProcSet \\ unread[i] \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ num[i] \ max[j]" - by simp - qed + case True + with empty unch after p2 type j show ?thesis + by (auto simp: p2_def TypeOK_def After_def) next - case False with empty selfi p2 type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i}: after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" "(j \ unread[i]) = FALSE" - with aft have "after(j)" by auto - with ii1 - show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed + case False + with self j p2 type have "pc'[j] = pc[j]" + by (auto simp: TypeOK_def p2_def) + with after unch show ?thesis + by (auto simp: After_def GG_def) qed qed have 4: "pc'[i] = ''p6'' @@ -695,79 +597,39 @@ proof auto with False empty selfi p2 type iinvi self i show ?thesis by (clarsimp simp: TypeOK_def IInv_def p2_def) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" - from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" - by (auto simp: TypeOK_def) + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" + unfolding IInv_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True - with empty selfi p2 type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "unread[j] = {}" - "\j \ ProcSet \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ num[i] \ max[j]" - by simp - qed + with empty unch after p2 type j show ?thesis + by (auto simp: p2_def TypeOK_def After_def) next case False - with empty selfi p2 type iinvi i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def p2_def nat_gt0_not0) - (** FIXME: similar problems **) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i}: - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i}: after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 - show "(pc[i] = ''p1'') = FALSE \ - (pc[i] = ''p2'') = FALSE \ - (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed + with self j p2 type have "pc'[j] = pc[j]" + by (auto simp: TypeOK_def p2_def) + with after unch show ?thesis + by (auto simp: After_def GG_def) qed qed from 1 2 3 4 5 show ?thesis unfolding IInv_def by blast next assume nempty: "unread[self] \ {}" - from type have "isAFcn(pc)" - (* FIXME: needed for proving that pc'=pc, but why doesn't "simp ..." suffice *) - by (force simp: TypeOK_def) - with nempty p2 obtain k where + with p2 type obtain k where k: "pc[self] = ''p2''" "k \ unread[self]" - "unread' = [unread EXCEPT ![self] = unread[self] \\ {k}]" + "unread' = [unread EXCEPT ![self] = unread[self] \ {k}]" "max' = (IF num[k] > max[self] THEN [max EXCEPT ![self] = num[k]] ELSE max)" "pc' = pc" "num' = num" "flag' = flag" "nxt' = nxt" - by (auto simp: p2_def) + by (auto simp: p2_def TypeOK_def) with type self have kproc: "k \ ProcSet" by (auto simp: TypeOK_def) + with type have numk: "num[k] \ Nat" + by (auto simp: TypeOK_def) from k selfi type iinvi self i have 1: "(num'[i] = 0) = (pc'[i] \ {''p1'', ''p2'', ''p3''})" by (clarsimp simp: TypeOK_def IInv_def) @@ -775,10 +637,12 @@ proof auto have 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" + from pc' i k 1 type have numi: "0 < num'[i]" + by (auto simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -790,38 +654,28 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj less type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[k] = ''p5'' \ pc[k] = ''p6''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj less type iinvi self i pc' j mx numi numk show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj less type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj less type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "pc[j] = ''p2''" "j \ ProcSet" + "pc[j] = ''p2''" "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" + "(i \ unread[j]) = FALSE" "\j \ ProcSet \\ unread[i] \\ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence ii2: "i \ unread[j] \ num[i] \ max[j]" + hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" "i \ ProcSet" - "max[j] \ Nat" - with ii2 kproc have "i \ unread[j] \ num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) - with ii1 - show "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ num[k])" - by simp + assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + with ii2 kproc mx i \pc[j] = ''p2''\ + show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed qed next @@ -830,94 +684,37 @@ proof auto proof (cases "i=k") case True with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse nat_not_less[simplified]) - assume "pc[k] = ''p5'' \ pc[k] = ''p6''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + by (auto simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse) + with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" + "j \ ProcSet" "(j \ unread[i]) = FALSE" + "(i \ unread[j]) = FALSE" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ unread[i] \\ {i} : + "\j \ ProcSet \ unread[i] \ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "i \ unread[j] \ num[i] \ max[j]" + thus "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ max[j])" - by simp qed qed qed next assume selfj: "self \ j" - show ?thesis - proof (cases "max[self] < num[k]") - case True - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ - (pc[i] = ''p2'') = FALSE \ - (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - next - case False - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 condElse) - assume "pc[i] = ''p5'' \ pc[i] = ''p6''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ unread[i] \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ unread[i] \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j \ unread[i]) = FALSE" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - qed + with selfi i j self type k + have unr: "unread'[i] = unread[i]" "unread'[j] = unread[j]" + "max'[j] = max[j]" + by (auto simp: TypeOK_def) + with iinvi \pc' = pc\ pc' j + have "After(j,i,unread,max,flag,pc,num,nxt)" + by (auto simp: IInv_def) + with unr \num' = num\ \pc' = pc\ show ?thesis + by (auto simp: After_def GG_def) qed qed have 4: "pc'[i] = ''p6'' @@ -926,21 +723,37 @@ proof auto \ max'[nxt'[i]] \ num'[i]" proof (cases "self = nxt[i]") assume nxt: "self = nxt[i]" - from type k self have mx: "max[self] \ Nat" + from type k kproc self have mx: "max[self] \ Nat" "num[k] \ Nat" by (auto simp: TypeOK_def) show ?thesis proof (cases "max[self] < num[k]") case True with k selfi type iinvi self i nxt mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def, cases "i=k", simp, simp) - assume "num \ [ProcSet \ Nat]" "num[i] \ max[nxt[i]]" "max[nxt[i]] < num[k]" - with mx kproc i nxt show "num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) + proof (clarsimp simp: TypeOK_def IInv_def, cases "i \ unread[nxt[i]]", simp) + assume "i \ unread[nxt[i]]" "num \ [ProcSet \ Nat]" + "i \ unread[nxt[i]] = FALSE \ num[i] \ max[nxt[i]]" + "max[nxt[i]] < num[k]" + with mx nxt kproc i show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed next case False with k selfi type iinvi self i nxt mx show ?thesis - by (clarsimp simp: TypeOK_def IInv_def nat_not_less[simplified]) + proof (clarsimp simp: TypeOK_def IInv_def) + from False mx nxt type + have mxexc: "(IF max[nxt[i]] < num[k] + THEN [max EXCEPT ![nxt[i]] = num[k]] + ELSE max) = max" + by (auto simp: TypeOK_def) + assume "i \ unread[nxt[i]] = FALSE \ num[i] \ max[nxt[i]]" + "i \ unread[nxt[i]] = FALSE \ i = k" + "k \ unread[nxt[i]]" + with mxexc mx False nxt + show "num[i] \ (IF max[nxt[i]] < num[k] + THEN [max EXCEPT ![nxt[i]] = num[k]] + ELSE max)[nxt[i]]" + by auto + qed qed next assume nxt: "self \ nxt[i]" @@ -957,10 +770,12 @@ proof auto by (clarsimp simp: TypeOK_def IInv_def condElse) qed qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" + from pc' i k 1 type have numi: "0 < num'[i]" + by (auto simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -972,38 +787,28 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj less type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[k] = ''p7'' \ pc[k] = ''p8''" - thus "(pc[k] = ''p1'') = FALSE - \ (pc[k] = ''p2'') = FALSE - \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj less type iinvi self i pc' j mx numi numk show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj less type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj less type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence ii2: "i \ unread[j] \ num[i] \ max[j]" + "pc[j] = ''p2''" "j \ ProcSet" + "(j = i) = FALSE" + "(i \ unread[j]) = FALSE" + "\j \ ProcSet \ {i} : + pc[j] = ''p1'' \ + pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ + pc[j] = ''p3'' \ num[i] \ max[j] \ + (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ + GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" + hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" "max[j] \ Nat" - with i kproc ii2 have "i \ unread[j] \ num[i] \ num[k]" - by (auto dest: nat_leq_less_trans) - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ num[k])" - by simp + assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + with ii2 kproc mx i \pc[j] = ''p2''\ + show "num[i] \ num[k]" + by (auto elim!: int_leq_trans) qed qed next @@ -1011,92 +816,38 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj nless type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0 nat_not_less[simplified]) - assume "pc[k] = ''p7'' \ pc[k] = ''p8''" - thus "(pc[k] = ''p1'') = FALSE \ (pc[k] = ''p2'') = FALSE \ (pc[k] = ''p3'') = FALSE" - by auto - qed + with k selfi selfj nless type iinvi self i pc' j mx show ?thesis + by (auto simp: TypeOK_def IInv_def After_def) next case False - with k selfi selfj nless type iinvi i pc' j mx show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) + with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis + proof (clarsimp simp: TypeOK_def IInv_def After_def) assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "j \ ProcSet" "(j = i) = FALSE" "pc[j] = ''p2''" - "\j \ ProcSet \\ {i} : + "j \ ProcSet" + "(i \ unread[j]) = FALSE" + "(j = i) = FALSE" "pc[j] = ''p2''" + "\j \ ProcSet \ {i} : pc[j] = ''p1'' \ pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ pc[j] = ''p3'' \ num[i] \ max[j] \ (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence "i \ unread[j] \ num[i] \ max[j]" + thus "num[i] \ max[j]" by auto - with ii1 show "(pc[i] = ''p1'') = FALSE - \ (pc[i] = ''p2'') = FALSE - \ (pc[i] = ''p3'') = FALSE - \ (i \ unread[j] \ num[i] \ max[j])" - by simp qed qed qed next assume selfj: "self \ j" - show ?thesis - proof (cases "max[self] < num[k]") - case True - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - next - case False - with k selfi selfj type iinvi self i pc' j show ?thesis - proof (clarsimp simp: TypeOK_def IInv_def After_def nat_gt0_not0) - assume "pc[i] = ''p7'' \ pc[i] = ''p8''" - hence ii1: "pc[i] \ ''p1'' \ pc[i] \ ''p2'' \ pc[i] \ ''p3''" - by auto - assume "\j \ ProcSet \\ {i} : - pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" - hence aft: "\j \ ProcSet \\ {i} : after(j)" - unfolding after_def . - assume "j \ ProcSet" "(j = i) = FALSE" - with aft have "after(j)" by auto - with ii1 show - "(pc[i] = ''p1'') = FALSE \ (pc[i] = ''p2'') = FALSE \ (pc[i] = ''p3'') = FALSE \ - (pc[j] = ''p1'' \ - pc[j] = ''p2'' \ (i \ unread[j] \ num[i] \ max[j]) \ - pc[j] = ''p3'' \ num[i] \ max[j] \ - (pc[j] = ''p4'' \ pc[j] = ''p5'' \ pc[j] = ''p6'') \ - GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j]))" - by (simp add: after_def) - qed - qed + with selfi i j self type k + have unr: "unread'[i] = unread[i]" "unread'[j] = unread[j]" + "max'[j] = max[j]" + by (auto simp: TypeOK_def) + with iinvi \pc' = pc\ pc' j + have "After(j,i,unread,max,flag,pc,num,nxt)" + by (auto simp: IInv_def) + with unr \num' = num\ \pc' = pc\ show ?thesis + by (auto simp: After_def GG_def) qed qed from 1 2 3 4 5 show ?thesis @@ -1109,7 +860,7 @@ proof auto proof (cases "self = i") case True with p3 type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p3_def) auto + by (clarsimp simp: TypeOK_def IInv_def p3_def) next assume selfi: "self \ i" from selfi p3 type iinvi self i @@ -1117,7 +868,7 @@ proof auto and 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p3_def)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" @@ -1125,7 +876,6 @@ proof auto proof (cases "self = j") case True with selfi type p3 self i j iinv3[of i j] pc' show ?thesis by (clarsimp simp: TypeOK_def p3_def After_def GG_def) auto - (** a bit slow, but it's not necessary to do case analysis on IF **) next case False with selfi p3 type self i j iinv3[of i j] pc' show ?thesis @@ -1177,10 +927,10 @@ proof auto and 2: "flag'[i] = (pc'[i] \ {''p2'', ''p3'', ''p4''})" by (clarsimp simp: TypeOK_def IInv_def p4_def)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" with selfi type p4 self i iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" unfolding TypeOK_def p4_def After_def GG_def @@ -1199,10 +949,10 @@ proof auto with selfi p4 type iinvi self i nxti show ?thesis by (clarsimp simp: TypeOK_def IInv_def p4_def) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" with selfi type p4 self i j pc' iinv5[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" by (cases "self = j") (clarsimp simp: TypeOK_def p4_def After_def)+ @@ -1224,10 +974,10 @@ proof auto unfolding TypeOK_def IInv_def p5_def by(cases "self = i", clarsimp+) have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" with empty type p5 self i j iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" unfolding TypeOK_def After_def GG_def p5_def @@ -1245,19 +995,16 @@ proof auto by (cases "self = nxt[i]", (clarsimp simp: nxti)+) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True with empty type p5 self i j iinv3[of i j] (** IInv part 3 is used, not part 5! **) show ?thesis - apply (clarsimp simp: IInv_def TypeOK_def p5_def After_def) - apply (cases "pc[j] = ''p2''", clarsimp) - apply (cases "pc[j] = ''p3''", clarsimp+) - done + by (clarsimp simp: IInv_def TypeOK_def p5_def After_def) auto next from empty have unreadj: "j \ unread[self]" by auto @@ -1282,18 +1029,28 @@ proof auto unfolding TypeOK_def IInv_def p5_def by(cases "self = i") clarsimp+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" - with empty type p5 self i j iinv3[of i j] show "After(j,i,unread',max',flag',pc',num',nxt')" - unfolding TypeOK_def After_def GG_def p5_def - apply(cases "self = j", clarsimp) - apply (cases "self = i", clarsimp) - apply (cases "pc[j] = ''p2''", clarsimp) - apply (cases "pc[j] = ''p3''", clarsimp+) - done + proof (cases "self=j") + case True + with empty type p5 self i j pc' iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) auto + next + case False + show ?thesis + proof (cases "self = i") + case True + with empty type p5 self i j pc' iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) auto + next + case False + with empty type p5 self i j pc' \self \ j\ iinv3[of i j] show ?thesis + by (clarsimp simp: TypeOK_def After_def GG_def p5_def) + qed + qed qed have 4: "pc'[i] = ''p6'' \ ( (pc'[nxt'[i]] = ''p2'') \ i \ unread'[nxt'[i]] @@ -1326,10 +1083,10 @@ proof auto by (cases "self = nxt[i]") clarsimp+ qed qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) from j iinv have iinvj: "IInv(j, unread, max, flag, pc, num, nxt)" by auto @@ -1349,10 +1106,10 @@ proof auto unfolding TypeOK_def IInv_def p6_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof(rule+) fix j - assume pc': "pc'[i] \ {''p5'', ''p6''}" and j: "j \ ProcSet \\ unread'[i] \\ {i}" + assume pc': "pc'[i] \ {''p5'', ''p6''}" and j: "j \ ProcSet \ unread'[i] \ {i}" from iinv j have iinvj: "IInv(j, unread, max, flag, pc, num, nxt)" by auto from type j @@ -1361,16 +1118,17 @@ proof auto show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = i") assume selfi: "self = i" + with iinvi p6 type i have numi: "0 < num[i]" "nxt[i] \ ProcSet" "nxt[i] \ i" + by (auto simp: IInv_def TypeOK_def p6_def) show ?thesis proof (cases "j \ unread[self]") assume junread: "j \ unread[self]" show ?thesis proof(cases "num[j] = 0") case True - with j selfi i pc' p6 type self iinvi junread iinvj (* FIXME: iinvj *) + with j selfi i pc' p6 type self iinvi junread iinvj show ?thesis - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def nat_gt0_not0) - by (cases "pc[nxt[i]] = ''p2''", clarsimp+) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def) auto next assume numj: "num[j] \ 0" show ?thesis @@ -1381,30 +1139,49 @@ proof auto case True with j selfi i pc' p6 type self iinvi junread numj pcj show ?thesis - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) auto next assume notunread: "i \ unread[j]" show ?thesis proof (cases "j < i") case True - from True j selfi i pc' p6 type self iinvi junread numj pcj iinv3[of j i] nxti notunread + with j selfi i pc' p6 type self iinvi junread numj numi pcj + iinv3[of j i] nxti notunread (*procInNat*) show ?thesis - using procInNat nat_less_antisym_false[of j i] - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) (force simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "nxt[i] < i" "j = nxt[i]" + with i j have nxti: "nxt[i] \ ProcSet" "\(i < nxt[i])" + by (auto dest: procInNat elim!: int_less_asym) + assume "num \ [ProcSet \ Nat]" + "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num[i] < num[nxt[i]]" + with i nxti + show "(pc[nxt[i]] = ''p5'') = FALSE + \ (pc[nxt[i]] = ''p6'') = FALSE" + by (auto simp: int_leq_not_less) + qed next case False - with j selfi i pc' p6 type self iinvi junread numj pcj iinv3[of j i] nxti notunread + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinv3[of j i] nxti notunread show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - proof (clarsimp simp: nat_gt0_not0 nat_not_less[simplified] nat_less_not_sym procInNat - dest!: leq_neq_trans'[of "i" "nxt[i]", simplified]) - assume - "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ num[nxt[i]] < num[i]" - "num[i] \ num[nxt[i]]" "nxt[i] \ ProcSet" "i \ ProcSet" - "num \ [ProcSet \ Nat]" - thus "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" - by (auto simp: nat_less_antisym_leq_false) - qed + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "(nxt[i] < i) = FALSE" + with i numi int_less_linear[of "i" "nxt[i]"] have "i < nxt[i]" + by (auto dest: procInNat) + assume + "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num[i] \ num[nxt[i]]" + "num \ [ProcSet \ Nat]" + with i numi \i < nxt[i]\ + show "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" + by (auto dest: procInNat simp: int_leq_not_less) + qed qed qed next @@ -1415,25 +1192,50 @@ proof auto show ?thesis proof (cases "j < i") case True - with j selfi i pc' p6 type self iinvi junread numj pcj iinvj nxti iinv5[of "nxt[i]" i] + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinvj nxti iinv5[of "nxt[i]" i] show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - using nat_less_antisym_false[of j i] - apply (clarsimp simp: nat_gt0_not0 procInNat) - by (clarsimp simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume nxti: "nxt[i] < i" + hence "i \ nxt[i]" by (auto simp: less_def) + assume + "i = nxt[i] = FALSE + \ 0 < num[nxt[i]] + \ (IF i < nxt[i] + THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i]))" + "num \ [ProcSet \ Nat]" + with \i \ nxt[i]\ have "num[nxt[i]] \ num[i]" + by (auto dest: procInNat elim!: int_less_asym[OF nxti]) + assume "num[i] < num[nxt[i]]" + with i numi \num \ [ProcSet \ Nat]\ \num[nxt[i]] \ num[i]\ + show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') + \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" + by (auto simp: int_less_not_leq) + qed next case False - with j selfi i pc' p6 type self iinvi junread numj pcj iinv5[of j i] nxti + with j selfi i pc' p6 type self iinvi junread numi numj pcj iinv5[of j i] nxti + int_less_linear[of "nxt[i]" "i"] show ?thesis - unfolding p6_def TypeOK_def IInv_def After_def GG_def - by (clarsimp simp: nat_gt0_not0 nat_not_less[simplified] procInNat) - (clarsimp dest!: leq_neq_trans'[simplified] simp: nat_less_antisym_leq_false nat_less_not_sym) + proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + assume "(nxt[i] < i) = FALSE" + with i numi int_less_linear[of "nxt[i]" "i"] have "i < nxt[i]" + by (auto dest: procInNat) + assume "num[i] \ num[nxt[i]]" + "IF i < nxt[i] THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i])" + "num \ [ProcSet \ Nat]" + with i numi \i < nxt[i]\ + show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') + \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" + by (auto simp: int_less_not_leq) + qed qed next assume pcj2: "pc[j] \ {''p7'', ''p8''}" with j selfi i pc' p6 type self iinvi junread numj pcj iinvj nxti show ?thesis - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) + apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) by (erule funcSetE[where f="pc" and x="nxt[i]"], simp+) qed qed @@ -1453,18 +1255,32 @@ proof auto show ?thesis proof (cases "j < nxt[j]") case True - with j pc' selfi selfj i p6 type self iinv3[of i j] + with j pc' selfi selfj i p6 type self iinv3[of i j] nxtj show ?thesis - apply (clarsimp simp: p6_def TypeOK_def After_def GG_def nat_gt0_not0) - by (auto dest: nat_less_trans nat_less_leq_trans) - (** slow **) + by (clarsimp simp: p6_def TypeOK_def After_def GG_def) + (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] + elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False with j pc' selfi selfj i p6 type self iinv3[of i j] show ?thesis - using nxtj nat_not_less procInNat - by (clarsimp simp: p6_def TypeOK_def After_def GG_def nat_gt0_not0 leq_neq_iff_less[simplified]) - (auto simp: nat_less_antisym_leq_false) + proof (clarsimp simp: p6_def TypeOK_def After_def GG_def) + assume + "pc[j] = ''p6''" + "(j < nxt[j]) = FALSE" + with i j nxtj int_less_linear[of "j" "nxt[j]"] have "nxt[j] < j" + by (force dest: procInNat) + + assume + "num[nxt[j]] = 0 + \ (IF nxt[j] < j THEN num[j] < num[nxt[j]] ELSE num[j] \ num[nxt[j]])" + "0 < num[i]" + "IF j < i THEN num[i] < num[j] ELSE num[i] \ num[j]" + "num \ [ProcSet \ Nat]" + with i j \nxt[j] < j\ show "(i = nxt[j]) = FALSE" + by (auto simp: int_less_not_leq[of "num[j]" "num[nxt[j]]"] + elim!: int_less_asym[of "num[nxt[j]]" "num[j]"]) + qed qed next case False @@ -1488,10 +1304,10 @@ proof auto apply (clarsimp simp: TypeOK_def IInv_def p6_def) by (cases "self = i", simp+) qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True @@ -1501,24 +1317,24 @@ proof auto case True with type p6 i pc' j self selfj iinv5[of i j] show ?thesis - using nat_less_antisym_false[of j i] procInNat - apply (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) - by (cases "i = nxt[j]") (auto simp: nat_less_antisym_leq_false) + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] + elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False from i j type have num: "num[i] \ Nat" "num[j] \ Nat" by (auto simp: TypeOK_def) from False type p6 i pc' j self selfj iinv5[of i j] show ?thesis - using nat_less_linear[of i j] procInNat - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def nat_gt0_not0) - (auto dest!: nat_leq_less_trans) + using int_less_linear[of i j] procInNat + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) + (auto dest!: int_leq_less_trans) qed next case False with type iinv5[of i j] p6 i pc' j show ?thesis apply (clarsimp simp: TypeOK_def IInv_def After_def p6_def) - by (cases "self = i", simp, simp) + by (cases "self = i", simp+) qed qed from 1 2 3 4 5 show ?thesis @@ -1532,10 +1348,10 @@ proof auto unfolding TypeOK_def IInv_def p7_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \\ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True @@ -1564,10 +1380,10 @@ proof auto unfolding TypeOK_def IInv_def p7_def by (cases "self = nxt[i]") clarsimp+ qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = j") case True @@ -1590,10 +1406,10 @@ proof auto unfolding p8_def TypeOK_def IInv_def by (cases "self = i", clarsimp, clarsimp)+ have 3: "pc'[i] \ {''p5'', ''p6''} \ - (\j \ (P \\ unread'[i]) \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + (\j \ (P \ unread'[i]) \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \\ unread'[i]) \\ {i}" + assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = i") case True @@ -1622,10 +1438,10 @@ proof auto unfolding p8_def TypeOK_def IInv_def by (cases "self = nxt[i]") clarsimp+ qed - have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \\ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" + have 5: "(pc'[i] \ {''p7'', ''p8''}) \ (\j \ P \ {i} : After(j,i,unread',max',flag',pc',num',nxt'))" proof (rule+) fix j - assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \\ {i}" + assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = i") case True diff --git a/isabelle/examples/ROOT b/isabelle/examples/ROOT new file mode 100644 index 00000000..cbcd3fa6 --- /dev/null +++ b/isabelle/examples/ROOT @@ -0,0 +1,18 @@ +(* Title: TLA+/examples/ROOT + Author: Stephan Merz, Inria Nancy + License: BSD + +Examples of using Isabelle/TLA+ +*) + +(* build the session (without a heap) using "isabelle build -D ." *) +chapter "TLA+ Examples" +session "TLA-Examples" = "TLA+" + + options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + + theories + Allocator + AtomicBakeryG + + document_files + "root.tex" diff --git a/isabelle/examples/ROOT.ML b/isabelle/examples/ROOT.ML deleted file mode 100644 index 6812d8d3..00000000 --- a/isabelle/examples/ROOT.ML +++ /dev/null @@ -1,13 +0,0 @@ -(* Title: TLA+/examples/ROOT.ML - Author: Stephan Merz, LORIA - License: BSD - -Examples of using Isabelle/TLA+ -*) - - -val banner = "TLA+ Examples"; -writeln banner; - -use_thy "Allocator"; -use_thy "AtomicBakeryG"; From 21bbf9f9adc24cdd17a4092d23a490547d88fa12 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 17:20:09 +0200 Subject: [PATCH 125/167] Make the new isabelle to work with the dune-based build. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 19 +++++++++---------- src/isabelle_keywords.ml | 1 - src/params.ml | 32 ++++++++++++++++++++++++++------ 3 files changed, 35 insertions(+), 17 deletions(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index aa78d541..73a6341f 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -1,19 +1,20 @@ OS_TYPE=$(patsubst CYGWIN%,Cygwin,$(shell uname)) HOST_CPU=$(shell uname -m) -ISABELLE_VSN=Isabelle2011-1 +ISABELLE_VSN=Isabelle2020 ifeq ($(OS_TYPE),Linux) - ISABELLE_ARCHIVE=$(ISABELLE_VSN)_bundle_x86-linux.tar.gz + ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz ISABELLE_ARCHIVE_TYPE=tgz FIND_EXEC=-executable endif ifeq ($(OS_TYPE),Darwin) - ISABELLE_ARCHIVE=$(ISABELLE_VSN).dmg + ISABELLE_ARCHIVE=$(ISABELLE_VSN)_macos.tar.gz ISABELLE_ARCHIVE_TYPE=dmg FIND_EXEC=-perm +111 endif ifeq ($(OS_TYPE),Cygwin) + # TODO: Fix this. ISABELLE_ARCHIVE=$(ISABELLE_VSN)_bundle_x86-cygwin.tar.gz ISABELLE_ARCHIVE_TYPE=tgz FIND_EXEC=-executable @@ -66,10 +67,9 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && echo "ISABELLE_OUTPUT=$$HEAPS_PATH" >> etc/settings mkdir -p $(ISABELLE_DIR)/src/TLA+ \ && cp -a theories/TLA+/* $(ISABELLE_DIR)/src/TLA+/ - cd $(ISABELLE_DIR)/src/Pure \ - && ../../bin/isabelle make - cd $(ISABELLE_DIR)/src/TLA+ \ - && ../../bin/isabelle usedir -b Pure TLA+ + cd $(ISABELLE_DIR)/ \ + && ./bin/isabelle build -b -v -d src/Pure -o system_heaps Pure \ + && ./bin/isabelle build -b -v -d src/TLA+ -o system_heaps TLA+ cd $(ISABELLE_DIR) \ && rm etc/settings \ && mv etc/settings.target etc/settings @@ -79,9 +79,8 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) .PRECIOUS: $(ISABELLE_DIR).no-links $(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ rm -rf $@ - cd $< && rm -f ./contrib/jre1.6.0_27_x86-linux/jre1.6.0_27/man/ja - cd $< && rm -f ./contrib/polyml-5.4.0/src - cd $< && rm -f ./contrib/polyml && mv ./contrib/polyml-5.4.0 ./contrib/polyml + cd $< && rm -f ./contrib/jdk-11.0.5+10/x86_64-linux/jre + cd $< && rm -f ./contrib/jdk-11.0.5+10/x86_64-linux/man/ja touch $@ # TODO: This is a workaround, because the dune install removes all the executable diff --git a/src/isabelle_keywords.ml b/src/isabelle_keywords.ml index 4a3f13fc..b58ddeeb 100644 --- a/src/isabelle_keywords.ml +++ b/src/isabelle_keywords.ml @@ -1,5 +1,4 @@ (* AUTOMATICALLY GENERATED by `update_isabelle_keywords.py` - DO NOT EDIT *) -Revision.f "$Rev$";; let v = [ "ML"; "ML_command"; diff --git a/src/params.ml b/src/params.ml index c98f136f..a1649bc9 100644 --- a/src/params.ml +++ b/src/params.ml @@ -61,12 +61,32 @@ let library_path = let d = Filename.concat d "tlaps" in d +(* The backends directory should be resolved via the dune sites mechanism. + If the sites are not available, then we assume this file layout to locate + the Isabelle installation. + - bin/tlapm + - lib/tlapm/backends/Isabelle + *) +let isabelle_base_path = + let rec find paths = + match paths with + | [] -> + let bin_dir = Filename.dirname Sys.executable_name in + let prefix = Filename.dirname bin_dir in + List.fold_left + Filename.concat + prefix + ["lib"; "tlapm"; "backends"; "Isabelle"] + | path :: other -> + let isabelle_base = Filename.concat path "Isabelle" in + match Sys.file_exists isabelle_base with + | true -> isabelle_base + | false -> find other + in find Setup_paths.Sites.backends + let isabelle_tla_path = - let d = Sys.executable_name in - let d = Filename.dirname (Filename.dirname d) in - let d = Filename.concat d "lib" in - let d = Filename.concat d "isabelle_tla" in - d + List.fold_left Filename.concat isabelle_base_path ["src"; "TLA+"] + type executable = | Unchecked of string * string * string (* exec, command, version_command *) @@ -155,7 +175,7 @@ let isabelle = isabelle_success_string isabelle_tla_path in - make_exec "isabelle process" cmd "isabelle version" + make_exec "isabelle" cmd "isabelle version" ;; let set_fast_isabelle () = From f270dd4ee208fde0b3b4f32868f718178843b3e8 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 18:11:32 +0200 Subject: [PATCH 126/167] Use tar.gz archive of Isabelle on MacOS. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 9 +-------- 1 file changed, 1 insertion(+), 8 deletions(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 73a6341f..3fe3b692 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -10,7 +10,7 @@ ifeq ($(OS_TYPE),Linux) endif ifeq ($(OS_TYPE),Darwin) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_macos.tar.gz - ISABELLE_ARCHIVE_TYPE=dmg + ISABELLE_ARCHIVE_TYPE=tgz FIND_EXEC=-perm +111 endif ifeq ($(OS_TYPE),Cygwin) @@ -47,13 +47,6 @@ ifeq ($(ISABELLE_ARCHIVE_TYPE),tgz) tar -xzf $< mv $(ISABELLE_VSN) $(ISABELLE_DIR) endif -ifeq ($(ISABELLE_ARCHIVE_TYPE),dmg) - rm -rf isabelle-dmg - hdiutil attach "$<" -quiet -readonly -nobrowse -mountpoint "$(shell pwd)/isabelle-dmg" - cp -a "isabelle-dmg/$(ISABELLE_VSN).app/Contents/Resources/$(ISABELLE_VSN)" "$@" - hdiutil detach "$(shell pwd)/isabelle-dmg" -quiet - rm -rf isabelle-dmg -endif # Build the TLA+ theory. .PRECIOUS: $(ISABELLE_DIR)/src/TLA+ From 2c5b99d8e4af060a349b5ef0000dfb9943632f68 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 21:00:51 +0200 Subject: [PATCH 127/167] It works with Isabelle2023, except 1 test. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 8 ++++---- deps/ls4/Makefile | 2 +- deps/z3/Makefile | 2 +- isabelle/.gitignore | 1 + isabelle/Makefile | 8 ++++---- 5 files changed, 11 insertions(+), 10 deletions(-) create mode 100644 isabelle/.gitignore diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 3fe3b692..a11c16e0 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -1,7 +1,7 @@ OS_TYPE=$(patsubst CYGWIN%,Cygwin,$(shell uname)) HOST_CPU=$(shell uname -m) -ISABELLE_VSN=Isabelle2020 +ISABELLE_VSN=Isabelle2023 ifeq ($(OS_TYPE),Linux) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz @@ -59,7 +59,8 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && cp etc/settings etc/settings.target \ && echo "ISABELLE_OUTPUT=$$HEAPS_PATH" >> etc/settings mkdir -p $(ISABELLE_DIR)/src/TLA+ \ - && cp -a theories/TLA+/* $(ISABELLE_DIR)/src/TLA+/ + && cp -a theories/TLA+/* $(ISABELLE_DIR)/src/TLA+/ \ + && chmod -R u+w $(ISABELLE_DIR)/src/TLA+/ cd $(ISABELLE_DIR)/ \ && ./bin/isabelle build -b -v -d src/Pure -o system_heaps Pure \ && ./bin/isabelle build -b -v -d src/TLA+ -o system_heaps TLA+ @@ -72,8 +73,7 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) .PRECIOUS: $(ISABELLE_DIR).no-links $(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ rm -rf $@ - cd $< && rm -f ./contrib/jdk-11.0.5+10/x86_64-linux/jre - cd $< && rm -f ./contrib/jdk-11.0.5+10/x86_64-linux/man/ja + cd $< && rm -rf ./contrib/e-2.6-1/src/lib/ touch $@ # TODO: This is a workaround, because the dune install removes all the executable diff --git a/deps/ls4/Makefile b/deps/ls4/Makefile index a4df8aab..ddfe1324 100644 --- a/deps/ls4/Makefile +++ b/deps/ls4/Makefile @@ -17,7 +17,7 @@ endif all: ls4 v$(LS4_VERS).zip: - wget $(LS4_URL_PREFIX)/$@ + wget -q $(LS4_URL_PREFIX)/$@ $(LS4_BASE_NAME): v$(LS4_VERS).zip rm -rf $@ diff --git a/deps/z3/Makefile b/deps/z3/Makefile index 91918cec..04005463 100644 --- a/deps/z3/Makefile +++ b/deps/z3/Makefile @@ -19,7 +19,7 @@ endif all: z3 $(Z3_BASE_NAME).zip: - wget $(Z3_URL_PREFIX)/$@ + wget -q $(Z3_URL_PREFIX)/$@ $(Z3_BASE_NAME): $(Z3_BASE_NAME).zip rm -rf $@ diff --git a/isabelle/.gitignore b/isabelle/.gitignore new file mode 100644 index 00000000..ea1472ec --- /dev/null +++ b/isabelle/.gitignore @@ -0,0 +1 @@ +output/ diff --git a/isabelle/Makefile b/isabelle/Makefile index cd752015..0c44755a 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -1,4 +1,4 @@ -.PHONY: rebuild heap-only +.PHONY: rebuild clean # Generate theory information (in HTML and PDF formats) as well as the main # heap file for the TLA+ object logic. By default these will be generated @@ -7,7 +7,7 @@ # file etc/settings beneath the Isabelle home (i.e., where the Isabelle # distribution is installed). rebuild: - isabelle usedir -b -i true -d pdf Pure TLA+ + isabelle build -b -v -D . -heap-only: - isabelle build -v -d . -b TLA+ +clean: + rm -rf output/ From 37ba3124fc015b4b3c21ff21cecd1afb35f43d37 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 21:19:57 +0200 Subject: [PATCH 128/167] `lualatex` is needed to build isabelle theory. Signed-off-by: Karolis Petrauskas --- .github/workflows/main.yml | 2 +- .github/workflows/pr.yml | 2 +- .github/workflows/release.yml | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/main.yml b/.github/workflows/main.yml index c01dc056..e1f90e2d 100644 --- a/.github/workflows/main.yml +++ b/.github/workflows/main.yml @@ -22,7 +22,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time + sudo apt-get install --yes time texlive-latex-base - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/pr.yml b/.github/workflows/pr.yml index baafc9af..351286db 100644 --- a/.github/workflows/pr.yml +++ b/.github/workflows/pr.yml @@ -26,7 +26,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time + sudo apt-get install --yes time texlive-latex-base - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 5de22a38..e2229b3b 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time + sudo apt-get install --yes time texlive-latex-base - name: Set up Python uses: actions/setup-python@v2 with: From 13bf0e35efef017f1366cfee9f542e0d72d55696 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 21:22:57 +0200 Subject: [PATCH 129/167] Attempt to fix the macos build. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index a11c16e0..fe7d2677 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -6,17 +6,20 @@ ISABELLE_VSN=Isabelle2023 ifeq ($(OS_TYPE),Linux) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz ISABELLE_ARCHIVE_TYPE=tgz + ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN) FIND_EXEC=-executable endif ifeq ($(OS_TYPE),Darwin) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_macos.tar.gz ISABELLE_ARCHIVE_TYPE=tgz + ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN).app FIND_EXEC=-perm +111 endif ifeq ($(OS_TYPE),Cygwin) # TODO: Fix this. ISABELLE_ARCHIVE=$(ISABELLE_VSN)_bundle_x86-cygwin.tar.gz ISABELLE_ARCHIVE_TYPE=tgz + ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN) FIND_EXEC=-executable endif @@ -45,7 +48,7 @@ $(ISABELLE_DIR): $(ISABELLE_ARCHIVE) ifeq ($(ISABELLE_ARCHIVE_TYPE),tgz) rm -rf $(ISABELLE_DIR) tar -xzf $< - mv $(ISABELLE_VSN) $(ISABELLE_DIR) + mv $(ISABELLE_ARCHIVE_DIR) $(ISABELLE_DIR) endif # Build the TLA+ theory. From 280559e82fe63d05b2b9095f74f60a965bfd6a60 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 21:46:16 +0200 Subject: [PATCH 130/167] Add missing latex packages. Signed-off-by: Karolis Petrauskas --- .github/workflows/main.yml | 2 +- .github/workflows/pr.yml | 2 +- .github/workflows/release.yml | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/main.yml b/.github/workflows/main.yml index e1f90e2d..8ff9c011 100644 --- a/.github/workflows/main.yml +++ b/.github/workflows/main.yml @@ -22,7 +22,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base + sudo apt-get install --yes time texlive-latex-base texlive-plain-generic - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/pr.yml b/.github/workflows/pr.yml index 351286db..f2422665 100644 --- a/.github/workflows/pr.yml +++ b/.github/workflows/pr.yml @@ -26,7 +26,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base + sudo apt-get install --yes time texlive-latex-base texlive-plain-generic - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index e2229b3b..22a3c6ff 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base + sudo apt-get install --yes time texlive-latex-base texlive-plain-generic - name: Set up Python uses: actions/setup-python@v2 with: From ce9a64bb7263bf36aeef8019e74697e153a11770 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 23 Dec 2023 23:00:55 +0200 Subject: [PATCH 131/167] Do not build the Isabelle document when building heaps. Signed-off-by: Karolis Petrauskas --- .github/workflows/main.yml | 2 +- .github/workflows/pr.yml | 2 +- .github/workflows/release.yml | 2 +- deps/isabelle/Makefile | 4 ++-- isabelle/Makefile | 7 +++++-- isabelle/ROOT | 2 +- 6 files changed, 11 insertions(+), 8 deletions(-) diff --git a/.github/workflows/main.yml b/.github/workflows/main.yml index 8ff9c011..c01dc056 100644 --- a/.github/workflows/main.yml +++ b/.github/workflows/main.yml @@ -22,7 +22,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-plain-generic + sudo apt-get install --yes time - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/pr.yml b/.github/workflows/pr.yml index f2422665..baafc9af 100644 --- a/.github/workflows/pr.yml +++ b/.github/workflows/pr.yml @@ -26,7 +26,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-plain-generic + sudo apt-get install --yes time - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 22a3c6ff..5de22a38 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-plain-generic + sudo apt-get install --yes time - name: Set up Python uses: actions/setup-python@v2 with: diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index fe7d2677..ba9f2144 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -65,8 +65,8 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && cp -a theories/TLA+/* $(ISABELLE_DIR)/src/TLA+/ \ && chmod -R u+w $(ISABELLE_DIR)/src/TLA+/ cd $(ISABELLE_DIR)/ \ - && ./bin/isabelle build -b -v -d src/Pure -o system_heaps Pure \ - && ./bin/isabelle build -b -v -d src/TLA+ -o system_heaps TLA+ + && ./bin/isabelle build -o system_heaps -o document=false -b -v -d src/Pure Pure \ + && ./bin/isabelle build -o system_heaps -o document=false -b -c -v -d src/TLA+ TLA+ cd $(ISABELLE_DIR) \ && rm etc/settings \ && mv etc/settings.target etc/settings diff --git a/isabelle/Makefile b/isabelle/Makefile index 0c44755a..6421e79b 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -1,4 +1,4 @@ -.PHONY: rebuild clean +.PHONY: rebuild heap-only clean # Generate theory information (in HTML and PDF formats) as well as the main # heap file for the TLA+ object logic. By default these will be generated @@ -7,7 +7,10 @@ # file etc/settings beneath the Isabelle home (i.e., where the Isabelle # distribution is installed). rebuild: - isabelle build -b -v -D . + isabelle build -o document -o browser_info -b -c -v -D . + +heap-only: + isabelle build -o document=false -o browser_info=false -b -c -v -D . clean: rm -rf output/ diff --git a/isabelle/ROOT b/isabelle/ROOT index 30960a90..c4124c97 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,7 +1,7 @@ (* build the session using "isabelle build -b -D ." *) chapter "TLA+" session "TLA+" = "Pure" + - options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + options [document_output = "output", document_variants="document:outline=/proof"] theories Constant (global) From 512d89d8141b29bff94a82bf4e5335f0b9ffd481 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 24 Dec 2023 00:01:49 +0200 Subject: [PATCH 132/167] Try to fix the MacOS build. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index ba9f2144..d05a6fbc 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -55,7 +55,7 @@ endif .PRECIOUS: $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) cd $(ISABELLE_DIR) \ - && rm -rf contrib/ProofGeneral* doc heaps/*/HOL lib/{classes,fonts,logo} + && rm -rf contrib/ProofGeneral* doc heaps/*/HOL cd $(ISABELLE_DIR) \ && HEAPS_PATH=$(shell pwd)/$(ISABELLE_DIR)/heaps \ && if [ "$(HOST_CPU)" = "x86_64" ] ; then sed -i -e 's/^ML_PLATFORM=.*$$/ML_PLATFORM="$${ISABELLE_PLATFORM64:-$$ISABELLE_PLATFORM}"/' etc/settings ; fi \ From 4a881d2812d64381b6e83fde8882113a5a6dc676 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 24 Dec 2023 00:35:44 +0200 Subject: [PATCH 133/167] Make build more portable. Signed-off-by: Karolis Petrauskas --- deps/Makefile.post-install | 2 +- deps/isabelle/Makefile | 9 ++++----- deps/isabelle/dune | 6 +++--- 3 files changed, 8 insertions(+), 9 deletions(-) diff --git a/deps/Makefile.post-install b/deps/Makefile.post-install index 700a9ebe..934b8608 100644 --- a/deps/Makefile.post-install +++ b/deps/Makefile.post-install @@ -1,3 +1,3 @@ all: chmod +x backends/bin/* - make -C backends -f Isabelle.post-install + cd backends && cat Isabelle.exec-files | xargs chmod +x diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index d05a6fbc..37c10717 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -31,7 +31,7 @@ PROJECT_ROOT=../.. CACHE_DIR=$(PROJECT_ROOT)/_build_cache -all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR).no-links Isabelle.post-install +all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR).no-links Isabelle.exec-files # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): @@ -81,11 +81,10 @@ $(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ # TODO: This is a workaround, because the dune install removes all the executable # flags (or sets on all the files). Here we generate a script to restore the flags. -Isabelle.post-install: $(ISABELLE_DIR).no-links - echo "FILES=$(shell find $(ISABELLE_DIR) -type f $(FIND_EXEC))" > $@ - echo "all:\n\t chmod +x \$$(FILES)" >> $@ +Isabelle.exec-files: $(ISABELLE_DIR).no-links + echo "$(shell find $(ISABELLE_DIR) -type f $(FIND_EXEC))" > $@ clean: - rm -rf $(ISABELLE_ARCHIVE) $(ISABELLE_DIR) $(ISABELLE_DIR).no-links Isabelle.post-install + rm -rf $(ISABELLE_ARCHIVE) $(ISABELLE_DIR) $(ISABELLE_DIR).no-links Isabelle.exec-files .PHONY: all clean diff --git a/deps/isabelle/dune b/deps/isabelle/dune index 963365f8..12ad3ce8 100644 --- a/deps/isabelle/dune +++ b/deps/isabelle/dune @@ -3,8 +3,8 @@ ; The generated heaps (Pure and TLA+) are in Isabelle/heaps/polyml-*/. (rule (deps "Makefile" (source_tree theories)) - (targets (dir Isabelle) "Isabelle.no-links" "Isabelle.post-install") - (action (run "make" "-C" "." "PROJECT_ROOT=%{project_root}" "Isabelle.no-links" "Isabelle.post-install"))) + (targets (dir Isabelle) "Isabelle.no-links" "Isabelle.exec-files") + (action (run "make" "-C" "." "PROJECT_ROOT=%{project_root}" "Isabelle.no-links" "Isabelle.exec-files"))) (install (section (site (tlapm backends))) @@ -12,4 +12,4 @@ (install (section (site (tlapm backends))) - (files ("Isabelle.post-install" as "Isabelle.post-install"))) + (files ("Isabelle.exec-files" as "Isabelle.exec-files"))) From 4159f54d16b2840dfa0f1e4df41e636443862a4f Mon Sep 17 00:00:00 2001 From: Dominykas Bartkus Date: Thu, 28 Dec 2023 20:56:04 +0200 Subject: [PATCH 134/167] Make macos build work again Signed-off-by: Dominykas Bartkus --- deps/isabelle/Makefile | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 37c10717..8956a1b4 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -77,6 +77,15 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) $(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ rm -rf $@ cd $< && rm -rf ./contrib/e-2.6-1/src/lib/ +ifeq ($(OS_TYPE),Darwin) + cd $< && cd contrib/jdk-17.0.7/arm64-darwin/ \ + && (find . -type link | xargs rm) \ + && mv zulu-17.jdk/Contents/Home/* ./ + cd $< && cd contrib/jdk-17.0.7/x86_64-darwin/ \ + && (find . -type link | xargs rm) \ + && mv zulu-17.jdk/Contents/Home/* ./ + cd $< && rm -rf contrib/vscodium-1.70.1 +endif touch $@ # TODO: This is a workaround, because the dune install removes all the executable From 4c4fc8fc86ffc64093fc95fabe16d8bd5e2110e3 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 13 Jan 2024 17:59:22 +0200 Subject: [PATCH 135/167] Use HTTPS to download the Isabelle/HOL. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 8956a1b4..61dfa01f 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -23,7 +23,7 @@ ifeq ($(OS_TYPE),Cygwin) FIND_EXEC=-executable endif -ISABELLE_URL=http://isabelle.in.tum.de/website-$(ISABELLE_VSN)/dist/$(ISABELLE_ARCHIVE) +ISABELLE_URL=https://isabelle.in.tum.de/website-$(ISABELLE_VSN)/dist/$(ISABELLE_ARCHIVE) ISABELLE_DIR=Isabelle # Some defaults, for the case if makefile is called not by the dune build system. From dcab4e9612dbd7c87e21425f4b6c4d7020856739 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Tue, 14 May 2024 16:00:31 +0300 Subject: [PATCH 136/167] Try use Isabelle2024-RC2. (#124) * Use Isabelle2024-RC2. Signed-off-by: Karolis Petrauskas * Isabelle build fixed. Signed-off-by: Karolis Petrauskas * Trying to fix the macos build. Signed-off-by: Karolis Petrauskas * Use a constant for the default isabelle tactic. Signed-off-by: Karolis Petrauskas * tentative fix of broken auto method * tentative fix of broken auto method Signed-off-by: Stephan Merz * arith_add is replaced by other definitions in the isabelle thory. Signed-off-by: Karolis Petrauskas * Use Isabelle2024-RC3, behaviour looks the same. Signed-off-by: Karolis Petrauskas * Notes. Signed-off-by: Karolis Petrauskas * An attempt to resolve regressions related to how Nat is reduced to Int. Signed-off-by: Karolis Petrauskas * Multiple regressions resolved in the examples repo. Signed-off-by: Karolis Petrauskas * Don't skip passing proofs in the examples repo. Signed-off-by: Karolis Petrauskas * tentative fix of broken auto method Signed-off-by: Stephan Merz * arith_add is replaced by other definitions in the isabelle thory. Signed-off-by: Karolis Petrauskas * Added Arch Linux instructions to INSTALL.md Signed-off-by: Andrew Helwer <2n8rn1w1f@mozmail.com> * Use tlaplus/examples for tests Unified main/pr workflows Signed-off-by: Andrew Helwer <2n8rn1w1f@mozmail.com> * Use Isabelle2024-RC3, behaviour looks the same. Signed-off-by: Karolis Petrauskas * Notes. Signed-off-by: Karolis Petrauskas * An attempt to resolve regressions related to how Nat is reduced to Int. Signed-off-by: Karolis Petrauskas * Multiple regressions resolved in the examples repo. Signed-off-by: Karolis Petrauskas * Don't skip passing proofs in the examples repo. Signed-off-by: Karolis Petrauskas --------- Signed-off-by: Karolis Petrauskas Signed-off-by: Stephan Merz Signed-off-by: Andrew Helwer <2n8rn1w1f@mozmail.com> Co-authored-by: Stephan Merz Co-authored-by: Andrew Helwer <2n8rn1w1f@mozmail.com> --- .github/workflows/ci.yml | 8 +-- deps/isabelle/Makefile | 12 ++--- deps/isabelle/README.md | 13 +++++ isabelle/Functions.thy | 2 +- isabelle/IntegerArithmetic.thy | 9 ++-- isabelle/Integers.thy | 98 +++++++++++++++++++++++----------- isabelle/PredicateLogic.thy | 18 ++++--- isabelle/ROOT | 1 + isabelle/SetTheory.thy | 92 +++++++++++++++++++++---------- isabelle/Tests.thy | 51 ++++++++++++++++++ isabelle/Tuples.thy | 47 ++++++++++++---- isabelle/Zenon.thy | 16 ++++-- src/backend/isabelle.ml | 5 +- src/backend/prep.ml | 2 +- 14 files changed, 270 insertions(+), 104 deletions(-) create mode 100644 deps/isabelle/README.md create mode 100644 isabelle/Tests.thy diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 6797264d..95b86003 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -73,9 +73,6 @@ jobs: rm "$DEPS_DIR/tlaps.tar.gz" mv $DEPS_DIR/tlaps* "$DEPS_DIR/tlapm-install" SKIP=( - # Missing operator ENABLEDrules; will be fixed after - # updated_enabled_cdot branch is merged. - "specifications/ewd998/AsyncTerminationDetection.tla" # General proof failure "specifications/Bakery-Boulangerie/Bakery.tla" "specifications/Bakery-Boulangerie/Boulanger.tla" @@ -88,12 +85,9 @@ jobs: "specifications/LoopInvariance/BinarySearch.tla" "specifications/LoopInvariance/Quicksort.tla" "specifications/LoopInvariance/SumSequence.tla" - "specifications/MisraReachability/ReachabilityProofs.tla" - "specifications/Paxos/Consensus.tla" - "specifications/PaxosHowToWinATuringAward/Consensus.tla" "specifications/lamport_mutex/LamportMutex_proofs.tla" "specifications/TeachingConcurrency/SimpleRegular.tla" - # Long-running + # Failing and long-running "specifications/byzpaxos/BPConProof.tla" # Takes about 30 minutes ) python "$SCRIPT_DIR/check_proofs.py" \ diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 61dfa01f..524a2642 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -1,7 +1,7 @@ OS_TYPE=$(patsubst CYGWIN%,Cygwin,$(shell uname)) HOST_CPU=$(shell uname -m) -ISABELLE_VSN=Isabelle2023 +ISABELLE_VSN=Isabelle2024-RC3 ifeq ($(OS_TYPE),Linux) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz @@ -76,14 +76,14 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) .PRECIOUS: $(ISABELLE_DIR).no-links $(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ rm -rf $@ - cd $< && rm -rf ./contrib/e-2.6-1/src/lib/ + cd $< && rm -rf ./contrib/e-3.0.03-1/src/lib/ ifeq ($(OS_TYPE),Darwin) - cd $< && cd contrib/jdk-17.0.7/arm64-darwin/ \ + cd $< && cd contrib/jdk-21.0.3/arm64-darwin/ \ && (find . -type link | xargs rm) \ - && mv zulu-17.jdk/Contents/Home/* ./ - cd $< && cd contrib/jdk-17.0.7/x86_64-darwin/ \ + && mv zulu-21.jdk/Contents/Home/* ./ + cd $< && cd contrib/jdk-21.0.3/x86_64-darwin/ \ && (find . -type link | xargs rm) \ - && mv zulu-17.jdk/Contents/Home/* ./ + && mv zulu-21.jdk/Contents/Home/* ./ cd $< && rm -rf contrib/vscodium-1.70.1 endif touch $@ diff --git a/deps/isabelle/README.md b/deps/isabelle/README.md new file mode 100644 index 00000000..692d4ad7 --- /dev/null +++ b/deps/isabelle/README.md @@ -0,0 +1,13 @@ +## Debugging Isabelle prover + +Run the `tlapm` with the `--debug=tempfiles` option, e.g.: + + (cd ../tlaplus-examples/specifications/MisraReachability/ \ + && rm -rf .tlacache/ && tlapm --toolbox 228 228 --debug=tempfiles ReachabilityProofs.tla) + +Then look for the corresponding `*.thy` files and open them with Isabelle, e.g. + + ./_build/default/deps/isabelle/Isabelle/bin/isabelle jedit \ + -d ./_build/default/deps/isabelle/Isabelle/src/TLA+/ \ + ../tlaplus-examples/specifications/MisraReachability/.tlacache/ReachabilityProofs.tlaps/tlapm_624cb2.thy + diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index ca831784..6ad4cf88 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -196,7 +196,7 @@ proof - proof auto fix x assume "x \ S" - with hyp have "\y : P(x,y)" .. + with hyp have "\y : P(x,y)" by auto thus "P(x, CHOOSE y : P(x,y))" by (rule chooseI_ex) qed thus ?thesis .. diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy index 07aaf7a6..6002a476 100644 --- a/isabelle/IntegerArithmetic.thy +++ b/isabelle/IntegerArithmetic.thy @@ -1076,7 +1076,8 @@ proof - next fix n assume "n \ Nat" "\m \ Nat : m \ n \ P(m)" - with step show "\m \ Nat : m \ succ[n] \ P(m)" by auto + with step show "\m \ Nat : m \ succ[n] \ P(m)" + by (auto simp del: nat_iff_int_geq0) qed thus ?thesis by auto qed @@ -1398,7 +1399,7 @@ proof (rule int_leq_lessE[OF ij i j]) thus ?thesis by auto next assume "i = j" - with i f show ?thesis by simp + with i f show ?thesis by auto qed lemma nat_less_mono_imp_leq_mono: @@ -1412,7 +1413,7 @@ proof (rule int_leq_lessE[OF ij]) thus ?thesis by auto next assume "i = j" - with i f show ?thesis by simp + with i f show ?thesis by (auto simp del: nat_iff_int_geq0) qed (auto simp: i[simplified] j[simplified]) lemma int_succ_lessI: @@ -2070,7 +2071,7 @@ next { assume eq: "i..j = k..l" and idx: "i \ k \ j \ l" from eq \i \ i..j\ intvlIsEmpty[OF k l] have "\(k > l)" - by blast + by auto with k l have "k \ k..l" "l \ k..l" by simp+ from idx have "FALSE" proof diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index bfbefc7f..c33f40d9 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -118,7 +118,7 @@ proof - next fix n assume n: "n \ S" - with Sc have "Sc[n] \ S" .. + with Sc have "Sc[n] \ S" by (rule bspec) moreover from n sub have "n \ N" by auto hence "Sc[n] = ?sc(n)" by (simp add: Sc_def) @@ -149,15 +149,14 @@ proof - show "FALSE" proof (cases "k \ N") case True - with e \k \ Z\ have "e \ k \ cp(k) = k \ {e}" by (simp add: cp_def) - with \cp(k) = k\ show ?thesis by blast + with e \k \ Z\ \cp(k) = k\ show ?thesis + by (auto simp: cp_def) next case False with \k \ I\ obtain n where "n \ N" "k = cp(n)" by (auto simp: I_def) with \k \ Z\ have "n \ Z" by (auto simp: cp_def) - with e \n \ N\ \k = cp(n)\ \k \ Z\ - have "k = n \ {e}" "cp(k) = k \ {e}" by (auto simp: cp_def) - with \cp(k) = k\ show ?thesis by auto + with e \n \ N\ \k = cp(n)\ \k \ Z\ \cp(k) = k\ show ?thesis + by (auto simp: cp_def) qed qed @@ -178,7 +177,7 @@ proof - case False with \k \ I\ obtain n where "n \ N" "k = cp(n)" by (auto simp: I_def) - with \cp(k) = Z\ cpN have "n = Z" by simp + with \cp(k) = Z\ cpN have "n = Z" by auto with \k = cp(n)\ show ?thesis by (simp add: cp_def) qed } @@ -210,8 +209,7 @@ proof - assume "l \ N" with \l \ I\ obtain n where "n \ N" "l = cp(n)" by (auto simp: I_def) - with cpk cpN eq have "n = k \ {e}" by simp - with \n \ N\ e show "FALSE" by blast + with cpk eq cpN \n \ N\ e show "FALSE" by force qed have "l \ Z" proof @@ -219,28 +217,51 @@ proof - with eq cpZ dom kl cpisZ show "FALSE" by auto qed with \l \ N\ e have "e \ l" "cp(l) = l \ {e}" by (auto simp: cp_def) - with cpk eq have "k \ l" "l \ k" by auto + with cpk eq have "k \ l" "l \ k" by force+ with kl show ?thesis by (auto dest: setEqual) next case False with dom obtain nk where nk: "nk \ N" "k = cp(nk)" by (auto simp: I_def) - with cpN eq have cpl: "cp(l) \ N" by simp + with cpN eq have cpl: "cp(l) \ N" by auto have "l \ N" proof assume "l \ N" - with cpl cpNN have "l = Z" by simp - with eq dom cpZ cpisZ False 1 show "FALSE" by simp + with cpl cpNN have "l = Z" by auto + with eq dom cpZ cpisZ False 1 show "FALSE" by auto qed with dom obtain nl where "nl \ N" "l = cp(nl)" by (auto simp: I_def) - with nk eq kl cpN show ?thesis by simp + with nk eq have "cp(cp(nk)) = cp(cp(nl))" by simp + moreover + from \nk \ N\ cpN have "cp(cp(nk)) = nk" by auto + moreover + from \nl \ N\ cpN have "cp(cp(nl)) = nl" by auto + ultimately have "k = l" + using \k = cp(nk)\ \l = cp(nl)\ by simp + with kl show ?thesis .. qed } hence 10: "\k,l \ I : ?cp[k] = ?cp[l] \ k = l" by auto - from cpI cpN have 11: "\k \ I : ?cp[?cp[k]] = k" - by (auto simp: I_def) + have 11: "\k \ I : ?cp[?cp[k]] = k" + proof + fix k + assume "k \ I" + with cpI have "?cp[?cp[k]] = cp(cp(k))" + by force + moreover have "cp(cp(k)) = k" + proof (cases "k \ N") + case True + with cpN show ?thesis by auto + next + case False + with \k \ I\ obtain nk where "nk \ N" "k = cp(nk)" + by (auto simp: I_def) + with cpN show ?thesis by auto + qed + ultimately show "?cp[?cp[k]] = k" by simp + qed from 1 2 3 4 5 6 7 8 9 10 11 have "ExtendedPeano(I,N,Z,Sc,?cp)" unfolding ExtendedPeano_def by blast @@ -261,10 +282,10 @@ definition Nat :: "c" where "Nat \ DOMAIN Succ" definition zero :: "c" ("0") where - "0 \ CHOOSE Z : \cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" + "zero \ CHOOSE Z : \cp,I : ExtendedPeano(I,Nat,Z,Succ,cp)" definition intCplt :: "c" where - "intCplt \ CHOOSE cp : \I : ExtendedPeano(I,Nat,0,Succ,cp)" + "intCplt \ CHOOSE cp : \I : ExtendedPeano(I,Nat,zero,Succ,cp)" definition Int :: "c" where "Int \ DOMAIN intCplt" @@ -277,6 +298,10 @@ lemmas setEqualI [where B = "Nat", intro!] setEqualI [where A = "Int", intro!] setEqualI [where B = "Int", intro!] + setEqualE [where A = "Nat", elim] + setEqualE [where B = "Nat", elim] + setEqualE [where A = "Int", elim] + setEqualE [where B = "Int", elim] lemma intExtendedPeano: "ExtendedPeano(Int,Nat,0,Succ,intCplt)" proof - @@ -369,8 +394,11 @@ proof - from z have "0 \ ?P" by simp moreover from sc have "\n \ ?P : Succ[n] \ ?P" by simp - ultimately have "Nat \ ?P" + moreover + have "\S \ SUBSET Nat : 0 \ S \ (\n \ S : Succ[n] \ S) \ Nat \ S" using intExtendedPeano by (simp add: ExtendedPeano_def) + ultimately have "Nat \ ?P" + by blast thus ?thesis by auto qed @@ -411,8 +439,14 @@ lemma uminusInjIff [simp]: by auto lemma uminusUminus [simp]: - "a \ Int \ --a = a" - using intExtendedPeano by (simp add: ExtendedPeano_def uminus_def) + assumes "a \ Int" + shows "--a = a" +proof - + from intExtendedPeano have "\k \ Int : intCplt[intCplt[k]] = k" + by (simp add: ExtendedPeano_def) + with assms show ?thesis + by (auto simp add: uminus_def) +qed lemma uminusSwap: assumes "a \ Int" and "b \ Int" @@ -727,7 +761,7 @@ next fix n assume n: "n \ Nat" "a = -Succ[n]" hence "a = succ[-succ[Succ[n]]]" by simp - with n show ?thesis by blast + with n show ?thesis by auto qed definition pred where @@ -856,7 +890,7 @@ next fix n assume "n \ Nat" thus "P(succ[m],n)" proof (cases) - case 0 with b1 m show ?thesis by auto + case 0 with b1 m show ?thesis by force next case succ with step ih m show ?thesis by auto qed @@ -975,6 +1009,8 @@ where "upto(n) \ lfp(Nat, \S. {n} \ { k \ Nat : succ[k lemmas setEqualI [where A = "upto(n)" for n, intro!] setEqualI [where B = "upto(n)" for n, intro!] + setEqualE [where A = "upto(n)" for n, elim] + setEqualE [where B = "upto(n)" for n, elim] lemma uptoNat: "upto(n) \ Nat" unfolding upto_def by (rule lfpSubsetDomain) @@ -1082,7 +1118,7 @@ lemma uptoLimit: proof - from m uptoNat have mNat: "m \ Nat" by blast from n have "\m\Nat: m \ upto(n) \ succ[m] \ upto(n) \ m=n" (is "?P(n)") - by (induct, auto simp: uptoZero uptoSucc) + by (induct, (force simp add: uptoZero uptoSucc)+) with mNat m suc show ?thesis by blast qed @@ -1188,12 +1224,12 @@ proof - assume kNat: "k \ Nat" and ih: "k \ upto(m) \ f[k] = g[k]" and sk: "succ[k] \ upto(m)" from sk kNat g have "g[succ[k]] = step(k, g[k])" - unfolding primrec_nat_upto_def by simp + unfolding primrec_nat_upto_def by auto moreover from sk m n have "succ[k] \ upto(n)" by (rule uptoTrans) with kNat f have "f[succ[k]] = step(k, f[k])" - unfolding primrec_nat_upto_def by simp + unfolding primrec_nat_upto_def by auto moreover from sk kNat mNat have "k \ upto(m)" by (rule uptoPred) @@ -1256,7 +1292,7 @@ proof - unfolding g_def by (simp+) moreover from F have "g[0] = base" - by (simp add: primrec_nat_upto_def g_def) + by (auto simp: primrec_nat_upto_def g_def) moreover have "\n \ Nat : g[succ[n]] = step(n, g[n])" proof @@ -1363,14 +1399,14 @@ proof - from sk kNat g have "g[succ[k]] = pstep(k, g[k], g[-k]) \ g[-succ[k]] = nstep(k, g[k], g[-k])" - unfolding primrec_int_upto_def by simp + unfolding primrec_int_upto_def by auto moreover from sk m n have "succ[k] \ upto(n)" by (rule uptoTrans) with kNat f have "f[succ[k]] = pstep(k, f[k], f[-k]) \ f[-succ[k]] = nstep(k, f[k], f[-k])" - unfolding primrec_int_upto_def by simp + unfolding primrec_int_upto_def by auto moreover from sk kNat mNat have "k \ upto(m)" by (rule uptoPred) @@ -1448,7 +1484,7 @@ proof - unfolding g_def by (simp+) moreover from F have "g[0] = base" - by (simp add: primrec_int_upto_def g_def) + by (auto simp: primrec_int_upto_def g_def) moreover have "\n \ Nat : g[succ[n]] = pstep(n, g[n], g[-n]) \ g[-succ[n]] = nstep(n, g[n], g[-n])" @@ -1490,7 +1526,7 @@ proof - and 5: "\n\Nat : f[-succ[n]] = nstep(n, f[n], f[-n])" by blast from 3 4 5 base pstep nstep have "\n\Int : f[n] \ S" - by (intro intInduct) auto + by (intro intInduct) force+ with 1 2 3 4 5 show ?thesis by blast qed diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index cd9411a9..6aaf351d 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -1428,15 +1428,18 @@ ML \ declaration \ fn _ => Induct.map_simpset (fn ss => ss addsimprocs - [Simplifier.make_simproc @{context} "swap_cases_false" - {lhss = [@{term "cases_false \ PROP P \ PROP Q"}], + [Simplifier.make_simproc @{context} + {name = "swap_cases_false", + lhss = [@{term "cases_false \ PROP P \ PROP Q"}], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (P as _ $ @{const cases_false}) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE - | _ => NONE)}, - Simplifier.make_simproc @{context} "cases_equal_conj_curry" - {lhss = [@{term "cases_conj(P, Q) \ PROP R"}], + | _ => NONE), + identifier = []}, + Simplifier.make_simproc @{context} + {name = "cases_equal_conj_curry", + lhss = [@{term "cases_conj(P, Q) \ PROP R"}], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (_ $ P) $ _ => @@ -1448,7 +1451,8 @@ declaration \ | is_conj @{const cases_false} = true | is_conj _ = false in if is_conj P then SOME @{thm cases_conj_curry} else NONE end - | _ => NONE)}] + | _ => NONE), + identifier = []}] |> Simplifier.set_mksimps (fn ctxt => Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> map (rewrite_rule ctxt (map Thm.symmetric @{thms cases_rulify_fallback})))) @@ -1787,7 +1791,7 @@ simproc_setup neq ("x = y") = \fn _ => eq $ lhs $ rhs => (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of SOME thm => SOME ((thm RS symEqLeft_to_symEQLeft) - handle _ => thm RS neq_to_EQ_False) + handle THM _ => thm RS neq_to_EQ_False) | NONE => NONE) | _ => NONE); in proc end diff --git a/isabelle/ROOT b/isabelle/ROOT index c4124c97..fb0a8202 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -7,6 +7,7 @@ session "TLA+" = "Pure" + Constant (global) Zenon (global) (* NewSMT (global) *) + Tests (* Test cases only. *) document_files "root.tex" diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index bbb6fbe0..9bfb1ac8 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -48,6 +48,7 @@ text \ axioms of set theory, in order to make them more readable. \ + text \ Concrete syntax: set comprehension \ (*** @@ -372,34 +373,26 @@ lemma bAllI [intro!]: shows "\x\A : P(x)" using assms unfolding bAll_def by blast -lemma bspec [dest?]: +lemma bspec: assumes "\x\A : P(x)" and "x \ A" shows "P(x)" using assms unfolding bAll_def by blast -(*** currently inactive -- causes several obscure warnings about renamed - variables and seems to slow down large examples such as AtomicBakery ***) -ML \ - structure Simpdata = - struct - - open Simpdata; - - val mksimps_pairs = [(@{const_name bAll}, (@{thms bspec}, false))] @ mksimps_pairs; +text \The following rule intentionally has single assumption (non-conditional), +because otherwise it interferes with how \Nat\ is reduced to \Int\.\ +lemma bspec' [dest]: + assumes "\x\A : P(x)" + shows "\x : (x\A) \ P(x)" + using assms unfolding bAll_def by blast - end; - - open Simpdata; -\ +lemma bAll_unb [simp] : "\T P. (\e : e \ T \ P(e)) \ (\e \ T : P(e))" + using bAll_def by simp -declaration \ - fn _ => Simplifier.map_ss (fn ss => ss - |> Simplifier.set_mksimps (fn ctxt => - Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +setup \ + map_theory_claset (fn ctxt => + ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) \ -(***) - lemma bAllE [elim]: assumes "\x\A : P(x)" and "x \ A \ Q" and "P(x) \ Q" shows "Q" @@ -491,6 +484,28 @@ unfolding bChoose_def proof (rule choose_det) from assms show "x \ A \ P(x) \ x \ B \ Q(x)" by blast qed +(*** currently inactive -- causes several obscure warnings about renamed + variables and seems to slow down large examples such as AtomicBakery *** +ML \ + structure Simpdata = + struct + + open Simpdata; + + val mksimps_pairs = [(@{const_name bAll}, (@{thms bspec}, false))] @ mksimps_pairs; + + end; + + open Simpdata; +\ + +declaration \ + fn _ => Simplifier.map_ss (fn ss => ss + |> Simplifier.set_mksimps (fn ctxt => + Simpdata.mksimps Simpdata.mksimps_pairs ctxt)) +\ +***) + (*** TYPE CHECKING CURRENTLY NOT INSTALLED @@ -689,17 +704,37 @@ lemma setEqualD1: "A = B \ A \ B" lemma setEqualD2: "A = B \ B \ A" by blast -text \ - We declare the elimination rule for set equations as an unsafe - rule to use with the classical reasoner, so it will be tried if - the more obvious uses of equality fail. -\ -lemma setEqualE [elim]: +lemma setEqualE: assumes "A = B" and "\ c \ A; c \ B \ \ P" and "\ c \ A; c \ B \ \ P" shows "P" using assms by (blast dest: setEqualD1 setEqualD2) +text \ + Again, we add instances of this theorem for obvious set constructors. +\ +lemmas + setEqualE [where A = "{}", elim] + setEqualE [where B = "{}", elim] + setEqualE [where A = "addElt(a,C)" for a C, elim] + setEqualE [where B = "addElt(a,C)" for a C, elim] + setEqualE [where A = "SUBSET C" for C, elim] + setEqualE [where B = "SUBSET C" for C, elim] + setEqualE [where A = "UNION C" for C, elim] + setEqualE [where B = "UNION C" for C, elim] + setEqualE [where A = "INTER C" for C, elim] + setEqualE [where B = "INTER C" for C, elim] + setEqualE [where A = "C \ D" for C D, elim] + setEqualE [where B = "C \ D" for C D, elim] + setEqualE [where A = "C \ D" for C D, elim] + setEqualE [where B = "C \ D" for C D, elim] + setEqualE [where A = "C \ D" for C D, elim] + setEqualE [where B = "C \ D" for C D, elim] + setEqualE [where A = "subsetOf(S,P)" for S P, elim] + setEqualE [where B = "subsetOf(S,P)" for S P, elim] + setEqualE [where A = "setOfAll(S,e)" for S e, elim] + setEqualE [where B = "setOfAll(S,e)" for S e, elim] + lemma setEqual_iff: "(A = B) = (\x : x \ A \ x \ B)" by (blast intro: setEqualI) @@ -1102,7 +1137,8 @@ lemma inAsym: shows "P" proof (rule contradiction) assume "\P" - with foundation[where A="{a,b}"] hyps show "FALSE" by blast + with foundation[where A="{a,b}"] hyps + show "FALSE" by blast qed lemma inIrrefl: @@ -1422,7 +1458,7 @@ lemma isEmptySimps [simp]: "({} = addElt(a,S)) = FALSE" "(SUBSET S = {}) = FALSE" "({} = SUBSET S) = FALSE" -by (blast+) +by blast+ subsection \ Generalized union: inclusions and equalities \ diff --git a/isabelle/Tests.thy b/isabelle/Tests.thy new file mode 100644 index 00000000..7ebe8ba3 --- /dev/null +++ b/isabelle/Tests.thy @@ -0,0 +1,51 @@ +theory Tests + imports Constant +begin +text \ +This theory only acts as a test suite to check for regressions. +The test cases should be proved by \auto\, because that's the +default method used by TLAPM. +\ + + +section \Interference of \Nat\ and \\\ elimination.\ + +text \After upgrading theories from Isabelle-2011 to Isabelle-202X +an interference of \x \ Nat \ x \ Int \ 0 \ x\ and \\\ elimination +was introduced. The former rule rewrites assumptions and then the +\bAllE\ elimination rule don't apply anymore.\ + +lemma + fixes P + assumes "\e : e \ Nat \ P(e)" + shows "\e. e \ Nat \ P(e)" + using assms + by auto + + +text \This lemma was passing with "blast" and failing with auto +until "lemma bspec' [dest]" was added. It was working with "auto" +in Isabelle-2011, so can be considered a regression.\ +lemma + fixes P Q e + assumes "e \ Nat" and "P(e)" and "\e \ Nat : (P(e) \ Q(e))" + shows "Q(e)" + using assms + by auto + + + +text \Only started to work after \bAll_unb [simp]\ was added.\ +lemma + assumes "RR (0)" + assumes "\ i :: c. i \ Nat \ RR(i) \ RR(F(i))" + assumes "\ PP :: c => c. + PP (0) \ + (\n\Nat : PP(n) \ PP(F(n))) \ + (\n\Nat : PP(n))" + shows "\ ii \ Nat : RR (ii)" + using assms + by auto + + +end diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5aa1bd04..1030ff1f 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -72,7 +72,13 @@ lemma isASeqE [elim]: assumes s: "isASeq(s)" and p: "\isAFcn(s); DOMAIN s = 1 .. Len(s); Len(s) \ Nat\ \ P" shows "P" - using assms unfolding isASeq_def by (blast dest: LenI) +proof - + from s obtain n where "isAFcn(s)" "n \ Nat" "DOMAIN s = 1 .. n" + unfolding isASeq_def by blast + with LenI have "isAFcn(s) \ DOMAIN s = 1 .. Len(s) \ Len(s) \ Nat" + by simp + thus ?thesis by (blast intro: p) +qed \ \The corresponding lemma for @{text "Seq(S)"} will be proved later.\ lemma SeqIsAFcn (*[elim!]*): @@ -156,7 +162,7 @@ lemma seqElt [elim]: lemma seqInSeqRange: assumes "isASeq(s)" shows "s \ Seq(Range(s))" - using assms by auto + using assms by force lemma isASeqInSeq: "isASeq(s) = (\S: s \ Seq(S))" by (auto elim: seqInSeqRange) @@ -273,8 +279,12 @@ lemmas appendElt1' [simp] = SeqIsASeq[THEN appendElt1] lemma appendElt2 [simp]: assumes "isASeq(s)" shows "Append(s,e)[succ[Len(s)]] = e" - using assms unfolding Append_def by force - +proof - + from assms have "succ[Len(s)] \ DOMAIN Append(s,e)" + by (auto simp: Append_def) + thus ?thesis by (auto simp: Append_def) +qed + lemmas appendElt2' [simp] = SeqIsASeq[THEN appendElt2] lemma isAppend [intro]: @@ -309,7 +319,7 @@ proof - fix k assume k: "k \ 1 .. Len(t)" with s ls have "s[k] = ?s1[k]" by (intro sym[OF appendElt1]) auto - also from k elt t have "\ = ?t1[k]" by auto + also from k elt t have "\ = ?t1[k]" by force also from t k have "... = t[k]" by (intro appendElt1) auto finally show "s[k] = t[k]" . qed @@ -413,10 +423,10 @@ proof - qed from sn n have 2: "so \ [1 .. n \ S]" unfolding so_def by force - with ih have 3: "P(so)" .. + with ih have 3: "P(so)" by blast from 2 n have 4: "so \ Seq(S)" unfolding Seq_def by auto - from sn n have "lst \ S" by (auto simp: lst_def) + from sn n have "lst \ S" by (force simp: lst_def) with 1 3 4 show "P(sn)" by (auto intro: step) qed qed @@ -690,11 +700,26 @@ lemma inProductLen: using assms unfolding Product_def by auto lemma inProductE [elim!]: - assumes "p \ Product(s)" and "isASeq(s)" - and "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; + assumes p: "p \ Product(s)" and s: "isASeq(s)" + and P: "\isASeq(p); Len(p) = Len(s); p \ [1 .. Len(s) \ UNION Range(s)]; \k \ 1 .. Len(s) : p[k] \ s[k] \ \ P" shows "P" - using assms unfolding Product_def by auto +proof - + from \p \ Product(s)\ have + 1: "p \ [1 .. Len(s) \ UNION {s[i] : i \ 1 .. Len(s)}]" and + 2: "\i \ 1 .. Len(s) : p[i] \ s[i]" + by (auto simp: Product_def) + from p s have "isASeq(p)" by (rule inProductIsASeq) + moreover + from p s have "Len(p) = Len(s)" by (rule inProductLen) + moreover + from s have "Range(s) = {s[i] : i \ 1 .. Len(s)}" + by auto + with 1 have "p \ [1 .. Len(s) \ UNION Range(s)]" + by simp + ultimately show ?thesis + using 2 by (rule P) +qed (*** examples *** lemma "Product(\\) = { \\ }" by auto @@ -1136,7 +1161,7 @@ lemma transitive_converse [simp]: lemma symmetric_iff_converse_eq: assumes r: "r \ A \ B" shows "symmetric(r) = (r^-1 = r)" - using assms by (auto simp: symmetric_def) + using assms by (auto simp: symmetric_def dest: converseI) subsubsection \ Identity relation over a set \ diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index 91cf9114..ebd642a7 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -19,7 +19,7 @@ proof thus "\ x \ S : P(x)" .. next assume "\ x \ S : P(x)" - thus "(\x. (x \ S \ P(x)))" .. + thus "(\x. (x \ S \ P(x)))" by blast qed lemma atomize_object_bAll [atomize]: @@ -2108,11 +2108,15 @@ next Append (zenon_seqify (es), e)[succ[Len(zenon_seqify(cs))]]) = CaseArm (c, e)" using h3 zenon_seqifyIsASeq by auto - with h5 + with h5 + have "x \\in CaseArm (Append (zenon_seqify (cs), c)[succ[Len(zenon_seqify(cs))]], + Append (zenon_seqify (es), e)[succ[Len(zenon_seqify(cs))]])" + by simp + with \succ[Len(zenon_seqify(cs))] \ DOMAIN Append (zenon_seqify (cs), c)\ have h6: "x \\in UNION {CaseArm (Append (zenon_seqify (cs), c)[i], Append (zenon_seqify (es), e)[i]) : i \\in DOMAIN Append (zenon_seqify (cs), c)}" - by blast + by blast show "?g" using h4 h6 by auto next @@ -2207,8 +2211,10 @@ proof (rule boolEqual, rule iffI) using h7 zenon_seqifyIsASeq by auto have h9: "zenon_seqify (cs) [i] = Append (zenon_seqify (cs), c)[i]" using zenon_case_append1 zenon_seqifyIsASeq h7 by auto - show "~ zenon_seqify(cs)[i]" - using h9 h8 h6 by auto + from h8 h6 have "~ Append(zenon_seqify(cs), c)[i]" + by blast + with h9 show "~ zenon_seqify(cs)[i]" + by simp qed next assume h6: "?f2" diff --git a/src/backend/isabelle.ml b/src/backend/isabelle.ml index 286d2a1d..6daefc2f 100644 --- a/src/backend/isabelle.ml +++ b/src/backend/isabelle.ml @@ -167,9 +167,8 @@ let rec pp_apply sd cx ff op args = match op.core with | B.Nat, [] -> atomic "Nat" | B.Int, [] -> atomic "Int" | B.Real, [] -> atomic (cook "Real") - | B.Plus, [e ; f] -> nonfix "arith_add" [e ; f] - | B.Minus, [e ; f] -> - nonfix "arith_add" [e ; Apply (Internal B.Uminus @@ f, [f]) @@ f] + | B.Plus, [e ; f] -> nonfix "addint" [e ; f] + | B.Minus, [e ; f] -> nonfix "subint" [e ; f] | B.Uminus, [e] -> nonfix "minus" [e] | B.Times, [e ; f] -> nonfix "mult" [e ; f] | B.Ratio, [e ; f] -> nonfix (cook "/") [e ; f] diff --git a/src/backend/prep.ml b/src/backend/prep.ml index c91a4f1e..097977c2 100644 --- a/src/backend/prep.ml +++ b/src/backend/prep.ml @@ -1193,7 +1193,7 @@ let compute_meth def args usept = Method.Zenon tmo | Some "isabelle" -> let tmo = Option.default Method.default_isabelle_timeout !timeout in - let tac = Option.default "auto" !tactic in + let tac = Option.default Method.default_isabelle_tactic !tactic in Method.Isabelle (tmo, tac) | Some "yices" -> let tmo = Option.default Method.default_yices_timeout !timeout in From 2e2bbff0718db822b8aa8bfd43f1e3f1ff764f9d Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 18 May 2024 10:31:24 +0300 Subject: [PATCH 137/167] Reorganize Isabelle sessions. Signed-off-by: Karolis Petrauskas --- isabelle/.gitignore | 2 ++ isabelle/ROOT | 16 +++++++++++++++- isabelle/examples/ROOT | 18 ------------------ isabelle/{ => tests}/Tests.thy | 0 4 files changed, 17 insertions(+), 19 deletions(-) delete mode 100644 isabelle/examples/ROOT rename isabelle/{ => tests}/Tests.thy (100%) diff --git a/isabelle/.gitignore b/isabelle/.gitignore index ea1472ec..78c2ceb9 100644 --- a/isabelle/.gitignore +++ b/isabelle/.gitignore @@ -1 +1,3 @@ output/ +*~ +*# diff --git a/isabelle/ROOT b/isabelle/ROOT index fb0a8202..dc2d6c1b 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -7,7 +7,21 @@ session "TLA+" = "Pure" + Constant (global) Zenon (global) (* NewSMT (global) *) - Tests (* Test cases only. *) + + document_files + "root.tex" + +session "TLA+Tests" in tests = "TLA+" + + options [document=false] + theories + Tests + +session "TLA+Examples" in examples = "TLA+" + + options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + + theories + Allocator + AtomicBakeryG document_files "root.tex" diff --git a/isabelle/examples/ROOT b/isabelle/examples/ROOT deleted file mode 100644 index cbcd3fa6..00000000 --- a/isabelle/examples/ROOT +++ /dev/null @@ -1,18 +0,0 @@ -(* Title: TLA+/examples/ROOT - Author: Stephan Merz, Inria Nancy - License: BSD - -Examples of using Isabelle/TLA+ -*) - -(* build the session (without a heap) using "isabelle build -D ." *) -chapter "TLA+ Examples" -session "TLA-Examples" = "TLA+" + - options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] - - theories - Allocator - AtomicBakeryG - - document_files - "root.tex" diff --git a/isabelle/Tests.thy b/isabelle/tests/Tests.thy similarity index 100% rename from isabelle/Tests.thy rename to isabelle/tests/Tests.thy From 71b228c92fb40c91fcb9e1832da8367785be422e Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 29 May 2024 18:45:31 +0300 Subject: [PATCH 138/167] Adapt to Isabelle2024. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 524a2642..c050e056 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -1,7 +1,7 @@ OS_TYPE=$(patsubst CYGWIN%,Cygwin,$(shell uname)) HOST_CPU=$(shell uname -m) -ISABELLE_VSN=Isabelle2024-RC3 +ISABELLE_VSN=Isabelle2024 ifeq ($(OS_TYPE),Linux) ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz @@ -55,7 +55,8 @@ endif .PRECIOUS: $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) cd $(ISABELLE_DIR) \ - && rm -rf contrib/ProofGeneral* doc heaps/*/HOL + && rm -rf contrib/ProofGeneral* doc heaps/*/HOL contrib/vscodium* contrib/vscode* \ + && sed -i -e 's@^\(contrib/vscode_extension\|contrib/vscodium\|src/Tools/Demo\)@#rm at TLA# \1@' etc/components cd $(ISABELLE_DIR) \ && HEAPS_PATH=$(shell pwd)/$(ISABELLE_DIR)/heaps \ && if [ "$(HOST_CPU)" = "x86_64" ] ; then sed -i -e 's/^ML_PLATFORM=.*$$/ML_PLATFORM="$${ISABELLE_PLATFORM64:-$$ISABELLE_PLATFORM}"/' etc/settings ; fi \ @@ -84,7 +85,6 @@ ifeq ($(OS_TYPE),Darwin) cd $< && cd contrib/jdk-21.0.3/x86_64-darwin/ \ && (find . -type link | xargs rm) \ && mv zulu-21.jdk/Contents/Home/* ./ - cd $< && rm -rf contrib/vscodium-1.70.1 endif touch $@ From 0f40bb1311f022a926fdb72068126bbb459ee5bf Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 29 May 2024 23:17:00 +0300 Subject: [PATCH 139/167] Cleanup Isabelle build a bit; run Isabelle examples/tests during dune tests. Signed-off-by: Karolis Petrauskas --- deps/isabelle/.gitignore | 3 +++ deps/isabelle/Makefile | 7 ++++--- deps/isabelle/dune | 6 ++++-- deps/isabelle/theories/TLA+ | 1 - dune-project | 2 +- isabelle/Makefile | 18 ++++++++++++++---- isabelle/ROOT | 2 +- isabelle/dune | 13 +++++++++++++ tlapm.opam | 2 +- 9 files changed, 41 insertions(+), 13 deletions(-) create mode 100644 deps/isabelle/.gitignore delete mode 120000 deps/isabelle/theories/TLA+ create mode 100644 isabelle/dune diff --git a/deps/isabelle/.gitignore b/deps/isabelle/.gitignore new file mode 100644 index 00000000..20068ef0 --- /dev/null +++ b/deps/isabelle/.gitignore @@ -0,0 +1,3 @@ +/Isabelle.exec-files +/Isabelle.no-links +/Isabelle/ diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index c050e056..28e0dae8 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -27,7 +27,7 @@ ISABELLE_URL=https://isabelle.in.tum.de/website-$(ISABELLE_VSN)/dist/$(ISABELLE_ ISABELLE_DIR=Isabelle # Some defaults, for the case if makefile is called not by the dune build system. -PROJECT_ROOT=../.. +PROJECT_ROOT=$(if $(DUNE_SOURCEROOT),$(DUNE_SOURCEROOT),../..) CACHE_DIR=$(PROJECT_ROOT)/_build_cache @@ -63,11 +63,12 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && cp etc/settings etc/settings.target \ && echo "ISABELLE_OUTPUT=$$HEAPS_PATH" >> etc/settings mkdir -p $(ISABELLE_DIR)/src/TLA+ \ - && cp -a theories/TLA+/* $(ISABELLE_DIR)/src/TLA+/ \ + && cp -a ../../isabelle/* $(ISABELLE_DIR)/src/TLA+/ \ && chmod -R u+w $(ISABELLE_DIR)/src/TLA+/ cd $(ISABELLE_DIR)/ \ && ./bin/isabelle build -o system_heaps -o document=false -b -v -d src/Pure Pure \ - && ./bin/isabelle build -o system_heaps -o document=false -b -c -v -d src/TLA+ TLA+ + && ./bin/isabelle build -o system_heaps -o document=false -b -c -v -d src/TLA+ TLA+ \ + && rm -rf ./heaps/polyml-*/log/* cd $(ISABELLE_DIR) \ && rm etc/settings \ && mv etc/settings.target etc/settings diff --git a/deps/isabelle/dune b/deps/isabelle/dune index 12ad3ce8..d9d492db 100644 --- a/deps/isabelle/dune +++ b/deps/isabelle/dune @@ -2,9 +2,11 @@ ; The source code for the TLA+ theory is in the $PROJECT_ROOT/isabelle directory. ; The generated heaps (Pure and TLA+) are in Isabelle/heaps/polyml-*/. (rule - (deps "Makefile" (source_tree theories)) + (alias default) + (mode (promote (until-clean))) + (deps "Makefile" (source_tree ../../isabelle)) (targets (dir Isabelle) "Isabelle.no-links" "Isabelle.exec-files") - (action (run "make" "-C" "." "PROJECT_ROOT=%{project_root}" "Isabelle.no-links" "Isabelle.exec-files"))) + (action (run "make" "-C" "." "Isabelle.no-links" "Isabelle.exec-files"))) (install (section (site (tlapm backends))) diff --git a/deps/isabelle/theories/TLA+ b/deps/isabelle/theories/TLA+ deleted file mode 120000 index 0dbabb87..00000000 --- a/deps/isabelle/theories/TLA+ +++ /dev/null @@ -1 +0,0 @@ -../../../isabelle \ No newline at end of file diff --git a/dune-project b/dune-project index c0503ecb..e75fec3a 100644 --- a/dune-project +++ b/dune-project @@ -1,4 +1,4 @@ -(lang dune 3.8) +(lang dune 3.12) (using dune_site 0.1) diff --git a/isabelle/Makefile b/isabelle/Makefile index 6421e79b..617e536e 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -1,4 +1,7 @@ -.PHONY: rebuild heap-only clean +ISABELLE=isabelle +SRC_ROOT=$(if $(DUNE_SOURCEROOT),$(DUNE_SOURCEROOT),..) + +.PHONY: rebuild heap-only dune-runtest clean # Generate theory information (in HTML and PDF formats) as well as the main # heap file for the TLA+ object logic. By default these will be generated @@ -7,10 +10,17 @@ # file etc/settings beneath the Isabelle home (i.e., where the Isabelle # distribution is installed). rebuild: - isabelle build -o document -o browser_info -b -c -v -D . + $(ISABELLE) build -o document -o browser_info -b -c -v -D . heap-only: - isabelle build -o document=false -o browser_info=false -b -c -v -D . + $(ISABELLE) build -o document=false -o browser_info=false -b -c -v -D . + +# This is invoked from the dune file to run the tests. +# The tex files might be readonly in the sandbox and Isabelle copies +# them several times, and thus fails by overwriting a readonly file. +dune-runtest: + chmod 644 examples/document/root.tex document/root.tex + $(SRC_ROOT)/deps/isabelle/Isabelle/bin/isabelle build -o document=false -o browser_info=false -c -v -D . clean: - rm -rf output/ + rm -rf output/ examples/output/ diff --git a/isabelle/ROOT b/isabelle/ROOT index dc2d6c1b..03f8d057 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -21,7 +21,7 @@ session "TLA+Examples" in examples = "TLA+" + theories Allocator - AtomicBakeryG + (* AtomicBakeryG TODO: Restore it here, when examples are fixed. *) document_files "root.tex" diff --git a/isabelle/dune b/isabelle/dune new file mode 100644 index 00000000..49f60647 --- /dev/null +++ b/isabelle/dune @@ -0,0 +1,13 @@ +; We use the promoted version of the Isabelle here to have proper file permissions. +; The promoted version is in the source tree, which is located based on ${DUNE_SOURCEROOT} environment variable. +; We assume the project will be built before running the tests, otherwise the Isabelle is always rebuilt. + +(rule + (alias runtest) + (deps + ; (alias_rec ../deps/isabelle/default) -- We don't use it here, because otherwise the isabelle is always rebuilt. + (source_tree ".")) + (action + (run make dune-runtest))) + +(data_only_dirs document examples tests) diff --git a/tlapm.opam b/tlapm.opam index b83e8eed..ed0fc086 100644 --- a/tlapm.opam +++ b/tlapm.opam @@ -25,7 +25,7 @@ authors: [ homepage: "https://github.com/tlaplus/tlapm" bug-reports: "https://github.com/tlaplus/tlapm/issues" depends: [ - "dune" {>= "3.8"} + "dune" {>= "3.12"} "ocaml" "dune-site" "dune-build-info" From 972a39417a87abd60e332338dd9660aa599029b1 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Fri, 31 May 2024 06:54:00 +0300 Subject: [PATCH 140/167] =?UTF-8?q?Rewrite=20Nat=20to=20{x=20=E2=88=88=20I?= =?UTF-8?q?nt=20:=200=20=E2=89=A4=20x}=20instead=20of=20n=20=E2=88=88=20Na?= =?UTF-8?q?t=20to=20(n=20=E2=88=88=20Int=20=E2=88=A7=200=20=E2=89=A4=20n).?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Signed-off-by: Karolis Petrauskas --- isabelle/IntegerArithmetic.thy | 22 +++++++++++++++++++--- isabelle/Tuples.thy | 2 +- isabelle/examples/AtomicBakeryG.thy | 10 +++++----- isabelle/tests/Tests.thy | 20 ++++++++++++++++++++ test/TOOLS/do_one_test | 2 +- 5 files changed, 46 insertions(+), 10 deletions(-) diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy index 6002a476..81d651a4 100644 --- a/isabelle/IntegerArithmetic.thy +++ b/isabelle/IntegerArithmetic.thy @@ -816,9 +816,25 @@ text \ From now on, we reduce the set @{text "Nat"} to the set of non-negative integers. \ -lemma nat_iff_int_geq0 [simp]: "n \ Nat = (n \ Int \ 0 \ n)" +lemma nat_iff_int_geq0' (*[simp]*) : "n \ Nat = (n \ Int \ 0 \ n)" by (auto simp: int_leq_def) +text \ +We will use \nat_iff_int_geq0\ for simplification instead +of \nat_iff_int_geq0'\ for rewriting \Nat\ to \Int\ to +handle also the cases where no \Nat\ is used without +the \\\ operator, e.g. in function sets. +\ +lemma nat_iff_int_geq0 [simp] : "Nat = {x \ Int : 0 \ x}" +proof + show "\x. x \ Nat \ x \ {x \ Int : 0 \ x}" + using nat_iff_int_geq0' by auto +next + show "\x. x \ {x \ Int : 0 \ x} \ x \ Nat" + using nat_iff_int_geq0' by auto +qed + + declare natIsInt [simp del] declare succInNat [simp del] declare predInNat [simp del] @@ -1696,12 +1712,12 @@ lemma trans_leq_diff_nat2 [simp]: lemma int_leq_iff_add: assumes "i \ Int" and "j \ Int" shows "(i \ j) = (\k \ Nat: j = i + k)" (is "?lhs = ?rhs") - using assms by (auto intro: int_leq_imp_add) + using assms by (auto intro: int_leq_imp_add simp del: nat_iff_int_geq0 simp add: nat_iff_int_geq0') lemma int_less_iff_succ_add: assumes "i \ Int" and "j \ Int" shows "(i < j) = (\k \ Nat: j = succ[i + k])" (is "?lhs = ?rhs") - using assms by (auto intro: int_less_imp_succ_add) + using assms by (auto intro: int_less_imp_succ_add simp del: nat_iff_int_geq0 simp add: nat_iff_int_geq0') lemma leq_add_self1 [simp]: assumes "i \ Int" and "j \ Int" diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 1030ff1f..5981db15 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -129,7 +129,7 @@ lemma SeqI [intro]: shows "s \ Seq(S)" proof - from assms have "s \ [1 .. Len(s) \ S]" by auto - with s show ?thesis unfolding Seq_def by auto + with s show ?thesis unfolding Seq_def by (auto simp del: nat_iff_int_geq0) qed lemma SeqI': \ \ closer to the definition but probably less useful \ diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index 8e0f0a91..265c37e6 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -434,7 +434,7 @@ proof auto pc[k] = ''p3'' \ num[i] \ max[k] \ (pc[k] = ''p4'' \ pc[k] = ''p5'' \ pc[k] = ''p6'') \ GG(i, k, num) \ (pc[k] = ''p5'' \ pc[k] = ''p6'' \ i \ unread[k])" - from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. + from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" by auto \ \ iinv3 and iinv5: particular parts of the invariant, taken to the meta-level for then being instantiated with the proper variables, @@ -447,7 +447,7 @@ proof auto fix i j assume pci: "pc[i] \ {''p5'', ''p6''}" and i: "i \ ProcSet" and j:"j \ ProcSet \ unread[i] \ {i}" - from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. + from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" by auto hence "pc[i] \ {''p5'', ''p6''} \ (\j \ ProcSet \ unread[i] \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" @@ -465,7 +465,7 @@ proof auto fix i j assume pci: "pc[i] \ {''p7'', ''p8''}" and i: "i \ ProcSet" and j:"j \ ProcSet \\ {i}" - from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" .. + from iinv i have iinvi: "IInv(i, unread, max, flag, pc, num, nxt)" by auto hence "pc[i] \ {''p7'', ''p8''} \ (\j \ ProcSet \ {i} : After(j, i, unread, max, flag, pc, num, nxt))" @@ -568,7 +568,7 @@ proof auto fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" - unfolding IInv_def by auto + unfolding IInv_def using After_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True @@ -602,7 +602,7 @@ proof auto fix j assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" from unch pc' i j iinvi have after: "After(j,i,unread,max,flag,pc,num,nxt)" - unfolding IInv_def by auto + unfolding IInv_def using After_def by auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True diff --git a/isabelle/tests/Tests.thy b/isabelle/tests/Tests.thy index 7ebe8ba3..4f19d7a9 100644 --- a/isabelle/tests/Tests.thy +++ b/isabelle/tests/Tests.thy @@ -48,4 +48,24 @@ lemma by auto +section \Nat vs Functions\ + +lemma \ \Works with abstract types, by auto.\ + fixes p :: "c" and f and i + assumes "f \ [p \ a]" and "i \ p" + shows "f[i] \ a" + using assms by auto + +text \ +The following only started to work after changing the \Nat\ to \Int\ +rewrite from \n \ Nat = (n \ Int \ 0 \ n)\ to \Nat = {x \ Int : 0 \ x}\ +see \nat_iff_int_geq0\. +\ +lemma + fixes p :: "c" and f and i + assumes "f \ [p \ Nat]" and "i \ p" + shows "f[i] \ Nat" + using assms by auto + + end diff --git a/test/TOOLS/do_one_test b/test/TOOLS/do_one_test index 8706f963..f0c45445 100755 --- a/test/TOOLS/do_one_test +++ b/test/TOOLS/do_one_test @@ -119,7 +119,7 @@ phase == 1 && FILENAME ~ /.*\/separator$/ { result = system (check_command); if (result != 0){ fail("explicit check failed"); } }else if (!command_done){ - run_command("${TLAPM} --toolbox 0 0 --nofp ${FILE} --stretch ${STRETCH}"); + run_command("${TLAPM} --toolbox 0 0 --nofp ${FILE} --stretch ${STRETCH} --debug=tempfiles "); } if (result_code != expected_result){ fail("wrong result code"); From 2b0aa197703e9a17f2a7e7bdd70a89d4f65429b2 Mon Sep 17 00:00:00 2001 From: Markus Alexander Kuppe Date: Fri, 7 Jun 2024 09:41:16 -0700 Subject: [PATCH 141/167] Also build on arm64 macos (M1). https://github.blog/changelog/2024-01-30-github-actions-introducing-the-new-m1-macos-runner-available-to-open-source/ Related to Github issue #97 https://github.com/tlaplus/tlapm/issues/97 [Build] Signed-off-by: Markus Alexander Kuppe --- .github/workflows/ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 95b86003..0edb7eb7 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -11,7 +11,7 @@ jobs: strategy: fail-fast: false matrix: - os: [ubuntu-latest, macos-latest] + os: [ubuntu-latest, macos-latest, macos-14] ocaml-compiler: ['0', '1'] env: EXAMPLES_DIR: "tlaplus-examples" From 70db8d02ca604b24b8877338abee5e7f89ad5062 Mon Sep 17 00:00:00 2001 From: Stephan Merz Date: Mon, 10 Jun 2024 18:57:11 +0200 Subject: [PATCH 142/167] changes to Isabelle setup and fixed AtomicBakeryG example Signed-off-by: Stephan Merz --- isabelle/CaseExpressions.thy | 4 +- isabelle/Constant.thy | 4 +- isabelle/FixedPoints.thy | 6 +- isabelle/Functions.thy | 6 +- isabelle/IntegerArithmetic.thy | 62 ++++++++------ isabelle/IntegerDivision.thy | 4 +- isabelle/Integers.thy | 6 +- isabelle/PredicateLogic.thy | 6 +- isabelle/ROOT | 3 +- isabelle/SetTheory.thy | 18 +++-- isabelle/Strings.thy | 6 +- isabelle/Tuples.thy | 4 +- isabelle/Zenon.thy | 4 +- isabelle/examples/Allocator.thy | 4 +- isabelle/examples/AtomicBakeryG.thy | 121 +++++++++++++++------------- 15 files changed, 144 insertions(+), 114 deletions(-) diff --git a/isabelle/CaseExpressions.thy b/isabelle/CaseExpressions.thy index f9c6cbeb..c0539ce4 100644 --- a/isabelle/CaseExpressions.thy +++ b/isabelle/CaseExpressions.thy @@ -1,8 +1,8 @@ (* Title: TLA+/CaseExpressions.thy Author: Stephan Merz, Inria Nancy - Copyright (C) 2009-2022 INRIA and Microsoft Corporation + Copyright (C) 2009-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ Case expressions \ diff --git a/isabelle/Constant.thy b/isabelle/Constant.thy index d085827c..4ce4c15f 100644 --- a/isabelle/Constant.thy +++ b/isabelle/Constant.thy @@ -1,8 +1,8 @@ (* Title: TLA+/Constant.thy Author: Stephan Merz, LORIA - Copyright (C) 2008-2022 INRIA and Microsoft Corporation + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ Main theory for constant-level Isabelle/\tlaplus{} \ diff --git a/isabelle/FixedPoints.thy b/isabelle/FixedPoints.thy index 28162aea..77237265 100644 --- a/isabelle/FixedPoints.thy +++ b/isabelle/FixedPoints.thy @@ -1,8 +1,8 @@ (* Title: TLA+/FixedPoints.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2024 *) section \Fixed points for set-theoretical constructions\ diff --git a/isabelle/Functions.thy b/isabelle/Functions.thy index 6ad4cf88..1ed74966 100644 --- a/isabelle/Functions.thy +++ b/isabelle/Functions.thy @@ -1,8 +1,8 @@ (* Title: TLA+/Functions.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2024 *) section \ \tlaplus{} Functions \ diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy index 81d651a4..12a7f692 100644 --- a/isabelle/IntegerArithmetic.thy +++ b/isabelle/IntegerArithmetic.thy @@ -1,8 +1,8 @@ (* Title: TLA+/IntegerArithmetic.thy Author: Stephan Merz, Inria Nancy - Copyright (C) 2009-2022 INRIA and Microsoft Corporation + Copyright (C) 2009-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ Arithmetic (except division) for the integers \ @@ -719,20 +719,33 @@ by simp lemma less_not_refl: "a < b \ a \ b" by auto -lemma neq_leq_trans [trans]: "a \ b \ a \ b \ a < b" +lemma neq_leq_trans [trans,intro]: "a \ b \ a \ b \ a < b" by (simp add: less_def) -declare neq_leq_trans[simplified,trans] +declare neq_leq_trans[simplified,trans,intro] + +lemma neq_leq_trans_sym [trans,intro]: "b \ a \ a \ b \ a < b" +by (simp add: less_def) + +declare neq_leq_trans_sym[simplified,trans,intro] lemma leq_neq_trans [trans,elim!]: "a \ b \ a \ b \ a < b" by (simp add: less_def) -declare leq_neq_trans[simplified,trans] +declare leq_neq_trans[simplified,trans,elim!] + +lemma leq_neq_trans_sym [trans,elim!]: "a \ b \ b \ a \ a < b" +by (simp add: less_def) + +declare leq_neq_trans_sym[simplified,trans,elim!] (* Don't add to [simp]: will be tried on all disequalities! *) lemma leq_neq_iff_less: "a \ b \ (a \ b) = (a < b)" by auto +lemma leq_neq_iff_less_sym: "a \ b \ (b \ a) = (a < b)" +by auto + subsection \ Facts about @{text "\"} over @{text Int} \ lemma int_leq_refl [intro,simp]: "a \ Int \ a \ a" @@ -825,7 +838,7 @@ of \nat_iff_int_geq0'\ for rewriting \Nat\ to \I handle also the cases where no \Nat\ is used without the \\\ operator, e.g. in function sets. \ -lemma nat_iff_int_geq0 [simp] : "Nat = {x \ Int : 0 \ x}" +lemma nat_is_int_geq0 [simp] : "Nat = {x \ Int : 0 \ x}" proof show "\x. x \ Nat \ x \ {x \ Int : 0 \ x}" using nat_iff_int_geq0' by auto @@ -1087,14 +1100,7 @@ lemma nat_leq_induct: \ \sometimes called ``complete induction'' shows "\n\Nat : P(n)" proof - have "\n\Nat : \m\Nat : m \ n \ P(m)" - proof (rule natInduct) - from base show "\m \ Nat : m \ 0 \ P(m)" by auto - next - fix n - assume "n \ Nat" "\m \ Nat : m \ n \ P(m)" - with step show "\m \ Nat : m \ succ[n] \ P(m)" - by (auto simp del: nat_iff_int_geq0) - qed + using assms by (intro natInduct) auto thus ?thesis by auto qed @@ -1179,12 +1185,12 @@ lemma int_lessE: by (simp add: int_less_iff_zero_less_diff[of "a" "b"]) lemma int_leq_not_less: - assumes "a \ Int" and "b \ Int" and "a \ b" + assumes "a \ b" and "a \ Int" and "b \ Int" shows "(b < a) = FALSE" using assms by (auto simp: less_def dest: int_leq_antisym) lemma int_less_not_leq: - assumes "a \ Int" and "b \ Int" and "a < b" + assumes "a < b" and "a \ Int" and "b \ Int" shows "(b \ a) = FALSE" using assms by (auto simp: less_def dest: int_leq_antisym) @@ -1429,7 +1435,7 @@ proof (rule int_leq_lessE[OF ij]) thus ?thesis by auto next assume "i = j" - with i f show ?thesis by (auto simp del: nat_iff_int_geq0) + with i f show ?thesis by auto qed (auto simp: i[simplified] j[simplified]) lemma int_succ_lessI: @@ -1629,12 +1635,12 @@ lemma leq_add_nat2 [iff]: shows "b \ a + b" using assms add_leq_right_mono [of 0 a b] by simp -lemma trans_leq_add_nat1 [simp] (* was [elim]*): +lemma trans_leq_add_nat1 [simp,intro]: assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" shows "a \ b+m" using assms by (simp add: int_leq_trans[of "a" "b" "b+m"]) -lemma trans_leq_add_nat2 [simp] (* was [elim] *): +lemma trans_leq_add_nat2 [simp,intro]: assumes "a \ b" and "a \ Int" and "b \ Int" and "m \ Int" and "0 \ m" shows "a \ m+b" using assms by (simp add: int_leq_trans[of "a" "b" "m+b"]) @@ -1712,12 +1718,12 @@ lemma trans_leq_diff_nat2 [simp]: lemma int_leq_iff_add: assumes "i \ Int" and "j \ Int" shows "(i \ j) = (\k \ Nat: j = i + k)" (is "?lhs = ?rhs") - using assms by (auto intro: int_leq_imp_add simp del: nat_iff_int_geq0 simp add: nat_iff_int_geq0') + using assms by (auto intro: int_leq_imp_add simp del: nat_is_int_geq0 simp add: nat_iff_int_geq0') lemma int_less_iff_succ_add: assumes "i \ Int" and "j \ Int" shows "(i < j) = (\k \ Nat: j = succ[i + k])" (is "?lhs = ?rhs") - using assms by (auto intro: int_less_imp_succ_add simp del: nat_iff_int_geq0 simp add: nat_iff_int_geq0') + using assms by (auto intro: int_less_imp_succ_add simp del: nat_is_int_geq0 simp add: nat_iff_int_geq0') lemma leq_add_self1 [simp]: assumes "i \ Int" and "j \ Int" @@ -1734,16 +1740,26 @@ lemma leq_add_self2 [simp]: shows "(j + i \ i) = (j \ 0)" using assms by (simp add: add_commute) -lemma trans_less_add_nat1 [simp]: +lemma trans_less_add_nat1 [simp,intro]: assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" shows "i < j + m" using assms by (simp add: int_less_leq_trans[of "i" "j" "j+m"]) -lemma trans_less_add_nat2 [simp]: +lemma trans_less_add_nat2 [simp,intro]: assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" shows "i < m + j" using assms by (simp add: int_less_leq_trans[of "i" "j" "m+j"]) +lemma trans_leq_less_add_nat1 [simp,intro]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 < m" + shows "i < j + m" + using assms add_leq_less_mono[of "i" "j" "0" "m"] by simp + +lemma trans_leq_less_add_nat2 [simp,intro]: + assumes "i \ j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 < m" + shows "i < m + j" + using assms add_leq_less_mono[of "i" "j" "0" "m"] by (simp add: add_commute) + lemma trans_less_diff_nat1: assumes "i < j" and "i \ Int" and "j \ Int" and "m \ Int" and "0 \ m" shows "i - m < j" diff --git a/isabelle/IntegerDivision.thy b/isabelle/IntegerDivision.thy index 2da1b4eb..ee603162 100644 --- a/isabelle/IntegerDivision.thy +++ b/isabelle/IntegerDivision.thy @@ -1,8 +1,8 @@ (* Title: TLA+/IntegerDivision.thy Author: Stephan Merz, Inria Nancy - Copyright (C) 2009-2022 INRIA and Microsoft Corporation + Copyright (C) 2009-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ The division operators div and mod on the integers \ diff --git a/isabelle/Integers.thy b/isabelle/Integers.thy index c33f40d9..a429ee7b 100644 --- a/isabelle/Integers.thy +++ b/isabelle/Integers.thy @@ -1,8 +1,8 @@ (* Title: TLA+/Integers.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ The set of integer numbers \ diff --git a/isabelle/PredicateLogic.thy b/isabelle/PredicateLogic.thy index 6aaf351d..b6f43ed9 100644 --- a/isabelle/PredicateLogic.thy +++ b/isabelle/PredicateLogic.thy @@ -1,8 +1,8 @@ (* Title: TLA+/PredicateLogic.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2024 *) section \First-Order Logic for TLA+\ diff --git a/isabelle/ROOT b/isabelle/ROOT index 03f8d057..38f64449 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -10,7 +10,7 @@ session "TLA+" = "Pure" + document_files "root.tex" - +(* session "TLA+Tests" in tests = "TLA+" + options [document=false] theories @@ -25,3 +25,4 @@ session "TLA+Examples" in examples = "TLA+" + document_files "root.tex" +*) \ No newline at end of file diff --git a/isabelle/SetTheory.thy b/isabelle/SetTheory.thy index 9bfb1ac8..2877d661 100644 --- a/isabelle/SetTheory.thy +++ b/isabelle/SetTheory.thy @@ -1,8 +1,8 @@ (* Title: TLA+/SetTheory.thy - Author: Stephan Merz, LORIA - Copyright (C) 2008-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2008-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2024 *) section \\tlaplus{} Set Theory\ @@ -378,21 +378,23 @@ lemma bspec: shows "P(x)" using assms unfolding bAll_def by blast +setup \ + map_theory_claset (fn ctxt => + ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) +\ + +(** text \The following rule intentionally has single assumption (non-conditional), because otherwise it interferes with how \Nat\ is reduced to \Int\.\ lemma bspec' [dest]: assumes "\x\A : P(x)" shows "\x : (x\A) \ P(x)" using assms unfolding bAll_def by blast +**) lemma bAll_unb [simp] : "\T P. (\e : e \ T \ P(e)) \ (\e \ T : P(e))" using bAll_def by simp -setup \ - map_theory_claset (fn ctxt => - ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) -\ - lemma bAllE [elim]: assumes "\x\A : P(x)" and "x \ A \ Q" and "P(x) \ Q" shows "Q" diff --git a/isabelle/Strings.thy b/isabelle/Strings.thy index da52b7a4..7d0cebf1 100644 --- a/isabelle/Strings.thy +++ b/isabelle/Strings.thy @@ -1,8 +1,8 @@ (* Title: TLA+/Strings.thy - Author: Stephan Merz, LORIA - Copyright (C) 2009-2021 INRIA and Microsoft Corporation + Author: Stephan Merz, Inria Nancy + Copyright (C) 2009-2024 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2020 + Version: Isabelle2024 *) section \ Characters and strings \ diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 5981db15..7f5e20f6 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -1,8 +1,8 @@ (* Title: TLA+/Tuples.thy Author: Stephan Merz, Inria Nancy - Copyright (C) 2008-2022 Inria and Microsoft Corporation + Copyright (C) 2008-2024 Inria and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \ Tuples and Relations in \tlaplus{} \ diff --git a/isabelle/Zenon.thy b/isabelle/Zenon.thy index ebd642a7..163cf4d3 100644 --- a/isabelle/Zenon.thy +++ b/isabelle/Zenon.thy @@ -1,7 +1,7 @@ (* Zenon.thy --- Support lemmas for Zenon proofs * Author: Damien Doligez - * Copyright (C) 2008-2022 INRIA and Microsoft Corporation - * Version: Isabelle2021-1 + * Copyright (C) 2008-2024 INRIA and Microsoft Corporation + * Version: Isabelle2024 *) (* isabelle usedir -b Pure TLA+ *) diff --git a/isabelle/examples/Allocator.thy b/isabelle/examples/Allocator.thy index 30542f61..d7debecf 100644 --- a/isabelle/examples/Allocator.thy +++ b/isabelle/examples/Allocator.thy @@ -2,13 +2,13 @@ Author: Stephan Merz, Inria Nancy Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \A Simple Resource Allocator\ theory Allocator -imports Constant +imports "../Constant" begin text \ diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index 265c37e6..340dc762 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -2,13 +2,13 @@ Author: Hernan Vanzetto, LORIA Copyright (C) 2009-2022 INRIA and Microsoft Corporation License: BSD - Version: Isabelle2021-1 + Version: Isabelle2024 *) section \Safety Proof of the Atomic Version of the Bakery Algorithm\ theory AtomicBakeryG -imports Constant +imports "../Constant" begin text \ @@ -294,8 +294,7 @@ next next case False from type have "isAFcn(pc)" - (* FIXME: needed for proving that pc'=pc, but why doesn't "simp ..." suffice *) - by (force simp add: TypeOK_def) + by (auto simp: TypeOK_def) with False p2 obtain i where i: "pc[self] = ''p2''" "i \ unread[self]" "unread' = [unread EXCEPT ![self] = unread[self] \\ {i}]" @@ -409,9 +408,6 @@ proof (auto simp: MutualExclusion_def) qed -lemma leq_neq_trans' (*[dest!]*): "a \ b \ b \ a \ a < b" - by (simp add: less_def) - theorem InvInvariant: assumes inv: "Inv(unread,max,flag,pc,num,nxt)" and nxt: "Next(unread,max,flag,pc,num,nxt,unread',max',flag',pc',num',nxt')" @@ -504,10 +500,15 @@ proof auto show "After(j,i,unread',max',flag',pc',num',nxt')" proof (cases "self = j") case True with selfi p1 type iinvi self i pc' show ?thesis - by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) + (force elim: int_leq_lessE) next - case False with selfi p1 type iinvi self i pc' j show ?thesis - by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + case False + from i type have "num[i] \ Nat" + by (auto simp: TypeOK_def) + with False selfi p1 type iinvi self i pc' j show ?thesis + by (clarsimp simp: TypeOK_def IInv_def After_def p1_def int_leq_less) + (force elim!: bspec) qed qed have 4: "pc'[i] = ''p6'' @@ -534,8 +535,10 @@ proof auto case True with selfi p1 type iinvi self i pc' show ?thesis by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto next - case False with selfi p1 type iinvi self i pc' j show ?thesis - by (clarsimp simp: TypeOK_def IInv_def After_def p1_def) auto + case False + with pc' self selfi p1 type iinvi i j show ?thesis + by (clarsimp simp: p1_def TypeOK_def IInv_def After_def) + (force elim!: bspec) qed qed from 1 2 3 4 5 show ?thesis @@ -642,7 +645,7 @@ proof auto fix j assume pc': "pc'[i] \ {''p5'',''p6''}" and j: "j \ (ProcSet \ unread'[i]) \ {i}" from pc' i k 1 type have numi: "0 < num'[i]" - by (auto simp: TypeOK_def) + by (force simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -672,7 +675,7 @@ proof auto GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + assume "max[j] < num[k]" "num \ [ProcSet \ {x \ Int : 0 \ x}]" with ii2 kproc mx i \pc[j] = ''p2''\ show "num[i] \ num[k]" by (auto elim!: int_leq_trans) @@ -683,8 +686,11 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - by (auto simp: TypeOK_def IInv_def After_def) + from i type have "num[i] \ Int \ 0 \ num[i]" + by (auto simp: TypeOK_def) + with True k selfi selfj nless type iinvi self i pc' j mx + show ?thesis + by (auto simp: TypeOK_def IInv_def After_def elim: int_leq_lessE) next case False with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis @@ -730,7 +736,7 @@ proof auto case True with k selfi type iinvi self i nxt mx show ?thesis proof (clarsimp simp: TypeOK_def IInv_def, cases "i \ unread[nxt[i]]", simp) - assume "i \ unread[nxt[i]]" "num \ [ProcSet \ Nat]" + assume "i \ unread[nxt[i]]" "num \ [ProcSet \ {x \ Int : 0 \ x}]" "i \ unread[nxt[i]] = FALSE \ num[i] \ max[nxt[i]]" "max[nxt[i]] < num[k]" with mx nxt kproc i show "num[i] \ num[k]" @@ -775,7 +781,7 @@ proof auto fix j assume pc': "pc'[i] \ {''p7'',''p8''}" and j: "j \ ProcSet \ {i}" from pc' i k 1 type have numi: "0 < num'[i]" - by (auto simp: TypeOK_def) + by (force simp: TypeOK_def) from j type have mx: "pc[j] = ''p2'' \ max[j] \ Nat" by (auto simp: TypeOK_def) show "After(j,i,unread',max',flag',pc',num',nxt')" @@ -805,7 +811,7 @@ proof auto GG(i, j, num) \ (pc[j] = ''p5'' \ pc[j] = ''p6'' \ i \ unread[j])" hence ii2: "num[i] \ max[j]" by auto - assume "max[j] < num[k]" "num \ [ProcSet \ Nat]" + assume "max[j] < num[k]" "num \ [ProcSet \ {x \ Int : 0 \ x}]" with ii2 kproc mx i \pc[j] = ''p2''\ show "num[i] \ num[k]" by (auto elim!: int_leq_trans) @@ -816,8 +822,11 @@ proof auto show ?thesis proof (cases "i=k") case True - with k selfi selfj nless type iinvi self i pc' j mx show ?thesis - by (auto simp: TypeOK_def IInv_def After_def) + from type i have "num[i] \ Int \ 0 \ num[i]" + by (auto simp: TypeOK_def) + with True k selfi selfj nless type iinvi self i pc' j mx + show ?thesis + by (auto simp: TypeOK_def IInv_def After_def elim: int_leq_lessE) next case False with k selfi selfj nless type iinvi self i pc' j mx numi show ?thesis @@ -860,7 +869,7 @@ proof auto proof (cases "self = i") case True with p3 type iinvi i show ?thesis - by (clarsimp simp: TypeOK_def IInv_def p3_def) + by (clarsimp simp: TypeOK_def IInv_def p3_def) auto next assume selfi: "self \ i" from selfi p3 type iinvi self i @@ -1066,7 +1075,8 @@ proof auto by(auto simp: p5_def) with type self have kproc: "k \ ProcSet" by (auto simp: TypeOK_def) - with iinv have iinvk: "IInv(k, unread, max, flag, pc, num, nxt)" .. + with iinv have iinvk: "IInv(k, unread, max, flag, pc, num, nxt)" + by (rule bspec) show ?thesis proof (cases "self = i") case True @@ -1118,7 +1128,9 @@ proof auto show "After(j, i, unread', max', flag', pc', num', nxt')" proof (cases "self = i") assume selfi: "self = i" - with iinvi p6 type i have numi: "0 < num[i]" "nxt[i] \ ProcSet" "nxt[i] \ i" + with iinvi p6 type i have numi: "0 < num[i]" + by (force simp: IInv_def TypeOK_def p6_def) + from selfi p6 type i have nxtip: "nxt[i] \ ProcSet \ {i}" by (auto simp: IInv_def TypeOK_def p6_def) show ?thesis proof (cases "j \ unread[self]") @@ -1126,9 +1138,12 @@ proof auto show ?thesis proof(cases "num[j] = 0") case True - with j selfi i pc' p6 type self iinvi junread iinvj + from i type have "num[i] \ Nat" + by (auto simp: TypeOK_def) + with j selfi i pc' p6 type self iinvi junread iinvj True show ?thesis - by (clarsimp simp: p6_def TypeOK_def IInv_def After_def) auto + by (clarsimp simp: p6_def TypeOK_def IInv_def After_def int_leq_less) + (force elim!: bspec) next assume numj: "num[j] \ 0" show ?thesis @@ -1146,13 +1161,13 @@ proof auto proof (cases "j < i") case True with j selfi i pc' p6 type self iinvi junread numj numi pcj - iinv3[of j i] nxti notunread (*procInNat*) + iinv3[of j i] nxti notunread show ?thesis proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) assume "nxt[i] < i" "j = nxt[i]" with i j have nxti: "nxt[i] \ ProcSet" "\(i < nxt[i])" by (auto dest: procInNat elim!: int_less_asym) - assume "num \ [ProcSet \ Nat]" + assume "num \ [ProcSet \ {x \ Int : 0 \ x}]" "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ 0 < num[nxt[i]] \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] @@ -1161,7 +1176,7 @@ proof auto with i nxti show "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" - by (auto simp: int_leq_not_less) + by (auto dest!: int_leq_not_less) auto qed next case False @@ -1169,18 +1184,18 @@ proof auto show ?thesis proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) assume "(nxt[i] < i) = FALSE" - with i numi int_less_linear[of "i" "nxt[i]"] have "i < nxt[i]" - by (auto dest: procInNat) + with i numi nxtip int_less_linear[of "i" "nxt[i]"] have "i < nxt[i]" + by (auto dest!: procInNat) assume "pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ 0 < num[nxt[i]] \ (IF i < nxt[i] THEN num[nxt[i]] < num[i] ELSE boolify(num[nxt[i]] \ num[i]))" "num[i] \ num[nxt[i]]" - "num \ [ProcSet \ Nat]" - with i numi \i < nxt[i]\ + "num \ [ProcSet \ {x \ Int : 0 \ x}]" + with i numi nxtip \i < nxt[i]\ show "(pc[nxt[i]] = ''p5'') = FALSE \ (pc[nxt[i]] = ''p6'') = FALSE" - by (auto dest: procInNat simp: int_leq_not_less) + by (auto dest!: int_leq_not_less) qed qed qed @@ -1198,19 +1213,18 @@ proof auto assume nxti: "nxt[i] < i" hence "i \ nxt[i]" by (auto simp: less_def) assume - "i = nxt[i] = FALSE - \ 0 < num[nxt[i]] - \ (IF i < nxt[i] - THEN num[nxt[i]] < num[i] - ELSE boolify(num[nxt[i]] \ num[i]))" - "num \ [ProcSet \ Nat]" + "0 < num[nxt[i]]" + "IF i < nxt[i] + THEN num[nxt[i]] < num[i] + ELSE boolify(num[nxt[i]] \ num[i])" + "num \ [ProcSet \ {x \ Int : 0 \ x}]" with \i \ nxt[i]\ have "num[nxt[i]] \ num[i]" by (auto dest: procInNat elim!: int_less_asym[OF nxti]) assume "num[i] < num[nxt[i]]" - with i numi \num \ [ProcSet \ Nat]\ \num[nxt[i]] \ num[i]\ + with i numi nxtip \num \ [ProcSet \ {x \ Int : 0 \ x}]\ \num[nxt[i]] \ num[i]\ show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" - by (auto simp: int_less_not_leq) + by (auto dest!: int_less_not_leq) qed next case False @@ -1219,16 +1233,16 @@ proof auto show ?thesis proof (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) assume "(nxt[i] < i) = FALSE" - with i numi int_less_linear[of "nxt[i]" "i"] have "i < nxt[i]" + with i numi nxtip int_less_linear[of "nxt[i]" "i"] have "i < nxt[i]" by (auto dest: procInNat) assume "num[i] \ num[nxt[i]]" "IF i < nxt[i] THEN num[nxt[i]] < num[i] ELSE boolify(num[nxt[i]] \ num[i])" - "num \ [ProcSet \ Nat]" - with i numi \i < nxt[i]\ + "num \ [ProcSet \ {x \ Int : 0 \ x}]" + with i numi nxtip \i < nxt[i]\ show "(pc[nxt[i]] = ''p4'' \ pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'') \ (pc[nxt[i]] = ''p5'' \ pc[nxt[i]] = ''p6'' \ i \ unread[nxt[i]])" - by (auto simp: int_less_not_leq) + by (auto dest!: int_less_not_leq) qed qed next @@ -1258,8 +1272,7 @@ proof auto with j pc' selfi selfj i p6 type self iinv3[of i j] nxtj show ?thesis by (clarsimp simp: p6_def TypeOK_def After_def GG_def) - (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] - elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) + (auto dest!: int_leq_not_less int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False with j pc' selfi selfj i p6 type self iinv3[of i j] @@ -1276,10 +1289,9 @@ proof auto \ (IF nxt[j] < j THEN num[j] < num[nxt[j]] ELSE num[j] \ num[nxt[j]])" "0 < num[i]" "IF j < i THEN num[i] < num[j] ELSE num[i] \ num[j]" - "num \ [ProcSet \ Nat]" + "num \ [ProcSet \ {x \ Int : 0 \ x}]" with i j \nxt[j] < j\ show "(i = nxt[j]) = FALSE" - by (auto simp: int_less_not_leq[of "num[j]" "num[nxt[j]]"] - elim!: int_less_asym[of "num[nxt[j]]" "num[j]"]) + by (auto dest!: int_leq_not_less int_less_asym[of "num[nxt[j]]" "num[j]"]) qed qed next @@ -1318,8 +1330,7 @@ proof auto with type p6 i pc' j self selfj iinv5[of i j] show ?thesis by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) - (auto simp: int_less_not_leq[of "num[nxt[j]]" "num[j]"] - elim!: int_less_asym[of "num[j]" "num[nxt[j]]"]) + (auto dest!: int_leq_not_less int_less_asym[of "num[j]" "num[nxt[j]]"]) next case False from i j type have num: "num[i] \ Nat" "num[j] \ Nat" @@ -1328,7 +1339,7 @@ proof auto show ?thesis using int_less_linear[of i j] procInNat by (clarsimp simp: p6_def TypeOK_def IInv_def After_def GG_def) - (auto dest!: int_leq_less_trans) + (auto dest!: int_leq_not_less) qed next case False @@ -1338,7 +1349,7 @@ proof auto qed qed from 1 2 3 4 5 show ?thesis - unfolding IInv_def by blast + unfolding IInv_def by simp next assume p7: "p7(self, unread, max, flag, pc, num, nxt, unread', max', flag', pc', num', nxt')" From 932c682b7918d20dd8baf1b0b34206e3931f078d Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Tue, 11 Jun 2024 08:37:12 +0300 Subject: [PATCH 143/167] A cleanup. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 3 ++- isabelle/IntegerArithmetic.thy | 2 +- isabelle/Makefile | 2 +- isabelle/ROOT | 5 ++--- isabelle/Tuples.thy | 2 +- isabelle/examples/Allocator.thy | 2 +- isabelle/examples/AtomicBakeryG.thy | 2 +- isabelle/tests/Tests.thy | 2 +- 8 files changed, 10 insertions(+), 10 deletions(-) diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 28e0dae8..0cb2265c 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -64,7 +64,8 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && echo "ISABELLE_OUTPUT=$$HEAPS_PATH" >> etc/settings mkdir -p $(ISABELLE_DIR)/src/TLA+ \ && cp -a ../../isabelle/* $(ISABELLE_DIR)/src/TLA+/ \ - && chmod -R u+w $(ISABELLE_DIR)/src/TLA+/ + && chmod -R u+w $(ISABELLE_DIR)/src/TLA+/ \ + && make -C $(ISABELLE_DIR)/src/TLA+/ clean cd $(ISABELLE_DIR)/ \ && ./bin/isabelle build -o system_heaps -o document=false -b -v -d src/Pure Pure \ && ./bin/isabelle build -o system_heaps -o document=false -b -c -v -d src/TLA+ TLA+ \ diff --git a/isabelle/IntegerArithmetic.thy b/isabelle/IntegerArithmetic.thy index 12a7f692..2ca45352 100644 --- a/isabelle/IntegerArithmetic.thy +++ b/isabelle/IntegerArithmetic.thy @@ -833,7 +833,7 @@ lemma nat_iff_int_geq0' (*[simp]*) : "n \ Nat = (n \ Int \ 0 \ by (auto simp: int_leq_def) text \ -We will use \nat_iff_int_geq0\ for simplification instead +We will use \nat_is_int_geq0\ for simplification instead of \nat_iff_int_geq0'\ for rewriting \Nat\ to \Int\ to handle also the cases where no \Nat\ is used without the \\\ operator, e.g. in function sets. diff --git a/isabelle/Makefile b/isabelle/Makefile index 617e536e..aa2b2fce 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -18,7 +18,7 @@ heap-only: # This is invoked from the dune file to run the tests. # The tex files might be readonly in the sandbox and Isabelle copies # them several times, and thus fails by overwriting a readonly file. -dune-runtest: +dune-runtest: clean chmod 644 examples/document/root.tex document/root.tex $(SRC_ROOT)/deps/isabelle/Isabelle/bin/isabelle build -o document=false -o browser_info=false -c -v -D . diff --git a/isabelle/ROOT b/isabelle/ROOT index 38f64449..dc2d6c1b 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -10,7 +10,7 @@ session "TLA+" = "Pure" + document_files "root.tex" -(* + session "TLA+Tests" in tests = "TLA+" + options [document=false] theories @@ -21,8 +21,7 @@ session "TLA+Examples" in examples = "TLA+" + theories Allocator - (* AtomicBakeryG TODO: Restore it here, when examples are fixed. *) + AtomicBakeryG document_files "root.tex" -*) \ No newline at end of file diff --git a/isabelle/Tuples.thy b/isabelle/Tuples.thy index 7f5e20f6..a6b39697 100644 --- a/isabelle/Tuples.thy +++ b/isabelle/Tuples.thy @@ -129,7 +129,7 @@ lemma SeqI [intro]: shows "s \ Seq(S)" proof - from assms have "s \ [1 .. Len(s) \ S]" by auto - with s show ?thesis unfolding Seq_def by (auto simp del: nat_iff_int_geq0) + with s show ?thesis unfolding Seq_def by (auto simp del: nat_is_int_geq0) qed lemma SeqI': \ \ closer to the definition but probably less useful \ diff --git a/isabelle/examples/Allocator.thy b/isabelle/examples/Allocator.thy index d7debecf..cb67e2dd 100644 --- a/isabelle/examples/Allocator.thy +++ b/isabelle/examples/Allocator.thy @@ -8,7 +8,7 @@ section \A Simple Resource Allocator\ theory Allocator -imports "../Constant" +imports Constant begin text \ diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index 340dc762..3e627e1b 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -8,7 +8,7 @@ section \Safety Proof of the Atomic Version of the Bakery Algorithm\ theory AtomicBakeryG -imports "../Constant" +imports Constant begin text \ diff --git a/isabelle/tests/Tests.thy b/isabelle/tests/Tests.thy index 4f19d7a9..6ff4a798 100644 --- a/isabelle/tests/Tests.thy +++ b/isabelle/tests/Tests.thy @@ -59,7 +59,7 @@ lemma \ \Works with abstract types, by auto.\ text \ The following only started to work after changing the \Nat\ to \Int\ rewrite from \n \ Nat = (n \ Int \ 0 \ n)\ to \Nat = {x \ Int : 0 \ x}\ -see \nat_iff_int_geq0\. +see \nat_is_int_geq0\. \ lemma fixes p :: "c" and f and i From 6e0ab8d1a999080265ad0620df9e06e65d1a8a08 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sat, 17 Aug 2024 18:56:59 +0300 Subject: [PATCH 144/167] Make examples/AtomicBakeryG.thy proof a bit faster. Signed-off-by: Karolis Petrauskas --- isabelle/examples/AtomicBakeryG.thy | 49 +++++++++++++++++++++++++++++ 1 file changed, 49 insertions(+) diff --git a/isabelle/examples/AtomicBakeryG.thy b/isabelle/examples/AtomicBakeryG.thy index 3e627e1b..20aa97ad 100644 --- a/isabelle/examples/AtomicBakeryG.thy +++ b/isabelle/examples/AtomicBakeryG.thy @@ -1140,10 +1140,59 @@ proof auto case True from i type have "num[i] \ Nat" by (auto simp: TypeOK_def) + (* Here \show ?thesis\ works, as shown here, but takes long time. + Thus tried to decompose the step. + with j selfi i pc' p6 type self iinvi junread iinvj True show ?thesis by (clarsimp simp: p6_def TypeOK_def IInv_def After_def int_leq_less) (force elim!: bspec) + *) + show ?thesis + proof - + have asd: "0 < num'[i]" + proof - + have "TypeOK (unread', max', flag', pc', num', nxt')" + using type self p TypeOKInvariant by auto + hence "num'[i] \ Nat" + using i TypeOK_def by auto + moreover have "num'[i] \ 0" + using selfi p6 iinvi by (simp add: p6_def IInv_def) + ultimately show ?thesis + by simp + qed + moreover have " + (pc'[j] = ''p1'') \ + (pc'[j] = ''p2'' \ (i \ unread'[j] \ num'[i] \ max'[j])) \ + (pc'[j] = ''p3'' \ num'[i] \ max'[j]) \ + (pc'[j] \ {''p4'', ''p5'', ''p6''} \ GG(i, j, num') \ (pc'[j] \ {''p5'', ''p6''} \ i \ unread'[j]))" + proof (rule,rule,rule) + assume aa1: "\(pc'[j] = ''p1'')" + and aa2: "\(pc'[j] = ''p2'' \ (i \ unread'[j] \ num'[i] \ max'[j]))" + and aa3: "\(pc'[j] = ''p3'' \ num'[i] \ max'[j])" + show "(pc'[j] \ {''p4'', ''p5'', ''p6''} \ GG(i, j, num') \ (pc'[j] \ {''p5'', ''p6''} \ i \ unread'[j]))" + proof - + have "pc'[j] \ {''p4'', ''p5'', ''p6''}" (* 35s *) + using aa1 aa2 aa3 + using j selfi i pc' p6 type iinvi junread iinvj True + by (clarsimp simp: p6_def TypeOK_def IInv_def) + (force elim!: bspec) + moreover have "GG(i, j, num')" (* 30s *) + using aa1 aa2 aa3 + using j selfi i pc' p6 type iinvi junread iinvj True + by (clarsimp simp: p6_def TypeOK_def IInv_def) + (force elim!: bspec) + moreover have "(pc'[j] \ {''p5'', ''p6''} \ i \ unread'[j])" (* 4s *) + using j selfi i pc' p6 type iinvi junread iinvj True + by (clarsimp simp: p6_def TypeOK_def IInv_def) + (force elim!: bspec) + ultimately show ?thesis by + auto + qed + qed + ultimately show ?thesis + unfolding After_def by simp + qed next assume numj: "num[j] \ 0" show ?thesis From 04996b1c33f6b981a80964b6604b408c7aa48f8b Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 11:47:40 +0300 Subject: [PATCH 145/167] Run Isabelle THY tests on a clear Isabelle installation. Signed-off-by: Karolis Petrauskas --- deps/isabelle/Makefile | 8 ++++++-- deps/isabelle/dune | 3 +-- dune-project | 2 +- isabelle/Makefile | 10 +++------- isabelle/dune | 6 +++--- isabelle/dune.mk | 10 ++++++++++ test/dune | 3 ++- tlapm.opam | 2 +- 8 files changed, 27 insertions(+), 17 deletions(-) create mode 100644 isabelle/dune.mk diff --git a/deps/isabelle/Makefile b/deps/isabelle/Makefile index 0cb2265c..f1b80590 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/Makefile @@ -44,19 +44,23 @@ $(ISABELLE_ARCHIVE): $(CACHE_DIR)/$(ISABELLE_ARCHIVE) ln -s $< $@ # Extract the isabelle archive. +# $(ISABELLE_DIR) is for the release and +# $(ISABELLE_DIR)-test is for running the isabelle TLA+ tests. $(ISABELLE_DIR): $(ISABELLE_ARCHIVE) + rm -rf $(ISABELLE_DIR) $(ISABELLE_DIR)-test ifeq ($(ISABELLE_ARCHIVE_TYPE),tgz) - rm -rf $(ISABELLE_DIR) tar -xzf $< mv $(ISABELLE_ARCHIVE_DIR) $(ISABELLE_DIR) endif + cp -r $(ISABELLE_DIR) $(ISABELLE_DIR)-test # Build the TLA+ theory. .PRECIOUS: $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) cd $(ISABELLE_DIR) \ && rm -rf contrib/ProofGeneral* doc heaps/*/HOL contrib/vscodium* contrib/vscode* \ - && sed -i -e 's@^\(contrib/vscode_extension\|contrib/vscodium\|src/Tools/Demo\)@#rm at TLA# \1@' etc/components + && sed -i -e 's@^\(contrib/vscode_extension\|contrib/vscodium\|src/Tools/Demo\)@#rm at TLA# \1@' etc/components \ + && echo 'src/TLA+' >> etc/components cd $(ISABELLE_DIR) \ && HEAPS_PATH=$(shell pwd)/$(ISABELLE_DIR)/heaps \ && if [ "$(HOST_CPU)" = "x86_64" ] ; then sed -i -e 's/^ML_PLATFORM=.*$$/ML_PLATFORM="$${ISABELLE_PLATFORM64:-$$ISABELLE_PLATFORM}"/' etc/settings ; fi \ diff --git a/deps/isabelle/dune b/deps/isabelle/dune index d9d492db..984135ab 100644 --- a/deps/isabelle/dune +++ b/deps/isabelle/dune @@ -3,8 +3,7 @@ ; The generated heaps (Pure and TLA+) are in Isabelle/heaps/polyml-*/. (rule (alias default) - (mode (promote (until-clean))) - (deps "Makefile" (source_tree ../../isabelle)) + (deps "Makefile" (source_tree ../../isabelle) (sandbox none)) (targets (dir Isabelle) "Isabelle.no-links" "Isabelle.exec-files") (action (run "make" "-C" "." "Isabelle.no-links" "Isabelle.exec-files"))) diff --git a/dune-project b/dune-project index e75fec3a..951ab239 100644 --- a/dune-project +++ b/dune-project @@ -1,4 +1,4 @@ -(lang dune 3.12) +(lang dune 3.15) (using dune_site 0.1) diff --git a/isabelle/Makefile b/isabelle/Makefile index aa2b2fce..215cc23f 100644 --- a/isabelle/Makefile +++ b/isabelle/Makefile @@ -1,5 +1,4 @@ ISABELLE=isabelle -SRC_ROOT=$(if $(DUNE_SOURCEROOT),$(DUNE_SOURCEROOT),..) .PHONY: rebuild heap-only dune-runtest clean @@ -15,12 +14,9 @@ rebuild: heap-only: $(ISABELLE) build -o document=false -o browser_info=false -b -c -v -D . -# This is invoked from the dune file to run the tests. -# The tex files might be readonly in the sandbox and Isabelle copies -# them several times, and thus fails by overwriting a readonly file. -dune-runtest: clean - chmod 644 examples/document/root.tex document/root.tex - $(SRC_ROOT)/deps/isabelle/Isabelle/bin/isabelle build -o document=false -o browser_info=false -c -v -D . +# Launch Isabelle / Jedit to edit the theory, examples, tests. +jedit: + $(ISABELLE) jedit -d . Constant.thy clean: rm -rf output/ examples/output/ diff --git a/isabelle/dune b/isabelle/dune index 49f60647..7958eed6 100644 --- a/isabelle/dune +++ b/isabelle/dune @@ -5,9 +5,9 @@ (rule (alias runtest) (deps - ; (alias_rec ../deps/isabelle/default) -- We don't use it here, because otherwise the isabelle is always rebuilt. - (source_tree ".")) + (source_tree ".") + (sandbox none)) (action - (run make dune-runtest))) + (run make -f dune.mk runtest))) (data_only_dirs document examples tests) diff --git a/isabelle/dune.mk b/isabelle/dune.mk new file mode 100644 index 00000000..ca9ec618 --- /dev/null +++ b/isabelle/dune.mk @@ -0,0 +1,10 @@ +## +## This Makefile is called from the dune script. +## We call the Isabelle via make to avoid dune attempts to +## find a rule for building $(ISABELLE). It is already built. +## + +ISABELLE=../deps/isabelle/Isabelle-test/bin/isabelle + +runtest: + $(ISABELLE) build -o document=false -o browser_info=false -c -v -D . diff --git a/test/dune b/test/dune index dff698f2..169b9b00 100644 --- a/test/dune +++ b/test/dune @@ -4,7 +4,8 @@ (alias_rec ../deps/all) (alias_rec ../src/all) (glob_files_rec "*.tla") - (source_tree "TOOLS")) + (source_tree "TOOLS") + (sandbox none)) (action (setenv USE_TLAPM %{exe:../src/tlapm.exe} (setenv USE_LIB ../library diff --git a/tlapm.opam b/tlapm.opam index ed0fc086..bf087f95 100644 --- a/tlapm.opam +++ b/tlapm.opam @@ -25,7 +25,7 @@ authors: [ homepage: "https://github.com/tlaplus/tlapm" bug-reports: "https://github.com/tlaplus/tlapm/issues" depends: [ - "dune" {>= "3.12"} + "dune" {>= "3.15"} "ocaml" "dune-site" "dune-build-info" From c28594c821e6decebf2a6c8f717630672c8ab884 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 15:06:40 +0300 Subject: [PATCH 146/167] Tweaking the build system. Signed-off-by: Karolis Petrauskas --- deps/isabelle/.gitignore | 2 +- deps/isabelle/dune | 12 ++++++-- deps/isabelle/{Makefile => dune.mk} | 44 +++++++++++++---------------- isabelle/dune.mk | 4 ++- 4 files changed, 32 insertions(+), 30 deletions(-) rename deps/isabelle/{Makefile => dune.mk} (80%) diff --git a/deps/isabelle/.gitignore b/deps/isabelle/.gitignore index 20068ef0..f1e395c1 100644 --- a/deps/isabelle/.gitignore +++ b/deps/isabelle/.gitignore @@ -1,3 +1,3 @@ /Isabelle.exec-files -/Isabelle.no-links /Isabelle/ +/Isabelle-test/ diff --git a/deps/isabelle/dune b/deps/isabelle/dune index 984135ab..fd95b6d3 100644 --- a/deps/isabelle/dune +++ b/deps/isabelle/dune @@ -3,9 +3,15 @@ ; The generated heaps (Pure and TLA+) are in Isabelle/heaps/polyml-*/. (rule (alias default) - (deps "Makefile" (source_tree ../../isabelle) (sandbox none)) - (targets (dir Isabelle) "Isabelle.no-links" "Isabelle.exec-files") - (action (run "make" "-C" "." "Isabelle.no-links" "Isabelle.exec-files"))) + (deps + "dune.mk" + (source_tree ../../isabelle) + (sandbox none)) + (targets + (dir "Isabelle") + (dir "Isabelle-test") + "Isabelle.exec-files") + (action (run "make" "-f" "dune.mk"))) (install (section (site (tlapm backends))) diff --git a/deps/isabelle/Makefile b/deps/isabelle/dune.mk similarity index 80% rename from deps/isabelle/Makefile rename to deps/isabelle/dune.mk index f1b80590..1654ec89 100644 --- a/deps/isabelle/Makefile +++ b/deps/isabelle/dune.mk @@ -25,13 +25,14 @@ endif ISABELLE_URL=https://isabelle.in.tum.de/website-$(ISABELLE_VSN)/dist/$(ISABELLE_ARCHIVE) ISABELLE_DIR=Isabelle +ISABELLE_TEST=Isabelle-test # Some defaults, for the case if makefile is called not by the dune build system. PROJECT_ROOT=$(if $(DUNE_SOURCEROOT),$(DUNE_SOURCEROOT),../..) CACHE_DIR=$(PROJECT_ROOT)/_build_cache -all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_DIR).no-links Isabelle.exec-files +all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_TEST) Isabelle.exec-files # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): @@ -43,16 +44,25 @@ $(ISABELLE_ARCHIVE): $(CACHE_DIR)/$(ISABELLE_ARCHIVE) rm -f $@ ln -s $< $@ -# Extract the isabelle archive. -# $(ISABELLE_DIR) is for the release and -# $(ISABELLE_DIR)-test is for running the isabelle TLA+ tests. -$(ISABELLE_DIR): $(ISABELLE_ARCHIVE) - rm -rf $(ISABELLE_DIR) $(ISABELLE_DIR)-test +# Extract the isabelle archive and remove the symlinks. +# TODO: This is is a workaround to eliminate symlinks to directories +# until https://github.com/ocaml/dune/issues/7831 is resolved. +$(ISABELLE_DIR) $(ISABELLE_TEST): $(ISABELLE_ARCHIVE) + rm -rf $(ISABELLE_DIR) ifeq ($(ISABELLE_ARCHIVE_TYPE),tgz) tar -xzf $< mv $(ISABELLE_ARCHIVE_DIR) $(ISABELLE_DIR) endif - cp -r $(ISABELLE_DIR) $(ISABELLE_DIR)-test + cd $(ISABELLE_DIR) && rm -rf ./contrib/e-3.0.03-1/src/lib/ +ifeq ($(OS_TYPE),Darwin) + cd $(ISABELLE_DIR) && cd contrib/jdk-21.0.3/arm64-darwin/ \ + && (find . -type link | xargs rm) \ + && mv zulu-21.jdk/Contents/Home/* ./ + cd $(ISABELLE_DIR) && cd contrib/jdk-21.0.3/x86_64-darwin/ \ + && (find . -type link | xargs rm) \ + && mv zulu-21.jdk/Contents/Home/* ./ +endif + cp -r $(ISABELLE_DIR) $(ISABELLE_TEST) # Build the TLA+ theory. .PRECIOUS: $(ISABELLE_DIR)/src/TLA+ @@ -78,28 +88,12 @@ $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) && rm etc/settings \ && mv etc/settings.target etc/settings -# TODO: This is is a workaround to eliminate symlinks to directories -# until https://github.com/ocaml/dune/issues/7831 is resolved. -.PRECIOUS: $(ISABELLE_DIR).no-links -$(ISABELLE_DIR).no-links: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ - rm -rf $@ - cd $< && rm -rf ./contrib/e-3.0.03-1/src/lib/ -ifeq ($(OS_TYPE),Darwin) - cd $< && cd contrib/jdk-21.0.3/arm64-darwin/ \ - && (find . -type link | xargs rm) \ - && mv zulu-21.jdk/Contents/Home/* ./ - cd $< && cd contrib/jdk-21.0.3/x86_64-darwin/ \ - && (find . -type link | xargs rm) \ - && mv zulu-21.jdk/Contents/Home/* ./ -endif - touch $@ - # TODO: This is a workaround, because the dune install removes all the executable # flags (or sets on all the files). Here we generate a script to restore the flags. -Isabelle.exec-files: $(ISABELLE_DIR).no-links +Isabelle.exec-files: $(ISABELLE_DIR) echo "$(shell find $(ISABELLE_DIR) -type f $(FIND_EXEC))" > $@ clean: - rm -rf $(ISABELLE_ARCHIVE) $(ISABELLE_DIR) $(ISABELLE_DIR).no-links Isabelle.exec-files + rm -rf $(ISABELLE_ARCHIVE) $(ISABELLE_DIR) $(ISABELLE_TEST) Isabelle.exec-files .PHONY: all clean diff --git a/isabelle/dune.mk b/isabelle/dune.mk index ca9ec618..a65a31dc 100644 --- a/isabelle/dune.mk +++ b/isabelle/dune.mk @@ -4,7 +4,9 @@ ## find a rule for building $(ISABELLE). It is already built. ## -ISABELLE=../deps/isabelle/Isabelle-test/bin/isabelle +ISABELLE_TEST=../deps/isabelle/Isabelle-test +ISABELLE=$(ISABELLE_TEST)/bin/isabelle runtest: + chmod -R ug+rw . $(ISABELLE_TEST)/lib/texinputs/ # Fighting the dune sandboxing... $(ISABELLE) build -o document=false -o browser_info=false -c -v -D . From 3eb56fb13c3e2f5f4e56a35ca78a632025545c7c Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 16:41:09 +0300 Subject: [PATCH 147/167] Trying to make CI run again. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 2c3fc665..a42b13b6 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time + sudo apt-get install --yes time texlive-latex-base - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 966afeea..14892550 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time + sudo apt-get install --yes time texlive-latex-base # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From 2cb14c268840daa51310b3347338c9c384023bf5 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 17:19:53 +0300 Subject: [PATCH 148/167] Fixing the CI. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index a42b13b6..5b49fc82 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base + sudo apt-get install --yes time texlive-latex-base texlive-plain-generic - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 14892550..95d111a7 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base + sudo apt-get install --yes time texlive-latex-base texlive-plain-generic # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From 85ea35a7128fce9d1eecf8fd5d162edecadedeef Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 17:35:34 +0300 Subject: [PATCH 149/167] Fixing the CI. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 5b49fc82..19934dd3 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-plain-generic + sudo apt-get install --yes time texlive-latex-base texlive-latex-extra - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 95d111a7..e0622205 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-plain-generic + sudo apt-get install --yes time texlive-latex-base texlive-latex-extra # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From 71215239ab526eace330a9d0b63f0afd52913257 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 17:56:45 +0300 Subject: [PATCH 150/167] Fixing the CI. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 19934dd3..07072ef3 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-latex-extra + sudo apt-get install --yes time texlive texlive-latex-extra - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index e0622205..602d2977 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive-latex-base texlive-latex-extra + sudo apt-get install --yes time texlive texlive-latex-extra # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From a7206271428736d7482caa9b52922c947eae4408 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 18:13:48 +0300 Subject: [PATCH 151/167] Fixing the CI. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 07072ef3..496bb96e 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra + sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 602d2977..e2e602af 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra + sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From 4f601e599dda775aaaeee1eb9cf3dcd391b03fc0 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 18:35:52 +0300 Subject: [PATCH 152/167] Fixing the CI. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 496bb96e..81db3cb3 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -25,7 +25,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra + sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index e2e602af..baa6d4ef 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -71,7 +71,7 @@ jobs: if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra + sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From c12628ab131c941d2bbe92877b3cf92da4673a00 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 19:05:07 +0300 Subject: [PATCH 153/167] Try to fix macos as well. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 6 +++++- .github/workflows/release.yml | 6 +++++- 2 files changed, 10 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 81db3cb3..1c0a9157 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -21,11 +21,15 @@ jobs: steps: - name: Clone repo uses: actions/checkout@v4 - - name: Install deps + - name: Install deps (ubuntu) if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex + - name: Install deps (macos) + if: startsWith(matrix.os, 'macos') + run: | + brew install texlive texlive-latex-extra texlive-fonts-extra texlive-luatex - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index baa6d4ef..ff662239 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -67,11 +67,15 @@ jobs: '4.13.1', ] steps: - - name: Install deps + - name: Install deps (ubuntu) if: matrix.operating-system == 'ubuntu-latest' run: | sudo apt-get update sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex + - name: Install deps (macos) + if: startsWith(matrix.operating-system, 'macos') + run: | + brew install texlive texlive-latex-extra texlive-fonts-extra texlive-luatex # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From 0c79fd71cc3163cce156f32a052e814d7b9b2596 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Sun, 18 Aug 2024 22:11:59 +0300 Subject: [PATCH 154/167] Try to fix macos as well. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 2 +- .github/workflows/release.yml | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 1c0a9157..aeb92d7d 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -29,7 +29,7 @@ jobs: - name: Install deps (macos) if: startsWith(matrix.os, 'macos') run: | - brew install texlive texlive-latex-extra texlive-fonts-extra texlive-luatex + brew install texlive - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index ff662239..07b4e1fe 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -75,7 +75,7 @@ jobs: - name: Install deps (macos) if: startsWith(matrix.operating-system, 'macos') run: | - brew install texlive texlive-latex-extra texlive-fonts-extra texlive-luatex + brew install texlive # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job From b2518186606446e47f17238d6c9574be25c07fb7 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Tue, 20 Aug 2024 16:41:31 +0300 Subject: [PATCH 155/167] Unify CI scripts. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 9 +++++++-- .github/workflows/release.yml | 17 ++++++++--------- 2 files changed, 15 insertions(+), 11 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index aeb92d7d..d73c8b19 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -11,8 +11,13 @@ jobs: strategy: fail-fast: false matrix: - os: [ubuntu-latest, macos-13, macos-14] - ocaml-compiler: ['4.13.1', '5.1.0'] + os: + - ubuntu-latest + - macos-13 + - macos-14 + ocaml-compiler: + - '4.13.1' + - '5.1.0' env: EXAMPLES_DIR: "tlaplus-examples" SCRIPT_DIR: "tlaplus-examples/.github/scripts" diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 07b4e1fe..9236798b 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -57,23 +57,22 @@ jobs: name: Create TLAPS installer, and upload it to GitHub release # fast testing *does* commence below needs: [release] - runs-on: ${{ matrix.operating-system }} + runs-on: ${{ matrix.os }} strategy: matrix: - operating-system: [ - macos-13, - ubuntu-latest] - ocaml-compiler: [ - '4.13.1', - ] + os: + - macos-13 + - ubuntu-latest + ocaml-compiler: + - '5.1.0' steps: - name: Install deps (ubuntu) - if: matrix.operating-system == 'ubuntu-latest' + if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex - name: Install deps (macos) - if: startsWith(matrix.operating-system, 'macos') + if: startsWith(matrix.os, 'macos') run: | brew install texlive # Read "Output Release URL and From c1758522355a65526875e9880b23a469da7f2005 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 00:24:58 +0300 Subject: [PATCH 156/167] Build/install instructions updated. Signed-off-by: Karolis Petrauskas --- INSTALL.md | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/INSTALL.md b/INSTALL.md index bbbcf666..6e5b0a29 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -33,11 +33,12 @@ Copyright (C) 2008-2010 INRIA and Microsoft Corporation Setup required OS packages; Debian/Ubuntu: ```{bash} -sudo apt install opam zlib1g-dev gawk time +sudo apt install opam zlib1g-dev gawk time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex ``` Arch Linux: ```{bash} -sudo pacman -S ocaml opam dune zlib time +sudo pacman -Sy git make gcc patch diffutils ocaml opam dune zlib wget fontconfig gnu-free-fonts # ocaml and tlapm build +sudo pacman -Sy time texlive texlive-latexextra texlive-fontsextra texlive-luatex # for tlapm tests ``` Initialize the OPAM. Add `--disable-sandboxing` option if running this on the docker or sandboxing is not supported for other reasons. @@ -105,7 +106,7 @@ Then view/edit the code e.g. using VSCode with the `OCaml Platform` extension in #### 1.1.5. Testing the build/install procedures -The above instructions were tested with +The above instructions were tested for Debian as follows: ```{bash} docker run -it --rm debian bash From b67ce91f49c23d2e831c59e73cb066f122292fd6 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 09:52:11 +0300 Subject: [PATCH 157/167] Install instructions for macOS, not tested. Signed-off-by: Karolis Petrauskas --- INSTALL.md | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/INSTALL.md b/INSTALL.md index 6e5b0a29..e17dc02c 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -27,9 +27,7 @@ Copyright (C) 2008-2010 INRIA and Microsoft Corporation ## 1. Installation -### 1.1. Linux - -#### 1.1.1. Setup the environment +### 1.1. Setup the environment Setup required OS packages; Debian/Ubuntu: ```{bash} @@ -41,6 +39,11 @@ sudo pacman -Sy git make gcc patch diffutils ocaml opam dune zlib wget fontconfi sudo pacman -Sy time texlive texlive-latexextra texlive-fontsextra texlive-luatex # for tlapm tests ``` +maxOS: +```{bash} +brew install texlive +``` + Initialize the OPAM. Add `--disable-sandboxing` option if running this on the docker or sandboxing is not supported for other reasons. ```{bash} @@ -54,7 +57,7 @@ opam switch create 5.1.0 eval $(opam env --switch=5.1.0) ``` -#### 1.1.2. Build and install TLAPM +### 1.2. Build and install TLAPM Clone the TLAPM source code @@ -87,13 +90,13 @@ Now you can invoke `tlapm` in either way: - `~/.opam/5.1.0/bin/tlapm --help`. -#### 1.1.3. Running the tests +### 1.3. Running the tests To run the test suite, invoke: ```{bash} make test ``` -#### 1.1.4. Development environment +### 1.4. Development environment To setup the development environment, run the following in addition to the above steps: @@ -104,7 +107,7 @@ make opam-deps-dev Then view/edit the code e.g. using VSCode with the `OCaml Platform` extension installed. -#### 1.1.5. Testing the build/install procedures +## Appendix: Testing the build/install procedures The above instructions were tested for Debian as follows: From 31cfa09ae8d63aa14d84ba0214808b87b72c4adf Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 09:54:59 +0300 Subject: [PATCH 158/167] Mistype. Signed-off-by: Karolis Petrauskas --- INSTALL.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/INSTALL.md b/INSTALL.md index e17dc02c..23c6c2ea 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -39,7 +39,7 @@ sudo pacman -Sy git make gcc patch diffutils ocaml opam dune zlib wget fontconfi sudo pacman -Sy time texlive texlive-latexextra texlive-fontsextra texlive-luatex # for tlapm tests ``` -maxOS: +macOS: ```{bash} brew install texlive ``` From f790a5f1b1a42475c7e8ba36463d4a0c23162c42 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 23:33:58 +0300 Subject: [PATCH 159/167] Don't force Isabelle document build in the ROOT file. Signed-off-by: Karolis Petrauskas --- isabelle/ROOT | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/isabelle/ROOT b/isabelle/ROOT index dc2d6c1b..d2f3e118 100644 --- a/isabelle/ROOT +++ b/isabelle/ROOT @@ -1,4 +1,4 @@ -(* build the session using "isabelle build -b -D ." *) +(* See also Makefile. *) chapter "TLA+" session "TLA+" = "Pure" + options [document_output = "output", document_variants="document:outline=/proof"] @@ -17,7 +17,7 @@ session "TLA+Tests" in tests = "TLA+" + Tests session "TLA+Examples" in examples = "TLA+" + - options [document=pdf, document_output = "output", document_variants="document:outline=/proof"] + options [document_output = "output", document_variants="document:outline=/proof"] theories Allocator From 1154124c6febd4aaeb3ce269025d65359520fec0 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 23:39:09 +0300 Subject: [PATCH 160/167] Don't depend on latex during the build. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 6 +----- .github/workflows/release.yml | 6 +----- INSTALL.md | 7 +++---- isabelle/dune.mk | 1 - 4 files changed, 5 insertions(+), 15 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index d73c8b19..fa6632fb 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -30,11 +30,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex - - name: Install deps (macos) - if: startsWith(matrix.os, 'macos') - run: | - brew install texlive + sudo apt-get install --yes time - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 9236798b..3d8a9a3f 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -70,11 +70,7 @@ jobs: if: matrix.os == 'ubuntu-latest' run: | sudo apt-get update - sudo apt-get install --yes time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex - - name: Install deps (macos) - if: startsWith(matrix.os, 'macos') - run: | - brew install texlive + sudo apt-get install --yes time # Read "Output Release URL and # ID File" of `release` job above. - name: Load Release URL File from release job diff --git a/INSTALL.md b/INSTALL.md index 23c6c2ea..db46d8cc 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -31,17 +31,16 @@ Copyright (C) 2008-2010 INRIA and Microsoft Corporation Setup required OS packages; Debian/Ubuntu: ```{bash} -sudo apt install opam zlib1g-dev gawk time texlive texlive-latex-extra texlive-fonts-extra texlive-luatex +sudo apt install opam zlib1g-dev gawk time ``` Arch Linux: ```{bash} -sudo pacman -Sy git make gcc patch diffutils ocaml opam dune zlib wget fontconfig gnu-free-fonts # ocaml and tlapm build -sudo pacman -Sy time texlive texlive-latexextra texlive-fontsextra texlive-luatex # for tlapm tests +sudo pacman -Sy time git make gcc patch diffutils ocaml opam dune zlib wget fontconfig gnu-free-fonts ``` macOS: ```{bash} -brew install texlive +# No additional packages needed. ``` Initialize the OPAM. Add `--disable-sandboxing` option if running this on the docker or sandboxing is not supported for other reasons. diff --git a/isabelle/dune.mk b/isabelle/dune.mk index a65a31dc..29f7518c 100644 --- a/isabelle/dune.mk +++ b/isabelle/dune.mk @@ -8,5 +8,4 @@ ISABELLE_TEST=../deps/isabelle/Isabelle-test ISABELLE=$(ISABELLE_TEST)/bin/isabelle runtest: - chmod -R ug+rw . $(ISABELLE_TEST)/lib/texinputs/ # Fighting the dune sandboxing... $(ISABELLE) build -o document=false -o browser_info=false -c -v -D . From aaf5a15d0e4b0a6974b63c51addac0b201f9f302 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Wed, 21 Aug 2024 23:52:03 +0300 Subject: [PATCH 161/167] Try to make `wget` less quiet. Signed-off-by: Karolis Petrauskas --- INSTALL.md | 1 + Makefile | 3 +++ deps/isabelle/dune.mk | 2 +- deps/ls4/Makefile | 2 +- deps/z3/Makefile | 2 +- 5 files changed, 7 insertions(+), 3 deletions(-) diff --git a/INSTALL.md b/INSTALL.md index db46d8cc..81a0836b 100644 --- a/INSTALL.md +++ b/INSTALL.md @@ -68,6 +68,7 @@ cd tlapm Install the dependencies via OPAM. A helper `make` target is present for that: ```{bash} +make opam-update make opam-deps ``` diff --git a/Makefile b/Makefile index 402f5f26..994192ff 100644 --- a/Makefile +++ b/Makefile @@ -17,6 +17,9 @@ PREFIX=$(OPAM_SWITCH_PREFIX) all: build +opam-update: # Update the package lists. + opam update + opam-deps: opam install ./ --deps-only --yes --working-dir diff --git a/deps/isabelle/dune.mk b/deps/isabelle/dune.mk index 1654ec89..36f03a01 100644 --- a/deps/isabelle/dune.mk +++ b/deps/isabelle/dune.mk @@ -37,7 +37,7 @@ all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_TEST) Isabelle.exec-fil # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): mkdir -p $(CACHE_DIR) - cd $(CACHE_DIR) && wget -q $(ISABELLE_URL) + cd $(CACHE_DIR) && wget -nv $(ISABELLE_URL) # Take the Isabelle archive from the cache. $(ISABELLE_ARCHIVE): $(CACHE_DIR)/$(ISABELLE_ARCHIVE) diff --git a/deps/ls4/Makefile b/deps/ls4/Makefile index ddfe1324..a63eb93b 100644 --- a/deps/ls4/Makefile +++ b/deps/ls4/Makefile @@ -17,7 +17,7 @@ endif all: ls4 v$(LS4_VERS).zip: - wget -q $(LS4_URL_PREFIX)/$@ + wget -nv $(LS4_URL_PREFIX)/$@ $(LS4_BASE_NAME): v$(LS4_VERS).zip rm -rf $@ diff --git a/deps/z3/Makefile b/deps/z3/Makefile index 04005463..e840e257 100644 --- a/deps/z3/Makefile +++ b/deps/z3/Makefile @@ -19,7 +19,7 @@ endif all: z3 $(Z3_BASE_NAME).zip: - wget -q $(Z3_URL_PREFIX)/$@ + wget -nv $(Z3_URL_PREFIX)/$@ $(Z3_BASE_NAME): $(Z3_BASE_NAME).zip rm -rf $@ From 7be12690f934e05ce5fd967fb6d8dd99505aa1a6 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Thu, 22 Aug 2024 00:03:46 +0300 Subject: [PATCH 162/167] Tweaking wget options. Signed-off-by: Karolis Petrauskas --- deps/isabelle/dune.mk | 2 +- deps/ls4/Makefile | 2 +- deps/z3/Makefile | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/deps/isabelle/dune.mk b/deps/isabelle/dune.mk index 36f03a01..38567072 100644 --- a/deps/isabelle/dune.mk +++ b/deps/isabelle/dune.mk @@ -37,7 +37,7 @@ all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_TEST) Isabelle.exec-fil # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): mkdir -p $(CACHE_DIR) - cd $(CACHE_DIR) && wget -nv $(ISABELLE_URL) + cd $(CACHE_DIR) && wget --progress=dot:giga $(ISABELLE_URL) # Take the Isabelle archive from the cache. $(ISABELLE_ARCHIVE): $(CACHE_DIR)/$(ISABELLE_ARCHIVE) diff --git a/deps/ls4/Makefile b/deps/ls4/Makefile index a63eb93b..6c409354 100644 --- a/deps/ls4/Makefile +++ b/deps/ls4/Makefile @@ -17,7 +17,7 @@ endif all: ls4 v$(LS4_VERS).zip: - wget -nv $(LS4_URL_PREFIX)/$@ + wget --progress=dot:giga $(LS4_URL_PREFIX)/$@ $(LS4_BASE_NAME): v$(LS4_VERS).zip rm -rf $@ diff --git a/deps/z3/Makefile b/deps/z3/Makefile index e840e257..fe68ce14 100644 --- a/deps/z3/Makefile +++ b/deps/z3/Makefile @@ -19,7 +19,7 @@ endif all: z3 $(Z3_BASE_NAME).zip: - wget -nv $(Z3_URL_PREFIX)/$@ + wget --progress=dot:giga $(Z3_URL_PREFIX)/$@ $(Z3_BASE_NAME): $(Z3_BASE_NAME).zip rm -rf $@ From 2c0566b54707c6eaff296dd206d2e1220958e429 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Thu, 22 Aug 2024 00:58:05 +0300 Subject: [PATCH 163/167] Try to use a cache for Isabelle downloads, doublecheck the checksums. Signed-off-by: Karolis Petrauskas --- .github/workflows/ci.yml | 10 +++++----- .github/workflows/release.yml | 10 +++++----- deps/isabelle/dune.mk | 13 ++++++++++++- 3 files changed, 22 insertions(+), 11 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index fa6632fb..2d10f3c6 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -34,11 +34,11 @@ jobs: - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} - # - uses: actions/cache@v3 - # id: cache - # with: - # path: _build_cache - # key: ${{ runner.os }}_build_cache + - uses: actions/cache@v4 + id: cache + with: + path: _build_cache + key: ${{ runner.os }}_build_cache - name: Build TLAPM run: | eval $(opam env) diff --git a/.github/workflows/release.yml b/.github/workflows/release.yml index 3d8a9a3f..51e19fe4 100644 --- a/.github/workflows/release.yml +++ b/.github/workflows/release.yml @@ -92,11 +92,11 @@ jobs: - uses: ocaml/setup-ocaml@v2 with: ocaml-compiler: ${{ matrix.ocaml-compiler }} - # - uses: actions/cache@v3 - # id: cache - # with: - # path: _build_cache - # key: ${{ runner.os }}_build_cache + - uses: actions/cache@v4 + id: cache + with: + path: _build_cache + key: ${{ runner.os }}_build_cache - name: Build installer of TLAPS run: | eval $(opam env) diff --git a/deps/isabelle/dune.mk b/deps/isabelle/dune.mk index 38567072..520cb83c 100644 --- a/deps/isabelle/dune.mk +++ b/deps/isabelle/dune.mk @@ -1,15 +1,20 @@ +# +# See https://isabelle.in.tum.de/dist/ for download files and their SHA256 sums. +# OS_TYPE=$(patsubst CYGWIN%,Cygwin,$(shell uname)) HOST_CPU=$(shell uname -m) ISABELLE_VSN=Isabelle2024 ifeq ($(OS_TYPE),Linux) + ISABELLE_SHA256=603aaaf8abea36597af3b0651d2c162a86c0a0dd4420766f47e5724039639267 ISABELLE_ARCHIVE=$(ISABELLE_VSN)_linux.tar.gz ISABELLE_ARCHIVE_TYPE=tgz ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN) FIND_EXEC=-executable endif ifeq ($(OS_TYPE),Darwin) + ISABELLE_SHA256=22035f996f71ea1f03063f6f144195eb6a04974d4d916ed0772cd79569a28bc7 ISABELLE_ARCHIVE=$(ISABELLE_VSN)_macos.tar.gz ISABELLE_ARCHIVE_TYPE=tgz ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN).app @@ -17,6 +22,7 @@ ifeq ($(OS_TYPE),Darwin) endif ifeq ($(OS_TYPE),Cygwin) # TODO: Fix this. + ISABELLE_SHA256=ab449a85c0f7c483027c2000889ec93b3f7df565d9d0c6902af2d666b3b58220 ISABELLE_ARCHIVE=$(ISABELLE_VSN)_bundle_x86-cygwin.tar.gz ISABELLE_ARCHIVE_TYPE=tgz ISABELLE_ARCHIVE_DIR=$(ISABELLE_VSN) @@ -37,7 +43,12 @@ all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_TEST) Isabelle.exec-fil # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): mkdir -p $(CACHE_DIR) - cd $(CACHE_DIR) && wget --progress=dot:giga $(ISABELLE_URL) + (echo "$(ISABELLE_SHA256) $@" | sha256sum -c -) || ( \ + (rm -f $@) && \ + (wget --progress=dot:giga --directory-prefix=$(CACHE_DIR) $(ISABELLE_URL)) && \ + (echo "$(ISABELLE_SHA256) $@" | sha256sum -c -) \ + ) +.PHONY: $(CACHE_DIR)/$(ISABELLE_ARCHIVE) # Have to double-check the checksum. # Take the Isabelle archive from the cache. $(ISABELLE_ARCHIVE): $(CACHE_DIR)/$(ISABELLE_ARCHIVE) From 72ce8effc3266184c4c9fa51844ee1ab5e412c0e Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Thu, 22 Aug 2024 01:16:17 +0300 Subject: [PATCH 164/167] try make shasum working on macos. Signed-off-by: Karolis Petrauskas --- deps/isabelle/dune.mk | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/deps/isabelle/dune.mk b/deps/isabelle/dune.mk index 520cb83c..c84eaf87 100644 --- a/deps/isabelle/dune.mk +++ b/deps/isabelle/dune.mk @@ -43,10 +43,10 @@ all: $(ISABELLE_DIR) $(ISABELLE_DIR)/src/TLA+ $(ISABELLE_TEST) Isabelle.exec-fil # Download the isabelle archive to the cache. $(CACHE_DIR)/$(ISABELLE_ARCHIVE): mkdir -p $(CACHE_DIR) - (echo "$(ISABELLE_SHA256) $@" | sha256sum -c -) || ( \ + (echo "$(ISABELLE_SHA256) *$@" | shasum -a 256 -c -) || ( \ (rm -f $@) && \ (wget --progress=dot:giga --directory-prefix=$(CACHE_DIR) $(ISABELLE_URL)) && \ - (echo "$(ISABELLE_SHA256) $@" | sha256sum -c -) \ + (echo "$(ISABELLE_SHA256) *$@" | shasum -a 256 -c -) \ ) .PHONY: $(CACHE_DIR)/$(ISABELLE_ARCHIVE) # Have to double-check the checksum. From d59401191f88e7bae458de694d04865aac183d82 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Thu, 22 Aug 2024 22:06:42 +0300 Subject: [PATCH 165/167] Try make build more portable. Signed-off-by: Karolis Petrauskas --- deps/isabelle/dune.mk | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/deps/isabelle/dune.mk b/deps/isabelle/dune.mk index c84eaf87..56759580 100644 --- a/deps/isabelle/dune.mk +++ b/deps/isabelle/dune.mk @@ -80,11 +80,10 @@ endif $(ISABELLE_DIR)/src/TLA+: $(ISABELLE_DIR) cd $(ISABELLE_DIR) \ && rm -rf contrib/ProofGeneral* doc heaps/*/HOL contrib/vscodium* contrib/vscode* \ - && sed -i -e 's@^\(contrib/vscode_extension\|contrib/vscodium\|src/Tools/Demo\)@#rm at TLA# \1@' etc/components \ - && echo 'src/TLA+' >> etc/components + && awk '/^((contrib\/(vscode_extension|vscodium))|(src\/Tools\/Demo))/{ print "#rm at TLA# " $$0; next } END { print "src/TLA+" } { print }' etc/components > etc/components.tmp \ + && rm etc/components && mv etc/components.tmp etc/components cd $(ISABELLE_DIR) \ && HEAPS_PATH=$(shell pwd)/$(ISABELLE_DIR)/heaps \ - && if [ "$(HOST_CPU)" = "x86_64" ] ; then sed -i -e 's/^ML_PLATFORM=.*$$/ML_PLATFORM="$${ISABELLE_PLATFORM64:-$$ISABELLE_PLATFORM}"/' etc/settings ; fi \ && cp etc/settings etc/settings.target \ && echo "ISABELLE_OUTPUT=$$HEAPS_PATH" >> etc/settings mkdir -p $(ISABELLE_DIR)/src/TLA+ \ From 07926530c70dc6c40fd752c2f79333b7e0e328cc Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Thu, 22 Aug 2024 22:07:14 +0300 Subject: [PATCH 166/167] Add upgrade to the build instructions. Signed-off-by: Karolis Petrauskas --- Makefile | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Makefile b/Makefile index 994192ff..e087cc7b 100644 --- a/Makefile +++ b/Makefile @@ -17,8 +17,9 @@ PREFIX=$(OPAM_SWITCH_PREFIX) all: build -opam-update: # Update the package lists. +opam-update: # Update the package lists and install updates. opam update + opam upgrade opam-deps: opam install ./ --deps-only --yes --working-dir From 9e75ea943c0064b30af4eeffee8e3c5c8553d295 Mon Sep 17 00:00:00 2001 From: Karolis Petrauskas Date: Fri, 23 Aug 2024 00:13:41 +0300 Subject: [PATCH 167/167] Remove compile outputs from gitignore. They are all in `_build/` now. Signed-off-by: Karolis Petrauskas --- .gitignore | 4 ---- 1 file changed, 4 deletions(-) diff --git a/.gitignore b/.gitignore index 52e72259..ab6fe08c 100644 --- a/.gitignore +++ b/.gitignore @@ -1,7 +1,3 @@ -*.o -*.cmi -*.cmx -*.annot *.exe *.err *.out