diff --git a/thys/Approximation_Algorithms/Approx_BP_Hoare.thy b/thys/Approximation_Algorithms/Approx_BP_Hoare.thy --- a/thys/Approximation_Algorithms/Approx_BP_Hoare.thy +++ b/thys/Approximation_Algorithms/Approx_BP_Hoare.thy @@ -1,1706 +1,1708 @@ section \Bin Packing\ theory Approx_BP_Hoare imports Complex_Main "HOL-Hoare.Hoare_Logic" "HOL-Library.Disjoint_Sets" begin text \The algorithm and proofs are based on the work by Berghammer and Reuter \<^cite>\BerghammerR03\.\ subsection \Formalization of a Correct Bin Packing\ text \Definition of the unary operator \\\\\ from the article. \B\ will only be wrapped into a set if it is non-empty.\ definition wrap :: "'a set \ 'a set set" where "wrap B = (if B = {} then {} else {B})" lemma wrap_card: "card (wrap B) \ 1" unfolding wrap_def by auto text \If \M\ and \N\ are pairwise disjoint with \V\ and not yet contained in V, then the union of \M\ and \N\ is also pairwise disjoint with \V\.\ lemma pairwise_disjnt_Un: assumes "pairwise disjnt ({M} \ {N} \ V)" "M \ V" "N \ V" shows "pairwise disjnt ({M \ N} \ V)" using assms unfolding pairwise_def by auto text \A Bin Packing Problem is defined like in the article:\ locale BinPacking = fixes U :: "'a set" \ \A finite, non-empty set of objects\ and w :: "'a \ real" \ \A mapping from objects to their respective weights (positive real numbers)\ and c :: nat \ \The maximum capacity of a bin (a natural number)\ and S :: "'a set" \ \The set of \small\ objects (weight no larger than \1/2\ of \c\)\ and L :: "'a set" \ \The set of \large\ objects (weight larger than \1/2\ of \c\)\ assumes weight: "\u \ U. 0 < w(u) \ w(u) \ c" and U_Finite: "finite U" and U_NE: "U \ {}" and S_def: "S = {u \ U. w(u) \ c / 2}" and L_def: "L = U - S" begin text \In the article, this is defined as \w\ as well. However, to avoid ambiguity, we will abbreviate the weight of a bin as \W\.\ abbreviation W :: "'a set \ real" where "W B \ (\u \ B. w(u))" text \\P\ constitutes as a correct bin packing if \P\ is a partition of \U\ (as defined in @{thm [source] partition_on_def}) and the weights of the bins do not exceed their maximum capacity \c\.\ definition bp :: "'a set set \ bool" where "bp P \ partition_on U P \ (\B \ P. W(B) \ c)" lemma bpE: assumes "bp P" shows "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" using assms unfolding bp_def partition_on_def by blast+ lemma bpI: assumes "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" shows "bp P" using assms unfolding bp_def partition_on_def by blast text \Although we assume the \S\ and \L\ sets as given, manually obtaining them from \U\ is trivial and can be achieved in linear time. Proposed by the article \<^cite>\"BerghammerR03"\.\ lemma S_L_set_generation: "VARS S L W u {True} S := {}; L := {}; W := U; WHILE W \ {} INV {W \ U \ S = {v \ U - W. w(v) \ c / 2} \ L = {v \ U - W. w(v) > c / 2}} DO u := (SOME u. u \ W); IF 2 * w(u) \ c THEN S := S \ {u} ELSE L := L \ {u} FI; W := W - {u} OD {S = {v \ U. w(v) \ c / 2} \ L = {v \ U. w(v) > c / 2}}" by vcg (auto simp: some_in_eq) subsection \The Proposed Approximation Algorithm\ subsubsection \Functional Correctness\ text \According to the article, \inv\<^sub>1\ holds if \P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}\ is a correct solution for the bin packing problem \<^cite>\BerghammerR03\. However, various assumptions made in the article seem to suggest that more information is demanded from this invariant and, indeed, mere correctness (as defined in @{thm [source] bp_def}) does not appear to suffice. To amend this, four additional conjuncts have been added to this invariant, whose necessity will be explained in the following proofs. It should be noted that there may be other (shorter) ways to amend this invariant. This approach, however, makes for rather straight-forward proofs, as these conjuncts can be utilized and proved in relatively few steps.\ definition inv\<^sub>1 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ \A correct solution to the bin packing problem\ \ \(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V \ \The partial solution does not contain objects that have not yet been assigned\ \ B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ \\B\<^sub>1\ is distinct from all the other bins\ \ B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2) \ \\B\<^sub>2\ is distinct from all the other bins\ \ (P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {} \ \The first and second partial solutions are disjoint from each other.\" (* lemma "partition_on U (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ u \ V \ partition_on U (P\<^sub>1 \ wrap (insert u B\<^sub>1) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ (V-{u})})" nitpick*) lemma inv\<^sub>1E: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using assms unfolding inv\<^sub>1_def by auto lemma inv\<^sub>1I: assumes "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>1_def by blast lemma wrap_Un [simp]: "wrap (M \ {x}) = {M \ {x}}" unfolding wrap_def by simp lemma wrap_empty [simp]: "wrap {} = {}" unfolding wrap_def by simp lemma wrap_not_empty [simp]: "M \ {} \ wrap M = {M}" unfolding wrap_def by simp text \If \inv\<^sub>1\ holds for the current partial solution, and the weight of an object \u \ V\ added to \B\<^sub>1\ does not exceed its capacity, then \inv\<^sub>1\ also holds if \B\<^sub>1\ and \{u}\ are replaced by \B\<^sub>1 \ {u}\.\ lemma inv\<^sub>1_stepA: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W(B\<^sub>1) + w(u) \ c" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] text \In the proof for \Theorem 3.2\ of the article it is erroneously argued that if \P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}\ is a partition of \U\, then the same holds if \B\<^sub>1\ is replaced by \B\<^sub>1 \ {u}\. This is, however, not necessarily the case if \B\<^sub>1\ or \{u}\ are already contained in the partial solution. Suppose \P\<^sub>1\ contains the non-empty bin \B\<^sub>1\, then \P\<^sub>1 \ wrap B\<^sub>1\ would still be pairwise disjoint, provided \P\<^sub>1\ was pairwise disjoint before, as the union simply ignores the duplicate \B\<^sub>1\. Now, if the algorithm modifies \B\<^sub>1\ by adding an element from \V\ such that \B\<^sub>1\ becomes some non-empty \B\<^sub>1'\ with \B\<^sub>1 \ B\<^sub>1' \ \\ and \B\<^sub>1' \ P\<^sub>1\, one can see that this property would no longer be preserved. To avoid such a situation, we will use the first additional conjunct in \inv\<^sub>1\ to ensure that \{u}\ is not yet contained in the partial solution, and the second additional conjunct to ensure that \B\<^sub>1\ is not yet contained in the partial solution.\ \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ ({{u}} \ {{v} |v. v \ V - {u}}))" using bprules(1) assms(2) by simp then have "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" by (simp add: Un_commute) then have assm: "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc) have "pairwise disjnt ({B\<^sub>1 \ {u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>1 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>1} \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,3) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>1" using bprules(3) U_Finite by (cases \B\<^sub>1 = {}\) auto then have "W (B\<^sub>1 \ {u}) \ W B\<^sub>1 + w u" using \0 < w u\ by (cases \u \ B\<^sub>1\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>1 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>1 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" by blast from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = U - (V - {u})" unfolding L R invrules(2) .. have 3: "B\<^sub>1 \ {u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2" using NOTIN by auto have 4: "B\<^sub>2 \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2" using invrules(4) NOTIN unfolding wrap_def by fastforce have 5: "(P\<^sub>1 \ wrap (B\<^sub>1 \ {u})) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) NOTIN unfolding wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \If \inv\<^sub>1\ holds for the current partial solution, and the weight of an object \u \ V\ added to \B\<^sub>2\ does not exceed its capacity, then \inv\<^sub>1\ also holds if \B\<^sub>2\ and \{u}\ are replaced by \B\<^sub>2 \ {u}\.\ lemma inv\<^sub>1_stepB: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W B\<^sub>2 + w u \ c" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] text \The argumentation here is similar to the one in @{thm [source] inv\<^sub>1_stepA} with \B\<^sub>1\ replaced with \B\<^sub>2\ and using the first and third additional conjuncts of \inv\<^sub>1\ to amend the issue, instead of the first and second.\ \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}})" using bprules(1) assms(2) by simp then have assm: "pairwise disjnt (wrap B\<^sub>2 \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc Un_commute) have "pairwise disjnt ({B\<^sub>2 \ {u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>2 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>2} \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,4) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>2" using bprules(3) U_Finite by (cases \B\<^sub>2 = {}\) auto then have "W (B\<^sub>2 \ {u}) \ W B\<^sub>2 + w u" using \0 < w u\ by (cases \u \ B\<^sub>2\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>2 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>2 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}. W B \ c" by auto from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = U - (V - {u})" unfolding L R using invrules(2) by simp have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})" using bpE(2)[OF 1] by simp have 4: "B\<^sub>2 \ {u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = {}" using invrules(5) NOTIN unfolding wrap_empty wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \If \inv\<^sub>1\ holds for the current partial solution, then \inv\<^sub>1\ also holds if \B\<^sub>1\ and \B\<^sub>2\ are added to \P\<^sub>1\ and \P\<^sub>2\ respectively, \B\<^sub>1\ is emptied and \B\<^sub>2\ initialized with \u \ V\.\ lemma inv\<^sub>1_stepC: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] \ \Rule 1-4: Correct Bin Packing\ have "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}" by (metis (no_types, lifting) Un_assoc Un_empty_right insert_not_empty wrap_empty wrap_not_empty) also have "... = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" using assms(2) by auto finally have EQ: "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" . from invrules(1) have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}})" unfolding EQ . \ \Auxiliary information is preserved\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "u \ U" using assms(2) bpE(3)[OF invrules(1)] by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = U - (V - {u})" unfolding L R using invrules(2) by auto have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}" using bpE(2)[OF 1] by simp have 4: "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}) = {}" using invrules(5) NOTIN unfolding wrap_def by force from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \A simplified version of the bin packing algorithm proposed in the article. It serves as an introduction into the approach taken, and, while it does not provide the desired approximation factor, it does ensure that \P\ is a correct solution of the bin packing problem.\ lemma simple_bp_correct: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u {True} P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; V := U; WHILE V \ S \ {} INV {inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V} DO u := (SOME u. u \ V); V := V - {u}; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} | v. v \ V} {bp P}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) show ?case unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" .. from 2 have "V \ {}" by blast then have IN: "(SOME u. u \ V) \ V" by (simp add: some_in_eq) from inv\<^sub>1_stepA[OF INV IN] inv\<^sub>1_stepB[OF INV IN] inv\<^sub>1_stepC[OF INV IN] show ?case by blast next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then show ?case unfolding inv\<^sub>1_def by blast qed subsubsection \Lower Bounds for the Bin Packing Problem\ lemma bp_bins_finite [simp]: assumes "bp P" shows "\B \ P. finite B" using bpE(3)[OF assms] U_Finite by (meson Sup_upper finite_subset) lemma bp_sol_finite [simp]: assumes "bp P" shows "finite P" using bpE(3)[OF assms] U_Finite by (simp add: finite_UnionD) text \If \P\ is a solution of the bin packing problem, then no bin in \P\ may contain more than one large object.\ lemma only_one_L_per_bin: assumes "bp P" "B \ P" shows "\x \ B. \y \ B. x \ y \ x \ L \ y \ L" proof (rule ccontr, simp) assume "\x\B. \y\B. x \ y \ x \ L \ y \ L" then obtain x y where *: "x \ B" "y \ B" "x \ y" "x \ L" "y \ L" by blast then have "c < w x + w y" using L_def S_def by force have "finite B" using assms by simp have "y \ B - {x}" using *(2,3) by blast have "W B = W (B - {x}) + w x" using *(1) \finite B\ by (simp add: sum.remove) also have "... = W (B - {x} - {y}) + w x + w y" using \y \ B - {x}\ \finite B\ by (simp add: sum.remove) finally have *: "W B = W (B - {x} - {y}) + w x + w y" . have "\u \ B. 0 < w u" using bpE(3)[OF assms(1)] assms(2) weight by blast then have "0 \ W (B - {x} - {y})" by (smt DiffD1 sum_nonneg) with * have "c < W B" using \c < w x + w y\ by simp then show False using bpE(4)[OF assms(1)] assms(2) by fastforce qed text \If \P\ is a solution of the bin packing problem, then the amount of large objects is a lower bound for the amount of bins in P.\ lemma L_lower_bound_card: assumes "bp P" shows "card L \ card P" proof - have "\x \ L. \B \ P. x \ B" using bpE(3)[OF assms] L_def by blast then obtain f where f_def: "\u \ L. u \ f u \ f u \ P" by metis then have "inj_on f L" unfolding inj_on_def using only_one_L_per_bin[OF assms] by blast then have card_eq: "card L = card (f ` L)" by (simp add: card_image) have "f ` L \ P" using f_def by blast moreover have "finite P" using assms by simp ultimately have "card (f ` L) \ card P" by (simp add: card_mono) then show ?thesis unfolding card_eq . qed text \If \P\ is a solution of the bin packing problem, then the amount of bins of a subset of P in which every bin contains a large object is a lower bound on the amount of large objects.\ lemma subset_bp_card: assumes "bp P" "M \ P" "\B \ M. B \ L \ {}" shows "card M \ card L" proof - have "\B \ M. \u \ L. u \ B" using assms(3) by fast then have "\f. \B \ M. f B \ L \ f B \ B" by metis then obtain f where f_def: "\B \ M. f B \ L \ f B \ B" .. have "inj_on f M" proof (rule ccontr) assume "\ inj_on f M" then have "\x \ M. \y \ M. x \ y \ f x = f y" unfolding inj_on_def by blast then obtain x y where *: "x \ M" "y \ M" "x \ y" "f x = f y" by blast then have "\u. u \ x \ u \ y" using f_def by metis then have "x \ y \ {}" by blast moreover have "pairwise disjnt M" using pairwise_subset[OF bpE(1)[OF assms(1)] assms(2)] . ultimately show False using * unfolding pairwise_def disjnt_def by simp qed moreover have "finite L" using L_def U_Finite by blast moreover have "f ` M \ L" using f_def by blast ultimately show ?thesis using card_inj_on_le by blast qed text \If \P\ is a correct solution of the bin packing problem, \inv\<^sub>1\ holds for the partial solution, and every bin in \P\<^sub>1 \ wrap B\<^sub>1\ contains a large object, then the amount of bins in \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}\ is a lower bound for the amount of bins in \P\.\ lemma L_bins_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}) \ card P" proof - note invrules = inv\<^sub>1E[OF assms(2)] have "\B \ {{v} |v. v \ V \ L}. B \ L \ {}" by blast with assms(3) have "P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" "\B \ P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}. B \ L \ {}" by blast+ from subset_bp_card[OF invrules(1) this] show ?thesis using L_lower_bound_card[OF assms(1)] by linarith qed text \If \P\ is a correct solution of the bin packing problem, then the sum of the weights of the objects is equal to the sum of the weights of the bins in \P\.\ lemma sum_Un_eq_sum_sum: assumes "bp P" shows "(\u \ U. w u) = (\B \ P. W B)" proof - have FINITE: "\B \ P. finite B" using assms by simp have DISJNT: "\A \ P. \B \ P. A \ B \ A \ B = {}" using bpE(1)[OF assms] unfolding pairwise_def disjnt_def . have "(\u \ (\P). w u) = (\B \ P. W B)" using sum.Union_disjoint[OF FINITE DISJNT] by auto then show ?thesis unfolding bpE(3)[OF assms] . qed text \If \P\ is a correct solution of the bin packing problem, then the sum of the weights of the items is a lower bound of amount of bins in \P\ multiplied by their maximum capacity.\ lemma sum_lower_bound_card: assumes "bp P" shows "(\u \ U. w u) \ c * card P" proof - have *: "\B \ P. 0 < W B \ W B \ c" using bpE(2-4)[OF assms] weight by (metis UnionI assms bp_bins_finite sum_pos) have "(\u \ U. w u) = (\B \ P. W B)" using sum_Un_eq_sum_sum[OF assms] . also have "... \ (\B \ P. c)" using sum_mono * by fastforce also have "... = c * card P" by simp finally show ?thesis . qed lemma bp_NE: assumes "bp P" shows "P \ {}" using U_NE bpE(3)[OF assms] by blast lemma sum_Un_ge: fixes f :: "_ \ real" assumes "finite M" "finite N" "\B \ M \ N. 0 < f B" shows "sum f M \ sum f (M \ N)" proof - have "0 \ sum f N - sum f (M \ N)" using assms by (smt DiffD1 inf.cobounded2 UnCI sum_mono2) then have "sum f M \ sum f M + sum f N - sum f (M \ N)" by simp also have "... = sum f (M \ N)" using sum_Un[OF assms(1,2), symmetric] . finally show ?thesis . qed text \If \bij_exists\ holds, one can obtain a function which is bijective between the bins in \P\ and the objects in \V\ such that an object returned by the function would cause the bin to exceed its capacity.\ definition bij_exists :: "'a set set \ 'a set \ bool" where "bij_exists P V = (\f. bij_betw f P V \ (\B \ P. W B + w (f B) > c))" text \If \P\ is a functionally correct solution of the bin packing problem, \inv\<^sub>1\ holds for the partial solution, and such a bijective function exists between the bins in \P\<^sub>1\ and the objects in @{term "P\<^sub>2 \ wrap B\<^sub>2"}, the following strict lower bound can be shown:\ lemma P\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card P\<^sub>1 + 1 \ card P" proof (cases \P\<^sub>1 = {}\) case True have "finite P" using assms(1) by simp then have "1 \ card P" using bp_NE[OF assms(1)] by (metis Nat.add_0_right Suc_diff_1 Suc_le_mono card_gt_0_iff le0 mult_Suc_right nat_mult_1) then show ?thesis unfolding True by simp next note invrules = inv\<^sub>1E[OF assms(2)] case False obtain f where f_def: "bij_betw f P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" "\B \ P\<^sub>1. W B + w (f B) > c" using assms(3) unfolding bij_exists_def by blast have FINITE: "finite P\<^sub>1" "finite (P\<^sub>2 \ wrap B\<^sub>2)" "finite (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" "finite (wrap B\<^sub>1 \ {{v} |v. v \ V})" using inv\<^sub>1E(1)[OF assms(2)] bp_sol_finite by blast+ have F: "\B \ P\<^sub>2 \ wrap B\<^sub>2. finite B" using invrules(1) by simp have D: "\A \ P\<^sub>2 \ wrap B\<^sub>2. \B \ P\<^sub>2 \ wrap B\<^sub>2. A \ B \ A \ B = {}" using bpE(1)[OF invrules(1)] unfolding pairwise_def disjnt_def by auto have sum_eq: "W (\ (P\<^sub>2 \ wrap B\<^sub>2)) = (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.Union_disjoint[OF F D] by auto have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}. 0 < W B" using bpE(2,3)[OF invrules(1)] weight by (metis (no_types, lifting) UnionI bp_bins_finite invrules(1) sum_pos) then have "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ (wrap B\<^sub>1 \ {{v} |v. v \ V}). W B)" using sum_Un_ge[OF FINITE(3,4), of W] by blast also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}. W B)" by (smt Un_assoc Un_commute) also have "... = W U" using sum_Un_eq_sum_sum[OF invrules(1), symmetric] . finally have *: "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ W U" . \ \This follows from the fourth and final additional conjunct of \inv\<^sub>1\ and is necessary to combine the sums of the bins of the two partial solutions. This does not inherently follow from the union being a correct solution, as this need not be the case if \P\<^sub>1\ and \P\<^sub>2 \ wrap B\<^sub>2\ happened to be equal.\ have DISJNT: "P\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) by blast \ \This part of the proof is based on the proof on page 72 of the article \<^cite>\BerghammerR03\.\ have "c * card P\<^sub>1 = (\B \ P\<^sub>1. c)" by simp also have "... < (\B \ P\<^sub>1. W B + w (f B))" using f_def(2) sum_strict_mono[OF FINITE(1) False] by fastforce also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>1. w (f B))" by (simp add: Groups_Big.comm_monoid_add_class.sum.distrib) also have "... = (\B \ P\<^sub>1. W B) + W (\ (P\<^sub>2 \ wrap B\<^sub>2))" unfolding sum.reindex_bij_betw[OF f_def(1), of w] .. also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" unfolding sum_eq .. also have "... = (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.union_disjoint[OF FINITE(1,2) DISJNT, of W] by (simp add: Un_assoc) also have "... \ (\u \ U. w u)" using * . also have "... \ c * card P" using sum_lower_bound_card[OF assms(1)] . - finally show ?thesis by (meson discrete nat_mult_less_cancel_disj of_nat_less_imp_less) + finally show ?thesis + by (simp flip: of_nat_mult) qed text \As @{thm wrap_card} holds, it follows that the amount of bins in \P\<^sub>1 \ wrap B\<^sub>1\ are a lower bound for the amount of bins in \P\.\ lemma P\<^sub>1_B\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P" proof - have "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P\<^sub>1 + card (wrap B\<^sub>1)" using card_Un_le by blast also have "... \ card P\<^sub>1 + 1" using wrap_card by simp also have "... \ card P" using P\<^sub>1_lower_bound_card[OF assms] . finally show ?thesis . qed text \If \inv\<^sub>1\ holds, there are at most half as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\, and we can again obtain a bijective function between the bins in \P\<^sub>1\ and the objects of the second partial solution, then the amount of bins in the second partial solution are a strict lower bound for half the bins of the first partial solution.\ lemma P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "2 * card P\<^sub>2 \ card (\P\<^sub>2)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card P\<^sub>1 + 1" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using invrules(4) by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using bp_bins_finite[OF invrules(1)] wrap_not_empty by blast have "finite P\<^sub>2" "finite (wrap B\<^sub>2)" using bp_sol_finite[OF invrules(1)] by blast+ have DISJNT2: "P\<^sub>2 \ wrap B\<^sub>2 = {}" unfolding wrap_def using \B\<^sub>2 \ P\<^sub>2\ by auto have "card (wrap B\<^sub>2) \ card B\<^sub>2" proof (cases \B\<^sub>2 = {}\) case False then have "1 \ card B\<^sub>2" by (simp add: leI \finite B\<^sub>2\) then show ?thesis using wrap_card[of B\<^sub>2] by linarith qed simp \ \This part of the proof is based on the proof on page 73 of the article \<^cite>\BerghammerR03\.\ from assms(2) have "2 * card P\<^sub>2 + 2 * card (wrap B\<^sub>2) \ card (\P\<^sub>2) + card (wrap B\<^sub>2) + 1" using wrap_card[of B\<^sub>2] by linarith then have "2 * (card P\<^sub>2 + card (wrap B\<^sub>2)) \ card (\P\<^sub>2) + card B\<^sub>2 + 1" using \card (wrap B\<^sub>2) \ card B\<^sub>2\ by simp then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\P\<^sub>2 \ B\<^sub>2) + 1" using card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\ DISJNT] and card_Un_disjoint[OF \finite P\<^sub>2\ \finite (wrap B\<^sub>2)\ DISJNT2] by argo then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\(P\<^sub>2 \ wrap B\<^sub>2)) + 1" by (cases \B\<^sub>2 = {}\) (auto simp: Un_commute) then show "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card P\<^sub>1 + 1" using assms(3) bij_betw_same_card unfolding bij_exists_def by metis qed subsubsection \Proving the Approximation Factor\ text \We define \inv\<^sub>2\ as it is defined in the article. These conjuncts allow us to prove the desired approximation factor.\ definition inv\<^sub>2 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ \\inv\<^sub>1\ holds for the partial solution\ \ (V \ L \ {} \ (\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {})) \ \If there are still large objects left, then every bin of the first partial solution must contain a large object\ \ bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2)) \ \There exists a bijective function between the bins of the first partial solution and the objects of the second one\ \ (2 * card P\<^sub>2 \ card (\P\<^sub>2)) \ \There are at most twice as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\\" lemma inv\<^sub>2E: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "V \ L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" using assms unfolding inv\<^sub>2_def by blast+ lemma inv\<^sub>2I: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "V \ L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>2_def by blast text \If \P\ is a correct solution of the bin packing problem, \inv\<^sub>2\ holds for the partial solution, and there are no more small objects left to be distributed, then the amount of bins of the partial solution is no larger than \3 / 2\ of the amount of bins in \P\. This proof strongly follows the proof in \Theorem 4.1\ of the article \<^cite>\BerghammerR03\.\ lemma bin_packing_lower_bound_card: assumes "V \ S = {}" "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "bp P" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ 3 / 2 * card P" proof (cases \V = {}\) note invrules = inv\<^sub>2E[OF assms(2)] case True then have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" by simp also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using P\<^sub>1_B\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by simp also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith next note invrules = inv\<^sub>2E[OF assms(2)] case False have "U = S \ L" using S_def L_def by blast then have *: "V = V \ L" using bpE(3)[OF inv\<^sub>1E(1)[OF invrules(1)]] and assms(1) by blast with False have NE: "V \ L \ {}" by simp have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L} \ P\<^sub>2 \ wrap B\<^sub>2)" using * by (simp add: Un_commute Un_assoc) also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ V \ L}\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using L_bins_lower_bound_card[OF assms(3) invrules(1) invrules(2)[OF NE]] by linarith also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith qed text \We define \inv\<^sub>3\ as it is defined in the article. This final conjunct allows us to prove that the invariant will be maintained by the algorithm.\ definition inv\<^sub>3 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ B\<^sub>2 \ S" lemma inv\<^sub>3E: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "B\<^sub>2 \ S" using assms unfolding inv\<^sub>3_def by blast+ lemma inv\<^sub>3I: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and "B\<^sub>2 \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>3_def by blast lemma loop_init: "inv\<^sub>3 {} {} {} {} U" proof - have *: "inv\<^sub>1 {} {} {} {} U" unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto have "bij_exists {} (\ ({} \ wrap {}))" using bij_betwI' unfolding bij_exists_def by fastforce from inv\<^sub>2I[OF * _ this] have "inv\<^sub>2 {} {} {} {} U" by auto from inv\<^sub>3I[OF this] show ?thesis by blast qed text \If \B\<^sub>1\ is empty and there are no large objects left, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is initialized with \u \ V \ S\.\ lemma loop_stepA: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 = {}" "V \ L = {}" "u \ V \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using S_def assms(2,4) by simp from assms(4) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by simp have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L \ {}" using assms(3) by blast from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is empty and there are large objects left, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is initialized with \u \ V \ L\.\ lemma loop_stepB: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 = {}" "u \ V \ L" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using L_def weight assms(2,3) by simp from assms(3) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by simp have "\B\P\<^sub>1. B \ L \ {}" using assms(3) invrules(2) by blast then have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L \ {}" using assms(3) by (metis Int_iff UnE empty_iff insertE singletonI wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ does not exceed its maximum capacity, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ and \{u}\ are replaced with \B\<^sub>1 \ {u}\.\ lemma loop_stepC: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) \ c" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepA[OF invrules(1) this assms(4)] have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" . have "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1. B \ L \ {}" using invrules(2) by blast then have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}). B \ L \ {}" by (metis Int_commute Un_empty_right Un_insert_right assms(2) disjoint_insert(2) insert_iff wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ does exceed its maximum capacity but not the capacity of \B\<^sub>2\, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ is added to \P\<^sub>1\ and emptied, and \B\<^sub>2\ and \{u}\ are replaced with \B\<^sub>2 \ {u}\.\ lemma loop_stepD: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) \ c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepB[OF invrules(1) this assms(5)] have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" . have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast then have "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" by (metis Sup_empty Un_assoc Union_Un_distrib ccpo_Sup_singleton wrap_empty wrap_not_empty) also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding bij_exists_def by blast from inv\<^sub>2I[OF 1 2 3] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" using invrules(4) by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] assms(3) by blast qed text \If the maximum capacity of \B\<^sub>2\ is exceeded by \u \ V \ S\, then \B\<^sub>2\ must contain at least two objects.\ lemma B\<^sub>2_at_least_two_objects: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V \ S" "W B\<^sub>2 + w(u) > c" shows "2 \ card B\<^sub>2" proof (rule ccontr, simp add: not_le) have FINITE: "finite B\<^sub>2" using inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF assms(1)]]] by (metis (no_types, lifting) Finite_Set.finite.simps U_Finite Union_Un_distrib bpE(3) ccpo_Sup_singleton finite_Un wrap_not_empty) assume "card B\<^sub>2 < 2" then consider (0) "card B\<^sub>2 = 0" | (1) "card B\<^sub>2 = 1" by linarith then show False proof cases case 0 then have "B\<^sub>2 = {}" using FINITE by simp then show ?thesis using assms(2,3) S_def by simp next case 1 then obtain v where "B\<^sub>2 = {v}" using card_1_singletonE by auto with inv\<^sub>3E(2)[OF assms(1)] have "2 * w v \ c" using S_def by simp moreover from \B\<^sub>2 = {v}\ have "W B\<^sub>2 = w v" by simp ultimately show ?thesis using assms(2,3) S_def by simp qed qed text \If \B\<^sub>1\ is not empty and \u \ V \ S\ exceeds the maximum capacity of both \B\<^sub>1\ and \B\<^sub>2\, then \inv\<^sub>3\ will be maintained if \B\<^sub>1\ and \B\<^sub>2\ are added to \P\<^sub>1\ and \P\<^sub>2\ respectively, emptied, and \B\<^sub>2\ initialized with \u\.\ lemma loop_stepE: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "B\<^sub>1 \ {}" "u \ V \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) > c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] from assms(3) have "u \ V" by blast from inv\<^sub>1_stepC[OF invrules(1) this] have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" . have 2: "(V - {u}) \ L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast have "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" unfolding wrap_def by simp also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding bij_exists_def by blast have 4: "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card (\ (P\<^sub>2 \ wrap B\<^sub>2))" proof - note bprules = bpE[OF inv\<^sub>1E(1)[OF invrules(1)]] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using inv\<^sub>1E(4)[OF invrules(1)] by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using inv\<^sub>1E(1)[OF invrules(1)] bp_bins_finite wrap_not_empty by blast have "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ 2 * (card P\<^sub>2 + card (wrap B\<^sub>2))" using card_Un_le[of P\<^sub>2 \wrap B\<^sub>2\] by simp also have "... \ 2 * card P\<^sub>2 + 2" using wrap_card by auto also have "... \ card (\ P\<^sub>2) + 2" using invrules(4) by simp also have "... \ card (\ P\<^sub>2) + card B\<^sub>2" using B\<^sub>2_at_least_two_objects[OF assms(1,3,5)] by simp also have "... = card (\ (P\<^sub>2 \ {B\<^sub>2}))" using DISJNT card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\] by (simp add: Un_commute) also have "... = card (\ (P\<^sub>2 \ wrap B\<^sub>2))" by (cases \B\<^sub>2 = {}\) auto finally show ?thesis . qed from inv\<^sub>2I[OF 1 2 3 4] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" . from inv\<^sub>3I[OF this] show ?thesis using assms(3) by blast qed text \The bin packing algorithm as it is proposed in the article \<^cite>\BerghammerR03\. \P\ will not only be a correct solution of the bin packing problem, but the amount of bins will be a lower bound for \3 / 2\ of the amount of bins of any correct solution \Q\, and thus guarantee an approximation factor of \3 / 2\ for the optimum.\ lemma bp_approx: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u {True} P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; V := U; WHILE V \ S \ {} INV {inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V} DO IF B\<^sub>1 \ {} THEN u := (SOME u. u \ V \ S) ELSE IF V \ L \ {} THEN u := (SOME u. u \ V \ L) ELSE u := (SOME u. u \ V \ S) FI FI; V := V - {u}; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} | v. v \ V} {bp P \ (\Q. bp Q \ card P \ 3 / 2 * card Q)}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then show ?case by (simp add: loop_init) next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" .. let ?s = "SOME u. u \ V \ S" let ?l = "SOME u. u \ V \ L" have LIN: "V \ L \ {} \ ?l \ V \ L" using some_in_eq by metis then have LWEIGHT: "V \ L \ {} \ w ?l \ c" using L_def weight by blast from 2 have "V \ S \ {}" .. then have IN: "?s \ V \ S" using some_in_eq by metis then have "w ?s \ c" using S_def by simp then show ?case using LWEIGHT loop_stepA[OF INV _ _ IN] loop_stepB[OF INV _ LIN] loop_stepC[OF INV _ IN] and loop_stepD[OF INV _ IN] loop_stepE[OF INV _ IN] by (cases \B\<^sub>1 = {}\, cases \V \ L = {}\) auto next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" and EMPTY: "V \ S = {}" by blast+ from inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF INV]]] and bin_packing_lower_bound_card[OF EMPTY inv\<^sub>3E(1)[OF INV]] show ?case by blast qed end (* BinPacking *) subsection \The Full Linear Time Version of the Proposed Algorithm\ text \Finally, we prove the Algorithm proposed on page 78 of the article \<^cite>\BerghammerR03\. This version generates the S and L sets beforehand and uses them directly to calculate the solution, thus removing the need for intersection operations, and ensuring linear time if we can perform \insertion, removal, and selection of an element, the union of two sets, and the emptiness test in constant time\ \<^cite>\BerghammerR03\.\ locale BinPacking_Complete = fixes U :: "'a set" \ \A finite, non-empty set of objects\ and w :: "'a \ real" \ \A mapping from objects to their respective weights (positive real numbers)\ and c :: nat \ \The maximum capacity of a bin (as a natural number)\ assumes weight: "\u \ U. 0 < w(u) \ w(u) \ c" and U_Finite: "finite U" and U_NE: "U \ {}" begin text \The correctness proofs will be identical to the ones of the simplified algorithm.\ abbreviation W :: "'a set \ real" where "W B \ (\u \ B. w(u))" definition bp :: "'a set set \ bool" where "bp P \ partition_on U P \ (\B \ P. W(B) \ c)" lemma bpE: assumes "bp P" shows "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" using assms unfolding bp_def partition_on_def by blast+ lemma bpI: assumes "pairwise disjnt P" "{} \ P" "\P = U" "\B \ P. W(B) \ c" shows "bp P" using assms unfolding bp_def partition_on_def by blast definition inv\<^sub>1 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V \ bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}) \ \A correct solution to the bin packing problem\ \ \(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V \ \The partial solution does not contain objects that have not yet been assigned\ \ B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ \\B\<^sub>1\ is distinct from all the other bins\ \ B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2) \ \\B\<^sub>2\ is distinct from all the other bins\ \ (P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {} \ \The first and second partial solutions are disjoint from each other.\" lemma inv\<^sub>1E: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" shows "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using assms unfolding inv\<^sub>1_def by auto lemma inv\<^sub>1I: assumes "bp (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V})" and "\(P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) = U - V" and "B\<^sub>1 \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" and "B\<^sub>2 \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2)" and "(P\<^sub>1 \ wrap B\<^sub>1) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" using assms unfolding inv\<^sub>1_def by blast lemma wrap_Un [simp]: "wrap (M \ {x}) = {M \ {x}}" unfolding wrap_def by simp lemma wrap_empty [simp]: "wrap {} = {}" unfolding wrap_def by simp lemma wrap_not_empty [simp]: "M \ {} \ wrap M = {M}" unfolding wrap_def by simp lemma inv\<^sub>1_stepA: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W(B\<^sub>1) + w(u) \ c" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] \ \Rule 1: Pairwise Disjoint\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ ({{u}} \ {{v} |v. v \ V - {u}}))" using bprules(1) assms(2) by simp then have "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" by (simp add: Un_commute) then have assm: "pairwise disjnt (wrap B\<^sub>1 \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc) have "pairwise disjnt ({B\<^sub>1 \ {u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>1 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>1} \ {{u}} \ (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,3) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>1" using bprules(3) U_Finite by (cases \B\<^sub>1 = {}\) auto then have "W (B\<^sub>1 \ {u}) \ W B\<^sub>1 + w u" using \0 < w u\ by (cases \u \ B\<^sub>1\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>1 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>1 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" by blast from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2 \ wrap B\<^sub>2) = U - (V - {u})" unfolding L R invrules(2) .. have 3: "B\<^sub>1 \ {u} \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2" using NOTIN by auto have 4: "B\<^sub>2 \ P\<^sub>1 \ wrap (B\<^sub>1 \ {u}) \ P\<^sub>2" using invrules(4) NOTIN unfolding wrap_def by fastforce have 5: "(P\<^sub>1 \ wrap (B\<^sub>1 \ {u})) \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) NOTIN unfolding wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed lemma inv\<^sub>1_stepB: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" "W B\<^sub>2 + w u \ c" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "{{v} |v. v \ V} = {{u}} \ {{v} |v. v \ V - {u}}" using assms(2) by blast then have "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}})" using bprules(1) assms(2) by simp then have assm: "pairwise disjnt (wrap B\<^sub>2 \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by (simp add: Un_assoc Un_commute) have "pairwise disjnt ({B\<^sub>2 \ {u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" proof (cases \B\<^sub>2 = {}\) case True with assm show ?thesis by simp next case False with assm have assm: "pairwise disjnt ({B\<^sub>2} \ {{u}} \ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}))" by simp from NOTIN have "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}" by blast from pairwise_disjnt_Un[OF assm _ this] invrules(2,4) show ?thesis using False by auto qed then have 1: "pairwise disjnt (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" unfolding wrap_Un by simp \ \Rule 2: No empty sets\ from bprules(2) have 2: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}" unfolding wrap_def by simp \ \Rule 3: Union preserved\ from bprules(3) have "\ (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}) = U" using assms(2) by blast then have 3: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}) = U" unfolding wrap_def by force \ \Rule 4: Weights below capacity\ have "0 < w u" using weight assms(2) bprules(3) by blast have "finite B\<^sub>2" using bprules(3) U_Finite by (cases \B\<^sub>2 = {}\) auto then have "W (B\<^sub>2 \ {u}) \ W B\<^sub>2 + w u" using \0 < w u\ by (cases \u \ B\<^sub>2\) (auto simp: insert_absorb) also have "... \ c" using assms(3) . finally have "W (B\<^sub>2 \ {u}) \ c" . then have "\B \ wrap (B\<^sub>2 \ {u}). W B \ c" unfolding wrap_Un by blast moreover have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ {{v} |v. v \ V - {u}}. W B \ c" using bprules(4) by blast ultimately have 4: "\B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}}. W B \ c" by auto from bpI[OF 1 2 3 4] have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u}) \ {{v} |v. v \ V - {u}})" . \ \Auxiliary information is preserved\ have "u \ U" using assms(2) bprules(3) by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap B\<^sub>2) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = U - (V - {u})" unfolding L R using invrules(2) by simp have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap (B\<^sub>2 \ {u})" using bpE(2)[OF 1] by simp have 4: "B\<^sub>2 \ {u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ P\<^sub>2" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})) = {}" using invrules(5) NOTIN unfolding wrap_empty wrap_Un by auto from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed lemma inv\<^sub>1_stepC: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V" "u \ V" shows "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (V - {u})" proof - note invrules = inv\<^sub>1E[OF assms(1)] \ \Rule 1-4: Correct Bin Packing\ have "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{u}} \ {{v} |v. v \ V - {u}}" by (metis (no_types, lifting) Un_assoc Un_empty_right insert_not_empty wrap_empty wrap_not_empty) also have "... = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" using assms(2) by auto finally have EQ: "P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V}" . from invrules(1) have 1: "bp (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u} \ {{v} |v. v \ V - {u}})" unfolding EQ . \ \Auxiliary information is preserved\ have NOTIN: "\M \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ V - {u}}. u \ M" using invrules(2) assms(2) by blast have "u \ U" using assms(2) bpE(3)[OF invrules(1)] by blast then have R: "U - (V - {u}) = U - V \ {u}" by blast have L: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = \ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" unfolding wrap_def using NOTIN by auto have 2: "\ (P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}) = U - (V - {u})" unfolding L R using invrules(2) by auto have 3: "{} \ P\<^sub>1 \ wrap B\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) \ wrap {u}" using bpE(2)[OF 1] by simp have 4: "{u} \ P\<^sub>1 \ wrap B\<^sub>1 \ wrap {} \ (P\<^sub>2 \ wrap B\<^sub>2)" using NOTIN by auto have 5: "(P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}) \ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}) = {}" using invrules(5) NOTIN unfolding wrap_def by force from inv\<^sub>1I[OF 1 2 3 4 5] show ?thesis . qed text \From this point onward, we will require a different approach for proving lower bounds. Instead of fixing and assuming the definitions of the \S\ and \L\ sets, we will introduce the abbreviations \S\<^sub>U\ and \L\<^sub>U\ for any occurrences of the original \S\ and \L\ sets. The union of \S\ and \L\ can be interpreted as \V\. As a result, occurrences of \V \ S\ become \(S \ L) \ S = S\, and \V \ L\ become \(S \ L) \ L = L\. Occurrences of these sets will have to be replaced appropriately.\ abbreviation S\<^sub>U where "S\<^sub>U \ {u \ U. w u \ c / 2}" abbreviation L\<^sub>U where "L\<^sub>U \ {u \ U. c / 2 < w u}" text \As we will remove elements from \S\ and \L\, we will only be able to show that they remain subsets of \S\<^sub>U\ and \L\<^sub>U\ respectively.\ abbreviation SL where "SL S L \ S \ S\<^sub>U \ L \ L\<^sub>U" lemma bp_bins_finite [simp]: assumes "bp P" shows "\B \ P. finite B" using bpE(3)[OF assms] U_Finite by (meson Sup_upper finite_subset) lemma bp_sol_finite [simp]: assumes "bp P" shows "finite P" using bpE(3)[OF assms] U_Finite by (simp add: finite_UnionD) lemma only_one_L_per_bin: assumes "bp P" "B \ P" shows "\x \ B. \y \ B. x \ y \ x \ L\<^sub>U \ y \ L\<^sub>U" proof (rule ccontr, simp) assume "\x\B. \y\B. x \ y \ y \ U \ x \ U \ real c < w x * 2 \ real c < w y * 2" then obtain x y where *: "x \ B" "y \ B" "x \ y" "x \ L\<^sub>U" "y \ L\<^sub>U" by auto then have "c < w x + w y" by force have "finite B" using assms by simp have "y \ B - {x}" using *(2,3) by blast have "W B = W (B - {x}) + w x" using *(1) \finite B\ by (simp add: sum.remove) also have "... = W (B - {x} - {y}) + w x + w y" using \y \ B - {x}\ \finite B\ by (simp add: sum.remove) finally have *: "W B = W (B - {x} - {y}) + w x + w y" . have "\u \ B. 0 < w u" using bpE(3)[OF assms(1)] assms(2) weight by blast then have "0 \ W (B - {x} - {y})" by (smt DiffD1 sum_nonneg) with * have "c < W B" using \c < w x + w y\ by simp then show False using bpE(4)[OF assms(1)] assms(2) by fastforce qed lemma L_lower_bound_card: assumes "bp P" shows "card L\<^sub>U \ card P" proof - have "\x \ L\<^sub>U. \B \ P. x \ B" using bpE(3)[OF assms] by blast then obtain f where f_def: "\u \ L\<^sub>U. u \ f u \ f u \ P" by metis then have "inj_on f L\<^sub>U" unfolding inj_on_def using only_one_L_per_bin[OF assms] by blast then have card_eq: "card L\<^sub>U = card (f ` L\<^sub>U)" by (simp add: card_image) have "f ` L\<^sub>U \ P" using f_def by blast moreover have "finite P" using assms by simp ultimately have "card (f ` L\<^sub>U) \ card P" by (simp add: card_mono) then show ?thesis unfolding card_eq . qed lemma subset_bp_card: assumes "bp P" "M \ P" "\B \ M. B \ L\<^sub>U \ {}" shows "card M \ card L\<^sub>U" proof - have "\B \ M. \u \ L\<^sub>U. u \ B" using assms(3) by fast then have "\f. \B \ M. f B \ L\<^sub>U \ f B \ B" by metis then obtain f where f_def: "\B \ M. f B \ L\<^sub>U \ f B \ B" .. have "inj_on f M" proof (rule ccontr) assume "\ inj_on f M" then have "\x \ M. \y \ M. x \ y \ f x = f y" unfolding inj_on_def by blast then obtain x y where *: "x \ M" "y \ M" "x \ y" "f x = f y" by blast then have "\u. u \ x \ u \ y" using f_def by metis then have "x \ y \ {}" by blast moreover have "pairwise disjnt M" using pairwise_subset[OF bpE(1)[OF assms(1)] assms(2)] . ultimately show False using * unfolding pairwise_def disjnt_def by simp qed moreover have "finite L\<^sub>U" using U_Finite by auto moreover have "f ` M \ L\<^sub>U" using f_def by blast ultimately show ?thesis using card_inj_on_le by blast qed lemma L_bins_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and SL_def: "SL S L" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}) \ card P" proof - note invrules = inv\<^sub>1E[OF assms(2)] have "\B \ {{v} |v. v \ L}. B \ L\<^sub>U \ {}" using SL_def by blast with assms(3) have "P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L} \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}" "\B \ P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}. B \ L\<^sub>U \ {}" by blast+ from subset_bp_card[OF invrules(1) this] show ?thesis using L_lower_bound_card[OF assms(1)] by linarith qed lemma sum_Un_eq_sum_sum: assumes "bp P" shows "(\u \ U. w u) = (\B \ P. W B)" proof - have FINITE: "\B \ P. finite B" using assms by simp have DISJNT: "\A \ P. \B \ P. A \ B \ A \ B = {}" using bpE(1)[OF assms] unfolding pairwise_def disjnt_def . have "(\u \ (\P). w u) = (\B \ P. W B)" using sum.Union_disjoint[OF FINITE DISJNT] by auto then show ?thesis unfolding bpE(3)[OF assms] . qed lemma sum_lower_bound_card: assumes "bp P" shows "(\u \ U. w u) \ c * card P" proof - have *: "\B \ P. 0 < W B \ W B \ c" using bpE(2-4)[OF assms] weight by (metis UnionI assms bp_bins_finite sum_pos) have "(\u \ U. w u) = (\B \ P. W B)" using sum_Un_eq_sum_sum[OF assms] . also have "... \ (\B \ P. c)" using sum_mono * by fastforce also have "... = c * card P" by simp finally show ?thesis . qed lemma bp_NE: assumes "bp P" shows "P \ {}" using U_NE bpE(3)[OF assms] by blast lemma sum_Un_ge: fixes f :: "_ \ real" assumes "finite M" "finite N" "\B \ M \ N. 0 < f B" shows "sum f M \ sum f (M \ N)" proof - have "0 \ sum f N - sum f (M \ N)" using assms by (smt DiffD1 inf.cobounded2 UnCI sum_mono2) then have "sum f M \ sum f M + sum f N - sum f (M \ N)" by simp also have "... = sum f (M \ N)" using sum_Un[OF assms(1,2), symmetric] . finally show ?thesis . qed definition bij_exists :: "'a set set \ 'a set \ bool" where "bij_exists P V = (\f. bij_betw f P V \ (\B \ P. W B + w (f B) > c))" lemma P\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card P\<^sub>1 + 1 \ card P" proof (cases \P\<^sub>1 = {}\) case True have "finite P" using assms(1) by simp then have "1 \ card P" using bp_NE[OF assms(1)] by (metis Nat.add_0_right Suc_diff_1 Suc_le_mono card_gt_0_iff le0 mult_Suc_right nat_mult_1) then show ?thesis unfolding True by simp next note invrules = inv\<^sub>1E[OF assms(2)] case False obtain f where f_def: "bij_betw f P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" "\B \ P\<^sub>1. W B + w (f B) > c" using assms(3) unfolding bij_exists_def by blast have FINITE: "finite P\<^sub>1" "finite (P\<^sub>2 \ wrap B\<^sub>2)" "finite (P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" "finite (wrap B\<^sub>1 \ {{v} |v. v \ S \ L})" using inv\<^sub>1E(1)[OF assms(2)] bp_sol_finite by blast+ have F: "\B \ P\<^sub>2 \ wrap B\<^sub>2. finite B" using invrules(1) by simp have D: "\A \ P\<^sub>2 \ wrap B\<^sub>2. \B \ P\<^sub>2 \ wrap B\<^sub>2. A \ B \ A \ B = {}" using bpE(1)[OF invrules(1)] unfolding pairwise_def disjnt_def by auto have sum_eq: "W (\ (P\<^sub>2 \ wrap B\<^sub>2)) = (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.Union_disjoint[OF F D] by auto have "\B\P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}. 0 < W B" using bpE(2,3)[OF invrules(1)] weight by (metis (no_types, lifting) UnionI bp_bins_finite invrules(1) sum_pos) then have "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ (wrap B\<^sub>1 \ {{v} |v. v \ S \ L}). W B)" using sum_Un_ge[OF FINITE(3,4), of W] by blast also have "... = (\B \ P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}. W B)" by (smt Un_assoc Un_commute) also have "... = W U" using sum_Un_eq_sum_sum[OF invrules(1), symmetric] . finally have *: "(\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B) \ W U" . have DISJNT: "P\<^sub>1 \ (P\<^sub>2 \ wrap B\<^sub>2) = {}" using invrules(5) by blast \ \This part of the proof is based on the proof on page 72 of the article \<^cite>\BerghammerR03\.\ have "c * card P\<^sub>1 = (\B \ P\<^sub>1. c)" by simp also have "... < (\B \ P\<^sub>1. W B + w (f B))" using f_def(2) sum_strict_mono[OF FINITE(1) False] by fastforce also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>1. w (f B))" by (simp add: Groups_Big.comm_monoid_add_class.sum.distrib) also have "... = (\B \ P\<^sub>1. W B) + W (\ (P\<^sub>2 \ wrap B\<^sub>2))" unfolding sum.reindex_bij_betw[OF f_def(1), of w] .. also have "... = (\B \ P\<^sub>1. W B) + (\B \ P\<^sub>2 \ wrap B\<^sub>2. W B)" unfolding sum_eq .. also have "... = (\B \ P\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2. W B)" using sum.union_disjoint[OF FINITE(1,2) DISJNT, of W] by (simp add: Un_assoc) also have "... \ (\u \ U. w u)" using * . also have "... \ c * card P" using sum_lower_bound_card[OF assms(1)] . - finally show ?thesis by (meson discrete nat_mult_less_cancel_disj of_nat_less_imp_less) + finally show ?thesis + by (simp flip: of_nat_mult) qed lemma P\<^sub>1_B\<^sub>1_lower_bound_card: assumes "bp P" "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P" proof - have "card (P\<^sub>1 \ wrap B\<^sub>1) \ card P\<^sub>1 + card (wrap B\<^sub>1)" using card_Un_le by blast also have "... \ card P\<^sub>1 + 1" using wrap_card by simp also have "... \ card P" using P\<^sub>1_lower_bound_card[OF assms] . finally show ?thesis . qed lemma P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" "2 * card P\<^sub>2 \ card (\P\<^sub>2)" "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" shows "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card P\<^sub>1 + 1" proof - note invrules = inv\<^sub>1E[OF assms(1)] and bprules = bpE[OF invrules(1)] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using invrules(4) by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using bp_bins_finite[OF invrules(1)] wrap_not_empty by blast have "finite P\<^sub>2" "finite (wrap B\<^sub>2)" using bp_sol_finite[OF invrules(1)] by blast+ have DISJNT2: "P\<^sub>2 \ wrap B\<^sub>2 = {}" unfolding wrap_def using \B\<^sub>2 \ P\<^sub>2\ by auto have "card (wrap B\<^sub>2) \ card B\<^sub>2" proof (cases \B\<^sub>2 = {}\) case False then have "1 \ card B\<^sub>2" by (simp add: leI \finite B\<^sub>2\) then show ?thesis using wrap_card[of B\<^sub>2] by linarith qed simp \ \This part of the proof is based on the proof on page 73 of the article \<^cite>\BerghammerR03\.\ from assms(2) have "2 * card P\<^sub>2 + 2 * card (wrap B\<^sub>2) \ card (\P\<^sub>2) + card (wrap B\<^sub>2) + 1" using wrap_card[of B\<^sub>2] by linarith then have "2 * (card P\<^sub>2 + card (wrap B\<^sub>2)) \ card (\P\<^sub>2) + card B\<^sub>2 + 1" using \card (wrap B\<^sub>2) \ card B\<^sub>2\ by simp then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\P\<^sub>2 \ B\<^sub>2) + 1" using card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\ DISJNT] and card_Un_disjoint[OF \finite P\<^sub>2\ \finite (wrap B\<^sub>2)\ DISJNT2] by argo then have "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card (\(P\<^sub>2 \ wrap B\<^sub>2)) + 1" by (cases \B\<^sub>2 = {}\) (auto simp: Un_commute) then show "2 * (card (P\<^sub>2 \ wrap B\<^sub>2)) \ card P\<^sub>1 + 1" using assms(3) bij_betw_same_card unfolding bij_exists_def by metis qed text \We add \SL S L\ to \inv\<^sub>2\ to ensure that the \S\ and \L\ sets only contain objects with correct weights.\ definition inv\<^sub>2 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L) \ \\inv\<^sub>1\ holds for the partial solution\ \ (L \ {} \ (\B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {})) \ \If there are still large objects left, then every bin of the first partial solution must contain a large object\ \ bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2)) \ \There exists a bijective function between the bins of the first partial solution and the objects of the second one\ \ (2 * card P\<^sub>2 \ card (\P\<^sub>2)) \ \There are at most twice as many bins in \P\<^sub>2\ as there are objects in \P\<^sub>2\\ \ SL S L \ \\S\ and \L\ are subsets of \S\<^sub>U\ and \L\<^sub>U\\" lemma inv\<^sub>2E: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" shows "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" and "L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" and "SL S L" using assms unfolding inv\<^sub>2_def by blast+ lemma inv\<^sub>2I: assumes "inv\<^sub>1 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 (S \ L)" and "L \ {} \ \B \ P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" and "bij_exists P\<^sub>1 (\(P\<^sub>2 \ wrap B\<^sub>2))" and "2 * card P\<^sub>2 \ card (\P\<^sub>2)" and "SL S L" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using assms unfolding inv\<^sub>2_def by blast lemma bin_packing_lower_bound_card: assumes "S = {}" "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "bp P" shows "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) \ 3 / 2 * card P" proof (cases \L = {}\) note invrules = inv\<^sub>2E[OF assms(2)] case True then have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2)" using assms(1) by simp also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using P\<^sub>1_B\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by simp also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith next note invrules = inv\<^sub>2E[OF assms(2)] case False have "card (P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ S \ L}) = card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L} \ P\<^sub>2 \ wrap B\<^sub>2)" using assms(1) by (simp add: Un_commute Un_assoc) also have "... \ card (P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}) + card (P\<^sub>2 \ wrap B\<^sub>2)" using card_Un_le[of \P\<^sub>1 \ wrap B\<^sub>1 \ {{v} |v. v \ L}\] by (simp add: Un_assoc) also have "... \ card P + card (P\<^sub>2 \ wrap B\<^sub>2)" using L_bins_lower_bound_card[OF assms(3) invrules(1) invrules(2)[OF False] invrules(5)] by linarith also have "... \ card P + card P / 2" using P\<^sub>2_B\<^sub>2_lower_bound_P\<^sub>1[OF invrules(1,4,3)] and P\<^sub>1_lower_bound_card[OF assms(3) invrules(1,3)] by linarith finally show ?thesis by linarith qed definition inv\<^sub>3 :: "'a set set \ 'a set set \ 'a set \ 'a set \ 'a set \ 'a set \ bool" where "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ B\<^sub>2 \ S\<^sub>U" lemma inv\<^sub>3E: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" shows "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" and "B\<^sub>2 \ S\<^sub>U" using assms unfolding inv\<^sub>3_def by blast+ lemma inv\<^sub>3I: assumes "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" and "B\<^sub>2 \ S\<^sub>U" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using assms unfolding inv\<^sub>3_def by blast lemma loop_init: "inv\<^sub>3 {} {} {} {} S\<^sub>U L\<^sub>U" proof - have "S\<^sub>U \ L\<^sub>U = U" by auto then have *: "inv\<^sub>1 {} {} {} {} (S\<^sub>U \ L\<^sub>U)" unfolding bp_def partition_on_def pairwise_def wrap_def inv\<^sub>1_def using weight by auto have "bij_exists {} (\ ({} \ wrap {}))" using bij_betwI' unfolding bij_exists_def by fastforce from inv\<^sub>2I[OF * _ this] have "inv\<^sub>2 {} {} {} {} S\<^sub>U L\<^sub>U" by auto from inv\<^sub>3I[OF this] show ?thesis by blast qed lemma loop_stepA: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 = {}" "L = {}" "u \ S" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using invrules(5) assms(2,4) by fastforce from assms(4) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2,3) have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L\<^sub>U \ {}" using assms(3) by blast from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S - {u}) L" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepB: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 = {}" "u \ L" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 S (L - {u})" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have WEIGHT: "W B\<^sub>1 + w u \ c" using weight invrules(5) assms(2,3) by fastforce \ \This observation follows from the fact that the \S\ and \L\ sets have to be disjoint from each other, and allows us to reuse our proofs of the preservation of \inv\<^sub>1\ by simply replacing \V\ with \S \ L\\ have *: "S \ L - {u} = S \ (L - {u})" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this WEIGHT] assms(2) * have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 (S \ (L - {u}))" by simp have "\B\P\<^sub>1. B \ L\<^sub>U \ {}" "{u} \ L\<^sub>U \ {}" using assms(3) invrules(2,5) by blast+ then have 2: "L \ {} \ \B\P\<^sub>1 \ wrap {u}. B \ L\<^sub>U \ {}" using assms(3) by (metis (full_types) Un_iff empty_iff insert_iff wrap_not_empty) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 {u} B\<^sub>2 S (L - {u})" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepC: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) \ c" shows "inv\<^sub>3 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] \ \Same approach, but removing \{u}\ from \S\ instead of \L\\ have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepA[OF invrules(1) this assms(4)] * have 1: "inv\<^sub>1 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u} \ L)" by simp have "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1. B \ L\<^sub>U \ {}" using invrules(2) by blast then have 2: "L \ {} \ \B\P\<^sub>1 \ wrap (B\<^sub>1 \ {u}). B \ L\<^sub>U \ {}" by (smt Int_insert_left Un_empty_right Un_iff Un_insert_right assms(2) insert_not_empty singletonD singletonI wrap_def) from inv\<^sub>2I[OF 1 2] invrules have "inv\<^sub>2 P\<^sub>1 P\<^sub>2 (B\<^sub>1 \ {u}) B\<^sub>2 (S - {u}) L" by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] . qed lemma loop_stepD: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) \ c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepB[OF invrules(1) this assms(5)] * have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L\<^sub>U \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast then have "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" by (metis Sup_empty Un_assoc Union_Un_distrib ccpo_Sup_singleton wrap_empty wrap_not_empty) also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u}))) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap (B\<^sub>2 \ {u})))" unfolding bij_exists_def by blast from inv\<^sub>2I[OF 1 2 3] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) P\<^sub>2 {} (B\<^sub>2 \ {u}) (S - {u}) L" using invrules(4,5) by blast from inv\<^sub>3I[OF this] show ?thesis using inv\<^sub>3E(2)[OF assms(1)] assms(3) invrules(5) by blast qed lemma B\<^sub>2_at_least_two_objects: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "u \ S" "W B\<^sub>2 + w(u) > c" shows "2 \ card B\<^sub>2" proof (rule ccontr, simp add: not_le) have FINITE: "finite B\<^sub>2" using inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF inv\<^sub>3E(1)[OF assms(1)]]] by (metis (no_types, lifting) Finite_Set.finite.simps U_Finite Union_Un_distrib bpE(3) ccpo_Sup_singleton finite_Un wrap_not_empty) assume "card B\<^sub>2 < 2" then consider (0) "card B\<^sub>2 = 0" | (1) "card B\<^sub>2 = 1" by linarith then show False proof cases case 0 then have "B\<^sub>2 = {}" using FINITE by simp then show ?thesis using assms(2,3) inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by force next case 1 then obtain v where "B\<^sub>2 = {v}" using card_1_singletonE by auto with inv\<^sub>3E(2)[OF assms(1)] have "2 * w v \ c" using inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by simp moreover from \B\<^sub>2 = {v}\ have "W B\<^sub>2 = w v" by simp ultimately show ?thesis using assms(2,3) inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF assms(1)]] by force qed qed lemma loop_stepE: assumes "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" "B\<^sub>1 \ {}" "u \ S" "W B\<^sub>1 + w(u) > c" "W B\<^sub>2 + w(u) > c" shows "inv\<^sub>3 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u}) L" proof - note invrules = inv\<^sub>2E[OF inv\<^sub>3E(1)[OF assms(1)]] have *: "S \ L - {u} = (S - {u}) \ L" using invrules(5) assms(3) by force from assms(3) have "u \ S \ L" by blast from inv\<^sub>1_stepC[OF invrules(1) this] * have 1: "inv\<^sub>1 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u} \ L)" by simp have 2: "L \ {} \ \B\P\<^sub>1 \ wrap B\<^sub>1 \ wrap {}. B \ L\<^sub>U \ {}" using invrules(2) unfolding wrap_empty by blast from invrules(3) obtain f where f_def: "bij_betw f P\<^sub>1 (\ (P\<^sub>2 \ wrap B\<^sub>2))" "\B\P\<^sub>1. c < W B + w (f B)" unfolding bij_exists_def by blast have "B\<^sub>1 \ P\<^sub>1" using inv\<^sub>1E(3)[OF invrules(1)] by blast have "u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))" using inv\<^sub>1E(2)[OF invrules(1)] assms(3) by blast have "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2 \ {{u}}))" unfolding wrap_def by simp also have "... = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" by simp finally have UN: "(\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) = (\ (P\<^sub>2 \ wrap B\<^sub>2)) \ {u}" . have "wrap B\<^sub>1 = {B\<^sub>1}" using wrap_not_empty[of B\<^sub>1] assms(2) by simp let ?f = "f (B\<^sub>1 := u)" have BIJ: "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding wrap_empty \wrap B\<^sub>1 = {B\<^sub>1}\ UN using f_def(1) \B\<^sub>1 \ P\<^sub>1\ \u \ (\ (P\<^sub>2 \ wrap B\<^sub>2))\ by (metis (no_types, lifting) bij_betw_cong fun_upd_other fun_upd_same notIn_Un_bij_betw3) have "c < W B\<^sub>1 + w (?f B\<^sub>1)" using assms(4) by simp then have "(\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" unfolding \wrap B\<^sub>1 = {B\<^sub>1}\ using f_def(2) by simp with BIJ have "bij_betw ?f (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u})) \ (\B\P\<^sub>1 \ wrap B\<^sub>1. c < W B + w (?f B))" by blast then have 3: "bij_exists (P\<^sub>1 \ wrap B\<^sub>1) (\ (P\<^sub>2 \ wrap B\<^sub>2 \ wrap {u}))" unfolding bij_exists_def by blast have 4: "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ card (\ (P\<^sub>2 \ wrap B\<^sub>2))" proof - note bprules = bpE[OF inv\<^sub>1E(1)[OF invrules(1)]] have "pairwise disjnt (P\<^sub>2 \ wrap B\<^sub>2)" using bprules(1) pairwise_subset by blast moreover have "B\<^sub>2 \ P\<^sub>2" using inv\<^sub>1E(4)[OF invrules(1)] by simp ultimately have DISJNT: "\P\<^sub>2 \ B\<^sub>2 = {}" by (auto, metis (no_types, opaque_lifting) sup_bot.right_neutral Un_insert_right disjnt_iff mk_disjoint_insert pairwise_insert wrap_Un) have "finite (\P\<^sub>2)" using U_Finite bprules(3) by auto have "finite B\<^sub>2" using inv\<^sub>1E(1)[OF invrules(1)] bp_bins_finite wrap_not_empty by blast have "2 * card (P\<^sub>2 \ wrap B\<^sub>2) \ 2 * (card P\<^sub>2 + card (wrap B\<^sub>2))" using card_Un_le[of P\<^sub>2 \wrap B\<^sub>2\] by simp also have "... \ 2 * card P\<^sub>2 + 2" using wrap_card by auto also have "... \ card (\ P\<^sub>2) + 2" using invrules(4) by simp also have "... \ card (\ P\<^sub>2) + card B\<^sub>2" using B\<^sub>2_at_least_two_objects[OF assms(1,3,5)] by simp also have "... = card (\ (P\<^sub>2 \ {B\<^sub>2}))" using DISJNT card_Un_disjoint[OF \finite (\P\<^sub>2)\ \finite B\<^sub>2\] by (simp add: Un_commute) also have "... = card (\ (P\<^sub>2 \ wrap B\<^sub>2))" by (cases \B\<^sub>2 = {}\) auto finally show ?thesis . qed from inv\<^sub>2I[OF 1 2 3 4] have "inv\<^sub>2 (P\<^sub>1 \ wrap B\<^sub>1) (P\<^sub>2 \ wrap B\<^sub>2) {} {u} (S - {u}) L" using invrules(5) by blast from inv\<^sub>3I[OF this] show ?thesis using assms(3) invrules(5) by blast qed text \The bin packing algorithm as it is proposed on page 78 of the article \<^cite>\BerghammerR03\. \P\ will not only be a correct solution of the bin packing problem, but the amount of bins will be a lower bound for \3 / 2\ of the amount of bins of any correct solution \Q\, and thus guarantee an approximation factor of \3 / 2\ for the optimum.\ lemma bp_approx: "VARS P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u {True} S := {}; L:= {}; V := U; WHILE V \ {} INV {V \ U \ S = {u \ U - V. w(u) \ c / 2} \ L = {u \ U - V. c / 2 < w(u)}} DO u := (SOME u. u \ V); IF w(u) \ c / 2 THEN S := S \ {u} ELSE L := L \ {u} FI; V := V - {u} OD; P\<^sub>1 := {}; P\<^sub>2 := {}; B\<^sub>1 := {}; B\<^sub>2 := {}; WHILE S \ {} INV {inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L} DO IF B\<^sub>1 \ {} THEN u := (SOME u. u \ S); S := S - {u} ELSE IF L \ {} THEN u := (SOME u. u \ L); L := L - {u} ELSE u := (SOME u. u \ S); S := S - {u} FI FI; IF W(B\<^sub>1) + w(u) \ c THEN B\<^sub>1 := B\<^sub>1 \ {u} ELSE IF W(B\<^sub>2) + w(u) \ c THEN B\<^sub>2 := B\<^sub>2 \ {u} ELSE P\<^sub>2 := P\<^sub>2 \ wrap B\<^sub>2; B\<^sub>2 := {u} FI; P\<^sub>1 := P\<^sub>1 \ wrap B\<^sub>1; B\<^sub>1 := {} FI OD; P := P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2; V := L; WHILE V \ {} INV {S = {} \ inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L \ V \ L \ P = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v}|v. v \ L - V}} DO u := (SOME u. u \ V); P := P \ {{u}}; V := V - {u} OD {bp P \ (\Q. bp Q \ card P \ 3 / 2 * card Q)}" proof (vcg, goal_cases) case (1 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by blast next case (2 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by (auto simp: some_in_eq) next case (3 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case using loop_init by force next case (4 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have INV: "inv\<^sub>3 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" .. let ?s = "SOME u. u \ S" let ?l = "SOME u. u \ L" note SL_def = inv\<^sub>2E(5)[OF inv\<^sub>3E(1)[OF INV]] have LIN: "L \ {} \ ?l \ L" using some_in_eq by metis then have LWEIGHT: "L \ {} \ w ?l \ c" using weight SL_def by blast from 4 have "S \ {}" .. then have IN: "?s \ S" using some_in_eq by metis then have "w ?s \ c" using SL_def by auto then show ?case using LWEIGHT loop_stepA[OF INV _ _ IN] loop_stepB[OF INV _ LIN] loop_stepC[OF INV _ IN] and loop_stepD[OF INV _ IN] loop_stepE[OF INV _ IN] by (cases \B\<^sub>1 = {}\, cases \L = {}\) auto next case (5 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then show ?case by blast next case (6 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have *: "(SOME u. u \ V) \ V" "(SOME u. u \ V) \ L" by (auto simp add: some_in_eq) then have "P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ L - (V - {SOME u. u \ V})} = P\<^sub>1 \ wrap B\<^sub>1 \ P\<^sub>2 \ wrap B\<^sub>2 \ {{v} |v. v \ L - V \ {SOME u. u \ V}}" by blast with 6 * show ?case by blast next case (7 P P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 V S L u) then have *: "inv\<^sub>2 P\<^sub>1 P\<^sub>2 B\<^sub>1 B\<^sub>2 S L" using inv\<^sub>3E(1) by blast from inv\<^sub>1E(1)[OF inv\<^sub>2E(1)[OF *]] 7 have "bp P" by fastforce with bin_packing_lower_bound_card[OF _ *] 7 show ?case by fastforce qed end (* BinPacking_Complete *) end (* Theory *) \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/BKR_Proofs.thy b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy --- a/thys/BenOr_Kozen_Reif/BKR_Proofs.thy +++ b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy @@ -1,2232 +1,2244 @@ theory BKR_Proofs imports Matrix_Equation_Construction begin definition well_def_signs:: "nat => rat list list \ bool" where "well_def_signs num_polys sign_conds \ \ signs \ set(sign_conds). (length signs = num_polys)" definition satisfies_properties:: "real poly \ real poly list \ nat list list \ rat list list \ rat mat \ bool" where "satisfies_properties p qs subsets signs matrix = ( all_list_constr subsets (length qs) \ well_def_signs (length qs) signs \ distinct signs \ satisfy_equation p qs subsets signs \ invertible_mat matrix \ matrix = matrix_A signs subsets \ set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs) )" section "Base Case" lemma mat_base_case: shows "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1, 1], [1, -1]])" unfolding matrix_A_def mtx_row_def z_def apply (auto) by (simp add: numeral_2_eq_2) lemma base_case_sgas: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto lemma base_case_sgas_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ {[1], [- 1]}" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_sgas nonzero rel_prime apply (auto) using characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto qed lemma base_case_satisfy_equation: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "satisfy_equation p [q] [[],[0]] [[1],[-1]]" proof - have h1: "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" using base_case_sgas assms by auto have h2: " \qa. List.member [q] qa \ coprime p qa" using rel_prime by (simp add: member_rec(1) member_rec(2)) have h3: "all_list_constr [[], [0]] (Suc 0)" unfolding all_list_constr_def by (simp add: list_constr_def member_def) then show ?thesis using matrix_equation[where p = "p", where qs = "[q]", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] nonzero h1 h2 h3 by auto qed lemma base_case_satisfy_equation_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "satisfy_equation p qs [[],[0]] [[1],[-1]]" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_satisfy_equation nonzero rel_prime apply (auto) by (simp add: member_rec(1)) qed lemma base_case_matrix_eq: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "(mult_mat_vec (mat_of_rows_list 2 [[1, 1], [1, -1]]) (construct_lhs_vector p [q] [[1],[-1]]) = (construct_rhs_vector p [q] [[],[0]]))" using mat_base_case base_case_satisfy_equation unfolding satisfy_equation_def by (simp add: nonzero rel_prime) lemma less_two: shows "j < Suc (Suc 0) \ j = 0 \ j = 1" by auto lemma inverse_mat_base_case: shows "inverts_mat (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat)" unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma inverse_mat_base_case_2: shows "inverts_mat (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat) (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) " unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma base_case_invertible_mat: shows "invertible_mat (matrix_A [[1], [- 1]] [[], [0]])" unfolding invertible_mat_def using inverse_mat_base_case inverse_mat_base_case_2 apply (auto) apply (simp add: mat_base_case mat_of_rows_list_def) using mat_base_case by auto section "Inductive Step" subsection "Lemmas on smashing subsets and signs" lemma signs_smash_property: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ n m :: nat. elem = ((nth signs1 n)@(nth signs2 m))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\n m. elem = signs1 ! n @ signs2 ! m" using assum unfolding signs_smash_def apply (auto) by (metis in_set_conv_nth) qed lemma signs_smash_property_set: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ (elem1::'a list). \ (elem2::'a list). (elem1 \ set(signs1) \ elem2 \ set(signs2) \ elem = (elem1@elem2))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\elem1. elem1 \ set signs1 \ (\elem2. elem2 \ set signs2 \ elem = elem1 @ elem2)" using assum unfolding signs_smash_def by auto qed lemma subsets_smash_property: fixes subsets1 subsets2 :: "nat list list" fixes n:: "nat" shows "\ (elem :: nat list). (List.member (subsets_smash n subsets1 subsets2) elem) \ (\ (elem1::nat list). \ (elem2::nat list). (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = (elem1 @ ((map ((+) n) elem2)))))" proof - let ?new_subsets = "(map (map ((+) n)) subsets2)" (* a slightly unorthodox use of signs_smash, but it closes the proof *) have "subsets_smash n subsets1 subsets2 = signs_smash subsets1 ?new_subsets" unfolding subsets_smash_def signs_smash_def apply (auto) by (simp add: comp_def) then show ?thesis by (smt imageE in_set_member set_map signs_smash_property_set) qed (* If subsets for smaller systems are well-defined, then subsets for combined systems are well-defined *) subsection "Well-defined subsets preserved when smashing" lemma list_constr_append: "list_constr a n \ list_constr b n \ list_constr (a@b) n" apply (auto) unfolding list_constr_def using list_all_append by blast lemma well_def_step: fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" assumes well_def_subsets1: "all_list_constr (subsets1) (length qs1)" assumes well_def_subsets2: "all_list_constr (subsets2) (length qs2)" shows "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1 @ qs2))" proof - have h1: "\elem. List.member (subsets_smash (length qs1) subsets1 subsets2) elem \ (\elem1 elem2. elem1 \ set subsets1 \ elem2 \ set subsets2 \ elem = elem1 @ map ((+) (length qs1)) elem2)" using subsets_smash_property by blast have h2: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr (elem1 @ map ((+) (length qs1)) elem2) (length (qs1 @ qs2))" proof clarsimp fix elem1 fix elem2 assume e1: "elem1 \ set subsets1" assume e2: "elem2 \ set subsets2" have h1: "list_constr elem1 (length qs1 + length qs2) " proof - have h1: "list_constr elem1 (length qs1)" using e1 well_def_subsets1 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list.pred_mono_strong) qed have h2: "list_constr (map ((+) (length qs1)) elem2) (length qs1 + length qs2)" proof - have h1: "list_constr elem2 (length qs2)" using e2 well_def_subsets2 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list_all_length) qed show "list_constr (elem1 @ map ((+) (length qs1)) elem2) (length qs1 + length qs2)" using h1 h2 list_constr_append by blast qed then show ?thesis using h1 unfolding all_list_constr_def by auto qed subsection "Well def signs preserved when smashing" lemma well_def_signs_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes "length qs1 > 0" assumes "length qs2 > 0" assumes well_def1: "well_def_signs (length qs1) signs1" assumes well_def2: "well_def_signs (length qs2) signs2" shows "well_def_signs (length (qs1@qs2)) (signs_smash signs1 signs2)" using well_def1 well_def2 unfolding well_def_signs_def using signs_smash_property_set[of signs1 signs2] length_append by auto subsection "Distinct signs preserved when smashing" lemma distinct_map_append: assumes "distinct ls" shows "distinct (map ((@) xs) ls)" unfolding distinct_map inj_on_def using assms by auto lemma length_eq_append: assumes "length y = length b" shows "(x @ y = a @ b) \ x=a \ y = b" by (simp add: assms) lemma distinct_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes well_def1: "well_def_signs n1 signs1" assumes well_def2: "well_def_signs n2 signs2" assumes distinct1: "distinct signs1" assumes distinct2: "distinct signs2" shows "distinct (signs_smash signs1 signs2)" unfolding signs_smash_def using distinct1 proof(induction signs1) case Nil then show ?case by auto next case (Cons a signs1) then show ?case apply (auto simp add: distinct2 distinct_map_append) using assms unfolding well_def_signs_def by (simp add: distinct1 distinct2 length_eq_append) qed (* In this section we will show that if signs1 contains all consistent sign assignments and signs2 contains all consistent sign assignments, then their smash contains all consistent sign assignments. Intuitively, this makes sense because signs1 and signs2 are carrying information about different sets of polynomials, and their smash contains all possible combinations of that information. *) subsection "Consistent sign assignments preserved when smashing" lemma subset_smash_signs: fixes a1 b1 a2 b2:: "rat list list" assumes sub1: "set a1 \ set a2" assumes sub2: "set b1 \ set b2" shows "set (signs_smash a1 b1) \ set (signs_smash a2 b2)" proof - have h1: "\x\set (signs_smash a1 b1). x\set (signs_smash a2 b2)" proof clarsimp fix x assume x_in: "x \ set (signs_smash a1 b1)" have h1: "\ n m :: nat. x = (nth a1 n)@(nth b1 m)" using signs_smash_property[of a1 b1] x_in by blast then have h2: "\ p q::nat. x = (nth a2 p)@(nth b2 q)" using sub1 sub2 apply (auto) by (metis nth_find_first signs_smash_property_set subset_code(1) x_in) then show "x \ set (signs_smash a2 b2)" unfolding signs_smash_def apply (auto) using signs_smash_property_set sub1 sub2 x_in by fastforce qed then show ?thesis by blast qed lemma subset_helper: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" shows "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). x = x1@x2" proof clarsimp fix x assume x_in: "x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2))" have x_in_csv: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p))" using x_in unfolding characterize_consistent_signs_at_roots_copr_def by simp have csv_hyp: "\r. consistent_sign_vec_copr (qs1 @ qs2) r = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r)" unfolding consistent_sign_vec_copr_def by auto have exr_iff: "(\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r) \ (\r \ set (characterize_root_list_p p). x = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r))" using csv_hyp by auto have exr: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p)) \ (\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r)" by auto have mem_list1: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs1 r) \ set(map (consistent_sign_vec_copr qs1) (characterize_root_list_p p))" by simp have mem_list2: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs2 r) \ set(map (consistent_sign_vec_copr qs2) (characterize_root_list_p p))" by simp then show "\x1\set (characterize_consistent_signs_at_roots_copr p qs1). \x2\set (characterize_consistent_signs_at_roots_copr p qs2). x = x1 @ x2" using x_in_csv exr mem_list1 mem_list2 characterize_consistent_signs_at_roots_copr_def exr_iff by auto qed lemma subset_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes csa1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set (signs1)" assumes csa2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set (signs2)" shows "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (signs_smash signs1 signs2)" proof - have h0: "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). (x = x1@x2)" using subset_helper[of p qs1 qs2] by blast then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash (characterize_consistent_signs_at_roots_copr p qs1) (characterize_consistent_signs_at_roots_copr p qs2))" unfolding signs_smash_def apply (auto) by fastforce then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash signs1 signs2)" using csa1 csa2 subset_smash_signs[of _ signs1 _ signs2] apply (auto) by blast thus ?thesis by blast qed subsection "Main Results" lemma dim_row_mat_of_rows_list[simp]: shows "dim_row (mat_of_rows_list nr ls) = length ls" unfolding mat_of_rows_list_def by auto lemma dim_col_mat_of_rows_list[simp]: shows "dim_col (mat_of_rows_list nr ls) = nr" unfolding mat_of_rows_list_def by auto lemma dim_row_matrix_A[simp]: shows "dim_row (matrix_A signs subsets) = length subsets" unfolding matrix_A_def by auto lemma dim_col_matrix_A[simp]: shows "dim_col (matrix_A signs subsets) = length signs" unfolding matrix_A_def by auto lemma length_signs_smash: shows "length (signs_smash signs1 signs2) = length signs1 * length signs2" unfolding signs_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_subsets_smash: shows "length (subsets_smash n subs1 subs2) = length subs1 * length subs2" unfolding subsets_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_eq_concat: assumes "\x. x \ set ls \ length x = n" assumes "i < n * length ls" shows "concat ls ! i = ls ! (i div n) ! (i mod n)" using assms proof (induct ls arbitrary: i) case Nil then show ?case by simp next + from assms have \n > 0\ + by (cases \n = 0\) simp_all case (Cons a ls) - then show ?case - apply (auto simp add: nth_append) - using div_if mod_geq by auto + show ?case + proof (cases \n \ i\) + case False + with Cons show ?thesis + by (simp add: nth_append) + next + case True + moreover define j where \j = i - n\ + ultimately have \i = n + j\ + by simp + with Cons \n > 0\ show ?thesis + by (simp add: nth_append div_add_self1) + qed qed lemma z_append: assumes "\i. i \ set xs \ i < length as" shows "z (xs @ (map ((+) (length as)) ys)) (as @ bs) = z xs as * z ys bs" proof - have 1: "map ((!) (as @ bs)) xs = map ((!) as) xs" unfolding list_eq_iff_nth_eq using assms by (auto simp add:nth_append) have 2: "map ((!) (as @ bs) \ (+) (length as)) ys = map ((!) bs) ys" unfolding list_eq_iff_nth_eq using assms by auto show ?thesis unfolding z_def apply auto unfolding 1 2 by auto qed (* Shows that the matrix of a smashed system is the Kronecker product of the matrices of the two systems being combined *) lemma matrix_construction_is_kronecker_product: fixes qs1 :: "real poly list" fixes subs1 subs2 :: "nat list list" fixes signs1 signs2 :: "rat list list" (* n1 is the number of polynomials in the "1" sets *) assumes "\l i. l \ set subs1 \ i \ set l \ i < n1" assumes "\j. j \ set signs1 \ length j = n1" shows " (matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2)) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2)" unfolding mat_eq_iff dim_row_matrix_A dim_col_matrix_A length_subsets_smash length_signs_smash dim_kronecker proof safe fix i j assume i: "i < length subs1 * length subs2" and j: "j < length signs1 * length signs2" have ld: "i div length subs2 < length subs1" "j div length signs2 < length signs1" using i j less_mult_imp_div_less by auto have lm: "i mod length subs2 < length subs2" "j mod length signs2 < length signs2" using i j less_mult_imp_mod_less by auto have n1: "n1 = length (signs1 ! (j div length signs2))" using assms(2) ld(2) nth_mem by blast have 1: "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = z (subsets_smash n1 subs1 subs2 ! i) (signs_smash signs1 signs2 ! j)" unfolding mat_of_rows_list_def matrix_A_def mtx_row_def using i j by (simp add: length_signs_smash length_subsets_smash) have 2: " ... = z (subs1 ! (i div length subs2) @ map ((+) n1) (subs2 ! (i mod length subs2))) (signs1 ! (j div length signs2) @ signs2 ! (j mod length signs2))" unfolding signs_smash_def subsets_smash_def apply (subst length_eq_concat) using i apply (auto simp add: mult.commute) apply (subst length_eq_concat) using j apply (auto simp add: mult.commute) using ld lm by auto have 3: "... = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding n1 apply (subst z_append) apply (auto simp add: n1[symmetric]) using assms(1) ld(1) nth_mem by blast have 4: "kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i,j) = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding kronecker_product_def matrix_A_def mat_of_rows_list_def mtx_row_def using i j apply (auto simp add: Let_def) apply (subst index_mat(1)[OF ld]) apply (subst index_mat(1)[OF lm]) using ld lm by auto show "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i, j)" using 1 2 3 4 by auto qed (* Given that two smaller systems satisfy some mild constraints, show that their combined system (after smashing the signs and subsets) satisfies the matrix equation, and that the matrix of the combined system is invertible. *) lemma inductive_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes welldefined_signs1: "well_def_signs (length qs1) signs1" assumes welldefined_signs2: "well_def_signs (length qs2) signs2" assumes distinct_signs1: "distinct signs1" assumes distinct_signs2: "distinct signs2" assumes all_info1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set(signs1)" assumes all_info2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set(signs2)" assumes welldefined_subsets1: "all_list_constr (subsets1) (length qs1)" assumes welldefined_subsets2: "all_list_constr (subsets2) (length qs2)" assumes invertibleMtx1: "invertible_mat (matrix_A signs1 subsets1)" assumes invertibleMtx2: "invertible_mat (matrix_A signs2 subsets2)" shows "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2) \ invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" proof - have h1: "invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" using matrix_construction_is_kronecker_product kronecker_invertible invertibleMtx1 invertibleMtx2 welldefined_subsets1 welldefined_subsets2 unfolding all_list_constr_def list_constr_def by (smt in_set_conv_nth in_set_member list_all_length well_def_signs_def welldefined_signs1) have h2a: "distinct (signs_smash signs1 signs2)" using distinct_signs1 distinct_signs2 welldefined_signs1 welldefined_signs2 nontriv1 nontriv2 distinct_step by auto have h2ba: "\ q. List.member (qs1 @ qs2) q \ (List.member qs1 q \ List.member qs2 q)" unfolding member_def by auto have h2b: "\q. ((List.member (qs1@qs2) q) \ (coprime p q))" using h2ba pairwise_rel_prime1 pairwise_rel_prime2 by auto have h2c: "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1@qs2))" using well_def_step welldefined_subsets1 welldefined_subsets2 by blast have h2d: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set(signs_smash signs1 signs2)" using subset_step all_info1 all_info2 by simp have h2: "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2)" using matrix_equation[where p="p", where qs="qs1@qs2", where subsets = "subsets_smash (length qs1) subsets1 subsets2", where signs = "signs_smash signs1 signs2"] h2a h2b h2c h2d apply (auto) using nonzero by blast show ?thesis using h1 h2 by blast qed section "Reduction Step Proofs" (* Some definitions *) definition get_matrix:: "(rat mat \ (nat list list \ rat list list)) \ rat mat" where "get_matrix data = fst(data)" definition get_subsets:: "(rat mat \ (nat list list \ rat list list)) \ nat list list" where "get_subsets data = fst(snd(data))" definition get_signs:: "(rat mat \ (nat list list \ rat list list)) \ rat list list" where "get_signs data = snd(snd(data))" definition reduction_signs:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ rat list list" where "reduction_signs p qs signs subsets matr = (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets matr)))" definition reduction_subsets:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ nat list list" where "reduction_subsets p qs signs subsets matr = (take_indices subsets (rows_to_keep (reduce_mat_cols matr (solve_for_lhs p qs subsets matr))))" (* Some basic lemmas *) lemma reduction_signs_is_get_signs: "reduction_signs p qs signs subsets m = get_signs (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_signs_def get_signs_def by (metis reduce_system.simps reduction_step.elims snd_conv) lemma reduction_subsets_is_get_subsets: "reduction_subsets p qs signs subsets m = get_subsets (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_subsets_def get_subsets_def by (metis fst_conv reduce_system.simps reduction_step.elims snd_conv) lemma find_zeros_from_vec_prop: fixes input_vec :: "rat vec" shows "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" proof - have h1: "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" unfolding find_nonzeros_from_input_vec_def List.member_def by auto have h2: "\n < (dim_vec input_vec). (List.member (find_nonzeros_from_input_vec input_vec) n \ (input_vec $ n \ 0) )" unfolding find_nonzeros_from_input_vec_def List.member_def by auto show ?thesis using h1 h2 by auto qed subsection "Showing sign conditions preserved when reducing" lemma take_indices_lem: fixes p:: "real poly" fixes qs :: "real poly list" fixes arb_list :: "'a list list" fixes index_list :: "nat list" fixes n:: "nat" assumes lest: "n < length (take_indices arb_list index_list)" assumes well_def: "\q.(List.member index_list q \ q < length arb_list)" shows "\km (dim_row A)" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis by (metis assms carrier_matI invertible_mat_def mat_inverse(2) matr_option.simps(2) square_mat.elims(2)) qed lemma dim_invertible: fixes A:: "'a::field mat" fixes n:: "nat" assumes assm: "invertible_mat A" assumes dim: "A \ carrier_mat n n" shows "matr_option (dim_row A) (mat_inverse (A)) \ carrier_mat n n" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis using dim mat_inverse(2) by auto qed lemma mult_assoc: fixes A B:: "rat mat" fixes v:: "rat vec" fixes n:: "nat" assumes "A \ carrier_mat n n" assumes "B \ carrier_mat n n" assumes "dim_vec v = n" shows "A *\<^sub>v (mult_mat_vec B v) = (A*B)*\<^sub>v v" using assms(1) assms(2) assms(3) by auto lemma size_of_mat: fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(matrix_A signs subsets) \ carrier_mat (length subsets) (length signs)" unfolding matrix_A_def carrier_mat_def by auto lemma size_of_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes signs :: "rat list list" shows "dim_vec (construct_lhs_vector p qs signs) = length signs" unfolding construct_lhs_vector_def by simp lemma size_of_rhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" shows "dim_vec (construct_rhs_vector p qs subsets) = length subsets" unfolding construct_rhs_vector_def by simp lemma same_size: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "length subsets = length signs" using invertible_mat unfolding invertible_mat_def using size_of_mat[of signs subsets] size_of_lhs[of p qs signs] size_of_rhs[of p qs subsets] by simp lemma mat_equal_list_lem: fixes A:: "'a::field mat" fixes B:: "'a::field mat" shows "(dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B) \ A = B" proof - assume hyp: "dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B" then have "\i j. i < dim_row A \ j < dim_col A \ B $$ (i, j) = A $$ (i, j)" unfolding mat_to_list_def by auto then show ?thesis using hyp unfolding mat_to_list_def using eq_matI[of A B] by metis qed lemma mat_inverse_same: "mat_inverse_var A = mat_inverse A" unfolding mat_inverse_var_def mat_inverse_def mat_equal_def using mat_equal_list_lem apply (simp) by (smt case_prod_beta index_one_mat(2) index_one_mat(3) mat_equal_list_lem) lemma construct_lhs_matches_solve_for_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" proof - have matrix_equation_hyp: "(mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs) = (construct_rhs_vector p qs subsets))" using match unfolding satisfy_equation_def by blast then have eqn_hyp: " ((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs)) = mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets))" using invertible_mat by (simp add: matrix_equation_hyp) have match_hyp: "length subsets = length signs" using same_size invertible_mat by auto then have dim_hyp1: "matrix_A signs subsets \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible by blast have mult_assoc_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs))" using mult_assoc dim_hyp1 dim_hyp2 size_of_lhs by auto have cancel_helper: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (1\<^sub>m (length signs)) *\<^sub>v (construct_lhs_vector p qs signs)" using invertible_means_mult_id[where A= "matrix_A signs subsets"] dim_hyp1 by (simp add: invertible_mat match_hyp) then have cancel_hyp: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (construct_lhs_vector p qs signs)" apply (auto) by (metis carrier_vec_dim_vec one_mult_mat_vec size_of_lhs) then have mult_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (construct_lhs_vector p qs signs)" using mult_assoc_hyp cancel_hyp by simp then have "(construct_lhs_vector p qs signs) = (mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets)) " using eqn_hyp by simp then show ?thesis unfolding solve_for_lhs_def by (simp add: mat_inverse_same) qed (* If set(A) is a subset of B then for all n, nth c n = 0 means nth a n not in B *) lemma reduction_signs_set_helper_lemma: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n = 0 \ (nth B n) \ A" shows "A \ set (take_indices B (find_nonzeros_from_input_vec C))" proof - have unfold: "\ x. (x \ A \ x \ set (take_indices B (find_nonzeros_from_input_vec C)))" proof - fix x assume in_a: "x \ A" have "x \ set (B)" using sub in_a by blast then have "\ n < length B. nth B n = x" by (simp add: in_set_conv_nth) then have "\ n < length B. (nth B n = x \ (List.member (find_nonzeros_from_input_vec C) n) = True)" using not_in_hyp find_zeros_from_vec_prop[of C] using assms(1) in_a by auto thus "x \ set (take_indices B (find_nonzeros_from_input_vec C))" unfolding take_indices_def using member_def by fastforce qed then show ?thesis by blast qed lemma reduction_signs_set_helper_lemma2: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes dist: "distinct B" assumes eq_len: "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n \ 0 \ (nth B n) \ A" shows "set (take_indices B (find_nonzeros_from_input_vec C)) \ A" proof - have unfold: "\ x. (x \ (set (take_indices B (find_nonzeros_from_input_vec C))) \ x \ A)" proof - fix x assume in_set: "x \ set (take_indices B (find_nonzeros_from_input_vec C))" have h: "\n< dim_vec C. (C $ n \ 0 \ (nth B n) = x)" proof - have h1: "\n < length B.(nth B n) = x" using in_set unfolding take_indices_def find_nonzeros_from_input_vec_def eq_len by auto have h2: "\n< length B. (nth B n = x \ List.member (find_nonzeros_from_input_vec C) n)" proof clarsimp fix n assume leq_hyp: "n < length B" assume x_eq: "x = B ! n" have h: "(B !n) \ set (map ((!) B) (find_nonzeros_from_input_vec C))" using x_eq in_set by (simp add: take_indices_def) then have h2: "List.member (map ((!) B) (find_nonzeros_from_input_vec C)) (B ! n)" using in_set by (meson in_set_member) then have h3: "\k< length B. (List.member (find_nonzeros_from_input_vec C) k \ ((B ! k) = (B ! n)))" by (smt atLeastLessThan_iff eq_len find_nonzeros_from_input_vec_def imageE in_set_member mem_Collect_eq set_filter set_map set_upt) have h4: "\v< length B. ((B ! v) = (B ! n) \ v = n)" using dist apply (auto) using leq_hyp nth_eq_iff_index_eq by blast then show "List.member (find_nonzeros_from_input_vec C) n" using h2 h3 h4 by auto qed then have "\n C $ n \ 0)" using find_zeros_from_vec_prop [of C] by (simp add: eq_len) then show ?thesis using h1 eq_len by auto qed thus "x \ A" using not_in_hyp by blast qed then show ?thesis by blast qed (* Show that dropping columns doesn't affect the consistent sign assignments *) lemma reduction_doesnt_break_things_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix x assume card_asm: "card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = 0" assume zero_asm: "poly p x = 0" have card_hyp: "{xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = {}" using card_eq_0_iff using card_asm nonzero poly_roots_finite by fastforce have "x \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x}" using zero_asm by auto thus "False" using card_hyp by blast qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\List.member (characterize_consistent_signs_at_roots_copr p qs) w))" by (smt (verit, best) characterize_consistent_signs_at_roots_copr_def characterize_root_list_p_def h1 imageE in_set_member mem_Collect_eq nonzero poly_roots_finite set_map set_remdups sorted_list_of_set(1)) then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by blast qed then have "set (characterize_consistent_signs_at_roots_copr p qs) \ set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using all_info reduction_signs_set_helper_lemma[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed lemma reduction_deletes_bad_sign_conds: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) = set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have supset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n \ 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix w assume card_asm: "0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" show "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" by (metis (mono_tags, lifting) Collect_empty_eq card_asm card_eq_0_iff gr_implies_not0) qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (List.member (characterize_consistent_signs_at_roots_copr p qs) w))" proof clarsimp fix w assume card_asm: " 0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" have h0: "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" using card_asm by (simp add: h1) then show "List.member (characterize_consistent_signs_at_roots_copr p qs) w" unfolding characterize_consistent_signs_at_roots_copr_def using in_set_member nonzero poly_roots_finite characterize_root_list_p_def by fastforce qed then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by (metis (no_types, lifting) \solve_for_lhs p qs subsets (matrix_A signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))\ dim_vec_of_list length_map nth_map vec_of_list_index) qed then have "set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (characterize_consistent_signs_at_roots_copr p qs)" using all_info reduction_signs_set_helper_lemma2[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using distinct_signs dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed have subset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms by blast then show ?thesis using supset subset by auto qed theorem reduce_system_sign_conditions: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using reduction_deletes_bad_sign_conds[of p qs signs subsets] apply (auto) apply (simp add: all_info distinct_signs match nonzero reduction_signs_def welldefined_signs1) using nonzero invertible_mat apply (metis snd_conv) by (metis all_info distinct_signs invertible_mat match nonzero reduction_signs_def snd_conv welldefined_signs1) subsection "Showing matrix equation preserved when reducing" lemma rows_to_keep_lem: fixes A:: "('a::field) mat" shows "\ell. ell \ set (rows_to_keep A) \ ell < dim_row A" unfolding rows_to_keep_def apply auto using rref_pivot_positions by (metis carrier_mat_triv gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(3)) lemma reduce_system_matrix_equation_preserved: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs: "well_def_signs (length qs) signs" assumes welldefined_subsets: "all_list_constr (subsets) (length qs)" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfy_equation p qs (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - have poly_type_hyp: "p \ 0" using nonzero by auto have distinct_signs_hyp: "distinct (snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" proof - let ?sym = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" have h1: "\ i < length (take_indices signs ?sym). \j < length(take_indices signs ?sym). i \ j \ nth (take_indices signs ?sym) i \ nth (take_indices signs ?sym) j" using distinct_signs unfolding take_indices_def proof clarsimp fix i fix j assume "distinct signs" assume "i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume "j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume neq_hyp: "i \ j" assume "signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)" have h1: "find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i \ find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j" unfolding find_nonzeros_from_input_vec_def using neq_hyp by (metis \i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ distinct_conv_nth distinct_filter distinct_upt find_nonzeros_from_input_vec_def) then show "False" using distinct_signs proof - have f1: "\p ns n. ((n::nat) \ {n \ set ns. p n}) = (n \ set ns \ n \ Collect p)" by simp then have f2: "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) have "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) then show ?thesis using f2 by (metis \signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)\ atLeastLessThan_iff distinct_conv_nth distinct_signs find_nonzeros_from_input_vec_def h1 set_upt) qed qed then have "distinct (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using distinct_conv_nth by blast then show ?thesis using get_signs_def reduction_signs_def reduction_signs_is_get_signs by auto qed have all_info_hyp: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" using reduce_system_sign_conditions assms unfolding get_signs_def by auto have pairwise_rel_prime_hyp: "\q. ((List.member qs q) \ (coprime p q))" using pairwise_rel_prime by auto have welldefined_hyp: "all_list_constr (fst (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))) (length qs)" using welldefined_subsets rows_to_keep_lem unfolding all_list_constr_def List.member_def list_constr_def list_all_def apply (auto simp add: Let_def take_indices_def take_cols_from_matrix_def) using nth_mem by fastforce then show ?thesis using poly_type_hyp distinct_signs_hyp all_info_hyp pairwise_rel_prime_hyp welldefined_hyp using matrix_equation unfolding get_subsets_def get_signs_def by blast qed (* Show that we are tracking the correct matrix in the algorithm *) subsection "Showing matrix preserved" lemma reduce_system_matrix_signs_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "alt_matrix_A (take_indices signs S) subsets = take_cols_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_col (take_cols_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices signs S)" unfolding take_cols_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices signs S). (col (alt_matrix_A (take_indices signs S) subsets ) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices signs S)" have i_lt: "i < length (map ((!) (cols (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: " vec (length subsets) (\j. z (subsets ! j) (map ((!) signs) S ! i)) = vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i)))" using nth_map by (metis \i < length (take_indices signs S)\ length_map take_indices_def) have dim: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_row (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) have well_d: "S ! i < length (signs)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i = nth (cols (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (cols (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_row (alt_matrix_A signs subsets))" using col_dim unfolding cols_def using nth_map well_d by (simp add: \dim_col (alt_matrix_A signs subsets) = length signs\) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i " by (simp add: take_cols_from_matrix_def take_indices_def) have h2: "col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i = nth (map ((!) (cols (alt_matrix_A signs subsets))) S) i " using dim i_lt asm col_mat_of_cols[where j = "i", where n = "(dim_row (alt_matrix_A signs subsets))", where vs = "(map ((!) (cols (alt_matrix_A signs subsets))) S)"] by blast have h3: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = (col (alt_matrix_A signs subsets) (S !i))" using h1 h2 apply (auto) by (metis alt_matrix_char asm cols_nth dim_col_mat(1) in_set_member length_map mat_of_rows_list_def matrix_A_def nth_map nth_mem take_indices_def well_def_h) have "vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i))) = (col (alt_matrix_A signs subsets) (S !i))" by (metis asm in_set_member length_map nth_mem signs_are_cols take_indices_def well_def_h) then have "vec (length subsets) (\j. z (subsets ! j) (take_indices signs S ! i)) = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using h0 h3 by (simp add: take_indices_def) then show "col (alt_matrix_A (take_indices signs S) subsets) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using asm signs_are_cols[where signs = "(take_indices signs S)", where subsets = "subsets"] by auto qed then show ?thesis unfolding alt_matrix_A_def take_cols_from_matrix_def apply (auto) using h0a mat_col_eqI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 mat_of_cols_def take_cols_from_matrix_def) qed lemma reduce_system_matrix_signs_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "matrix_A (take_indices signs S) subsets = take_cols_from_matrix (matrix_A signs subsets) S" using reduce_system_matrix_signs_helper_aux alt_matrix_char assms by auto lemma reduce_system_matrix_subsets_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" assumes nonzero: "p \ 0" shows "alt_matrix_A signs (take_indices subsets S) = take_rows_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_row (take_rows_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices subsets S)" unfolding take_rows_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices subsets S). (row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices subsets S)" have i_lt: "i < length (map ((!) (rows (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: "row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i = row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i " unfolding take_rows_from_matrix_def take_indices_def by auto have dim: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_col (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) then have lenh: "dim_col (alt_matrix_A signs subsets) \ length (subsets)" using assms by auto have well_d: "S ! i < length (subsets)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i = nth (rows (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (rows (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_col (alt_matrix_A signs subsets))" using col_dim unfolding rows_def using nth_map well_d using lenh by (simp add: alt_matrix_A_def) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: " row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i = (row (alt_matrix_A signs subsets) (S ! i)) " using dim i_lt mat_of_rows_row[where i = "i", where n = "(dim_col (alt_matrix_A signs subsets))", where vs = "(map ((!) (rows (alt_matrix_A signs subsets))) S)"] using alt_matrix_char in_set_member nth_mem well_def_h by fastforce have "row (alt_matrix_A signs (take_indices subsets S) ) i = (row (alt_matrix_A signs subsets) (S ! i))" using subsets_are_rows proof - have f1: "i < length S" by (metis (no_types) asm length_map take_indices_def) then have "List.member S (S ! i)" by (meson in_set_member nth_mem) then show ?thesis using f1 by (simp add: \\subsets signs. \ij. z (subsets ! i) (signs ! j))\ take_indices_def well_def_h) qed then show "(row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" using h0 h1 unfolding take_indices_def by auto qed then show ?thesis unfolding alt_matrix_A_def take_rows_from_matrix_def apply (auto) using eq_rowI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 length_map mat_of_rows_def take_indices_def take_rows_from_matrix_def) qed lemma reduce_system_matrix_subsets_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes nonzero: "p \ 0" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" shows "matrix_A signs (take_indices subsets S) = take_rows_from_matrix (matrix_A signs subsets) S" using assms reduce_system_matrix_subsets_helper_aux alt_matrix_char by auto lemma reduce_system_matrix_match: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes inv: "invertible_mat (matrix_A signs subsets)" shows "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?A = "(matrix_A signs subsets)" let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" let ?red_mtx = "(take_rows_from_matrix (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec) (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec)))" have h1: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (reduction_signs p qs signs subsets (matrix_A signs subsets)) (reduction_subsets p qs signs subsets (matrix_A signs subsets))" using reduction_signs_is_get_signs reduction_subsets_is_get_subsets by auto have h1_var: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)))" using h1 reduction_signs_def reduction_subsets_def by auto have h2: "?red_mtx = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" by simp have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3b_a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (subsets)" using assms h30 size_of_lhs same_size unfolding find_nonzeros_from_input_vec_def apply (auto) by (simp add: find_nonzeros_from_input_vec_def h3a) have h3ba: "length (filter (\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using length_filter_le[where P = "(\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0)", where xs = "[0..i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using h3ba by auto then have h3b: "length subsets \ length (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec))" unfolding take_indices_def find_nonzeros_from_input_vec_def by auto have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" using f4 by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def by (metis (no_types) dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed have h3: "matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec))) = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" using inv h3a h3b h3c reduce_system_matrix_subsets_helper reduce_system_matrix_signs_helper assms by auto show ?thesis using h1 h2 h3 by (metis fst_conv get_matrix_def h1_var reduce_system.simps reduction_step.simps) qed (* gauss_jordan_single_rank is crucial in this section *) subsection "Showing invertibility preserved when reducing" (* Overall: Start with a matrix equation. Input a matrix, subsets, and signs. Drop columns of the matrix based on the 0's on the LHS---so extract a list of 0's. Reduce signs accordingly. Then find a list of rows to delete based on using rank (use the transpose result, pivot positions!), and delete those rows. Reduce subsets accordingly. End with a reduced system! *) lemma well_def_find_zeros_from_lhs_vec: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" proof - fix i fix j assume j_in: " j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) " let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) = (length signs)" using size_of_lhs construct_lhs_matches_solve_for_lhs assms by (metis (full_types)) then have h: "j < (length signs)" using j_in unfolding find_nonzeros_from_input_vec_def by simp then show "j < length (cols (matrix_A signs subsets))" using mat_size by auto qed lemma take_cols_subsets_og_cols: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" proof - have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have "\x. x \ set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ x \ set (cols (matrix_A signs subsets))" proof clarsimp fix x let ?og_list = "(cols (matrix_A signs subsets))" let ?ind_list = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume x_in: "x \ set (take_indices ?og_list ?ind_list)" show "x \ set (cols (matrix_A signs subsets))" using x_in unfolding take_indices_def using in_set_member apply (auto) using in_set_conv_nth well_def by fastforce qed then show ?thesis by blast qed lemma reduction_doesnt_break_things_invertibility_step1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "vec_space.rank (length signs) (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets))) = (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have mat_size_alt: "?og_mat \ carrier_mat (length subsets) (length subsets)" using size_of_mat same_size assms by auto have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff mat_size by auto then have dist_cols: "distinct (cols ?og_mat)" using mat_size vec_space.non_distinct_low_rank[where A = ?og_mat, where n = "length signs"] by auto have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have dist1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def by auto have clear: "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" using assms take_cols_subsets_og_cols by auto then have "distinct (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" unfolding take_indices_def using dist1 dist_cols well_def conjugatable_vec_space.distinct_map_nth[where ls = "cols (matrix_A signs subsets)", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by auto then have unfold_thesis: "vec_space.rank (length signs) (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) (find_nonzeros_from_input_vec ?lhs))) = (length (find_nonzeros_from_input_vec ?lhs))" using clear inv conjugatable_vec_space.rank_invertible_subset_cols[where A= "matrix_A signs subsets", where B = "(take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))"] by (simp add: len_eq mat_size take_indices_def) then show ?thesis apply (simp) unfolding take_cols_from_matrix_def by auto qed lemma rechar_take_cols: "take_cols_var A B = take_cols_from_matrix A B" unfolding take_cols_var_def take_cols_from_matrix_def take_indices_def by auto lemma rows_and_cols_transpose: "rows M = cols M\<^sup>T" using row_transpose by simp lemma take_rows_and_take_cols: "(take_rows_from_matrix M r) = (take_cols_from_matrix M\<^sup>T r)\<^sup>T" unfolding take_rows_from_matrix_def take_cols_from_matrix_def using transpose_carrier_mat rows_and_cols_transpose apply (auto) by (simp add: transpose_mat_of_cols) lemma reduction_doesnt_break_things_invertibility: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "invertible_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?step1_mat = "(reduce_mat_cols ?og_mat ?lhs)" let ?new_mat = "take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)" let ?ind_list = "(find_nonzeros_from_input_vec ?lhs)" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have og_mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_col (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list)) = (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_indices_def) then have "mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list) \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: len_eq mat_of_cols_def) then have step1_mat_size: "?step1_mat \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_cols_from_matrix_def) then have "dim_row ?step1_mat \ dim_col ?step1_mat" by (metis carrier_matD(1) carrier_matD(2) construct_lhs_matches_solve_for_lhs diff_zero find_nonzeros_from_input_vec_def inv length_filter_le length_upt match size_of_lhs) then have gt_eq_assm: "dim_col ?step1_mat\<^sup>T \ dim_row ?step1_mat\<^sup>T" by simp have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] og_mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff og_mat_size by auto have rank_drop_cols: "vec_space.rank (length signs) ?step1_mat = (dim_col ?step1_mat)" using assms reduction_doesnt_break_things_invertibility_step1 step1_mat_size by auto let ?step1_T = "?step1_mat\<^sup>T" have rank_transpose: "vec_space.rank (length signs) ?step1_mat = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using transpose_rank[of ?step1_mat] using step1_mat_size by auto have obv: "?step1_T \ carrier_mat (dim_row ?step1_T) (dim_col ?step1_T)" by auto have should_have_this:"vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using obv gt_eq_assm conjugatable_vec_space.gauss_jordan_single_rank[where A = "?step1_T", where n = "dim_row ?step1_T", where nc = "dim_col ?step1_T"] by (simp add: take_cols_from_matrix_def take_indices_def) then have "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_drop_cols rank_transpose by auto then have with_take_cols_from_matrix: "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols_from_matrix ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_transpose rechar_take_cols conjugatable_vec_space.gjs_and_take_cols_var apply (auto) by (smt conjugatable_vec_space.gjs_and_take_cols_var gt_eq_assm obv rechar_take_cols reduce_mat_cols.simps) have "(take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)) = (take_cols_from_matrix ?step1_T (rows_to_keep ?step1_mat))\<^sup>T" using take_rows_and_take_cols by blast then have rank_new_mat: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_col ?step1_mat" using with_take_cols_from_matrix transpose_rank apply (auto) proof - assume a1: "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))) = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" have f2: "\m. vec_space.rank (dim_row (m::rat mat)) m = vec_space.rank (dim_row m\<^sup>T) m\<^sup>T" by (simp add: transpose_rank) have f3: "dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ by force have "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using a1 by (simp add: take_cols_from_matrix_def) then have "vec_space.rank (dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using f3 f2 by presburger then show "vec_space.rank (dim_col (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))\<^sup>T = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" by (simp add: rows_to_keep_def take_cols_from_matrix_def) qed have "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ (length (find_nonzeros_from_input_vec ?lhs))" using obv length_pivot_positions_dim_row[where A = "(gauss_jordan_single (?step1_mat\<^sup>T))"] by (metis carrier_matD(1) carrier_matD(2) gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(2) step1_mat_size) then have len_lt_eq: "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ dim_col ?step1_mat" using step1_mat_size by simp have len_gt_false: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat) \ False" proof - assume length_lt: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat)" have h: "dim_row ?new_mat < (dim_col ?step1_mat)" by (metis Matrix.transpose_transpose index_transpose_mat(3) length_lt length_map mat_of_cols_carrier(3) take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) then show "False" using rank_new_mat by (metis Matrix.transpose_transpose carrier_matI index_transpose_mat(2) nat_less_le not_less_iff_gr_or_eq transpose_rank vec_space.rank_le_nc) qed then have len_gt_eq: "length (rows_to_keep ?step1_mat) \ (dim_col ?step1_mat)" using not_less by blast have len_match: "length (rows_to_keep ?step1_mat) = (dim_col ?step1_mat)" using len_lt_eq len_gt_eq by (simp add: rows_to_keep_def) have mat_match: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using reduce_system_matrix_match[of p qs signs subsets] assms by blast have f2: "length (get_subsets (take_rows_from_matrix (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) subsets) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) signs) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" by (metis (no_types) \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \length (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) = dim_col (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ length_map reduce_mat_cols.simps reduce_system.simps reduction_step.simps reduction_subsets_def reduction_subsets_is_get_subsets take_cols_from_matrix_def take_indices_def) have f3: "\p ps rss nss m. map ((!) rss) (find_nonzeros_from_input_vec (solve_for_lhs p ps nss m)) = get_signs (reduce_system p (ps, m, nss, rss))" by (metis (no_types) reduction_signs_def reduction_signs_is_get_signs take_indices_def) have square_final_mat: "square_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using mat_match assms size_of_mat same_size apply (auto) using f2 f3 by (metis (no_types, lifting) Matrix.transpose_transpose fst_conv get_matrix_def index_transpose_mat(2) len_match length_map mat_of_cols_carrier(2) mat_of_cols_carrier(3) reduce_mat_cols.simps take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have rank_match: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_row ?new_mat" using len_match rank_new_mat by (simp add: take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have "invertible_mat ?new_mat" using invertible_det og_mat_size using vec_space.det_rank_iff square_final_mat by (metis dim_col_matrix_A dim_row_matrix_A fst_conv get_matrix_def mat_match rank_match reduce_system.simps reduction_step.simps size_of_mat square_mat.elims(2)) then show ?thesis apply (simp) by (metis fst_conv get_matrix_def) qed subsection "Well def signs preserved when reducing" lemma reduction_doesnt_break_length_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes length_init : "\ x\ set(signs). length x = length qs" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "\x \ set(reduction_signs p qs signs subsets (matrix_A signs subsets)). length x = length qs" proof clarsimp fix x assume x_in_set: "x \ set (reduction_signs p qs signs subsets (matrix_A signs subsets))" have "List.member (reduction_signs p qs signs subsets (matrix_A signs subsets)) x" using x_in_set by (simp add: in_set_member) then have take_ind: "List.member (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x" unfolding reduction_signs_def by auto have find_nz_len: "\y. List.member (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) y \ y < length signs" using size_of_lhs sat_eq inv_mat construct_lhs_matches_solve_for_lhs[of p qs subsets signs] unfolding find_nonzeros_from_input_vec_def by (metis atLeastLessThan_iff filter_is_subset in_set_member set_upt subset_code(1)) then have "\ n < length signs. x = signs ! n" using take_ind by (metis in_set_conv_nth reduction_signs_def take_indices_lem x_in_set) then show "length x = length qs" using length_init take_indices_lem using nth_mem by blast qed subsection "Distinct signs preserved when reducing" lemma reduction_signs_are_distinct: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" assumes distinct_init: "distinct signs" shows "distinct (reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have h1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def using distinct_filter using distinct_upt by blast have h2: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length signs)" proof - fix j assume "j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" show "j < length signs" using solve_construct size_of_lhs by (metis \j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ atLeastLessThan_iff filter_is_subset find_nonzeros_from_input_vec_def set_upt subset_iff) qed then show ?thesis unfolding reduction_signs_def unfolding take_indices_def using distinct_init h1 h2 conjugatable_vec_space.distinct_map_nth[where ls = "signs", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by blast qed subsection "Well def subsets preserved when reducing" lemma reduction_doesnt_break_subsets: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes length_init : "all_list_constr subsets (length qs)" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "all_list_constr (reduction_subsets p qs signs subsets (matrix_A signs subsets)) (length qs)" unfolding all_list_constr_def proof clarsimp fix x assume in_red_subsets: "List.member (reduction_subsets p qs signs subsets (matrix_A signs subsets)) x " have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have rows_to_keep_hyp: "\y. y \ set (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ y < length subsets" proof clarsimp fix y assume in_set: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))" have in_set_2: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (construct_lhs_vector p qs signs))))" using in_set solve_construct by simp let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def using pivot_positions using dim_row_matrix_A h3a in_set_member nonzero reduce_system_matrix_signs_helper rows_to_keep_def rows_to_keep_lem apply (auto) by (smt (z3) \M_mat (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets)))) subsets = take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))))\<^sup>T)))\ dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed then show "y < length subsets" using in_set_member apply (auto) using in_set_2 solve_construct by fastforce qed show "list_constr x (length qs)" using in_red_subsets unfolding reduction_subsets_def using take_indices_lem[of _ subsets] rows_to_keep_hyp by (metis all_list_constr_def in_set_conv_nth in_set_member length_init) qed section "Overall Lemmas" lemma combining_to_smash: "combine_systems p (qs1, m1, (sub1, sgn1)) (qs2, m2, (sub2, sgn2)) = smash_systems p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2" by simp lemma getter_functions: "calculate_data p qs = (get_matrix (calculate_data p qs), (get_subsets (calculate_data p qs), get_signs (calculate_data p qs))) " unfolding get_matrix_def get_subsets_def get_signs_def by auto subsection "Key properties preserved" subsubsection "Properties preserved when combining and reducing systems" lemma combining_sys_satisfies_properties_helper: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" fixes matrix1 matrix2:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 subsets1 signs1 matrix1" assumes satisfies_properties_sys2: "satisfies_properties p qs2 subsets2 signs2 matrix2" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_signs (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_matrix (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2)))))))" proof - let ?subsets = "(get_subsets (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?signs = "(get_signs (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?matrix = "(get_matrix (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" have h1: "all_list_constr ?subsets (length (qs1 @ qs2))" using well_def_step[of subsets1 qs1 subsets2 qs2] assms by (simp add: nontriv2 get_subsets_def satisfies_properties_def smash_systems_def) have h2: "well_def_signs (length (qs1 @ qs2)) ?signs" using well_def_signs_step[of qs1 qs2 signs1 signs2] using get_signs_def nontriv1 nontriv2 satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 smash_systems_def by auto have h3: "distinct ?signs" using distinct_step[of _ signs1 _ signs2] assms using combine_systems.simps get_signs_def satisfies_properties_def smash_systems_def snd_conv by auto have h4: "satisfy_equation p (qs1 @ qs2) ?subsets ?signs" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] using get_signs_def get_subsets_def satisfies_properties_def smash_systems_def by auto have h5: " invertible_mat ?matrix" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] by (metis combining_to_smash fst_conv get_matrix_def kronecker_invertible satisfies_properties_def smash_systems_def snd_conv) have h6: "?matrix = matrix_A ?signs ?subsets" unfolding get_matrix_def combine_systems.simps smash_systems_def get_signs_def get_subsets_def apply simp apply (subst matrix_construction_is_kronecker_product[of subsets1 _ signs1 signs2 subsets2]) apply (metis Ball_set all_list_constr_def in_set_member list_constr_def satisfies_properties_def satisfies_properties_sys1) using satisfies_properties_def satisfies_properties_sys1 well_def_signs_def apply blast using satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 by auto have h7: "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (?signs)" using subset_step[of p qs1 signs1 qs2 signs2] assms by (simp add: nonzero get_signs_def satisfies_properties_def smash_systems_def) then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed lemma combining_sys_satisfies_properties: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 (get_subsets (calculate_data p qs1)) (get_signs (calculate_data p qs1)) (get_matrix (calculate_data p qs1))" assumes satisfies_properties_sys2: "satisfies_properties p qs2 (get_subsets (calculate_data p qs2)) (get_signs (calculate_data p qs2)) (get_matrix (calculate_data p qs2))" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using combining_sys_satisfies_properties_helper by (metis getter_functions nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_sys1 satisfies_properties_sys2) lemma reducing_sys_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes matrix:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv: "length qs > 0" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" assumes satisfies_properties_sys: "satisfies_properties p qs subsets signs matrix" shows "satisfies_properties p qs (get_subsets (reduce_system p (qs,matrix,subsets,signs))) (get_signs (reduce_system p (qs,matrix,subsets,signs))) (get_matrix (reduce_system p (qs,matrix,subsets,signs)))" proof - have h1: " all_list_constr (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (length qs)" using reduction_doesnt_break_subsets assms reduction_subsets_is_get_subsets satisfies_properties_def satisfies_properties_sys by auto have h2: "well_def_signs (length qs) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_length_signs[of signs qs p subsets] assms reduction_signs_is_get_signs satisfies_properties_def well_def_signs_def by auto have h3: "distinct (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_signs_are_distinct[of p qs subsets signs] assms reduction_signs_is_get_signs satisfies_properties_def by auto have h4: "satisfy_equation p qs (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_equation_preserved[of p qs signs subsets] assms satisfies_properties_def by auto have h5: "invertible_mat (get_matrix (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_invertibility assms same_size satisfies_properties_def by auto have h6: "get_matrix (reduce_system p (qs, matrix, subsets, signs)) = matrix_A (get_signs (reduce_system p (qs, matrix, subsets, signs))) (get_subsets (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_match[of p qs signs subsets] assms satisfies_properties_def by auto have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms reduction_signs_is_get_signs satisfies_properties_def by auto then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed subsubsection "For length 1 qs" lemma length_1_calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof - have h1: "all_list_constr [[],[0]] (length qs)" using len1 unfolding all_list_constr_def list_constr_def apply (auto) by (metis (full_types) length_Cons less_Suc0 list.size(3) list_all_length list_all_simps(2) member_rec(1) member_rec(2) nth_Cons_0) have h2: "well_def_signs (length qs) [[1],[-1]]" unfolding well_def_signs_def using len1 in_set_member by auto have h3: "distinct ([[1],[-1]]::rat list list)" unfolding distinct_def using in_set_member by auto have h4: "satisfy_equation p qs [[],[0]] [[1],[-1]]" using assms base_case_satisfy_equation_alt[of qs p] by auto have h6: "(mat_of_rows_list 2 [[1,1], [1,-1]]::rat mat) = (matrix_A ([[1],[-1]]) ([[],[0]]) :: rat mat)" using mat_base_case by auto then have h5: "invertible_mat (mat_of_rows_list 2 [[1,1], [1,-1]]:: rat mat)" using base_case_invertible_mat by simp have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set ([[1],[-1]])" using assms base_case_sgas_alt[of qs p] by simp have base_case_hyp: "satisfies_properties p qs [[],[0]] [[1],[-1]] (mat_of_rows_list 2 [[1,1], [1,-1]])" using h1 h2 h3 h4 h5 h6 h7 using satisfies_properties_def by blast then have key_hyp: "satisfies_properties p qs (get_subsets (reduce_system p (qs,base_case_info))) (get_signs (reduce_system p (qs,base_case_info))) (get_matrix (reduce_system p (qs,base_case_info)))" using reducing_sys_satisfies_properties by (metis base_case_info_def len1 nonzero pairwise_rel_prime nonzero zero_less_one_class.zero_less_one) show ?thesis by (simp add: key_hyp len1) qed subsubsection "For arbitrary qs" lemma append_not_distinct_helper: "(List.member l1 m \ List.member l2 m) \ (distinct (l1@l2) = False)" proof - have h1: "List.member l1 m \ (\ n. n < length l1 \ (nth l1 n) = m)" using member_def nth_find_first by (simp add: member_def in_set_conv_nth) have h2: "\n. n < length l1 \ (nth l1 n) = m \ (nth (l1@l2) n) = m" proof clarsimp fix n assume lt: "n < length l1" assume nth_l1: "m = l1 ! n" show "(l1 @ l2) ! n = l1 ! n" proof (induct l2) case Nil then show ?case by simp next case (Cons a l2) then show ?case by (simp add: lt nth_append) qed qed have h3: "List.member l1 m \ (\n. n < length l1 \ (nth (l1@l2) n) = m)" using h1 h2 by auto have h4: "List.member l2 m \ (\ n. (nth l2 n) = m)" by (meson member_def nth_find_first) have h5: "\n. (nth l2 n) = m \ (nth (l1@l2) (n + length l1)) = m" proof clarsimp fix n assume nth_l2: "m = l2 ! n" show "(l1 @ l2) ! (n + length l1) = l2 ! n" proof (induct l2) case Nil then show ?case by (metis add.commute nth_append_length_plus) next case (Cons a l2) then show ?case by (simp add: nth_append) qed qed have h6: "List.member l2 m \ (\n. (nth (l1@l2) (n + length l1)) = m)" using h4 h5 by blast show ?thesis using h3 h6 by (metis distinct_append equalityI insert_disjoint(1) insert_subset member_def order_refl) qed lemma calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof (induction "length qs" arbitrary: qs rule: less_induct) case less have len1_h: "length qs = 1 \ ( p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" using length_1_calculate_data_satisfies_properties by blast let ?len = "length qs" let ?q1 = "take (?len div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop (?len div 2) qs" let ?right = "calculate_data p ?q2" let ?comb = "combine_systems p (?q1,?left) (?q2,?right)" let ?red = "reduce_system p ?comb" have h_q1_len: "length qs > 1 \ (length ?q1 > 0)" by auto have h_q2_len: "length qs > 1 \ (length ?q2 > 0)" by auto have h_q1_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q1 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q1 q) \ (List.member qs q)))" by (meson in_set_member in_set_takeD) then show ?thesis by blast qed have h_q2_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q2 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q2 q) \ (List.member qs q)))" by (meson in_set_dropD in_set_member) then show ?thesis by blast qed have q1_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q1 (get_subsets (calculate_data p ?q1)) (get_signs (calculate_data p ?q1)) (get_matrix (calculate_data p ?q1))" using less.hyps[of ?q1] h_q1_len h_q1_copr by (auto simp add: Let_def) have q2_help: "length qs > 1 \ length (drop (length qs div 2) qs) < length qs" using h_q1_len by auto then have q2_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q2 (get_subsets (calculate_data p ?q2)) (get_signs (calculate_data p ?q2)) (get_matrix (calculate_data p ?q2))" using less.hyps[of ?q2] h_q2_len h_q2_copr by blast have put_tog: "?q1@?q2 = qs" using append_take_drop_id by blast then have comb_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_signs (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_matrix (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))))" using q1_sat_props q2_sat_props combining_sys_satisfies_properties using h_q1_copr h_q1_len h_q2_copr h_q2_len put_tog by metis then have comb_sat: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ?comb)) (get_signs (snd ?comb)) (get_matrix (snd ?comb)))" by blast have red_char: "?red = (reduce_system p (qs,(get_matrix (snd ?comb)),(get_subsets (snd ?comb)),(get_signs (snd ?comb))))" using getter_functions by (smt Pair_inject combining_to_smash get_matrix_def get_signs_def get_subsets_def prod.collapse put_tog smash_systems_def) then have "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p qs (get_subsets ?red) (get_signs ?red) (get_matrix ?red))" using reducing_sys_satisfies_properties comb_sat by presburger have len_gt1: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" by (smt \1 < length qs \ p \ 0 \ 0 < length qs \ (\q. List.member qs q \ coprime p q) \ satisfies_properties p qs (get_subsets (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_signs (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_matrix (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs)))))\ calculate_data.simps neq0_conv not_less) then show ?case using len1_h len_gt1 by (metis One_nat_def Suc_lessI) qed subsection "Some key results on consistent sign assignments" lemma find_consistent_signs_at_roots_len1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - let ?signs = "[[1],[-1]]::rat list list" let ?subsets = "[[],[0]]::nat list list" have mat_help: "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1,1], [1,-1]])" using mat_base_case by auto have well_def_signs: "well_def_signs (length qs) ?signs" unfolding well_def_signs_def using len1 by auto have distinct_signs: "distinct ?signs" unfolding distinct_def by auto have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then have all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(?signs)" using assms base_case_sgas apply (auto) by (meson in_mono insertE insert_absorb insert_not_empty member_rec(1)) have match: "satisfy_equation p qs ?subsets ?signs" using ex_q base_case_satisfy_equation nonzero pairwise_rel_prime apply (auto) by (simp add: member_rec(1)) have invertible_mat: "invertible_mat (matrix_A ?signs ?subsets)" using inverse_mat_base_case inverse_mat_base_case_2 unfolding invertible_mat_def using mat_base_case by auto have h: "set (get_signs (reduce_system p (qs, ((matrix_A ?signs ?subsets), (?subsets, ?signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" using nonzero nonzero well_def_signs distinct_signs all_info match invertible_mat reduce_system_sign_conditions[where p = "p", where qs = "qs", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] by blast then have "set (snd (snd (reduce_system p (qs, ((mat_of_rows_list 2 [[1,1], [1,-1]]), ([[],[0]], [[1],[-1]])))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using mat_help by auto then have "set (snd (snd (reduce_system p (qs, base_case_info)))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding base_case_info_def by auto then show ?thesis using len1 by (simp add: find_consistent_signs_at_roots_thm) qed lemma smaller_sys_are_good: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes "set(find_consistent_signs_at_roots p qs1) = set(characterize_consistent_signs_at_roots_copr p qs1)" assumes "set(find_consistent_signs_at_roots p qs2) = set(characterize_consistent_signs_at_roots_copr p qs2)" shows "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" proof - let ?signs = "(get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " let ?subsets = "(get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " have h0: "satisfies_properties p (qs1@qs2) ?subsets ?signs (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using calculate_data_satisfies_properties combining_sys_satisfies_properties using nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero by simp then have h1: "set(characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set ?signs" unfolding satisfies_properties_def by linarith have h2: "well_def_signs (length (qs1@qs2)) ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h3: "distinct ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h4: "satisfy_equation p (qs1@qs2) ?subsets ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h5: "invertible_mat (matrix_A ?signs ?subsets) " using calculate_data_satisfies_properties using h0 satisfies_properties_def by auto have h: "set (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" using h1 h2 h3 h4 h5 reduction_deletes_bad_sign_conds using nonzero nonzero reduction_signs_def by auto then have h: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" unfolding reduction_signs_def get_signs_def by blast have help_h: "reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets) = (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets))))" unfolding reduction_signs_def by auto have clear_signs: "(signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) = (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt combining_to_smash get_signs_def getter_functions smash_systems_def snd_conv) have clear_subsets: "(subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt Pair_inject combining_to_smash get_subsets_def prod.collapse smash_systems_def) have "well_def_signs (length qs1) (get_signs (calculate_data p qs1))" using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_signs1: "(\j. j \ set (get_signs (calculate_data p qs1)) \ length j = (length qs1))" using well_def_signs_def by blast have "all_list_constr (get_subsets(calculate_data p qs1)) (length qs1) " using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_subsets1: "(\l i. l \ set (get_subsets(calculate_data p qs1)) \ i \ set l \ i < (length qs1))" unfolding all_list_constr_def list_constr_def using in_set_member by (metis in_set_conv_nth list_all_length) have extra_matrix_same: "matrix_A (signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) (subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using well_def_signs1 well_def_subsets1 using matrix_construction_is_kronecker_product using calculate_data_satisfies_properties nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_def by auto then have matrix_same: "matrix_A ?signs ?subsets = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using clear_signs clear_subsets by simp have comb_sys_h: "snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))) = snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs)))))" unfolding get_signs_def get_subsets_def using matrix_same by (smt combining_to_smash get_signs_def get_subsets_def getter_functions prod.collapse prod.inject smash_systems_def) then have extra_h: " snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs))))) = snd(snd(reduction_step (matrix_A ?signs ?subsets) ?signs ?subsets (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) " by simp then have same_h: "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" using comb_sys_h unfolding reduction_signs_def by (metis get_signs_def help_h reduction_signs_is_get_signs) then show ?thesis using h by blast qed lemma find_consistent_signs_at_roots_1: fixes p:: "real poly" fixes qs :: "real poly list" shows "(p \ 0 \ length qs > 0 \ (\q. ((List.member qs q) \ (coprime p q)))) \ set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" proof (induction "length qs" arbitrary: qs rule: less_induct) case less then show ?case proof clarsimp assume ind_hyp: "(\qsa. length qsa < length qs \ qsa \ [] \ (\q. List.member qsa q \ coprime p q) \ set (find_consistent_signs_at_roots p qsa) = set (characterize_consistent_signs_at_roots_copr p qsa))" assume nonzero: "p \ 0" assume nontriv: "qs \ []" assume copr:" \q. List.member qs q \ coprime p q" let ?len = "length qs" let ?q1 = "take ((?len) div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop ((?len) div 2) qs" let ?right = "calculate_data p ?q2" have nontriv_q1: "length qs>1 \ length ?q1 > 0" by auto have qs_more_q1: "length qs>1 \ length qs > length ?q1" by auto have pairwise_rel_prime_q1: "\q. ((List.member ?q1 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (take (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (take (length qs div 2) qs) q\ in_set_member in_set_takeD) qed then show "coprime p q" using copr by blast qed have nontriv_q2: "length qs>1 \length ?q2 > 0" by auto have qs_more_q2: "length qs>1 \ length qs > length ?q2" by auto have pairwise_rel_prime_q2: "\q. ((List.member ?q2 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (drop (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (drop (length qs div 2) qs) q\ in_set_dropD in_set_member) qed then show "coprime p q" using copr by blast qed have key_h: "set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_len1 : "?len = 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" using find_consistent_signs_at_roots_len1[of p qs] copr nonzero nontriv by (simp add: find_consistent_signs_at_roots_thm) have h_len_gt1 : "?len > 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_imp_a: "?len > 1 \ set (snd (snd (reduce_system p (combine_systems p (?q1, ?left) (?q2, ?right))))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h1: "?len > 1 \ set(snd(snd(?left))) = set (characterize_consistent_signs_at_roots_copr p ?q1)" using nontriv_q1 pairwise_rel_prime_q1 ind_hyp[of ?q1] qs_more_q1 by (metis find_consistent_signs_at_roots_thm less_numeral_extra(3) list.size(3)) have h2: "?len > 1 \ set(snd(snd(?right))) = set (characterize_consistent_signs_at_roots_copr p ?q2)" using nontriv_q2 pairwise_rel_prime_q2 ind_hyp[of ?q2] qs_more_q2 by (metis (full_types) find_consistent_signs_at_roots_thm list.size(3) not_less_zero) show ?thesis using nonzero nontriv_q1 pairwise_rel_prime_q1 nontriv_q2 pairwise_rel_prime_q2 h1 h2 smaller_sys_are_good by (metis append_take_drop_id find_consistent_signs_at_roots_thm) qed then have h_imp: "?len > 1 \ set (snd (snd (Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" by auto then show ?thesis by auto qed show ?thesis using h_len1 h_len_gt1 by (meson \qs \ []\ length_0_conv less_one nat_neq_iff) qed then show "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" by (smt One_nat_def calculate_data.simps find_consistent_signs_at_roots_thm length_0_conv nontriv) qed qed lemma find_consistent_signs_at_roots_0: fixes p:: "real poly" assumes "p \ 0" shows "set(find_consistent_signs_at_roots p []) = set(characterize_consistent_signs_at_roots_copr p [])" proof - obtain a b c where abc: "reduce_system p ([1], base_case_info) = (a,b,c)" using prod_cases3 by blast have "find_consistent_signs_at_roots p [1] = c" using abc by (simp add: find_consistent_signs_at_roots_thm) have *: "set (find_consistent_signs_at_roots p [1]) = set (characterize_consistent_signs_at_roots_copr p [1])" apply (subst find_consistent_signs_at_roots_1) using assms apply auto by (simp add: member_rec(1) member_rec(2)) have "set(characterize_consistent_signs_at_roots_copr p []) = drop 1 ` set(characterize_consistent_signs_at_roots_copr p [1])" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def apply simp by (smt drop0 drop_Suc_Cons image_cong image_image) thus ?thesis using abc * apply (auto) apply (simp add: find_consistent_signs_at_roots_thm) by (simp add: find_consistent_signs_at_roots_thm) qed lemma find_consistent_signs_at_roots_copr: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" by (metis assms find_consistent_signs_at_roots_0 find_consistent_signs_at_roots_1 in_set_member length_greater_0_conv) lemma find_consistent_signs_at_roots: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots p qs)" using assms find_consistent_signs_at_roots_copr csa_list_copr_rel by (simp add: in_set_member) (* Prettifying theorem *) theorem find_consistent_signs_at_roots_alt: assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set (find_consistent_signs_at_roots p qs) = consistent_signs_at_roots p qs" using consistent_signs_at_roots_eq assms(1) assms(2) find_consistent_signs_at_roots by auto end \ No newline at end of file diff --git a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy --- a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy @@ -1,3145 +1,3147 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) section \The Berlekamp Algorithm\ theory Berlekamp_Type_Based imports Jordan_Normal_Form.Matrix_Kernel Jordan_Normal_Form.Gauss_Jordan_Elimination Jordan_Normal_Form.Missing_VectorSpace Polynomial_Factorization.Square_Free_Factorization Polynomial_Factorization.Missing_Multiset Finite_Field Chinese_Remainder_Poly Poly_Mod_Finite_Field "HOL-Computational_Algebra.Field_as_Ring" begin hide_const (open) up_ring.coeff up_ring.monom Modules.module subspace Modules.module_hom subsection \Auxiliary lemmas\ context fixes g :: "'b \ 'a :: comm_monoid_mult" begin lemma prod_list_map_filter: "prod_list (map g (filter f xs)) * prod_list (map g (filter (\ x. \ f x) xs)) = prod_list (map g xs)" by (induct xs, auto simp: ac_simps) lemma prod_list_map_partition: assumes "List.partition f xs = (ys, zs)" shows "prod_list (map g xs) = prod_list (map g ys) * prod_list (map g zs)" using assms by (subst prod_list_map_filter[symmetric, of _ f], auto simp: o_def) end lemma coprime_id_is_unit: fixes a::"'b::semiring_gcd" shows "coprime a a \ is_unit a" using dvd_unit_imp_unit by auto lemma dim_vec_of_list[simp]: "dim_vec (vec_of_list x) = length x" by (transfer, auto) lemma length_list_of_vec[simp]: "length (list_of_vec A) = dim_vec A" by (transfer', auto) lemma list_of_vec_vec_of_list[simp]: "list_of_vec (vec_of_list a) = a" proof - { fix aa :: "'a list" have "map (\n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0..n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0.. list_of_vec) {v. dim_vec v = n}" proof (rule comp_inj_on) show "inj_on list_of_vec {v. dim_vec v = n}" by (auto simp add: inj_on_def, transfer, auto simp add: mk_vec_def) show "inj_on Poly (list_of_vec ` {v. dim_vec v = n})" proof (auto simp add: inj_on_def) fix x y::"'c vec" assume "n = dim_vec x" and dim_xy: "dim_vec y = dim_vec x" and Poly_eq: "Poly (list_of_vec x) = Poly (list_of_vec y)" note [simp del] = nth_list_of_vec show "list_of_vec x = list_of_vec y" proof (rule nth_equalityI, auto simp: dim_xy) have length_eq: "length (list_of_vec x ) = length (list_of_vec y)" using dim_xy by (transfer, auto) fix i assume "i < dim_vec x" thus "list_of_vec x ! i = list_of_vec y ! i" using Poly_eq unfolding poly_eq_iff coeff_Poly_eq using dim_xy unfolding nth_default_def by (auto, presburger) qed qed qed corollary inj_Poly_list_of_vec: "inj_on (Poly \ list_of_vec) (carrier_vec n)" using inj_Poly_list_of_vec' unfolding carrier_vec_def . lemma list_of_vec_rw_map: "list_of_vec m = map (\n. m $ n) [0.. []" shows "degree (Poly xs) < length xs" using xs by (induct xs, auto intro: Poly.simps(1)) lemma vec_of_list_list_of_vec[simp]: "vec_of_list (list_of_vec a) = a" by (transfer, auto simp add: mk_vec_def) lemma row_mat_of_rows_list: assumes b: "b < length A" and nc: "\i. i < length A \ length (A ! i) = nc" shows "(row (mat_of_rows_list nc A) b) = vec_of_list (A ! b)" proof (auto simp add: vec_eq_iff) show "dim_col (mat_of_rows_list nc A) = length (A ! b)" unfolding mat_of_rows_list_def using b nc by auto fix i assume i: "i < length (A ! b)" show "row (mat_of_rows_list nc A) b $ i = vec_of_list (A ! b) $ i" using i b nc unfolding mat_of_rows_list_def row_def by (transfer, auto simp add: mk_vec_def mk_mat_def) qed lemma degree_Poly_list_of_vec: assumes n: "x \ carrier_vec n" and n0: "n > 0" shows "degree (Poly (list_of_vec x)) < n" proof - have x_dim: "dim_vec x = n" using n by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] n0 n x_dim) have "degree (Poly (list_of_vec x)) < length (list_of_vec x)" by (rule degree_Poly'[OF l]) also have "... = n" using x_dim by auto finally show ?thesis . qed lemma list_of_vec_nth: assumes i: "i < dim_vec x" shows "list_of_vec x ! i = x $ i" using i by (transfer, auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth': assumes i: "i < dim_vec x" shows "coeff (Poly (list_of_vec x)) i = x $ i" using i by (auto simp add: list_of_vec_nth nth_default_def) lemma list_of_vec_row_nth: assumes x: "x < dim_col A" shows "list_of_vec (row A i) ! x = A $$ (i, x)" using x unfolding row_def by (transfer', auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth: assumes x: "x < dim_col A" shows "coeff (Poly (list_of_vec (row A i))) x = A $$ (i, x)" proof - have "coeff (Poly (list_of_vec (row A i))) x = nth_default 0 (list_of_vec (row A i)) x" unfolding coeff_Poly_eq by simp also have "... = A $$ (i, x)" using x list_of_vec_row_nth unfolding nth_default_def by (auto simp del: nth_list_of_vec) finally show ?thesis . qed lemma inj_on_list_of_vec: "inj_on list_of_vec (carrier_vec n)" unfolding inj_on_def unfolding list_of_vec_rw_map by auto lemma vec_of_list_carrier[simp]: "vec_of_list x \ carrier_vec (length x)" unfolding carrier_vec_def by simp lemma card_carrier_vec: "card (carrier_vec n:: 'b::finite vec set) = CARD('b) ^ n" proof - let ?A = "UNIV::'b set" let ?B = "{xs. set xs \ ?A \ length xs = n}" let ?C = "(carrier_vec n:: 'b::finite vec set)" have "card ?C = card ?B" proof - have "bij_betw (list_of_vec) ?C ?B" proof (unfold bij_betw_def, auto) show "inj_on list_of_vec (carrier_vec n)" by (rule inj_on_list_of_vec) fix x::"'b list" assume n: "n = length x" thus "x \ list_of_vec ` carrier_vec (length x)" unfolding image_def by auto (rule bexI[of _ "vec_of_list x"], auto) qed thus ?thesis using bij_betw_same_card by blast qed also have "... = card ?A ^ n" by (rule card_lists_length_eq, simp) finally show ?thesis . qed lemma finite_carrier_vec[simp]: "finite (carrier_vec n:: 'b::finite vec set)" by (rule card_ge_0_finite, unfold card_carrier_vec, auto) lemma row_echelon_form_dim0_row: assumes "A \ carrier_mat 0 n" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_dim0_col: assumes "A \ carrier_mat n 0" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_one_dim0[simp]: "row_echelon_form (1\<^sub>m 0)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma Poly_list_of_vec_0[simp]: "Poly (list_of_vec (0\<^sub>v 0)) = [:0:]" by (simp add: poly_eq_iff nth_default_def) lemma monic_normalize: assumes "(p :: 'b :: {field,euclidean_ring_gcd} poly) \ 0" shows "monic (normalize p)" by (simp add: assms normalize_poly_old_def) lemma exists_factorization_prod_list: fixes P::"'b::field poly list" assumes "degree (prod_list P) > 0" and "\ u. u \ set P \ degree u > 0 \ monic u" and "square_free (prod_list P)" shows "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" using assms proof (induct P) case Nil thus ?case by auto next case (Cons x P) have sf_P: "square_free (prod_list P)" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons mult.commute square_free_factor) have deg_x: "degree x > 0" using Cons.prems by auto have distinct_P: "distinct P" by (meson Cons.prems(2) Cons.prems(3) distinct.simps(2) square_free_prod_list_distinct) have "\A. finite A \ x = \A \ A \ {q. irreducible q \ monic q}" proof (rule monic_square_free_irreducible_factorization) show "monic x" by (simp add: Cons.prems(2)) show "square_free x" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons square_free_factor) qed from this obtain A where fin_A: "finite A" and xA: "x = \A" and A: "A \ {q. irreducible\<^sub>d q \ monic q}" by auto obtain A' where s: "set A' = A" and length_A': "length A' = card A" using \finite A\ distinct_card finite_distinct_list by force have A_not_empty: "A \ {}" using xA deg_x by auto have x_prod_list_A': "x = prod_list A'" proof - have "x = \A" using xA by simp also have "... = prod id A" by simp also have "... = prod id (set A')" unfolding s by simp also have "... = prod_list (map id A')" by (rule prod.distinct_set_conv_list, simp add: card_distinct length_A' s) also have "... = prod_list A'" by auto finally show ?thesis . qed show ?case proof (cases "P = []") case True show ?thesis proof (rule exI[of _ "A'"], auto simp add: True) show "prod_list A' = x" using x_prod_list_A' by simp show "Suc 0 \ length A'" using A_not_empty using s length_A' by (simp add: Suc_leI card_gt_0_iff fin_A) show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed next case False have hyp: "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule Cons.hyps[OF _ _ sf_P]) have set_P: "set P \ {}" using False by auto have "prod_list P = prod_list (map id P)" by simp also have "... = prod id (set P)" using prod.distinct_set_conv_list[OF distinct_P, of id] by simp also have "... = \(set P)" by simp finally have "prod_list P = \(set P)" . hence "degree (prod_list P) = degree (\(set P))" by simp also have "... = degree (prod id (set P))" by simp also have "... = (\i\(set P). degree (id i))" proof (rule degree_prod_eq_sum_degree) show "\i\set P. id i \ 0" using Cons.prems(2) by force qed also have "... > 0" by (metis Cons.prems(2) List.finite_set set_P gr0I id_apply insert_iff list.set(2) sum_pos) finally show "degree (prod_list P) > 0" by simp show "\u. u \ set P \ degree u > 0 \ monic u" using Cons.prems by auto qed from this obtain Q where QP: "prod_list Q = prod_list P" and length_PQ: "length P \ length Q" and monic_irr_Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast show ?thesis proof (rule exI[of _ "A' @ Q"], auto simp add: monic_irr_Q) show "prod_list A' * prod_list Q = x * prod_list P" unfolding QP x_prod_list_A' by auto have "length A' \ 0" using A_not_empty using s length_A' by auto thus "Suc (length P) \ length A' + length Q" using QP length_PQ by linarith show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed qed qed lemma normalize_eq_imp_smult: fixes p :: "'b :: {euclidean_ring_gcd} poly" assumes n: "normalize p = normalize q" shows "\ c. c \ 0 \ q = smult c p" proof(cases "p = 0") case True with n show ?thesis by (auto intro:exI[of _ 1]) next case p0: False have degree_eq: "degree p = degree q" using n degree_normalize by metis hence q0: "q \ 0" using p0 n by auto have p_dvd_q: "p dvd q" using n by (simp add: associatedD1) from p_dvd_q obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with q0 have "k \ 0" by auto then have "degree k = 0" using degree_eq degree_mult_eq p0 q by fastforce then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed lemma prod_list_normalize: fixes P :: "'b :: {idom_divide,normalization_semidom_multiplicative} poly list" shows "normalize (prod_list P) = prod_list (map normalize P)" proof (induct P) case Nil show ?case by auto next case (Cons p P) have "normalize (prod_list (p # P)) = normalize p * normalize (prod_list P)" using normalize_mult by auto also have "... = normalize p * prod_list (map normalize P)" using Cons.hyps by auto also have "... = prod_list (normalize p # (map normalize P))" by auto also have "... = prod_list (map normalize (p # P))" by auto finally show ?case . qed lemma prod_list_dvd_prod_list_subset: fixes A::"'b::comm_monoid_mult list" assumes dA: "distinct A" and dB: "distinct B" (*Maybe this condition could be avoided*) and s: "set A \ set B" shows "prod_list A dvd prod_list B" proof - have "prod_list A = prod_list (map id A)" by auto also have "... = prod id (set A)" by (rule prod.distinct_set_conv_list[symmetric, OF dA]) also have "... dvd prod id (set B)" by (rule prod_dvd_prod_subset[OF _ s], auto) also have "... = prod_list (map id B)" by (rule prod.distinct_set_conv_list[OF dB]) also have "... = prod_list B" by simp finally show ?thesis . qed end lemma gcd_monic_constant: "gcd f g \ {1, f}" if "monic f" and "degree g = 0" for f g :: "'a :: {field_gcd} poly" proof (cases "g = 0") case True moreover from \monic f\ have "normalize f = f" by (rule normalize_monic) ultimately show ?thesis by simp next case False with \degree g = 0\ have "is_unit g" by simp then have "Rings.coprime f g" by (rule is_unit_right_imp_coprime) then show ?thesis by simp qed lemma distinct_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "distinct (find_base_vectors A)" proof - note non_pivot_base = non_pivot_base[OF ref A] let ?pp = "set (pivot_positions A)" from A have dim: "dim_row A = nr" "dim_col A = nc" by auto { fix j j' assume j: "j < nc" "j \ snd ` ?pp" and j': "j' < nc" "j' \ snd ` ?pp" and neq: "j' \ j" from non_pivot_base(2)[OF j] non_pivot_base(4)[OF j' j neq] have "non_pivot_base A (pivot_positions A) j \ non_pivot_base A (pivot_positions A) j'" by auto } hence inj: "inj_on (non_pivot_base A (pivot_positions A)) (set [j\[0.. snd ` ?pp])" unfolding inj_on_def by auto thus ?thesis unfolding find_base_vectors_def Let_def unfolding distinct_map dim by auto qed lemma length_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "length (find_base_vectors A) = card (set (find_base_vectors A))" using distinct_card[OF distinct_find_base_vectors[OF ref A]] by auto subsection \Previous Results\ definition power_poly_f_mod :: "'a::field poly \ 'a poly \ nat \ 'a poly" where "power_poly_f_mod modulus = (\a n. a ^ n mod modulus)" lemma power_poly_f_mod_binary: "power_poly_f_mod m a n = (if n = 0 then 1 mod m else let (d, r) = Euclidean_Rings.divmod_nat n 2; rec = power_poly_f_mod m ((a * a) mod m) d in if r = 0 then rec else (rec * a) mod m)" for m a :: "'a :: {field_gcd} poly" proof - note d = power_poly_f_mod_def show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding d by simp next case False obtain q r where div: "Euclidean_Rings.divmod_nat n 2 = (q,r)" by force hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding Euclidean_Rings.divmod_nat_def by auto have id: "a ^ (2 * q) = (a * a) ^ q" by (simp add: power_mult_distrib semiring_normalization_rules) show ?thesis proof (cases "r = 0") case True show ?thesis using power_mod [of "a * a" m q] by (auto simp add: Euclidean_Rings.divmod_nat_def Let_def True n d div id) next case False with r have r: "r = 1" by simp show ?thesis by (auto simp add: d r div Let_def mod_simps) (simp add: n r mod_simps ac_simps power_mult_distrib power_mult power2_eq_square) qed qed qed fun power_polys where "power_polys mul_p u curr_p (Suc i) = curr_p # power_polys mul_p u ((curr_p * mul_p) mod u) i" | "power_polys mul_p u curr_p 0 = []" context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma fermat_theorem_mod_ring [simp]: fixes a::"'a mod_ring" shows "a ^ CARD('a) = a" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis proof transfer fix a assume "a \ {0.. 0" then have a: "1 \ a" "a < int CARD('a)" by simp_all then have [simp]: "a mod int CARD('a) = a" by simp from a have "\ int CARD('a) dvd a" by (auto simp add: zdvd_not_zless) then have "\ CARD('a) dvd nat \a\" by simp with a have "\ CARD('a) dvd nat a" by simp with prime_card have "[nat a ^ (CARD('a) - 1) = 1] (mod CARD('a))" by (rule fermat_theorem) with a have "int (nat a ^ (CARD('a) - 1) mod CARD('a)) = 1" by (simp add: cong_def) with a have "a ^ (CARD('a) - 1) mod CARD('a) = 1" by (simp add: of_nat_mod) then have "a * (a ^ (CARD('a) - 1) mod int CARD('a)) = a" by simp then have "(a * (a ^ (CARD('a) - 1) mod int CARD('a))) mod int CARD('a) = a mod int CARD('a)" by (simp only:) then show "a ^ CARD('a) mod int CARD('a) = a" by (simp add: mod_simps semiring_normalization_rules(27)) qed qed lemma mod_eq_dvd_iff_poly: "((x::'a mod_ring poly) mod n = y mod n) = (n dvd x - y)" proof assume H: "x mod n = y mod n" hence "x mod n - y mod n = 0" by simp hence "(x mod n - y mod n) mod n = 0" by simp hence "(x - y) mod n = 0" by (simp add: mod_diff_eq) thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) next assume H: "n dvd x - y" then obtain k where k: "x-y = n*k" unfolding dvd_def by blast hence "x = n*k + y" using diff_eq_eq by blast hence "x mod n = (n*k + y) mod n" by simp thus "x mod n = y mod n" by (simp add: mod_add_left_eq) qed lemma cong_gcd_eq_poly: "gcd a m = gcd b m" if "[(a::'a mod_ring poly) = b] (mod m)" using that by (simp add: cong_def) (metis gcd_mod_left mod_by_0) lemma coprime_h_c_poly: fixes h::"'a mod_ring poly" assumes "c1 \ c2" shows "coprime (h - [:c1:]) (h - [:c2:])" proof (intro coprimeI) fix d assume "d dvd h - [:c1:]" and "d dvd h - [:c2:]" hence "h mod d = [:c1:] mod d" and "h mod d = [:c2:] mod d" using mod_eq_dvd_iff_poly by simp+ hence "[:c1:] mod d = [:c2:] mod d" by simp hence "d dvd [:c2 - c1:]" by (metis (no_types) mod_eq_dvd_iff_poly diff_pCons right_minus_eq) thus "is_unit d" by (metis (no_types) assms dvd_trans is_unit_monom_0 monom_0 right_minus_eq) qed lemma coprime_h_c_poly2: fixes h::"'a mod_ring poly" assumes "coprime (h - [:c1:]) (h - [:c2:])" and "\ is_unit (h - [:c1:])" shows "c1 \ c2" using assms coprime_id_is_unit by blast lemma degree_minus_eq_right: fixes p::"'b::ab_group_add poly" shows "degree q < degree p \ degree (p - q) = degree p" using degree_add_eq_left[of "-q" p] degree_minus by auto lemma coprime_prod: fixes A::"'a mod_ring set" and g::"'a mod_ring \ 'a mod_ring poly" assumes "\x\A. coprime (g a) (g x)" shows "coprime (g a) (prod (\x. g x) A)" proof - have f: "finite A" by simp show ?thesis using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = (g x) * (\c\A. g c)" by (simp add: insert.hyps(2)) with insert.prems show ?case by (auto simp: insert.hyps(3) prod_coprime_right) qed auto qed lemma coprime_prod2: fixes A::"'b::semiring_gcd set" assumes "\x\A. coprime (a) (x)" and f: "finite A" shows "coprime (a) (prod (\x. x) A)" using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. c) = (x) * (\c\A. c)" by (simp add: insert.hyps) with insert.prems show ?case by (simp add: insert.hyps prod_coprime_right) qed auto lemma divides_prod: fixes g::"'a mod_ring \ 'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2 \ coprime (g c1) (g c2)" assumes "\c\ A. g c dvd f" shows "(\c\A. g c) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = g x * (\c\ A. g c)" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "g x dvd f" using insert.prems by auto show "prod g A dvd f" using insert.hyps(3) insert.prems by auto from insert show "Rings.coprime (g x) (prod g A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto qed (* Proof of equation 9 in the book by Knuth x^p - x = (x-0)(x-1)...(x-(p-1)) (mod p) *) lemma poly_monom_identity_mod_p: "monom (1::'a mod_ring) (CARD('a)) - monom 1 1 = prod (\x. [:0,1:] - [:x:]) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof - let ?f="(\x::'a mod_ring. [:0,1:] - [:x:])" have "?rhs dvd ?lhs" proof (rule divides_prod) { fix a::"'a mod_ring" have "poly ?lhs a = 0" by (simp add: poly_monom) hence "([:0,1:] - [:a:]) dvd ?lhs" using poly_eq_0_iff_dvd by fastforce } thus "\x\UNIV::'a mod_ring set. [:0, 1:] - [:x:] dvd monom 1 CARD('a) - monom 1 1" by fast show "\c1 c2. c1 \ UNIV \ c2 \ UNIV \ c1 \ (c2 :: 'a mod_ring) \ coprime ([:0, 1:] - [:c1:]) ([:0, 1:] - [:c2:])" by (auto dest!: coprime_h_c_poly[of _ _ "[:0,1:]"]) qed from this obtain g where g: "?lhs = ?rhs * g" using dvdE by blast have degree_lhs_card: "degree ?lhs = CARD('a)" proof - have "degree (monom (1::'a mod_ring) 1) = 1" by (simp add: degree_monom_eq) moreover have d_c: "degree (monom (1::'a mod_ring) CARD('a)) = CARD('a)" by (simp add: degree_monom_eq) ultimately have "degree (monom (1::'a mod_ring) 1) < degree (monom (1::'a mod_ring) CARD('a))" using prime_card unfolding prime_nat_iff by auto hence "degree ?lhs = degree (monom (1::'a mod_ring) CARD('a))" by (rule degree_minus_eq_right) thus ?thesis unfolding d_c . qed have degree_rhs_card: "degree ?rhs = CARD('a)" proof - have "degree (prod ?f UNIV) = sum (degree \ ?f) UNIV \ coeff (prod ?f UNIV) (sum (degree \ ?f) UNIV) = 1" by (rule degree_prod_sum_monic, auto) moreover have "sum (degree \ ?f) UNIV = CARD('a)" by auto ultimately show ?thesis by presburger qed have monic_lhs: "monic ?lhs" using degree_lhs_card by auto have monic_rhs: "monic ?rhs" by (rule monic_prod, simp) have degree_eq: "degree ?rhs = degree ?lhs" unfolding degree_lhs_card degree_rhs_card .. have g_not_0: "g \ 0" using g monic_lhs by auto have degree_g0: "degree g = 0" proof - have "degree (?rhs * g) = degree ?rhs + degree g" by (rule degree_monic_mult[OF monic_rhs g_not_0]) thus ?thesis using degree_eq g by simp qed have monic_g: "monic g" using monic_factor g monic_lhs monic_rhs by auto have "g = 1" using monic_degree_0[OF monic_g] degree_g0 by simp thus ?thesis using g by auto qed (* Proof of equation 10 in the book by Knuth v(x)^p - v(x) = (v(x)-0)(v(x)-1)...(v(x)-(p-1)) (mod p) *) lemma poly_identity_mod_p: "v^(CARD('a)) - v = prod (\x. v - [:x:]) (UNIV::'a mod_ring set)" proof - have id: "monom 1 1 \\<^sub>p v = v" "[:0, 1:] \\<^sub>p v = v" unfolding pcompose_def apply (auto) by (simp add: fold_coeffs_def) have id2: "monom 1 (CARD('a)) \\<^sub>p v = v ^ (CARD('a))" by (metis id(1) pcompose_hom.hom_power x_pow_n) show ?thesis using arg_cong[OF poly_monom_identity_mod_p, of "\ f. f \\<^sub>p v"] unfolding pcompose_hom.hom_minus pcompose_hom.hom_prod id pcompose_const id2 . qed lemma coprime_gcd: fixes h::"'a mod_ring poly" assumes "Rings.coprime (h-[:c1:]) (h-[:c2:])" shows "Rings.coprime (gcd f(h-[:c1:])) (gcd f (h-[:c2:]))" using assms coprime_divisors by blast lemma divides_prod_gcd: fixes h::"'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2\ coprime (h-[:c1:]) (h-[:c2:])" shows "(\c\A. gcd f (h - [:c:])) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "gcd f (h - [:x:]) dvd f" by simp show "(\c\A. gcd f (h - [:c:])) dvd f" using insert.hyps(3) insert.prems by auto show "Rings.coprime (gcd f (h - [:x:])) (\c\A. gcd f (h - [:c:]))" by (rule prod_coprime_right) (metis Berlekamp_Type_Based.coprime_h_c_poly coprime_gcd coprime_iff_coprime insert.hyps(2)) qed finally show ?case . qed auto qed lemma monic_prod_gcd: assumes f: "finite A" and f0: "(f :: 'b :: {field_gcd} poly) \ 0" shows "monic (\c\A. gcd f (h - [:c:]))" using f proof (induct A) case (insert x A) have rw: "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps) show ?case proof (unfold rw, rule monic_mult) show "monic (gcd f (h - [:x:]))" using poly_gcd_monic[of f] f0 using insert.prems insert_iff by blast show "monic (\c\A. gcd f (h - [:c:]))" using insert.hyps(3) insert.prems by blast qed qed auto lemma coprime_not_unit_not_dvd: fixes a::"'b::semiring_gcd" assumes "a dvd b" and "coprime b c" and "\ is_unit a" shows "\ a dvd c" using assms coprime_divisors coprime_id_is_unit by fastforce lemma divides_prod2: fixes A::"'b::semiring_gcd set" assumes f: "finite A" and "\a\A. a dvd c" and "\a1 a2. a1 \ A \ a2 \ A \ a1 \ a2 \ coprime a1 a2" shows "\A dvd c" using assms proof (induct A) case (insert x A) have "\(insert x A) = x * \A" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... dvd c" proof (rule divides_mult) show "x dvd c" by (simp add: insert.prems) show "\A dvd c" using insert by auto from insert show "Rings.coprime x (\A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto lemma coprime_polynomial_factorization: fixes a1 :: "'b :: {field_gcd} poly" assumes irr: "as \ {q. irreducible q \ monic q}" and "finite as" and a1: "a1 \ as" and a2: "a2 \ as" and a1_not_a2: "a1 \ a2" shows "coprime a1 a2" proof (rule ccontr) assume not_coprime: "\ coprime a1 a2" let ?b= "gcd a1 a2" have b_dvd_a1: "?b dvd a1" and b_dvd_a2: "?b dvd a2" by simp+ have irr_a1: "irreducible a1" using a1 irr by blast have irr_a2: "irreducible a2" using a2 irr by blast have a2_not0: "a2 \ 0" using a2 irr by auto have degree_a1: "degree a1 \ 0" using irr_a1 by auto have degree_a2: "degree a2 \ 0" using irr_a2 by auto have not_a2_dvd_a1: "\ a2 dvd a1" proof (rule ccontr, simp) assume a2_dvd_a1: "a2 dvd a1" from this obtain k where k: "a1 = a2 * k" unfolding dvd_def by auto have k_not0: "k \ 0" using degree_a1 k by auto show False proof (cases "degree a2 = degree a1") case False have "degree a2 < degree a1" using False a2_dvd_a1 degree_a1 divides_degree by fastforce hence "\ irreducible a1" using degree_a2 a2_dvd_a1 degree_a2 by (metis degree_a1 irreducible\<^sub>dD(2) irreducible\<^sub>d_multD irreducible_connect_field k neq0_conv) thus False using irr_a1 by contradiction next case True have "degree a1 = degree a2 + degree k" unfolding k using degree_mult_eq[OF a2_not0 k_not0] by simp hence "degree k = 0" using True by simp hence "k = 1" using monic_factor a1 a2 irr k monic_degree_0 by auto hence "a1 = a2" using k by simp thus False using a1_not_a2 by contradiction qed qed have b_not0: "?b \ 0" by (simp add: a2_not0) have degree_b: "degree ?b > 0" using not_coprime[simplified] b_not0 is_unit_gcd is_unit_iff_degree by blast have "degree ?b < degree a2" by (meson b_dvd_a1 b_dvd_a2 irreducibleD' dvd_trans gcd_dvd_1 irr_a2 not_a2_dvd_a1 not_coprime) hence "\ irreducible\<^sub>d a2" using degree_a2 b_dvd_a2 degree_b by (metis degree_smult_eq irreducible\<^sub>d_dvd_smult less_not_refl3) thus False using irr_a2 by auto qed (* Proof of equation 14 in the book by Knuth *) theorem Berlekamp_gcd_step: fixes f::"'a mod_ring poly" and h::"'a mod_ring poly" assumes hq_mod_f: "[h^(CARD('a)) = h] (mod f)" and monic_f: "monic f" and sf_f: "square_free f" shows "f = prod (\c. gcd f (h - [:c:])) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof (cases "f=0") case True thus ?thesis using coeff_0 monic_f zero_neq_one by auto next case False note f_not_0 = False show ?thesis proof (rule poly_dvd_antisym) show "?rhs dvd f" using coprime_h_c_poly by (intro divides_prod_gcd, auto) have "monic ?rhs" by (rule monic_prod_gcd[OF _ f_not_0], simp) thus "coeff f (degree f) = coeff ?rhs (degree ?rhs)" using monic_f by auto next show "f dvd ?rhs" proof - let ?p = "CARD('a)" obtain P where finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" using monic_square_free_irreducible_factorization[OF monic_f sf_f] by auto have f_dvd_hqh: "f dvd (h^?p - h)" using hq_mod_f unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "f = \P" using f_desc_square_free by simp also have "... dvd ?rhs" proof (rule divides_prod2[OF finite_P]) show "\a1 a2. a1 \ P \ a2 \ P \ a1 \ a2 \ coprime a1 a2" using coprime_polynomial_factorization[OF P finite_P] by simp show "\a\P. a dvd (\c\UNIV. gcd f (h - [:c:]))" proof fix fi assume fi_P: "fi \ P" show "fi dvd ?rhs" proof (rule dvd_prod, auto) show "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible (fi)" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) have fi_dvd_hc: "\c\UNIV::'a mod_ring set. fi dvd (h-[:c:])" by (rule irreducible_dvd_prod[OF _ fi_dvd_prod_hc], simp add: irr_fi) thus "\c. fi dvd h - [:c:]" by simp qed qed qed finally show "f dvd ?rhs" . qed qed qed (******* Implementation of Berlekamp's algorithm (type-based version) *******) subsection \Definitions\ definition berlekamp_mat :: "'a mod_ring poly \ 'a mod_ring mat" where "berlekamp_mat u = (let n = degree u; mul_p = power_poly_f_mod u [:0,1:] (CARD('a)); xks = power_polys mul_p u 1 n in mat_of_rows_list n (map (\ cs. let coeffs_cs = (coeffs cs); k = n - length (coeffs cs) in (coeffs cs) @ replicate k 0) xks))" definition berlekamp_resulting_mat :: "('a mod_ring) poly \ 'a mod_ring mat" where "berlekamp_resulting_mat u = (let Q = berlekamp_mat u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then Q $$ (i,j) - 1 else Q $$ (i,j)) in (gauss_jordan_single (transpose_mat QI)))" definition berlekamp_basis :: "'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_basis u = (map (Poly o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" lemma berlekamp_basis_code[code]: "berlekamp_basis u = (map (poly_of_list o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" unfolding berlekamp_basis_def poly_of_list_def .. primrec berlekamp_factorization_main :: "nat \ 'a mod_ring poly list \ 'a mod_ring poly list \ nat \ 'a mod_ring poly list" where "berlekamp_factorization_main i divs (v # vs) n = (if v = 1 then berlekamp_factorization_main i divs vs n else if length divs = n then divs else let facts = [ w . u \ divs, s \ [0 ..< CARD('a)], w \ [gcd u (v - [:of_int s:])], w \ 1]; (lin,nonlin) = List.partition (\ q. degree q = i) facts in lin @ berlekamp_factorization_main i nonlin vs (n - length lin))" | "berlekamp_factorization_main i divs [] n = divs" definition berlekamp_monic_factorization :: "nat \ 'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_monic_factorization d f = (let vs = berlekamp_basis f; n = length vs; fs = berlekamp_factorization_main d [f] vs n in fs)" subsection \Properties\ lemma power_polys_works: fixes u::"'b::unique_euclidean_semiring" assumes i: "i < n" and c: "curr_p = curr_p mod u" (*Equivalent to degree curr_p < degree u*) shows "power_polys mult_p u curr_p n ! i = curr_p * mult_p ^ i mod u" using i c proof (induct n arbitrary: curr_p i) case 0 thus ?case by simp next case (Suc n) have p_rw: "power_polys mult_p u curr_p (Suc n) ! i = (curr_p # power_polys mult_p u (curr_p * mult_p mod u) n) ! i" by simp show ?case proof (cases "i=0") case True show ?thesis using Suc.prems unfolding p_rw True by auto next case False note i_not_0 = False show ?thesis proof (cases "i < n") case True note i_less_n = True have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (i - 1)" unfolding p_rw using nth_Cons_pos False by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (i-1) mod u" by (rule Suc.hyps) (auto simp add: i_less_n less_imp_diff_less) also have "... = curr_p * mult_p ^ i mod u" using False by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . next case False hence i_n: "i = n" using Suc.prems by auto have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (n - 1)" unfolding p_rw using nth_Cons_pos i_n i_not_0 by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (n-1) mod u" proof (rule Suc.hyps) show "n - 1 < n" using i_n i_not_0 by linarith show "curr_p * mult_p mod u = curr_p * mult_p mod u mod u" by simp qed also have "... = curr_p * mult_p ^ i mod u" using i_n [symmetric] i_not_0 by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . qed qed qed lemma length_power_polys[simp]: "length (power_polys mult_p u curr_p n) = n" by (induct n arbitrary: curr_p, auto) (* Equation 12 *) lemma Poly_berlekamp_mat: assumes k: "k < degree u" shows "Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k) mod u" proof - let ?map ="(map (\cs. coeffs cs @ replicate (degree u - length (coeffs cs)) 0) (power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)))" have "row (berlekamp_mat u) k = row (mat_of_rows_list (degree u) ?map) k" by (simp add: berlekamp_mat_def Let_def) also have "... = vec_of_list (?map ! k)" proof- { fix i assume i: "i < degree u" + then have \u \ 0\ + by auto let ?c= "power_polys (power_poly_f_mod u [:0, 1:] CARD('a)) u 1 (degree u) ! i" let ?coeffs_c="(coeffs ?c)" have "?c = 1*([:0, 1:] ^ CARD('a) mod u)^i mod u" proof (unfold power_poly_f_mod_def, rule power_polys_works[OF i]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = [:0, 1:] ^ (CARD('a) * i) mod u" by (simp add: power_mod power_mult) finally have c_rw: "?c = [:0, 1:] ^ (CARD('a) * i) mod u" . have "length ?coeffs_c \ degree u" proof - show ?thesis proof (cases "?c = 0") case True thus ?thesis by auto next case False have "length ?coeffs_c = degree (?c) + 1" by (rule length_coeffs[OF False]) also have "... = degree ([:0, 1:] ^ (CARD('a) * i) mod u) + 1" using c_rw by simp also have "... \ degree u" - by (metis One_nat_def add.right_neutral add_Suc_right c_rw calculation coeffs_def degree_0 - degree_mod_less discrete gr_implies_not0 k list.size(3) one_neq_zero) + using \i < degree u\ \u \ 0\ degree_mod_less [of u \pCons 0 1 ^ (CARD('a) * i)\] + by auto finally show ?thesis . qed qed then have "length ?coeffs_c + (degree u - length ?coeffs_c) = degree u" by auto } with k show ?thesis by (intro row_mat_of_rows_list, auto) qed finally have row_rw: "row (berlekamp_mat u) k = vec_of_list (?map ! k)" . have "Poly (list_of_vec (row (berlekamp_mat u) k)) = Poly (list_of_vec (vec_of_list (?map ! k)))" unfolding row_rw .. also have "... = Poly (?map ! k)" by simp also have "... = [:0,1:]^(CARD('a) * k) mod u" proof - let ?cs = "(power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)) ! k" let ?c = "coeffs ?cs @ replicate (degree u - length (coeffs ?cs)) 0" have map_k_c: "?map ! k = ?c" by (rule nth_map, simp add: k) have "(Poly (?map ! k)) = Poly (coeffs ?cs)" unfolding map_k_c Poly_append_replicate_0 .. also have "... = ?cs" by simp also have "... = power_polys ([:0, 1:] ^ CARD('a) mod u) u 1 (degree u) ! k" by (simp add: power_poly_f_mod_def) also have "... = 1* ([:0,1:]^(CARD('a)) mod u) ^ k mod u" proof (rule power_polys_works[OF k]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = ([:0,1:]^(CARD('a)) mod u) ^ k mod u" by auto also have "... = [:0,1:]^(CARD('a) * k) mod u" by (simp add: power_mod power_mult) finally show ?thesis . qed finally show ?thesis . qed corollary Poly_berlekamp_cong_mat: assumes k: "k < degree u" shows "[Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k)] (mod u)" using Poly_berlekamp_mat[OF k] unfolding cong_def by auto lemma mat_of_rows_list_dim[simp]: "mat_of_rows_list n vs \ carrier_mat (length vs) n" "dim_row (mat_of_rows_list n vs) = length vs" "dim_col (mat_of_rows_list n vs) = n" unfolding mat_of_rows_list_def by auto lemma berlekamp_mat_closed[simp]: "berlekamp_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_mat u) = degree u" "dim_col (berlekamp_mat u) = degree u" unfolding carrier_mat_def berlekamp_mat_def Let_def by auto lemma vec_of_list_coeffs_nth: assumes i: "i \ {..degree h}" and h_not0: "h \ 0" shows "vec_of_list (coeffs h) $ i = coeff h i" proof - have "vec_of_list (map (coeff h) [0..i. f i mod z) A" using f by (induct, auto simp add: poly_mod_add_left) lemma prime_not_dvd_fact: assumes kn: "k < n" and prime_n: "prime n" shows "\ n dvd fact k" using kn proof (induct k) case 0 thus ?case using prime_n unfolding prime_nat_iff by auto next case (Suc k) show ?case proof (rule ccontr, unfold not_not) assume "n dvd fact (Suc k)" also have "... = Suc k * \{1..k}" unfolding fact_Suc unfolding fact_prod by simp finally have "n dvd Suc k * \{1..k}" . hence "n dvd Suc k \ n dvd \{1..k}" using prime_dvd_mult_eq_nat[OF prime_n] by blast moreover have "\ n dvd Suc k" by (simp add: Suc.prems(1) nat_dvd_not_less) moreover hence "\ n dvd \{1..k}" using Suc.hyps Suc.prems using Suc_lessD fact_prod[of k] by (metis of_nat_id) ultimately show False by simp qed qed lemma dvd_choose_prime: assumes kn: "k < n" and k: "k \ 0" and n: "n \ 0" and prime_n: "prime n" shows "n dvd (n choose k)" proof - have "n dvd (fact n)" by (simp add: fact_num_eq_if n) moreover have "\ n dvd (fact k * fact (n-k))" proof (rule ccontr, simp) assume "n dvd fact k * fact (n - k)" hence "n dvd fact k \ n dvd fact (n - k)" using prime_dvd_mult_eq_nat[OF prime_n] by simp moreover have "\ n dvd (fact k)" by (rule prime_not_dvd_fact[OF kn prime_n]) moreover have "\ n dvd fact (n - k)" using prime_not_dvd_fact[OF _ prime_n] kn k by simp ultimately show False by simp qed moreover have "(fact n::nat) = fact k * fact (n-k) * (n choose k)" using binomial_fact_lemma kn by auto ultimately show ?thesis using prime_n by (auto simp add: prime_dvd_mult_iff) qed lemma add_power_poly_mod_ring: fixes x :: "'a mod_ring poly" shows "(x + y) ^ CARD('a) = x ^ CARD('a) + y ^ CARD('a)" proof - let ?A="{0..CARD('a)}" let ?f="\k. of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k)" have A_rw: "?A = insert CARD('a) (insert 0 (?A - {0} - {CARD('a)}))" by fastforce have sum0: "sum ?f (?A - {0} - {CARD('a)}) = 0" proof (rule sum.neutral, rule) fix xa assume xa: "xa \ {0..CARD('a)} - {0} - {CARD('a)}" have card_dvd_choose: "CARD('a) dvd (CARD('a) choose xa)" proof (rule dvd_choose_prime) show "xa < CARD('a)" using xa by simp show "xa \ 0" using xa by simp show "CARD('a) \ 0" by simp show "prime CARD('a)" by (rule prime_card) qed hence rw0: "of_int (CARD('a) choose xa) = (0 :: 'a mod_ring)" by transfer simp have "of_nat (CARD('a) choose xa) = [:of_int (CARD('a) choose xa) :: 'a mod_ring:]" by (simp add: of_nat_poly) also have "... = [:0:]" using rw0 by simp finally show "of_nat (CARD('a) choose xa) * x ^ xa * y ^ (CARD('a) - xa) = 0" by auto qed have "(x + y)^CARD('a) = (\k = 0..CARD('a). of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k))" unfolding binomial_ring by (rule sum.cong, auto) also have "... = sum ?f (insert CARD('a) (insert 0 (?A - {0} - {CARD('a)})))" using A_rw by simp also have "... = ?f 0 + ?f CARD('a) + sum ?f (?A - {0} - {CARD('a)})" by auto also have "... = x^CARD('a) + y^CARD('a)" unfolding sum0 by auto finally show ?thesis . qed lemma power_poly_sum_mod_ring: fixes f :: "'b \ 'a mod_ring poly" assumes f: "finite A" shows "(sum f A) ^ CARD('a) = sum (\i. (f i) ^ CARD('a)) A" using f by (induct, auto simp add: add_power_poly_mod_ring) lemma poly_power_card_as_sum_of_monoms: fixes h :: "'a mod_ring poly" shows "h ^ CARD('a) = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof - have "h ^ CARD('a) = (\i\degree h. monom (coeff h i) i) ^ CARD('a)" by (simp add: poly_as_sum_of_monoms) also have "... = (\i\degree h. (monom (coeff h i) i) ^ CARD('a))" by (simp add: power_poly_sum_mod_ring) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" show "monom (coeff h x) x ^ CARD('a) = monom (coeff h x) (CARD('a) * x)" by (unfold poly_eq_iff, auto simp add: monom_power) qed finally show ?thesis . qed lemma degree_Poly_berlekamp_le: assumes i: "i < degree u" shows "degree (Poly (list_of_vec (row (berlekamp_mat u) i))) < degree u" by (metis Poly_berlekamp_mat degree_0 degree_mod_less gr_implies_not0 i linorder_neqE_nat) (* Equation 12: alternative statement. *) lemma monom_card_pow_mod_sum_berlekamp: assumes i: "i < degree u" shows "monom 1 (CARD('a) * i) mod u = (\j 0" using i by simp hence set_rw: "{..degree u - 1} = {.. degree u - 1" by auto have "monom 1 (CARD('a) * i) mod u = [:0, 1:] ^ (CARD('a) * i) mod u" using x_as_monom x_pow_n by metis also have "... = ?p" unfolding Poly_berlekamp_mat[OF i] by simp also have "... = (\i\degree u - 1. monom (coeff ?p i) i)" using degree_le2 poly_as_sum_of_monoms' by fastforce also have "... = (\ij {.. v = (\i = 0.. v = (\i = 0.. v = col A j \ v" using j row_transpose by auto also have "... = (\i = 0..iiiiiiiiiii degree u" by (metis Suc_leI assms coeffs_0_eq_Nil degree_0 length_coeffs_degree list.size(3) not_le_imp_less order.asym) thus ?thesis by simp qed lemma vec_of_list_coeffs_nth': assumes i: "i \ {..degree h}" and h_not0: "h \ 0" assumes "degree h < degree u" shows "vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) $ i = coeff h i" using assms by (transfer', auto simp add: mk_vec_def coeffs_nth length_coeffs_degree nth_append) lemma vec_of_list_coeffs_replicate_nth_0: assumes i: "i \ {.. {..{..degree h}") case True thus ?thesis using assms vec_of_list_coeffs_nth' h_not0 by simp next case False have c0: "coeff h i = 0" using False le_degree by auto thus ?thesis using assms False h_not0 by (transfer', auto simp add: mk_vec_def length_coeffs_degree nth_append c0) qed qed (* Equation 13 *) lemma equation_13: fixes u h defines H: "H \ vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0)" assumes deg_le: "degree h < degree u" (*Mandatory from equation 8*) shows "[h^CARD('a) = h] (mod u) \ (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H" (is "?lhs = ?rhs") proof - have f: "finite {..degree u}" by auto have [simp]: "dim_vec H = degree u" unfolding H using dim_vec_of_list_h deg_le by simp let ?B = "(berlekamp_mat u)" let ?f = "\i. (transpose_mat ?B *\<^sub>v H) $ i" show ?thesis proof assume rhs: ?rhs have dimv_h_dimr_B: "dim_vec H = dim_row ?B" by (metis berlekamp_mat_closed(2) berlekamp_mat_closed(3) dim_mult_mat_vec index_transpose_mat(2) rhs) have degree_h_less_dim_H: "degree h < dim_vec H" by (auto simp add: deg_le) have set_rw: "{..degree u - 1} = {.. degree u - 1" using deg_le by simp hence "h = (\j\degree u - 1. monom (coeff h j) j)" using poly_as_sum_of_monoms' by fastforce also have "... = (\jj {..j H) j)" by (rule sum.cong, auto) also have "... = (\ji = 0.. {.. H) x = monom (\i = 0..ji = 0..i = 0..ji = 0..ji = 0..ji = 0.. {0..jji = 0..i = 0..i = 0..x. x mod u"], rule sum.cong, rule) fix x assume x: "x \ {0.. {..i\degree h. monom (coeff h i) (CARD('a) * i)) mod u" proof (rule arg_cong[of _ _ "\x. x mod u"]) let ?f="\i. monom (coeff h i) (CARD('a) * i)" have ss0: "(\i = degree h + 1 ..< dim_vec H. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0..< dim_vec H} = {0..degree h} \ {degree h + 1 ..< dim_vec H}" using degree_h_less_dim_H by auto have "(\i = 0..i = 0..degree h. ?f i) + (\i = degree h + 1 ..< dim_vec H. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i = 0..degree h. ?f i)" unfolding ss0 by auto finally show "(\i = 0..i\degree h. ?f i)" by (simp add: atLeast0AtMost) qed also have "... = h^CARD('a) mod u" using poly_power_card_as_sum_of_monoms by auto finally show ?lhs unfolding cong_def using deg_le by (simp add: mod_poly_less) next assume lhs: ?lhs have deg_le': "degree h \ degree u - 1" using deg_le by auto have set_rw: "{..ii \ degree u - 1. monom (coeff h i) i)" by simp also have "... = (\i\degree h. monom (coeff h i) i)" unfolding poly_as_sum_of_monoms using poly_as_sum_of_monoms' deg_le' by auto also have "... = (\i\degree h. monom (coeff h i) i) mod u" by (simp add: deg_le mod_poly_less poly_as_sum_of_monoms) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i)) mod u" using lhs unfolding cong_def poly_as_sum_of_monoms poly_power_card_as_sum_of_monoms by auto also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i)) mod u" by (rule arg_cong[of _ _ "\x. x mod u"], rule sum.cong, simp_all add: mult_monom) also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i) mod u)" by (simp add: poly_mod_sum) also have "... = (\i\degree h. monom (coeff h i) 0 * (monom 1 (CARD('a)*i) mod u))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" have h_rw: "monom (coeff h x) 0 mod u = monom (coeff h x) 0" by (metis deg_le degree_pCons_eq_if gr_implies_not_zero linorder_neqE_nat mod_poly_less monom_0) have "monom (coeff h x) 0 * monom 1 (CARD('a) * x) = smult (coeff h x) (monom 1 (CARD('a) * x))" by (simp add: monom_0) also have "... mod u = Polynomial.smult (coeff h x) (monom 1 (CARD('a) * x) mod u)" using mod_smult_left by auto also have "... = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" by (simp add: monom_0) finally show "monom (coeff h x) 0 * monom 1 (CARD('a) * x) mod u = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" . qed also have "... = (\i\degree h. monom (coeff h i) 0 * (\j {..degree h}" have "(monom 1 (CARD('a) * x) mod u) = (\jjiji=degree h+1 ..< degree u. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0.. {degree h+1..i=0..i=0..degree h. ?f i) + (\i=degree h+1 ..< degree u. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i=0..degree h. ?f i)" using ss0 by simp finally show ?thesis by (simp add: atLeast0AtMost atLeast0LessThan) qed also have "... = (\ijijjijiijii. i < degree u \ coeff h i = (\ji. i < degree u \ H $ i = (\jjjjv H = H" proof (rule eq_vecI) fix i show "dim_vec (transpose_mat ?B *\<^sub>v H) = dim_vec (H)" by auto assume i: "i < dim_vec (H)" have "(transpose_mat ?B *\<^sub>v H) $ i = row (transpose_mat ?B) i \ H" using i by simp also have "... = (\j = 0..jv H) $ i = H $ i" . qed qed qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma exists_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\s. fi dvd (h - [:s:])" proof - let ?p = "CARD('a)" have f_dvd_hqh: "f dvd (h^?p - h)" using h unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis using irreducible_dvd_prod[OF _ fi_dvd_prod_hc] irr_fi by auto qed corollary exists_unique_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\!s. fi dvd (h - [:s:])" proof - obtain c where fi_dvd: "fi dvd (h - [:c:])" using assms exists_s_factor_dvd_h_s by blast have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" by (simp add: irr_fi irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis proof (rule ex1I[of _ c], auto simp add: fi_dvd) fix c2 assume fi_dvd_hc2: "fi dvd h - [:c2:]" have *: "fi dvd (h - [:c:]) * (h - [:c2:])" using fi_dvd by auto hence "fi dvd (h - [:c:]) \ fi dvd (h - [:c2:])" using irr_fi by auto thus "c2 = c" using coprime_h_c_poly coprime_not_unit_not_dvd fi_dvd fi_dvd_hc2 fi_not_unit by blast qed qed lemma exists_two_distint: "\a b::'a mod_ring. a \ b" by (rule exI[of _ 0], rule exI[of _ 1], auto) lemma coprime_cong_mult_factorization_poly: fixes f::"'b::{field} poly" and a b p :: "'c :: {field_gcd} poly" assumes finite_P: "finite P" and P: "P \ {q. irreducible q}" and p: "\p\P. [a=b] (mod p)" and coprime_P: "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" shows "[a = b] (mod (\a\P. a))" using finite_P P p coprime_P proof (induct P) case empty thus ?case by simp next case (insert p P) have ab_mod_pP: "[a=b] (mod (p * \P))" proof (rule coprime_cong_mult_poly) show "[a = b] (mod p)" using insert.prems by auto show "[a = b] (mod \P)" using insert.prems insert.hyps by auto from insert show "Rings.coprime p (\P)" by (auto intro: prod_coprime_right) qed thus ?case by (simp add: insert.hyps(1) insert.hyps(2)) qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma W_eq_berlekamp_mat: fixes u::"'a mod_ring poly" shows "{v. [v^CARD('a) = v] (mod u) \ degree v < degree u} = {h. let H = vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) in (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H \ degree h < degree u}" using equation_13 by (auto simp add: Let_def) lemma transpose_minus_1: assumes "dim_row Q = dim_col Q" shows "transpose_mat (Q - (1\<^sub>m (dim_row Q))) = (transpose_mat Q - (1\<^sub>m (dim_row Q)))" using assms unfolding mat_eq_iff by auto lemma system_iff: fixes v::"'b::comm_ring_1 vec" assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" proof - have t1:"transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v) \ (transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v)" by (subst minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have t2:"(transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v) \ transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v)" by (subst (asm) minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have "transpose_mat Q *\<^sub>v v - v = v - v \ transpose_mat Q *\<^sub>v v = v" proof - assume a1: "transpose_mat Q *\<^sub>v v - v = v - v" have f2: "transpose_mat Q *\<^sub>v v \ carrier_vec (dim_vec v)" by (metis dim_mult_mat_vec index_transpose_mat(2) sq_Q v carrier_vec_dim_vec) then have f3: "0\<^sub>v (dim_vec v) + transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v" by (meson left_zero_vec) have f4: "0\<^sub>v (dim_vec v) = transpose_mat Q *\<^sub>v v - v" using a1 by auto have f5: "- v \ carrier_vec (dim_vec v)" by simp then have f6: "- v + transpose_mat Q *\<^sub>v v = v - v" using f2 a1 using comm_add_vec minus_add_uminus_vec by fastforce have "v - v = - v + v" by auto then have "transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v - v + v" using f6 f4 f3 f2 by (metis (no_types, lifting) a1 assoc_add_vec comm_add_vec f5 carrier_vec_dim_vec) then show ?thesis using a1 by auto qed hence "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q *\<^sub>v v) - v = v - v)" by auto also have "... = ((transpose_mat Q *\<^sub>v v) - v = 0\<^sub>v (dim_vec v))" by auto also have "... = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using t1 t2 by auto finally show ?thesis. qed lemma system_if_mat_kernel: assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q))))" proof - have "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using assms system_iff by blast also have "... = (v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q)))))" unfolding mat_kernel_def unfolding transpose_minus_1[OF sq_Q] unfolding v by auto finally show ?thesis . qed lemma degree_u_mod_irreducible\<^sub>d_factor_0: fixes v and u::"'a mod_ring poly" defines W: "W \ {v. [v ^ CARD('a) = v] (mod u)}" assumes v: "v \ W" and finite_U: "finite U" and u_U: "u = \U" and U_irr_monic: "U \ {q. irreducible q \ monic q}" and fi_U: "fi \ U" shows "degree (v mod fi) = 0" proof - have deg_fi: "degree fi > 0" using U_irr_monic using fi_U irreducible\<^sub>dD[of fi] by auto have "fi dvd u" using u_U U_irr_monic finite_U dvd_prod_eqI fi_U by blast moreover have "u dvd (v^CARD('a) - v)" using v unfolding W cong_def by (simp add: mod_eq_dvd_iff_poly) ultimately have "fi dvd (v^CARD('a) - v)" by (rule dvd_trans) then have fi_dvd_prod_vc: "fi dvd prod (\c. v - [:c:]) (UNIV::'a mod_ring set)" by (simp add: poly_identity_mod_p) have irr_fi: "irreducible fi" using fi_U U_irr_monic by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (auto simp: poly_dvd_1) have fi_dvd_vc: "\c. fi dvd v - [:c:]" using irreducible_dvd_prod[OF _ fi_dvd_prod_vc] irr_fi by auto from this obtain a where "fi dvd v - [:a:]" by blast hence "v mod fi = [:a:] mod fi" using mod_eq_dvd_iff_poly by blast also have "... = [:a:]" by (simp add: deg_fi mod_poly_less) finally show ?thesis by simp qed (* Also polynomials over a field as a vector space in HOL-Algebra.*) definition "poly_abelian_monoid = \carrier = UNIV::'a mod_ring poly set, monoid.mult = ((*)), one = 1, zero = 0, add = (+), module.smult = smult\" interpretation vector_space_poly: vectorspace class_ring poly_abelian_monoid rewrites [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 0" and [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 1" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (+)" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (*)" and [simp]: "carrier poly_abelian_monoid = UNIV" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = smult" apply unfold_locales apply (auto simp: poly_abelian_monoid_def class_field_def smult_add_left smult_add_right Units_def) by (metis add.commute add.right_inverse) lemma subspace_Berlekamp: assumes f: "degree f \ 0" shows "subspace (class_ring :: 'a mod_ring ring) {v. [v^(CARD('a)) = v] (mod f) \ (degree v < degree f)} poly_abelian_monoid" proof - { fix v :: "'a mod_ring poly" and w :: "'a mod_ring poly" assume a1: "v ^ card (UNIV::'a set) mod f = v mod f" assume "w ^ card (UNIV::'a set) mod f = w mod f" then have "(v ^ card (UNIV::'a set) + w ^ card (UNIV::'a set)) mod f = (v + w) mod f" using a1 by (meson mod_add_cong) then have "(v + w) ^ card (UNIV::'a set) mod f = (v + w) mod f" by (simp add: add_power_poly_mod_ring) } note r=this thus ?thesis using f by (unfold_locales, auto simp: zero_power mod_smult_left smult_power cong_def degree_add_less) qed lemma berlekamp_resulting_mat_closed[simp]: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" proof - let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" let ?G="(gauss_jordan_single ?A)" have "?G \carrier_mat (degree u) (degree u)" by (rule gauss_jordan_single(2)[of ?A], auto) thus "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" unfolding berlekamp_resulting_mat_def Let_def by auto qed (*find_base vectors returns a basis:*) lemma berlekamp_resulting_mat_basis: "kernel.basis (degree u) (berlekamp_resulting_mat u) (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule find_base_vectors(3)) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" have "row_echelon_form (gauss_jordan_single ?A)" by (rule gauss_jordan_single(3)[of ?A], auto) thus "row_echelon_form (berlekamp_resulting_mat u)" unfolding berlekamp_resulting_mat_def Let_def by auto qed lemma set_berlekamp_basis_eq: "(set (berlekamp_basis u)) = (Poly \ list_of_vec)` (set (find_base_vectors (berlekamp_resulting_mat u)))" by (auto simp add: image_def o_def berlekamp_basis_def) lemma berlekamp_resulting_mat_constant: assumes deg_u: "degree u = 0" shows "berlekamp_resulting_mat u = 1\<^sub>m 0" by (unfold mat_eq_iff, auto simp add: deg_u) context fixes u::"'a::prime_card mod_ring poly" begin lemma set_berlekamp_basis_constant: assumes deg_u: "degree u = 0" shows "set (berlekamp_basis u) = {}" proof - have one_carrier: "1\<^sub>m 0 \ carrier_mat 0 0" by auto have m: "mat_kernel (1\<^sub>m 0) = {(0\<^sub>v 0) :: 'a mod_ring vec}" unfolding mat_kernel_def by auto have r: "row_echelon_form (1\<^sub>m 0 :: 'a mod_ring mat)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto have "set (find_base_vectors (1\<^sub>m 0)) \ {0\<^sub>v 0 :: 'a mod_ring vec}" using find_base_vectors(1)[OF r one_carrier] unfolding m . hence "set (find_base_vectors (1\<^sub>m 0) :: 'a mod_ring vec list) = {}" using find_base_vectors(2)[OF r one_carrier] using subset_singletonD by fastforce thus ?thesis unfolding set_berlekamp_basis_eq unfolding berlekamp_resulting_mat_constant[OF deg_u] by auto qed (*Maybe [simp]*) lemma row_echelon_form_berlekamp_resulting_mat: "row_echelon_form (berlekamp_resulting_mat u)" by (rule gauss_jordan_single(3), auto simp add: berlekamp_resulting_mat_def Let_def) lemma mat_kernel_berlekamp_resulting_mat_degree_0: assumes d: "degree u = 0" shows "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (auto simp add: mat_kernel_def mult_mat_vec_def d) lemma in_mat_kernel_berlekamp_resulting_mat: assumes x: "transpose_mat (berlekamp_mat u) *\<^sub>v x = x" and x_dim: "x \ carrier_vec (degree u)" shows "x \ mat_kernel (berlekamp_resulting_mat u)" proof - let ?QI="(mat(dim_row (berlekamp_mat u)) (dim_row (berlekamp_mat u)) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" have *: "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (degree u)) = transpose_mat ?QI" by auto have "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (dim_row (berlekamp_mat u))) *\<^sub>v x = 0\<^sub>v (dim_vec x)" using system_iff[of "berlekamp_mat u" x] x_dim x by auto hence "transpose_mat ?QI *\<^sub>v x = 0\<^sub>v (degree u)" using x_dim * by auto hence "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" unfolding berlekamp_resulting_mat_def Let_def using gauss_jordan_single(1)[of "transpose_mat ?QI" "degree u" "degree u" _ x] x_dim by auto thus ?thesis by (auto simp add: mat_kernel_def x_dim) qed private abbreviation "V \ kernel.VK (degree u) (berlekamp_resulting_mat u)" private abbreviation "W \ vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u) \ (degree v < degree u)}" interpretation V: vectorspace class_ring V proof - interpret k: kernel "(degree u)" "(degree u)" "(berlekamp_resulting_mat u)" by (unfold_locales; auto) show "vectorspace class_ring V" by intro_locales qed lemma linear_Poly_list_of_vec: shows "(Poly \ list_of_vec) \ module_hom class_ring V (vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u)})" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed qed lemma linear_Poly_list_of_vec': assumes "degree u > 0" shows "(Poly \ list_of_vec) \ module_hom R V W" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "degree (Poly (list_of_vec x)) < degree u" by (rule degree_Poly_list_of_vec, insert assms x, auto simp: mat_kernel_def) qed lemma berlekamp_basis_eq_8: assumes v: "v \ set (berlekamp_basis u)" shows "[v ^ CARD('a) = v] (mod u)" proof - { fix x assume x: "x \ set (find_base_vectors (berlekamp_resulting_mat u))" have "set (find_base_vectors (berlekamp_resulting_mat u)) \ mat_kernel (berlekamp_resulting_mat u)" proof (rule find_base_vectors(1)) show "row_echelon_form (berlekamp_resulting_mat u)" by (rule row_echelon_form_berlekamp_resulting_mat) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp qed hence "x \ mat_kernel (berlekamp_resulting_mat u)" using x by auto hence "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto } thus "[v ^ CARD('a) = v] (mod u)" using v unfolding set_berlekamp_basis_eq by auto qed lemma surj_Poly_list_of_vec: assumes deg_u: "degree u > 0" shows "(Poly \ list_of_vec)` (carrier V) = carrier W" proof (auto simp add: image_def) fix xa assume xa: "xa \ mat_kernel (berlekamp_resulting_mat u)" thus "[Poly (list_of_vec xa) ^ CARD('a) = Poly (list_of_vec xa)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto show "degree (Poly (list_of_vec xa)) < degree u" proof (rule degree_Poly_list_of_vec[OF _ deg_u]) show "xa \ carrier_vec (degree u)" using xa unfolding mat_kernel_def by simp qed next fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" show "\xa \ mat_kernel (berlekamp_resulting_mat u). x = Poly (list_of_vec xa)" proof (rule bexI[of _ "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)"]) let ?X = "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)" show "x = Poly (list_of_vec (vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)))" by auto have X: "?X \ carrier_vec (degree u)" unfolding carrier_vec_def by (auto, metis Suc_leI coeffs_0_eq_Nil deg_x degree_0 le_add_diff_inverse length_coeffs_degree linordered_semidom_class.add_diff_inverse list.size(3) order.asym) have t: "transpose_mat (berlekamp_mat u) *\<^sub>v ?X = ?X" using equation_13[OF deg_x] x by auto show "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0) \ mat_kernel (berlekamp_resulting_mat u)" by (rule in_mat_kernel_berlekamp_resulting_mat[OF t X]) qed qed lemma card_set_berlekamp_basis: "card (set (berlekamp_basis u)) = length (berlekamp_basis u)" proof - have b: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by auto have "(set (berlekamp_basis u)) = (Poly \ list_of_vec) ` set (find_base_vectors (berlekamp_resulting_mat u))" unfolding set_berlekamp_basis_eq .. also have " card ... = card (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule card_image, rule subset_inj_on[OF inj_Poly_list_of_vec]) show "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier_vec (degree u)" using find_base_vectors(1)[OF row_echelon_form_berlekamp_resulting_mat b] unfolding carrier_vec_def mat_kernel_def by auto qed also have "... = length (find_base_vectors (berlekamp_resulting_mat u))" by (rule length_find_base_vectors[symmetric, OF row_echelon_form_berlekamp_resulting_mat b]) finally show ?thesis unfolding berlekamp_basis_def by auto qed context assumes deg_u0[simp]: "degree u > 0" begin interpretation Berlekamp_subspace: vectorspace class_ring W by (rule vector_space_poly.subspace_is_vs[OF subspace_Berlekamp], simp) lemma linear_map_Poly_list_of_vec': "linear_map class_ring V W (Poly \ list_of_vec)" proof (auto simp add: linear_map_def) show "vectorspace class_ring V" by intro_locales show "vectorspace class_ring W" by (rule Berlekamp_subspace.vectorspace_axioms) show "mod_hom class_ring V W (Poly \ list_of_vec)" proof (rule mod_hom.intro, unfold mod_hom_axioms_def) show "module class_ring V" by intro_locales show "module class_ring W" using Berlekamp_subspace.vectorspace_axioms by intro_locales show "Poly \ list_of_vec \ module_hom class_ring V W" by (rule linear_Poly_list_of_vec'[OF deg_u0]) qed qed lemma berlekamp_basis_basis: "Berlekamp_subspace.basis (set (berlekamp_basis u))" proof (unfold set_berlekamp_basis_eq, rule linear_map.linear_inj_image_is_basis) show "linear_map class_ring V W (Poly \ list_of_vec)" by (rule linear_map_Poly_list_of_vec') show "inj_on (Poly \ list_of_vec) (carrier V)" proof (rule subset_inj_on[OF inj_Poly_list_of_vec]) show "carrier V \ carrier_vec (degree u)" by (auto simp add: mat_kernel_def) qed show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show b: "V.basis (set (find_base_vectors (berlekamp_resulting_mat u)))" by (rule berlekamp_resulting_mat_basis) show "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using b unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed qed lemma finsum_sum: fixes f::"'a mod_ring poly" assumes f: "finite B" and a_Pi: "a \ B \ carrier R" and V: "B \ carrier W" shows "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" using f a_Pi V proof (induct B) case empty thus ?case unfolding Berlekamp_subspace.module.M.finsum_empty by auto next case (insert x V) have hyp: "(\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) V" proof (rule insert.hyps) show "a \ V \ carrier R" using insert.prems unfolding class_field_def by auto show "V \ carrier W" using insert.prems by simp qed have "(\\<^bsub>W\<^esub>v\insert x V. a v \\<^bsub>W\<^esub> v) = (a x \\<^bsub>W\<^esub> x) \\<^bsub>W\<^esub> (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" proof (rule abelian_monoid.finsum_insert) show "abelian_monoid W" by (unfold_locales) show "finite V" by fact show "x \ V" by fact show "(\v. a v \\<^bsub>W\<^esub> v) \ V \ carrier W" proof (unfold Pi_def, rule, rule allI, rule impI) fix v assume v: "v\V" show "a v \\<^bsub>W\<^esub> v \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a v \ carrier class_ring" using insert.prems v unfolding Pi_def by (simp add: class_field_def) show "v \ carrier W" using v insert.prems by auto qed qed show "a x \\<^bsub>W\<^esub> x \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a x \ carrier class_ring" using insert.prems unfolding Pi_def by (simp add: class_field_def) show "x \ carrier W" using insert.prems by auto qed qed also have "... = (a x \\<^bsub>W\<^esub> x) + (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" by auto also have "... = (a x \\<^bsub>W\<^esub> x) + sum (\v. smult (a v) v) V" unfolding hyp by simp also have "... = (smult (a x) x) + sum (\v. smult (a v) v) V" by simp also have "... = sum (\v. smult (a v) v) (insert x V)" by (simp add: insert.hyps(1) insert.hyps(2)) finally show ?case . qed lemma exists_vector_in_Berlekamp_subspace_dvd: fixes p_i::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v. v \ {h. [h^(CARD('a)) = h] (mod u) \ degree h < degree u} \ v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1)" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by auto have irr_pj: "irreducible p_j" using pj P by auto obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m f_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m f_desc_square_free f_not_0 by auto obtain i where mi: "m i = p_i" and i: "i < n" using P_m pi by blast obtain j where mj: "m j = p_j" and j: "j < n" using P_m pj by blast have ij: "i \ j" using mi mj pi_pj by auto obtain s_i and s_j::"'a mod_ring" where si_sj: "s_i \ s_j" using exists_two_distint by blast let ?u="\x. if x = i then [:s_i:] else if x = j then [:s_j:] else [:0:]" have degree_si: "degree [:s_i:] = 0" by auto have degree_sj: "degree [:s_j:] = 0" by auto have "\!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\a\{i. i < n}. [v = ?u a] (mod m a))" proof (rule chinese_remainder_unique_poly) show "\a\{i. i < n}. \b\{i. i < n}. a \ b \ Rings.coprime (m a) (m b)" proof (rule+) fix a b assume "a \ {i. i < n}" and "b \ {i. i < n}" and "a \ b" thus "Rings.coprime (m a) (m b)" using coprime_polynomial_factorization[OF P finite_P, simplified] P_m by (metis image_eqI inj_onD inj_on_m) qed show "\i\{i. i < n}. m i \ 0" by (rule not_zero) show "0 < degree (prod m {i. i < n})" unfolding degree_prod using deg_u0 by blast qed from this obtain v where v: "\a\{i. i < n}. [v = ?u a] (mod m a)" and degree_v: "degree v < (\i\{i. i < n}. degree (m i))" by blast show ?thesis proof (rule exI[of _ v], auto) show vp_v_mod: "[v ^ CARD('a) = v] (mod u)" proof (unfold f_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P]) show "P \ {q. irreducible q}" using P by blast show "\p\P. [v ^ CARD('a) = v] (mod p)" proof (rule ballI) fix p assume p: "p \ P" hence irr_p: "irreducible\<^sub>d p" using P by auto obtain k where mk: "m k = p" and k: "k < n" using P_m p by blast have "[v = ?u k] (mod p)" using v mk k by auto moreover have "?u k mod p = ?u k" apply (rule mod_poly_less) using irreducible\<^sub>dD(1)[OF irr_p] by auto ultimately obtain s where v_mod_p: "v mod p = [:s:]" unfolding cong_def by force hence deg_v_p: "degree (v mod p) = 0" by auto have "v mod p = [:s:]" by (rule v_mod_p) also have "... = [:s:]^CARD('a)" unfolding poly_const_pow by auto also have "... = (v mod p) ^ CARD('a)" using v_mod_p by auto also have "... = (v mod p) ^ CARD('a) mod p" using calculation by auto also have "... = v^CARD('a) mod p" using power_mod by blast finally show "[v ^ CARD('a) = v] (mod p)" unfolding cong_def .. qed show "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" using P coprime_polynomial_factorization finite_P by auto qed have "[v = ?u i] (mod m i)" using v i by auto hence v_pi_si_mod: "v mod p_i = [:s_i:] mod p_i" unfolding cong_def mi by auto also have "... = [:s_i:]" apply (rule mod_poly_less) using irr_pi by auto finally have v_pi_si: "v mod p_i = [:s_i:]" . have "[v = ?u j] (mod m j)" using v j by auto hence v_pj_sj_mod: "v mod p_j = [:s_j:] mod p_j" unfolding cong_def mj using ij by auto also have "... = [:s_j:]" apply (rule mod_poly_less) using irr_pj by auto finally have v_pj_sj: "v mod p_j = [:s_j:]" . show "v mod p_i = v mod p_j \ False" using si_sj v_pi_si v_pj_sj by auto show "degree (v mod p_i) = 0" unfolding v_pi_si by simp show "degree (v mod p_j) = 0" unfolding v_pj_sj by simp show "\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" using v_pi_si_mod mod_eq_dvd_iff_poly by blast have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" using v_pj_sj_mod mod_eq_dvd_iff_poly by blast have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using vp_v_mod cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" proof (rule ccontr, simp) assume gcd_w: "gcd w (v - [:s_i:]) = w" show False apply (rule \v mod p_i = v mod p_j \ False\) by (metis irreducibleE \degree (v mod p_i) = 0\ gcd_greatest_iff gcd_w irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_si) qed show "gcd w (v - [:s_i:]) \ 1" by (metis irreducibleE gcd_greatest_iff irr_pi pi_dvd_v_si pi_dvd_w) qed show "degree v < degree u" proof - have "(\i | i < n. degree (m i)) = degree (prod m {i. i < n})" by (rule degree_prod_eq_sum_degree[symmetric, OF not_zero]) thus ?thesis using degree_v unfolding degree_prod by auto qed qed qed lemma exists_vector_in_Berlekamp_basis_dvd_aux: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j" proof (rule ccontr, auto) have V_in_carrier: "B \ carrier W" using basis_V unfolding Berlekamp_subspace.basis_def by auto assume all_eq: "\v\B. v mod p_i = v mod p_j" obtain x where x: "x \ {h. [h ^ CARD('a) = h] (mod u) \ degree h < degree u}" and x_pi_pj: "x mod p_i \ x mod p_j" and "degree (x mod p_i) = 0" and "degree (x mod p_j) = 0" "(\s. gcd w (x - [:s:]) \ w \ gcd w (x - [:s:]) \ 1)" using exists_vector_in_Berlekamp_subspace_dvd[OF _ _ _ pi pj _ _ _ _ w_dvd_f monic_w pi_dvd_w] assms by meson have x_in: "x \ carrier W" using x by auto hence "(\!a. a \ B \\<^sub>E carrier class_ring \ Berlekamp_subspace.lincomb a B = x)" using Berlekamp_subspace.basis_criterion[OF finite_V V_in_carrier] using basis_V by (simp add: class_field_def) from this obtain a where a_Pi: "a \ B \\<^sub>E carrier class_ring" and lincomb_x: "Berlekamp_subspace.lincomb a B = x" by blast have fs_ss: "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" proof (rule finsum_sum) show "finite B" by fact show "a \ B \ carrier class_ring" using a_Pi by auto show "B \ carrier W" by (rule V_in_carrier) qed have "x mod p_i = Berlekamp_subspace.lincomb a B mod p_i" using lincomb_x by simp also have " ... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_i" unfolding Berlekamp_subspace.lincomb_def .. also have "... = (sum (\v. smult (a v) v) B) mod p_i" unfolding fs_ss .. also have "... = sum (\v. smult (a v) v mod p_i) B" using finite_V poly_mod_sum by blast also have "... = sum (\v. smult (a v) (v mod p_i)) B" by (meson mod_smult_left) also have "... = sum (\v. smult (a v) (v mod p_j)) B" using all_eq by auto also have "... = sum (\v. smult (a v) v mod p_j) B" by (metis mod_smult_left) also have "... = (sum (\v. smult (a v) v) B) mod p_j" by (metis (mono_tags, lifting) finite_V poly_mod_sum sum.cong) also have "... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_j" unfolding fs_ss .. also have "... = Berlekamp_subspace.lincomb a B mod p_j" unfolding Berlekamp_subspace.lincomb_def .. also have "... = x mod p_j" using lincomb_x by simp finally have "x mod p_i = x mod p_j" . thus False using x_pi_pj by contradiction qed lemma exists_vector_in_Berlekamp_basis_dvd: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ \ coprime w (v - [:s:]))" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by fast have irr_pj: "irreducible p_j" using pj P by fast obtain v where vV: "v \ B" and v_pi_pj: "v mod p_i \ v mod p_j" using assms exists_vector_in_Berlekamp_basis_dvd_aux by blast have v: "v \ {v. [v ^ CARD('a) = v] (mod u)}" using basis_V vV unfolding Berlekamp_subspace.basis_def by auto have deg_v_pi: "degree (v mod p_i) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pi]) from this obtain s_i where v_pi_si: "v mod p_i = [:s_i:]" using degree_eq_zeroE by blast have deg_v_pj: "degree (v mod p_j) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pj]) from this obtain s_j where v_pj_sj: "v mod p_j = [:s_j:]" using degree_eq_zeroE by blast have si_sj: "s_i \ s_j" using v_pi_si v_pj_sj v_pi_pj by auto have "(\s. gcd w (v - [:s:]) \ w \ \ Rings.coprime w (v - [:s:]))" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pi_si) have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pj_sj) have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using v cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" by (metis irreducibleE deg_v_pi gcd_greatest_iff irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_pj v_pi_si) show "\ Rings.coprime w (v - [:s_i:])" using irr_pi pi_dvd_v_si pi_dvd_w by (simp add: irreducible\<^sub>dD(1) not_coprimeI) qed thus ?thesis using v_pi_pj vV deg_v_pi deg_v_pj by auto qed lemma exists_bijective_linear_map_W_vec: assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" shows "\f. linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f \ bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" proof - let ?B="carrier_vec (card P)::'a mod_ring vec set" have u_not_0: "u \ 0" using deg_u0 degree_0 by force obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence n: "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m u_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m u_desc_square_free u_not_0 by auto have deg_sum_eq: "(\i\{i. i < n}. degree (m i)) = degree u" by (metis degree_prod degree_prod_eq_sum_degree not_zero) have coprime_mi_mj:"\i\{i. i < n}. \j\{i. i < n}. i \ j \ coprime (m i) (m j)" proof (rule+) fix i j assume i: "i \ {i. i < n}" and j: "j \ {i. i < n}" and ij: "i \ j" show "coprime (m i) (m j)" proof (rule coprime_polynomial_factorization[OF P finite_P]) show "m i \ P" using i P_m by auto show "m j \ P" using j P_m by auto show "m i \ m j" using inj_on_m i ij j unfolding inj_on_def by blast qed qed let ?f = "\v. vec n (\i. coeff (v mod (m i)) 0)" interpret vec_VS: vectorspace class_ring "(module_vec TYPE('a mod_ring) n)" by (rule VS_Connect.vec_vs) interpret linear_map class_ring W "(module_vec TYPE('a mod_ring) n)" ?f by (intro_locales, unfold mod_hom_axioms_def LinearCombinations.module_hom_def, auto simp add: vec_eq_iff module_vec_def mod_smult_left poly_mod_add_left) have "linear_map class_ring W (module_vec TYPE('a mod_ring) n) ?f" by (intro_locales) moreover have inj_f: "inj_on ?f (carrier W)" proof (rule Ke0_imp_inj, auto simp add: mod_hom.ker_def) show "[0 ^ CARD('a) = 0] (mod u)" by (simp add: cong_def zero_power) show "vec n (\i. 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" by (auto simp add: module_vec_def) fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" and v: "vec n (\i. coeff (x mod m i) 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" have cong_0: "\i\{i. i < n}. [x = (\i. 0) i] (mod m i)" proof (rule, unfold cong_def) fix i assume i: "i \ {i. i < n}" have deg_x_mod_mi: "degree (x mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "x \ {v. [v ^ CARD('a) = v] (mod u)}" using x by auto show "m i \ P" using P_m i by auto qed thus "x mod m i = 0 mod m i" using v unfolding module_vec_def by (auto, metis i leading_coeff_neq_0 mem_Collect_eq index_vec index_zero_vec(1)) qed moreover have deg_x2: "degree x < (\i\{i. i < n}. degree (m i))" using deg_sum_eq deg_x by simp moreover have "\i\{i. i < n}. [0 = (\i. 0) i] (mod m i)" by (auto simp add: cong_def) moreover have "degree 0 < (\i\{i. i < n}. degree (m i))" using degree_prod deg_sum_eq deg_u0 by force moreover have "\!x. degree x < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [x = (\i. 0) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) ultimately show "x = 0" by blast qed moreover have "?f ` (carrier W) = ?B" proof (auto simp add: image_def) fix xa show "n = card P" by (auto simp add: n) next fix x::"'a mod_ring vec" assume x: "x \ carrier_vec (card P)" have " \!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) from this obtain v where deg_v: "degree v < (\i\{i. i < n}. degree (m i))" and v_x_cong: "(\i \ {i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" by auto show "\xa. [xa ^ CARD('a) = xa] (mod u) \ degree xa < degree u \ x = vec n (\i. coeff (xa mod m i) 0)" proof (rule exI[of _ v], auto) show v: "[v ^ CARD('a) = v] (mod u)" proof (unfold u_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P], auto) fix y assume y: "y \ P" thus "irreducible y" using P by blast obtain i where i: "i \ {i. i < n}" and mi: "y = m i" using P_m y by blast have "irreducible (m i)" using i P_m P by auto moreover have "[v = [:x $ i:]] (mod m i)" using v_x_cong i by auto ultimately have v_mi_eq_xi: "v mod m i = [:x $ i:]" by (auto simp: cong_def intro!: mod_poly_less) have xi_pow_xi: "[:x $ i:]^CARD('a) = [:x $ i:]" by (simp add: poly_const_pow) hence "(v mod m i)^CARD('a) = v mod m i" using v_mi_eq_xi by auto hence "(v mod m i)^CARD('a) = (v^CARD('a) mod m i)" by (metis mod_mod_trivial power_mod) thus "[v ^ CARD('a) = v] (mod y)" unfolding mi cong_def v_mi_eq_xi xi_pow_xi by simp next fix p1 p2 assume "p1 \ P" and "p2 \ P" and "p1 \ p2" then show "Rings.coprime p1 p2" using coprime_polynomial_factorization[OF P finite_P] by auto qed show "degree v < degree u" using deg_v deg_sum_eq degree_prod by presburger show "x = vec n (\i. coeff (v mod m i) 0)" proof (unfold vec_eq_iff, rule conjI) show "dim_vec x = dim_vec (vec n (\i. coeff (v mod m i) 0))" using x n by simp show "\ii. coeff (v mod m i) 0)). x $ i = vec n (\i. coeff (v mod m i) 0) $ i" proof (auto) fix i assume i: "i < n" have deg_mi: "irreducible (m i)" using i P_m P by auto have deg_v_mi: "degree (v mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "v \ {v. [v ^ CARD('a) = v] (mod u)}" using v by fast show "m i \ P" using P_m i by auto qed have "v mod m i = [:x $ i:] mod m i" using v_x_cong i unfolding cong_def by auto also have "... = [:x $ i:]" using deg_mi by (auto intro!: mod_poly_less) finally show "x $ i = coeff (v mod m i) 0" by simp qed qed qed qed ultimately show ?thesis unfolding bij_betw_def n by auto qed lemma fin_dim_kernel_berlekamp: "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using berlekamp_resulting_mat_basis[of u] unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed lemma Berlekamp_subspace_fin_dim: "Berlekamp_subspace.fin_dim" proof (rule linear_map.surj_fin_dim[OF linear_map_Poly_list_of_vec']) show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show "V.fin_dim" by (rule fin_dim_kernel_berlekamp) qed context fixes P assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" begin interpretation RV: vec_space "TYPE('a mod_ring)" "card P" . lemma Berlekamp_subspace_eq_dim_vec: "Berlekamp_subspace.dim = RV.dim" proof - obtain f where lm_f: "linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f" and bij_f: "bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" using exists_bijective_linear_map_W_vec[OF finite_P u_desc_square_free P] by blast show ?thesis proof (rule linear_map.dim_eq[OF lm_f Berlekamp_subspace_fin_dim]) show "inj_on f (carrier W)" by (rule bij_betw_imp_inj_on[OF bij_f]) show " f ` carrier W = carrier RV.V" using bij_f unfolding bij_betw_def by auto qed qed lemma Berlekamp_subspace_dim: "Berlekamp_subspace.dim = card P" using Berlekamp_subspace_eq_dim_vec RV.dim_is_n by simp corollary card_berlekamp_basis_number_factors: "card (set (berlekamp_basis u)) = card P" unfolding Berlekamp_subspace_dim[symmetric] by (rule Berlekamp_subspace.dim_basis[symmetric], auto simp add: berlekamp_basis_basis) lemma length_berlekamp_basis_numbers_factors: "length (berlekamp_basis u) = card P" using card_set_berlekamp_basis card_berlekamp_basis_number_factors by auto end end end end context assumes "SORT_CONSTRAINT('a :: prime_card)" begin context fixes f :: "'a mod_ring poly" and n assumes sf: "square_free f" and n: "n = length (berlekamp_basis f)" and monic_f: "monic f" begin lemma berlekamp_basis_length_factorization: assumes f: "f = prod_list us" and d: "\ u. u \ set us \ degree u > 0" shows "length us \ n" proof (cases "degree f = 0") case True have "us = []" proof (rule ccontr) assume "us \ []" from this obtain u where u: "u \ set us" by fastforce hence deg_u: "degree u > 0" using d by auto have "degree f = degree (prod_list us)" unfolding f .. also have "... = sum_list (map degree us)" proof (rule degree_prod_list_eq) fix p assume p: "p \ set us" show "p \ 0" using d[OF p] degree_0 by auto qed also have " ... \ degree u" by (simp add: member_le_sum_list u) finally have "degree f > 0" using deg_u by auto thus False using True by auto qed thus ?thesis by simp next case False hence f_not_0: "f \ 0" using degree_0 by fastforce obtain P where fin_P: "finite P" and f_P: "f = \P" and P: "P \ {p. irreducible p \ monic p}" using monic_square_free_irreducible_factorization[OF monic_f sf] by auto have n_card_P: "n = card P" using P False f_P fin_P length_berlekamp_basis_numbers_factors n by blast have distinct_us: "distinct us" using d f sf square_free_prod_list_distinct by blast let ?us'="(map normalize us)" have distinct_us': "distinct ?us'" proof (auto simp add: distinct_map) show "distinct us" by (rule distinct_us) show "inj_on normalize (set us)" proof (auto simp add: inj_on_def, rule ccontr) fix x y assume x: "x \ set us" and y: "y \ set us" and n: "normalize x = normalize y" and x_not_y: "x \ y" from normalize_eq_imp_smult[OF n] obtain c where c0: "c \ 0" and y_smult: "y = smult c x" by blast have sf_xy: "square_free (x*y)" proof (rule square_free_factor[OF _ sf]) have "x*y = prod_list [x,y]" by simp also have "... dvd prod_list us" by (rule prod_list_dvd_prod_list_subset, auto simp add: x y x_not_y distinct_us) also have "... = f" unfolding f .. finally show "x * y dvd f" . qed have "x * y = smult c (x*x)" using y_smult mult_smult_right by auto hence sf_smult: "square_free (smult c (x*x))" using sf_xy by auto have "x*x dvd (smult c (x*x))" by (simp add: dvd_smult) hence "\ square_free (smult c (x*x))" by (metis d square_free_def x) thus "False" using sf_smult by contradiction qed qed have length_us_us': "length us = length ?us'" by simp have f_us': "f = prod_list ?us'" proof - have "f = normalize f" using monic_f f_not_0 by (simp add: normalize_monic) also have "... = prod_list ?us'" by (unfold f, rule prod_list_normalize[of us]) finally show ?thesis . qed have "\Q. prod_list Q = prod_list ?us' \ length ?us' \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule exists_factorization_prod_list) show "degree (prod_list ?us') > 0" using False f_us' by auto show "square_free (prod_list ?us')" using f_us' sf by auto fix u assume u: "u \ set ?us'" have u_not0: "u \ 0" using d u degree_0 by fastforce have "degree u > 0" using d u by auto moreover have "monic u" using u monic_normalize[OF u_not0] by auto ultimately show "degree u > 0 \ monic u" by simp qed from this obtain Q where Q_us': "prod_list Q = prod_list ?us'" and length_us'_Q: "length ?us' \ length Q" and Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast have distinct_Q: "distinct Q" proof (rule square_free_prod_list_distinct) show "square_free (prod_list Q)" using Q_us' f_us' sf by auto show "\u. u \ set Q \ degree u > 0" using Q irreducible_degree_field by auto qed have set_Q_P: "set Q = P" proof (rule monic_factorization_uniqueness) show "\(set Q) = \P" using Q_us' by (metis distinct_Q f_P f_us' list.map_ident prod.distinct_set_conv_list) qed (insert P Q fin_P, auto) hence "length Q = card P" using distinct_Q distinct_card by fastforce have "length us = length ?us'" by (rule length_us_us') also have "... \ length Q" using length_us'_Q by auto also have "... = card (set Q)" using distinct_card[OF distinct_Q] by simp also have "... = card P" using set_Q_P by simp finally show ?thesis using n_card_P by simp qed lemma berlekamp_basis_irreducible: assumes f: "f = prod_list us" and n_us: "length us = n" and us: "\ u. u \ set us \ degree u > 0" and u: "u \ set us" shows "irreducible u" proof (fold irreducible_connect_field, intro irreducible\<^sub>dI[OF us[OF u]]) fix q r :: "'a mod_ring poly" assume dq: "degree q > 0" and qu: "degree q < degree u" and dr: "degree r > 0" and uqr: "u = q * r" with us[OF u] have q: "q \ 0" and r: "r \ 0" by auto from split_list[OF u] obtain xs ys where id: "us = xs @ u # ys" by auto let ?us = "xs @ q # r # ys" have f: "f = prod_list ?us" unfolding f id uqr by simp { fix x assume "x \ set ?us" with us[unfolded id] dr dq have "degree x > 0" by auto } from berlekamp_basis_length_factorization[OF f this] have "length ?us \ n" by simp also have "\ = length us" unfolding n_us by simp also have "\ < length ?us" unfolding id by simp finally show False by simp qed end lemma not_irreducible_factor_yields_prime_factors: assumes uf: "u dvd (f :: 'b :: {field_gcd} poly)" and fin: "finite P" and fP: "f = \P" and P: "P \ {q. irreducible q \ monic q}" and u: "degree u > 0" "\ irreducible u" shows "\ pi pj. pi \ P \ pj \ P \ pi \ pj \ pi dvd u \ pj dvd u" proof - from finite_distinct_list[OF fin] obtain ps where Pps: "P = set ps" and dist: "distinct ps" by auto have fP: "f = prod_list ps" unfolding fP Pps using dist by (simp add: prod.distinct_set_conv_list) note P = P[unfolded Pps] have "set ps \ P" unfolding Pps by auto from uf[unfolded fP] P dist this show ?thesis proof (induct ps) case Nil with u show ?case using divides_degree[of u 1] by auto next case (Cons p ps) from Cons(3) have ps: "set ps \ {q. irreducible q \ monic q}" by auto from Cons(2) have dvd: "u dvd p * prod_list ps" by simp obtain k where gcd: "u = gcd p u * k" by (meson dvd_def gcd_dvd2) from Cons(3) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of u] *(2) have "gcd p u = 1 \ gcd p u = p" by auto thus ?case proof assume "gcd p u = 1" then have "Rings.coprime p u" by (rule gcd_eq_1_imp_coprime) with dvd have "u dvd prod_list ps" using coprime_dvd_mult_right_iff coprime_imp_coprime by blast from Cons(1)[OF this ps] Cons(4-5) show ?thesis by auto next assume "gcd p u = p" with gcd have upk: "u = p * k" by auto hence p: "p dvd u" by auto from dvd[unfolded upk] *(3) have kps: "k dvd prod_list ps" by auto from dvd u * have dk: "degree k > 0" by (metis gr0I irreducible_mult_unit_right is_unit_iff_degree mult_zero_right upk) from ps kps have "\ q \ set ps. q dvd k" proof (induct ps) case Nil with dk show ?case using divides_degree[of k 1] by auto next case (Cons p ps) from Cons(3) have dvd: "k dvd p * prod_list ps" by simp obtain l where gcd: "k = gcd p k * l" by (meson dvd_def gcd_dvd2) from Cons(2) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of k] *(2) have "gcd p k = 1 \ gcd p k = p" by auto thus ?case proof assume "gcd p k = 1" with dvd have "k dvd prod_list ps" by (metis dvd_triv_left gcd_greatest_mult mult.left_neutral) from Cons(1)[OF _ this] Cons(2) show ?thesis by auto next assume "gcd p k = p" with gcd have upk: "k = p * l" by auto hence p: "p dvd k" by auto thus ?thesis by auto qed qed then obtain q where q: "q \ set ps" and dvd: "q dvd k" by auto from dvd upk have qu: "q dvd u" by auto from Cons(4) q have "p \ q" by auto thus ?thesis using q p qu Cons(5) by auto qed qed qed lemma berlekamp_factorization_main: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and vs: "vs = vs1 @ vs2" and vsf: "vs = berlekamp_basis f" and n_bb: "n = length (berlekamp_basis f)" and n: "n = length us1 + n2" and us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" and us1: "\ u. u \ set us1 \ monic u \ irreducible u" and divs: "\ d. d \ set divs \ monic d \ degree d > 0" and vs1: "\ u v i. v \ set vs1 \ u \ set us1 \ set divs \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1,u}" and f: "f = prod_list (us1 @ divs)" and deg_f: "degree f > 0" and d: "\ g. g dvd f \ degree g = d \ irreducible g" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - have mon_f: "monic f" unfolding f by (rule monic_prod_list, insert divs us1, auto) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \ P" "P \ {q. irreducible q \ monic q}" by auto hence f0: "f \ 0" by auto show ?thesis using vs n us divs f us1 vs1 proof (induct vs2 arbitrary: divs n2 us1 vs1) case (Cons v vs2) show ?case proof (cases "v = 1") case False from Cons(2) vsf have v: "v \ set (berlekamp_basis f)" by auto from berlekamp_basis_eq_8[OF this] have vf: "[v ^ CARD('a) = v] (mod f)" . let ?gcd = "\ u i. gcd u (v - [:of_int i:])" let ?gcdn = "\ u i. gcd u (v - [:of_nat i:])" let ?map = "\ u. (map (\ i. ?gcd u i) [0 ..< CARD('a)])" define udivs where "udivs \ \ u. filter (\ w. w \ 1) (?map u)" { obtain xs where xs: "[0.. u. [w. i \ [0 ..< CARD('a)], w \ [?gcd u i], w \ 1])" unfolding udivs_def xs by (intro ext, auto simp: o_def, induct xs, auto) } note udivs_def' = this define facts where "facts \ [ w . u \ divs, w \ udivs u]" { fix u assume u: "u \ set divs" then obtain bef aft where divs: "divs = bef @ u # aft" by (meson split_list) from Cons(5)[OF u] have mon_u: "monic u" by simp have uf: "u dvd f" unfolding Cons(6) divs by auto from vf uf have vu: "[v ^ CARD('a) = v] (mod u)" by (rule cong_dvd_modulus_poly) from square_free_factor[OF uf sf_f] have sf_u: "square_free u" . let ?g = "?gcd u" from mon_u have u0: "u \ 0" by auto have "u = (\c\UNIV. gcd u (v - [:c:]))" using Berlekamp_gcd_step[OF vu mon_u sf_u] . also have "\ = (\i \ {0..< int CARD('a)}. ?g i)" by (rule sym, rule prod.reindex_cong[OF to_int_mod_ring_hom.inj_f range_to_int_mod_ring[symmetric]], simp add: of_int_of_int_mod_ring) finally have u_prod: "u = (\i \ {0..< int CARD('a)}. ?g i)" . let ?S = "{0.. ?S" hence "?g i \ 1" by auto moreover have mgi: "monic (?g i)" by (rule poly_gcd_monic, insert u0, auto) ultimately have "degree (?g i) > 0" using monic_degree_0 by blast note this mgi } note gS = this have int_set: "int ` set [0.. ?S" and j: "j \ ?S" and gij: "?g i = ?g j" show "i = j" proof (rule ccontr) define S where "S = {0.. S" "j \ S" "finite S" using i j unfolding S_def by auto assume ij: "i \ j" have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = ?g i * ?g j * (\i \ S. ?g i)" unfolding id using S ij by auto also have "\ = ?g i * ?g i * (\i \ S. ?g i)" unfolding gij by simp finally have dvd: "?g i * ?g i dvd u" unfolding dvd_def by auto with sf_u[unfolded square_free_def, THEN conjunct2, rule_format, OF gS(1)[OF i]] show False by simp qed qed have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = (\i \ ?S. ?g i)" by (rule sym, rule prod.setdiff_irrelevant, auto) also have "\ = \ (set (udivs u))" unfolding udivs_def set_filter set_map by (rule sym, rule prod.reindex_cong[of ?g, OF inj _ refl], auto simp: int_set[symmetric]) finally have u_udivs: "u = \(set (udivs u))" . { fix w assume mem: "w \ set (udivs u)" then obtain i where w: "w = ?g i" and i: "i \ ?S" unfolding udivs_def set_filter set_map int_set by auto have wu: "w dvd u" by (simp add: w) let ?v = "\ j. v - [:of_nat j:]" define j where "j = nat i" from i have j: "of_int i = (of_nat j :: 'a mod_ring)" "j < CARD('a)" unfolding j_def by auto from gS[OF i, folded w] have *: "degree w > 0" "monic w" "w \ 0" by auto from w have "w dvd ?v j" using j by simp hence gcdj: "?gcdn w j = w" by (metis gcd.commute gcd_left_idem j(1) w) { fix j' assume j': "j' < CARD('a)" have "?gcdn w j' \ {1,w}" proof (rule ccontr) assume not: "?gcdn w j' \ {1,w}" with gcdj have neq: "int j' \ int j" by auto (* next step will yield contradiction to square_free u *) let ?h = "?gcdn w j'" from *(3) not have deg: "degree ?h > 0" using monic_degree_0 poly_gcd_monic by auto have hw: "?h dvd w" by auto have "?h dvd ?gcdn u j'" using wu using dvd_trans by auto also have "?gcdn u j' = ?g j'" by simp finally have hj': "?h dvd ?g j'" by auto from divides_degree[OF this] deg u0 have degj': "degree (?g j') > 0" by auto hence j'1: "?g j' \ 1" by auto with j' have mem': "?g j' \ set (udivs u)" unfolding udivs_def by auto from degj' j' have j'S: "int j' \ ?S" by auto from i j have jS: "int j \ ?S" by auto from inj_on_contraD[OF inj neq j'S jS] have neq: "w \ ?g j'" using w j by auto have cop: "\ coprime w (?g j')" using hj' hw deg by (metis coprime_not_unit_not_dvd poly_dvd_1 Nat.neq0_conv) obtain w' where w': "?g j' = w'" by auto from u_udivs sf_u have "square_free (\(set (udivs u)))" by simp from square_free_prodD[OF this finite_set mem mem'] cop neq show False by simp qed } from gS[OF i, folded w] i this have "degree w > 0" "monic w" "\ j. j < CARD('a) \ ?gcdn w j \ {1,w}" by auto } note udivs = this let ?is = "filter (\ i. ?g i \ 1) (map int [0 ..< CARD('a)])" have id: "udivs u = map ?g ?is" unfolding udivs_def filter_map o_def .. have dist: "distinct (udivs u)" unfolding id distinct_map proof (rule conjI[OF distinct_filter], unfold distinct_map) have "?S = set ?is" unfolding int_set[symmetric] by auto thus "inj_on ?g (set ?is)" using inj by auto qed (auto simp: inj_on_def) from u_udivs prod.distinct_set_conv_list[OF dist, of id] have "prod_list (udivs u) = u" by auto note udivs this dist } note udivs = this have facts: "facts = concat (map udivs divs)" unfolding facts_def by auto obtain lin nonlin where part: "List.partition (\ q. degree q = d) facts = (lin,nonlin)" by force from Cons(6) have "f = prod_list us1 * prod_list divs" by auto also have "prod_list divs = prod_list facts" unfolding facts using udivs(4) by (induct divs, auto) finally have f: "f = prod_list us1 * prod_list facts" . note facts' = facts { fix u assume u: "u \ set facts" from u[unfolded facts] obtain u' where u': "u' \ set divs" and u: "u \ set (udivs u')" by auto from u' udivs(1-2)[OF u' u] prod_list_dvd[OF u, unfolded udivs(4)[OF u']] have "degree u > 0" "monic u" "\ u' \ set divs. u dvd u'" by auto } note facts = this have not1: "(v = 1) = False" using False by auto have "us = us1 @ (if length divs = n2 then divs else let (lin, nonlin) = List.partition (\q. degree q = d) facts in lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding Cons(4) facts_def udivs_def' berlekamp_factorization_main.simps Let_def not1 if_False by (rule arg_cong[where f = "\ x. us1 @ x"], rule if_cong, simp_all) (* takes time *) hence res: "us = us1 @ (if length divs = n2 then divs else lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding part by auto show ?thesis proof (cases "length divs = n2") case False with res have us: "us = (us1 @ lin) @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin)" by auto from Cons(2) have vs: "vs = (vs1 @ [v]) @ vs2" by auto have f: "f = prod_list ((us1 @ lin) @ nonlin)" unfolding f using prod_list_partition[OF part] by simp { fix u assume "u \ set ((us1 @ lin) @ nonlin)" with part have "u \ set facts \ set us1" by auto with facts Cons(7) have "degree u > 0" by (auto simp: irreducible_degree_field) } note deg = this from berlekamp_basis_length_factorization[OF sf_f n_bb mon_f f deg, unfolded Cons(3)] have "n2 \ length lin" by auto hence n: "n = length (us1 @ lin) + (n2 - length lin)" unfolding Cons(3) by auto show ?thesis proof (rule Cons(1)[OF vs n us _ f]) fix u assume "u \ set nonlin" with part have "u \ set facts" by auto from facts[OF this] show "monic u \ degree u > 0" by auto next fix u assume u: "u \ set (us1 @ lin)" { assume *: "\ (monic u \ irreducible\<^sub>d u)" with Cons(7) u have "u \ set lin" by auto with part have uf: "u \ set facts" and deg: "degree u = d" by auto from facts[OF uf] obtain u' where "u' \ set divs" and uu': "u dvd u'" by auto from this(1) have "u' dvd f" unfolding Cons(6) using prod_list_dvd[of u'] by auto with uu' have "u dvd f" by (rule dvd_trans) from facts[OF uf] d[OF this deg] * have False by auto } thus "monic u \ irreducible u" by auto next fix w u i assume w: "w \ set (vs1 @ [v])" and u: "u \ set (us1 @ lin) \ set nonlin" and i: "i < CARD('a)" from u part have u: "u \ set us1 \ set facts" by auto show "gcd u (w - [:of_nat i:]) \ {1, u}" proof (cases "u \ set us1") case True from Cons(7)[OF this] have "monic u" "irreducible u" by auto thus ?thesis by (rule monic_irreducible_gcd) next case False with u have u: "u \ set facts" by auto show ?thesis proof (cases "w = v") case True from u[unfolded facts'] obtain u' where u: "u \ set (udivs u')" and u': "u' \ set divs" by auto from udivs(3)[OF u' u i] show ?thesis unfolding True . next case False with w have w: "w \ set vs1" by auto from u obtain u' where u': "u' \ set divs" and dvd: "u dvd u'" using facts(3)[of u] dvd_refl[of u] by blast from w have "w \ set vs1 \ w = v" by auto from facts(1-2)[OF u] have u: "monic u" by auto from Cons(8)[OF w _ i] u' have "gcd u' (w - [:of_nat i:]) \ {1, u'}" by auto with dvd u show ?thesis by (rule monic_gcd_dvd) qed qed qed next case True with res have us: "us = us1 @ divs" by auto from Cons(3) True have n: "n = length us" unfolding us by auto show ?thesis unfolding us[symmetric] proof (intro conjI ballI) show f: "f = prod_list us" unfolding us using Cons(6) by simp { fix u assume "u \ set us" hence "degree u > 0" using Cons(5) Cons(7)[unfolded irreducible\<^sub>d_def] unfolding us by (auto simp: irreducible_degree_field) } note deg = this fix u assume u: "u \ set us" thus "monic u" unfolding us using Cons(5) Cons(7) by auto show "irreducible u" by (rule berlekamp_basis_irreducible[OF sf_f n_bb mon_f f n[symmetric] deg u]) qed qed next case True (* v = 1 *) with Cons(4) have us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" by simp from Cons(2) True have vs: "vs = (vs1 @ [1]) @ vs2" by auto show ?thesis proof (rule Cons(1)[OF vs Cons(3) us Cons(5-7)], goal_cases) case (3 v u i) show ?case proof (cases "v = 1") case False with 3 Cons(8)[of v u i] show ?thesis by auto next case True hence deg: "degree (v - [: of_nat i :]) = 0" by (metis (no_types, opaque_lifting) degree_pCons_0 diff_pCons diff_zero pCons_one) from 3(2) Cons(5,7)[of u] have "monic u" by auto from gcd_monic_constant[OF this deg] show ?thesis . qed qed qed next case Nil with vsf have vs1: "vs1 = berlekamp_basis f" by auto from Nil(3) have us: "us = us1 @ divs" by auto from Nil(4,6) have md: "\ u. u \ set us \ monic u \ degree u > 0" unfolding us by (auto simp: irreducible_degree_field) from Nil(7)[unfolded vs1] us have no_further_splitting_possible: "\ u v i. v \ set (berlekamp_basis f) \ u \ set us \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1, u}" by auto from Nil(5) us have prod: "f = prod_list us" by simp show ?case proof (intro conjI ballI) fix u assume u: "u \ set us" from md[OF this] have mon_u: "monic u" and deg_u: "degree u > 0" by auto from prod u have uf: "u dvd f" by (simp add: prod_list_dvd) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \P" "P \ {q. irreducible q \ monic q}" by auto show "irreducible u" proof (rule ccontr) assume irr_u: "\ irreducible u" from not_irreducible_factor_yields_prime_factors[OF uf P deg_u this] obtain pi pj where pij: "pi \ P" "pj \ P" "pi \ pj" "pi dvd u" "pj dvd u" by blast from exists_vector_in_Berlekamp_basis_dvd[OF deg_f berlekamp_basis_basis[OF deg_f, folded vs1] finite_set P pij(1-3) mon_f sf_f irr_u uf mon_u pij(4-5), unfolded vs1] obtain v s where v: "v \ set (berlekamp_basis f)" and gcd: "gcd u (v - [:s:]) \ {1,u}" using is_unit_gcd by auto from surj_of_nat_mod_ring[of s] obtain i where i: "i < CARD('a)" and s: "s = of_nat i" by auto from no_further_splitting_possible[OF v u i] gcd[unfolded s] show False by auto qed qed (insert prod md, auto) qed qed lemma berlekamp_monic_factorization: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and us: "berlekamp_monic_factorization d f = us" and d: "\ g. g dvd f \ degree g = d \ irreducible g" and deg: "degree f > 0" and mon: "monic f" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - from us[unfolded berlekamp_monic_factorization_def Let_def] deg have us: "us = [] @ berlekamp_factorization_main d [f] (berlekamp_basis f) (length (berlekamp_basis f))" by (auto) have id: "berlekamp_basis f = [] @ berlekamp_basis f" "length (berlekamp_basis f) = length [] + length (berlekamp_basis f)" "f = prod_list ([] @ [f])" by auto show "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" by (rule berlekamp_factorization_main[OF sf_f id(1) refl refl id(2) us _ _ _ id(3)], insert mon deg d, auto) qed end end diff --git a/thys/CVP_Hardness/BHLE.thy b/thys/CVP_Hardness/BHLE.thy --- a/thys/CVP_Hardness/BHLE.thy +++ b/thys/CVP_Hardness/BHLE.thy @@ -1,1707 +1,1706 @@ theory BHLE imports Main Additional_Lemmas begin section \Bounded Homogeneous Linear Equation Problem\ definition bhle :: "(int vec * int) set" where "bhle \ {(a,k). \x. a \ x = 0 \ dim_vec x = dim_vec a \ dim_vec a > 0 \ x \ 0\<^sub>v (dim_vec x) \ \x\\<^sub>\ \ k}" text \Reduction of bounded homogeneous linear equation to partition problem\ text \For the reduction function, one defines five-tuples for every element in a. The last tuple plays an important role. It consists only of four elements in order to add constraints to the other tuples. These values are formed in a way such that they form all constraints needed for the reductions. Note, that some indices have been changed with respect to \cite{EmBo81} to enable better indexing in the vectors/lists.\ definition b1 :: "nat \ int \ int \ int" where "b1 i M a = a + M * (5^(4*i-4) + 5^(4*i-3) + 5^(4*i-1))" definition b2 :: "nat \ int \ int" where "b2 i M = M * (5^(4*i-3) + 5^(4*i))" definition b2_last :: "nat \ int \ int" where "b2_last i M = M * (5^(4*i-3) + 1)" definition b3 :: "nat \ int \ int" where "b3 i M = M * (5^(4*i-4) + 5^(4*i-2))" definition b4 :: "nat \ int \ int \ int" where "b4 i M a = a + M * (5^(4*i-2) + 5^(4*i-1) + 5^(4*i))" definition b4_last :: "nat \ int \ int \ int" where "b4_last i M a = a + M * (5^(4*i-2) + 5^(4*i-1) + 1)" definition b5 :: "nat \ int \ int" where "b5 i M = M * (5^(4*i-1))" text \Change order of indices in \cite{EmBo81} such that $b3$ is in last place and can be omitted in the last entry. This ensures that the weight of the solution is $1$ or $-1$, essential for the proof of NP-hardnes.\ definition b_list :: "int list \ nat \ int \ int list" where "b_list as i M = [b1 (i+1) M (as!i), b2 (i+1) M, b4 (i+1) M (as!i), b5 (i+1) M, b3 (i+1) M]" definition b_list_last :: "int list \ nat \ int \ int list" where "b_list_last as n M = [b1 n M (last as), b2_last n M, b4_last n M (last as), b5 n M]" definition gen_bhle :: "int list \ int vec" where "gen_bhle as = (let M = 2*(\ias!i\)+1; n = length as in vec_of_list (concat (map (\i. b_list as i M) [0..0 then b_list_last as n M else [])))" text \The reduction function.\ definition reduce_bhle_partition:: "(int list) \ ((int vec) * int)" where "reduce_bhle_partition \ (\ a. (gen_bhle a, 1))" text \Lemmas for proof\ lemma dim_vec_gen_bhle: assumes "as\[]" shows "dim_vec (gen_bhle as) = 5 * (length as) - 1" using assms proof (induct as rule: list_nonempty_induct) case (single x) then show ?case unfolding gen_bhle_def Let_def b_list_def b_list_last_def by auto next case (cons x xs) define M where "M = (2 * (\i(x # xs) ! i\) + 1)" then show ?case using cons unfolding gen_bhle_def b_list_def b_list_last_def Let_def M_def[symmetric] by (subst dim_vec_of_list)+ (use length_concat_map_5[of "(\i. b1 (i + 1) M ((x#xs) ! i))" "(\i. b2 (i + 1) M )" "(\i. b4 (i + 1) M ((x#xs) ! i))" "(\i. b5 (i + 1) M )" "(\i. b3 (i + 1) M )"] in \simp\) qed lemma dim_vec_gen_bhle_empty: "dim_vec (gen_bhle []) = 0" unfolding gen_bhle_def by auto text \Lemmas about length\ lemma length_b_list: "length (b_list a i M) = 5" unfolding b_list_def by auto lemma length_b_list_last: "length (b_list_last a n M) = 4" unfolding b_list_last_def by auto lemma length_concat_map_b_list: "length (concat (map (\i. b_list as i M) [0..Last values of \gen_bhle\\ lemma gen_bhle_last0: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5) = b1 (length as) (2*(\ias!i\)+1) (last as)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_splits, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed lemma gen_bhle_last1: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5 + 1) = b2_last (length as) (2*(\ias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed text \The third entry of the last tuple is omitted, thus we skip one lemma\ lemma gen_bhle_last3: assumes "length as > 0" shows "(gen_bhle as) $ ((length as -1) * 5 + 2) = b4_last (length as) (2*(\ias!i\)+1) (last as)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed lemma gen_bhle_last4: assumes "length as > 0" shows "(gen_bhle as) $ ((length as-1) * 5 + 3) = b5 (length as) (2*(\ias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 then show ?case using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (auto split: if_splits simp add: b_list_last_def) qed text \Up to last values of \gen_bhle\\ lemma b_list_nth: assumes "ii. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [0..i. b_list as i M) [i..i. b_list as i M) [i..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [0..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [i+1..i. b_list as i M) [0..i. b_list as i M) [i + 1..i. b_list as i M) [i+1..i. b_list as i M) [i+1..i. b_list as i M) [0..ias!i\)+1) (as!i)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth0[OF assms, of M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_1: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 1 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_4: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 4 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_2: assumes "iias!i\)+1) (as!i)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 2 M], auto split: if_splits, simp add: b_list_def) qed lemma gen_bhle_3: assumes "iias!i\)+1)" unfolding gen_bhle_def Let_def proof (subst vec_of_list_append, subst index_append_vec(1), goal_cases) case 1 then show ?case using assms by (subst dim_vec_of_list)+ (split if_split, subst length_b_list_last, subst length_concat_map_b_list, auto) next case 2 define M where "M = (2 * (\ias ! i\) + 1)" then show ?case unfolding M_def[symmetric] using assms by (subst dim_vec_of_list, subst length_concat_map_b_list, subst vec_index_vec_of_list)+ (subst b_list_nth[OF assms, of 3 M], auto split: if_splits, simp add: b_list_def) qed text \Well-definedness of reduction function\ lemma well_defined_reduction_subset_sum: assumes "a \ partition_problem_nonzero" shows "reduce_bhle_partition a \ bhle" using assms unfolding partition_problem_nonzero_def reduce_bhle_partition_def bhle_def proof (safe, goal_cases) case (1 I) text \Given a subset $I$, we must construct a vector $x$ such that the properties of \bhle\ hold.\ have "finite I" using 1 by (meson subset_eq_atLeast0_lessThan_finite) have "length a > 0" using \a\[]\ by auto define n where "n = length a" define minus_x::"int list" where "minus_x = [0,0,-1,1,1]" define plus_x::"int list" where "plus_x = [1,-1,0,-1,0]" define plus_x_last::"int list" where "plus_x_last = [1,-1,0,-1]" define plus_minus::"int list" where "plus_minus = (if n-1\I then plus_x else minus_x)" define minus_plus::"int list" where "minus_plus = (if n-1\I then minus_x else plus_x)" define x::"int vec" where "x = vec_of_list(concat (map (\i. if i\I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [0..length a > 0\ unfolding n_def[symmetric] by auto have "0 < dim_vec x" unfolding dimx_eq_5dima using \length a > 0\ by linarith define M where "M = 2*(\ia!i\)+1" text \Some conditional lemmas for the proof.\ have x_nth: "x $ (i*5+j) = (if i\I then plus_minus ! j else minus_plus ! j)" if "ii. if i \ I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [0..i. if i\I then plus_minus else minus_plus) = (\i. [(if i\I then plus_minus!0 else minus_plus!0), (if i\I then plus_minus!1 else minus_plus!1), (if i\I then plus_minus!2 else minus_plus!2), (if i\I then plus_minus!3 else minus_plus!3), (if i\I then plus_minus!4 else minus_plus!4)])" unfolding plus_minus_def minus_plus_def plus_x_def minus_x_def by auto then show ?thesis unfolding if_rew length_concat_map_5[of "(\i. if i\I then plus_minus!0 else minus_plus!0)" "(\i. if i\I then plus_minus!1 else minus_plus!1)" "(\i. if i\I then plus_minus!2 else minus_plus!2)" "(\i. if i\I then plus_minus!3 else minus_plus!3)" "(\i. if i\I then plus_minus!4 else minus_plus!4)" "[0.. int list" using that(1) by (metis append_Nil map_append not_gr_zero upt_0 upt_append) have "concat (map (\i. if i \ I then plus_minus else minus_plus) [0..i. if i \ I then plus_minus else minus_plus) [i..i. if i \ I then plus_minus else minus_plus) [0.. = (if i \ I then plus_minus!j else minus_plus!j)" proof - have concat_rewr: "concat (map f [i.. int list" using that(1) upt_conv_Cons by force have length_if: "length (if i \ I then plus_minus else minus_plus) = 5" using length_plus_minus length_minus_plus by auto show ?thesis unfolding concat_rewr nth_append length_if using \j<5\ by auto qed finally show ?thesis unfolding x_def by (subst vec_index_vec_of_list) (subst nth_append, use lt in \auto\) qed have x_nth0: "x $ (i*5) = (if i\I then plus_minus ! 0 else minus_plus ! 0)" if "ii. if i \ I then plus_minus else minus_plus) [0..j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" if "i\I-{length a-1}" "n-1\I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def plus_minus_def minus_plus_def plus_x_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_in_I_minus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" if "i\I-{length a-1}" "n-1\I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def plus_minus_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_not_in_I_plus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" if "i\{0..I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def minus_plus_def minus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_not_in_I_minus: "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" if "i\{0..I" for i proof - have "(\j=0..<5. (gen_bhle a) $ (i*5+j) * x $ (i*5+j)) = (gen_bhle a) $ (i*5) * x $ (i*5) + (gen_bhle a) $ (i*5+1) * x $ (i*5+1) + (gen_bhle a) $ (i*5+2) * x $ (i*5+2) + (gen_bhle a) $ (i*5+3) * x $ (i*5+3) + (gen_bhle a) $ (i*5+4) * x $ (i*5+4)" by (simp add: eval_nat_numeral) also have "\ = (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)" using that 1 n_def \length a > 0\ apply (subst gen_bhle_0[of i a], fastforce) apply (subst gen_bhle_1[of i a], fastforce) apply (subst gen_bhle_2[of i a], fastforce) apply (subst gen_bhle_3[of i a], fastforce) apply (subst gen_bhle_4[of i a], fastforce) apply (subst x_nth[of i], fastforce, fastforce)+ apply (subst x_nth0[of i], fastforce) apply (unfold M_def minus_plus_def plus_x_def) apply (simp add: eval_nat_numeral) done finally show ?thesis by auto qed have gen_bhle_last: "(\j=0..<4. (gen_bhle a) $ ((n-1)*5+j) * x $ ((n-1)*5+j)) = (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" proof - have "(\j=0..<4. (gen_bhle a) $ ((n-1)*5+j) * x $ ((n-1)*5+j)) = (gen_bhle a) $ ((n-1)*5) * x $ ((n-1)*5) + (gen_bhle a) $ ((n-1)*5+1) * x $ ((n-1)*5+1) + (gen_bhle a) $ ((n-1)*5+2) * x $ ((n-1)*5+2) + (gen_bhle a) $ ((n-1)*5+3) * x $ ((n-1)*5+3)" by (simp add: eval_nat_numeral) also have "\ = (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" using 1 n_def \length a > 0\ unfolding n_def apply (subst gen_bhle_last0[of a, OF \length a > 0\]) apply (subst gen_bhle_last1[of a, OF \length a > 0\]) apply (subst gen_bhle_last3[of a, OF \length a > 0\]) apply (subst gen_bhle_last4[of a, OF \length a > 0\]) apply (subst x_nth_last, simp)+ apply (subst x_nth0_last, simp add: n_def) apply (unfold M_def plus_x_last_def) apply (auto simp add: eval_nat_numeral last_conv_nth) done finally show ?thesis by auto qed text \The actual proof. \ have "(gen_bhle a) \ x = 0" proof - define f where "f = (\i. (\j = 0..<5. gen_bhle a $ (i*5+j) * x $ (i*5+j)))" have "(gen_bhle a) \ x = (\i = (\i<(n-1)*5. (gen_bhle a) $ i * x $ i) + (\i = (n-1)*5..<(n-1)*5 +4. (gen_bhle a) $ i * x $ i)" proof (subst split_sum_mid_less[of "(n-1)*5" "n*5-1"], goal_cases) case 1 then show ?case unfolding n_def using \0 < length a\ by linarith next case 2 have "n * 5 - 1 = (n-1) * 5 + 4" unfolding n_def using \0 < length a\ by linarith then show ?case by auto qed also have "\ = (\i = 0..j=0..<4. gen_bhle a $ ((n-1)*5+j) * x $ ((n-1)*5+j))" proof - have *: "(+) ((n - 1) * 5) ` {0..<4} = {(n-1)*5..<(n-1)*5+4}" by auto have "(\i = (n - 1) * 5..<(n - 1) * 5 + 4. gen_bhle a $ i * x $ i) = (\j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j))" using sum.reindex[of "(\j. (n-1)*5+j)" "{0..<4}" "(\i. gen_bhle a $ i * x $ i)"] unfolding comp_def * by auto then show ?thesis unfolding f_def lessThan_atLeast0 by (subst sum_split_idx_prod[of "(\i. (gen_bhle a) $ i * x $ i)" "n-1" 5], auto) qed also have "\ = (\i\I-{n-1}. f i) + (\i\{0..j=0..<4. gen_bhle a $ ((n-1)*5+j) * x $ ((n-1)*5+j))" proof - have "I - {n - 1} \ (({0..finite I\) qed also have "\ = M * (\i\{0..I") case True have "sum f (I - {n - 1}) + sum f ({0..j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j)) = (\i\I-{n-1}. (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)) + (\i\{0..i\I-{n-1}. f i) = (\i\I-{n-1}. (b1 (i+1) M (a!i)) - (b2 (i+1) M) - (b5 (i+1) M)) " unfolding f_def using gen_bhle_in_I_plus[OF _ True] by (simp add: n_def) moreover have "(\i\{0..i\{0.. = (\i\I-{n-1}. (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))) + (\i\{0..I-{n-1}" for i unfolding b1_def b2_def b5_def by (smt (verit, best) add_uminus_conv_diff right_diff_distrib) moreover have "b3 (i + 1) M - b4 (i + 1) M (a ! i) + b5 (i + 1) M = - (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))" if "i\{0.. = (\i\I-{n-1}. (a!i)) + (\i\{0..i\{0.. I = I - {n -1}" using "1"(1) n_def by auto have sets2: "{0..i\I-{n-1}. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1)) + (\i\{0..i = 0..i. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1))" "I"] \finite I\ unfolding sets1 sets2 by auto show ?thesis apply (subst distrib_left)+ apply (subst sum.distrib)+ apply (subst sum_distrib_left) apply (subst right_diff_distrib)+ apply (subst subs[symmetric]) apply auto done qed also have "\ = (\i\I. (a!i)) + (\i\{0..i\{0.. {n-1})" using True by auto have "sum ((!) a) (I - {n-1}) + a ! (n-1) = sum ((!) a) I" by (subst sum.Int_Diff[of "I" _ "{n-1}"], unfold *, auto simp add: \finite I\) then show ?thesis using True by auto qed also have "\ = M * (\i\{0..I*) next case False have "sum f (I - {n - 1}) + sum f ({0..j = 0..<4. gen_bhle a $ ((n - 1) * 5 + j) * x $ ((n - 1) * 5 + j)) = (\i\{0..i\I-{n-1}. (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)) + (b1 n M (a!(n-1))) - (b2_last n M) - (b5 n M)" proof - have "(\i\{0..i\{0..i\I-{n-1}. f i) = (\i\I-{n-1}. (b3 (i+1) M) - (b4 (i+1) M (a!i)) + (b5 (i+1) M)) " unfolding f_def using gen_bhle_in_I_minus[OF _ False] by (simp add: n_def) ultimately show ?thesis unfolding f_def using gen_bhle_last by auto qed also have "\ = (\i\{0..i\I-{n-1}. - (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))) + (a!(n-1)) + M * 5^(4*n-4) - M*1" proof - have "b1 (i + 1) M (a ! i) - b2 (i + 1) M - b5 (i + 1) M = (a!i) + (M * 5^(4*(i+1)-4) - M * 5^(4*(i+1)))" if "i\{0..I-{n-1}" for i unfolding b3_def b4_def b5_def by (smt (verit, best) add_uminus_conv_diff right_diff_distrib) moreover have "b1 n M (a ! (n - 1)) - b2_last n M - b5 n M = (a!(n-1)) + M * 5^(4*n-4) - M" unfolding b1_def b2_last_def b5_def by (simp add: distrib_left) moreover have "- b4_last n M (a ! (n - 1)) + b5 n M = -(a!(n-1)) - M * 5^(4*n-2) - M" unfolding b4_last_def b5_def by (simp add: distrib_left) ultimately show ?thesis by auto qed also have "\ = (\i\{0..i\I-{n-1}. - (a!i)) + M * (\i\{0.. I = I - {n -1}" using "1"(1) n_def by auto have sets2: "{0..i\{0..i\I-{n-1}. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1)) = (\i = 0..i. M * 5^(4*i+4*1-4) - M * 5^(4*i+4*1))" "I"] \finite I\ unfolding sets1 sets2 by auto show ?thesis apply (subst distrib_left)+ apply (subst sum.distrib)+ apply (subst sum_distrib_left) apply (subst right_diff_distrib)+ apply (subst subs[symmetric]) apply auto done qed also have "\ = (\i\{0..i\I. - (a!i)) + M * (\i\{0.. {n-1} = {n-1}" using False n_def \length a > 0\ by auto then have **: "a ! (n-1) = sum ((!) a) (({0.. {n-1})" using False by auto have "sum ((!) a) (({0..finite I\) then show ?thesis using False by auto qed also have "\ = M * (\i\{0.. = M * ((\i\{1..i = Suc 0..i. 5 ^ (4*(i+1) - 4) - 5 ^ (4*(i+1)))) {0..i. 5^(4*i-4) - 5^(4*i))" 0 "n-1"] unfolding comp_def by auto show ?thesis by (subst distrib_left[symmetric], subst right_diff_distrib, subst mult_1_right) (subst sums[symmetric], use \0 < length a\ n_def in \force\) qed also have "\ = M * (((\i\{1..i\{1..i. 5^(4*i-4))" "(\i. (-1) * 5^(4*i))" "{1.. = M * ((\i\{1..i\{0..i. 5^(4*i-4))"] sum.atLeast_Suc_lessThan[of 0 n "(\i. 5^(4*i))"] by (smt (verit, best) One_nat_def Suc_eq_plus1 Suc_leI \0 < length a\ mult_eq_0_iff n_def power_0) also have "\ = M * ((\i\{0..i\{0..i. 5^(4*i-4))" 0 n] by auto also have "\ = 0" by auto finally show ?thesis by blast qed moreover have "dim_vec x = dim_vec (gen_bhle a)" using dim_vec_gen_bhle[OF \a\[]\] dimx_eq_5dima by simp moreover have "x \ 0\<^sub>v (dim_vec x)" proof (rule ccontr) assume "\ x \ 0\<^sub>v (dim_vec x)" then have "x = 0\<^sub>v (dim_vec x)" by auto have "dim_vec x > 3" using \0 < length a\ dimx_eq_5dima by linarith have "(n - Suc 0) * 5 + 3 < dim_vec x" unfolding dimx_eq_5dima n_def using \length a > 0\ by linarith then have "x $ ((n-1)*5 + 3) = 0" using \dim_vec x > 3\ by (subst \x=0\<^sub>v (dim_vec x)\, subst index_zero_vec(1)[of "(n-1)*5+3"], auto) moreover have "x $ ((n-1)*5+3) \ 0" proof - have "\ ((n - 1) * 5 + 3 < (n - 1) * 5)" by auto then show ?thesis unfolding x_def vec_of_list_index nth_append length_concat_map plus_x_last_def by auto qed ultimately show False by auto qed moreover have "\x\\<^sub>\ \ 1" proof - let ?x_list = "(concat (map (\i. if i \ I then plus_minus else minus_plus) [0.. {-1,0,1::int}" using \length a > 0\ 1 unfolding n_def plus_minus_def minus_plus_def plus_x_def minus_x_def by (subst set_concat, subst set_map)(auto simp add: atLeast0LessThan) have "?x_list ! i \ {-1,0,1::int}" if "i< length ?x_list" for i using nth_mem[OF that] set by auto then have *:"?x_list ! i \ {-1,0,1::int}" if "i< (n - 1) * 5" for i using that unfolding length_concat_map by auto have **: "plus_x_last ! (i - (n - 1) * 5)\{-1,0,1::int}" if "\ (i<(n-1)*5)" "ilength a > 0\ unfolding dimx_eq_5dima n_def by auto then show ?thesis unfolding plus_x_last_def by (smt (z3) add.assoc add_diff_cancel_right' diff_is_0_eq insertCI less_Suc_eq not_le not_less_iff_gr_or_eq nth_Cons' numeral_1_eq_Suc_0 numeral_Bit0 plus_1_eq_Suc) qed have "x$i\{-1,0,1::int}" if "ix$i\\1" if "ii. dim_vec x > i)" 0] \dim_vec x > 0\) qed ultimately show ?case by (subst exI[of _ x], auto) qed text \NP-hardness of reduction function\ lemma NP_hardness_reduction_subset_sum: assumes "reduce_bhle_partition a \ bhle" shows "a \ partition_problem_nonzero" using assms unfolding reduce_bhle_partition_def bhle_def partition_problem_nonzero_def proof (safe, goal_cases) case (1 x) text \Given a vector $x$ from \bhle\ find the subset $I$ such that it has the same sum as its complement.\ have "length a > 0" using 1(3) dim_vec_gen_bhle dim_vec_gen_bhle_empty by auto define I where "I = {i\{0..0}" define n where "n = length a" then have "n > 0" using \length a>0\ by auto have dim_vec_x_5n: "dim_vec x = 5 * n - 1" unfolding n_def using 1 by (metis dim_vec_gen_bhle dim_vec_gen_bhle_empty less_not_refl2) have "I\{0..This is the trickiest part in the proof. One first has to generate equations from $x$ which form conditions on the entries of $x$. To do so, we consider the formula \gen_bhle a \ x = 0\ and rewrite it in basis $5$. Every digit then gives us an arithmetic expression in entries of $x$ equals to zero. From these equations, we can deduce the wanted properties.\ moreover have "sum ((!) a) I = sum ((!) a) ({0..ia ! i\) + 1" text \Rewriting the first formula in a huge sum.\ define a0 where "a0 = (\i. if i mod (5::nat) \ {0,2} then a!(i div 5) else 0)" define a1 where "a1 = (\i. if i mod (5::nat) \ {0,4} then 1 else 0::int)" define a1_last where "a1_last = (\i. if i mod (5::nat) \ {0} then 1 else 0::int)" define a2 where "a2 = (\i. if i mod (5::nat) \ {0,1} then 1 else 0::int)" define a3 where "a3 = (\i. if i mod (5::nat) \ {4,2} then 1 else 0::int)" define a3_last where "a3_last = (\i. if i mod (5::nat) \ {2} then 1 else 0::int)" define a4 where "a4 = (\i. if i mod (5::nat) \ {0,2,3} then 1 else 0::int)" define a5 where "a5 = (\i. if i mod (5::nat) \ {1,2} then (if i div 5 < n-1 then 5^(4*(i div 5 +1)) else 1) else 0::int)" define a0_rest' where "a0_rest' = (\i. a1 i * 5 ^ (4 * (i div 5)) + a2 i * 5 ^ (4 * (i div 5) + 1) + a3 i * 5 ^ (4 * (i div 5) + 2) + a4 i * 5 ^ (4 * (i div 5) + 3) + a5 i)" define a0_last where "a0_last = (\i. a1_last i * 5 ^ (4 * (i div 5)) + a2 i * 5 ^ (4 * (i div 5) + 1) + a3_last i * 5 ^ (4 * (i div 5) + 2) + a4 i * 5 ^ (4 * (i div 5) + 3) + a5 i)" define a0_rest where "a0_rest = (\i. if i div 5 < n-1 then a0_rest' i else a0_last i)" let ?P0 = "(\i. b1 (i div 5 + 1) M (a!(i div 5)))" let ?P1 = "(\i. (if i div 5 < n-1 then b2 (i div 5 + 1) M else b2_last (i div 5 + 1) M))" let ?P4 = "(\i. (if i div 5 < n-1 then b3 (i div 5 + 1) M else 0))" let ?P2 = "(\i. (if i div 5 < n-1 then b4 (i div 5 + 1) M (a!(i div 5)) else b4_last (i div 5 + 1) M (a!(i div 5))))" let ?P3 = "(\i. b5 (i div 5 + 1) M)" have cases_a0: "(a0 i + M * (a0_rest i)) = (if i mod 5 = 0 then ?P0 i else (if i mod 5 = 1 then ?P1 i else (if i mod 5 = 2 then ?P2 i else (if i mod 5 = 3 then ?P3 i else ?P4 i))))" if "i<5*n" for i proof (cases "i div 5 < n-1") case True then show ?thesis by (subst mod_5_if_split[of i "a0 i + M * (a0_rest i)" ?P0 ?P1 ?P2 ?P3 ?P4]) (auto simp add: a0_def a0_rest_def a0_rest'_def a1_def a2_def a3_def a4_def a5_def b1_def b2_def b3_def b4_def b5_def) next case False then have "i div 5 = n-1" using that by auto then show ?thesis by (subst mod_5_if_split[of i "a0 i + M * (a0_rest i)" ?P0 ?P1 ?P2 ?P3 ?P4]) (auto simp add: False a0_def a0_rest_def a0_last_def a1_def a1_last_def a2_def a3_def a3_last_def a4_def a5_def b1_def b2_def b2_last_def b3_def b4_def b4_last_def b5_def) qed text \Splitting of the first part of the sum containing the $a_i$.\ have gen_bhle_nth: "gen_bhle a $ i = a0 i + M * (a0_rest i)" if "ilength a >0\ dim_vec_gen_bhle by auto have gen_bhle_subst: "gen_bhle a $ i = (concat (map (\i. b_list a i M) [0..length a > 0\ by auto have len_concat_map: "length ?concat_map = 5 * (n-1)"using length_concat_map_b_list . show ?thesis proof (cases "i div 5 < n-1") case True then have "i<5*(n-1)" by auto then have j_th: "(?concat_map @ ?last)! i = ?concat_map ! i" by (subst nth_append, subst len_concat_map, auto) have "concat (map (\i. b_list a i M) [0..i. b_list a i M) [0.. = (concat (map (\i. b_list a i M) [0..i. b_list a i M) [i div 5.. = concat (map (\i. b_list a i M) [i div 5..i. b_list a i M) [0..i. b_list a i M) [i div 5.. = (b_list a (i div 5) M @ concat (map (\i. b_list a i M) [i div 5 +1.. = b_list a (i div 5) M ! (i mod 5)" unfolding nth_append b_list_def by auto finally have i_th: "?concat_map ! i = b_list a (i div 5) M ! (i mod 5)" by auto show ?thesis apply(subst gen_bhle_subst, subst j_th, subst i_th, subst cases_a0) subgoal apply (use that dim_vec_gen_bhle n_def \length a > 0\ in \auto\) done subgoal apply (subst mod_5_if_split[of i "b_list a (i div 5) M ! (i mod 5)" ?P0 ?P1 ?P2 ?P3 ?P4]) apply (use True in \auto simp add: b_list_def that\) done done next case False then have "i \ 5*(n-1)" by auto then obtain j where "i = 5*(n-1) + j" and "j<4" using \i \length a > 0\ unfolding dim_gen_bhle n_def by (subst split_lower_plus_diff[of i "5*(length a-1)" "5*(length a)"], auto) have j_th: "(?concat_map @ ?last)! i = ?last ! j" using \i nth_append_length_plus[of ?concat_map ?last j] unfolding len_concat_map by (auto simp add: \i = 5*(n-1) + j\ \j<4\) have "j = i mod 5" using \i = 5*(n-1) + j\ \j<4\ by auto have "n = i div 5 + 1" using \i = 5*(n-1) + j\ \j<4\ \length a > 0\ n_def by auto then have "i div 5 = n-1" by auto have "last a = a ! (i div 5)" unfolding \i div 5 = n-1\ by (subst last_conv_nth[of a]) (use \length a > 0\ n_def in \auto\) have *: "i mod 5 = 4 \ [] ! 0 = 0" by (use \j < 4\ \j = i mod 5\ in \presburger\) show ?thesis apply(subst gen_bhle_subst, subst j_th, subst cases_a0) subgoal apply (use that dim_vec_gen_bhle n_def \length a > 0\ in \auto\) done subgoal apply (subst mod_5_if_split[of i "b_list_last a n M ! j" ?P0 ?P1 ?P2 ?P3 ?P4]) apply (auto simp add: b_list_last_def that \j = i mod 5\ \n = i div 5 + 1\ \last a = a ! (i div 5)\ *) done done qed qed have "gen_bhle a \ x = (\i<5*n-1. x $ i * (a0 i + M * a0_rest i))" using 1(1) gen_bhle_nth unfolding scalar_prod_def Let_def dim_vec_x_5n 1(2)[symmetric] by (smt (verit, best) lessThan_atLeast0 lessThan_iff mult.commute sum.cong) then have sum_gen_bhle: "(\i<5 * n-1. x $ i * (a0 i + M * a0_rest i)) = 0" using 1(1) by simp text \The first equation containing the $a_i$\ have eq_0: "(\ii<5*n-1. x$i * (a0_rest i)) = 0" proof - have *: "(\i<5*n-1. \a0 i\) < M" proof - have "5 * n - 1 = Suc (5 * n - 2)" using \n > 0\ by auto have "5 * n - 2 = Suc (5 * n - 3)" using \n > 0\ by auto have "5 * n - 3 = Suc (5 * n - 4)" using \n > 0\ by auto have "5 * n - 4 = Suc (5 * n - 5)" using \n > 0\ by auto have "(\i<5*n-1. \a0 i\) = (\i<5*(n-1). \a0 i\) + (\j<4.\a0 ((n-1)*5+j)\)" proof - have "5*n-5 = 5*(n-1)" by auto have "(\i<5*n-1. \a0 i\) = (\i<5*n-5. \a0 i\) + \a0 (5*(n-1))\ + \a0 (5*(n-1)+1)\ + \a0 (5*(n-1)+2)\ + \a0 (5*(n-1)+3)\" unfolding \5 * n - 1 = Suc (5 * n - 2)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 2"] unfolding \5 * n - 2 = Suc (5 * n - 3)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 3"] unfolding \5 * n - 3 = Suc (5 * n - 4)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 4"] unfolding \5 * n - 4 = Suc (5 * n - 5)\ sum.lessThan_Suc[of "(\i.\a0 i\)" "5 * n - 5"] unfolding \5*n-5 = 5*(n-1)\ by (auto simp add: Suc3_eq_add_3 add.commute) moreover have "\a0 (5*(n-1))\ + \a0 (5*(n-1)+1)\ + \a0 (5*(n-1)+2)\ + \a0 (5*(n-1)+3)\ = (\j<4.\a0 ((n-1)*5+j)\)" by (simp add:eval_nat_numeral) ultimately show ?thesis using \5*n-5 = 5*(n-1)\ by auto qed also have "\ = (\ij<5. \a0 (i*5+j)\)) + (\j<4.\a0 ((n-1)*5+j)\)" using sum_split_idx_prod[of "(\i. \a0 i\)" "n-1" 5] by (simp add: lessThan_atLeast0 mult.commute) also have "\ = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\))" proof - have "(5::nat) = Suc 4" by auto have "(\ij<5::nat. \a0 (i*5+j)\)) = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\))" by(rule sum.cong[OF refl], rule sum.cong[OF refl], auto simp add: a0_def div_mult_self3[of 5]) moreover have "(\j<4::nat.\a0 ((n-1)*5+j)\) = (\j<4::nat.\if j \ {0, 2} then a ! (n-1) else 0\)" by(rule sum.cong[OF refl], auto simp add: a0_def div_mult_self3[of 5]) moreover have "(\j<4::nat.\if j \ {0, 2} then a ! (n-1) else 0\) = (\j<5::nat.\if j \ {0, 2} then a ! (n-1) else 0\)" unfolding \5=Suc 4\ using sum.lessThan_Suc[of "(\j.\if j \ {0, 2} then a ! (n-1) else 0\)" "4::nat"] by auto ultimately have *: "(\ij<5. \a0 (i*5+j)\)) + (\j<4.\a0 ((n-1)*5+j)\) = (\ij<5::nat. \if j \ {0, 2} then a ! i else 0\)) + (\j<5::nat. \if j \ {0, 2} then a ! (n-1) else 0\)" by auto show ?thesis by (subst *, subst sum.lessThan_Suc[symmetric], auto simp add: \n>0\) qed also have "\ = (\ia ! i\)" by (auto simp add: eval_nat_numeral) also have "\ = 2 * (\ia ! i\)" by (simp add: sum_distrib_left) finally show ?thesis unfolding M_def n_def by linarith qed have **: "\i<5*n-1. \x $ i\ \ 1" using 1(5) by (metis dim_vec_x_5n order_trans vec_index_le_linf_norm) have "(\i<5*n-1. x $ i * a0 i) = 0 \ (\i<5*n-1. x$i * (a0_rest i)) = 0" using split_eq_system[OF * ** sum_gen_bhle] by auto moreover have "(\i<5*n-1. x $ i * a0 i) = (\ij = i*5..j = i*5.. = (\j<5. x$(i*5 + j) * (if j \ {0, 2} then a!i else 0))" using mod_mult_self3[of i 5] div_rule by (subst sum.reindex[of "(\j. i*5+j)" "{0..<5}"], simp, unfold comp_def) (metis (no_types, lifting) One_nat_def lessThan_atLeast0 lessThan_iff nat_mod_lem not_less_eq not_numeral_less_one sum.cong) also have "\ = x$(i*5 + 0) * a!i + x$(i*5 + 2) * a!i" by (auto simp add: eval_nat_numeral split: if_splits) finally show ?thesis by (simp add: distrib_left mult.commute) qed then have *: "(\i<(n-1)*5. x $ i * a0 i) = (\ij = 5*(n-1)..<5*(n-1)+4. x $ j * a0 j) = ?h (n-1)" proof - have div_rule: "((n-1) * 5 + xa) div 5 = (n-1)" if "xa <4" for xa using that by auto have "(\j = (n-1)*5..<(n-1)*5+4. ?g j) = sum ?g ((+) ((n-1) * 5) ` {0..<4})" by (simp add: add.commute) also have "\ = (\j<4. x$((n-1)*5 + j) * (if j \ {0, 2} then a!(n-1) else 0))" proof - have *: "(\xa = 0..<4.?g ((n - 1) * 5 + xa)) = (\j = 0..<4. x $ ((n - 1) * 5 + j) * (if j \ {0, 2} then a ! (n - 1) else 0))" (is "(\xa = 0..<4. ?g' xa) = (\j = 0..<4. ?h' j)") by (rule sum.cong) auto show ?thesis using mod_mult_self3[of "n-1" 4] div_rule by (subst sum.reindex[of "(\j. (n-1)*5+j)" "{0..<4}"], simp, unfold comp_def lessThan_atLeast0) (use * in \auto\) qed also have "\ = x$((n-1)*5 + 0) * a!(n-1) + x$((n-1)*5 + 2) * a!(n-1)" by (auto simp add: eval_nat_numeral split: if_splits) finally show ?thesis unfolding a0_def by (simp add: distrib_left mult.commute) qed have "5 * (n - 1) < 5 * n - 1" using \n> 0\ by auto then have ***: "(\i<5*n-1. x $ i * a0 i) = (\i<5*(n-1). x $ i * a0 i) + (\i=5*(n-1)..<5*n-1. x $ i * a0 i)" by (subst split_sum_mid_less[of "5*(n-1)" "5*n-1"], auto) have ****: "5 * n - 1 = 5 * (n - 1) + 4" using \n> 0\ by auto show ?thesis by (unfold ***, subst mult.commute[of 5 "n-1"], unfold * **** **) (subst sum.lessThan_Suc[of ?h "n-1", symmetric], auto simp add: \n>0\) qed ultimately show "(\ii<5*n-1. x$i * (a0_rest i)) = 0" by auto qed let ?eq_0'_left = "(\i<5*n-1. x$i * (a0_rest i))" interpret digits 5 by (simp add: digits_def) have digit_a0_rest: "digit ?eq_0'_left k = 0" for k using eq_0' by (simp add: eq_0' digit_altdef) text \Define the digits in basis 5.\ define d1 where "d1 = (\i. x$(i*5) + (if ii. x$(i*5) + x$(i*5+1))" define d3 where "d3 = (\i. (if ii. x$(i*5) + x$(i*5+2) + x$(i*5+3))" define d5 where "d5 = (\i. x$(5*i+1) + x$(5*i+2))" define d where "d = (\k. (if k mod 4 = 0 then (if k = 0 then d5 (n-1) + d1 0 else (d1 (k div 4) + d5 (k div 4 -1))) else (if k mod 4 = 1 then d2 (k div 4) else (if k mod 4 = 2 then d3 (k div 4) else d4 (k div 4)))))" text \Rewrite the sum in basis 5.\ have rewrite_digits: "(\i<5*n-1. x$i * (a0_rest i)) = (\k<4*n. d k * 5^k)" proof - define f1::"nat \ nat \ int" where "f1 = (\i j. ((if i {0, 4} then 1 else 0) else (if j\{0} then 1 else 0)) * 5 ^ (4 * i) + (if j \ {0, 1} then 1 else 0) * 5 ^ (4 * i + 1) + (if i {4, 2} then 1 else 0) else (if j\{2} then 1 else 0)) * 5 ^ (4 * i + 2) + (if j \ {0, 2, 3} then 1 else 0) * 5 ^ (4 * i + 3) + (if j \ {1, 2} then if i < n - 1 then 5 ^ (4 * (i + 1)) else 1 else 0)))" have "f1 (n-1) 4 = 0" unfolding f1_def by auto define f2 where "f2 = (\i. x $ (i * 5) * (5 ^ (4 * i) + 5 ^ (4 * i + 1) + 5 ^ (4 * i + 3)) + x $ (i * 5 + 1) * (5 ^ (4 * i + 1) + (if i < n - 1 then 5 ^ (4 * (i + 1)) else 1)) + x $ (i * 5 + 4) * (if ii. d1 i * (5 ^ (4 * i)) + d5 i * (if i < n - 1 then 5 ^ (4 * (i + 1)) else 1) + d2 i * 5 ^ (4 * i + 1) + d3 i * 5 ^ (4 * i + 2) + d4 i * 5 ^ (4 * i + 3))" have f2_f3: "f2 i = f3 i" if "i = f3 i" unfolding f3_def d1_def d2_def d3_def d4_def d5_def using True by (subst distrib_right)+ (simp add: mult.commute[of 5 i]) ultimately show ?thesis by auto next case False then have "f2 i = x $ (i * 5) * 5 ^ (4 * i) + x $ (i * 5) * 5 ^ (4 * i + 1) + x $ (i * 5) * 5 ^ (4 * i + 3) + x $ (i * 5 + 1) * 5 ^ (4 * i + 1) + x $ (i * 5 + 1) + x $ (i * 5 + 2) * 5 ^ (4 * i + 2) + x $ (i * 5 + 2) * 5 ^ (4 * i + 3) + x $ (i * 5 + 2) + x $ (i * 5 + 3) * 5 ^ (4 * i + 3)" unfolding f2_def by (subst ring_distribs)+ simp also have "\ = f3 i" unfolding f3_def d1_def d2_def d3_def d4_def d5_def using False by(simp add: algebra_simps) ultimately show ?thesis by auto qed define x_pad where "x_pad = (\i. if i<5*n-1 then x$i else 1)" have pad: "(\i<5*n-1. x$i * (a0_rest i)) = (\i<5*n. x_pad i * (a0_rest i))" proof - have "Suc (5 * n - 1) = 5*n" using \n>0\ by auto have "(\i<5 * n - 1. x_pad i * a0_rest i) = (\i<5 * n - 1. x$i * a0_rest i)" by(rule sum.cong)(auto simp: x_pad_def) moreover have "x_pad (5 * n - 1) * a0_rest (5 * n - 1) = 0" proof - have "\((5 * n - 1) div 5 < n - 1)" using \n>0\ by auto moreover have "(5 * n - 1) mod 5 = 4" proof - have "5 * n - 1 = 5*(n-1)+4" using \n>0\ by auto show ?thesis unfolding \5 * n - 1 = 5*(n-1)+4\ by auto qed ultimately show ?thesis unfolding a0_rest_def a0_last_def a1_last_def a2_def a3_last_def a4_def a5_def by auto qed ultimately show ?thesis using sum.lessThan_Suc[of "(\i. x_pad i * (a0_rest i))" "5 * n - 1"] unfolding \Suc (5 * n - 1) = 5*n\ by auto qed have *: "(\i<5*n. x_pad i * (a0_rest i)) = (\ij<5. x_pad (i*5+j) * (a0_rest (i*5+j))))" (is "\ = (\ij<5. ?f0 i j))") using sum_split_idx_prod[of "(\i. x_pad i * a0_rest i)" n 5] unfolding mult.commute[of n 5] using atLeast0LessThan by auto also have "\ = (\ij<5. x_pad (i * 5 + j) * f1 i j)" proof (subst sum.cong[of _ _ "(\i. (\j<(5::nat). ?f0 i j))" "(\i. (\j<(5::nat). x_pad (i * 5 + j) * f1 i j))", symmetric], simp, goal_cases) case (1 i) have **: "a0_rest (i * 5 + j) = f1 i j" if "j<5" for j using that lt_5_split[of j] 1 unfolding f1_def a0_rest_def a0_rest'_def a0_last_def a1_def a1_last_def a2_def a3_def a3_last_def a4_def a5_def by auto show ?case by (rule sum.cong) (use ** 1 in auto) next case 2 then show ?case using * by auto qed also have "\ = (\ij<5. x_pad (i * 5 + j) * f1 i j) = x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 + x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 + x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 + x_pad ((n-1) * 5 + 3) * f1 (n-1) 3 + x_pad ((n-1) * 5 + 4) * f1 (n-1) 4" by (simp add: eval_nat_numeral) also have "\ = x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 + x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 + x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 + x_pad ((n-1) * 5 + 3) * f1 (n-1) 3" using \f1 (n-1) 4 = 0\ by auto also have "\ = f2 i" proof - have "x_pad ((n-1) * 5 + 0) * f1 (n-1) 0 = x $ ((n-1)*5) * (5^(4 * (n - 1)) + 5 ^ (4 * (n - 1) + 1) + 5 ^ (4 * (n - 1) + 3))" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 1) * f1 (n-1) 1 = x $ ((n-1)*5 + 1) * (5^(4 * (n - 1) + 1) + 1)" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 2) * f1 (n-1) 2 = x $ ((n-1)*5 + 2) * (5^(4 * (n - 1) + 2) + 5 ^ (4 * (n - 1) + 3) + 1)" unfolding f1_def x_pad_def using \n>0\ by auto moreover have "x_pad ((n-1) * 5 + 3) * f1 (n-1) 3 = x $ ((n-1)*5 + 3) * 5^(4 * (n - 1) + 3)" unfolding f1_def x_pad_def using \n>0\ by auto ultimately show ?thesis unfolding f2_def using \i=n-1\ by auto qed finally show ?thesis by auto qed qed also have "\ = (\i0 < length a\ n_def in \presburger\) also have "\ = (\ilength a > 0\ n_def in \auto\) also have "\ = (\i = (\iiiii = (\iiiii = ?f4") using sum.lessThan_Suc[of "(\i. d1 i * 5 ^ (4 * i))" "n-1"] using sum.lessThan_Suc[of "(\i. d2 i * 5 ^ (4 * i + 1))" "n-1"] using sum.lessThan_Suc[of "(\i. d3 i * 5 ^ (4 * i + 2))" "n-1"] using sum.lessThan_Suc[of "(\i. d4 i * 5 ^ (4 * i + 3))" "n-1"] using \0 < length a\ n_def by force finally have "(\i<5*n. x_pad i * (a0_rest i)) = ?f4" using * by auto then have "(\i<5*n-1. x$i * (a0_rest i)) = ?f4" using pad by auto moreover have "(\iii. d1 i * 5 ^ (4 * i))" "n-1"] using \0 < length a\ n_def by force moreover have "(\iiii. d1 (i + 1) * 5 ^ (4 * (i + 1)))" "(\i. d5 i * 5 ^ (4 * (i + 1)))" "{..ii\{1..i. i - 1) {1..0 < length a\ add_diff_cancel_left' atLeastLessThan_iff image_eqI n_def plus_1_eq_Suc zero_less_Suc) then show ?thesis by (subst sum.reindex_bij_betw[of "(\i. i-1)" "{1..i. (d1 (i + 1) + d5 i) * 5 ^ (4 * (i + 1)))", symmetric]) auto qed moreover have "(\i\{1..i0, of "(\i. (if i = 0 then d5 (n - 1) + d1 0 else d1 i + d5 (i - 1)) * 5 ^ (4 * i))"] unfolding atLeast0LessThan n_def by (auto) ultimately have "(\i<5*n-1. x$i * (a0_rest i)) = (\ii<5*n-1. x$i * (a0_rest i)) = (\i = (\ij<4. d (i*4+j) * 5^(i*4+j)))" proof (rule sum.cong, goal_cases) case (2 i) have d_rew: "(\j<4. d (i * 4 + j) * 5 ^ (i * 4 + j)) = d (i * 4) * 5 ^ (i * 4) + d (i * 4 + 1) * 5 ^ (i * 4 + 1) + d (i * 4 + 2) * 5 ^ (i * 4 + 2) + d (i * 4 + 3) * 5 ^ (i * 4 + 3)" by (simp add: eval_nat_numeral) have d1_rew: "d (i * 4) = (if i = 0 then d5 (n - 1) + d1 0 else d1 i + d5 (i - 1))" unfolding d_def by auto have d2_rew: "d (i*4+1) = d2 i" unfolding d_def by (metis add.commute add.right_neutral div_mult_self3 mod_Suc mod_div_trivial mod_mult_self2_is_0 one_eq_numeral_iff plus_1_eq_Suc semiring_norm(85) zero_neq_numeral) have d3_rew: "d (i*4+2) = d3 i" unfolding d_def by (metis add.commute add_2_eq_Suc' add_cancel_right_left div_less div_mult_self1 less_Suc_eq mod_mult_self1 nat_mod_lem numeral_Bit0 one_add_one zero_neq_numeral) have d4_rew: "d (i*4+3) = d4 i" unfolding d_def by auto show ?case by (subst d_rew, subst d1_rew, subst d2_rew, subst d3_rew, subst d4_rew) (auto simp add: mult.commute) qed auto moreover have "\ = (\k<4*n. d k * 5^k)" using sum_split_idx_prod[of "(\k. d k * 5^k)" n 4, symmetric] by (simp add: lessThan_atLeast0 mult.commute) ultimately show ?thesis by auto qed text \Some helping lemmas to get equations.\ have xi_le_1: "\x$i\\1" if "i< dim_vec x" for i using 1(5) that unfolding linf_norm_vec_Max by auto have xs_le_2: "\x$i + x$j\\2" if "i< dim_vec x" "j< dim_vec x" for i j proof - have "\x$i + x$j\ \ \x$i\ + \x$j\" by (auto simp add: abs_triangle_ineq) then show ?thesis using xi_le_1[OF that(1)] xi_le_1[OF that(2)] by auto qed have if_xi_le_1: "\(if i < n - 1 then x $ (i * 5 + 4) else 0)\ \ 1" if "i< n" for i using that xi_le_1[of "i*5+4"] unfolding dim_vec_x_5n by auto have prec_0: "i * 5 < dim_vec x" if "id1 i\\2" if "id2 i\\2" if "id3 i\\2" if "id4 i\\3" if "id5 i\\2" if "id5 (n - Suc 0) + d1 0\ < 5" using abs_d5[of "n-1"] abs_d1[of 0] \0 < length a\ n_def by fastforce moreover have "\d1 (i div 4) + d5 (i div 4 - Suc 0)\ < 5" if "0 < i" and "i<4*n" and "i mod 4 = 0" for i using that abs_d1[of "i div 4"] abs_d5[of "i div 4 - 1"] \0 < length a\ n_def by fastforce moreover have "\d2 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 1" for i using that abs_d2[of "i div 4"] \0 < length a\ n_def by fastforce moreover have "\d3 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 2" for i using that abs_d3[of "i div 4"] \0 < length a\ n_def by fastforce moreover have "\d4 (i div 4)\ < 5" if "i<4*n" and "i mod 4 = 3" for i using that abs_d4[of "i div 4"] \0 < length a\ n_def by fastforce ultimately have d_lt_5: "\d i\ < 5" if "i < 4 * n" for i using that by (subst mod_4_choices[of i], unfold d_def, auto) have sum_zero: "(\k<4*n. d k * 5^k) = 0" using eq_0' rewrite_digits by auto then have d_eq_0: "d k = 0" if "k<4*n" for k using respresentation_in_basis_eq_zero[OF sum_zero _ _ that] d_lt_5 by auto text \These are the main equations.\ have eq_1: "x$(i*5) + (if i {1..0 < length a\ add_cancel_right_left bits_mod_0 bot_nat_0.not_eq_extremum mult.commute mult_is_0 n_def zero_neq_numeral) have eq_3: "x$(i*5) + x$(i*5+1) = 0" if "i\{0..{0..0 < length a\ add.right_neutral add_self_div_2 add_self_mod_2 atLeastLessThan_iff bits_one_mod_two_eq_one div_mult2_eq div_mult_self4 mod_mult2_eq mod_mult_self3 mult.commute mult_2 mult_is_0 n_def nat_1_add_1 nat_mod_lem neq0_conv numeral_Bit0 one_div_two_eq_zero) qed have eq_5: "x$(i*5) + x$(i*5+2) + x$(i*5+3) = 0" if "i\{0..{0..{0..{1.. {0..{0..This defines the weight of the solution, since $x_{i,0} + x_{i,4}$ does not depend on the index i. We take $x_{n-1,0} + x_{n-1,4}$, since we omitted the last element (thus $x_{n-1,4} = 0$) to ensure that the weight has absolute value at most $1$.\ define w where "w = x$((n-1)*5)" have w_eq_02: "w = x$(i*5) + (if i{0..n-1" using that by auto then show ?thesis proof (induct rule: Nat.inc_induct) case (step m) then show ?case unfolding w_def using eq_1'[of "Suc m"] by auto qed (unfold w_def, auto) qed have "\w\ \ 1" using xi_le_1[of "(n-1)*5"] \n > 0\ unfolding w_def dim_vec_x_5n by auto text \Rule out the all zero solution.\ moreover have "w\0" proof (rule ccontr) assume "\ w \ 0" then have "w = 0" by blast then have "x$((n-1)*5) = 0" unfolding w_def by auto have zero_eq_min2: "x$(i*5) = - (if i{0..w=0\ by auto have two0_plus_4: "2 * x$(i*5) + x$(i*5+3) = 0" if "i \ {0..x$((n-1)*5) = 0\ by auto next case False then have "i \ {0.. 0" then have "\2 * x $ (i * 5) + x $ (i * 5 + 3)\\1" using xi_le_1[OF prec_0[where i=i]] xi_le_1[OF prec_i[where i=i and j=3]] by (auto simp add: \i) then show False using two0_plus_4[OF \i \ {0..] by auto qed qed have "x$j = 0" if "j<5*n" "j mod 5 = 0" for j proof - - obtain i where "i*5 = j" "ij < 5 * n\ \j mod 5 = 0\ mod_eq_0D mult.commute nat_mult_less_cancel1 - zero_less_numeral) + from \j mod 5 = 0\ \j < 5 * n\ obtain i where "i*5 = j" "ii*5 = j\[symmetric] using p_zero[OF \i] by auto qed moreover have p_one: "x$(i*5+1) = 0" if "ij<5*n\ \j mod 5 = 1\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+1 = j\[symmetric] using p_one[OF \i] by auto qed moreover have p_four: "x$(i*5+4) = 0" if "iw=0\ by auto then have "x$j = 0" if "j<5*n-1" "j mod 5 = 4" for j proof - obtain i where "i*5+4 = j" "ij<5*n-1\ \j mod 5 = 4\ by (metis add.assoc add.commute less_Suc_eq less_diff_conv less_mult_imp_div_less mod_mult_div_eq mult.commute mult_Suc_right not_less_eq numeral_nat(3) plus_1_eq_Suc) then show ?thesis unfolding \i*5+4 = j\[symmetric] using p_four by auto qed moreover have p_two: "x$(i*5+2) = 0" if "ii] eq_4' by auto next case False then have "i=n-1" using that by auto show ?thesis using eq_4[of "n-1"] unfolding \i=n-1\ using \n>0\ by auto qed then have "x$j = 0" if "j<5*n" "j mod 5 = 2" for j proof - obtain i where "i*5+2 = j" "ij<5*n\ \j mod 5 = 2\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+2 = j\[symmetric] using p_two[OF \i] by auto qed moreover have p_three: "x$(i*5+3) = 0" if "ij<5*n\ \j mod 5 = 3\ by (metis add.commute less_mult_imp_div_less mod_mult_div_eq mult.commute) then show ?thesis unfolding \i*5+3 = j\[symmetric] using p_three[OF \i] by auto qed ultimately have "x$j = 0" if "j<5*n-1" for j using mod_5_choices[of j "(\j. x $ j = 0)"] that by auto then have "x = 0\<^sub>v (dim_vec x)" unfolding dim_vec_x_5n[symmetric] by auto then show False using 1(4) by auto qed text \Then we can deduce the wanted property for both cases $w=1$ and $w=-1$. The only differences between the two cases is the switch of signs.\ ultimately have "w=1 \ w = -1" by auto then consider (pos) "w=1" | (neg) "w=-1" by blast then show ?thesis proof cases case pos have 01:"(x$(i*5) = 0 \ x$(i*5+4) = 1) \ (x$(i*5) = 1 \ x$(i*5+4) = 0)" if "in>0\ unfolding dim_vec_x_5n by auto then have "x$(i*5) \ {-1,0,1}" using xi_le_1[of "i*5"] by auto then consider (minus) "x$(i*5) = -1" | (zero) "x$(i*5) = 0" | (plus) "x$(i*5) = 1" by blast then show ?thesis proof (cases) case minus then have "2 = x $ (i * 5 + 4)" using w_eq_02[of i] that \w=1\ by auto then have False using xi_le_1[of "i*5+4"] that unfolding dim_vec_x_5n by linarith then show ?thesis by auto next case zero then have "x $ (i * 5 + 4) = 1" using w_eq_02[of i] that unfolding \w=1\ by auto then show ?thesis using zero by auto next case plus then have "x $ (i * 5 + 4) = 0" using w_eq_02[of i] that unfolding \w=1\ by auto then show ?thesis using plus by auto qed qed have "n-1\I" using w_def unfolding \w=1\ I_def n_def[symmetric] by (auto simp add: \n>0\ mult.commute) have I_0: "x$(i*5) = 1" if "i\I" for i proof (cases "iw=1\ by auto qed have I_2: "x$(i*5+2) = 0" if "i\I" for i proof (cases "i0 < n\) moreover have "1 + x $ (i * 5 + 1) + x $ (i * 5 + 2) = 0" using eq_2 w_eq_02[of 0] unfolding \i=n-1\ by (metis \0 < n\ add_cancel_right_left atLeast0LessThan lessThan_iff mult_zero_left pos) ultimately show ?thesis by auto qed have notI_0: "x$(i*5) = 0" if "i\{0..n-1\I\ that by auto qed have notI_2: "x$(i*5+2) = -1" if "i\{0..n-1\I\ that by auto qed have "0 = (\i\I. (x $ (i * 5) + x $ (i * 5 + 2)) * a ! i) + (\i\{0..I" for i using I_0 I_2 that by auto moreover have "(x $ (i * 5) + x $ (i * 5 + 2)) = -1" if "i\{0..i\I. a ! i) - (\i\{0.. x$(i*5+4) = -1) \ (x$(i*5) = -1 \ x$(i*5+4) = 0)" if "in>0\ unfolding dim_vec_x_5n by auto then have "x$(i*5) \ {-1,0,1}" using xi_le_1[of "i*5"] by auto then consider (minus) "x$(i*5) = -1" | (zero) "x$(i*5) = 0" | (plus) "x$(i*5) = 1" by blast then show ?thesis proof (cases) case plus then have "-2 = x $ (i * 5 + 4)" using w_eq_02[of i] that \w=-1\ by force then have False using xi_le_1[of "i*5+4"] that unfolding dim_vec_x_5n by linarith then show ?thesis by auto next case zero then have "x $ (i * 5 + 4) = -1" using w_eq_02[of i] that unfolding \w=-1\ by auto then show ?thesis using zero by auto next case minus then have "x $ (i * 5 + 4) = 0" using w_eq_02[of i] that unfolding \w=-1\ by auto then show ?thesis using minus by auto qed qed have "n-1\I" using w_def unfolding \w=-1\ I_def n_def[symmetric] by (auto simp add: \n>0\ mult.commute) have I_0: "x$(i*5) = -1" if "i\I" for i proof (cases "iw=-1\ by auto qed have I_2: "x$(i*5+2) = 0" if "i\I" for i proof (cases "i0 < n\) moreover have "-1 + x $ (i * 5 + 1) + x $ (i * 5 + 2) = 0" using eq_2 w_eq_02[of 0] unfolding \i=n-1\ by (metis \0 < n\ add_cancel_right_left lessThan_atLeast0 lessThan_iff mult_zero_left neg) ultimately show ?thesis by auto qed have notI_0: "x$(i*5) = 0" if "i\{0..n-1\I\ that by auto qed have notI_2: "x$(i*5+2) = 1" if "i\{0..n-1\I\ that by auto qed have "0 = (\i\I. (x $ (i * 5) + x $ (i * 5 + 2)) * a ! i) + (\i\{0..I" for i using I_0 I_2 that by auto moreover have "(x $ (i * 5) + x $ (i * 5 + 2)) = 1" if "i\{0..i\I. a ! i) - (\i\{0..length a > 0\ by auto qed text \The Gap-SVP is NP-hard.\ lemma "is_reduction reduce_bhle_partition partition_problem_nonzero bhle" unfolding is_reduction_def proof (safe, goal_cases) case (1 a) then show ?case using well_defined_reduction_subset_sum by auto next case (2 a) then show ?case using NP_hardness_reduction_subset_sum by auto qed end \ No newline at end of file diff --git a/thys/CVP_Hardness/SVP_vec.thy b/thys/CVP_Hardness/SVP_vec.thy --- a/thys/CVP_Hardness/SVP_vec.thy +++ b/thys/CVP_Hardness/SVP_vec.thy @@ -1,563 +1,564 @@ theory SVP_vec imports BHLE begin section \SVP in $\ell_\infty$\ text \The reduction of SVP to BHLE in $l_\infty$ norm\ text \The shortest vector problem.\ definition is_shortest_vec :: "int_lattice \ int vec \ bool" where "is_shortest_vec L v \ (is_lattice L) \ (\x\L. \x\\<^sub>\ \ \v\\<^sub>\ \ v\L) " text \The decision problem associated with solving SVP exactly.\ definition gap_svp :: "(int_lattice \ int) set" where "gap_svp \ {(L, r). (is_lattice L) \ (dim_lattice L \ 2) \ (\v\L. \v\\<^sub>\ \ r \ v \ 0\<^sub>v (dim_vec v))}" text \Generate a lattice to solve SVP for reduction.\ text \Here, the factor $K''$ from the paper \cite{EmBo81} was changed to be $2\cdot\mathbf{k}\cdot(k+1)\cdot \sum_i \mathbf{| a_i |}$ in order for the proofs to finish.\ definition gen_svp_basis :: "int vec \ int \ int mat" where "gen_svp_basis a k = mat (dim_vec a + 1) (dim_vec a + 1) (\ (i, j). if (i < dim_vec a) \ (j< dim_vec a) then (if i = j then 1 else 0) else (if (i < dim_vec a) \ (j \ dim_vec a) then 0 else (if (i \ dim_vec a) \ (j < dim_vec a) then (k+1) * (a $ j) else 2*k*(k+1)* (\ i \ {0 ..< dim_vec a}. \a $ i\) +1 )))" text \Reduction SVP to bounded homogeneous linear equation problem in $\ell_\infty$ norm.\ definition reduce_svp_bhle:: "(int vec \ int) \ (int_lattice \ int)" where "reduce_svp_bhle \ (\ (a, k). (gen_lattice (gen_svp_basis a k), k))" text \Lemmas for proof\ lemma gen_svp_basis_mult: assumes "dim_vec z = dim_vec a + 1" shows "(gen_svp_basis a k) *\<^sub>v z = vec (dim_vec a + 1) (\i. if i < dim_vec a then z$i else (k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i) + (2*k*(k+1)* (\ i \ {0 ..< dim_vec a}. \a $ i\) +1) * (z$(dim_vec a)))" using assms proof (subst vec_eq_iff, safe, goal_cases) case 1 then show ?case using assms unfolding gen_svp_basis_def by auto next case (2 i) then show ?case proof (cases "iia = 0..ia \ insert i {0.. = z$i" by (subst sum.insert_remove, auto) finally have "(\ia = 0..v z = vec (dim_vec a + 1) (\i. if i < dim_vec a then z$i else (k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i) + (2*k*(k+1)* (\ i \ {0 ..< dim_vec a}. \a $ i\) +1) * (z$(dim_vec a)))" using assms proof (subst vec_eq_iff, safe, goal_cases) case 1 then show ?case using assms unfolding gen_svp_basis_def by auto next case (2 i) then show ?case proof (cases "iia = 0..ia \ insert i {0.. = z$i" by (subst sum.insert_remove, auto) finally have "(\ia = 0..ia = 0..v z) $ (dim_vec a) = (k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i) + (2*k*(k+1)* (\ i \ {0 ..< dim_vec a}. \a $ i\) +1) * (z$(dim_vec a))" using gen_svp_basis_mult[OF assms] by auto text \The set generated by \gen_svp_basis\ is indeed linearly independent.\ lemma is_indep_gen_svp_basis: assumes "k>0" shows "is_indep (gen_svp_basis a k)" unfolding is_indep_int_def proof (safe, goal_cases) case (1 z) have dim_row_dim_vec: "dim_row (gen_svp_basis a k) = dim_vec z" using 1 unfolding gen_svp_basis_def by auto then have suc_dim_a_dim_z: "dim_vec z = dim_vec a + 1" unfolding gen_svp_basis_def by auto have alt1: "(real_of_int_mat (gen_svp_basis a k) *\<^sub>v z) $ i = 0" if "i< dim_vec a +1"for i using 1(1) that unfolding gen_svp_basis_def by auto have z_upto_last: "z$i = 0" if "i < dim_vec a" for i proof - have elem: "(real_of_int_mat (gen_svp_basis a k) *\<^sub>v z) $ i = z $ i" using gen_svp_basis_mult_real[OF suc_dim_a_dim_z] that by auto show ?thesis by (subst elem[symmetric]) (use alt1[of i] that in \auto\) qed moreover have "z $ (dim_vec a) = 0" proof - have "0 = (real_of_int_mat (gen_svp_basis a k) *\<^sub>v z) $ (dim_vec a)" using alt1 by auto also have "\ = (real_of_int k + 1) * (\i = 0..x = 0..a $ x\)) + 1) * z $ dim_vec a" using gen_svp_basis_mult_real[OF suc_dim_a_dim_z] by auto also have "\ = real_of_int (2 * k * (k + 1) * (\x = 0..a $ x\) + 1 ) * (z$dim_vec a)" using suc_dim_a_dim_z using z_upto_last by auto finally have "0 = real_of_int (2 * k * (k + 1) * (\x = 0..a $ x\) + 1 ) * (z$dim_vec a)" by blast moreover have "real_of_int (2 * k * (k + 1) * (\x = 0..a $ x\) + 1 ) \ 0" proof (rule ccontr, goal_cases) case 1 then have eq: "real_of_int (\x = 0..a $ x\) = - 1 / real_of_int (k * (k + 1))" using assms by (smt (verit, best) mult_minus_right of_int_hom.hom_0 pos_zmult_eq_1_iff pos_zmult_eq_1_iff_lemma) have "k\1" using assms by auto have "real_of_int (\x = 0..a $ x\) \ \" by auto moreover have "- 1 / real_of_int (k * (k + 1)) \ \" using \k\1\ by (smt (verit, del_insts) "1" linordered_semiring_strict_class.mult_pos_pos mult_minus_right of_int_hom.hom_0 pos_zmult_eq_1_iff) ultimately show False using eq by auto qed (*Problems here when changing 2*(k+1) to k*(k+1). This is necessary since k'a only is smaller than k'' under the assumtion that z$i\k, not z$i\2.*) ultimately show ?thesis by auto qed ultimately have "z$i = 0" if "i < dim_vec z" for i using that suc_dim_a_dim_z - by (metis discrete le_eq_less_or_eq verit_comp_simplify1(3)) - then show ?case by auto + by (cases \dim_vec a = i\) simp_all + then show ?case + by auto qed lemma insert_set_comprehension: "insert (f x) {f i |i. i<(x::nat)} = {f i | i. i bhle" shows "k>0" using assms unfolding bhle_def proof (safe, goal_cases) case (1 v) have "\iv $ i\ > 0" using 1 by auto then have "\v\\<^sub>\ > 0" unfolding linf_norm_vec_Max using 1 by (subst Max_gr_iff, auto) then show ?thesis using 1 by auto qed lemma svp_k_pos: assumes "reduce_svp_bhle (a, k) \ gap_svp" shows "k>0" proof - obtain v where v_in_lattice: "v\gen_lattice (gen_svp_basis a k)" and infnorm_v: "\v\\<^sub>\ \ k" and v_nonzero: "v \ 0\<^sub>v (dim_vec v)" using assms unfolding reduce_svp_bhle_def gap_svp_def by force have "\ i < dim_vec v. \v $ i\ > 0" using v_nonzero by auto then have "\v\\<^sub>\ > 0" unfolding linf_norm_vec_Max by (subst Max_gr_iff, auto) then show ?thesis using infnorm_v by auto qed lemma svp_dim_vec_a: assumes "reduce_svp_bhle (a, k) \ gap_svp" shows "dim_vec a > 0" proof - have "2 \ dim_lattice (gen_lattice (gen_svp_basis a k))" using assms unfolding reduce_svp_bhle_def gap_svp_def by auto then have "2 \ dim_col (gen_svp_basis a k)" using dim_lattice_gen_lattice[of "gen_svp_basis a k",OF is_indep_gen_svp_basis] svp_k_pos[OF assms] by auto then show ?thesis unfolding gen_svp_basis_def by auto qed lemma bhle_dim_vec_a: assumes "(a, k) \ bhle" shows "dim_vec a > 0" using assms unfolding bhle_def by auto lemma first_lt_second: assumes "k>0" and z_le_k:"\i. i< dim_vec a \ \z $ i\ \ k" shows "2 * \(k + 1) * (\i = 0.. < (\2 * k * (k + 1) * (\i = 0..a $ i\) + 1\::int)" proof - have take_k1_out: "\(k + 1) * (\i = 0.. = (k+1) * \\i=0.." using \k>0\ by (smt (verit, best) mult_minus_right mult_nonneg_nonneg) have "\\i = 0.. \ (\i = 0..z $ i * a $ i\)" by (subst sum_abs[of "(\i. z$i * a$i)" "{0.. = (\i=0..z $ i\ * \a $ i\)" by (meson abs_mult) also have "\ \ (\i=0..a $ i\)" by (subst sum_mono) (auto simp add: z_le_k mult_right_mono) also have "\ = k * (\i=0..a $ i\)" by (metis mult_hom.hom_sum) finally have "\(\i=0.. \ k * (\i=0..a $ i\)" by blast then show ?thesis using take_k1_out using \0 < k\ by auto qed text \Well-definedness of reduction function\ lemma well_defined_reduction_svp: assumes "(a,k) \ bhle" shows "reduce_svp_bhle (a,k) \ gap_svp" using assms unfolding reduce_svp_bhle_def gap_svp_def bhle_def proof (safe, goal_cases) case (1 x) have "k>0" using assms bhle_k_pos by auto then show ?case using is_lattice_gen_lattice is_indep_gen_svp_basis by auto next case (2 x) have "dim_lattice (gen_lattice (gen_svp_basis a k)) = dim_col (gen_svp_basis a k)" using dim_lattice_gen_lattice assms bhle_k_pos is_indep_gen_svp_basis by presburger also have "\ = dim_vec a + 1" unfolding gen_svp_basis_def by auto also have "\ \ 2" using bhle_dim_vec_a[OF assms] by auto finally show ?case by auto next case (3 x) let ?x = "vec (dim_vec x + 1) (\i. if i< dim_vec x then x$i else 0)" define v where "v = (gen_svp_basis a k) *\<^sub>v ?x" have eigen_v: "v = ?x" unfolding v_def apply (subst gen_svp_basis_mult[where a= a and k=k and z=" ?x"], auto simp add: 3(2) comp_def) apply (subst vec_eq_iff, auto simp add: 3(1) scalar_prod_def) proof (goal_cases one two) case (one i) then show ?case by (auto simp add: comp_def 3(2)) next case (two i) have "(\i = 0.. gen_lattice (gen_svp_basis a k)" unfolding v_def gen_lattice_def by auto moreover have "\v\\<^sub>\ \ k" proof - have "\v\\<^sub>\ \ \x\\<^sub>\" unfolding eigen_v linf_norm_vec_Max proof (rule Max.boundedI, goal_cases _ _ three) case (three b) have dim_x_nonzero: "dim_vec x \ 0" using 3(3) 3(2) by auto then have nonempty: "(\a\{\x $ i\ |i. i < dim_vec x}. 0 \ a)" using abs_ge_zero by blast have " \x $ i\ \ Max (insert 0 {\x $ j\ |j. j < dim_vec x})" if "i < dim_vec x" for i using that by (subst Max_ge, auto) moreover have "0 \ Max (insert 0 {\x $ i\ |i. i < dim_vec x})" using 3 nonempty by (subst Max_ge_iff, auto) ultimately show ?case using three by auto qed auto then show ?thesis using 3 by auto qed moreover have "v \ 0\<^sub>v (dim_vec v)" using 3(3) proof (safe) assume "0 < dim_vec a" assume "v = 0\<^sub>v (dim_vec v)" have fst: "x \ 0\<^sub>v (dim_vec x)" using 3(4) by blast have snd: "?x = 0\<^sub>v (dim_vec v)" using \v = 0\<^sub>v (dim_vec v)\ eigen_v by auto have "x$i = 0" if "i< dim_vec x" for i using snd by (smt (z3) add.commute dim_vec eigen_v index_map_vec(2) index_vec index_zero_vec(1) of_int_eq_iff of_int_hom.hom_zero that trans_less_add2) then show False using fst by auto qed ultimately show ?case by blast qed text \NP-hardness of reduction function\ lemma NP_hardness_reduction_svp: assumes "reduce_svp_bhle (a,k) \ gap_svp" shows "(a,k) \ bhle" proof (cases "(\i 0" using svp_dim_vec_a[OF assms] by auto have "k>0" using svp_k_pos[OF assms] by auto define x where "x = vec (dim_vec a) (\i. k)" have "a \ x = 0" unfolding x_def scalar_prod_def by (auto simp add: True sum_distrib_right[symmetric]) (metis True lessThan_atLeast0) moreover have "dim_vec x = dim_vec a" unfolding x_def by auto moreover have "x \ 0\<^sub>v (dim_vec x)" proof - have "\i< dim_vec x. x $ i \ 0" unfolding x_def using \k>0\ a_nonempty by auto then show ?thesis using vec_eq_iff[of "x" "0\<^sub>v (dim_vec x)"] by auto qed moreover have "real_of_int (\x\\<^sub>\) \ k" proof - have "vec (dim_vec a) (\i. k) $ j = k" if "j < dim_vec a" for j using that by auto then have "Max {\vec (dim_vec a) (\i. k) $ i\ |i. i < dim_vec a} = Max {\k\ |i. i < dim_vec a}" by metis also have "\ = Max {k}" using \k>0\ a_nonempty by (smt (verit, best) Collect_cong singleton_conv) also have "\ = k" by auto finally have "Max {\vec (dim_vec a) (\i. k) $ i\ |i. i < dim_vec a} = k" by blast then show ?thesis unfolding x_def linf_norm_vec_Max using \k>0\ by auto qed ultimately show ?thesis using assms unfolding reduce_svp_bhle_def gap_svp_def bhle_def by fastforce next case False show ?thesis using assms unfolding reduce_svp_bhle_def gap_svp_def bhle_def proof (safe, goal_cases) case (1 v) have a_nonempty: "dim_vec a > 0" using svp_dim_vec_a[OF assms] by auto have "k>0" unfolding linf_norm_vec_Max proof - have "\iv $ i\ > 0" using 1 by auto then have "\v\\<^sub>\ > 0" unfolding linf_norm_vec_Max by (subst Max_gr_iff, auto) then show ?thesis using 1 by auto qed then have "k\1" by auto obtain z where z_def: "v = (gen_svp_basis a k) *\<^sub>v z" "dim_vec z = dim_vec a + 1" using 1 unfolding gen_lattice_def gen_svp_basis_def by auto have dim_v_dim_a:"dim_vec v = dim_vec a + 1" using 1 unfolding gen_lattice_def gen_svp_basis_def by auto define x where "x = vec (dim_vec a) (($) z)" have v_eq_z_upto_last: "v $ i = z $ i" if "i< dim_vec a" for i by (subst z_def) (subst gen_svp_basis_mult; use that z_def in \auto\) have v_le_k: "\v $ i\ \ k" if "i< dim_vec a + 1" for i using 1 dim_v_dim_a that unfolding linf_norm_vec_Max using Max_le_iff[of "{\v $ i\ |i. i < dim_vec v}" k] by fastforce then have z_le_k: "\z $ i\ \ k" if "i< dim_vec a" for i using v_eq_z_upto_last[OF that] that by (metis less_add_one less_trans) have v_last_zero: "v$(dim_vec a) = 0" proof (rule ccontr) assume ccontr_assms:"v $ dim_vec a \ 0" have v_last: "v$(dim_vec a) = (k+1) * (\i = 0..x = 0..a $ x\) + 1) * (z $ dim_vec a) " (is "v$(dim_vec a) = ?first + ?second") by (auto simp: z_def gen_svp_basis_mult_last) then have "?first \ 0 \ ?second \ 0" using ccontr_assms by auto then consider (first) "?first \ 0 \ ?second = 0" | (second) "?second \ 0 \ ?first = 0" | (both) "?first \ 0 \ ?second \ 0" by blast then show False proof (cases) case first then have gr_1: "\\i = 0.. \ 1" using \k>0\ by auto have "\v$ dim_vec a\ = \?first\" using first v_last by auto also have "\ = (k+1) * \\i = 0.." using \k>0\ by (smt (z3) mult_le_0_iff mult_minus_right) also have "\ > k" using first gr_1 \k>0\ by (smt (verit, best) mult_le_cancel_left1) finally have "\v$ dim_vec a\ > k" by auto then show ?thesis using v_le_k[of "dim_vec a"] by auto next case second have "\z $ dim_vec a\ \ 1" using second by auto have sum_a_ge_1:"(\xa $ x\) \1" using False atLeast0LessThan by (metis abs_sum_abs int_one_le_iff_zero_less not_less sum_abs zero_less_abs_iff) then have "2*k*(k+1)*(\xa $ x\) + 1 > k" proof - have "2*(k+1)*(\xa $ x\)\1" using \k>0\ by (smt (verit, ccfv_SIG) int_distrib(2) sum_a_ge_1 zmult_zless_mono2) then show ?thesis using \k>0\ by (smt (verit, best) pos_mult_pos_ge sum_a_ge_1) qed moreover have "\?second\ \ \2*k*(k+1)*(\xa $ x\) + 1\" using \\z $ dim_vec a\ \ 1\ by (subst abs_mult) (simp add: lessThan_atLeast0 mult_le_cancel_left1) ultimately have "\v $ dim_vec a\ > k" using second v_last by linarith then show ?thesis using v_le_k[of "dim_vec a"] by auto next case both have z_gr_one:"\real_of_int (z$dim_vec a)\\1" using both by auto let ?second' = "2 * k * (k + 1) * (\i=0..a $ i\) + 1" have "(\i = 0.. 0" using both by auto then have one: "k < \?first\" by (smt (z3) mult_less_cancel_left2 mult_minus_right) then have first_nonzero: "\?first\ \ 0" using \k>0\ by auto have two'':"2 * \?first\ < \?second'\" using first_lt_second[OF \k>0\ z_le_k] by auto then have two': "\real_of_int ?second' / real_of_int ?first\ > 2" proof - have "0?first\" using first_nonzero by presburger have "2 * real_of_int \?first\ < real_of_int \?second'\" using two'' by linarith then have "2 < real_of_int \?second'\ / real_of_int \?first\" by (subst pos_less_divide_eq[OF \0?first\\], auto) also have "\ = \real_of_int ?second' / real_of_int ?first\" by simp finally show ?thesis by presburger qed then have two:"\(real_of_int ?second' / real_of_int ?first) * real_of_int (z$dim_vec a)\ > 2" proof - have "\(real_of_int ?second' / real_of_int ?first) * real_of_int (z$dim_vec a)\ = \real_of_int ?second' / real_of_int ?first\ * \real_of_int (z$dim_vec a)\" by (subst abs_mult, auto) also have ineq: "\ \ \real_of_int ?second' / real_of_int ?first\" using z_gr_one by (smt (z3) mult_less_cancel_left2) finally have "\(real_of_int ?second' / real_of_int ?first) * real_of_int (z$dim_vec a)\\ \real_of_int ?second' / real_of_int ?first\" by blast then show ?thesis using two' by linarith qed have "real_of_int \v $ dim_vec a\ / real_of_int \?first\ \ 1" proof - have "real_of_int \v $ dim_vec a\ / real_of_int \?first\ = \real_of_int (v $ dim_vec a)\ / \real_of_int ?first\" using of_int_abs[of "v $ dim_vec a"] of_int_abs[of "?first"] by auto also have "\ = \real_of_int (v $ dim_vec a) / real_of_int ?first\" using abs_divide[symmetric] by auto also have "\ = \(real_of_int ?first + real_of_int ?second' * real_of_int (z$dim_vec a)) / real_of_int ?first\" using v_last by auto also have "\ = \real_of_int ?first / real_of_int ?first + (real_of_int ?second' / real_of_int ?first)* real_of_int (z$dim_vec a)\" by (metis (no_types, lifting) add_divide_distrib times_divide_eq_left) also have "\ = \1 + (real_of_int ?second' / real_of_int ?first) * real_of_int (z$dim_vec a)\" using first_nonzero by (smt (verit, del_insts) of_int_eq_0_iff one_eq_divide_iff) also have "\ \ 1" using two by linarith finally show "real_of_int \v $ dim_vec a\ / real_of_int \?first\ \ 1" by blast qed then have "real_of_int \v $ dim_vec a\ \ real_of_int \?first\" by (simp add: le_divide_eq_1) then have "\v $ dim_vec a\ \ \?first\" using of_int_le_iff by blast then have "\v $ dim_vec a\ > k" using one by auto then show ?thesis using v_le_k[of "dim_vec a"] by auto qed qed have z_last_zero: "z$dim_vec a = 0" proof (rule ccontr) assume ccontr_assms:"z $ dim_vec a \ 0" then have "\z $ dim_vec a\ \ 1" by auto have "(k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i) + (2*k*(k+1)* (\ i \ {0 ..< dim_vec a}. \a $ i\) +1) * (z$(dim_vec a)) = 0" (is "?first + ?second * (z$(dim_vec a)) = 0") using v_last_zero z_def gen_svp_basis_mult_last by auto then have abs_eq: "\?first\ = \?second * (z$(dim_vec a))\" by linarith moreover have "\?first\ < \?second * (z$(dim_vec a))\" proof - have "\?first\ \ 2*\?first\" by auto then have "\?first\ < \?second\" using first_lt_second[OF \k>0\ z_le_k, of a] by auto moreover have "\?second\ \ \?second\*\z$dim_vec a\" using \\z $ dim_vec a\ \ 1\ by (simp add: mult_le_cancel_left1) ultimately have "\?first\ < \?second\*\z$dim_vec a\" by linarith then show ?thesis by linarith qed ultimately show False by auto qed have v_real_z: "v = z" using v_eq_z_upto_last v_last_zero z_last_zero by (metis Suc_eq_plus1 dim_v_dim_a less_Suc_eq vec_eq_iff z_def(2)) have "a \ x = 0" proof - have "0 = ((gen_svp_basis a k) *\<^sub>v z) $ (dim_vec a)" using v_last_zero z_def by auto also have "\ = (k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i)" by (subst gen_svp_basis_mult, auto simp add: z_def z_last_zero) finally have zero: "(k+1) * (\ i \ {0 ..< dim_vec a}. z $ i * a $ i) = 0" by auto then show ?thesis unfolding scalar_prod_def using x_def \k>0\ by (smt (verit, ccfv_SIG) atLeastLessThan_iff dim_vec index_vec mult.commute mult_eq_0_iff of_int_hom.hom_0 sum.cong) qed moreover have "dim_vec x = dim_vec a" unfolding x_def by auto moreover have "x \ 0\<^sub>v (dim_vec x)" proof - - have "\i< dim_vec a + 1. v$i \ 0" using 1 unfolding gen_lattice_def gen_svp_basis_def by auto - then have "\i< dim_vec a. v$i \ 0" using v_last_zero - by (metis add_le_cancel_right discrete nat_less_le) - then obtain i where "i0" by blast - then have "z$i \ 0" using v_real_z z_def - by (auto) + have "\i\dim_vec a. v$i \ 0" using 1 unfolding gen_lattice_def gen_svp_basis_def by auto + then obtain i where \i \ dim_vec a\ \v $ i \ 0\ by blast + then have \dim_vec a \ i\ using v_last_zero by auto + with \i \ dim_vec a\ have \i < dim_vec a\ by simp + with \v $ i \ 0\ have "z $ i \ 0" using v_real_z z_def + by auto then have "\i< dim_vec a. x$i \ 0" using x_def \i by auto then show ?thesis using x_def by auto qed moreover have "\x\\<^sub>\ \ k" proof - have "\z$i\ = \vec (dim_vec a) (($) z) $ i\" if "i < dim_vec a" for i using that by auto then have "Max (insert 0 {\vec (dim_vec a) (($) z) $ i\ |i. i < dim_vec a}) = Max (insert 0 {\z $ i\ |i. i < dim_vec a})" by (smt (z3) Collect_cong Setcompr_eq_image mem_Collect_eq) also have "\ = Max (insert 0 {\z $ i\ |i. i < dim_vec a + 1})" proof - have "Max (insert 0 {\z $ i\ |i. i < dim_vec a}) = Max (insert 0 (insert ( \z$dim_vec a\) {\z $ i\ |i. i < dim_vec a}))" proof - have "Max {\z $ i\ |i. i < dim_vec a} \ 0" using \dim_vec a >0\ by (subst Max_ge_iff, auto) then have "Max {\z $ i\ |i. i < dim_vec a} \ \z$dim_vec a\" using z_last_zero by simp then have max_subst: "Max {\z $ i\ |i. i < dim_vec a} = Max (insert \z $ dim_vec a\ {\z $ i\ |i. i < dim_vec a})" using \dim_vec a > 0\ by (subst Max_insert)+ (auto) then show ?thesis using \dim_vec a > 0\ by (subst Max_insert)+ (auto) qed then show ?thesis using insert_set_comprehension[of "(\i. \z $ i\)" "dim_vec a"] by auto qed finally have "Max (insert 0 {\vec (dim_vec a) (($) z) $ i\ |i. i < dim_vec a}) = Max (insert 0 {\z $ i\ |i. i < dim_vec a +1})" by blast then have "\x\\<^sub>\ = \v\\<^sub>\" unfolding linf_norm_vec_Max using x_def z_def v_real_z by auto then show ?thesis using 1 by auto qed ultimately show ?case by force qed qed text \The SVP is NP-hard in $\ell_\infty$.\ lemma "is_reduction reduce_svp_bhle bhle gap_svp" unfolding is_reduction_def proof (safe, goal_cases) case (1 as s) then show ?case using well_defined_reduction_svp by auto next case (2 as s) then show ?case using NP_hardness_reduction_svp by auto qed end \ No newline at end of file diff --git a/thys/Chandy_Lamport/Snapshot.thy b/thys/Chandy_Lamport/Snapshot.thy --- a/thys/Chandy_Lamport/Snapshot.thy +++ b/thys/Chandy_Lamport/Snapshot.thy @@ -1,5271 +1,5285 @@ section \The Chandy--Lamport algorithm\ theory Snapshot imports "HOL-Library.Sublist" Distributed_System Trace Util Swap begin subsection \The computation locale\ text \We extend the distributed system locale presented earlier: Now we are given a trace t of the distributed system between two configurations, the initial and final configuartions of t. Our objective is to show that the Chandy--Lamport algorithm terminated successfully and exhibits the same properties as claimed in~\<^cite>\"chandy"\. In the initial state no snapshotting must have taken place yet, however the computation itself may have progressed arbitrarily far already. We assume that there exists at least one process, that the total number of processes in the system is finite, and that there are only finitely many channels between the processes. The process graph is strongly connected. Finally there are Chandy and Lamport's core assumptions: every process snapshots at some time and no marker may remain in a channel forever.\ locale computation = distributed_system + fixes init final :: "('a, 'b, 'c) configuration" assumes finite_channels: "finite {i. \p q. channel i = Some (p, q)}" and strongly_connected_raw: "\p q. (p \ q) \ (tranclp (\p q. (\i. channel i = Some (p, q)))) p q" and at_least_two_processes: "card (UNIV :: 'a set) > 1" and finite_processes: "finite (UNIV :: 'a set)" and no_initial_Marker: "\i. (\p q. channel i = Some (p, q)) \ Marker \ set (msgs init i)" and no_msgs_if_no_channel: "\i. channel i = None \ msgs init i = []" and no_initial_process_snapshot: "\p. ~ has_snapshotted init p" and no_initial_channel_snapshot: "\i. channel_snapshot init i = ([], NotStarted)" and valid: "\t. trace init t final" and l1: "\t i cid. trace init t final \ Marker \ set (msgs (s init t i) cid) \ (\j. j \ i \ Marker \ set (msgs (s init t j) cid))" and l2: "\t p. trace init t final \ (\i. has_snapshotted (s init t i) p \ i \ length t)" begin definition has_channel where "has_channel p q \ (\i. channel i = Some (p, q))" lemmas strongly_connected = strongly_connected_raw[folded has_channel_def] lemma exists_some_channel: shows "\i p q. channel i = Some (p, q)" proof - obtain p q where "p : (UNIV :: 'a set) \ q : (UNIV :: 'a set) \ p \ q" by (metis (mono_tags) One_nat_def UNIV_eq_I all_not_in_conv at_least_two_processes card_Suc_Diff1 card.empty finite_processes insert_iff iso_tuple_UNIV_I less_numeral_extra(4) n_not_Suc_n) then have "(tranclp has_channel) p q" using strongly_connected by simp then obtain r s where "has_channel r s" by (meson tranclpD) then show ?thesis using has_channel_def by auto qed abbreviation S where "S \ s init" lemma no_messages_if_no_channel: assumes "trace init t final" shows "channel cid = None \ msgs (s init t i) cid = []" using no_messages_introduced_if_no_channel[OF assms no_msgs_if_no_channel] by blast lemma S_induct [consumes 3, case_names S_init S_step]: "\ trace init t final; i \ j; j \ length t; \i. P i i; \i j. i < j \ j \ length t \ (S t i) \ (t ! i) \ (S t (Suc i)) \ P (Suc i) j \ P i j \ \ P i j" proof (induct "j - i" arbitrary: i) case 0 then show ?case by simp next case (Suc n) then have "(S t i) \ (t ! i) \ (S t (Suc i))" using Suc step_Suc by simp then show ?case using Suc by simp qed lemma exists_index: assumes "trace init t final" and "ev \ set (take (j - i) (drop i t))" shows "\k. i \ k \ k < j \ ev = t ! k" proof - have "trace (S t i) (take (j - i) (drop i t)) (S t j)" by (metis assms(1) assms(2) diff_is_0_eq' exists_trace_for_any_i_j list.distinct(1) list.set_cases nat_le_linear take_eq_Nil) obtain l where "ev = (take (j - i) (drop i t)) ! l" "l < length (take (j - i) (drop i t))" by (metis assms(2) in_set_conv_nth) let ?k = "l + i" have "(take (j - i) (drop i t)) ! l = drop i t ! l" using \l < length (take (j - i) (drop i t))\ by auto also have "... = t ! ?k" by (metis add.commute assms(2) drop_all empty_iff list.set(1) nat_le_linear nth_drop take_Nil) finally have "ev = t ! ?k" using \ev = take (j - i) (drop i t) ! l\ by blast moreover have "i \ ?k \ ?k < j" using \l < length (take (j - i) (drop i t))\ by auto ultimately show ?thesis by blast qed lemma no_change_if_ge_length_t: assumes "trace init t final" and "i \ length t" and "j \ i" shows "S t i = S t j" proof - have "trace (S t i) (take (j - i) (drop i t)) (S t j)" using assms(1) assms(3) exists_trace_for_any_i_j by blast moreover have "(take (j - i) (drop i t)) = Nil" by (simp add: assms(2)) ultimately show ?thesis by (metis tr_init trace_and_start_determines_end) qed lemma no_marker_if_no_snapshot: shows "\ trace init t final; channel cid = Some (p, q); ~ has_snapshotted (S t i) p \ \ Marker \ set (msgs (S t i) cid)" proof (induct i) case 0 then show ?case by (metis exists_trace_for_any_i no_initial_Marker take_eq_Nil tr_init trace_and_start_determines_end) next case (Suc n) then have IH: "Marker \ set (msgs (S t n) cid)" by (meson distributed_system.exists_trace_for_any_i_j distributed_system.snapshot_stable_2 distributed_system_axioms eq_iff le_Suc_eq) then obtain tr where decomp: "trace (S t n) tr (S t (Suc n))" "tr = take (Suc n - n) (drop n t)" using Suc exists_trace_for_any_i_j le_Suc_eq by blast have "Marker \ set (msgs (S t (Suc n)) cid)" proof (cases "tr = []") case True then show ?thesis by (metis IH decomp(1) tr_init trace_and_start_determines_end) next case False then obtain ev where step: "tr = [ev]" "(S t n) \ ev \ (S t (Suc n))" by (metis One_nat_def Suc_eq_plus1 Suc_leI \tr = take (Suc n - n) (drop n t)\ \trace (S t n) tr (S t (Suc n))\ add_diff_cancel_left' append.simps(1) butlast_take cancel_comm_monoid_add_class.diff_cancel length_greater_0_conv list.distinct(1) list.sel(3) snoc_eq_iff_butlast take0 take_Nil trace.cases) then show ?thesis proof (cases ev) case (Snapshot p') then show ?thesis by (metis IH Suc.prems(2) Suc.prems(3) local.step(2) new_Marker_in_set_implies_nonregular_occurence new_msg_in_set_implies_occurrence nonregular_event_induces_snapshot snapshot_state_unchanged) next case (RecvMarker cid' p' q') have "p' \ p" proof (rule ccontr) assume asm: "~ p' \ p" then have "has_snapshotted (S t (Suc n)) p" proof - have "~ regular_event ev" using RecvMarker by auto moreover have "occurs_on ev = p" using asm RecvMarker by auto ultimately show ?thesis using step(2) Suc.hyps Suc.prems by (metis nonregular_event_induces_snapshot snapshot_state_unchanged) qed then show False using Suc.prems by blast qed moreover have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "hd (msgs (S t n) cid) = Marker \ length (msgs (S t n) cid) > 0" using step RecvMarker can_occur_def by auto then have "Marker : set (msgs (S t n) cid)" using list.set_sel(1) by fastforce then show False using IH by simp qed ultimately have "msgs (S t (Suc n)) cid = msgs (S t n) cid" proof - have "\r. channel cid = Some (p', r)" using Suc.prems(2) \p' \ p\ by auto with \cid \ cid'\ RecvMarker step show ?thesis by (cases "has_snapshotted (S t n) p'", auto) qed then show ?thesis by (simp add: IH) next case (Trans p' s s') then show ?thesis using IH local.step(2) by force next case (Send cid' p' q' s s' m) with step IH show ?thesis by (cases "cid' = cid", auto) next case (Recv cid' p' q' s s' m) with step IH show ?thesis by (cases "cid' = cid", auto) qed qed then show ?case by blast qed subsection \Termination\ text \We prove that the snapshot algorithm terminates, as exhibited by lemma \texttt{snapshot\_algorithm\_must\_terminate}. In the final configuration all processes have snapshotted, and no markers remain in the channels.\ lemma must_exist_snapshot: assumes "trace init t final" shows "\p i. Snapshot p = t ! i" proof (rule ccontr) assume "\p i. Snapshot p = t ! i" have "\i p. ~ has_snapshotted (S t i) p" proof (rule allI) fix i show "\p. ~ has_snapshotted (S t i) p" proof (induct i) case 0 then show ?case by (metis assms distributed_system.trace_and_start_determines_end distributed_system_axioms exists_trace_for_any_i computation.no_initial_process_snapshot computation_axioms take0 tr_init) next case (Suc n) then have IH: "\p. ~ has_snapshotted (S t n) p" by auto then obtain tr where "trace (S t n) tr (S t (Suc n))" "tr = take (Suc n - n) (drop n t)" using assms exists_trace_for_any_i_j le_Suc_eq by blast show "\p. ~ has_snapshotted (S t (Suc n)) p" proof (cases "tr = []") case True then show ?thesis by (metis IH \trace (S t n) tr (S t (Suc n))\ tr_init trace_and_start_determines_end) next case False then obtain ev where step: "tr = [ev]" "(S t n) \ ev \ (S t (Suc n))" by (metis One_nat_def Suc_eq_plus1 Suc_leI \tr = take (Suc n - n) (drop n t)\ \trace (S t n) tr (S t (Suc n))\ add_diff_cancel_left' append.simps(1) butlast_take cancel_comm_monoid_add_class.diff_cancel length_greater_0_conv list.distinct(1) list.sel(3) snoc_eq_iff_butlast take0 take_Nil trace.cases) then show ?thesis using step Suc.hyps proof (cases ev) case (Snapshot q) then show ?thesis by (metis \\p i. Snapshot p = t ! i\ \tr = [ev]\ \tr = take (Suc n - n) (drop n t)\ append_Cons append_take_drop_id nth_append_length) next case (RecvMarker cid' q r) then have m: "Marker \ set (msgs (S t n) cid')" using RecvMarker_implies_Marker_in_set step by blast have "~ has_snapshotted (S t n) q" using Suc by auto then have "Marker \ set (msgs (S t n) cid')" proof - have "channel cid' = Some (r, q)" using step can_occur_def RecvMarker by auto then show ?thesis using IH assms no_marker_if_no_snapshot by blast qed then show ?thesis using m by auto qed auto qed qed qed obtain j p where "has_snapshotted (S t j) p" using l2 assms by blast then show False using \\i p. \ has_snapshotted (S t i) p\ by blast qed lemma recv_marker_means_snapshotted: assumes "trace init t final" and "ev = RecvMarker cid p q" and "(S t i) \ ev \ (S t (Suc i))" shows "has_snapshotted (S t i) q" proof - have "Marker = hd (msgs (S t i) cid) \ length (msgs (S t i) cid) > 0" proof - have "Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid" using assms(2) assms(3) next_recv_marker by blast then show ?thesis by (metis length_greater_0_conv list.discI list.sel(1)) qed then have "Marker \ set (msgs (S t i) cid)" using hd_in_set by fastforce then show "has_snapshotted (S t i) q" proof - have "channel cid = Some (q, p)" using assms can_occur_def by auto then show ?thesis using \Marker \ set (msgs (S t i) cid)\ assms(1) no_marker_if_no_snapshot by blast qed qed lemma recv_marker_means_cs_Done: assumes "trace init t final" and "t ! i = RecvMarker cid p q" and "i < length t" shows "snd (cs (S t (i+1)) cid) = Done" proof - have "(S t i) \ (t ! i) \ (S t (i+1))" using assms(1) assms(3) step_Suc by auto then show ?thesis by (simp add: assms(2)) qed lemma snapshot_produces_marker: assumes "trace init t final" and "~ has_snapshotted (S t i) p" and "has_snapshotted (S t (Suc i)) p" and "channel cid = Some (p, q)" shows "Marker : set (msgs (S t (Suc i)) cid) \ has_snapshotted (S t i) q" proof - obtain ev where ex_ev: "(S t i) \ ev \ (S t (Suc i))" by (metis append_Nil2 append_take_drop_id assms(1) assms(2) assms(3) distributed_system.step_Suc distributed_system_axioms drop_eq_Nil less_Suc_eq_le nat_le_linear not_less_eq s_def) then have "occurs_on ev = p" using assms(2) assms(3) no_state_change_if_no_event by force then show ?thesis using assms ex_ev proof (cases ev) case (Snapshot r) then have "Marker \ set (msgs (S t (Suc i)) cid)" using ex_ev assms(2) assms(3) assms(4) by fastforce then show ?thesis by simp next case (RecvMarker cid' r s) have "r = p" using \occurs_on ev = p\ by (simp add: RecvMarker) then show ?thesis proof (cases "cid = cid'") case True then have "has_snapshotted (S t i) q" using RecvMarker RecvMarker_implies_Marker_in_set assms(1) assms(2) assms(4) ex_ev no_marker_if_no_snapshot by blast then show ?thesis by simp next case False then have "\s. channel cid = Some (r, s)" using RecvMarker assms can_occur_def \r = p\ by simp then have "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Marker]" using RecvMarker assms ex_ev \r = p\ False by simp then show ?thesis by simp qed qed auto qed lemma exists_snapshot_for_all_p: assumes "trace init t final" shows "\i. ~ has_snapshotted (S t i) p \ has_snapshotted (S t (Suc i)) p" (is ?Q) proof - obtain i where "has_snapshotted (S t i) p" using l2 assms by blast let ?j = "LEAST j. has_snapshotted (S t j) p" have "?j \ 0" proof - have "~ has_snapshotted (S t 0) p" by (metis exists_trace_for_any_i list.discI no_initial_process_snapshot s_def take_eq_Nil trace.simps) then show ?thesis by (metis (mono_tags, lifting) \has_snapshotted (S t i) p\ wellorder_Least_lemma(1)) qed have "?j \ i" by (meson Least_le \has_snapshotted (S t i) p\) have "\ has_snapshotted (S t (?j - 1)) p" (is ?P) proof (rule ccontr) assume "\ ?P" then have "has_snapshotted (S t (?j - 1)) p" by simp then have "\j. j < ?j \ has_snapshotted (S t j) p" by (metis One_nat_def \(LEAST j. ps (S t j) p \ None) \ 0\ diff_less lessI not_gr_zero) then show False using not_less_Least by blast qed show ?thesis proof (rule ccontr) assume "\ ?Q" have "\i. \ has_snapshotted (S t i) p" proof (rule allI) fix i' show "\ has_snapshotted (S t i') p" proof (induct i') case 0 then show ?case using \(LEAST j. ps (S t j) p \ None) \ 0\ by force next case (Suc i'') then show ?case using \\i. \ ps (S t i) p \ None \ ps (S t (Suc i)) p \ None\ by blast qed qed then show False using \ps (S t i) p \ None\ by blast qed qed lemma all_processes_snapshotted_in_final_state: assumes "trace init t final" shows "has_snapshotted final p" proof - obtain i where "has_snapshotted (S t i) p \ i \ length t" using assms l2 by blast moreover have "final = (S t (length t))" by (metis (no_types, lifting) assms exists_trace_for_any_i le_Suc_eq length_Cons take_Nil take_all trace.simps trace_and_start_determines_end) ultimately show ?thesis using assms exists_trace_for_any_i_j snapshot_stable by blast qed definition next_marker_free_state where "next_marker_free_state t i cid = (LEAST j. j \ i \ Marker \ set (msgs (S t j) cid))" lemma exists_next_marker_free_state: assumes "channel cid = Some (p, q)" "trace init t final" shows "\!j. next_marker_free_state t i cid = j \ j \ i \ Marker \ set (msgs (S t j) cid)" proof (cases "Marker \ set (msgs (S t i) cid)") case False then have "next_marker_free_state t i cid = i" unfolding next_marker_free_state_def by (metis (no_types, lifting) Least_equality order_refl) then show ?thesis using False assms by blast next case True then obtain j where "j \ i" "Marker \ set (msgs (S t j) cid)" using l1 assms by blast then show ?thesis by (metis (no_types, lifting) LeastI_ex next_marker_free_state_def) qed theorem snapshot_algorithm_must_terminate: assumes "trace init t final" shows "\phi. ((\p. has_snapshotted (S t phi) p) \ (\cid. Marker \ set (msgs (S t phi) cid)))" proof - let ?i = "{i. i \ length t \ (\p. has_snapshotted (S t i) p)}" have fin_i: "finite ?i" by auto moreover have "?i \ empty" proof - have "\p. has_snapshotted (S t (length t)) p" by (meson assms exists_trace_for_any_i_j l2 snapshot_stable_2) then show ?thesis by blast qed then obtain i where asm: "\p. has_snapshotted (S t i) p" by blast have f: "\j. j \ i \ (\p. has_snapshotted (S t j) p)" using snapshot_stable asm exists_trace_for_any_i_j valid assms by blast let ?s = "(\cid. (next_marker_free_state t i cid)) ` { cid. channel cid \ None }" have "?s \ empty" using exists_some_channel by auto have fin_s: "finite ?s" using finite_channels by simp let ?phi = "Max ?s" have "?phi \ i" proof (rule ccontr) assume asm: "\ ?phi \ i" obtain cid p q where g: "channel cid = Some (p, q)" using exists_some_channel by auto then have "next_marker_free_state t i cid \ i" using exists_next_marker_free_state assms by blast then have "Max ?s \ i" using Max_ge_iff g fin_s by fast then show False using asm by simp qed then have "\cid. Marker \ set (msgs (S t ?phi) cid)" proof - fix cid show "Marker \ set (msgs (S t ?phi) cid)" proof (cases "Marker : set (msgs (S t i) cid)") case False then show ?thesis using \i \ Max ?s\ asm assms exists_trace_for_any_i_j no_markers_if_all_snapshotted by blast next case True then have cpq: "channel cid \ None" using no_messages_if_no_channel assms by fastforce then obtain p q where chan: "channel cid = Some (p, q)" by auto then obtain j where i: "j = next_marker_free_state t i cid" "Marker \ set (msgs (S t j) cid)" using exists_next_marker_free_state assms by fast have "j \ ?phi" using cpq fin_s i(1) pair_imageI by simp then show "Marker \ set (msgs (S t ?phi) cid)" proof - have "trace (S t j) (take (?phi - j) (drop j t)) (S t ?phi)" using \j \ ?phi\ assms exists_trace_for_any_i_j by blast moreover have "\p. has_snapshotted (S t j) p" by (metis assms chan f computation.exists_next_marker_free_state computation_axioms i(1)) ultimately show ?thesis using i(2) no_markers_if_all_snapshotted by blast qed qed qed thus ?thesis using f \?phi \ i\ by blast qed subsection \Correctness\ text \The greatest part of this work is spent on the correctness of the Chandy-Lamport algorithm. We prove that the snapshot is consistent, i.e.\ there exists a permutation $t'$ of the trace $t$ and an intermediate configuration $c'$ of $t'$ such that the configuration recorded in the snapshot corresponds to the snapshot taken during execution of $t$, which is given as Theorem 1 in~\<^cite>\"chandy"\.\ lemma snapshot_stable_ver_2: shows "trace init t final \ has_snapshotted (S t i) p \ j \ i \ has_snapshotted (S t j) p" using exists_trace_for_any_i_j snapshot_stable by blast lemma snapshot_stable_ver_3: shows "trace init t final \ ~ has_snapshotted (S t i) p \ i \ j \ ~ has_snapshotted (S t j) p" using snapshot_stable_ver_2 by blast lemma marker_must_stay_if_no_snapshot: assumes "trace init t final" and "has_snapshotted (S t i) p" and "~ has_snapshotted (S t i) q" and "channel cid = Some (p, q)" shows "Marker : set (msgs (S t i) cid)" proof - obtain j where "~ has_snapshotted (S t j) p \ has_snapshotted (S t (Suc j)) p" using exists_snapshot_for_all_p assms by blast have "j \ i" proof (rule ccontr) assume asm: "~ j \ i" then have "~ has_snapshotted (S t i) p" using \\ has_snapshotted (S t j) p \ has_snapshotted (S t (Suc j)) p\ assms(1) less_imp_le_nat snapshot_stable_ver_3 by (meson nat_le_linear) then show False using assms(2) by simp qed have "i \ length t" proof (rule ccontr) assume "~ i \ length t" then have "i > length t" using not_less by blast obtain i' where a: "\p. has_snapshotted (S t i') p" using assms snapshot_algorithm_must_terminate by blast have "i' \ i" using \\p. has_snapshotted (S t i') p\ assms(1) assms(3) nat_le_linear snapshot_stable_ver_3 by blast have "(S t i') \ (S t i)" using assms a by force then have "i \ length t" using \i \ i'\ assms(1) computation.no_change_if_ge_length_t computation_axioms nat_le_linear by fastforce then show False using \~ i \ length t\ by simp qed have marker_in_set: "Marker : set (msgs (S t (Suc j)) cid)" using \\ has_snapshotted (S t j) p \ has_snapshotted (S t (Suc j)) p\ \j \ i\ assms(1) assms(3) assms(4) snapshot_produces_marker snapshot_stable_ver_3 by blast show ?thesis proof (rule ccontr) assume asm: "Marker \ set (msgs (S t i) cid)" then have range: "(Suc j) < i" by (metis Suc_lessI \\ ps (S t j) p \ None \ ps (S t (Suc j)) p \ None\ \j \ i\ assms(2) marker_in_set order.order_iff_strict) let ?k = "LEAST k. k \ (Suc j) \ Marker \ set (msgs (S t k) cid)" have range_k: "(Suc j) < ?k \ ?k \ i" proof - have "j < (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid))" by (metis (full_types) Suc_le_eq assms(1) assms(4) exists_next_marker_free_state next_marker_free_state_def) then show ?thesis proof - assume a1: "j < (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid))" have "j < i" using local.range by linarith (* 4 ms *) then have "(Suc j \ i \ Marker \ set (msgs (S t i) cid)) \ (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) \ Suc j" by (metis (lifting) Suc_leI asm marker_in_set wellorder_Least_lemma(1)) (* 64 ms *) then show ?thesis using a1 by (simp add: wellorder_Least_lemma(2)) (* 16 ms *) qed qed have a: "Marker : set (msgs (S t (?k-1)) cid)" proof - obtain nn :: "nat \ nat \ nat" where "\x0 x1. (\v2. x0 = Suc (x1 + v2)) = (x0 = Suc (x1 + nn x0 x1))" by moura then have f1: "(LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) = Suc (Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j))" using \Suc j < (LEAST k. Suc j \ k \ Marker \ set (msgs (S t k) cid)) \ (LEAST k. Suc j \ k \ Marker \ set (msgs (S t k) cid)) \ i\ less_iff_Suc_add by fastforce have f2: "Suc j \ Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j)" by simp have f3: "\p n. \ p (n::nat) \ Least p \ n" by (meson wellorder_Least_lemma(2)) have "\ (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) \ Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j)" using f1 by linarith then have f4: "\ (Suc j \ Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j) \ Marker \ set (msgs (S t (Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j))) cid))" using f3 by force have "Suc j + nn (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) (Suc j) = (LEAST n. Suc j \ n \ Marker \ set (msgs (S t n) cid)) - 1" using f1 by linarith then show ?thesis using f4 f2 by presburger qed have b: "Marker \ set (msgs (S t ?k) cid)" using assms(1) assms(4) exists_next_marker_free_state next_marker_free_state_def by fastforce have "?k - 1 < i" using range_k by auto then obtain ev where step: "(S t (?k-1)) \ ev \ (S t (Suc (?k-1)))" by (meson Suc_le_eq \i \ length t\ assms(1) le_trans step_Suc) then show False using a assms(1) assms(3) assms(4) b computation.snapshot_stable_ver_3 computation_axioms less_iff_Suc_add range_k recv_marker_means_snapshotted_2 by fastforce qed qed subsubsection \Pre- and postrecording events\ definition prerecording_event: "prerecording_event t i \ i < length t \ regular_event (t ! i) \ ~ has_snapshotted (S t i) (occurs_on (t ! i))" definition postrecording_event: "postrecording_event t i \ i < length t \ regular_event (t ! i) \ has_snapshotted (S t i) (occurs_on (t ! i))" abbreviation neighboring where "neighboring t i j \ i < j \ j < length t \ regular_event (t ! i) \ regular_event (t ! j) \ (\k. i < k \ k < j \ ~ regular_event (t ! k))" lemma pre_if_regular_and_not_post: assumes "regular_event (t ! i)" and "~ postrecording_event t i" and "i < length t" shows "prerecording_event t i" using assms computation.postrecording_event computation_axioms prerecording_event by metis lemma post_if_regular_and_not_pre: assumes "regular_event (t ! i)" and "~ prerecording_event t i" and "i < length t" shows "postrecording_event t i" using assms computation.postrecording_event computation_axioms prerecording_event by metis lemma post_before_pre_different_processes: assumes "i < j" and "j < length t" and neighboring: "\k. (i < k \ k < j) \ ~ regular_event (t ! k)" and post_ei: "postrecording_event t i" and pre_ej: "prerecording_event t j" and valid: "trace init t final" shows "occurs_on (t ! i) \ occurs_on (t ! j)" proof - let ?p = "occurs_on (t ! i)" let ?q = "occurs_on (t ! j)" have sp: "has_snapshotted (S t i) ?p" using assms postrecording_event prerecording_event by blast have nsq: "~ has_snapshotted (S t j) ?q" using assms postrecording_event prerecording_event by blast show "?p \ ?q" proof - have "~ has_snapshotted (S t i) ?q" proof (rule ccontr) assume sq: "~ ~ has_snapshotted (S t i) ?q" from \i < j\ have "i \ j" using less_imp_le by blast then obtain tr where ex_trace: "trace (S t i) tr (S t j)" using exists_trace_for_any_i_j valid by blast then have "has_snapshotted (S t j) ?q" using ex_trace snapshot_stable sq by blast then show False using nsq by simp qed then show ?thesis using sp by auto qed qed lemma post_before_pre_neighbors: assumes "i < j" and "j < length t" and neighboring: "\k. (i < k \ k < j) \ ~ regular_event (t ! k)" and post_ei: "postrecording_event t i" and pre_ej: "prerecording_event t j" and valid: "trace init t final" shows "Ball (set (take (j - (i+1)) (drop (i+1) t))) (%ev. ~ regular_event ev \ ~ occurs_on ev = occurs_on (t ! j))" proof - let ?p = "occurs_on (t ! i)" let ?q = "occurs_on (t ! j)" let ?between = "take (j - (i+1)) (drop (i+1) t)" show ?thesis proof (unfold Ball_def, rule allI, rule impI) fix ev assume "ev : set ?between" have len_nr: "length ?between = (j - (i+1))" using assms(2) by auto then obtain l where "?between ! l = ev" and range_l: "0 \ l \ l < (j - (i+1))" by (metis \ev \ set (take (j - (i + 1)) (drop (i + 1) t))\ gr_zeroI in_set_conv_nth le_numeral_extra(3) less_le) let ?k = "l + (i+1)" have "?between ! l = (t ! ?k)" proof - have "j < length t" by (metis assms(2)) then show ?thesis by (metis (no_types) Suc_eq_plus1 Suc_leI add.commute assms(1) drop_take length_take less_diff_conv less_imp_le_nat min.absorb2 nth_drop nth_take range_l) qed have "~ regular_event ev" by (metis (no_types, lifting) assms(3) range_l One_nat_def Suc_eq_plus1 \take (j - (i + 1)) (drop (i + 1) t) ! l = ev\ \take (j - (i + 1)) (drop (i + 1) t) ! l = t ! (l + (i + 1))\ add.left_commute add_lessD1 lessI less_add_same_cancel2 less_diff_conv order_le_less) have step_ev: "(S t ?k) \ ev \ (S t (?k+1))" proof - have "j \ length t" by (metis assms(2) less_or_eq_imp_le) then have "l + (i + 1) < length t" by (meson less_diff_conv less_le_trans range_l) then show ?thesis by (metis (no_types) Suc_eq_plus1 \take (j - (i + 1)) (drop (i + 1) t) ! l = ev\ \take (j - (i + 1)) (drop (i + 1) t) ! l = t ! (l + (i + 1))\ distributed_system.step_Suc distributed_system_axioms valid) qed obtain cid s r where f: "ev = RecvMarker cid s r \ ev = Snapshot r" using \~ regular_event ev\ by (meson isRecvMarker_def isSnapshot_def nonregular_event) from f have "occurs_on ev \ ?q" proof (elim disjE) assume snapshot: "ev = Snapshot r" show ?thesis proof (rule ccontr) assume occurs_on_q: "~ occurs_on ev \ ?q" then have "?q = r" using snapshot by auto then have q_snapshotted: "has_snapshotted (S t (?k+1)) ?q" using snapshot step_ev by auto then show False proof - have "l + (i + 1) < j" by (meson less_diff_conv range_l) then show ?thesis by (metis (no_types) Suc_eq_plus1 Suc_le_eq computation.snapshot_stable_ver_2 computation_axioms pre_ej prerecording_event q_snapshotted valid) qed qed next assume RecvMarker: "ev = RecvMarker cid s r" show ?thesis proof (rule ccontr) assume occurs_on_q: "~ occurs_on ev \ ?q" then have "s = ?q" using RecvMarker by auto then have q_snapshotted: "has_snapshotted (S t (?k+1)) ?q" proof (cases "has_snapshotted (S t ?k) ?q") case True then show ?thesis using snapshot_stable_ver_2 step_Suc step_ev valid by auto next case False then show "has_snapshotted (S t (?k+1)) ?q" using \s = ?q\ next_recv_marker RecvMarker step_ev by auto qed then show False proof - have "l + (i + 1) < j" using less_diff_conv range_l by blast then show ?thesis by (metis (no_types) Suc_eq_plus1 Suc_le_eq computation.snapshot_stable_ver_2 computation_axioms pre_ej prerecording_event q_snapshotted valid) qed qed qed then show "\ regular_event ev \ occurs_on ev \ ?q" using \~ regular_event ev\ by simp qed qed lemma can_swap_neighboring_pre_and_postrecording_events: assumes "i < j" and "j < length t" and "occurs_on (t ! i) = p" and "occurs_on (t ! j) = q" and neighboring: "\k. (i < k \ k < j) \ ~ regular_event (t ! k)" and post_ei: "postrecording_event t i" and pre_ej: "prerecording_event t j" and valid: "trace init t final" shows "can_occur (t ! j) (S t i)" proof - have "p \ q" using post_before_pre_different_processes assms by auto have sp: "has_snapshotted (S t i) p" using assms(3) post_ei postrecording_event prerecording_event by blast have nsq: "~ has_snapshotted (S t j) q" using assms(4) pre_ej prerecording_event by auto let ?nr = "take (j - (Suc i)) (drop (Suc i) t)" have valid_subtrace: "trace (S t (Suc i)) ?nr (S t j)" using assms(1) exists_trace_for_any_i_j valid by fastforce have "Ball (set ?nr) (%ev. ~ occurs_on ev = q \ ~ regular_event ev)" proof - have "?nr = take (j - (i+1)) (drop (i+1) t)" by auto then show ?thesis by (metis assms(1) assms(2) assms(4) neighboring post_ei pre_ej valid post_before_pre_neighbors) qed then have la: "list_all (%ev. ~ occurs_on ev = q) ?nr" by (meson list_all_length nth_mem) have tj_to_tSi: "can_occur (t ! j) (S t (Suc i))" proof - have "list_all (%ev. ~ isSend ev) ?nr" proof - have "list_all (%ev. ~ regular_event ev) ?nr" using \\ev\set (take (j - (Suc i)) (drop (Suc i) t)). occurs_on ev \ q \ \ regular_event ev\ \list_all (\ev. occurs_on ev \ q) (take (j - (Suc i)) (drop (Suc i) t))\ list.pred_mono_strong by fastforce then show ?thesis by (simp add: list.pred_mono_strong) qed moreover have "~ isRecvMarker (t ! j)" using prerecording_event assms by auto moreover have "can_occur (t ! j) (S t j)" proof - have "(S t j) \ (t ! j) \ (S t (Suc j))" using assms(2) step_Suc valid by auto then show ?thesis using happen_implies_can_occur by blast qed ultimately show "can_occur (t ! j) (S t (Suc i))" using assms(4) event_can_go_back_if_no_sender_trace valid_subtrace la by blast qed show "can_occur (t ! j) (S t i)" proof (cases "isSend (t ! i)") case False have "~ isRecvMarker (t ! j)" using assms prerecording_event by auto moreover have "~ isSend (t ! i)" using False by simp ultimately show ?thesis by (metis \p \ q\ assms(3) assms(4) event_can_go_back_if_no_sender post_ei postrecording_event step_Suc tj_to_tSi valid) next case True obtain cid s u u' m where Send: "t ! i = Send cid p s u u' m" by (metis True isSend_def assms(3) event.sel(2)) have chan: "channel cid = Some (p, s)" proof - have "can_occur (t ! i) (S t i)" by (meson computation.postrecording_event computation_axioms happen_implies_can_occur post_ei step_Suc valid) then show ?thesis using can_occur_def Send by simp qed have n: "(S t i) \ (t ! i) \ (S t (Suc i))" using assms(1) assms(2) step_Suc valid True by auto have st: "states (S t i) q = states (S t (Suc i)) q" using Send \p \ q\ n by auto have "isTrans (t ! j) \ isSend (t ! j) \ isRecv (t ! j)" using assms(7) computation.prerecording_event computation_axioms regular_event by blast then show ?thesis proof (elim disjE) assume "isTrans (t ! j)" then show ?thesis by (metis (no_types, lifting) tj_to_tSi st can_occur_def assms(4) event.case(1) event.collapse(1)) next assume "isSend (t ! j)" then obtain cid' s' u'' u''' m' where Send: "t ! j = Send cid' q s' u'' u''' m'" by (metis (no_types, lifting) assms(4) event.sel(2) isSend_def) have co_tSi: "can_occur (Send cid' q s' u'' u''' m') (S t (Suc i))" using Send tj_to_tSi by auto then have "channel cid' = Some (q, s') \ send cid' q s' u'' u''' m'" using Send can_occur_def by simp then show ?thesis using can_occur_def st Send assms co_tSi by auto next assume "isRecv (t ! j)" then obtain cid' s' u'' u''' m' where Recv: "t ! j = Recv cid' q s' u'' u''' m'" by (metis assms(4) event.sel(3) isRecv_def) have co_tSi: "can_occur (Recv cid' q s' u'' u''' m') (S t (Suc i))" using Recv tj_to_tSi by auto then have a: "channel cid' = Some (s', q) \ length (msgs (S t (Suc i)) cid') > 0 \ hd (msgs (S t (Suc i)) cid') = Msg m'" using can_occur_def co_tSi by fastforce show "can_occur (t ! j) (S t i)" proof (cases "cid = cid'") case False with Send n have "msgs (S t (Suc i)) cid' = msgs (S t i) cid'" by auto then have b: "length (msgs (S t i) cid') > 0 \ hd (msgs (S t i) cid') = Msg m'" using a by simp with can_occur_Recv co_tSi st a Recv show ?thesis unfolding can_occur_def by auto next case True (* This is the interesting case *) have stu: "states (S t i) q = u''" using can_occur_Recv co_tSi st by blast show ?thesis proof (rule ccontr) have marker_in_set: "Marker \ set (msgs (S t i) cid)" proof - have "(s', q) = (p, q)" using True a chan by auto then show ?thesis by (metis (no_types, lifting) True \p \ q\ a assms(3) marker_must_stay_if_no_snapshot n no_state_change_if_no_event nsq snapshot_stable_2 sp valid valid_subtrace) qed assume asm: "~ can_occur (t ! j) (S t i)" then show False proof (unfold can_occur_def, (auto simp add: marker_in_set True Recv stu)) assume "msgs (S t i) cid' = []" then show False using marker_in_set by (simp add: True) next assume "hd (msgs (S t i) cid') \ Msg m'" have "msgs (S t i) cid \ []" using marker_in_set by auto then have "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Msg m]" using Send True n chan by auto then have "hd (msgs (S t (Suc i)) cid) \ Msg m'" using True \hd (msgs (S t i) cid') \ Msg m'\ \msgs (S t i) cid \ []\ by auto then have "~ can_occur (t ! j) (S t (Suc i))" using True a by blast then show False using tj_to_tSi by blast next assume "~ recv cid' q s' u'' u''' m'" then show False using can_occur_Recv co_tSi by blast next assume "channel cid' \ Some (s', q)" then show False using can_occur_def tj_to_tSi Recv by simp qed qed qed qed qed qed subsubsection \Event swapping\ lemma swap_events: shows "\ i < j; j < length t; \k. (i < k \ k < j) \ ~ regular_event (t ! k); postrecording_event t i; prerecording_event t j; trace init t final \ \ trace init (swap_events i j t) final \ (\k. k \ j + 1 \ S (swap_events i j t) k = S t k) \ (\k. k \ i \ S (swap_events i j t) k = S t k) \ prerecording_event (swap_events i j t) i \ postrecording_event (swap_events i j t) (i+1) \ (\k. k > i+1 \ k < j+1 \ ~ regular_event ((swap_events i j t) ! k))" proof (induct "j - (i+1)" arbitrary: j t) case 0 let ?p = "occurs_on (t ! i)" let ?q = "occurs_on (t ! j)" have "j = (i+1)" using "0.prems" "0.hyps" by linarith let ?subt = "take (j - (i+1)) (drop (i+1) t)" have "t = take i t @ [t ! i] @ ?subt @ [t ! j] @ drop (j+1) t" proof - have "take (Suc i) t = take i t @ [t ! i]" using "0.prems"(2) \j = i + 1\ add_lessD1 take_Suc_conv_app_nth by blast then show ?thesis by (metis (no_types) "0.hyps" "0.prems"(2) Suc_eq_plus1 \j = i + 1\ append_assoc append_take_drop_id self_append_conv2 take_Suc_conv_app_nth take_eq_Nil) qed have sp: "has_snapshotted (S t i) ?p" using "0.prems" postrecording_event prerecording_event by blast have nsq: "~ has_snapshotted (S t j) ?q" using "0.prems" postrecording_event prerecording_event by blast have "?p \ ?q" using "0.prems" computation.post_before_pre_different_processes computation_axioms by blast have "?subt = Nil" by (simp add: \j = i + 1\) have reg_step_1: "(S t i) \ (t ! i) \ (S t j)" by (metis "0.prems"(2) "0.prems"(6) Suc_eq_plus1 \j = i + 1\ add_lessD1 step_Suc) have reg_step_2: "(S t j) \ (t ! j) \ (S t (j+1))" using "0.prems"(2) "0.prems"(6) step_Suc by auto have "can_occur (t ! j) (S t i)" using "0.prems" can_swap_neighboring_pre_and_postrecording_events by blast then obtain d' where new_step1: "(S t i) \ (t ! j) \ d'" using exists_next_if_can_occur by blast have st: "states d' ?p = states (S t i) ?p" using \(S t i) \ t ! j \ d'\ \occurs_on (t ! i) \ occurs_on (t ! j)\ no_state_change_if_no_event by auto then have "can_occur (t ! i) d'" using \occurs_on (t ! i) \ occurs_on (t ! j)\ event_stays_valid_if_no_occurrence happen_implies_can_occur new_step1 reg_step_1 by auto then obtain e where new_step2: "d' \ (t ! i) \ e" using exists_next_if_can_occur by blast have "states e = states (S t (j+1))" proof (rule ext) fix p show "states e p = states (S t (j+1)) p" proof (cases "p = ?p \ p = ?q") case True then show ?thesis proof (elim disjE) assume "p = ?p" then have "states d' p = states (S t i) p" by (simp add: st) thm same_state_implies_same_result_state then have "states e p = states (S t j) p" using "0.prems"(2) "0.prems"(6) new_step2 reg_step_1 by (blast intro:same_state_implies_same_result_state[symmetric]) moreover have "states (S t j) p = states (S t (j+1)) p" using \occurs_on (t ! i) \ occurs_on (t ! j)\ \p = occurs_on (t ! i)\ no_state_change_if_no_event reg_step_2 by auto ultimately show ?thesis by simp next assume "p = ?q" then have "states (S t j) p = states (S t i) p" using reg_step_1 \occurs_on (t ! i) \ occurs_on (t ! j)\ no_state_change_if_no_event by auto then have "states d' p = states (S t (j+1)) p" using "0.prems"(5) prerecording_event computation_axioms new_step1 reg_step_2 same_state_implies_same_result_state by blast moreover have "states e p = states (S t (j+1)) p" using \occurs_on (t ! i) \ occurs_on (t ! j)\ \p = occurs_on (t ! j)\ calculation new_step2 no_state_change_if_no_event by auto ultimately show ?thesis by simp qed next case False then have "states (S t i) p = states (S t j) p" using no_state_change_if_no_event reg_step_1 by auto moreover have "... = states (S t (j+1)) p" using False no_state_change_if_no_event reg_step_2 by auto moreover have "... = states d' p" using False calculation new_step1 no_state_change_if_no_event by auto moreover have "... = states e p" using False new_step2 no_state_change_if_no_event by auto ultimately show ?thesis by simp qed qed moreover have "msgs e = msgs (S t (j+1))" proof (rule ext) fix cid have "isTrans (t ! i) \ isSend (t ! i) \ isRecv (t ! i)" using "0.prems"(4) computation.postrecording_event computation_axioms regular_event by blast moreover have "isTrans (t ! j) \ isSend (t ! j) \ isRecv (t ! j)" using "0.prems"(5) computation.prerecording_event computation_axioms regular_event by blast ultimately show "msgs e cid = msgs (S t (j+1)) cid" proof (elim disjE, goal_cases) case 1 then have "msgs d' cid = msgs (S t j) cid" by (metis Trans_msg new_step1 reg_step_1) then show ?thesis using Trans_msg \isTrans (t ! i)\ \isTrans (t ! j)\ new_step2 reg_step_2 by auto next case 2 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Trans_Send by auto next case 3 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Trans_Recv by auto next case 4 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Trans by auto next case 5 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Recv_Trans by auto next case 6 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 by (blast intro:swap_msgs_Send_Send[symmetric]) next case 7 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Recv by auto next case 8 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Recv by simp next case 9 then show ?thesis using \occurs_on (t ! i) \ occurs_on (t ! j)\ new_step1 new_step2 reg_step_1 reg_step_2 by (blast intro:swap_msgs_Recv_Recv[symmetric]) qed qed moreover have "process_snapshot e = process_snapshot (S t (j+1))" proof (rule ext) fix p have "process_snapshot d' p = process_snapshot (S t j) p" by (metis "0.prems"(4) "0.prems"(5) computation.postrecording_event computation.prerecording_event computation_axioms new_step1 reg_step_1 regular_event_preserves_process_snapshots) then show "process_snapshot e p = process_snapshot (S t (j+1)) p" by (metis "0.prems"(4) "0.prems"(5) computation.postrecording_event computation.prerecording_event computation_axioms new_step2 reg_step_2 regular_event_preserves_process_snapshots) qed moreover have "channel_snapshot e = channel_snapshot (S t (j+1))" proof (rule ext) fix cid show "cs e cid = cs (S t (j+1)) cid" proof (cases "isRecv (t ! i)"; cases "isRecv (t ! j)", goal_cases) case 1 then show ?thesis using \?p \ ?q\ new_step1 new_step2 reg_step_1 reg_step_2 by (blast intro:regular_event_implies_same_channel_snapshot_Recv_Recv[symmetric]) next case 2 moreover have "regular_event (t ! j)" using prerecording_event 0 by simp ultimately show ?thesis using \?p \ ?q\ new_step1 new_step2 reg_step_1 reg_step_2 regular_event_implies_same_channel_snapshot_Recv by auto next assume 3: "~ isRecv (t ! i)" "isRecv (t ! j)" moreover have "regular_event (t ! i)" using postrecording_event 0 by simp ultimately show ?thesis using \?p \ ?q\ new_step1 new_step2 reg_step_1 reg_step_2 regular_event_implies_same_channel_snapshot_Recv by auto next assume 4: "~ isRecv (t ! i)" "~ isRecv (t ! j)" moreover have "regular_event (t ! j)" using prerecording_event 0 by simp moreover have "regular_event (t ! i)" using postrecording_event 0 by simp ultimately show ?thesis using \?p \ ?q\ new_step1 new_step2 reg_step_1 reg_step_2 by (metis no_cs_change_if_no_event) qed qed ultimately have "e = S t (j+1)" by simp then have "(S t i) \ (t ! j) \ d' \ d' \ (t ! i) \ (S t (j+1))" using new_step1 new_step2 by blast then have swap: "trace (S t i) [t ! j, t ! i] (S t (j+1))" by (meson trace.simps) have "take (j-1) t @ [t ! j, t ! i] = ((take (j+1) t)[i := t ! j])[j := t ! i]" proof - have "i = j - 1" by (simp add: \j = i + 1\) show ?thesis proof (subst (1 2 3) \i = j - 1\) have "j < length t" using "0.prems" by auto then have "take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t = t[j - 1 := t ! j, j := t ! (j - 1)]" by (metis Suc_eq_plus1 \i = j - 1\ \j = i + 1\ add_Suc_right arith_special(3) swap_neighbors) then show "take (j - 1) t @ [t ! j, t ! (j - 1)] = (take (j+1) t)[j - 1 := t ! j, j := t ! (j - 1)]" proof - assume a1: "take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t = t [j - 1 := t ! j, j := t ! (j - 1)]" have f2: "t[j - 1 := t ! j, j := t ! (j - 1)] = take j (t[j - 1 := t ! j]) @ t ! (j - 1) # drop (Suc j) (t[j - 1 := t ! j])" by (metis (no_types) "0.prems"(2) length_list_update upd_conv_take_nth_drop) (* 32 ms *) have f3: "\n na. \ n < na \ Suc n \ na" using Suc_leI by blast (* 0.0 ms *) then have "min (length t) (j + 1) = j + 1" by (metis (no_types) "0.prems"(2) Suc_eq_plus1 min.absorb2) (* 16 ms *) then have f4: "length ((take (j + 1) t)[j - 1 := t ! j]) = j + 1" by simp (* 4 ms *) have f5: "j + 1 \ length (t[j - 1 := t ! j])" using f3 by (metis (no_types) "0.prems"(2) Suc_eq_plus1 length_list_update) (* 8 ms *) have "Suc j \ j + 1" by linarith (* 0.0 ms *) then have "(take (j + 1) (t[j - 1 := t ! j]))[j := t ! (j - 1)] = take j (t[j - 1 := t ! j]) @ t ! (j - 1) # [] @ []" using f5 f4 by (metis (no_types) Suc_eq_plus1 add_diff_cancel_right' butlast_conv_take butlast_take drop_eq_Nil lessI self_append_conv2 take_update_swap upd_conv_take_nth_drop) (* 180 ms *) then show ?thesis using f2 a1 by (simp add: take_update_swap) (* 120 ms *) qed qed qed have s: "trace init (take i t) (S t i)" using "0.prems"(6) exists_trace_for_any_i by blast have e: "trace (S t (j+1)) (take (length t - (j+1)) (drop (j+1) t)) final" proof - have "trace init (take (length t) t) final" by (simp add: "0.prems"(6)) then show ?thesis by (metis "0.prems"(2) Suc_eq_plus1 Suc_leI exists_trace_for_any_i exists_trace_for_any_i_j nat_le_linear take_all trace_and_start_determines_end) qed have "trace init (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) final" proof - from s swap have "trace init (take i t @ [t ! j,t ! i]) (S t (j+1))" using trace_trans by blast then have "trace init (take i t @ [t ! j, t ! i] @ (take (length t - (j+1)) (drop (j+1) t))) final" using e trace_trans by fastforce moreover have "take (length t - (j+1)) (drop (j+1) t) = drop (j+1) t" by simp ultimately show ?thesis by simp qed moreover have "take i t @ [t ! j] @ [t ! i] @ drop (j+1) t = (t[i := t ! j])[j := t ! i]" proof - have "length (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) = length ((t[i := t ! j])[j := t ! i])" by (metis (mono_tags, lifting) \t = take i t @ [t ! i] @ take (j - (i + 1)) (drop (i + 1) t) @ [t ! j] @ drop (j + 1) t\ \take (j - (i + 1)) (drop (i + 1) t) = []\ length_append length_list_update list.size(4) self_append_conv2) moreover have "\k. k < length ((t[i := t ! j])[j := t ! i]) \ (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = ((t[i := t ! j])[j := t ! i]) ! k" proof - fix k assume "k < length ((t[i := t ! j])[j := t ! i])" show "(take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = ((t[i := t ! j])[j := t ! i]) ! k" proof (cases "k = i \ k = j") case True then show ?thesis proof (elim disjE) assume "k = i" then show ?thesis by (metis (no_types, lifting) \k < length (t[i := t ! j, j := t ! i])\ append_Cons le_eq_less_or_eq length_list_update length_take min.absorb2 nth_append_length nth_list_update_eq nth_list_update_neq) next assume "k = j" then show ?thesis by (metis (no_types, lifting) "0.prems"(4) Suc_eq_plus1 \j = i + 1\ \k < length (t[i := t ! j, j := t ! i])\ append.assoc append_Cons le_eq_less_or_eq length_append_singleton length_list_update length_take min.absorb2 nth_append_length nth_list_update postrecording_event) qed next case knij: False then show ?thesis proof (cases "k < i") case True then show ?thesis by (metis (no_types, lifting) "0.prems"(2) \j = i + 1\ add_lessD1 length_take less_imp_le_nat min.absorb2 not_less nth_append nth_list_update_neq nth_take) next case False then have "k > j" using \j = i + 1\ knij by linarith then have "(take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = drop (j+1) t ! (k-(j+1))" proof - assume a1: "j < k" have f2: "\n na. ((n::nat) < na) = (n \ na \ n \ na)" using nat_less_le by blast (* 0.0 ms *) have f3: "i + 0 = min (length t) i + (0 + 0)" using "0.prems"(2) \j = i + 1\ by linarith (* 8 ms *) have f4: "min (length t) i + Suc (0 + 0) = length (take i t) + length [t ! j]" by force (* 4 ms *) have f5: "take i t @ [t ! j] @ [] = take i t @ [t ! j]" by auto (* 0.0 ms *) have "j = length (take i t @ [t ! j] @ [])" using f3 by (simp add: \j = i + 1\) (* 4 ms *) then have "j + 1 = length (take i t @ [t ! j] @ [t ! i])" by fastforce (* 4 ms *) then show ?thesis using f5 f4 f3 f2 a1 by (metis (no_types) One_nat_def \j = i + 1\ add_Suc_right append.assoc length_append less_antisym list.size(3) not_less nth_append) (* 284 ms *) qed moreover have "(t[i := t ! j])[j := t ! i] ! k = drop (j+1) ((t[i := t ! j])[j := t ! i]) ! (k-(j+1))" using "0.prems"(2) \j < k\ by auto moreover have "drop (j+1) ((t[i := t ! j])[j := t ! i]) = drop (j+1) t" using "0.prems"(1) by auto ultimately show ?thesis by simp qed qed qed ultimately show ?thesis by (simp add: list_eq_iff_nth_eq) qed moreover have "\k. k \ j + 1 \ S t k = S ((t[i := t ! j])[j := t ! i]) k" proof (rule allI, rule impI) fix k assume "k \ j + 1" let ?newt = "((t[i := t ! j])[j := t ! i])" have "trace init (take k ?newt) (S ?newt k)" using calculation(1) calculation(2) exists_trace_for_any_i by auto have "take k ?newt = take (j+1) ?newt @ take (k - (j+1)) (drop (j+1) ?newt)" by (metis \j + 1 \ k\ le_add_diff_inverse take_add) have same_traces: "drop (j+1) t = drop (j+1) ?newt" by (metis "0.prems"(1) Suc_eq_plus1 \j = i + 1\ drop_update_cancel less_SucI less_add_same_cancel1) have "trace init (take (j+1) ((t[i := t ! j])[j := t ! i])) (S t (j+1))" by (metis (no_types, lifting) \j = i + 1\ \take (j - 1) t @ [t ! j, t ! i] = (take (j + 1) t)[i := t ! j, j := t ! i]\ add_diff_cancel_right' local.swap s take_update_swap trace_trans) moreover have "trace init (take (j+1) ?newt) (S ?newt (j+1))" using \take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\ \trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\ exists_trace_for_any_i by auto ultimately have "S ?newt (j+1) = S t (j+1)" using trace_and_start_determines_end by blast have "trace (S t (j+1)) (take (k - (j+1)) (drop (j+1) t)) (S t k)" using "0.prems"(6) \j + 1 \ k\ exists_trace_for_any_i_j by blast moreover have "trace (S ?newt (j+1)) (take (k - (j+1)) (drop (j+1) ?newt)) (S ?newt k)" using \j + 1 \ k\ \take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\ \trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\ exists_trace_for_any_i_j by fastforce ultimately show "S t k = S ?newt k" using \S (t[i := t ! j, j := t ! i]) (j + 1) = S t (j + 1)\ same_traces trace_and_start_determines_end by auto qed moreover have "\k. k \ i \ S t k = S ((t[i := t ! j])[j := t ! i]) k" proof (rule allI, rule impI) fix k assume "k \ i" let ?newt = "((t[i := t ! j])[j := t ! i])" have "trace init (take k t) (S t k)" using "0.prems"(6) exists_trace_for_any_i by blast moreover have "trace init (take k ?newt) (S ?newt k)" using \take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\ \trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\ exists_trace_for_any_i by auto moreover have "take k t = take k ?newt" using "0.prems"(1) \k \ i\ by auto ultimately show "S t k = S ?newt k" by (simp add: trace_and_start_determines_end) qed moreover have "prerecording_event (swap_events i j t) i" proof - have "~ has_snapshotted (S ((t[i := t ! j])[j := t ! i]) i) ?q" by (metis "0.prems"(6) \j = i + 1\ add.right_neutral calculation(4) le_add1 nsq snapshot_stable_ver_3) moreover have "regular_event ((t[i := t ! j])[j := t ! i] ! i)" by (metis "0.prems"(4) "0.prems"(5) \occurs_on (t ! i) \ occurs_on (t ! j)\ nth_list_update_eq nth_list_update_neq postrecording_event prerecording_event) moreover have "i < length ((t[i := t ! j])[j := t ! i])" using "0.prems"(1) "0.prems"(2) by auto ultimately show ?thesis unfolding prerecording_event by (metis (no_types, opaque_lifting) "0.prems"(1) \take (j - (i + 1)) (drop (i + 1) t) = []\ \take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\ append_Cons length_list_update nat_less_le nth_list_update_eq nth_list_update_neq self_append_conv2) qed moreover have "postrecording_event (swap_events i j t) (i+1)" proof - have "has_snapshotted (S ((t[i := t ! j])[j := t ! i]) (i+1)) ?p" by (metis "0.prems"(4) add.right_neutral calculation(1) calculation(2) calculation(4) le_add1 postrecording_event snapshot_stable_ver_3) moreover have "regular_event ((t[i := t ! j])[j := t ! i] ! j)" using "0.prems"(2) "0.prems"(4) length_list_update postrecording_event by auto moreover have "j < length t" using "0.prems" by auto ultimately show ?thesis unfolding postrecording_event by (metis \j = i + 1\ length_list_update nth_list_update_eq swap_neighbors_2) qed moreover have "\k. k > i+1 \ k < j+1 \ ~ regular_event ((swap_events i j t) ! k)" using "0" by force ultimately show ?case using \j = i + 1\ by force next case (Suc n) let ?p = "occurs_on (t ! i)" let ?q = "occurs_on (t ! j)" let ?t = "take ((j+1) - i) (drop i t)" let ?subt = "take (j - (i+1)) (drop (i+1) t)" let ?subt' = "take ((j-1) - (i+1)) (drop (i+1) t)" have sp: "has_snapshotted (S t i) ?p" using Suc.prems postrecording_event prerecording_event by blast have nsq: "~ has_snapshotted (S t j) ?q" using Suc.prems postrecording_event prerecording_event by blast have "?p \ ?q" using Suc.prems computation.post_before_pre_different_processes computation_axioms by blast have "?subt \ Nil" using Suc.hyps(2) Suc.prems(1) Suc.prems(2) by auto have "?subt' = butlast ?subt" by (metis Suc.prems(2) Suc_eq_plus1 butlast_drop butlast_take drop_take less_imp_le_nat) + have \i < length t\ + using \postrecording_event t i\ postrecording_event [of t i] by simp + then have step: \S t i \ t ! i \ S t (Suc i)\ + using \trace init t final\ by (rule step_Suc) have "?t = t ! i # ?subt @ [t ! j]" proof - have f1: "Suc j - i = Suc (j - i)" using Suc.prems(1) Suc_diff_le le_simps(1) by presburger have f2: "t ! i # drop (Suc i) t = drop i t" by (meson Cons_nth_drop_Suc Suc.prems(1) Suc.prems(2) less_trans) have f3: "t ! j # drop (Suc j) t = drop j t" using Cons_nth_drop_Suc Suc.prems(2) by blast have f4: "j - (i + 1) + (i + 1) = j" using Suc.prems(1) by force have "j - (i + 1) + Suc 0 = j - i" using Suc.prems(1) Suc_diff_Suc by presburger then show ?thesis using f4 f3 f2 f1 by (metis One_nat_def Suc.hyps(2) Suc_eq_plus1 drop_drop take_Suc_Cons take_add take_eq_Nil) qed then have "trace (S t i) ?t (S t (j+1))" by (metis Suc.prems(1) Suc.prems(6) Suc_eq_plus1 exists_trace_for_any_i_j less_SucI nat_less_le) then have reg_tr_1: "trace (S t i) (t ! i # ?subt) (S t j)" - by (metis (no_types, opaque_lifting) Suc.hyps(2) Suc.prems(1) Suc.prems(4) Suc.prems(6) Suc_eq_plus1 discrete exists_trace_for_any_i_j postrecording_event step_Suc tr_step) + using \i < j\ \i < length t\ \trace init t final\ step + by simp (meson exists_trace_for_any_i_j less_eq_Suc_le trace.simps) have reg_st_2: "(S t j) \ (t ! j) \ (S t (j+1))" using Suc.prems(2) Suc.prems(6) step_Suc by auto have "?subt = ?subt' @ [t ! (j-1)]" proof - have f1: "\n es. \ n < length es \ take n es @ [hd (drop n es)::('a, 'b, 'c) event] = take (Suc n) es" by (meson take_hd_drop) (* 0.0 ms *) have f2: "j - 1 - (i + 1) = n" by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 diff_Suc_1 diff_diff_left plus_1_eq_Suc) (* 28 ms *) have f3: "\n na. \ n < na \ Suc n \ na" using Suc_leI by blast (* 0.0 ms *) then have f4: "Suc i \ j - 1" by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 diff_diff_left plus_1_eq_Suc zero_less_Suc zero_less_diff) (* 12 ms *) have f5: "i + 1 < j" by (metis Suc.hyps(2) zero_less_Suc zero_less_diff) (* 4 ms *) then have f6: "t ! (j - 1) = hd (drop n (drop (i + 1) t))" using f4 f3 by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_eq_plus1 Suc_lessD add_Suc_right diff_Suc_1 drop_drop hd_drop_conv_nth le_add_diff_inverse2 plus_1_eq_Suc) (* 140 ms *) have "n < length (drop (i + 1) t)" using f5 f3 by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_eq_plus1 Suc_lessD drop_drop le_add_diff_inverse2 length_drop zero_less_diff) (* 144 ms *) then show ?thesis using f6 f2 f1 Suc.hyps(2) by presburger (* 4 ms *) qed then have reg_tr: "trace (S t i) (t ! i # ?subt') (S t (j-1))" proof - have f1: "j - Suc i = Suc n" using Suc.hyps(2) by presburger have f2: "length (take j t) = j" by (metis (no_types) Suc.prems(2) length_take min.absorb2 nat_le_linear not_less) have f3: "(t ! i # drop (Suc i) (take j t)) @ [t ! j] = drop i (take (Suc j) t)" by (metis (no_types) Suc_eq_plus1 \take (j + 1 - i) (drop i t) = t ! i # take (j - (i + 1)) (drop (i + 1) t) @ [t ! j]\ append_Cons drop_take) have f4: "Suc (i + n) = j - 1" using f1 by (metis (no_types) Suc.prems(1) Suc_diff_Suc add_Suc_right diff_Suc_1 le_add_diff_inverse nat_le_linear not_less) have "Suc (j - 1) = j" using f1 by simp then have f5: "butlast (take (Suc j) t) = take j t" using f4 f3 f2 f1 by (metis (no_types) Groups.add_ac(2) One_nat_def append_eq_conv_conj append_take_drop_id butlast_take diff_Suc_1 drop_drop length_append length_drop list.size(3) list.size(4) order_refl plus_1_eq_Suc plus_nat.simps(2) take_add take_all) have f6: "butlast (take j t) = take (j - 1) t" by (meson Suc.prems(2) butlast_take nat_le_linear not_less) have "drop (Suc i) (take j t) \ []" by (metis (no_types) Nil_is_append_conv Suc_eq_plus1 \take (j - (i + 1)) (drop (i + 1) t) = take (j - 1 - (i + 1)) (drop (i + 1) t) @ [t ! (j - 1)]\ drop_take list.distinct(1)) then show ?thesis using f6 f5 f4 f3 by (metis (no_types) Suc.prems(6) Suc_eq_plus1 butlast.simps(2) butlast_drop butlast_snoc drop_take exists_trace_for_any_i_j less_add_Suc1 nat_le_linear not_less) qed have reg_st_1: "(S t (j-1)) \ (t ! (j-1)) \ (S t j)" by (metis Suc.prems(1) Suc.prems(2) Suc.prems(6) Suc_lessD diff_Suc_1 less_imp_Suc_add step_Suc) have "~ regular_event (t ! (j-1))" using Suc.prems(3) \take (j - (i + 1)) (drop (i + 1) t) \ []\ less_diff_conv by auto moreover have "regular_event (t ! j)" using Suc.prems(5) computation.prerecording_event computation_axioms by blast moreover have "can_occur (t ! j) (S t j)" using happen_implies_can_occur reg_tr_1 reg_st_2 by blast moreover have njmiq: "occurs_on (t ! (j-1)) \ ?q" proof (rule ccontr) assume "~ occurs_on (t ! (j-1)) \ ?q" then have "occurs_on (t ! (j-1)) = ?q" by simp then have "has_snapshotted (S t j) ?q" using Suc.prems(6) calculation(1) diff_le_self nonregular_event_induces_snapshot reg_st_1 snapshot_stable_ver_2 by blast then show False using nsq by simp qed ultimately have "can_occur (t ! j) (S t (j-1))" using reg_tr reg_st_1 event_can_go_back_if_no_sender by auto then obtain d where new_st_1: "(S t (j-1)) \ (t ! j) \ d" using exists_next_if_can_occur by blast then have "trace (S t i) (t ! i # ?subt' @ [t ! j]) d" using reg_tr trace_snoc by fastforce moreover have "can_occur (t ! (j-1)) d" using \(S t (j-1)) \ t ! j \ d\ \occurs_on (t ! (j - 1)) \ occurs_on (t ! j)\ event_stays_valid_if_no_occurrence happen_implies_can_occur reg_st_1 by auto moreover obtain e where new_st_2: "d \ (t ! (j-1)) \ e" using calculation(2) exists_next_if_can_occur by blast have pre_swap: "e = (S t (j+1))" proof - have "states e = states (S t (j+1))" proof (rule ext) fix p have "states (S t (j-1)) p = states (S t j) p" using no_state_change_if_nonregular_event\~ regular_event (t ! (j-1))\ reg_st_1 by auto moreover have "states d p = states e p" using no_state_change_if_nonregular_event\~ regular_event (t ! (j-1))\ new_st_2 by auto moreover have "states d p = states (S t (j+1)) p" proof - have "\a. states (S t (j + 1)) a = states d a" by (meson \\ regular_event (t ! (j - 1))\ new_st_1 no_state_change_if_nonregular_event reg_st_1 reg_st_2 same_state_implies_same_result_state) then show ?thesis by presburger qed ultimately show "states e p = states (S t (j+1)) p" by simp qed moreover have "msgs e = msgs (S t (j+1))" proof (rule ext) fix cid have "isTrans (t ! j) \ isSend (t ! j) \ isRecv (t ! j)" using \regular_event (t ! j)\ by auto moreover have "isSnapshot (t ! (j-1)) \ isRecvMarker (t ! (j-1))" using nonregular_event \~ regular_event (t ! (j-1))\ by auto ultimately show "msgs e cid = msgs (S t (j+1)) cid" proof (elim disjE, goal_cases) case 1 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Trans_Snapshot by auto next case 2 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Trans_RecvMarker by auto next case 3 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Send_Snapshot by auto next case 4 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Recv_Snapshot by auto next case 5 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Send_RecvMarker by auto next case 6 then show ?case using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Recv_RecvMarker by auto qed qed moreover have "process_snapshot e = process_snapshot (S t (j+1))" proof (rule ext) fix p have "process_snapshot (S t j) p = process_snapshot (S t (j+1)) p" using \regular_event (t ! j)\ reg_st_2 regular_event_preserves_process_snapshots by blast moreover have "process_snapshot (S t (j-1)) p = process_snapshot d p" using \regular_event (t ! j)\ new_st_1 regular_event_preserves_process_snapshots by blast moreover have "process_snapshot e p = process_snapshot (S t j) p" proof - have "occurs_on (t ! j) = p \ ps e p = ps (S t j) p" using calculation(2) new_st_2 njmiq no_state_change_if_no_event reg_st_1 by force then show ?thesis by (meson new_st_1 new_st_2 no_state_change_if_no_event reg_st_1 same_snapshot_state_implies_same_result_snapshot_state) qed ultimately show "process_snapshot e p = process_snapshot (S t (j+1)) p" by simp qed moreover have "cs e = cs (S t (j+1))" proof (rule ext) fix cid have "isTrans (t ! j) \ isSend (t ! j) \ isRecv (t ! j)" using \regular_event (t ! j)\ by auto moreover have "isSnapshot (t ! (j-1)) \ isRecvMarker (t ! (j-1))" using nonregular_event \~ regular_event (t ! (j-1))\ by auto ultimately show "cs e cid = cs (S t (j+1)) cid" proof (elim disjE, goal_cases) case 1 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Trans_Snapshot by auto next case 2 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Trans_RecvMarker by auto next case 3 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Send_Snapshot by auto next case 4 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Recv_Snapshot njmiq by auto next case 5 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Send_RecvMarker by auto next case 6 then show ?case using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Recv_RecvMarker njmiq by auto qed qed ultimately show ?thesis by auto qed let ?it = "(t[j-1 := t ! j])[j := t ! (j-1)]" have same_prefix: "take (j-1) ?it = take (j-1) t" by simp have same_suffix: "drop (j+1) ?it = drop (j+1) t" by simp have trace_prefix: "trace init (take (j-1) ?it) (S t (j-1))" using Suc.prems(6) exists_trace_for_any_i by auto have "?it = take (j-1) t @ [t ! j, t ! (j-1)] @ drop (j+1) t" proof - have "1 < j" by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 add_lessD1 plus_1_eq_Suc zero_less_Suc zero_less_diff) (* 12 ms *) then have "j - 1 + 1 = j" by (metis (no_types) le_add_diff_inverse2 nat_less_le) (* 4 ms *) then show ?thesis by (metis (no_types) Suc.prems(2) Suc_eq_plus1 add_Suc_right one_add_one swap_neighbors) (* 76 ms *) qed have "trace (S t (j-1)) [t ! j, t ! (j-1)] (S t (j+1))" by (metis new_st_1 new_st_2 pre_swap trace.simps) have "trace init (take (j+1) t @ drop (j+1) t) final" by (simp add: Suc.prems(6)) then have "trace init (take (j+1) t) (S t (j+1)) \ trace (S t (j+1)) (drop (j+1) t) final" using Suc.prems(6) exists_trace_for_any_i split_trace trace_and_start_determines_end by blast then have trace_suffix: "trace (S t (j+1)) (drop (j+1) ?it) final" using same_suffix by simp have "trace init ?it final" by (metis (no_types, lifting) \t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\ \trace (S t (j + 1)) (drop (j + 1) (t[j - 1 := t ! j, j := t ! (j - 1)])) final\ \trace (S t (j - 1)) [t ! j, t ! (j - 1)] (S t (j + 1))\ \trace init (take (j - 1) (t[j - 1 := t ! j, j := t ! (j - 1)])) (S t (j - 1))\ same_prefix same_suffix trace_trans) have suffix_same_states: "\k. k > j \ S t k = S ?it k" proof (rule allI, rule impI) fix k assume "k > j" have eq_trace: "drop (j+1) t = drop (j+1) ?it" by simp have "trace init (take (j+1) ?it) (S ?it (j+1))" using \trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\ exists_trace_for_any_i by blast moreover have "trace init (take (j+1) ?it) (S t (j+1))" proof - have f1: "\es esa esb esc. (esb::('a, 'b, 'c) event list) @ es \ esa @ esc @ es \ esa @ esc = esb" by auto have f2: "take (j + 1) (t[j - 1 := t ! j, j := t ! (j - 1)]) @ drop (j + 1) t = t [j - 1 := t ! j, j := t ! (j - 1)]" by (metis append_take_drop_id same_suffix) have "trace init (take (j - 1) t @ [t ! j, t ! (j - 1)]) (S t (j + 1))" using \trace (S t (j - 1)) [t ! j, t ! (j - 1)] (S t (j + 1))\ same_prefix trace_prefix trace_trans by presburger then show ?thesis using f2 f1 by (metis (no_types) \t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\) qed ultimately have eq_start: "S ?it (j+1) = S t (j+1)" using trace_and_start_determines_end by blast then have "take k ?it = take (j+1) ?it @ take (k - (j+1)) (drop (j+1) ?it)" by (metis Suc_eq_plus1 Suc_leI \j < k\ le_add_diff_inverse take_add) have "trace (S ?it (j+1)) (take (k - (j+1)) (drop (j+1) ?it)) (S ?it k)" by (metis Suc_eq_plus1 Suc_leI \j < k\ \trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\ exists_trace_for_any_i_j) moreover have "trace (S t (j+1)) (take (k - (j+1)) (drop (j+1) t)) (S t k)" using Suc.prems(6) \j < k\ exists_trace_for_any_i_j by fastforce ultimately show "S t k = S ?it k" using eq_start trace_and_start_determines_end by auto qed have prefix_same_states: "\k. k < j \ S t k = S ?it k" proof (rule allI, rule impI) fix k assume "k < j" have "trace init (take k t) (S t k)" using Suc.prems(6) exists_trace_for_any_i by blast moreover have "trace init (take k ?it) (S ?it k)" by (meson \trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\ exists_trace_for_any_i) ultimately show "S t k = S ?it k" using \k < j\ s_def by auto qed moreover have "j - 1 < length ?it" using Suc.prems(2) by auto moreover have "prerecording_event ?it (j-1)" proof - have f1: "t[j - 1 := t ! j, j := t ! (j - 1)] ! (j - 1) = t[j - 1 := t ! j] ! (j - 1)" by (metis (no_types) njmiq nth_list_update_neq) (* 28 ms *) have "j \ 0" by (metis (no_types) Suc.prems(1) not_less_zero) (* 0.0 ms *) then have "\ j < 1" by blast (* 0.0 ms *) then have "S t (j - 1) = S (t[j - 1 := t ! j, j := t ! (j - 1)]) (j - 1)" by (simp add: prefix_same_states) (* 8 ms *) then show ?thesis using f1 by (metis \regular_event (t ! j)\ calculation(4) computation.prerecording_event computation_axioms length_list_update njmiq no_state_change_if_no_event nsq nth_list_update_eq reg_st_1) (* 456 ms *) qed moreover have "postrecording_event ?it i" proof - have "i < length ?it" using Suc.prems(4) postrecording_event by auto then show ?thesis proof - assume "i < length (t[j - 1 := t ! j, j := t ! (j - 1)])" have "i < j - 1" by (metis (no_types) Suc.hyps(2) cancel_ab_semigroup_add_class.diff_right_commute diff_diff_left zero_less_Suc zero_less_diff) then show ?thesis using Suc.prems(1) Suc.prems(4) postrecording_event prefix_same_states by auto qed qed moreover have "i < j - 1" using Suc.hyps(2) by auto moreover have "\k. i < k \ k < (j-1) \ ~ regular_event (?it ! k)" proof (rule allI, rule impI) fix k assume "i < k \ k < (j-1)" show "~ regular_event (?it ! k)" using Suc.prems(3) \i < k \ k < j - 1\ by force qed moreover have "(j-1) - (i+1) = n" using Suc.prems Suc.hyps by auto ultimately have ind: "trace init (swap_events i (j-1) ?it) final \ (\k. k \ (j-1)+1 \ S (swap_events i (j-1) ?it) k = S ?it k) \ (\k. k \ i \ S (swap_events i (j-1) ?it) k = S ?it k) \ prerecording_event (swap_events i (j-1) ?it) i \ postrecording_event (swap_events i (j-1) ?it) (i+1) \ (\k. k > i+1 \ k < (j-1)+1 \ ~ regular_event ((swap_events i (j-1) ?it) ! k))" using Suc.hyps \trace init ?it final\ by blast then have new_trace: "trace init (swap_events i (j-1) ?it) final" by blast have equal_suffix_states: "\k. k \ j \ S (swap_events i (j-1) ?it) k = S ?it k" using Suc.prems(1) ind by simp have equal_prefix_states: "\k. k \ i \ S (swap_events i (j-1) ?it) k = S ?it k" using ind by blast have neighboring_events_shifted: "\k. k > i+1 \ k < j \ ~ regular_event ((swap_events i (j-1) ?it) ! k)" using ind by force let ?itn = "swap_events i (j-1) ?it" have "?itn = swap_events i j t" proof - have f1: "i \ j - 1" using \i < j - 1\ less_imp_le_nat by blast have "t ! j # [t ! (j - 1)] @ drop (j + 1) t = drop (j - 1) (take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t)" using \t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\ same_prefix by force then have f2: "t[j - 1 := t ! j, j := t ! (j - 1)] ! (j - 1) = t ! j \ drop (j - 1 + 1) (t[j - 1 := t ! j, j := t ! (j - 1)]) = t ! (j - 1) # [] @ drop (j + 1) t" by (metis (no_types) Cons_nth_drop_Suc Suc_eq_plus1 \j - 1 < length (t[j - 1 := t ! j, j := t ! (j - 1)])\ \t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\ append_Cons list.inject) have "t ! i = t[j - 1 := t ! j, j := t ! (j - 1)] ! i" by (metis (no_types) Suc.prems(1) \i < j - 1\ nat_neq_iff nth_list_update_neq) then show ?thesis using f2 f1 by (metis (no_types) Suc.prems(1) \take (j - (i + 1)) (drop (i + 1) t) = take (j - 1 - (i + 1)) (drop (i + 1) t) @ [t ! (j - 1)]\ append.assoc append_Cons drop_take less_imp_le_nat same_prefix take_update_cancel) qed moreover have "\k. k \ i \ S t k = S ?itn k" using Suc.prems(1) equal_prefix_states prefix_same_states by auto moreover have "\k. k \ j + 1 \ S t k = S ?itn k" by (metis (no_types, lifting) Suc_eq_plus1 add_lessD1 equal_suffix_states lessI nat_less_le suffix_same_states) moreover have "\k. k > i+1 \ k < j+1 \ ~ regular_event (?itn ! k)" proof - have "~ regular_event (?itn ! j)" proof - have f1: "j - 1 < length t" using \j - 1 < length (t[j - 1 := t ! j, j := t ! (j - 1)])\ by force have f2: "\n na es. \ n < na \ \ na < length es \ drop (Suc na) (take n es @ [hd (drop na es), es ! n::('a, 'b, 'c) event] @ take (na - Suc n) (drop (Suc n) es) @ drop (Suc na) es) = drop (Suc na) es" by (metis Suc_eq_plus1 hd_drop_conv_nth swap_identical_tails) have f3: "t ! j = hd (drop j t)" by (simp add: Suc.prems(2) hd_drop_conv_nth) have "\ j < 1" using Suc.prems(1) by blast then have "\ regular_event (hd (drop j (take i (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)]) @ [hd (drop (j - 1) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)])), t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)] ! i] @ take (j - 1 - Suc i) (drop (Suc i) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)])) @ drop (Suc (j - 1)) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)]))))" using f2 f1 by (metis (no_types) Suc.prems(2) \\ regular_event (t ! (j - 1))\ \i < j - 1\ add_diff_inverse_nat hd_drop_conv_nth length_list_update nth_list_update_eq plus_1_eq_Suc) then show ?thesis using f3 f1 by (metis Suc.prems(2) Suc_eq_plus1 \i < j - 1\ hd_drop_conv_nth length_list_update swap_identical_length) qed then show ?thesis by (metis Suc_eq_plus1 less_Suc_eq neighboring_events_shifted) qed ultimately show ?case using ind by presburger qed subsubsection \Relating configurations and the computed snapshot\ definition ps_equal_to_snapshot where "ps_equal_to_snapshot c c' \ \p. Some (states c p) = process_snapshot c' p" definition cs_equal_to_snapshot where "cs_equal_to_snapshot c c' \ \cid. channel cid \ None \ filter ((\) Marker) (msgs c cid) = map Msg (fst (channel_snapshot c' cid))" definition state_equal_to_snapshot where "state_equal_to_snapshot c c' \ ps_equal_to_snapshot c c' \ cs_equal_to_snapshot c c'" lemma init_is_s_t_0: assumes "trace init t final" shows "init = (S t 0)" by (metis assms exists_trace_for_any_i take_eq_Nil tr_init trace_and_start_determines_end) lemma final_is_s_t_len_t: assumes "trace init t final" shows "final = S t (length t)" by (metis assms exists_trace_for_any_i order_refl take_all trace_and_start_determines_end) lemma snapshot_event: assumes "trace init t final" and "~ has_snapshotted (S t i) p" and "has_snapshotted (S t (i+1)) p" shows "isSnapshot (t ! i) \ isRecvMarker (t ! i)" proof - have "(S t i) \ (t ! i) \ (S t (i+1))" by (metis Suc_eq_plus1 assms(1) assms(2) assms(3) distributed_system.step_Suc computation_axioms computation_def nat_less_le not_less not_less_eq s_def take_all) then show ?thesis using assms(2) assms(3) nonregular_event regular_event_cannot_induce_snapshot by blast qed lemma snapshot_state: assumes "trace init t final" and "states (S t i) p = u" and "~ has_snapshotted (S t i) p" and "has_snapshotted (S t (i+1)) p" shows "ps (S t (i+1)) p = Some u" proof - have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis add.commute assms(1) assms(3) assms(4) le_SucI le_eq_less_or_eq le_refl nat_neq_iff no_change_if_ge_length_t plus_1_eq_Suc step_Suc) let ?q = "occurs_on (t ! i)" have qp: "?q = p" proof (rule ccontr) assume "?q \ p" then have "has_snapshotted (S t (i+1)) p = has_snapshotted (S t i) p" using local.step no_state_change_if_no_event by auto then show False using assms by simp qed have "isSnapshot (t ! i) \ isRecvMarker (t ! i)" using assms snapshot_event by auto then show ?thesis proof (elim disjE, goal_cases) case 1 then have "t ! i = Snapshot p" by (metis event.collapse(4) qp) then show ?thesis using assms(2) local.step by auto next case 2 then obtain cid' q where "t ! i = RecvMarker cid' p q" by (metis event.collapse(5) qp) then show ?thesis using assms step by auto qed qed lemma snapshot_state_unchanged_trace_2: shows "\ trace init t final; i \ j; j \ length t; ps (S t i) p = Some u \ \ ps (S t j) p = Some u" proof (induct i j rule:S_induct) case S_init then show ?case by simp next case S_step then show ?case using snapshot_state_unchanged by auto qed lemma no_recording_cs_if_not_snapshotted: shows "\ trace init t final; ~ has_snapshotted (S t i) p; channel cid = Some (q, p) \ \ cs (S t i) cid = cs init cid" proof (induct i) case 0 then show ?case by (metis exists_trace_for_any_i list.discI take_eq_Nil trace.simps) next case (Suc i) have "Suc i < length t" proof - have "has_snapshotted final p" using all_processes_snapshotted_in_final_state valid by blast show ?thesis proof (rule ccontr) assume "~ Suc i < length t" then have "Suc i \ length t" by simp then have "has_snapshotted (S t (Suc i)) p" using Suc.prems(1) \ps final p \ None\ final_is_s_t_len_t snapshot_stable_ver_3 by blast then show False using Suc by simp qed qed then have t_dec: "trace init (take i t) (S t i) \ (S t i) \ (t ! i) \ (S t (Suc i))" using Suc.prems(1) exists_trace_for_any_i step_Suc by auto moreover have step: "(S t i) \ (t ! i) \ (S t (Suc i))" using calculation by simp ultimately have IH: "cs (S t i) cid = cs init cid" using Suc.hyps Suc.prems(1) Suc.prems(2) Suc.prems(3) snapshot_state_unchanged by fastforce then show ?case proof (cases "t ! i") case (Snapshot r) have "r \ p" proof (rule ccontr) assume "~ r \ p" then have "r = p" by simp then have "has_snapshotted (S t (Suc i)) p" using Snapshot step by auto then show False using Suc by simp qed then have "cs (S t i) cid = cs (S t (Suc i)) cid" using Snapshot Suc.prems(3) local.step by auto then show ?thesis using IH by simp next case (RecvMarker cid' r s) have "r \ p" proof (rule ccontr) assume "~ r \ p" then have "r = p" by simp then have "has_snapshotted (S t (Suc i)) p" using RecvMarker t_dec recv_marker_means_snapshotted_1 by blast then show False using Suc by simp qed have "cid' \ cid" proof (rule ccontr) assume "~ cid' \ cid" then have "channel cid' = Some (s, r)" using t_dec can_occur_def RecvMarker by simp then show False using Suc.prems(3) \\ cid' \ cid\ \r \ p\ by auto qed then have "cs (S t i) cid = cs (S t (Suc i)) cid" proof - have "\s. channel cid = Some (s, r)" using \r \ p\ Suc by simp with RecvMarker t_dec \cid' \ cid\ \r \ p\ Suc.prems(3) show ?thesis by (cases "has_snapshotted (S t i) r", auto) qed then show ?thesis using IH by simp next case (Trans r u u') then show ?thesis using IH t_dec by auto next case (Send cid' r s u u' m) then show ?thesis using IH local.step by auto next case (Recv cid' r s u u' m) then have "snd (cs (S t i) cid) = NotStarted" by (simp add: IH no_initial_channel_snapshot) with Recv step Suc show ?thesis by (cases "cid' = cid", auto) qed qed lemma cs_done_implies_has_snapshotted: assumes "trace init t final" and "snd (cs (S t i) cid) = Done" and "channel cid = Some (p, q)" shows "has_snapshotted (S t i) q" proof - show ?thesis using assms no_initial_channel_snapshot no_recording_cs_if_not_snapshotted by fastforce qed lemma exactly_one_snapshot: assumes "trace init t final" shows "\!i. ~ has_snapshotted (S t i) p \ has_snapshotted (S t (i+1)) p" (is ?P) proof - have "~ has_snapshotted init p" using no_initial_process_snapshot by auto moreover have "has_snapshotted final p" using all_processes_snapshotted_in_final_state valid by blast moreover have "trace (S t 0) t (S t (length t))" using assms final_is_s_t_len_t init_is_s_t_0 by auto ultimately have ex_snap: "\i. ~ has_snapshotted (S t i) p \ has_snapshotted (S t (i+1)) p" using assms exists_snapshot_for_all_p by auto show ?thesis proof (rule ccontr) assume "~ ?P" then have "\i j. (i \ j) \ ~ has_snapshotted (S t i) p \ has_snapshotted (S t (i+1)) p \ ~ has_snapshotted (S t j) p \ has_snapshotted (S t (j+1)) p" using ex_snap by blast then have "\i j. (i < j) \ ~ has_snapshotted (S t i) p \ has_snapshotted (S t (i+1)) p \ ~ has_snapshotted (S t j) p \ has_snapshotted (S t (j+1)) p" by (meson linorder_neqE_nat) then obtain i j where "i < j" "~ has_snapshotted (S t i) p" "has_snapshotted (S t (i+1)) p" "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p" by blast have "trace (S t (i+1)) (take (j - (i+1)) (drop (i+1) t)) (S t j)" using \i < j\ assms exists_trace_for_any_i_j by fastforce then have "has_snapshotted (S t j) p" using \ps (S t (i + 1)) p \ None\ snapshot_stable by blast then show False using \~ has_snapshotted (S t j) p\ by simp qed qed lemma initial_cs_changes_implies_nonregular_event: assumes "trace init t final" and "snd (cs (S t i) cid) = NotStarted" and "snd (cs (S t (i+1)) cid) \ NotStarted" and "channel cid = Some (p, q)" shows "~ regular_event (t ! i)" proof - have "i < length t" proof (rule ccontr) assume "~ i < length t" then have "S t i = S t (i+1)" using assms(1) no_change_if_ge_length_t by auto then show False using assms by presburger qed then have step: "(S t i) \ (t ! i) \ (S t (i+1))" using assms(1) step_Suc by auto show ?thesis proof (rule ccontr) assume "~ ~ regular_event (t ! i)" then have "regular_event (t ! i)" by simp then have "cs (S t i) cid = cs (S t (i+1)) cid" proof (cases "isRecv (t ! i)") case False then show ?thesis using \regular_event (t ! i)\ local.step no_cs_change_if_no_event by blast next case True then obtain cid' r s u u' m where Recv: "t ! i = Recv cid' r s u u' m" by (meson isRecv_def) with assms step show ?thesis proof (cases "cid = cid'") case True then show ?thesis using assms step Recv by simp next case False then show ?thesis using assms step Recv by simp qed qed then show False using assms by simp qed qed lemma cs_in_initial_state_implies_not_snapshotted: assumes "trace init t final" and "snd (cs (S t i) cid) = NotStarted" and "channel cid = Some (p, q)" shows "~ has_snapshotted (S t i) q" proof (rule ccontr) assume "~ ~ has_snapshotted (S t i) q" then obtain j where "j < i" "~ has_snapshotted (S t j) q" "has_snapshotted (S t (j+1)) q" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p computation.snapshot_stable_ver_3 computation_axioms nat_le_linear order_le_less) have step_j: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis \\ \ ps (S t i) q \ None\ \\ ps (S t j) q \ None\ \j < i\ add.commute assms(1) linorder_neqE_nat no_change_if_ge_length_t order_le_less order_refl plus_1_eq_Suc step_Suc) have tr_j_i: "trace (S t (j+1)) (take (i - (j+1)) (drop (j+1) t)) (S t i)" using \j < i\ assms(1) exists_trace_for_any_i_j by fastforce have "~ regular_event (t ! j)" using step_j \\ ps (S t j) q \ None\ \ps (S t (j + 1)) q \ None\ regular_event_cannot_induce_snapshot by blast then have "isSnapshot (t ! j) \ isRecvMarker (t ! j)" using nonregular_event by auto then have "snd (cs (S t (j+1)) cid) \ NotStarted" proof (elim disjE, goal_cases) case 1 have "occurs_on (t ! j) = q" using \\ ps (S t j) q \ None\ \ps (S t (j + 1)) q \ None\ distributed_system.no_state_change_if_no_event distributed_system_axioms step_j by fastforce with 1 have "t ! j = Snapshot q" using isSnapshot_def by auto then show ?thesis using step_j assms by simp next case 2 have "occurs_on (t ! j) = q" using \\ ps (S t j) q \ None\ \ps (S t (j + 1)) q \ None\ distributed_system.no_state_change_if_no_event distributed_system_axioms step_j by fastforce with 2 obtain cid' s where RecvMarker: "t ! j = RecvMarker cid' q s" by (metis event.collapse(5)) then show ?thesis proof (cases "cid' = cid") case True then show ?thesis using RecvMarker step_j assms by simp next case False have "~ has_snapshotted (S t j) q" using \\ ps (S t j) q \ None\ by auto moreover have "\r. channel cid = Some (r, q)" by (simp add: assms(3)) ultimately show ?thesis using RecvMarker step_j assms False by simp qed qed then have "snd (cs (S t i) cid) \ NotStarted" using tr_j_i cs_not_not_started_stable_trace assms by blast then show False using assms by simp qed lemma nonregular_event_in_initial_state_implies_cs_changed: assumes "trace init t final" and "snd (cs (S t i) cid) = NotStarted" and "~ regular_event (t ! i)" and "occurs_on (t ! i) = q" and "channel cid = Some (p, q)" and "i < length t" shows "snd (cs (S t (i+1)) cid) \ NotStarted" proof - have step: "(S t i) \ (t ! i) \ (S t (i+1))" using step_Suc assms by auto have "isSnapshot (t ! i) \ isRecvMarker (t ! i)" using assms(3) nonregular_event by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then show ?thesis using assms cs_in_initial_state_implies_not_snapshotted local.step nonregular_event_induces_snapshot by blast next case 2 then show ?thesis by (metis assms(1) assms(2) assms(3) assms(4) assms(5) cs_in_initial_state_implies_not_snapshotted local.step nonregular_event_induces_snapshot) qed qed lemma cs_recording_implies_snapshot: assumes "trace init t final" and "snd (cs (S t i) cid) = Recording" and "channel cid = Some (p, q)" shows "has_snapshotted (S t i) q" proof (rule ccontr) assume "~ has_snapshotted (S t i) q" have "\ trace init t final; ~ has_snapshotted (S t i) p; channel cid = Some (p, q) \ \ snd (cs (S t i) cid) = NotStarted" proof (induct i) case 0 then show ?case using init_is_s_t_0 no_initial_channel_snapshot by auto next case (Suc n) have step: "(S t n) \ (t ! n) \ (S t (n+1))" by (metis Suc.prems(2) Suc_eq_plus1 all_processes_snapshotted_in_final_state assms(1) distributed_system.step_Suc distributed_system_axioms final_is_s_t_len_t le_add1 not_less snapshot_stable_ver_3) have "snd (cs (S t n) cid) = NotStarted" using Suc.hyps Suc.prems(2) assms snapshot_state_unchanged computation_axioms local.step by fastforce then show ?case by (metis Suc.prems(1) \\ ps (S t i) q \ None\ assms(2) assms(3) cs_not_not_started_stable_trace exists_trace_for_any_i no_recording_cs_if_not_snapshotted recording_state.simps(2)) qed then show False using \\ ps (S t i) q \ None\ assms computation.no_initial_channel_snapshot computation_axioms no_recording_cs_if_not_snapshotted by fastforce qed lemma cs_done_implies_both_snapshotted: assumes "trace init t final" and "snd (cs (S t i) cid) = Done" and "i < length t" and "channel cid = Some (p, q)" shows "has_snapshotted (S t i) p" "has_snapshotted (S t i) q" proof - have "trace init (take i t) (S t i)" using assms(1) exists_trace_for_any_i by blast then have "RecvMarker cid q p : set (take i t)" by (metis assms(1,2,4) cs_done_implies_has_snapshotted done_only_from_recv_marker_trace computation.no_initial_process_snapshot computation_axioms init_is_s_t_0 list.discI trace.simps) then obtain k where "t ! k = RecvMarker cid q p" "0 \ k" "k < i" by (metis add.right_neutral add_diff_cancel_right' append_Nil append_take_drop_id assms(1) exists_index take0) then have "has_snapshotted (S t (k+1)) q" by (metis (no_types, lifting) Suc_eq_plus1 Suc_leI assms(1,2,4) computation.cs_done_implies_has_snapshotted computation.no_change_if_ge_length_t computation_axioms less_le not_less_eq recv_marker_means_cs_Done) then show "has_snapshotted (S t i) q" using assms cs_done_implies_has_snapshotted by blast have step_k: "(S t k) \ (t ! k) \ (S t (k+1))" by (metis Suc_eq_plus1 \k < i\ add_lessD1 assms(1) assms(3) distributed_system.step_Suc distributed_system_axioms less_imp_add_positive) then have "Marker : set (msgs (S t k) cid)" proof - have "can_occur (t ! k) (S t k)" using happen_implies_can_occur step_k by blast then show ?thesis unfolding can_occur_def \t ! k = RecvMarker cid q p\ using hd_in_set by fastforce qed then have "has_snapshotted (S t k) p" using assms(1,4) no_marker_if_no_snapshot by blast then show "has_snapshotted (S t i) p" using \k < i\ assms(1) less_imp_le_nat snapshot_stable_ver_3 by blast qed lemma cs_done_implies_same_snapshots: assumes "trace init t final" "i \ j" "j \ length t" shows "snd (cs (S t i) cid) = Done \ channel cid = Some (p, q) \ cs (S t i) cid = cs (S t j) cid" using assms proof (induct i j rule: S_induct) case (S_init i) then show ?case by auto next case (S_step i j) have snap_p: "has_snapshotted (S t i) p" using S_step.hyps(1) S_step.hyps(2) S_step.prems(1,2) assms(1) cs_done_implies_both_snapshotted(1) by auto have snap_q: "has_snapshotted (S t i) q" using S_step.prems(1,2) assms(1) cs_done_implies_has_snapshotted by blast from S_step have "cs (S t i) cid = cs (S t (Suc i)) cid" proof (cases "t ! i") case (Snapshot r) from Snapshot S_step.hyps(3) snap_p have False if "r = p" using that by (auto simp: can_occur_def) moreover from Snapshot S_step.hyps(3) snap_q have False if "r = q" using that by (auto simp: can_occur_def) ultimately show ?thesis using Snapshot S_step by force next case (RecvMarker cid' r s) then show ?thesis proof (cases "has_snapshotted (S t i) r") case True with RecvMarker S_step show ?thesis proof (cases "cid = cid'") case True then have "cs (S t (Suc i)) cid = (fst (cs (S t i) cid), Done)" using RecvMarker S_step by simp then show ?thesis by (metis S_step.prems(1) prod.collapse) qed auto next case no_snap: False then show ?thesis proof (cases "cid = cid'") case True then have "cs (S t (Suc i)) cid = (fst (cs (S t i) cid), Done)" using RecvMarker S_step by simp then show ?thesis by (metis S_step.prems(1) prod.collapse) next case False then have "r \ p" using no_snap snap_p by auto moreover have "\s. channel cid = Some (s, r)" using S_step(5) assms(1) cs_done_implies_has_snapshotted no_snap by blast ultimately show ?thesis using RecvMarker S_step False no_snap by simp qed qed next case (Recv cid' r s u u' m) with S_step show ?thesis by (cases "cid = cid'", auto) qed auto with S_step show ?case by auto qed lemma snapshotted_and_not_done_implies_marker_in_channel: assumes "trace init t final" and "has_snapshotted (S t i) p" and "snd (cs (S t i) cid) \ Done" and "i \ length t" and "channel cid = Some (p, q)" shows "Marker : set (msgs (S t i) cid)" proof - obtain j where jj: "j < i" "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 assms(1) assms(2) exists_snapshot_for_all_p computation.snapshot_stable_ver_2 computation_axioms le_eq_less_or_eq nat_neq_iff) have step: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis \\ ps (S t j) p \ None\ \j < i\ add.commute assms(1) assms(2) linorder_neqE_nat no_change_if_ge_length_t order_le_less order_refl plus_1_eq_Suc step_Suc) then have "Marker : set (msgs (S t (j+1)) cid)" proof - have "~ regular_event (t ! j)" by (meson \\ ps (S t j) p \ None\ \ps (S t (j + 1)) p \ None\ distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms local.step) then have "isSnapshot (t ! j) \ isRecvMarker (t ! j)" using nonregular_event by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then obtain r where Snapshot: "t ! j = Snapshot r" by (meson isSnapshot_def) then have "r = p" using jj(2) jj(3) local.step by auto then show ?thesis using Snapshot assms step by simp next case 2 then obtain cid' s where RecvMarker: "t ! j = RecvMarker cid' p s" by (metis jj(2,3) distributed_system.no_state_change_if_no_event distributed_system_axioms event.sel(5) isRecvMarker_def local.step) moreover have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "snd (cs (S t (j+1)) cid) = Done" using RecvMarker step by simp then have "snd (cs (S t i) cid) = Done" proof - assume a1: "snd (cs (S t (j + 1)) cid) = Done" have f2: "ps (S t j) p = None" using jj(2) by blast have "j < length t" using assms(4) jj(1) by linarith then have "t ! j = RecvMarker cid q p" using f2 a1 assms(1) assms(5) cs_done_implies_both_snapshotted(1) done_only_from_recv_marker local.step by blast then show ?thesis using f2 by (metis (no_types) Suc_eq_plus1 assms(1) local.step recv_marker_means_snapshotted) qed then show False using assms by simp qed ultimately show ?thesis using jj assms step by auto qed qed show ?thesis proof (rule ccontr) let ?t = "take (i - (j+1)) (drop (j+1) t)" - have tr_j: "trace (S t (j+1)) ?t (S t i)" - by (metis \j < i\ assms(1) discrete exists_trace_for_any_i_j) + have tr_j: "trace (S t (j+1)) ?t (S t i)" + using assms(1) exists_trace_for_any_i_j jj(1) by fastforce assume "~ Marker : set (msgs (S t i) cid)" then obtain ev where "ev \ set ?t" "\p q. ev = RecvMarker cid p q" using \Marker \ set (msgs (S t (j + 1)) cid)\ marker_must_be_delivered_2_trace tr_j assms by blast obtain k where "t ! k = ev" "j < k" "k < i" using \ev \ set (take (i - (j + 1)) (drop (j + 1) t))\ assms(1) exists_index by fastforce have step_k: "(S t k) \ (t ! k) \ (S t (k+1))" proof - have "k < length t" using \k < i\ assms(4) by auto then show ?thesis using step_Suc assms by simp qed have "ev = RecvMarker cid q p" using assms step_k can_occur_def using \\p q. ev = RecvMarker cid p q\ \t ! k = ev\ by auto then have "snd (cs (S t (k+1)) cid) = Done" using \k < i\ \t ! k = ev\ assms(1) assms(4) recv_marker_means_cs_Done by auto - moreover have "trace (S t (k+1)) (take (i - (k+1)) (drop (k+1) t)) (S t i)" - by (meson \k < i\ assms(1) discrete exists_trace_for_any_i_j) - ultimately have "snd (cs (S t i) cid) = Done" - by (metis \k < i\ assms(1) assms(4) assms(5) cs_done_implies_same_snapshots discrete) - then show False using assms by simp + moreover have "trace (S t (k+1)) (take (i - (k+1)) (drop (k+1) t)) (S t i)" + using \k < i\ \trace init t final\ exists_trace_for_any_i_j by fastforce + ultimately have \snd (cs (S t i) cid) = Done\ + using cs_done_implies_same_snapshots [of t \Suc k\ i cid p q] \k < i\ assms(1) assms(4) assms(5) + by simp + then show False + using assms by simp qed qed lemma no_marker_left_in_final_state: assumes "trace init t final" shows "Marker \ set (msgs final cid)" (is ?P) proof (rule ccontr) assume "~ ?P" then obtain i where "i > length t" "Marker \ set (msgs (S t i) cid)" using assms l1 by (metis final_is_s_t_len_t le_neq_implies_less) then have "S t (length t) \ S t i" proof - have "msgs (S t i) cid \ msgs final cid" using \Marker \ set (msgs (S t i) cid)\ \~ ?P\ by auto then show ?thesis using final_is_s_t_len_t assms by auto qed moreover have "S t (length t) = S t i" using assms \i > length t\ less_imp_le no_change_if_ge_length_t by simp ultimately show False by simp qed lemma all_channels_done_in_final_state: assumes "trace init t final" and "channel cid = Some (p, q)" shows "snd (cs final cid) = Done" proof (rule ccontr) assume cs_not_done: "~ snd (cs final cid) = Done" obtain i where snap_p: "~ has_snapshotted (S t i) p" "has_snapshotted (S t (i+1)) p" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have "i < length t" proof - have "S t i \ S t (i+1)" using snap_p by auto then show ?thesis by (meson assms(1) computation.no_change_if_ge_length_t computation_axioms le_add1 not_less) qed let ?t = "take (length t - (i+1)) (drop (i+1) t)" - have tr: "trace (S t (i+1)) ?t (S t (length t))" - by (meson \i < length t\ assms(1) discrete exists_trace_for_any_i_j) + have tr: "trace (S t (i+1)) ?t (S t (length t))" + using \i < length t\ assms(1) exists_trace_for_any_i_j less_eq_Suc_le by fastforce have "Marker \ set (msgs (S t (i+1)) cid)" proof - have n_done: "snd (cs (S t (i+1)) cid) \ Done" proof (rule ccontr) assume "~ snd (cs (S t (i+1)) cid) \ Done" then have "snd (cs final cid) = Done" by (metis Suc_eq_plus1 Suc_leI \i < length t\ assms final_is_s_t_len_t computation.cs_done_implies_same_snapshots computation_axioms order_refl) then show False using cs_not_done by simp qed then show ?thesis using snapshotted_and_not_done_implies_marker_in_channel snap_p assms proof - have "i+1 \ length t" using \i < length t\ by auto then show ?thesis using snapshotted_and_not_done_implies_marker_in_channel snap_p assms n_done by simp qed qed moreover have "Marker \ set (msgs (S t (length t)) cid)" using final_is_s_t_len_t no_marker_left_in_final_state assms by blast ultimately have rm_prov: "\ev \ set ?t. (\q p. ev = RecvMarker cid q p)" using tr message_must_be_delivered_2_trace assms by (simp add: marker_must_be_delivered_2_trace) then obtain k where "\q p. t ! k = RecvMarker cid q p" "i+1 \ k" "k < length t" by (metis assms(1) exists_index) then have step: "(S t k) \ (t ! k) \ (S t (k+1))" by (metis Suc_eq_plus1_left add.commute assms(1) step_Suc) then have RecvMarker: "t ! k = RecvMarker cid q p" by (metis RecvMarker_given_channel \\q p. t ! k = RecvMarker cid q p\ assms(2) event.disc(25) event.sel(10) happen_implies_can_occur) then have "snd (cs (S t (k+1)) cid) = Done" using step \k < length t\ assms(1) recv_marker_means_cs_Done by blast then have "snd (cs final cid) = Done" using \Marker \ set (msgs (S t (length t)) cid)\ all_processes_snapshotted_in_final_state assms(1) assms(2) final_is_s_t_len_t snapshotted_and_not_done_implies_marker_in_channel by fastforce then show False using cs_not_done by simp qed lemma cs_NotStarted_implies_empty_cs: shows "\ trace init t final; channel cid = Some (p, q); i < length t; ~ has_snapshotted (S t i) q \ \ cs (S t i) cid = ([], NotStarted)" by (simp add: no_initial_channel_snapshot no_recording_cs_if_not_snapshotted) lemma fst_changed_by_recv_recording_trace: assumes "i < j" and "j \ length t" and "trace init t final" and "fst (cs (S t i) cid) \ fst (cs (S t j) cid)" and "channel cid = Some (p, q)" shows "\k. i \ k \ k < j \ (\p q u u' m. t ! k = Recv cid q p u u' m) \ (snd (cs (S t k) cid) = Recording)" (is ?P) proof (rule ccontr) assume "~ ?P" have "\ i < j; j \ length t; ~ ?P; trace init t final; channel cid = Some (p, q) \ \ fst (cs (S t i) cid) = fst (cs (S t j) cid)" proof (induct "j - i" arbitrary: i) case 0 then show ?case by linarith next case (Suc n) then have step: "(S t i) \ t ! i \ (S t (Suc i))" using step_Suc by auto then have "fst (cs (S t (Suc i)) cid) = fst (cs (S t i) cid)" by (metis Suc.prems(1) Suc.prems(3) assms(5) fst_cs_changed_by_recv_recording le_eq_less_or_eq) also have "fst (cs (S t (Suc i)) cid) = fst (cs (S t j) cid)" proof - have "j - Suc i = n" using Suc by simp moreover have "~ (\k. (Suc i) \ k \ k < j \ (\p q u u' m. t ! k = Recv cid q p u u' m) \ (snd (cs (S t k) cid) = Recording))" using \~ ?P\ Suc.prems(3) Suc_leD by blast ultimately show ?thesis using Suc by (metis Suc_lessI) qed finally show ?case by simp qed then show False using assms \~ ?P\ by blast qed lemma cs_not_nil_implies_postrecording_event: assumes "trace init t final" and "fst (cs (S t i) cid) \ []" and "i \ length t" and "channel cid = Some (p, q)" shows "\j. j < i \ postrecording_event t j" proof - have "fst (cs init cid) = []" using no_initial_channel_snapshot by auto then have diff_cs: "fst (cs (S t 0) cid) \ fst (cs (S t i) cid)" using assms(1) assms(2) init_is_s_t_0 by auto moreover have "0 < i" proof (rule ccontr) assume "~ 0 < i" then have "0 = i" by auto then have "fst (cs (S t 0) cid) = fst (cs (S t i) cid)" by blast then show False using diff_cs by simp qed ultimately obtain j where "j < i" and Recv: "\p q u u' m. t ! j = Recv cid q p u u' m" "snd (cs (S t j) cid) = Recording" using assms(1) assms(3) assms(4) fst_changed_by_recv_recording_trace by blast then have "has_snapshotted (S t j) q" using assms(1) assms(4) cs_recording_implies_snapshot by blast moreover have "regular_event (t ! j)" using Recv by auto moreover have "occurs_on (t ! j) = q" proof - have "can_occur (t ! j) (S t j)" by (meson Suc_le_eq \j < i\ assms(1) assms(3) happen_implies_can_occur le_trans step_Suc) then show ?thesis using Recv Recv_given_channel assms(4) by force qed ultimately have "postrecording_event t j" unfolding postrecording_event using \j < i\ assms(3) by simp then show ?thesis using \j < i\ by auto qed subsubsection \Relating process states\ lemma snapshot_state_must_have_been_reached: assumes "trace init t final" and "ps final p = Some u" and "~ has_snapshotted (S t i) p" and "has_snapshotted (S t (i+1)) p" and "i < length t" shows "states (S t i) p = u" proof (rule ccontr) assume "states (S t i) p \ u" then have "ps (S t (i+1)) p \ Some u" using assms(1) assms(3) snapshot_state by force then have "ps final p \ Some u" by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right assms(1) assms(3) assms(4) assms(5) final_is_s_t_len_t order_refl snapshot_state snapshot_state_unchanged_trace_2) then show False using assms by simp qed lemma ps_after_all_prerecording_events: assumes "trace init t final" and "\i'. i' \ i \ ~ prerecording_event t i'" and "\j'. j' < i \ ~ postrecording_event t j'" shows "ps_equal_to_snapshot (S t i) final" proof (unfold ps_equal_to_snapshot_def, rule allI) fix p show "Some (states (S t i) p) = ps final p" proof (rule ccontr) obtain s where "ps final p = Some s \ ps final p = None" by auto moreover assume "Some (states (S t i) p) \ ps final p" ultimately have "ps final p = None \ states (S t i) p \ s" by auto then show False proof (elim disjE) assume "ps final p = None" then show False using assms all_processes_snapshotted_in_final_state by blast next assume st: "states (S t i) p \ s" then obtain j where "~ has_snapshotted (S t j) p \ has_snapshotted (S t (j+1)) p" using Suc_eq_plus1 assms(1) exists_snapshot_for_all_p by presburger then show False proof (cases "has_snapshotted (S t i) p") case False then have "j \ i" by (metis Suc_eq_plus1 \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ assms(1) not_less_eq_eq snapshot_stable_ver_3) let ?t = "take (j-i) (drop i t)" have "\ev. ev \ set ?t \ regular_event ev \ occurs_on ev = p" proof (rule ccontr) assume "~ (\ev. ev \ set ?t \ regular_event ev \ occurs_on ev = p)" moreover have "trace (S t i) ?t (S t j)" using \i \ j\ assms(1) exists_trace_for_any_i_j by blast ultimately have "states (S t j) p = states (S t i) p" using no_state_change_if_only_nonregular_events st by blast then show False by (metis \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ \ps final p = Some s \ ps final p = None\ assms(1) final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation.snapshot_stable_ver_3 computation_axioms linorder_not_le snapshot_state_must_have_been_reached st) qed then obtain ev where "ev \ set ?t \ regular_event ev \ occurs_on ev = p" by blast then obtain k where t_ind: "0 \ k \ k < length ?t \ ev = ?t ! k" by (metis in_set_conv_nth not_le not_less_zero) moreover have "length ?t \ j - i" by simp ultimately have "?t ! k = (drop i t) ! k" using less_le_trans nth_take by blast also have "... = t ! (k+i)" by (metis \ev \ set (take (j - i) (drop i t)) \ regular_event ev \ occurs_on ev = p\ add.commute drop_eq_Nil length_greater_0_conv length_pos_if_in_set nat_le_linear nth_drop take_eq_Nil) finally have "?t ! k = t ! (k+i)" by simp have "prerecording_event t (k+i)" proof - have "regular_event (?t ! k)" using \ev \ set (take (j - i) (drop i t)) \ regular_event ev \ occurs_on ev = p\ t_ind by blast moreover have "occurs_on (?t ! k) = p" using \ev \ set (take (j - i) (drop i t)) \ regular_event ev \ occurs_on ev = p\ t_ind by blast moreover have "~ has_snapshotted (S t (k+i)) p" proof - have "k+i \ j" using \length (take (j - i) (drop i t)) \ j - i\ t_ind by linarith show ?thesis using \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ \k+i \ j\ assms(1) snapshot_stable_ver_3 by blast qed ultimately show ?thesis using \take (j - i) (drop i t) ! k = t ! (k + i)\ prerecording_event t_ind by auto qed then show False using assms by auto next case True have "j < i" proof (rule ccontr) assume "~ j < i" then have "j \ i" by simp moreover have "~ has_snapshotted (S t j) p" using \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ by blast moreover have "trace (S t i) (take (j - i) (drop i t)) (S t j)" using assms(1) calculation(1) exists_trace_for_any_i_j by blast ultimately have "~ has_snapshotted (S t i) p" using snapshot_stable by blast then show False using True by simp qed let ?t = "take (i-j) (drop j t)" have "\ev. ev \ set ?t \ regular_event ev \ occurs_on ev = p" proof (rule ccontr) assume "~ (\ev. ev \ set ?t \ regular_event ev \ occurs_on ev = p)" moreover have "trace (S t j) ?t (S t i)" using \j < i\ assms(1) exists_trace_for_any_i_j less_imp_le by blast ultimately have "states (S t j) p = states (S t i) p" using no_state_change_if_only_nonregular_events by auto moreover have "states (S t j) p = s" by (metis \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ \ps final p = Some s \ ps final p = None\ assms(1) final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation.snapshot_stable_ver_3 computation_axioms linorder_not_le snapshot_state_must_have_been_reached) ultimately show False using \states (S t i) p \ s\ by simp qed then obtain ev where ev: "ev \ set ?t \ regular_event ev \ occurs_on ev = p" by blast then obtain k where t_ind: "0 \ k \ k < length ?t \ ev = ?t ! k" by (metis in_set_conv_nth le0) have "length ?t \ i - j" by simp have "?t ! k = (drop j t) ! k" using t_ind by auto also have "... = t ! (k + j)" by (metis \ev \ set (take (i - j) (drop j t)) \ regular_event ev \ occurs_on ev = p\ add.commute drop_eq_Nil length_greater_0_conv length_pos_if_in_set nat_le_linear nth_drop take_eq_Nil) finally have "?t ! k = t ! (k+j)" by simp have "postrecording_event t (k+j)" proof - have "trace (S t j) (take k (drop j t)) (S t (k+j))" by (metis add_diff_cancel_right' assms(1) exists_trace_for_any_i_j le_add_same_cancel2 t_ind) then have "has_snapshotted (S t (k+j)) p" by (metis Suc_eq_plus1 Suc_leI \\ ps (S t j) p \ None \ ps (S t (j + 1)) p \ None\ \take (i - j) (drop j t) ! k = t ! (k + j)\ assms(1) drop_eq_Nil ev computation.snapshot_stable_ver_3 computation_axioms le_add_same_cancel2 length_greater_0_conv length_pos_if_in_set linorder_not_le order_le_less regular_event_preserves_process_snapshots step_Suc t_ind take_eq_Nil) then show ?thesis using \take (i - j) (drop j t) ! k = t ! (k + j)\ ev postrecording_event t_ind by auto qed moreover have "k + j < i" using \length (take (i - j) (drop j t)) \ i - j\ t_ind by linarith ultimately show False using assms(3) by simp qed qed qed qed subsubsection \Relating channel states\ lemma cs_when_recording: shows "\ i < j; j \ length t; trace init t final; has_snapshotted (S t i) p; snd (cs (S t i) cid) = Recording; snd (cs (S t j) cid) = Done; channel cid = Some (p, q) \ \ map Msg (fst (cs (S t j) cid)) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid)" proof (induct "j - (i+1)" arbitrary: i) case 0 then have "j = i+1" by simp then have step: "(S t i) \ (t ! i) \ (S t j)" using "0.prems" step_Suc by simp then have rm: "\q p. t ! i = RecvMarker cid q p" using done_only_from_recv_marker "0.prems" by force then have RecvMarker: "t ! i = RecvMarker cid q p" by (metis "0.prems"(7) RecvMarker_given_channel event.collapse(5) event.disc(25) event.inject(5) happen_implies_can_occur local.step) then have "takeWhile ((\) Marker) (msgs (S t i) cid) = []" proof - have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by simp then show ?thesis proof - have "\p ms. takeWhile p ms = [] \ p (hd ms::'c message)" by (metis (no_types) hd_append2 hd_in_set set_takeWhileD takeWhile_dropWhile_id) then show ?thesis using \can_occur (t ! i) (S t i)\ can_occur_def rm by fastforce qed qed then show ?case using local.step rm by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis Suc_eq_plus1 less_SucI nat_induct_at_least step_Suc) have ib: "i+1 < j \ j \ length t \ has_snapshotted (S t (i+1)) p \ snd (cs (S t j) cid) = Done" using Suc.hyps(2) Suc.prems(2) Suc.prems(4) Suc.prems(6) local.step snapshot_state_unchanged by auto have snap_q: "has_snapshotted (S t i) q" using Suc(7) Suc.prems(3) Suc cs_recording_implies_snapshot by blast then show ?case proof (cases "t ! i") case (Snapshot r) then have "r \ p" using Suc.prems(4) can_occur_def local.step by auto then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using Snapshot local.step Suc.prems(7) by auto moreover have "cs (S t (i+1)) cid = cs (S t i) cid" proof - have "r \ q" using Snapshot can_occur_def snap_q step by auto then show ?thesis using Snapshot local.step Suc.prems(7) by auto qed ultimately show ?thesis using Suc ib by force next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "takeWhile ((\) Marker) (msgs (S t i) cid) = []" proof - have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by simp then show ?thesis proof - have "\p ms. takeWhile p ms = [] \ p (hd ms::'c message)" by (metis (no_types) hd_append2 hd_in_set set_takeWhileD takeWhile_dropWhile_id) then show ?thesis using RecvMarker True \can_occur (t ! i) (S t i)\ can_occur_def by fastforce qed qed moreover have "snd (cs (S t (i+1)) cid) = Done" using RecvMarker Suc.prems(1) Suc.prems(2) Suc.prems(3) True recv_marker_means_cs_Done by auto moreover have "fst (cs (S t i) cid) = fst (cs (S t (i+1)) cid)" using RecvMarker True local.step by auto ultimately show ?thesis by (metis Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(7) Suc_eq_plus1 Suc_leI append_Nil2 cs_done_implies_same_snapshots) next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step Suc False by auto next case False with RecvMarker step Suc \cid \ cid'\ show ?thesis by (cases "s = p", auto) qed moreover have "cs (S t i) cid = cs (S t (i+1)) cid" proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step Suc False by auto next case no_snap: False then show ?thesis proof (cases "r = q") case True then show ?thesis using snap_q no_snap \r = q\ by simp next case False then show ?thesis using RecvMarker step Suc no_snap False \cid \ cid'\ by simp qed qed ultimately show ?thesis using Suc ib by simp qed next case (Trans r u u') then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Trans by auto ultimately show ?thesis using Suc ib by simp next case (Send cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True have marker_in_msgs: "Marker \ set (msgs (S t i) cid)" proof - have "has_snapshotted (S t i) p" using Suc by simp moreover have "i < length t" using Suc.prems(1) Suc.prems(2) less_le_trans by blast moreover have "snd (cs (S t i) cid) \ Done" using Suc by simp ultimately show ?thesis using snapshotted_and_not_done_implies_marker_in_channel less_imp_le using Suc by blast qed then have "takeWhile ((\) Marker) (msgs (S t i) cid) = takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" proof - have "butlast (msgs (S t (i+1)) cid) = msgs (S t i) cid" using step True Send by auto then show ?thesis proof - have "takeWhile ((\) Marker) (msgs (S t i) cid @ [last (msgs (S t (i + 1)) cid)]) = takeWhile ((\) Marker) (msgs (S t i) cid)" using marker_in_msgs by force then show ?thesis by (metis (no_types) \butlast (msgs (S t (i + 1)) cid) = msgs (S t i) cid\ append_butlast_last_id in_set_butlastD length_greater_0_conv length_pos_if_in_set marker_in_msgs) qed qed moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto ultimately show ?thesis using ib Suc by simp next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send by auto moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto ultimately show ?thesis using Suc ib by simp qed next case (Recv cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have msgs_ip1: "Msg m # msgs (S t (i+1)) cid = msgs (S t i) cid" using Suc Recv step by auto moreover have cs_ip1: "cs (S t (i+1)) cid = (fst (cs (S t i) cid) @ [m], Recording)" using True Suc Recv step by auto ultimately show ?thesis proof - have "map Msg (fst (cs (S t j) cid)) = map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" using Suc ib cs_ip1 by force moreover have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid) = map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" proof - have "takeWhile ((\) Marker) (Msg m # msgs (S t (i+1)) cid) = Msg m # takeWhile ((\) Marker) (msgs (S t (i + 1)) cid)" by auto then have "takeWhile ((\) Marker) (msgs (S t i) cid) = Msg m # takeWhile ((\) Marker) (msgs (S t (i + 1)) cid)" by (metis msgs_ip1) then show ?thesis using cs_ip1 by auto qed ultimately show ?thesis by simp qed next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv by auto moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Recv False by auto ultimately show ?thesis using Suc ib by simp qed qed qed lemma cs_when_recording_2: shows "\ i \ j; trace init t final; ~ has_snapshotted (S t i) p; \k. i \ k \ k < j \ ~ occurs_on (t ! k) = p; snd (cs (S t i) cid) = Recording; channel cid = Some (p, q) \ \ map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid) \ snd (cs (S t j) cid) = Recording" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis Suc_eq_plus1 all_processes_snapshotted_in_final_state distributed_system.step_Suc distributed_system_axioms computation.final_is_s_t_len_t computation_axioms linorder_not_le snapshot_stable_ver_3) have ib: "i+1 \ j \ ~ has_snapshotted (S t (i+1)) p \ (\k. (i+1) \ k \ k < j \ ~ occurs_on (t ! k) = p) \ j - (i+1) = n" by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(4) diff_Suc_1 diff_diff_left Suc_eq_plus1 Suc_leD Suc_le_eq Suc_neq_Zero cancel_comm_monoid_add_class.diff_cancel le_neq_implies_less le_refl local.step no_state_change_if_no_event) have snap_q: "has_snapshotted (S t i) q" using Suc.prems(5,6) Suc.prems(2) cs_recording_implies_snapshot by blast then show ?case proof (cases "t ! i") case (Snapshot r) then have "r \ p" using Suc by auto then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using Snapshot local.step Suc.prems(6) by auto moreover have "cs (S t (i+1)) cid = cs (S t i) cid" proof - have "r \ q" using step can_occur_def Snapshot snap_q by auto then show ?thesis using Snapshot step Suc by simp qed ultimately show ?thesis using Suc ib by auto next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "Marker \ set (msgs (S t i) cid)" using RecvMarker RecvMarker_implies_Marker_in_set local.step by blast then have "has_snapshotted (S t i) p" using Suc.prems(2) no_marker_if_no_snapshot Suc by blast then show ?thesis using Suc.prems by simp next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step Suc False by auto next case False with RecvMarker step Suc \cid \ cid'\ show ?thesis by (cases "s = p", auto) qed moreover have "cs (S t i) cid = cs (S t (i+1)) cid" proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step Suc False by auto next case no_snap: False then show ?thesis proof (cases "r = q") case True then show ?thesis using snap_q no_snap \r = q\ by simp next case False then show ?thesis using RecvMarker step Suc no_snap False \cid \ cid'\ by simp qed qed ultimately show ?thesis using Suc ib by auto qed next case (Trans r u u') then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Trans by auto ultimately show ?thesis using Suc ib by auto next case (Send cid' r s u u' m) then have "r \ p" using Suc.hyps(2) Suc.prems(4) Suc by auto have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "channel cid = channel cid'" by auto then have "(p, q) = (r, s)" using can_occur_def step Send Suc by auto then show False using \r \ p\ by simp qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send by simp moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto ultimately show ?thesis using Suc ib by auto next case (Recv cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have msgs_ip1: "Msg m # msgs (S t (i+1)) cid = msgs (S t i) cid" using Suc Recv step by auto moreover have cs_ip1: "cs (S t (i+1)) cid = (fst (cs (S t i) cid) @ [m], Recording)" using True Suc Recv step by auto ultimately show ?thesis proof - have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (i+1)) cid) \ snd (cs (S t j) cid) = Recording" using Suc ib cs_ip1 by auto moreover have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid) = map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" proof - have "takeWhile ((\) Marker) (Msg m # msgs (S t (i + 1)) cid) = Msg m # takeWhile ((\) Marker) (msgs (S t (i + 1)) cid)" by fastforce then have "takeWhile ((\) Marker) (msgs (S t i) cid) = Msg m # takeWhile ((\) Marker) (msgs (S t (i + 1)) cid)" by (metis msgs_ip1) then show ?thesis using cs_ip1 by force qed ultimately show ?thesis using cs_ip1 by simp qed next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv by auto moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Recv False by auto ultimately show ?thesis using Suc ib by auto qed qed qed lemma cs_when_recording_3: shows "\ i \ j; trace init t final; ~ has_snapshotted (S t i) q; \k. i \ k \ k < j \ ~ occurs_on (t ! k) = q; snd (cs (S t i) cid) = NotStarted; has_snapshotted (S t i) p; Marker : set (msgs (S t i) cid); channel cid = Some (p, q) \ \ map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid) \ snd (cs (S t j) cid) = NotStarted" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis Suc_eq_plus1 all_processes_snapshotted_in_final_state distributed_system.step_Suc distributed_system_axioms computation.final_is_s_t_len_t computation_axioms linorder_not_le snapshot_stable_ver_3) have ib: "i+1 \ j \ ~ has_snapshotted (S t (i+1)) q \ has_snapshotted (S t (i+1)) p \ (\k. (i+1) \ k \ k < j \ ~ occurs_on (t ! k) = q) \ j - (i+1) = n \ Marker : set (msgs (S t (i+1)) cid) \ cs (S t i) cid = cs (S t (i+1)) cid" proof - have "i+1 \ j \ ~ has_snapshotted (S t (i+1)) q \ (\k. (i+1) \ k \ k < j \ ~ occurs_on (t ! k) = q) \ j - (i+1) = n" by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(4) diff_Suc_1 diff_diff_left Suc_eq_plus1 Suc_leD Suc_le_eq Suc_neq_Zero cancel_comm_monoid_add_class.diff_cancel le_neq_implies_less le_refl local.step no_state_change_if_no_event) moreover have "has_snapshotted (S t (i+1)) p" using Suc.prems(6) local.step snapshot_state_unchanged by auto moreover have "Marker : set (msgs (S t (i+1)) cid)" using Suc calculation(1) local.step recv_marker_means_snapshotted_2 by blast moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using Suc calculation(1) no_recording_cs_if_not_snapshotted by auto ultimately show ?thesis by simp qed then show ?case proof (cases "t ! i") case (Snapshot r) then have "r \ q" using Suc by auto then have "takeWhile ((\) Marker) (msgs (S t (i+1)) cid) = takeWhile ((\) Marker) (msgs (S t i) cid)" proof (cases "occurs_on (t ! i) = p") case True then show ?thesis by (metis (mono_tags, lifting) Snapshot Suc.prems(6) distributed_system.can_occur_def event.sel(4) event.simps(29) computation_axioms computation_def happen_implies_can_occur local.step) next case False then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using Snapshot local.step Suc by auto then show ?thesis by simp qed then show ?thesis using Suc ib by metis next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "snd (cs (S t i) cid) = Done" by (metis RecvMarker Suc.prems(2) Suc_eq_plus1 Suc_le_eq exactly_one_snapshot computation.no_change_if_ge_length_t computation.recv_marker_means_cs_Done computation.snapshot_stable_ver_2 computation_axioms ib nat_le_linear) then show ?thesis using Suc.prems by simp next case False then have "takeWhile ((\) Marker) (msgs (S t i) cid) = takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" proof (cases "has_snapshotted (S t i) r") case True with RecvMarker False step show ?thesis by auto next case no_snap: False then have "r \ p" using Suc.prems(6) by auto then show ?thesis using no_snap RecvMarker step Suc.prems False by auto qed then show ?thesis using Suc ib by metis qed next case (Trans r u u') then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto then show ?thesis using Suc ib by auto next case (Send cid' r s u u' m) then have "r \ q" using Suc.hyps(2) Suc.prems(4) by auto have marker: "Marker \ set (msgs (S t i) cid)" using Suc by simp with step Send marker have "takeWhile ((\) Marker) (msgs (S t i) cid) = takeWhile ((\) Marker) (msgs (S t (i+1)) cid)" by (cases "cid = cid'", auto) then show ?thesis using Suc ib by auto next case (Recv cid' r s u u' m) then have "cid' \ cid" by (metis Suc.hyps(2) Suc.prems(4) Suc.prems(8) distributed_system.can_occur_Recv distributed_system_axioms event.sel(3) happen_implies_can_occur local.step option.inject order_refl prod.inject zero_less_Suc zero_less_diff) then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv Suc by simp then show ?thesis using Suc ib by auto qed qed lemma at_most_one_marker: shows "\ trace init t final; channel cid = Some (p, q) \ \ Marker \ set (msgs (S t i) cid) \ (\!j. j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j = Marker)" proof (induct i) case 0 then show ?case using no_initial_Marker init_is_s_t_0 by auto next case (Suc i) then show ?case proof (cases "i < length t") case False then show ?thesis by (metis Suc.prems(1) final_is_s_t_len_t computation.no_change_if_ge_length_t computation_axioms le_refl less_imp_le_nat no_marker_left_in_final_state not_less_eq) next case True then have step: "(S t i) \ (t ! i) \ (S t (Suc i))" using step_Suc Suc.prems by blast moreover have "Marker \ set (msgs (S t i) cid) \ (\!j. j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j = Marker)" using Suc.hyps Suc.prems(1) Suc.prems(2) by linarith moreover have "Marker \ set (msgs (S t (Suc i)) cid) \ (\!j. j < length (msgs (S t (Suc i)) cid) \ msgs (S t (Suc i)) cid ! j = Marker)" proof (cases "Marker \ set (msgs (S t i) cid)") case no_marker: True then show ?thesis proof (cases "t ! i") case (Snapshot r) then show ?thesis proof (cases "r = p") case True then have new_msgs: "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Marker]" using step Snapshot Suc by auto then show ?thesis using util_exactly_one_element no_marker by fastforce next case False then show ?thesis using Snapshot local.step no_marker Suc by auto qed next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then show ?thesis using RecvMarker RecvMarker_implies_Marker_in_set local.step no_marker by blast next case False then show ?thesis proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step Suc False by simp next case no_snap: False then show ?thesis proof (cases "r = p") case True then have "msgs (S t (i+1)) cid = msgs (S t i) cid @ [Marker]" using RecvMarker step Suc.prems no_snap \cid \ cid'\ by simp then show ?thesis proof - assume a1: "msgs (S t (i + 1)) cid = msgs (S t i) cid @ [Marker]" { fix nn :: "nat \ nat" have "\ms m. \n. \na. ((m::'c message) \ set ms \ n < length (ms @ [m])) \ (m \ set ms \ (ms @ [m]) ! n = m) \ (\ na < length (ms @ [m]) \ (ms @ [m]) ! na \ m \ m \ set ms \ na = n)" by (metis (no_types) util_exactly_one_element) then have "\n. n < length (msgs (S t (Suc i)) cid) \ nn n = n \ msgs (S t (Suc i)) cid ! n = Marker \ n < length (msgs (S t (Suc i)) cid) \ msgs (S t (Suc i)) cid ! n = Marker \ \ nn n < length (msgs (S t (Suc i)) cid) \ n < length (msgs (S t (Suc i)) cid) \ msgs (S t (Suc i)) cid ! n = Marker \ msgs (S t (Suc i)) cid ! nn n \ Marker" using a1 by (metis Suc_eq_plus1 no_marker) } then show ?thesis by (metis (no_types)) qed next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using RecvMarker step Suc.prems \cid \ cid'\ no_snap by simp then show ?thesis using Suc by simp qed qed qed next case (Trans r u u') then show ?thesis using no_marker step by auto next case (Send cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have "Marker \ set (msgs (S t (Suc i)) cid)" using step no_marker Send by auto then show ?thesis by simp next case False then have "Marker \ set (msgs (S t (Suc i)) cid)" using step no_marker Send by auto then show ?thesis by simp qed next case (Recv cid' r s u u' m) with step no_marker Recv show ?thesis by (cases "cid = cid'", auto) qed next case False then have asm: "\!j. j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j = Marker" using Suc by simp have len_filter: "length (filter ((=) Marker) (msgs (S t i) cid)) = 1" by (metis False \Marker \ set (msgs (S t i) cid) \ (\!j. j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j = Marker)\ exists_one_iff_filter_one) have snap_p: "has_snapshotted (S t i) p" using False Suc.prems no_marker_if_no_snapshot by blast show ?thesis proof (cases "t ! i") case (Snapshot r) have "r \ p" proof (rule ccontr) assume "~ r \ p" moreover have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by blast ultimately show False using snap_p can_occur_def Snapshot by auto qed then have "msgs (S t (Suc i)) cid = msgs (S t i) cid" using step Snapshot Suc by auto then show ?thesis using asm by simp next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid" using RecvMarker step by auto then have "Marker \ set (msgs (S t (Suc i)) cid)" proof - have "\j. j \ 0 \ j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j \ Marker" by (metis False \Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\ asm length_pos_if_in_set nth_Cons_0) then show ?thesis proof - assume a1: "\j. j \ 0 \ j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j \ Marker" have "\ms n. ms \ msgs (S t i) cid \ length (msgs (S t (Suc i)) cid) \ n \ length ms = Suc n" by (metis \Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\ length_Suc_conv) then show ?thesis using a1 by (metis (no_types) Suc_mono Zero_not_Suc \Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\ in_set_conv_nth nth_Cons_Suc) qed qed then show ?thesis by simp next case cid_neq_cid': False then show ?thesis proof (cases "has_snapshotted (S t i) r") case True then have "msgs (S t (Suc i)) cid = msgs (S t i) cid" using cid_neq_cid' RecvMarker local.step snap_p by auto then show ?thesis using asm by simp next case False then have "r \ p" using snap_p by blast then have "msgs (S t (Suc i)) cid = msgs (S t i) cid" using cid_neq_cid' RecvMarker step False Suc by auto then show ?thesis using asm by simp qed qed next case (Trans r u u') then show ?thesis using step asm by auto next case (Send cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have new_messages: "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Msg m]" using step Send by auto then have "\!j. j < length (msgs (S t (Suc i)) cid) \ msgs (S t (Suc i)) cid ! j = Marker" proof - have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid)) = length (filter ((=) Marker) (msgs (S t i) cid)) + length (filter ((=) Marker) [Msg m])" by (simp add: new_messages) then have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid)) = 1" using len_filter by simp then show ?thesis using exists_one_iff_filter_one by metis qed then show ?thesis by simp next case False then show ?thesis using step Send asm by auto qed next case (Recv cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have new_msgs: "Msg m # msgs (S t (Suc i)) cid = msgs (S t i) cid" using step Recv by auto then show ?thesis proof - have "length (filter ((=) Marker) (msgs (S t i) cid)) = length (filter ((=) Marker) [Msg m]) + length (filter ((=) Marker) (msgs (S t (Suc i)) cid))" by (metis append_Cons append_Nil filter_append len_filter length_append new_msgs) then have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid)) = 1" using len_filter by simp then show ?thesis using exists_one_iff_filter_one by metis qed next case False then show ?thesis using step Recv asm by auto qed qed qed then show ?thesis by simp qed qed lemma last_changes_implies_send_when_msgs_nonempty: assumes "trace init t final" and "msgs (S t i) cid \ []" and "msgs (S t (i+1)) cid \ []" and "last (msgs (S t i) cid) = Marker" and "last (msgs (S t (i+1)) cid) \ Marker" and "channel cid = Some (p, q)" shows "(\u u' m. t ! i = Send cid p q u u' m)" proof - have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis Suc_eq_plus1_left add.commute assms(1) assms(4) assms(5) le_Suc_eq nat_le_linear nat_less_le no_change_if_ge_length_t step_Suc) then show ?thesis proof (cases "t ! i") case (Snapshot r) then show ?thesis by (metis assms(4) assms(5) last_snoc local.step next_snapshot) next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)" proof - have "Marker # msgs (S t (i + 1)) cid = msgs (S t i) cid" using RecvMarker local.step True by auto then show ?thesis by (metis assms(3) last_ConsR) qed then show ?thesis using assms by simp next case no_snap: False then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)" proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using RecvMarker step no_snap by simp next case False with RecvMarker step no_snap \cid \ cid'\ assms show ?thesis by (cases "p = r", auto) qed then show ?thesis using assms by simp qed next case (Trans r u u') then show ?thesis using assms(4) assms(5) local.step by auto next case (Send cid' r s u u' m) then have "cid = cid'" by (metis (no_types, opaque_lifting) assms(4) assms(5) local.step next_send) moreover have "(p, q) = (r, s)" proof - have "channel cid = channel cid'" using \cid = cid'\ by simp moreover have "channel cid = Some (p, q)" using assms by simp moreover have "channel cid' = Some (r, s)" using Send step can_occur_def by auto ultimately show ?thesis by simp qed ultimately show ?thesis using Send by auto next case (Recv cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)" by (metis (no_types, lifting) Recv assms(3) assms(4) last_ConsR local.step next_recv) then show ?thesis using assms by simp next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto then show ?thesis using assms by simp qed qed qed lemma no_marker_after_RecvMarker: assumes "trace init t final" and "(S t i) \ RecvMarker cid p q \ (S t (i+1))" and "channel cid = Some (q, p)" shows "Marker \ set (msgs (S t (i+1)) cid)" proof - have new_msgs: "msgs (S t i) cid = Marker # msgs (S t (i+1)) cid" using assms(2) by auto have one_marker: "\!j. j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j = Marker" by (metis assms(1,3) at_most_one_marker list.set_intros(1) new_msgs) then obtain j where "j < length (msgs (S t i) cid)" "msgs (S t i) cid ! j = Marker" by blast then have "j = 0" using one_marker new_msgs by auto then have "\j. j \ 0 \ j < length (msgs (S t i) cid) \ msgs (S t i) cid ! j \ Marker" using one_marker using \j < length (msgs (S t i) cid)\ \msgs (S t i) cid ! j = Marker\ by blast then have "\j. j < length (msgs (S t (i+1)) cid) \ msgs (S t (i+1)) cid ! j \ Marker" by (metis One_nat_def Suc_eq_plus1 Suc_le_eq Suc_mono le_zero_eq list.size(4) new_msgs not_less0 nth_Cons_Suc) then show ?thesis by (simp add: in_set_conv_nth) qed lemma no_marker_and_snapshotted_implies_no_more_markers_trace: shows "\ trace init t final; i \ j; j \ length t; has_snapshotted (S t i) p; Marker \ set (msgs (S t i) cid); channel cid = Some (p, q) \ \ Marker \ set (msgs (S t j) cid)" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis (no_types, opaque_lifting) Suc_eq_plus1 cancel_comm_monoid_add_class.diff_cancel distributed_system.step_Suc distributed_system_axioms less_le_trans linorder_not_less old.nat.distinct(2) order_eq_iff) then have "Marker \ set (msgs (S t (i+1)) cid)" using no_marker_and_snapshotted_implies_no_more_markers Suc step by blast moreover have "has_snapshotted (S t (i+1)) p" using Suc.prems(4) local.step snapshot_state_unchanged by auto ultimately show ?case proof - assume a1: "ps (S t (i + 1)) p \ None" assume a2: "Marker \ set (msgs (S t (i + 1)) cid)" have f3: "j \ i" using Suc.hyps(2) by force have "j - Suc i = n" by (metis (no_types) Suc.hyps(2) Suc.prems(2) add.commute add_Suc_right add_diff_cancel_left' le_add_diff_inverse) then show ?thesis using f3 a2 a1 by (metis Suc.hyps(1) Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(6) Suc_eq_plus1_left add.commute less_Suc_eq linorder_not_less) qed qed lemma marker_not_vanishing_means_always_present: shows "\ trace init t final; i \ j; j \ length t; Marker : set (msgs (S t i) cid); Marker : set (msgs (S t j) cid); channel cid = Some (p, q) \ \ \k. i \ k \ k \ j \ Marker : set (msgs (S t k) cid)" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis (no_types, lifting) Suc_eq_plus1 add_lessD1 distributed_system.step_Suc distributed_system_axioms le_add_diff_inverse order_le_less zero_less_Suc zero_less_diff) have "Marker : set (msgs (S t (i+1)) cid)" proof (rule ccontr) assume asm: "~ Marker : set (msgs (S t (i+1)) cid)" have snap_p: "has_snapshotted (S t i) p" using Suc.prems(1) Suc.prems(4,6) no_marker_if_no_snapshot by blast then have "has_snapshotted (S t (i+1)) p" using local.step snapshot_state_unchanged by auto then have "Marker \ set (msgs (S t j) cid)" - by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(6) asm discrete no_marker_and_snapshotted_implies_no_more_markers_trace zero_less_Suc zero_less_diff) + using Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(6) asm + no_marker_and_snapshotted_implies_no_more_markers_trace [of t \Suc i\ j p cid q] + by simp then show False using Suc.prems by simp qed then show ?case by (meson Suc.prems(1) Suc.prems(3) Suc.prems(4) Suc.prems(5) Suc.prems(6) computation.snapshot_stable_ver_3 computation_axioms no_marker_and_snapshotted_implies_no_more_markers_trace no_marker_if_no_snapshot) qed lemma last_stays_if_no_recv_marker_and_no_send: shows "\ trace init t final; i < j; j \ length t; last (msgs (S t i) cid) = Marker; Marker : set (msgs (S t i) cid); Marker : set (msgs (S t j) cid); \k. i \ k \ k < j \ ~ (\u u' m. t ! k = Send cid p q u u' m); channel cid = Some (p, q) \ \ last (msgs (S t j) cid) = Marker" proof (induct "j - (i+1)" arbitrary: i) case 0 then have "j = i+1" by simp then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis "0"(2) "0.prems"(2) "0.prems"(3) Suc_eq_plus1 distributed_system.step_Suc distributed_system_axioms less_le_trans) have "Marker = last (msgs (S t (i+1)) cid)" proof (rule ccontr) assume "~ Marker = last (msgs (S t (i+1)) cid)" then have "\u u' m. t ! i = Send cid p q u u' m" proof - have "msgs (S t (i+1)) cid \ []" using "0" \j = i+1\ by auto moreover have "msgs (S t i) cid \ []" using "0" by auto ultimately show ?thesis using "0.prems"(1) "0.prems"(4) "0.prems"(8) \Marker \ last (msgs (S t (i + 1)) cid)\ last_changes_implies_send_when_msgs_nonempty by auto qed then show False using 0 by auto qed then show ?case using \j = i+1\ by simp next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis (no_types, opaque_lifting) Suc_eq_plus1 distributed_system.step_Suc distributed_system_axioms less_le_trans) have marker_present: "Marker : set (msgs (S t (i+1)) cid)" - by (meson Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(5) Suc.prems(6) Suc.prems(8) discrete le_add1 less_imp_le_nat marker_not_vanishing_means_always_present) + using Suc.prems + marker_not_vanishing_means_always_present [of t i j cid p q] by simp moreover have "Marker = last (msgs (S t (i+1)) cid)" proof (rule ccontr) assume asm: "~ Marker = last (msgs (S t (i+1)) cid)" then have "\u u' m. t ! i = Send cid p q u u' m" proof - have "msgs (S t (i+1)) cid \ []" using marker_present by auto moreover have "msgs (S t i) cid \ []" using Suc by auto ultimately show ?thesis using Suc.prems(1) Suc.prems(4) Suc.prems(8) asm last_changes_implies_send_when_msgs_nonempty by auto qed then show False using Suc by auto qed moreover have "\k. i+1 \ k \ k < j \ ~ (\u u' m. t ! k = Send cid p q u u' m)" using Suc.prems by force moreover have "i+1 < j" using Suc by auto moreover have "j \ length t" using Suc by auto moreover have "trace init t final" using Suc by auto moreover have "Marker : set (msgs (S t j) cid)" using Suc by auto ultimately show ?case using Suc by (metis diff_Suc_1 diff_diff_left) qed lemma last_changes_implies_send_when_msgs_nonempty_trace: assumes "trace init t final" "i < j" "j \ length t" "Marker : set (msgs (S t i) cid)" "Marker : set (msgs (S t j) cid)" "last (msgs (S t i) cid) = Marker" "last (msgs (S t j) cid) \ Marker" "channel cid = Some (p, q)" shows "\k u u' m. i \ k \ k < j \ t ! k = Send cid p q u u' m" proof (rule ccontr) assume "~ (\k u u' m. i \ k \ k < j \ t ! k = Send cid p q u u' m)" then have "\k. i \ k \ k < j \ ~ (\u u' m. t ! k = Send cid p q u u' m)" by blast then have "last (msgs (S t j) cid) = Marker" using assms last_stays_if_no_recv_marker_and_no_send by blast then show False using assms by simp qed lemma msg_after_marker_and_nonempty_implies_postrecording_event: assumes "trace init t final" and "Marker : set (msgs (S t i) cid)" and "Marker \ last (msgs (S t i) cid)" and "i \ length t" and "channel cid = Some (p, q)" shows "\j. j < i \ postrecording_event t j" (is ?P) proof - let ?len = "length (msgs (S t i) cid)" have "?len \ 0" using assms(2) by auto have snap_p_i: "has_snapshotted (S t i) p" using assms no_marker_if_no_snapshot by blast obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have "j < i" by (meson assms(1) computation.snapshot_stable_ver_2 computation_axioms not_less snap_p(1) snap_p_i) have step_snap: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis Suc_eq_plus1 assms(1) l2 nat_le_linear nat_less_le snap_p(1) snapshot_stable_ver_2 step_Suc) have re: "~ regular_event (t ! j)" by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_snap) have op: "occurs_on (t ! j) = p" using no_state_change_if_no_event snap_p(1) snap_p(2) step_snap by force have marker_last: "Marker = last (msgs (S t (j+1)) cid) \ msgs (S t (j+1)) cid \ []" proof - have "isSnapshot (t ! j) \ isRecvMarker (t ! j)" using re nonregular_event by auto then show ?thesis proof (elim disjE, goal_cases) case 1 then have "t ! j = Snapshot p" using op by auto then show ?thesis using step_snap assms by auto next case 2 then obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' p r" by (metis event.collapse(5) op) then have "cid \ cid'" using RecvMarker_implies_Marker_in_set assms(1) assms(5) no_marker_if_no_snapshot snap_p(1) step_snap by blast then show ?thesis using assms snap_p(1) step_snap RecvMarker by auto qed qed then have "\k u u' m. j+1 \ k \ k < i \ t ! k = Send cid p q u u' m" proof - have "j+1 < i" proof - have "(S t (j+1)) \ (S t i)" using assms(3) marker_last by auto then have "j+1 \ i" by auto moreover have "j+1 \ i" using \j < i\ by simp ultimately show ?thesis by simp qed moreover have "trace init t final" using assms by simp moreover have "Marker = last (msgs (S t (j+1)) cid)" using marker_last by simp moreover have "Marker : set (msgs (S t (j+1)) cid)" using marker_last by (simp add: marker_last) ultimately show ?thesis using assms last_changes_implies_send_when_msgs_nonempty_trace by simp qed then obtain k where Send: "\u u' m. j+1 \ k \ k < i \ t ! k = Send cid p q u u' m" by blast then have "postrecording_event t k" proof - have "k < length t" using Send assms by simp moreover have "isSend (t ! k)" using Send by auto moreover have "has_snapshotted (S t k) p" using Send snap_p using assms(1) snapshot_stable_ver_3 by blast moreover have "occurs_on (t ! k) = p" using Send by auto ultimately show ?thesis unfolding postrecording_event by simp qed then show ?thesis using Send by auto qed lemma same_messages_if_no_occurrence_trace: shows "\ trace init t final; i \ j; j \ length t; (\k. i \ k \ k < j \ occurs_on (t ! k) \ p \ occurs_on (t ! k) \ q); channel cid = Some (p, q) \ \ msgs (S t i) cid = msgs (S t j) cid \ cs (S t i) cid = cs (S t j) cid" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" by (metis (no_types, opaque_lifting) Suc_eq_plus1 Suc_n_not_le_n diff_self_eq_0 distributed_system.step_Suc distributed_system_axioms le0 le_eq_less_or_eq less_le_trans) then have "msgs (S t i) cid = msgs (S t (i+1)) cid \ cs (S t i) cid = cs (S t (i+1)) cid" proof - have "~ occurs_on (t ! i) = p" using Suc by simp moreover have "~ occurs_on (t ! i) = q" using Suc by simp ultimately show ?thesis using step Suc same_messages_if_no_occurrence by blast qed moreover have "msgs (S t (i+1)) cid = msgs (S t j) cid \ cs (S t (i+1)) cid = cs (S t j) cid" proof - have "i+1 \ j" using Suc by linarith moreover have "\k. i+1 \ k \ k < j \ occurs_on (t ! k) \ p \ occurs_on (t ! k) \ q" using Suc by force ultimately show ?thesis using Suc by auto qed ultimately show ?case by simp qed lemma snapshot_step_cs_preservation_p: assumes "c \ ev \ c'" and "~ regular_event ev" and "occurs_on ev = p" and "channel cid = Some (p, q)" shows "map Msg (fst (cs c cid)) @ takeWhile ((\) Marker) (msgs c cid) = map Msg (fst (cs c' cid)) @ takeWhile ((\) Marker) (msgs c' cid) \ snd (cs c cid) = snd (cs c' cid)" proof - have "isSnapshot ev \ isRecvMarker ev" using assms nonregular_event by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then have Snap: "ev = Snapshot p" by (metis event.collapse(4) assms(3)) then have "fst (cs c cid) = fst (cs c' cid)" using assms(1) assms(2) regular_event same_cs_if_not_recv by blast moreover have "takeWhile ((\) Marker) (msgs c cid) = takeWhile ((\) Marker) (msgs c' cid)" proof - have "msgs c' cid = msgs c cid @ [Marker]" using assms Snap by auto then show ?thesis by (simp add: takeWhile_tail) qed moreover have "snd (cs c cid) = snd (cs c' cid)" using Snap assms no_self_channel by fastforce ultimately show ?thesis by simp next case 2 then obtain cid' r where RecvMarker: "ev = RecvMarker cid' p r" by (metis event.collapse(5) assms(3)) have "cid \ cid'" by (metis "2" RecvMarker assms(1) assms(4) distributed_system.RecvMarker_given_channel distributed_system.happen_implies_can_occur distributed_system_axioms event.sel(5,10) no_self_channel) then have "fst (cs c cid) = fst (cs c' cid)" using RecvMarker assms(1) assms(2) regular_event same_cs_if_not_recv by blast moreover have "takeWhile ((\) Marker) (msgs c cid) = takeWhile ((\) Marker) (msgs c' cid)" proof (cases "has_snapshotted c p") case True then have "msgs c cid = msgs c' cid" using RecvMarker \cid \ cid'\ assms by auto then show ?thesis by simp next case False then have "msgs c' cid = msgs c cid @ [Marker]" using RecvMarker \cid \ cid'\ assms by auto then show ?thesis by (simp add: takeWhile_tail) qed moreover have "snd (cs c cid) = snd (cs c' cid)" proof (cases "has_snapshotted c p") case True then have "cs c cid = cs c' cid" using RecvMarker \cid \ cid'\ assms by simp then show ?thesis by simp next case False then show ?thesis using RecvMarker \cid \ cid'\ assms(1) assms(4) no_self_channel by auto qed ultimately show ?thesis by simp qed qed lemma snapshot_step_cs_preservation_q: assumes "c \ ev \ c'" and "~ regular_event ev" and "occurs_on ev = q" and "channel cid = Some (p, q)" and "Marker \ set (msgs c cid)" and "~ has_snapshotted c q" shows "map Msg (fst (cs c cid)) @ takeWhile ((\) Marker) (msgs c cid) = map Msg (fst (cs c' cid)) @ takeWhile ((\) Marker) (msgs c' cid) \ snd (cs c' cid) = Recording" proof - have "isSnapshot ev \ isRecvMarker ev" using assms nonregular_event by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then have Snapshot: "ev = Snapshot q" by (metis event.collapse(4) assms(3)) then have "fst (cs c cid) = fst (cs c' cid)" using assms(1) assms(2) regular_event same_cs_if_not_recv by blast moreover have "takeWhile ((\) Marker) (msgs c cid) = takeWhile ((\) Marker) (msgs c' cid)" proof - have "msgs c' cid = msgs c cid" using assms Snapshot by (metis distributed_system.next_snapshot distributed_system_axioms eq_fst_iff no_self_channel option.inject) then show ?thesis by simp qed moreover have "snd (cs c' cid) = Recording" using assms Snapshot by auto ultimately show ?thesis by simp next case 2 then obtain cid' r where RecvMarker: "ev = RecvMarker cid' q r" by (metis event.collapse(5) assms(3)) have "cid \ cid'" using RecvMarker RecvMarker_implies_Marker_in_set assms(1) assms(5) by blast have "fst (cs c cid) = fst (cs c' cid)" using assms(1) assms(2) regular_event same_cs_if_not_recv by blast moreover have "takeWhile ((\) Marker) (msgs c cid) = takeWhile ((\) Marker) (msgs c' cid)" proof - have "\r. channel cid = Some (q, r)" using assms(4) no_self_channel by auto with RecvMarker assms \cid \ cid'\ have "msgs c cid = msgs c' cid" by (cases "has_snapshotted c r", auto) then show ?thesis by simp qed moreover have "snd (cs c' cid) = Recording" using assms RecvMarker \cid \ cid'\ by simp ultimately show ?thesis by simp qed qed lemma Marker_in_channel_implies_not_done: assumes "trace init t final" and "Marker : set (msgs (S t i) cid)" and "channel cid = Some (p, q)" and "i \ length t" shows "snd (cs (S t i) cid) \ Done" proof (rule ccontr) assume is_done: "~ snd (cs (S t i) cid) \ Done" let ?t = "take i t" have tr: "trace init ?t (S t i)" using assms(1) exists_trace_for_any_i by blast have "\q p. RecvMarker cid q p \ set ?t" by (metis (mono_tags, lifting) assms(3) distributed_system.trace.simps distributed_system_axioms done_only_from_recv_marker_trace computation.no_initial_channel_snapshot computation_axioms is_done list.discI recording_state.simps(4) snd_conv tr) then obtain j where RecvMarker: "\q p. t ! j = RecvMarker cid q p" "j < i" by (metis (no_types, lifting) assms(4) in_set_conv_nth length_take min.absorb2 nth_take) then have step_j: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis Suc_eq_plus1 assms(1) distributed_system.step_Suc distributed_system_axioms assms(4) less_le_trans) then have "t ! j = RecvMarker cid q p" by (metis RecvMarker(1) RecvMarker_given_channel assms(3) event.disc(25) event.sel(10) happen_implies_can_occur) then have "Marker \ set (msgs (S t (j+1)) cid)" using assms(1) assms(3) no_marker_after_RecvMarker step_j by presburger moreover have "has_snapshotted (S t (j+1)) p" using Suc_eq_plus1 \t ! j = RecvMarker cid q p\ assms(1) recv_marker_means_snapshotted snapshot_state_unchanged step_j by presburger ultimately have "Marker \ set (msgs (S t i) cid)" by (metis RecvMarker(2) Suc_eq_plus1 Suc_leI assms(1,3) assms(4) no_marker_and_snapshotted_implies_no_more_markers_trace) then show False using assms by simp qed lemma keep_empty_if_no_events: shows "\ trace init t final; i \ j; j \ length t; msgs (S t i) cid = []; has_snapshotted (S t i) p; channel cid = Some (p, q); \k. i \ k \ k < j \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p \ \ msgs (S t j) cid = []" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" proof - have "i < length t" using Suc.hyps(2) Suc.prems(3) by linarith then show ?thesis by (metis (full_types) Suc.prems(1) Suc_eq_plus1 step_Suc) qed have "msgs (S t (i+1)) cid = []" proof (cases "t ! i") case (Snapshot r) have "r \ p" proof (rule ccontr) assume "~ r \ p" moreover have "can_occur (t ! i) (S t i)" using happen_implies_can_occur local.step by blast ultimately show False using can_occur_def Snapshot Suc by simp qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Snapshot local.step Suc by auto then show ?thesis using Suc by simp next case (RecvMarker cid' r s) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "msgs (S t i) cid \ []" by (metis RecvMarker length_greater_0_conv linorder_not_less list.size(3) local.step nat_less_le recv_marker_other_channels_not_shrinking) then show False using Suc by simp qed then show ?thesis proof (cases "has_snapshotted (S t i) r") case True then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using RecvMarker Suc step \cid \ cid'\ by auto then show ?thesis using Suc by simp next case False have "r \ p" using False Suc.prems(5) by blast then show ?thesis using RecvMarker Suc step \cid \ cid'\ False by simp qed next case (Trans r u u') then show ?thesis using Suc step by simp next case (Send cid' r s u u' m) have "r \ p" proof (rule ccontr) assume "~ r \ p" then have "occurs_on (t ! i) = p \ regular_event (t ! i)" using Send by simp moreover have "i \ i \ i < j" using Suc by simp ultimately show False using Suc.prems by blast qed have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "channel cid = channel cid'" by auto then have "channel cid' = Some (r, s)" using Send step can_occur_def by simp then show False using Suc \r \ p\ \~ cid \ cid'\ by auto qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send Suc by simp then show ?thesis using Suc by simp next case (Recv cid' r s u u' m) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "msgs (S t i) cid \ []" using Recv local.step by auto then show False using Suc by simp qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto then show ?thesis using Suc by simp qed moreover have "\k. i+1 \ k \ k < j \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p" using Suc by simp moreover have "has_snapshotted (S t (i+1)) p" - by (meson Suc.prems(1) Suc.prems(5) discrete less_not_refl nat_le_linear snapshot_stable_ver_3) - moreover have "i+1 \ j" using Suc by simp - moreover have "j \ length t" using Suc by simp + using Suc.prems(5) local.step snapshot_state_unchanged [of \S t i\ \t ! i\ \S t (Suc i)\] + by simp moreover have "j - (i+1) = n" using Suc by linarith - ultimately show ?case using Suc by blast + ultimately show ?case using Suc by auto qed lemma last_unchanged_or_empty_if_no_events: shows "\ trace init t final; i \ j; j \ length t; msgs (S t i) cid \ []; last (msgs (S t i) cid) = Marker; has_snapshotted (S t i) p; channel cid = Some (p, q); \k. i \ k \ k < j \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p \ \ msgs (S t j) cid = [] \ (msgs (S t j) cid \ [] \ last (msgs (S t j) cid) = Marker)" proof (induct "j - i" arbitrary: i) case 0 then show ?case by auto next case (Suc n) then have step: "(S t i) \ (t ! i) \ (S t (i+1))" proof - have "i < length t" using Suc.hyps(2) Suc.prems(3) by linarith then show ?thesis by (metis (full_types) Suc.prems(1) Suc_eq_plus1 step_Suc) qed have msgs_s_ip1: "msgs (S t (i+1)) cid = [] \ (msgs (S t (i+1)) cid \ [] \ last (msgs (S t (i+1)) cid) = Marker)" proof (cases "t ! i") case (Snapshot r) have "r \ p" proof (rule ccontr) assume "~ r \ p" moreover have "can_occur (t ! i) (S t i)" using happen_implies_can_occur local.step by blast ultimately show False using can_occur_def Snapshot Suc by simp qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Snapshot local.step Suc by auto then show ?thesis using Suc by simp next case (RecvMarker cid' r s) then show ?thesis proof (cases "cid = cid'") case True then have "msgs (S t (i+1)) cid = []" proof - have "Marker # msgs (S t (i+1)) cid = msgs (S t i) cid" using RecvMarker True local.step by auto then show ?thesis proof - assume a1: "Marker # msgs (S t (i + 1)) cid = msgs (S t i) cid" have "i < j" by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_neq_Zero diff_is_0_eq le_neq_implies_less) then have "i < length t" using Suc.prems(3) less_le_trans by blast then show ?thesis using a1 by (metis (no_types) Marker_in_channel_implies_not_done RecvMarker Suc.prems(1) Suc.prems(5) Suc.prems(7) Suc_eq_plus1 Suc_le_eq True last_ConsR last_in_set recv_marker_means_cs_Done) qed qed then show ?thesis by simp next case False then show ?thesis proof (cases "has_snapshotted (S t i) r") case True then show ?thesis using False RecvMarker Suc.prems(5) local.step by auto next case False then have "r \ p" using Suc.prems(6) by blast with RecvMarker False Suc.prems step \cid \ cid'\ show ?thesis by auto qed qed next case (Trans r u u') then show ?thesis using Suc step by simp next case (Send cid' r s u u' m) have "r \ p" proof (rule ccontr) assume "~ r \ p" then have "occurs_on (t ! i) = p \ regular_event (t ! i)" using Send by simp moreover have "i \ i \ i < j" using Suc by simp ultimately show False using Suc.prems by blast qed have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "channel cid = channel cid'" by auto then have "channel cid' = Some (r, s)" using Send step can_occur_def by simp then show False using Suc \r \ p\ \~ cid \ cid'\ by auto qed then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send by simp then show ?thesis using Suc by simp next case (Recv cid' r s u u' m) then show ?thesis proof (cases "cid = cid'") case True then have "msgs (S t i) cid = Msg m # msgs (S t (i+1)) cid" using Recv local.step by auto then have "last (msgs (S t (i+1)) cid) = Marker" by (metis Suc.prems(5) last.simps message.simps(3)) then show ?thesis by blast next case False then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto then show ?thesis using Suc by simp qed qed then show ?case proof (elim disjE, goal_cases) case 1 moreover have "trace init t final" using Suc by simp moreover have "i+1 \ j" using Suc by simp moreover have "j \ length t" using Suc by simp moreover have "has_snapshotted (S t (i+1)) p" using Suc.prems(6) local.step snapshot_state_unchanged by auto moreover have "j - (i+1) = n" using Suc by linarith moreover have "\k. i+1 \ k \ k < j \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p" using Suc by auto ultimately have "msgs (S t j) cid = []" using keep_empty_if_no_events Suc.prems(7) by blast then show ?thesis by simp next case 2 moreover have "trace init t final" using Suc by simp moreover have "i+1 \ j" using Suc by simp moreover have "j \ length t" using Suc by simp moreover have "has_snapshotted (S t (i+1)) p" using Suc.prems(6) local.step snapshot_state_unchanged by auto moreover have "j - (i+1) = n" using Suc by linarith moreover have "\k. i+1 \ k \ k < j \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p" using Suc by auto ultimately show ?thesis using Suc.prems(7) Suc.hyps by blast qed qed lemma cs_after_all_prerecording_events: assumes "trace init t final" and "\i'. i' \ i \ ~ prerecording_event t i'" and "\j'. j' < i \ ~ postrecording_event t j'" and "i \ length t" shows "cs_equal_to_snapshot (S t i) final" proof (unfold cs_equal_to_snapshot_def, rule allI, rule impI) fix cid assume "channel cid \ None" then obtain p q where chan: "channel cid = Some (p, q)" by auto have cs_done: "snd (cs (S t (length t)) cid) = Done" using chan all_channels_done_in_final_state assms(1) final_is_s_t_len_t by blast have "filter ((\) Marker) (msgs (S t i) cid) = takeWhile ((\) Marker) (msgs (S t i) cid)" (is ?B) proof (rule ccontr) let ?m = "msgs (S t i) cid" assume "~ ?B" then obtain j k where range: "j < k" "k < length ?m" and "?m ! j = Marker" "?m ! k \ Marker" using filter_neq_takeWhile by metis then have "Marker \ set ?m" by (metis less_trans nth_mem) moreover have "last ?m \ Marker" proof - have "\l. l < length ?m \ l \ j \ ?m ! l \ Marker" using chan \j < k\ \k < length (msgs (S t i) cid)\ \msgs (S t i) cid ! j = Marker\ assms(1) at_most_one_marker calculation less_trans by blast moreover have "last ?m = ?m ! (length ?m - 1)" by (metis \Marker \ set (msgs (S t i) cid)\ empty_iff last_conv_nth list.set(1)) moreover have "length ?m - 1 \ j" using range by auto ultimately show ?thesis using range by auto qed moreover have "i \ length t" using chan assms(1) calculation(1) computation.exists_next_marker_free_state computation.no_change_if_ge_length_t computation_axioms nat_le_linear by fastforce ultimately have "\j. j < i \ postrecording_event t j" using chan assms(1) msg_after_marker_and_nonempty_implies_postrecording_event by auto then show False using assms by auto qed moreover have "takeWhile ((\) Marker) (msgs (S t i) cid) = map Msg (fst (cs final cid))" proof (cases "snd (cs (S t i) cid)") case NotStarted text \show that q and p have to snapshot, and then reduce it to the case below depending on the order they snapshotted in\ have nsq: "~ has_snapshotted (S t i) q" using NotStarted chan assms(1) cs_in_initial_state_implies_not_snapshotted by auto obtain j where snap_q: "~ has_snapshotted (S t j) q" "has_snapshotted (S t (j+1)) q" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have step_q: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis \\ ps (S t j) q \ None\ add.commute assms(1) le_SucI le_eq_less_or_eq le_refl linorder_neqE_nat no_change_if_ge_length_t plus_1_eq_Suc snap_q step_Suc) obtain k where snap_p: "~ has_snapshotted (S t k) p" "has_snapshotted (S t (k+1)) p" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have bound: "i \ j" proof (rule ccontr) assume "~ i \ j" then have "i \ j+1" by simp then have "has_snapshotted (S t i) q" by (meson assms(1) computation.snapshot_stable_ver_3 computation_axioms snap_q(2)) then show False using nsq by simp qed have step_p: "(S t k) \ (t ! k) \ (S t (k+1))" by (metis \\ ps (S t k) p \ None\ add.commute assms(1) le_SucI le_eq_less_or_eq le_refl linorder_neqE_nat no_change_if_ge_length_t plus_1_eq_Suc snap_p step_Suc) have oq: "occurs_on (t ! j) = q" proof (rule ccontr) assume "~ occurs_on (t ! j) = q" then have "has_snapshotted (S t j) q = has_snapshotted (S t (j+1)) q" using no_state_change_if_no_event step_q by auto then show False using snap_q by blast qed have op: "occurs_on (t ! k) = p" proof (rule ccontr) assume "~ occurs_on (t ! k) = p" then have "has_snapshotted (S t k) p = has_snapshotted (S t (k+1)) p" using no_state_change_if_no_event step_p by auto then show False using snap_p by blast qed have "p \ q" using chan no_self_channel by blast then have "j \ k" using oq op event_occurs_on_unique by blast show ?thesis proof (cases "j < k") case True then have "msgs (S t i) cid = msgs (S t j) cid \ cs (S t i) cid = cs (S t j) cid" proof - have "\k. i \ k \ k < j \ occurs_on (t ! k) \ p \ occurs_on (t ! k) \ q" (is ?Q) proof (rule ccontr) assume "~ ?Q" then obtain l where range: "i \ l" "l < j" and "occurs_on (t ! l) = p \ occurs_on (t ! l) = q" by blast then show False proof (elim disjE, goal_cases) case 1 then show ?thesis proof (cases "regular_event (t ! l)") case True have "l < k" using range \j < k\ by simp have "~ has_snapshotted (S t l) p" using snap_p(1) range \j < k\ snapshot_stable_ver_3 assms(1) by simp then have "prerecording_event t l" using True "1" prerecording_event using s_def snap_q(1) snap_q(2) by fastforce then show False using assms range by blast next case False then have step_l: "(S t l) \ t ! l \ (S t (l+1))" by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD True assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all) then have "has_snapshotted (S t (l+1)) p" using False nonregular_event_induces_snapshot by (metis "1"(3) snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI True assms(1) less_imp_le_nat local.range(2) snap_p(1) snapshot_stable_ver_3) qed next case 2 then show ?thesis proof (cases "regular_event (t ! l)") case True have "~ has_snapshotted (S t l) q" using snap_q(1) range \j < k\ snapshot_stable_ver_3 assms(1) by simp then have "prerecording_event t l" using True "2" prerecording_event using s_def snap_q(2) by fastforce then show False using assms range by blast next case False then have step_l: "(S t l) \ t ! l \ (S t (l+1))" by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD True assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all) then have "has_snapshotted (S t (l+1)) q" using False nonregular_event_induces_snapshot by (metis "2"(3) snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI assms(1) range(2) snap_q(1) snapshot_stable_ver_3) qed qed qed moreover have "j \ length t" proof (rule ccontr) assume "~ j \ length t" then have "S t j = S t (j+1)" using no_change_if_ge_length_t assms by simp then show False using snap_q by auto qed ultimately show ?thesis using chan same_messages_if_no_occurrence_trace assms less_imp_le bound by blast qed moreover have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid) \ snd (cs (S t (j+1)) cid) = Recording" proof - have "~ regular_event (t ! j)" using snap_q using regular_event_cannot_induce_snapshot step_q by blast moreover have "Marker \ set (msgs (S t j) cid)" by (meson chan True assms(1) computation.no_marker_if_no_snapshot computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_p(1)) ultimately show ?thesis using oq snapshot_step_cs_preservation_q step_q chan snap_q(1) by blast qed moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid)" proof - have "snd (cs (S t (j+1)) cid) = Recording" using calculation by simp moreover have "\a. j+1 \ a \ a < k \ ~ occurs_on (t ! a) = p" (is ?R) proof (rule ccontr) assume "~ ?R" then obtain a where "j+1 \ a" "a < k" and ocp: "occurs_on (t ! a) = p" by blast have "a < length t" proof - have "k < length t" proof (rule ccontr) assume "~ k < length t" then have "S t k = S t (k+1)" using assms(1) no_change_if_ge_length_t by auto then show False using snap_p by auto qed then show ?thesis using \a < k\ by simp qed then show False proof (cases "regular_event (t ! a)") case True have "~ has_snapshotted (S t a) p" by (meson \a < k\ assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_p(1)) then have "prerecording_event t a" using \a < length t\ ocp True prerecording_event by simp then show False using \j+1 \ a\ \j \ i\ assms by auto next case False then have "(S t a) \ (t ! a) \ (S t (a+1))" using \a < length t\ assms(1) step_Suc by auto then have "has_snapshotted (S t (a+1)) p" by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI \a < k\ assms(1) snap_p(1) snapshot_stable_ver_3) qed qed moreover have "~ has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 Suc_le_eq True assms(1) computation.snapshot_stable_ver_2 computation_axioms snap_p(1)) ultimately show ?thesis using chan cs_when_recording_2 True assms(1) by auto qed moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k+1)) cid)" proof - have "\ regular_event (t ! k)" using regular_event_preserves_process_snapshots snap_p(1) snap_p(2) step_p by force then show ?thesis using chan computation.snapshot_step_cs_preservation_p computation_axioms op step_p by fastforce qed moreover have "map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k+1)) cid) = map Msg (fst (cs final cid))" proof - have f1: "\f p pa pb c ca es n a na. \ computation f p pa pb (c::('a, 'b, 'c) configuration) ca \ \ distributed_system.trace f p pa pb c es ca \ ps (distributed_system.s f p pa pb c es n) a = None \ \ n \ na \ ps (distributed_system.s f p pa pb c es na) a \ None" by (meson computation.snapshot_stable_ver_2) have f2: "computation channel trans send recv init (S t (length t))" using assms(1) final_is_s_t_len_t computation_axioms by blast have f3: "trace init t (S t (length t))" using assms(1) final_is_s_t_len_t by blast have f4: "ps (S t k) p = None" by (meson snap_p(1)) then have f5: "k < length t" using f3 f2 f1 by (metis le_eq_less_or_eq not_le s_def snap_p(2) take_all) have "\ regular_event (t ! k)" using f4 by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(2) step_p) then have f6: "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k + 1)) cid) \ snd (cs (S t k) cid) = snd (cs (S t (k + 1)) cid)" using chan computation.snapshot_step_cs_preservation_p computation_axioms op step_p by fastforce then have f7: "snd (cs (S t (k + 1)) cid) \ Done" using f5 f4 by (metis (no_types) assms(1) chan cs_done_implies_both_snapshotted(1)) have "j + 1 \ k + 1" using True by linarith then have "snd (cs (S t (k + 1)) cid) = Recording" using f7 f3 f2 f1 by (meson chan computation.cs_in_initial_state_implies_not_snapshotted recording_state.exhaust snap_q(2)) then show ?thesis using f6 f5 by (metis (no_types) Suc_eq_plus1 Suc_leI assms(1) chan cs_done cs_done_implies_both_snapshotted(1) cs_when_recording final_is_s_t_len_t le_eq_less_or_eq snap_p(2)) qed ultimately show ?thesis by (metis (no_types, lifting) chan Nil_is_map_conv assms(1) computation.no_initial_channel_snapshot computation_axioms fst_conv no_recording_cs_if_not_snapshotted self_append_conv2 snap_q(1)) next case False then have "k < j" using \j \ k\ False by simp then have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid) = map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid)" proof (cases "i \ k") case True then have "msgs (S t i) cid = msgs (S t k) cid \ cs (S t i) cid = cs (S t k) cid" proof - have "\j. i \ j \ j < k \ occurs_on (t ! j) \ p \ occurs_on (t ! j) \ q" (is ?Q) proof (rule ccontr) assume "~ ?Q" then obtain l where range: "i \ l" "l < k" and "occurs_on (t ! l) = p \ occurs_on (t ! l) = q" by blast then show False proof (elim disjE, goal_cases) case 1 then show ?thesis proof (cases "regular_event (t ! l)") case True have "l < k" using range \k < j\ by simp have "~ has_snapshotted (S t l) p" using snap_p(1) range \k < j\ snapshot_stable_ver_3 assms(1) by simp then have "prerecording_event t l" using True "1" prerecording_event using s_def snap_p by fastforce then show False using assms range by blast next case False then have step_l: "(S t l) \ t ! l \ (S t (l+1))" by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all) then have "has_snapshotted (S t (l+1)) p" using False nonregular_event_induces_snapshot by (metis "1"(3) snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI assms(1) local.range(2) snap_p(1) snapshot_stable_ver_3) qed next case 2 then show ?thesis proof (cases "regular_event (t ! l)") case True have "~ has_snapshotted (S t l) p" using snap_p(1) range \k < j\ snapshot_stable_ver_3 assms(1) by simp moreover have "l < length t" using \k < j\ local.range(2) s_def snap_q(1) snap_q(2) by force ultimately have "prerecording_event t l" using True "2" prerecording_event proof - have "l \ j" by (meson False \l < k\ less_trans not_less) then show ?thesis by (metis (no_types) True \l < length t\ \occurs_on (t ! l) = q\ assms(1) computation.prerecording_event computation.snapshot_stable_ver_2 computation_axioms snap_q(1)) qed then show False using assms range by blast next case False then have step_l: "(S t l) \ t ! l \ (S t (l+1))" by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all) then have "has_snapshotted (S t (l+1)) q" using False nonregular_event_induces_snapshot by (metis "2"(3) snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI \k < j\ assms(1) less_imp_le_nat local.range(2) snap_q(1) snapshot_stable_ver_3) qed qed qed moreover have "k \ length t" proof (rule ccontr) assume "~ k \ length t" then have "S t k = S t (k+1)" using no_change_if_ge_length_t assms by simp then show False using snap_p by auto qed ultimately show ?thesis using chan same_messages_if_no_occurrence_trace assms True less_imp_le by auto qed moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k+1)) cid) \ snd (cs (S t (k+1)) cid) = NotStarted" proof - have "~ regular_event (t ! k)" using snap_p using regular_event_cannot_induce_snapshot step_p by blast then show ?thesis using calculation op snapshot_step_cs_preservation_p step_p chan NotStarted by auto qed moreover have "map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k+1)) cid) = map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid)" proof - have "\a. k+1 \ a \ a < j \ ~ occurs_on (t ! a) = q" (is ?R) proof (rule ccontr) assume "~ ?R" then obtain a where "k+1 \ a" "a < j" and ocp: "occurs_on (t ! a) = q" by blast have "a < length t" proof - have "j < length t" proof (rule ccontr) assume "~ j < length t" then have "S t j = S t (j+1)" using assms(1) no_change_if_ge_length_t by auto then show False using snap_q by auto qed then show ?thesis using \a < j\ by simp qed then show False proof (cases "regular_event (t ! a)") case True have "~ has_snapshotted (S t a) q" by (meson \a < j\ assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_q(1)) then have "prerecording_event t a" using \a < length t\ ocp True prerecording_event by simp then show False using \k+1 \ a\ \k \ i\ assms by auto next case False then have "(S t a) \ (t ! a) \ (S t (a+1))" using \a < length t\ assms(1) step_Suc by auto then have "has_snapshotted (S t (a+1)) q" by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI \a < j\ assms(1) snap_q(1) snapshot_stable_ver_3) qed qed moreover have "Marker : set (msgs (S t (k+1)) cid)" using chan \map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k + 1)) cid) \ snd (cs (S t (k + 1)) cid) = NotStarted\ assms(1) cs_in_initial_state_implies_not_snapshotted marker_must_stay_if_no_snapshot snap_p(2) by blast moreover have "has_snapshotted (S t (k+1)) p" using snap_p(2) by blast moreover have "~ has_snapshotted (S t (k+1)) q" using chan \map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k + 1)) cid) \ snd (cs (S t (k + 1)) cid) = NotStarted\ assms(1) cs_in_initial_state_implies_not_snapshotted by blast moreover have "k+1 \ j" using \k < j\ by auto moreover have "trace init t final" using assms by simp moreover have "snd (cs (S t (k+1)) cid) = NotStarted" using \map Msg (fst (cs (S t k) cid)) @ takeWhile ((\) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (k + 1)) cid) \ snd (cs (S t (k + 1)) cid) = NotStarted\ by blast ultimately show ?thesis using cs_when_recording_3 chan by simp qed ultimately show ?thesis by simp next case False show ?thesis proof - have "has_snapshotted (S t i) p" by (metis False Suc_eq_plus1 assms(1) not_less_eq_eq snap_p(2) snapshot_stable_ver_3) moreover have "~ has_snapshotted (S t i) q" using nsq by auto moreover have "Marker : set (msgs (S t i) cid)" using chan assms(1) calculation(1) marker_must_stay_if_no_snapshot nsq by blast moreover have "\k. i \ k \ k < j \ ~ occurs_on (t ! k) = q" (is ?R) proof (rule ccontr) assume "~ ?R" then obtain k where "i \ k" "k < j" and ocp: "occurs_on (t ! k) = q" by blast have "k < length t" proof - have "j < length t" proof (rule ccontr) assume "~ j < length t" then have "S t j = S t (j+1)" using assms(1) no_change_if_ge_length_t by auto then show False using snap_q by auto qed then show ?thesis using \k < j\ by simp qed then show False proof (cases "regular_event (t ! k)") case True have "~ has_snapshotted (S t k) q" by (meson \k < j\ assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_q(1)) then have "prerecording_event t k" using \k < length t\ ocp True prerecording_event by simp then show False using \i \ j\ \k \ i\ assms by auto next case False then have "(S t k) \ (t ! k) \ (S t (k+1))" using \k < length t\ assms(1) step_Suc by auto then have "has_snapshotted (S t (k+1)) q" by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged) then show False by (metis Suc_eq_plus1 Suc_leI \k < j\ assms(1) snap_q(1) snapshot_stable_ver_3) qed qed ultimately show ?thesis using cs_when_recording_3 using NotStarted assms(1) bound chan by auto qed qed moreover have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs final cid))" proof (cases "\q p. t ! j = RecvMarker cid q p") case True then have "fst (cs (S t j) cid) = fst (cs (S t (j+1)) cid)" using step_q by auto moreover have RecvMarker: "t ! j = RecvMarker cid q p" proof - have "can_occur (t ! j) (S t j)" using happen_implies_can_occur step_q by simp then show ?thesis using RecvMarker_given_channel True chan by force qed moreover have "takeWhile ((\) Marker) (msgs (S t j) cid) = []" proof - have "can_occur (t ! j) (S t j)" using happen_implies_can_occur step_q by simp then have "length (msgs (S t j) cid) > 0 \ hd (msgs (S t j) cid) = Marker" using RecvMarker can_occur_def by auto then show ?thesis by (metis (mono_tags, lifting) hd_conv_nth length_greater_0_conv nth_mem set_takeWhileD takeWhile_nth) qed moreover have "snd (cs (S t (j+1)) cid) = Done" using step_q True by auto moreover have "cs (S t (j+1)) cid = cs final cid" using chan calculation cs_done_implies_same_snapshots assms(1) by (metis final_is_s_t_len_t nat_le_linear no_change_if_ge_length_t) ultimately show ?thesis by simp next case False have "~ regular_event (t ! j)" using regular_event_preserves_process_snapshots snap_q(1) snap_q(2) step_q by auto then have "isSnapshot (t ! j) \ isRecvMarker (t ! j)" using nonregular_event by auto then have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid) \ snd (cs (S t (j+1)) cid) = Recording" proof (elim disjE, goal_cases) case 1 have Snapshot: "t ! j = Snapshot q" using "1" oq by auto then have "msgs (S t j) cid = msgs (S t (j+1)) cid" using \p \ q\ step_q chan by auto moreover have "cs (S t (j+1)) cid = (fst (cs (S t j) cid), Recording)" using step_q Snapshot chan by simp ultimately show ?thesis by simp next case 2 obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' q r" by (metis "2" event.collapse(5) oq) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "channel cid = channel cid'" by simp then have "channel cid' = Some (r, q)" using False RecvMarker \\ cid \ cid'\ by blast then show False using False RecvMarker \\ cid \ cid'\ by blast qed then have "msgs (S t j) cid = msgs (S t (j+1)) cid" using \cid \ cid'\ step_q snap_q RecvMarker chan \p \ q\ by simp moreover have "cs (S t (j+1)) cid = (fst (cs (S t j) cid), Recording)" using \p \ q\ \cid \ cid'\step_q snap_q RecvMarker chan by auto ultimately show ?thesis by simp qed moreover have "map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid) = map Msg (fst (cs final cid))" proof - have "snd (cs (S t (j+1)) cid) = Recording" using calculation by blast moreover have "has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 Suc_leI \k < j\ assms(1) le_add1 snap_p(2) snapshot_stable_ver_3) moreover have "has_snapshotted (S t (j+1)) q" using snap_q by auto moreover have "j < length t" by (metis (no_types, lifting) chan Suc_eq_plus1 assms(1) cs_done cs_done_implies_both_snapshotted(2) computation.no_change_if_ge_length_t computation.snapshot_stable_ver_3 computation_axioms leI le_Suc_eq snap_q(1) snap_q(2)) ultimately show ?thesis using cs_when_recording assms(1) cs_done final_is_s_t_len_t proof - assume a1: "j < length t" assume a2: "trace init t final" assume a3: "snd (cs (S t (length t)) cid) = Done" assume a4: "snd (cs (S t (j + 1)) cid) = Recording" assume a5: "ps (S t (j + 1)) p \ None" assume a6: "\t. trace init t final \ final = S t (length t)" assume a7: "\i j t p cid q. \i < j; j \ length t; trace init t final; ps (S t i) p \ None; snd (cs (S t i) cid) = Recording; snd (cs (S t j) cid) = Done; channel cid = Some (p, q)\ \ map Msg (fst (cs (S t j) cid)) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid)" have "Suc j < length t" using a3 a2 a1 by (metis (no_types) False Suc_eq_plus1 Suc_lessI chan cs_done_implies_has_snapshotted done_only_from_recv_marker snap_q(1) step_q) then show ?thesis using a7 a6 a5 a4 a3 a2 by (metis (no_types) Suc_eq_plus1 chan nat_le_linear) qed qed ultimately show ?thesis by simp qed ultimately show ?thesis by (metis (no_types, lifting) Nil_is_map_conv assms(1) assms(3) chan cs_done cs_done_implies_has_snapshotted cs_not_nil_implies_postrecording_event nat_le_linear nsq self_append_conv2 snapshot_stable_ver_3) qed next case Recording then obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have snap_q: "has_snapshotted (S t i) q" using Recording assms(1) chan cs_recording_implies_snapshot by blast have fst_cs_empty: "cs (S t i) cid = ([], Recording)" (is ?P) proof (rule ccontr) assume "~ ?P" have "snd (cs (S t i) cid) = Recording" using Recording by simp moreover have "fst (cs (S t i) cid) \ []" using \~ ?P\ prod.collapse calculation by metis ultimately have "\j. j < i \ postrecording_event t j" using assms(1) assms(4) chan cs_not_nil_implies_postrecording_event by blast then show False using assms by auto qed then show ?thesis proof - have i_less_len_t: "i < length t" proof (rule ccontr) assume "~ i < length t" then have "snd (cs (S t i) cid) = Done" by (metis assms(1) cs_done le_eq_less_or_eq nat_le_linear no_change_if_ge_length_t) then show False using Recording by simp qed then have "map Msg (fst (cs final cid)) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid)" proof (cases "j < i") case True then have "has_snapshotted (S t i) p" by (metis Suc_eq_plus1 Suc_leI assms(1) snap_p(2) snapshot_stable_ver_3) moreover have "length t \ length t" by simp ultimately show ?thesis using Recording chan assms(1) cs_done cs_when_recording final_is_s_t_len_t i_less_len_t by blast next case False text \need to show that next message that comes into the channel must be marker\ have "\k. i \ k \ k < j \ ~ occurs_on (t ! k) = p" (is ?P) proof (rule ccontr) assume "~ ?P" then obtain k where "i \ k" "k < j" "occurs_on (t ! k) = p" by blast then show False proof (cases "regular_event (t ! k)") case True then have "prerecording_event t k" by (metis (no_types, opaque_lifting) \k < j\ \occurs_on (t ! k) = p\ all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t computation.prerecording_event computation_axioms less_trans nat_le_linear not_less snap_p(1) snapshot_stable_ver_2) then show ?thesis using assms \i \ k\ by auto next case False then have step_k: "(S t k) \ (t ! k) \ (S t (Suc k))" by (metis (no_types, lifting) Suc_leI \k < j\ all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t le_Suc_eq less_imp_Suc_add linorder_not_less no_change_if_ge_length_t snap_p(1) step_Suc) then have "has_snapshotted (S t (Suc k)) p" by (metis False \occurs_on (t ! k) = p\ nonregular_event_induces_snapshot snapshot_state_unchanged) then have "k \ j" by (metis Suc_leI \k < j\ assms(1) snap_p(1) snapshot_stable_ver_3) then show False using \k < j\ by simp qed qed moreover have "~ has_snapshotted (S t i) p" using False assms(1) snap_p(1) snapshot_stable_ver_3 by auto ultimately have to_snapshot: "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\) Marker) (msgs (S t i) cid)" using False chan Recording assms(1) cs_when_recording_2 by auto have step_j: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis Suc_eq_plus1 Suc_le_eq assms(1) distributed_system.step_Suc distributed_system_axioms computation.no_change_if_ge_length_t computation_axioms le_add1 not_less_eq_eq snap_p(1) snap_p(2)) then have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\) Marker) (msgs (S t j) cid) = map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid)" proof - have o: "~ regular_event (t ! j) \ occurs_on (t ! j) = p" by (metis (no_types, opaque_lifting) distributed_system.no_state_change_if_no_event distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_j) then show ?thesis using chan snapshot_step_cs_preservation_p step_j by blast qed moreover have "map Msg (fst (cs final cid)) = map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\) Marker) (msgs (S t (j+1)) cid)" proof - have "snd (cs (S t (j+1)) cid) = Recording" proof - have f1: "ps (S t j) p = None" by (meson snap_p(1)) then have f2: "j < length t" by (metis (no_types) all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t linorder_not_le snapshot_stable_ver_3) have "t ! j \ RecvMarker cid q p" using f1 by (metis (no_types) Suc_eq_plus1 assms(1) recv_marker_means_snapshotted step_j) then show ?thesis using f2 f1 by (meson False assms(1) chan cs_done_implies_both_snapshotted(1) cs_in_initial_state_implies_not_snapshotted cs_not_not_started_stable done_only_from_recv_marker linorder_not_le recording_state.exhaust snap_q snapshot_stable_ver_3 step_j) qed moreover have "j+1 < length t" proof (rule ccontr) assume "~ j+1 < length t" then have "snd (cs (S t (j+1)) cid) = Done" by (metis assms(1) cs_done le_Suc_eq less_imp_Suc_add linorder_not_le no_change_if_ge_length_t) then show False using calculation by auto qed ultimately show ?thesis using chan snap_p(2) assms final_is_s_t_len_t cs_when_recording cs_done by blast qed ultimately show ?thesis using to_snapshot by simp qed then show ?thesis using fst_cs_empty by simp qed next case Done text \msgs must be empty, and cs must also be empty\ have fst_cs_empty: "fst (cs (S t i) cid) = []" proof (rule ccontr) assume "~ fst (cs (S t i) cid) = []" then have "fst (cs (S t 0) cid) \ fst (cs (S t i) cid)" by (metis chan assms(1) cs_not_nil_implies_postrecording_event gr_implies_not0 le0) then have "\j. j < i \ postrecording_event t j" using chan \fst (cs (S t i) cid) \ []\ assms(1) assms(4) cs_not_nil_implies_postrecording_event by blast then show False using assms by auto qed moreover have "msgs (S t i) cid = []" proof - have no_marker: "Marker \ set (msgs (S t i) cid)" (is ?P) proof (rule ccontr) assume "~ ?P" then have "Marker : set (msgs (S t i) cid)" by simp then have "snd (cs (S t i) cid) \ Done" by (metis Marker_in_channel_implies_not_done chan assms(1) nat_le_linear s_def take_all) then show False using Done by simp qed have snap_both: "has_snapshotted (S t i) p \ has_snapshotted (S t i) q" by (metis chan Done assms(1) cs_done_implies_both_snapshotted(1) cs_done_implies_has_snapshotted final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation_axioms le_refl not_less s_def take_all) obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p" by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p) have "j < i" by (meson assms(1) not_le_imp_less snap_both snap_p(1) snapshot_stable_ver_2) have step_j: "(S t j) \ (t ! j) \ (S t (j+1))" by (metis Suc_eq_plus1 assms(1) distributed_system.step_Suc distributed_system_axioms computation.no_change_if_ge_length_t computation_axioms le_add1 linorder_not_less snap_p(1) snap_p(2)) have nonreg_j: "~ regular_event (t ! j)" by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_j) have oc_j: "occurs_on (t ! j) = p" using no_state_change_if_no_event snap_p(1) snap_p(2) step_j by force have "msgs (S t i) cid = [] \ (msgs (S t i) cid \ [] \ last (msgs (S t i) cid) = Marker)" proof - have "msgs (S t (j+1)) cid \ [] \ last (msgs (S t (j+1)) cid) = Marker" proof - have "msgs (S t (j+1)) cid = msgs (S t j) cid @ [Marker]" proof - have "isSnapshot (t ! j) \ isRecvMarker (t ! j)" using nonregular_event nonreg_j by blast then show ?thesis proof (elim disjE, goal_cases) case 1 then have "t ! j = Snapshot p" using oc_j by auto then show ?thesis using step_j chan by auto next case 2 then obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' p r" by (metis event.collapse(5) oc_j) have "cid \ cid'" proof (rule ccontr) assume "~ cid \ cid'" then have "channel cid = channel cid'" by auto then have "Some (p, q) = Some (r, p)" by (metis RecvMarker RecvMarker_implies_Marker_in_set assms(1) chan computation.no_marker_if_no_snapshot computation_axioms snap_p(1) step_j) then show False using no_self_channel chan by simp qed then show ?thesis using oc_j snap_p step_j chan RecvMarker by auto qed qed then show ?thesis by auto qed moreover have "i \ length t" using assms by simp moreover have "j+1 \ i" using \j < i\ by simp moreover have "\k. j+1 \ k \ k < i \ regular_event (t ! k) \ ~ occurs_on (t ! k) = p" (is ?R) proof (rule ccontr) assume "~ ?R" then obtain k where range: "j+1 \ k" "k < i" and "regular_event (t ! k)" "occurs_on (t ! k) = p" by blast then have "postrecording_event t k" using snap_p by (meson assms(1) calculation(2) le_trans linorder_not_less pre_if_regular_and_not_post prerecording_event snapshot_stable_ver_2) then show False using assms range by auto qed ultimately show ?thesis using assms(1) chan last_unchanged_or_empty_if_no_events snap_p(2) by auto qed then show ?thesis using no_marker last_in_set by fastforce qed ultimately show ?thesis using chan Done assms(1) assms(4) final_is_s_t_len_t computation.cs_done_implies_same_snapshots computation_axioms by fastforce qed ultimately show "filter ((\) Marker) (msgs (S t i) cid) = map Msg (fst (cs final cid))" by simp qed lemma snapshot_after_all_prerecording_events: assumes "trace init t final" and "\i'. i' \ i \ ~ prerecording_event t i'" and "\j'. j' < i \ ~ postrecording_event t j'" and "i \ length t" shows "state_equal_to_snapshot (S t i) final" proof (unfold state_equal_to_snapshot_def, rule conjI) show "ps_equal_to_snapshot (S t i) final" using assms ps_after_all_prerecording_events by auto show "cs_equal_to_snapshot (S t i) final" using assms cs_after_all_prerecording_events by auto qed subsection \Obtaining the desired traces\ abbreviation all_prerecording_before_postrecording where "all_prerecording_before_postrecording t \ \i. (\j. j < i \ ~ postrecording_event t j) \ (\j. j \ i \ ~ prerecording_event t j) \ i \ length t \ trace init t final" definition count_violations :: "('a, 'b, 'c) trace \ nat" where "count_violations t = sum (\i. if postrecording_event t i then card {j \ {i+1.. {i+1.. 0" by simp lemma count_violations_ge_0: shows "count_violations t \ 0" by simp lemma violations_0_implies_all_subterms_0: assumes "count_violations t = 0" shows "\i \ {0.. {i+1..i. if postrecording_event t i then card {j \ {i+1..i \ {0.. {i+1.. 0" shows "\i. postrecording_event t i \ card {j \ {i+1.. 0" (is ?P) proof (rule ccontr) assume "~ ?P" then have "\i. ~ postrecording_event t i \ card {j \ {i+1..i. (if postrecording_event t i then card {j \ {i+1..\i. \ postrecording_event t i \ card {j \ {i + 1.. by auto then show "sum (\i. if postrecording_event t i then card {j \ {i+1.. {i+1.. 0" shows "\j \ {i+1.. {i+1.. 0" proof - have "j < length t" using prerecording_event assms by auto have "{j \ {i+1.. empty" using Suc_eq_plus1 \j < length t\ assms(1) assms(3) less_imp_triv by auto then show ?thesis by fastforce qed lemma exists_neighboring_violation_pair: assumes "trace init t final" and "count_violations t > 0" shows "\i j. i < j \ postrecording_event t i \ prerecording_event t j \ (\k. (i < k \ k < j) \ ~ regular_event (t ! k)) \ j < length t" proof - let ?I = "{i. postrecording_event t i \ card {j \ {i+1.. 0}" have nonempty_I: "?I \ empty" using assms exists_postrecording_violation_if_count_greater_0 by blast have fin_I: "finite ?I" proof (rule ccontr) assume "~ finite ?I" then obtain i where "i > length t" "postrecording_event t i" by (simp add: postrecording_event) then show False using postrecording_event by simp qed let ?i = "Max ?I" have no_greater_postrec_violation: "\i. i > ?i \ ~ (postrecording_event t i \ card {j \ {i+1.. 0)" using Max_gr_iff fin_I by blast have post_i: "postrecording_event t ?i" using Max_in fin_I nonempty_I by blast have "card {j \ {?i+1.. 0" proof - have "?i \ ?I" using Max_in fin_I nonempty_I by blast then show ?thesis by simp qed let ?J = "{j \ {?i+1.. empty" using \card {j \ {?i+1.. 0\ exists_prerecording_violation_when_card_greater_0 by blast have fin_J: "finite ?J" by auto let ?j = "Min ?J" have pre_j: "prerecording_event t ?j" using Min_in fin_J nonempty_J by blast have no_smaller_prerec_violation: "\j \ {?i+1.. ~ prerecording_event t j" using Min_less_iff fin_J by blast have j_less_len_t: "?j < length t" using pre_j prerecording_event by blast have "\k. (?i < k \ k < ?j) \ ~ regular_event (t ! k)" proof (rule allI, rule impI) fix k assume asm: "?i < k \ k < ?j" then show "~ regular_event (t ! k)" proof - have "0_le_k": "0 \ k" by simp have k_less_len_t: "k < length t" using j_less_len_t asm by auto show ?thesis proof (rule ccontr) assume reg_event: "~ ~ regular_event (t ! k)" then show False proof (cases "has_snapshotted (S t k) (occurs_on (t ! k))") case True then have post_k: "postrecording_event t k" using reg_event k_less_len_t postrecording_event by simp moreover have "card {j \ {k+1.. 0" using post_k pre_j card_greater_0_if_post_after_pre asm pre_j by blast ultimately show False using no_greater_postrec_violation asm by blast next case False then have pre_k: "prerecording_event t k" using reg_event k_less_len_t prerecording_event by simp moreover have "k \ {?i+1..k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "trace init t final" shows "{k \ {0.. {0..k. k < i \ t ! k = ?t ! k" by (metis nth_take same_begin) then have "\k. k < i \ prerecording_event t k = prerecording_event ?t k" proof - have "\k. k < i \ S t k = S ?t k" using assms swap_events by simp then show ?thesis unfolding prerecording_event using a same_length by presburger qed then show ?thesis by auto qed lemma same_cardinality_post_swap_2: assumes "prerecording_event t j" and "postrecording_event t i" and "i < j" and "j < length t" and "count_violations t = Suc n" and "\k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "trace init t final" shows "card {k \ {i.. {i.. {i..k. i \ k \ k < j \ ~ prerecording_event t k" proof (rule allI, rule impI) fix k assume asm: "i \ k \ k < j" then show "~ prerecording_event t k" proof (cases "k = i") case True then have "postrecording_event t k" using assms by simp then show ?thesis by (meson computation.postrecording_event computation.prerecording_event computation_axioms) next case False then have "i < k \ k < j" using asm by force then have "~ regular_event (t ! k)" using assms by simp then show ?thesis unfolding prerecording_event by simp qed qed then have "{k \ {i.. {j.. {i.. {i.. postrecording_event ?t (i+1) \ (\k. k > i+1 \ k < j+1 \ ~ regular_event (?t ! k))" using assms swap_events by blast have "\k. i+1 \ k \ k < j+1 \ ~ prerecording_event ?t k" proof (rule allI, rule impI) fix k assume asm: "i+1 \ k \ k < j+1" then show "~ prerecording_event ?t k" proof (cases "k = i+1") case True then have "postrecording_event ?t k" using swap_ind by blast then show ?thesis by (meson computation.postrecording_event computation.prerecording_event computation_axioms) next case False then have "i+1 < k \ k < j+1" using asm by linarith then have "~ regular_event (?t ! k)" using asm assms swap_ind by blast then show ?thesis unfolding prerecording_event by simp qed qed then have "{k \ {i+1.. {i.. {i..k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "trace init t final" shows "{k \ {j+1.. {j+1..k. j+1 \ k \ k < length t \ ?t ! k = t ! k" proof (rule allI, rule impI) fix k assume "j+1 \ k \ k < length t" then have "?t ! k = drop (j+1) (swap_events i j t) ! (k-(j+1))" by (metis (no_types, lifting) Suc_eq_plus1 Suc_leI assms(4) le_add_diff_inverse nth_drop same_length) moreover have "t ! k = drop (j+1) t ! (k-(j+1))" using \j + 1 \ k \ k < length t\ by auto ultimately have "drop (j+1) ?t ! (k-(j+1)) = drop (j+1) t ! (k-(j+1))" using assms swap_identical_tails by metis then show "?t ! k = t ! k" using \?t ! k = drop (j + 1) ?t ! (k - (j + 1))\ \t ! k = drop (j + 1) t ! (k - (j + 1))\ by auto qed then have "\k. j+1 \ k \ k < length t \ prerecording_event t k = prerecording_event ?t k" proof - have "\k. k \ (j+1) \ S t k = S ?t k" using assms swap_events by simp then show ?thesis unfolding prerecording_event using a by auto qed then have "{k \ {j+1.. {j+1..k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "count_violations t = Suc n" and "trace init t final" shows "card {k \ {i+1..k. i < k \ k < j \ ~ prerecording_event t k" proof (rule allI, rule impI) fix k assume asm: "i < k \ k < j" then show "~ prerecording_event t k" proof - have "~ regular_event (t ! k)" using asm assms by blast then show ?thesis unfolding prerecording_event by simp qed qed then have "{k \ {i+1.. {j.. {i+1..k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "count_violations t = Suc n" and "trace init t final" shows "card {k \ {i+1..k. i+1 < k \ k < j+1 \ ~ regular_event (?t ! k)" using assms swap_events by blast have "\k. i+1 \ k \ k < j+1 \ ~ prerecording_event ?t k" proof (rule allI, rule impI) fix k assume asm: "i+1 \ k \ k < j+1" then show "~ prerecording_event ?t k" proof (cases "k = i+1") case True then show ?thesis using postrec_ip1 by (meson computation.postrecording_event computation.prerecording_event computation_axioms) next case False then have "i+1 < k \ k < j+1" using asm by simp then have "~ regular_event (?t ! k)" using neigh_shift by blast then show ?thesis unfolding prerecording_event by simp qed qed then have "{k \ {i+1..k. (i < k \ k < j) \ ~ regular_event (t ! k)" and "count_violations t = Suc n" and "trace init t final" shows "count_violations (swap_events i j t) = n" proof - let ?t = "swap_events i j t" let ?f = "(\i. if postrecording_event t i then card {j \ {i+1..k. k < i \ postrecording_event t k = postrecording_event ?t k" proof - have "\k. k < i \ S t k = S ?t k" using assms swap_events by auto then show ?thesis unfolding postrecording_event proof - assume a1: "\kn na es nb. \ n < na \ \ na < length es \ \ nb < n \ swap_events n na es ! nb = (es ! nb::('a, 'b, 'c) event)" by (metis (no_types) nth_take swap_identical_heads) then have "\ nn < i \ \ nn < length t \ \ nn < length (swap_events i j t) \ \ regular_event (t ! nn) \ \ regular_event (swap_events i j t ! nn) \ ps (S t nn) (occurs_on (t ! nn)) = None \ ps (S (swap_events i j t) nn) (occurs_on (swap_events i j t ! nn)) = None \ regular_event (t ! nn) \ regular_event (swap_events i j t ! nn) \ nn < length t \ nn < length (swap_events i j t) \ ps (S t nn) (occurs_on (t ! nn)) \ None \ ps (S (swap_events i j t) nn) (occurs_on (swap_events i j t ! nn)) \ None" using a1 by (metis (no_types) assms(3) assms(4) swap_identical_length) } then show "\n regular_event (t ! n) \ ps (S t n) (occurs_on (t ! n)) \ None) = (n < length (swap_events i j t) \ regular_event (swap_events i j t ! n) \ ps (S (swap_events i j t) n) (occurs_on (swap_events i j t ! n)) \ None)" by (metis (no_types)) qed qed have same_postrec_suffix: "\k. k \ j+1 \ postrecording_event t k = postrecording_event ?t k" proof - have post_equal_states: "\k. k \ j+1 \ S t k = S ?t k" using assms swap_events by presburger show ?thesis proof (rule allI, rule impI) fix k assume "j+1 \ k" then show "postrecording_event t k = postrecording_event ?t k" proof (cases "k < length t") case False then have "~ postrecording_event t k" using postrecording_event by simp moreover have "~ postrecording_event ?t k" using postrecording_event swap_identical_length False assms by metis ultimately show ?thesis by simp next case True then show "postrecording_event t k = postrecording_event ?t k" using post_equal_states proof - assume a1: "\k\j + 1. S t k = S (swap_events i j t) k" assume a2: "k < length t" have f3: "length t = length (swap_events i j t)" using assms(3) assms(4) swap_identical_length by blast have f4: "k - (j + 1) + (j + 1) = k" using \j + 1 \ k\ le_add_diff_inverse2 by blast have "drop (j + 1) t = drop (j + 1) (swap_events i j t)" using assms(3) assms(4) swap_identical_tails by blast then have "swap_events i j t ! k = t ! k" using f4 f3 a2 by (metis (no_types) drop_drop hd_drop_conv_nth) then show ?thesis using f3 a1 \j + 1 \ k\ postrecording_event by presburger qed qed qed qed - have sum_decomp_f: "sum ?f {0..sum g {0.. + for g :: \nat \ nat\ + using sum.atLeastLessThan_concat [of 0 i \j + 1\ g] sum.atLeastLessThan_concat [of 0 \j + 1\ \length t\ g] assms + by simp + from sum_decomp_g [of ?f] + have sum_decomp_f: \sum ?f {0.. . + from sum_decomp_g [of ?f'] + have sum_decomp_f': \sum ?f' {0.. . have prefix_sum: "sum ?f {0..l. 0 \ l \ l < i \ ?f l = ?f' l" proof (rule allI, rule impI) fix l assume "0 \ l \ l < i" then have "l < i" by simp show "?f l = ?f' l" proof (cases "postrecording_event t l") case True let ?G = "{k \ {l+1.. (?B \ ?C)" using assms \l < i\ by auto then have "card ?G = card (?A \ (?B \ ?C))" by simp also have "card (?A \ (?B \ ?C)) = card ?A + card (?B \ ?C)" proof - have "?A \ (?B \ ?C) = {}" using \l < i\ assms by auto then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed also have "card ?A + card (?B \ ?C) = card ?A + card ?B + card ?C" proof - have "?B \ ?C = {}" by auto then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed finally show ?thesis by simp qed have card_G': "card ?G' = card ?A' + card ?B' + card ?C'" proof - have "?G' = ?A' \ (?B' \ ?C')" using assms \l < i\ by auto then have "card ?G' = card (?A' \ (?B' \ ?C'))" by simp also have "card (?A' \ (?B' \ ?C')) = card ?A' + card (?B' \ ?C')" proof - have "?A' \ (?B' \ ?C') = {}" using \l < i\ assms by auto then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed also have "card ?A' + card (?B' \ ?C') = card ?A' + card ?B' + card ?C'" proof - have "?B' \ ?C' = {}" by auto then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed finally show ?thesis by simp qed have "card ?G = card ?G'" proof - have "card ?A = card ?A'" proof - have "{k \ {0.. {0..l < i\ by blast moreover have "length ?t = length t" using assms(3) assms(4) by fastforce ultimately show ?thesis using True by presburger next case False then have "~ postrecording_event ?t l" using same_postrec_prefix \l < i\ by blast then show ?thesis using False by simp qed qed then show ?thesis using sum_eq_if_same_subterms by auto qed have infix_sum: "sum ?f {i..i < j\ in simp_all) have sum_decomp_f': "sum ?f' {i..i < j\ in simp_all) have "sum ?f {i+2..l. i+2 \ l \ l < j+1 \ ?f l = ?f' l" proof (rule allI, rule impI) fix l assume asm: "i+2 \ l \ l < j+1" have "?f l = 0" proof (cases "l = j") case True then have "~ postrecording_event t l" using assms(1) postrecording_event prerecording_event by auto then show ?thesis by simp next case False then have "i < l \ l < j" using assms asm by simp then have "~ regular_event (t ! l)" using assms by blast then have "~ postrecording_event t l" unfolding postrecording_event by simp then show ?thesis by simp qed moreover have "?f' l = 0" proof - have "\k. i+1 < k \ k < j+1 \ ~ regular_event (?t ! k)" using assms swap_events by blast then have "~ regular_event (?t ! l)" using asm by simp then have "~ postrecording_event ?t l" unfolding postrecording_event by simp then show ?thesis by simp qed ultimately show "?f l = ?f' l" by simp qed then show ?thesis using sum_eq_if_same_subterms by simp qed moreover have "sum ?f {i.. ?B" using assms by auto moreover have "?A \ ?B = {}" by auto ultimately show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed have card_G': "card ?G' = card ?A' + card ?B'" proof - have "?G' = ?A' \ ?B'" using assms by auto moreover have "?A' \ ?B' = {}" by auto ultimately show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal) qed have "card ?G = card ?G' + 1" proof - have "card ?A = card ?A' + 1" proof - have "card ?A = 1" using assms card_ip1_to_j_is_1_in_normal_events by blast moreover have "card ?A' = 0" using assms card_ip1_to_j_is_0_in_swapped_events by force ultimately show ?thesis by algebra qed moreover have "card ?B = card ?B'" using assms same_cardinality_post_swap_3 by force ultimately show ?thesis using card_G card_G' by presburger qed moreover have "card ?G = ?f i" using pi by simp moreover have "card ?G' = ?f' (i+1)" using pip1 by simp ultimately show ?thesis by argo qed ultimately show ?thesis by auto qed ultimately show ?thesis using sum_decomp_f sum_decomp_f' by linarith qed have suffix_sum: "sum ?f {j+1..l. l > j \ ?f l = ?f' l" proof (rule allI, rule impI) fix l assume "l > j" then show "?f l = ?f' l" proof (cases "postrecording_event t l") case True let ?G = "{k \ {l+1..l > j\ by fastforce then show ?thesis by simp qed moreover have "postrecording_event ?t l" using True same_postrec_suffix \l > j\ by simp moreover have "length ?t = length t" using assms(3) assms(4) by fastforce ultimately show ?thesis using True by presburger next case False then have "~ postrecording_event ?t l" using same_postrec_suffix \l > j\ by simp then show ?thesis using False by simp qed qed then have "\k. j+1 \ k \ k < length t \ ?f k = ?f' k" - using discrete by blast + by simp moreover have "length t = length ?t" using assms(3) assms(4) swap_identical_length by blast ultimately show ?thesis by (blast intro:sum_eq_if_same_subterms) qed have "sum ?f {0..t'. mset t' = mset t \ all_prerecording_before_postrecording t'" using assms proof (induct "count_violations t" arbitrary: t) case 0 then show ?case proof (cases "\i. prerecording_event t i") case False then have "\j. ~ prerecording_event t j" by auto then have "\j. j \ 0 \ ~ postrecording_event t j" using "0.prems" init_is_s_t_0 no_initial_process_snapshot postrecording_event by auto moreover have "\j. j > 0 \ ~ prerecording_event t j" using False by auto moreover have "length t > 0" by (metis "0.prems" all_processes_snapshotted_in_final_state length_greater_0_conv no_initial_process_snapshot tr_init trace_and_start_determines_end) ultimately show ?thesis using "0.prems" False by auto next case True let ?Is = "{i. prerecording_event t i}" have "?Is \ empty" by (simp add: True) moreover have fin_Is: "finite ?Is" proof (rule ccontr) assume "~ finite ?Is" then obtain i where "i > length t" "prerecording_event t i" by (simp add: prerecording_event) then show False using prerecording_event by auto qed let ?i = "Max ?Is" have pi: "prerecording_event t ?i" using Max_in calculation fin_Is by blast have "?i < length t" proof (rule ccontr) assume "~ ?i < length t" then show False using calculation fin_Is computation.prerecording_event computation_axioms by force qed moreover have "\j. j \ ?i+1 \ ~ prerecording_event t j" proof - have "\j. j > ?i \ ~ prerecording_event t j" using Max_less_iff fin_Is by auto then show ?thesis by auto qed moreover have "\j. j < ?i+1 \ ~ postrecording_event t j" proof - have "\j. j \ ?i \ ~ postrecording_event t j" proof (rule allI, rule impI, rule ccontr) fix j assume "j \ ?i" "~ ~ postrecording_event t j" then have "j < ?i" by (metis add_diff_inverse_nat dual_order.antisym le_add1 pi postrecording_event prerecording_event) then have "count_violations t > 0" proof - have "(if postrecording_event t j then card {l \ {j+1.. {j+1..~ ~ postrecording_event t j\ by simp moreover have "card {l \ {j+1.. 0" proof - have "j + 1 \ ?i \ ?i < length t" - using \Max {i. prerecording_event t i} < length t\ \j < Max {i. prerecording_event t i}\ discrete by blast + using \Max {i. prerecording_event t i} < length t\ \j < Max {i. prerecording_event t i}\ + by simp moreover have "prerecording_event t ?i" using pi by simp ultimately have "{l \ {j+1.. empty" by fastforce then show ?thesis by fastforce qed ultimately show ?thesis by (metis (no_types, lifting) violations_0_implies_all_subterms_0 \Max {i. prerecording_event t i} < length t\ \j < Max {i. prerecording_event t i}\ atLeastLessThan_iff less_trans linorder_not_le neq0_conv) qed then show False using "0" by simp qed then show ?thesis by auto qed moreover have "?i+1 \ length t" - using calculation(2) discrete by blast + using calculation(2) by simp ultimately show ?thesis using "0.prems" by blast qed next case (Suc n) then obtain i j where ind: "postrecording_event t i" "prerecording_event t j" "\k. (i < k \ k < j) \ ~ regular_event (t ! k)" "i < j" "j < length t" using exists_neighboring_violation_pair Suc by force then have "trace init (swap_events i j t) final \ (\k. k \ j + 1 \ S (swap_events i j t) k = S t k) \ (\k. k \ i \ S (swap_events i j t) k = S t k)" using Suc swap_events by presburger moreover have "mset (swap_events i j t) = mset t" using swap_events_perm ind by blast moreover have "count_violations (swap_events i j t) = n" using count_violations_swap Suc ind by simp ultimately show ?case using Suc.hyps by metis qed theorem snapshot_algorithm_is_correct: assumes "trace init t final" shows "\t' i. trace init t' final \ mset t' = mset t \ state_equal_to_snapshot (S t' i) final \ i \ length t'" proof - obtain t' where "mset t' = mset t" and "all_prerecording_before_postrecording t'" using assms desired_trace_always_exists by blast then show ?thesis using snapshot_after_all_prerecording_events by blast qed subsection \Stable property detection\ text \Finally, we show that the computed snapshot is indeed suitable for stable property detection, as claimed in~\<^cite>\"chandy"\.\ definition stable where "stable p \ (\c. p c \ (\t c'. trace c t c' \ p c'))" lemma has_snapshot_stable: assumes "trace init t final" shows "stable (\c. has_snapshotted c p)" using snapshot_stable stable_def by auto definition some_snapshot_state where "some_snapshot_state t \ SOME (t', i). trace init t final \ trace init t' final \ mset t' = mset t \ state_equal_to_snapshot (S t' i) final" lemma split_S: assumes "trace init t final" shows "trace (S t i) (drop i t) final" proof - have "t = take i t @ drop i t" by simp then show ?thesis by (metis split_trace assms exists_trace_for_any_i trace_and_start_determines_end) qed theorem Stable_Property_Detection: assumes "stable p" and "trace init t final" and "(t', i) = some_snapshot_state t" and "p (S t' i)" shows "p final" proof - have "\t' i. trace init t final \ trace init t' final \ mset t' = mset t \ state_equal_to_snapshot (S t' i) final" using snapshot_algorithm_is_correct assms(2) by blast then have "trace init t' final" using assms unfolding some_snapshot_state_def by auto (metis (mono_tags, lifting) case_prod_conv tfl_some) then show ?thesis using assms stable_def split_S by metis qed end (* locale computation *) end (* theory Snapshot *) diff --git a/thys/Chandy_Lamport/Util.thy b/thys/Chandy_Lamport/Util.thy --- a/thys/Chandy_Lamport/Util.thy +++ b/thys/Chandy_Lamport/Util.thy @@ -1,363 +1,366 @@ section \Utilties\ theory Util imports Main "HOL-Library.Sublist" "HOL-Library.Multiset" begin abbreviation swap_events where "swap_events i j t \ take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t) @ drop (j+1) t" lemma swap_neighbors_2: shows "i+1 < length t \ swap_events i (i+1) t = (t[i := t ! (i+1)])[i+1 := t ! i]" proof (induct i arbitrary: t) case 0 then show ?case by (metis One_nat_def Suc_eq_plus1 add_lessD1 append.left_neutral append_Cons cancel_comm_monoid_add_class.diff_cancel drop_update_cancel length_list_update numeral_One take_0 take_Cons_numeral upd_conv_take_nth_drop zero_less_Suc) next case (Suc n) let ?t = "tl t" have "t = hd t # ?t" by (metis Suc.prems hd_Cons_tl list.size(3) not_less_zero) moreover have "swap_events n (n+1) ?t = (?t[n := ?t ! (n+1)])[n+1 := ?t ! n]" by (metis Suc.hyps Suc.prems Suc_eq_plus1 length_tl less_diff_conv) ultimately show ?case by (metis Suc_eq_plus1 append_Cons diff_self_eq_0 drop_Suc_Cons list_update_code(3) nth_Cons_Suc take_Suc_Cons) qed lemma swap_identical_length: assumes "i < j" and "j < length t" shows "length t = length (swap_events i j t)" proof - have "length (take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i+1)) (drop (i+1) t))" by simp then have "... = j+1" using assms(1) assms(2) by auto then show ?thesis using assms(2) by auto qed lemma swap_identical_heads: assumes "i < j" and "j < length t" shows "take i t = take i (swap_events i j t)" using assms by auto lemma swap_identical_tails: assumes "i < j" and "j < length t" shows "drop (j+1) t = drop (j+1) (swap_events i j t)" proof - have "length (take i t @ [t ! j, t ! i] @ take (j - (i+1)) (drop (i+1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i+1)) (drop (i+1) t))" by simp then have "... = j+1" using assms(1) assms(2) by auto then show ?thesis by (metis \length (take i t @ [t ! j, t ! i] @ take (j - (i + 1)) (drop (i + 1) t)) = length (take i t) + length [t ! j, t ! i] + length (take (j - (i + 1)) (drop (i + 1) t))\ append.assoc append_eq_conv_conj) qed lemma swap_neighbors: shows "i+1 < length l \ (l[i := l ! (i+1)])[i+1 := l ! i] = take i l @ [l ! (i+1), l ! i] @ drop (i+2) l" proof (induct i arbitrary: l) case 0 then show ?case by (metis One_nat_def add.left_neutral add_lessD1 append_Cons append_Nil drop_update_cancel length_list_update one_add_one plus_1_eq_Suc take0 take_Suc_Cons upd_conv_take_nth_drop zero_less_two) next case (Suc n) let ?l = "tl l" have "(l[Suc n := l ! (Suc n + 1)])[Suc n + 1 := l ! Suc n] = hd l # (?l[n := ?l ! (n+1)])[n+1 := ?l ! n]" by (metis Suc.prems add.commute add_less_same_cancel2 list.collapse list.size(3) list_update_code(3) not_add_less2 nth_Cons_Suc plus_1_eq_Suc) have "n + 1 < length ?l" using Suc.prems by auto then have "(?l[n := ?l ! (n+1)])[n+1 := ?l ! n] = take n ?l @ [?l ! (n+1), ?l ! n] @ drop (n+2) ?l" using Suc.hyps by simp then show ?case by (cases l) auto qed lemma swap_events_perm: assumes "i < j" and "j < length t" shows "mset (swap_events i j t) = mset t" proof - have swap: "swap_events i j t = take i t @ [t ! j, t ! i] @ (take (j - (i+1)) (drop (i+1) t)) @ (drop (j+1) t)" by auto have reg: "t = take i t @ (take ((j+1) - i) (drop i t)) @ drop (j+1) t" by (metis add_diff_inverse_nat add_lessD1 append.assoc append_take_drop_id assms(1) less_imp_add_positive less_not_refl take_add) have "mset (take i t) = mset (take i t)" by simp moreover have "mset (drop (j+1) t) = mset (drop (j+1) t)" by simp moreover have "mset ([t ! j, t ! i] @ (take (j - (i+1)) (drop (i+1) t))) = mset (take ((j+1) - i) (drop i t))" proof - let ?l = "take (j - (i+1)) (drop (i+1) t)" have "take ((j+1) - i) (drop i t) = t ! i # ?l @ [t ! j]" proof - have f1: "Suc (j - Suc i) = j - i" by (meson Suc_diff_Suc assms(1)) have f2: "i < length t" using assms(1) assms(2) by linarith have f3: "j - (i + 1) + (i + 1) = j" - by (meson assms(1) discrete le_add_diff_inverse2) + using \i < j\ by simp then have "drop (j - (i + 1)) (drop (i + 1) t) = drop j t" by (metis drop_drop) then show ?thesis using f3 f2 f1 by (metis (no_types) Cons_nth_drop_Suc Suc_diff_le Suc_eq_plus1 assms(1) assms(2) hd_drop_conv_nth length_drop less_diff_conv nat_less_le take_Suc_Cons take_hd_drop) qed then show ?thesis by fastforce qed ultimately show ?thesis using swap reg by simp (metis mset_append union_mset_add_mset_left union_mset_add_mset_right) qed lemma sum_eq_if_same_subterms: fixes i :: nat shows "\k. i \ k \ k < j \ f k = f' k \ sum f {i..) a) l \ takeWhile ((\) a) l" shows "\i j. i < j \ j < length l \ l ! i = a \ l ! j \ a" (is ?P) proof (rule ccontr) assume "~ ?P" then have asm: "\i j. i < j \ j < length l \ l ! i \ a \ l ! j = a" (is ?Q) by simp then have "filter ((\) a) l = takeWhile ((\) a) l" proof (cases "a : set l") case False then have "\i. i < length l \ l ! i \ a" by auto then show ?thesis by (metis (mono_tags, lifting) False filter_True takeWhile_eq_all_conv) next case True then have ex_j: "\j. j < length l \ l ! j = a" by (simp add: in_set_conv_nth) let ?j = "Min {j. j < length l \ l ! j = a}" have fin_j: "finite {j. j < length l \ l ! j = a}" using finite_nat_set_iff_bounded by blast moreover have "{j. j < length l \ l ! j = a} \ empty" using ex_j by blast ultimately have "?j < length l" using Min_less_iff by blast have tail_all_a: "\j. j < length l \ j \ ?j \ l ! j = a" proof (rule allI, rule impI) fix j assume "j < length l \ j \ ?j" moreover have "\ ?Q; j < length l \ j \ ?j \ \ \k. k \ ?j \ k \ j \ l ! j = a" proof (induct "j - ?j") case 0 then have "j = ?j" using 0 by simp then show ?case using Min_in \{j. j < length l \ l ! j = a} \ {}\ fin_j by blast next case (Suc n) then have "\k. k \ ?j \ k < j \ l ! j = a" by (metis (mono_tags, lifting) Min_in \{j. j < length l \ l ! j = a} \ {}\ fin_j le_eq_less_or_eq mem_Collect_eq) then show ?case using Suc.hyps(2) by auto qed ultimately show "l ! j = a" using asm by blast qed moreover have head_all_not_a: "\i. i < ?j \ l ! i \ a" using asm calculation by (metis (mono_tags, lifting) Min_le \Min {j. j < length l \ l ! j = a} < length l\ fin_j leD less_trans mem_Collect_eq) ultimately have "takeWhile ((\) a) l = take ?j l" proof - have "length (takeWhile ((\) a) l) = ?j" proof - have "length (takeWhile ((\) a) l) \ ?j" (is ?S) proof (rule ccontr) assume "\ ?S" then have "l ! ?j \ a" by (metis (mono_tags, lifting) not_le_imp_less nth_mem set_takeWhileD takeWhile_nth) then show False using \Min {j. j < length l \ l ! j = a} < length l\ tail_all_a by blast qed moreover have "length (takeWhile ((\) a) l) \ ?j" (is ?T) proof (rule ccontr) assume "\ ?T" then have "\j. j < ?j \ l ! j = a" by (metis (mono_tags, lifting) \Min {j. j < length l \ l ! j = a} < length l\ calculation le_less_trans not_le_imp_less nth_length_takeWhile) then show False using head_all_not_a by blast qed ultimately show ?thesis by simp qed moreover have "length (take ?j l) = ?j" by (metis calculation takeWhile_eq_take) ultimately show ?thesis by (metis takeWhile_eq_take) qed moreover have "filter ((\) a) l = take ?j l" proof - have "filter ((\) a) l = filter ((\) a) (take ?j l) @ filter ((\) a) (drop ?j l)" by (metis append_take_drop_id filter_append) moreover have "filter ((\) a) (take ?j l) = take ?j l" using head_all_not_a by (metis \takeWhile ((\) a) l = take (Min {j. j < length l \ l ! j = a}) l\ filter_id_conv set_takeWhileD) moreover have "filter ((\) a) (drop ?j l) = []" proof - have "\j. j \ ?j \ j < length l \ l ! j = drop ?j l ! (j - ?j)" by simp then have "\j. j < length l - ?j \ drop ?j l ! j = a" using tail_all_a by (metis (no_types, lifting) Groups.add_ac(2) \Min {j. j < length l \ l ! j = a} < length l\ less_diff_conv less_imp_le_nat not_add_less2 not_le nth_drop) then show ?thesis proof - obtain aa :: "('a \ bool) \ 'a list \ 'a" where "\x0 x1. (\v2. v2 \ set x1 \ x0 v2) = (aa x0 x1 \ set x1 \ x0 (aa x0 x1))" by moura then have f1: "\as p. aa p as \ set as \ p (aa p as) \ filter p as = []" by (metis (full_types) filter_False) obtain nn :: "'a list \ 'a \ nat" where f2: "\x0 x1. (\v2 x0 ! nn x0 x1 = x1)" by moura { assume "drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) = a" then have "filter ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l) = [] \ \ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" using f1 by (metis (full_types)) } moreover { assume "\ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length l - Min {n. n < length l \ l ! n = a}" then have "\ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" by simp } ultimately have "filter ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l) = [] \ \ nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) < length (drop (Min {n. n < length l \ l ! n = a}) l) \ drop (Min {n. n < length l \ l ! n = a}) l ! nn (drop (Min {n. n < length l \ l ! n = a}) l) (aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)) \ aa ((\) a) (drop (Min {n. n < length l \ l ! n = a}) l)" using \\j l ! j = a}. drop (Min {j. j < length l \ l ! j = a}) l ! j = a\ by blast then show ?thesis using f2 f1 by (meson in_set_conv_nth) qed qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed then show False using assms by simp qed lemma util_exactly_one_element: assumes "m \ set l" and "l' = l @ [m]" shows "\!j. j < length l' \ l' ! j = m" (is ?P) proof - have "\j. j < length l' - 1 \ l' ! j \ m" by (metis assms(1) assms(2) butlast_snoc length_butlast nth_append nth_mem) then have one_j: "\j. j < length l' \ l' ! j = m \ j = (length l' - 1)" by (metis (no_types, opaque_lifting) diff_Suc_1 lessE) show ?thesis proof (rule ccontr) assume "~ ?P" then obtain i j where "i \ j" "i < length l'" "j < length l'" "l' ! i = m" "l' ! j = m" using assms by auto then show False using one_j by blast qed qed lemma exists_one_iff_filter_one: shows "(\!j. j < length l \ l ! j = a) \ length (filter ((=) a) l) = 1" proof (rule iffI) assume "\!j. j < length l \ l ! j = a" then obtain j where "j < length l" "l ! j = a" by blast moreover have "\k. k \ j \ k < length l \ l ! k \ a" using \\!j. j < length l \ l ! j = a\ \j < length l\ \l ! j = a\ by blast moreover have "l = take j l @ [l ! j] @ drop (j+1) l" by (metis Cons_eq_appendI Cons_nth_drop_Suc Suc_eq_plus1 append_self_conv2 append_take_drop_id calculation(1) calculation(2)) moreover have "filter ((=) a) (take j l) = []" proof - have "\k. k < length (take j l) \ (take j l) ! k \ a" using calculation(3) by auto then show ?thesis by (metis (full_types) filter_False in_set_conv_nth) qed moreover have "filter ((=) a) (drop (j+1) l) = []" proof - have "\k. k < length (drop (j+1) l) \ (drop (j+1) l) ! k \ a" using calculation(3) by auto then show ?thesis by (metis (full_types) filter_False in_set_conv_nth) qed ultimately show "length (filter ((=) a) l) = 1" by (metis (mono_tags, lifting) One_nat_def Suc_eq_plus1 append_Cons append_self_conv2 filter.simps(2) filter_append list.size(3) list.size(4)) next assume asm: "length (filter ((=) a) l) = 1" then have "filter ((=) a) l = [a]" proof - let ?xs = "filter ((=) a) l" have "length ?xs = 1" using asm by blast then show ?thesis by (metis (full_types) Cons_eq_filterD One_nat_def length_0_conv length_Suc_conv) qed then have "\j. j < length l \ l ! j = a" by (metis (full_types) filter_False in_set_conv_nth list.discI) then obtain j where j: "j < length l" "l ! j = a" by blast moreover have "\k. k < length l \ k \ j \ l ! k \ a" proof (rule allI, rule impI) fix k assume assm: "k < length l \ k \ j" + then have \k < length l\ .. show "l ! k \ a" proof (rule ccontr) assume lka: "\ l ! k \ a" + then have \l ! k = a\ + by simp show False proof (cases "k < j") let ?xs = "take (k+1) l" let ?ys = "drop (k+1) l" case True - then have "length (filter ((=) a) ?xs) > 0" - by (metis (full_types, opaque_lifting) add.commute assm discrete filter_empty_conv length_greater_0_conv length_take less_add_one lka min.absorb2 nth_mem nth_take) + then have "length (filter ((=) a) ?xs) > 0" + using \k < length l\ \l ! k = a\ by (auto simp add: filter_empty_conv in_set_conv_nth) moreover have "length (filter ((=) a) ?ys) > 0" proof - have "?ys ! (j - (k+1)) = l ! j" using True assm by auto moreover have "j - (k+1) < length ?ys" using True \j < length l\ by auto ultimately show ?thesis by (metis (full_types) \l ! j = a\ filter_empty_conv length_greater_0_conv nth_mem) qed moreover have "?xs @ ?ys = l" using append_take_drop_id by blast ultimately have "length (filter ((=) a) l) > 1" by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 asm filter_append length_append less_add_eq_less less_one nat_neq_iff) then show False using asm by simp next let ?xs = "take (j+1) l" let ?ys = "drop (j+1) l" case False then have "length (filter ((=) a) ?xs) > 0" - by (metis (full_types, opaque_lifting) add.commute j discrete filter_empty_conv length_greater_0_conv length_take less_add_one min.absorb2 nth_mem nth_take) + using \k < length l\ \l ! j = a\ by (auto simp add: filter_empty_conv in_set_conv_nth) moreover have "length (filter ((=) a) ?ys) > 0" proof - have "?ys ! (k - (j+1)) = l ! k" using False assm by auto moreover have "k - (j+1) < length ?ys" using False assm by auto ultimately show ?thesis by (metis (full_types) filter_empty_conv length_greater_0_conv lka nth_mem) qed moreover have "?xs @ ?ys = l" using append_take_drop_id by blast ultimately have "length (filter ((=) a) l) > 1" by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 asm filter_append length_append less_add_eq_less less_one nat_neq_iff) then show False using asm by simp qed qed qed ultimately show "\!j. j < length l \ l ! j = a" by blast qed end diff --git a/thys/Combinatorics_Words/Periodicity_Lemma.thy b/thys/Combinatorics_Words/Periodicity_Lemma.thy --- a/thys/Combinatorics_Words/Periodicity_Lemma.thy +++ b/thys/Combinatorics_Words/Periodicity_Lemma.thy @@ -1,530 +1,530 @@ (* Title: CoW/Periodicity_Lemma.thy Author: Štěpán Holub, Charles University Part of Combinatorics on Words Formalized. See https://gitlab.com/formalcow/combinatorics-on-words-formalized/ *) theory Periodicity_Lemma imports CoWBasic begin chapter "The Periodicity Lemma" text\The Periodicity Lemma says that if a sufficiently long word has two periods p and q, then the period can be refined to @{term "gcd p q"}. The consequence is equivalent to the fact that the corresponding periodic roots commute. ``Sufficiently long'' here means at least @{term "p + q - gcd p q"}. It is also known as the Fine and Wilf theorem due to its authors @{cite FineWilf}.\ text\ If we relax the requirement to @{term "p + q"}, then the claim becomes easy, and it is proved in @{theory Combinatorics_Words.CoWBasic} as @{term two_pers_root}: @{thm[names_long] two_pers_root[no_vars]}. \ theorem per_lemma_relaxed: assumes "period w p" and "period w q" and "p + q \ \<^bold>|w\<^bold>|" shows "(take p w)\(take q w) = (take q w)\(take p w)" using two_pers_root[OF \period w p\[unfolded period_def[of w p]] \period w q\[unfolded period_def[of w q]], unfolded take_len[OF add_leD1[OF \p + q \ \<^bold>|w\<^bold>|\]] take_len[OF add_leD2[OF \p + q \ \<^bold>|w\<^bold>|\]], OF \p + q \ \<^bold>|w\<^bold>|\]. text\Also in terms of the numeric period:\ thm two_periods section \Main claim\ text\We first formulate the claim of the Periodicity lemma in terms of commutation of two periodic roots. For trivial reasons we can also drop the requirement that the roots are nonempty. \ theorem per_lemma_comm: assumes "w \p r \ w" and "w \p s \ w" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - (gcd \<^bold>|r\<^bold>| \<^bold>|s\<^bold>|) \ \<^bold>|w\<^bold>|" shows "r \ s = s \ r" using assms proof (induction "\<^bold>|s\<^bold>| + \<^bold>|s\<^bold>| + \<^bold>|r\<^bold>|" arbitrary: w r s rule: less_induct) case less consider (empty) "s = \" | (short) "\<^bold>|r\<^bold>| < \<^bold>|s\<^bold>|" | (step) "s \ \ \ \<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|" by force then show ?case proof (cases) case (empty) thus "r \ s = s \ r" by fastforce next case (short) thus "r \ s = s \ r" using "less.hyps"[OF _ \ w \p s \ w\ \ w \p r \ w\ \\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - (gcd \<^bold>|r\<^bold>| \<^bold>|s\<^bold>|) \ \<^bold>|w\<^bold>|\[unfolded gcd.commute[of "\<^bold>|r\<^bold>|"] add.commute[of "\<^bold>|r\<^bold>|"]]] by fastforce next case (step) hence "s \ \" and "\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|" by blast+ from le_add_diff[OF gcd_le2_nat[OF \s \ \\[folded length_0_conv], of "\<^bold>|r\<^bold>|"], unfolded gcd.commute[of "\<^bold>|r\<^bold>|"], of "\<^bold>|r\<^bold>|"] have "\<^bold>|r\<^bold>| \ \<^bold>|w\<^bold>|" using \\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - (gcd \<^bold>|r\<^bold>| \<^bold>|s\<^bold>|) \ \<^bold>|w\<^bold>|\[unfolded gcd.commute[of "\<^bold>|r\<^bold>|"] add.commute[of "\<^bold>|r\<^bold>|"]] order.trans by blast hence "\<^bold>|s\<^bold>| \ \<^bold>|w\<^bold>|" using \\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|\ order.trans by blast from pref_prod_long[OF \w \p s \ w\ this] have "s \p w". obtain w' where "s \ w' = w" and "\<^bold>|w'\<^bold>| < \<^bold>|w\<^bold>|" using \s \ \\ \s \p w\[unfolded prefix_def] by force have "w' \p w" using \w \p s \ w\ unfolding \s \ w' = w\[symmetric] pref_cancel_conv. from this[folded \s \ w' = w\] have "w' \p s\w'". have "s \p r" using pref_prod_le[OF prefix_order.trans[OF \s \p w\ \w \p r \ w\] \\<^bold>|s\<^bold>| \ \<^bold>|r\<^bold>|\]. hence "w' \p (s\\<^sup>>r) \ w'" using \w \p r \ w\ \s \ w' = w\ pref_prod_pref[OF _ \w' \p w\, of "s\\<^sup>>r"] unfolding prefix_def by fastforce have ind_len: "\<^bold>|s\\<^sup>>r\<^bold>| + \<^bold>|s\<^bold>| - (gcd \<^bold>|s\\<^sup>>r\<^bold>| \<^bold>|s\<^bold>|) \ \<^bold>|w'\<^bold>|" using \\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - (gcd \<^bold>|r\<^bold>| \<^bold>|s\<^bold>|) \ \<^bold>|w\<^bold>|\[folded \s \ w' = w\] unfolding pref_gcd_lq[OF \s \p r\, unfolded gcd.commute[of "\<^bold>|s\<^bold>|"]] lenmorph lq_short_len[OF \s \p r\, unfolded add.commute[of "\<^bold>|s\<^bold>|"]] by force have "s \ s\\<^sup>>r = s\\<^sup>>r \ s" using "less.hyps"[OF _ \w' \p (s\\<^sup>>r) \ w'\ \w' \p s \ w'\ ind_len] \s \p r\ \\<^bold>|w'\<^bold>| < \<^bold>|w\<^bold>|\ unfolding prefix_def by force thus "r \ s = s \ r" using \s \p r\ by (fastforce simp add: prefix_def) qed qed lemma per_lemma_comm_pref: assumes "u \p r\<^sup>@k" "u \p s\<^sup>@l" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|) \ \<^bold>|u\<^bold>|" shows "r \ s = s \ r" using pref_prod_root[OF assms(2)] pref_prod_root[OF assms(1)] per_lemma_comm[OF _ _ len] by blast text\We can now prove the numeric version.\ theorem per_lemma: assumes "period w p" and "period w q" and len: "p + q - gcd p q \ \<^bold>|w\<^bold>|" shows "period w (gcd p q)" proof- have takep: "w \p (take p w) \ w" and takeq: "w \p (take q w) \ w" using \period w p\ \period w q\ per_pref by blast+ have "p \ \<^bold>|w\<^bold>|" using per_lemma_len_le[OF len] per_not_zero[OF \period w q\]. have lenp: "\<^bold>|take p w\<^bold>| = p" using gcd_le2_pos[OF per_not_zero[OF \period w q\], of p] len take_len by auto have lenq: "\<^bold>|take q w\<^bold>| = q" using gcd_le1_pos[OF per_not_zero[OF \period w p\], of q] len take_len by simp obtain t where "take p w \ t*" and "take q w \ t*" using comm_rootE[OF per_lemma_comm[OF takep takeq, unfolded lenp lenq, OF len], of thesis] by blast have "w

w" using \period w p\[unfolded period_def, THEN per_root_trans, OF \take p w \ t*\]. with per_nemp[OF \period w q\] have "period w \<^bold>|t\<^bold>|" by (rule periodI) have "\<^bold>|t\<^bold>| dvd (gcd p q)" using common_root_len_gcd[OF \take p w \ t*\ \take q w \ t*\] unfolding lenp lenq. from dvd_div_mult_self[OF this] have "gcd p q div \<^bold>|t\<^bold>| * \<^bold>|t\<^bold>| = gcd p q". have "gcd p q \ 0" using \period w p\ by auto from this[folded dvd_div_eq_0_iff[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\]] show "period w (gcd p q)" using per_mult[OF \period w \<^bold>|t\<^bold>|\, of "gcd p q div \<^bold>|t\<^bold>|", unfolded dvd_div_mult_self[OF \\<^bold>|t\<^bold>| dvd (gcd p q)\]] by blast qed section \Optimality\ text\@{term "FW_word"} (where FW stands for Fine and Wilf) yields a word which show the optimality of the bound in the Periodicity lemma. Moreover, the obtained word has maximum possible letters (each equality of letters is forced by periods). The latter is not proved here.\ term "butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q)))\[gcd p q]\(butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q))))" \ \an auxiliary claim\ lemma ext_per_sum: assumes "period w p" and "period w q" and "p \ \<^bold>|w\<^bold>|" shows "period ((take p w) \ w) (p+q)" proof- have nemp: "take p w \ take q w \ \" using \period w p\ by auto have "take (p + q) (take p w \ w) = take p (take p w \ w) \ take q (drop p (take p w \ w))" using take_add by blast also have "... = take p w \ take q w" by (simp add: \p \ \<^bold>|w\<^bold>|\) ultimately have sum: "take (p + q) (take p w \ w) = take p w \ take q w" by presburger note assms[unfolded period_def] show ?thesis unfolding period_def sum rassoc using pref_spref_prolong[OF self_pref spref_spref_prolong[OF \w

w\ \w

w\]]. qed definition "fw_p_per p q \ butlast ([0..<(gcd p q)]\<^sup>@(p div (gcd p q)))" definition "fw_base p q \ fw_p_per p q \ [gcd p q]\ fw_p_per p q" fun FW_word :: "nat \ nat \ nat list" where FW_word_def: "FW_word p q = \\symmetry\ (if q < p then FW_word q p else \\artificial value\ if p = 0 then \ else \\artificial value\ if p = q then \ else \\base case\ if gcd p q = q - p then fw_base p q \\step\ else (take p (FW_word p (q-p))) \ FW_word p (q-p))" lemma FW_sym: "FW_word p q = FW_word q p" by (cases rule: linorder_cases[of p q]) simp+ theorem fw_word': "\ p dvd q \ \ q dvd p \ \<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1 \ period (FW_word p q) p \ period (FW_word p q) q \ \ period (FW_word p q) (gcd p q)" proof (induction "p + p + q" arbitrary: p q rule: less_induct) case less have "p \ 0" using \\ q dvd p\ dvd_0_right[of q] by meson have "p \ q" using \\ p dvd q\ by auto then consider "q < p" | "p < q" by linarith then show ?case proof (cases) assume "q < p" have less: "q + q + p < p + p + q" by (simp add: \q < p\) thus ?case using "less.hyps"[OF _ \\ q dvd p\ \\ p dvd q\] unfolding FW_sym[of p q] gcd.commute[of p q] add.commute[of p q] by blast next let ?d = "gcd p q" let ?dw = "[0..<(gcd p q)]" let ?pd = "p div (gcd p q)" assume "p < q" thus ?thesis proof (cases "?d = q - p") assume "?d = q - p" hence "p + ?d = q" using \p < q\ by auto hence "p \ q" and "\ q < p" using \p \ 0\ \p < q\ by fastforce+ hence fw: "FW_word p q = fw_base p q" unfolding FW_word_def[of p q] using \p \ 0\ \gcd p q = q - p\ by presburger have "\<^bold>|[0..| = gcd p q" by simp hence *: "p div gcd p q * \<^bold>|[0..| = p" by fastforce have ppref: "\<^bold>|butlast (?dw\<^sup>@?pd)\[?d]\<^bold>| = p" using \p \ 0\ unfolding lenmorph pow_len length_butlast sing_len * by fastforce note ppref' = this[unfolded lenmorph] have qpref: "\<^bold>|butlast (?dw\<^sup>@?pd)\[?d]\?dw\<^bold>| = q" unfolding lassoc lenmorph ppref' using \p + gcd p q = q\ by simp have "butlast (?dw\<^sup>@?pd)\[?d] \p FW_word p q" unfolding fw fw_base_def fw_p_per_def lassoc using triv_pref. from pref_take[OF this] have takep: "take p (FW_word p q) = butlast (?dw\<^sup>@?pd)\[?d]" unfolding ppref. have "?dw \ \" and "\<^bold>|?dw\<^bold>| = ?d" using \p \ 0\ by auto have "?pd \ 0" by (simp add: \p \ 0\ dvd_div_eq_0_iff) from not0_implies_Suc[OF this] obtain e where "?pd = Suc e" by blast have "gcd p q \ p" using \\ p dvd q\ gcd_dvd2[of p q] by force hence "Suc e \ 1" using dvd_mult_div_cancel[OF gcd_dvd1[of p q], unfolded \?pd = Suc e\] by force hence "e \ 0" by simp have "[0..@ e \ \" using \[0.. \\ \e \ 0\ nonzero_pow_emp by blast hence but_dec: "butlast (?dw\<^sup>@?pd) = ?dw \ butlast (?dw\<^sup>@e)" unfolding \?pd = Suc e\ pow_Suc butlast_append if_not_P[OF \[0..@ e \ \\] by blast have but_dec_b: "butlast (?dw\<^sup>@?pd) = ?dw\<^sup>@e \ butlast ?dw" unfolding \?pd = Suc e\ pow_Suc' butlast_append if_not_P[OF \?dw \ \\] by blast have "butlast (?dw\<^sup>@?pd)\[?d]\?dw \p FW_word p q" unfolding fw but_dec lassoc fw_base_def fw_p_per_def by blast note takeq = pref_take[OF this, unfolded qpref] have "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" proof- have "p + q - (q - p) = p + p" using \p + gcd p q = q\ by auto hence "\<^bold>|?dw\<^sup>@?pd\<^bold>| = p" unfolding pow_len \\<^bold>|[0..| = gcd p q\ by force hence "\<^bold>|butlast (?dw\<^sup>@?pd)\<^bold>| = p - 1" unfolding length_butlast by argo hence "\<^bold>|FW_word p q\<^bold>| = (p - 1) + 1 + (p - 1)" unfolding fw lenmorph sing_len fw_base_def fw_p_per_def by presburger thus "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" unfolding \gcd p q = q - p\ \p + q - (q - p) = p + p\ using \p \ 0\ by fastforce qed have "period (FW_word p q) p" unfolding period_def proof (rule per_rootI) show "take p (FW_word p q) \ \" using \p \ 0\ unfolding length_0_conv[symmetric] ppref[folded takep]. have "fw_base p q \p fw_p_per p q \ [gcd p q] \ fw_base p q" unfolding rassoc pref_cancel_conv fw_base_def fw_p_per_def by blast thus "FW_word p q \p take p (FW_word p q) \ FW_word p q" unfolding fw rassoc fw_p_per_def takep[unfolded fw]. qed have "period (FW_word p q) q" unfolding period_def proof (rule per_rootI) show "take q (FW_word p q) \ \" unfolding length_0_conv[symmetric] qpref[folded takeq] using \p \ 0\ \p < q\ by linarith have "butlast ([0..@ (p div gcd p q)) \p [0.. butlast ([0..@ (p div gcd p q))" using pref_prod_root[OF prefixeq_butlast[of "[0..@ (p div gcd p q)"]]. from pref_ext[OF this, unfolded rassoc] have "fw_base p q \p fw_p_per p q \ [gcd p q] \ [0.. fw_base p q" unfolding rassoc pref_cancel_conv fw_base_def fw_p_per_def. thus "FW_word p q \p take q (FW_word p q) \ FW_word p q" unfolding fw rassoc fw_p_per_def takeq[unfolded fw]. qed have "\ period (FW_word p q) ?d" proof- have last_a: "last (take p (FW_word p q)) = ?d" unfolding takep nth_append_length[of "butlast (?dw \<^sup>@ ?pd)" "?d" \] last_snoc by blast have "?dw \p FW_word p q" unfolding fw but_dec rassoc fw_base_def fw_p_per_def by blast from pref_take[OF this, unfolded \\<^bold>|?dw\<^bold>| = ?d\] have takegcd: "take (gcd p q) (FW_word p q) = [0..@e \ butlast ([0.. [?d] \ (butlast ([0..@(p div gcd p q)))" unfolding fw but_dec_b rassoc fw_base_def fw_p_per_def .. have gcdepref: "[0..@ Suc e \p take (gcd p q) (FW_word p q) \ FW_word p q" unfolding takegcd pow_Suc pref_cancel_conv unfolding fw_dec_b by blast have "\<^bold>|[0..@ Suc e\<^bold>| = p" unfolding pow_len \\<^bold>|?dw\<^bold>| = ?d\ \?pd = Suc e\[symmetric] using dvd_div_mult_self[OF gcd_dvd1]. from pref_take[OF gcdepref, unfolded this] have takegcdp: "take p (take (gcd p q) (FW_word p q) \ (FW_word p q)) = [0..@e \ [0..p \ 0\) from last_upt[OF this] have last_b: "last (take p (take (gcd p q) (FW_word p q) \ (FW_word p q))) = gcd p q - 1" unfolding takegcdp last_appendR[of "[0..@e", OF \[0.. \\]. have "p \ \<^bold>|FW_word p q\<^bold>|" unfolding \\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1\ \gcd p q = q - p\ using \p < q\ by auto have "gcd p q \ gcd p q - 1" using \gcd p q = q - p\ \p < q\ by linarith hence "take p (FW_word p q) \ take p (take (gcd p q) (FW_word p q) \ (FW_word p q))" unfolding last_b[symmetric] unfolding last_a[symmetric] using arg_cong[of _ _ last] by blast hence "\ FW_word p q \p take (gcd p q) (FW_word p q) \ FW_word p q " using pref_share_take[OF _ \p \ \<^bold>|FW_word p q\<^bold>|\, of "take (gcd p q) (FW_word p q) \ FW_word p q"] by blast thus "\ period (FW_word p q) (gcd p q)" unfolding period_def by blast qed show ?thesis using \period (FW_word p q) p\ \period (FW_word p q) q\ \\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1\ \\ period (FW_word p q) (gcd p q)\ by blast next let ?d' = "gcd p (q-p)" assume "gcd p q \ q - p" hence fw: "FW_word p q = (take p (FW_word p (q-p))) \ FW_word p (q-p)" using FW_word_def \p \ 0\ \p \ q\ \p < q\ by (meson less_Suc_eq not_less_eq) have divhyp1: "\ p dvd q - p" using \\ p dvd q\ \p < q\ dvd_minus_self by auto have divhyp2: "\ q - p dvd p" proof (rule notI) assume "q - p dvd p" have "q = p + (q - p)" by (simp add: \p < q\ less_or_eq_imp_le) from gcd_add2[of p "q - p", folded this, unfolded gcd_nat.absorb2[of "q - p" p, OF \q - p dvd p\]] show "False" using \gcd p q \ q - p\ by blast qed have lenhyp: "p + p + (q - p) < p + p + q" using \p < q\ \p \ 0\ by linarith \ \induction assumption\ have "\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1" and "period (FW_word p (q-p)) p" and "period (FW_word p (q-p)) (q-p)" and "\ period (FW_word p (q-p)) (gcd p (q-p))" using "less.hyps"[OF _ divhyp1 divhyp2] lenhyp by blast+ \ \auxiliary facts\ have "p + (q - p) = q" using divhyp1 dvd_minus_self by auto have "?d = ?d'" using gcd_add2[of p "q-p", unfolded le_add_diff_inverse[OF less_imp_le[OF \p < q\]]]. have "?d \ q" using \\ q dvd p\ gcd_dvd2[of q p, unfolded gcd.commute[of q]] by force from this[unfolded nat_neq_iff] have "?d < q" using gr_implies_not0 \p < q\ nat_dvd_not_less by blast hence "1 \ q - ?d" by linarith have "?d' < q - p" using gcd_le2_pos[OF per_not_zero[OF \period (FW_word p (q - p)) (q - p)\], of p] divhyp2[unfolded gcd_nat.absorb_iff2] nat_less_le by blast hence "p \ \<^bold>|(FW_word p (q - p))\<^bold>|" - unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ diff_diff_left discrete by linarith + unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ by linarith have "FW_word p (q - p) \ \" unfolding length_0_conv[symmetric] using \p \ \<^bold>|FW_word p (q - p)\<^bold>|\ \p \ 0\[folded le_zero_eq] by linarith \ \claim 1\ have "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" proof- have "p + (q - p) = q" using less_imp_le[OF \p < q\] by fastforce have "\<^bold>|FW_word p q\<^bold>| = \<^bold>|take p (FW_word p (q - p))\<^bold>| + \<^bold>|FW_word p (q - p)\<^bold>|" using fw lenmorph[of "take p (FW_word p (q - p))" "FW_word p (q - p)"] by presburger also have "... = p + (p + (q - p) - ?d' - 1)" unfolding \\<^bold>|FW_word p (q - p)\<^bold>| = p + (q - p) - ?d' - 1\ take_len[OF \p \ \<^bold>|FW_word p (q - p)\<^bold>|\] by blast also have "... = p + (q - ?d - 1)" unfolding \?d = ?d'\ \p + (q - p) = q\.. also have "... = p + (q - ?d) - 1" using Nat.add_diff_assoc[OF \1 \ q - ?d\]. also have "... = p + q - ?d - 1" by (simp add: \?d < q\ less_or_eq_imp_le) finally show "\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1" by presburger qed \ \claim 2\ have "period (FW_word p q) p" using fw ext_per_left[OF \period (FW_word p (q-p)) p\ \p \ \<^bold>|FW_word p (q - p)\<^bold>|\] by presburger \ \claim 3\ have "period (FW_word p q) q" using ext_per_sum[OF \period (FW_word p (q - p)) p\ \period (FW_word p (q - p)) (q - p)\ \p \ \<^bold>|FW_word p (q - p)\<^bold>|\, folded fw, unfolded \p + (q-p) = q\]. \ \claim 4\ have "\ period (FW_word p q) ?d" using \\ period (FW_word p (q -p)) (gcd p (q-p))\ unfolding \?d = ?d'\[symmetric] using period_fac[of "take p (FW_word p (q - p))" "FW_word p (q - p)" \ "?d", unfolded append_Nil2, OF _ \FW_word p (q - p) \ \\, folded fw] by blast thus ?thesis using \period (FW_word p q) p\ \period (FW_word p q) q\ \\<^bold>|FW_word p q\<^bold>| = p + q - ?d - 1\ by blast qed qed qed theorem fw_word: assumes "\ p dvd q" "\ q dvd p" shows "\<^bold>|FW_word p q\<^bold>| = p + q - gcd p q - 1" and "period (FW_word p q) p" and "period (FW_word p q) q" and "\ period (FW_word p q) (gcd p q)" using fw_word'[OF assms] by blast+ text\Calculation examples\ section "Other variants of the periodicity lemma" text \Periodicity lemma is one of the most frequent tools in Combinatorics on words. Here are some useful variants.\ text\Note that the following lemmas are stronger versions of @{thm per_lemma_pref_suf fac_two_conjug_primroot fac_two_conjug_primroot' fac_two_conjug_primroot'' fac_two_prim_conjug} that have a relaxed length assumption @{term "\<^bold>|p\<^bold>| + \<^bold>|q\<^bold>| \ \<^bold>|w\<^bold>|"} instead of @{term "\<^bold>|p\<^bold>| + \<^bold>|q\<^bold>| - (gcd \<^bold>|p\<^bold>| \<^bold>|q\<^bold>|) \ \<^bold>|w\<^bold>|"} (and which follow from the relaxed version of periodicity lemma @{thm two_pers}.\ lemma per_lemma_pref_suf_gcd: assumes "w

w" and "w q" and fw: "\<^bold>|p\<^bold>| + \<^bold>|q\<^bold>| - (gcd \<^bold>|p\<^bold>| \<^bold>|q\<^bold>|) \ \<^bold>|w\<^bold>|" obtains r s k l m where "p = (r \ s)\<^sup>@k" and "q = (s \ r)\<^sup>@l" and "w = (r \ s)\<^sup>@m \ r" and "primitive (r\s)" proof- let ?q = "(w \ q)\<^sup><\w" have "w

w" using ssufD1[OF \w q\] rq_suf[symmetric, THEN per_rootI[OF prefI rq_ssuf[OF \w q\]]] by argo have "q \ ?q" by (meson assms(2) conjugI1 conjug_sym rq_suf suffix_order.less_imp_le) have nemps': "p \ \" "?q \ \" using assms(1) \w

w\ by fastforce+ have "\<^bold>|p\<^bold>| + \<^bold>|?q\<^bold>| - gcd (\<^bold>|p\<^bold>|) (\<^bold>|?q\<^bold>|) \ \<^bold>|w\<^bold>|" using fw unfolding conjug_len[OF \q \ ?q\]. from per_lemma_comm[OF sprefD1[OF \w

w\] sprefD1[OF \w

w\] this] have "p \ ?q = ?q \ p". then have "\ p = \ ?q" using comm_primroots[OF nemps'] by force hence [symmetric]: "\ q \ \ p" using conjug_primroot[OF \q \ (w \ q)\<^sup><\w\] by argo from conjug_primrootsE[OF this] obtain r s k l where "p = (r \ s) \<^sup>@ k" and "q = (s \ r) \<^sup>@ l" and "primitive (r \ s)". have "w \p (r\s)\w" using assms per_root_drop_exp sprefD1 \p = (r \ s) \<^sup>@ k\ by meson have "w \s w\(s\r)" using assms(2) per_root_drop_exp[reversed] ssufD1 \q = (s \ r) \<^sup>@ l\ by meson have "\<^bold>|r \ s\<^bold>| \ \<^bold>|w\<^bold>|" using conjug_nemp_iff[OF \q \ ?q\] dual_order.trans length_0_conv nemps' per_lemma_len_le[OF fw] primroot_len_le[OF nemps'(1)] unfolding primroot_unique[OF nemps'(1) \primitive (r \ s)\ \p = (r \ s) \<^sup>@ k\] by blast from root_suf_conjug[OF \primitive (r \ s)\ \w \p (r\s)\w\ \w \s w\(s\r)\ this] obtain m where "w = (r \ s) \<^sup>@ m \ r". from that[OF \p = (r \ s) \<^sup>@ k\ \q = (s \ r) \<^sup>@ l\ this \primitive (r \ s)\] show ?thesis. qed lemma fac_two_conjug_primroot_gcd: assumes facs: "w \f p\<^sup>@k" "w \f q\<^sup>@l" and nemps: "p \ \" "q \ \" and len: "\<^bold>|p\<^bold>| + \<^bold>|q\<^bold>| - gcd (\<^bold>|p\<^bold>|) (\<^bold>|q\<^bold>|) \ \<^bold>|w\<^bold>|" obtains r s m where "\ p \ r \ s" and "\ q \ r \ s" and "w = (r \ s)\<^sup>@m \ r" and "primitive (r\s)" proof - obtain p' where "w

w" "p \ p'" "p' \ \" using conjug_nemp_iff fac_pow_pref_conjug[OF facs(1)] nemps(1) per_rootI' by metis obtain q' where "w q'" "q \ q'" "q' \ \" using fac_pow_pref_conjug[reversed, OF \w \f q\<^sup>@l\] conjug_nemp_iff nemps(2) per_rootI'[reversed] by metis from per_lemma_pref_suf_gcd[OF \w

w\ \w q'\] obtain r s k l m where "p' = (r \ s) \<^sup>@ k" and "q' = (s \ r) \<^sup>@ l" and "w = (r \ s) \<^sup>@ m \ r" and "primitive (r \ s)" using len[unfolded conjug_len[OF \p \ p'\] conjug_len[OF \q \ q'\]] by blast moreover have "\ p' = r\s" using \p' = (r \ s) \<^sup>@ k\ \primitive (r \ s)\ \p' \ \\ primroot_unique by blast hence "\ p \ r\s" using conjug_primroot[OF \p \ p'\] by simp moreover have "\ q' = s\r" using \q' = (s \ r) \<^sup>@ l\ \primitive (r \ s)\[unfolded conjug_prim_iff'[of r]] \q' \ \\ primroot_unique by blast hence "\ q \ s\r" using conjug_primroot[OF \q \ q'\] by simp hence "\ q \ r\s" using conjug_trans[OF _ conjugI'] by meson ultimately show ?thesis using that by blast qed corollary fac_two_conjug_primroot'_gcd: assumes facs: "u \f r\<^sup>@k" "u \f s\<^sup>@l" and nemps: "r \ \" "s \ \" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|) \ \<^bold>|u\<^bold>|" shows "\ r \ \ s" using fac_two_conjug_primroot_gcd[OF assms] conjug_trans[OF _ conjug_sym[of "\ s"]]. lemma fac_two_conjug_primroot''_gcd: assumes facs: "u \f r\<^sup>@k" "u \f s\<^sup>@l" and "u \ \" and len: "\<^bold>|r\<^bold>| + \<^bold>|s\<^bold>| - gcd (\<^bold>|r\<^bold>|) (\<^bold>|s\<^bold>|) \ \<^bold>|u\<^bold>|" shows "\ r \ \ s" proof - have nemps: "r \ \" "s \ \" using facs \u \ \\ by auto show "conjugate (\ r) (\ s)" using fac_two_conjug_primroot'_gcd[OF facs nemps len]. qed lemma fac_two_prim_conjug_gcd: assumes "w \f u\<^sup>@n" "w \f v\<^sup>@m" "primitive u" "primitive v" "\<^bold>|u\<^bold>| + \<^bold>|v\<^bold>| - gcd (\<^bold>|u\<^bold>|) (\<^bold>|v\<^bold>|) \ \<^bold>|w\<^bold>|" shows "u \ v" using fac_two_conjug_primroot'_gcd[OF assms(1-2) _ _ assms(5)] prim_nemp[OF \primitive u\] prim_nemp[OF \primitive v\] unfolding prim_self_root[OF \primitive u\] prim_self_root[OF \primitive v\]. lemma two_pers_1: assumes pu: "w \p u \ w" and pv: "w \p v \ w" and len: "\<^bold>|u\<^bold>| + \<^bold>|v\<^bold>| - 1 \ \<^bold>|w\<^bold>|" shows "u \ v = v \ u" proof assume "u \ \" "v \ \" hence "1 \ gcd \<^bold>|u\<^bold>| \<^bold>|v\<^bold>|" using nemp_len by (simp add: Suc_leI) thus ?thesis using per_lemma_comm[OF pu pv] len by linarith qed end diff --git a/thys/CommCSL/CommCSL.thy b/thys/CommCSL/CommCSL.thy --- a/thys/CommCSL/CommCSL.thy +++ b/thys/CommCSL/CommCSL.thy @@ -1,781 +1,782 @@ section \CommCSL\ theory CommCSL imports Lang StateModel begin definition no_guard :: "('i, 'a) heap \ bool" where "no_guard h \ get_gs h = None \ (\k. get_gu h k = None)" typedef 'a precondition = "{ pre :: ('a \ 'a \ bool) |pre. \a b. pre a b \ (pre b a \ pre a a) }" using Sup2_E by auto lemma charact_rep_prec: assumes "Rep_precondition pre a b" shows "Rep_precondition pre b a \ Rep_precondition pre a a" using Rep_precondition assms by fastforce typedef ('i, 'a) indexed_precondition = "{ pre :: ('i \ 'a \ 'a \ bool) |pre. \a b k. pre k a b \ (pre k b a \ pre k a a) }" using Sup2_E by auto lemma charact_rep_indexed_prec: assumes "Rep_indexed_precondition pre k a b" shows "Rep_indexed_precondition pre k b a \ Rep_indexed_precondition pre k a a" by (metis (no_types, lifting) Abs_indexed_precondition_cases Rep_indexed_precondition_cases Rep_indexed_precondition_inverse assms mem_Collect_eq) type_synonym 'a list_exp = "store \ 'a list" subsection \Assertion Language\ datatype ('i, 'a, 'v) assertion = Bool bexp | Emp | And "('i, 'a, 'v) assertion" "('i, 'a, 'v) assertion" | Star "('i, 'a, 'v) assertion" "('i, 'a, 'v) assertion" ("_ * _" 70) | Low bexp | LowExp exp | PointsTo exp prat exp | Exists var "('i, 'a, 'v) assertion" | EmptyFullGuards | PreSharedGuards "'a precondition" | PreUniqueGuards "('i, 'a) indexed_precondition" | View "normal_heap \ 'v" "('i, 'a, 'v) assertion" "store \ 'v" | SharedGuard prat "store \ 'a multiset" | UniqueGuard 'i "'a list_exp" | Imp bexp "('i, 'a, 'v) assertion" | NoGuard inductive PRE_shared_simpler :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where Empty: "PRE_shared_simpler spre {#} {#}" | Step: "\ PRE_shared_simpler spre a b ; spre xa xb \ \ PRE_shared_simpler spre (a + {# xa #}) (b + {# xb #})" definition PRE_unique :: "('b \ 'b \ bool) \ 'b list \ 'b list \ bool" where "PRE_unique upre uargs uargs' \ length uargs = length uargs' \ (\i. i \ 0 \ i < length uargs' \ upre (uargs ! i) (uargs' ! i))" fun hyper_sat :: "(store \ ('i, 'a) heap) \ (store \ ('i, 'a) heap) \ ('i, 'a, nat) assertion \ bool" ("_, _ \ _" [51, 65, 66] 50) where "(s, _), (s', _) \ Bool b \ bdenot b s \ bdenot b s'" | "(_, h), (_, h') \ Emp \ dom (get_fh h) = {} \ dom (get_fh h') = {}" | "\, \' \ And A B \ \, \' \ A \ \, \' \ B" | "(s, h), (s', h') \ Star A B \ (\h1 h2 h1' h2'. Some h = Some h1 \ Some h2 \ Some h' = Some h1' \ Some h2' \ (s, h1), (s', h1') \ A \ (s, h2), (s', h2') \ B)" | "(s, h), (s', h') \ Low e \ bdenot e s = bdenot e s'" | "(s, h), (s', h') \ PointsTo loc p x \ get_fh h (edenot loc s) = Some (p, edenot x s) \ get_fh h' (edenot loc s') = Some (p, edenot x s') \ dom (get_fh h) = {edenot loc s} \ dom (get_fh h') = {edenot loc s'}" | "(s, h), (s', h') \ Exists x A \ (\v v'. (s(x := v), h), (s'(x := v'), h') \ A)" | "(s, h), (s', h') \ EmptyFullGuards \ (get_gs h = Some (pwrite, {#}) \ (\k. get_gu h k = Some []) \ get_gs h' = Some (pwrite, {#}) \ (\k. get_gu h' k = Some []))" | "(s, h), (s', h') \ PreSharedGuards spre \ (\sargs sargs'. get_gs h = Some (pwrite, sargs) \ get_gs h' = Some (pwrite, sargs') \ PRE_shared_simpler (Rep_precondition spre) sargs sargs' \ get_fh h = Map.empty \ get_fh h' = Map.empty)" | "(s, h), (s', h') \ PreUniqueGuards upre \ (\uargs uargs'. (\k. get_gu h k = Some (uargs k)) \ (\k. get_gu h' k = Some (uargs' k)) \ (\k. PRE_unique (Rep_indexed_precondition upre k) (uargs k) (uargs' k)) \ get_fh h = Map.empty \ get_fh h' = Map.empty)" | "(s, h), (s', h') \ View f J e \ ((s, h), (s', h') \ J \ f (normalize (get_fh h)) = e s \ f (normalize (get_fh h')) = e s')" | "(s, h), (s', h') \ SharedGuard \ ms \ ((\k. get_gu h k = None \ get_gu h' k = None) \ get_gs h = Some (\, ms s) \ get_gs h' = Some (\, ms s') \ get_fh h = Map.empty \ get_fh h' = Map.empty)" | "(s, h), (s', h') \ UniqueGuard k lexp \ (get_gs h = None \ get_gu h k = Some (lexp s) \ get_gu h' k = Some (lexp s') \ get_gs h' = None \ get_fh h = Map.empty \ get_fh h' = Map.empty \ (\k'. k' \ k \ get_gu h k' = None \ get_gu h' k' = None))" | "(s, h), (s', h') \ LowExp e \ edenot e s = edenot e s'" | "(s, h), (s', h') \ Imp b A \ bdenot b s = bdenot b s' \ (bdenot b s \ (s, h), (s', h') \ A)" | "(s, h), (s', h') \ NoGuard \ (get_gs h = None \ (\k. get_gu h k = None) \ get_gs h' = None \ (\k. get_gu h' k = None))" lemma sat_PreUniqueE: assumes "(s, h), (s', h') \ PreUniqueGuards upre" shows "\uargs uargs'. (\k. get_gu h k = Some (uargs k)) \ (\k. get_gu h' k = Some (uargs' k)) \ (\k. PRE_unique (Rep_indexed_precondition upre k) (uargs k) (uargs' k))" using assms by auto lemma decompose_heap_triple: "h = (get_fh h, get_gs h, get_gu h)" by simp definition depends_only_on :: "(store \ 'v) \ var set \ bool" where "depends_only_on e S \ (\s s'. agrees S s s' \ e s = e s')" lemma depends_only_onI: assumes "\s s' :: store. agrees S s s' \ e s = e s'" shows "depends_only_on e S" using assms depends_only_on_def by blast definition fvS :: "(store \ 'v) \ var set" where "fvS e = (SOME S. depends_only_on e S)" lemma fvSE: assumes "agrees (fvS e) s s'" shows "e s = e s'" proof - have "depends_only_on e UNIV" proof (rule depends_only_onI) fix s s' :: store assume "agrees UNIV s s'" have "s = s'" proof (rule ext) fix x :: var have "x \ UNIV" by auto then show "s x = s' x" by (meson \agrees UNIV s s'\ agrees_def) qed then show "e s = e s'" by simp qed then show ?thesis by (metis assms depends_only_on_def fvS_def someI_ex) qed fun fvA :: "('i, 'a, 'v) assertion \ var set" where "fvA (Bool b) = fvB b" | "fvA (And A B) = fvA A \ fvA B" | "fvA (Star A B) = fvA A \ fvA B" | "fvA (Low e) = fvB e" | "fvA Emp = {}" | "fvA (PointsTo v va vb) = fvE v \ fvE vb" | "fvA (Exists x A) = fvA A - {x}" | "fvA (SharedGuard _ e) = fvS e" | "fvA (UniqueGuard _ e) = fvS e" | "fvA (View view A e) = fvA A \ fvS e" | "fvA (PreSharedGuards _) = {}" | "fvA (PreUniqueGuards _) = {}" | "fvA EmptyFullGuards = {}" | "fvA (LowExp x) = fvE x" | "fvA (Imp b A) = fvB b \ fvA A" definition subS :: "var \ exp \ (store \ 'v) \ (store \ 'v)" where "subS x E e = (\s. e (s(x := edenot E s)))" lemma subS_assign: "subS x E e s \ e (s(x := edenot E s))" by (simp add: subS_def) fun collect_existentials :: "('i, 'a, nat) assertion \ var set" where "collect_existentials (And A B) = collect_existentials A \ collect_existentials B" | "collect_existentials (Star A B) = collect_existentials A \ collect_existentials B" | "collect_existentials (Exists x A) = collect_existentials A \ {x}" | "collect_existentials (View view A e) = collect_existentials A" | "collect_existentials (Imp _ A) = collect_existentials A" | "collect_existentials _ = {}" fun subA :: "var \ exp \ ('i, 'a, nat) assertion \ ('i, 'a, nat) assertion" where "subA x E (And A B) = And (subA x E A) (subA x E B)" | "subA x E (Star A B) = Star (subA x E A) (subA x E B)" | "subA x E (Bool B) = Bool (subB x E B)" | "subA x E (Low e) = Low (subB x E e)" | "subA x E (LowExp e) = LowExp (subE x E e)" | "subA x E (UniqueGuard k e) = UniqueGuard k (subS x E e)" | "subA x E (SharedGuard \ e) = SharedGuard \ (subS x E e)" | "subA x E (View view A e) = View view (subA x E A) (subS x E e)" | "subA x E (PointsTo loc \ e) = PointsTo (subE x E loc) \ (subE x E e)" | "subA x E (Exists y A) = (if x = y then Exists y A else Exists y (subA x E A))" | "subA x E (Imp b A) = Imp (subB x E b) (subA x E A)" | "subA _ _ A = A" lemma subA_assign: assumes "collect_existentials A \ fvE E = {}" shows "(s, h), (s', h') \ subA x E A \ (s(x := edenot E s), h), (s'(x := edenot E s'), h') \ A" using assms proof (induct A arbitrary: s h s' h') case (And A1 A2) then show ?case by (simp add: disjoint_iff_not_equal) next case (Star A1 A2) then show ?case by (simp add: disjoint_iff_not_equal) next case (Bool x) then show ?case by (metis hyper_sat.simps(1) subA.simps(3) subB_assign) next case (Low x2) then show ?case by (metis (no_types, lifting) hyper_sat.simps(5) subA.simps(4) subB_assign) next case (LowExp x2) then show ?case by (metis (no_types, lifting) hyper_sat.simps(14) subA.simps(5) subE_assign) next case (PointsTo x1a x2 x3) then show ?case by (metis (no_types, lifting) hyper_sat.simps(6) subA.simps(9) subE_assign) next case (Exists y A) then have asm0: "collect_existentials A \ fvE E = {}" by auto show ?case (is "?A \ ?B") proof show "?A \ ?B" proof - assume ?A show ?B proof (cases "x = y") case True then show ?thesis by (metis (no_types, opaque_lifting) \?A\ fun_upd_upd hyper_sat.simps(7) subA.simps(10)) next case False then obtain v v' where "(s(y := v), h), (s'(y := v'), h') \ subA x E A" using \(s, h), (s', h') \ subA x E (Exists y A)\ by auto then have "((s(y := v))(x := edenot E (s(y := v))), h), ((s'(y := v'))(x := edenot E (s'(y := v'))), h') \ A" using Exists asm0 by blast moreover have "y \ fvE E" using Exists.prems by force then have "edenot E (s(y := v)) = edenot E s \ edenot E (s'(y := v')) = edenot E s'" by (metis agrees_update exp_agrees) moreover have "(s(y := v))(x := edenot E s) = (s(x := edenot E s))(y := v) \ (s'(y := v'))(x := edenot E s') = (s'(x := edenot E s'))(y := v')" by (simp add: False fun_upd_twist) ultimately show ?thesis using False hyper_sat.simps(7) by metis qed qed assume ?B show ?A proof (cases "x = y") case True then show ?thesis using \(s(x := edenot E s), h), (s'(x := edenot E s'), h') \ Exists y A\ by fastforce next case False then obtain v v' where "((s(x := edenot E s))(y := v), h), ((s'(x := edenot E s'))(y := v'), h') \ A" using \(s(x := edenot E s), h), (s'(x := edenot E s'), h') \ Exists y A\ hyper_sat.simps(7) by blast moreover have "(s(y := v))(x := edenot E s) = (s(x := edenot E s))(y := v) \ (s'(y := v'))(x := edenot E s') = (s'(x := edenot E s'))(y := v')" by (simp add: False fun_upd_twist) then have "((s(y := v))(x := edenot E (s(y := v))), h), ((s'(y := v'))(x := edenot E (s'(y := v'))), h') \ A" using Exists.prems calculation by auto then show ?thesis by (metis Exists.hyps False asm0 hyper_sat.simps(7) subA.simps(10)) qed qed next case (View x1a A x3) then show ?case by (metis (mono_tags, lifting) collect_existentials.simps(4) hyper_sat.simps(11) subA.simps(8) subS_def) next case (SharedGuard x1a x2) then show ?case by (metis (mono_tags, lifting) hyper_sat.simps(12) subA.simps(7) subS_def) next case (UniqueGuard x) then show ?case by (metis (mono_tags, lifting) hyper_sat.simps(13) subA.simps(6) subS_def) next case (Imp b A) then show ?case by (metis collect_existentials.simps(5) hyper_sat.simps(15) subA.simps(11) subB_assign) qed (auto) lemma PRE_uniqueI: assumes "length uargs = length uargs'" and "\i. i \ 0 \ i < length uargs' \ upre (uargs ! i) (uargs' ! i)" shows "PRE_unique upre uargs uargs'" using assms PRE_unique_def by blast lemma PRE_unique_implies_tl: assumes "PRE_unique upre (ta # qa) (tb # qb)" shows "PRE_unique upre qa qb" proof (rule PRE_uniqueI) show "length qa = length qb" by (metis PRE_unique_def assms diff_Suc_1 length_Cons) fix i assume "0 \ i \ i < length qb" then have "upre ((ta # qa) ! (i + 1)) ((tb # qb) ! (i + 1))" - by (metis PRE_unique_def add_nonneg_nonneg assms discrete le_imp_less_Suc length_Cons zero_less_one_class.zero_le_one) + using assms PRE_unique_def [of upre \ta # qa\ \tb # qb\] + by (auto simp add: less_Suc_eq_le dest: spec [of _ \Suc i\]) then show "upre (qa ! i) (qb ! i)" by simp qed lemma charact_PRE_unique: assumes "PRE_unique (Rep_indexed_precondition pre k) a b" shows "PRE_unique (Rep_indexed_precondition pre k) b a \ PRE_unique (Rep_indexed_precondition pre k) a a" using assms proof (induct "length a" arbitrary: a b) case 0 then show ?case by (simp add: PRE_unique_def) next case (Suc n) then obtain ha ta hb tb where "a = ha # ta" "b = hb # tb" by (metis One_nat_def PRE_unique_def Suc_le_length_iff le_add1 plus_1_eq_Suc) then have "n = length ta" using Suc.hyps(2) by auto then have "PRE_unique (Rep_indexed_precondition pre k) tb ta \ PRE_unique (Rep_indexed_precondition pre k) ta ta" by (metis PRE_unique_implies_tl Suc.hyps(1) Suc.prems \a = ha # ta\ \b = hb # tb\) then show ?case by (metis PRE_unique_def Suc.prems charact_rep_indexed_prec) qed lemma charact_PRE_shared_simpler: assumes "PRE_shared_simpler rpre a b" and "Rep_precondition pre = rpre" shows "PRE_shared_simpler (Rep_precondition pre) b a \ PRE_shared_simpler (Rep_precondition pre) a a" using assms proof (induct arbitrary: pre rule: PRE_shared_simpler.induct) case (Empty spre) then show ?case by (simp add: PRE_shared_simpler.Empty) next case (Step spre a b xa xb) then have "spre xb xa \ spre xa xa" using charact_rep_prec by metis then show ?case by (metis PRE_shared_simpler.Step Step.hyps(2) Step.prems) qed lemma always_sat_refl_aux: assumes "(s, h), (s', h') \ A" shows "(s, h), (s, h) \ A" using assms proof (induct A arbitrary: s h s' h') case (Star A B) then obtain ha hb ha' hb' where "Some h = Some ha \ Some hb" "Some h' = Some ha' \ Some hb'" "(s, ha), (s', ha') \ A" "(s, hb), (s', hb') \ B" by auto then have "(s, ha), (s, ha) \ A \ (s, hb), (s, hb) \ B" using Star.hyps(1) Star.hyps(2) by blast then show ?case using \Some h = Some ha \ Some hb\ hyper_sat.simps(4) by blast next case (Exists x A) then show ?case by (meson hyper_sat.simps(7)) next case (PreSharedGuards x) then show ?case using charact_PRE_shared_simpler by force next case (PreUniqueGuards upre) then obtain uargs uargs' where "(\k. get_gu h k = Some (uargs k)) \ (\k. get_gu h' k = Some (uargs' k)) \ (\k. PRE_unique (Rep_indexed_precondition upre k) (uargs k) (uargs' k)) \ get_fh h = Map.empty \ get_fh h' = Map.empty" using hyper_sat.simps(10)[of s h s' h' upre] by blast then show "(s, h), (s, h) \ PreUniqueGuards upre" using charact_PRE_unique hyper_sat.simps(10)[of s h s h upre] by metis qed (auto) lemma always_sat_refl: assumes "\, \' \ A" shows "\, \ \ A" by (metis always_sat_refl_aux assms prod.exhaust_sel) lemma agrees_same_aux: assumes "agrees (fvA A) s s''" and "(s, h), (s', h') \ A" shows "(s'', h), (s', h') \ A" using assms proof (induct A arbitrary: s s' s'' h h') case (Bool b) then show ?case by (simp add: bexp_agrees) next case (And A1 A2) then show ?case using fvA.simps(2) hyper_sat.simps(3) by (metis (mono_tags, lifting) UnCI agrees_def) next case (Star A B) then obtain ha hb ha' hb' where "Some h = Some ha \ Some hb" "Some h' = Some ha' \ Some hb'" "(s, ha), (s', ha') \ A" "(s, hb), (s', hb') \ B" by auto then have "(s'', ha), (s', ha') \ A \ (s'', hb), (s', hb') \ B" using Star.hyps[of s s'' _ s' _] Star.prems(1) by (simp add: agrees_def) then show ?case using \Some h = Some ha \ Some hb\ \Some h' = Some ha' \ Some hb'\ hyper_sat.simps(4) by blast next case (Low e) then have "bdenot e s = bdenot e s''" by (metis bexp_agrees fvA.simps(4)) then show ?case using Low by simp next case (LowExp e) then have "edenot e s = edenot e s''" by (metis exp_agrees fvA.simps(14)) then show ?case using LowExp by simp next case (PointsTo l \ v) then have "edenot l s = edenot l s'' \ edenot v s = edenot v s''" by (simp add: agrees_def exp_agrees) then show ?case using PointsTo by auto next case (Exists x A) then obtain v v' where "(s(x := v), h), (s'(x := v'), h') \ A" by auto moreover have "agrees (fvA A) (s(x := v)) (s''(x := v))" proof (rule agreesI) fix y show "y \ fvA A \ (s(x := v)) y = (s''(x := v)) y" apply (cases "y = x") apply simp using Diff_iff[of y "fvA A" "{x}"] Exists.prems(1) agrees_def empty_iff fun_upd_apply[of s x v] fun_upd_apply[of s'' x v] fvA.simps(7) insert_iff by metis qed ultimately have "(s''(x := v), h), (s'(x := v'), h') \ A" using Exists.hyps by blast then show ?case by auto next case (View x1a A e) then have "(s'', h), (s', h') \ A \ e s = e s''" using fvA.simps(10) fvSE hyper_sat.simps(11) agrees_union by metis then show ?case using View.prems(2) by auto next case (SharedGuard x1a x2) then show ?case using fvSE by auto next case (UniqueGuard x) then show ?case using fvSE by auto next case (Imp b A) then show ?case by (metis agrees_union bexp_agrees fvA.simps(15) hyper_sat.simps(15)) qed (auto) lemma agrees_same: assumes "agrees (fvA A) s s''" shows "(s, h), (s', h') \ A \ (s'', h), (s', h') \ A" by (metis (mono_tags, lifting) agrees_def agrees_same_aux assms) lemma sat_comm_aux: "(s, h), (s', h') \ A \ (s', h'), (s, h) \ A" proof (induct A arbitrary: s h s' h') case (Star A B) then obtain ha hb ha' hb' where "Some h = Some ha \ Some hb" "Some h' = Some ha' \ Some hb'" "(s, ha), (s', ha') \ A" "(s, hb), (s', hb') \ B" by auto then have "(s', ha'), (s, ha) \ A \ (s', hb'), (s, hb) \ B" using Star.hyps(1) Star.hyps(2) by presburger then show ?case using \Some h = Some ha \ Some hb\ \Some h' = Some ha' \ Some hb'\ hyper_sat.simps(4) by blast next case (Exists x A) then obtain v v' where "(s(x := v), h), (s'(x := v'), h') \ A" by auto then have "(s'(x := v'), h'), (s(x := v), h) \ A" using Exists.hyps by blast then show ?case by auto next case (PreSharedGuards x) then show ?case by (meson charact_PRE_shared_simpler hyper_sat.simps(9)) next case (PreUniqueGuards upre) then obtain uargs uargs' where "(\k. get_gu h k = Some (uargs k)) \ (\k. get_gu h' k = Some (uargs' k)) \ (\k. PRE_unique (Rep_indexed_precondition upre k) (uargs k) (uargs' k)) \ get_fh h = Map.empty \ get_fh h' = Map.empty" using hyper_sat.simps(10)[of s h s' h' upre] by blast then show "(s', h'), (s, h) \ PreUniqueGuards upre" using charact_PRE_unique hyper_sat.simps(10)[of s' h' s h upre] by metis qed (auto) lemma sat_comm: "\, \' \ A \ \', \ \ A" using sat_comm_aux surj_pair by metis definition precise where "precise J \ (\s1 H1 h1 h1' s2 H2 h2 h2'. H1 \ h1 \ H1 \ h1' \ H2 \ h2 \ H2 \ h2' \ (s1, h1), (s2, h2) \ J \ (s1, h1'), (s2, h2') \ J \ h1' = h1 \ h2' = h2)" lemma preciseI: assumes "\s1 H1 h1 h1' s2 H2 h2 h2'. H1 \ h1 \ H1 \ h1' \ H2 \ h2 \ H2 \ h2' \ (s1, h1), (s2, h2) \ J \ (s1, h1'), (s2, h2') \ J \ h1' = h1 \ h2' = h2" shows "precise J" using assms precise_def by blast lemma preciseE: assumes "precise J" and "H1 \ h1 \ H1 \ h1' \ H2 \ h2 \ H2 \ h2'" and "(s1, h1), (s2, h2) \ J \ (s1, h1'), (s2, h2') \ J" shows "h1' = h1 \ h2' = h2" using assms(1) assms(2) assms(3) precise_def by blast definition unary where "unary J \ (\s h s' h'. (s, h), (s, h) \ J \ (s', h'), (s', h') \ J \ (s, h), (s', h') \ J)" lemma unaryI: assumes "\s1 h1 s2 h2. (s1, h1), (s1, h1) \ J \ (s2, h2), (s2, h2) \ J \ (s1, h1), (s2, h2) \ J" shows "unary J" using assms unary_def by blast lemma unary_smallerI: assumes "\\1 \2. \1, \1 \ J \ \2, \2 \ J \ \1, \2 \ J" shows "unary J" using assms unary_def by blast lemma unaryE: assumes "unary J" and "(s, h), (s, h) \ J \ (s', h'), (s', h') \ J" shows "(s, h), (s', h') \ J" using assms(1) assms(2) unary_def by blast definition entails :: "('i, 'a, nat) assertion \ ('i, 'a, nat) assertion \ bool" where "entails A B \ (\\ \'. \, \' \ A \ \, \' \ B)" lemma entailsI: assumes "\x y. x, y \ A \ x, y \ B" shows "entails A B" using assms entails_def by blast lemma sat_points_to: assumes "(s, h :: ('i, 'a) heap), (s, h) \ PointsTo a \ e" shows "get_fh h = [edenot a s \ (\, edenot e s)]" proof - have "get_fh h (edenot a s) = Some (\, edenot e s) \ dom (get_fh h) = {edenot a s}" using assms by auto then show ?thesis by fastforce qed lemma unary_inv_then_view: assumes "unary J" shows "unary (View f J e)" proof (rule unaryI) fix s h s' h' assume asm0: "(s, h), (s, h) \ View f J e \ (s', h'), (s', h') \ View f J e" then show "(s, h), (s', h') \ View f J e" by (meson assms hyper_sat.simps(11) unaryE) qed lemma precise_inv_then_view: assumes "precise J" shows "precise (View f J e)" proof (rule preciseI) fix s1 H1 h1 h1' s2 H2 h2 h2' assume asm0: "H1 \ h1 \ H1 \ h1' \ H2 \ h2 \ H2 \ h2'" "(s1, h1), (s2, h2) \ View f J e" "(s1, h1'), (s2, h2') \ View f J e" then show "h1' = h1 \ h2' = h2" by (meson assms hyper_sat.simps(11) preciseE) qed fun syntactic_unary :: "('i, 'a, nat) assertion \ bool" where "syntactic_unary (Bool b) \ True" | "syntactic_unary (And A B) \ syntactic_unary A \ syntactic_unary B" | "syntactic_unary (Star A B) \ syntactic_unary A \ syntactic_unary B" | "syntactic_unary (Low e) \ False" | "syntactic_unary Emp \ True" | "syntactic_unary (PointsTo v va vb) \ True" | "syntactic_unary (Exists x A) \ syntactic_unary A" | "syntactic_unary (SharedGuard _ e) \ True" | "syntactic_unary (UniqueGuard _ e) \ True" | "syntactic_unary (View view A e) \ syntactic_unary A" | "syntactic_unary (PreSharedGuards _) \ False" | "syntactic_unary (PreUniqueGuards _) \ False" | "syntactic_unary EmptyFullGuards \ True" | "syntactic_unary (LowExp x) \ False" | "syntactic_unary (Imp b A) \ False" lemma syntactic_unary_implies_unary: assumes "syntactic_unary A" shows "unary A" using assms proof (induct A) case (And A1 A2) show ?case proof (rule unary_smallerI) fix \1 \2 assume "\1, \1 \ And A1 A2 \ \2, \2 \ And A1 A2" then have "\1, \2 \ A1 \ \1, \2 \ A2" using And unary_def by (metis hyper_sat.simps(3) prod.exhaust_sel syntactic_unary.simps(2)) then show "\1, \2 \ And A1 A2" using hyper_sat.simps(3) by blast qed next case (Star A B) then have "unary A \ unary B" by simp show ?case proof (rule unaryI) fix s1 h1 s2 h2 assume asm0: "(s1, h1), (s1, h1) \ Star A B \ (s2, h2), (s2, h2) \ Star A B" then obtain a1 b1 a2 b2 where "Some h1 = Some a1 \ Some b1" "(s1, a1), (s1, a1) \ A" "(s1, b1), (s1, b1) \ B" "Some h2 = Some a2 \ Some b2" "(s2, a2), (s2, a2) \ A" "(s2, b2), (s2, b2) \ B" by (meson always_sat_refl hyper_sat.simps(4)) then have "(s1, a1), (s2, a2) \ A \ (s1, b1), (s2, b2) \ B" using \unary A \ unary B\ unaryE by blast then show "(s1, h1), (s2, h2) \ Star A B" using \Some h1 = Some a1 \ Some b1\ \Some h2 = Some a2 \ Some b2\ hyper_sat.simps(4) by blast qed next case (Exists x A) then have "unary A" by simp show ?case proof (rule unaryI) fix s1 h1 s2 h2 assume "(s1, h1), (s1, h1) \ Exists x A \ (s2, h2), (s2, h2) \ Exists x A" then obtain v1 v2 where "(s1(x := v1), h1), (s1(x := v1), h1) \ A \ (s2(x := v2), h2), (s2(x := v2), h2) \ A" by (meson always_sat_refl hyper_sat.simps(7)) then have "(s1(x := v1), h1), (s2(x := v2), h2) \ A" using \unary A\ unary_def by blast then show "(s1, h1), (s2, h2) \ Exists x A" by auto qed next case (View view A x) then have "unary A" by simp show ?case proof (rule unaryI) fix s1 h1 s2 h2 assume asm0: "(s1, h1), (s1, h1) \ View view A x \ (s2, h2), (s2, h2) \ View view A x" then have "(s1, h1), (s2, h2) \ A" by (meson \unary A\ hyper_sat.simps(11) unaryE) then show "(s1, h1), (s2, h2) \ View view A x" using asm0 by fastforce qed qed (auto simp add: unary_def) record ('i, 'a, 'v) single_context = view :: "(loc \ val) \ 'v" abstract_view :: "'v \ 'v" saction :: "'v \ 'a \ 'v" uaction :: "'i \ 'v \ 'a \ 'v" invariant :: "('i, 'a, 'v) assertion" type_synonym ('i, 'a, 'v) cont = "('i, 'a, 'v) single_context option" definition no_guard_assertion where "no_guard_assertion A \ (\s1 h1 s2 h2. (s1, h1), (s2, h2) \ A \ no_guard h1 \ no_guard h2)" text \Axiom that says that view only depends on the part of the heap described by inv\ definition view_function_of_inv :: "('i, 'a, nat) single_context \ bool" where "view_function_of_inv \ \ (\(h :: ('i, 'a) heap) (h' :: ('i, 'a) heap) s. (s, h), (s, h) \ invariant \ \ (h' \ h) \ view \ (normalize (get_fh h)) = view \ (normalize (get_fh h')))" definition wf_indexed_precondition :: "('i \ 'a \ 'a \ bool) \ bool" where "wf_indexed_precondition pre \ (\a b k. pre k a b \ (pre k b a \ pre k a a))" definition wf_precondition :: "('a \ 'a \ bool) \ bool" where "wf_precondition pre \ (\a b. pre a b \ (pre b a \ pre a a))" lemma wf_precondition_rep_prec: assumes "wf_precondition pre" shows "Rep_precondition (Abs_precondition pre) = pre" using Abs_precondition_inverse[of pre] assms mem_Collect_eq wf_precondition_def[of pre] by blast lemma wf_indexed_precondition_rep_prec: assumes "wf_indexed_precondition pre" shows "Rep_indexed_precondition (Abs_indexed_precondition pre) = pre" using Abs_indexed_precondition_inverse[of pre] assms mem_Collect_eq wf_indexed_precondition_def[of pre] by blast definition LowView where "LowView f A x = (Exists x (And (View f A (\s. s x)) (LowExp (Evar x))))" lemma LowViewE: assumes "(s, h), (s', h') \ LowView f A x" and "x \ fvA A" shows "(s, h), (s', h') \ A \ f (normalize (get_fh h)) = f (normalize (get_fh h'))" proof - obtain v v' where "(s(x := v), h), (s'(x := v'), h') \ And (View f A (\s. s x)) (LowExp (Evar x))" by (metis LowView_def assms(1) hyper_sat.simps(7)) then obtain "(s(x := v), h), (s'(x := v'), h') \ View f A (\s. s x)" "(s(x := v), h), (s'(x := v'), h') \ LowExp (Evar x)" using hyper_sat.simps(3) by blast then obtain "(s(x := v), h), (s'(x := v'), h') \ A" "v = v'" "f (normalize (get_fh h)) = f (normalize (get_fh h'))" by simp moreover have "(s, h), (s', h') \ A" by (meson agrees_same agrees_update assms(2) calculation(1) sat_comm_aux) ultimately show ?thesis by blast qed lemma LowViewI: assumes "(s, h), (s', h') \ A" and "f (normalize (get_fh h)) = f (normalize (get_fh h'))" and "x \ fvA A" shows "(s, h), (s', h') \ LowView f A x" proof - let ?s = "s(x := f (normalize (get_fh h)))" let ?s' = "s'(x := f (normalize (get_fh h')))" have "(?s, h), (?s', h') \ A" by (meson agrees_same_aux agrees_update assms(1) assms(3) sat_comm_aux) then have "(?s, h), (?s', h') \ And (View f A (\s. s x)) (LowExp (Evar x))" using assms(2) by auto then show ?thesis using LowView_def by (metis hyper_sat.simps(7)) qed definition disjoint :: "('a set) \ ('a set) \ bool" where "disjoint h1 h2 = (h1 \ h2 = {})" definition unambiguous where "unambiguous A x \ (\s1 h1 s2 h2 v1 v2 v1' v2'. (s1(x := v1), h1), (s2(x := v2), h2) \ A \ (s1(x := v1'), h1), (s2(x := v2'), h2) \ A \ v1 = v1' \ v2 = v2')" definition all_axioms :: "('v \ 'w) \ ('v \ 'a \ 'v) \ ('a \ 'a \ bool) \ ('i \ 'v \ 'b \ 'v) \ ('i \ 'b \ 'b \ bool) \ bool" where "all_axioms \ sact spre uact upre \ \\Every action's relational precondition is sufficient to preserve the low-ness of the abstract view of the resource value:\ (\v v' sarg sarg'. \ v = \ v' \ spre sarg sarg' \ \ (sact v sarg) = \ (sact v' sarg')) \ (\v v' uarg uarg' i. \ v = \ v' \ upre i uarg uarg' \ \ (uact i v uarg) = \ (uact i v' uarg')) \ (\sarg sarg'. spre sarg sarg' \ spre sarg' sarg') \ (\uarg uarg' i. upre i uarg uarg' \ upre i uarg' uarg') \ \\All relevant pairs of actions commute w.r.t. the abstract view:\ (\v v' sarg sarg'. \ v = \ v' \ spre sarg sarg \ spre sarg' sarg' \ \ (sact (sact v sarg) sarg') = \ (sact (sact v' sarg') sarg)) \ (\v v' sarg uarg i. \ v = \ v' \ spre sarg sarg \ upre i uarg uarg \ \ (sact (uact i v uarg) sarg) = \ (uact i (sact v' sarg) uarg)) \ (\v v' uarg uarg' i i'. i \ i' \ \ v = \ v' \ upre i uarg uarg \ upre i' uarg' uarg' \ \ (uact i' (uact i v uarg) uarg') = \ (uact i (uact i' v' uarg') uarg))" subsection \Rules of the Logic\ inductive CommCSL :: "('i, 'a, nat) cont \ ('i, 'a, nat) assertion \ cmd \ ('i, 'a, nat) assertion \ bool" ("_ \ {_} _ {_}" [51,0,0] 81) where RuleSkip: "\ \ {P} Cskip {P}" | RuleAssign: "\ \\. \ = Some \ \ x \ fvA (invariant \) ; collect_existentials P \ fvE E = {} \ \ \ \ {subA x E P} Cassign x E {P} " | RuleNew: "\ x \ fvE E; \\. \ = Some \ \ x \ fvA (invariant \) \ view_function_of_inv \ \ \ \ \ {Emp} Calloc x E {PointsTo (Evar x) pwrite E}" | RuleWrite: "\ \\. \ = Some \ \ view_function_of_inv \ ; v \ fvE loc \ \ \ \ {Exists v (PointsTo loc pwrite (Evar v))} Cwrite loc E {PointsTo loc pwrite E}" | "\ \\. \ = Some \ \ x \ fvA (invariant \) \ view_function_of_inv \ ; x \ fvE E \ fvE e \ \ \ \ {PointsTo E \ e} Cread x E {And (PointsTo E \ e) (Bool (Beq (Evar x) e))}" | RuleShare: "\ \ = \ view = f, abstract_view = \, saction = sact, uaction = uact, invariant = J \ ; all_axioms \ sact spre uact upre ; Some \ \ {Star P EmptyFullGuards} C {Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre)))}; view_function_of_inv \ ; unary J ; precise J ; wf_indexed_precondition upre ; wf_precondition spre ; x \ fvA J ; no_guard_assertion (Star P (LowView (\ \ f) J x)) \ \ None \ {Star P (LowView (\ \ f) J x)} C {Star Q (LowView (\ \ f) J x)}" | RuleAtomicUnique: "\ \ = \ view = f, abstract_view = \, saction = sact, uaction = uact, invariant = J \ ; no_guard_assertion P ; no_guard_assertion Q ; None \ {Star P (View f J (\s. s x))} C {Star Q (View f J (\s. uact index (s x) (map_to_arg (s uarg))))} ; precise J ; unary J ; view_function_of_inv \ ; x \ fvC C \ fvA P \ fvA Q \ fvA J ; uarg \ fvC C ; l \ fvC C ; x \ fvS (\s. map_to_list (s l)) ; x \ fvS (\s. map_to_arg (s uarg) # map_to_list (s l)) \ \ Some \ \ {Star P (UniqueGuard index (\s. map_to_list (s l)))} Catomic C {Star Q (UniqueGuard index (\s. map_to_arg (s uarg) # map_to_list (s l)))}" | RuleAtomicShared: "\ \ = \ view = f, abstract_view = \, saction = sact, uaction = uact, invariant = J \ ; no_guard_assertion P ; no_guard_assertion Q ; None \ {Star P (View f J (\s. s x))} C {Star Q (View f J (\s. sact (s x) (map_to_arg (s sarg))))} ; precise J ; unary J ; view_function_of_inv \ ; x \ fvC C \ fvA P \ fvA Q \ fvA J ; sarg \ fvC C ; ms \ fvC C ; x \ fvS (\s. map_to_multiset (s ms)) ; x \ fvS (\s. {# map_to_arg (s sarg) #} + map_to_multiset (s ms)) \ \ Some \ \ {Star P (SharedGuard \ (\s. map_to_multiset (s ms)))} Catomic C {Star Q (SharedGuard \ (\s. {# map_to_arg (s sarg) #} + map_to_multiset (s ms)))}" | RulePar: "\ \ \ {P1} C1 {Q1} ; \ \ {P2} C2 {Q2} ; disjoint (fvA P1 \ fvC C1 \ fvA Q1) (wrC C2) ; disjoint (fvA P2 \ fvC C2 \ fvA Q2) (wrC C1) ; \\. \ = Some \ \ disjoint (fvA (invariant \)) (wrC C2) ; \\. \ = Some \ \ disjoint (fvA (invariant \)) (wrC C1) ; precise P1 \ precise P2 \ \ \ \ {Star P1 P2} Cpar C1 C2 {Star Q1 Q2}" | RuleIf1: "\ \ \ {And P (Bool b)} C1 {Q} ; \ \ {And P (Bool (Bnot b))} C2 {Q} \ \ \ \ {And P (Low b)} Cif b C1 C2 {Q}" | RuleIf2: "\ \ \ {And P (Bool b)} C1 {Q} ; \ \ {And P (Bool (Bnot b))} C2 {Q} ; unary Q \ \ \ \ {P} Cif b C1 C2 {Q}" | RuleSeq: "\ \ \ {P} C1 {R} ; \ \ {R} C2 {Q} \ \ \ \ {P} Cseq C1 C2 {Q}" | RuleFrame: "\ \ \ {P} C {Q} ; disjoint (fvA R) (wrC C) ; precise P \ precise R \ \ \ \ {Star P R} C {Star Q R}" | RuleCons: "\ \ \ {P'} C {Q'} ; entails P P' ; entails Q' Q \ \ \ \ {P} C {Q}" | RuleExists: "\ \ \ {P} C {Q} ; x \ fvC C ; \\. \ = Some \ \ x \ fvA (invariant \) ; unambiguous P x \ \ \ \ {Exists x P} C {Exists x Q}" | RuleWhile1: "\ \ {And I (Bool b)} C {And I (Low b)} \ \ \ {And I (Low b)} Cwhile b C {And I (Bool (Bnot b))}" | RuleWhile2: "\ unary I ; \ \ {And I (Bool b)} C {I} \ \ \ \ {I} Cwhile b C {And I (Bool (Bnot b))}" end \ No newline at end of file diff --git a/thys/Crypto_Standards/EC_Common.thy b/thys/Crypto_Standards/EC_Common.thy --- a/thys/Crypto_Standards/EC_Common.thy +++ b/thys/Crypto_Standards/EC_Common.thy @@ -1,1117 +1,1118 @@ theory EC_Common imports "Elliptic_Curves_Group_Law.Elliptic_Test" "More_Residues" begin section \Edits to Elliptic_Test\ text \There are a few edits we need to make to the AFP entry we build on here. First there are several variables defined in its last section that we need to ignore. In particular, we want to ignore the definitions of "m", "a", etc. so that we can use those names otherwise. To be entirely safe, and because we don't use these definitions anywhere, we scrub them entirely.\ hide_const Elliptic_Test.m hide_fact Elliptic_Test.m_def hide_const Elliptic_Test.a hide_fact Elliptic_Test.a_def hide_const Elliptic_Test.b hide_fact Elliptic_Test.b_def hide_const Elliptic_Test.gx hide_fact Elliptic_Test.gx_def hide_const Elliptic_Test.gy hide_fact Elliptic_Test.gy_def hide_const Elliptic_Test.priv hide_fact Elliptic_Test.priv_def hide_const Elliptic_Test.pubx hide_fact Elliptic_Test.pubx_def hide_const Elliptic_Test.puby hide_fact Elliptic_Test.puby_def hide_const Elliptic_Test.order hide_fact Elliptic_Test.order_def text \The second thing we need to "edit" is this. We need to be able to use this lemma for checking test vectors. So we need it to be designated with [code].\ lemmas (in residues_prime_gt2) [code] = on_curve_residues_eq section \EC_Common\ text \This section has facts about the arithmetic of points on an elliptic curve.\ lemma (in euclidean_semiring_cancel) mod_pow_cong: "a mod c = b mod c \ (a ^ n) mod c = (b ^ n) mod c" by (induction n) (auto intro!: mod_mult_cong) declare (in domain) integral_iff[simp] declare (in field) divide_self[simp] declare (in field) divide_eq_0_iff[simp] lemma (in cring) power2_sum: assumes [simp]: "x \ carrier R" "y \ carrier R" shows "(x \ y)[^](2::nat) = x[^](2::nat) \ y[^](2::nat) \ \2\ \ x \ y" proof - have 1: "(2::int) = 1 + 1" by arith have 2: "\2\ = \ \ \" unfolding 1 of_int_add by simp show ?thesis by (simp add: power2_eq_square 2 minus_eq l_distr r_distr a_ac m_ac minus_add r_minus) qed lemma (in ring) diff_same[simp]: "a \ carrier R \ a \ a = \" using eq_diff0 by force lemma (in cring) power2_diff: assumes [simp]: "x \ carrier R" "y \ carrier R" shows "(x \ y)[^](2::nat) = x[^](2::nat) \ y[^](2::nat) \ \2\ \ x \ y" proof - have 1: "(2::int) = 1 + 1" by arith have 2: "\2\ = \ \ \" unfolding 1 of_int_add by simp show ?thesis by (simp add: power2_eq_square 2 minus_eq l_distr r_distr a_ac m_ac minus_add r_minus) qed lemma (in ring) power3_eq_cube: "x \ carrier R \ x[^]\<^bsub>R\<^esub>(3::nat) = x \\<^bsub>R\<^esub> x \\<^bsub>R\<^esub> x" by (simp add: numeral_3_eq_3) lemma (in ring) zero_pow_nat[simp]: "n \ 0 \ \\<^bsub>R\<^esub> [^]\<^bsub>R\<^esub> (n::nat) = \\<^bsub>R\<^esub>" using nat_pow_zero by blast lemma (in comm_group) m_one_iff: "p \ carrier G \ q \ carrier G \ p \ q = \ \ p = inv q" using local.inv_equality by auto lemma (in residues) res_diff_eq: "x \ y = (x - y) mod m" unfolding a_minus_def diff_conv_add_uminus res_neg_eq res_add_eq by (simp add: mod_simps) lemma (in field) nonzero_mult_divide_mult_cancel_left' [simp]: assumes [simp]: "a \ carrier R" "b \ carrier R" "c \ carrier R" "b \ \" "c \ \" shows "(a \ c) \ (b \ c) = a \ b" by (subst (1 2) m_comm) simp_all lemma (in field) nonzero_mult_divide_divide_cancel_left [simp]: assumes [simp]: "a \ carrier R" "b \ carrier R" "c \ carrier R" "b \ \" "c \ \" shows "(a \ c) \ (b \ c) = a \ b" by (subst (1 3) m_div_def) (simp add: nonzero_imp_inverse_nonzero) lemma (in field) l_diff_distr: "x \ carrier R \ y \ carrier R \ z \ carrier R \ (x \ y) \ z = x \ z \ y \ z" using r_diff_distr[of x y z] by (simp add: m_ac) lemma (in field) l_diff_div_distr: "x \ carrier R \ y \ carrier R \ z \ carrier R \ (x \ y) \ z = (x \ z) \ (y \ z)" by (auto simp: m_div_def l_diff_distr) lemma (in field) of_natural_nat_pow: "\n\\<^sub>\ [^] (p::nat) = \n^p\\<^sub>\" by (induction p) (auto simp: m_ac) lemma (in field) of_integer_int[simp]: "\int n\ = \n\\<^sub>\" by (induction n) auto lemma (in field) of_integer_numeral: "\numeral n\ = \numeral n\\<^sub>\" apply (subst semiring_1_class.of_nat_numeral[symmetric]) apply (subst of_integer_int) .. lemma (in field) divide_mult_comm: "a \ carrier R \ b \ carrier R \ c \ carrier R \ b \ \ \ (a \ b) \ c = a \ (c \ b)" by (subst times_divide_eq_left) (auto simp flip: times_divide_eq_right) lemma (in field) div_cancel: assumes [simp]: "a \ carrier R" "b \ carrier R" "c \ carrier R" "c \ \" shows "a \ c = b \ c \ a = b" proof assume "a \ c = b \ c" then have "(a \ c) \ inv c = (b \ c) \ inv c" by simp then show "a = b" by (subst (asm) (1 2) nonzero_divide_divide_eq_left) (auto simp: nonzero_imp_inverse_nonzero) qed simp lemma (in group) pow_dbl_eq_pow_of_prime_ord_gt_2: "e \ carrier G \ prime (ord e) \ ord e > 2 \ e [^] (2 * n :: nat) = \ \ e [^] n = \" by (subst (1 2) pow_eq_id) (auto simp: prime_dvd_mult_nat) lemma (in ring) add_eq_iff_eq_minus: "a \ carrier R \ b \ carrier R \ c \ carrier R \ a \ b = c \ a = c \ b" by (metis a_minus_def add.inv_solve_right) lemma (in field) mult_eq_iff_eq_divide: "a \ carrier R \ b \ carrier R \ c \ carrier R \ b \ \ \ a \ b = c \ a = c \ b" by (metis (full_types) cring_simprules(14) div_closed divide_mult_comm l_one local.divide_self m_lcancel) lemma (in field) eq_mult_iff_divide_eq: "a \ carrier R \ b \ carrier R \ c \ carrier R \ b \ \ \ c = a \ b \ c \ b = a" by (metis mult_eq_iff_eq_divide) lemma (in cring) r_distr_diff: assumes "x \ carrier R" "y \ carrier R" "z \ carrier R" shows "x \ (y \ z) = x \ y \ x \ z" using assms by (simp add: cring.cring_simprules(15) is_cring r_distr r_minus) lemma (in field) divide_eq_divide_iff_mult_eq_mult: assumes [simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" "w \ carrier R" "y \ \" "w \ \" shows "x \ y = z \ w \ x \ w = z \ y" by (simp add: mult_eq_iff_eq_divide flip: eq_mult_iff_divide_eq) lemma (in field) inv_pow_distr: assumes "x \ carrier R" "x \ \" shows "inv (x[^](k::nat)) = (inv x)[^]k" by (metis assms comm_inv_char inv_closed nat_pow_closed nat_pow_one r_inv nat_pow_distrib) lemma (in field) inv_Suc_pow_cancel: assumes "x \ carrier R" "x \ \" shows "x \ (inv x)[^](Suc (k::nat)) = (inv x)[^]k" using assms r_inv nat_pow_Suc m_lcomm by fastforce lemma (in ell_field) xz_coord_padd_simp: assumes "on_curvep a b (x1, y1, z)" and "on_curvep a b (x2, y2, z)" and "(x3, y3, z3) = padd a (x1, y1, z) (x2, y2, z)" and ab_in_carrier[simp]: "a \ carrier R" "b \ carrier R" and "(x2 \ x1) \ z \ \" shows "x3 = (x2 \ x1) \ z[^](4::nat) \ ((x1 \ x2 \ a \ z[^](2::nat)) \ (x1 \ x2) \ \2\ \ b \ z[^](3::nat) \ \2\ \ y1 \ y2 \ z) \ z3 = (x2 \ x1)[^](3::nat) \ z[^](5::nat)" proof - define u where "u = x2 \ z \ x1 \ z" have carrier: "x1 \ carrier R" "y1 \ carrier R" "x2 \ carrier R" "y2 \ carrier R" "z \ carrier R" using assms(1) assms(2) unfolding on_curvep_def by blast+ then have z_nz: "z \ \" using assms(1) assms(6) domain.integral_iff unfolding on_curvep_def by force have u_nz: "u \ \" using assms(1) assms(6) carrier u_def unfolding on_curvep_def by (simp add: cring_axioms cring.cring_simprules(14) cring.r_distr_diff) have [simp]: "in_carrierp (x1, y1, z)" "in_carrierp (x2, y2, z)" using assms(1) assms(2) on_curvep_imp_in_carrierp by blast+ then have on_curvep3: "on_curvep a b (x3, y3, z3)" by (simp add: assms(3) padd_closed assms(1) assms(2)) then have carrier3: "x3 \ carrier R" "y3 \ carrier R" "z3 \ carrier R" using on_curvep_imp_in_carrierp in_carrierp_def by auto have p1_ec_eqn: "y1[^](2::nat) \ z = x1[^](3::nat) \ a \ x1 \ z[^](2::nat) \ b \ z[^](3::nat)" using assms(1) z_nz prod_cases3 carrier unfolding on_curvep_def by (auto split: prod.splits) have p2_ec_eqn: "y2[^](2::nat) \ z = x2[^](3::nat) \ a \ x2 \ z[^](2::nat) \ b \ z[^](3::nat)" using assms(2) z_nz prod_cases3 unfolding on_curvep_def by (auto split: prod.splits) have z3: "z3 = (x2 \ x1)[^](3::nat) \ z[^](5::nat)" using assms(3) u_nz z_nz carrier padd_def p1_ec_eqn p2_ec_eqn unfolding u_def Let_def by (auto split: prod.splits, field, simp) have x3: "x3 = (x2 \ x1) \ z[^](4::nat) \ ((y2[^](2::nat) \ z) \ (y1[^](2::nat) \ z) \ \2\ \ y1 \ y2 \ z \ (x1 \ x2) \ (x2 \ x1)[^](2::nat))" using assms(3) u_nz z_nz carrier padd_def unfolding u_def Let_def by (auto split: prod.splits, field, simp) then have x3': "\ = (x2 \ x1) \ z[^](4::nat) \ ((x1 \ x2 \ a \ z[^](2::nat)) \ (x1 \ x2) \ \2\ \ b \ z[^](3::nat) \ \2\ \ y1 \ y2 \ z)" apply (subst p1_ec_eqn, subst p2_ec_eqn) using carrier ab_in_carrier by (field, simp) then show ?thesis using x3 x3' z3 by simp qed lemma (in ell_field) pdouble_remove_y_eqn: assumes on_curvep: "on_curvep \-3\ b (x, y, z)" and [simp]: "b \ carrier R" "z \ \" "y \ \" and double: "Point qx qy = make_affine (pdouble \-3\ (x, y, z))" shows "qx \ (\2\ \ y \ z) [^](2::nat) = (x [^](2::nat) \ \3\ \ z [^](2::nat)) [^](2::nat) \ \8\ \ b \ x \ z [^](3::nat)" proof - have in_carrierp[simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" using assms(1) on_curvep unfolding on_curvep_def by fast+ then have "on_curve \-3\ b (Point qx qy)" using on_curvep double pdouble_closed assms(1) assms(2) assms(3) on_curvep_iff_on_curve of_int_closed pdouble_in_carrierp in_carrierp on_curvep_imp_in_carrierp by metis define l where "l = \2\ \ y \ z" define m where "m = \3\ \ x [^] (2::nat) \ \-3\ \ z[^](2::nat)" define t where "t = inv (l[^](3::nat))" have 0: "l \ \" by (simp add: l_def) have 1: "l[^](3::nat) \ \" by (simp add: l_def) then have 2: "t \ l \ l[^](2::nat) = \" using t_def l_def by (simp add: m_assoc numeral_2_eq_2 numeral_3_eq_3) have [simp]: "l \ carrier R" "t \ carrier R" "m \ carrier R" using l_def inv_closed l_def m_def t_def m_closed a_closed nat_pow_closed 1 by auto then have 3: "qx = (m[^](2::nat) \ \4\ \ x \ y \ l) \ t \ l" using assms(3) assms(4) double in_carrierp m_comm int_pow_int 0 1 m_div_def t_def unfolding make_affine_def pdouble_def Let_def l_def m_def by (simp add: 0 1 m_div_def m_ac) then have 4: "qx \ l[^](2::nat) = m[^](2::nat) \ \4\ \ x \ y \ l" using 2 by (simp add: 2 m_assoc numeral_2_eq_2 numeral_3_eq_3) have 4: "y[^](2::nat) \ z = x[^](3::nat) \ \-3\ \ x \ z[^](2::nat) \ b \ z[^](3::nat)" using assms unfolding on_curvep_def by fast then have 5: "m[^](2::nat) \ \4\ \ x \ y \ l = (x[^](2::nat) \ \3\ \ z[^](2::nat))[^](2::nat) \ \8\ \ b \ x \ z[^](3::nat)" unfolding m_def l_def apply (field 4) by (rule TrueI) then show ?thesis using 4 l_def 2 3 m_assoc by auto qed lemma (in cring) on_curvep_nz: "\on_curvep a b (x, y, z); z \ \\ \ y [^] (2::nat) \ z = x [^] (3::nat) \ a \ x \ z [^] (2::nat) \ b \ z [^] (3::nat)" by (simp add: on_curvep_def) lemma (in field) on_curvep_nz_identity: assumes "on_curvep a b (x, y, z)" "z \ \" "a \ carrier R" "b \ carrier R" shows "(\4\ \ x \ (x[^](2::nat) \ a \ z[^](2::nat)) \ \4\ \ b \ z[^](3::nat)) \ z = (\2\ \ y \ z)[^](2::nat)" proof - have x: "x \ carrier R" "y \ carrier R" "z \ carrier R" using assms(1) on_curvep_def by simp_all have 0: "x [^] (3::nat) \ a \ x \ z [^] (2::nat) \ b \ z [^] (3::nat) = y [^] (2::nat) \ z" using assms on_curvep_nz by simp have 1: "(\4\ \ x \ (x[^](2::nat) \ a \ z[^](2::nat)) \ \4\ \ b \ z[^](3::nat)) \ z = \4\ \ (x[^](3::nat) \ a \ x \ z[^](2::nat) \ b \ z[^](3::nat)) \ z" by (field, simp) have 2: "\ = \4\ \ y [^] (2::nat) \ z \ z" by (simp add: 0 x m_assoc) show ?thesis using assms x by (simp add: 1 2, field, simp) qed lemma (in residues_prime) res_inv_mult: assumes "z \ 0" "z \ carrier R" shows "z ^ (p - 2) * z mod p = (inv z) * z mod p" proof - have "\ p dvd (nat z)" using assms R_def residues.res_carrier_eq residues_axioms nat_dvd_not_less by auto then have "z ^ (p - 1) mod p = 1" by (metis cong_def dvd_eq_mod_eq_0 int_nat_eq euler_theorem mod_0 nat_int of_nat_dvd_iff one_cong p_coprime_right_int res_one_eq prime_totient_eq) then show ?thesis by (metis R_def Suc_1 Suc_diff_Suc assms l_inv p_prime prime_gt_1_nat res_one_eq residues.res_mult_eq residues_axioms semiring_normalization_rules(28) zero_cong) qed lemma (in residues_prime) res_inv_one: assumes "z \ 0" "z \ carrier R" shows "z ^ (p - 2) * z mod p = 1" using assms res_inv_mult l_inv res_mult_eq res_one_eq res_zero_eq by auto lemma (in residues_prime) res_mult_div: assumes "z \ 0" "z \ carrier R" shows "x * z ^ (p - 2) mod p = x \ z" proof - have "z \ inv z = \" using assms r_inv zero_cong by blast then have "z ^ (p - 2) mod p = (inv z) mod p" by (metis assms comm_inv_char mod_in_carrier mult_cong res_mult_eq semiring_normalization_rules(7) res_inv_mult) then show ?thesis using assms by (metis m_div_def mod_mult_right_eq res_mult_eq zero_cong) qed lemma (in residues_prime_gt2) add_pnts_eqn_x: assumes p1_on_curve: "on_curve a b (Point x1 y1)" and p2_on_curve: "on_curve a b (Point x2 y2)" and p1_plus_p2: "Point x3 y3 = add a (Point x1 y1) (Point x2 y2)" and nz_cond: "x1 = x2 \ y1 = y2" and ab_in_carrier: "a \ carrier R" "b \ carrier R" shows "x3 \ (x2 \ x1)[^](2::nat) = (a \ x1 \ x2) \ (x1 \ x2) \ \2\ \ b \ \2\ \ y1 \ y2" proof cases assume "x1 = x2" with nz_cond p1_on_curve have [simp]: "y1 = y2" "x1 = x2" "x2 \ carrier R" "y2 \ carrier R" and *: "y2 [^] (2::nat) = x2 [^] (3::nat) \ a \ x2 \ b" by (auto simp add: on_curve_def) have "on_curve a b (Point x3 y3)" unfolding p1_plus_p2 by (intro add_closed assms) then have [simp]: "x3 \ carrier R" "y3 \ carrier R" by (auto simp add: on_curve_def) show ?thesis using * apply (simp add: local.power2_eq_square local.m_assoc) apply field apply simp done next assume nz_cond: "x1 \ x2" have carrier[simp]: "x1 \ carrier R" "y1 \ carrier R" "x2 \ carrier R" "y2 \ carrier R" using assms(1) assms(2) unfolding on_curve_def by simp_all have y1sq: "y1[^](2::nat) = x1[^](3::nat) \ a \ x1 \ b" using p1_on_curve unfolding on_curve_def by simp have y2sq: "y2[^](2::nat) = x2[^](3::nat) \ a \ x2 \ b" using p2_on_curve unfolding on_curve_def by simp have denom_nz: "(x2 \ x1)[^](2::nat) \ \ \ (x2 \ x1)[^](2::nat) \ carrier R" using nz_cond by auto then have r_cancel: "\ z \ carrier R. z \ (x2 \ x1)[^](2::nat) \ (x2 \ x1)[^](2::nat) = z" using denom_nz divide_self m_assoc by (metis local.times_divide_eq_right r_one) have 0: "x3 = ((y2 \ y1) \ (x2 \ x1))[^](2::nat) \ (x1 \ x2)" using p1_plus_p2 nz_cond add.m_assoc nz_cond cring_simprules(15) eq_diff0 local.minus_add unfolding add_def Let_def by force then have 1: "\ = (y2 \ y1)[^](2::nat) \ (x2 \ x1) [^](2::nat) \ (x1 \ x2)" using nonzero_power_divide nz_cond by auto then have "\ = ((y2 \ y1)[^](2::nat) \ (x1 \ x2) \ (x2 \ x1)[^](2::nat)) \ (x2 \ x1)[^](2::nat)" using l_diff_distr denom_nz carrier r_cancel by (simp add: m_div_def) then have 2: "x3 \ (x2 \ x1)[^](2::nat) = (y2 \ y1)[^](2::nat) \ (x1 \ x2) \ (x2 \ x1)[^](2::nat)" using r_cancel 0 1 by (simp add: denom_nz) then have "\ = (a \ x1 \ x2) \ (x1 \ x2) \ \2\ \ b \ \2\ \ y1 \ y2" using p1_on_curve p2_on_curve ab_in_carrier unfolding on_curve_def by (field y1sq y2sq, simp) then show ?thesis using 2 by simp qed lemma (in residues_prime_gt2) add_eliminate_ys_simp: assumes p1_on_curve: "on_curve a b (Point x1 y1)" and p2_on_curve: "on_curve a b (Point x2 y2)" and p1_plus_p2: "Point x3 y3 = add a (Point x1 y1) (Point x2 y2)" and mp1_plus_p2: "Point xd yd = add a (opp (Point x1 y1)) (Point x2 y2)" and x1_neq_x2: "x1 \ x2" and ab_in_carrier: "a \ carrier R" "b \ carrier R" shows "(x3 \ xd) \ (x2 \ x1)[^](2::nat) = \2\ \ (a \ x1 \ x2) \ (x1 \ x2) \ \4\ \ b" proof - have carrier[simp]: "x1 \ carrier R" "y1 \ carrier R" "x2 \ carrier R" "y2 \ carrier R" using assms(1) assms(2) unfolding on_curve_def by simp_all have "on_curve a b (Point x3 y3)" "on_curve a b (Point xd yd)" using p1_plus_p2 mp1_plus_p2 p1_on_curve p2_on_curve opp_closed add_closed by (simp add: ab_in_carrier)+ then have x3_xd_carrier[simp]: "x3 \ carrier R" "xd \ carrier R" unfolding on_curve_def by simp_all have x3: "x3 \ (x2 \ x1)[^](2::nat) = (a \ x1 \ x2) \ (x1 \ x2) \ \2\ \ b \ \2\ \ y1 \ y2" using assms add_pnts_eqn_x by blast have 0: "on_curve a b (Point x1 (\ y1))" using p1_on_curve opp_closed opp_def by fastforce have "Point xd yd = add a (Point x1 (\ y1)) (Point x2 y2)" using mp1_plus_p2 unfolding opp_def by simp then have "xd \ (x2 \ x1)[^](2::nat) = (a \ x1 \ x2) \ (x1 \ x2) \ \2\ \ b \ \2\ \ (\ y1) \ y2" using carrier p1_on_curve p2_on_curve x1_neq_x2 add_pnts_eqn_x minus_minus 0 cring_simprules(15) ab_in_carrier by blast then have xd: "xd \ (x2 \ x1)[^](2::nat) = (a \ x1 \ x2) \ (x1 \ x2) \ \2\ \ b \ \2\ \ y1 \ y2" by (simp add: cring_simprules(28) cring_simprules(29) local.minus_eq) have sum: "(x3 \ xd) \ (x2 \ x1)[^](2::nat) = x3 \ (x2 \ x1)[^](2::nat) \ xd \ (x2 \ x1)[^](2::nat)" using x3_xd_carrier carrier by (field, simp) have "\ = \2\ \ (a \ x1 \ x2) \ (x1 \ x2) \ \4\ \ b" unfolding x3 xd by (field, simp) then show ?thesis using sum by presburger qed named_theorems remove_mod lemma mmult_mod_mod[remove_mod]: "(a mod m) **\<^bsub>m\<^esub> (b mod m) = (a * b) mod m" unfolding mmult_def mod_mult_eq .. lemma madd_mod_mod[remove_mod]: "(a mod m) ++\<^bsub>m\<^esub> (b mod m) = (a + b) mod m" unfolding madd_def mod_add_eq .. lemma msub_mod_mod[remove_mod]: "(a mod m) --\<^bsub>m\<^esub> (b mod m) = (a - b) mod m" unfolding msub_def mod_diff_eq .. lemma mpow_mod[remove_mod]: "(a mod m) ^^^\<^bsub>m\<^esub> n = (a ^ n) mod m" unfolding mpow_def power_mod .. lemma madd_mod: "(x ++\<^bsub>m\<^esub> y) mod m = x ++\<^bsub>m\<^esub> y" by (simp add: madd_def) lemma msub_same_eq_zero[simp]: "x --\<^bsub>m\<^esub> y = 0 \ x mod m = y mod m" by (metis (no_types, opaque_lifting) cancel_ab_semigroup_add_class.diff_right_commute eq_iff_diff_eq_0 mod_0 msub_def msub_mod_mod) definition mpdouble_nz :: "int \ int \ int ppoint \ int ppoint" where "mpdouble_nz m a = (\(x, y, z). let l = 2 mod m **\<^bsub>m\<^esub> y **\<^bsub>m\<^esub> z; n = 3 mod m **\<^bsub>m\<^esub> x ^^^\<^bsub>m\<^esub> 2 ++\<^bsub>m\<^esub> a **\<^bsub>m\<^esub> z ^^^\<^bsub>m\<^esub> 2 in (l **\<^bsub>m\<^esub> (n ^^^\<^bsub>m\<^esub> 2 --\<^bsub>m\<^esub> 4 mod m **\<^bsub>m\<^esub> x **\<^bsub>m\<^esub> y **\<^bsub>m\<^esub> l), n **\<^bsub>m\<^esub> (6 mod m **\<^bsub>m\<^esub> x **\<^bsub>m\<^esub> y **\<^bsub>m\<^esub> l --\<^bsub>m\<^esub> n ^^^\<^bsub>m\<^esub> 2) --\<^bsub>m\<^esub> 2 mod m **\<^bsub>m\<^esub> y ^^^\<^bsub>m\<^esub> 2 **\<^bsub>m\<^esub> l ^^^\<^bsub>m\<^esub> 2, l ^^^\<^bsub>m\<^esub> 3))" lemma mpdouble_eq_mpdouble_nz: "snd (snd p) \ 0 \ mpdouble_nz m a p = mpdouble m a p" by (auto simp add: mpdouble_def mpdouble_nz_def split: prod.split) locale ell_prime_field = residues_prime_gt2 + fixes A B :: nat assumes AB_in_carrier[simp]: "A \ carrier R" "B \ carrier R" and non_singular: "(4 * A ^ 3 + 27 * B ^ 2) mod p \ 0" begin definition curve :: "int point monoid" where "curve = \ carrier = ({ P. on_curve A B P }), monoid.mult = add A, monoid.one = Infinity \" lemma curve_simps: shows in_carrier_curve: "P \ carrier curve \ on_curve A B P" and add_curve: "x \\<^bsub>curve\<^esub> y = add A x y" and one_curve: "\\<^bsub>curve\<^esub> = Infinity" by (simp_all add: curve_def) lemma assumes "Point x y \ carrier curve" shows Point_in_carrierD1: "x \ carrier R" and Point_in_carrierD2: "y \ carrier R" using assms by (auto simp: in_carrier_curve on_curve_def) text \Just a few basic facts for casting a nat as an int.\ lemma nat_int_eq: "a = b \ int a = int b" by simp lemma nat_int_less: "a < b \ int a < int b" by simp lemma nat_int_le: "a \ b \ int a \ int b" by simp lemma nonsingular_in_bf: "nonsingular A B" using non_singular unfolding nonsingular_def res_of_natural_eq res_of_integer_eq res_mult_eq res_add_eq res_pow_eq res_zero_eq nat_int_eq zmod_int by (simp add: mod_simps) sublocale curve: comm_group curve apply (intro comm_groupI) using nonsingular_in_bf apply (auto simp: curve_simps add_0_l add_0_r Units_def add_assoc[of "int A" "int B" _ _ _, symmetric] add_comm[of A B] opp_closed intro!: add_closed) apply (intro bexI[of _ "opp _"]) apply (auto simp: add_opp curve_simps intro!: opp_closed) done lemma inv_curve: "x \ carrier curve \ inv\<^bsub>curve\<^esub> x = opp x" by (intro curve.inv_unique'[symmetric]) (auto simp: curve_simps add_opp add_comm[of A B _ x] opp_closed) lemma curve_square_is_mult: "P [^]\<^bsub>curve\<^esub> (2::nat) = P \\<^bsub>curve\<^esub> P" unfolding nat_pow_def apply (simp add: curve_def curve_simps) by (simp add: ell_field_axioms ell_field.add_0_l) lemma finite_curve[simp]: "finite (carrier curve)" proof (rule finite_subset) show "carrier curve \ case_prod Point ` (carrier R \ carrier R) \ { Infinity }" by (auto simp: in_carrier_curve on_curve_def split: point.splits) qed auto lemma Point_ne_one[simp]: "Point x y \ \\<^bsub>curve\<^esub>" "\\<^bsub>curve\<^esub> \ Point x y" by (simp_all add: one_curve) lemma odd_p: "odd p" using p_prime gt2 using prime_odd_nat by auto lemma mod_of_in_carrier: "x \ carrier R \ x mod int p = x" by (simp add: res_carrier_eq) lemma nz_in_carrier_gcd: assumes "x \ carrier R" "x \ 0" shows "gcd x (int p) = 1" proof - have "0 < x \ x < int p" using assms res_carrier_eq by auto then show ?thesis using p_prime by (simp add: p_coprime_right_int zdvd_not_zless) qed text \xy_to_pnt is for implementations that use (0,0) to represent Infinity. Below are many lemmas that relate pairs of ints and the point type. Similarly, xyz_to_pnt is useful for implementations that use triples of integers to represent points on the curve in projective form.\ definition xy_to_pnt :: "int \ int \ int point" where "xy_to_pnt = (\(x, y). if x mod int p = 0 \ y mod int p = 0 then Infinity else Point (x mod p) (y mod p))" lemma xy_to_pnt_eq: "x \ carrier R \ y \ carrier R \ xy_to_pnt (x, y) = (if x = \\<^bsub>R\<^esub> \ y = \\<^bsub>R\<^esub> then \\<^bsub>curve\<^esub> else Point x y)" by (simp add: xy_to_pnt_def mod_of_in_carrier one_curve res_zero_eq) lemma xy_to_pnt_Point: assumes "xy_to_pnt (x, y) = Point x' y'" shows "x mod p = x' \ y mod p = y'" proof - have "x mod int p \ 0 \ y mod int p \ 0" using assms unfolding xy_to_pnt_def by force then have "xy_to_pnt (x, y) = Point (x mod p) (y mod p)" using assms unfolding xy_to_pnt_def by fastforce then show ?thesis using assms by simp qed lemma xy_to_pnt_Infinity: assumes "xy_to_pnt (x, y) = Infinity" shows "x mod p = 0 \ y mod p = 0" proof - have "Infinity = (if x mod int p = 0 \ y mod int p = 0 then Infinity else Point (x mod int p) (y mod int p))" using assms xy_to_pnt_def by force then show ?thesis by (metis (no_types) point.simps(3)) qed definition xyz_to_pnt :: "int \ int \ int \ int point" where "xyz_to_pnt = (\(x, y, z). if z mod int p = 0 then Infinity else Point ((x mod int p) \\<^bsub>R\<^esub> (z mod int p)) ((y mod int p) \\<^bsub>R\<^esub> (z mod int p)))" lemma xyz_to_pnt_zero_z[simp]: "xyz_to_pnt (x, y, \\<^bsub>R\<^esub>) = \\<^bsub>curve\<^esub>" by (simp add: xyz_to_pnt_def mod_of_in_carrier one_curve zero_cong) lemma xyz_to_pnt_zero[simp]: "xyz_to_pnt (x, y, 0) = Infinity" by (simp add: xyz_to_pnt_def res_zero_eq) lemma xyz_to_pnt_eq: "x \ carrier R \ y \ carrier R \ z \ carrier R \ xyz_to_pnt (x, y, z) = (if z = \\<^bsub>R\<^esub> then \\<^bsub>curve\<^esub> else Point (x \\<^bsub>R\<^esub> z) (y \\<^bsub>R\<^esub> z))" using xyz_to_pnt_def mod_of_in_carrier res_zero_eq one_curve by simp lemma xyz_to_pnt_z_1: assumes [simp]: "x \ carrier R" "y \ carrier R" shows "xyz_to_pnt (x, y, \\<^bsub>R\<^esub>) = Point x y" proof - have "(1::int) \ carrier R" using res_carrier_eq gt2 by auto then show ?thesis unfolding xyz_to_pnt_def using mod_of_in_carrier m_gt_one assms divide_1 res_one_eq by auto qed lemma xyz_to_pnt_eq_make_affine: "x \ carrier R \ y \ carrier R \ z \ carrier R \ xyz_to_pnt (x, y, z) = make_affine (x, y, z)" by (simp add: xyz_to_pnt_def make_affine_def mod_of_in_carrier res_zero_eq) lemma xyz_to_pnt_in_carrier_on_curvep: assumes [simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" "xyz_to_pnt (x, y, z) \ carrier curve" shows "on_curvep A B (x, y, z)" proof - have "on_curve A B (make_affine (x, y, z))" by (simp flip: xyz_to_pnt_eq_make_affine in_carrier_curve) then show ?thesis by (subst on_curvep_iff_on_curve) (simp_all add: in_carrierp_def) qed lemma xyz_to_pnt_scale[symmetric]: assumes [simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" "c \ carrier R" "c \ \\<^bsub>R\<^esub>" shows "xyz_to_pnt (x, y, z) = xyz_to_pnt (x \\<^bsub>R\<^esub> c, y \\<^bsub>R\<^esub> c, z \\<^bsub>R\<^esub> c)" by (simp add: xyz_to_pnt_eq) lemma xyz_to_pnt_scale'[symmetric]: assumes [simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" "c \ carrier R" and zc: "z \ \\<^bsub>R\<^esub> \ c \ \\<^bsub>R\<^esub>" shows "xyz_to_pnt (x, y, z) = xyz_to_pnt (x \\<^bsub>R\<^esub> c, y \\<^bsub>R\<^esub> c, z \\<^bsub>R\<^esub> c)" proof cases assume "z = \\<^bsub>R\<^esub>" then show ?thesis using assms by (simp add: xyz_to_pnt_eq in_carrierp_def) next assume "z \ \\<^bsub>R\<^esub>" with zc show ?thesis using assms by (simp add: xyz_to_pnt_scale in_carrierp_def) qed lemma xyz_to_pnt_mod: "xyz_to_pnt (x mod int p, y mod int p, z mod int p) = xyz_to_pnt (x, y, z)" by (simp add: mod_of_in_carrier xyz_to_pnt_def) lemma inv_xyz_to_pnt: "x \ carrier R \ y \ carrier R \ z \ carrier R \ xyz_to_pnt (x, y, z) \ carrier curve \ inv\<^bsub>curve\<^esub> (xyz_to_pnt (x, y, z)) = xyz_to_pnt (x, \ y, z)" unfolding inv_curve opp_def by (auto simp add: xyz_to_pnt_eq one_curve) lemma xyz_to_pnt_dbl: assumes [simp]: "x \ carrier R" "y \ carrier R" "z \ carrier R" shows "xyz_to_pnt (x, y, z) \\<^bsub>curve\<^esub> xyz_to_pnt (x, y, z) = (let l = \2\ \ y \ z; m = \3\ \ x [^] (2::nat) \ A \ z [^] (2::nat) in xyz_to_pnt ( l \ (m [^] (2::nat) \ \4\ \ x \ y \ l), m \ (\6\ \ x \ y \ l \ m [^] (2::nat)) \ \2\ \ y [^] (2::nat) \ l [^] (2::nat), l [^] (3::nat)))" (is "?l = ?r") proof - have "?l = make_affine (pdouble A (x, y, z))" by (simp add: pdouble_correct in_carrierp_def add_curve xyz_to_pnt_eq_make_affine) also have "\ = ?r" by (simp add: pdouble_def Let_def one_curve xyz_to_pnt_eq_make_affine) finally show ?thesis . qed lemma xyz_to_pnt_square: assumes "x \ carrier R" "y \ carrier R" "z \ carrier R" shows "xyz_to_pnt (x, y, z) [^]\<^bsub>curve\<^esub> (2::nat) = (let l = \2\ \ y \ z; m = \3\ \ x [^] (2::nat) \ A \ z [^] (2::nat) in xyz_to_pnt ( l \ (m [^] (2::nat) \ \4\ \ x \ y \ l), m \ (\6\ \ x \ y \ l \ m [^] (2::nat)) \ \2\ \ y [^] (2::nat) \ l [^] (2::nat), l [^] (3::nat)))" (is "?l = ?r") unfolding curve_square_is_mult by (rule xyz_to_pnt_dbl) fact+ lemma xyz_to_pnt_add: assumes [simp]: "x\<^sub>1 \ carrier R" "y\<^sub>1 \ carrier R" "z\<^sub>1 \ carrier R" assumes [simp]: "x\<^sub>2 \ carrier R" "y\<^sub>2 \ carrier R" "z\<^sub>2 \ carrier R" assumes in_curve: "xyz_to_pnt (x\<^sub>1, y\<^sub>1, z\<^sub>1) \ carrier curve" "xyz_to_pnt (x\<^sub>2, y\<^sub>2, z\<^sub>2) \ carrier curve" assumes ne: "xyz_to_pnt (x\<^sub>1, y\<^sub>1, z\<^sub>1) \ xyz_to_pnt (x\<^sub>2, y\<^sub>2, z\<^sub>2)" and [simp]: "z\<^sub>1 \ \" "z\<^sub>2 \ \" shows "xyz_to_pnt (x\<^sub>1, y\<^sub>1, z\<^sub>1) \\<^bsub>curve\<^esub> xyz_to_pnt (x\<^sub>2, y\<^sub>2, z\<^sub>2) = (let d\<^sub>1 = x\<^sub>2 \ z\<^sub>1; d\<^sub>2 = x\<^sub>1 \ z\<^sub>2; l = d\<^sub>1 \ d\<^sub>2; m = y\<^sub>2 \ z\<^sub>1 \ y\<^sub>1 \ z\<^sub>2; h = m [^] (2::nat) \ z\<^sub>1 \ z\<^sub>2 \ (d\<^sub>1 \ d\<^sub>2) \ l [^] (2::nat) in xyz_to_pnt (l \ h, (d\<^sub>2 \ l [^] (2::nat) \ h) \ m \ l [^] (3::nat) \ y\<^sub>1 \ z\<^sub>2, l [^] (3::nat) \ z\<^sub>1 \ z\<^sub>2))" (is "?l = ?r") proof - from ne have *: "\ (y\<^sub>2 \ z\<^sub>1 = y\<^sub>1 \ z\<^sub>2 \ x\<^sub>2 \ z\<^sub>1 = x\<^sub>1 \ z\<^sub>2)" by (auto simp add: xyz_to_pnt_eq simp flip: mult_eq_iff_eq_divide eq_mult_iff_divide_eq) have "?l = make_affine (padd A (x\<^sub>1, y\<^sub>1, z\<^sub>1) (x\<^sub>2, y\<^sub>2, z\<^sub>2))" using in_curve by (simp add: padd_correct[where b=B] in_carrierp_def curve_simps xyz_to_pnt_eq_make_affine on_curvep_iff_on_curve) also have "\ = ?r" using * by (auto simp add: padd_def Let_def xyz_to_pnt_eq_make_affine simp flip: one_curve) finally show ?thesis . qed lemma xyz_to_pnt_padd: assumes [simp]: "x\<^sub>1 \ carrier R" "y\<^sub>1 \ carrier R" "z\<^sub>1 \ carrier R" assumes [simp]: "x\<^sub>2 \ carrier R" "y\<^sub>2 \ carrier R" "z\<^sub>2 \ carrier R" assumes in_curve: "xyz_to_pnt (x\<^sub>1, y\<^sub>1, z\<^sub>1) \ carrier curve" "xyz_to_pnt (x\<^sub>2, y\<^sub>2, z\<^sub>2) \ carrier curve" shows "xyz_to_pnt (padd A (x\<^sub>1, y\<^sub>1, z\<^sub>1) (x\<^sub>2, y\<^sub>2, z\<^sub>2)) = xyz_to_pnt (x\<^sub>1, y\<^sub>1, z\<^sub>1) \\<^bsub>curve\<^esub> xyz_to_pnt (x\<^sub>2, y\<^sub>2, z\<^sub>2)" apply (subst (1 2) xyz_to_pnt_eq_make_affine) apply (simp_all add: curve_simps) apply (subst padd_correct[symmetric, OF AB_in_carrier]) using in_curve apply (simp_all add: xyz_to_pnt_eq_make_affine padd_def pdouble_def Let_def curve_simps on_curvep_iff_on_curve in_carrierp_def) done lemma Point_Point_eq_one_iff: assumes 1: "Point x1 y1 \ carrier curve" and 2: "Point x2 y2 \ carrier curve" shows "Point x1 y1 \\<^bsub>curve\<^esub> Point x2 y2 = \\<^bsub>curve\<^esub> \ (x1 = x2 \ y1 = \ y2)" by (subst curve.m_one_iff[OF 1 2]) (simp add: inv_curve 1 2 opp_def) lemma y_coord_eq_or_eq_neg: "Point x y1 \ carrier curve \ Point x y2 \ carrier curve \ y1 = y2 \ y1 = \ y2" by (auto simp: in_carrier_curve on_curve_def eq_commute[of _ "_ \ _"] power2_eq_square square_eq_iff) lemma yz_coord_eq_or_eq_neg: assumes [simp]: "x \ carrier R" "y1 \ carrier R" "y2 \ carrier R" "z \ carrier R" shows "xyz_to_pnt (x, y1, z) \ carrier curve \ xyz_to_pnt (x, y2, z) \ carrier curve \ z \ \\<^bsub>R\<^esub> \ y1 = y2 \ y1 = \ y2" using y_coord_eq_or_eq_neg[of "x \ z" "y1 \ z" "y2 \ z"] by (auto simp add: xyz_to_pnt_eq div_cancel) lemma xyz_coord_eq_or_eq_neg: assumes [simp]: "x1 \ carrier R" "y1 \ carrier R" "z1 \ carrier R" "x2 \ carrier R" "y2 \ carrier R" "z2 \ carrier R" assumes 1: "xyz_to_pnt (x1, y1, z1) \ carrier curve" assumes 2: "xyz_to_pnt (x2, y2, z2) \ carrier curve" and [simp]: "z1 \ \" "z2 \ \" and x: "x1 \ z2 = x2 \ z1" shows "y1 \ z2 = y2 \ z1 \ y1 \ z2 = \ y2 \ z1" apply (subst (1 2) divide_eq_divide_iff_mult_eq_mult[symmetric]) apply (simp_all flip: minus_divide_left) using 1 2 x y_coord_eq_or_eq_neg[of "x1 \ z1" "y1 \ z1" "y2 \ z2"] apply (subst (asm) divide_eq_divide_iff_mult_eq_mult[symmetric]) apply (auto simp: xyz_to_pnt_eq split: if_splits) done end (* of ell_prime_field locale *) section \Added for SEC1\ text \This section adds more facts about an elliptic curve as a group. For example, we examine the order of points on the curve and the order of the curve itself.\ context ell_prime_field begin lemma multEqPowInCurve: assumes "on_curve A B P" shows "point_mult A x P = P [^]\<^bsub>curve\<^esub> x" proof (induct x) case 0 then show ?case by (simp add: one_curve) next case (Suc x) then show ?case using add_curve curve.nat_pow_Suc2 in_carrier_curve assms point_mult.simps(2) by presburger qed text \If P is not the point at infinity, n is a prime, and nP = the point at infinity, then n is actually the order of P. So if 0 < x < n, xP \ the point at infinity.\ lemma order_n: assumes "x < n" "0 < x" "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "point_mult A x P \ Infinity" proof - have 1: "gcd n x = 1" by (metis assms(1,2,6) gcd_dvd1 gcd_le2_nat neq0_conv not_le prime_nat_iff) obtain i and j where 2: "i*(int n) + j*(int x) = 1" by (metis bezw_aux 1 int_ops(2)) have 3: "P = P [^]\<^bsub>curve\<^esub>(i*(int n) + j*(int x))" by (simp add: 2 assms(3) in_carrier_curve) have 4: "P = P [^]\<^bsub>curve\<^esub>(i*(int n)) \\<^bsub>curve\<^esub> P [^]\<^bsub>curve\<^esub>(j*(int x))" using assms(3) in_carrier_curve 3 curve.int_pow_mult by auto have 5: "P = (P [^]\<^bsub>curve\<^esub>(int n))[^]\<^bsub>curve\<^esub>i \\<^bsub>curve\<^esub> (P [^]\<^bsub>curve\<^esub>(int x))[^]\<^bsub>curve\<^esub>j" by (metis assms(3) in_carrier_curve 4 curve.int_pow_pow mult_of_nat_commute) have 6: "P = (P [^]\<^bsub>curve\<^esub>(int x))[^]\<^bsub>curve\<^esub>j" by (metis assms(3,5) in_carrier_curve 5 curve.int_pow_closed curve.int_pow_one one_curve curve.l_one int_pow_int multEqPowInCurve) have 7: "P [^]\<^bsub>curve\<^esub> x = \\<^bsub>curve\<^esub> \ P = \\<^bsub>curve\<^esub>" by (metis 6 curve.int_pow_one int_pow_int) show ?thesis using assms(3,4) 7 one_curve multEqPowInCurve by auto qed lemma order_n': assumes "x < n" "0 < x" "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "P [^]\<^bsub>curve\<^esub> x \ \\<^bsub>curve\<^esub>" using assms order_n one_curve multEqPowInCurve in_carrier_curve by algebra text \The idea for the next two lemmas is that every point on the cycle of order n are the "same." So if you start at one point on the cycle (not Infinity), then you can get to any other point on the cycle by multiplying by some scalar x. Then the point you land on will also be on the curve, have order n, and not be infinity (as long as n does not divide x.)\ lemma order_n_cycle: assumes "on_curve A B P" "point_mult A n P = Infinity" "Q = point_mult A x P" shows "point_mult A n Q = Infinity" by (metis AB_in_carrier(1,2) assms(1,2,3) curve.int_pow_one curve.int_pow_pow curve_simps(1) int_pow_int mult.commute multEqPowInCurve one_curve point_mult_closed) lemma order_n_cycle': assumes "x < n" "0 < x" "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" "Q = point_mult A x P" shows "on_curve A B Q \ Q \ Infinity \ point_mult A n Q = Infinity" using AB_in_carrier(1,2) assms order_n order_n_cycle point_mult_closed by presburger lemma multQmodn: assumes "on_curve A B Q" "point_mult A n Q = Infinity" shows "point_mult A x Q = point_mult A (x mod n) Q" proof - let ?d = "x div n" let ?m = "x mod n" have 1: "x = ?d*n + ?m" using div_mod_decomp by blast have 2: "point_mult A x Q = Q [^]\<^bsub>curve\<^esub> (?d*n + ?m)" using in_carrier_curve assms(1) multEqPowInCurve by presburger have 3: "Q [^]\<^bsub>curve\<^esub> (?d*n + ?m) = Q [^]\<^bsub>curve\<^esub>(?d*n)\\<^bsub>curve\<^esub>Q [^]\<^bsub>curve\<^esub>?m" using curve.nat_pow_mult assms(1) in_carrier_curve by presburger have 4: "Q [^]\<^bsub>curve\<^esub>(?d*n) = (Q [^]\<^bsub>curve\<^esub>n)[^]\<^bsub>curve\<^esub>?d" by (simp add: curve.nat_pow_pow assms(1) in_carrier_curve mult.commute) have 5: "Q [^]\<^bsub>curve\<^esub>(?d*n) = \\<^bsub>curve\<^esub>" using 4 assms(1,2) curve.nat_pow_one one_curve multEqPowInCurve by algebra show ?thesis using 3 5 multEqPowInCurve curve.nat_pow_closed in_carrier_curve assms(1) by force qed lemma multQmodn': assumes "on_curve A B Q" "point_mult A n Q = Infinity" shows "Q[^]\<^bsub>curve\<^esub>x = Q[^]\<^bsub>curve\<^esub>(x mod n)" by (metis assms(1,2) multEqPowInCurve multQmodn) lemma multQmodn'int_pos: assumes "on_curve A B Q" "point_mult A n Q = Infinity" "0 \ (x::int)" shows "Q[^]\<^bsub>curve\<^esub>x = Q[^]\<^bsub>curve\<^esub>(x mod n)" by (metis assms int_pow_int multQmodn' zero_le_imp_eq_int zmod_int) lemma multQmodn'int_neg: assumes "on_curve A B Q" "point_mult A n Q = Infinity" "(x::int) < 0" "1 < n" shows "Q[^]\<^bsub>curve\<^esub>(x::int) = Q[^]\<^bsub>curve\<^esub>(x mod n)" proof - let ?y = "-x" have 1: "Q[^]\<^bsub>curve\<^esub>(?y*n) = \\<^bsub>curve\<^esub>" by (metis assms(1,2) curve.int_pow_one curve.int_pow_pow one_curve int_pow_int multEqPowInCurve mult_of_nat_commute in_carrier_curve) have 2: "0 \ x + ?y*n" using assms(3,4) by auto have 3: "x mod n = (x + ?y*n) mod n" by presburger have 4: "Q[^]\<^bsub>curve\<^esub>(x + ?y*n) = Q[^]\<^bsub>curve\<^esub>((x + ?y*n) mod n)" using 2 assms(1,2) multQmodn'int_pos by fast have 5: "Q[^]\<^bsub>curve\<^esub>(x + ?y*n) = Q[^]\<^bsub>curve\<^esub>x \\<^bsub>curve\<^esub> Q[^]\<^bsub>curve\<^esub>(?y*n)" using curve.int_pow_mult assms(1) in_carrier_curve by presburger have 6: "Q[^]\<^bsub>curve\<^esub>(x + ?y*n) = Q[^]\<^bsub>curve\<^esub>x" using 1 5 assms(1) in_carrier_curve by force show ?thesis using 3 4 6 by argo qed lemma multQmodn'int: assumes "on_curve A B Q" "point_mult A n Q = Infinity" "1 < n" shows "Q[^]\<^bsub>curve\<^esub>(x::int) = Q[^]\<^bsub>curve\<^esub>(x mod n)" apply (cases "0 \ x") using assms(1,2) multQmodn'int_pos apply fast using assms(1,2,3) multQmodn'int_neg by simp text \We use the above to relate to the definition of "ord" for a group.\ lemma curve_ord_n1: assumes "on_curve A B P" "point_mult A n P = Infinity" "n dvd x" shows "point_mult A x P = Infinity" by (metis assms(1,2,3) dvd_eq_mod_eq_0 multQmodn point_mult.simps(1)) lemma curve_ord_n2: assumes "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" "\ n dvd x" shows "point_mult A x P \ Infinity" proof - let ?m = "x mod n" have "0 < ?m \ ?m < n" by (simp add: assms(4,5) mod_greater_zero_iff_not_dvd prime_gt_0_nat) then show ?thesis by (metis assms(1,2,3,4) order_n multQmodn) qed lemma curve_ord_n3: assumes "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "(n dvd x) = (point_mult A x P = Infinity)" by (meson assms(1,2,3,4) curve_ord_n1 curve_ord_n2) lemma curve_ord_n4: assumes "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "(n dvd x) = (P[^]\<^bsub>curve\<^esub>x = \\<^bsub>curve\<^esub>)" using assms curve_ord_n3 multEqPowInCurve one_curve by presburger lemma curve_ord_n5: assumes "on_curve A B P" "P \ Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "curve.ord P = n" using assms curve_ord_n4 curve.ord_def curve.ord_unique in_carrier_curve by presburger lemma curve_ord_n6: assumes "on_curve A B P" "P \ \\<^bsub>curve\<^esub>" "point_mult A n P = \\<^bsub>curve\<^esub>" "prime (n::nat)" "d1 < n" "d2 < n" "d1 < d2" shows "P [^]\<^bsub>curve\<^esub> d1 \ P [^]\<^bsub>curve\<^esub> d2" proof - let ?d = "d2 - d1" have d1: "0 < ?d \ ?d < n" using assms(5,6,7) by linarith have d2: "\ n dvd ?d" using d1 by auto have A1: "(P [^]\<^bsub>curve\<^esub> d1 = P [^]\<^bsub>curve\<^esub> d2) \ P [^]\<^bsub>curve\<^esub> ?d = \\<^bsub>curve\<^esub>" using assms(1) curve.pow_eq_div2 curve_simps(1) by presburger show ?thesis by (metis d2 A1 assms(1,2,3,4) curve_ord_n4 curve_simps(3)) qed lemma curve_ord_n7: assumes "on_curve A B P" "P \ \\<^bsub>curve\<^esub>" "point_mult A n P = \\<^bsub>curve\<^esub>" "prime (n::nat)" "d1 < n" "d2 < n" "d1 \ d2" shows "P [^]\<^bsub>curve\<^esub> d1 \ P [^]\<^bsub>curve\<^esub> d2" apply (cases "d1 < d2") using assms curve_ord_n6 apply blast by (metis assms curve_ord_n6 verit_comp_simplify(3) verit_la_disequality) lemma curve_cycle_n1: assumes "on_curve A B P" "P \ \\<^bsub>curve\<^esub>" "point_mult A n P = \\<^bsub>curve\<^esub>" "prime (n::nat)" shows "card {Q. (\dcurve\<^esub> d)} = n" proof - let ?S1 = "{d. d Infinity" "point_mult A n P = Infinity" "prime (n::nat)" shows "card {Q. (\dAdditions to Elliptic_test\ text \Elliptic_Test only defines addition and scalar multiplication for integer points that are in projective form. We want to have these defined for affine (integer) points. This becomes important when we are running test vectors.\ context residues_prime_gt2 begin definition point_madd :: "nat \ int point \ int point \ int point" where "point_madd a p\<^sub>1 p\<^sub>2 = (case p\<^sub>1 of Infinity \ p\<^sub>2 | Point x\<^sub>1 y\<^sub>1 \ (case p\<^sub>2 of Infinity \ p\<^sub>1 | Point x\<^sub>2 y\<^sub>2 \ if x\<^sub>1 = x\<^sub>2 then if (y\<^sub>1 = - y\<^sub>2 mod (int p)) then Infinity else let twoy1 = nat ((2*y\<^sub>1) mod (int p)); inv_2y1 = int (inv_mod p twoy1); l = ((3 * x\<^sub>1^2 + (int a)) * inv_2y1) mod p; x\<^sub>3 = (l^2 - 2*x\<^sub>1) mod (int p); y\<^sub>3 = (- y\<^sub>1 - l * (x\<^sub>3 - x\<^sub>1)) mod (int p) in Point x\<^sub>3 y\<^sub>3 else let x2_x1 = nat ((x\<^sub>2 - x\<^sub>1) mod (int p)); inv_x2_x1 = int (inv_mod p x2_x1); l = ((y\<^sub>2 - y\<^sub>1) * inv_x2_x1) mod (int p); x\<^sub>3 = (l^2 - x\<^sub>1 - x\<^sub>2) mod (int p); y\<^sub>3 = (- y\<^sub>1 - l * (x\<^sub>3 - x\<^sub>1)) mod (int p) in Point x\<^sub>3 y\<^sub>3 ) )" fun point_mmult :: "nat \ nat \ int point \ int point" where "point_mmult a 0 P = Infinity" | "point_mmult a (Suc n) P = point_madd a P (point_mmult a n P)" lemma add_eq_h1: assumes "p\<^sub>1 = Point x1 y1" "p\<^sub>2 = Point x2 y2" "x1 = x2" "y1 \ \\<^bsub>R\<^esub> y2" "l = (\3\ \\<^bsub>R\<^esub> x1 [^] (2::nat) \\<^bsub>R\<^esub> a) \\<^bsub>R\<^esub> (\2\ \\<^bsub>R\<^esub> y1)" "x3 = l [^] (2::nat) \\<^bsub>R\<^esub> \2\ \\<^bsub>R\<^esub> x1" "y3 = (\\<^bsub>R\<^esub> y1 \\<^bsub>R\<^esub> l \\<^bsub>R\<^esub> (x3 \\<^bsub>R\<^esub> x1))" shows "add a p\<^sub>1 p\<^sub>2 = Point x3 y3" unfolding add_def Let_def using assms point.distinct by simp lemma add_eq_h2: assumes "p\<^sub>1 = Point x1 y1" "p\<^sub>2 = Point x2 y2" "x1 = x2" "(y1 \ - y2 mod (int p))" "twoy1 = nat ((2*y1) mod (int p))" "inv_2y1 = int (inv_mod p twoy1)" "l = ((3 * x1^2 + (int a)) * inv_2y1) mod p" "x3 = (l^2 - 2*x1) mod (int p)" "y3 = (- y1 - l * (x3 - x1)) mod (int p)" - shows "point_madd a p\<^sub>1 p\<^sub>2 = Point x3 y3" - unfolding point_madd_def Let_def using assms point.distinct by simp + shows "point_madd a p\<^sub>1 p\<^sub>2 = Point x3 y3" + by (simp add: assms(1-2) assms(4) assms(8-9) point_madd_def Let_def + flip: assms(3) assms(5) assms(6) assms(7)) lemma point_add_eq [code]: "add a p\<^sub>1 p\<^sub>2 = point_madd a p\<^sub>1 p\<^sub>2" proof (cases p\<^sub>1) case Infinity then show ?thesis by (simp add: point_madd_def add_0_l) next case P1: (Point x1 y1) show ?thesis proof (cases p\<^sub>2) case Infinity then show ?thesis using P1 point_madd_def add_0_r by auto next case P2: (Point x2 y2) then show ?thesis proof (cases "x1 = x2") case x_eq: True then show ?thesis proof (cases "y1 = \\<^bsub>R\<^esub> y2") case y_neg: True have A1: "(y1 = - y2 mod (int p))" using res_neg_eq y_neg by blast have A2: "add a p\<^sub>1 p\<^sub>2 = Infinity" using add_def P1 P2 x_eq y_neg by fastforce have A3: "point_madd a p\<^sub>1 p\<^sub>2 = Infinity" using point_madd_def P1 P2 x_eq A1 by force show ?thesis using A2 A3 by presburger next case y_nneg: False have B0: "(y1 \ - y2 mod (int p))" using y_nneg res_neg_eq by algebra let ?twoy1 = "((2*y1) mod (int p))" let ?inv_2y1 = "int (inv_mod p (nat ?twoy1))" let ?l = "((3 * x1^2 + (int a)) * ?inv_2y1) mod p" let ?l' = "(\3\ \\<^bsub>R\<^esub> x1 [^] (2::nat) \\<^bsub>R\<^esub> a) \\<^bsub>R\<^esub> (\2\ \\<^bsub>R\<^esub> y1)" have B1: "\2\ \\<^bsub>R\<^esub> y1 = (2*y1) mod (int p)" by (simp add: mod_mult_right_eq mult.commute res_mult_eq res_of_integer_eq) have B2: "0 \ (2*y1) mod (int p)" using gt2 by auto have B3: "(\3\ \\<^bsub>R\<^esub> x1 [^] (2::nat) \\<^bsub>R\<^esub> a) = (3 * x1^2 + (int a)) mod p" by (metis mod_mod_trivial res_add_eq local.of_int_add local.of_int_mult res_of_integer_eq res_pow_eq) have B4: "?inv_2y1 mod p = ?inv_2y1" by (metis inv_mod_def mod_mod_trivial zmod_int) have B5: "?l = (((3 * x1^2 + (int a)) mod p) * ?inv_2y1) mod p" by (metis mod_mult_left_eq) have B6: "?l = ((3 * x1^2 + (int a)) mod p) \\<^bsub>R\<^esub> ?twoy1" by (smt (verit, ccfv_SIG) B2 B4 R_def cring_class.of_nat_0 gt2 int_distrib(4) int_nat_eq inv_mod_0' m_div_def m_gt_one mod_in_carrier nat_int of_nat_0_less_iff res_to_cong_simps(2) res_zero_eq residues_prime.inv_mod_p_inv_in_ring residues_prime.p_prime residues_prime_axioms) have B7: "?l = (\3\ \\<^bsub>R\<^esub> x1 [^] (2::nat) \\<^bsub>R\<^esub> a) \\<^bsub>R\<^esub> ?twoy1" using B6 B3 by argo have B8: "?l = ?l'" using B7 B1 by argo let ?x3 = "(?l^2 - 2*x1) mod (int p)" let ?x3' = "?l' [^] (2::nat) \\<^bsub>R\<^esub> \2\ \\<^bsub>R\<^esub> x1" have B9: "?x3 = ?x3'" by (metis B8 of_int_diff mod_mult_left_eq res_mult_eq res_of_integer_eq res_pow_eq) let ?y3 = "(- y1 - ?l * (?x3 - x1)) mod (int p)" let ?y3' = "(\\<^bsub>R\<^esub> y1 \\<^bsub>R\<^esub> ?l' \\<^bsub>R\<^esub> (?x3 \\<^bsub>R\<^esub> x1))" have B10: "?y3 = ?y3'" by (smt (verit) B8 B9 mod_diff_eq mod_mod_trivial mult_cong res_diff_eq res_neg_eq) have B11: "add a p\<^sub>1 p\<^sub>2 = Point ?x3 ?y3" using y_nneg add_eq_h1 B9 B10 P1 P2 x_eq by blast have B12: "point_madd a p\<^sub>1 p\<^sub>2 = Point ?x3 ?y3" using B0 add_eq_h2 B9 B10 P1 P2 x_eq by algebra show ?thesis using B11 B12 by argo qed next case x_neq: False let ?x2_x1 = "((x2 - x1) mod (int p))" let ?inv_x2_x1 = "int (inv_mod p (nat ?x2_x1))" let ?l = "((y2 - y1) * ?inv_x2_x1) mod (int p)" let ?x3 = "(?l^2 - x1 - x2) mod (int p)" let ?y3 = "(- y1 - ?l * (?x3 - x1)) mod (int p)" have C1: "point_madd a p\<^sub>1 p\<^sub>2 = Point ?x3 ?y3" unfolding point_madd_def Let_def using x_neq P1 P2 by auto let ?l' = "(y2 \\<^bsub>R\<^esub> y1) \\<^bsub>R\<^esub> (x2 \\<^bsub>R\<^esub> x1)" let ?x3' = "?l' [^] (2::nat) \\<^bsub>R\<^esub> x1 \\<^bsub>R\<^esub> x2" let ?y3' = "(\\<^bsub>R\<^esub> y1 \\<^bsub>R\<^esub> ?l' \\<^bsub>R\<^esub> (?x3' \\<^bsub>R\<^esub> x1))" have C2: "add a p\<^sub>1 p\<^sub>2 = Point ?x3' ?y3'" unfolding add_def Let_def using x_neq P1 P2 by auto have C3: "?l = ?l'" by (smt (verit, ccfv_threshold) Euclidean_Rings.pos_mod_sign R_m_def gt2 int_nat_eq int_ops(1) integral_iff inv_mod_0' m_div_def mod_in_carrier mod_mod_trivial mult_cong nat_int nat_less_iff res_diff_eq residues.res_mult_eq residues_axioms zero_cong residues_prime.inv_mod_p_inv_in_ring residues_prime.p_prime residues_prime_axioms) have C4: "?x3 = ?x3'" by (simp add: C3 mod_diff_left_eq res_diff_eq res_pow_eq) have C5: "?y3 = ?y3'" by (smt (verit) C3 C4 mod_diff_eq mod_mod_trivial mult_cong res_diff_eq res_neg_eq) show ?thesis using C1 C2 C4 C5 by argo qed qed qed lemma point_mult_eq [code]: "point_mult a n P = point_mmult a n P" apply (induct n) apply simp using point_add_eq by force text \Also in Elliptic_Test, they provide methods for converting a point in projective coordinates to a point in affine coordinates, but not the other way around. Here we provide a few more tools for going between affine and projective coordinates.\ definition get_x :: "int point \ int" where "get_x P = (case P of Infinity \ 0 | Point x y \ x)" definition get_y :: "int point \ int" where "get_y P = (case P of Infinity \ 0 | Point x y \ y)" lemma get_coords_correct: assumes "P = Point x y" shows "P = Point (get_x P) (get_y P)" by (simp add: assms get_x_def get_y_def) definition mmake_proj :: "int point \ int ppoint" where "mmake_proj P = (if P = Infinity then (1,1,0) else (get_x P, get_y P, 1))" lemma bezout_coeff_1: assumes "1 < (x::nat)" shows "bezout_coefficients 1 x = (1,0)" proof - have 10: "1 div x = 0" using assms(1) by fastforce have 11: "x \ 0" using assms(1) by fastforce have 12: "euclid_ext_aux 1 0 0 1 1 x = euclid_ext_aux 0 (1 - 0 * 0) 1 (0 - 0 * 1) x (1 mod x)" by (smt (verit, ccfv_threshold) 10 11 euclid_ext_aux.simps mult_cancel_left1 of_nat_1 of_nat_eq_0_iff of_nat_mod unique_euclidean_semiring_with_nat_class.of_nat_div) have 1: "euclid_ext_aux 1 0 0 1 1 x = euclid_ext_aux 0 1 1 0 x 1" using 12 assms(1) by force have 20: "x div 1 = x" by simp have 21: "(1::int) \ 0" by simp have 22: "euclid_ext_aux 0 1 1 0 x 1 = euclid_ext_aux 1 (0 - x * 1) 0 (1 - x * 0) 1 (x mod 1)" by (smt (verit, ccfv_threshold)20 21 euclid_ext_aux.simps mult_cancel_left1 of_nat_1 of_nat_eq_0_iff of_nat_mod unique_euclidean_semiring_with_nat_class.of_nat_div mod_by_1) have 2: "euclid_ext_aux 0 1 1 0 x 1 = euclid_ext_aux 1 (0 - x) 0 1 1 0" using 22 assms(1) by simp have 3: "euclid_ext_aux 1 (0 - x) 0 1 1 0 = ((1,0), 1)" by (metis (no_types, lifting) euclid_ext_aux.simps div_by_1 mult_1 mult_eq_0_iff normalize_div unit_factor_1) have 4: "euclid_ext 1 x = ((1,0), 1)" using 1 2 3 by presburger show ?thesis using 4 by simp qed lemma to_proj_to_aff: assumes "0 \ get_x P" "get_x P < p" "0 \ get_y P" "get_y P < p" shows "mmake_affine (int p) (mmake_proj P) = P" proof (cases P) case Infinity then show ?thesis using mmake_proj_def mmake_affine_def by simp next case C: (Point x y) have Cx: "x = get_x P" by (simp add: C get_x_def) have Cy: "y = get_y P" by (simp add: C get_y_def) have C1: "mmake_proj P = (x,y,1)" using C Cx Cy mmake_proj_def by fastforce have C2: "bezout_coefficients 1 (int p) = (1,0)" using bezout_coeff_1 m_gt_one by fastforce have C3: "1 **\<^bsub>(int p)\<^esub> x = x" using assms(1,2) mmult_def Cx by simp have C4: "1 **\<^bsub>(int p)\<^esub> y = y" using assms(3,4) mmult_def Cy by simp have C5: "mmake_affine (int p) (x,y,1) = Point x y" using mmake_affine_def C2 C3 C4 by force show ?thesis using C C1 C5 by presburger qed end (* residues_prime_gt2 context *) end diff --git a/thys/Crypto_Standards/FIPS180_4.thy b/thys/Crypto_Standards/FIPS180_4.thy --- a/thys/Crypto_Standards/FIPS180_4.thy +++ b/thys/Crypto_Standards/FIPS180_4.thy @@ -1,2231 +1,2230 @@ theory FIPS180_4 imports Words begin text \ https://nvlpubs.nist.gov/nistpubs/FIPS/NIST.FIPS.180-4.pdf In this theory, we translate the NIST standard 180-4, the Secure Hash Standard, to Isabelle. We aim to adhere as closely to the written standard as possible, including function and variable names and overall document structure. It should be clear to any reader, regardless of their experience with Isabelle, that this exactly matches the standard. We will include direct quotations from the standard to aid the reader. We have translated the Secure Hash Standard so that each of the Secure Hash Algorithms is a function from natural numbers to natural numbers. This allows us the flexibility to apply the SHA definition to any implementation, whether that implementation acts on natural numbers, bit strings, octet (i.e., byte) strings, or (generically) n-bit word strings. See the final section below, called Octets, for versions of the secure hash algorithms as functions on octets. "This Standard specifies secure hash algorithms, SHA-1, SHA-224, SHA-256, SHA-384, SHA-512, SHA-512/224 and SHA-512/256. All of the algorithms are iterative, one-way hash functions that can process a message to produce a condensed representation called a message digest. These algorithms enable the determination of a message's integrity: any change to the message will, with a very high probability, result in a different message digest. This property is useful in the generation and verification of digital signatures and message authentication codes, and in the generation of random numbers or bits. Each algorithm can be described in two stages: preprocessing and hash computation. Preprocessing involves padding a message, parsing the padded message into m-bit blocks, and setting initialization values to be used in the hash computation. The hash computation generates a message schedule from the padded message and uses that schedule, along with functions, constants, and word operations to iteratively generate a series of hash values. The final hash value generated by the hash computation is used to determine the message digest. The algorithms differ most significantly in the security strengths that are provided for the data being hashed. The security strengths of these hash functions and the system as a whole when each of them is used with other cryptographic algorithms, such as digital signature algorithms and keyed-hash message authentication codes, can be found in [SP 800-57] and [SP 800-107]. Additionally, the algorithms differ in terms of the size of the blocks and words of data that are used during hashing or message digest sizes. Figure 1 presents the basic properties of these hash algorithms. Algorithm Message Size Block Size Word Size Message Digest Size SHA-1 < 2^64 512 32 160 SHA-224 < 2^64 512 32 224 SHA-256 < 2^64 512 32 256 SHA-384 < 2^128 1024 64 384 SHA-512 < 2^128 1024 64 512 SHA-512/224 < 2^128 1024 64 224 SHA-512/256 < 2^128 1024 64 256 "\ section \2. Definitions\ text \"Left-shift operation, where x << n is obtained by discarding the left-most n bits of the word x and then padding the result with n zeroes on the right." So we need to know the word size w for x in order to discard the top n out of w bits. For this standard, the word size is either w = 32 or w = 64.\ definition SHL :: "nat \ nat \ nat \ nat" where "SHL w n x = (x mod 2^(w-n)) * 2^n" abbreviation SHL32 :: "nat \ nat \ nat" where "SHL32 n x \ SHL 32 n x" abbreviation SHL64 :: "nat \ nat \ nat" where "SHL64 n x \ SHL 64 n x" lemma SHL_0: assumes "x < 2^w" shows "SHL w 0 x = x" by (simp add: SHL_def assms) lemma SHL_w: "SHL w w x = 0" by (simp add: SHL_def) lemma SHL_n_ge_w: assumes "w \ n" shows "SHL w n x = 0" by (simp add: SHL_def assms) lemma SHL_bnd: assumes "n < w" shows "SHL w n x < 2^w" - by (metis SHL_def assms less_imp_le mult_exp_mod_exp_eq pos_mod_bound zero_less_numeral - zero_less_power) + using assms by (simp add: SHL_def flip: mult_exp_mod_exp_eq) lemma SHL_mod: "(SHL w n x) mod 2^n = 0" by (simp add: SHL_def) lemma SHL_div: "(SHL w n x) div 2^n < 2^(w-n)" by (simp add: SHL_def) text \"Right-shift operation, where x >> n is obtained by discarding the right-most n bits of the word x and then padding the result with n zeroes on the left." Here we do not need to know the word size w.\ definition SHR :: "nat \ nat \ nat" where "SHR n x = x div 2^n" lemma SHR_0: "SHR 0 x = x" by (simp add: SHR_def) lemma SHR_w: assumes "x < 2^w" shows "SHR w x = 0" by (simp add: SHR_def assms) lemma SHR_n_ge_w: assumes "x < 2^w" "w \ n" shows "SHR n x = 0" by (metis SHR_w assms div_less dual_order.strict_trans2 le_less_linear power_diff power_not_zero zero_neq_numeral) lemma SHR_less: "SHR n x \ x" by (simp add: SHR_def) lemma SHR_bnd1: assumes "x < 2^w" shows "SHR n x < 2^w" by (meson SHR_less assms dual_order.trans leD leI) lemma SHR_bnd2: assumes "x \ 2^w" shows "SHR n x \ 2^(w-n)" proof - have 1: "x div 2^n \ 2^w div 2^n" using assms(1) div_le_mono by presburger have 2: "(2::nat)^w div 2^n \ 2^(w-n)" by (simp add: exp_div_exp_eq) show ?thesis using SHR_def 1 2 by auto qed text \"The rotate left (circular left shift) operation, where x is a w-bit word and n is an integer with 0 <= n < w, is defined b ROTL n x =(x << n) \/ (x >> w - n)." \ definition ROTL :: "nat \ nat \ nat \ nat" where "ROTL w n x = (SHL w n x) + (SHR (w-n) x)" abbreviation ROTL32 :: "nat \ nat \ nat" where "ROTL32 n x \ ROTL 32 n x" abbreviation ROTL64 :: "nat \ nat \ nat" where "ROTL64 n x \ ROTL 64 n x" lemma ROTL_0: assumes "x < 2^w" shows "ROTL w 0 x = x" by (simp add: ROTL_def SHL_0 SHR_w assms) lemma ROTL_OR: assumes "x < 2^w" "n < w" shows "ROTL w n x = (SHL w n x) OR (SHR (w-n) x)" proof - have 1: "SHR (w-n) x < 2^n" by (metis SHR_def assms le_add_diff_inverse less_imp_le less_mult_imp_div_less power_add) let ?m = "(x mod 2^(w-n))" have 2: "SHL w n x = ?m * 2^n" using SHL_def by fast have 3: "?m * 2^n + (SHR (w-n) x) = ?m * 2^n OR (SHR (w-n) x)" by (simp add: 1 OR_sum_nat_hilo_2 mult.commute) show ?thesis using 2 3 ROTL_def by presburger qed lemma ROTL_bnd: assumes "n < w" "x < 2^w" shows "ROTL w n x < 2^w" using ROTL_OR SHL_bnd SHR_bnd1 assms nat_OR_upper by presburger text \"The rotate right (circular right shift) operation, where x is a w-bit word and n is an integer with 0 <= n < w, is defined by ROTR n x =(x >> n) \/ (x << w - n)." \ definition ROTR :: "nat \ nat \ nat \ nat" where "ROTR w n x = (SHR n x) + (SHL w (w-n) x)" abbreviation ROTR32 :: "nat \ nat \ nat" where "ROTR32 n x \ ROTR 32 n x" abbreviation ROTR64 :: "nat \ nat \ nat" where "ROTR64 n x \ ROTR 64 n x" lemma ROTL_ROTR_eq: assumes "n \ w" shows "ROTL w n x = ROTR w (w-n) x" using ROTL_def ROTR_def assms by force lemma ROTR_ROTL_eq: assumes "n \ w" shows "ROTR w n x = ROTL w (w-n) x" using ROTL_def ROTR_def assms by force lemma ROTR_0: "ROTR w 0 x = x" by (simp add: ROTR_def SHL_w SHR_0) lemma ROTR_OR: assumes "x < 2^w" "n < w" shows "ROTR w n x = (SHR n x) OR (SHL w (w-n) x)" proof - have 0: "w - (w - n) = n" using assms(2) by fastforce have 1: "ROTR w n x = ROTL w (w-n) x" by (meson ROTR_ROTL_eq assms(2) nat_less_le) have 2: "ROTL w (w-n) x = (SHL w (w-n) x) OR (SHR n x)" by (metis 0 ROTL_OR ROTL_def SHL_w add_cancel_left_left assms(1) diff_is_0_eq' diff_zero not_le or.commute or.right_neutral) show ?thesis using 1 2 or.commute by auto qed lemma ROTR_bnd: assumes "n < w" "x < 2^w" shows "ROTR w n x < 2^w" by (metis ROTL_bnd ROTR_0 ROTR_ROTL_eq assms diff_less less_nat_zero_code less_or_eq_imp_le nat_neq_iff) section \4. Functions and Constants\ subsection \4.1 Functions\ text \NOT for a natural number (nat) with a bit length w or less.\ definition wNOT :: "nat \ nat \ nat" where "wNOT w x = (2^w - 1) XOR x" abbreviation NOT32 :: "nat \ nat" where "NOT32 x \ wNOT 32 x" abbreviation NOT64 :: "nat \ nat" where "NOT64 x \ wNOT 64 x" lemma wNOT_wNOT: "wNOT w (wNOT w x) = x" by (metis wNOT_def nat_xor_inv xor.commute) lemma wNOT_bnd: assumes "x < 2^w" shows "wNOT w x < 2^w" by (simp add: assms nat_XOR_upper wNOT_def) lemma wNOT_XOR: assumes "x < 2^w" shows "x XOR (wNOT w x) = 2^w - 1" using nat_xor_inv wNOT_def by presburger text \Choose\ definition Ch :: "nat \ nat \ nat \ nat \ nat" where "Ch w x y z = (x AND y) XOR ((wNOT w x) AND z)" abbreviation Ch32 :: "nat \ nat \ nat \ nat" where "Ch32 x y z \ Ch 32 x y z" abbreviation Ch64 :: "nat \ nat \ nat \ nat" where "Ch64 x y z \ Ch 64 x y z" lemma Ch_asymm: "Ch w x y z = Ch w (wNOT w x) z y" by (simp add: Ch_def wNOT_wNOT xor.commute) lemma Ch_bnd: assumes "x < 2^w" "y < 2^w" "z < 2^w" shows "Ch w x y z < 2^w" by (simp add: Ch_def and_nat_def assms(1) nat_XOR_upper wNOT_bnd) lemma Ch_fst: assumes "y < 2^w" "z < 2^w" shows "Ch w (2^w-1) y z = y" by (metis Ch_asymm Ch_def and.commute assms(1) mask_nat_def mod_less take_bit_eq_mask take_bit_nat_def wNOT_def xor.right_neutral zero_and_eq) lemma Ch_snd: assumes "y < 2^w" "z < 2^w" shows "Ch w 0 y z = z" by (metis (no_types) Ch_asymm Ch_fst assms wNOT_def xor.comm_neutral) text \Parity\ definition Parity :: "nat \ nat \ nat \ nat" where "Parity x y z = x XOR y XOR z" lemma Parity_symm: "Parity x y z = Parity x z y \ Parity x y z = Parity z y x \ Parity x y z = Parity y x z \ Parity x y z = Parity z x y \ Parity x y z = Parity y z x" by (simp add: Parity_def xor.commute xor.left_commute) lemma Parity_bnd: assumes "x < 2^w" "y < 2^w" "z < 2^w" shows "Parity x y z < 2^w" by (simp add: Parity_def assms nat_XOR_upper) lemma Parity_2same: "Parity x y y = x" using Parity_def Parity_symm nat_xor_inv by presburger lemma Parity_wNOT: assumes "x < 2^w" "y < 2^w" shows "Parity x y (wNOT w y) = wNOT w x" using Parity_def assms(2) wNOT_XOR wNOT_def xor.commute by force text \Majority vote\ definition Maj :: "nat \ nat \ nat \ nat" where "Maj x y z = (x AND y) XOR (x AND z) XOR (y AND z)" lemma Maj_symm: "Maj x y z = Maj x z y \ Maj x y z = Maj z y x \ Maj x y z = Maj y x z \ Maj x y z = Maj z x y \ Maj x y z = Maj y z x" by (simp add: Maj_def and.commute xor.commute xor.left_commute) lemma Maj_bnd: assumes "x < 2^w" "y < 2^w" shows "Maj x y z < 2^w" by (simp add: Maj_def and_nat_def assms nat_XOR_upper) lemma Maj_3same: "Maj x x x = x" by (metis (no_types) Maj_def and.idem nat_xor_inv) lemma Maj_2same: "Maj x y y = y" by (simp add: Maj_def xor_nat_def) subsubsection \4.1.1 SHA-1 Functions\ text \"SHA-1 uses a sequence of logical functions, f0, f1,..., f79. Each function ft, where 0 \ t \ 79, operates on three 32-bit words, x, y, and z, and produces a 32-bit word as output."\ definition SHA1_ft :: "nat \ nat \ nat \ nat \ nat" where "SHA1_ft t x y z = ( if (t \ 19) then (Ch32 x y z) else if (t \ 39) then (Parity x y z) else if (t \ 59) then (Maj x z y) else (Parity x y z) )" lemma SHA1_ft_bnd: assumes "x < 2^32" "y < 2^32" "z < 2^32" shows "SHA1_ft t x y z < 2^32" using Ch_bnd Maj_bnd Parity_bnd SHA1_ft_def assms by presburger subsubsection \4.1.2 SHA-224 and SHA-256 Functions\ text \"SHA-224 and SHA-256 both use six logical functions, where each function operates on 32-bit words, which are represented as x, y, and z. The result of each function is a new 32-bit word."\ definition Sigma256_0 :: "nat \ nat" where "Sigma256_0 x = (ROTR32 2 x) XOR (ROTR32 13 x) XOR (ROTR32 22 x)" definition Sigma256_1 :: "nat \ nat" where "Sigma256_1 x = (ROTR32 6 x) XOR (ROTR32 11 x) XOR (ROTR32 25 x)" definition sigma256_0 :: "nat \ nat" where "sigma256_0 x = (ROTR32 7 x) XOR (ROTR32 18 x) XOR (SHR 3 x)" definition sigma256_1 :: "nat \ nat" where "sigma256_1 x = (ROTR32 17 x) XOR (ROTR32 19 x) XOR (SHR 10 x)" lemma Sigma256s_bnd: assumes "x < 2^32" "a < 32" "b < 32" "c < 32" shows "(ROTR32 a x) XOR (ROTR32 b x) XOR (ROTR32 c x) < 2^32" by (meson ROTR_bnd assms nat_XOR_upper) lemma Sigma256_0_bnd: assumes "x < 2^32" shows "Sigma256_0 x < 2^32" by (metis Sigma256s_bnd Sigma256_0_def assms numeral_less_iff semiring_norm(76,78,81)) lemma Sigma256_1_bnd: assumes "x < 2^32" shows "Sigma256_1 x < 2^32" by (metis Sigma256s_bnd Sigma256_1_def assms numeral_less_iff semiring_norm(76,78,81)) lemma sigma256s_bnd: assumes "x < 2^32" "a < 32" "b < 32" "c < 32" shows "(ROTR32 a x) XOR (ROTR32 b x) XOR (SHR c x) < 2^32" using ROTR_bnd SHR_bnd1 assms nat_XOR_upper by presburger lemma sigma256_0_bnd: assumes "x < 2^32" shows "sigma256_0 x < 2^32" by (metis sigma256s_bnd sigma256_0_def assms numeral_less_iff semiring_norm(76,78,81)) lemma sigma256_1_bnd: assumes "x < 2^32" shows "sigma256_1 x < 2^32" by (metis sigma256s_bnd sigma256_1_def assms numeral_less_iff semiring_norm(76,78,81)) subsubsection \4.1.3 SHA-384, SHA-512, SHA-512/224 and SHA-512/256 Functions\ text \"SHA-384, SHA-512, SHA-512/224 and SHA-512/256 use six logical functions, where each function operates on 64-bit words, which are represented as x, y, and z. The result of each function is a new 64-bit word."\ definition Sigma512_0 :: "nat \ nat" where "Sigma512_0 x = (ROTR64 28 x) XOR (ROTR64 34 x) XOR (ROTR64 39 x)" definition Sigma512_1 :: "nat \ nat" where "Sigma512_1 x = (ROTR64 14 x) XOR (ROTR64 18 x) XOR (ROTR64 41 x)" definition sigma512_0 :: "nat \ nat" where "sigma512_0 x = (ROTR64 1 x) XOR (ROTR64 8 x) XOR (SHR 7 x)" definition sigma512_1 :: "nat \ nat" where "sigma512_1 x = (ROTR64 19 x) XOR (ROTR64 61 x) XOR (SHR 6 x)" lemma Sigma512s_bnd: assumes "x < 2^64" "a < 64" "b < 64" "c < 64" shows "(ROTR64 a x) XOR (ROTR64 b x) XOR (ROTR64 c x) < 2^64" by (meson ROTR_bnd assms nat_XOR_upper) lemma Sigma512_0_bnd: assumes "x < 2^64" shows "Sigma512_0 x < 2^64" using Sigma512s_bnd Sigma512_0_def assms numeral_less_iff semiring_norm(76,78,81) by auto lemma Sigma512_1_bnd: assumes "x < 2^64" shows "Sigma512_1 x < 2^64" using Sigma512s_bnd Sigma512_1_def assms numeral_less_iff semiring_norm(76,78,81) by simp lemma sigma512s_bnd: assumes "x < 2^64" "a < 64" "b < 64" "c < 64" shows "(ROTR64 a x) XOR (ROTR64 b x) XOR (SHR c x) < 2^64" using ROTR_bnd SHR_bnd1 assms nat_XOR_upper by presburger lemma sigma512_0_bnd: assumes "x < 2^64" shows "sigma512_0 x < 2^64" using sigma512s_bnd sigma512_0_def assms numeral_less_iff semiring_norm(76,78,81) by auto lemma sigma512_1_bnd: assumes "x < 2^64" shows "sigma512_1 x < 2^64" by (metis sigma512s_bnd sigma512_1_def assms numeral_less_iff semiring_norm(76,78,81)) subsection \4.2 Constants\ subsubsection \4.2.1 SHA-1 Constants\ text \"SHA-1 uses a sequence of eighty constant 32-bit words, K0, K1,..., K79, which are given by"\ definition SHA1_Kt :: "nat \ nat" where "SHA1_Kt t = ( if (t \ 19) then (0x5a827999) else if (t \ 39) then (0x6ed9eba1) else if (t \ 59) then (0x8f1bbcdc) else (0xca62c1d6) )" lemma SHA1_Kt_bnd: "SHA1_Kt t < 2^32" unfolding SHA1_Kt_def apply (cases t) by auto subsubsection \4.2.2 SHA-224 and SHA-256 Constants\ text \"SHA-224 and SHA-256 use the same sequence of sixty-four constant 32-bit words, K_0^{256}, K_1^{256}, ..., K_63^{256}. These words represent the first thirty-two bits of the fractional parts of the cube roots of the first sixty-four prime numbers. In hex, these constant words are (from left to right)"\ definition K256list :: "nat list" where "K256list = [0x428a2f98, 0x71374491, 0xb5c0fbcf, 0xe9b5dba5, 0x3956c25b, 0x59f111f1, 0x923f82a4, 0xab1c5ed5, 0xd807aa98, 0x12835b01, 0x243185be, 0x550c7dc3, 0x72be5d74, 0x80deb1fe, 0x9bdc06a7, 0xc19bf174, 0xe49b69c1, 0xefbe4786, 0x0fc19dc6, 0x240ca1cc, 0x2de92c6f, 0x4a7484aa, 0x5cb0a9dc, 0x76f988da, 0x983e5152, 0xa831c66d, 0xb00327c8, 0xbf597fc7, 0xc6e00bf3, 0xd5a79147, 0x06ca6351, 0x14292967, 0x27b70a85, 0x2e1b2138, 0x4d2c6dfc, 0x53380d13, 0x650a7354, 0x766a0abb, 0x81c2c92e, 0x92722c85, 0xa2bfe8a1, 0xa81a664b, 0xc24b8b70, 0xc76c51a3, 0xd192e819, 0xd6990624, 0xf40e3585, 0x106aa070, 0x19a4c116, 0x1e376c08, 0x2748774c, 0x34b0bcb5, 0x391c0cb3, 0x4ed8aa4a, 0x5b9cca4f, 0x682e6ff3, 0x748f82ee, 0x78a5636f, 0x84c87814, 0x8cc70208, 0x90befffa, 0xa4506ceb, 0xbef9a3f7, 0xc67178f2]" lemma K256_bnd1: assumes "k \ set K256list" shows "k < 2^32" using assms unfolding K256list_def apply (cases k) by auto lemma K256_bnd2: assumes "i < length K256list" shows "K256list ! i < 2^32" using K256_bnd1 assms nth_mem by blast lemma K256_len [simp]: "length K256list = 64" unfolding K256list_def by simp subsubsection \4.2.3 SHA-384, SHA-512, SHA-512/224 and SHA-512/256 Constants\ text \SHA-384, SHA-512, SHA-512/224 and SHA-512/256 use the same sequence of eighty constant 64-bit words, K_0^{512}, K_1^{512}, ..., K_79^{512}. These words represent the first sixty-four bits of the fractional parts of the cube roots of the first eighty prime numbers. In hex, these constant words are (from left to right)\ definition K512list :: "nat list" where "K512list = [0x428a2f98d728ae22, 0x7137449123ef65cd, 0xb5c0fbcfec4d3b2f, 0xe9b5dba58189dbbc, 0x3956c25bf348b538, 0x59f111f1b605d019, 0x923f82a4af194f9b, 0xab1c5ed5da6d8118, 0xd807aa98a3030242, 0x12835b0145706fbe, 0x243185be4ee4b28c, 0x550c7dc3d5ffb4e2, 0x72be5d74f27b896f, 0x80deb1fe3b1696b1, 0x9bdc06a725c71235, 0xc19bf174cf692694, 0xe49b69c19ef14ad2, 0xefbe4786384f25e3, 0x0fc19dc68b8cd5b5, 0x240ca1cc77ac9c65, 0x2de92c6f592b0275, 0x4a7484aa6ea6e483, 0x5cb0a9dcbd41fbd4, 0x76f988da831153b5, 0x983e5152ee66dfab, 0xa831c66d2db43210, 0xb00327c898fb213f, 0xbf597fc7beef0ee4, 0xc6e00bf33da88fc2, 0xd5a79147930aa725, 0x06ca6351e003826f, 0x142929670a0e6e70, 0x27b70a8546d22ffc, 0x2e1b21385c26c926, 0x4d2c6dfc5ac42aed, 0x53380d139d95b3df, 0x650a73548baf63de, 0x766a0abb3c77b2a8, 0x81c2c92e47edaee6, 0x92722c851482353b, 0xa2bfe8a14cf10364, 0xa81a664bbc423001, 0xc24b8b70d0f89791, 0xc76c51a30654be30, 0xd192e819d6ef5218, 0xd69906245565a910, 0xf40e35855771202a, 0x106aa07032bbd1b8, 0x19a4c116b8d2d0c8, 0x1e376c085141ab53, 0x2748774cdf8eeb99, 0x34b0bcb5e19b48a8, 0x391c0cb3c5c95a63, 0x4ed8aa4ae3418acb, 0x5b9cca4f7763e373, 0x682e6ff3d6b2b8a3, 0x748f82ee5defb2fc, 0x78a5636f43172f60, 0x84c87814a1f0ab72, 0x8cc702081a6439ec, 0x90befffa23631e28, 0xa4506cebde82bde9, 0xbef9a3f7b2c67915, 0xc67178f2e372532b, 0xca273eceea26619c, 0xd186b8c721c0c207, 0xeada7dd6cde0eb1e, 0xf57d4f7fee6ed178, 0x06f067aa72176fba, 0x0a637dc5a2c898a6, 0x113f9804bef90dae, 0x1b710b35131c471b, 0x28db77f523047d84, 0x32caab7b40c72493, 0x3c9ebe0a15c9bebc, 0x431d67c49c100d4c, 0x4cc5d4becb3e42b6, 0x597f299cfc657e2a, 0x5fcb6fab3ad6faec, 0x6c44198c4a475817]" lemma K512_bnd1: assumes "k \ set K512list" shows "k < 2^64" using assms unfolding K512list_def apply (cases k) by auto lemma K512_bnd2: assumes "i < length K512list" shows "K512list ! i < 2^64" using K512_bnd1 assms nth_mem by blast lemma K512_len [simp]: "length K512list = 80" unfolding K512list_def by simp section \5. Preprocessing\ subsection \5.1/2 Padding and Parsing the Message\ text \Note for reasons of convenience, we combine sections 5.1 "Padding the Message" and 5.2 "Parsing the Message" of the FIPS standard into this combined Isabelle section. From section 5.1: "The purpose of this padding is to ensure that the padded message is a multiple of 512 or 1024 bits, depending on the algorithm. Padding can be inserted before hash computation begins on a message, or at any other time during the hash computation prior to processing the block(s) that will contain the padding." The padding for the 32-bit-word SHAs and the padding for the 64-bit-word SHAs are very similar. We define it generically here. From the standard sections 5.1.1 and 5.1.2 only differ in the values of X and Y: "Suppose that the length of the message, M, is l bits. Append the bit ``1'' to the end of the message, followed by k zero bits, where k is the smallest, non-negative solution to the equation l + 1 + k = [X] mod [X+Y]. Then append the [Y]-bit block that is equal to the number l expressed using a binary representation." For us, the message M is viewed as a natural number. We don't know how many leading zero bits are intended in the bit string version of the message. So we need to be told l, the number of bits of the message, where bit_length M <= l. Then section 5.2 states only: "The message and its padding must be parsed into N m-bit blocks." where N is the number of blocks in the padded message and m is (X+Y). For the SHA1 group of hash algorithms, m = X+Y = 512. For the SHA512 group, m = X+Y = 1024. Each block is further parsed into 16 words. For the SHA1 group, the word size is 32 bits, and 32 * 16 = 512. For the SHA512 group, the word size is 64 and 64 * 16 = 1024. So we introduce the variable W in the following locale to represent the word size and the only thing we need to know about W at the moment is that W divides X+Y. Here we skip the step of parsing into blocks. The output of this parsing is a list of W-bit numbers (words). \ locale SHA_PadParse = fixes X Y W :: nat assumes XYWpos : "0 < X" "0 < Y" "0 < W" and Wdvd : "W dvd (X+Y)" begin definition SHApadding_k :: "nat \ nat" where "SHApadding_k l = ( let x = (l + 1) mod (X+Y) in ( if x \ X then X - x else X + (X+Y) - x) )" text \First we implement "padding the message" as natural-number arithmetic.\ definition SHApadded :: "nat \ nat \ nat" where "SHApadded M l = ( let k = (SHApadding_k l) in (2*M + 1)*2^(k+Y) + l )" definition SHApadded_len :: "nat \ nat" where "SHApadded_len l = l + 1 + (SHApadding_k l) + Y" definition SHApadded_numBlocks :: "nat \ nat" where "SHApadded_numBlocks l = (SHApadded_len l) div (X+Y)" text \We can also define "padding the message" as concatenating strings of bits. Below we prove that this corresponds to the above arithmetic definition. We include this definition, and the proof of equality, merely to demonstrate that these definitions exactly match the standard.\ definition SHApadded_asBits :: "nat \ nat \ bits" where "SHApadded_asBits M l = ( let k = (SHApadding_k l) in (nat_to_bits_len M l) @ [1] @ (replicate k 0) @ (nat_to_bits_len l Y) )" definition SHApadded_numWords :: "nat \ nat" where "SHApadded_numWords l = (SHApadded_numBlocks l) * ((X+Y) div W)" definition SHA_PaddedParsed :: "nat \ nat \ words" where "SHA_PaddedParsed M l = nat_to_words_len W (SHApadded M l) (SHApadded_numWords l)" text \The input message M and bit length l only make sense if M < 2^l. Also the padding only works correctly when l has at most Y bits.\ definition SHA_inputValid :: "nat \ nat \ bool" where "SHA_inputValid M l \ (M < 2^l) \ (l < 2^Y)" lemma SHApadding_len_mod1: "(l + 1 + (SHApadding_k l)) mod (X+Y) = X" proof - let ?x = "(l + 1) mod (X+Y)" have 0: "?x < (X+Y)" by (meson add_pos_pos XYWpos(1,2) mod_less_divisor) show ?thesis proof (cases "?x \ X") case T0: True have T1: "SHApadding_k l = X - ?x" using SHApadding_k_def T0 by presburger have T2: "(l + 1 + (SHApadding_k l)) mod (X+Y) = (?x + (X - ?x)) mod (X+Y)" using T1 by presburger have T3: "(?x + (X - ?x)) mod (X+Y) = X mod (X+Y)" by (metis T0 le_add_diff_inverse) show ?thesis by (metis T2 T3 mod_less XYWpos(2) less_add_same_cancel1) next case F0: False have F1: "SHApadding_k l = X + (X+Y) - ?x" using F0 SHApadding_k_def by presburger have F2: "(l + 1 + (SHApadding_k l)) mod (X+Y) = (?x + (X + (X+Y) - ?x)) mod (X+Y)" using F1 by presburger have F3: "(?x + (X + (X+Y) - ?x)) = X + (X+Y)" using 0 by simp have F4: "(?x + (X + (X+Y) - ?x)) mod (X+Y) = X mod (X+Y)" using F3 by force show ?thesis by (metis F2 F4 mod_less XYWpos(2) less_add_same_cancel1) qed qed lemma SHApadding_len_mod2: "(l + 1 + (SHApadding_k l) + Y) mod (X+Y) = 0" by (metis SHApadding_len_mod1 mod_add_left_eq mod_self) lemma SHApadded_len_mod: "SHApadded_len l mod (X+Y) = 0" using SHApadded_len_def SHApadding_len_mod2 by presburger lemma SHApadded_OR: assumes "l < 2^Y" "k = SHApadding_k l" shows "SHApadded M l = ((2*M + 1)*2^(k+Y)) OR l" proof - let ?x = "(2*M + 1)*2^k" have "SHApadded M l = (?x*2^Y) + l" by (metis assms(2) SHApadded_def mult.assoc power_add) then show ?thesis by (smt (z3) assms(1) OR_sum_nat_hilo_2 mult.assoc mult.commute power_add) qed lemma SHApadded_bitlen: assumes "bit_length M \ l" "l < 2^Y" shows "bit_length (SHApadded M l) \ SHApadded_len l" proof - have 1: "bit_length (2*M+1) = (bit_length M) + 1" by (metis add.commute add_cancel_left_left add_diff_cancel_left' bit_len_exact3 bit_len_zero_eq le_add2 less_exp mult.commute nat_0_less_mult_iff nat_less_le power_0 power_one_right bit_len_shift_add) let ?k = "SHApadding_k l" have 2: "0 < 2*M + 1" by simp have 3: "bit_length ((2*M + 1)*2^?k) = (bit_length M) + 1 + ?k" using 1 2 bit_len_shift by presburger have 4: "bit_length (((2*M + 1)*2^(?k+Y)) + l) = (bit_length M) + 1 + ?k + Y" by (smt (z3) 2 3 assms(2) bit_len_shift_add mult.assoc nat_0_less_mult_iff power_add zero_less_numeral zero_less_power) have 5: "bit_length (SHApadded M l) = (bit_length M) + 1 + ?k + Y" using 4 SHApadded_def by presburger show ?thesis using assms(1) SHApadded_len_def 5 add_le_mono1 by presburger qed lemma SHApadded_bitlen2: assumes "SHA_inputValid M l" shows "bit_length (SHApadded M l) \ SHApadded_len l" using SHApadded_bitlen SHA_inputValid_def assms less_bit_len2 by presburger lemma SHApadded_l: assumes "P = SHApadded M l" "l < 2^Y" shows "l = P mod 2^Y" proof - let ?k = "SHApadding_k l" let ?x = "(2*M + 1)*2^?k" have "P = ?x*2^Y + l" by (metis SHApadded_def assms(1) mult.assoc power_add) then show ?thesis by (metis assms(2) mod_less mod_mult_self3) qed lemma SHApadded_M: assumes "P = SHApadded M l" "l < 2^Y" "k = SHApadding_k l" shows "M = P div 2^(k+Y+1)" proof - have 1: "P = ((2*M + 1)*2^(k+Y)) + l" using assms(1,3) SHApadded_def by meson have 2: "l < 2^(k+Y)" by (metis assms(2) Suc_eq_plus1 le_add2 lessI less_le_trans one_add_one power_increasing_iff) have 3: "P div 2^(k+Y) = 2*M+1" by (simp add: 1 2 ab_semigroup_add_class.add_ac(1)) show ?thesis by (metis 3 div_exp_eq div_mult_self1_is_m dvd_triv_left even_succ_div_two pos2 power_one_right) qed lemma SHApadded_asBits_valid: "bits_valid (SHApadded_asBits M l)" by (simp add: SHApadded_asBits_def nat_to_words_len_valid words_valid_concat words_valid_cons words_valid_zeros) lemma SHApadded_asBits_len: assumes "l < 2^Y" "M < 2^l" shows "length (SHApadded_asBits M l) = SHApadded_len l" by (simp add: SHApadded_asBits_def SHApadded_len_def assms nat_to_words_len_upbnd) lemma SHApadded_asBits_len2: assumes "SHA_inputValid M l" shows "length (SHApadded_asBits M l) = SHApadded_len l" by (meson SHA_inputValid_def SHApadded_asBits_len assms) lemma SHApadded_asBits_to_nat: assumes "l < 2^Y" shows "bits_to_nat (SHApadded_asBits M l) = SHApadded M l" proof - let ?k = "SHApadding_k l" have 1: "SHApadded_asBits M l = (nat_to_bits_len M l)@[1]@(replicate ?k 0)@(nat_to_bits_len l Y)" using SHApadded_asBits_def by presburger have l1: "bits_to_nat (nat_to_bits_len l Y) = l" by (simp add: nat_to_words_len_to_nat) have l2: "length (nat_to_bits_len l Y) = Y" by (simp add: assms(1) nat_to_words_len_upbnd) have k1: "bits_to_nat (replicate ?k 0) = 0" by (simp add: words_to_zero_intro) have M1: "bits_to_nat (nat_to_bits_len M l) = M" by (simp add: nat_to_words_len_to_nat) have 2: "bits_to_nat ((nat_to_bits_len M l)@[1]) = 2*M+1" using M1 words_to_nat_append by force have 3: "bits_to_nat ((nat_to_bits_len M l)@[1]@(replicate ?k 0)) = (2*M+1)*2^?k" using 2 k1 words_to_nat_concat by force let ?X = "(nat_to_bits_len M l)@[1]@(replicate ?k 0)" have 4: "bits_to_nat (?X@(nat_to_bits_len l Y)) = ((2*M+1)*2^?k)*2^Y + l" by (metis 3 l1 l2 power_one_right words_to_nat_concat) show ?thesis by (metis 1 4 SHApadded_def append.assoc mult.assoc power_add) qed lemma SHApadded_toBits: assumes "l < 2^Y" "M < 2^l" shows "nat_to_bits_len (SHApadded M l) (SHApadded_len l) = SHApadded_asBits M l" by (metis SHApadded_asBits_len SHApadded_asBits_to_nat SHApadded_asBits_valid assms words_to_nat_to_words_len2) lemma SHApadded_toBits2: assumes "SHA_inputValid M l" shows "nat_to_bits_len (SHApadded M l) (SHApadded_len l) = SHApadded_asBits M l" using SHA_inputValid_def SHApadded_toBits assms by presburger lemma SHApadded_WdvdLen: "W dvd (SHApadded_len l)" by (meson SHApadded_len_mod Wdvd dvd_eq_mod_eq_0 gcd_nat.trans) lemma SHApadded_wordLen: "SHApadded_numWords l = (SHApadded_len l) div W" by (simp add: SHApadded_numWords_def SHApadded_len_mod SHApadded_numBlocks_def Wdvd div_mult_swap mod_0_imp_dvd) lemma SHA_parsed_valid: "words_valid W (SHA_PaddedParsed M l)" by (simp add: SHA_PaddedParsed_def nat_to_words_len_valid) end (* SHA_PadParse locale *) subsubsection \5.1/2.1 SHA-1, SHA-224 and SHA-256\ text \"Suppose that the length of the message, M, is l bits. Append the bit ``1'' to the end of the message, followed by k zero bits, where k is the smallest, non-negative solution to the equation l + 1 + k = 448 mod 512. Then append the 64-bit block that is equal to the number l expressed using a binary representation." This matches our generic version above with X = 448 and Y = 64. Then the block size is X+Y = 512. Then the word size is W = 32, noting that 32 * 16 = 512, so that each block splits into 16 words.\ lemma SHA1parsing_h: "(32::nat) dvd (448 + 64)" by force global_interpretation SHA1_PadParse: SHA_PadParse 448 64 32 defines SHA1padding_k = "SHA1_PadParse.SHApadding_k" and SHA1padded = "SHA1_PadParse.SHApadded" and SHA1padded_len = "SHA1_PadParse.SHApadded_len" and SHA1padded_numBlocks = "SHA1_PadParse.SHApadded_numBlocks" and SHA1padded_asBits = "SHA1_PadParse.SHApadded_asBits" and SHA1padded_numWords = "SHA1_PadParse.SHApadded_numWords" and SHA1_PaddedParsed = "SHA1_PadParse.SHA_PaddedParsed" and SHA1_inputValid = "SHA1_PadParse.SHA_inputValid" by (simp add: SHA_PadParse_def SHA1parsing_h) subsubsection \5.1/2.2 SHA-384, SHA-512, SHA-512/224 and SHA-512/256\ text \"Suppose the length of the message M, in bits, is l bits. Append the bit ``1'' to the end of the message, followed by k zero bits, where k is the smallest non-negative solution to the equation l + 1 + k = 896 mod 1024. Then append the 128-bit block that is equal to the number expressed using a binary representation." This matches our generic definition with X = 896 and Y = 128. Then the block size is X+Y = 1024. Here the word size is W = 64, so each block is parsed into 16 64-bit numbers.\ lemma SHA512parsing_h: "(64::nat) dvd (896 + 128)" by force global_interpretation SHA512_PadParse: SHA_PadParse 896 128 64 defines SHA512padding_k = "SHA512_PadParse.SHApadding_k" and SHA512padded = "SHA512_PadParse.SHApadded" and SHA512padded_len = "SHA512_PadParse.SHApadded_len" and SHA512padded_numBlocks = "SHA512_PadParse.SHApadded_numBlocks" and SHA512padded_asBits = "SHA512_PadParse.SHApadded_asBits" and SHA512padded_numWords = "SHA512_PadParse.SHApadded_numWords" and SHA512_PaddedParsed = "SHA512_PadParse.SHA_PaddedParsed" and SHA512_inputValid = "SHA512_PadParse.SHA_inputValid" by (simp add: SHA_PadParse_def SHA512parsing_h) subsection \5.3 Setting the Initial Hash Value H^(0)\ text \Before hash computation begins for each of the secure hash algorithms, the initial hash value, H(0), must be set. The size and number of words in H(0) depends on the message digest size.\ subsubsection \5.3.1 SHA-1\ definition SHA1_H0 :: "nat list" where "SHA1_H0 = [0x67452301, 0xefcdab89, 0x98badcfe, 0x10325476, 0xc3d2e1f0]" lemma SHA1_H0_bnd1: assumes "h \ set SHA1_H0" shows "h < 2^32" using assms unfolding SHA1_H0_def apply (cases h) by auto lemma SHA1_H0_bnd2: assumes "i < length SHA1_H0" shows "SHA1_H0 ! i < 2^32" using SHA1_H0_bnd1 assms nth_mem by blast lemma SHA1_H0_valid: "word32s_valid SHA1_H0" using SHA1_H0_bnd1 words_valid_def by blast lemma SHA1_H0_len [simp]: "length SHA1_H0 = 5" unfolding SHA1_H0_def by simp subsubsection \5.3.2 SHA-224\ definition SHA224_H0 :: "nat list" where "SHA224_H0 = [0xc1059ed8, 0x367cd507, 0x3070dd17, 0xf70e5939, 0xffc00b31, 0x68581511, 0x64f98fa7, 0xbefa4fa4]" lemma SHA224_H0_bnd1: assumes "h \ set SHA224_H0" shows "h < 2^32" using assms unfolding SHA224_H0_def apply (cases h) by auto lemma SHA224_H0_bnd2: assumes "i < length SHA224_H0" shows "SHA224_H0 ! i < 2^32" using SHA224_H0_bnd1 assms nth_mem by blast lemma SHA224_H0_valid: "word32s_valid SHA224_H0" using SHA224_H0_bnd1 words_valid_def by blast lemma SHA224_H0_len [simp]: "length SHA224_H0 = 8" unfolding SHA224_H0_def by simp subsubsection \5.3.3 SHA-256\ definition SHA256_H0 :: "nat list" where "SHA256_H0 = [0x6a09e667, 0xbb67ae85, 0x3c6ef372, 0xa54ff53a, 0x510e527f, 0x9b05688c, 0x1f83d9ab, 0x5be0cd19]" lemma SHA256_H0_bnd1: assumes "h \ set SHA256_H0" shows "h < 2^32" using assms unfolding SHA256_H0_def apply (cases h) by auto lemma SHA256_H0_bnd2: assumes "i < length SHA256_H0" shows "SHA256_H0 ! i < 2^32" using SHA256_H0_bnd1 assms nth_mem by blast lemma SHA256_H0_valid: "word32s_valid SHA256_H0" using SHA256_H0_bnd1 words_valid_def by blast lemma SHA256_H0_len [simp]: "length SHA256_H0 = 8" unfolding SHA256_H0_def by simp subsubsection \5.3.4 SHA-384\ definition SHA384_H0 :: "nat list" where "SHA384_H0 = [0xcbbb9d5dc1059ed8, 0x629a292a367cd507, 0x9159015a3070dd17, 0x152fecd8f70e5939, 0x67332667ffc00b31, 0x8eb44a8768581511, 0xdb0c2e0d64f98fa7, 0x47b5481dbefa4fa4]" lemma SHA384_H0_bnd1: assumes "h \ set SHA384_H0" shows "h < 2^64" using assms unfolding SHA384_H0_def apply (cases h) by auto lemma SHA384_H0_bnd2: assumes "i < length SHA384_H0" shows "SHA384_H0 ! i < 2^64" using SHA384_H0_bnd1 assms nth_mem by blast lemma SHA384_H0_valid: "word64s_valid SHA384_H0" using SHA384_H0_bnd1 words_valid_def by blast lemma SHA384_H0_len [simp]: "length SHA384_H0 = 8" unfolding SHA384_H0_def by simp subsubsection \5.3.5 SHA-512\ definition SHA512_H0 :: "nat list" where "SHA512_H0 = [0x6a09e667f3bcc908, 0xbb67ae8584caa73b, 0x3c6ef372fe94f82b, 0xa54ff53a5f1d36f1, 0x510e527fade682d1, 0x9b05688c2b3e6c1f, 0x1f83d9abfb41bd6b, 0x5be0cd19137e2179]" lemma SHA512_H0_bnd1: assumes "h \ set SHA512_H0" shows "h < 2^64" using assms unfolding SHA512_H0_def apply (cases h) by auto lemma SHA512_H0_bnd2: assumes "i < length SHA512_H0" shows "SHA512_H0 ! i < 2^64" using SHA512_H0_bnd1 assms nth_mem by blast lemma SHA512_H0_valid: "word64s_valid SHA512_H0" using SHA512_H0_bnd1 words_valid_def by blast lemma SHA512_H0_len [simp]: "length SHA512_H0 = 8" unfolding SHA512_H0_def by simp subsubsection \5.3.6 SHA-512/t\ text \"``SHA-512/t'' is the general name for a t-bit hash function based on SHA-512 whose output is truncated to t bits. Each hash function requires a distinct initial hash value. This section provides a procedure for determining the initial value for SHA-512/t for a given value of t. ... SHA-512/224 (t = 224) and SHA-512/256 (t = 256) are approved hash algorithms. Other SHA- 512/t hash algorithms with different t values may be specified in [SP 800-107] in the future as the need arises. Below are the IVs for SHA-512/224 and SHA-512/256."\ text \5.3.6.1 SHA-512/224\ definition SHA512_224_H0 :: "nat list" where "SHA512_224_H0 = [0x8C3D37C819544DA2, 0x73E1996689DCD4D6, 0x1DFAB7AE32FF9C82, 0x679DD514582F9FCF, 0x0F6D2B697BD44DA8, 0x77E36F7304C48942, 0x3F9D85A86A1D36C8, 0x1112E6AD91D692A1]" lemma SHA512_224_H0_bnd1: assumes "h \ set SHA512_224_H0" shows "h < 2^64" using assms unfolding SHA512_224_H0_def apply (cases h) by auto lemma SHA512_224_H0_bnd2: assumes "i < length SHA512_224_H0" shows "SHA512_224_H0 ! i < 2^64" using SHA512_224_H0_bnd1 assms nth_mem by blast lemma SHA512_224_H0_valid: "word64s_valid SHA512_224_H0" using SHA512_224_H0_bnd1 words_valid_def by blast lemma SHA512_224_H0_len [simp]: "length SHA512_224_H0 = 8" unfolding SHA512_224_H0_def by simp text \5.3.6.2 SHA-512/256\ definition SHA512_256_H0 :: "nat list" where "SHA512_256_H0 = [0x22312194FC2BF72C, 0x9F555FA3C84C64C2, 0x2393B86B6F53B151, 0x963877195940EABD, 0x96283EE2A88EFFE3, 0xBE5E1E2553863992, 0x2B0199FC2C85B8AA, 0x0EB72DDC81C52CA2]" lemma SHA512_256_H0_bnd1: assumes "h \ set SHA512_256_H0" shows "h < 2^64" using assms unfolding SHA512_256_H0_def apply (cases h) by auto lemma SHA512_256_H0_bnd2: assumes "i < length SHA512_256_H0" shows "SHA512_256_H0 ! i < 2^64" using SHA512_256_H0_bnd1 assms nth_mem by blast lemma SHA512_256_H0_valid: "word64s_valid SHA512_256_H0" using SHA512_256_H0_bnd1 words_valid_def by blast lemma SHA512_256_H0_len [simp]: "length SHA512_256_H0 = 8" unfolding SHA512_256_H0_def by simp section \6. Secure Hash Algorithms\ text \"In the following sections, the hash algorithms are not described in ascending order of size. SHA-256 is described before SHA-224 because the specification for SHA-224 is identical to SHA-256, except that different initial hash values are used, and the final hash value is truncated to 224 bits for SHA-224. The same is true for SHA-512, SHA-384, SHA-512/224 and SHA-512/256, except that the final hash value is truncated to 224 bits for SHA-512/224, 256 bits for SHA-512/256 or 384 bits for SHA-384." It is important to remember that in the standard, addition is always done modulo the pertinent word size. So for SHA-1, they may write a + b and suppress the (mod 2^32).\ subsection \6.1 SHA-1\ subsubsection \Message Schedule\ definition SHA1_MessageSchedule_1 :: "words \ words" where "SHA1_MessageSchedule_1 W = ( let t = length W in W @ [ROTL32 1 ((W ! (t-3)) XOR (W ! (t-8)) XOR (W ! (t-14)) XOR (W ! (t-16)))] )" fun SHA1_MessageSchedule_rec :: "nat \ words \ words" where "SHA1_MessageSchedule_rec n W = ( let t = length W in if t < 16 then W else ( if n = 0 then W else SHA1_MessageSchedule_rec (n-1) (SHA1_MessageSchedule_1 W) ) )" lemma SHA1_MessageSchedule_1_len: "length (SHA1_MessageSchedule_1 W) = (length W) + 1" by (metis SHA1_MessageSchedule_1_def Suc_eq_plus1 length_append_singleton) lemma SHA1_MessageSchedule_1_valid: assumes "word32s_valid W" "t = length W" "16 \ t" shows "word32s_valid (SHA1_MessageSchedule_1 W)" proof - have 10: "W ! (t-3) \ set W" by (metis assms(2,3) diff_less length_greater_0_conv list.size(3) not_le nth_mem zero_less_numeral) have 11: "W ! (t-3) < 2^32" using 10 assms(1) words_valid_def by blast have 20: "W ! (t-8) \ set W" by (metis assms(2,3) diff_less length_greater_0_conv list.size(3) not_le nth_mem zero_less_numeral) have 21: "W ! (t-8) < 2^32" using 20 assms(1) words_valid_def by blast have 30: "W ! (t-14) \ set W" by (metis assms(2,3) diff_less length_greater_0_conv list.size(3) not_le nth_mem zero_less_numeral) have 31: "W ! (t-14) < 2^32" using 30 assms(1) words_valid_def by blast have 40: "W ! (t-16) \ set W" by (metis assms(2,3) diff_less length_greater_0_conv list.size(3) not_le nth_mem zero_less_numeral) have 41: "W ! (t-16) < 2^32" using 40 assms(1) words_valid_def by blast have 50: "(W ! (t-3)) XOR (W ! (t-8)) XOR (W ! (t-14)) XOR (W ! (t-16)) < 2^32" using 11 21 31 41 nat_XOR_upper by presburger have 51: "ROTL32 1 ((W ! (t-3)) XOR (W ! (t-8)) XOR (W ! (t-14)) XOR (W ! (t-16))) < 2^32" by (meson 50 ROTL_bnd one_less_numeral_iff semiring_norm(76)) have 52: "word32s_valid (W @ [ROTL32 1 ((W ! (t-3)) XOR (W ! (t-8)) XOR (W ! (t-14)) XOR (W ! (t-16)))])" using assms(1) 51 words_valid_def words_valid_concat words_valid_cons words_valid_nil by blast show ?thesis by (metis assms(2) 52 SHA1_MessageSchedule_1_def) qed lemma SHA1_MessageSchedule_rec_valid: assumes "word32s_valid W" shows "word32s_valid (SHA1_MessageSchedule_rec n W)" proof (cases "length W < 16") case True then have "SHA1_MessageSchedule_rec n W = W" by simp then show ?thesis using assms by simp next case F: False then show ?thesis using assms proof (induction n arbitrary: W) case 0 then have "SHA1_MessageSchedule_rec 0 W = W" by simp then show ?case using 0(2) by simp next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA1_MessageSchedule_1 W" have 1: "word32s_valid ?W1" by (metis C.prems(1,2) SHA1_MessageSchedule_1_valid leI) have 2: "length ?W1 = ?t + 1" by (simp add: SHA1_MessageSchedule_1_len) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 4: "word32s_valid (SHA1_MessageSchedule_rec n ?W1)" using 1 3 C.IH by blast have 5: "(Suc n) - 1 = n" by simp have 6: "SHA1_MessageSchedule_rec (Suc n) W = (SHA1_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA1_MessageSchedule_rec.simps) show ?case using 4 6 by presburger qed qed lemma SHA1_MessageSchedule_rec_len: assumes "\ length W < 16" shows "length (SHA1_MessageSchedule_rec n W) = (length W) + n" using assms proof (induction n arbitrary: W) case 0 then have "SHA1_MessageSchedule_rec 0 W = W" by simp then show ?case by presburger next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA1_MessageSchedule_1 W" have 2: "length ?W1 = ?t + 1" by (simp add: SHA1_MessageSchedule_1_len) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 5: "(Suc n) - 1 = n" by simp have 6: "SHA1_MessageSchedule_rec (Suc n) W = (SHA1_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA1_MessageSchedule_rec.simps) show ?case using 2 3 6 C.IH by presburger qed definition SHA1_MessageSchedule :: "words \ words" where "SHA1_MessageSchedule MessageBlock = SHA1_MessageSchedule_rec (80-16) MessageBlock" lemma SHA1_MessageSchedule_len: assumes "length MessageBlock = 16" shows "length (SHA1_MessageSchedule MessageBlock) = 80" using SHA1_MessageSchedule_def SHA1_MessageSchedule_rec_len assms less_not_refl by presburger lemma SHA1_MessageSchedule_valid: assumes "word32s_valid MessageBlock" shows "word32s_valid (SHA1_MessageSchedule MessageBlock)" using SHA1_MessageSchedule_rec_valid SHA1_MessageSchedule_def assms by presburger subsubsection \Round Function\ definition SHA1_RoundFunction :: "nat \ words \ nat \ words" where "SHA1_RoundFunction t ABCDE Wt = (let a = ABCDE ! 0; b = ABCDE ! 1; c = ABCDE ! 2; d = ABCDE ! 3; e = ABCDE ! 4; T = ((ROTL32 5 a) + (SHA1_ft t b c d) + e + (SHA1_Kt t) + Wt) mod (2^32) in [ T, a, ROTL32 30 b, c, d ] )" text \We need to iterate over the round function for t=0 to 79. For the definition of this recursive function, we start s at 80 and decrement in each loop, so t = 80 - s. The message schedule W starts at length 80. We use the head of W in each round and drop it before starting the next round. So when s=0, W = []. Could replace s with (length W).\ fun SHA1_RoundFunction_rec :: "nat \ words \ words \ words" where "SHA1_RoundFunction_rec 0 ABCDE W = ABCDE" | "SHA1_RoundFunction_rec s ABCDE W = SHA1_RoundFunction_rec (s-1) (SHA1_RoundFunction (80-s) ABCDE (hd W)) (drop 1 W)" lemma SHA1_RoundFunction_Valid: assumes "5 \ length ABCDE" "word32s_valid ABCDE" shows "word32s_valid (SHA1_RoundFunction t ABCDE Wt)" proof - let ?a = "ABCDE ! 0" let ?b = "ABCDE ! 1" let ?c = "ABCDE ! 2" let ?d = "ABCDE ! 3" let ?e = "ABCDE ! 4" have 0: "0 < length ABCDE \ 1 < length ABCDE \ 2 < length ABCDE \ 3 < length ABCDE \ 4 < length ABCDE" using assms(1) by linarith have 1: "?a < 2^32 \ ?b < 2^32 \ ?c < 2^32 \ ?d < 2^32 \ ?e < 2^32" using 0 assms(2) words_valid_ith by presburger let ?T = "((ROTL32 5 ?a) + (SHA1_ft t ?b ?c ?d) + ?e + (SHA1_Kt t) + Wt) mod (2^32)" have 2: "?T < 2^32" by force have 3: "ROTL32 30 ?b < 2^32" by (meson 1 ROTL_bnd numeral_less_iff semiring_norm(76,78,81)) have 4: "word32s_valid [ ?T, ?a, ROTL32 30 ?b, ?c, ?d ]" using 1 2 3 words_valid_def by auto show ?thesis using 4 SHA1_RoundFunction_def by metis qed lemma SHA1_RoundFunction_len: "length (SHA1_RoundFunction t ABCDE Wt) = 5" proof - let ?a = "ABCDE ! 0" let ?b = "ABCDE ! 1" let ?c = "ABCDE ! 2" let ?d = "ABCDE ! 3" let ?e = "ABCDE ! 4" let ?T = "((ROTL32 5 ?a) + (SHA1_ft t ?b ?c ?d) + ?e + (SHA1_Kt t) + Wt) mod (2^32)" have 4: "length [ ?T, ?a, ROTL32 30 ?b, ?c, ?d ] = 5" by auto show ?thesis using 4 SHA1_RoundFunction_def by metis qed lemma SHA1_RoundFunction_rec_Valid: assumes "5 \ length ABCDE" "word32s_valid ABCDE" shows "word32s_valid (SHA1_RoundFunction_rec s ABCDE W)" using assms proof (induction s arbitrary: ABCDE W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA1_RoundFunction (80-(Suc s)) ABCDE (hd W))" have X1: "length ?X = 5" using SHA1_RoundFunction_len by blast have X2: "word32s_valid ?X" by (simp add: C.prems(1,2) SHA1_RoundFunction_Valid) have X3: "5 \ length ?X" using X1 by simp have "SHA1_RoundFunction_rec (Suc s) ABCDE W = SHA1_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X2 X3 by presburger qed lemma SHA1_RoundFunction_rec_len: assumes "length ABCDE = 5" shows "length (SHA1_RoundFunction_rec s ABCDE W) = 5" using assms proof (induction s arbitrary: ABCDE W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA1_RoundFunction (80-(Suc s)) ABCDE (hd W))" have X1: "length ?X = 5" using SHA1_RoundFunction_len by blast have "SHA1_RoundFunction_rec (Suc s) ABCDE W = SHA1_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X1 by presburger qed subsubsection \SHA-1\ fun SHA1_rec :: "words \ nat \ words \ words" where "SHA1_rec LastHashValue 0 PaddedParsedM = LastHashValue" | "SHA1_rec LastHashValue numBlocks PaddedParsedM = ( let MessageBlock = take 16 PaddedParsedM; MessageSchedule = SHA1_MessageSchedule MessageBlock; ABCDE = SHA1_RoundFunction_rec 80 LastHashValue MessageSchedule; NewHashValue = map2 (\x y. (x+y) mod (2^32)) ABCDE LastHashValue in SHA1_rec NewHashValue (numBlocks-1) (drop 16 PaddedParsedM) )" definition SHA1 :: "nat \ nat \ nat" where "SHA1 M l = ( let PaddedParsedM = SHA1_PaddedParsed M l; numBlocks = SHA1padded_numBlocks l in word32s_to_nat (SHA1_rec SHA1_H0 numBlocks PaddedParsedM) )" lemma SHA1_rec_len: assumes "length H = 5" shows "length (SHA1_rec H n PPM) = 5" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA1_MessageSchedule ?MessageBlock" let ?ABCDE = "SHA1_RoundFunction_rec 80 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^32)) ?ABCDE H" have 1: "length ?ABCDE = 5" by (simp add: C.prems SHA1_RoundFunction_rec_len) have 2: "length ?NewHashValue = 5" using C(2) 1 by simp have 3: "SHA1_rec H (Suc n) PPM = SHA1_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 2 3 by presburger qed lemma SHA1_H0_rec_len: "length (SHA1_rec SHA1_H0 n PPM) = 5" using SHA1_rec_len by simp lemma SHA1_rec_valid: assumes "word32s_valid H" "word32s_valid PPM" "length H = 5" shows "word32s_valid (SHA1_rec H n PPM)" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA1_MessageSchedule ?MessageBlock" let ?ABCDE = "SHA1_RoundFunction_rec 80 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^32)) ?ABCDE H" have 1: "word32s_valid ?MessageBlock" using C(3) words_valid_take by blast have 2: "word32s_valid ?MessageSchedule" using 1 SHA1_MessageSchedule_valid by fast have 3: "length ?ABCDE = 5" using C(4) SHA1_RoundFunction_rec_len by fast have 4: "length ?NewHashValue = 5" using C(4) 3 by simp have 5: "word32s_valid ?NewHashValue" using words_valid_sum_mod by fast have 6: "word32s_valid (drop 16 PPM)" using words_valid_drop C(3) by presburger have 7: "SHA1_rec H (Suc n) PPM = SHA1_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 4 5 6 7 by presburger qed lemma SHA1_H0_rec_valid: assumes "PPM = SHA1_PaddedParsed M l" shows "word32s_valid (SHA1_rec SHA1_H0 n PPM)" by (simp add: SHA1_H0_valid SHA1_rec_valid SHA1_PadParse.SHA_parsed_valid assms) lemma SHA1_bnd: "SHA1 M l < 2^160" proof - let ?PPM = "SHA1_PaddedParsed M l" let ?n = "SHA1padded_numBlocks l" have 1: "word32s_valid (SHA1_rec SHA1_H0 ?n ?PPM)" by (meson SHA1_H0_rec_valid) have 2: "length (SHA1_rec SHA1_H0 ?n ?PPM) = 5" using SHA1_H0_rec_len by presburger have 3: "SHA1 M l = word32s_to_nat (SHA1_rec SHA1_H0 ?n ?PPM)" by (meson SHA1_def) have 4: "SHA1 M l < (2^32)^5" by (metis 1 2 3 words_valid_def zero_less_numeral words_to_nat_len_bnd words_valid_def) show ?thesis using 4 by simp qed subsection \6.2 SHA-256\ subsubsection \Message Schedule\ definition SHA256_MessageSchedule_1 :: "words \ words" where "SHA256_MessageSchedule_1 W = ( let t = length W in W @ [(sigma256_1 (W ! (t-2)) + (W ! (t-7)) + sigma256_0 (W ! (t-15)) + (W ! (t-16))) mod 2^32])" fun SHA256_MessageSchedule_rec :: "nat \ words \ words" where "SHA256_MessageSchedule_rec n W = ( let t = length W in if t < 16 then W else ( if n = 0 then W else SHA256_MessageSchedule_rec (n-1) (SHA256_MessageSchedule_1 W) ) )" lemma SHA256_MessageSchedule_1_valid: assumes "word32s_valid W" shows "word32s_valid (SHA256_MessageSchedule_1 W)" by (metis SHA256_MessageSchedule_1_def assms mod_less_divisor words_valid_concat words_valid_cons words_valid_nil zero_less_numeral zero_less_power) lemma SHA256_MessageSchedule_rec_valid: assumes "word32s_valid W" shows "word32s_valid (SHA256_MessageSchedule_rec n W)" proof (cases "length W < 16") case True then have "SHA256_MessageSchedule_rec n W = W" by simp then show ?thesis using assms by simp next case F: False then show ?thesis using assms proof (induction n arbitrary: W) case 0 then have "SHA256_MessageSchedule_rec 0 W = W" by simp then show ?case using 0(2) by simp next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA256_MessageSchedule_1 W" have 1: "word32s_valid ?W1" by (metis C(3) SHA256_MessageSchedule_1_valid) have 2: "length ?W1 = ?t + 1" by (metis SHA256_MessageSchedule_1_def Suc_eq_plus1 length_append_singleton) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 4: "word32s_valid (SHA256_MessageSchedule_rec n ?W1)" using 1 3 C.IH by blast have 5: "(Suc n) - 1 = n" by simp have 6: "SHA256_MessageSchedule_rec (Suc n) W = (SHA256_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA256_MessageSchedule_rec.simps) show ?case using 4 6 by presburger qed qed lemma SHA256_MessageSchedule_rec_len: assumes "\ length W < 16" shows "length (SHA256_MessageSchedule_rec n W) = (length W) + n" using assms proof (induction n arbitrary: W) case 0 then have "SHA256_MessageSchedule_rec 0 W = W" by simp then show ?case by presburger next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA256_MessageSchedule_1 W" have 2: "length ?W1 = ?t + 1" by (metis SHA256_MessageSchedule_1_def Suc_eq_plus1 length_append_singleton) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 5: "(Suc n) - 1 = n" by simp have 6: "SHA256_MessageSchedule_rec (Suc n) W = (SHA256_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA256_MessageSchedule_rec.simps[of "Suc n" W]) show ?case using 2 3 6 C.IH by presburger qed definition SHA256_MessageSchedule :: "words \ words" where "SHA256_MessageSchedule MessageBlock = SHA256_MessageSchedule_rec (64-16) MessageBlock" lemma SHA256_MessageSchedule_len: assumes "length MessageBlock = 16" shows "length (SHA256_MessageSchedule MessageBlock) = 64" using SHA256_MessageSchedule_def SHA256_MessageSchedule_rec_len assms less_not_refl by presburger lemma SHA256_MessageSchedule_valid: assumes "word32s_valid MessageBlock" shows "word32s_valid (SHA256_MessageSchedule MessageBlock)" using SHA256_MessageSchedule_rec_valid SHA256_MessageSchedule_def assms by presburger subsubsection \Round Function\ definition SHA256_RoundFunction :: "nat \ words \ nat \ words" where "SHA256_RoundFunction t ABCDEFGH Wt = (let a = ABCDEFGH ! 0; b = ABCDEFGH ! 1; c = ABCDEFGH ! 2; d = ABCDEFGH ! 3; e = ABCDEFGH ! 4; f = ABCDEFGH ! 5; g = ABCDEFGH ! 6; h = ABCDEFGH ! 7; T1 = (h + (Sigma256_1 e) + (Ch32 e f g) + (K256list ! t) + Wt) mod (2^32); T2 = ((Sigma256_0 a) + (Maj a b c)) mod (2^32); a' = (T1 + T2) mod (2^32); e' = ( d + T1) mod (2^32) in [a', a, b, c, e', e, f, g] )" fun SHA256_RoundFunction_rec :: "nat \ words \ words \ words" where "SHA256_RoundFunction_rec 0 ABCDEFGH W = ABCDEFGH" | "SHA256_RoundFunction_rec s ABCDEFGH W = SHA256_RoundFunction_rec (s-1) (SHA256_RoundFunction (64-s) ABCDEFGH (hd W)) (drop 1 W)" lemma SHA256_RoundFunction_Valid: assumes "8 \ length ABCDEFGH" "word32s_valid ABCDEFGH" shows "word32s_valid (SHA256_RoundFunction t ABCDEFGH Wt)" proof - let ?a = "ABCDEFGH ! 0" let ?b = "ABCDEFGH ! 1" let ?c = "ABCDEFGH ! 2" let ?d = "ABCDEFGH ! 3" let ?e = "ABCDEFGH ! 4" let ?f = "ABCDEFGH ! 5" let ?g = "ABCDEFGH ! 6" let ?h = "ABCDEFGH ! 7" have 0: "0 < length ABCDEFGH \ 1 < length ABCDEFGH \ 2 < length ABCDEFGH \ 3 < length ABCDEFGH \ 4 < length ABCDEFGH \ 5 < length ABCDEFGH \ 6 < length ABCDEFGH \ 7 < length ABCDEFGH" using assms(1) by linarith have 1: "?a < 2^32 \ ?b < 2^32 \ ?c < 2^32 \ ?d < 2^32 \ ?e < 2^32 \ ?f < 2^32 \ ?g < 2^32 \ ?h < 2^32" using 0 assms(2) words_valid_ith by presburger let ?T1 = "(?h + (Sigma256_1 ?e) + (Ch32 ?e ?f ?g) + (K256list ! t) + Wt) mod (2^32)" have 2: "?T1 < 2^32" by force let ?T2 = "((Sigma256_0 ?a) + (Maj ?a ?b ?c)) mod (2^32)" have 3: "?T2 < 2^32" by force let ?a' = "(?T1 + ?T2) mod (2^32)" have 4: "?a' < 2^32" by force let ?e' = "(?d + ?T1) mod (2^32)" have 5: "?e' < 2^32" by force have 6: "word32s_valid [?a', ?a, ?b, ?c, ?e', ?e, ?f, ?g]" using 1 2 3 4 5 words_valid_cons words_valid_nil by presburger show ?thesis using 6 SHA256_RoundFunction_def by metis qed lemma SHA256_RoundFunction_len: "length (SHA256_RoundFunction t ABCDEFGH Wt) = 8" proof - let ?a = "ABCDEFGH ! 0" let ?b = "ABCDEFGH ! 1" let ?c = "ABCDEFGH ! 2" let ?d = "ABCDEFGH ! 3" let ?e = "ABCDEFGH ! 4" let ?f = "ABCDEFGH ! 5" let ?g = "ABCDEFGH ! 6" let ?h = "ABCDEFGH ! 7" let ?T1 = "(?h + (Sigma256_1 ?e) + (Ch32 ?e ?f ?g) + (K256list ! t) + Wt) mod (2^32)" let ?T2 = "((Sigma256_0 ?a) + (Maj ?a ?b ?c)) mod (2^32)" let ?a' = "(?T1 + ?T2) mod (2^32)" let ?e' = " (?d + ?T1) mod (2^32)" have "length [?a', ?a, ?b, ?c, ?e', ?e, ?f, ?g] = 8" by force then show ?thesis using SHA256_RoundFunction_def by metis qed lemma SHA256_RoundFunction_rec_Valid: assumes "8 \ length ABCDEFGH" "word32s_valid ABCDEFGH" shows "word32s_valid (SHA256_RoundFunction_rec s ABCDEFGH W)" using assms proof (induction s arbitrary: ABCDEFGH W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA256_RoundFunction (64-(Suc s)) ABCDEFGH (hd W))" have X1: "length ?X = 8" using SHA256_RoundFunction_len by blast have X2: "word32s_valid ?X" by (simp add: C.prems(1,2) SHA256_RoundFunction_Valid) have X3: "8 \ length ?X" using X1 by simp have "SHA256_RoundFunction_rec (Suc s) ABCDEFGH W = SHA256_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X2 X3 by presburger qed lemma SHA256_RoundFunction_rec_len: assumes "length ABCDEFGH = 8" shows "length (SHA256_RoundFunction_rec s ABCDEFGH W) = 8" using assms proof (induction s arbitrary: ABCDEFGH W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA256_RoundFunction (64-(Suc s)) ABCDEFGH (hd W))" have X1: "length ?X = 8" using SHA256_RoundFunction_len by blast have "SHA256_RoundFunction_rec (Suc s) ABCDEFGH W = SHA256_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X1 by presburger qed subsubsection \SHA-256\ fun SHA256_rec :: "words \ nat \ words \ words" where "SHA256_rec LastHashValue 0 PaddedParsedM = LastHashValue" | "SHA256_rec LastHashValue numBlocks PaddedParsedM = ( let MessageBlock = take 16 PaddedParsedM; MessageSchedule = SHA256_MessageSchedule MessageBlock; ABCDEFGH = SHA256_RoundFunction_rec 64 LastHashValue MessageSchedule; NewHashValue = map2 (\x y. (x+y) mod (2^32)) ABCDEFGH LastHashValue in SHA256_rec NewHashValue (numBlocks-1) (drop 16 PaddedParsedM) )" definition SHA256 :: "nat \ nat \ nat" where "SHA256 M l = ( let PaddedParsedM = SHA1_PaddedParsed M l; numBlocks = SHA1padded_numBlocks l in word32s_to_nat (SHA256_rec SHA256_H0 numBlocks PaddedParsedM) )" text \SHA-1, -256, and -224 all allow messages up to length 2^64.\ abbreviation SHA256_inputValid :: "nat \ nat \ bool" where "SHA256_inputValid M l \ SHA1_inputValid M l" lemma SHA256_rec_len: assumes "length H = 8" shows "length (SHA256_rec H n PPM) = 8" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA256_MessageSchedule ?MessageBlock" let ?ABCDEFGH = "SHA256_RoundFunction_rec 64 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^32)) ?ABCDEFGH H" have 1: "length ?ABCDEFGH = 8" by (simp add: C.prems SHA256_RoundFunction_rec_len) have 2: "length ?NewHashValue = 8" using C(2) 1 by simp have 3: "SHA256_rec H (Suc n) PPM = SHA256_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 2 3 by presburger qed lemma SHA256_H0_rec_len: "length (SHA256_rec SHA256_H0 n PPM) = 8" using SHA256_rec_len by simp lemma SHA256_rec_valid: assumes "word32s_valid H" "word32s_valid PPM" "length H = 8" shows "word32s_valid (SHA256_rec H n PPM)" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA256_MessageSchedule ?MessageBlock" let ?ABCDEFGH = "SHA256_RoundFunction_rec 64 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^32)) ?ABCDEFGH H" have 1: "word32s_valid ?MessageBlock" using C(3) words_valid_take by blast have 2: "word32s_valid ?MessageSchedule" using 1 SHA256_MessageSchedule_valid by fast have 3: "length ?ABCDEFGH = 8" using C(4) SHA256_RoundFunction_rec_len by fast have 4: "length ?NewHashValue = 8" using C(4) 3 by simp have 5: "word32s_valid ?NewHashValue" using words_valid_sum_mod by fast have 6: "word32s_valid (drop 16 PPM)" using words_valid_drop C(3) by presburger have 7: "SHA256_rec H (Suc n) PPM = SHA256_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 4 5 6 7 by presburger qed lemma SHA256_H0_rec_valid: assumes "PPM = SHA1_PaddedParsed M l" shows "word32s_valid (SHA256_rec SHA256_H0 n PPM)" by (simp add: SHA1_PadParse.SHA_parsed_valid SHA256_H0_valid SHA256_rec_valid assms) lemma SHA256_bnd: "SHA256 M l < 2^256" proof - let ?PPM = "SHA1_PaddedParsed M l" let ?n = "SHA1padded_numBlocks l" have 1: "word32s_valid (SHA256_rec SHA256_H0 ?n ?PPM)" by (meson SHA256_H0_rec_valid) have 2: "length (SHA256_rec SHA256_H0 ?n ?PPM) = 8" using SHA256_H0_rec_len by presburger have 3: "SHA256 M l = word32s_to_nat (SHA256_rec SHA256_H0 ?n ?PPM)" by (meson SHA256_def) have 4: "SHA256 M l < (2^32)^8" by (metis 1 2 3 words_valid_def zero_less_numeral words_to_nat_len_bnd words_valid_def) show ?thesis using 4 by simp qed subsection \6.3 SHA-224\ definition SHA224 :: "nat \ nat \ nat" where "SHA224 M l = ( let PaddedParsedM = SHA1_PaddedParsed M l; numBlocks = SHA1padded_numBlocks l; H256 = SHA256_rec SHA224_H0 numBlocks PaddedParsedM in word32s_to_nat (butlast H256) )" text \SHA-1, -256, and -224 all allow messages up to length 2^64.\ abbreviation SHA224_inputValid :: "nat \ nat \ bool" where "SHA224_inputValid M l \ SHA1_inputValid M l" lemma SHA224_H0_rec_len: "length (SHA256_rec SHA224_H0 n PPM) = 8" using SHA256_rec_len by simp lemma SHA224_H0_rec_valid: assumes "PPM = SHA1_PaddedParsed M l" shows "word32s_valid (SHA256_rec SHA224_H0 n PPM)" by (simp add: SHA1_PadParse.SHA_parsed_valid SHA224_H0_valid SHA256_rec_valid assms) lemma SHA224_bnd: "SHA224 M l < 2^224" proof - let ?PPM = "SHA1_PaddedParsed M l" let ?n = "SHA1padded_numBlocks l" let ?H256 = "SHA256_rec SHA224_H0 ?n ?PPM" have 1: "word32s_valid ?H256" by (meson SHA224_H0_rec_valid) have 2: "length ?H256 = 8" using SHA224_H0_rec_len by presburger have 3: "SHA224 M l = word32s_to_nat (butlast ?H256)" by (meson SHA224_def) have 4: "length (butlast ?H256) = 7" using 2 by simp have 5: "word32s_valid (butlast ?H256)" using 1 words_valid_butlast by fast have 6: "SHA224 M l < (2^32)^7" by (metis 3 4 5 words_to_nat_len_bnd words_valid_def zero_less_numeral) show ?thesis using 6 by simp qed subsection \6.4 SHA-512\ subsubsection \Message Schedule\ definition SHA512_MessageSchedule_1 :: "words \ words" where "SHA512_MessageSchedule_1 W = ( let t = length W in W @ [(sigma512_1 (W ! (t-2)) + (W ! (t-7)) + sigma512_0 (W ! (t-15)) + (W ! (t-16))) mod 2^64] )" fun SHA512_MessageSchedule_rec :: "nat \ words \ words" where "SHA512_MessageSchedule_rec n W = ( let t = length W in if t < 16 then W else ( if n = 0 then W else SHA512_MessageSchedule_rec (n-1) (SHA512_MessageSchedule_1 W) ) )" lemma SHA512_MessageSchedule_1_valid: assumes "word64s_valid W" shows "word64s_valid (SHA512_MessageSchedule_1 W)" by (metis SHA512_MessageSchedule_1_def assms mod_less_divisor words_valid_concat words_valid_cons words_valid_nil zero_less_numeral zero_less_power) lemma SHA512_MessageSchedule_rec_valid: assumes "word64s_valid W" shows "word64s_valid (SHA512_MessageSchedule_rec n W)" proof (cases "length W < 16") case True then have "SHA512_MessageSchedule_rec n W = W" by simp then show ?thesis using assms by simp next case F: False then show ?thesis using assms proof (induction n arbitrary: W) case 0 then have "SHA512_MessageSchedule_rec 0 W = W" by simp then show ?case using 0(2) by simp next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA512_MessageSchedule_1 W" have 1: "word64s_valid ?W1" by (metis C(3) SHA512_MessageSchedule_1_valid) have 2: "length ?W1 = ?t + 1" by (metis SHA512_MessageSchedule_1_def Suc_eq_plus1 length_append_singleton) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 4: "word64s_valid (SHA512_MessageSchedule_rec n ?W1)" using 1 3 C.IH by blast have 5: "(Suc n) - 1 = n" by simp have 6: "SHA512_MessageSchedule_rec (Suc n) W = (SHA512_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA512_MessageSchedule_rec.simps) show ?case using 4 6 by presburger qed qed lemma SHA512_MessageSchedule_rec_len: assumes "\ length W < 16" shows "length (SHA512_MessageSchedule_rec n W) = (length W) + n" using assms proof (induction n arbitrary: W) case 0 then have "SHA512_MessageSchedule_rec 0 W = W" by simp then show ?case by presburger next case C: (Suc n) have 0: "\ 0 = Suc n" by simp let ?t = "length W" let ?W1 = "SHA512_MessageSchedule_1 W" have 2: "length ?W1 = ?t + 1" by (metis SHA512_MessageSchedule_1_def Suc_eq_plus1 length_append_singleton) have 3: "\ length ?W1 < 16" by (metis 2 C.prems(1) add_lessD1) have 5: "(Suc n) - 1 = n" by simp have 6: "SHA512_MessageSchedule_rec (Suc n) W = (SHA512_MessageSchedule_rec n ?W1)" by (metis C(2) 0 5 SHA512_MessageSchedule_rec.simps[of "Suc n" W]) show ?case using 2 3 6 C.IH by presburger qed definition SHA512_MessageSchedule :: "words \ words" where "SHA512_MessageSchedule MessageBlock = SHA512_MessageSchedule_rec (80-16) MessageBlock" lemma SHA512_MessageSchedule_len: assumes "length MessageBlock = 16" shows "length (SHA512_MessageSchedule MessageBlock) = 80" using SHA512_MessageSchedule_def SHA512_MessageSchedule_rec_len assms less_not_refl by presburger lemma SHA512_MessageSchedule_valid: assumes "word64s_valid MessageBlock" shows "word64s_valid (SHA512_MessageSchedule MessageBlock)" using SHA512_MessageSchedule_rec_valid SHA512_MessageSchedule_def assms by presburger subsubsection \Round Function\ definition SHA512_RoundFunction :: "nat \ words \ nat \ words" where "SHA512_RoundFunction t ABCDEFGH Wt = (let a = ABCDEFGH ! 0; b = ABCDEFGH ! 1; c = ABCDEFGH ! 2; d = ABCDEFGH ! 3; e = ABCDEFGH ! 4; f = ABCDEFGH ! 5; g = ABCDEFGH ! 6; h = ABCDEFGH ! 7; T1 = (h + (Sigma512_1 e) + (Ch64 e f g) + (K512list ! t) + Wt) mod (2^64); T2 = ((Sigma512_0 a) + (Maj a b c)) mod (2^64); a' = (T1 + T2) mod (2^64); e' = ( d + T1) mod (2^64) in [a', a, b, c, e', e, f, g] )" fun SHA512_RoundFunction_rec :: "nat \ words \ words \ words" where "SHA512_RoundFunction_rec 0 ABCDEFGH W = ABCDEFGH" | "SHA512_RoundFunction_rec s ABCDEFGH W = SHA512_RoundFunction_rec (s-1) (SHA512_RoundFunction (80-s) ABCDEFGH (hd W)) (drop 1 W)" lemma SHA512_RoundFunction_Valid: assumes "8 \ length ABCDEFGH" "word64s_valid ABCDEFGH" shows "word64s_valid (SHA512_RoundFunction t ABCDEFGH Wt)" proof - let ?a = "ABCDEFGH ! 0" let ?b = "ABCDEFGH ! 1" let ?c = "ABCDEFGH ! 2" let ?d = "ABCDEFGH ! 3" let ?e = "ABCDEFGH ! 4" let ?f = "ABCDEFGH ! 5" let ?g = "ABCDEFGH ! 6" let ?h = "ABCDEFGH ! 7" have 0: "0 < length ABCDEFGH \ 1 < length ABCDEFGH \ 2 < length ABCDEFGH \ 3 < length ABCDEFGH \ 4 < length ABCDEFGH \ 5 < length ABCDEFGH \ 6 < length ABCDEFGH \ 7 < length ABCDEFGH" using assms(1) by linarith have 1: "?a < 2^64 \ ?b < 2^64 \ ?c < 2^64 \ ?d < 2^64 \ ?e < 2^64 \ ?f < 2^64 \ ?g < 2^64 \ ?h < 2^64" using 0 assms(2) words_valid_ith by presburger let ?T1 = "(?h + (Sigma512_1 ?e) + (Ch64 ?e ?f ?g) + (K512list ! t) + Wt) mod (2^64)" have 2: "?T1 < 2^64" by force let ?T2 = "((Sigma512_0 ?a) + (Maj ?a ?b ?c)) mod (2^64)" have 3: "?T2 < 2^64" by force let ?a' = "(?T1 + ?T2) mod (2^64)" have 4: "?a' < 2^64" by force let ?e' = "(?d + ?T1) mod (2^64)" have 5: "?e' < 2^64" by force have 6: "word64s_valid [?a', ?a, ?b, ?c, ?e', ?e, ?f, ?g]" using 1 2 3 4 5 words_valid_cons words_valid_nil by presburger show ?thesis using 6 SHA512_RoundFunction_def by metis qed lemma SHA512_RoundFunction_len: "length (SHA512_RoundFunction t ABCDEFGH Wt) = 8" proof - let ?a = "ABCDEFGH ! 0" let ?b = "ABCDEFGH ! 1" let ?c = "ABCDEFGH ! 2" let ?d = "ABCDEFGH ! 3" let ?e = "ABCDEFGH ! 4" let ?f = "ABCDEFGH ! 5" let ?g = "ABCDEFGH ! 6" let ?h = "ABCDEFGH ! 7" let ?T1 = "(?h + (Sigma512_1 ?e) + (Ch64 ?e ?f ?g) + (K512list ! t) + Wt) mod (2^64)" let ?T2 = "((Sigma512_0 ?a) + (Maj ?a ?b ?c)) mod (2^64)" let ?a' = "(?T1 + ?T2) mod (2^64)" let ?e' = "(?d + ?T1) mod (2^64)" have "length [?a', ?a, ?b, ?c, ?e', ?e, ?f, ?g] = 8" by force then show ?thesis using SHA512_RoundFunction_def by metis qed lemma SHA512_RoundFunction_rec_Valid: assumes "8 \ length ABCDEFGH" "word64s_valid ABCDEFGH" shows "word64s_valid (SHA512_RoundFunction_rec s ABCDEFGH W)" using assms proof (induction s arbitrary: ABCDEFGH W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA512_RoundFunction (80-(Suc s)) ABCDEFGH (hd W))" have X1: "length ?X = 8" using SHA512_RoundFunction_len by blast have X2: "word64s_valid ?X" by (simp add: C.prems(1,2) SHA512_RoundFunction_Valid) have X3: "8 \ length ?X" using X1 by simp have "SHA512_RoundFunction_rec (Suc s) ABCDEFGH W = SHA512_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X2 X3 by presburger qed lemma SHA512_RoundFunction_rec_len: assumes "length ABCDEFGH = 8" shows "length (SHA512_RoundFunction_rec s ABCDEFGH W) = 8" using assms proof (induction s arbitrary: ABCDEFGH W) case 0 then show ?case by simp next case C: (Suc s) let ?d1W = "drop 1 W" let ?X = "(SHA512_RoundFunction (80-(Suc s)) ABCDEFGH (hd W))" have X1: "length ?X = 8" using SHA512_RoundFunction_len by blast have "SHA512_RoundFunction_rec (Suc s) ABCDEFGH W = SHA512_RoundFunction_rec s ?X ?d1W" by simp then show ?case using C.IH X1 by presburger qed subsubsection \SHA-512\ fun SHA512_rec :: "words \ nat \ words \ words" where "SHA512_rec LastHashValue 0 PaddedParsedM = LastHashValue" | "SHA512_rec LastHashValue numBlocks PaddedParsedM = ( let MessageBlock = take 16 PaddedParsedM; MessageSchedule = SHA512_MessageSchedule MessageBlock; ABCDEFGH = SHA512_RoundFunction_rec 80 LastHashValue MessageSchedule; NewHashValue = map2 (\x y. (x+y) mod (2^64)) ABCDEFGH LastHashValue in SHA512_rec NewHashValue (numBlocks-1) (drop 16 PaddedParsedM) )" definition SHA512 :: "nat \ nat \ nat" where "SHA512 M l = ( let PaddedParsedM = SHA512_PaddedParsed M l; numBlocks = SHA512padded_numBlocks l in word64s_to_nat (SHA512_rec SHA512_H0 numBlocks PaddedParsedM) )" lemma SHA512_rec_len: assumes "length H = 8" shows "length (SHA512_rec H n PPM) = 8" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA512_MessageSchedule ?MessageBlock" let ?ABCDEFGH = "SHA512_RoundFunction_rec 80 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^64)) ?ABCDEFGH H" have 1: "length ?ABCDEFGH = 8" by (simp add: C.prems SHA512_RoundFunction_rec_len) have 2: "length ?NewHashValue = 8" using C(2) 1 by simp have 3: "SHA512_rec H (Suc n) PPM = SHA512_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 2 3 by presburger qed lemma SHA512_H0_rec_len: "length (SHA512_rec SHA512_H0 n PPM) = 8" using SHA512_rec_len by simp lemma SHA512_rec_valid: assumes "word64s_valid H" "word64s_valid PPM" "length H = 8" shows "word64s_valid (SHA512_rec H n PPM)" using assms proof (induction n arbitrary: H PPM) case 0 then show ?case by simp next case C: (Suc n) let ?MessageBlock = "take 16 PPM" let ?MessageSchedule = "SHA512_MessageSchedule ?MessageBlock" let ?ABCDEFGH = "SHA512_RoundFunction_rec 80 H ?MessageSchedule" let ?NewHashValue = "map2 (\x y. (x+y) mod (2^64)) ?ABCDEFGH H" have 1: "word64s_valid ?MessageBlock" using C(3) words_valid_take by blast have 2: "word64s_valid ?MessageSchedule" using 1 SHA512_MessageSchedule_valid by fast have 3: "length ?ABCDEFGH = 8" using C(4) SHA512_RoundFunction_rec_len by fast have 4: "length ?NewHashValue = 8" using C(4) 3 by simp have 5: "word64s_valid ?NewHashValue" using words_valid_sum_mod by fast have 6: "word64s_valid (drop 16 PPM)" using words_valid_drop C(3) by presburger have 7: "SHA512_rec H (Suc n) PPM = SHA512_rec ?NewHashValue n (drop 16 PPM)" by force show ?case using C(1) 4 5 6 7 by presburger qed lemma SHA512_H0_rec_valid: assumes "PPM = SHA512_PaddedParsed M l" shows "word64s_valid (SHA512_rec SHA512_H0 n PPM)" by (simp add: SHA512_PadParse.SHA_parsed_valid SHA512_H0_valid SHA512_rec_valid assms) lemma SHA512_bnd: "SHA512 M l < 2^512" proof - let ?PPM = "SHA512_PaddedParsed M l" let ?n = "SHA512padded_numBlocks l" have 1: "word64s_valid (SHA512_rec SHA512_H0 ?n ?PPM)" by (meson SHA512_H0_rec_valid) have 2: "length (SHA512_rec SHA512_H0 ?n ?PPM) = 8" using SHA512_H0_rec_len by presburger have 3: "SHA512 M l = word64s_to_nat (SHA512_rec SHA512_H0 ?n ?PPM)" by (meson SHA512_def) have 4: "SHA512 M l < (2^64)^8" by (metis 1 2 3 words_valid_def words_to_nat_len_bnd zero_less_numeral) show ?thesis using 4 by simp qed subsection \6.5 SHA-384\ definition SHA384 :: "nat \ nat \ nat" where "SHA384 M l = ( let PaddedParsedM = SHA512_PaddedParsed M l; numBlocks = SHA512padded_numBlocks l; H512 = SHA512_rec SHA384_H0 numBlocks PaddedParsedM in word64s_to_nat (take 6 H512) )" text \SHA-512, -384, and -512/t all allow messages up to length 2^128.\ abbreviation SHA384_inputValid :: "nat \ nat \ bool" where "SHA384_inputValid M l \ SHA512_inputValid M l" lemma SHA384_H0_rec_len: "length (SHA512_rec SHA384_H0 n PPM) = 8" using SHA512_rec_len by simp lemma SHA384_H0_rec_valid: assumes "PPM = SHA512_PaddedParsed M l" shows "word64s_valid (SHA512_rec SHA384_H0 n PPM)" by (simp add: SHA512_PadParse.SHA_parsed_valid SHA384_H0_valid SHA512_rec_valid assms) lemma SHA384_bnd: "SHA384 M l < 2^384" proof - let ?PPM = "SHA512_PaddedParsed M l" let ?n = "SHA512padded_numBlocks l" let ?H512 = "SHA512_rec SHA384_H0 ?n ?PPM" have 1: "word64s_valid ?H512" by (meson SHA384_H0_rec_valid) have 10: "word64s_valid (take 6 ?H512)" using 1 words_valid_take by blast have 2: "length (?H512) = 8" using SHA384_H0_rec_len by presburger have 20: "length (take 6 ?H512) = 6" using 2 by simp have 3: "SHA384 M l = word64s_to_nat (take 6 ?H512)" by (meson SHA384_def) have 4: "SHA384 M l < (2^64)^6" by (metis 10 20 3 SHA512_PadParse.XYWpos(3) words_to_nat_len_bnd words_valid_def) show ?thesis using 4 by simp qed subsection \6.6 SHA-512/224\ definition SHA512_224 :: "nat \ nat \ nat" where "SHA512_224 M l = ( let PaddedParsedM = SHA512_PaddedParsed M l; numBlocks = SHA512padded_numBlocks l; H512 = word64s_to_nat (SHA512_rec SHA512_224_H0 numBlocks PaddedParsedM) in H512 div 2^288 )" text \SHA-512, -384, and -512/t all allow messages up to length 2^128.\ abbreviation SHA512_224_inputValid :: "nat \ nat \ bool" where "SHA512_224_inputValid M l \ SHA512_inputValid M l" lemma SHA512_224_H0_rec_len: "length (SHA512_rec SHA512_224_H0 n PPM) = 8" using SHA512_rec_len by simp lemma SHA512_224_H0_rec_valid: assumes "PPM = SHA512_PaddedParsed M l" shows "word64s_valid (SHA512_rec SHA512_224_H0 n PPM)" by (simp add: SHA512_PadParse.SHA_parsed_valid SHA512_224_H0_valid SHA512_rec_valid assms) lemma SHA512_224_bnd: "SHA512_224 M l < 2^224" proof - let ?PPM = "SHA512_PaddedParsed M l" let ?n = "SHA512padded_numBlocks l" let ?H512 = "SHA512_rec SHA512_224_H0 ?n ?PPM" let ?H512_nat = "word64s_to_nat ?H512" have 1: "word64s_valid ?H512" by (meson SHA512_224_H0_rec_valid) have 2: "length (?H512) = 8" using SHA512_224_H0_rec_len by presburger have 3: "?H512_nat < (2^64)^8" by (metis 1 2 SHA512_PadParse.XYWpos(3) words_to_nat_len_bnd words_valid_def) have 4: "?H512_nat < 2^512" using 3 by simp show ?thesis using 4 SHA512_224_def by simp qed subsection \6.6 SHA-512/256\ definition SHA512_256 :: "nat \ nat \ nat" where "SHA512_256 M l = ( let PaddedParsedM = SHA512_PaddedParsed M l; numBlocks = SHA512padded_numBlocks l; H512 = word64s_to_nat (SHA512_rec SHA512_256_H0 numBlocks PaddedParsedM) in H512 div 2^256 )" text \SHA-512, -384, and -512/t all allow messages up to length 2^128.\ abbreviation SHA512_256_inputValid :: "nat \ nat \ bool" where "SHA512_256_inputValid M l \ SHA512_inputValid M l" lemma SHA512_256_H0_rec_len: "length (SHA512_rec SHA512_256_H0 n PPM) = 8" using SHA512_rec_len by simp lemma SHA512_256_H0_rec_valid: assumes "PPM = SHA512_PaddedParsed M l" shows "word64s_valid (SHA512_rec SHA512_256_H0 n PPM)" by (simp add: SHA512_PadParse.SHA_parsed_valid SHA512_256_H0_valid SHA512_rec_valid assms) lemma SHA512_256_bnd: "SHA512_256 M l < 2^256" proof - let ?PPM = "SHA512_PaddedParsed M l" let ?n = "SHA512padded_numBlocks l" let ?H512 = "SHA512_rec SHA512_256_H0 ?n ?PPM" let ?H512_nat = "word64s_to_nat ?H512" have 1: "word64s_valid ?H512" by (meson SHA512_256_H0_rec_valid) have 2: "length (?H512) = 8" using SHA512_256_H0_rec_len by presburger have 3: "?H512_nat < (2^64)^8" by (metis 1 2 SHA512_PadParse.XYWpos(3) words_to_nat_len_bnd words_valid_def) have 4: "?H512_nat < 2^512" using 3 by simp show ?thesis using 4 SHA512_256_def by simp qed section \Octets\ text \We have translated the Secure Hash Standard so that each of the Secure Hash Algorithms is a function from natural numbers to natural numbers. This allows us the flexibility to apply the SHA definition to any implementation, whether that implementation acts on natural numbers, bit strings, octet (i.e., byte) strings, or (generically) n-bit word strings. (Words.thy contains all the conversion functions needed between natural numbers and n-bit words with abbreviations for conversions to and from bit strings, octet strings, 32-bit word strings, and 64-bit word strings.) Some NIST standards rely on an underlying hash function and they assume that that hash function takes as input a string of 8-bit values, that is octets, and produces a string of octets. So here we provide the "octet version" of each of the secure hash algorithms above. For each of these, we prove the basic things that one would like to know: that the output octet string is valid (meaning each octet is < 256) and that we know how many octets is output by each of the hash algorithms. Note that these functions assume that the message length (in bits) is exactly 8 times the number of input octets. For a particular implementation, you will need to decide if a message bit-length must be 0 mod 8. If not, you may need to provide a message length as an additional input and deal with right- or left-alignment of the message within the words. The following as a guide, along with the conversions available in Words.thy, you can easily form any wrapper you may need to apply to the above nat-to-nat SHA functions defined above.\ subsection \SHA-1\ definition SHA1octets :: "octets \ octets" where "SHA1octets os = nat_to_octets_len (SHA1 (octets_to_nat os) (8*(length os))) 20" definition SHA1_hLen :: nat where "SHA1_hLen = 20" lemma SHA1octets_len: "length (SHA1octets os) = 20" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA1 ?M ?l < 2^160" using SHA1_bnd by blast have 2: "(160::nat) = 8*20" by force have 3: "SHA1 ?M ?l < (2^8)^20" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA1 ?M ?l) 20 ) = 20" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA1octets_def by presburger qed lemma SHA1octets_len2: "\x. length (SHA1octets x) = 20" using SHA1octets_len by blast lemma SHA1octets_valid: "octets_valid (SHA1octets os)" using nat_to_words_len_valid SHA1octets_def by presburger lemma SHA1octets_valid2: "\x. octets_valid (SHA1octets x)" using SHA1octets_valid by satx subsection \SHA-224\ definition SHA224octets :: "octets \ octets" where "SHA224octets os = nat_to_octets_len (SHA224 (octets_to_nat os) (8*(length os))) 28" definition SHA224_hLen :: nat where "SHA224_hLen = 28" lemma SHA224octets_len: "length (SHA224octets os) = 28" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA224 ?M ?l < 2^224" using SHA224_bnd by blast have 2: "(224::nat) = 8*28" by force have 3: "SHA224 ?M ?l < (2^8)^28" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA224 ?M ?l) 28 ) = 28" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA224octets_def by presburger qed lemma SHA224octets_len2: "\x. length (SHA224octets x) = 28" using SHA224octets_len by blast lemma SHA224octets_valid: "octets_valid (SHA224octets os)" using nat_to_words_len_valid SHA224octets_def by presburger lemma SHA224octets_valid2: "\x. octets_valid (SHA224octets x)" using SHA224octets_valid by satx subsection \SHA-256\ definition SHA256octets :: "octets \ octets" where "SHA256octets os = nat_to_octets_len (SHA256 (octets_to_nat os) (8*(length os))) 32" definition SHA256_hLen :: nat where "SHA256_hLen = 32" lemma SHA256octets_len: "length (SHA256octets os) = 32" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA256 ?M ?l < 2^256" using SHA256_bnd by blast have 2: "(256::nat) = 8*32" by force have 3: "SHA256 ?M ?l < (2^8)^32" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA256 ?M ?l) 32 ) = 32" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA256octets_def by presburger qed lemma SHA256octets_len2: "\x. length (SHA256octets x) = 32" using SHA256octets_len by blast lemma SHA256octets_valid: "octets_valid (SHA256octets os)" using nat_to_words_len_valid SHA256octets_def by presburger lemma SHA256octets_valid2: "\x. octets_valid (SHA256octets x)" using SHA256octets_valid by satx subsection \SHA-384\ definition SHA384octets :: "octets \ octets" where "SHA384octets os = nat_to_octets_len (SHA384 (octets_to_nat os) (8*(length os))) 48" definition SHA384_hLen :: nat where "SHA384_hLen = 48" lemma SHA384octets_len: "length (SHA384octets os) = 48" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA384 ?M ?l < 2^384" using SHA384_bnd by blast have 2: "(384::nat) = 8*48" by force have 3: "SHA384 ?M ?l < (2^8)^48" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA384 ?M ?l) 48 ) = 48" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA384octets_def by presburger qed lemma SHA384octets_len2: "\x. length (SHA384octets x) = 48" using SHA384octets_len by blast lemma SHA384octets_valid: "octets_valid (SHA384octets os)" using nat_to_words_len_valid SHA384octets_def by presburger lemma SHA384octets_valid2: "\x. octets_valid (SHA384octets x)" using SHA384octets_valid by satx subsection \SHA-512\ definition SHA512octets :: "octets \ octets" where "SHA512octets os = nat_to_octets_len (SHA512 (octets_to_nat os) (8*(length os))) 64" definition SHA512_hLen :: nat where "SHA512_hLen = 64" lemma SHA512octets_len: "length (SHA512octets os) = 64" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA512 ?M ?l < 2^512" using SHA512_bnd by blast have 2: "(512::nat) = 8*64" by force have 3: "SHA512 ?M ?l < (2^8)^64" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA512 ?M ?l) 64 ) = 64" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA512octets_def by presburger qed lemma SHA512octets_len2: "\x. length (SHA512octets x) = 64" using SHA512octets_len by blast lemma SHA512octets_valid: "octets_valid (SHA512octets os)" using nat_to_words_len_valid SHA512octets_def by presburger lemma SHA512octets_valid2: "\x. octets_valid (SHA512octets x)" using SHA512octets_valid by satx subsection \SHA-512/224\ definition SHA512_224octets :: "octets \ octets" where "SHA512_224octets os = nat_to_octets_len (SHA512_224 (octets_to_nat os) (8*(length os))) 28" definition SHA512_224_hLen :: nat where "SHA512_224_hLen = 28" lemma SHA512_224octets_len: "length (SHA512_224octets os) = 28" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA512_224 ?M ?l < 2^224" using SHA512_224_bnd by blast have 2: "(224::nat) = 8*28" by force have 3: "SHA512_224 ?M ?l < (2^8)^28" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA512_224 ?M ?l) 28 ) = 28" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA512_224octets_def by presburger qed lemma SHA512_224octets_len2: "\x. length (SHA512_224octets x) = 28" using SHA512_224octets_len by blast lemma SHA512_224octets_valid: "octets_valid (SHA512_224octets os)" using nat_to_words_len_valid SHA512_224octets_def by presburger lemma SHA512_224octets_valid2: "\x. octets_valid (SHA512_224octets x)" using SHA512_224octets_valid by satx subsection \SHA-512/256\ definition SHA512_256octets :: "octets \ octets" where "SHA512_256octets os = nat_to_octets_len (SHA512_256 (octets_to_nat os) (8*(length os))) 32" definition SHA512_256_hLen :: nat where "SHA512_256_hLen = 32" lemma SHA512_256octets_len: "length (SHA512_256octets os) = 32" proof - let ?M = "(octets_to_nat os)" let ?l = "8*(length os)" have 1: "SHA512_256 ?M ?l < 2^256" using SHA512_256_bnd by blast have 2: "(256::nat) = 8*32" by force have 3: "SHA512_256 ?M ?l < (2^8)^32" by (metis 1 2 power_mult) have 4: "length ( nat_to_octets_len (SHA512_256 ?M ?l) 32 ) = 32" using 3 nat_to_words_len_upbnd by force show ?thesis using 4 SHA512_256octets_def by presburger qed lemma SHA512_256octets_len2: "\x. length (SHA512_256octets x) = 32" using SHA512_256octets_len by blast lemma SHA512_256octets_valid: "octets_valid (SHA512_256octets os)" using nat_to_words_len_valid SHA512_256octets_def by presburger lemma SHA512_256octets_valid2: "\x. octets_valid (SHA512_256octets x)" using SHA512_256octets_valid by satx end diff --git a/thys/Crypto_Standards/SEC1v2_0.thy b/thys/Crypto_Standards/SEC1v2_0.thy --- a/thys/Crypto_Standards/SEC1v2_0.thy +++ b/thys/Crypto_Standards/SEC1v2_0.thy @@ -1,3409 +1,3407 @@ theory SEC1v2_0 imports "Words" "EC_Common" begin text \ https://www.secg.org/sec1-v2.pdf SEC 1 v2.0 is the current standard from the Standards for Efficient Cryptography Group (SECG) for elliptic curve cryptography (ECC). The text of the standard is available at the above link. In this theory, we translate the functions defined in SEC 1 v2.0 to HOL. We adhere as closely to the written standard as possible, including function and variable names. It should be clear to anyone, regardless of their experience with HOL, that this translation exactly matches the standard. Variances from the written standard are (over-)explained in comments. Note: We only consider curves of the form \y\<^sup>2 \ x\<^sup>3 + ax + b (mod p)\ over a finite prime field \F\<^sub>p\, 2 < p, where \4a\<^sup>3 + 27b\<^sup>2 \ 0 (mod p)\. The standard also considers curves over finite fields of characteristic 2. The residues_prime_gt2 locale fixes a prime p, where 2 < p, and the corresponding residue ring labeled R. That is, \R = \\<^sub>p\, the integers modulo p. Note that SEC 1 sometimes uses the letter R to represent points on an elliptic curve. Because R is used for the residue ring here, we will need to choose another letter when SEC 1 uses R. Otherwise we strive to keep the same variable names as the standard.\ context residues_prime_gt2 begin section \2. Mathematical Foundations\ text \The majority of this chapter of the standard is, as the title suggests, background information about finite fields and elliptic curves. HOL/Isabelle has well-built theories of finite fields and elliptic curves. We don't need to translate any of that now. We use what HOL already has. There are a few things in Section 2.3 Data Types and Conversions which need translation. See below.\ subsection \2.1-2.2 Finite Fields and Elliptic Curves\ text \The language used in the HOL formalization of finite fields might be opaque to non-mathematicians. We use this section only to point out things to the reader who might appreciate some reminders on the language used for finite fields and elliptic curves. First we state a few simple things about the prime p that will be useful in some proof below. For example, that p is odd and that the number of octets (8-bit bytes) needed to represent p is greater than 0. The third lemma here is useful when thinking about the functions that convert between points in the elliptic curve and a string of octets.\ lemma p_mod2: "p mod 2 = 1" using gt2 p_prime prime_odd_nat by presburger lemma octet_len_p: "0 < octet_length p" by (metis gt2 nat_to_words_nil_iff_zero2 neq0_conv not_less_iff_gr_or_eq zero_less_numeral) lemma octet_len_p': assumes "l = octet_length p" shows "1 \ l + 1 \ 1 \ 2*l + 1 \ l + 1 \ 2*l + 1" using assms octet_len_p by linarith text \The term "carrier" might be unknown to the non-math reader. For the finite field with p elements, where p is prime, the carrier is integers mod p, or simply the interval [0, p-1].\ lemma inCarrier: "(a \ carrier R) = (0 \ a \ a < p)" - by (metis int_ops(1) mod_pos_pos_trivial not_one_less_zero of_int_closed of_nat_le_0_iff - pos_mod_conj r_one res_mult_eq res_of_integer_eq residues_axioms residues_def - verit_comp_simplify1(3)) + by (simp add: res_carrier_eq) text \The carrier R is a ring of integers. But also, it's only the integers in [0,p-1]. Isabelle/HOL is a typed language, so we might need to convert an integer (int) in the carrier to a natural number (nat). This is not a problem, since all elements of the carrier are non-negative. So we can convert to a nat and back to an int and we are back to where we started.\ lemma inCarrierNatInt: assumes "a \ carrier R" shows "int (nat a) = a" - by (metis assms local.minus_minus m_gt_one nat_eq_iff2 nat_int of_nat_1 of_nat_less_iff - pos_mod_conj res_neg_eq zless_nat_conj) + using assms by (simp add: inCarrier) text \The term nonsingular is defined in Elliptic_Locale.thy and is done for a generic type of curve. Then Elliptic_Test.thy builds on Elliptic_Locale.thy builds on that theory but specifically for curves over the integers. So just to make it more clear for the case of integers, the definition of nonsingular here is exactly the familiar notion for elliptic curves as found in SEC 1, page 16.\ lemma nonsingularEq: assumes "nonsingular a b" shows "(4*a^3 + 27*b^2) mod p \ 0" by (metis nonsingular_def add_cong assms mult_cong res_of_integer_eq res_pow_eq zero_cong) lemma nonsingularEq_nat: fixes a b :: nat assumes "nonsingular (int a) (int b)" shows "(4*a^3 + 27*b^2) mod p \ 0" proof - have "(4*(int a)^3 + 27*(int b)^2) mod p \ 0" using assms nonsingularEq by meson then show ?thesis by (metis (mono_tags, opaque_lifting) cring_class.of_nat_add cring_class.of_nat_mult of_nat_mod of_nat_numeral of_nat_power semiring_1_class.of_nat_0) qed text \As we did for "nonsingular", we see that the definition "on_curve" is exactly the usual definition for curves over the integers, for example on page 6 of SEC 1.\ lemma onCurveEq: assumes "on_curve a b P" "P = Point x y" shows "y^2 mod p = (x^3 + a*x + b) mod p" proof - have "y [^] (2::nat) = x [^] (3::nat) \ a \ x \ b" using assms(1,2) on_curve_def by simp then show ?thesis by (metis add.commute mod_add_right_eq res_add_eq res_mult_eq res_pow_eq) qed lemma onCurveEq2: assumes "on_curve a b (Point x y)" shows "(x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)" proof - have 1: "y^2 mod p = (x^3 + a*x + b) mod p" using assms onCurveEq by blast have 2: "x \ carrier R \ y \ carrier R" using assms on_curve_def by auto show ?thesis using 1 2 on_curve_def by fast qed lemma onCurveEq3: assumes "(x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)" shows "on_curve a b (Point x y)" by (smt (z3) assms on_curve_def mod_add_right_eq point.case(2) res_add_eq res_mult_eq res_pow_eq) lemma onCurveEq4: "on_curve a b (Point x y) = ((x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p))" using onCurveEq2 onCurveEq3 by blast text \The number of points on an elliptic curve over a finite field is finite.\ lemma CurveFinite: "finite {P. on_curve a b P}" proof - let ?S = "{P. on_curve a b P}" have 1: "Infinity \ ?S" using on_curve_def by force let ?S' = "{P. P \ Infinity \ on_curve a b P}" have 2: "?S = {Infinity} \ ?S'" using 1 by fast have 3: "?S' = {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R) \ (y^2 mod p = (x^3 + a*x + b) mod p)}" by (metis (no_types, lifting) onCurveEq2 onCurveEq3 point.distinct(2) point.exhaust) have 4: "?S' \ {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R)}" using 3 by auto have 5: "finite {P. \x y. (P = Point x y) \ (x \ carrier R) \ (y \ carrier R)}" by (simp add: finite_image_set2) have 6: "finite ?S'" using 4 5 finite_subset by blast show ?thesis using 2 6 by force qed text \The points on a non-singular elliptic curve form a group. So if the number of points is a prime n, every point on the curve has order n.\ lemma CurveOrderPrime: assumes "a \ carrier R" "b \ carrier R" "nonsingular a b" "C = {P. on_curve a b P}" "prime n" "card C = n" "Q \ C" shows "point_mult a n Q = Infinity" proof - have 1: "ell_prime_field p (nat a) (nat b)" using assms(1,2,3) ell_prime_field_def R_def ell_prime_field_axioms_def inCarrierNatInt nonsingularEq_nat residues_prime_gt2_axioms by algebra interpret EPF: ell_prime_field p R "(nat a)" "(nat b)" using 1 apply blast using R_def by presburger have 2: "C = carrier EPF.curve" using assms(1,2,4) EPF.curve_def inCarrierNatInt by simp have 3: "Q [^]\<^bsub>EPF.curve\<^esub>n = \\<^bsub>EPF.curve\<^esub>" using 2 assms(5,6,7) EPF.curve.power_order_eq_one EPF.finite_curve by presburger show ?thesis by (metis 2 3 EPF.in_carrier_curve EPF.multEqPowInCurve EPF.one_curve assms(1,7) inCarrierNatInt) qed text \"opp" is defined in Elliptic_Locale. This is the "opposite" of a point on the curve, meaning you just negate the y coordinate (modulo p).\ lemma oppEq: "opp (Point x y) = Point x ((-y) mod p)" using opp_Point res_neg_eq by presburger lemma oppEq': fixes x y :: int assumes "0 < y" "y < p" shows "opp (Point x y) = Point x (p-y)" proof - let ?y = "(-y) mod p" have "?y = p - y" using assms(1,2) zmod_zminus1_eq_if by auto then show ?thesis using oppEq by presburger qed subsection \2.3 Data Types and Conversions\ text \We have translated routines for converting natural numbers from/to words of a given number of bits for many crypto standards. This can be found in Words.thy. We have abbreviations for converting natural numbers to/from octet string, and to/from bit strings. For ease of reference, we will discuss in comments the existing conversion methods for each mentioned in section 2.3 of SEC 1. New in this standard is converting between elliptic curve points and octet strings. We define those primitives below.\ subsubsection \2.3.1 Bit-String-to-Octet-String Conversion\ text\"Bit strings should be converted to octet strings as described in this section. Informally the idea is to pad the bit string with 0's on the left to make its length a multiple of 8, then chop the result up into octets." This conversion in Words.thy is named bits_to_octets_len. Note that there is another routine called bits_to_octets. This second routine will convert to octets but will "wash away" any "extra" 0 bits on the left so that the leftmost byte of the output is not zero. Whereas, the output of bits_to_octets_len will always be \l/8\, where l is the length of the input bit string, given that the input bit string is valid. A bit string is valid when each value in the list is either 0 or 1. Note bit strings are represented as lists, as are octet strings, so bits_to_octets_len [0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] = [0,9] bits_to_octets [0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] = [9] Bottom line: when this standard calls for conversion from a bit string to an octet string, we use bits_to_octets_len. \ subsubsection \2.3.2 Octet-String-to-Bit-String Conversion\ text \Similar to above, to convert from a valid octet string to a bit string, we use octets_to_bits_len in Words.thy. Again, there is a second routine called octets_to_bits that will "wash away" any high 0s. So octets_to_bits_len [0,9] = [0,0,0,0,0,0,0,0,0,0,0,0,1,0,0,1] octets_to_bits [0,9] = [1,0,0,1] The output of octets_to_bits_len will always be 8x as long as the input, when the input is a valid string of octets. Octets are valid when each octet is in the list is a value in [0, 255]. The output of octets_to_bits will always start with 1. Bottom line: when the standard calls for conversion from an octet string to a bit string, we use octets_to_bits_len. \ subsubsection \2.3.3 Elliptic-Curve-Point-to-Octet-String Conversion\ text \We have not dealt with storing elliptic curve points as octet strings in other standards. Thus this is a new data conversion primitive. Note that the standard writes ceiling(log2 p / 8). This is the same as the octet_length when p is an odd prime. (See Words/octetLenLog2Prime.) So we write octet_length instead. Also the standard says: Assign the value 0x02 to the single octet Y if (y mod 2) = 0, or the value 0x03 if (y mod 2) = 1. Here we write that as Y = [2 + (y mod 2)]. Remember that octets are modeled in Isabelle as a list of natural numbers. So Y will be [2] if y is even and will be [3] when y is odd. We hope that this is clear to the reader that it exactly matches the standard as written.\ definition point_to_octets_nocomp :: "nat \ nat \ octets" where "point_to_octets_nocomp x y = ( let l = octet_length p; X = nat_to_octets_len x l; Y = nat_to_octets_len y l in 4 # X @ Y )" definition point_to_octets_comp :: "nat \ nat \ octets" where "point_to_octets_comp x y = (let l = octet_length p; X = nat_to_octets_len x l; Y = [2 + (y mod 2)] in Y @ X )" definition point_to_octets :: "int point \ bool \ octets" where "point_to_octets P Compress = ( case P of Infinity \ [0] | Point x y \ ( if Compress then (point_to_octets_comp (nat x) (nat y)) else (point_to_octets_nocomp (nat x) (nat y)) ) )" text \Now a few lemmas about converting a point to octets. The first set of lemmas shows that we know the length of the output when the input is a point on the curve. For points not on the curve, we have a bound on the length of the resulting octets.\ lemma point_to_octets_len_Inf: "length (point_to_octets Infinity C) = 1" using point_to_octets_def by simp lemma point_to_octets_nocomp_len_bnd: assumes "l = octet_length p" shows "2*l + 1 \ length (point_to_octets_nocomp x y)" proof - let ?X = "nat_to_octets_len x l" have 1: "l \ length ?X" using assms(1) nat_to_words_len_len by auto let ?Y = "nat_to_octets_len y l" have 2: "l \ length ?Y" using assms(1) nat_to_words_len_len by auto have 3: "point_to_octets_nocomp x y = 4 # ?X @ ?Y" using point_to_octets_nocomp_def assms(1) by meson then show ?thesis using 1 2 by simp qed lemma point_to_octets_nocomp_len: assumes "x < p" "y < p" "l = octet_length p" shows "length (point_to_octets_nocomp x y) = 2*l + 1" proof - let ?X = "nat_to_octets_len x l" have 1: "length ?X = l" using assms(1,3) nat_to_word_len_mono' zero_less_numeral by blast let ?Y = "nat_to_octets_len y l" have 2: "length ?Y = l" using assms(2,3) nat_to_word_len_mono' zero_less_numeral by blast have 3: "point_to_octets_nocomp x y = 4 # ?X @ ?Y" using point_to_octets_nocomp_def assms(3) by meson show ?thesis using 1 2 3 by auto qed lemma point_to_octets_len_F_bnd: assumes "P = Point x y" "l = octet_length p" shows "2*l + 1 \ length (point_to_octets P False)" using assms point_to_octets_nocomp_len_bnd point_to_octets_def by auto lemma point_to_octets_len_F': assumes "P = Point (x::int) (y::int)" "x < p" "y < p" "l = octet_length p" shows "length (point_to_octets P False) = 2*l + 1" proof - have x2: "(nat x) < p" using assms(2) m_gt_one by linarith have y2: "(nat y) < p" using assms(3) m_gt_one by linarith have "length (point_to_octets P False) = length (point_to_octets_nocomp (nat x) (nat y))" using assms(1) point_to_octets_def by simp then show ?thesis using assms(4) point_to_octets_nocomp_len x2 y2 by presburger qed lemma point_to_octets_len_F: assumes "on_curve a b P" "l = octet_length p" "P \ Infinity" shows "length (point_to_octets P False) = 2*l + 1" proof - obtain x and y where 1: "P = Point x y" using assms(3) by (meson point.exhaust) have x1: "x < p" using 1 assms(1) on_curve_def inCarrier by auto have y1: "y < p" using 1 assms(1) on_curve_def inCarrier by auto show ?thesis using 1 x1 y1 assms(2) point_to_octets_len_F' by blast qed lemma point_to_octets_comp_len_bnd: assumes "l = octet_length p" shows "l + 1 \ length (point_to_octets_comp x y)" proof - let ?X = "nat_to_octets_len x l" have 1: "l \ length ?X" using assms(1) nat_to_words_len_len by auto let ?Y = "[2 + (y mod 2)]" have 2: "length ?Y = 1" by simp have 3: "point_to_octets_comp x y = ?Y @ ?X" using point_to_octets_comp_def assms(1) by meson then show ?thesis using 1 2 by simp qed lemma point_to_octets_comp_len: assumes "x < p" "l = octet_length p" shows "length (point_to_octets_comp x y) = l + 1" proof - let ?X = "nat_to_octets_len x l" have 1: "length ?X = l" using assms(1,2) nat_to_word_len_mono' zero_less_numeral by blast then show ?thesis by (metis 1 assms(2) point_to_octets_comp_def One_nat_def length_append list.size(3,4) Suc_eq_plus1 plus_1_eq_Suc) qed lemma point_to_octets_len_T_bnd: assumes "P = Point x y" "l = octet_length p" shows "l + 1 \ length (point_to_octets P True)" using assms point_to_octets_comp_len_bnd point_to_octets_def by auto lemma point_to_octets_len_T': assumes "P = Point (x::int) (y::int)" "x < p" "l = octet_length p" shows "length (point_to_octets P True) = l + 1" proof - have 1: "(nat x) < p" using assms(2) m_gt_one by linarith have "length (point_to_octets P True) = length (point_to_octets_comp (nat x) (nat y))" using assms(1) point_to_octets_def by simp then show ?thesis using assms(3) point_to_octets_comp_len 1 by presburger qed lemma point_to_octets_len_T: assumes "on_curve a b P" "l = octet_length p" "P \ Infinity" shows "length (point_to_octets P True) = l + 1" proof - obtain x and y where 1: "P = Point x y" using assms(3) by (meson point.exhaust) have x1: "x < p" using 1 assms(1) on_curve_def inCarrier by auto show ?thesis using 1 x1 assms(2) point_to_octets_len_T' by blast qed lemma point_to_octets_len1: assumes "length (point_to_octets P C) = 1" shows "P = Infinity" apply (cases P) apply simp by (metis (full_types) add_diff_cancel_right' assms diff_is_0_eq' mult_is_0 neq0_conv octet_len_p point_to_octets_len_F_bnd point_to_octets_len_T_bnd zero_neq_numeral) lemma point_to_octets_len_oncurve: assumes "on_curve a b P" "l = octet_length p" "L = length (point_to_octets P C)" shows "L = 1 \ L = l+1 \ L = 2*l+1" apply (cases P) using point_to_octets_len_Inf assms(3) apply presburger apply (cases C) using point_to_octets_len_T assms(1,2,3) apply simp using point_to_octets_len_F assms(1,2,3) by simp text \Next set of lemmas: We know the output of point_to_octets is always a valid string of octets, no matter what is input. Recall that an octet is valid if its value is < 256 = 2^8.\ lemma point_to_octets_comp_valid: "octets_valid (point_to_octets_comp x y)" proof - have "2 + (y mod 2) < 256" by linarith then show ?thesis using point_to_octets_comp_def nat_to_words_len_valid words_valid_cons words_valid_concat by force qed lemma point_to_octets_nocomp_valid: "octets_valid (point_to_octets_nocomp x y)" proof - have "(4::nat) < 256" by fastforce then show ?thesis by (metis point_to_octets_nocomp_def nat_to_words_len_valid words_valid_cons words_valid_concat Twoto8) qed lemma point_to_octets_valid: "octets_valid (point_to_octets P C)" apply (cases P) apply (simp add: point_to_octets_def words_valid_cons words_valid_nil) apply (cases C) apply (simp add: point_to_octets_comp_valid point_to_octets_def) by (simp add: point_to_octets_nocomp_valid point_to_octets_def) text \Next up: converting a point to octets without compression is injective. A similar lemma when compression is used is a bit more complicated and is shown below.\ lemma point_to_octets_nocomp_inj: assumes "point_to_octets_nocomp x y = point_to_octets_nocomp x' y'" "x < p" "x' < p" shows "x = x' \ y = y'" proof - let ?l = "octet_length p" let ?M = "point_to_octets_nocomp x y" let ?M' = "point_to_octets_nocomp x' y'" let ?X = "nat_to_octets_len x ?l" let ?X' = "nat_to_octets_len x' ?l" let ?Y = "nat_to_octets_len y ?l" let ?Y' = "nat_to_octets_len y' ?l" have l: "length ?X = ?l \ length ?X' = ?l" by (meson assms(2,3) le_trans less_or_eq_imp_le nat_lowbnd_word_len2 nat_to_words_len_upbnd not_le zero_less_numeral) have m1: "?M = 4 # ?X @ ?Y" using point_to_octets_nocomp_def by meson have m2: "?M' = 4 # ?X' @ ?Y'" using point_to_octets_nocomp_def by meson have x1: "?X = take ?l (drop 1 ?M)" using m1 l by simp have x2: "?X' = take ?l (drop 1 ?M')" using m2 l by simp have x3: "?X = ?X'" using x1 x2 assms(1) by argo have x4: "x = x'" using nat_to_words_len_inj x3 by auto have y1: "?Y = drop (1+?l) ?M" using m1 l by simp have y2: "?Y' = drop (1+?l) ?M'" using m2 l by simp have y3: "?Y = ?Y'" using y1 y2 assms(1) by argo have y4: "y = y'" using nat_to_words_len_inj y3 by auto show ?thesis using x4 y4 by fast qed lemma point_to_octets_nocomp_inj': assumes "point_to_octets_nocomp x y = point_to_octets_nocomp x' y'" "x < p" "y < p" shows "x = x' \ y = y'" proof - let ?l = "octet_length p" let ?M = "point_to_octets_nocomp x y" let ?M' = "point_to_octets_nocomp x' y'" let ?X = "nat_to_octets_len x ?l" let ?X' = "nat_to_octets_len x' ?l" let ?Y = "nat_to_octets_len y ?l" let ?Y' = "nat_to_octets_len y' ?l" have l1: "length ?X = ?l \ length ?Y = ?l" by (meson assms(2,3) le_trans less_or_eq_imp_le nat_lowbnd_word_len2 nat_to_words_len_upbnd not_le zero_less_numeral) have l2: "?l \ length ?X' \ ?l \ length ?Y'" using nat_to_words_len_len by auto have m1: "?M = 4 # ?X @ ?Y" using point_to_octets_nocomp_def by meson have m2: "?M' = 4 # ?X' @ ?Y'" using point_to_octets_nocomp_def by meson have l3: "length ?M = 1 + length ?X + length ?Y" using m1 by force have l4: "length ?M' = 1 + length ?X' + length ?Y'" using m2 by force have l5: "length ?X + length ?Y = length ?X' + length ?Y'" using assms(1) l3 l4 by simp have l6: "length ?X' = ?l \ length ?Y' = ?l" using l5 l2 l1 by simp have x1: "?X = take ?l (drop 1 ?M)" using m1 l1 by simp have x2: "?X' = take ?l (drop 1 ?M')" using m2 l6 by simp have x3: "?X = ?X'" using x1 x2 assms(1) by argo have x4: "x = x'" using nat_to_words_len_inj x3 by auto have y1: "?Y = drop (1+?l) ?M" using m1 l1 by simp have y2: "?Y' = drop (1+?l) ?M'" using m2 l6 by simp have y3: "?Y = ?Y'" using y1 y2 assms(1) by argo have y4: "y = y'" using nat_to_words_len_inj y3 by auto show ?thesis using x4 y4 by fast qed text \Now a few lemmas about converting the "opposite" of a point to octets using compression.\ definition FlipY :: "nat \ nat" where "FlipY Y = (if Y = 2 then 3 else 2)" definition FlipYhd :: "octets \ octets" where "FlipYhd M = (FlipY (hd M)) # (tl M)" lemma point_to_octets_comp_opp: assumes "M = point_to_octets_comp x y" "M' = point_to_octets_comp x (p-y)" "y < p" "0 < y" shows "M = FlipYhd M' \ M' = FlipYhd M" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len x ?l" let ?Y = "2 + (y mod 2)" let ?y = "(p-y)" let ?Y' = "2 + (?y mod 2)" have 1: "M = ?Y # ?X \ M' = ?Y' # ?X" using assms(1,2) point_to_octets_comp_def by force have 2: "(y mod 2 = 0) = (?y mod 2 = 1)" by (metis p_mod2 assms(3) dvd_imp_mod_0 le_add_diff_inverse less_imp_le_nat odd_add odd_iff_mod_2_eq_one) have 3: "(?Y = 2) = (?Y' = 3)" using 2 by auto have 4: "(?Y = FlipY ?Y') \ (?Y' = FlipY ?Y)" using FlipY_def 3 by fastforce show ?thesis using FlipYhd_def 4 1 by force qed lemma point_to_octets_comp_opp': assumes "P = Point x y" "0 < y" "y < p" "M = point_to_octets P True" "M' = point_to_octets (opp P) True" shows "M = FlipYhd M' \ M' = FlipYhd M" by (simp add: assms point_to_octets_comp_opp oppEq' point_to_octets_def nat_minus_as_int) subsubsection \2.3.4 Octet-String-to-Elliptic-Curve-Point Conversion\ text \This conversion function needs to know the elliptic curve coefficients a and b to check if the recovered point is really on the curve. We also check that the high octet is 4. And we insist that the octet string is valid (meaning every octet is a value < 256) and the number of octets must be 2*l+ 1, where l is the length of the modulus p in octets.\ definition octets_to_point_nocomp_validInput :: "int \ int \ octets \ bool" where "octets_to_point_nocomp_validInput a b M = ( let W = hd M; M' = drop 1 M; l = octet_length p; X = take l M'; x = int (octets_to_nat X); Y = drop l M'; y = int (octets_to_nat Y); P = Point x y in (octets_valid M) \ (W = 4) \ (on_curve a b P) \ (length M = 2*l + 1) )" definition octets_to_point_nocomp :: "int \ int \ octets \ int point option" where "octets_to_point_nocomp a b M = ( let W = hd M; M' = drop 1 M; l = octet_length p; X = take l M'; x = int (octets_to_nat X); Y = drop l M'; y = int (octets_to_nat Y); P = Point x y in if (octets_to_point_nocomp_validInput a b M) then Some P else None )" text \Recovering a point when compression was used is not as straight forward. We need to find a square root mod p, which, depending on p, is not always simple to do. If p mod 4 = 3, then it is easy to write down the square root of a quadratic residue, but the SEC 1 standard does not require that p meet that restriction. In this definition, we check if \ is a quadratic residue and let HOL return one of its square roots, \. There is an issue that is not discussed in the standard. In the case when \ = 0, there is only one square root < p, namely 0. If \ = 0 and y = 1, the standard says to set y' = p. This will solve the elliptic curve equation but p is not a field element modulo p. To handle this, we add a check that is not described in the standard. If \ = 0, we insist that y must be 0 (correspondingly Y must be 2). Note that we define a variable T below that does not appear in the standard explicitly. T is for "test" and it just tests that the high octet is either 2 or 3. (See step 2.3 in section 2.3.4 in the standard.)\ definition octets_to_point_comp_validInput :: "int \ int \ octets \ bool" where "octets_to_point_comp_validInput a b M = ( let Y = hd M; T = (Y = 2) \ (Y = 3); l = octet_length p; X = drop 1 M; x = int (octets_to_nat X); \ = (x^3 + a*x + b) mod p in (octets_valid M) \ (x < p) \ T \ (\ = 0 \ Y = 2) \ (QuadRes' p \) \ (length M = l + 1) )" definition octets_to_point_comp :: "int \ int \ octets \ int point option" where "octets_to_point_comp a b M = ( let Y = hd M; y = Y - 2; X = drop 1 M; x = int (octets_to_nat X); \ = (x^3 + a*x + b) mod p; \ = some_root_nat p \; y' = int (if \ mod 2 = y then \ else (p-\)); P = Point x y' in if (octets_to_point_comp_validInput a b M) then Some P else None )" definition octets_to_point_validInput :: "int \ int \ octets \ bool" where "octets_to_point_validInput a b M \ (M = [0]) \ (octets_to_point_comp_validInput a b M) \ (octets_to_point_nocomp_validInput a b M)" definition octets_to_point :: "int \ int \ octets \ int point option" where "octets_to_point a b M = ( let lp = octet_length p; lM = length M in if M = [0] then Some Infinity else ( if lM = lp + 1 then octets_to_point_comp a b M else ( if lM = 2*lp + 1 then octets_to_point_nocomp a b M else None ) ) )" text \octets_to_point produces the point at infinity in only one way.\ lemma octets2PointNoCompNotInf: assumes "octets_to_point_nocomp a b M = Some P" shows "P \ Infinity" by (smt (z3) assms octets_to_point_nocomp_def option.distinct(1) option.inject point.distinct(1)) lemma octets2PointCompNotInf: assumes "octets_to_point_comp a b M = Some P" shows "P \ Infinity" by (smt (z3) assms octets_to_point_comp_def option.distinct(1) option.inject point.distinct(1)) lemma octets2PointInf1: "octets_to_point a b [0] = Some Infinity" using octets_to_point_def by presburger lemma octets2PointInf2: assumes "octets_to_point a b M = Some Infinity" shows "M = [0]" by (smt (z3) assms octets_to_point_def octets2PointNoCompNotInf octets2PointCompNotInf option.distinct(1)) lemma octets2PointInf: "(octets_to_point a b M = Some Infinity) = (M = [0])" using octets2PointInf1 octets2PointInf2 by blast text \octets_to_point produces a point on the curve or None.\ lemma octets2PointNoCompOnCurve: assumes "octets_to_point_nocomp a b M = Some P" shows "on_curve a b P" by (smt (z3) assms octets_to_point_nocomp_def octets_to_point_nocomp_validInput_def option.distinct(1) option.inject) lemma octets2PointCompOnCurve: assumes "octets_to_point_comp a b M = Some P" shows "on_curve a b P" proof - let ?Y = "hd M" let ?T = "(?Y = 2) \ (?Y = 3)" let ?y = "?Y - 2" let ?X = "drop 1 M" let ?x = "int (octets_to_nat ?X)" let ?\ = "(?x^3 + a*?x + b) mod p" let ?\ = "some_root_nat p ?\" let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" have 1: "P = Point ?x ?y'" by (smt (z3) assms(1) octets_to_point_comp_def option.distinct(1) option.inject) have 2: "octets_to_point_comp_validInput a b M" by (smt (verit, best) assms(1) octets_to_point_comp_def option.simps(3)) have 3: "QuadRes' p ?\" using 2 octets_to_point_comp_validInput_def by meson have a1: "?\ = 0 \ ?y = 0" by (metis 2 octets_to_point_comp_validInput_def diff_self_eq_0) have a2: "?\ \ 0 \ 0 < ?\" by (metis (mono_tags, opaque_lifting) 3 QuadRes'HasNatRoot bits_mod_0 mod_mod_trivial bot_nat_0.not_eq_extremum cong_def int_ops(1) power_zero_numeral) have x1: "0 \ ?x" using of_nat_0_le_iff by blast have x2: "?x < p" using 2 octets_to_point_comp_validInput_def by meson have x3: "?x \ carrier R" using x1 x2 inCarrier by blast have y1: "0 \ ?y'" using of_nat_0_le_iff by blast have y2: "?y' < p" by (smt (verit, ccfv_threshold) 3 a1 a2 QuadRes'HasNatRoot dvd_imp_mod_0 even_diff_nat gr0I int_ops(6) of_nat_0_less_iff zero_less_diff) have y3: "?y' \ carrier R" using y1 y2 inCarrier by blast have y4: "?y'^2 mod p = (?x^3 + a*?x + b) mod p" by (smt (z3) 3 QuadRes'HasNatRoot QuadRes'HasTwoNatRoots cong_def mod_mod_trivial of_nat_power) show ?thesis using 1 y4 x3 y3 onCurveEq3 by presburger qed lemma octets2PointOnCurve: assumes "octets_to_point a b M = Some P" shows "on_curve a b P" by (smt (verit) assms octets2PointCompOnCurve octets2PointNoCompOnCurve octets_to_point_def on_curve_infinity option.inject option.simps(3)) text \Now we have some lemmas about converting a point to octets then back to a point. First when compression is not used.\ lemma assumes "on_curve a b P" "P = Point x y" "M = point_to_octets_nocomp (nat x) (nat y)" shows point2OctetsNoComp_validInput: "octets_to_point_nocomp_validInput a b M" and point2Octets2PointNoComp: "octets_to_point_nocomp a b M = Some P" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len (nat x) ?l" let ?Y = "nat_to_octets_len (nat y) ?l" have x1: "nat x < p" using assms(1,2) on_curve_def inCarrier by fastforce have x2: "length ?X = ?l" using nat_to_word_len_mono' x1 zero_less_numeral by blast have y1: "nat y < p" using assms(1,2) on_curve_def inCarrier by fastforce have y2: "length ?Y = ?l" using nat_to_word_len_mono' y1 zero_less_numeral by blast have m1: "M = 4 # ?X @ ?Y" using assms(3) point_to_octets_nocomp_def by meson have m2: "octets_valid M" using assms(3) point_to_octets_nocomp_valid by blast let ?M' = "drop 1 M" let ?X' = "take ?l ?M'" let ?Y' = "drop ?l ?M'" have x3: "?X = ?X'" using m1 x2 by fastforce have y3: "?Y = ?Y'" using m1 x2 by fastforce let ?x = "int (octets_to_nat ?X')" let ?y = "int (octets_to_nat ?Y')" have x4: "?x = int (nat x)" by (metis x3 nat_to_words_len_to_nat zero_less_numeral) have x5: "?x = x" using assms(1,2) on_curve_def inCarrierNatInt x4 by auto have y4: "?y = int (nat y)" by (metis y3 nat_to_words_len_to_nat zero_less_numeral) have y5: "?y = y" using assms(1,2) on_curve_def inCarrierNatInt y4 by auto show "octets_to_point_nocomp_validInput a b M" using m2 m1 octets_to_point_nocomp_validInput_def assms(1,2) x2 x5 y2 y5 by auto then show "octets_to_point_nocomp a b M = Some P" using assms(2) octets_to_point_nocomp_def x5 y5 by presburger qed text \Continuing lemmas about converting a point to octets then back to a point: Now for when compression is used.\ lemma assumes "on_curve a b P" "P = Point x y" "M = point_to_octets_comp (nat x) (nat y)" shows point2OctetsValidInputComp: "octets_to_point_comp_validInput a b M" and point2Octets2PointComp: "octets_to_point_comp a b M = Some P" proof - let ?l = "octet_length p" let ?X = "nat_to_octets_len (nat x) ?l" let ?Y = "[2 + ((nat y) mod 2)]" have m1: "M = ?Y @ ?X" using assms(3) point_to_octets_comp_def by presburger let ?Y' = "hd M" have y1: "?Y' = 2 + (nat y) mod 2" using m1 by fastforce let ?T = "?Y' = 2 \ ?Y' = 3" have T1: "?T" using y1 by force let ?X' = "drop 1 M" have x1: "?X' = ?X" using m1 by auto let ?x = "int (octets_to_nat ?X')" have x2: "?x = int (nat x)" by (metis x1 nat_to_words_len_to_nat zero_less_numeral) have x3: "?x = x" using assms(1,2) on_curve_def inCarrierNatInt x2 by auto have x4: "?x < p" using x3 assms(1,2) inCarrier on_curve_def by auto let ?\ = "((?x^3 + a*?x + b) mod p)" have y2: "y^2 mod p = ?\" using assms(1,2) onCurveEq x3 by presburger have y3: "0 \ y \ y < p" using assms(1,2) inCarrier onCurveEq2 by blast have a1: "QuadRes' p ?\" by (metis y2 QuadRes'_def QuadResImplQuadRes'2 QuadRes_def cong_def mod_mod_trivial p_prime) have y4: "?\ = 0 \ y^2 mod p = 0" using y2 by argo have y5: "?\ = 0 \ y = 0" by (meson y4 y3 ZeroSqrtMod p_prime prime_nat_int_transfer) have y6: "?\ = 0 \ ?Y' = 2" using y1 y5 by fastforce have l1: "length ?X = ?l" using nat_to_word_len_mono' x3 x4 by force have l2: "length M = ?l + 1" using l1 m1 by auto show A1: "octets_to_point_comp_validInput a b M" by (metis assms(3) x4 T1 y6 a1 l2 octets_to_point_comp_validInput_def point_to_octets_comp_valid) let ?y = "?Y' - 2" have y7: "?y = (nat y) mod 2" using y1 by force let ?\ = "some_root_nat p ?\" have b1: "?\ < p" using QuadRes'HasNatRoot a1 by blast let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" let ?P = "Point ?x ?y'" have A2: "octets_to_point_comp a b M = Some ?P" by (smt (verit) A1 octets_to_point_comp_def) - show "octets_to_point_comp a b M = Some P" proof (cases "?\ = 0") + show "octets_to_point_comp a b M = Some P" + proof (cases "?\ = 0") case T0: True have T1: "y = 0" by (metis ZeroSqrtMod y4 y3 p_prime T0 prime_nat_int_transfer) show ?thesis using A2 T0 T1 ZeroSqrtMod' assms(2) p_prime x3 y7 by force next case False - then have "0 < ?\" by (smt (verit, best) m_gt_one pos_mod_conj) + then have "0 < ?\" using m_gt_one by (auto intro: le_neq_trans) then have "(nat y) = ?\ \ (nat y) = p - ?\" by (metis Euclidean_Rings.pos_mod_bound QuadResHasExactlyTwoRoots2 a1 gt2 int_ops(1) nat_int not_less of_nat_le_0_iff p_prime y2 y3) then show ?thesis by (smt (verit) A2 assms(2) b1 diff_is_0_eq' even_diff_nat int_nat_eq less_imp_le_nat less_numeral_extra(3) odd_add odd_iff_mod_2_eq_one p_mod2 x3 y7 y3 zero_less_diff) qed qed text \Now it's easy to show that for points on the curve, point_to_octets_comp is injective.\ lemma point_to_octets_comp_inj: assumes "point_to_octets_comp x y = point_to_octets_comp x' y'" "on_curve a b (Point (int x) (int y))" "on_curve a b (Point (int x') (int y'))" shows "x = x' \ y = y'" by (metis assms(1,2,3) nat_int option.inject point.inject point2Octets2PointComp) lemma point_to_octets_comp_inj': assumes "point_to_octets_comp (nat x) (nat y) = point_to_octets_comp (nat x') (nat y')" "on_curve a b (Point x y)" "on_curve a b (Point x' y')" shows "x = x' \ y = y'" by (metis assms(1,2,3) option.inject point.inject point2Octets2PointComp) text \In the Public Key Recovery Operation, you have "forgotten" y. So you just try to recover the whole point by guessing that y is even. So we have a few lemmas here that will be useful for that recovery operation.\ lemma point2Octets2PointComp_forgot_Y2: assumes "on_curve a b P" "P = Point x y" "y mod 2 = 0" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some P" proof - have "M = point_to_octets_comp (nat x) (nat y)" by (metis assms(3,4,5,6) point_to_octets_comp_def add.right_neutral append_Cons append_Nil dvd_imp_mod_0 even_mod_2_iff even_of_nat int_nat_eq odd_add) then show ?thesis using assms(1,2) point2Octets2PointComp by blast qed lemma point2Octets2PointComp_forgot_Y2': assumes "on_curve a b P" "P = Point x y" "y mod 2 = 1" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some (opp P)" proof - have 0: "y \ 0" using assms(3) by force have 1: "0 < y" using 0 assms(1,2) on_curve_def inCarrier by force have 2: "y < p" using assms(1,2) on_curve_def inCarrier by force let ?y = "p - y" let ?P = "Point x ?y" have 3: "?P = opp P" using oppEq'[OF 1 2] assms(2) by presburger have 4: "(nat ?y) mod 2 = 0" by (metis assms(3) cancel_comm_monoid_add_class.diff_cancel even_of_nat int_nat_eq mod_0 mod_diff_right_eq not_mod_2_eq_1_eq_0 odd_iff_mod_2_eq_one p_mod2) have 5: "M = point_to_octets_comp (nat x) (nat ?y)" using 4 assms(4,5,6) point_to_octets_comp_def by fastforce show ?thesis using 3 5 assms(1,2) point2Octets2PointComp opp_closed by auto qed lemma point2Octets2PointComp_PoppP: assumes "on_curve a b P" "P = Point x y" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point_comp a b M = Some P \ octets_to_point_comp a b M = Some (opp P)" apply (cases "y mod 2 = 0") using assms(1,2,3,4,5) point2Octets2PointComp_forgot_Y2 apply presburger using assms(1,2,3,4,5) point2Octets2PointComp_forgot_Y2' by presburger lemma point2Octets2PointComp_PoppP': assumes "on_curve a b P" "P = Point x y" "l = octet_length p" "X = nat_to_octets_len (nat x) l" "M = 2 # X" shows "octets_to_point a b M = Some P \ octets_to_point a b M = Some (opp P)" proof - have "x < p" using assms(1,2) inCarrier onCurveEq2 by blast then have "nat x < p" using m_gt_one by linarith then have "length X = l" using assms(3,4) nat_to_word_len_mono' zero_less_numeral by blast then have "octets_to_point a b M = octets_to_point_comp a b M" by (simp add: octets_to_point_def assms(3,5)) then show ?thesis using assms point2Octets2PointComp_PoppP by algebra qed text \Putting together what we have proved above, if P is on the curve, then converting P to octets and then back to a point, whether or not compression is used, will get back the original P.\ lemma point2Octets2Point: assumes "on_curve a b P" "M = point_to_octets P C" shows "octets_to_point a b M = Some P" proof (cases P) case Infinity then show ?thesis by (simp add: assms(2) octets_to_point_def point_to_octets_def) next case P: (Point x y) then show ?thesis proof (cases C) case T0: True let ?lp = "octet_length p" have T1: "length M = ?lp + 1" using P assms(1,2) T0 point_to_octets_len_T by fastforce have T2: "octets_to_point a b M = octets_to_point_comp a b M" by (smt (verit, best) octet_len_p' T1 octets_to_point_def One_nat_def Suc_eq_plus1 list.size(3,4)) show ?thesis using T2 P T0 assms(1,2) point2Octets2PointComp point_to_octets_def by auto next case F0: False let ?lp = "octet_length p" have l: "0 < ?lp" by (metis gr_implies_not_zero gt2 nat_to_words_nil_iff_zero2 neq0_conv zero_less_numeral) have F1: "length M = 2*?lp + 1" using P assms(1,2) F0 point_to_octets_len_F by fastforce have F2: "octets_to_point a b M = octets_to_point_nocomp a b M" by (smt (verit, best) F1 octet_len_p' octets_to_point_def One_nat_def Suc_eq_plus1 list.size(3,4)) then show ?thesis using F2 P F0 assms(1,2) point2Octets2PointNoComp point_to_octets_def by simp qed qed text \Now we think about the other direction. We start with octets. If the octets are a valid input for when compression is or is not used, then if we convert the octets to a point and then back again, we will get back the original octets.\ lemma octets2PointNoCompValidInIFF1: "(octets_to_point_nocomp a b M \ None) = (octets_to_point_nocomp_validInput a b M)" by (smt (z3) not_None_eq octets_to_point_nocomp_def) lemma octets2PointNoCompalidInIFF2: "(\P. octets_to_point_nocomp a b M = Some P) = (octets_to_point_nocomp_validInput a b M)" using octets2PointNoCompValidInIFF1 by force lemma octets2Point2OctetsNoComp: assumes "octets_to_point_nocomp a b M = Some P" shows "point_to_octets P False = M" proof - let ?W = "hd M" let ?M' = "drop 1 M" let ?l = "octet_length p" let ?X = "take ?l ?M'" let ?x = "int (octets_to_nat ?X)" let ?Y = "drop ?l ?M'" let ?y = "int (octets_to_nat ?Y)" let ?P = "Point ?x ?y" have V: "octets_to_point_nocomp_validInput a b M" using assms octets2PointNoCompValidInIFF1 by blast have P1: "?P = P" by (smt (z3) assms(1) V octets_to_point_nocomp_def option.inject) have M1: "length M = 2*?l + 1" using V octets_to_point_nocomp_validInput_def by meson have M2: "octets_valid M" using V octets_to_point_nocomp_validInput_def by meson have W1: "?W = 4" using V octets_to_point_nocomp_validInput_def by meson have X1: "length ?X = ?l" using M1 by force have X2: "octets_valid ?X" using M2 words_valid_take words_valid_drop by blast have Y1: "length ?Y = ?l" using M1 by force have Y2: "octets_valid ?Y" using M2 words_valid_drop by blast have M3: "M = ?W # ?X @ ?Y" by (metis M1 One_nat_def append_take_drop_id drop_0 drop_Suc_Cons hd_Cons_tl leD le_add2 list.size(3) zero_less_one) have O1: "point_to_octets P False = point_to_octets_nocomp (nat ?x) (nat ?y)" using P1 point_to_octets_def by force have x1: "nat ?x = octets_to_nat ?X" by (meson nat_int) have y1: "nat ?y = octets_to_nat ?Y" by (meson nat_int) let ?X' = "nat_to_octets_len (nat ?x) ?l" have X3: "?X' = ?X" by (metis X1 X2 x1 words_to_nat_to_words_len2) let ?Y' = "nat_to_octets_len (nat ?y) ?l" have Y3: "?Y' = ?Y" by (metis Y1 Y2 y1 words_to_nat_to_words_len2) show ?thesis using M3 O1 X3 Y3 W1 point_to_octets_nocomp_def by algebra qed lemma octets2PointCompValidInIFF1: "(octets_to_point_comp a b M \ None) = (octets_to_point_comp_validInput a b M)" by (smt (z3) not_None_eq octets_to_point_comp_def) lemma octets2PointCompValidInIFF2: "(\P. octets_to_point_comp a b M = Some P) = (octets_to_point_comp_validInput a b M)" using octets2PointCompValidInIFF1 by force lemma octets2Point2OctetsComp: assumes "octets_to_point_comp a b M = Some P" shows "point_to_octets P True = M" proof - let ?Y = "hd M" let ?T = "(?Y = 2) \ (?Y = 3)" let ?y = "?Y - 2" let ?l = "octet_length p" let ?X = "drop 1 M" let ?x = "int (octets_to_nat ?X)" let ?\ = "(?x^3 + a*?x + b) mod p" let ?\ = "some_root_nat p ?\" let ?y' = "int (if ?\ mod 2 = ?y then ?\ else (p-?\))" let ?P = "Point ?x ?y'" have V: "octets_to_point_comp_validInput a b M" using assms octets2PointCompValidInIFF1 by blast have P1: "?P = P" by (smt (z3) assms(1) V octets_to_point_comp_def option.inject) let ?X' = "nat_to_octets_len (nat ?x) ?l" let ?Y' = "[2 + ((nat ?y') mod 2)]" let ?M = "?Y' @ ?X'" have O1: "point_to_octets P True = ?M" using P1 point_to_octets_def point_to_octets_comp_def by force have x1: "?X = ?X'" by (metis V add_diff_cancel_right' length_drop nat_int octets_to_point_comp_validInput_def words_to_nat_to_words_len2 words_valid_drop) have a1: "QuadRes' p ?\" by (meson V octets_to_point_comp_validInput_def) have M1: "length M = ?l + 1" by (meson V octets_to_point_comp_validInput_def) have y1: "?y = 0 \ ?y = 1" by (metis V add_diff_cancel_right' diff_self_eq_0 numeral_2_eq_2 numeral_3_eq_3 octets_to_point_comp_validInput_def plus_1_eq_Suc) have y3: "(nat ?y') mod 2 = ?y" by (smt (verit, ccfv_threshold) QuadRes'HasNatRoot a1 even_add even_diff_nat less_imp_of_nat_less mod2_eq_if nat_int p_mod2 y1) have y4: "[?Y] = ?Y'" by (metis V add.commute add.left_neutral add_diff_cancel_right' numeral_2_eq_2 numeral_3_eq_3 octets_to_point_comp_validInput_def plus_1_eq_Suc y3) have M2: "?M = M" by (metis (mono_tags, opaque_lifting) x1 y4 M1 One_nat_def Suc_eq_plus1 append_Cons append_Nil drop0 drop_Suc list.exhaust_sel list.size(3) nat.simps(3)) show ?thesis using O1 M2 by argo qed lemma octets2PointValidInIFF1: "(octets_to_point a b M \ None) = (octets_to_point_validInput a b M)" by (smt (verit, best) R_def octets2Point2OctetsNoComp octets2PointCompValidInIFF1 octets2PointNoCompOnCurve octets_to_point_comp_validInput_def octets_to_point_def octets_to_point_nocomp_def octets_to_point_validInput_def option.simps(3) point2Octets2Point residues_prime_gt2_axioms) lemma octets2PointValidInIFF2: "(\P. octets_to_point a b M = Some P) = (octets_to_point_validInput a b M)" using octets2PointValidInIFF1 by auto lemma octets2Point2Octets: assumes "octets_to_point a b M = Some P" "lM = length M" "lp = octet_length p" "C = (if lM = lp + 1 then True else False)" shows "point_to_octets P C = M" proof - have 1: "(M = [0]) \ (lM = lp + 1) \ (lM = 2*lp + 1)" by (smt (z3) assms(1,2,3) octets_to_point_def option.simps(3)) have 2: "M = [0] \ point_to_octets P C = M" using assms(1) octets_to_point_def point_to_octets_def by force have 3: "lM = lp + 1 \ point_to_octets P C = M" using 2 assms(1,2,3,4) octets_to_point_def octets2Point2OctetsComp by force have 4: "lM = 2*lp + 1 \ point_to_octets P C = M" using 2 3 assms(1,2,3,4) octets_to_point_def octets2Point2OctetsNoComp by force show ?thesis using 1 2 3 4 by fast qed subsubsection \2.3.5 Field-Element-to-Octet-String Conversion\ text \In this translation of SEC 1, we only consider prime fields. So a "field element" is simply a natural number in the range [0, p-1]. So converting field elements to or from octets is the same as converting between nats and octets. The only detail to keep in mind is that the resulting octet string should have the same number of octets as the prime p. So if x is a field element (natural number in [0,p-1]), then we convert x to an octet string M with: M = nat_to_octets_len x (octet_length p) \ subsubsection \2.3.6 Octet-String-to-Field-Element Conversion\ text \Again, we only consider prime fields in this translation. So a field element is a natural number in the range [0,p-1]. We can convert any octet string to a nat using octets_to_nat, defined in Words.thy. If the resulting nat is p or greater, it should be rejected. Note that you might need to check if the input octets were valid using octets_valid, which checks that each octet is a value no greater than 255.\ subsubsection \2.3.7 Integer-to-Octet-String Conversion\ text \Again, there are two conversion primitives for converting a non-negative integer, a.k.a. a natural number, to an octet string. nat_to_octets will convert a nat to an octet string so that the high octet is non-zero. nat_to_octets_len will produce an octet string that has at least a minimum number of octets. In either case, if you translate back to a natural number, you will get back the original value. For example: nat_to_octets 256 = [1,0] nat_to_octets_len 256 5 = [0,0,0,1,0] nat_to_octets_len 256 1 = [1,0] and octets_to_nat [1,0] = 256 octets_to_nat [0,0,0,1,0] = 256 \ subsubsection \2.3.8 Octet-String-to-Integer Conversion\ text \As described above, the routine to convert an octet string to a natural number is octets_to_nat. Of note, octets_valid ensures that a string of octets is valid, meaning each octet has the value 255 or less. So octets_valid [0,1,255,3] = True and octets_valid [0,1,256,3] = False.\ subsubsection \2.3.9 Field-Element-to-Integer Conversion\ text \Again, we only consider prime fields in this translation of SEC 1. So a field element is simply an integer in [0,p-1]. So there is no conversion required.\ section \3. Cryptographic Components\ subsection \3.1 Elliptic Curve Domain Parameters\ text \3.1.1.2.1 Elliptic Curve Domain Parameters over \F\<^sub>p\ Validation Primitive SEC 1 only allows a security level t of 80, 112, 128, 192, or 256. Then "twice the security level" will actually be 2*t except for t = 80 and 256, where it's merely close to 2*t.\ definition twiceSecurityLevel :: "nat \ nat" where "twiceSecurityLevel t = ( if (t = 80) then 192 else ( if (t = 256) then 521 else 2*t ) )" text \The standard writes \log 2 p\. Note that because p is odd (thus not a power of 2), \log 2 p\ = bit_length p. See Words.bitLenLog2Prime. Also, the standard explains that the number of points on the elliptic curve is n*h. (h is called the "cofactor".) It doesn't make a lot of sense to consider an elliptic curve that has zero points on it. So we must have 0 < h. (We know that 0 < n because the standard insists that n is prime.) So while not explicitly written in the standard, we include the requirement that 0 < h.\ definition ECdomainParametersValid :: "nat \ nat \ int point \ nat \ nat \ nat \ bool" where "ECdomainParametersValid a b G n h t \ ( (t \ {80, 112, 128, 192, 256}) \ (\log 2 p\ = twiceSecurityLevel t) \ (a \ carrier R) \ (b \ carrier R) \ (nonsingular a b) \ (n*h \ p) \ (prime n) \ card {P. on_curve a b P} = n*h \ (on_curve a b G) \ (G \ Infinity) \ (point_mult a n G = Infinity) \ (h \ 2^(t div 8)) \ (h = \((sqrt p + 1)^2) div n\) \ 0 < h \ (\B\{1..99}. (p^B mod n \ 1)) )" text \The security level t is not always explicitly listed in elliptic curve parameters.\ definition ECdomainParametersValid' :: "nat \ nat \ int point \ nat \ nat \ bool" where "ECdomainParametersValid' a b G n h \ (\t. ECdomainParametersValid a b G n h t)" lemma ECparamsValidImplies': "ECdomainParametersValid a b G n h t \ ECdomainParametersValid' a b G n h" using ECdomainParametersValid'_def by blast definition twiceSecurityLevelInv :: "nat \ nat" where "twiceSecurityLevelInv twoT = ( if (twoT = 192) then 80 else ( if (twoT = 521) then 256 else (twoT div 2) ) )" lemma twiceSecurityLevelInv: assumes "t \ {80, 112, 128, 192, 256}" shows "twiceSecurityLevelInv (twiceSecurityLevel t) = t" using assms twiceSecurityLevelInv_def twiceSecurityLevel_def by fastforce text \Even when not explicitly stated, the security level can be recovered from the bit length of the prime modulus p.\ lemma ECparamsValid'Implies: assumes "ECdomainParametersValid' a b G n h" "t = twiceSecurityLevelInv (nat \log 2 p\)" shows "ECdomainParametersValid a b G n h t" proof - obtain t' where t1: "ECdomainParametersValid a b G n h t'" using ECdomainParametersValid'_def assms(1) by presburger have t2: "t' \ {80, 112, 128, 192, 256}" using ECdomainParametersValid_def t1 by presburger have t3: "\log 2 p\ = twiceSecurityLevel t'" using ECdomainParametersValid_def t1 by presburger have t4: "t' = twiceSecurityLevelInv (nat \log 2 p\)" using nat_int t2 t3 twiceSecurityLevelInv by presburger show ?thesis using assms(2) t1 t4 by argo qed lemma ECparamsValid'Implies2: assumes "ECdomainParametersValid' a b G n h" "t = twiceSecurityLevelInv (bit_length p)" shows "ECdomainParametersValid a b G n h t" by (metis ECparamsValid'Implies assms(1,2) bitLenLog2Prime gt2 nat_int p_prime) lemma ECparamsValid_min_p: assumes "ECdomainParametersValid a b G n h t" shows "2^191 \ p" proof - have 1: "t \ {80, 112, 128, 192, 256}" using assms ECdomainParametersValid_def by algebra have 2: "bit_length p = twiceSecurityLevel t" by (metis assms bitLenLog2Prime gt2 nat_int p_prime ECdomainParametersValid_def) have 3: "192 \ twiceSecurityLevel t" using 1 twiceSecurityLevel_def by fastforce have 4: "192 \ bit_length p" using 2 3 by fastforce show ?thesis by (meson 4 less_bit_len less_le_trans numeral_less_iff semiring_norm(68,79,81)) qed text \The following lemma is convenient below when showing that valid EC domain parameters meet the definition of "Elliptic Prime Field" defined in EC_Common.\ lemma paramsValidEllPrimeField: assumes "ECdomainParametersValid a b G n h t" shows "a \ carrier R \ b \ carrier R \ (4*a^3 + 27*b^2) mod p \ 0" using ECdomainParametersValid_def assms nonsingularEq_nat by presburger lemma paramsValidEllPrimeField': assumes "ECdomainParametersValid a b G n h t" shows "ell_prime_field p a b" using assms ell_prime_field.intro ell_prime_field_axioms_def residues_prime_gt2.paramsValidEllPrimeField residues_prime_gt2_axioms by blast text \The following lemmas are useful for the Public Key Recovery Operation below.\ lemma p_div_n: assumes "ECdomainParametersValid a b G n h t" shows "p div n \ h" proof - let ?h = "((sqrt p + 1)^2) / n" have h1: "h = nat (floor ?h)" using assms(1) ECdomainParametersValid_def by fastforce have h2: "?h < h + 1" using h1 by linarith have f1: "(sqrt p + 1)^2 = p + 1 + 2*(sqrt p)" by (simp add: comm_semiring_1_class.power2_sum) have h3: "?h = (p + 2*(sqrt p) + 1) / n" using f1 by force have h4: "?h = (real p / real n) + ((2*(sqrt p) + 1)/n)" by (simp add: add_divide_distrib h3) have h5: "(real p / real n) \ ?h" using h4 by auto have h6: "(real p / real n) < h + 1" using h5 h2 by linarith have a1: "real (p div n) \ (real p / real n)" using real_of_nat_div4 by blast have a2: "p div n < h + 1" using h6 a1 by auto show ?thesis using a2 by simp qed lemma less_p_div_n: assumes "ECdomainParametersValid a b G n h t" "x < p" shows "x div n \ h" by (meson assms p_div_n div_le_mono le_trans less_imp_le_nat) lemma less_p_div_n': assumes "ECdomainParametersValid a b G n h t" "(int x) \ carrier R" shows "x div n \ h" by (meson assms(1,2) inCarrier less_p_div_n of_nat_less_iff) text \Next up we show that h < n. The number of points on the elliptic curve is n*h. For recommended curves, the cofactor h is typically small, like h < 10. The prime n is very large, around the size of p. So in some sense, it's not surprising that h < n. The only issue here is to show that given the definition of valid parameters.\ lemma h_less_n: assumes "ECdomainParametersValid a b G n h t" shows "h < n" proof - let ?t2 = "twiceSecurityLevel t" have t: "t \ {80, 112, 128, 192, 256}" using assms ECdomainParametersValid_def by algebra have t2: "2*t \ ?t2" using t twiceSecurityLevel_def by force have h: "h \ 2^(t div 8)" using assms ECdomainParametersValid_def by algebra have h1: "real h \ real 2^(t div 8)" using h by simp have h2: "0 < real (h + 1)" by auto have h3: "real (h+1) \ real (2^(t div 8) + 1)" using h by fastforce have p1: "\log 2 p\ = ?t2" using assms ECdomainParametersValid_def by algebra have p2: "bit_length p = ?t2" by (metis p1 bitLenLog2Prime gt2 of_nat_eq_iff p_prime) have p3: "2^(?t2-1) \ p" using bit_len_exact3 gt2 p2 by force have p4: "2^(2*t-1) \ p" by (meson p3 t2 diff_le_mono le_trans one_le_numeral power_increasing) have n1: "p div n \ h" using assms p_div_n by fast have n2: "\(real p / real n)\ \ h" by (simp add: n1 floor_divide_of_nat_eq) have n3: "(real p / real n) < h + 1" using n2 by linarith have n4: "prime n" using assms ECdomainParametersValid_def by algebra have n5: "0 < n" by (simp add: n4 prime_gt_0_nat) have n6: "0 < real n" using n5 by fastforce have n7: "real p < real n * (h+1)" using n3 n6 by (simp add: mult.commute pos_divide_less_eq) have n: "real p / real (h+1) < real n" using n7 h2 pos_divide_less_eq by blast have a1: "real 2^(2*t-1) / real (h+1) < real n" by (metis n p4 h2 less_le_trans not_less of_nat_power_less_of_nat_cancel_iff pos_divide_less_eq) have a2: "real 2^(2*t-1) / real (2^(t div 8)+1) \ real 2^(2*t-1) / real (h+1)" using h3 by (simp add: frac_le) have a3: "real 2^(2*t-1) / real (2^(t div 8)+1) < real n" using a1 a2 by argo have a4: "real 2^(t div 8) < real 2^(2*t-1) / real (2^(t div 8)+1)" using t by fastforce have "real h < real n" using a3 a4 h1 by argo then show ?thesis by simp qed text \It is sometimes helpful to know that n is an odd prime. In the following proof, we see if n = 2, then h = 1 and p = 2. We already know that p cannot be that small, so we are done. But also, if n = 2, the n*h = p, which also contradicts the definition of valid elliptic curve parameters.\ lemma n_not_2: assumes "ECdomainParametersValid a b G n h t" shows "2 < n" proof - have n0: "prime n" using assms ECdomainParametersValid_def by algebra have h0: "0 < h" using assms ECdomainParametersValid_def by algebra have p0: "2^191 \ p" using assms ECparamsValid_min_p by fast have n2: "n = 2 \ p \ 2" proof assume A: "n = 2" have A1: "h = 1" using A h0 h_less_n assms by fastforce have A2: "1 = \((sqrt p + 1)^2) div 2\" using A A1 assms ECdomainParametersValid_def by simp have A3: "((sqrt p + 1)^2) div 2 \ 2" using A2 by linarith have A4: "((sqrt p + 1)^2) \ 4" using A3 by simp show "p \ 2" using A4 real_le_rsqrt by force qed have "n \ 2" using p0 n2 by fastforce then show ?thesis using n0 le_neq_implies_less prime_ge_2_nat by presburger qed lemma n_odd: assumes "ECdomainParametersValid a b G n h t" shows "odd n" by (metis assms n_not_2 ECdomainParametersValid_def prime_odd_nat) text \In fact, n must be a great deal larger than 2.\ lemma ECparamsValid_min_n: assumes "ECdomainParametersValid a b G n h t" shows "2^95 < n" proof - have n1: "p div n \ h" using assms p_div_n by fast have n2: "\(real p / real n)\ \ h" by (simp add: n1 floor_divide_of_nat_eq) have n3: "(real p / real n) < h + 1" using n2 by linarith have n4: "prime n" using assms ECdomainParametersValid_def by algebra have n5: "0 < n" by (simp add: n4 prime_gt_0_nat) have n6: "0 < real n" using n5 by fastforce have n7: "real p < real n * (h+1)" using n3 n6 by (simp add: mult.commute pos_divide_less_eq) have n8: "p < n *(h+1)" by (metis n7 of_nat_less_imp_less semiring_1_class.of_nat_mult) have h1: "h+1 \ n" using assms h_less_n by fastforce have a1: "p < n^2" by (metis n8 h1 less_le_trans nat_1_add_1 nat_mult_le_cancel_disj power_add power_one_right) have a2: "2^191 < n^2" using a1 assms ECparamsValid_min_p by fastforce have a3: "((2::nat)^95)^2 = 2^190" by auto have a4: "((2::nat)^95)^2 < n^2" using a2 a3 by force show ?thesis using a4 power_less_imp_less_base by blast qed end (*residues_prime_gt2 context*) text \Now we know what valid elliptic curve parameters are. So we fix a set of parameters and define the cryptographic primitives below using those valid parameters.\ locale SEC1 = residues_prime_gt2 + fixes a b n h t :: nat and G :: "int point" assumes ECparamsValid: "ECdomainParametersValid a b G n h t" begin text \Elliptic_Locale defines elliptic curve operations. In this locale, we have fixed a set of elliptic curve parameters. We use the coefficients a and b to make convenient abbreviations for definitions found in Elliptic_Locale as well as a data conversion primitive defined above.\ abbreviation on_curve' :: "int point \ bool" where "on_curve' Q \ on_curve a b Q" abbreviation point_mult' :: "nat \ int point \ int point" where "point_mult' x Q \ point_mult a x Q" abbreviation add' :: "int point \ int point \ int point" where "add' P Q \ add a P Q" abbreviation octets_to_point' :: "octets \ int point option" where "octets_to_point' M \ octets_to_point a b M" text \We show that the valid elliptic curve parameters meet the assumptions of the ell_prime_field locale in EC_Common. We then write some properties in terms of "curve" which is defined in ell_prime_field. Of note: the definitions in Elliptic_Locale use additive notation for the elliptic curve group. For example, \add' P Q\, or "P + Q". In contrast, ell_prime_field uses multiplicative notation, so that the same operation is denoted \P\\<^bsub>EPF.curve\<^esub>Q\. \ sublocale EPF: ell_prime_field p R a b using ECparamsValid paramsValidEllPrimeField' apply blast by (simp add: R_def) lemma order_EPF_curve: "order EPF.curve = n*h" using order_def ECparamsValid ECdomainParametersValid_def EPF.curve_def by (metis partial_object.select_convs(1)) lemma order_EPF_curve_h1: assumes "h = 1" shows "order EPF.curve = n" using assms order_EPF_curve by simp lemma order_EPF_curve_h1': assumes "h = 1" shows "prime (order EPF.curve)" using assms order_EPF_curve_h1 ECparamsValid ECdomainParametersValid_def by blast text \When h=1, then all points on the curve can be written as some power of the generator G. (And any point on the curve can be picked as the generator, for that matter. It's just a cyclic group.)\ lemma EC_h1_cyclic: assumes "h = 1" shows "carrier EPF.curve = {Q. (\d G \ Infinity \ point_mult' n G = Infinity\ prime (n::nat)" using ECparamsValid ECdomainParametersValid_def by blast let ?S1 = "carrier EPF.curve" let ?S2 = "{Q. (\dd. on_curve' (point_mult' d G)" by (simp add: H point_mult_closed) have 4: "?S2 \ ?S1" using 3 EPF.curve_def by force show ?thesis using 1 2 4 EPF.finite_curve card_subset_eq by blast qed lemma EC_h1_cyclic': assumes "h = 1" "on_curve' P" shows "\dNext up, we write down some facts about the generator G, and its multiples P = d*G (or powers when written with multiplicative notation, P = G^d).\ lemma GonEPFcurve: "G \ carrier EPF.curve" using ECdomainParametersValid_def ECparamsValid EPF.in_carrier_curve by presburger lemma dGonEPFcurve: assumes "P = point_mult' d G" shows "P \ carrier EPF.curve" using GonEPFcurve EPF.multEqPowInCurve EPF.curve.nat_pow_closed EPF.in_carrier_curve assms by force lemma Gnot1onEPFcurve: "G \ \\<^bsub>EPF.curve\<^esub>" using ECdomainParametersValid_def ECparamsValid EPF.one_curve by presburger lemma nGeq1onEPFcurve: "G [^]\<^bsub>EPF.curve\<^esub> n = \\<^bsub>EPF.curve\<^esub>" proof - have "point_mult' n G = Infinity" using ECdomainParametersValid_def ECparamsValid by presburger then show ?thesis using GonEPFcurve EPF.in_carrier_curve EPF.one_curve EPF.multEqPowInCurve by auto qed lemma invG: "inv\<^bsub>EPF.curve\<^esub> G = G [^]\<^bsub>EPF.curve\<^esub> (n-1)" proof - have "prime n" using ECparamsValid ECdomainParametersValid_def by blast then have "1 < n" using prime_gt_1_nat by simp then have "1 + (n-1) = n" by auto then have "G \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub> (n-1)) = \\<^bsub>EPF.curve\<^esub>" by (metis nGeq1onEPFcurve EPF.curve.nat_pow_Suc2 GonEPFcurve plus_1_eq_Suc) then show ?thesis using EPF.curve.comm_inv_char EPF.curve.nat_pow_closed GonEPFcurve by presburger qed lemma ordGisn: assumes "x < n" "0 < x" shows "point_mult' x G \ Infinity" using ECdomainParametersValid_def ECparamsValid assms(1,2) EPF.order_n by presburger lemma ordGisn': assumes "x < n" "0 < x" shows "G [^]\<^bsub>EPF.curve\<^esub> x \ \\<^bsub>EPF.curve\<^esub>" using assms ordGisn EPF.one_curve EPF.multEqPowInCurve GonEPFcurve EPF.in_carrier_curve by algebra lemma curve_ord_G: "EPF.curve.ord G = n" using ECdomainParametersValid_def ECparamsValid EPF.curve_ord_n5 by presburger lemma multGmodn: "point_mult' x G = point_mult' (x mod n) G" by (metis ECdomainParametersValid_def ECparamsValid EPF.multQmodn) lemma multGmodn': "G[^]\<^bsub>EPF.curve\<^esub>x = G[^]\<^bsub>EPF.curve\<^esub>(x mod n)" using multGmodn EPF.in_carrier_curve GonEPFcurve EPF.multEqPowInCurve by fastforce lemma multdGmodn: assumes "P = point_mult' d G" shows "point_mult' x P = point_mult' (x mod n) P" by (metis ECdomainParametersValid_def ECparamsValid EPF.in_carrier_curve EPF.multQmodn EPF.order_n_cycle assms dGonEPFcurve) lemma multdGmodn': assumes "P = point_mult' d G" shows "P[^]\<^bsub>EPF.curve\<^esub>x = P[^]\<^bsub>EPF.curve\<^esub>(x mod n)" using assms multdGmodn EPF.in_carrier_curve dGonEPFcurve EPF.multEqPowInCurve by fastforce lemma multGmodn'int: "G[^]\<^bsub>EPF.curve\<^esub>(x::int) = G[^]\<^bsub>EPF.curve\<^esub>(x mod n)" by (metis EPF.in_carrier_curve EPF.one_curve GonEPFcurve EPF.multEqPowInCurve EPF.multQmodn'int nGeq1onEPFcurve prime_gt_1_nat ECparamsValid ECdomainParametersValid_def) lemma multdGmodn'int: assumes "P = point_mult' d G" shows "P[^]\<^bsub>EPF.curve\<^esub>(x::int) = P[^]\<^bsub>EPF.curve\<^esub>(x mod n)" by (metis assms ECparamsValid EPF.in_carrier_curve EPF.multQmodn'int EPF.order_n_cycle dGonEPFcurve prime_gt_1_nat ECdomainParametersValid_def) text \Computations done on the components of elliptic curve points are done modulo p. The order of the curve, however, is n, a prime different from p. So, for example, the scalar multiplication (n+a)G = aG, because the generator G has order n. So when computations are done on scalars intended for a scalar multiplication of a point on the curve, those should be reduced modulo n. In particular, we will need to compute inverses of scalars modulo n.\ definition "Rn = residue_ring (int n)" sublocale N: residues_prime n Rn using ECdomainParametersValid_def ECparamsValid residues_prime_def apply presburger using Rn_def by presburger abbreviation inv_mod_n :: "nat \ nat" where "inv_mod_n x \ inv_mod n x" lemma inv_mod_n_inv: assumes "x \ carrier Rn" "x \ 0" shows "inv\<^bsub>Rn\<^esub> x = inv_mod_n x" using N.inv_mod_p_inv_in_ring assms(1,2) by presburger lemma inv_mod_n_inv': assumes "0 < x" "x < n" shows "inv\<^bsub>Rn\<^esub> x = inv_mod_n x" using assms inv_mod_n_inv by (simp add: N.res_carrier_eq) subsection \3.2 Elliptic Curve Key Pairs\ text \3.2.1 Elliptic Curve Key Pair Generation Primitive This section doesn't exactly define a validation primitive. But from the Key Pair Generation Primitive, we can write down what it means to be a valid EC key pair. d is the private key and Q is the public key. We can also write down separately what it means for d to be a valid private key. Also note that the key generation method in 3.2.1 says to randomly select d. As we said above, we do not model random selection here, so we include d as an input.\ definition ECkeyGen :: "nat \ int point" where "ECkeyGen d = point_mult' d G" definition ECkeyPairValid :: "nat \ int point \ bool" where "ECkeyPairValid d Q \ (0 < d) \ (d < n) \ (point_mult' d G = Q)" definition ECprivateKeyValid :: "nat \ bool" where "ECprivateKeyValid d \ (0 < d) \ (d < n)" lemma ECkeyPairImpliesKeyGen: assumes "ECkeyPairValid d Q" shows "Q = ECkeyGen d" using ECkeyGen_def ECkeyPairValid_def assms by presburger lemma ECkeyPairEqKeyGen: "(ECkeyPairValid d Q) = ((Q = ECkeyGen d) \ ECprivateKeyValid d)" using ECkeyGen_def ECkeyPairValid_def ECprivateKeyValid_def by auto lemma ECkeyPairImpliesPrivateKeyValid: assumes "ECkeyPairValid d Q" shows "ECprivateKeyValid d" using assms ECkeyPairValid_def ECprivateKeyValid_def by auto lemma ECkeyPairNotInf: assumes "ECkeyPairValid d Q" shows "Q \ Infinity" by (metis assms ECkeyPairValid_def ordGisn) lemma ECkeyPairOnCurve: assumes "ECkeyPairValid d Q" shows "on_curve' Q" by (metis ECdomainParametersValid_def ECkeyPairValid_def ECparamsValid assms point_mult_closed) lemma ECkeyPairOrd_n: assumes "ECkeyPairValid d Q" shows "point_mult' n Q = Infinity" by (metis ECkeyPairValid_def assms mod_self multdGmodn point_mult.simps(1)) lemma ECkeyPair_dInRn': assumes "ECprivateKeyValid d" shows "d \ carrier Rn" using ECprivateKeyValid_def N.res_carrier_eq assms by force lemma ECkeyPair_dInRn: assumes "ECkeyPairValid d Q" shows "d \ carrier Rn" using ECkeyPairValid_def N.res_carrier_eq assms by force lemma ECkeyPair_invRn': assumes "ECprivateKeyValid d" shows "inv\<^bsub>Rn\<^esub>d \ carrier Rn \ d\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>d = \\<^bsub>Rn\<^esub>" using ECprivateKeyValid_def N.inv_closed N.res_carrier_eq N.zero_cong assms by force lemma ECkeyPair_invRn: assumes "ECkeyPairValid d Q" shows "inv\<^bsub>Rn\<^esub>d \ carrier Rn \ d\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>d = \\<^bsub>Rn\<^esub>" using ECkeyPairValid_def N.inv_closed N.res_carrier_eq N.zero_cong assms by force lemma ECkeyPair_inv_d': assumes "ECprivateKeyValid d" shows "inv_mod_n d = inv\<^bsub>Rn\<^esub>d" using ECprivateKeyValid_def N.res_carrier_eq assms inv_mod_n_inv by force lemma ECkeyPair_inv_d: assumes "ECkeyPairValid d Q" shows "inv_mod_n d = inv\<^bsub>Rn\<^esub>d" using ECkeyPairValid_def N.res_carrier_eq assms inv_mod_n_inv by force lemma ECkeyPair_curveMult: assumes "ECkeyPairValid d Q" shows "G[^]\<^bsub>EPF.curve\<^esub>d = Q" using ECkeyPairValid_def EPF.in_carrier_curve GonEPFcurve assms EPF.multEqPowInCurve by auto lemma ECkeyGen_mod_n: "ECkeyGen d = ECkeyGen (d mod n)" using ECkeyGen_def multGmodn by presburger lemma ECkeyGen_valid_mod_n: "(d mod n \ 0) = ECprivateKeyValid (d mod n)" by (simp add: ECprivateKeyValid_def N.p_prime prime_gt_0_nat) lemma ECkeyGen_int: fixes d :: int shows "G[^]\<^bsub>EPF.curve\<^esub>d = ECkeyGen (nat (d mod int n))" using EPF.ell_prime_field_axioms GonEPFcurve N.p_prime SEC1.ECkeyGen_def SEC1_axioms ell_prime_field.in_carrier_curve ell_prime_field.multEqPowInCurve multGmodn'int prime_gt_0_nat by force text \3.2.2.1 Elliptic Curve Public Key Validation Primitive\ definition ECpublicKeyValid :: "int point \ bool" where "ECpublicKeyValid Q \ (Q \ Infinity) \ (on_curve' Q) \ (point_mult' n Q = Infinity)" text \3.2.3.1 Elliptic Curve Public Key Partial Validation Primitive\ definition ECpublicKeyPartialValid :: "int point \ bool" where "ECpublicKeyPartialValid Q \ (Q \ Infinity) \ (on_curve' Q)" lemma validImpliesPartialValid: "ECpublicKeyValid Q \ ECpublicKeyPartialValid Q" using ECpublicKeyPartialValid_def ECpublicKeyValid_def by blast text \In general, if we have a partially valid public key, it is not necessarily (fully) valid. However, when h = 1, then this is true.\ lemma partValidImpliesValidIFheq1: assumes "h = 1" "ECpublicKeyPartialValid Q" shows "ECpublicKeyValid Q" proof - have 1: "Q \ Infinity" using assms(2) ECpublicKeyPartialValid_def by fast have 2: "on_curve' Q" using assms(2) ECpublicKeyPartialValid_def by fast have 3: "card {P. on_curve a b P} = n" using ECparamsValid ECdomainParametersValid_def assms(1) by simp have 4: "point_mult' n Q = Infinity" using 2 3 CurveOrderPrime EPF.AB_in_carrier EPF.nonsingular_in_bf N.p_prime by blast show ?thesis using 1 2 4 ECpublicKeyValid_def by presburger qed lemma partValidImpliesValidIFheq1': assumes "h = 1" "Q \ Infinity" "on_curve' Q" shows "point_mult' n Q = Infinity" using partValidImpliesValidIFheq1 assms ECpublicKeyPartialValid_def ECpublicKeyValid_def by blast lemma keyPairValidImpliespublicKeyValid: assumes "ECkeyPairValid d Q" shows "ECpublicKeyValid Q" using ECkeyPairNotInf ECkeyPairOnCurve ECkeyPairOrd_n ECpublicKeyValid_def assms by blast lemma ECpublicKeyValid_curve: assumes "ECpublicKeyValid Q" shows "Q \ \\<^bsub>EPF.curve\<^esub> \ Q \ carrier EPF.curve \ Q[^]\<^bsub>EPF.curve\<^esub> n = \\<^bsub>EPF.curve\<^esub>" using assms ECpublicKeyValid_def EPF.in_carrier_curve EPF.one_curve EPF.multEqPowInCurve by auto lemma ECpublicKeyPartValid_curve: assumes "ECpublicKeyPartialValid Q" shows "Q \ \\<^bsub>EPF.curve\<^esub> \ Q \ carrier EPF.curve" using assms ECpublicKeyPartialValid_def EPF.in_carrier_curve EPF.one_curve by auto lemma ECkeyGenValid: assumes "ECprivateKeyValid d" shows "ECpublicKeyValid (ECkeyGen d)" using ECkeyGen_def ECkeyPairValid_def ECprivateKeyValid_def assms keyPairValidImpliespublicKeyValid by auto lemma ECKeyGenValidPair: assumes "ECprivateKeyValid d" shows "ECkeyPairValid d (ECkeyGen d)" using assms ECprivateKeyValid_def ECkeyGen_def ECkeyPairValid_def by simp subsection \3.3 Elliptic Curve Diffie-Hellman Primitives\ text \3.3.1 Elliptic Curve Diffie-Hellman Primitive\ definition ECDHprim :: "nat \ int point \ int option" where "ECDHprim d\<^sub>U Q\<^sub>V = ( let P = point_mult' d\<^sub>U Q\<^sub>V in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ) )" lemma ECDHinCarrier: assumes "ECDHprim d Q = Some z" "on_curve' Q" shows "z \ carrier R" proof - let ?P = "point_mult' d Q" have 1: "?P \ Infinity" using assms(1) ECDHprim_def by fastforce obtain x and y where 2: "?P = Point x y" by (meson 1 point.exhaust) have 3: "z = x" using 1 2 assms(1) ECDHprim_def by auto have 4: "on_curve' ?P" using assms(2) point_mult_closed by auto show ?thesis using 2 3 4 on_curve_def by auto qed lemma ECDH_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECDHprim d Q = Some z" proof - have 1: "0 < d \ d < n" using assms(1) ECprivateKeyValid_def by simp have 2: "on_curve' Q \ Q \ Infinity \ point_mult' n Q = Infinity" using assms(2) ECpublicKeyValid_def by blast let ?P = "point_mult' d Q" have 3: "?P \ Infinity" using 1 2 EPF.order_n by blast obtain x and y where 4: "?P = Point x y" using 3 point.exhaust by blast have 5: "ECDHprim d Q = Some x" using 4 ECDHprim_def by force show ?thesis using 5 by fast qed lemma ECDH_curveMult: assumes "on_curve' Q" shows "ECDHprim d Q = ( let P = Q [^]\<^bsub>EPF.curve\<^esub> d in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ))" using assms ECDHprim_def EPF.multEqPowInCurve by presburger text \The important thing about Diffie-Hellman is that two entities can compute the same value using only their own (private) key and the other's public information. This lemma shows that.\ lemma ECDH_2ValidKeyPairs: assumes "ECkeyPairValid d1 P1" "ECkeyPairValid d2 P2" shows "ECDHprim d1 P2 = ECDHprim d2 P1" by (metis (no_types, lifting) assms(1,2) ECDHprim_def ECkeyPairValid_def EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute point_mult_mult) text \3.3.2 Elliptic Curve Cofactor Diffie-Hellman Primitive\ definition ECcofDHprim :: "nat \ int point \ int option" where "ECcofDHprim d\<^sub>U Q\<^sub>V = ( let P = point_mult' (h*d\<^sub>U) Q\<^sub>V in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ) )" lemma ECcofDHinCarrier: assumes "ECcofDHprim d Q = Some z" "on_curve' Q" shows "z \ carrier R" proof - let ?P = "point_mult' (h*d) Q" have 1: "?P \ Infinity" using assms(1) ECcofDHprim_def by fastforce obtain x and y where 2: "?P = Point x y" by (meson 1 point.exhaust) have 3: "z = x" using 1 2 assms(1) ECcofDHprim_def by auto have 4: "on_curve' ?P" using assms(2) point_mult_closed by auto show ?thesis using 2 3 4 on_curve_def by auto qed lemma ECcofDH_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECcofDHprim d Q = Some z" proof - have 1: "0 < d \ d < n" using assms(1) ECprivateKeyValid_def by simp have 2: "on_curve' Q \ Q \ Infinity \ point_mult' n Q = Infinity" using assms(2) ECpublicKeyValid_def by blast have 3: "h < n" using h_less_n ECparamsValid by auto have 4: "0 < h" using ECdomainParametersValid_def ECparamsValid by presburger have 5: "\ n dvd d" using 1 by fastforce have 6: "\ n dvd h" using 3 4 by fastforce have 7: "\ n dvd (h*d)" by (simp add: 5 6 N.p_coprime_left coprime_dvd_mult_left_iff) let ?m = "(h*d) mod n" have 8: "0 < ?m" using 7 mod_greater_zero_iff_not_dvd by blast have 9: "?m < n" using 3 by fastforce let ?P = "point_mult' (h*d) Q" have 10: "?P = point_mult' ?m Q" using EPF.multQmodn assms(2) ECpublicKeyValid_def by blast have 11: "?P \ Infinity" using 8 9 2 10 EPF.order_n by auto obtain x and y where 12: "?P = Point x y" using 11 point.exhaust by blast have "ECcofDHprim d Q = Some x" using 12 ECcofDHprim_def by force then show ?thesis by fast qed lemma ECcoDH_curveMult: assumes "on_curve' Q" shows "ECcofDHprim d Q = ( let P = Q [^]\<^bsub>EPF.curve\<^esub> (h*d) in ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ Some x\<^sub>P ))" using assms ECcofDHprim_def EPF.multEqPowInCurve by presburger text \The important thing about Diffie-Hellman is that two entities can compute the same value using only their own (private) key and the other's public information. This lemma shows that.\ lemma ECcofDH_2ValidKeyPairs: assumes "ECkeyPairValid k1 P1" "ECkeyPairValid k2 P2" shows "ECcofDHprim k1 P2 = ECcofDHprim k2 P1" proof - have 1: "P1 = point_mult' k1 G" using assms(1) by (simp add: ECkeyPairValid_def) have 2: "P2 = point_mult' k2 G" using assms(2) by (simp add: ECkeyPairValid_def) have 3: "point_mult' (h*k2) P1 = point_mult' (k1*(h*k2)) G" using 1 EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve point_mult_mult by presburger have 4: "point_mult' (h*k1) P2 = point_mult' (k2*(h*k1)) G" using 2 EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve point_mult_mult by presburger show ?thesis by (metis 3 4 mult.commute mult.left_commute ECcofDHprim_def) qed text \3.3.* Choose either "normal" or cofactor Diffie-Hellman. In some of the schemes below, the entities U and V must agree to use either "normal" or cofactor Diffie-Hellman. This primitive is convenient to use in those definitions.\ definition ECDHprimChoose :: "bool \ nat \ int point \ int option" where "ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V = (if useCoDH then (ECcofDHprim d\<^sub>U Q\<^sub>V) else (ECDHprim d\<^sub>U Q\<^sub>V) )" lemma ECDHchooseinCarrier: assumes "ECDHprimChoose useCo d Q = Some z" "on_curve' Q" shows "z \ carrier R" using ECDHprimChoose_def ECDHinCarrier ECcofDHinCarrier assms by presburger lemma ECDHchoose_validKeys: assumes "ECprivateKeyValid d" "ECpublicKeyValid Q" shows "\z. ECDHprimChoose useCo d Q = Some z" using ECDHprimChoose_def ECDH_validKeys ECcofDH_validKeys assms by presburger text \The important thing is that two entities can compute the same value using only their own (private) key together with the other's public information.\ lemma ECDHch_2ValidKeyPairs: assumes "ECkeyPairValid k1 P1" "ECkeyPairValid k2 P2" shows "ECDHprimChoose useCoDH k1 P2 = ECDHprimChoose useCoDH k2 P1" by (simp add: ECDHprimChoose_def ECDH_2ValidKeyPairs ECcofDH_2ValidKeyPairs assms(1,2)) subsection \3.4 Elliptic Curve MQV Primitive Compared to the Diffie-Hellman primitive, the MQV primitive is involved. So we break the definition into smaller chunks. MQVcomputeQbar corresponds to steps 2 and 4 in the standard. MQVcompute_s is step 3. MQVcomputeP is step 5.\ definition MQVcomputeQbar :: "int point \ nat option" where "MQVcomputeQbar Q = (case Q of Infinity \ None | Point x y \ ( let z = (2::nat)^(nat (\(log 2 n)/2\)); xBar = nat (x mod z) in Some (xBar + z) ))" definition MQVcompute_s :: "nat \ nat option \ nat \ nat option" where "MQVcompute_s d2U Q2Ubar d1U = ( case Q2Ubar of None \ None | Some Q2Ubar' \ Some ((d2U + Q2Ubar'*d1U) mod n))" definition MQVcomputeP :: "nat option \ int point \ nat option \ int point \ int point option" where "MQVcomputeP s Q2V Q2Vbar Q1V = ( case s of None \ None | Some s' \ (case Q2Vbar of None \ None | Some Q2Vbar' \ Some ( point_mult' (h*s') (add' Q2V (point_mult' Q2Vbar' Q1V)) ) ) )" definition ECMQVprim_validInput :: "nat \ int point \ nat \ int point \ int point \ int point \ bool" where "ECMQVprim_validInput d1U Q1U d2U Q2U Q1V Q2V \ (ECkeyPairValid d1U Q1U) \ (ECkeyPairValid d2U Q2U) \ (ECpublicKeyPartialValid Q1V) \ (ECpublicKeyPartialValid Q2V)" text \The standard has (d1U, Q1U) as one of the inputs, but the curve point Q1U does not appear in the computation. We write this function to be consistent with the text of the standard.\ definition ECMQVprim :: "nat \ int point \ nat \ int point \ int point \ int point \ int option" where "ECMQVprim d1U Q1U d2U Q2U Q1V Q2V = ( let Q2Ubar = MQVcomputeQbar Q2U; s = MQVcompute_s d2U Q2Ubar d1U; Q2Vbar = MQVcomputeQbar Q2V; P = MQVcomputeP s Q2V Q2Vbar Q1V in (case P of None \ None | Some Infinity \ None | Some (Point x y) \ Some x ) )" lemma MQVcomputeQbar_notInf: assumes "Q \ Infinity" shows "\Qbar. MQVcomputeQbar Q = Some Qbar" by (metis (no_types, lifting) MQVcomputeQbar_def assms point.case(2) point.exhaust) lemma MQVcomputeQbar_validPair: assumes "ECkeyPairValid d2U Q2U" shows "\Q2Ubar. MQVcomputeQbar Q2U = Some Q2Ubar" using ECkeyPairNotInf MQVcomputeQbar_notInf assms by blast lemma MQVcomputeQbar_validPub: assumes "ECpublicKeyPartialValid Q2V" shows "\Q2Vbar. MQVcomputeQbar Q2V = Some Q2Vbar" using ECpublicKeyPartialValid_def MQVcomputeQbar_notInf assms by presburger lemma MQVcompute_s_validIn: assumes "ECkeyPairValid d2U Q2U" "Q2Ubar = MQVcomputeQbar Q2U" shows "\s. MQVcompute_s d2U Q2Ubar d1U = Some s" using MQVcomputeQbar_validPair MQVcompute_s_def assms(1,2) by fastforce lemma MQVcomputeP_validIn: assumes "ECkeyPairValid d2U Q2U" "ECpublicKeyPartialValid Q2V" "Q2Ubar = MQVcomputeQbar Q2U" "s = MQVcompute_s d2U Q2Ubar d1U" "Q2Vbar = MQVcomputeQbar Q2V" shows "\P. MQVcomputeP s Q2V Q2Vbar Q1V = Some P" using MQVcomputeP_def MQVcomputeQbar_validPair MQVcomputeQbar_validPub MQVcompute_s_def assms by fastforce text \The important thing is that both U and V can compute the same value using only their own key pairs and the other's public keys. This lemma shows that is true.\ lemma MQV_reverseUV: assumes "ECkeyPairValid d1U Q1U" "ECkeyPairValid d2U Q2U" "ECkeyPairValid d1V Q1V" "ECkeyPairValid d2V Q2V" shows "ECMQVprim d1U Q1U d2U Q2U Q1V Q2V = ECMQVprim d1V Q1V d2V Q2V Q1U Q2U" proof - have U11: "Q1U = point_mult' d1U G" using assms(1) by (simp add: ECkeyPairValid_def) have U21: "Q2U = point_mult' d2U G" using assms(2) by (simp add: ECkeyPairValid_def) have V11: "Q1V = point_mult' d1V G" using assms(3) by (simp add: ECkeyPairValid_def) have V21: "Q2V = point_mult' d2V G" using assms(4) by (simp add: ECkeyPairValid_def) obtain Q2Ubar where Q2Ubar: "MQVcomputeQbar Q2U = Some Q2Ubar" using MQVcomputeQbar_validPair assms(2) by presburger obtain Q2Vbar where Q2Vbar: "MQVcomputeQbar Q2V = Some Q2Vbar" using MQVcomputeQbar_validPair assms(4) by presburger obtain sU where sU: "MQVcompute_s d2U (Some Q2Ubar) d1U = Some sU" by (simp add: MQVcompute_s_def) have sU1: "sU = (d2U + Q2Ubar*d1U) mod n" using sU MQVcompute_s_def by simp obtain sV where sV: "MQVcompute_s d2V (Some Q2Vbar) d1V = Some sV" by (simp add: MQVcompute_s_def) have sV1: "sV = (d2V + Q2Vbar*d1V) mod n" using sV MQVcompute_s_def by simp obtain PU where PU: "MQVcomputeP (Some sU) Q2V (Some Q2Vbar) Q1V = Some PU" by (simp add: MQVcomputeP_def) have U1: "PU = point_mult' (h*sU) (add' Q2V (point_mult' Q2Vbar Q1V))" using MQVcomputeP_def PU by auto have U2: "PU = point_mult' (h*sU) (add' (point_mult' d2V G) (point_mult' Q2Vbar (point_mult' d1V G)))" using U1 V11 V21 by fast have U3: "PU = point_mult' (h*sU) (point_mult' (d2V + Q2Vbar*d1V) G)" by (metis U2 point_mult_add EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve V11 mult.commute point_mult_mult) have U4: "PU = point_mult' (h*sU*(d2V + Q2Vbar*d1V)) G" by (metis U3 point_mult_mult EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute) let ?xU = "(h*sU*(d2V + Q2Vbar*d1V)) mod n" have U5: "PU = point_mult' ?xU G" using U4 multGmodn by blast obtain PV where PV: "MQVcomputeP (Some sV) Q2U (Some Q2Ubar) Q1U = Some PV" by (simp add: MQVcomputeP_def) have V1: "PV = point_mult' (h*sV) (add' Q2U (point_mult' Q2Ubar Q1U))" using MQVcomputeP_def PV by auto have V2: "PV = point_mult' (h*sV) (add' (point_mult' d2U G) (point_mult' Q2Ubar (point_mult' d1U G)))" using V1 U11 U21 by fast have V3: "PV = point_mult' (h*sV) (point_mult' (d2U + Q2Ubar*d1U) G)" by (metis V2 point_mult_add EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve U11 mult.commute point_mult_mult) have V4: "PV = point_mult' (h*sV*(d2U + Q2Ubar*d1U)) G" by (metis V3 point_mult_mult EPF.AB_in_carrier(1,2) EPF.in_carrier_curve EPF.nonsingular_in_bf GonEPFcurve mult.commute) let ?xV = "(h*sV*(d2U + Q2Ubar*d1U)) mod n" have V5: "PV = point_mult' ?xV G" using V4 multGmodn by blast have x1: "?xU = ?xV" by (metis sU1 sV1 mult.commute mod_mult_right_eq mult.assoc) have x2: "PU = PV" using x1 U5 V5 by argo show ?thesis using x2 PU PV Q2Ubar Q2Vbar sU sV ECMQVprim_def by presburger qed end (* of SEC1 locale *) subsection \3.5 Hash Functions\ text \Section 3.5 of SEC 1 covers the HASH functions that are supported by this standard. The supported hash functions are SHA-1, 224, 256, 384, and 512, all defined in FIPS 180, currently in version 4. We have translated that standard to HOL in FIPS180_4.thy. In this theory, we don't define any hash function. We only need to know the type of a hash function. For that, we use the type synonym below. The key fact we will use for a hash function is that for any input, the length of the output has hLen octets, for some fixed hLen. For example, SHA-256 always outputs hLen = 32 octets (256 bits). There are limits on the length of a message that should be given to a SHA function, however as we show in FIPS180_4.thy, you can still compute a hash output for any length of input message and that output will have the correct length. If Hash is hash_type, we can insist \\x. octets_valid (Hash x)\ and \\x. length (Hash x) = hLen\. \ type_synonym hash_type = "octets \ octets" subsection \3.6 Key Derivation Functions\ text \This sections lists the key derivation functions (KDFs) that are supported by this standard. Those are ANSI-X9.63-KDF, IKEv2-KDF, TLS-KDF, NIST-800-56-Concatenation-KDF. The specification states "The NIST-800-56-Catenation-KDF should be used, except for backwards compatability with implementations already using one of the three other key derivation functions." Again, we don't define any key derivation functions here. We need only a type for key derivation functions. The example in 3.6.1 shows that the input is an octet string Z and a non-negative integer keydatalen. It has an optional input called SharedInfo which is also octets. If the optional SharedInfo is not used, that input will be the empty list []. It produces an octet string of length keydatalen. So we have the type of a key derivation function. We also introduce the type of a function which indicates if the input is valid for the key derivation function. So if KDF is kdf_type and KDF_ValidIn is kdf_validin_type, then we can insist that \\Z l I. KDF_ValidIn Z l I \ length (KDF Z l I) = l \ octets_valid (KDF Z l I)\\ type_synonym kdf_type = "octets \ nat \ octets \ octets" type_synonym kdf_validin_type = "octets \ nat \ octets \ bool" subsection \3.7 MAC Schemes\ text \This section lists the MAC Schemes that are supported by this standard. There are many, based on HMAC (using various SHAs) and CMAC (using AES of various sizes). A MAC scheme will be used by ECIES in Section 5.1 below. Again we need only to define a type for MAC functions. The example given in 3.7 says that the entities need to agree on a MAC scheme plus mackeylen and maclen. Then they establish a shared secret key K, an octet string of length mackeylen. Then the tagging function takes an input an octet string M and outputs an octet string D of length maclen. Since mackeylen = length K, we don't need to list it separately as an input. And maclen is the length of the output of MAC, so it's not an input as much as a feature of the MAC scheme. So our type synonym for a generic MAC function, using the variable names in the standard, follows MAC(K, M) = D. We also give a mac_validin_type. So if MAC is mac_type, and MAC_VI is mac_validin_type, We could insist that \\K M. MAC_VI K M \ length K = mackeylen\ and \\K M. MAC_VI K M \ length (MAC K M) = maclen\ and \\K M. MAC_VI K M \ octets_valid (MAC K M)\. \ type_synonym mac_type = "octets \ octets \ octets" type_synonym mac_validin_type = "octets \ octets \ bool" subsection \3.8 Symmetric Encryption Schemes\ text \This standard supports 3-key TDES, XOR, and AES as symmetric encryption schemes. And encryption scheme, labelled ENC later in this standard, takes a shared private key K and a message M and produces a ciphertext C. All of these are octet strings. The type of the decryption is the same: it takes K and C and produces M. If ENC is enc_type, ENC_VI is enc_validin_type, DEC is dec_type, and DEC_VI is dec_valid_in_type, we might insist \ENC_VI K M \ DEC_VI K (ENC K M) \ DEC K (ENC K M) = M\ and \DEC_VI K C \ ENC_VI K (DEC K C) \ ENC K (DEC K C) = C\ \ type_synonym enc_type = "octets \ octets \ octets" type_synonym enc_validin_type = "octets \ octets \ bool" type_synonym dec_type = "octets \ octets \ octets" type_synonym dec_validin_type = "octets \ octets \ bool" subsection \3.9 Key Wrap Schemes\ text \This subsection specifies that either the NIST AES key wrap algorithm or the CMS TDES key wrap algorithm • must be used as the key wrap scheme in the Wrapped Key Transport Scheme, and • should be used more generally when wrapping an existing symmetric key with another sym- metric key. A key wrap scheme takes as input a key-encryption key K, an octet string of length wrapkeylen, and an octet string C which is the key that needs wrapping. It produces an octet string W, which is the wrapped key. If WRAP is wrap_type, WRAP_VI is wrap_validin_type, UNWRAP is unwrap_type, and UNWRAP_VI is unwrap_validin_type, we might insist \WRAP_VI K C \ UNWRAP_VI K (WRAP K C) \ UNWRAP K (WRAP K C) = C\ and \UNWRAP_VI K W \ WRAP_VI K (UNWRAP K W) \ WRAP K (UNWRAP K W) = W\ \ type_synonym wrap_type = "octets \ octets \ octets" type_synonym wrap_validin_type = "octets \ octets \ bool" type_synonym unwrap_type = "octets \ octets \ octets" type_synonym unwrap_validin_type = "octets \ octets \ bool" subsection \3.10 Random Number Generation\ text \Cryptographic keys must be generated in a way that prevents an adversary from guessing the private key. Keys should be generated with the help of a random number generator. Random number generators must comply with ANS X9.82 [X9.82] or corresponding NIST publi- cation [800-90]. We don't model random number generation in this theory. When the standard says "pick a random number", we simply include that random value as an input to the function.\ subsection \3.11 Security Levels and Protection Lifetimes\ text \"Data protected with cryptography today may continue to need protection in the future. Advances in cryptanalysis can be predicted, at least approximately. Based on current approximations, this document requires that data that needs protection beyond the year 2010 must be protected with 112-bit security or higher. Data that needs protection beyond the year 2030 must be protected with 128-bit security or higher. Data that needs protection beyond the year 2040 should be protected with 192-bit security or higher. Data that needs protection beyond 2080 should be protected with 256-bit security or higher."\ section \4 Signature Schemes\ type_synonym sig_type = "nat \ nat" subsection \4.1 Elliptic Curve Digital Signature Algorithm\ text \In the setup (Section 4.1.1) for the Elliptic Curve Digital Signature Algorithm (ECDSA), the "entities" U and V need to agree on a few things. They have to pick a hash algorithm. Above we discuss that the only approved hash algorithms are found in FIPS 180, The Secure Hash Standard. This hash is used only in the computation of the message digest e. They must also agree on an elliptic curve. Also (Section 4.1.2) Entity U should generate a key pair \(d\<^sub>U, Q\<^sub>U)\. Entity V should obtain \Q\<^sub>U\.\ subsubsection \4.1.3 Signing Operation\ text \Note in the standard, they convert the binary string Ebar to an octet string E and then convert E to an integer e. We do not need to compute E. We can go straight from a bit string to a nat. You can see this in Words.bits_to_octets_to_nat, which shows that converting to a nat from the bits is the same as going to octets then to a nat. As an aside, this "e" is a "message digest". As noted above for p, because n is odd (n is an odd prime), and thus not a power of 2, \\log 2 n\ = bit_length n\. I feel that using \bit_length n\ is more understandable, and for Isabelle to compute these values for given test vectors I must use \bit_length n\.\ context SEC1 begin definition ECDSA_compute_e :: "hash_type \ octets \ nat" where "ECDSA_compute_e Hash M = ( let H = Hash M; Hbar = octets_to_bits_len H; x = bit_length n; Ebar = take x Hbar in bits_to_nat Ebar )" lemma ECDSA_e_bnd: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "e < 2^(bit_length n)" proof - let ?H = "Hash M" let ?Hbar = "octets_to_bits_len ?H" let ?x = "bit_length n" let ?Ebar = "take ?x ?Hbar" have 1: "e = bits_to_nat ?Ebar" using assms(1) ECDSA_compute_e_def by presburger have 2: "length ?Ebar \ ?x" by fastforce have 3: "bits_valid ?Hbar" using assms(2) octets_valid_to_bits_len_valid by fast have 4: "bits_valid ?Ebar" using 3 words_valid_take by fast have 5: "e < 2^(length ?Ebar)" by (metis 1 4 words_to_nat_len_bnd words_valid_def less_numeral_extra(1) power_one_right) show ?thesis by (meson 2 5 le_trans linorder_not_less one_le_numeral power_increasing) qed lemma ECDSA_e_bitlen: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "bit_length e \ bit_length n" using assms less_bit_len2 ECDSA_e_bnd by blast lemma ECDSA_e_bitlen': assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" shows "bit_length e \ nat \log 2 n\" by (metis N.p_prime ECparamsValid n_not_2 ECDSA_e_bitlen bitLenLog2Prime assms nat_int) text \It could be that the hash function output has the same number of bits as the order n. For example, the NIST curve P-256 has a 256-bit n. Then if the hash used is SHA-256, the hash output always has 256 bits. In this case, the computation of e is simplified.\ lemma ECDSA_e_hlen_eq: assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" "length (Hash M) = hlen" "bit_length n = 8*hlen" shows "e = octets_to_nat (Hash M)" by (metis ECDSA_compute_e_def assms octets_to_bits_len_len octets_to_bits_len_to_nat order.refl take_all) lemma ECDSA_e_hlen_eq': assumes "e = ECDSA_compute_e Hash M" "octets_valid (Hash M)" "length (Hash M) = hlen" "nat \log 2 n\ = 8*hlen" shows "e = octets_to_nat (Hash M)" by (metis ECparamsValid n_not_2 N.p_prime bitLenLog2Prime nat_int assms ECDSA_e_hlen_eq) text \Note that in the standard, the ephemeral key pair is denoted (k, R). However the letter R is used in this theory to refer to the ring of integers modulo the prime p. We need to use a different letter, so we use P instead of R. Also, the procedure for singing a message is to first compute the message digest e and then sign that digest. For ease of application, we clearly delineate these two steps.\ definition ECDSA_Sign_e :: "nat \ nat \ nat \ int point \ sig_type option" where "ECDSA_Sign_e d\<^sub>U e k P = ( case P of Infinity \ None | Point x\<^sub>P y\<^sub>P \ ( let r = nat (x\<^sub>P mod n); kinv = inv_mod_n k; s = (kinv*(e + r*d\<^sub>U)) mod n in if r = 0 then None else ( if s = 0 then None else Some (r,s) ) ) )" abbreviation ECDSA_Sign :: "hash_type \ nat \ octets \ nat \ int point \ sig_type option" where "ECDSA_Sign Hash d\<^sub>U M k P \ ECDSA_Sign_e d\<^sub>U (ECDSA_compute_e Hash M) k P" text \The standard says to pick a key pair (k,P) for signing. But as P is a function of k, we can abbreviate these definitions.\ abbreviation ECDSA_Sign_e' :: "nat \ nat \ nat \ sig_type option" where "ECDSA_Sign_e' d\<^sub>U e k \ ECDSA_Sign_e d\<^sub>U e k (ECkeyGen k)" abbreviation ECDSA_Sign' :: "hash_type \ nat \ octets \ nat \ sig_type option" where "ECDSA_Sign' Hash d\<^sub>U M k \ ECDSA_Sign Hash d\<^sub>U M k (ECkeyGen k)" text \Now a few facts about the signing operation. Some of these are helpful in the public key recovery operation which is defined in 4.1.6\ lemma ECDSA_Sign_e_Inf: assumes "P = Infinity" shows "ECDSA_Sign_e d\<^sub>U e k P = None" by (simp add: ECDSA_Sign_e_def assms) lemma ECDSA_Sign_Inf: assumes "P = Infinity" shows "ECDSA_Sign Hash d\<^sub>U M k P = None" using assms ECDSA_Sign_e_Inf ECDSA_Sign_e_def by simp lemma ECDSA_Sign_e_Some: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" shows "P \ Infinity" using ECDSA_Sign_e_Inf assms by force lemma ECDSA_Sign_Some: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" shows "P \ Infinity" using ECDSA_Sign_e_Some assms ECDSA_Sign_e_def by simp lemma ECDSA_Sign_e_SomeOut: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv_mod_n k" "s = (kinv*(e + r*d\<^sub>U)) mod n" shows "S = (r, s)" by (smt (z3) assms ECDSA_Sign_e_def option.distinct(1) option.inject point.simps(5)) lemma ECDSA_Sign_SomeOut: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv_mod_n k" "s = (kinv*(e + r*d\<^sub>U)) mod n" "e = ECDSA_compute_e Hash M" shows "S = (r, s)" using ECDSA_Sign_e_SomeOut assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_inRn: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "r = fst S" "s = snd S" shows "0 < r \ r < n \ r \ carrier Rn \ 0 < s \ s < n \ s \ carrier Rn" proof - obtain x y where P: "P = Point x y" by (metis assms(1) ECDSA_Sign_e_Some point.exhaust) let ?r = "nat (x mod n)" let ?kinv = "inv_mod_n k" let ?s = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have r1: "0 < ?r" using assms(1) ECDSA_Sign_e_def P gr_zeroI by fastforce have r2: "?r < n" using nat_less_iff prime_gt_1_nat r1 by fastforce have r3: "?r \ carrier Rn" using r2 by fastforce have s1: "0 < ?s" using assms(1) ECDSA_Sign_e_def P gr_zeroI r1 by fastforce have s2: "?s < n" using nat_less_iff prime_gt_1_nat by fastforce have s3: "?s \ carrier Rn" by (metis N.mod_in_carrier of_nat_mod) have S1: "S = (?r, ?s)" using ECDSA_Sign_e_SomeOut P assms(1) by blast show ?thesis using S1 r1 r2 r3 s1 s2 s3 assms(2,3) by force qed lemma ECDSA_Sign_inRn: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "r = fst S" "s = snd S" shows "0 < r \ r < n \ r \ carrier Rn \ 0 < s \ s < n \ s \ carrier Rn" using ECDSA_Sign_e_inRn assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_invRn: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "r = fst S" "s = snd S" shows "inv_mod_n r = inv\<^bsub>Rn\<^esub>r \ r\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>r = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>r \ carrier Rn \ inv_mod_n s = inv\<^bsub>Rn\<^esub>s \ s\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>s = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>s \ carrier Rn" using ECDSA_Sign_e_inRn ECkeyPair_invRn' ECkeyPair_inv_d' ECprivateKeyValid_def assms by presburger lemma ECDSA_Sign_invRn: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "r = fst S" "s = snd S" shows "inv_mod_n r = inv\<^bsub>Rn\<^esub>r \ r\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>r = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>r \ carrier Rn \ inv_mod_n s = inv\<^bsub>Rn\<^esub>s \ s\\<^bsub>Rn\<^esub>inv\<^bsub>Rn\<^esub>s = \\<^bsub>Rn\<^esub> \ inv\<^bsub>Rn\<^esub>s \ carrier Rn" using ECDSA_Sign_e_invRn assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_e_SomeOut_Curve: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyPairValid k P" shows "S = (r, s)" by (smt (verit) assms ECDSA_Sign_e_SomeOut ECDSA_Sign_e_invRn ECkeyPair_inv_d split_beta N.res_mult_eq fst_eqD mod_mult_right_eq of_nat_mod semiring_1_class.of_nat_mult snd_eqD) lemma ECDSA_Sign_e_SomeOut_Curve': assumes "ECDSA_Sign_e' d\<^sub>U e k = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyGen k = P" "ECprivateKeyValid k" shows "S = (r, s)" using assms ECkeyPairEqKeyGen ECDSA_Sign_e_SomeOut_Curve by blast lemma ECDSA_Sign_SomeOut_Curve: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyPairValid k P" "e = ECDSA_compute_e Hash M" shows "S = (r, s)" using ECDSA_Sign_e_SomeOut_Curve assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_SomeOut_Curve': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "s = kinv \\<^bsub>Rn\<^esub> c" "ECkeyGen k = P" "e = ECDSA_compute_e Hash M" "ECprivateKeyValid k" shows "S = (r, s)" using assms ECkeyPairEqKeyGen ECDSA_Sign_SomeOut_Curve by presburger lemma ECDSA_Sign_e_SomeOut_r: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "r = fst S" shows "r = nat (x mod n)" using assms ECDSA_Sign_e_SomeOut by fastforce lemma ECDSA_Sign_SomeOut_r: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "r = fst S" shows "r = nat (x mod n)" using ECDSA_Sign_e_SomeOut_r assms ECDSA_Sign_e_def by blast lemma ECDSA_Sign_e_SomeOut_s: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv_mod_n k" shows "s = (kinv*(e + r*d\<^sub>U)) mod n" using assms ECDSA_Sign_e_SomeOut by fastforce lemma ECDSA_Sign_SomeOut_s: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv_mod_n k" "e = ECDSA_compute_e Hash M" shows "s = (kinv*(e + r*d\<^sub>U)) mod n" using ECDSA_Sign_e_SomeOut_s assms ECDSA_Sign_e_def by blast lemma ECDSA_Sign_e_SomeOut_curve_s: fixes x :: int assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyPairValid k P" shows "s = kinv \\<^bsub>Rn\<^esub> c" by (smt (verit) assms ECDSA_Sign_e_SomeOut_s ECkeyPair_inv_d N.mult_cong N.res_mult_eq mod_mod_trivial of_nat_mod semiring_1_class.of_nat_mult) lemma ECDSA_Sign_e_SomeOut_curve_s': fixes x :: int assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyGen k = P" "ECprivateKeyValid k" shows "s = kinv \\<^bsub>Rn\<^esub> c" using assms ECkeyPairEqKeyGen ECDSA_Sign_e_SomeOut_curve_s by algebra lemma ECDSA_Sign_SomeOut_curve_s: fixes x :: int assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyPairValid k P" "e = ECDSA_compute_e Hash M" shows "s = kinv \\<^bsub>Rn\<^esub> c" using ECDSA_Sign_e_SomeOut_curve_s assms ECDSA_Sign_e_def by algebra lemma ECDSA_Sign_SomeOut_curve_s': fixes x :: int assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "P = Point x y" "s = snd S" "r = nat (x mod n)" "kinv = inv\<^bsub>Rn\<^esub> k" "c = (e + r*d\<^sub>U) mod n" "ECkeyGen k = P" "ECprivateKeyValid k" "e = ECDSA_compute_e Hash M" shows "s = kinv \\<^bsub>Rn\<^esub> c" using assms ECkeyPairEqKeyGen ECDSA_Sign_SomeOut_curve_s by algebra text \Now a few facts about signing with (k mod n) as the private ephemeral key versus k.\ lemma ECDSA_Sign_e_kmodn: "ECDSA_Sign_e d\<^sub>U e k P = ECDSA_Sign_e d\<^sub>U e (k mod n) P" using inv_mod_mod ECDSA_Sign_e_def by presburger lemma ECDSA_Sign_kmodn: "ECDSA_Sign Hash d\<^sub>U M k P = ECDSA_Sign Hash d\<^sub>U M (k mod n) P" using ECDSA_Sign_e_kmodn ECDSA_Sign_e_def by fast lemma ECDSA_Sign_e'_kmodn: "ECDSA_Sign_e' d\<^sub>U e k = ECDSA_Sign_e' d\<^sub>U e (k mod n)" using ECDSA_Sign_e_kmodn ECkeyGen_mod_n by presburger lemma ECDSA_Sign'_kmodn: "ECDSA_Sign' Hash d\<^sub>U M k = ECDSA_Sign' Hash d\<^sub>U M (k mod n)" using ECDSA_Sign_e'_kmodn ECDSA_Sign_e_def by fast subsubsection \4.1.4 Verifying Operation\ text \Again, the standard uses the letter R to represent a curve point. Here R is used to denote the ring of integers mod p. So use a different letter (P) to avoid confusion. As above, the way to verify that a signature S matches a message M is to first compute the message digest e and verify that S matches e. So we again delineate the two steps of calculating the digest and verifying S against e. Recall that \(d\<^sub>U,Q\<^sub>U)\ is the (long-term) key pair for the signer U.\ definition ECDSA_Verify_e :: "int point \ nat \ sig_type \ bool" where "ECDSA_Verify_e Q\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = (e*sinv) mod n; u2 = (r*sinv) mod n; P = add' (point_mult' u1 G) (point_mult' u2 Q\<^sub>U) in (case P of Infinity \ False | Point x y \ ( (0 < r) \ (r < n) \ (0 < s) \ (s < n) \ (nat (x mod n) = r) ) ) )" abbreviation ECDSA_Verify :: "hash_type \ int point \ octets \ sig_type \ bool" where "ECDSA_Verify Hash Q\<^sub>U M S \ ECDSA_Verify_e Q\<^sub>U (ECDSA_compute_e Hash M) S" text \If M is signed with a valid ephemeral key pair (k, P), then the resulting signature S will pass the verification routine for M.\ lemma ECDSA_Sign_e_Verify: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e Q\<^sub>U e S" proof - obtain x y where P: "P = Point x y" by (metis assms(1) ECDSA_Sign_e_Some point.exhaust) let ?r = "nat (x mod n)" let ?kinv = "inv_mod_n k" let ?s = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have S1: "S = (?r, ?s)" using ECDSA_Sign_e_SomeOut P assms(1) by blast have S2: "fst S = ?r \ snd S = ?s" using S1 by force have r1: "0 < ?r" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have r2: "?r < n" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have s1: "0 < ?s" by (metis assms(1) S2 ECDSA_Sign_e_inRn) have s2: "?s < n" by (metis assms(1) S2 ECDSA_Sign_e_inRn) let ?x = "(e + ?r*d\<^sub>U) mod n" have x0: "?s = ?kinv \\<^bsub>Rn\<^esub> ?x" by (metis ECDSA_Sign_e_SomeOut_curve_s ECkeyPair_inv_d P S2 assms(1,2)) have x1: "0 < ?x" by (metis s1 mod_mod_trivial mod_mult_right_eq mult_0_right not_gr0) have x2: "?x \ carrier Rn" by (metis N.mod_in_carrier of_nat_mod) have k1: "k \ carrier Rn" using assms(2) ECkeyPair_dInRn by fast have k2: "k \ 0" using ECkeyPairValid_def assms(2) by blast have k3: "?kinv = inv\<^bsub>Rn\<^esub>k" using ECkeyPair_inv_d assms(2) by auto have k4: "inv\<^bsub>Rn\<^esub>?kinv = k" using N.nonzero_inverse_inverse_eq N.zero_cong k1 k2 k3 by force let ?sinv = "inv_mod_n ?s" have s3: "?sinv = inv\<^bsub>Rn\<^esub>?s" by (metis s2 s1 inv_mod_n_inv') have s4: "?sinv = inv\<^bsub>Rn\<^esub>?kinv \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" using N.m_comm N.nonzero_imp_inverse_nonzero N.nonzero_inverse_mult_distrib N.zero_cong k1 k2 k3 s3 x0 x1 x2 by auto have s5: "?sinv = k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" using s4 k4 by presburger let ?u1 = " (e*?sinv) mod n" let ?u2 = "(?r*?sinv) mod n" let ?P = "add' (point_mult' ?u1 G) (point_mult' ?u2 Q\<^sub>U)" have kP2: "G [^]\<^bsub>EPF.curve\<^esub>k = P" using assms(2) ECkeyPair_curveMult by fast have P1: "?P = G [^]\<^bsub>EPF.curve\<^esub>?u1 \\<^bsub>EPF.curve\<^esub> Q\<^sub>U [^]\<^bsub>EPF.curve\<^esub>?u2" using ECkeyPairOnCurve EPF.add_curve EPF.in_carrier_curve GonEPFcurve assms(3) EPF.multEqPowInCurve by presburger have P2: "Q\<^sub>U [^]\<^bsub>EPF.curve\<^esub>?u2 = G [^]\<^bsub>EPF.curve\<^esub>(d\<^sub>U*?u2)" using ECkeyPair_curveMult EPF.curve.nat_pow_pow GonEPFcurve assms(3) by blast have P3: "?P = G [^]\<^bsub>EPF.curve\<^esub>(?u1 + d\<^sub>U*?u2)" using P1 P2 EPF.curve.nat_pow_mult GonEPFcurve by presburger have k5: "(?u1 + d\<^sub>U*?u2) mod n = ?x*?sinv mod n" by (smt (z3) add_mult_distrib mod_add_cong mod_add_left_eq mod_mult_right_eq mult.assoc mult.commute) have k6: "?x*?sinv mod n = ?x\\<^bsub>Rn\<^esub>?sinv" by (simp add: N.res_mult_eq of_nat_mod) have k7: "?x\\<^bsub>Rn\<^esub>?sinv = ?x \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x" by (metis N.inv_closed N.m_assoc N.zero_cong gr_implies_not0 k1 of_nat_eq_0_iff x1 x2 k6 s5) have k8: "?x \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?x = k" by (metis N.inv_closed N.m_lcomm N.r_inv N.r_one N.zero_cong k1 k7 less_not_refl2 of_nat_eq_0_iff s5 x1 x2) have k9: "(?u1 + d\<^sub>U*?u2) mod n = k" by (metis k5 k6 k7 k8 nat_int) have P4: "?P = P" by (metis P3 k9 kP2 multGmodn') show ?thesis using P4 S2 P r1 r2 s1 s2 ECDSA_Verify_e_def by auto qed lemma ECDSA_Sign_e_Verify': assumes "ECDSA_Sign_e' d\<^sub>U e k = Some S" "ECprivateKeyValid k" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e Q\<^sub>U e S" using assms ECDSA_Sign_e_Verify ECkeyPairEqKeyGen by blast lemma ECDSA_Sign_Verify: assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify Hash Q\<^sub>U M S" using ECDSA_Sign_e_def ECDSA_Sign_e_Verify ECDSA_Verify_e_def assms by fast lemma ECDSA_Sign_Verify': assumes "ECDSA_Sign' Hash d\<^sub>U M k = Some S" "ECprivateKeyValid k" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify Hash Q\<^sub>U M S" using assms ECDSA_Sign_Verify ECkeyPairEqKeyGen by blast subsubsection \4.1.5 Alternative Verifying Operation\ text \If the verifier knows U's private key \d\<^sub>U\, they can use this alternate form of the verification routine. In short, if the verifier knows \d\<^sub>U\, they can compute the (private) ephemeral signing k labeled k above. Note that this alternate verifying operation needs only one elliptic curve operation, as opposed to the two needed above.\ definition ECDSA_Verify_e_Alt :: "nat \ nat \ sig_type \ bool" where "ECDSA_Verify_e_Alt d\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = nat ((e*sinv) mod n); u2 = nat ((r*sinv) mod n); P = point_mult' (u1 + u2*d\<^sub>U) G in (case P of Infinity \ False | Point x y \ ( (0 < r) \ (r < n) \ (0 < s) \ (s < n) \ (nat (x mod n) = r) ) ) )" definition ECDSA_Verify_Alt :: "hash_type \ nat \ octets \ sig_type \ bool" where "ECDSA_Verify_Alt Hash d\<^sub>U M S = ECDSA_Verify_e_Alt d\<^sub>U (ECDSA_compute_e Hash M) S" lemma ECDSA_Verify_e_eq: assumes "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_e_Alt d\<^sub>U e S = ECDSA_Verify_e Q\<^sub>U e S" by (smt (z3) ECDSA_Verify_e_Alt_def ECDSA_Verify_e_def ECkeyPairOnCurve ECkeyPairValid_def EPF.add_curve EPF.curve.nat_pow_mult EPF.curve.nat_pow_pow EPF.in_carrier_curve GonEPFcurve assms(1) mult.commute EPF.multEqPowInCurve nat_int) lemma ECDSA_Verify_eq: assumes "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_Verify_Alt Hash d\<^sub>U M S = ECDSA_Verify Hash Q\<^sub>U M S" using ECDSA_Verify_Alt_def ECDSA_Verify_e_eq assms by presburger text \If you can recover the ephemeral signing key (because you know the long-term secret key) from a paired message (digest) and verified signature, then if you sign that message with that recovered signing key, you will get back that signature. This is the inverse of the lemma above which says that if you sign a message and get S, then S will be verified as a correct signature for the message.\ definition ECDSA_Verify_Alt_recover_k :: "nat \ nat \ sig_type \ nat" where "ECDSA_Verify_Alt_recover_k d\<^sub>U e S = ( let r = fst S; s = snd S; sinv = inv_mod_n s; u1 = nat ((e*sinv) mod n); u2 = nat ((r*sinv) mod n) in (u1 + u2*d\<^sub>U) mod n )" lemma recovered_k_less_n: "ECDSA_Verify_Alt_recover_k d\<^sub>U e S < n" by (metis ECDSA_Verify_Alt_recover_k_def N.p_prime mod_less_divisor prime_gt_0_nat) lemma ECDSA_Verify_e_Sign: assumes "ECDSA_Verify_e Q\<^sub>U e S" "P = point_mult' k G" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "k = ECDSA_Verify_Alt_recover_k d\<^sub>U e S" shows "ECDSA_Sign_e d\<^sub>U e k P = Some S" proof - have 1: "ECDSA_Verify_e_Alt d\<^sub>U e S" using assms(1,3) ECDSA_Verify_e_eq by fast let ?r = "fst S" let ?s = "snd S" let ?sinv = "inv_mod_n ?s" let ?u1 = "nat (( e*?sinv) mod n)" let ?u2 = "nat ((?r*?sinv) mod n)" let ?P = "point_mult' (?u1 + ?u2*d\<^sub>U) G" have P1: "?P = point_mult' ((?u1 + ?u2*d\<^sub>U) mod n) G" using multGmodn by presburger have P2: "?P = point_mult' k G" by (metis P1 assms(4) ECDSA_Verify_Alt_recover_k_def) have P3: "?P = P" using assms(2) ECkeyPairValid_def P2 by algebra have P4: "?P \ Infinity" using 1 ECDSA_Verify_e_Alt_def by fastforce obtain x and y where P5: "?P = Point x y" using P4 point.exhaust by blast have 2: "(0 < ?r) \ (?r < n) \ (0 < ?s) \ (?s < n) \ (nat (x mod n) = ?r)" by (smt (verit) 1 ECDSA_Verify_e_Alt_def P5 point.case(2)) have s1: "?sinv = inv\<^bsub>Rn\<^esub>?s" using 2 inv_mod_n_inv' by fastforce have s2: "inv\<^bsub>Rn\<^esub>?sinv = ?s" by (simp add: 2 s1 ECkeyPair_dInRn' ECprivateKeyValid_def N.nonzero_inverse_inverse_eq N.zero_cong) have k1: "0 < k" by (metis P2 P4 assms(2) gr0I point_mult.simps(1)) have k2: "k < n" using assms(4) recovered_k_less_n by simp let ?c = "(e + ?r*d\<^sub>U) mod n" have k3: "k = ?c*?sinv mod n" by (metis (no_types, lifting) ECDSA_Verify_Alt_recover_k_def assms(4) distrib_right mod_add_left_eq mod_add_right_eq mod_mult_left_eq mult.assoc mult.commute nat_int) have c1: "0 < ?c \ ?c < n" by (metis bot_nat_0.not_eq_extremum k1 k2 k3 less_nat_zero_code mod_0 mod_less_divisor mult_0) let ?kinv = "inv_mod_n k" have k4: "?kinv = inv\<^bsub>Rn\<^esub> k" using k1 k2 inv_mod_n_inv' by fastforce have k5: "k = ?c \\<^bsub>Rn\<^esub> inv\<^bsub>Rn\<^esub>?s" by (simp add: N.res_mult_eq k3 of_nat_mod s1) have k6: "inv\<^bsub>Rn\<^esub> k = (inv\<^bsub>Rn\<^esub>?c) \\<^bsub>Rn\<^esub> ?s" by (metis 2 k5 s1 s2 c1 ECkeyGen_valid_mod_n ECkeyPair_dInRn' ECkeyPair_invRn' N.m_comm N.nonzero_imp_inverse_nonzero N.nonzero_inverse_mult_distrib N.zero_cong bot_nat_0.not_eq_extremum mod_if of_nat_eq_0_iff) let ?s' = "(?kinv*(e + ?r*d\<^sub>U)) mod n" have s3: "?s' = (inv\<^bsub>Rn\<^esub> k) \\<^bsub>Rn\<^esub> ?c" by (simp add: N.res_mult_eq k4 mod_mult_right_eq of_nat_mod) have s4: "?s' = ((inv\<^bsub>Rn\<^esub>?c) \\<^bsub>Rn\<^esub> ?c) \\<^bsub>Rn\<^esub> ?s" by (metis (no_types, lifting) 2 k4 k6 s3 ECkeyPair_invRn' ECprivateKeyValid_def ECkeyPair_dInRn' N.cring_simprules(11) N.res_mult_eq c1 mult.commute) have s5: "?s = ?s'" by (metis 2 s4 ECkeyGen_valid_mod_n ECkeyPair_dInRn' ECkeyPair_invRn' EPF.nat_int_eq N.m_closed N.m_comm N.r_one c1 mod_if nat_neq_iff) have A1: "ECDSA_Sign_e d\<^sub>U e k P = ( let r = nat (x mod n); kinv = inv_mod_n k; s = (kinv*(e + r*d\<^sub>U)) mod n in if r = 0 then None else (if s = 0 then None else Some (r,s)))" using ECDSA_Sign_e_def P3 P4 P5 by force have A2: "ECDSA_Sign_e d\<^sub>U e k P = (if ?r = 0 then None else (if ?s = 0 then None else Some (?r,?s)))" by (smt (verit) 2 A1 s5) show ?thesis using A2 2 by auto qed lemma ECDSA_Verify_Sign: assumes "ECDSA_Verify Hash Q\<^sub>U M S" "P = point_mult' k G" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "k = ECDSA_Verify_Alt_recover_k d\<^sub>U (ECDSA_compute_e Hash M) S" shows "ECDSA_Sign Hash d\<^sub>U M k P = Some S" by (metis ECDSA_Verify_e_Sign assms) text \One last variation on this theme. Suppose that the cofactor h = 1. Then you know that every point on the curve is a power of the generator G. So if you know that \Q\<^sub>U\ is a partially valid public key, then we know that there is a corresponding private key \d\<^sub>U\. If you could find that private key (which, in practice, you can't), then you can use it to recover the ephemeral signing key.\ lemma ECDSA_Verify_Sign_h1: assumes "ECpublicKeyPartialValid Q\<^sub>U" "ECDSA_Verify_e Q\<^sub>U e S" "h = 1" shows "\d\<^sub>U. (Q\<^sub>U = point_mult' d\<^sub>U G \ ECDSA_Sign_e' d\<^sub>U e (ECDSA_Verify_Alt_recover_k d\<^sub>U e S) = Some S)" proof - have 1: "on_curve' Q\<^sub>U" using assms(1) ECpublicKeyPartialValid_def by blast obtain d\<^sub>U where 2: "d\<^sub>U Q\<^sub>U = point_mult' d\<^sub>U G" using assms(3) 1 EC_h1_cyclic' by blast have 3: "Q\<^sub>U \ Infinity" using assms(1) ECpublicKeyPartialValid_def by blast have 4: "0 < d\<^sub>U" by (metis 2 3 not_gr0 point_mult.simps(1)) have 5: "ECkeyPairValid d\<^sub>U Q\<^sub>U" using 2 4 ECkeyPairValid_def by blast have 6: "ECDSA_Sign_e' d\<^sub>U e (ECDSA_Verify_Alt_recover_k d\<^sub>U e S) = Some S" using 5 ECDSA_Verify_e_Sign ECkeyGen_def assms(2) by blast show ?thesis using 2 6 by fast qed subsubsection \4.1.6 Public Key Recovery Operation\ text \The idea is that if you have a signature, you can generate a small list of putative values for the public key. Then if you have an oracle, you can test each of these putative values. This is a convenient option in bandwidth-constrained environments. What might this oracle be? For example, if you have a second message (M2) and the signature for that message (S2), then you can check that (ECDSA_Verify Hash Q M2 S2) returns True when the putative public key Q is used. We don't have access to the oracle needed in step 1.6.2 of the public key recovery operation. So here we simply call it the ValidationOracle, which takes a point and returns True or False. Note that the public key recovery operation loops over j. In HOL/Isabelle, we translate the loop as a recursive function ECDSA_PublicKeyRecovery_rec. The "meat" of each iteration is contained in the function ECDSA_PublicKeyRecovery_1. ECDSA_PublicKeyRecovery first computes e (because it is the same in every loop iteration) and kicks off the loop by calling ECDSA_PublicKeyRecovery_rec. Lastly, we note again that the standard writes \(log 2 p)/8\. We know because p is odd (and thus not a power of two) that octet_length p = \(log 2 p)/8\. The proof of this may be found in Words.thy/octetLenLog2Prime.\ definition ECDSA_PublicKeyRecovery_1 :: "(int point \ bool) \ nat \ sig_type \ nat \ int point option" where "ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig j = ( let r = fst Sig; s = snd Sig; x = r + j*n; mlen = octet_length p; X = nat_to_octets_len x mlen; P = octets_to_point' (2 # X) in ( case P of None \ None | Some R' \ ( let nR = point_mult' n R'; rInv = inv_mod_n r; eG = opp (point_mult' e G); sR1 = point_mult' s R'; sR2 = point_mult' s (opp R'); Q1 = point_mult' rInv (add' sR1 eG); Q2 = point_mult' rInv (add' sR2 eG) in if (nR = Infinity) then ( if (ValidationOracle Q1) then (Some Q1) else if (ValidationOracle Q2) then (Some Q2) else None ) else None ) ) )" text \It is a little more convenient to have a loop counter that decreases until it reaches 0. So the loop index i here satisfies j = h + 1 - i, where j is the loop counter from step 1 in the standard. i starts at h+1, which corresponds to j = 0. When i = 1, j = h, so when i = 0, there are no more iterations to perform. If the signature is valid, the public key should be recovered so that i = 0 is never reached.\ fun ECDSA_PublicKeyRecovery_rec :: "(int point \ bool) \ nat \ sig_type \ nat \ int point option" where "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig 0 = None" | "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig i = ( let P = ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig (h+1-i) in ( if (P = None) then (ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig (i-1)) else P ) )" definition ECDSA_PublicKeyRecovery_e :: "(int point \ bool) \ nat \ sig_type \ int point option" where "ECDSA_PublicKeyRecovery_e ValidationOracle e S = ECDSA_PublicKeyRecovery_rec ValidationOracle e S (h+1)" definition ECDSA_PublicKeyRecovery :: "hash_type \ (int point \ bool) \ octets \ sig_type \ int point option" where "ECDSA_PublicKeyRecovery Hash ValidationOracle M S = ECDSA_PublicKeyRecovery_e ValidationOracle (ECDSA_compute_e Hash M) S" text \Knowing nothing about the hash function or the oracle, at the very least we know that if the key recovery operation returns a point, then that point is on the elliptic curve.\ lemma ECDSA_PublicKeyRecovery_1_onCurve: assumes "ECDSA_PublicKeyRecovery_1 ValidationOracle e Sig j = Some Q" shows "on_curve' Q" proof - let ?r = "fst Sig" let ?s = "snd Sig" let ?x = "?r + j*n" let ?mlen = "octet_length p" let ?X = "nat_to_octets_len ?x ?mlen" let ?P = "octets_to_point' (2 # ?X)" have 1: "?P \ None" by (metis (no_types, lifting) ECDSA_PublicKeyRecovery_1_def assms option.case(1) option.distinct(1)) obtain R' where 2: "?P = Some R'" using 1 by blast have 3: "on_curve' R'" using 2 octets2PointOnCurve by fast let ?rInv = "inv_mod_n ?r" let ?eG = "opp (point_mult' e G)" let ?sR1 = "point_mult' ?s R'" let ?sR2 = "point_mult' ?s (opp R')" let ?Q1 = "point_mult' ?rInv (add' ?sR1 ?eG)" let ?Q2 = "point_mult' ?rInv (add' ?sR2 ?eG)" have 4: "Q = ?Q1 \ Q = ?Q2" by (smt (verit, best) 2 assms ECDSA_PublicKeyRecovery_1_def not_None_eq not_Some_eq option.distinct(1) option.inject option.simps(5)) have 5: "on_curve' ?eG \ on_curve' ?sR1 \ on_curve' ?sR2" using 3 ECparamsValid ECdomainParametersValid_def opp_closed point_mult_closed by simp have 6: "on_curve' ?Q1 \ on_curve' ?Q2" using 5 point_mult_closed add_closed by auto show ?thesis using 4 6 by fast qed lemma ECDSA_PublicKeyRecovery_rec_onCurve: assumes "ECDSA_PublicKeyRecovery_rec ValidationOracle e Sig i = Some Q" shows "on_curve' Q" using assms apply (induct i) apply simp using ECDSA_PublicKeyRecovery_1_onCurve ECDSA_PublicKeyRecovery_rec.simps(2) diff_Suc_1 by (smt (verit, best)) text \The public key recovery operation is guaranteed to return the correct public key \Q\<^sub>U\, given a correct validation oracle. In the next several lemmas, we work toward proving this. In the ECDSA locale, U has signed a message with its key pair. So the correct oracle will only return "true" when the putative public key matches U's public key.\ definition UOracle :: "int point \ int point \ bool" where "UOracle Q\<^sub>U possibleQ = (Q\<^sub>U = possibleQ)" lemma OracleCorrect_1: assumes "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e Sig j = Some P" shows "P = Q\<^sub>U" proof - let ?r = "fst Sig" let ?s = "snd Sig" let ?x = "?r + j*n" let ?mlen = "octet_length p" let ?X = "nat_to_octets_len ?x ?mlen" let ?P = "octets_to_point' (2 # ?X)" have 1: "?P \ None" by (metis (no_types, lifting) ECDSA_PublicKeyRecovery_1_def assms option.case(1) option.distinct(1)) obtain R' where 2: "?P = Some R'" using 1 by blast show ?thesis by (smt (verit) 2 ECDSA_PublicKeyRecovery_1_def UOracle_def assms option.distinct(1) option.inject option.simps(5)) qed lemma OracleCorrect_rec: assumes "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e Sig i = Some P" shows "P = Q\<^sub>U" using assms apply (induct i) apply simp by (smt (verit, del_insts) ECDSA_PublicKeyRecovery_rec.simps(2) OracleCorrect_1 diff_Suc_1) lemma KeyRecoveryWorks_1: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" proof - have x0: "0 \ x \ x < p \ x \ carrier R" using ECkeyPairOnCurve assms(2,3) inCarrier onCurveEq2 by blast have j0: "0 \ x div n" using x0 div_int_pos_iff of_nat_0_le_iff by blast have j1: "j = nat (x div n)" using assms(4) nat_div_distrib nat_int x0 by presburger let ?r = "fst S" let ?s = "snd S" have r1: "?r = nat (x mod n)" by (metis ECDSA_Sign_e_SomeOut_r assms(1,3)) have r2: "?r \ carrier Rn" using ECDSA_Sign_e_inRn assms(1) by blast have r3: "0 < ?r" using ECDSA_Sign_e_inRn assms(1) by blast let ?x = "?r + j*n" let ?kinv = "inv\<^bsub>Rn\<^esub> k" have k0: "0 < k" using assms(2) ECkeyPairValid_def by simp have k1: "k \ carrier Rn" using assms(2) ECkeyPair_dInRn by blast have k2: "?kinv \ carrier Rn" by (metis k1 k0 N.inv_closed N.zero_cong bot_nat_0.not_eq_extremum dual_order.refl of_nat_le_0_iff) have k3: "k \\<^bsub>Rn\<^esub> ?kinv = \\<^bsub>Rn\<^esub>" using ECkeyPairValid_def N.zero_cong assms(2) k1 by force let ?c = "(e + ?r*d\<^sub>U) mod n" have c1: "?c \ carrier Rn" by (simp add: of_nat_mod) have s1: "?s = ?kinv \\<^bsub>Rn\<^esub> ?c" using ECDSA_Sign_e_SomeOut_curve_s assms(1,2,3,5) r1 by blast have x1: "?x = x" using j0 j1 r1 r3 by auto let ?l = "octet_length p" let ?X = "nat_to_octets_len ?x ?l" let ?P = "octets_to_point' (2 # ?X)" have P1: "on_curve' P" using assms(2) ECkeyPairOnCurve by blast have P2: "?P = Some P \ ?P = Some (opp P)" by (metis P1 assms(3) x1 point2Octets2PointComp_PoppP' nat_int) have P3: "point_mult' n P = Infinity" using assms(2) by (simp add: ECkeyPairOrd_n) have P4: "P \ carrier EPF.curve" by (simp add: EPF.in_carrier_curve P1) have P5: "opp P = inv\<^bsub>EPF.curve\<^esub> P" using EPF.inv_curve P4 by presburger have P6: "point_mult' n (opp P) = Infinity" by (metis EPF.curve.inv_one EPF.curve.nat_pow_inv EPF.one_curve P1 P3 P4 P5 EPF.multEqPowInCurve opp_closed) have P7: "point_mult' k G = P" using assms(2) ECkeyPairValid_def by blast have P8: "G[^]\<^bsub>EPF.curve\<^esub>k = P" using P7 EPF.in_carrier_curve GonEPFcurve EPF.multEqPowInCurve by force obtain R' where R1: "?P = Some R'" using P2 by blast have R2: "R' = P \ R' = opp P" using P2 R1 by auto have R3: "(R' = opp P) = (opp R' = P)" using P1 R2 opp_opp by auto let ?nR = "point_mult' n R'" have nR1: "?nR = point_mult' n P \ ?nR = point_mult' n (opp P)" using R2 by fast have nR2: "?nR = Infinity" using nR1 P3 P6 by argo let ?rInv = "inv_mod_n ?r" have rInv: "?rInv = inv\<^bsub>Rn\<^esub> ?r" using ECDSA_Sign_e_inRn assms(1) bot_nat_0.not_eq_extremum inv_mod_n_inv by presburger have rInv2: "?r \\<^bsub>Rn\<^esub> ?rInv = \\<^bsub>Rn\<^esub>" by (simp add: N.zero_cong r2 r3 rInv) let ?eG = "opp (point_mult' e G)" have eG1: "?eG = G[^]\<^bsub>EPF.curve\<^esub>(-(int e))" using EPF.curve.int_pow_neg_int EPF.in_carrier_curve EPF.inv_curve GonEPFcurve dGonEPFcurve EPF.multEqPowInCurve by auto have eG2: "?eG = G[^]\<^bsub>EPF.curve\<^esub>(-(int e) mod n)" using eG1 multGmodn'int by presburger let ?em = "(-(int e) mod n)" have em1: "?em = \\<^bsub>Rn\<^esub>e" using N.res_neg_eq by presburger let ?sR1 = "point_mult' ?s R'" let ?sR2 = "point_mult' ?s (opp R')" have sR1: "{?sR1, ?sR2} = {point_mult' ?s P, point_mult' ?s (opp P)}" using R2 R3 by fastforce have sR2: "{?sR1, ?sR2} = {P[^]\<^bsub>EPF.curve\<^esub>?s, P[^]\<^bsub>EPF.curve\<^esub>(-?s)}" by (metis sR1 P1 P4 P5 EPF.curve.int_pow_neg_int EPF.curve.nat_pow_inv opp_closed EPF.multEqPowInCurve) have sR3: "(R' = opp P) \ (?sR1 = P[^]\<^bsub>EPF.curve\<^esub>(-?s) \ ?sR2 = P[^]\<^bsub>EPF.curve\<^esub>?s)" by (metis P1 doubleton_eq_iff EPF.multEqPowInCurve sR2 opp_opp) let ?Q1 = "point_mult' ?rInv (add' ?sR1 ?eG)" let ?Q2 = "point_mult' ?rInv (add' ?sR2 ?eG)" let ?Q = "((P[^]\<^bsub>EPF.curve\<^esub>?s) \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub>?em))[^]\<^bsub>EPF.curve\<^esub>?rInv" have Q1: "(R' = P) \ ?Q1 = ?Q" using P4 eG2 EPF.add_curve GonEPFcurve EPF.curve.int_pow_closed EPF.curve.m_closed EPF.curve.nat_pow_closed EPF.in_carrier_curve EPF.multEqPowInCurve by auto have Q2: "(R' = opp P) \ ?Q2 = ?Q" using P4 sR3 eG2 GonEPFcurve EPF.add_curve EPF.curve.int_pow_closed EPF.curve.m_closed EPF.curve.nat_pow_closed EPF.in_carrier_curve EPF.multEqPowInCurve by auto have Q3: "P[^]\<^bsub>EPF.curve\<^esub>?s = G[^]\<^bsub>EPF.curve\<^esub>(k*?s)" using P8 EPF.curve.nat_pow_pow GonEPFcurve by blast have Q4: "(G[^]\<^bsub>EPF.curve\<^esub>(k*?s) \\<^bsub>EPF.curve\<^esub> (G [^]\<^bsub>EPF.curve\<^esub>?em)) = G[^]\<^bsub>EPF.curve\<^esub>(k*?s+?em)" by (metis EPF.curve.is_group GonEPFcurve group.int_pow_mult int_pow_int) have Q5: "?Q = G[^]\<^bsub>EPF.curve\<^esub>((k*?s+?em)*?rInv)" by (metis Q3 Q4 EPF.curve.int_pow_pow GonEPFcurve int_pow_int) have Q6: "?Q = G[^]\<^bsub>EPF.curve\<^esub>((k*?s+?em)*?rInv mod n)" using Q5 multGmodn'int by algebra have ex1: "(k*?s+?em)*?rInv mod n = (k\\<^bsub>Rn\<^esub>?s \\<^bsub>Rn\<^esub> ?em)\\<^bsub>Rn\<^esub>?rInv" by (simp add: N.res_add_eq N.res_mult_eq mod_add_left_eq mod_mult_right_eq mult.commute) have ex2: "(k*?s+?em)*?rInv mod n = (k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) \\<^bsub>Rn\<^esub>e)\\<^bsub>Rn\<^esub>(inv\<^bsub>Rn\<^esub> ?r)" using ex1 s1 em1 rInv by algebra have ex3: "k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) = k\\<^bsub>Rn\<^esub>?kinv \\<^bsub>Rn\<^esub> ?c" using m_assoc k1 k2 c1 by algebra have ex4: "k\\<^bsub>Rn\<^esub>(?kinv \\<^bsub>Rn\<^esub> ?c) = ?c" using ex3 k3 N.l_one c1 by presburger have ex5: "?c \\<^bsub>Rn\<^esub>e = e \\<^bsub>Rn\<^esub> (?r\\<^bsub>Rn\<^esub>d\<^sub>U) \\<^bsub>Rn\<^esub>e" by (simp add: N.res_add_eq N.res_mult_eq mod_add_right_eq of_nat_mod) have ex6: "?c \\<^bsub>Rn\<^esub>e = ?r\\<^bsub>Rn\<^esub>d\<^sub>U" using ex5 by (simp add: N.res_add_eq N.res_diff_eq N.res_mult_eq mod_diff_left_eq) have ex7: "(k*?s+?em)*?rInv mod n = ?r\\<^bsub>Rn\<^esub>d\<^sub>U\\<^bsub>Rn\<^esub>?rInv" using ex2 ex4 ex6 rInv by argo have ex8: "(k*?s+?em)*?rInv mod n = d\<^sub>U mod n" by (metis (mono_tags, opaque_lifting) ex7 rInv2 N.res_mult_eq N.res_one_eq mod_mult_right_eq mult.left_commute mult.right_neutral of_nat_mod semiring_1_class.of_nat_mult) have ex9: "(k*?s+?em)*?rInv mod n = d\<^sub>U" using ex8 assms(5) ECkeyPairValid_def by force have Q7: "?Q = G[^]\<^bsub>EPF.curve\<^esub>d\<^sub>U" by (metis Q6 ex9 int_pow_int) have Q8: "?Q = Q\<^sub>U" using Q7 ECkeyPair_curveMult assms(5) by presburger have A1: "(R' = P) \ ?Q1 = Q\<^sub>U" using Q8 Q1 by argo have A11: "(R' = P) \ UOracle Q\<^sub>U ?Q1" using A1 UOracle_def by blast have A12: "(R' = P) \ ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using A11 ECDSA_PublicKeyRecovery_1_def P2 nR2 A1 R1 by auto have A2: "(R' = opp P) \ ?Q2 = Q\<^sub>U" using Q8 Q2 by argo have A21: "(R' = opp P) \ UOracle Q\<^sub>U ?Q2" using A2 UOracle_def by blast have A22: "(R' = opp P) \ ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" by (smt (verit) nR2 A21 R1 UOracle_def ECDSA_PublicKeyRecovery_1_def option.simps(5)) show ?thesis using A12 A22 R2 by fast qed lemma KeyRecoveryWorks_1': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using KeyRecoveryWorks_1 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec_h1: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1-j) = Some Q\<^sub>U" proof - have 1: "x \ carrier R" using assms(2,3) ECkeyPairOnCurve onCurveEq2 by blast have 2: "j \ h" by (metis 1 assms(4) less_p_div_n' ECparamsValid inCarrierNatInt) have 3: "h+1-(h+1-j) = j" using 2 by fastforce have 4: "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S j = Some Q\<^sub>U" using KeyRecoveryWorks_1 assms(1,2,3,4,5) by blast show ?thesis by (smt (verit, ccfv_SIG) 2 3 4 ECDSA_PublicKeyRecovery_rec.simps(2) Suc_diff_le Suc_eq_plus1 option.distinct(1)) qed lemma KeyRecoveryWorks_rec_h1': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1-j) = Some Q\<^sub>U" using KeyRecoveryWorks_rec_h1 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec_h2: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "h+1-j \ L" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L = Some Q\<^sub>U" using assms proof (induct L) case 0 then show ?case by (metis KeyRecoveryWorks_rec_h1 bot_nat_0.extremum_uniqueI) next case (Suc L) then show ?case proof (cases "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L)) = None") case True then have "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (Suc L) = ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L" using ECDSA_PublicKeyRecovery_rec.simps(2) diff_Suc_1 by presburger then show ?thesis by (metis KeyRecoveryWorks_rec_h1 Suc.hyps Suc.prems(1,5,6) assms(2,3,4) le_eq_less_or_eq less_Suc_eq_le) next case F0: False have F1: "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (Suc L) = ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L))" using ECDSA_PublicKeyRecovery_rec.simps(2) F0 by presburger have F2: "ECDSA_PublicKeyRecovery_1 (UOracle Q\<^sub>U) e S (h+1-(Suc L)) = Some Q\<^sub>U" using F0 OracleCorrect_1 by blast then show ?thesis using F1 F2 by presburger qed qed lemma KeyRecoveryWorks_rec_h2': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "P = Point x y" "j = (nat x) div n" "h+1-j \ L" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S L = Some Q\<^sub>U" using KeyRecoveryWorks_rec_h2 assms ECDSA_Sign_e_def by algebra lemma KeyRecoveryWorks_rec: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1) = Some Q\<^sub>U" by (metis assms(1,2,3) ECDSA_Sign_e_Some KeyRecoveryWorks_rec_h2 diff_le_self point.exhaust) lemma KeyRecoveryWorks_rec': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" "e = ECDSA_compute_e Hash M" shows "ECDSA_PublicKeyRecovery_rec (UOracle Q\<^sub>U) e S (h+1) = Some Q\<^sub>U" using KeyRecoveryWorks_rec assms ECDSA_Sign_e_def by algebra text \Here finally is the lemma that we have been working toward. The entity U has signed the message M with the ephemeral key pair (k,P). Given a correct oracle for \Q\<^sub>U\, the key recovery operation will recover \Q\<^sub>U\ given the message M and the signature S. Again, what could the oracle be? Typically, the person trying to recover \Q\<^sub>U\ will have at least 2 messages signed by U. They can run the key recovery operation with one (message, signature) pair and use a second (message, key) pair to to check any putative \Q\<^sub>U\ that the key recovery operation presents.\ lemma KeyRecoveryWorks: assumes "ECDSA_Sign_e d\<^sub>U e k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery_e (UOracle Q\<^sub>U) e S = Some Q\<^sub>U" using ECDSA_PublicKeyRecovery_e_def KeyRecoveryWorks_rec assms(1,2,3) by presburger lemma KeyRecoveryWorks': assumes "ECDSA_Sign Hash d\<^sub>U M k P = Some S" "ECkeyPairValid k P" "ECkeyPairValid d\<^sub>U Q\<^sub>U" shows "ECDSA_PublicKeyRecovery Hash (UOracle Q\<^sub>U) M S = Some Q\<^sub>U" using ECDSA_PublicKeyRecovery_def ECDSA_Sign_e_def KeyRecoveryWorks assms by presburger subsubsection \4.1.7 Self-Signing Operation\ text \The statement of the self-signing operation in the standard is vague in two ways and overly complex in another. First, it takes as input "information", labeled I. The format of this "information" is not specified. It is only said that I "should include the identity of the signer", without indicating what "identity" means. Second, step 4 in the standard says "Form a message M containing both I and (r,s)." It does not indicate how such a message should be formed. Isabelle is a typed language so we need to specify the type of I. We specify the type of I as 'a, which is the generic term for "some type". And we assume the existence of a "form a message" function that takes "information" type and a signature type and produces octets. Note that the standard lists as input the curve parameters and the "information" I. But also needed as input are an ephemeral key, which we label (k, P), and a random integer s in the range [1,n-1]. As noted above, we cannot use (k, R) to represent the ephemeral key because R is already used in this theory to refer to the integer ring R of integers modulo p. Finally, let's examine step 5 in the standard. It says to "Use the Public Key Recovery Operation" to recover a public key Q. In that operation, we know only (x mod n), where x is the first coordinate of the ephemeral key point P. So we must loop over all possible x and try to reconstruct P. In the self-signing operation, we know the ephemeral key is (k,P). No need to loop to recover P. Then we get the public key is Q = r^(-1)*(s*P - e*G). Because (k,P) is a valid key pair, we know that P = k*G, so Q = r^(-1)(s*k - e)*G. But we already know this if we look at step 6 in the standard which tells us that the private key is d = r^(-1)(s*k - e). If (d,Q) is to be a valid key pair, we must have Q = d*G. In summary, there is no need to actually use the public key recovery operation. We can simply defined d as shown in step 6 and set Q = d*G. Special note for computing d: As mentioned above, Isabelle/HOL is a typed language. We need to make sure that the difference computed for d is done on integers, in case that s*k < e.\ text \These are the conditions in the standard for the user of the self-signing operation. They must pick a valid ephemeral key (k,P) and a value s in [1, n-1].\ definition ECDSA_SelfSign_validIn' :: "nat \ int point \ nat \ bool" where "ECDSA_SelfSign_validIn' k P s = (ECkeyPairValid k P \ 0 < s \ s < n)" text \These are the conditions that guarantee that the self-signing operation will produce some output.\ definition ECDSA_SelfSign_validIn :: "hash_type \ ('a \ sig_type \ octets) \ 'a \ nat \ int point \ nat \ bool" where "ECDSA_SelfSign_validIn Hash FormMessage I k P s = ( case P of Infinity \ False | Point x y \ ( let r = nat (x mod n); M = FormMessage I (r,s); e = ECDSA_compute_e Hash M; rInv = inv_mod_n r; d = nat ((rInv*(s*(int k) - (int e)) ) mod n) in 0 < r \ 0 < d ) )" definition ECDSA_SelfSign :: "hash_type \ ('a \ sig_type \ octets) \ 'a \ nat \ int point \ nat \ (nat \ int point \ octets) option" where "ECDSA_SelfSign Hash FormMessage I k P s = ( case P of Infinity \ None | Point x y \ ( let r = nat (x mod n); M = FormMessage I (r,s); e = ECDSA_compute_e Hash M; rInv = inv_mod_n r; d = nat ((rInv*(s*(int k) - (int e)) ) mod n); Q = point_mult' d G in ( if (r=0) then None else ( if (d=0) then None else Some (d, Q, M) )) ) )" lemma ECDSA_SelfSign_ValidIn: assumes "ECDSA_SelfSign_validIn Hash FormMessage I k P s" shows "\d Q M. ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" proof - have 1: "P \ Infinity" using assms ECDSA_SelfSign_validIn_def by force obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" have 3: "0 < ?r" by (smt (verit) assms 1 2 ECDSA_SelfSign_validIn_def point.simps(5)) have 4: "0 < ?d" by (smt (verit) assms 1 2 ECDSA_SelfSign_validIn_def point.simps(5)) have 5: "?r \ 0" using 3 by fast have 6: "?d \ 0" using 4 by fast show ?thesis using 1 2 5 6 ECDSA_SelfSign_def option.discI point.simps(5) by (smt (z3)) qed lemma ECDSA_SelfSign_PnotInf: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "P \ Infinity" by (metis (no_types, lifting) assms(1) ECDSA_SelfSign_def option.discI point.simps(4)) lemma ECDSA_SelfSign_Some: fixes x :: int assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "M' = FormMessage I (r,s)" "e = ECDSA_compute_e Hash M'" "rInv = inv_mod_n r" "d' = nat ((rInv*(s*(int k) - (int e)) ) mod n)" "Q' = point_mult' d' G" shows "d = d' \ Q = Q' \ M = M' \ r \0 \ d \ 0" by (smt (verit, del_insts) assms ECDSA_SelfSign_def option.simps(1,3) point.case(2) prod.sel(1,2)) lemma ECDSA_SelfSign_SomeValid: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "ECDSA_SelfSign_validIn Hash FormMessage I k P s" proof - have 1: "P \ Infinity" by (metis ECDSA_SelfSign_PnotInf assms) obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' ?d G" have 3: "d = ?d \ Q = ?Q \ M = ?M \ ?r \0 \ d \ 0" by (metis 2 ECDSA_SelfSign_Some assms) have 4: "0 < ?r \ 0 < ?d" using 3 by force show ?thesis using 1 2 4 ECDSA_SelfSign_validIn_def by fastforce qed text \The output (d,Q) of the self-signing operation is a valid key pair.\ lemma ECDSA_SelfSign_ValidKeyPair: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" shows "ECkeyPairValid d Q" proof - have 1: "P \ Infinity" by (metis ECDSA_SelfSign_PnotInf assms) obtain x and y where 2: "P = Point x y" by (meson 1 point.exhaust) let ?r = "nat (x mod n)" let ?M = "FormMessage I (?r,s)" let ?e = "ECDSA_compute_e Hash ?M" let ?rInv = "inv_mod_n ?r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' ?d G" have 3: "d = ?d \ Q = ?Q \ M = ?M \ ?r \0 \ d \ 0" by (metis 2 ECDSA_SelfSign_Some assms) have 4: "d < n" by (metis 3 Euclidean_Rings.pos_mod_bound N.residues_prime_axioms nat_int prime_gt_1_nat residues_prime.p_prime zless_nat_conj) show ?thesis using 3 4 ECkeyPairValid_def by blast qed text \Here is the main result about the self-signing operation. Namely, if you follow the self-signing operation and get (d, Q, M), then the message M is formed from the information I and (r,s) and when M is signed using (d,Q) as the signer's key pair (together with the ephemeral key (k,P), then the resulting signature is (r,s).\ lemma ECDSA_SelfSign_SignWorks: assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "ECkeyPairValid k P" "s ECDSA_Sign Hash d M k P = Some (r,s)" proof - let ?M = "FormMessage I (r,s)" let ?e = "ECDSA_compute_e Hash M" let ?rInv = "inv_mod_n r" let ?d = "nat ((?rInv*(s*(int k) - (int ?e)) ) mod n)" let ?Q = "point_mult' d G" have 1: "d = ?d \ Q = ?Q \ M = ?M \ r \0 \ d \ 0" by (metis ECDSA_SelfSign_Some assms(1,2,3)) have r1: "r \ carrier Rn" by (metis 1 N.mod_in_carrier assms(3) nat_eq_iff2) have r2: "?rInv = inv\<^bsub>Rn\<^esub> r" using 1 r1 inv_mod_n_inv by presburger have d1: "?d = inv\<^bsub>Rn\<^esub>r \\<^bsub>Rn\<^esub> (s \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> ?e)" by (metis 1 r2 N.res_diff_eq N.res_mult_eq int_nat_eq int_ops(1) mod_diff_left_eq mod_mult_right_eq nat_int) let ?kinv = "inv_mod_n k" have k1: "0 < k \ k < n" using assms(4) ECkeyPairValid_def by blast have k2: "?kinv = inv\<^bsub>Rn\<^esub> k" using inv_mod_n_inv' k1 by presburger have k3: "k \ carrier Rn" using k1 ECkeyPair_dInRn' ECprivateKeyValid_def by blast let ?s = "(?kinv*(?e + r*?d)) mod n" have s1: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> (?e \\<^bsub>Rn\<^esub> r \\<^bsub>Rn\<^esub> ?d)" by (simp add: k2 N.res_add_eq N.res_mult_eq mod_add_right_eq mod_mult_right_eq of_nat_mod) have d2: "r \\<^bsub>Rn\<^esub> ?d = s \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> ?e" by (metis 1 r1 d1 EPF.nat_int_eq N.inv_closed N.l_one N.m_assoc N.mod_in_carrier N.r_inv N.res_diff_eq N.zero_cong int_ops(1)) have s2: "(?e \\<^bsub>Rn\<^esub> r \\<^bsub>Rn\<^esub> ?d) = s \\<^bsub>Rn\<^esub> k" using d2 by (simp add: N.res_add_eq N.res_diff_eq N.res_mult_eq mod_add_right_eq) have s3: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> (s \\<^bsub>Rn\<^esub> k)" using s1 s2 by argo have s4: "?s = inv\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> k \\<^bsub>Rn\<^esub> s" using s3 by (simp add: N.res_mult_eq mod_mult_right_eq mult.commute mult.left_commute) have s5: "?s = s" by (metis s4 k3 assms(4,5,6) ECkeyPair_dInRn' ECkeyPair_invRn ECprivateKeyValid_def N.l_one N.m_comm nat_int) show ?thesis by (smt (verit) 1 s5 ECDSA_Sign_e_def assms(2,3,6) bot_nat_0.not_eq_extremum point.simps(5)) qed text \This is a restatement of the lemma above where we use "ECDSA_SelfSign_validIn k P s", which again means that (k,P) is a valid key pair and s is in [1,n-1].\ lemma ECDSA_SelfSign_SignWorks': assumes "ECDSA_SelfSign Hash FormMessage I k P s = Some (d, Q, M)" "P = Point x y" "r = nat (x mod n)" "ECDSA_SelfSign_validIn' k P s" shows "M = FormMessage I (r,s) \ ECDSA_Sign Hash d M k P = Some (r,s)" using ECDSA_SelfSign_SignWorks assms ECDSA_SelfSign_validIn'_def by fastforce end (* SEC1 locale *) section \5 Encryption and Key Transport Schemes\ text \"This section specifies the public-key encryption and key transport schemes based on ECC supported in this document. Public-key encryption schemes are designed to be used by two entities - a sender U and a recipient V - when U wants to send a message M to V confidentially, and V wants to recover M. Key transport schemes are a special class of public-key encryption schemes where the message M is restricted to be a cryptographic key, usually a symmetric key. Except for this restriction, most of the discussion below about public-key encryption schemes also applies to key transport schemes. Here, public-key encryption schemes are described in terms of an encryption operation, a decryption operation, and associated setup and key deployment procedures. Entities U and V should use the scheme as follows when they want to communicate. First U and V should use the setup procedure to establish which options to use the scheme with, then V should use the key deployment procedure to select a key pair and U should obtain V's public key - U will use V's public key to control the encryption procedure, and V will use its key pair to control the decryption operation. Then each time U wants to send a message M to V, U should apply the encryption operation to M under V's public key to compute an encryption or ciphertext C of M , and convey C to V . Finally when V receives C, entity V should apply the decryption operation to C under its key pair to recover the message M. Loosely speaking, public-key encryption schemes are designed so that it is hard for an adversary who does not possess V 's secret key to recover messages from their ciphertexts. Thereby, public-key encryption schemes provide data confidentiality."\ subsection \5.1 Elliptic Curve Integrated Encryption Scheme\ locale ECIES = SEC1 + fixes KDF :: kdf_type and KDF_VI :: kdf_validin_type and MAC :: mac_type and MAC_VI :: mac_validin_type and mackeylen maclen :: nat and ENC :: enc_type and ENC_VI :: enc_validin_type and DEC :: dec_type and DEC_VI :: dec_validin_type and enckeylen :: nat and useCoDH useCompress :: bool and d\<^sub>V :: nat and Q\<^sub>V :: "int point" assumes VkeyPairValid : "ECkeyPairValid d\<^sub>V Q\<^sub>V" and ENC_valid : "\K M. ENC_VI K M \ DEC K (ENC K M) = M" begin subsubsection \5.1.3 Encryption Operation\ text\Note for Step 6: This definition doesn't cover the case when ENC = XOR and backward compatibility mode is not selected.\ definition ECIES_Encryption :: "octets \ octets \ octets \ nat \ int point \ (octets \ octets \ octets) option" where "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = ( case (ECDHprimChoose useCoDH k Q\<^sub>V) of None \ None | Some z \ ( let Pbar = point_to_octets P useCompress; mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; EM = ENC EK M; D = MAC MK (EM @ SharedInfo2) in Some (Pbar, EM, D) ) )" definition ECIES_Encryption_validIn :: "octets \ octets \ octets \ nat \ int point \ bool" where "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P = ( case (ECDHprimChoose useCoDH k Q\<^sub>V) of None \ False | Some z \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; EM = ENC EK M in (ECkeyPairValid k P) \ (ENC_VI EK M) \ (KDF_VI Z (enckeylen + mackeylen) SharedInfo1) \ (MAC_VI MK (EM @ SharedInfo2)) ) )" lemma ECIES_Encryption_validIn_someOut: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" shows "ECIES_Encryption M SharedInfo1 SharedInfo2 k P \ None" by (smt (z3) assms ECIES_Encryption_validIn_def ECIES_Encryption_def case_optionE option.discI option.simps(5)) lemma ECIES_Encryption_validIn_someZ: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" shows "\z. ECDHprimChoose useCoDH k Q\<^sub>V = Some z" using assms ECIES_Encryption_validIn_def case_optionE by blast lemma ECIES_Encryption_someZout: assumes "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" shows "\Pbar EM D. ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" by (metis (no_types, lifting) assms ECIES_Encryption_def option.simps(5)) lemma ECIES_Encryption_SomeOutEq: assumes "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" "Pbar = point_to_octets P useCompress" "mlen = octet_length p" "Z = nat_to_octets_len (nat z) mlen" "K = KDF Z (enckeylen + mackeylen) SharedInfo1" "EK = take enckeylen K" "MK = drop enckeylen K" "EM = ENC EK M" "D = MAC MK (EM @ SharedInfo2)" shows "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" by (metis (mono_tags, lifting) assms ECIES_Encryption_def option.simps(5)) subsubsection \5.1.4 Decryption Operation\ definition ECIES_Decryption :: "(octets \ octets \ octets) \ octets \ octets \ octets option" where "ECIES_Decryption C SharedInfo1 SharedInfo2 = ( let Pbar = fst C; EM = fst (snd C); D = snd (snd C) in ( case (octets_to_point' Pbar) of None \ None | Some P \ ( let T = (if useCoDH then (ECpublicKeyPartialValid P) else (ECpublicKeyValid P) ) in ( case (ECDHprimChoose useCoDH d\<^sub>V P) of None \ None | Some z \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z) mlen; K = KDF Z (enckeylen + mackeylen) SharedInfo1; EK = take enckeylen K; MK = drop enckeylen K; D' = MAC MK (EM @ SharedInfo2); M = DEC EK EM in if (T \ D = D') then (Some M) else None ) ) ) ) )" text \This is the main thing we would like to know. If you encrypt a message with ECIES_Encryption, then you can get back that message with ECIES_Decryption.\ lemma ECIES_Decryption_validIn: assumes "ECIES_Encryption_validIn M SharedInfo1 SharedInfo2 k P" "ECIES_Encryption M SharedInfo1 SharedInfo2 k P = Some (Pbar, EM, D)" shows "ECIES_Decryption (Pbar, EM, D) SharedInfo1 SharedInfo2 = Some M" proof - obtain z where z1: "ECDHprimChoose useCoDH k Q\<^sub>V = Some z" using ECIES_Encryption_validIn_someZ assms(1) by blast let ?mlen = "octet_length p" let ?Z = "nat_to_octets_len (nat z) ?mlen" let ?K = "KDF ?Z (enckeylen + mackeylen) SharedInfo1" let ?EK = "take enckeylen ?K" let ?MK = "drop enckeylen ?K" have P1: "Pbar = point_to_octets P useCompress" by (metis (mono_tags, lifting) assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1)) have P2: "ECkeyPairValid k P" by (metis (no_types, lifting) assms(1) z1 ECIES_Encryption_validIn_def option.simps(5)) have P3: "on_curve' P" using P2 ECkeyPairOnCurve by blast have P4: "octets_to_point' Pbar = Some P" using P1 P3 point2Octets2Point by fast have P5: "ECpublicKeyValid P \ ECpublicKeyPartialValid P" using P2 keyPairValidImpliespublicKeyValid validImpliesPartialValid by blast let ?T = "(if useCoDH then (ECpublicKeyPartialValid P) else (ECpublicKeyValid P) )" have T1: "?T" using P5 by presburger have M1: "EM = ENC ?EK M" using assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1) by simp let ?M = "DEC ?EK EM" have M2: "ENC_VI ?EK M" using assms(1) ECIES_Encryption_validIn_def by (smt (verit, best) assms(1) ECIES_Encryption_validIn_def z1 option.simps(5)) have M3: "?M = M" using M1 M2 ENC_valid by blast have D1: "D = MAC ?MK (EM @ SharedInfo2)" using assms(2) z1 ECIES_Encryption_SomeOutEq option.simps(1) prod.sel(1) by force have z2: "ECDHprimChoose useCoDH k Q\<^sub>V = ECDHprimChoose useCoDH d\<^sub>V P" using P2 VkeyPairValid ECDHch_2ValidKeyPairs by blast show ?thesis using z1 z2 P4 T1 M3 D1 ECIES_Decryption_def by force qed subsection \5.2 Wrapped Key Transport Scheme\ text \"The wrapped key transport scheme uses a combination of a key wrap scheme and a key agreement scheme." This section of the standard discusses how a wrapped key transport scheme may work without making any specific definitions.\ end (* of ECIES locale *) section \6 Key Agreement Schemes\ text \"Key agreement schemes are designed to be used by two entities --- an entity U and an entity V --- when U and V want to establish shared keying data that they can later use to control the operation of a symmetric cryptographic scheme. ... Note that this document does not address how specific keys should be derived from keying data established using a key agreement scheme. This detail is left to be determined on an application by application basis." \ subsection \6.1 Elliptic Curve Diffie-Hellman Scheme\ locale ECDH = SEC1 + fixes KDF :: kdf_type and KDF_VI :: kdf_validin_type and useCoDH :: bool and d\<^sub>V d\<^sub>U :: nat and Q\<^sub>V Q\<^sub>U :: "int point" assumes VkeyPairValid : "ECkeyPairValid d\<^sub>V Q\<^sub>V" and UkeyPairValid : "ECkeyPairValid d\<^sub>U Q\<^sub>U" begin subsubsection \6.1.3 Key Agreement Operation\ definition ECDH_KeyAgreement :: "nat \ octets \ octets option" where "ECDH_KeyAgreement keydatalen SharedInfo = ( let z = ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" text \Remember ECDHchoose_validKeys: Because these are both valid key pairs, we know that \\z. ECDHprimChoose useCo d Q = Some z\.\ definition ECDH_KeyAgreement_validIn :: "nat \ octets \ bool" where "ECDH_KeyAgreement_validIn keydatalen SharedInfo \ \z. (ECDHprimChoose useCoDH d\<^sub>U Q\<^sub>V = Some z) \ (KDF_VI (nat_to_octets_len (nat z) (octet_length p)) keydatalen SharedInfo)" text \This lemma shows that entity V can get the same key that U produced. U produces the key with U's private key and V's public key. V produces the same key with V's private key and U's public key.\ lemma ECDH_KeyAgreement_eq: "ECDH_KeyAgreement keydatalen SharedInfo = ( let z = ECDHprimChoose useCoDH d\<^sub>V Q\<^sub>U in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" by (metis ECDH_KeyAgreement_def ECDHch_2ValidKeyPairs UkeyPairValid VkeyPairValid option.simps(5) ECDHchoose_validKeys ECkeyPairImpliesPrivateKeyValid keyPairValidImpliespublicKeyValid) end (* of ECDH locale *) subsection \6.2 Elliptic Curve MQV Scheme\ locale ECMQV = SEC1 + fixes KDF :: kdf_type and d1U d2U d1V d2V :: nat and Q1U Q2U Q1V Q2V :: "int point" assumes VkeyPairValid : "ECkeyPairValid d1V Q1V" "ECkeyPairValid d2V Q2V" and UkeyPairValid : "ECkeyPairValid d1U Q1U" "ECkeyPairValid d2U Q2U" begin subsubsection \6.2.3 Key Agreement Operation\ definition ECMQV_KeyAgreement :: "nat \ octets \ octets option" where "ECMQV_KeyAgreement keydatalen SharedInfo = ( let z = ECMQVprim d1U Q1U d2U Q2U Q1V Q2V in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" text \As with the Diffie-Hellman Key Agreement operation, the important thing to note with the ECMQV Key Agreement operation is that both U and V can compute the same key using their own keys together with the public keys of the other party. This follows from MQV_reverseUV above.\ lemma ECMQV_KeyAgreement_eq: "ECMQV_KeyAgreement keydatalen SharedInfo = ( let z = ECMQVprim d1V Q1V d2V Q2V Q1U Q2U in case z of None \ None | Some z' \ ( let mlen = octet_length p; Z = nat_to_octets_len (nat z') mlen in Some (KDF Z keydatalen SharedInfo) ) )" by (metis ECMQV_KeyAgreement_def MQV_reverseUV UkeyPairValid VkeyPairValid) end (* of ECMQV locale *) end \ No newline at end of file diff --git a/thys/Deep_Learning/Tensor.thy b/thys/Deep_Learning/Tensor.thy --- a/thys/Deep_Learning/Tensor.thy +++ b/thys/Deep_Learning/Tensor.thy @@ -1,253 +1,257 @@ (* Author: Alexander Bentkamp, Universität des Saarlandes *) section \Tensor\ theory Tensor imports Main begin typedef 'a tensor = "{t::nat list \ 'a list. length (snd t) = prod_list (fst t)}" by (simp add: Ex_list_of_length) definition dims::"'a tensor \ nat list" where "dims A = fst (Rep_tensor A)" definition vec::"'a tensor \ 'a list" where "vec A = snd (Rep_tensor A)" definition tensor_from_vec::"nat list \ 'a list \ 'a tensor" where "tensor_from_vec d v = Abs_tensor (d,v)" lemma assumes "length v = prod_list d" shows dims_tensor[simp]: "dims (tensor_from_vec d v) = d" and vec_tensor[simp]: "vec (tensor_from_vec d v) = v" by (simp add: Abs_tensor_inverse assms dims_def tensor_from_vec_def vec_def)+ lemma tensor_from_vec_simp[simp]: "tensor_from_vec (dims A) (vec A) = A" by (simp add: Rep_tensor_inverse Tensor.vec_def dims_def tensor_from_vec_def) lemma length_vec: "length (vec A) = prod_list (dims A)" by (metis (mono_tags, lifting) Rep_tensor Tensor.vec_def dims_def mem_Collect_eq) lemma tensor_eqI[intro]: assumes "dims A = dims B" and "vec A = vec B" shows "A=B" by (metis assms tensor_from_vec_simp) abbreviation order::"'a tensor \ nat" where "order t == length (dims t)" inductive valid_index::"nat list \ nat list \ bool" (infix "\" 50) where Nil: "[] \ []" | Cons: "is \ ds \ i i#is \ d#ds" inductive_cases valid_indexE[elim]: "is \ ds" inductive_cases valid_index_dimsE[elim]: "is \ dims A" lemma valid_index_length: "is \ ds \ length is = length ds" by (induction rule:valid_index.induct; auto) lemma valid_index_lt: "is \ ds \ m is!m < ds!m" proof (induction arbitrary:m rule:valid_index.induct) case Nil then show ?case by auto next case Cons then show ?case by (metis gr0_conv_Suc length_Cons linorder_neqE_nat not_less_eq nth_Cons' nth_Cons_Suc) qed lemma valid_indexI: assumes "length is = length ds" and "\m. m is!m < ds!m" shows "is \ ds" using assms proof (induction "is" arbitrary:ds) case Nil then show ?case by (metis length_0_conv valid_index.simps) next case (Cons a "is" ds) then obtain d ds' where "ds = d # ds'" by (metis length_Suc_conv) then have "is \ ds'" using Cons by (metis length_Cons less_irrefl linorder_neqE_nat not_less_eq nth_Cons_Suc) then show ?case using Cons.prems(2) \ds = d # ds'\ valid_index.Cons by fastforce qed lemma valid_index_append: assumes is1_valid:"is1 \ ds1" and is2_valid:"is2 \ ds2" shows "is1 @ is2 \ ds1 @ ds2" apply (rule valid_indexI[of "is1 @ is2" "ds1 @ ds2"]) unfolding nth_append using valid_index_lt[OF is2_valid] valid_index_lt[OF is1_valid] valid_index_length[OF is1_valid] valid_index_length[OF is2_valid] length_append by (auto simp add: \length is1 = length ds1\) lemma valid_index_list_all2_iff: "is \ ds \ list_all2 (<) is ds" by (metis list_all2_conv_all_nth list_all2_nthD valid_indexI valid_index_length valid_index_lt) definition fixed_length_sublist::"'a list \ nat \ nat \ 'a list" where "fixed_length_sublist xs l i = (take l (drop (l*i) xs))" fun lookup_base::"nat list \ 'a list \ nat list \ 'a" where lookup_base_Nil: "lookup_base [] v [] = hd v" | lookup_base_Cons: "lookup_base (d # ds) v (i # is) = lookup_base ds (fixed_length_sublist v (prod_list ds) i) is" definition lookup::"'a tensor \ nat list \ 'a" where "lookup A = lookup_base (dims A) (vec A)" fun tensor_vec_from_lookup::"nat list \ (nat list \ 'a) \ 'a list" where tensor_vec_from_lookup_Nil: "tensor_vec_from_lookup [] e = [e []]" | tensor_vec_from_lookup_Cons: "tensor_vec_from_lookup (d # ds) e = concat (map (\i. tensor_vec_from_lookup ds (\is. e (i # is))) [0.. (nat list \ 'a) \ 'a tensor" where "tensor_from_lookup ds e = tensor_from_vec ds (tensor_vec_from_lookup ds e)" lemma concat_parts_leq: assumes "a * d \ length v" shows "concat (map (fixed_length_sublist v d) [0..is. is \ ds \ lookup_base ds v is = e is" shows "tensor_vec_from_lookup ds e = v" using assms proof (induction ds arbitrary:v e) case Nil then show ?case unfolding tensor_vec_from_lookup.simps by (metis One_nat_def Tensor.lookup_base_Nil length_0_conv length_Suc_conv list.sel(1) prod_list.Nil valid_index.Nil) next case (Cons a ds) - then have "a * prod_list ds = length v" by auto + then have \length v = prod_list ds * a\ by auto { - fix i assume "i length v" using \a * prod_list ds = length v\ using discrete mult.commute mult_le_mono1 by metis + fix i assume \i < a\ + then have \Suc i \ a\ + by simp + then have "prod_list ds * Suc i \ length v" + using \length v = prod_list ds * a\ + by (simp only: mult_le_mono2) have "\is'. is' \ ds \ e (i # is') = lookup_base ds (fixed_length_sublist v (prod_list ds) i) is'" using \i by (metis Cons.prems(2) Tensor.lookup_base_Cons valid_index.simps) then have "tensor_vec_from_lookup ds (\is'. e (i # is')) = fixed_length_sublist v (prod_list ds) i" - using Cons using \prod_list ds * (i + 1) \ length v\ by (simp add: Cons.IH fixed_length_sublist_def) + using Cons using \prod_list ds * Suc i \ length v\ by (simp add: Cons.IH fixed_length_sublist_def) } then show ?case unfolding tensor_vec_from_lookup_Cons lookup_base_Cons - using concat_parts_eq[OF \a * prod_list ds = length v\] - atLeastLessThan_iff map_eq_conv set_upt Cons by (metis (no_types, lifting)) + using atLeastLessThan_iff map_eq_conv set_upt Cons concat_parts_eq prod_list.Cons + by (metis (no_types, lifting)) qed lemma tensor_lookup: assumes "\is. is \ dims A \ lookup A is = e is" shows "tensor_from_lookup (dims A) e = A" using tensor_lookup_base lookup_def length_vec tensor_from_lookup_def by (metis assms tensor_from_vec_simp) lemma concat_equal_length: assumes "\xs. xs\set xss \ length xs = l" shows "length (concat xss) = length xss*l" using assms by (induction xss;auto) lemma concat_equal_length_map: assumes "\i. i length (f i) = d" shows "length (concat (map (\i. f i) [0..xs. xs\set xss \ length xs = d" and "ii. i length (f i) = d" and "ii. f i) [0..i. i < a \ length (f i) = d)" by auto then have "length (concat (map f [0..length (concat (map f [0.. fixed_length_sublist_def) then show ?case using \i=a\ by auto next assume "i\a" then have "fixed_length_sublist (concat (map f [0..concat (map f [0.. unfolding fixed_length_sublist_def drop_append using \length (concat (map f [0.. \fixed_length_sublist (concat (map f [0.. using append_assoc append_eq_conv_conj append_take_drop_id assms(1) assms(2) fixed_length_sublist_def by metis qed qed lemma length_tensor_vec_from_lookup: "length (tensor_vec_from_lookup ds e) = prod_list ds" by (induction ds arbitrary:e; auto simp add: concat_equal_length_map) lemma lookup_tensor_vec: assumes "is\ds" shows "lookup_base ds (tensor_vec_from_lookup ds e) is = e is" using assms proof (induction arbitrary:e rule:valid_index.induct) case Nil then show ?case by simp next case (Cons "is" ds i d e) then show ?case unfolding tensor_vec_from_lookup_Cons lookup_base_Cons by (simp add: length_tensor_vec_from_lookup concat_parts'[of d "\i. tensor_vec_from_lookup ds (\is. e (i # is))" "prod_list ds" i] \i < d\) qed lemma lookup_tensor_from_lookup: assumes "is\ds" shows "lookup (tensor_from_lookup ds e) is = e is" unfolding lookup_def tensor_from_lookup_def by (simp add: lookup_tensor_vec assms length_tensor_vec_from_lookup) lemma dims_tensor_from_lookup: "dims (tensor_from_lookup ds e) = ds" unfolding tensor_from_lookup_def by (simp add: length_tensor_vec_from_lookup) lemma tensor_lookup_cong: assumes "tensor_from_lookup ds e\<^sub>1 = tensor_from_lookup ds e\<^sub>2" and "is\ds" shows "e\<^sub>1 is = e\<^sub>2 is" using assms lookup_tensor_from_lookup by metis lemma tensor_from_lookup_eqI: assumes "\is. is\ds \ e\<^sub>1 is = e\<^sub>2 is" shows "tensor_from_lookup ds e\<^sub>1 = tensor_from_lookup ds e\<^sub>2" by (metis assms lookup_tensor_vec length_tensor_vec_from_lookup tensor_lookup_base tensor_from_lookup_def) lemma tensor_lookup_eqI: assumes "dims A = dims B" and "\is. is\(dims A) \ lookup A is = lookup B is" shows "A = B" by (metis assms(1) assms(2) tensor_lookup) end diff --git a/thys/Earley_Parser/Earley.thy b/thys/Earley_Parser/Earley.thy --- a/thys/Earley_Parser/Earley.thy +++ b/thys/Earley_Parser/Earley.thy @@ -1,493 +1,493 @@ theory Earley imports Derivations begin section \Slices\ fun slice :: "nat \ nat \ 'a list \ 'a list" where "slice _ _ [] = []" | "slice _ 0 (x#xs) = []" | "slice 0 (Suc b) (x#xs) = x # slice 0 b xs" | "slice (Suc a) (Suc b) (x#xs) = slice a b xs" lemma slice_drop_take: "slice a b xs = drop a (take b xs)" by (induction a b xs rule: slice.induct) auto lemma slice_append_aux: "Suc b \ c \ slice (Suc b) c (x # xs) = slice b (c-1) xs" using Suc_le_D by fastforce lemma slice_concat: "a \ b \ b \ c \ slice a b xs @ slice b c xs = slice a c xs" proof (induction a b xs arbitrary: c rule: slice.induct) case (3 b x xs) then show ?case using Suc_le_D by(fastforce simp: slice_append_aux) qed (auto simp: slice_append_aux) lemma slice_concat_Ex: "a \ c \ slice a c xs = ys @ zs \ \b. ys = slice a b xs \ zs = slice b c xs \ a \ b \ b \ c" proof (induction a c xs arbitrary: ys zs rule: slice.induct) case (3 b x xs) show ?case proof (cases ys) case Nil then obtain zs' where "x # slice 0 b xs = x # zs'" "x # zs' = zs" using "3.prems"(2) by auto thus ?thesis using Nil by force next case (Cons y ys') then obtain ys' where "x # slice 0 b xs = x # ys' @ zs" "x # ys' = ys" using "3.prems"(2) by auto thus ?thesis using "3.IH"[of ys' zs] by force qed next case (4 a b x xs) thus ?case by (auto, metis slice.simps(4) Suc_le_mono) qed auto lemma slice_nth: "a < length xs \ slice a (a+1) xs = [xs!a]" unfolding slice_drop_take by (metis Cons_nth_drop_Suc One_nat_def diff_add_inverse drop_take take_Suc_Cons take_eq_Nil) lemma slice_append_nth: "a \ b \ b < length xs \ slice a b xs @ [xs!b] = slice a (b+1) xs" by (metis le_add1 slice_concat slice_nth) lemma slice_empty: "b \ a \ slice a b xs = []" by (simp add: slice_drop_take) lemma slice_id[simp]: "slice 0 (length xs) xs = xs" by (simp add: slice_drop_take) lemma slice_singleton: "b \ length xs \ [x] = slice a b xs \ b = a + 1" by (induction a b xs rule: slice.induct) (auto simp: slice_drop_take) section \Earley recognizer\ subsection \Earley items\ definition rule_head :: "'a rule \ 'a" where "rule_head \ fst" definition rule_body :: "'a rule \ 'a list" where "rule_body \ snd" datatype 'a item = Item (item_rule: "'a rule") (item_dot : nat) (item_origin : nat) (item_end : nat) definition item_rule_head :: "'a item \ 'a" where "item_rule_head x \ rule_head (item_rule x)" definition item_rule_body :: "'a item \ 'a sentence" where "item_rule_body x \ rule_body (item_rule x)" definition item_\ :: "'a item \ 'a sentence" where "item_\ x \ take (item_dot x) (item_rule_body x)" definition item_\ :: "'a item \ 'a sentence" where "item_\ x \ drop (item_dot x) (item_rule_body x)" definition is_complete :: "'a item \ bool" where "is_complete x \ item_dot x \ length (item_rule_body x)" definition next_symbol :: "'a item \ 'a option" where "next_symbol x \ if is_complete x then None else Some (item_rule_body x ! item_dot x)" lemmas item_defs = item_rule_head_def item_rule_body_def item_\_def item_\_def rule_head_def rule_body_def definition is_finished :: "'a cfg \ 'a sentence \ 'a item \ bool" where "is_finished \ \ x \ item_rule_head x = \ \ \ item_origin x = 0 \ item_end x = length \ \ is_complete x" definition recognizing :: "'a item set \ 'a cfg \ 'a sentence \ bool" where "recognizing I \ \ \ \x \ I. is_finished \ \ x" inductive_set Earley :: "'a cfg \ 'a sentence \ 'a item set" for \ :: "'a cfg" and \ :: "'a sentence" where Init: "r \ set (\ \) \ fst r = \ \ \ Item r 0 0 0 \ Earley \ \" | Scan: "x = Item r b i j \ x \ Earley \ \ \ \!j = a \ j < length \ \ next_symbol x = Some a \ Item r (b + 1) i (j + 1) \ Earley \ \" | Predict: "x = Item r b i j \ x \ Earley \ \ \ r' \ set (\ \) \ next_symbol x = Some (rule_head r') \ Item r' 0 j j \ Earley \ \" | Complete: "x = Item r\<^sub>x b\<^sub>x i j \ x \ Earley \ \ \ y = Item r\<^sub>y b\<^sub>y j k \ y \ Earley \ \ \ is_complete y \ next_symbol x = Some (item_rule_head y) \ Item r\<^sub>x (b\<^sub>x + 1) i k \ Earley \ \" subsection \Well-formedness\ definition wf_item :: "'a cfg \ 'a sentence => 'a item \ bool" where "wf_item \ \ x \ item_rule x \ set (\ \) \ item_dot x \ length (item_rule_body x) \ item_origin x \ item_end x \ item_end x \ length \" lemma wf_Init: assumes "r \ set (\ \)" "fst r = \ \" shows "wf_item \ \ (Item r 0 0 0)" using assms unfolding wf_item_def by simp lemma wf_Scan: assumes "x = Item r b i j" "wf_item \ \ x" "\!j = a" "j < length \" "next_symbol x = Some a" shows "wf_item \ \ (Item r (b + 1) i (j+1))" using assms unfolding wf_item_def by (auto simp: item_defs is_complete_def next_symbol_def split: if_splits) lemma wf_Predict: assumes "x = Item r b i j" "wf_item \ \ x" "r' \ set (\ \)" "next_symbol x = Some (rule_head r')" shows "wf_item \ \ (Item r' 0 j j)" using assms unfolding wf_item_def by simp lemma wf_Complete: assumes "x = Item r\<^sub>x b\<^sub>x i j" "wf_item \ \ x" "y = Item r\<^sub>y b\<^sub>y j k" "wf_item \ \ y" assumes "is_complete y" "next_symbol x = Some (item_rule_head y)" shows "wf_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" using assms unfolding wf_item_def is_complete_def next_symbol_def item_rule_body_def by (auto split: if_splits) lemma wf_Earley: assumes "x \ Earley \ \" shows "wf_item \ \ x" using assms wf_Init wf_Scan wf_Predict wf_Complete by (induction rule: Earley.induct) fast+ subsection \Soundness\ definition sound_item :: "'a cfg \ 'a sentence \ 'a item \ bool" where "sound_item \ \ x \ derives \ [item_rule_head x] (slice (item_origin x) (item_end x) \ @ item_\ x)" lemma sound_Init: assumes "r \ set (\ \)" "fst r = \ \" shows "sound_item \ \ (Item r 0 0 0)" proof - let ?x = "Item r 0 0 0" have "(item_rule_head ?x, item_\ ?x) \ set (\ \)" using assms(1) by (simp add: item_defs) hence "derives \ [item_rule_head ?x] (item_\ ?x)" using derives_if_valid_rule by metis thus "sound_item \ \ ?x" unfolding sound_item_def by (simp add: slice_empty) qed lemma sound_Scan: assumes "x = Item r b i j" "wf_item \ \ x" "sound_item \ \ x" assumes "\!j = a" "j < length \" "next_symbol x = Some a" shows "sound_item \ \ (Item r (b+1) i (j+1))" proof - define x' where [simp]: "x' = Item r (b+1) i (j+1)" obtain item_\' where *: "item_\ x = a # item_\'" "item_\ x' = item_\'" using assms(1,6) apply (auto simp: item_defs next_symbol_def is_complete_def split: if_splits) by (metis Cons_nth_drop_Suc leI) have "slice i j \ @ item_\ x = slice i (j+1) \ @ item_\'" using * assms(1,2,4,5) by (auto simp: slice_append_nth wf_item_def) moreover have "derives \ [item_rule_head x] (slice i j \ @ item_\ x)" using assms(1,3) sound_item_def by force ultimately show ?thesis using assms(1) * by (auto simp: item_defs sound_item_def) qed lemma sound_Predict: assumes "x = Item r b i j" "wf_item \ \ x" "sound_item \ \ x" assumes "r' \ set (\ \)" "next_symbol x = Some (rule_head r')" shows "sound_item \ \ (Item r' 0 j j)" using assms by (auto simp: sound_item_def derives_if_valid_rule slice_empty item_defs) lemma sound_Complete: assumes "x = Item r\<^sub>x b\<^sub>x i j" "wf_item \ \ x" "sound_item \ \ x" assumes "y = Item r\<^sub>y b\<^sub>y j k" "wf_item \ \ y" "sound_item \ \ y" assumes "is_complete y" "next_symbol x = Some (item_rule_head y)" shows "sound_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" proof - have "derives \ [item_rule_head y] (slice j k \)" using assms(4,6,7) by (auto simp: sound_item_def is_complete_def item_defs) then obtain E where E: "Derivation \ [item_rule_head y] E (slice j k \)" using derives_implies_Derivation by blast have "derives \ [item_rule_head x] (slice i j \ @ item_\ x)" using assms(1,3,4) by (auto simp: sound_item_def) moreover have 0: "item_\ x = (item_rule_head y) # tl (item_\ x)" using assms(8) apply (auto simp: next_symbol_def is_complete_def item_defs split: if_splits) by (metis drop_eq_Nil hd_drop_conv_nth leI list.collapse) ultimately obtain D where D: "Derivation \ [item_rule_head x] D (slice i j \ @ [item_rule_head y] @ (tl (item_\ x)))" using derives_implies_Derivation by (metis append_Cons append_Nil) obtain F where F: "Derivation \ [item_rule_head x] F (slice i j \ @ slice j k \ @ tl (item_\ x))" using Derivation_append_rewrite D E by blast moreover have "i \ j" using assms(1,2) wf_item_def by force moreover have "j \ k" using assms(4,5) wf_item_def by force ultimately have "derives \ [item_rule_head x] (slice i k \ @ tl (item_\ x))" by (metis Derivation_implies_derives append.assoc slice_concat) thus "sound_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" using assms(1,4) by (auto simp: sound_item_def item_defs drop_Suc tl_drop) qed lemma sound_Earley: assumes "x \ Earley \ \" "wf_item \ \ x" shows "sound_item \ \ x" using assms proof (induction rule: Earley.induct) case (Init r) thus ?case using sound_Init by blast next case (Scan x r b i j a) thus ?case using wf_Earley sound_Scan by fast next case (Predict x r b i j r') thus ?case using wf_Earley sound_Predict by blast next case (Complete x r\<^sub>x b\<^sub>x i j y r\<^sub>y b\<^sub>y k) thus ?case using wf_Earley sound_Complete by metis qed theorem soundness_Earley: assumes "recognizing (Earley \ \) \ \" shows "derives \ [\ \] \" proof - obtain x where x: "x \ Earley \ \" "is_finished \ \ x" using assms recognizing_def by blast hence "sound_item \ \ x" using wf_Earley sound_Earley by blast thus ?thesis unfolding sound_item_def using x by (auto simp: is_finished_def is_complete_def item_defs) qed subsection \Completeness\ definition partially_completed :: "nat \ 'a cfg \ 'a sentence \ 'a item set \ ('a derivation \ bool) \ bool" where "partially_completed k \ \ E P \ \r b i' i j x a D. i \ j \ j \ k \ k \ length \ \ x = Item r b i' i \ x \ E \ next_symbol x = Some a \ Derivation \ [a] D (slice i j \) \ P D \ Item r (b+1) i' j \ E" lemma partially_completed_upto: assumes "j \ k" "k \ length \" assumes "x = Item (N,\) d i j" "x \ I" "\x \ I. wf_item \ \ x" assumes "Derivation \ (item_\ x) D (slice j k \)" assumes "partially_completed k \ \ I (\D'. length D' \ length D)" shows "Item (N,\) (length \) i k \ I" using assms proof (induction "item_\ x" arbitrary: d i j k N \ x D) case Nil have "item_\ x = \" using Nil(1,4) unfolding item_\_def item_\_def item_rule_body_def rule_body_def by simp hence "x = Item (N,\) (length \) i j" using Nil.hyps Nil.prems(3-5) unfolding wf_item_def item_defs by auto have "Derivation \ [] D (slice j k \)" using Nil.hyps Nil.prems(6) by auto hence "slice j k \ = []" using Derivation_from_empty by blast hence "j = k" unfolding slice_drop_take using Nil.prems(1,2) by simp thus ?case using \x = Item (N, \) (length \) i j\ Nil.prems(4) by blast next case (Cons b bs) obtain j' E F where *: "Derivation \ [b] E (slice j j' \)" "Derivation \ bs F (slice j' k \)" "j \ j'" "j' \ k" "length E \ length D" "length F \ length D" using Derivation_concat_split[of \ "[b]" bs D "slice j k \"] slice_concat_Ex using Cons.hyps(2) Cons.prems(1,6) by (smt (verit, ccfv_threshold) Cons_eq_appendI append_self_conv2) have "next_symbol x = Some b" using Cons.hyps(2) unfolding item_defs(4) next_symbol_def is_complete_def by (auto, metis nth_via_drop) hence "Item (N, \) (d+1) i j' \ I" using Cons.prems(7) unfolding partially_completed_def using Cons.prems(2,3,4) *(1,3-5) by blast moreover have "partially_completed k \ \ I (\D'. length D' \ length F)" using Cons.prems(7) *(6) unfolding partially_completed_def by fastforce moreover have "bs = item_\ (Item (N,\) (d+1) i j')" using Cons.hyps(2) Cons.prems(3) unfolding item_defs(4) item_rule_body_def by (auto, metis List.list.sel(3) drop_Suc drop_tl) ultimately show ?case using Cons.hyps(1) *(2,4) Cons.prems(2,3,5) wf_item_def by blast qed lemma partially_completed_Earley_k: assumes "wf_\ \" shows "partially_completed k \ \ (Earley \ \) (\_. True)" unfolding partially_completed_def proof (standard, standard, standard, standard, standard, standard, standard, standard, standard) fix r b i' i j x a D assume "i \ j \ j \ k \ k \ length \ \ x = Item r b i' i \ x \ Earley \ \ \ next_symbol x = Some a \ Derivation \ [a] D (slice i j \) \ True" thus "Item r (b + 1) i' j \ Earley \ \" proof (induction "length D" arbitrary: r b i' i j x a D rule: nat_less_induct) case 1 show ?case proof cases assume "D = []" hence "[a] = slice i j \" using "1.prems" by force moreover have "j \ length \" using le_trans "1.prems" by blast ultimately have "j = i+1" using slice_singleton by metis hence "i < length \" - using \j \ length \\ discrete by blast + using \j \ length \\ by simp hence "\!i = a" using slice_nth \[a] = slice i j \\ \j = i + 1\ by fastforce hence "Item r (b + 1) i' j \ Earley \ \" using Earley.Scan "1.prems" \i < length \\ \j = i + 1\ by metis thus ?thesis by (simp add: \j = i + 1\) next assume "\ D = []" then obtain d D' where "D = d # D'" by (meson List.list.exhaust) then obtain \ where *: "Derives1 \ [a] (fst d) (snd d) \" "Derivation \ \ D' (slice i j \)" using "1.prems" by auto hence rule: "(a, \) \ set (\ \)" "fst d = 0" "snd d = (a ,\)" using *(1) unfolding Derives1_def by (simp add: Cons_eq_append_conv)+ show ?thesis proof cases assume "is_terminal \ a" have "is_nonterminal \ a" using rule by (simp add: assms) thus ?thesis using \is_terminal \ a\ is_terminal_nonterminal by (metis assms) next assume "\ is_terminal \ a" define y where y_def: "y = Item (a ,\) 0 i i" have "length D' < length D" using \D = d # D'\ by fastforce hence "partially_completed k \ \ (Earley \ \) (\E. length E \ length D')" unfolding partially_completed_def using "1.hyps" order_le_less_trans by (smt (verit, best)) hence "partially_completed j \ \ (Earley \ \) (\E. length E \ length D')" unfolding partially_completed_def using "1.prems" by force moreover have "Derivation \ (item_\ y) D' (slice i j \)" using *(2) by (auto simp: item_defs y_def) moreover have "y \ Earley \ \" using y_def "1.prems" rule by (auto simp: item_defs Earley.Predict) moreover have "j \ length \" using "1.prems" by simp ultimately have "Item (a,\) (length \) i j \ Earley \ \" using partially_completed_upto "1.prems" wf_Earley y_def by metis moreover have x: "x = Item r b i' i" "x \ Earley \ \" using "1.prems" by blast+ moreover have "next_symbol x = Some a" using "1.prems" by linarith ultimately show ?thesis using Earley.Complete[OF x] by (auto simp: is_complete_def item_defs) qed qed qed qed lemma partially_completed_Earley: "wf_\ \ \ partially_completed (length \) \ \ (Earley \ \) (\_. True)" by (simp add: partially_completed_Earley_k) theorem completeness_Earley: assumes "derives \ [\ \] \" "is_word \ \" "wf_\ \" shows "recognizing (Earley \ \) \ \" proof - obtain \ D where *: "(\ \ ,\) \ set (\ \)" "Derivation \ \ D \" using Derivation_\1 assms derives_implies_Derivation by metis define x where x_def: "x = Item (\ \, \) 0 0 0" have "partially_completed (length \) \ \ (Earley \ \) (\_. True)" using assms(3) partially_completed_Earley by blast hence 0: "partially_completed (length \) \ \ (Earley \ \) (\D'. length D' \ length D)" unfolding partially_completed_def by blast have 1: "x \ Earley \ \" using x_def Earley.Init *(1) by fastforce have 2: "Derivation \ (item_\ x) D (slice 0 (length \) \)" using *(2) x_def by (simp add: item_defs) have "Item (\ \,\) (length \) 0 (length \) \ Earley \ \" using partially_completed_upto[OF _ _ _ _ _ 2 0] wf_Earley 1 x_def by auto then show ?thesis unfolding recognizing_def is_finished_def by (auto simp: is_complete_def item_defs, force) qed subsection \Correctness\ theorem correctness_Earley: assumes "wf_\ \" "is_word \ \" shows "recognizing (Earley \ \) \ \ \ derives \ [\ \] \" using assms soundness_Earley completeness_Earley by blast subsection \Finiteness\ lemma finiteness_empty: "set (\ \) = {} \ finite { x | x. wf_item \ \ x }" unfolding wf_item_def by simp fun item_intro :: "'a rule \ nat \ nat \ nat \ 'a item" where "item_intro (rule, dot, origin, ends) = Item rule dot origin ends" lemma finiteness_nonempty: assumes "set (\ \) \ {}" shows "finite { x | x. wf_item \ \ x }" proof - define M where "M = Max { length (rule_body r) | r. r \ set (\ \) }" define Top where "Top = (set (\ \) \ {0..M} \ {0..length \} \ {0..length \})" hence "finite Top" using finite_cartesian_product finite by blast have "inj_on item_intro Top" unfolding Top_def inj_on_def by simp hence "finite (item_intro ` Top)" using finite_image_iff \finite Top\ by auto have "{ x | x. wf_item \ \ x } \ item_intro ` Top" proof standard fix x assume "x \ { x | x. wf_item \ \ x }" then obtain rule dot origin endp where *: "x = Item rule dot origin endp" "rule \ set (\ \)" "dot \ length (item_rule_body x)" "origin \ length \" "endp \ length \" unfolding wf_item_def using item.exhaust_sel le_trans by blast hence "length (rule_body rule) \ { length (rule_body r) | r. r \ set (\ \) }" using *(1,2) item_rule_body_def by blast moreover have "finite { length (rule_body r) | r. r \ set (\ \) }" using finite finite_image_set[of "\x. x \ set (\ \)"] by fastforce ultimately have "M \ length (rule_body rule)" unfolding M_def by simp hence "dot \ M" using *(1,3) item_rule_body_def by (metis item.sel(1) le_trans) hence "(rule, dot, origin, endp) \ Top" using *(2,4,5) unfolding Top_def by simp thus "x \ item_intro ` Top" using *(1) by force qed thus ?thesis using \finite (item_intro ` Top)\ rev_finite_subset by auto qed lemma finiteness_UNIV_wf_item: "finite { x | x. wf_item \ \ x }" using finiteness_empty finiteness_nonempty by fastforce theorem finiteness_Earley: "finite (Earley \ \)" using finiteness_UNIV_wf_item wf_Earley rev_finite_subset by (metis mem_Collect_eq subsetI) end \ No newline at end of file diff --git a/thys/Earley_Parser/Earley_Parser.thy b/thys/Earley_Parser/Earley_Parser.thy --- a/thys/Earley_Parser/Earley_Parser.thy +++ b/thys/Earley_Parser/Earley_Parser.thy @@ -1,2528 +1,2529 @@ theory Earley_Parser imports Earley_Recognizer "HOL-Library.Monad_Syntax" begin section \Earley parser\ subsection \Pointer lemmas\ definition predicts :: "'a item \ bool" where "predicts x \ item_origin x = item_end x \ item_dot x = 0" definition scans :: "'a sentence \ nat \ 'a item \ 'a item \ bool" where "scans \ k x y \ y = inc_item x k \ (\a. next_symbol x = Some a \ \!(k-1) = a)" definition completes :: "nat \ 'a item \ 'a item \ 'a item \ bool" where "completes k x y z \ y = inc_item x k \ is_complete z \ item_origin z = item_end x \ (\N. next_symbol x = Some N \ N = item_rule_head z)" definition sound_null_ptr :: "'a entry \ bool" where "sound_null_ptr e \ (pointer e = Null \ predicts (item e))" definition sound_pre_ptr :: "'a sentence \ 'a bins \ nat \ 'a entry \ bool" where "sound_pre_ptr \ bs k e \ \pre. pointer e = Pre pre \ k > 0 \ pre < length (bs!(k-1)) \ scans \ k (item (bs!(k-1)!pre)) (item e)" definition sound_prered_ptr :: "'a bins \ nat \ 'a entry \ bool" where "sound_prered_ptr bs k e \ \p ps k' pre red. pointer e = PreRed p ps \ (k', pre, red) \ set (p#ps) \ k' < k \ pre < length (bs!k') \ red < length (bs!k) \ completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" definition sound_ptrs :: "'a sentence \ 'a bins \ bool" where "sound_ptrs \ bs \ \k < length bs. \e \ set (bs!k). sound_null_ptr e \ sound_pre_ptr \ bs k e \ sound_prered_ptr bs k e" definition mono_red_ptr :: "'a bins \ bool" where "mono_red_ptr bs \ \k < length bs. \i < length (bs!k). \k' pre red ps. pointer (bs!k!i) = PreRed (k', pre, red) ps \ red < i" lemma nth_item_bin_upd: "n < length es \ item (bin_upd e es ! n) = item (es!n)" by (induction es arbitrary: e n) (auto simp: less_Suc_eq_0_disj split: entry.splits pointer.splits) lemma bin_upd_append: "item e \ set (items es) \ bin_upd e es = es @ [e]" by (induction es arbitrary: e) (auto simp: items_def split: entry.splits pointer.splits) lemma bin_upd_null_pre: "item e \ set (items es) \ pointer e = Null \ pointer e = Pre pre \ bin_upd e es = es" by (induction es arbitrary: e) (auto simp: items_def split: entry.splits) lemma bin_upd_prered_nop: assumes "distinct (items es)" "i < length es" assumes "item e = item (es!i)" "pointer e = PreRed p ps" "\p ps. pointer (es!i) = PreRed p ps" shows "bin_upd e es = es" using assms by (induction es arbitrary: e i) (auto simp: less_Suc_eq_0_disj items_def split: entry.splits pointer.splits) lemma bin_upd_prered_upd: assumes "distinct (items es)" "i < length es" assumes "item e = item (es!i)" "pointer e = PreRed p rs" "pointer (es!i) = PreRed p' rs'" "bin_upd e es = es'" shows "pointer (es'!i) = PreRed p' (p#rs@rs') \ (\j < length es'. i\j \ es'!j = es!j) \ length (bin_upd e es) = length es" using assms proof (induction es arbitrary: e i es') case (Cons e' es) show ?case proof cases assume *: "item e = item e'" show ?thesis proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ e' = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where ee': "e = Entry x (PreRed xp xs)" "e' = Entry y (PreRed yp ys)" "x = y" using * by auto have simp: "bin_upd e (e' # es') = Entry x (PreRed yp (xp # xs @ ys)) # es'" using True ee' by simp show ?thesis using Cons simp ee' apply (auto simp: items_def) using less_Suc_eq_0_disj by fastforce+ next case False hence "bin_upd e (e' # es') = e' # es'" using * by (auto split: pointer.splits entry.splits) thus ?thesis using False * Cons.prems(1,2,3,4,5) by (auto simp: less_Suc_eq_0_disj items_def split: entry.splits) qed next assume *: "item e \ item e'" have simp: "bin_upd e (e' # es) = e' # bin_upd e es" using * by (auto split: pointer.splits entry.splits) have 0: "distinct (items es)" using Cons.prems(1) unfolding items_def by simp have 1: "i-1 < length es" using Cons.prems(2,3) * by (metis One_nat_def leI less_diff_conv2 less_one list.size(4) nth_Cons_0) have 2: "item e = item (es!(i-1))" using Cons.prems(3) * by (metis nth_Cons') have 3: "pointer e = PreRed p rs" using Cons.prems(4) by simp have 4: "pointer (es!(i-1)) = PreRed p' rs' " using Cons.prems(3,5) * by (metis nth_Cons') have "pointer (bin_upd e es!(i-1)) = PreRed p' (p # rs @ rs') \ (\j < length (bin_upd e es). i-1 \ j \ (bin_upd e es) ! j = es ! j)" using Cons.IH[OF 0 1 2 3 4] by blast hence "pointer ((e' # bin_upd e es) ! i) = PreRed p' (p # rs @ rs') \ (\j < length (e' # bin_upd e es). i \ j \ (e' # bin_upd e es) ! j = (e' # es) ! j)" using * Cons.prems(2,3) less_Suc_eq_0_disj by auto moreover have "e' # bin_upd e es = es'" using Cons.prems(6) simp by auto ultimately show ?thesis by (metis 0 1 2 3 4 Cons.IH Cons.prems(6) length_Cons) qed qed simp lemma sound_ptrs_bin_upd: assumes "sound_ptrs \ bs" "k < length bs" "es = bs!k" "distinct (items es)" assumes "sound_null_ptr e" "sound_pre_ptr \ bs k e" "sound_prered_ptr bs k e" shows "sound_ptrs \ (bs[k := bin_upd e es])" unfolding sound_ptrs_def proof (standard, standard, standard) fix idx elem let ?bs = "bs[k := bin_upd e es]" assume a0: "idx < length ?bs" assume a1: "elem \ set (?bs ! idx)" show "sound_null_ptr elem \ sound_pre_ptr \ ?bs idx elem \ sound_prered_ptr ?bs idx elem" proof cases assume a2: "idx = k" have "elem \ set es \ sound_pre_ptr \ bs idx elem" using a0 a2 assms(1-3) sound_ptrs_def by blast hence pre_es: "elem \ set es \ sound_pre_ptr \ ?bs idx elem" using a2 unfolding sound_pre_ptr_def by force have "elem = e \ sound_pre_ptr \ bs idx elem" using a2 assms(6) by auto hence pre_e: "elem = e \ sound_pre_ptr \ ?bs idx elem" using a2 unfolding sound_pre_ptr_def by force have "elem \ set es \ sound_prered_ptr bs idx elem" using a0 a2 assms(1-3) sound_ptrs_def by blast hence prered_es: "elem \ set es \ sound_prered_ptr (bs[k := bin_upd e es]) idx elem" using a2 assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def by (smt (verit, ccfv_SIG) dual_order.strict_trans1 nth_list_update) have "elem = e \ sound_prered_ptr bs idx elem" using a2 assms(7) by auto hence prered_e: "elem = e \ sound_prered_ptr ?bs idx elem" using a2 assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def by (smt (verit, best) dual_order.strict_trans1 nth_list_update) consider (A) "item e \ set (items es)" | (B) "item e \ set (items es) \ (\pre. pointer e = Null \ pointer e = Pre pre)" | (C) "item e \ set (items es) \ \ (\pre. pointer e = Null \ pointer e = Pre pre)" by blast thus ?thesis proof cases case A hence "elem \ set (es @ [e])" using a1 a2 bin_upd_append assms(2) by force thus ?thesis using assms(1-3,5) pre_e pre_es prered_e prered_es sound_ptrs_def by auto next case B hence "elem \ set es" using a1 a2 bin_upd_null_pre assms(2) by force thus ?thesis using assms(1-3) pre_es prered_es sound_ptrs_def by blast next case C then obtain i p ps where C: "i < length es \ item e = item (es!i) \ pointer e = PreRed p ps" by (metis assms(4) distinct_Ex1 items_def length_map nth_map pointer.exhaust) show ?thesis proof cases assume "\p' ps'. pointer (es!i) = PreRed p' ps'" hence C: "elem \ set es" using a1 a2 C bin_upd_prered_nop assms(2,4) by (metis nth_list_update_eq) thus ?thesis using assms(1-3) sound_ptrs_def pre_es prered_es by blast next assume "\ (\p' ps'. pointer (es!i) = PreRed p' ps')" then obtain p' ps' where D: "pointer (es!i) = PreRed p' ps'" by blast hence 0: "pointer (bin_upd e es!i) = PreRed p' (p#ps@ps') \ (\j < length (bin_upd e es). i\j \ bin_upd e es!j = es!j)" using C assms(4) bin_upd_prered_upd by blast obtain j where 1: "j < length es \ elem = bin_upd e es!j" using a1 a2 assms(2) C items_def bin_eq_items_bin_upd by (metis in_set_conv_nth length_map nth_list_update_eq nth_map) show ?thesis proof cases assume a3: "i=j" hence a3: "pointer elem = PreRed p' (p#ps@ps')" using 0 1 by blast have "sound_null_ptr elem" using a3 unfolding sound_null_ptr_def by simp moreover have "sound_pre_ptr \ ?bs idx elem" using a3 unfolding sound_pre_ptr_def by simp moreover have "sound_prered_ptr ?bs idx elem" unfolding sound_prered_ptr_def proof (standard, standard, standard, standard, standard, standard) fix P PS k' pre red assume a4: "pointer elem = PreRed P PS \ (k', pre, red) \ set (P#PS)" show "k' < idx \ pre < length (bs[k := bin_upd e es]!k') \ red < length (bs[k := bin_upd e es]!idx) \ completes idx (item (bs[k := bin_upd e es]!k'!pre)) (item elem) (item (bs[k := bin_upd e es]!idx!red))" proof cases assume a5: "(k', pre, red) \ set (p#ps)" show ?thesis using 0 1 C a2 a4 a5 prered_es assms(2,3,7) sound_prered_ptr_def length_bin_upd nth_item_bin_upd by (smt (verit) dual_order.strict_trans1 nth_list_update_eq nth_list_update_neq nth_mem) next assume a5: "(k', pre, red) \ set (p#ps)" hence a5: "(k', pre, red) \ set (p'#ps')" using a3 a4 by auto have "k' < idx \ pre < length (bs!k') \ red < length (bs!idx) \ completes idx (item (bs!k'!pre)) (item e) (item (bs!idx!red))" using assms(1-3) C D a2 a5 unfolding sound_ptrs_def sound_prered_ptr_def by (metis nth_mem) thus ?thesis using 0 1 C a4 assms(2,3) length_bin_upd nth_item_bin_upd prered_es sound_prered_ptr_def by (smt (verit, best) dual_order.strict_trans1 nth_list_update_eq nth_list_update_neq nth_mem) qed qed ultimately show ?thesis by blast next assume a3: "i\j" hence "elem \ set es" using 0 1 by (metis length_bin_upd nth_mem order_less_le_trans) thus ?thesis using assms(1-3) pre_es prered_es sound_ptrs_def by blast qed qed qed next assume a2: "idx \ k" have null: "sound_null_ptr elem" using a0 a1 a2 assms(1) sound_ptrs_def by auto have "sound_pre_ptr \ bs idx elem" using a0 a1 a2 assms(1,2) unfolding sound_ptrs_def by simp hence pre: "sound_pre_ptr \ ?bs idx elem" using assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_pre_ptr_def using dual_order.strict_trans1 nth_list_update by fastforce have "sound_prered_ptr bs idx elem" using a0 a1 a2 assms(1,2) unfolding sound_ptrs_def by simp hence prered: "sound_prered_ptr ?bs idx elem" using assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def by (smt (verit, best) dual_order.strict_trans1 nth_list_update) show ?thesis using null pre prered by blast qed qed lemma mono_red_ptr_bin_upd: assumes "mono_red_ptr bs" "k < length bs" "es = bs!k" "distinct (items es)" assumes "\k' pre red ps. pointer e = PreRed (k', pre, red) ps \ red < length es" shows "mono_red_ptr (bs[k := bin_upd e es])" unfolding mono_red_ptr_def proof (standard, standard) fix idx let ?bs = "bs[k := bin_upd e es]" assume a0: "idx < length ?bs" show "\i < length (?bs!idx). \k' pre red ps. pointer (?bs!idx!i) = PreRed (k', pre, red) ps \ red < i" proof cases assume a1: "idx=k" consider (A) "item e \ set (items es)" | (B) "item e \ set (items es) \ (\pre. pointer e = Null \ pointer e = Pre pre)" | (C) "item e \ set (items es) \ \ (\pre. pointer e = Null \ pointer e = Pre pre)" by blast thus ?thesis proof cases case A hence "bin_upd e es = es @ [e]" using bin_upd_append by blast thus ?thesis using a1 assms(1-3,5) mono_red_ptr_def by (metis length_append_singleton less_antisym nth_append nth_append_length nth_list_update_eq) next case B hence "bin_upd e es = es" using bin_upd_null_pre by blast thus ?thesis using a1 assms(1-3) mono_red_ptr_def by force next case C then obtain i p ps where C: "i < length es" "item e = item (es!i)" "pointer e = PreRed p ps" by (metis in_set_conv_nth items_def length_map nth_map pointer.exhaust) show ?thesis proof cases assume "\p' ps'. pointer (es!i) = PreRed p' ps'" hence "bin_upd e es = es" using bin_upd_prered_nop C assms(4) by blast thus ?thesis using a1 assms(1-3) mono_red_ptr_def by (metis nth_list_update_eq) next assume "\ (\p' ps'. pointer (es!i) = PreRed p' ps')" then obtain p' ps' where D: "pointer (es!i) = PreRed p' ps'" by blast have 0: "pointer (bin_upd e es!i) = PreRed p' (p#ps@ps') \ (\j < length (bin_upd e es). i \ j \ bin_upd e es!j = es!j) \ length (bin_upd e es) = length es" using C D assms(4) bin_upd_prered_upd by blast show ?thesis proof (standard, standard, standard, standard, standard, standard, standard) fix j k' pre red PS assume a2: "j < length (?bs!idx)" assume a3: "pointer (?bs!idx!j) = PreRed (k', pre, red) PS" have 1: "?bs!idx = bin_upd e es" by (simp add: a1 assms(2)) show "red < j" proof cases assume a4: "i=j" show ?thesis using 0 1 C(1) D a3 a4 assms(1-3) unfolding mono_red_ptr_def by (metis pointer.inject(2)) next assume a4: "i\j" thus ?thesis using 0 1 a2 a3 assms(1) assms(2) assms(3) mono_red_ptr_def by force qed qed qed qed next assume a1: "idx\k" show ?thesis using a0 a1 assms(1) mono_red_ptr_def by fastforce qed qed lemma sound_mono_ptrs_bin_upds: assumes "sound_ptrs \ bs" "mono_red_ptr bs" "k < length bs" "b = bs!k" "distinct (items b)" "distinct (items es)" assumes "\e \ set es. sound_null_ptr e \ sound_pre_ptr \ bs k e \ sound_prered_ptr bs k e" assumes "\e \ set es. \k' pre red ps. pointer e = PreRed (k', pre, red) ps \ red < length b" shows "sound_ptrs \ (bs[k := bin_upds es b]) \ mono_red_ptr (bs[k := bin_upds es b])" using assms proof (induction es arbitrary: b bs) case (Cons e es) let ?bs = "bs[k := bin_upd e b]" have 0: "sound_ptrs \ ?bs" using sound_ptrs_bin_upd Cons.prems(1,3-5,7) by (metis list.set_intros(1)) have 1: "mono_red_ptr ?bs" using mono_red_ptr_bin_upd Cons.prems(2-5,8) by auto have 2: "k < length ?bs" using Cons.prems(3) by simp have 3: "bin_upd e b = ?bs!k" using Cons.prems(3) by simp have 4: "\e' \ set es. sound_null_ptr e' \ sound_pre_ptr \ ?bs k e' \ sound_prered_ptr ?bs k e'" using Cons.prems(3,4,7) length_bin_upd nth_item_bin_upd sound_pre_ptr_def sound_prered_ptr_def by (smt (verit, ccfv_threshold) list.set_intros(2) nth_list_update order_less_le_trans) have 5: "\e' \ set es. \k' pre red ps. pointer e' = PreRed (k', pre, red) ps \ red < length (bin_upd e b)" by (meson Cons.prems(8) length_bin_upd order_less_le_trans set_subset_Cons subsetD) have "sound_ptrs \ ((bs[k := bin_upd e b])[k := bin_upds es (bin_upd e b)]) \ mono_red_ptr (bs[k := bin_upd e b, k := bin_upds es (bin_upd e b)])" using Cons.IH[OF 0 1 2 3 _ _ 4 5] distinct_bin_upd Cons.prems(4,5,6) items_def by (metis distinct.simps(2) list.simps(9)) thus ?case by simp qed simp lemma sound_mono_ptrs_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" assumes "sound_ptrs \ bs" "\x \ bins bs. sound_item \ \ x" assumes "mono_red_ptr bs" assumes "nonempty_derives \" "wf_\ \" shows "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ bs i) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ bs i)" using assms proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have x: "x \ set (items (bs ! k))" using Complete\<^sub>F.hyps(1,2) by force hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x by (metis dual_order.refl) hence sound: "\x \ bins ?bs'. sound_item \ \ x" by (metis Complete\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim) have 0: "k < length bs" using Complete\<^sub>F.prems(1) wf_earley_input_elim by auto have 1: "\e \ set (Complete\<^sub>L k x bs i). sound_null_ptr e" unfolding Complete\<^sub>L_def sound_null_ptr_def by auto have 2: "\e \ set (Complete\<^sub>L k x bs i). sound_pre_ptr \ bs k e" unfolding Complete\<^sub>L_def sound_pre_ptr_def by auto { fix e assume a0: "e \ set (Complete\<^sub>L k x bs i)" fix p ps k' pre red assume a1: "pointer e = PreRed p ps" "(k', pre, red) \ set (p#ps)" have "k' = item_origin x" using a0 a1 unfolding Complete\<^sub>L_def by auto moreover have "wf_item \ \ x" "item_end x = k" using Complete\<^sub>F.prems(1) x wf_earley_input_elim wf_bins_kth_bin by blast+ ultimately have 0: "k' \ k" using wf_item_def by blast have 1: "k' \ k" proof (rule ccontr) assume "\ k' \ k" have "sound_item \ \ x" using Complete\<^sub>F.prems(1,3) x kth_bin_sub_bins wf_earley_input_elim by (metis subset_eq) moreover have "is_complete x" using Complete\<^sub>F.hyps(3) by (auto simp: next_symbol_def split: if_splits) moreover have "item_origin x = k" using \\ k' \ k\ \k' = item_origin x\ by auto ultimately show False using impossible_complete_item Complete\<^sub>F.prems(1,5) wf_earley_input_elim \item_end x = k\ \wf_item \ \ x\ by blast qed have 2: "pre < length (bs!k')" using a0 a1 index_filter_with_index_lt_length unfolding Complete\<^sub>L_def by (auto simp: items_def; fastforce) have 3: "red < i+1" using a0 a1 unfolding Complete\<^sub>L_def by auto have "item e = inc_item (item (bs!k'!pre)) k" using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def by (auto simp: items_def, metis filter_with_index_nth nth_map) moreover have "is_complete (item (bs!k!red))" using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def by (auto simp: next_symbol_def items_def split: if_splits) moreover have "item_origin (item (bs!k!red)) = item_end (item (bs!k'!pre))" using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def apply (clarsimp simp: items_def) by (metis dual_order.strict_trans index_filter_with_index_lt_length items_def le_neq_implies_less nth_map nth_mem wf_bins_kth_bin wf_earley_input_elim) moreover have "(\N. next_symbol (item (bs ! k' ! pre)) = Some N \ N = item_rule_head (item (bs ! k ! red)))" using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def by (auto simp: items_def, metis (mono_tags, lifting) filter_with_index_P filter_with_index_nth nth_map) ultimately have 4: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" unfolding completes_def by blast have "k' < k" "pre < length (bs!k')" "red < i+1" "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" using 0 1 2 3 4 by simp_all } hence "\e \ set (Complete\<^sub>L k x bs i). \p ps k' pre red. pointer e = PreRed p ps \ (k', pre, red) \ set (p#ps) \ k' < k \ pre < length (bs!k') \ red < i+1 \ completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" by force hence 3: "\e \ set (Complete\<^sub>L k x bs i). sound_prered_ptr bs k e" - unfolding sound_prered_ptr_def using Complete\<^sub>F.hyps(1) items_def by (smt (verit) discrete dual_order.strict_trans1 leI length_map) + using Complete\<^sub>F.hyps(1) + by (simp add: not_le length_items_eq sound_prered_ptr_def) fastforce have 4: "distinct (items (Complete\<^sub>L k x bs i))" using distinct_Complete\<^sub>L x Complete\<^sub>F.prems(1) wf_earley_input_elim wf_bin_def wf_bin_items_def wf_bins_def wf_item_def by (metis order_le_less_trans) have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" using sound_mono_ptrs_bin_upds[OF Complete\<^sub>F.prems(2) Complete\<^sub>F.prems(4) 0] 1 2 3 4 sound_prered_ptr_def Complete\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def by (smt (verit, ccfv_SIG) list.set_intros(1)) moreover have "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Complete\<^sub>F.IH Complete\<^sub>F.prems(4-6) sound by blast thus ?case using Complete\<^sub>F.hyps by simp next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "x \ set (items (bs ! k))" using Scan\<^sub>F.hyps(1,2) by force hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items wf_bins_impl_wf_items by fast hence sound: "\x \ bins ?bs'. sound_item \ \ x" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,3) bins_bins_upd wf_earley_input_elim by (metis UnE add_less_cancel_right) have 0: "k+1 < length bs" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim by force have 1: "\e \ set (Scan\<^sub>L k \ a x i). sound_null_ptr e" unfolding Scan\<^sub>L_def sound_null_ptr_def by auto have 2: "\e \ set (Scan\<^sub>L k \ a x i). sound_pre_ptr \ bs (k+1) e" using Scan\<^sub>F.hyps(1,2,3) unfolding sound_pre_ptr_def Scan\<^sub>L_def scans_def items_def by auto have 3: "\e \ set (Scan\<^sub>L k \ a x i). sound_prered_ptr bs (k+1) e" unfolding Scan\<^sub>L_def sound_prered_ptr_def by simp have 4: "distinct (items (Scan\<^sub>L k \ a x i))" using distinct_Scan\<^sub>L by fast have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" using sound_mono_ptrs_bin_upds[OF Scan\<^sub>F.prems(2) Scan\<^sub>F.prems(4) 0] 0 1 2 3 4 sound_prered_ptr_def Scan\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def by (smt (verit, ccfv_threshold) list.set_intros(1)) moreover have "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Scan\<^sub>F.IH Scan\<^sub>F.prems(4-6) sound by blast thus ?case using Scan\<^sub>F.hyps by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "x \ set (items (bs ! k))" using Predict\<^sub>F.hyps(1,2) by force hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast hence sound: "\x \ bins ?bs'. sound_item \ \ x" using Predict\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim by metis have 0: "k < length bs" using Predict\<^sub>F.prems(1) wf_earley_input_elim by force have 1: "\e \ set (Predict\<^sub>L k \ a). sound_null_ptr e" unfolding sound_null_ptr_def Predict\<^sub>L_def predicts_def by (auto simp: init_item_def) have 2: "\e \ set (Predict\<^sub>L k \ a). sound_pre_ptr \ bs k e" unfolding sound_pre_ptr_def Predict\<^sub>L_def by simp have 3: "\e \ set (Predict\<^sub>L k \ a). sound_prered_ptr bs k e" unfolding sound_prered_ptr_def Predict\<^sub>L_def by simp have 4: "distinct (items (Predict\<^sub>L k \ a))" using Predict\<^sub>F.prems(6) distinct_Predict\<^sub>L by fast have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" using sound_mono_ptrs_bin_upds[OF Predict\<^sub>F.prems(2) Predict\<^sub>F.prems(4) 0] 0 1 2 3 4 sound_prered_ptr_def Predict\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def by (smt (verit, ccfv_threshold) list.set_intros(1)) moreover have "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Predict\<^sub>F.IH Predict\<^sub>F.prems(4-6) sound by blast thus ?case using Predict\<^sub>F.hyps by simp qed simp_all lemma sound_mono_ptrs_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "sound_ptrs \ bs" "\x \ bins bs. sound_item \ \ x" assumes "mono_red_ptr bs" assumes "nonempty_derives \" "wf_\ \" shows "sound_ptrs \ (Earley\<^sub>L_bin k \ \ bs) \ mono_red_ptr (Earley\<^sub>L_bin k \ \ bs)" using assms sound_mono_ptrs_Earley\<^sub>L_bin' Earley\<^sub>L_bin_def by metis lemma sound_ptrs_Init\<^sub>L: "sound_ptrs \ (Init\<^sub>L \ \)" unfolding sound_ptrs_def sound_null_ptr_def sound_pre_ptr_def sound_prered_ptr_def predicts_def scans_def completes_def Init\<^sub>L_def by (auto simp: init_item_def less_Suc_eq_0_disj) lemma mono_red_ptr_Init\<^sub>L: "mono_red_ptr (Init\<^sub>L \ \)" unfolding mono_red_ptr_def Init\<^sub>L_def by (auto simp: init_item_def less_Suc_eq_0_disj) lemma sound_mono_ptrs_Earley\<^sub>L_bins: assumes "k \ length \" "wf_\ \" "nonempty_derives \" "wf_\ \" shows "sound_ptrs \ (Earley\<^sub>L_bins k \ \) \ mono_red_ptr (Earley\<^sub>L_bins k \ \)" using assms proof (induction k) case 0 have "(0, \, \, (Init\<^sub>L \ \)) \ wf_earley_input" using assms(2) wf_earley_input_Init\<^sub>L by blast moreover have "\x \ bins (Init\<^sub>L \ \). sound_item \ \ x" by (metis Init\<^sub>L_eq_Init\<^sub>F Init\<^sub>F_sub_Earley sound_Earley subsetD wf_Earley) ultimately show ?case using sound_mono_ptrs_Earley\<^sub>L_bin sound_ptrs_Init\<^sub>L mono_red_ptr_Init\<^sub>L "0.prems"(2,3) by fastforce next case (Suc k) have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) moreover have "sound_ptrs \ (Earley\<^sub>L_bins k \ \)" using Suc by simp moreover have "\x \ bins (Earley\<^sub>L_bins k \ \). sound_item \ \ x" by (meson Suc.prems(1) Suc_leD Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins Earley\<^sub>F_bins_sub_Earley assms(2) sound_Earley subsetD wf_bins_Earley\<^sub>L_bins wf_bins_impl_wf_items) ultimately show ?case using Suc.prems(1,3,4) sound_mono_ptrs_Earley\<^sub>L_bin Suc.IH by fastforce qed lemma sound_mono_ptrs_Earley\<^sub>L: assumes "wf_\ \" "nonempty_derives \" shows "sound_ptrs \ (Earley\<^sub>L \ \) \ mono_red_ptr (Earley\<^sub>L \ \)" using assms sound_mono_ptrs_Earley\<^sub>L_bins Earley\<^sub>L_def by (metis dual_order.refl) subsection \Common Definitions\ datatype 'a tree = Leaf 'a | Branch 'a "'a tree list" fun yield_tree :: "'a tree \ 'a sentence" where "yield_tree (Leaf a) = [a]" | "yield_tree (Branch _ ts) = concat (map yield_tree ts)" fun root_tree :: "'a tree \ 'a" where "root_tree (Leaf a) = a" | "root_tree (Branch N _) = N" fun wf_rule_tree :: "'a cfg \ 'a tree \ bool" where "wf_rule_tree _ (Leaf a) \ True" | "wf_rule_tree \ (Branch N ts) \ ( (\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r) \ (\t \ set ts. wf_rule_tree \ t))" fun wf_item_tree :: "'a cfg \ 'a item \ 'a tree \ bool" where "wf_item_tree \ _ (Leaf a) \ True" | "wf_item_tree \ x (Branch N ts) \ ( N = item_rule_head x \ map root_tree ts = take (item_dot x) (item_rule_body x) \ (\t \ set ts. wf_rule_tree \ t))" definition wf_yield_tree :: "'a sentence \ 'a item \ 'a tree \ bool" where "wf_yield_tree \ x t \ yield_tree t = slice (item_origin x) (item_end x) \" datatype 'a forest = FLeaf 'a | FBranch 'a "'a forest list list" fun combinations :: "'a list list \ 'a list list" where "combinations [] = [[]]" | "combinations (xs#xss) = [ x#cs . x <- xs, cs <- combinations xss ]" fun trees :: "'a forest \ 'a tree list" where "trees (FLeaf a) = [Leaf a]" | "trees (FBranch N fss) = ( let tss = (map (\fs. concat (map (\f. trees f) fs)) fss) in map (\ts. Branch N ts) (combinations tss) )" lemma list_comp_flatten: "[ f xs . xs <- [ g xs ys . xs <- as, ys <- bs ] ] = [ f (g xs ys) . xs <- as, ys <- bs ]" by (induction as) auto lemma list_comp_flatten_Cons: "[ x#xs . x <- as, xs <- [ xs @ ys. xs <- bs, ys <- cs ] ] = [ x#xs@ys. x <- as, xs <- bs, ys <- cs ]" by (induction as) (auto simp: list_comp_flatten) lemma list_comp_flatten_append: "[ xs@ys . xs <- [ x#xs . x <- as, xs <- bs ], ys <- cs ] = [ x#xs@ys . x <- as, xs <- bs, ys <- cs ]" by (induction as) (auto simp: o_def, meson append_Cons map_eq_conv) lemma combinations_append: "combinations (xss @ yss) = [ xs @ ys . xs <- combinations xss, ys <- combinations yss ]" by (induction xss) (auto simp: list_comp_flatten_Cons list_comp_flatten_append map_idI) lemma trees_append: "trees (FBranch N (xss @ yss)) = ( let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in let ytss = (map (\ys. concat (map (\f. trees f) ys)) yss) in map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- combinations ytss ])" using combinations_append by (metis map_append trees.simps(2)) lemma trees_append_singleton: "trees (FBranch N (xss @ [ys])) = ( let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in let ytss = [concat (map trees ys)] in map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- combinations ytss ])" by (subst trees_append, simp) lemma trees_append_single_singleton: "trees (FBranch N (xss @ [[y]])) = ( let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- [ [t] . t <- trees y ] ])" by (subst trees_append_singleton, auto) subsection \foldl lemmas\ lemma foldl_add_nth: "k < length xs \ foldl (+) z (map length (take k xs)) + length (xs!k) = foldl (+) z (map length (take (k+1) xs))" proof (induction xs arbitrary: k z) case (Cons x xs) then show ?case proof (cases "k = 0") case False thus ?thesis using Cons by (auto simp add: take_Cons') qed simp qed simp lemma foldl_acc_mono: "a \ b \ foldl (+) a xs \ foldl (+) b xs" for a :: nat by (induction xs arbitrary: a b) auto lemma foldl_ge_z_nth: "j < length xs \ z + length (xs!j) \ foldl (+) z (map length (take (j+1) xs))" proof (induction xs arbitrary: j z) case (Cons x xs) show ?case proof (cases "j = 0") case False have "z + length ((x # xs) ! j) = z + length (xs!(j-1))" using False by simp also have "... \ foldl (+) z (map length (take (j-1+1) xs))" using Cons False by (metis add_diff_inverse_nat length_Cons less_one nat_add_left_cancel_less plus_1_eq_Suc) also have "... = foldl (+) z (map length (take j xs))" using False by simp also have "... \ foldl (+) (z + length x) (map length (take j xs))" using foldl_acc_mono by force also have "... = foldl (+) z (map length (take (j+1) (x#xs)))" by simp finally show ?thesis by blast qed simp qed simp lemma foldl_add_nth_ge: "i \ j \ j < length xs \ foldl (+) z (map length (take i xs)) + length (xs!j) \ foldl (+) z (map length (take (j+1) xs))" proof (induction xs arbitrary: i j z) case (Cons x xs) show ?case proof (cases "i = 0") case True have "foldl (+) z (map length (take i (x # xs))) + length ((x # xs) ! j) = z + length ((x # xs) ! j)" using True by simp also have "... \ foldl (+) z (map length (take (j+1) (x#xs)))" using foldl_ge_z_nth Cons.prems(2) by blast finally show ?thesis by blast next case False have "i-1 \ j-1" by (simp add: Cons.prems(1) diff_le_mono) have "j-1 < length xs" using Cons.prems(1,2) False by fastforce have "foldl (+) z (map length (take i (x # xs))) + length ((x # xs) ! j) = foldl (+) (z + length x) (map length (take (i-1) xs)) + length ((x#xs)!j)" using False by (simp add: take_Cons') also have "... = foldl (+) (z + length x) (map length (take (i-1) xs)) + length (xs!(j-1))" using Cons.prems(1) False by auto also have "... \ foldl (+) (z + length x) (map length (take (j-1+1) xs))" using Cons.IH \i - 1 \ j - 1\ \j - 1 < length xs\ by blast also have "... = foldl (+) (z + length x) (map length (take j xs))" using Cons.prems(1) False by fastforce also have "... = foldl (+) z (map length (take (j+1) (x#xs)))" by fastforce finally show ?thesis by blast qed qed simp lemma foldl_ge_acc: "foldl (+) z (map length xs) \ z" by (induction xs arbitrary: z) (auto elim: add_leE) lemma foldl_take_mono: "i \ j \ foldl (+) z (map length (take i xs)) \ foldl (+) z (map length (take j xs))" proof (induction xs arbitrary: z i j) case (Cons x xs) show ?case proof (cases "i = 0") case True have "foldl (+) z (map length (take i (x # xs))) = z" using True by simp also have "... \ foldl (+) z (map length (take j (x # xs)))" by (simp add: foldl_ge_acc) ultimately show ?thesis by simp next case False then show ?thesis using Cons by (simp add: take_Cons') qed qed simp subsection \Parse tree\ partial_function (option) build_tree' :: "'a bins \ 'a sentence \ nat \ nat \ 'a tree option" where "build_tree' bs \ k i = ( let e = bs!k!i in ( case pointer e of Null \ Some (Branch (item_rule_head (item e)) []) \\start building sub-tree\ | Pre pre \ ( \\add sub-tree starting from terminal\ do { t \ build_tree' bs \ (k-1) pre; case t of Branch N ts \ Some (Branch N (ts @ [Leaf (\!(k-1))])) | _ \ undefined \\impossible case\ }) | PreRed (k', pre, red) _ \ ( \\add sub-tree starting from non-terminal\ do { t \ build_tree' bs \ k' pre; case t of Branch N ts \ do { t \ build_tree' bs \ k red; Some (Branch N (ts @ [t])) } | _ \ undefined \\impossible case\ }) ))" declare build_tree'.simps [code] definition build_tree :: "'a cfg \ 'a sentence \ 'a bins \ 'a tree option" where "build_tree \ \ bs = ( let k = length bs - 1 in ( case filter_with_index (\x. is_finished \ \ x) (items (bs!k)) of [] \ None | (_, i)#_ \ build_tree' bs \ k i ))" lemma build_tree'_simps[simp]: "e = bs!k!i \ pointer e = Null \ build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = None \ build_tree' bs \ k i = None" "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = Some (Branch N ts) \ build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = Some (Leaf a) \ build_tree' bs \ k i = undefined" "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = None \ build_tree' bs \ k i = None" "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Branch N ts) \ build_tree' bs \ k red = None \ build_tree' bs \ k i = None" "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Leaf a) \ build_tree' bs \ k i = undefined" "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Branch N ts) \ build_tree' bs \ k red = Some t \ build_tree' bs \ k i = Some (Branch N (ts @ [t]))" by (subst build_tree'.simps, simp)+ definition wf_tree_input :: "('a bins \ 'a sentence \ nat \ nat) set" where "wf_tree_input = { (bs, \, k, i) | bs \ k i. sound_ptrs \ bs \ mono_red_ptr bs \ k < length bs \ i < length (bs!k) }" fun build_tree'_measure :: "('a bins \ 'a sentence \ nat \ nat) \ nat" where "build_tree'_measure (bs, \, k, i) = foldl (+) 0 (map length (take k bs)) + i" lemma wf_tree_input_pre: assumes "(bs, \, k, i) \ wf_tree_input" assumes "e = bs!k!i" "pointer e = Pre pre" shows "(bs, \, (k-1), pre) \ wf_tree_input" using assms unfolding wf_tree_input_def using less_imp_diff_less nth_mem by (fastforce simp: sound_ptrs_def sound_pre_ptr_def) lemma wf_tree_input_prered_pre: assumes "(bs, \, k, i) \ wf_tree_input" assumes "e = bs!k!i" "pointer e = PreRed (k', pre, red) ps" shows "(bs, \, k', pre) \ wf_tree_input" using assms unfolding wf_tree_input_def apply (auto simp: sound_ptrs_def sound_prered_ptr_def) apply metis+ apply (metis dual_order.strict_trans nth_mem) by (metis nth_mem) lemma wf_tree_input_prered_red: assumes "(bs, \, k, i) \ wf_tree_input" assumes "e = bs!k!i" "pointer e = PreRed (k', pre, red) ps" shows "(bs, \, k, red) \ wf_tree_input" using assms unfolding wf_tree_input_def apply (auto simp add: sound_ptrs_def sound_prered_ptr_def) apply (metis nth_mem)+ done lemma build_tree'_induct: assumes "(bs, \, k, i) \ wf_tree_input" assumes "\bs \ k i. (\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre) \ (\e k' pre red ps. e = bs!k!i \ pointer e = PreRed (k', pre, red) ps \ P bs \ k' pre) \ (\e k' pre red ps. e = bs!k!i \ pointer e = PreRed (k', pre, red) ps \ P bs \ k red) \ P bs \ k i" shows "P bs \ k i" using assms(1) proof (induction n\"build_tree'_measure (bs, \, k, i)" arbitrary: k i rule: nat_less_induct) case 1 obtain e where entry: "e = bs!k!i" by simp consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" by (metis pointer.exhaust surj_pair) thus ?case proof cases case Null thus ?thesis using assms(2) entry by fastforce next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast define n where n: "n = build_tree'_measure (bs, \, (k-1), pre)" have "0 < k" "pre < length (bs!(k-1))" using 1(2) entry pre unfolding wf_tree_input_def sound_ptrs_def sound_pre_ptr_def by (smt (verit) mem_Collect_eq nth_mem prod.inject)+ have "k < length bs" using 1(2) unfolding wf_tree_input_def by blast+ have "foldl (+) 0 (map length (take k bs)) + i - (foldl (+) 0 (map length (take (k-1) bs)) + pre) = foldl (+) 0 (map length (take (k-1) bs)) + length (bs!(k-1)) + i - (foldl (+) 0 (map length (take (k-1) bs)) + pre)" using foldl_add_nth[of \k-1\ bs 0] by (simp add: \0 < k\ \k < length bs\ less_imp_diff_less) also have "... = length (bs!(k-1)) + i - pre" by simp also have "... > 0" using \pre < length (bs!(k-1))\ by auto finally have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, (k-1), pre) > 0" by simp hence "P bs \ (k-1) pre" using 1 n wf_tree_input_pre entry pre zero_less_diff by blast thus ?thesis using assms(2) entry pre pointer.distinct(5) pointer.inject(1) by presburger next case PreRed then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" by blast have "k' < k" "pre < length (bs!k')" using 1(2) entry prered unfolding wf_tree_input_def sound_ptrs_def sound_prered_ptr_def apply simp_all apply (metis nth_mem)+ done have "red < i" using 1(2) entry prered unfolding wf_tree_input_def mono_red_ptr_def by blast have "k < length bs" "i < length (bs!k)" using 1(2) unfolding wf_tree_input_def by blast+ define n_pre where n_pre: "n_pre = build_tree'_measure (bs, \, k', pre)" have "0 < length (bs!k') + i - pre" by (simp add: \pre < length (bs!k')\ add.commute trans_less_add2) also have "... = foldl (+) 0 (map length (take k' bs)) + length (bs!k') + i - (foldl (+) 0 (map length (take k' bs)) + pre)" by simp also have "... \ foldl (+) 0 (map length (take (k'+1) bs)) + i - (foldl (+) 0 (map length (take k' bs)) + pre)" using foldl_add_nth_ge[of k' k' bs 0] \k < length bs\ \k' < k\ by simp also have "... \ foldl (+) 0 (map length (take k bs)) + i - (foldl (+) 0 (map length (take k' bs)) + pre)" using foldl_take_mono by (metis Suc_eq_plus1 Suc_leI \k' < k\ add.commute add_le_cancel_left diff_le_mono) finally have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, k', pre) > 0" by simp hence x: "P bs \ k' pre" using 1(1) zero_less_diff by (metis "1.prems" entry prered wf_tree_input_prered_pre) define n_red where n_red: "n_red = build_tree'_measure (bs, \, k, red)" have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, k, red) > 0" using \red < i\ by simp hence y: "P bs \ k red" using "1.hyps" "1.prems" entry prered wf_tree_input_prered_red zero_less_diff by blast show ?thesis using assms(2) x y entry prered by (smt (verit, best) Pair_inject filter_cong pointer.distinct(5) pointer.inject(2)) qed qed lemma build_tree'_termination: assumes "(bs, \, k, i) \ wf_tree_input" shows "\N ts. build_tree' bs \ k i = Some (Branch N ts)" proof - have "\N ts. build_tree' bs \ k i = Some (Branch N ts)" apply (induction rule: build_tree'_induct[OF assms(1)]) subgoal premises IH for bs \ k i proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\k' pre red ps. pointer e = PreRed (k', pre, red) ps" by (metis pointer.exhaust surj_pair) thus ?thesis proof cases case Null thus ?thesis using build_tree'_simps(1) entry by simp next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" using IH(1) entry pre by blast have "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" using build_tree'_simps(3) entry pre Nts by simp thus ?thesis by simp next case PreRed then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" by blast then obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" using IH(2) entry prered by blast obtain t_red where t_red: "build_tree' bs \ k red = Some t_red" using IH(3) entry prered Nts by (metis option.exhaust) have "build_tree' bs \ k i = Some (Branch N (ts @ [t_red]))" using build_tree'_simps(8) entry prered Nts t_red by auto thus ?thesis by blast qed qed done thus ?thesis by blast qed lemma wf_item_tree_build_tree': assumes "(bs, \, k, i) \ wf_tree_input" assumes "wf_bins \ \ bs" assumes "k < length bs" "i < length (bs!k)" assumes "build_tree' bs \ k i = Some t" shows "wf_item_tree \ (item (bs!k!i)) t" proof - have "wf_item_tree \ (item (bs!k!i)) t" using assms apply (induction arbitrary: t rule: build_tree'_induct[OF assms(1)]) subgoal premises prems for bs \ k i t proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\k' pre red ps. pointer e = PreRed (k', pre, red) ps" by (metis pointer.exhaust surj_pair) thus ?thesis proof cases case Null hence "build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" using entry by simp have simp: "t = Branch (item_rule_head (item e)) []" using build_tree'_simps(1) Null prems(8) entry by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_tree_input_def by blast hence "predicts (item e)" using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast hence "item_dot (item e) = 0" unfolding predicts_def by blast thus ?thesis using simp entry by simp next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" using build_tree'_termination entry pre prems(4) wf_tree_input_pre by blast have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" using build_tree'_simps(3) entry pre Nts by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_tree_input_def by blast hence "pre < length (bs!(k-1))" using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) moreover have "k-1 < length bs" by (simp add: prems(6) less_imp_diff_less) ultimately have IH: "wf_item_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" using prems(1,2,4,5) entry pre Nts by (meson wf_tree_input_pre) have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp hence *: "item_rule_head (item (bs!(k-1)!pre)) = item_rule_head (item e)" "item_rule_body (item (bs!(k-1)!pre)) = item_rule_body (item e)" "item_dot (item (bs!(k-1)!pre)) + 1 = item_dot (item e)" "next_symbol (item (bs!(k-1)!pre)) = Some (\!(k-1))" unfolding scans_def inc_item_def by (simp_all add: item_rule_head_def item_rule_body_def) have "map root_tree (ts @ [Leaf (\!(k-1))]) = map root_tree ts @ [\!(k-1)]" by simp also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item (bs!(k-1)!pre))) @ [\!(k-1)]" using IH by simp also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item e)) @ [\!(k-1)]" using *(2) by simp also have "... = take (item_dot (item e)) (item_rule_body (item e))" using *(2-4) by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) finally have "map root_tree (ts @ [Leaf (\!(k-1))]) = take (item_dot (item e)) (item_rule_body (item e))" . hence "wf_item_tree \ (item e) (Branch N (ts @ [Leaf (\!(k-1))]))" using IH *(1) by simp thus ?thesis using entry prems(8) simp by auto next case PreRed then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" by blast obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" using build_tree'_termination entry prems(4) prered wf_tree_input_prered_pre by blast obtain N_red ts_red where Nts_red: "build_tree' bs \ k red = Some (Branch N_red ts_red)" using build_tree'_termination entry prems(4) prered wf_tree_input_prered_red by blast have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Branch N_red ts_red]))" using build_tree'_simps(8) entry prered Nts Nts_red by auto have "sound_ptrs \ bs" using prems(4) wf_tree_input_def by fastforce have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_prered_ptr_def sound_ptrs_def by fastforce+ have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by fastforce have *: "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" "is_complete (item (bs!k!red))" using completes unfolding completes_def inc_item_def by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) have "(bs, \, k', pre) \ wf_tree_input" using wf_tree_input_prered_pre[OF prems(4) entry prered] by blast hence IH_pre: "wf_item_tree \ (item (bs!k'!pre)) (Branch N ts)" using prems(2)[OF entry prered _ prems(5)] Nts bounds(1,2) order_less_trans prems(6) by blast have "(bs, \, k, red) \ wf_tree_input" using wf_tree_input_prered_red[OF prems(4) entry prered] by blast hence IH_r: "wf_item_tree \ (item (bs!k!red)) (Branch N_red ts_red)" using bounds(3) entry prems(3,5,6) prered Nts_red by blast have "map root_tree (ts @ [Branch N_red ts_red]) = map root_tree ts @ [root_tree (Branch N_red ts)]" by simp also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [root_tree (Branch N_red ts)]" using IH_pre by simp also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [item_rule_head (item (bs!k!red))]" using IH_r by simp also have "... = take (item_dot (item e)) (item_rule_body (item e))" using * by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) finally have roots: "map root_tree (ts @ [Branch N_red ts]) = take (item_dot (item e)) (item_rule_body (item e))" by simp have "wf_item \ \ (item (bs!k!red))" using prems(5,6) bounds(3) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (auto simp: items_def) moreover have "N_red = item_rule_head (item (bs!k!red))" using IH_r by fastforce moreover have "map root_tree ts_red = item_rule_body (item (bs!k!red))" using IH_r *(5) by (auto simp: is_complete_def) ultimately have "\r \ set (\ \). N_red = rule_head r \ map root_tree ts_red = rule_body r" unfolding wf_item_def item_rule_body_def item_rule_head_def by blast hence "wf_rule_tree \ (Branch N_red ts_red)" using IH_r by simp hence "wf_item_tree \ (item (bs!k!i)) (Branch N (ts @ [Branch N_red ts_red]))" using "*"(1) roots IH_pre entry by simp thus ?thesis using Nts_red prems(8) simp by auto qed qed done thus ?thesis using assms(2) by blast qed lemma wf_yield_tree_build_tree': assumes "(bs, \, k, i) \ wf_tree_input" assumes "wf_bins \ \ bs" assumes "k < length bs" "i < length (bs!k)" "k \ length \" assumes "build_tree' bs \ k i = Some t" shows "wf_yield_tree \ (item (bs!k!i)) t" proof - have "wf_yield_tree \ (item (bs!k!i)) t" using assms apply (induction arbitrary: t rule: build_tree'_induct[OF assms(1)]) subgoal premises prems for bs \ k i t proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" by (metis pointer.exhaust surj_pair) thus ?thesis proof cases case Null hence "build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" using entry by simp have simp: "t = Branch (item_rule_head (item e)) []" using build_tree'_simps(1) Null prems(9) entry by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_tree_input_def by blast hence "predicts (item e)" using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast thus ?thesis unfolding wf_yield_tree_def predicts_def using simp entry by (auto simp: slice_empty) next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" using build_tree'_termination entry pre prems(4) wf_tree_input_pre by blast have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" using build_tree'_simps(3) entry pre Nts by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_tree_input_def by blast hence bounds: "k > 0" "pre < length (bs!(k-1))" using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem)+ moreover have "k-1 < length bs" by (simp add: prems(6) less_imp_diff_less) ultimately have IH: "wf_yield_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" using prems(1) entry pre Nts wf_tree_input_pre prems(4,5,8) by fastforce have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp have wf: "item_origin (item (bs!(k-1)!pre)) \ item_end (item (bs!(k-1)!pre))" "item_end (item (bs!(k-1)!pre)) = k-1" "item_end (item e) = k" using entry prems(5,6,7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def by (auto, meson less_imp_diff_less nth_mem) have "yield_tree (Branch N (ts @ [Leaf (\!(k-1))])) = concat (map yield_tree (ts @ [Leaf (\!(k-1))]))" by simp also have "... = concat (map yield_tree ts) @ [\!(k-1)]" by simp also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre))) \ @ [\!(k-1)]" using IH by (simp add: wf_yield_tree_def) also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre)) + 1) \" using slice_append_nth wf \k > 0\ prems(8) by (metis diff_less le_eq_less_or_eq less_imp_diff_less less_numeral_extra(1)) also have "... = slice (item_origin (item e)) (item_end (item (bs!(k-1)!pre)) + 1) \" using scans unfolding scans_def inc_item_def by simp also have "... = slice (item_origin (item e)) k \" using scans wf unfolding scans_def by (metis Suc_diff_1 Suc_eq_plus1 bounds(1)) also have "... = slice (item_origin (item e)) (item_end (item e)) \" using wf by auto finally show ?thesis using wf_yield_tree_def entry prems(9) simp by force next case PreRed then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" by blast obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" using build_tree'_termination entry prems(4) prered wf_tree_input_prered_pre by blast obtain N_red ts_red where Nts_red: "build_tree' bs \ k red = Some (Branch N_red ts_red)" using build_tree'_termination entry prems(4) prered wf_tree_input_prered_red by blast have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Branch N_red ts_red]))" using build_tree'_simps(8) entry prered Nts Nts_red by auto have "sound_ptrs \ bs" using prems(4) wf_tree_input_def by fastforce have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by fastforce+ have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by fastforce have "(bs, \, k', pre) \ wf_tree_input" using wf_tree_input_prered_pre[OF prems(4) entry prered] by blast hence IH_pre: "wf_yield_tree \ (item (bs!k'!pre)) (Branch N ts)" using prems(2)[OF entry prered _ prems(5)] Nts bounds(1,2) prems(6-8) by (meson dual_order.strict_trans1 nat_less_le) have "(bs, \, k, red) \ wf_tree_input" using wf_tree_input_prered_red[OF prems(4) entry prered] by blast hence IH_r: "wf_yield_tree \ (item (bs!k!red)) (Branch N_red ts_red)" using bounds(3) entry prems(3,5,6,8) prered Nts_red by blast have wf1: "item_origin (item (bs!k'!pre)) \ item_end (item (bs!k'!pre))" "item_origin (item (bs!k!red)) \ item_end (item (bs!k!red))" using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def by (metis length_map nth_map nth_mem order_less_trans)+ have wf2: "item_end (item (bs!k!red)) = k" "item_end (item (bs!k!i)) = k" using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def by simp_all have "yield_tree (Branch N (ts @ [Branch N_red ts_red])) = concat (map yield_tree (ts @ [Branch N_red ts_red]))" by (simp add: Nts_red) also have "... = concat (map yield_tree ts) @ yield_tree (Branch N_red ts_red)" by simp also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k'!pre))) \ @ slice (item_origin (item (bs!k!red))) (item_end (item (bs!k!red))) \" using IH_pre IH_r by (simp add: wf_yield_tree_def) also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k!red))) \" using slice_concat wf1 completes_def completes by (metis (no_types, lifting)) also have "... = slice (item_origin (item e)) (item_end (item (bs!k!red))) \" using completes unfolding completes_def inc_item_def by simp also have "... = slice (item_origin (item e)) (item_end (item e)) \" using wf2 entry by presburger finally show ?thesis using wf_yield_tree_def entry prems(9) simp by force qed qed done thus ?thesis using assms(2) by blast qed theorem wf_rule_root_yield_tree_build_forest: assumes "wf_bins \ \ bs" "sound_ptrs \ bs" "mono_red_ptr bs" "length bs = length \ + 1" assumes "build_tree \ \ bs = Some t" shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" proof - let ?k = "length bs - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items (bs!?k))" then obtain x i where *: "(x,i) \ set finished" "Some t = build_tree' bs \ ?k i" using assms(5) unfolding finished_def build_tree_def by (auto simp: Let_def split: list.splits, presburger) have k: "?k < length bs" "?k \ length \" using assms(4) by simp_all have i: "i < length (bs!?k)" using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) have x: "x = item (bs!?k!i)" using * i filter_with_index_nth items_def nth_map finished_def by metis have finished: "is_finished \ \ x" using * filter_with_index_P finished_def by metis have wf_trees_input: "(bs, \, ?k, i) \ wf_tree_input" unfolding wf_tree_input_def using assms(2,3) i k(1) by blast hence wf_item_tree: "wf_item_tree \ x t" using wf_item_tree_build_tree' assms(1,2) i k(1) x *(2) by metis have wf_item: "wf_item \ \ (item (bs!?k!i))" using k(1) i assms(1) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (simp add: items_def) obtain N ts where t: "t = Branch N ts" by (metis *(2) build_tree'_termination option.inject wf_trees_input) hence "N = item_rule_head x" "map root_tree ts = item_rule_body x" using finished wf_item_tree by (auto simp: is_finished_def is_complete_def) hence "\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r" using wf_item x unfolding wf_item_def item_rule_body_def item_rule_head_def by blast hence wf_rule: "wf_rule_tree \ t" using wf_item_tree t by simp have root: "root_tree t = \ \" using finished t \N = item_rule_head x\ by (auto simp: is_finished_def) have "yield_tree t = slice (item_origin (item (bs!?k!i))) (item_end (item (bs!?k!i))) \" using k i assms(1) wf_trees_input wf_yield_tree_build_tree' wf_yield_tree_def *(2) by (metis (no_types, lifting)) hence yield: "yield_tree t = \" using finished x unfolding is_finished_def by simp show ?thesis using * wf_rule root yield assms(4) unfolding build_tree_def by simp qed corollary wf_rule_root_yield_tree_build_tree_Earley\<^sub>L: assumes "wf_\ \" "nonempty_derives \" assumes "build_tree \ \ (Earley\<^sub>L \ \) = Some t" shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" using assms wf_rule_root_yield_tree_build_forest wf_bins_Earley\<^sub>L sound_mono_ptrs_Earley\<^sub>L Earley\<^sub>L_def length_Earley\<^sub>L_bins length_bins_Init\<^sub>L by (metis nle_le) theorem correctness_build_tree_Earley\<^sub>L: assumes "wf_\ \" "is_word \ \" "nonempty_derives \" shows "(\t. build_tree \ \ (Earley\<^sub>L \ \) = Some t) \ derives \ [\ \] \" (is "?L \ ?R") proof standard assume *: ?L let ?k = "length (Earley\<^sub>L \ \) - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" then obtain t x i where *: "(x,i) \ set finished" "Some t = build_tree' (Earley\<^sub>L \ \) \ ?k i" using * unfolding finished_def build_tree_def by (auto simp: Let_def split: list.splits, presburger) have k: "?k < length (Earley\<^sub>L \ \)" "?k \ length \" by (simp_all add: Earley\<^sub>L_def assms(1)) have i: "i < length ((Earley\<^sub>L \ \) ! ?k)" using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) have x: "x = item ((Earley\<^sub>L \ \)!?k!i)" using * i filter_with_index_nth items_def nth_map finished_def by metis have finished: "is_finished \ \ x" using * filter_with_index_P finished_def by metis moreover have "x \ set (items ((Earley\<^sub>L \ \) ! ?k))" using x by (auto simp: items_def; metis One_nat_def i imageI nth_mem) ultimately have "recognizing (bins (Earley\<^sub>L \ \)) \ \" by (meson k(1) kth_bin_sub_bins recognizing_def subsetD) thus ?R using correctness_Earley\<^sub>L assms by blast next assume *: ?R let ?k = "length (Earley\<^sub>L \ \) - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" have "recognizing (bins (Earley\<^sub>L \ \)) \ \" using assms * correctness_Earley\<^sub>L by blast moreover have "?k = length \" by (simp add: Earley\<^sub>L_def assms(1)) ultimately have "\x \ set (items ((Earley\<^sub>L \ \)!?k)). is_finished \ \ x" unfolding recognizing_def using assms(1) is_finished_def wf_bins_Earley\<^sub>L wf_item_in_kth_bin by metis then obtain x i xs where xis: "finished = (x,i)#xs" using filter_with_index_Ex_first by (metis finished_def) hence simp: "build_tree \ \ (Earley\<^sub>L \ \) = build_tree' (Earley\<^sub>L \ \) \ ?k i" unfolding build_tree_def finished_def by auto have "(x,i) \ set finished" using xis by simp hence "i < length ((Earley\<^sub>L \ \)!?k)" using index_filter_with_index_lt_length by (metis finished_def items_def length_map) moreover have "?k < length (Earley\<^sub>L \ \)" by (simp add: Earley\<^sub>L_def assms(1)) ultimately have "(Earley\<^sub>L \ \, \, ?k, i) \ wf_tree_input" unfolding wf_tree_input_def using sound_mono_ptrs_Earley\<^sub>L assms(1,3) by blast then obtain N ts where "build_tree' (Earley\<^sub>L \ \) \ ?k i = Some (Branch N ts)" using build_tree'_termination by blast thus ?L using simp by simp qed subsection \those, map, map option lemmas\ lemma those_map_exists: "Some ys = those (map f xs) \ y \ set ys \ \x. x \ set xs \ Some y \ set (map f xs)" proof (induction xs arbitrary: ys) case (Cons a xs) then show ?case apply (clarsimp split: option.splits) by (smt (verit, best) map_option_eq_Some set_ConsD) qed auto lemma those_Some: "(\x \ set xs. \a. x = Some a) \ (\ys. those xs = Some ys)" by (induction xs) (auto split: option.splits) lemma those_Some_P: assumes "\x \ set xs. \ys. x = Some ys \ (\y \ set ys. P y)" shows "\yss. those xs = Some yss \ (\ys \ set yss. \y \ set ys. P y)" using assms by (induction xs) auto lemma map_Some_P: assumes "z \ set (map f xs)" assumes "\x \ set xs. \ys. f x = Some ys \ (\y \ set ys. P y)" shows "\ys. z = Some ys \ (\y \ set ys. P y)" using assms by (induction xs) auto lemma those_map_FBranch_only: assumes "g = (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None)" assumes "Some fs = those (map g pres)" "f \ set fs" assumes "\f \ set pres. \N fss. f = FBranch N fss" shows "\f_pre N fss. f = FBranch N (fss @ [[FLeaf (\!(k-1))]]) \ f_pre = FBranch N fss \ f_pre \ set pres" using assms apply (induction pres arbitrary: fs f) apply (auto) by (smt (verit, best) list.inject list.set_cases map_option_eq_Some) lemma those_map_Some_concat_exists: assumes "y \ set (concat ys)" assumes "Some ys = those (map f xs)" shows "\ys x. Some ys = f x \ y \ set ys \ x \ set xs" using assms apply (induction xs arbitrary: ys y) apply (auto split: option.splits) by (smt (verit, ccfv_threshold) list.inject list.set_cases map_option_eq_Some) lemma map_option_concat_those_map_exists: assumes "Some fs = map_option concat (those (map F xs))" assumes "f \ set fs" shows "\fss fs'. Some fss = those (map F xs) \ fs' \ set fss \ f \ set fs'" using assms apply (induction xs arbitrary: fs f) apply (auto split: option.splits) by (smt (verit, best) UN_E map_option_eq_Some set_concat) lemma [partial_function_mono]: "monotone option.le_fun option_ord (\f. map_option concat (those (map (\((k', pre), reds). f ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. f ((((r, s), t), red), b \ {red})) reds) \ (\rss. those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None) pres)))) xs)))" proof - let ?f = " (\f. map_option concat (those (map (\((k', pre), reds). f ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. f ((((r, s), t), red), b \ {red})) reds) \ (\rss. those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None) pres)))) xs)))" have 0: "\x y. option.le_fun x y \ option_ord (?f x) (?f y)" apply (auto simp: flat_ord_def fun_ord_def option.leq_refl split: option.splits forest.splits) subgoal premises prems for x y proof - let ?t = "those (map (\((k', pre), reds). x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) xs) = None" show ?t proof (rule ccontr) assume a: "\?t" obtain fss where fss: "those (map (\((k', pre), reds). x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) xs) = Some fss" using a by blast { fix k' pre reds assume *: "((k', pre), reds) \ set xs" obtain pres where pres: "x ((((r, s), k'), pre), {pre}) = Some pres" using fss * those_Some by force have "\fs. Some fs = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" proof (rule ccontr) assume "\fs. Some fs = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" hence "None = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" by (smt (verit) not_None_eq) hence "None = x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" by (simp add: pres) hence "\((k', pre), reds) \ set xs. None = x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" using * by blast thus False using fss those_Some by force qed then obtain fs where fs: "Some fs = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" by blast obtain rss where rss: "those (map (\red. x ((((r, s), t), red), insert red b)) reds) = Some rss" using fs by force have "x ((((r, s), k'), pre), {pre}) = y ((((r, s), k'), pre), {pre})" using pres prems(1) by (metis option.distinct(1)) moreover have "those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)) = those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" proof - have "\red \ set reds. x ((((r, s), t), red), insert red b) = y ((((r, s), t), red), insert red b)" proof standard fix red assume "red \ set reds" have "\x\set (map (\red. x ((((r, s), t), red), insert red b)) reds) . \a. x = Some a" using rss those_Some by blast then obtain f where "x ((((r, s), t), red), insert red b) = Some f" using \red \ set reds\ by auto thus "x ((((r, s), t), red), insert red b) = y ((((r, s), t), red), insert red b)" using prems(1) by (metis option.distinct(1)) qed thus ?thesis by (smt (verit, best) map_eq_conv) qed ultimately have " x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))) = y ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" by (metis bind.bind_lunit pres) } hence "\((k', pre), reds) \ set xs. x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))) = y ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" by blast hence "those (map (\((k', pre), reds). x ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) xs) = those (map (\((k', pre), reds). y ((((r, s), k'), pre), {pre}) \ (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) xs)" using prems(1) by (smt (verit, best) case_prod_conv map_eq_conv split_cong) thus False using prems(2) by simp qed qed done show ?thesis using monotoneI[of option.le_fun option_ord ?f, OF 0] by blast qed subsection \Parse trees\ fun insert_group :: "('a \ 'k) \ ('a \ 'v) \ 'a \ ('k \ 'v list) list \ ('k \ 'v list) list" where "insert_group K V a [] = [(K a, [V a])]" | "insert_group K V a ((k, vs)#xs) = ( if K a = k then (k, V a # vs) # xs else (k, vs) # insert_group K V a xs )" fun group_by :: "('a \ 'k) \ ('a \ 'v) \ 'a list \ ('k \ 'v list) list" where "group_by K V [] = []" | "group_by K V (x#xs) = insert_group K V x (group_by K V xs)" lemma insert_group_cases: assumes "(k, vs) \ set (insert_group K V a xs)" shows "(k = K a \ vs = [V a]) \ (k, vs) \ set xs \ (\(k', vs') \ set xs. k' = k \ k = K a \ vs = V a # vs')" using assms by (induction xs) (auto split: if_splits) lemma group_by_exists_kv: "(k, vs) \ set (group_by K V xs) \ \x \ set xs. k = K x \ (\v \ set vs. v = V x)" using insert_group_cases by (induction xs) (simp, force) lemma group_by_forall_v_exists_k: "(k, vs) \ set (group_by K V xs) \ v \ set vs \ \x \ set xs. k = K x \ v = V x" proof (induction xs arbitrary: vs) case (Cons x xs) show ?case proof (cases "(k, vs) \ set (group_by K V xs)") case True thus ?thesis using Cons by simp next case False hence "(k, vs) \ set (insert_group K V x (group_by K V xs))" using Cons.prems(1) by force then consider (A) "(k = K x \ vs = [V x])" | (B) "(k, vs) \ set (group_by K V xs)" | (C) "(\(k', vs') \ set (group_by K V xs). k' = k \ k = K x \ vs = V x # vs')" using insert_group_cases by fastforce thus ?thesis proof cases case A thus ?thesis using Cons.prems(2) by auto next case B then show ?thesis using False by linarith next case C then show ?thesis using Cons.IH Cons.prems(2) by fastforce qed qed qed simp partial_function (option) build_trees' :: "'a bins \ 'a sentence \ nat \ nat \ nat set \ 'a forest list option" where "build_trees' bs \ k i I = ( let e = bs!k!i in ( case pointer e of Null \ Some ([FBranch (item_rule_head (item e)) []]) \\start building sub-trees\ | Pre pre \ ( \\add sub-trees starting from terminal\ do { pres \ build_trees' bs \ (k-1) pre {pre}; those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None \\impossible case\ ) pres) }) | PreRed p ps \ ( \\add sub-trees starting from non-terminal\ let ps' = filter (\(k', pre, red). red \ I) (p#ps) in let gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' in map_option concat (those (map (\((k', pre), reds). do { pres \ build_trees' bs \ k' pre {pre}; rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None \\impossible case\ ) pres) } ) gs)) ) ))" declare build_trees'.simps [code] definition build_trees :: "'a cfg \ 'a sentence \ 'a bins \ 'a forest list option" where "build_trees \ \ bs = ( let k = length bs - 1 in let finished = filter_with_index (\x. is_finished \ \ x) (items (bs!k)) in map_option concat (those (map (\(_, i). build_trees' bs \ k i {i}) finished)) )" lemma build_forest'_simps[simp]: "e = bs!k!i \ pointer e = Null \ build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" "e = bs!k!i \ pointer e = Pre pre \ build_trees' bs \ (k-1) pre {pre} = None \ build_trees' bs \ k i I = None" "e = bs!k!i \ pointer e = Pre pre \ build_trees' bs \ (k-1) pre {pre} = Some pres \ build_trees' bs \ k i I = those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None) pres)" by (subst build_trees'.simps, simp)+ definition wf_trees_input :: "('a bins \ 'a sentence \ nat \ nat \ nat set) set" where "wf_trees_input = { (bs, \, k, i, I) | bs \ k i I. sound_ptrs \ bs \ k < length bs \ i < length (bs!k) \ I \ {0.. i \ I }" fun build_forest'_measure :: "('a bins \ 'a sentence \ nat \ nat \ nat set) \ nat" where "build_forest'_measure (bs, \, k, i, I) = foldl (+) 0 (map length (take (k+1) bs)) - card I" lemma wf_trees_input_pre: assumes "(bs, \, k, i, I) \ wf_trees_input" assumes "e = bs!k!i" "pointer e = Pre pre" shows "(bs, \, (k-1), pre, {pre}) \ wf_trees_input" using assms unfolding wf_trees_input_def apply (auto simp: sound_ptrs_def sound_pre_ptr_def) apply (metis nth_mem) done lemma wf_trees_input_prered_pre: assumes "(bs, \, k, i, I) \ wf_trees_input" assumes "e = bs!k!i" "pointer e = PreRed p ps" assumes "ps' = filter (\(k', pre, red). red \ I) (p#ps)" assumes "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" assumes "((k', pre), reds) \ set gs" shows "(bs, \, k', pre, {pre}) \ wf_trees_input" proof - obtain red where "(k', pre, red) \ set ps'" using assms(5,6) group_by_exists_kv by fast hence *: "(k', pre, red) \ set (p#ps)" using assms(4) by (meson filter_is_subset in_mono) have "k < length bs" "e \ set (bs!k)" using assms(1,2) unfolding wf_trees_input_def by auto hence "k' < k" "pre < length (bs!k')" using assms(1,3) * unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def by blast+ thus ?thesis using assms by (simp add: wf_trees_input_def) qed lemma wf_trees_input_prered_red: assumes "(bs, \, k, i, I) \ wf_trees_input" assumes "e = bs!k!i" "pointer e = PreRed p ps" assumes "ps' = filter (\(k', pre, red). red \ I) (p#ps)" assumes "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" assumes "((k', pre), reds) \ set gs" "red \ set reds" shows "(bs, \, k, red, I \ {red}) \ wf_trees_input" proof - have "(k', pre, red) \ set ps'" using assms(5,6,7) group_by_forall_v_exists_k by fastforce hence *: "(k', pre, red) \ set (p#ps)" using assms(4) by (meson filter_is_subset in_mono) have "k < length bs" "e \ set (bs!k)" using assms(1,2) unfolding wf_trees_input_def by auto hence 0: "red < length (bs!k)" using assms(1,3) * unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def by blast moreover have "I \ {0.. {red} \ {0.., k, i, I) \ wf_trees_input" assumes "\bs \ k i I. (\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre {pre}) \ (\e p ps ps' gs k' pre reds. e = bs!k!i \ pointer e = PreRed p ps \ ps' = filter (\(k', pre, red). red \ I) (p#ps) \ gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' \ ((k', pre), reds) \ set gs \ P bs \ k' pre {pre}) \ (\e p ps ps' gs k' pre red reds reds'. e = bs!k!i \ pointer e = PreRed p ps \ ps' = filter (\(k', pre, red). red \ I) (p#ps) \ gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' \ ((k', pre), reds) \ set gs \ red \ set reds \ P bs \ k red (I \ {red})) \ P bs \ k i I" shows "P bs \ k i I" using assms(1) proof (induction n\"build_forest'_measure (bs, \, k, i, I)" arbitrary: k i I rule: nat_less_induct) case 1 obtain e where entry: "e = bs!k!i" by simp consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\p ps. pointer e = PreRed p ps" by (metis pointer.exhaust) thus ?case proof cases case Null thus ?thesis using assms(2) entry by fastforce next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast define n where n: "n = build_forest'_measure (bs, \, (k-1), pre, {pre})" have "0 < k" "pre < length (bs!(k-1))" using 1(2) entry pre unfolding wf_trees_input_def sound_ptrs_def sound_pre_ptr_def by (smt (verit) mem_Collect_eq nth_mem prod.inject)+ have "k < length bs" "i < length (bs!k)" "I \ {0.. I" using 1(2) unfolding wf_trees_input_def by blast+ have "length (bs!(k-1)) > 0" using \pre < length (bs!(k-1))\ by force hence "foldl (+) 0 (map length (take k bs)) > 0" by (smt (verit, del_insts) foldl_add_nth \0 < k\ \k < length bs\ add.commute add_diff_inverse_nat less_imp_diff_less less_one zero_eq_add_iff_both_eq_0) have "card I \ length (bs!k)" by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) have "card I + (foldl (+) 0 (map length (take (Suc (k - Suc 0)) bs)) - Suc 0) = card I + (foldl (+) 0 (map length (take k bs)) - 1)" using \0 < k\ by simp also have "... = card I + foldl (+) 0 (map length (take k bs)) - 1" using \0 < foldl (+) 0 (map length (take k bs))\ by auto also have "... < card I + foldl (+) 0 (map length (take k bs))" by (simp add: \0 < foldl (+) 0 (map length (take k bs))\) also have "... \ foldl (+) 0 (map length (take k bs)) + length (bs!k)" by (simp add: \card I \ length (bs ! k)\) also have "... = foldl (+) 0 (map length (take (k+1) bs))" using foldl_add_nth \k < length bs\ by blast finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, (k-1), pre, {pre}) > 0" by simp hence "P bs \ (k-1) pre {pre}" using 1 n wf_trees_input_pre entry pre zero_less_diff by blast thus ?thesis using assms(2) entry pre pointer.distinct(5) pointer.inject(1) by presburger next case PreRed then obtain p ps where pps: "pointer e = PreRed p ps" by blast define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" have 0: "\(k', pre, red) \ set ps'. k' < k \ pre < length (bs!k') \ red < length (bs!k)" using entry pps ps' 1(2) unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def apply (auto simp del: filter.simps) apply (metis nth_mem prod_cases3)+ done hence sound_gs: "\((k', pre), reds) \ set gs. k' < k \ pre < length (bs!k')" using gs group_by_exists_kv by fast have sound_gs2: "\((k', pre), reds) \ set gs. \red \ set reds. red < length (bs!k)" proof (standard, standard, standard, standard) fix x a b k' pre red assume "x \ set gs" "x = (a, b)" "(k', pre) = a" "red \ set b" hence "\x \ set ps'. red = (\(k', pre, red). red) x" using group_by_forall_v_exists_k gs ps' by meson thus "red < length (bs!k)" using 0 by fast qed { fix k' pre reds red assume a0: "((k', pre), reds) \ set gs" define n_pre where n_pre: "n_pre = build_forest'_measure (bs, \, k', pre, {pre})" have "k < length bs" "i < length (bs!k)" "I \ {0.. I" using 1(2) unfolding wf_trees_input_def by blast+ have "k' < k" "pre < length (bs!k')" using sound_gs a0 by fastforce+ have "length (bs!k') > 0" using \pre < length (bs!k')\ by force hence gt0: "foldl (+) 0 (map length (take (k'+1) bs)) > 0" by (smt (verit, del_insts) foldl_add_nth \k < length bs\ \k' < k\ add_gr_0 order.strict_trans) have card_bound: "card I \ length (bs!k)" by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) have "card I + (foldl (+) 0 (map length (take (Suc k') bs)) - Suc 0) = card I + foldl (+) 0 (map length (take (Suc k') bs)) - 1" by (metis Nat.add_diff_assoc One_nat_def Suc_eq_plus1 Suc_leI \0 < foldl (+) 0 (map length (take (k' + 1) bs))\) also have "... < card I + foldl (+) 0 (map length (take (Suc k') bs))" using gt0 by auto also have "... \ foldl (+) 0 (map length (take (Suc k') bs)) + length (bs!k)" using card_bound by simp also have "... \ foldl (+) 0 (map length (take (k+1) bs))" using foldl_add_nth_ge Suc_leI \k < length bs\ \k' < k\ by blast finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, k', pre, {pre}) > 0" by simp hence "P bs \ k' pre {pre}" using 1(1) wf_trees_input_prered_pre[OF "1.prems"(1) entry pps ps' gs a0] zero_less_diff by blast } moreover { fix k' pre reds red assume a0: "((k', pre), reds) \ set gs" assume a1: "red \ set reds" define n_red where n_red: "n_red = build_forest'_measure (bs, \, k, red, I \ {red})" have "k < length bs" "i < length (bs!k)" "I \ {0.. I" using 1(2) unfolding wf_trees_input_def by blast+ have "k' < k" "pre < length (bs!k')" "red < length (bs!k)" using a0 a1 sound_gs sound_gs2 by fastforce+ have "red \ I" using a0 a1 unfolding gs ps' by (smt (verit, best) group_by_forall_v_exists_k case_prodE case_prod_conv mem_Collect_eq set_filter) have card_bound: "card I \ length (bs!k)" by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) have "length (bs!k') > 0" using \pre < length (bs!k')\ by force hence gt0: "foldl (+) 0 (map length (take (k'+1) bs)) > 0" by (smt (verit, del_insts) foldl_add_nth \k < length bs\ \k' < k\ add_gr_0 order.strict_trans) have lt: "foldl (+) 0 (map length (take (Suc k') bs)) + length (bs!k) \ foldl (+) 0 (map length (take (k+1) bs))" using foldl_add_nth_ge Suc_leI \k < length bs\ \k' < k\ by blast have "card I + (foldl (+) 0 (map length (take (Suc k) bs)) - card (insert red I)) = card I + (foldl (+) 0 (map length (take (Suc k) bs)) - card I - 1)" using \I \ {0.. \red \ I\ finite_subset by fastforce also have "... = foldl (+) 0 (map length (take (Suc k) bs)) - 1" using gt0 card_bound lt by force also have "... < foldl (+) 0 (map length (take (Suc k) bs))" using gt0 lt by auto finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, k, red, I \ {red}) > 0" by simp moreover have "(bs, \, k, red, I \ {red}) \ wf_trees_input" using wf_trees_input_prered_red[OF "1"(2) entry pps ps' gs] a0 a1 by blast ultimately have "P bs \ k red (I \ {red})" using 1(1) zero_less_diff by blast } moreover have "(\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre {pre})" using entry pps by fastforce ultimately show ?thesis using assms(2) entry gs pointer.inject(2) pps ps' by presburger qed qed lemma build_trees'_termination: assumes "(bs, \, k, i, I) \ wf_trees_input" shows "\fs. build_trees' bs \ k i I = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" proof - have "\fs. build_trees' bs \ k i I = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" apply (induction rule: build_trees'_induct[OF assms(1)]) subgoal premises IH for bs \ k i I proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" by (metis pointer.exhaust surj_pair) thus ?thesis proof cases case Null have "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" using build_forest'_simps(1) Null entry by simp thus ?thesis by simp next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast obtain fs where fs: "build_trees' bs \ (k-1) pre {pre} = Some fs" "\f \ set fs. \N fss. f = FBranch N fss" using IH(1) entry pre by blast let ?g = "\f. case f of FLeaf a \ None | FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]]))" have simp: "build_trees' bs \ k i I = those (map ?g fs)" using build_forest'_simps(3) entry pre fs by blast moreover have "\f \ set (map ?g fs). \a. f = Some a" using fs(2) by auto ultimately obtain fs' where fs': "build_trees' bs \ k i I = Some fs'" using those_Some by (smt (verit, best)) moreover have "\f \ set fs'. \N fss. f = FBranch N fss" proof standard fix f assume "f \ set fs'" then obtain x where "x \ set fs" "Some f \ set (map ?g fs)" using those_map_exists by (metis (no_types, lifting) fs' simp) thus "\N fss. f = FBranch N fss" using fs(2) by auto qed ultimately show ?thesis by blast next case PreRed then obtain p ps where pps: "pointer e = PreRed p ps" by blast define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" let ?g = "\((k', pre), reds). do { pres \ build_trees' bs \ k' pre {pre}; rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None \\impossible case\ ) pres) }" have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" using entry pps ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" proof standard fix fso assume "fso \ set (map ?g gs)" moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" proof standard fix ps assume "ps \ set gs" then obtain k' pre reds where *: "((k', pre), reds) \ set gs" "((k', pre), reds) = ps" by (metis surj_pair) then obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" "\f \ set pres. \N fss. f = FBranch N fss" using IH(2) entry pps ps' gs by blast have "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" using IH(3)[OF entry pps ps' gs *(1)] by auto then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" using those_Some by (metis (full_types)) let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" have "\x \ set (map ?h pres). \a. x = Some a" using pres(2) by auto then obtain fs where fs: "Some fs = those (map ?h pres)" using those_Some by (smt (verit, best)) have "\f \ set fs. \N fss. f = FBranch N fss" proof standard fix f assume *: "f \ set fs" hence "\x. x \ set pres \ Some f \ set (map ?h pres)" using those_map_exists[OF fs *] by blast then obtain x where x: "x \ set pres \ Some f \ set (map ?h pres)" by blast thus "\N fss. f = FBranch N fss" using pres(2) by auto qed moreover have "?g ps = Some fs" using fs pres rss * by (auto, metis bind.bind_lunit) ultimately show "\fs. ?g ps = Some fs \ (\f\set fs. \N fss. f = FBranch N fss)" by blast qed ultimately show "\fs. fso = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" using map_Some_P by auto qed then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \N fss. f = FBranch N fss" using those_Some_P by blast hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \N fss. f = FBranch N fss" using simp by auto thus ?thesis by blast qed qed done thus ?thesis by blast qed lemma wf_item_tree_build_trees': assumes "(bs, \, k, i, I) \ wf_trees_input" assumes "wf_bins \ \ bs" assumes "k < length bs" "i < length (bs!k)" assumes "build_trees' bs \ k i I = Some fs" assumes "f \ set fs" assumes "t \ set (trees f)" shows "wf_item_tree \ (item (bs!k!i)) t" proof - have "wf_item_tree \ (item (bs!k!i)) t" using assms apply (induction arbitrary: fs f t rule: build_trees'_induct[OF assms(1)]) subgoal premises prems for bs \ k i I fs f t proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\p ps. pointer e = PreRed p ps" by (metis pointer.exhaust) thus ?thesis proof cases case Null hence simp: "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" using entry by simp moreover have "f = FBranch (item_rule_head (item e)) []" using build_forest'_simps(1) Null prems(8,9) entry by auto ultimately have simp: "t = Branch (item_rule_head (item e)) []" using prems(10) by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast hence "predicts (item e)" using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast hence "item_dot (item e) = 0" unfolding predicts_def by blast thus ?thesis using simp entry by simp next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast have sound: "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp hence *: "item_rule_head (item (bs!(k-1)!pre)) = item_rule_head (item e)" "item_rule_body (item (bs!(k-1)!pre)) = item_rule_body (item e)" "item_dot (item (bs!(k-1)!pre)) + 1 = item_dot (item e)" "next_symbol (item (bs!(k-1)!pre)) = Some (\!(k-1))" unfolding scans_def inc_item_def by (simp_all add: item_rule_head_def item_rule_body_def) have wf: "(bs, \, k-1, pre, {pre}) \ wf_trees_input" using entry pre prems(4) wf_trees_input_pre by blast then obtain pres where pres: "build_trees' bs \ (k-1) pre {pre} = Some pres" "\f \ set pres. \N fss. f = FBranch N fss" using build_trees'_termination wf by blast let ?g = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None" have "build_trees' bs \ k i I = those (map ?g pres)" using entry pre pres by simp hence fs: "Some fs = those (map ?g pres)" using prems(8) by simp then obtain f_pre N fss where Nfss: "f = FBranch N (fss @ [[FLeaf (\!(k-1))]])" "f_pre = FBranch N fss" "f_pre \ set pres" using those_map_FBranch_only fs pres(2) prems(9) by blast define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" have "trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])) = map (\ts. Branch N ts) [ ts @ [Leaf (\!(k-1))] . ts <- combinations tss ]" by (subst tss, subst trees_append_single_singleton, simp) moreover have "t \ set (trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])))" using Nfss(1) prems(10) by blast ultimately obtain ts where ts: "t = Branch N (ts @ [Leaf (\!(k-1))]) \ ts \ set (combinations tss)" by auto have "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast hence "pre < length (bs!(k-1))" using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) moreover have "k - 1 < length bs" by (simp add: prems(6) less_imp_diff_less) moreover have "Branch N ts \ set (trees (FBranch N fss))" using ts tss by simp ultimately have IH: "wf_item_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" using prems(1,2,4,5) entry pre Nfss(2,3) wf pres(1) by blast have "map root_tree (ts @ [Leaf (\!(k-1))]) = map root_tree ts @ [\!(k-1)]" by simp also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item (bs!(k-1)!pre))) @ [\!(k-1)]" using IH by simp also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item e)) @ [\!(k-1)]" using *(2) by simp also have "... = take (item_dot (item e)) (item_rule_body (item e))" using *(2-4) by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) finally have "map root_tree (ts @ [Leaf (\!(k-1))]) = take (item_dot (item e)) (item_rule_body (item e))" . hence "wf_item_tree \ (item e) (Branch N (ts @ [Leaf (\!(k-1))]))" using IH *(1) by simp thus ?thesis using ts entry by fastforce next case PreRed then obtain p ps where prered: "pointer e = PreRed p ps" by blast define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" let ?g = "\((k', pre), reds). do { pres \ build_trees' bs \ k' pre {pre}; rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None \\impossible case\ ) pres) }" have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" using entry prered ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" proof standard fix fso assume "fso \ set (map ?g gs)" moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" proof standard fix g assume "g \ set gs" then obtain k' pre reds where g: "((k', pre), reds) \ set gs" "((k', pre), reds) = g" by (metis surj_pair) moreover have wf_pre: "(bs, \, k', pre, {pre}) \ wf_trees_input" using wf_trees_input_prered_pre[OF prems(4) entry prered ps' gs g(1)] by blast ultimately obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" "\f_pre \ set pres. \N fss. f_pre = FBranch N fss" using build_trees'_termination by blast have wf_reds: "\red \ set reds. (bs, \, k, red, I \ {red}) \ wf_trees_input" using wf_trees_input_prered_red[OF prems(4) entry prered ps' gs g(1)] by blast hence "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" using build_trees'_termination by fastforce then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" using those_Some by (metis (full_types)) let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" have "\x \ set (map ?h pres). \a. x = Some a" using pres(2) by auto then obtain fs where fs: "Some fs = those (map ?h pres)" using those_Some by (smt (verit, best)) have "\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" proof (standard, standard) fix f t assume ft: "f \ set fs" "t \ set (trees f)" hence "\x. x \ set pres \ Some f \ set (map ?h pres)" using those_map_exists[OF fs ft(1)] by blast then obtain f_pre N fss where f_pre: "f_pre \ set pres" "f_pre = FBranch N fss" "f = FBranch N (fss @ [concat rss])" using pres(2) by force define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" have "trees (FBranch N (fss @ [concat rss])) = map (\ts. Branch N ts) [ ts0 @ ts1 . ts0 <- combinations tss, ts1 <- combinations [concat (map (\f. trees f) (concat rss)) ] ]" by (subst tss, subst trees_append_singleton, simp) moreover have "t \ set (trees (FBranch N (fss @ [concat rss])))" using ft(2) f_pre(3) by blast ultimately obtain ts0 ts1 where tsx: "t = Branch N (ts0 @ [ts1])" "ts0 \ set (combinations tss)" "ts1 \ set (concat (map (\f. trees f) (concat rss)))" by fastforce then obtain f_red where f_red: "f_red \ set (concat rss)" "ts1 \ set (trees f_red)" by auto obtain fs_red red where red: "Some fs_red = build_trees' bs \ k red (I \ {red})" "f_red \ set fs_red" "red \ set reds" using f_red(1) rss those_map_Some_concat_exists by fast then obtain N_red fss_red where "f_red = FBranch N_red fss_red" using build_trees'_termination wf_reds by (metis option.inject) then obtain ts where ts: "Branch N_red ts = ts1" using tsx(3) f_red by auto have "(k', pre, red) \ set ps'" using group_by_forall_v_exists_k \((k', pre), reds) \ set gs\ gs \red \ set reds\ by fast hence mem: "(k', pre, red) \ set (p#ps)" using ps' by (metis filter_set member_filter) have "sound_ptrs \ bs" using prems(4) wf_trees_input_def by fastforce have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by (meson mem nth_mem)+ have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by (metis mem nth_mem) have transform: "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" "is_complete (item (bs!k!red))" using completes unfolding completes_def inc_item_def by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) have "Branch N ts0 \ set (trees (FBranch N fss))" using tss tsx(2) by simp hence IH_pre: "wf_item_tree \ (item (bs!k'!pre)) (Branch N ts0)" using prems(2)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ wf_pre prems(5)] pres(1) f_pre f_pre(3) bounds(1,2) prems(6) by fastforce have IH_r: "wf_item_tree \ (item (bs!k!red)) (Branch N_red ts)" using prems(3)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ \red \ set reds\ _ prems(5)] bounds(3) f_red(2) red ts wf_reds prems(6) by metis have "map root_tree (ts0 @ [Branch N_red ts]) = map root_tree ts0 @ [root_tree (Branch N_red ts)]" by simp also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [root_tree (Branch N_red ts)]" using IH_pre by simp also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [item_rule_head (item (bs!k!red))]" using IH_r by simp also have "... = take (item_dot (item e)) (item_rule_body (item e))" using transform by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) finally have roots: "map root_tree (ts0 @ [Branch N_red ts]) = take (item_dot (item e)) (item_rule_body (item e))" . have "wf_item \ \ (item (bs!k!red))" using prems(5,6) bounds(3) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (auto simp: items_def) moreover have "N_red = item_rule_head (item (bs!k!red))" using IH_r by fastforce moreover have "map root_tree ts = item_rule_body (item (bs!k!red))" using IH_r transform(5) by (auto simp: is_complete_def) ultimately have "\r \ set (\ \). N_red = rule_head r \ map root_tree ts = rule_body r" unfolding wf_item_def item_rule_body_def item_rule_head_def by blast hence "wf_rule_tree \ (Branch N_red ts)" using IH_r by simp hence "wf_item_tree \ (item (bs!k!i)) (Branch N (ts0 @ [Branch N_red ts]))" using transform(1) roots IH_pre entry by simp thus "wf_item_tree \ (item (bs!k!i)) t" using tsx(1) red ts by blast qed moreover have "?g g = Some fs" using fs pres rss g by (auto, metis bind.bind_lunit) ultimately show "\fs. ?g g = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" by blast qed ultimately show "\fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" using map_Some_P by auto qed then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" using those_Some_P by blast hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" using simp by auto thus ?thesis using prems(8-10) by auto qed qed done thus ?thesis by blast qed lemma wf_yield_tree_build_trees': assumes "(bs, \, k, i, I) \ wf_trees_input" assumes "wf_bins \ \ bs" assumes "k < length bs" "i < length (bs!k)" "k \ length \" assumes "build_trees' bs \ k i I = Some fs" assumes "f \ set fs" assumes "t \ set (trees f)" shows "wf_yield_tree \ (item (bs!k!i)) t" proof - have "wf_yield_tree \ (item (bs!k!i)) t" using assms apply (induction arbitrary: fs f t rule: build_trees'_induct[OF assms(1)]) subgoal premises prems for bs \ k i I fs f t proof - define e where entry: "e = bs!k!i" consider (Null) "pointer e = Null" | (Pre) "\pre. pointer e = Pre pre" | (PreRed) "\p ps. pointer e = PreRed p ps" by (metis pointer.exhaust) thus ?thesis proof cases case Null hence simp: "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" using entry by simp moreover have "f = FBranch (item_rule_head (item e)) []" using build_forest'_simps(1) Null prems(9,10) entry by auto ultimately have simp: "t = Branch (item_rule_head (item e)) []" using prems(11) by simp have "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast hence "predicts (item e)" using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast thus ?thesis unfolding wf_yield_tree_def predicts_def using simp entry by (auto simp: slice_empty) next case Pre then obtain pre where pre: "pointer e = Pre pre" by blast have sound: "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast hence bounds: "k > 0" "pre < length (bs!(k-1))" using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem)+ have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp have wf: "(bs, \, k-1, pre, {pre}) \ wf_trees_input" using entry pre prems(4) wf_trees_input_pre by blast then obtain pres where pres: "build_trees' bs \ (k-1) pre {pre} = Some pres" "\f \ set pres. \N fss. f = FBranch N fss" using build_trees'_termination wf by blast let ?g = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None" have "build_trees' bs \ k i I = those (map ?g pres)" using entry pre pres by simp hence fs: "Some fs = those (map ?g pres)" using prems(9) by simp then obtain f_pre N fss where Nfss: "f = FBranch N (fss @ [[FLeaf (\!(k-1))]])" "f_pre = FBranch N fss" "f_pre \ set pres" using those_map_FBranch_only fs pres(2) prems(10) by blast define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" have "trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])) = map (\ts. Branch N ts) [ ts @ [Leaf (\!(k-1))] . ts <- combinations tss ]" by (subst tss, subst trees_append_single_singleton, simp) moreover have "t \ set (trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])))" using Nfss(1) prems(11) by blast ultimately obtain ts where ts: "t = Branch N (ts @ [Leaf (\!(k-1))]) \ ts \ set (combinations tss)" by auto have "sound_ptrs \ bs" using prems(4) unfolding wf_trees_input_def by blast hence "pre < length (bs!(k-1))" using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) moreover have "k-1 < length bs" by (simp add: prems(6) less_imp_diff_less) moreover have "Branch N ts \ set (trees (FBranch N fss))" using ts tss by simp ultimately have IH: "wf_yield_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" using prems(1,2,4,5,8) entry pre Nfss(2,3) wf pres(1) by simp have transform: "item_origin (item (bs!(k-1)!pre)) \ item_end (item (bs!(k-1)!pre))" "item_end (item (bs!(k-1)!pre)) = k-1" "item_end (item e) = k" using entry prems(5,6,7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def by (auto, meson less_imp_diff_less nth_mem) have "yield_tree t = concat (map yield_tree (ts @ [Leaf (\!(k-1))]))" by (simp add: ts) also have "... = concat (map yield_tree ts) @ [\!(k-1)]" by simp also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre))) \ @ [\!(k-1)]" using IH by (simp add: wf_yield_tree_def) also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre)) + 1) \" using slice_append_nth transform \k > 0\ prems(8) by (metis diff_less le_eq_less_or_eq less_imp_diff_less less_numeral_extra(1)) also have "... = slice (item_origin (item e)) (item_end (item (bs!(k-1)!pre)) + 1) \" using scans unfolding scans_def inc_item_def by simp also have "... = slice (item_origin (item e)) k \" using scans transform unfolding scans_def by (metis Suc_diff_1 Suc_eq_plus1 bounds(1)) also have "... = slice (item_origin (item e)) (item_end (item e)) \" using transform by auto finally show ?thesis using wf_yield_tree_def entry by blast next case PreRed then obtain p ps where prered: "pointer e = PreRed p ps" by blast define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" let ?g = "\((k', pre), reds). do { pres \ build_trees' bs \ k' pre {pre}; rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None \\impossible case\ ) pres) }" have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" using entry prered ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" proof standard fix fso assume "fso \ set (map ?g gs)" moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" proof standard fix g assume "g \ set gs" then obtain k' pre reds where g: "((k', pre), reds) \ set gs" "((k', pre), reds) = g" by (metis surj_pair) moreover have wf_pre: "(bs, \, k', pre, {pre}) \ wf_trees_input" using wf_trees_input_prered_pre[OF prems(4) entry prered ps' gs g(1)] by blast ultimately obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" "\f_pre \ set pres. \N fss. f_pre = FBranch N fss" using build_trees'_termination by blast have wf_reds: "\red \ set reds. (bs, \, k, red, I \ {red}) \ wf_trees_input" using wf_trees_input_prered_red[OF prems(4) entry prered ps' gs g(1)] by blast hence "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" using build_trees'_termination by fastforce then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" using those_Some by (metis (full_types)) let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" have "\x \ set (map ?h pres). \a. x = Some a" using pres(2) by auto then obtain fs where fs: "Some fs = those (map ?h pres)" using those_Some by (smt (verit, best)) have "\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" proof (standard, standard) fix f t assume ft: "f \ set fs" "t \ set (trees f)" hence "\x. x \ set pres \ Some f \ set (map ?h pres)" using those_map_exists[OF fs ft(1)] by blast then obtain f_pre N fss where f_pre: "f_pre \ set pres" "f_pre = FBranch N fss" "f = FBranch N (fss @ [concat rss])" using pres(2) by force define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" have "trees (FBranch N (fss @ [concat rss])) = map (\ts. Branch N ts) [ ts0 @ ts1 . ts0 <- combinations tss, ts1 <- combinations [concat (map (\f. trees f) (concat rss)) ] ]" by (subst tss, subst trees_append_singleton, simp) moreover have "t \ set (trees (FBranch N (fss @ [concat rss])))" using ft(2) f_pre(3) by blast ultimately obtain ts0 ts1 where tsx: "t = Branch N (ts0 @ [ts1])" "ts0 \ set (combinations tss)" "ts1 \ set (concat (map (\f. trees f) (concat rss)))" by fastforce then obtain f_red where f_red: "f_red \ set (concat rss)" "ts1 \ set (trees f_red)" by auto obtain fs_red red where red: "Some fs_red = build_trees' bs \ k red (I \ {red})" "f_red \ set fs_red" "red \ set reds" using f_red(1) rss those_map_Some_concat_exists by fast then obtain N_red fss_red where "f_red = FBranch N_red fss_red" using build_trees'_termination wf_reds by (metis option.inject) then obtain ts where ts: "Branch N_red ts = ts1" using tsx(3) f_red by auto have "(k', pre, red) \ set ps'" using group_by_forall_v_exists_k \((k', pre), reds) \ set gs\ gs \red \ set reds\ by fast hence mem: "(k', pre, red) \ set (p#ps)" using ps' by (metis filter_set member_filter) have "sound_ptrs \ bs" using prems(4) wf_trees_input_def by fastforce have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by (meson mem nth_mem)+ have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" using prered entry prems(6,7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_prered_ptr_def by (metis mem nth_mem) have transform: "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" "is_complete (item (bs!k!red))" using completes unfolding completes_def inc_item_def by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) have "Branch N ts0 \ set (trees (FBranch N fss))" using tss tsx(2) by simp hence IH_pre: "wf_yield_tree \ (item (bs!k'!pre)) (Branch N ts0)" using prems(2)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ wf_pre prems(5)] pres(1) f_pre f_pre(3) bounds(1,2) prems(6,8) by simp have IH_r: "wf_yield_tree \ (item (bs!k!red)) (Branch N_red ts)" using prems(3)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ \red \ set reds\ _ prems(5)] bounds(3) f_red(2) red ts wf_reds prems(6,8) by metis have wf1: "item_origin (item (bs!k'!pre)) \ item_end (item (bs!k'!pre))" "item_origin (item (bs!k!red)) \ item_end (item (bs!k!red))" using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def by (metis length_map nth_map nth_mem order_less_trans)+ have wf2: "item_end (item (bs!k!red)) = k" "item_end (item (bs!k!i)) = k" using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def by simp_all have "yield_tree t = concat (map yield_tree (ts0 @ [Branch N_red ts]))" by (simp add: ts tsx(1)) also have "... = concat (map yield_tree ts0) @ yield_tree (Branch N_red ts)" by simp also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k'!pre))) \ @ slice (item_origin (item (bs!k!red))) (item_end (item (bs!k!red))) \" using IH_pre IH_r by (simp add: wf_yield_tree_def) also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k!red))) \" using slice_concat wf1 completes_def completes by (metis (no_types, lifting)) also have "... = slice (item_origin (item e)) (item_end (item (bs!k!red))) \" using completes unfolding completes_def inc_item_def by simp also have "... = slice (item_origin (item e)) (item_end (item e)) \" using wf2 entry by presburger finally show "wf_yield_tree \ (item (bs!k!i)) t" using wf_yield_tree_def entry by blast qed moreover have "?g g = Some fs" using fs pres rss g by (auto, metis bind.bind_lunit) ultimately show "\fs. ?g g = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" by blast qed ultimately show "\fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" using map_Some_P by auto qed then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" using those_Some_P by blast hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" using simp by auto thus ?thesis using prems(9-11) by auto qed qed done thus ?thesis using assms(2) by blast qed theorem wf_rule_root_yield_tree_build_trees: assumes "wf_bins \ \ bs" "sound_ptrs \ bs" "length bs = length \ + 1" assumes "build_trees \ \ bs = Some fs" "f \ set fs" "t \ set (trees f)" shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" proof - let ?k = "length bs - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items (bs!?k))" have #: "Some fs = map_option concat (those (map (\(_, i). build_trees' bs \ ?k i {i}) finished))" using assms(4) build_trees_def finished_def by (metis (full_types)) then obtain fss fs' where fss: "Some fss = those (map (\(_, i). build_trees' bs \ ?k i {i}) finished)" "fs' \ set fss" "f \ set fs'" using map_option_concat_those_map_exists assms(5) by fastforce then obtain x i where *: "(x,i) \ set finished" "Some fs' = build_trees' bs \ (length bs - 1) i {i}" using those_map_exists[OF fss(1,2)] by auto have k: "?k < length bs" "?k \ length \" using assms(3) by simp_all have i: "i < length (bs!?k)" using index_filter_with_index_lt_length * items_def finished_def by (metis (no_types, opaque_lifting) length_map) have x: "x = item (bs!?k!i)" using * i filter_with_index_nth items_def nth_map finished_def assms(3) by metis have finished: "is_finished \ \ x" using * filter_with_index_P finished_def by metis have "{i} \ {0.., ?k, i, {i}) \ wf_trees_input" unfolding wf_trees_input_def using assms(2) i k(1) by simp hence wf_item_tree: "wf_item_tree \ (item (bs!?k!i)) t" using wf_item_tree_build_trees' assms(1,2,5,6) i k(1) x *(2) fss(3) by metis have wf_item: "wf_item \ \ (item (bs!?k!i))" using k(1) i assms(1) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (simp add: items_def) obtain N fss where Nfss: "f = FBranch N fss" using build_trees'_termination[OF wf] by (metis "*"(2) fss(3) option.inject) then obtain ts where ts: "t = Branch N ts" using assms(6) by auto hence "N = item_rule_head x" "map root_tree ts = item_rule_body x" using finished wf_item_tree x by (auto simp: is_finished_def is_complete_def) hence "\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r" using wf_item x unfolding wf_item_def item_rule_body_def item_rule_head_def by blast hence wf_rule: "wf_rule_tree \ t" using wf_item_tree ts by simp have root: "root_tree t = \ \" using finished ts \N = item_rule_head x\ by (auto simp: is_finished_def) have "yield_tree t = slice (item_origin (item (bs!?k!i))) (item_end (item (bs!?k!i))) \" using k i assms(1,6) wf wf_yield_tree_build_trees' wf_yield_tree_def *(2) fss(3) by (smt (verit, best)) hence yield: "yield_tree t = \" using finished x unfolding is_finished_def by simp show ?thesis using * wf_rule root yield assms(4) unfolding build_trees_def by simp qed corollary wf_rule_root_yield_tree_build_trees_Earley\<^sub>L: assumes "wf_\ \" "nonempty_derives \" assumes "build_trees \ \ (Earley\<^sub>L \ \) = Some fs" "f \ set fs" "t \ set (trees f)" shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" using assms wf_rule_root_yield_tree_build_trees wf_bins_Earley\<^sub>L Earley\<^sub>L_def length_Earley\<^sub>L_bins length_bins_Init\<^sub>L sound_mono_ptrs_Earley\<^sub>L by (metis dual_order.eq_iff ) theorem soundness_build_trees_Earley\<^sub>L: assumes "wf_\ \" "is_word \ \" "nonempty_derives \" assumes "build_trees \ \ (Earley\<^sub>L \ \) = Some fs" "f \ set fs" "t \ set (trees f)" shows "derives \ [\ \] \" proof - let ?k = "length (Earley\<^sub>L \ \) - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" have #: "Some fs = map_option concat (those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished))" using assms(4) build_trees_def finished_def by (metis (full_types)) then obtain fss fs' where fss: "Some fss = those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished)" "fs' \ set fss" "f \ set fs'" using map_option_concat_those_map_exists assms(5) by fastforce then obtain x i where *: "(x,i) \ set finished" "Some fs' = build_trees' (Earley\<^sub>L \ \) \ ?k i {i}" using those_map_exists[OF fss(1,2)] by auto have k: "?k < length (Earley\<^sub>L \ \)" "?k \ length \" by (simp_all add: Earley\<^sub>L_def assms(1)) have i: "i < length ((Earley\<^sub>L \ \) ! ?k)" using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) have x: "x = item ((Earley\<^sub>L \ \)!?k!i)" using * i filter_with_index_nth items_def nth_map finished_def by metis have finished: "is_finished \ \ x" using * filter_with_index_P finished_def by metis moreover have "x \ set (items ((Earley\<^sub>L \ \) ! ?k))" using x by (auto simp: items_def; metis One_nat_def i imageI nth_mem) ultimately have "recognizing (bins (Earley\<^sub>L \ \)) \ \" by (meson k(1) kth_bin_sub_bins recognizing_def subsetD) thus ?thesis using correctness_Earley\<^sub>L assms by blast qed theorem termination_build_tree_Earley\<^sub>L: assumes "wf_\ \" "nonempty_derives \" "derives \ [\ \] \" shows "\fs. build_trees \ \ (Earley\<^sub>L \ \) = Some fs" proof - let ?k = "length (Earley\<^sub>L \ \) - 1" define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" have "\f \ set finished. (Earley\<^sub>L \ \, \, ?k, snd f, {snd f}) \ wf_trees_input" proof standard fix f assume a: "f \ set finished" then obtain x i where *: "(x,i) = f" by (metis surj_pair) have "sound_ptrs \ (Earley\<^sub>L \ \)" using sound_mono_ptrs_Earley\<^sub>L assms by blast moreover have "?k < length (Earley\<^sub>L \ \)" by (simp add: Earley\<^sub>L_def assms(1)) moreover have "i < length ((Earley\<^sub>L \ \)!?k)" using index_filter_with_index_lt_length a * items_def finished_def by (metis length_map) ultimately show "(Earley\<^sub>L \ \, \, ?k, snd f, {snd f}) \ wf_trees_input" using * unfolding wf_trees_input_def by auto qed hence "\fso \ set (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished). \fs. fso = Some fs" using build_trees'_termination by fastforce then obtain fss where fss: "Some fss = those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished)" by (smt (verit, best) those_Some) then obtain fs where fs: "Some fs = map_option concat (those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished))" by (metis map_option_eq_Some) show ?thesis using finished_def fss fs build_trees_def by (metis (full_types)) qed end \ No newline at end of file diff --git a/thys/Earley_Parser/Earley_Recognizer.thy b/thys/Earley_Parser/Earley_Recognizer.thy --- a/thys/Earley_Parser/Earley_Recognizer.thy +++ b/thys/Earley_Parser/Earley_Recognizer.thy @@ -1,2331 +1,2335 @@ theory Earley_Recognizer imports Earley_Fixpoint begin section \Earley recognizer\ subsection \List auxilaries\ fun filter_with_index' :: "nat \ ('a \ bool) \ 'a list \ ('a \ nat) list" where "filter_with_index' _ _ [] = []" | "filter_with_index' i P (x#xs) = ( if P x then (x,i) # filter_with_index' (i+1) P xs else filter_with_index' (i+1) P xs)" definition filter_with_index :: "('a \ bool) \ 'a list \ ('a \ nat) list" where "filter_with_index P xs = filter_with_index' 0 P xs" lemma filter_with_index'_P: "(x, n) \ set (filter_with_index' i P xs) \ P x" by (induction xs arbitrary: i) (auto split: if_splits) lemma filter_with_index_P: "(x, n) \ set (filter_with_index P xs) \ P x" by (metis filter_with_index'_P filter_with_index_def) lemma filter_with_index'_cong_filter: "map fst (filter_with_index' i P xs) = filter P xs" by (induction xs arbitrary: i) auto lemma filter_with_index_cong_filter: "map fst (filter_with_index P xs) = filter P xs" by (simp add: filter_with_index'_cong_filter filter_with_index_def) lemma size_index_filter_with_index': "(x, n) \ set (filter_with_index' i P xs) \ n \ i" by (induction xs arbitrary: i) (auto simp: Suc_leD split: if_splits) lemma index_filter_with_index'_lt_length: "(x, n) \ set (filter_with_index' i P xs) \ n-i < length xs" by (induction xs arbitrary: i)(auto simp: less_Suc_eq_0_disj split: if_splits; metis Suc_diff_Suc leI)+ lemma index_filter_with_index_lt_length: "(x, n) \ set (filter_with_index P xs) \ n < length xs" by (metis filter_with_index_def index_filter_with_index'_lt_length minus_nat.diff_0) lemma filter_with_index'_nth: "(x, n) \ set (filter_with_index' i P xs) \ xs ! (n-i) = x" proof (induction xs arbitrary: i) case (Cons y xs) show ?case proof (cases "x = y") case True thus ?thesis using Cons by (auto simp: nth_Cons' split: if_splits) next case False hence "(x, n) \ set (filter_with_index' (i+1) P xs)" using Cons.prems by (cases xs) (auto split: if_splits) hence "n \ i + 1" "xs ! (n - i - 1) = x" by (auto simp: size_index_filter_with_index' Cons.IH) thus ?thesis by simp qed qed simp lemma filter_with_index_nth: "(x, n) \ set (filter_with_index P xs) \ xs ! n = x" by (metis diff_zero filter_with_index'_nth filter_with_index_def) lemma filter_with_index_nonempty: "x \ set xs \ P x \ filter_with_index P xs \ []" by (metis filter_empty_conv filter_with_index_cong_filter list.map(1)) lemma filter_with_index'_Ex_first: "(\x i xs'. filter_with_index' n P xs = (x, i)#xs') \ (\x \ set xs. P x)" by (induction xs arbitrary: n) auto lemma filter_with_index_Ex_first: "(\x i xs'. filter_with_index P xs = (x, i)#xs') \ (\x \ set xs. P x)" using filter_with_index'_Ex_first filter_with_index_def by metis subsection \Definitions\ datatype pointer = Null | Pre nat \\pre\ | PreRed "nat \ nat \ nat" "(nat \ nat \ nat) list" \\k', pre, red\ datatype 'a entry = Entry (item : "'a item") (pointer : pointer) type_synonym 'a bin = "'a entry list" type_synonym 'a bins = "'a bin list" definition items :: "'a bin \ 'a item list" where "items b \ map item b" +lemma length_items_eq: + \length (items b) = length b\ + by (simp add: items_def) + definition pointers :: "'a bin \ pointer list" where "pointers b \ map pointer b" definition bins_eq_items :: "'a bins \ 'a bins \ bool" where "bins_eq_items bs0 bs1 \ map items bs0 = map items bs1" definition bins :: "'a bins \ 'a item set" where "bins bs \ \ { set (items (bs!k)) | k. k < length bs }" definition bin_upto :: "'a bin \ nat \ 'a item set" where "bin_upto b i \ { items b ! j | j. j < i \ j < length (items b) }" definition bins_upto :: "'a bins \ nat \ nat \ 'a item set" where "bins_upto bs k i \ \ { set (items (bs ! l)) | l. l < k } \ bin_upto (bs ! k) i" definition wf_bin_items :: "'a cfg \ 'a sentence \ nat \ 'a item list \ bool" where "wf_bin_items \ \ k xs \ \x \ set xs. wf_item \ \ x \ item_end x = k" definition wf_bin :: "'a cfg \ 'a sentence \ nat \ 'a bin \ bool" where "wf_bin \ \ k b \ distinct (items b) \ wf_bin_items \ \ k (items b)" definition wf_bins :: "'a cfg \ 'a list \ 'a bins \ bool" where "wf_bins \ \ bs \ \k < length bs. wf_bin \ \ k (bs!k)" definition nonempty_derives :: "'a cfg \ bool" where "nonempty_derives \ \ \N. N \ set (\ \) \ \ derives \ [N] []" definition Init\<^sub>L :: "'a cfg \ 'a sentence \ 'a bins" where "Init\<^sub>L \ \ \ let rs = filter (\r. rule_head r = \ \) (\ \) in let b0 = map (\r. (Entry (init_item r 0) Null)) rs in let bs = replicate (length \ + 1) ([]) in bs[0 := b0]" definition Scan\<^sub>L :: "nat \ 'a sentence \ 'a \ 'a item \ nat \ 'a entry list" where "Scan\<^sub>L k \ a x pre \ if \!k = a then let x' = inc_item x (k+1) in [Entry x' (Pre pre)] else []" definition Predict\<^sub>L :: "nat \ 'a cfg \ 'a \ 'a entry list" where "Predict\<^sub>L k \ X \ let rs = filter (\r. rule_head r = X) (\ \) in map (\r. (Entry (init_item r k) Null)) rs" definition Complete\<^sub>L :: "nat \ 'a item \ 'a bins \ nat \ 'a entry list" where "Complete\<^sub>L k y bs red \ let orig = bs ! (item_origin y) in let is = filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items orig) in map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) is" fun bin_upd :: "'a entry \ 'a bin \ 'a bin" where "bin_upd e' [] = [e']" | "bin_upd e' (e#es) = ( case (e', e) of (Entry x (PreRed px xs), Entry y (PreRed py ys)) \ if x = y then Entry x (PreRed py (px#xs@ys)) # es else e # bin_upd e' es | _ \ if item e' = item e then e # es else e # bin_upd e' es)" fun bin_upds :: "'a entry list \ 'a bin \ 'a bin" where "bin_upds [] b = b" | "bin_upds (e#es) b = bin_upds es (bin_upd e b)" definition bins_upd :: "'a bins \ nat \ 'a entry list \ 'a bins" where "bins_upd bs k es \ bs[k := bin_upds es (bs!k)]" partial_function (tailrec) Earley\<^sub>L_bin' :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ nat \ 'a bins" where "Earley\<^sub>L_bin' k \ \ bs i = ( if i \ length (items (bs ! k)) then bs else let x = items (bs!k) ! i in let bs' = case next_symbol x of Some a \ if is_terminal \ a then if k < length \ then bins_upd bs (k+1) (Scan\<^sub>L k \ a x i) else bs else bins_upd bs k (Predict\<^sub>L k \ a) | None \ bins_upd bs k (Complete\<^sub>L k x bs i) in Earley\<^sub>L_bin' k \ \ bs' (i+1))" declare Earley\<^sub>L_bin'.simps[code] definition Earley\<^sub>L_bin :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ 'a bins" where "Earley\<^sub>L_bin k \ \ bs \ Earley\<^sub>L_bin' k \ \ bs 0" fun Earley\<^sub>L_bins :: "nat \ 'a cfg \ 'a sentence \ 'a bins" where "Earley\<^sub>L_bins 0 \ \ = Earley\<^sub>L_bin 0 \ \ (Init\<^sub>L \ \)" | "Earley\<^sub>L_bins (Suc n) \ \ = Earley\<^sub>L_bin (Suc n) \ \ (Earley\<^sub>L_bins n \ \)" definition Earley\<^sub>L :: "'a cfg \ 'a sentence \ 'a bins" where "Earley\<^sub>L \ \ \ Earley\<^sub>L_bins (length \) \ \" subsection \Bin lemmas\ lemma length_bins_upd[simp]: "length (bins_upd bs k es) = length bs" unfolding bins_upd_def by simp lemma length_bin_upd: "length (bin_upd e b) \ length b" by (induction e b rule: bin_upd.induct) (auto split: pointer.splits entry.splits) lemma length_bin_upds: "length (bin_upds es b) \ length b" by (induction es arbitrary: b) (auto, meson le_trans length_bin_upd) lemma length_nth_bin_bins_upd: "length (bins_upd bs k es ! n) \ length (bs ! n)" unfolding bins_upd_def using length_bin_upds by (metis linorder_not_le list_update_beyond nth_list_update_eq nth_list_update_neq order_refl) lemma nth_idem_bins_upd: "k \ n \ bins_upd bs k es ! n = bs ! n" unfolding bins_upd_def by simp lemma items_nth_idem_bin_upd: "n < length b \ items (bin_upd e b) ! n = items b ! n" by (induction b arbitrary: e n) (auto simp: items_def less_Suc_eq_0_disj split!: entry.split pointer.split) lemma items_nth_idem_bin_upds: "n < length b \ items (bin_upds es b) ! n = items b ! n" by (induction es arbitrary: b) (auto, metis items_def items_nth_idem_bin_upd length_bin_upd nth_map order.strict_trans2) lemma items_nth_idem_bins_upd: "n < length (bs ! k) \ items (bins_upd bs k es ! k) ! n = items (bs ! k) ! n" unfolding bins_upd_def using items_nth_idem_bin_upds by (metis linorder_not_less list_update_beyond nth_list_update_eq) lemma bin_upto_eq_set_items: "i \ length b \ bin_upto b i = set (items b)" by (auto simp: bin_upto_def items_def, metis in_set_conv_nth nth_map order_le_less order_less_trans) lemma bins_upto_empty: "bins_upto bs 0 0 = {}" unfolding bins_upto_def bin_upto_def by simp lemma set_items_bin_upd: "set (items (bin_upd e b)) = set (items b) \ {item e}" proof (induction b arbitrary: e) case (Cons b bs) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" by blast thus ?thesis using Cons.IH by (auto simp: items_def) next case False then show ?thesis proof cases assume *: "item e = item b" hence "bin_upd e (b # bs) = b # bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * by (auto simp: items_def) next assume *: "\ item e = item b" hence "bin_upd e (b # bs) = b # bin_upd e bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons.IH by (auto simp: items_def) qed qed qed (auto simp: items_def) lemma set_items_bin_upds: "set (items (bin_upds es b)) = set (items b) \ set (items es)" using set_items_bin_upd by (induction es arbitrary: b) (auto simp: items_def, blast, force+) lemma bins_bins_upd: assumes "k < length bs" shows "bins (bins_upd bs k es) = bins bs \ set (items es)" proof - let ?bs = "bins_upd bs k es" have "bins (bins_upd bs k es) = \ {set (items (?bs ! k)) |k. k < length ?bs}" unfolding bins_def by blast also have "... = \ {set (items (bs ! l)) |l. l < length bs \ l \ k} \ set (items (?bs ! k))" unfolding bins_upd_def using assms by (auto, metis nth_list_update) also have "... = \ {set (items (bs ! l)) |l. l < length bs \ l \ k} \ set (items (bs ! k)) \ set (items es)" using set_items_bin_upds[of es "bs!k"] by (simp add: assms bins_upd_def sup_assoc) also have "... = \ {set (items (bs ! k)) |k. k < length bs} \ set (items es)" using assms by blast also have "... = bins bs \ set (items es)" unfolding bins_def by blast finally show ?thesis . qed lemma kth_bin_sub_bins: "k < length bs \ set (items (bs ! k)) \ bins bs" unfolding bins_def bins_upto_def bin_upto_def by blast+ lemma bin_upto_Cons_0: "bin_upto (e#es) 0 = {}" by (auto simp: bin_upto_def) lemma bin_upto_Cons: assumes "0 < n" shows "bin_upto (e#es) n = { item e } \ bin_upto es (n-1)" proof - have "bin_upto (e#es) n = { items (e#es) ! j | j. j < n \ j < length (items (e#es)) }" unfolding bin_upto_def by blast also have "... = { item e } \ { items es ! j | j. j < (n-1) \ j < length (items es) }" using assms by (cases n) (auto simp: items_def nth_Cons', metis One_nat_def Zero_not_Suc diff_Suc_1 not_less_eq nth_map) also have "... = { item e } \ bin_upto es (n-1)" unfolding bin_upto_def by blast finally show ?thesis . qed lemma bin_upto_nth_idem_bin_upd: "n < length b \ bin_upto (bin_upd e b) n = bin_upto b n" proof (induction b arbitrary: e n) case (Cons b bs) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" by blast thus ?thesis using Cons bin_upto_Cons_0 by (cases n) (auto simp: items_def bin_upto_Cons, blast+) next case False then show ?thesis proof cases assume *: "item e = item b" hence "bin_upd e (b # bs) = b # bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * by (auto simp: items_def) next assume *: "\ item e = item b" hence "bin_upd e (b # bs) = b # bin_upd e bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons bin_upto_Cons_0 by (cases n) (auto simp: items_def bin_upto_Cons, blast+) qed qed qed (auto simp: items_def) lemma bin_upto_nth_idem_bin_upds: "n < length b \ bin_upto (bin_upds es b) n = bin_upto b n" using bin_upto_nth_idem_bin_upd length_bin_upd apply (induction es arbitrary: b) apply auto using order.strict_trans2 order.strict_trans1 by blast+ lemma bins_upto_kth_nth_idem: assumes "l < length bs" "k \ l" "n < length (bs ! k)" shows "bins_upto (bins_upd bs l es) k n = bins_upto bs k n" proof - let ?bs = "bins_upd bs l es" have "bins_upto ?bs k n = \ {set (items (?bs ! l)) |l. l < k} \ bin_upto (?bs ! k) n" unfolding bins_upto_def by blast also have "... = \ {set (items (bs ! l)) |l. l < k} \ bin_upto (?bs ! k) n" unfolding bins_upd_def using assms(1,2) by auto also have "... = \ {set (items (bs ! l)) |l. l < k} \ bin_upto (bs ! k) n" unfolding bins_upd_def using assms(1,3) bin_upto_nth_idem_bin_upds by (metis (no_types, lifting) nth_list_update) also have "... = bins_upto bs k n" unfolding bins_upto_def by blast finally show ?thesis . qed lemma bins_upto_sub_bins: "k < length bs \ bins_upto bs k n \ bins bs" unfolding bins_def bins_upto_def bin_upto_def using less_trans by (auto, blast) lemma bins_upto_Suc_Un: "n < length (bs ! k) \ bins_upto bs k (n+1) = bins_upto bs k n \ { items (bs ! k) ! n }" unfolding bins_upto_def bin_upto_def using less_Suc_eq by (auto simp: items_def, metis nth_map) lemma bins_bin_exists: "x \ bins bs \ \k < length bs. x \ set (items (bs ! k))" unfolding bins_def by blast lemma distinct_bin_upd: "distinct (items b) \ distinct (items (bin_upd e b))" proof (induction b arbitrary: e) case (Cons b bs) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" by blast thus ?thesis using Cons apply (auto simp: items_def) by (metis Un_insert_right entry.sel(1) imageI items_def list.set_map list.simps(15) set_ConsD set_items_bin_upd sup_bot_right) next case False then show ?thesis proof cases assume *: "item e = item b" hence "bin_upd e (b # bs) = b # bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons.prems by (auto simp: items_def) next assume *: "\ item e = item b" hence "bin_upd e (b # bs) = b # bin_upd e bs" using False by (auto split: pointer.splits entry.splits) moreover have "distinct (items (bin_upd e bs))" using Cons by (auto simp: items_def) ultimately show ?thesis using * Cons.prems set_items_bin_upd by (metis Un_insert_right distinct.simps(2) insertE items_def list.simps(9) sup_bot_right) qed qed qed (auto simp: items_def) lemma wf_bins_kth_bin: "wf_bins \ \ bs \ k < length bs \ x \ set (items (bs ! k)) \ wf_item \ \ x \ item_end x = k" using wf_bin_def wf_bins_def wf_bin_items_def by blast lemma wf_bin_bin_upd: assumes "wf_bin \ \ k b" "wf_item \ \ (item e) \ item_end (item e) = k" shows "wf_bin \ \ k (bin_upd e b)" using assms proof (induction b arbitrary: e) case (Cons b bs) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" by blast thus ?thesis using Cons distinct_bin_upd wf_bin_def wf_bin_items_def set_items_bin_upd by (smt (verit, best) Un_insert_right insertE sup_bot.right_neutral) next case False then show ?thesis proof cases assume *: "item e = item b" hence "bin_upd e (b # bs) = b # bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons.prems by (auto simp: items_def) next assume *: "\ item e = item b" hence "bin_upd e (b # bs) = b # bin_upd e bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons.prems set_items_bin_upd distinct_bin_upd wf_bin_def wf_bin_items_def by (smt (verit, best) Un_insert_right insertE sup_bot_right) qed qed qed (auto simp: items_def wf_bin_def wf_bin_items_def) lemma wf_bin_bin_upds: assumes "wf_bin \ \ k b" "distinct (items es)" assumes "\x \ set (items es). wf_item \ \ x \ item_end x = k" shows "wf_bin \ \ k (bin_upds es b)" using assms by (induction es arbitrary: b) (auto simp: wf_bin_bin_upd items_def) lemma wf_bins_bins_upd: assumes "wf_bins \ \ bs" "distinct (items es)" assumes "\x \ set (items es). wf_item \ \ x \ item_end x = k" shows "wf_bins \ \ (bins_upd bs k es)" unfolding bins_upd_def using assms wf_bin_bin_upds wf_bins_def by (metis length_list_update nth_list_update_eq nth_list_update_neq) lemma wf_bins_impl_wf_items: "wf_bins \ \ bs \ \x \ (bins bs). wf_item \ \ x" unfolding wf_bins_def wf_bin_def wf_bin_items_def bins_def by auto lemma bin_upds_eq_items: "set (items es) \ set (items b) \ set (items (bin_upds es b)) = set (items b)" apply (induction es arbitrary: b) apply (auto simp: set_items_bin_upd set_items_bin_upds) apply (simp add: items_def) by (metis Un_iff Un_subset_iff items_def list.simps(9) set_subset_Cons) lemma bin_eq_items_bin_upd: "item e \ set (items b) \ items (bin_upd e b) = items b" proof (induction b arbitrary: e) case (Cons b bs) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" by blast thus ?thesis using Cons by (auto simp: items_def) next case False then show ?thesis proof cases assume *: "item e = item b" hence "bin_upd e (b # bs) = b # bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons.prems by (auto simp: items_def) next assume *: "\ item e = item b" hence "bin_upd e (b # bs) = b # bin_upd e bs" using False by (auto split: pointer.splits entry.splits) thus ?thesis using * Cons by (auto simp: items_def) qed qed qed (auto simp: items_def) lemma bin_eq_items_bin_upds: assumes "set (items es) \ set (items b)" shows "items (bin_upds es b) = items b" using assms proof (induction es arbitrary: b) case (Cons e es) have "items (bin_upds es (bin_upd e b)) = items (bin_upd e b)" using Cons bin_upds_eq_items set_items_bin_upd set_items_bin_upds by (metis Un_upper2 bin_upds.simps(2) sup.coboundedI1) moreover have "items (bin_upd e b) = items b" using bin_eq_items_bin_upd Cons.prems by (auto simp: items_def) ultimately show ?case by simp qed (auto simp: items_def) lemma bins_eq_items_bins_upd: assumes "set (items es) \ set (items (bs!k))" shows "bins_eq_items (bins_upd bs k es) bs" unfolding bins_upd_def using assms bin_eq_items_bin_upds bins_eq_items_def by (metis list_update_id map_update) lemma bins_eq_items_imp_eq_bins: "bins_eq_items bs bs' \ bins bs = bins bs'" unfolding bins_eq_items_def bins_def items_def by (metis (no_types, lifting) length_map nth_map) lemma bin_eq_items_dist_bin_upd_bin: assumes "items a = items b" shows "items (bin_upd e a) = items (bin_upd e b)" using assms proof (induction a arbitrary: e b) case (Cons a as) obtain b' bs where bs: "b = b' # bs" "item a = item b'" "items as = items bs" using Cons.prems by (auto simp: items_def) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ a = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where #: "e = Entry x (PreRed xp xs)" "a = Entry y (PreRed yp ys)" by blast show ?thesis proof cases assume *: "x = y" hence "items (bin_upd e (a # as)) = x # items as" using # by (auto simp: items_def) moreover have "items (bin_upd e (b' # bs)) = x # items bs" using bs # * by (auto simp: items_def split: pointer.splits entry.splits) ultimately show ?thesis using bs by simp next assume *: "\ x = y" hence "items (bin_upd e (a # as)) = y # items (bin_upd e as)" using # by (auto simp: items_def) moreover have "items (bin_upd e (b' # bs)) = y # items (bin_upd e bs)" using bs # * by (auto simp: items_def split: pointer.splits entry.splits) ultimately show ?thesis using bs Cons.IH by simp qed next case False then show ?thesis proof cases assume *: "item e = item a" hence "items (bin_upd e (a # as)) = item a # items as" using False by (auto simp: items_def split: pointer.splits entry.splits) moreover have "items (bin_upd e (b' # bs)) = item b' # items bs" using bs False * by (auto simp: items_def split: pointer.splits entry.splits) ultimately show ?thesis using bs by simp next assume *: "\ item e = item a" hence "items (bin_upd e (a # as)) = item a # items (bin_upd e as)" using False by (auto simp: items_def split: pointer.splits entry.splits) moreover have "items (bin_upd e (b' # bs)) = item b' # items (bin_upd e bs)" using bs False * by (auto simp: items_def split: pointer.splits entry.splits) ultimately show ?thesis using bs Cons by simp qed qed qed (auto simp: items_def) lemma bin_eq_items_dist_bin_upds_bin: assumes "items a = items b" shows "items (bin_upds es a) = items (bin_upds es b)" using assms proof (induction es arbitrary: a b) case (Cons e es) hence "items (bin_upds es (bin_upd e a)) = items (bin_upds es (bin_upd e b))" using bin_eq_items_dist_bin_upd_bin by blast thus ?case by simp qed simp lemma bin_eq_items_dist_bin_upd_entry: assumes "item e = item e'" shows "items (bin_upd e b) = items (bin_upd e' b)" using assms proof (induction b arbitrary: e e') case (Cons a as) show ?case proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ a = Entry y (PreRed yp ys)") case True then obtain x xp xs y yp ys where #: "e = Entry x (PreRed xp xs)" "a = Entry y (PreRed yp ys)" by blast show ?thesis proof cases assume *: "x = y" thus ?thesis using # Cons.prems by (auto simp: items_def split: pointer.splits entry.splits) next assume *: "\ x = y" thus ?thesis using # Cons.prems by (auto simp: items_def split!: pointer.splits entry.splits, metis Cons.IH Cons.prems items_def)+ qed next case False then show ?thesis proof cases assume *: "item e = item a" thus ?thesis using Cons.prems by (auto simp: items_def split: pointer.splits entry.splits) next assume *: "\ item e = item a" thus ?thesis using Cons.prems by (auto simp: items_def split!: pointer.splits entry.splits, metis Cons.IH Cons.prems items_def)+ qed qed qed (auto simp: items_def) lemma bin_eq_items_dist_bin_upds_entries: assumes "items es = items es'" shows "items (bin_upds es b) = items (bin_upds es' b)" using assms proof (induction es arbitrary: es' b) case (Cons e es) then obtain e' es'' where "item e = item e'" "items es = items es''" "es' = e' # es''" by (auto simp: items_def) hence "items (bin_upds es (bin_upd e b)) = items (bin_upds es'' (bin_upd e' b))" using Cons.IH by (metis bin_eq_items_dist_bin_upd_entry bin_eq_items_dist_bin_upds_bin) thus ?case by (simp add: \es' = e' # es''\) qed (auto simp: items_def) lemma bins_eq_items_dist_bins_upd: assumes "bins_eq_items as bs" "items aes = items bes" "k < length as" shows "bins_eq_items (bins_upd as k aes) (bins_upd bs k bes)" proof - have "k < length bs" using assms(1,3) bins_eq_items_def map_eq_imp_length_eq by metis hence "items (bin_upds (as!k) aes) = items (bin_upds (bs!k) bes)" using bin_eq_items_dist_bin_upds_entries bin_eq_items_dist_bin_upds_bin bins_eq_items_def assms by (metis (no_types, lifting) nth_map) thus ?thesis using \k < length bs\ assms bin_eq_items_dist_bin_upds_bin bin_eq_items_dist_bin_upds_entries bins_eq_items_def bins_upd_def by (smt (verit) map_update nth_map) qed subsection \Well-formed bins\ lemma distinct_Scan\<^sub>L: "distinct (items (Scan\<^sub>L k \ a x pre))" unfolding Scan\<^sub>L_def by (auto simp: items_def) lemma distinct_Predict\<^sub>L: "wf_\ \ \ distinct (items (Predict\<^sub>L k \ X))" unfolding Predict\<^sub>L_def wf_\_defs by (auto simp: init_item_def rule_head_def distinct_map inj_on_def items_def) lemma inj_on_inc_item: "\x \ A. item_end x = l \ inj_on (\x. inc_item x k) A" unfolding inj_on_def inc_item_def by (simp add: item.expand) lemma distinct_Complete\<^sub>L: assumes "wf_bins \ \ bs" "item_origin y < length bs" shows "distinct (items (Complete\<^sub>L k y bs red))" proof - let ?orig = "bs ! (item_origin y)" let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" let ?is' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?is" have wf: "wf_bin \ \ (item_origin y) ?orig" using assms wf_bins_def by blast have 0: "\x \ set (map fst ?is). item_end x = (item_origin y)" using wf wf_bin_def wf_bin_items_def filter_is_subset filter_with_index_cong_filter by (metis in_mono) hence "distinct (items ?orig)" using wf unfolding wf_bin_def by blast hence "distinct (map fst ?is)" using filter_with_index_cong_filter distinct_filter by metis moreover have "items ?is' = map (\x. inc_item x k) (map fst ?is)" by (induction ?is) (auto simp: items_def) moreover have "inj_on (\x. inc_item x k) (set (map fst ?is))" using inj_on_inc_item 0 by blast ultimately have "distinct (items ?is')" using distinct_map by metis thus ?thesis unfolding Complete\<^sub>L_def by simp qed lemma wf_bins_Scan\<^sub>L': assumes "wf_bins \ \ bs" "k < length bs" "x \ set (items (bs ! k))" assumes "k < length \" "next_symbol x \ None" "y = inc_item x (k+1)" shows "wf_item \ \ y \ item_end y = k+1" using assms wf_bins_kth_bin[OF assms(1-3)] unfolding wf_item_def inc_item_def next_symbol_def is_complete_def item_rule_body_def by (auto split: if_splits) lemma wf_bins_Scan\<^sub>L: assumes "wf_bins \ \ bs" "k < length bs" "x \ set (items (bs ! k))" "k < length \" "next_symbol x \ None" shows "\y \ set (items (Scan\<^sub>L k \ a x pre)). wf_item \ \ y \ item_end y = (k+1)" using wf_bins_Scan\<^sub>L'[OF assms] by (simp add: Scan\<^sub>L_def items_def) lemma wf_bins_Predict\<^sub>L: assumes "wf_bins \ \ bs" "k < length bs" "k \ length \" "wf_\ \" shows "\y \ set (items (Predict\<^sub>L k \ X)). wf_item \ \ y \ item_end y = k" using assms by (auto simp: Predict\<^sub>L_def wf_item_def wf_bins_def wf_bin_def init_item_def wf_\_defs items_def) lemma wf_item_inc_item: assumes "wf_item \ \ x" "next_symbol x = Some a" "item_origin x \ k" "k \ length \" shows "wf_item \ \ (inc_item x k) \ item_end (inc_item x k) = k" using assms by (auto simp: wf_item_def inc_item_def item_rule_body_def next_symbol_def is_complete_def split: if_splits) lemma wf_bins_Complete\<^sub>L: assumes "wf_bins \ \ bs" "k < length bs" "y \ set (items (bs ! k))" shows "\x \ set (items (Complete\<^sub>L k y bs red)). wf_item \ \ x \ item_end x = k" proof - let ?orig = "bs ! (item_origin y)" let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" let ?is' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?is" { fix x assume *: "x \ set (map fst ?is)" have "item_end x = item_origin y" using * assms wf_bins_kth_bin wf_item_def filter_with_index_cong_filter by (metis dual_order.strict_trans2 filter_is_subset subsetD) have "wf_item \ \ x" using * assms wf_bins_kth_bin wf_item_def filter_with_index_cong_filter by (metis dual_order.strict_trans2 filter_is_subset subsetD) moreover have "next_symbol x = Some (item_rule_head y)" using * filter_set filter_with_index_cong_filter member_filter by metis moreover have "item_origin x \ k" using \item_end x = item_origin y\ \wf_item \ \ x\ assms wf_bins_kth_bin wf_item_def by (metis dual_order.order_iff_strict dual_order.strict_trans1) moreover have "k \ length \" using assms wf_bins_kth_bin wf_item_def by blast ultimately have "wf_item \ \ (inc_item x k)" "item_end (inc_item x k) = k" by (simp_all add: wf_item_inc_item) } hence "\x \ set (items ?is'). wf_item \ \ x \ item_end x = k" by (auto simp: items_def rev_image_eqI) thus ?thesis unfolding Complete\<^sub>L_def by presburger qed lemma Ex_wf_bins: "\n bs \ \. n \ length \ \ length bs = Suc (length \) \ wf_\ \ \ wf_bins \ \ bs" apply (rule exI[where x="0"]) apply (rule exI[where x="[[]]"]) apply (rule exI[where x="[]"]) apply (auto simp: wf_bins_def wf_bin_def wf_\_defs wf_bin_items_def items_def split: prod.splits) by (metis cfg.sel distinct.simps(1) empty_iff empty_set inf_bot_right list.set_intros(1)) definition wf_earley_input :: "(nat \ 'a cfg \ 'a sentence \ 'a bins) set" where "wf_earley_input = { (k, \, \, bs) | k \ \ bs. k \ length \ \ length bs = length \ + 1 \ wf_\ \ \ wf_bins \ \ bs }" typedef 'a wf_bins = "wf_earley_input::(nat \ 'a cfg \ 'a sentence \ 'a bins) set" morphisms from_wf_bins to_wf_bins using Ex_wf_bins by (auto simp: wf_earley_input_def) lemma wf_earley_input_elim: assumes "(k, \, \, bs) \ wf_earley_input" shows "k \ length \ \ k < length bs \ length bs = length \ + 1 \ wf_\ \ \ wf_bins \ \ bs" using assms(1) from_wf_bins wf_earley_input_def by (smt (verit) Suc_eq_plus1 less_Suc_eq_le mem_Collect_eq prod.sel(1) snd_conv) lemma wf_earley_input_intro: assumes "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" shows "(k, \, \, bs) \ wf_earley_input" by (simp add: assms wf_earley_input_def) lemma wf_earley_input_Complete\<^sub>L: assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" assumes "x = items (bs ! k) ! i" "next_symbol x = None" shows "(k, \, \, bins_upd bs k (Complete\<^sub>L k x bs red)) \ wf_earley_input" proof - have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" using wf_earley_input_elim assms(1) by metis+ have x: "x \ set (items (bs ! k))" using assms(2,3) by simp have "item_origin x < length bs" using x wf_bins_kth_bin *(1,2,4) wf_item_def by (metis One_nat_def add.right_neutral add_Suc_right dual_order.trans le_imp_less_Suc) hence "wf_bins \ \ (bins_upd bs k (Complete\<^sub>L k x bs red))" using *(1,2,4) Suc_eq_plus1 distinct_Complete\<^sub>L le_imp_less_Suc wf_bins_Complete\<^sub>L wf_bins_bins_upd x by metis thus ?thesis by (simp add: *(1-3) wf_earley_input_def) qed lemma wf_earley_input_Scan\<^sub>L: assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" assumes "x = items (bs ! k) ! i" "next_symbol x = Some a" assumes "is_terminal \ a" "k < length \" shows "(k, \, \, bins_upd bs (k+1) (Scan\<^sub>L k \ a x pre)) \ wf_earley_input" proof - have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" using wf_earley_input_elim assms(1) by metis+ have x: "x \ set (items(bs ! k))" using assms(2,3) by simp have "wf_bins \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x pre))" using * x assms(1,4,6) distinct_Scan\<^sub>L wf_bins_Scan\<^sub>L wf_bins_bins_upd wf_earley_input_elim by (metis option.discI) thus ?thesis by (simp add: *(1-3) wf_earley_input_def) qed lemma wf_earley_input_Predict\<^sub>L: assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" assumes "x = items (bs ! k) ! i" "next_symbol x = Some a" "\ is_terminal \ a" shows "(k, \, \, bins_upd bs k (Predict\<^sub>L k \ a)) \ wf_earley_input" proof - have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" using wf_earley_input_elim assms(1) by metis+ have x: "x \ set (items (bs ! k))" using assms(2,3) by simp hence "wf_bins \ \ (bins_upd bs k (Predict\<^sub>L k \ a))" using * x assms(1,4) distinct_Predict\<^sub>L wf_bins_Predict\<^sub>L wf_bins_bins_upd wf_earley_input_elim by metis thus ?thesis by (simp add: *(1-3) wf_earley_input_def) qed fun earley_measure :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ nat \ nat" where "earley_measure (k, \, \, bs) i = card { x | x. wf_item \ \ x \ item_end x = k } - i" lemma Earley\<^sub>L_bin'_simps[simp]: "i \ length (items (bs ! k)) \ Earley\<^sub>L_bin' k \ \ bs i = bs" "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = None \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs k (Complete\<^sub>L k x bs i)) (i+1)" "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ is_terminal \ a \ k < length \ \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)) (i+1)" "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ is_terminal \ a \ \ k < length \ \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ bs (i+1)" "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ \ is_terminal \ a \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs k (Predict\<^sub>L k \ a)) (i+1)" by (subst Earley\<^sub>L_bin'.simps, simp)+ lemma Earley\<^sub>L_bin'_induct[case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]: assumes "(k, \, \, bs) \ wf_earley_input" assumes base: "\k \ \ bs i. i \ length (items (bs ! k)) \ P k \ \ bs i" assumes complete: "\k \ \ bs i x. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ next_symbol x = None \ P k \ \ (bins_upd bs k (Complete\<^sub>L k x bs i)) (i+1) \ P k \ \ bs i" assumes scan: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ next_symbol x = Some a \ is_terminal \ a \ k < length \ \ P k \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)) (i+1) \ P k \ \ bs i" assumes pass: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ next_symbol x = Some a \ is_terminal \ a \ \ k < length \ \ P k \ \ bs (i+1) \ P k \ \ bs i" assumes predict: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ next_symbol x = Some a \ \ is_terminal \ a \ P k \ \ (bins_upd bs k (Predict\<^sub>L k \ a)) (i+1) \ P k \ \ bs i" shows "P k \ \ bs i" using assms(1) proof (induction n\"earley_measure (k, \, \, bs) i" arbitrary: bs i rule: nat_less_induct) case 1 have wf: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" using "1.prems" wf_earley_input_elim by metis+ hence k: "k < length bs" by simp have fin: "finite { x | x. wf_item \ \ x \ item_end x = k }" using finiteness_UNIV_wf_item by fastforce show ?case proof cases assume "i \ length (items (bs ! k))" then show ?thesis by (simp add: base) next assume a1: "\ i \ length (items (bs ! k))" let ?x = "items (bs ! k) ! i" have x: "?x \ set (items (bs ! k))" using a1 by fastforce show ?thesis proof cases assume a2: "next_symbol ?x = None" let ?bs' = "bins_upd bs k (Complete\<^sub>L k ?x bs i)" have "item_origin ?x < length bs" using wf(4) k wf_bins_kth_bin wf_item_def x by (metis order_le_less_trans) hence wf_bins': "wf_bins \ \ ?bs'" using wf_bins_Complete\<^sub>L distinct_Complete\<^sub>L wf(4) wf_bins_bins_upd k x by metis hence wf': "(k, \, \, ?bs') \ wf_earley_input" using wf(1,2,3) wf_earley_input_intro by fastforce have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto have "i < length (items (?bs' ! k))" using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) also have "... = card (set (items (?bs' ! k)))" using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def by (metis k length_bins_upd) also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" using card_mono fin sub by blast finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" by blast hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" by simp thus ?thesis using 1 a1 a2 complete wf' by simp next assume a2: "\ next_symbol ?x = None" then obtain a where a_def: "next_symbol ?x = Some a" by blast show ?thesis proof cases assume a3: "is_terminal \ a" show ?thesis proof cases assume a4: "k < length \" let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a ?x i)" have wf_bins': "wf_bins \ \ ?bs'" using wf_bins_Scan\<^sub>L distinct_Scan\<^sub>L wf(1,4) wf_bins_bins_upd a2 a4 k x by metis hence wf': "(k, \, \, ?bs') \ wf_earley_input" using wf(1,2,3) wf_earley_input_intro by fastforce have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto have "i < length (items (?bs' ! k))" using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) also have "... = card (set (items (?bs' ! k)))" using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def by (metis Suc_eq_plus1 le_imp_less_Suc length_bins_upd) also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" using card_mono fin sub by blast finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" by blast hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" by simp thus ?thesis using 1 a1 a_def a3 a4 scan wf' by simp next assume a4: "\ k < length \" have sub: "set (items (bs ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" using wf(1,2,4) unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto have "i < length (items (bs ! k))" using a1 by simp also have "... = card (set (items (bs ! k)))" using wf(1,2,4) distinct_card wf_bins_def wf_bin_def by (metis Suc_eq_plus1 le_imp_less_Suc) also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" using card_mono fin sub by blast finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" by blast hence "earley_measure (k, \, \, bs) (Suc i) < earley_measure (k, \, \, bs) i" by simp thus ?thesis using 1 a1 a3 a4 a_def pass by simp qed next assume a3: "\ is_terminal \ a" let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have wf_bins': "wf_bins \ \ ?bs'" using wf_bins_Predict\<^sub>L distinct_Predict\<^sub>L wf(1,3,4) wf_bins_bins_upd k x by metis hence wf': "(k, \, \, ?bs') \ wf_earley_input" using wf(1,2,3) wf_earley_input_intro by fastforce have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto have "i < length (items (?bs' ! k))" using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) also have "... = card (set (items (?bs' ! k)))" using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def by (metis Suc_eq_plus1 le_imp_less_Suc length_bins_upd) also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" using card_mono fin sub by blast finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" by blast hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" by simp thus ?thesis using 1 a1 a_def a3 a_def predict wf' by simp qed qed qed qed lemma wf_earley_input_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" shows "(k, \, \, Earley\<^sub>L_bin' k \ \ bs i) \ wf_earley_input" using assms proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems wf_earley_input_Complete\<^sub>L by blast thus ?case using Complete\<^sub>F.IH Complete\<^sub>F.hyps by simp next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems wf_earley_input_Scan\<^sub>L by metis thus ?case using Scan\<^sub>F.IH Scan\<^sub>F.hyps by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems wf_earley_input_Predict\<^sub>L by metis thus ?case using Predict\<^sub>F.IH Predict\<^sub>F.hyps by simp qed simp_all lemma wf_earley_input_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" shows "(k, \, \, Earley\<^sub>L_bin k \ \ bs) \ wf_earley_input" using assms by (simp add: Earley\<^sub>L_bin_def wf_earley_input_Earley\<^sub>L_bin') lemma length_bins_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" shows "length (Earley\<^sub>L_bin' k \ \ bs i) = length bs" by (metis assms wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_elim) lemma length_nth_bin_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" shows "length (items (Earley\<^sub>L_bin' k \ \ bs i ! l)) \ length (items (bs ! l))" using length_nth_bin_bins_upd order_trans by (induction i rule: Earley\<^sub>L_bin'_induct[OF assms]) (auto simp: items_def, blast+) lemma wf_bins_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" shows "wf_bins \ \ (Earley\<^sub>L_bin' k \ \ bs i)" using assms wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_elim by blast lemma wf_bins_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" shows "wf_bins \ \ (Earley\<^sub>L_bin k \ \ bs)" using assms Earley\<^sub>L_bin_def wf_bins_Earley\<^sub>L_bin' by metis lemma kth_Earley\<^sub>L_bin'_bins: assumes "(k, \, \, bs) \ wf_earley_input" assumes "j < length (items (bs ! l))" shows "items (Earley\<^sub>L_bin' k \ \ bs i ! l) ! j = items (bs ! l) ! j" using assms(2) proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" using Complete\<^sub>F.IH Complete\<^sub>F.prems length_nth_bin_bins_upd items_def order.strict_trans2 by (metis length_map) also have "... = items (bs ! l) ! j" using Complete\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis finally show ?case using Complete\<^sub>F.hyps by simp next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" using Scan\<^sub>F.IH Scan\<^sub>F.prems length_nth_bin_bins_upd order.strict_trans2 items_def by (metis length_map) also have "... = items (bs ! l) ! j" using Scan\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis finally show ?case using Scan\<^sub>F.hyps by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" using Predict\<^sub>F.IH Predict\<^sub>F.prems length_nth_bin_bins_upd order.strict_trans2 items_def by (metis length_map) also have "... = items (bs ! l) ! j" using Predict\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis finally show ?case using Predict\<^sub>F.hyps by simp qed simp_all lemma nth_bin_sub_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" shows "set (items (bs ! l)) \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" proof standard fix x assume "x \ set (items (bs ! l))" then obtain j where *: "j < length (items (bs ! l))" "items (bs ! l) ! j = x" using in_set_conv_nth by metis have "x = items (Earley\<^sub>L_bin' k \ \ bs i ! l) ! j" using kth_Earley\<^sub>L_bin'_bins assms * by metis moreover have "j < length (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" using assms *(1) length_nth_bin_Earley\<^sub>L_bin' less_le_trans by blast ultimately show "x \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" by simp qed lemma nth_Earley\<^sub>L_bin'_eq: assumes "(k, \, \, bs) \ wf_earley_input" shows "l < k \ Earley\<^sub>L_bin' k \ \ bs i ! l = bs ! l" by (induction i rule: Earley\<^sub>L_bin'_induct[OF assms]) (auto simp: bins_upd_def) lemma set_items_Earley\<^sub>L_bin'_eq: assumes "(k, \, \, bs) \ wf_earley_input" shows "l < k \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l)) = set (items (bs ! l))" by (simp add: assms nth_Earley\<^sub>L_bin'_eq) lemma bins_upto_k0_Earley\<^sub>L_bin'_eq: assumes "(k, \, \, bs) \ wf_earley_input" shows "bins_upto (Earley\<^sub>L_bin k \ \ bs) k 0 = bins_upto bs k 0" unfolding bins_upto_def bin_upto_def Earley\<^sub>L_bin_def using set_items_Earley\<^sub>L_bin'_eq assms nth_Earley\<^sub>L_bin'_eq by fastforce lemma wf_earley_input_Init\<^sub>L: assumes "k \ length \" "wf_\ \" shows "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" proof - let ?rs = "filter (\r. rule_head r = \ \) (\ \)" let ?b0 = "map (\r. (Entry (init_item r 0) Null)) ?rs" let ?bs = "replicate (length \ + 1) ([])" have "distinct (items ?b0)" using assms unfolding wf_bin_def wf_item_def wf_\_def distinct_rules_def items_def by (auto simp: init_item_def distinct_map inj_on_def) moreover have "\x \ set (items ?b0). wf_item \ \ x \ item_end x = 0" using assms unfolding wf_bin_def wf_item_def by (auto simp: init_item_def items_def) moreover have "wf_bins \ \ ?bs" unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def using less_Suc_eq_0_disj by force ultimately show ?thesis using assms length_replicate wf_earley_input_intro unfolding wf_bin_def Init\<^sub>L_def wf_bin_def wf_bin_items_def wf_bins_def by (metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq) qed lemma length_bins_Init\<^sub>L[simp]: "length (Init\<^sub>L \ \) = length \ + 1" by (simp add: Init\<^sub>L_def) lemma wf_earley_input_Earley\<^sub>L_bins[simp]: assumes "k \ length \" "wf_\ \" shows "(k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" using assms proof (induction k) case 0 have "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" using assms wf_earley_input_Init\<^sub>L by blast thus ?case by (simp add: assms(2) wf_earley_input_Init\<^sub>L wf_earley_input_Earley\<^sub>L_bin) next case (Suc k) have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" using Suc.IH Suc.prems(1) Suc_leD assms(2) wf_earley_input_elim wf_earley_input_intro by metis thus ?case by (simp add: wf_earley_input_Earley\<^sub>L_bin) qed lemma length_Earley\<^sub>L_bins[simp]: assumes "k \ length \" "wf_\ \" shows "length (Earley\<^sub>L_bins k \ \) = length (Init\<^sub>L \ \)" using assms wf_earley_input_Earley\<^sub>L_bins wf_earley_input_elim by fastforce lemma wf_bins_Earley\<^sub>L_bins[simp]: assumes "k \ length \" "wf_\ \" shows "wf_bins \ \ (Earley\<^sub>L_bins k \ \)" using assms wf_earley_input_Earley\<^sub>L_bins wf_earley_input_elim by fastforce lemma wf_bins_Earley\<^sub>L: "wf_\ \ \ wf_bins \ \ (Earley\<^sub>L \ \)" by (simp add: Earley\<^sub>L_def) subsection \Soundness\ lemma Init\<^sub>L_eq_Init\<^sub>F: "bins (Init\<^sub>L \ \) = Init\<^sub>F \" proof - let ?rs = "filter (\r. rule_head r = \ \) (\ \)" let ?b0 = "map (\r. (Entry (init_item r 0) Null)) ?rs" let ?bs = "replicate (length \ + 1) ([])" have "bins (?bs[0 := ?b0]) = set (items ?b0)" proof - have "bins (?bs[0 := ?b0]) = \ {set (items ((?bs[0 := ?b0]) ! k)) |k. k < length (?bs[0 := ?b0])}" unfolding bins_def by blast also have "... = set (items ((?bs[0 := ?b0]) ! 0)) \ \ {set (items ((?bs[0 := ?b0]) ! k)) |k. k < length (?bs[0 := ?b0]) \ k \ 0}" by fastforce also have "... = set (items (?b0))" by (auto simp: items_def) finally show ?thesis . qed also have "... = Init\<^sub>F \" by (auto simp: Init\<^sub>F_def items_def rule_head_def) finally show ?thesis by (auto simp: Init\<^sub>L_def) qed lemma Scan\<^sub>L_sub_Scan\<^sub>F: assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs ! k))" "k < length bs" "k < length \" assumes "next_symbol x = Some a" shows "set (items (Scan\<^sub>L k \ a x pre)) \ Scan\<^sub>F k \ I" proof standard fix y assume *: "y \ set (items (Scan\<^sub>L k \ a x pre))" have "x \ bin I k" using kth_bin_sub_bins assms(1-4) items_def wf_bin_def wf_bins_def wf_bin_items_def bin_def by fastforce { assume #: "k < length \" "\!k = a" hence "y = inc_item x (k+1)" using * unfolding Scan\<^sub>L_def by (simp add: items_def) hence "y \ Scan\<^sub>F k \ I" using \x \ bin I k\ # assms(6) unfolding Scan\<^sub>F_def by blast } thus "y \ Scan\<^sub>F k \ I" using * assms(5) unfolding Scan\<^sub>L_def by (auto simp: items_def) qed lemma Predict\<^sub>L_sub_Predict\<^sub>F: assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs ! k))" "k < length bs" assumes "next_symbol x = Some X" shows "set (items (Predict\<^sub>L k \ X)) \ Predict\<^sub>F k \ I" proof standard fix y assume *: "y \ set (items (Predict\<^sub>L k \ X))" have "x \ bin I k" using kth_bin_sub_bins assms(1-4) items_def wf_bin_def wf_bins_def bin_def wf_bin_items_def by fast let ?rs = "filter (\r. rule_head r = X) (\ \)" let ?xs = "map (\r. init_item r k) ?rs" have "y \ set ?xs" using * unfolding Predict\<^sub>L_def items_def by simp then obtain r where "y = init_item r k" "rule_head r = X" "r \ set (\ \)" "next_symbol x = Some (rule_head r)" using assms(5) by auto thus "y \ Predict\<^sub>F k \ I" unfolding Predict\<^sub>F_def using \x \ bin I k\ by blast qed lemma Complete\<^sub>L_sub_Complete\<^sub>F: assumes "wf_bins \ \ bs" "bins bs \ I" "y \ set (items (bs ! k))" "k < length bs" assumes "next_symbol y = None" shows "set (items (Complete\<^sub>L k y bs red)) \ Complete\<^sub>F k I" proof standard fix x assume *: "x \ set (items (Complete\<^sub>L k y bs red))" have "y \ bin I k" using kth_bin_sub_bins assms items_def wf_bin_def wf_bins_def bin_def wf_bin_items_def by fast let ?orig = "bs ! item_origin y" let ?xs = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" let ?xs' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?xs" have 0: "item_origin y < length bs" using wf_bins_def wf_bin_def wf_item_def wf_bin_items_def assms(1,3,4) by (metis Orderings.preorder_class.dual_order.strict_trans1 leD not_le_imp_less) { fix z assume *: "z \ set (map fst ?xs)" have "next_symbol z = Some (item_rule_head y)" using * by (simp add: filter_with_index_cong_filter) moreover have "z \ bin I (item_origin y)" using 0 * assms(1,2) bin_def kth_bin_sub_bins wf_bins_kth_bin filter_with_index_cong_filter by (metis (mono_tags, lifting) filter_is_subset in_mono mem_Collect_eq) ultimately have "next_symbol z = Some (item_rule_head y)" "z \ bin I (item_origin y)" by simp_all } hence 1: "\z \ set (map fst ?xs). next_symbol z = Some (item_rule_head y) \ z \ bin I (item_origin y)" by blast obtain z where z: "x = inc_item z k" "z \ set (map fst ?xs)" using * unfolding Complete\<^sub>L_def by (auto simp: rev_image_eqI items_def) moreover have "next_symbol z = Some (item_rule_head y)" "z \ bin I (item_origin y)" using 1 z by blast+ ultimately show "x \ Complete\<^sub>F k I" using \y \ bin I k\ assms(5) unfolding Complete\<^sub>F_def next_symbol_def by (auto split: if_splits) qed lemma sound_Scan\<^sub>L: assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs!k))" "k < length bs" "k < length \" assumes "next_symbol x = Some a" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" shows "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" proof standard fix y assume "y \ set (items (Scan\<^sub>L k \ a x i))" hence "y \ Scan\<^sub>F k \ I" by (meson Scan\<^sub>L_sub_Scan\<^sub>F assms(1-6) in_mono) thus "sound_item \ \ y" using sound_Scan assms(7,8) unfolding Scan\<^sub>F_def inc_item_def bin_def by (smt (verit, best) item.exhaust_sel mem_Collect_eq) qed lemma sound_Predict\<^sub>L: assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs!k))" "k < length bs" assumes "next_symbol x = Some X" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" shows "\x \ set (items (Predict\<^sub>L k \ X)). sound_item \ \ x" proof standard fix y assume "y \ set (items (Predict\<^sub>L k \ X))" hence "y \ Predict\<^sub>F k \ I" by (meson Predict\<^sub>L_sub_Predict\<^sub>F assms(1-5) subsetD) thus "sound_item \ \ y" using sound_Predict assms(6,7) unfolding Predict\<^sub>F_def init_item_def bin_def by (smt (verit, del_insts) item.exhaust_sel mem_Collect_eq) qed lemma sound_Complete\<^sub>L: assumes "wf_bins \ \ bs" "bins bs \ I" "y \ set (items (bs!k))" "k < length bs" assumes "next_symbol y = None" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" shows "\x \ set (items (Complete\<^sub>L k y bs i)). sound_item \ \ x" proof standard fix x assume "x \ set (items (Complete\<^sub>L k y bs i))" hence "x \ Complete\<^sub>F k I" using Complete\<^sub>L_sub_Complete\<^sub>F assms(1-5) by blast thus "sound_item \ \ x" using sound_Complete assms(6,7) unfolding Complete\<^sub>F_def inc_item_def bin_def by (smt (verit, del_insts) item.exhaust_sel mem_Collect_eq) qed lemma sound_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" assumes "\x \ bins bs. sound_item \ \ x" shows "\x \ bins (Earley\<^sub>L_bin' k \ \ bs i). sound_item \ \ x" using assms proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have "x \ set (items (bs ! k))" using Complete\<^sub>F.hyps(1,2) by force hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fastforce moreover have "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" using Complete\<^sub>F.IH Complete\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim Suc_eq_plus1 Un_iff le_imp_less_Suc by metis thus ?case using Complete\<^sub>F.hyps by simp next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "x \ set (items (bs ! k))" using Scan\<^sub>F.hyps(1,2) by force hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2) wf_earley_input_elim wf_bins_impl_wf_items by fast moreover have "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" using Scan\<^sub>F.IH Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim by (metis UnE add_less_cancel_right) thus ?case using Scan\<^sub>F.hyps by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "x \ set (items (bs ! k))" using Predict\<^sub>F.hyps(1,2) by force hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast moreover have "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" using Predict\<^sub>F.IH Predict\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim by (metis Suc_eq_plus1 UnE) thus ?case using Predict\<^sub>F.hyps by simp qed simp_all lemma sound_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "\x \ bins bs. sound_item \ \ x" shows "\x \ bins (Earley\<^sub>L_bin k \ \ bs). sound_item \ \ x" using sound_Earley\<^sub>L_bin' assms Earley\<^sub>L_bin_def by metis lemma Earley\<^sub>L_bin'_sub_Earley\<^sub>F_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "bins bs \ I" shows "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ I" using assms proof (induction i arbitrary: I rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Base k \ \ bs i) thus ?case using Earley\<^sub>F_bin_mono by fastforce next case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have "x \ set (items (bs ! k))" using Complete\<^sub>F.hyps(1,2) by force hence "bins ?bs' \ I \ Complete\<^sub>F k I" using Complete\<^sub>L_sub_Complete\<^sub>F Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems(1,2) bins_bins_upd wf_earley_input_elim by blast moreover have "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Complete\<^sub>F k I)" using Complete\<^sub>F.IH Complete\<^sub>F.hyps by simp also have "... \ Earley\<^sub>F_bin k \ \ (Earley\<^sub>F_bin k \ \ I)" using Complete\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono by (metis Un_subset_iff) finally show ?case using Earley\<^sub>F_bin_idem by blast next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "x \ set (items (bs ! k))" using Scan\<^sub>F.hyps(1,2) by force hence "bins ?bs' \ I \ Scan\<^sub>F k \ I" using Scan\<^sub>L_sub_Scan\<^sub>F Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems bins_bins_upd wf_earley_input_elim by (metis add_mono1 sup_mono) moreover have "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Scan\<^sub>F k \ I)" using Scan\<^sub>F.IH Scan\<^sub>F.hyps by simp thus ?case using Scan\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono Earley\<^sub>F_bin_idem by (metis le_supI order_trans) next case (Pass k \ \ bs i x a) thus ?case by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "x \ set (items (bs ! k))" using Predict\<^sub>F.hyps(1,2) by force hence "bins ?bs' \ I \ Predict\<^sub>F k \ I" using Predict\<^sub>L_sub_Predict\<^sub>F Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems bins_bins_upd wf_earley_input_elim by (metis sup_mono) moreover have "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Predict\<^sub>F k \ I)" using Predict\<^sub>F.IH Predict\<^sub>F.hyps by simp thus ?case using Predict\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono Earley\<^sub>F_bin_idem by (metis le_supI order_trans) qed lemma Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "bins bs \ I" shows "bins (Earley\<^sub>L_bin k \ \ bs) \ Earley\<^sub>F_bin k \ \ I" using assms Earley\<^sub>L_bin'_sub_Earley\<^sub>F_bin Earley\<^sub>L_bin_def by metis lemma Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins: assumes "k \ length \" "wf_\ \" shows "bins (Earley\<^sub>L_bins k \ \) \ Earley\<^sub>F_bins k \ \" using assms proof (induction k) case 0 have "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" using assms(1) assms(2) wf_earley_input_Init\<^sub>L by blast thus ?case by (simp add: Init\<^sub>L_eq_Init\<^sub>F Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin assms(2) wf_earley_input_Init\<^sub>L) next case (Suc k) have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) thus ?case by (simp add: Suc.IH Suc.prems(1) Suc_leD Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin assms(2)) qed lemma Earley\<^sub>L_sub_Earley\<^sub>F: "wf_\ \ \ bins (Earley\<^sub>L \ \) \ Earley\<^sub>F \ \" using Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins Earley\<^sub>F_def Earley\<^sub>L_def by (metis dual_order.refl) theorem soundness_Earley\<^sub>L: assumes "wf_\ \" "recognizing (bins (Earley\<^sub>L \ \)) \ \" shows "derives \ [\ \] \" using assms Earley\<^sub>L_sub_Earley\<^sub>F recognizing_def soundness_Earley\<^sub>F by (meson subsetD) subsection \Completeness\ lemma bin_bins_upto_bins_eq: assumes "wf_bins \ \ bs" "k < length bs" "i \ length (items (bs ! k))" "l \ k" shows "bin (bins_upto bs k i) l = bin (bins bs) l" unfolding bins_upto_def bins_def bin_def using assms nat_less_le apply (auto simp: nth_list_update bin_upto_eq_set_items wf_bins_kth_bin items_def) apply (metis imageI nle_le order_trans, fast) done lemma impossible_complete_item: assumes "wf_\ \" "wf_item \ \ x" "sound_item \ \ x" assumes "is_complete x" "item_origin x = k" "item_end x = k" "nonempty_derives \" shows False proof - have "derives \ [item_rule_head x] []" using assms(3-6) by (simp add: slice_empty is_complete_def sound_item_def item_\_def ) moreover have "is_nonterminal \ (item_rule_head x)" using assms(1,2) unfolding wf_item_def item_rule_head_def rule_head_def by (metis prod.collapse rule_nonterminal_type) ultimately show ?thesis using assms(7) nonempty_derives_def is_nonterminal_def by metis qed lemma Complete\<^sub>F_Un_eq_terminal: assumes "next_symbol z = Some a" "is_terminal \ a" "\x \ I. wf_item \ \ x" "wf_item \ \ z" "wf_\ \" shows "Complete\<^sub>F k (I \ {z}) = Complete\<^sub>F k I" proof (rule ccontr) assume "Complete\<^sub>F k (I \ {z}) \ Complete\<^sub>F k I" hence "Complete\<^sub>F k I \ Complete\<^sub>F k (I \ {z})" using Complete\<^sub>F_sub_mono by blast then obtain w x y where *: "w \ Complete\<^sub>F k (I \ {z})" "w \ Complete\<^sub>F k I" "w = inc_item x k" "x \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" "is_complete y" "next_symbol x = Some (item_rule_head y)" unfolding Complete\<^sub>F_def by fast show False proof (cases "x = z") case True have "is_nonterminal \ (item_rule_head y)" using *(5,6) assms(1,3-5) apply (clarsimp simp: wf_item_def bin_def item_rule_head_def rule_head_def next_symbol_def) by (metis prod.exhaust_sel rule_nonterminal_type) thus ?thesis using True *(7) assms(1,2,5) is_terminal_nonterminal by fastforce next case False thus ?thesis using * assms(1) by (auto simp: next_symbol_def Complete\<^sub>F_def bin_def) qed qed lemma Complete\<^sub>F_Un_eq_nonterminal: assumes "wf_\ \" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" assumes "nonempty_derives \" "wf_item \ \ z" assumes "item_end z = k" "next_symbol z \ None" shows "Complete\<^sub>F k (I \ {z}) = Complete\<^sub>F k I" proof (rule ccontr) assume "Complete\<^sub>F k (I \ {z}) \ Complete\<^sub>F k I" hence "Complete\<^sub>F k I \ Complete\<^sub>F k (I \ {z})" using Complete\<^sub>F_sub_mono by blast then obtain x x' y where *: "x \ Complete\<^sub>F k (I \ {z})" "x \ Complete\<^sub>F k I" "x = inc_item x' k" "x' \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" "is_complete y" "next_symbol x' = Some (item_rule_head y)" unfolding Complete\<^sub>F_def by fast consider (A) "x' = z" | (B) "y = z" using *(2-7) Complete\<^sub>F_def by (auto simp: bin_def; blast) thus False proof cases case A have "item_origin y = k" using *(4) A bin_def assms(6) by (metis (mono_tags, lifting) mem_Collect_eq) moreover have "item_end y = k" using *(5) bin_def by blast moreover have "sound_item \ \ y" using *(5,6) assms(3,7) by (auto simp: bin_def next_symbol_def sound_item_def) moreover have "wf_item \ \ y" using *(5) assms(2,5) wf_item_def by (auto simp: bin_def) ultimately show ?thesis using impossible_complete_item *(6) assms(1,4) by blast next case B thus ?thesis using *(6) assms(7) by (auto simp: next_symbol_def) qed qed lemma wf_item_in_kth_bin: "wf_bins \ \ bs \ x \ bins bs \ item_end x = k \ x \ set (items (bs ! k))" using bins_bin_exists wf_bins_kth_bin wf_bins_def by blast lemma Complete\<^sub>F_bins_upto_eq_bins: assumes "wf_bins \ \ bs" "k < length bs" "i \ length (items (bs ! k))" shows "Complete\<^sub>F k (bins_upto bs k i) = Complete\<^sub>F k (bins bs)" proof - have "\l. l \ k \ bin (bins_upto bs k i) l = bin (bins bs) l" using bin_bins_upto_bins_eq[OF assms] by blast moreover have "\x \ bins bs. wf_item \ \ x" using assms(1) wf_bins_impl_wf_items by metis ultimately show ?thesis unfolding Complete\<^sub>F_def bin_def wf_item_def wf_item_def by auto qed lemma Complete\<^sub>F_sub_bins_Un_Complete\<^sub>L: assumes "Complete\<^sub>F k I \ bins bs" "I \ bins bs" "is_complete z" "wf_bins \ \ bs" "wf_item \ \ z" shows "Complete\<^sub>F k (I \ {z}) \ bins bs \ set (items (Complete\<^sub>L k z bs red))" proof standard fix w assume "w \ Complete\<^sub>F k (I \ {z})" then obtain x y where *: "w = inc_item x k" "x \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" "is_complete y" "next_symbol x = Some (item_rule_head y)" unfolding Complete\<^sub>F_def by blast consider (A) "x = z" | (B) "y = z" | "\ (x = z \ y = z)" by blast thus "w \ bins bs \ set (items (Complete\<^sub>L k z bs red))" proof cases case A thus ?thesis using *(5) assms(3) by (auto simp: next_symbol_def) next case B let ?orig = "bs ! item_origin z" let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head z)) (items ?orig)" have "x \ bin I (item_origin y)" using B *(2) *(5) assms(3) by (auto simp: next_symbol_def bin_def) moreover have "bin I (item_origin z) \ set (items (bs ! item_origin z))" using wf_item_in_kth_bin assms(2,4) bin_def by blast ultimately have "x \ set (map fst ?is)" using *(5) B by (simp add: filter_with_index_cong_filter in_mono) thus ?thesis unfolding Complete\<^sub>L_def *(1) by (auto simp: rev_image_eqI items_def) next case 3 thus ?thesis using * assms(1) Complete\<^sub>F_def by (auto simp: bin_def; blast) qed qed lemma Complete\<^sub>L_eq_item_origin: "bs ! item_origin y = bs' ! item_origin y \ Complete\<^sub>L k y bs red = Complete\<^sub>L k y bs' red" by (auto simp: Complete\<^sub>L_def) lemma kth_bin_bins_upto_empty: assumes "wf_bins \ \ bs" "k < length bs" shows "bin (bins_upto bs k 0) k = {}" proof - { fix x assume "x \ bins_upto bs k 0" then obtain l where "x \ set (items (bs ! l))" "l < k" unfolding bins_upto_def bin_upto_def by blast hence "item_end x = l" using wf_bins_kth_bin assms by fastforce hence "item_end x < k" using \l < k\ by blast } thus ?thesis by (auto simp: bin_def) qed lemma Earley\<^sub>L_bin'_mono: assumes "(k, \, \, bs) \ wf_earley_input" shows "bins bs \ bins (Earley\<^sub>L_bin' k \ \ bs i)" using assms proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have wf: "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast hence "bins bs \ bins ?bs'" using length_bins_upd bins_bins_upd wf_earley_input_elim by (metis Un_upper1) also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using wf Complete\<^sub>F.IH by blast finally show ?case using Complete\<^sub>F.hyps by simp next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have wf: "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis hence "bins bs \ bins ?bs'" using Scan\<^sub>F.hyps(5) length_bins_upd bins_bins_upd wf_earley_input_elim by (metis add_mono1 sup_ge1) also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using wf Scan\<^sub>F.IH by blast finally show ?case using Scan\<^sub>F.hyps by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have wf: "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis hence "bins bs \ bins ?bs'" using length_bins_upd bins_bins_upd wf_earley_input_elim by (metis sup_ge1) also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using wf Predict\<^sub>F.IH by blast finally show ?case using Predict\<^sub>F.hyps by simp qed simp_all lemma Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin': assumes "(k, \, \, bs) \ wf_earley_input" assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k i) \ bins bs" assumes "\x \ bins bs. sound_item \ \ x" "is_word \ \" "nonempty_derives \" shows "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin' k \ \ bs i)" using assms proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Base k \ \ bs i) have "bin (bins bs) k = bin (bins_upto bs k i) k" using Base.hyps Base.prems(1) bin_bins_upto_bins_eq wf_earley_input_elim by blast thus ?case using Scan\<^sub>F_bin_absorb Predict\<^sub>F_bin_absorb Complete\<^sub>F_bins_upto_eq_bins wf_earley_input_elim Base.hyps Base.prems(1,2,3,5) Earley\<^sub>F_bin_step_def Complete\<^sub>F_Earley\<^sub>F_bin_step_mono Predict\<^sub>F_Earley\<^sub>F_bin_step_mono Scan\<^sub>F_Earley\<^sub>F_bin_step_mono Earley\<^sub>L_bin'_mono by (metis (no_types, lifting) Un_assoc sup.orderE) next case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have x: "x \ set (items (bs ! k))" using Complete\<^sub>F.hyps(1,2) by auto have wf: "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast hence sound: "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x by (metis dual_order.refl) have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using Complete\<^sub>F.hyps(1) bins_upto_Suc_Un length_nth_bin_bins_upd items_def by (metis length_map linorder_not_less sup.boundedE sup.order_iff) also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... \ bins bs \ Scan\<^sub>F k \ {x}" using Complete\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce also have "... = bins bs" using Complete\<^sub>F.hyps(3) by (auto simp: Scan\<^sub>F_def bin_def) finally show ?thesis using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast qed moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using Complete\<^sub>F.hyps(1) bins_upto_Suc_Un length_nth_bin_bins_upd by (metis dual_order.strict_trans1 items_def length_map not_le_imp_less) also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... \ bins bs \ Predict\<^sub>F k \ {x}" using Complete\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by blast also have "... = bins bs" using Complete\<^sub>F.hyps(3) by (auto simp: Predict\<^sub>F_def bin_def) finally show ?thesis using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast qed moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using bins_upto_Suc_Un length_nth_bin_bins_upd Complete\<^sub>F.hyps(1) by (metis (no_types, opaque_lifting) dual_order.trans items_def length_map not_le_imp_less) also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" using items_nth_idem_bins_upd Complete\<^sub>F.hyps(1,2) bins_upto_kth_nth_idem Complete\<^sub>F.prems(1) wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... \ bins bs \ set (items (Complete\<^sub>L k x bs i))" using Complete\<^sub>F_sub_bins_Un_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems(1,2,3) next_symbol_def bins_upto_sub_bins wf_bins_kth_bin x Complete\<^sub>F_Earley\<^sub>F_bin_step_mono wf_earley_input_elim by (smt (verit, best) option.distinct(1) subset_trans) finally show ?thesis using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast qed ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Complete\<^sub>F.IH Complete\<^sub>F.prems sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins wf_earley_input_elim bins_bins_upd by (metis UnE sup.boundedI) thus ?case using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(2) Earley\<^sub>F_bin_step_sub_mono bins_bins_upd wf_earley_input_elim by (smt (verit, best) sup.coboundedI2 sup.orderE sup_ge1) next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have x: "x \ set (items (bs ! k))" using Scan\<^sub>F.hyps(1,2) by auto hence sound: "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items x by (metis dual_order.refl) have wf: "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd by (metis Suc_eq_plus1 items_def length_map lessI less_not_refl not_le_imp_less) also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis add_mono_thms_linordered_field(1) items_def length_map less_add_one linorder_le_less_linear not_add_less1) also have "... \ bins bs \ Scan\<^sub>F k \ {x}" using Scan\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce finally have *: "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs \ Scan\<^sub>F k \ {x}" . show ?thesis proof cases assume a1: "\!k = a" hence "Scan\<^sub>F k \ {x} = {inc_item x (k+1)}" using Scan\<^sub>F.hyps(1-3,5) Scan\<^sub>F.prems(1,2) wf_earley_input_elim apply (auto simp: Scan\<^sub>F_def bin_def) using wf_bins_kth_bin x by blast hence "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs \ {inc_item x (k+1)}" using * by blast also have "... = bins bs \ set (items (Scan\<^sub>L k \ a x i))" using a1 Scan\<^sub>F.hyps(5) by (auto simp: Scan\<^sub>L_def items_def) also have "... = bins ?bs'" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (metis add_mono1) finally show ?thesis . next assume a1: "\ \!k = a" hence "Scan\<^sub>F k \ {x} = {}" using Scan\<^sub>F.hyps(3) by (auto simp: Scan\<^sub>F_def bin_def) hence "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs" using * by blast also have "... \ bins ?bs'" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (metis Un_left_absorb add_strict_right_mono subset_Un_eq) finally show ?thesis . qed qed moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd by (metis Suc_eq_plus1 dual_order.refl items_def length_map lessI linorder_not_less) also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis add_strict_right_mono items_def le_add1 length_map less_add_one linorder_not_le) also have "... \ bins bs \ Predict\<^sub>F k \ {x}" using Scan\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce also have "... = bins bs" using Scan\<^sub>F.hyps(3,4) Scan\<^sub>F.prems(1) is_terminal_nonterminal wf_earley_input_elim by (auto simp: Predict\<^sub>F_def bin_def rule_head_def, fastforce) finally show ?thesis using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) by (simp add: bins_bins_upd sup.coboundedI1 wf_earley_input_elim) qed moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd by (metis Suc_eq_plus1 items_def length_map lessI less_not_refl not_le_imp_less) also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis add_mono1 items_def length_map less_add_one linorder_not_le not_add_less1) also have "... = Complete\<^sub>F k (bins_upto bs k i)" using Complete\<^sub>F_Un_eq_terminal Scan\<^sub>F.hyps(3,4) Scan\<^sub>F.prems bins_upto_sub_bins subset_iff wf_bins_impl_wf_items wf_bins_kth_bin wf_item_def x wf_earley_input_elim by (smt (verit, ccfv_threshold)) finally show ?thesis using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,2,3) Complete\<^sub>F_Earley\<^sub>F_bin_step_mono by (auto simp: bins_bins_upd wf_earley_input_elim, blast) qed ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Scan\<^sub>F.IH Scan\<^sub>F.prems Scan\<^sub>F.hyps(5) sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins wf_earley_input_elim bins_bins_upd by (metis UnE add_mono1 le_supI) thus ?case using Earley\<^sub>F_bin_step_sub_mono Earley\<^sub>L_bin'_simps(3) Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (smt (verit, ccfv_SIG) add_mono1 sup.cobounded1 sup.coboundedI2 sup.orderE) next case (Pass k \ \ bs i x a) have x: "x \ set (items (bs ! k))" using Pass.hyps(1,2) by auto have "Scan\<^sub>F k \ (bins_upto bs k (i + 1)) \ bins bs" using Scan\<^sub>F_def Pass.hyps(5) by auto moreover have "Predict\<^sub>F k \ (bins_upto bs k (i + 1)) \ bins bs" proof - have "Predict\<^sub>F k \ (bins_upto bs k (i + 1)) = Predict\<^sub>F k \ (bins_upto bs k i \ {items (bs ! k) ! i})" using bins_upto_Suc_Un Pass.hyps(1) by (metis items_def length_map not_le_imp_less) also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" using Pass.hyps(1,2,5) nth_idem_bins_upd bins_upto_kth_nth_idem by simp also have "... \ bins bs \ Predict\<^sub>F k \ {x}" using Pass.prems(2) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by blast also have "... = bins bs" using Pass.hyps(3,4) Pass.prems(1) is_terminal_nonterminal wf_earley_input_elim by (auto simp: Predict\<^sub>F_def bin_def rule_head_def, fastforce) finally show ?thesis using bins_bins_upd Pass.hyps(5) Pass.prems(3) by auto qed moreover have "Complete\<^sub>F k (bins_upto bs k (i + 1)) \ bins bs" proof - have "Complete\<^sub>F k (bins_upto bs k (i + 1)) = Complete\<^sub>F k (bins_upto bs k i \ {x})" using bins_upto_Suc_Un Pass.hyps(1,2) by (metis items_def length_map not_le_imp_less) also have "... = Complete\<^sub>F k (bins_upto bs k i)" using Complete\<^sub>F_Un_eq_terminal Pass.hyps Pass.prems bins_upto_sub_bins subset_iff wf_bins_impl_wf_items wf_item_def wf_bins_kth_bin x wf_earley_input_elim by (smt (verit, best)) finally show ?thesis using Pass.prems(1,2) Complete\<^sub>F_Earley\<^sub>F_bin_step_mono wf_earley_input_elim by blast qed ultimately have "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin' k \ \ bs (i+1))" using Pass.IH Pass.prems Earley\<^sub>F_bin_step_def bins_upto_sub_bins wf_earley_input_elim by (metis le_sup_iff) thus ?case using bins_bins_upd Pass.hyps Pass.prems by simp next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "k \ length \ \ \ \!k = a" using Predict\<^sub>F.hyps(4) Predict\<^sub>F.prems(4) is_word_is_terminal leI by blast have x: "x \ set (items (bs ! k))" using Predict\<^sub>F.hyps(1,2) by auto hence sound: "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast have wf: "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis have len: "i < length (items (?bs' ! k))" using length_nth_bin_bins_upd Predict\<^sub>F.hyps(1) by (metis dual_order.strict_trans1 items_def length_map linorder_not_less) have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using Predict\<^sub>F.hyps(1) bins_upto_Suc_Un by (metis items_def len length_map) also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" using Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... \ bins bs \ Scan\<^sub>F k \ {x}" using Predict\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce also have "... = bins bs" using Predict\<^sub>F.hyps(3) \length \ \ k \ \ ! k \ a\ by (auto simp: Scan\<^sub>F_def bin_def) finally show ?thesis using Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast qed moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using Predict\<^sub>F.hyps(1) bins_upto_Suc_Un by (metis items_def len length_map) also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" using Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... \ bins bs \ Predict\<^sub>F k \ {x}" using Predict\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce also have "... = bins bs \ set (items (Predict\<^sub>L k \ a))" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1-3) wf_earley_input_elim apply (auto simp: Predict\<^sub>F_def Predict\<^sub>L_def bin_def items_def) using wf_bins_kth_bin x by blast finally show ?thesis using Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast qed moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" proof - have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" using bins_upto_Suc_Un len by (metis items_def length_map) also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" using items_nth_idem_bins_upd Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) bins_upto_kth_nth_idem wf_earley_input_elim by (metis dual_order.refl items_def length_map not_le_imp_less) also have "... = Complete\<^sub>F k (bins_upto bs k i)" using Complete\<^sub>F_Un_eq_nonterminal Predict\<^sub>F.prems bins_upto_sub_bins Predict\<^sub>F.hyps(3) subset_eq wf_bins_kth_bin x wf_bins_impl_wf_items wf_item_def wf_earley_input_elim by (smt (verit, ccfv_SIG) option.simps(3)) also have "... \ bins bs" using Complete\<^sub>F_Earley\<^sub>F_bin_step_mono Predict\<^sub>F.prems(2) by blast finally show ?thesis using bins_bins_upd Predict\<^sub>F.prems(1,2,3) wf_earley_input_elim by (metis Un_upper1 dual_order.trans) qed ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" using Predict\<^sub>F.IH Predict\<^sub>F.prems sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins bins_bins_upd wf_earley_input_elim by (metis UnE le_supI) hence "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ bs i)" using Predict\<^sub>F.hyps Earley\<^sub>L_bin'_simps(5) by simp moreover have "Earley\<^sub>F_bin_step k \ \ (bins bs) \ Earley\<^sub>F_bin_step k \ \ (bins ?bs')" using Earley\<^sub>F_bin_step_sub_mono Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (metis Un_upper1) ultimately show ?case by blast qed lemma Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" assumes "\x \ bins bs. sound_item \ \ x" "is_word \ \" "nonempty_derives \" shows "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" using assms Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin' Earley\<^sub>L_bin_def by metis lemma bins_eq_items_Complete\<^sub>L: assumes "bins_eq_items as bs" "item_origin x < length as" shows "items (Complete\<^sub>L k x as i) = items (Complete\<^sub>L k x bs i)" proof - let ?orig_a = "as ! item_origin x" let ?orig_b = "bs ! item_origin x" have "items ?orig_a = items ?orig_b" using assms by (metis (no_types, opaque_lifting) bins_eq_items_def length_map nth_map) thus ?thesis unfolding Complete\<^sub>L_def by simp qed lemma Earley\<^sub>L_bin'_bins_eq: assumes "(k, \, \, as) \ wf_earley_input" assumes "bins_eq_items as bs" "wf_bins \ \ as" shows "bins_eq_items (Earley\<^sub>L_bin' k \ \ as i) (Earley\<^sub>L_bin' k \ \ bs i)" using assms proof (induction i arbitrary: bs rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Base k \ \ as i) have "Earley\<^sub>L_bin' k \ \ as i = as" by (simp add: Base.hyps) moreover have "Earley\<^sub>L_bin' k \ \ bs i = bs" using Base.hyps Base.prems(1,2) unfolding bins_eq_items_def by (metis Earley\<^sub>L_bin'_simps(1) length_map nth_map wf_earley_input_elim) ultimately show ?case using Base.prems(2) by presburger next case (Complete\<^sub>F k \ \ as i x) let ?as' = "bins_upd as k (Complete\<^sub>L k x as i)" let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have k: "k < length as" using Complete\<^sub>F.prems(1) wf_earley_input_elim by blast hence wf_x: "wf_item \ \ x" using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(3) wf_bins_kth_bin by fastforce have "(k, \, \, ?as') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast moreover have "bins_eq_items ?as' ?bs'" using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(2,3) bins_eq_items_dist_bins_upd bins_eq_items_Complete\<^sub>L k wf_x wf_bins_kth_bin wf_item_def by (metis dual_order.strict_trans2 leI nth_mem) ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using Complete\<^sub>F.IH wf_earley_input_elim by blast moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" using Complete\<^sub>F.hyps by simp moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" using Complete\<^sub>F.hyps Complete\<^sub>F.prems unfolding bins_eq_items_def by (metis Earley\<^sub>L_bin'_simps(2) map_eq_imp_length_eq nth_map wf_earley_input_elim) ultimately show ?case by argo next case (Scan\<^sub>F k \ \ as i x a) let ?as' = "bins_upd as (k+1) (Scan\<^sub>L k \ a x i)" let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have "(k, \, \, ?as') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by fast moreover have "bins_eq_items ?as' ?bs'" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,2) bins_eq_items_dist_bins_upd add_mono1 wf_earley_input_elim by metis ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using Scan\<^sub>F.IH wf_earley_input_elim by blast moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" using Scan\<^sub>F.hyps by simp moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" using Scan\<^sub>F.hyps Scan\<^sub>F.prems unfolding bins_eq_items_def by (smt (verit, ccfv_threshold) Earley\<^sub>L_bin'_simps(3) length_map nth_map wf_earley_input_elim) ultimately show ?case by argo next case (Pass k \ \ as i x a) have "bins_eq_items (Earley\<^sub>L_bin' k \ \ as (i + 1)) (Earley\<^sub>L_bin' k \ \ bs (i + 1))" using Pass.prems Pass.IH by blast moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ as (i+1)" using Pass.hyps by simp moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ bs (i+1)" using Pass.hyps Pass.prems unfolding bins_eq_items_def by (metis Earley\<^sub>L_bin'_simps(4) map_eq_imp_length_eq nth_map wf_earley_input_elim) ultimately show ?case by argo next case (Predict\<^sub>F k \ \ as i x a) let ?as' = "bins_upd as k (Predict\<^sub>L k \ a)" let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have "(k, \, \, ?as') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by fast moreover have "bins_eq_items ?as' ?bs'" using Predict\<^sub>F.prems(1,2) bins_eq_items_dist_bins_upd wf_earley_input_elim by blast ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using Predict\<^sub>F.IH wf_earley_input_elim by blast moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" using Predict\<^sub>F.hyps by simp moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" using Predict\<^sub>F.hyps Predict\<^sub>F.prems unfolding bins_eq_items_def by (metis Earley\<^sub>L_bin'_simps(5) length_map nth_map wf_earley_input_elim) ultimately show ?case by argo qed lemma Earley\<^sub>L_bin'_idem: assumes "(k, \, \, bs) \ wf_earley_input" assumes "i \ j" "\x \ bins bs. sound_item \ \ x" "nonempty_derives \" shows "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ bs i)" using assms proof (induction i arbitrary: j rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) case (Complete\<^sub>F k \ \ bs i x) let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" have x: "x \ set (items (bs ! k))" using Complete\<^sub>F.hyps(1,2) by auto have wf: "(k, \, \, ?bs') \ wf_earley_input" using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x by (metis dual_order.refl) hence sound: "\x \ bins ?bs'. sound_item \ \ x" by (metis Complete\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim) show ?case proof cases assume "i+1 \ j" thus ?thesis using wf sound Complete\<^sub>F Earley\<^sub>L_bin'_simps(2) by metis next assume "\ i+1 \ j" hence "i = j" using Complete\<^sub>F.prems(2) by simp have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" using Earley\<^sub>L_bin'_simps(2) Complete\<^sub>F.hyps(1-3) by simp also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" proof - let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" have "length (items (?bs'' ! k)) \ length (items (bs ! k))" using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans wf Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) by (smt (verit, ccfv_threshold) Earley\<^sub>L_bin'_simps(2)) hence 0: "\ length (items (?bs'' ! k)) \ j" using \i = j\ Complete\<^sub>F.hyps(1) by linarith have "x = items (?bs' ! k) ! j" using \i = j\ items_nth_idem_bins_upd Complete\<^sub>F.hyps(1,2) by (metis items_def length_map not_le_imp_less) hence 1: "x = items (?bs'' ! k) ! j" using \i = j\ kth_Earley\<^sub>L_bin'_bins Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(2) leI by metis have "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) (j+1))" using Earley\<^sub>L_bin'_simps(2) 0 1 Complete\<^sub>F.hyps(1,3) Complete\<^sub>F.prems(2) \i = j\ by auto moreover have "bins_eq_items (bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) ?bs''" proof - have "k < length bs" using Complete\<^sub>F.prems(1) wf_earley_input_elim by blast have 0: "set (Complete\<^sub>L k x bs i) = set (Complete\<^sub>L k x ?bs'' i)" proof (cases "item_origin x = k") case True thus ?thesis using impossible_complete_item kth_bin_sub_bins Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_kth_bin x next_symbol_def by (metis option.distinct(1) subsetD) next case False hence "item_origin x < k" using x Complete\<^sub>F.prems(1) wf_bins_kth_bin wf_item_def nat_less_le by (metis wf_earley_input_elim) hence "bs ! item_origin x = ?bs'' ! item_origin x" using False nth_idem_bins_upd nth_Earley\<^sub>L_bin'_eq wf by metis thus ?thesis using Complete\<^sub>L_eq_item_origin by metis qed have "set (items (Complete\<^sub>L k x bs i)) \ set (items (?bs' ! k))" by (simp add: \k < length bs\ bins_upd_def set_items_bin_upds) hence "set (items (Complete\<^sub>L k x ?bs'' i)) \ set (items (?bs' ! k))" using 0 by (simp add: items_def) also have "... \ set (items (?bs'' ! k))" by (simp add: wf nth_bin_sub_Earley\<^sub>L_bin') finally show ?thesis using bins_eq_items_bins_upd by blast qed moreover have "(k, \, \, bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) \ wf_earley_input" using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Complete\<^sub>L Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins 0 1 by blast ultimately show ?thesis using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast qed also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using Complete\<^sub>F.IH[OF wf _ sound Complete\<^sub>F.prems(4)] \i = j\ by blast finally show ?thesis using Complete\<^sub>F.hyps by simp qed next case (Scan\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" have x: "x \ set (items (bs ! k))" using Scan\<^sub>F.hyps(1,2) by auto hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items x by (metis dual_order.refl) hence sound: "\x \ bins ?bs'. sound_item \ \ x" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,3) bins_bins_upd wf_earley_input_elim by (metis UnE add_less_cancel_right) have wf: "(k, \, \, ?bs') \ wf_earley_input" using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis show ?case proof cases assume "i+1 \ j" thus ?thesis using sound Scan\<^sub>F by (metis Earley\<^sub>L_bin'_simps(3) wf_earley_input_Scan\<^sub>L) next assume "\ i+1 \ j" hence "i = j" using Scan\<^sub>F.prems(2) by auto have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" using Scan\<^sub>F.hyps by simp also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" proof - let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" have "length (items (?bs'' ! k)) \ length (items (bs ! k))" using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(3) by (smt (verit, ccfv_SIG)) hence "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) (j+1))" using \i = j\ kth_Earley\<^sub>L_bin'_bins nth_idem_bins_upd Earley\<^sub>L_bin'_simps(3) Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) by (smt (verit, best) leI le_trans) moreover have "bins_eq_items (bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) ?bs''" proof - have "k+1 < length bs" using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems wf_earley_input_elim by fastforce+ hence "set (items (Scan\<^sub>L k \ a x i)) \ set (items (?bs' ! (k+1)))" by (simp add: bins_upd_def set_items_bin_upds) also have "... \ set (items (?bs'' ! (k+1)))" using wf nth_bin_sub_Earley\<^sub>L_bin' by blast finally show ?thesis using bins_eq_items_bins_upd by blast qed moreover have "(k, \, \, bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) \ wf_earley_input" using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Scan\<^sub>L Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins by (smt (verit, ccfv_SIG) Earley\<^sub>L_bin'_simps(3) linorder_not_le order.trans) ultimately show ?thesis using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast qed also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using \i = j\ Scan\<^sub>F.IH Scan\<^sub>F.prems Scan\<^sub>F.hyps sound wf_earley_input_Scan\<^sub>L by fast finally show ?thesis using Scan\<^sub>F.hyps by simp qed next case (Pass k \ \ bs i x a) show ?case proof cases assume "i+1 \ j" thus ?thesis using Pass by (metis Earley\<^sub>L_bin'_simps(4)) next assume "\ i+1 \ j" show ?thesis using Pass Earley\<^sub>L_bin'_simps(1,4) kth_Earley\<^sub>L_bin'_bins by (metis Suc_eq_plus1 Suc_leI antisym_conv2 not_le_imp_less) qed next case (Predict\<^sub>F k \ \ bs i x a) let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" have x: "x \ set (items (bs ! k))" using Predict\<^sub>F.hyps(1,2) by auto hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast hence sound: "\x \ bins ?bs'. sound_item \ \ x" using Predict\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim by metis have wf: "(k, \, \, ?bs') \ wf_earley_input" using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis have len: "i < length (items (?bs' ! k))" using length_nth_bin_bins_upd Predict\<^sub>F.hyps(1) Orderings.preorder_class.dual_order.strict_trans1 linorder_not_less by (metis items_def length_map) show ?case proof cases assume "i+1 \ j" thus ?thesis using sound wf Predict\<^sub>F by (metis Earley\<^sub>L_bin'_simps(5)) next assume "\ i+1 \ j" hence "i = j" using Predict\<^sub>F.prems(2) by auto have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" using Predict\<^sub>F.hyps by simp also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" proof - let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" have "length (items (?bs'' ! k)) \ length (items (bs ! k))" using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans wf by (metis (no_types, lifting) items_def length_map) hence "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' k (Predict\<^sub>L k \ a)) (j+1))" using \i = j\ kth_Earley\<^sub>L_bin'_bins nth_idem_bins_upd Earley\<^sub>L_bin'_simps(5) Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) length_bins_Earley\<^sub>L_bin' wf_bins_Earley\<^sub>L_bin' wf_bins_kth_bin wf_item_def x by (smt (verit, ccfv_SIG) linorder_not_le order.trans) moreover have "bins_eq_items (bins_upd ?bs'' k (Predict\<^sub>L k \ a)) ?bs''" proof - have "k < length bs" using wf_earley_input_elim[OF Predict\<^sub>F.prems(1)] by blast hence "set (items (Predict\<^sub>L k \ a)) \ set (items (?bs' ! k))" by (simp add: bins_upd_def set_items_bin_upds) also have "... \ set (items (?bs'' ! k))" using wf nth_bin_sub_Earley\<^sub>L_bin' by blast finally show ?thesis using bins_eq_items_bins_upd by blast qed moreover have "(k, \, \, bins_upd ?bs'' k (Predict\<^sub>L k \ a)) \ wf_earley_input" using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Predict\<^sub>L Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins by (smt (verit, best) Earley\<^sub>L_bin'_simps(5) dual_order.trans not_le_imp_less) ultimately show ?thesis using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast qed also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" using \i = j\ Predict\<^sub>F.IH Predict\<^sub>F.prems sound wf by (metis order_refl) finally show ?thesis using Predict\<^sub>F.hyps by simp qed qed simp lemma Earley\<^sub>L_bin_idem: assumes "(k, \, \, bs) \ wf_earley_input" assumes "\x \ bins bs. sound_item \ \ x" "nonempty_derives \" shows "bins (Earley\<^sub>L_bin k \ \ (Earley\<^sub>L_bin k \ \ bs)) = bins (Earley\<^sub>L_bin k \ \ bs)" using assms Earley\<^sub>L_bin'_idem Earley\<^sub>L_bin_def le0 by metis lemma funpower_Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" "\x \ bins bs. sound_item \ \ x" assumes "is_word \ \" "nonempty_derives \" shows "funpower (Earley\<^sub>F_bin_step k \ \) n (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" using assms proof (induction n) case 0 thus ?case using Earley\<^sub>L_bin'_mono Earley\<^sub>L_bin_def by (simp add: Earley\<^sub>L_bin'_mono Earley\<^sub>L_bin_def) next case (Suc n) have 0: "Earley\<^sub>F_bin_step k \ \ (bins_upto (Earley\<^sub>L_bin k \ \ bs) k 0) \ bins (Earley\<^sub>L_bin k \ \ bs)" using Earley\<^sub>L_bin'_mono bins_upto_k0_Earley\<^sub>L_bin'_eq assms(1,2) Earley\<^sub>L_bin_def order_trans by (metis (no_types, lifting)) have "funpower (Earley\<^sub>F_bin_step k \ \) (Suc n) (bins bs) \ Earley\<^sub>F_bin_step k \ \ (bins (Earley\<^sub>L_bin k \ \ bs))" using Earley\<^sub>F_bin_step_sub_mono Suc by (metis funpower.simps(2)) also have "... \ bins (Earley\<^sub>L_bin k \ \ (Earley\<^sub>L_bin k \ \ bs))" using Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin Suc.prems wf_bins_Earley\<^sub>L_bin sound_Earley\<^sub>L_bin 0 wf_earley_input_Earley\<^sub>L_bin by blast also have "... \ bins (Earley\<^sub>L_bin k \ \ bs)" using Earley\<^sub>L_bin_idem Suc.prems by blast finally show ?case . qed lemma Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin: assumes "(k, \, \, bs) \ wf_earley_input" assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" "\x \ bins bs. sound_item \ \ x" assumes "is_word \ \" "nonempty_derives \" shows "Earley\<^sub>F_bin k \ \ (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" using assms funpower_Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin Earley\<^sub>F_bin_def elem_limit_simp by fastforce lemma Earley\<^sub>F_bins_sub_Earley\<^sub>L_bins: assumes "k \ length \" "wf_\ \" assumes "is_word \ \" "nonempty_derives \" shows "Earley\<^sub>F_bins k \ \ \ bins (Earley\<^sub>L_bins k \ \)" using assms proof (induction k) case 0 hence "Earley\<^sub>F_bin 0 \ \ (Init\<^sub>F \) \ bins (Earley\<^sub>L_bin 0 \ \ (Init\<^sub>L \ \))" using Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin Init\<^sub>L_eq_Init\<^sub>F length_bins_Init\<^sub>L Init\<^sub>L_eq_Init\<^sub>F sound_Init bins_upto_empty Earley\<^sub>F_bin_step_empty bins_upto_sub_bins wf_earley_input_Init\<^sub>L wf_earley_input_elim by (smt (verit, ccfv_threshold) Init\<^sub>F_sub_Earley basic_trans_rules(31) sound_Earley wf_bins_impl_wf_items) thus ?case by simp next case (Suc k) have wf: "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) have sub: "Earley\<^sub>F_bin_step (Suc k) \ \ (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) \ bins (Earley\<^sub>L_bins k \ \)" proof - have "bin (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) (Suc k) = {}" using kth_bin_bins_upto_empty wf Suc.prems wf_earley_input_elim by blast hence "Earley\<^sub>F_bin_step (Suc k) \ \ (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) = bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0" unfolding Earley\<^sub>F_bin_step_def Scan\<^sub>F_def Complete\<^sub>F_def Predict\<^sub>F_def bin_def by blast also have "... \ bins (Earley\<^sub>L_bins k \ \)" using wf Suc.prems bins_upto_sub_bins wf_earley_input_elim by blast finally show ?thesis . qed have sound: "\x \ bins (Earley\<^sub>L_bins k \ \). sound_item \ \ x" using Suc Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins by (metis Suc_leD Earley\<^sub>F_bins_sub_Earley in_mono sound_Earley wf_Earley) have "Earley\<^sub>F_bins (Suc k) \ \ \ Earley\<^sub>F_bin (Suc k) \ \ (bins (Earley\<^sub>L_bins k \ \))" using Suc Earley\<^sub>F_bin_sub_mono by simp also have "... \ bins (Earley\<^sub>L_bin (Suc k) \ \ (Earley\<^sub>L_bins k \ \))" using Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin wf sub sound Suc.prems by fastforce finally show ?case by simp qed lemma Earley\<^sub>F_sub_Earley\<^sub>L: assumes "wf_\ \" "is_word \ \" "nonempty_derives \" shows "Earley\<^sub>F \ \ \ bins (Earley\<^sub>L \ \)" using assms Earley\<^sub>F_bins_sub_Earley\<^sub>L_bins Earley\<^sub>F_def Earley\<^sub>L_def by (metis le_refl) theorem completeness_Earley\<^sub>L: assumes "derives \ [\ \] \" "is_word \ \" "wf_\ \" "nonempty_derives \" shows "recognizing (bins (Earley\<^sub>L \ \)) \ \" using assms Earley\<^sub>F_sub_Earley\<^sub>L Earley\<^sub>L_sub_Earley\<^sub>F completeness_Earley\<^sub>F by (metis subset_antisym) subsection \Correctness\ theorem Earley_eq_Earley\<^sub>L: assumes "wf_\ \" "is_word \ \" "nonempty_derives \" shows "Earley \ \ = bins (Earley\<^sub>L \ \)" using assms Earley\<^sub>F_sub_Earley\<^sub>L Earley\<^sub>L_sub_Earley\<^sub>F Earley_eq_Earley\<^sub>F by blast theorem correctness_Earley\<^sub>L: assumes "wf_\ \" "is_word \ \" "nonempty_derives \" shows "recognizing (bins (Earley\<^sub>L \ \)) \ \ \ derives \ [\ \] \" using assms Earley_eq_Earley\<^sub>L correctness_Earley by fastforce end diff --git a/thys/Echelon_Form/Echelon_Form.thy b/thys/Echelon_Form/Echelon_Form.thy --- a/thys/Echelon_Form/Echelon_Form.thy +++ b/thys/Echelon_Form/Echelon_Form.thy @@ -1,3005 +1,3002 @@ (* Title: Echelon_Form.thy Author: Jose Divasón Author: Jesús Aransay *) section\Echelon Form\ theory Echelon_Form imports Rings2 Gauss_Jordan.Determinants2 Cayley_Hamilton_Compatible begin subsection\Definition of Echelon Form\ text\Echelon form up to column k (NOT INCLUDED).\ definition echelon_form_upt_k :: "'a::{bezout_ring}^'cols::{mod_type}^'rows::{finite, ord} \ nat \ bool" where "echelon_form_upt_k A k = ( (\i. is_zero_row_upt_k i k A \ \ (\j. j>i \ \ is_zero_row_upt_k j k A)) \ (\i j. i \ (is_zero_row_upt_k i k A) \ \ (is_zero_row_upt_k j k A) \ ((LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0))))" definition "echelon_form A = echelon_form_upt_k A (ncols A)" text\Some properties of matrices in echelon form.\ lemma echelon_form_upt_k_intro: assumes "(\i. is_zero_row_upt_k i k A \ \ (\j. j>i \ \ is_zero_row_upt_k j k A))" and "(\i j. i \ (is_zero_row_upt_k i k A) \ \ (is_zero_row_upt_k j k A) \ ((LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)))" shows "echelon_form_upt_k A k" using assms unfolding echelon_form_upt_k_def by fast lemma echelon_form_upt_k_condition1: assumes "echelon_form_upt_k A k" "is_zero_row_upt_k i k A" shows "\ (\j. j>i \ \ is_zero_row_upt_k j k A)" using assms unfolding echelon_form_upt_k_def by auto lemma echelon_form_upt_k_condition1': assumes "echelon_form_upt_k A k" "is_zero_row_upt_k i k A" and "i (is_zero_row_upt_k i k A)" "\ (is_zero_row_upt_k j k A)" shows "(LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using assms unfolding echelon_form_upt_k_def by auto lemma echelon_form_upt_k_if_equal: assumes e: "echelon_form_upt_k A k" and eq: "\a. \b is_zero_row_upt_k i k B" and not_zero_jB: "\ is_zero_row_upt_k j k B" obtain a where Bia: "B $ i $ a \ 0" and ak: "a 0" by (metis ak Bia eq) obtain b where Bjb: "B $ j $ b \ 0" and bk: "b 0" by (metis bk Bjb eq) have not_zero_iA: "\ is_zero_row_upt_k i k A" by (metis (full_types) Aia ak is_zero_row_upt_k_def to_nat_le) have not_zero_jA: "\ is_zero_row_upt_k j k A" by (metis (full_types) Ajb bk is_zero_row_upt_k_def to_nat_le) have "(LEAST n. A $ i $ n \ 0) = (LEAST n. B $ i $ n \ 0)" proof (rule Least_equality) have "(LEAST n. B $ i $ n \ 0) \ a" by (rule Least_le, simp add: Bia) hence least_bi_less: "(LEAST n. B $ i $ n \ 0) < from_nat k" using ak by simp thus "A $ i $ (LEAST n. B $ i $ n \ 0) \ 0" by (metis (mono_tags, lifting) LeastI eq is_zero_row_upt_k_def not_zero_iB) fix y assume "A $ i $ y \ 0" thus "(LEAST n. B $ i $ n \ 0) \ y" by (metis (mono_tags, lifting) Least_le dual_order.strict_trans2 eq least_bi_less linear) qed moreover have "(LEAST n. A $ j $ n \ 0) = (LEAST n. B $ j $ n \ 0)" proof (rule Least_equality) have "(LEAST n. B $ j $ n \ 0) \ b" by (rule Least_le, simp add: Bjb) hence least_bi_less: "(LEAST n. B $ j $ n \ 0) < from_nat k" using bk by simp thus "A $ j $ (LEAST n. B $ j $ n \ 0) \ 0" by (metis (mono_tags, lifting) LeastI eq is_zero_row_upt_k_def not_zero_jB) fix y assume "A $ j $ y \ 0" thus "(LEAST n. B $ j $ n \ 0) \ y" by (metis (mono_tags, lifting) Least_le dual_order.strict_trans2 eq least_bi_less linear) qed moreover have "(LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" by (rule echelon_form_upt_k_condition2[OF e ij not_zero_iA not_zero_jA]) ultimately show "(LEAST n. B $ i $ n \ 0) < (LEAST n. B $ j $ n \ 0)" by auto qed lemma echelon_form_upt_k_0: "echelon_form_upt_k A 0" unfolding echelon_form_upt_k_def is_zero_row_upt_k_def by auto lemma echelon_form_condition1: assumes r: "echelon_form A" shows "(\i. is_zero_row i A \ \ (\j. j>i \ \ is_zero_row j A))" using r unfolding echelon_form_def by (metis echelon_form_upt_k_condition1' is_zero_row_def) lemma echelon_form_condition2: assumes r: "echelon_form A" shows "(\i. i \ (is_zero_row i A) \ \ (is_zero_row j A) \ ((LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)))" using r unfolding echelon_form_def by (metis echelon_form_upt_k_condition2 is_zero_row_def) lemma echelon_form_condition2_explicit: assumes rref_A: "echelon_form A" and i_le: "i < j" and "\ is_zero_row i A" and "\ is_zero_row j A" shows "(LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using echelon_form_condition2 assms by blast lemma echelon_form_intro: assumes 1: "(\i. is_zero_row i A \ \ (\j. j>i \ \ is_zero_row j A))" and 2: "(\i j. i \ (is_zero_row i A) \ \ (is_zero_row j A) \ ((LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)))" shows "echelon_form A" proof (unfold echelon_form_def, rule echelon_form_upt_k_intro, auto) fix i j assume "is_zero_row_upt_k i (ncols A) A" and "i < j" thus "is_zero_row_upt_k j (ncols A) A" using 1 is_zero_row_imp_is_zero_row_upt by (metis is_zero_row_def) next fix i j assume "i < j" and "\ is_zero_row_upt_k i (ncols A) A" and "\ is_zero_row_upt_k j (ncols A) A" thus "(LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using 2 by (metis is_zero_row_imp_is_zero_row_upt) qed lemma echelon_form_implies_echelon_form_upt: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes rref: "echelon_form A" shows "echelon_form_upt_k A k" proof (rule echelon_form_upt_k_intro) show "\i. is_zero_row_upt_k i k A \ \ (\j>i. \ is_zero_row_upt_k j k A)" proof (auto, rule ccontr) fix i j assume zero_i_k: "is_zero_row_upt_k i k A" and i_less_j: "i < j" and not_zero_j_k:"\ is_zero_row_upt_k j k A" have not_zero_j: "\ is_zero_row j A" using is_zero_row_imp_is_zero_row_upt not_zero_j_k by blast hence not_zero_i: "\ is_zero_row i A" using echelon_form_condition1[OF rref] i_less_j by blast have Least_less: "(LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" by (rule echelon_form_condition2_explicit[OF rref i_less_j not_zero_i not_zero_j]) moreover have "(LEAST n. A $ j $ n \ 0) < (LEAST n. A $ i $ n \ 0)" proof (rule LeastI2_ex) show "\a. A $ i $ a \ 0" using not_zero_i unfolding is_zero_row_def is_zero_row_upt_k_def by fast fix x assume Aix_not_0: "A $ i $ x \ 0" have k_less_x: "k \ to_nat x" using zero_i_k Aix_not_0 unfolding is_zero_row_upt_k_def by force hence k_less_ncols: "k < ncols A" unfolding ncols_def using to_nat_less_card[of x] by simp obtain s where Ajs_not_zero: "A $ j $ s \ 0" and s_less_k: "to_nat s < k" using not_zero_j_k unfolding is_zero_row_upt_k_def by blast have "(LEAST n. A $ j $ n \ 0) \ s" using Ajs_not_zero Least_le by fast also have "... = from_nat (to_nat s)" unfolding from_nat_to_nat_id .. also have "... < from_nat k" by (rule from_nat_mono[OF s_less_k k_less_ncols[unfolded ncols_def]]) also have "... \ x" using k_less_x leD not_le_imp_less to_nat_le by fast finally show "(LEAST n. A $ j $ n \ 0) < x" . qed ultimately show False by fastforce qed show "\i j. i < j \ \ is_zero_row_upt_k i k A \ \ is_zero_row_upt_k j k A \ (LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" using echelon_form_condition2[OF rref] is_zero_row_imp_is_zero_row_upt by blast qed lemma upper_triangular_upt_k_def': assumes "\i j. to_nat j \ k \ A $ i $ j \ 0 \ j\i" shows "upper_triangular_upt_k A k" using assms unfolding upper_triangular_upt_k_def by (metis leD linear) lemma echelon_form_imp_upper_triagular_upt: fixes A::"'a::{bezout_ring}^'n::{mod_type}^'n::{mod_type}" assumes "echelon_form A" shows "upper_triangular_upt_k A k" proof (induct k) case 0 show ?case unfolding upper_triangular_upt_k_def by simp next case (Suc k) show ?case unfolding upper_triangular_upt_k_def proof (clarify) fix i j::'n assume j_less_i: "j < i" and j_less_suc_k: "to_nat j < Suc k" show "A $ i $ j = 0" proof (cases "to_nat j < k") case True thus ?thesis using Suc.hyps unfolding upper_triangular_upt_k_def using j_less_i True by auto next case False hence j_eq_k: "to_nat j = k" using j_less_suc_k by simp hence j_eq_k2: "from_nat k = j" by (metis from_nat_to_nat_id) have rref_suc: "echelon_form_upt_k A (Suc k)" by (metis assms echelon_form_implies_echelon_form_upt) have zero_j_k: "is_zero_row_upt_k j k A" unfolding is_zero_row_upt_k_def by (metis (opaque_lifting, mono_tags) Suc.hyps leD le_less_linear j_eq_k to_nat_mono' upper_triangular_upt_k_def) hence zero_i_k: "is_zero_row_upt_k i k A" by (metis (poly_guards_query) assms echelon_form_implies_echelon_form_upt echelon_form_upt_k_condition1' j_less_i) show ?thesis proof (cases "A $ j $ j = 0") case True have "is_zero_row_upt_k j (Suc k) A" by (rule is_zero_row_upt_k_suc[OF zero_j_k], simp add: True j_eq_k2) hence "is_zero_row_upt_k i (Suc k) A" by (metis echelon_form_upt_k_condition1' j_less_i rref_suc) thus ?thesis by (metis is_zero_row_upt_k_def j_eq_k lessI) next case False note Ajj_not_zero=False hence not_zero_j:"\ is_zero_row_upt_k j (Suc k) A" by (metis is_zero_row_upt_k_def j_eq_k lessI) show ?thesis proof (rule ccontr) assume Aij_not_zero: "A $ i $ j \ 0" hence not_zero_i: "\ is_zero_row_upt_k i (Suc k) A" by (metis is_zero_row_upt_k_def j_eq_k lessI) have Least_eq: "(LEAST n. A $ i $ n \ 0) = from_nat k" proof (rule Least_equality) show "A $ i $ from_nat k \ 0" using Aij_not_zero j_eq_k2 by simp show "\y. A $ i $ y \ 0 \ from_nat k \ y" by (metis (full_types) is_zero_row_upt_k_def not_le_imp_less to_nat_le zero_i_k) qed moreover have Least_eq2: "(LEAST n. A $ j $ n \ 0) = from_nat k" proof (rule Least_equality) show "A $ j $ from_nat k \ 0" using Ajj_not_zero j_eq_k2 by simp show "\y. A $ j $ y \ 0 \ from_nat k \ y" by (metis (full_types) is_zero_row_upt_k_def not_le_imp_less to_nat_le zero_j_k) qed ultimately show False using echelon_form_upt_k_condition2[OF rref_suc j_less_i not_zero_j not_zero_i] by simp qed qed qed qed qed text\A matrix in echelon form is upper triangular.\ lemma echelon_form_imp_upper_triagular: fixes A::"'a::{bezout_ring}^'n::{mod_type}^'n::{mod_type}" assumes "echelon_form A" shows "upper_triangular A" using echelon_form_imp_upper_triagular_upt[OF assms] by (metis upper_triangular_upt_imp_upper_triangular) lemma echelon_form_upt_k_interchange: fixes A::"'a::{bezout_ring}^'c::{mod_type}^'b::{mod_type}" assumes e: "echelon_form_upt_k A k" and zero_ikA: "is_zero_row_upt_k (from_nat i) k A" and Amk_not_0: "A $ m $ from_nat k \ 0" and i_le_m: "(from_nat i)\m" and k: "k 0 \ (from_nat i) \ n)) k" proof (rule echelon_form_upt_k_if_equal[OF e _ k], auto) fix a and b::'c assume b: "b < from_nat k" let ?least = "(LEAST n. A $ n $ from_nat k \ 0 \ (from_nat i) \ n)" let ?interchange = "(interchange_rows A (from_nat i) ?least)" have "(from_nat i)\?least" by (metis (mono_tags, lifting) Amk_not_0 LeastI_ex i_le_m) hence zero_leastkA: "is_zero_row_upt_k ?least k A" using echelon_form_upt_k_condition1[OF e zero_ikA] by (metis (poly_guards_query) dual_order.strict_iff_order zero_ikA) show "A $ a $ b = ?interchange $ a $ b" proof (cases "a=from_nat i") case True hence "?interchange $ a $ b = A $ ?least $ b" unfolding interchange_rows_def by auto also have "... = 0" using zero_leastkA unfolding is_zero_row_upt_k_def by (metis (mono_tags) b to_nat_le) finally have "?interchange $ a $ b = 0" . moreover have "A $ a $ b = 0" by (metis True b is_zero_row_upt_k_def to_nat_le zero_ikA) ultimately show ?thesis by simp next case False note a_not_i=False show ?thesis proof (cases "a=?least") case True hence "?interchange $ a $ b = A $ (from_nat i) $ b" unfolding interchange_rows_def by auto also have "... = 0" using zero_ikA unfolding is_zero_row_upt_k_def by (metis (poly_guards_query) b to_nat_le) finally have "?interchange $ a $ b = 0" . moreover have "A $ a $ b = 0" by (metis True b is_zero_row_upt_k_def to_nat_le zero_leastkA) ultimately show ?thesis by simp next case False thus ?thesis using a_not_i unfolding interchange_rows_def by auto qed qed qed text\There are similar theorems to the following ones in the Gauss-Jordan developments, but for matrices in reduced row echelon form. It is possible to prove that reduced row echelon form implies echelon form. Then the theorems in the Gauss-Jordan development could be obtained with ease.\ lemma greatest_less_zero_row: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{finite, wellorder}" assumes r: "echelon_form_upt_k A k" and zero_i: "is_zero_row_upt_k i k A" and not_all_zero: "\ (\a. is_zero_row_upt_k a k A)" shows "(GREATEST m. \ is_zero_row_upt_k m k A) < i" proof (rule ccontr) assume not_less_i: "\ (GREATEST m. \ is_zero_row_upt_k m k A) < i" have i_less_greatest: "i < (GREATEST m. \ is_zero_row_upt_k m k A)" by (metis not_less_i neq_iff GreatestI not_all_zero zero_i) have "is_zero_row_upt_k (GREATEST m. \ is_zero_row_upt_k m k A) k A" using r zero_i i_less_greatest unfolding echelon_form_upt_k_def by blast thus False using GreatestI_ex not_all_zero by fast qed lemma greatest_ge_nonzero_row': fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes r: "echelon_form_upt_k A k" and i: "i \ (GREATEST m. \ is_zero_row_upt_k m k A)" and not_all_zero: "\ (\a. is_zero_row_upt_k a k A)" shows "\ is_zero_row_upt_k i k A" using greatest_less_zero_row[OF r] i not_all_zero by fastforce lemma rref_imp_ef: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes rref: "reduced_row_echelon_form A" shows "echelon_form A" proof (rule echelon_form_intro) show "\i. is_zero_row i A \ \ (\j>i. \ is_zero_row j A)" by (simp add: rref rref_condition1) show "\i j. i < j \ \ is_zero_row i A \ \ is_zero_row j A \ (LEAST n. A $ i $ n \ 0) < (LEAST n. A $ j $ n \ 0)" by (simp add: rref_condition3_equiv rref) qed subsection\Computing the echelon form of a matrix\ subsubsection\Demonstration over principal ideal rings\ text\Important remark: We want to prove that there exist the echelon form of any matrix whose elements belong to a bezout domain. In addition, we want to compute the echelon form, so we will need computable gcd and bezout operations which is possible over euclidean domains. Our approach consists of demonstrating the correctness over bezout domains and executing over euclidean domains. To do that, we have studied several options: \begin{enumerate} \item We could define a gcd in bezout rings (\bezout_ring_gcd\) as follows: \gcd_bezout_ring a b = (SOME d. d dvd a \ d dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd d))\ And then define an algorithm that computes the Echelon Form using such a definition to the gcd. This would allow us to prove the correctness over bezout rings, but we would not be able to execute over euclidean rings because it is not possible to demonstrate a (code) lemma stating that \(gcd_bezout_ring a b) = gcd_eucl a b\ (the gcd is not unique over bezout rings and GCD rings). \item Create a \bezout_ring_norm\ class and define a gcd normalized over bezout rings: \definition gcd_bezout_ring_norm a b = gcd_bezout_ring a b div normalisation_factor (gcd_bezout_ring a b)\ Then, one could demonstrate a (code) lemma stating that: \(gcd_bezout_ring_norm a b) = gcd_eucl a b\ This allows us to execute the gcd function, but with bezout it is not possible. \item The third option (and the chosen one) consists of defining the algorithm over bezout domains and parametrizing the algorithm by a \bezout\ operation which must satisfy suitable properties (i.e @{term "is_bezout_ext bezout"}). Then we can prove the correctness over bezout domains and we will execute over euclidean domains, since we can prove that the operation @{term "euclid_ext2"} is an executable operation which satisfies @{term "is_bezout_ext euclid_ext2"}. \end{enumerate} \ subsubsection\Definition of the algorithm\ context bezout_ring begin definition bezout_matrix :: "'a^'cols^'rows \ 'rows \ 'rows \ 'cols \ ('a \ 'a \ ('a \ 'a \ 'a \ 'a \ 'a)) \ 'a^'rows^'rows" where "bezout_matrix A a b j bezout = (\ x y. (let (p, q, u, v, d) = bezout (A $ a $ j) (A $ b $ j) in if x = a \ y = a then p else if x = a \ y = b then q else if x = b \ y = a then u else if x = b \ y = b then v else if x = y then 1 else 0))" end primrec bezout_iterate :: "'a::{bezout_ring}^'cols^'rows::{mod_type} \ nat \ 'rows::{mod_type} \ 'cols \ ('a \'a \ ('a \ 'a \ 'a \ 'a \ 'a)) \ 'a^'cols^'rows::{mod_type}" where "bezout_iterate A 0 i j bezout = A" | "bezout_iterate A (Suc n) i j bezout = (if (Suc n) \ to_nat i then A else bezout_iterate (bezout_matrix A i (from_nat (Suc n)) j bezout ** A) n i j bezout)" text\If every element in column @{term "k::nat"} over index @{term "i::nat"} are equal to zero, the same input is returned. If every element over @{term "i::nat"} is equal to zero, except the pivot, the algorithm does nothing, but pivot @{term "i::nat"} is increased in a unit. Finally, if there is a position @{term "n::nat"} whose coefficient is different from zero, its row is interchanged with row @{term "i::nat"} and the bezout coefficients are used to produce a zero in its position.\ definition "echelon_form_of_column_k bezout A' k = (let (A, i) = A' in if (\m\from_nat i. A $ m $ from_nat k = 0) \ (i = nrows A) then (A, i) else if (\m>from_nat i. A $ m $ from_nat k = 0) then (A, i + 1) else let n = (LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n); interchange_A = interchange_rows A (from_nat i) n in (bezout_iterate (interchange_A) (nrows A - 1) (from_nat i) (from_nat k) bezout, i + 1))" definition "echelon_form_of_upt_k A k bezout = (fst (foldl (echelon_form_of_column_k bezout) (A,0) [0..The executable definition:\ context euclidean_space begin definition [code_unfold]: "echelon_form_of_euclidean A = echelon_form_of A euclid_ext2" end subsubsection\Properties of the bezout matrix\ lemma bezout_matrix_works1: assumes ib: "is_bezout_ext bezout" and a_not_b: "a \ b" shows "(bezout_matrix A a b j bezout ** A) $ a $ j = snd (snd (snd (snd (bezout (A $ a $ j) (A $ b $ j)))))" proof (unfold matrix_matrix_mult_def bezout_matrix_def Let_def, simp) let ?a = "(A $ a $ j)" let ?b = "(A $ b $ j)" let ?z = "bezout (A $ a $ j) (A $ b $ j)" obtain p q u v d where bz: "(p, q, u, v, d) = ?z" by (cases ?z, auto) from ib have foo: "(\a b. let (p, q, u, v, gcd_a_b) = bezout a b in p * a + q * b = gcd_a_b \ gcd_a_b dvd a \ gcd_a_b dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd gcd_a_b) \ gcd_a_b * u = - b \ gcd_a_b * v = a)" using is_bezout_ext_def [of bezout] by simp have foo: "p * ?a + q * ?b = d \ d dvd ?a \ d dvd ?b \ (\d'. d' dvd ?a \ d' dvd ?b \ d' dvd d) \ d * u = - ?b \ d * v = ?a" using ib using is_bezout_ext_def using bz [symmetric] using foo [of ?a ?b] by fastforce have pa_bq_d: "p * ?a + ?b * q = d" using foo by (auto simp add: mult.commute) define f where "f k = (if k = a then p else if a = a \ k = b then q else if a = b \ k = a then u else if a = b \ k = b then v else if a = k then 1 else 0) * A $ k $ j" for k have UNIV_rw: "UNIV = insert b (insert a (UNIV - {a} - {b}))" by auto have sum_rw: "sum f (insert a (UNIV - {a} - {b})) = f a + sum f (UNIV - {a} - {b})" by (rule sum.insert, auto) have sum0: "sum f (UNIV - {a} - {b}) = 0" by (rule sum.neutral, simp add: f_def) have "(\k\UNIV. (case bezout (A $ a $ j) (A $ b $ j) of (p, q, u, v, d) \ if k = a then p else if a = a \ k = b then q else if a = b \ k = a then u else if a = b \ k = b then v else if a = k then 1 else 0) * A $ k $ j) = (\k\UNIV. (if k = a then p else if a = a \ k = b then q else if a = b \ k = a then u else if a = b \ k = b then v else if a = k then 1 else 0) * A $ k $ j)" unfolding bz [symmetric] by auto also have "... = sum f UNIV" unfolding f_def .. also have "sum f UNIV = sum f (insert b (insert a (UNIV - {a} - {b})))" using UNIV_rw by simp also have "... = f b + sum f (insert a (UNIV - {a} - {b}))" by (rule sum.insert, auto, metis a_not_b) also have "... = f b + f a" unfolding sum_rw sum0 by simp also have "... = d" unfolding f_def using a_not_b bz [symmetric] by (auto, metis add.commute mult.commute pa_bq_d) also have "... = snd (snd (snd (snd (bezout (A $ a $ j) (A $ b $ j)))))" using bz by (metis snd_conv) finally show "(\k\UNIV. (case bezout (A $ a $ j) (A $ b $ j) of (p, q, u, v, d) \ if k = a then p else if a = a \ k = b then q else if a = b \ k = a then u else if a = b \ k = b then v else if a = k then 1 else 0) * A $ k $ j) = snd (snd (snd (snd (bezout (A $ a $ j) (A $ b $ j)))))" unfolding f_def by simp qed lemma bezout_matrix_not_zero: assumes ib: "is_bezout_ext bezout" and a_not_b: "a \ b" and Aaj: "A $ a $ j \ 0" shows "(bezout_matrix A a b j bezout ** A) $ a $ j \ 0" proof - have "(bezout_matrix A a b j bezout ** A) $ a $ j = snd (snd (snd(snd (bezout (A $ a $ j) (A $ b $ j)))))" using bezout_matrix_works1[OF ib a_not_b] . also have "... = (\a b. (case bezout a b of (_, _,_ ,_,gcd') \ (gcd'))) (A $ a $ j) (A $ b $ j)" by (simp add: split_beta) also have "... \ 0" using gcd'_zero[OF is_gcd_is_bezout_ext[OF ib]] Aaj by simp finally show ?thesis . qed lemma ua_vb_0: fixes a::"'a::bezout_domain" assumes ib: "is_bezout_ext bezout" and nz: "snd (snd (snd (snd (bezout a b)))) \ 0" shows "fst (snd (snd (bezout a b))) * a + fst (snd (snd (snd (bezout a b)))) * b = 0" proof- obtain p q u v d where bz: "(p, q, u, v, d) = bezout a b" by (cases "bezout a b", auto) from ib have foo: "(\a b. let (p, q, u, v, gcd_a_b) = bezout a b in p * a + q * b = gcd_a_b \ gcd_a_b dvd a \ gcd_a_b dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd gcd_a_b) \ gcd_a_b * u = - b \ gcd_a_b * v = a)" using is_bezout_ext_def [of bezout] by simp have "p * a + q * b = d \ d dvd a \ d dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd d) \ d * u = - b \ d * v = a" using foo [of a b] using bz by fastforce hence dub: "d * u = - b" and dva: "d * v = a" by (simp_all) hence "d * u * a + d * v * b = 0" using eq_neg_iff_add_eq_0 mult.commute mult_minus_left by auto hence "u * a + v * b = 0" by (metis (no_types, lifting) dub dva minus_minus mult_minus_left neg_eq_iff_add_eq_0 semiring_normalization_rules(18) semiring_normalization_rules(7)) thus ?thesis using bz [symmetric] by simp qed lemma bezout_matrix_works2: fixes A::"'a::bezout_domain^'cols^'rows" assumes ib: "is_bezout_ext bezout" and a_not_b: "a \ b" and not_0: "A $ a $ j \ 0 \ A $ b $ j \ 0" shows "(bezout_matrix A a b j bezout ** A) $ b $ j = 0" proof (unfold matrix_matrix_mult_def bezout_matrix_def Let_def, auto) let ?a = "(A $ a $ j)" let ?b = "(A $ b $ j)" let ?z = "bezout (A $ a $ j) (A $ b $ j)" from ib have foo: "(\a b. let (p, q, u, v, gcd_a_b) = bezout a b in p * a + q * b = gcd_a_b \ gcd_a_b dvd a \ gcd_a_b dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd gcd_a_b) \ gcd_a_b * u = - b \ gcd_a_b * v = a)" using is_bezout_ext_def [of bezout] by simp obtain p q u v d where bz: "(p, q, u, v, d) = ?z" by (cases ?z, auto) hence pib: "p * ?a + q * ?b = d \ d dvd ?a \ d dvd ?b \ (\d'. d' dvd ?a \ d' dvd ?b \ d' dvd d) \ d * u = - ?b \ d * v = ?a" using foo [of ?a ?b] by fastforce hence pa_bq_d: "p * ?a + ?b * q = d" by (simp add: mult.commute) have d_dvd_a: "d dvd ?a" using pib by auto have d_dvd_b: "d dvd -?b" using pib by auto have pa_bq_d: "p * ?a + ?b * q = d" using pa_bq_d by (simp add: mult.commute) define f where "f k = (if b = a \ k = a then p else if b = a \ k = b then q else if b = b \ k = a then u else if b = b \ k = b then v else if b = k then 1 else 0) * A $ k $ j" for k have UNIV_rw: "UNIV = insert b (insert a (UNIV - {a} - {b}))" by auto have sum_rw: "sum f (insert a (UNIV - {a} - {b})) = f a + sum f (UNIV - {a} - {b})" by (rule sum.insert, auto) have sum0: "sum f (UNIV - {a} - {b}) = 0" by (rule sum.neutral, simp add: f_def) have "(\k\UNIV. (case bezout (A $ a $ j) (A $ b $ j) of (p, q, u, v, d) \ if b = a \ k = a then p else if b = a \ k = b then q else if b = b \ k = a then u else if b = b \ k = b then v else if b = k then 1 else 0) * A $ k $ j) = sum f UNIV" unfolding f_def bz [symmetric] by simp also have "sum f UNIV = sum f (insert b (insert a (UNIV - {a} - {b})))" using UNIV_rw by simp also have "... = f b + sum f (insert a (UNIV - {a} - {b}))" by (rule sum.insert, auto, metis a_not_b) also have "... = f b + f a" unfolding sum_rw sum0 by simp also have "... = v * ?b + u * ?a" unfolding f_def using a_not_b by auto also have "... = u * ?a + v * ?b" by auto also have "... = 0" using ua_vb_0 [OF ib] bz by (metis fst_conv minus_minus minus_zero mult_eq_0_iff pib snd_conv) finally show "(\k\UNIV. (case bezout (A $ a $ j) (A $ b $ j) of (p, q, u, v, d) \ if b = a \ k = a then p else if b = a \ k = b then q else if b = b \ k = a then u else if b = b \ k = b then v else if b = k then 1 else 0) * A $ k $ j) = 0" . qed lemma bezout_matrix_preserves_previous_columns: assumes ib: "is_bezout_ext bezout" and i_not_j: "i \ j" and Aik: "A $ i $ k \ 0" and b_k: "ba b. let (p, q, u, v, gcd_a_b) = bezout a b in p * a + q * b = gcd_a_b \ gcd_a_b dvd a \ gcd_a_b dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd gcd_a_b) \ gcd_a_b * u = - b \ gcd_a_b * v = a)" using is_bezout_ext_def [of bezout] by simp obtain p q u v d where bz: "(p, q, u, v, d) = ?z" by (cases ?z, auto) have Aib: "A $ i $ b = 0" by (metis b_k i is_zero_row_upt_k_def to_nat_mono) have Ajb: "A $ j $ b = 0" by (metis b_k j is_zero_row_upt_k_def to_nat_mono) define f where "f ka = (if a = i \ ka = i then p else if a = i \ ka = j then q else if a = j \ ka = i then u else if a = j \ ka = j then v else if a = ka then 1 else 0) * A $ ka $ b" for ka show "(\ka\UNIV. (case bezout (A $ i $ k) (A $ j $ k) of (p, q, u, v, d) \ if a = i \ ka = i then p else if a = i \ ka = j then q else if a = j \ ka = i then u else if a = j \ ka = j then v else if a = ka then 1 else 0) * A $ ka $ b) = A $ a $ b" proof (cases "a=i") case True have "(\ka\UNIV. (case bezout (A $ i $ k) (A $ j $ k) of (p, q, u, v, d) \ if a = i \ ka = i then p else if a = i \ ka = j then q else if a = j \ ka = i then u else if a = j \ ka = j then v else if a = ka then 1 else 0) * A $ ka $ b) = sum f UNIV" unfolding f_def bz [symmetric] by simp also have "sum f UNIV = 0" by (rule sum.neutral, auto simp add: Aib Ajb f_def True i_not_j) also have "... = A $ a $ b" unfolding True using Aib by simp finally show ?thesis . next case False note a_not_i=False show ?thesis proof (cases "a=j") case True have "(\ka\UNIV. (case bezout (A $ i $ k) (A $ j $ k) of (p, q, u, v, d) \ if a = i \ ka = i then p else if a = i \ ka = j then q else if a = j \ ka = i then u else if a = j \ ka = j then v else if a = ka then 1 else 0) * A $ ka $ b) = sum f UNIV" unfolding f_def bz [symmetric] by simp also have "sum f UNIV = 0" by (rule sum.neutral, auto simp add: Aib Ajb f_def True i_not_j) also have "... = A $ a $ b" unfolding True using Ajb by simp finally show ?thesis . next case False have UNIV_rw: "UNIV = insert j (insert i (UNIV - {i} - {j}))" by auto have UNIV_rw2: "UNIV - {i} - {j} = insert a (UNIV - {i} - {j}-{a})" using False a_not_i by auto have sum0: "sum f (UNIV - {i} - {j} - {a}) = 0" by (rule sum.neutral, simp add: f_def) have "(\ka\UNIV. (case bezout (A $ i $ k) (A $ j $ k) of (p, q, u, v, d) \ if a = i \ ka = i then p else if a = i \ ka = j then q else if a = j \ ka = i then u else if a = j \ ka = j then v else if a = ka then 1 else 0) * A $ ka $ b) = sum f UNIV" unfolding f_def bz [symmetric] by simp also have "sum f UNIV = sum f (insert j (insert i (UNIV - {i} - {j})))" using UNIV_rw by simp also have "... = f j + sum f (insert i (UNIV - {i} - {j}))" by (rule sum.insert, auto, metis i_not_j) also have "... = sum f (insert i (UNIV - {i} - {j}))" unfolding f_def using False a_not_i by auto also have "... = f i + sum f (UNIV - {i} - {j})" by (rule sum.insert, auto) also have "... = sum f (UNIV - {i} - {j})" unfolding f_def using False a_not_i by auto also have "... = sum f (insert a (UNIV - {i} - {j} - {a}))" using UNIV_rw2 by simp also have "... = f a + sum f (UNIV - {i} - {j} - {a})" by (rule sum.insert, auto) also have "... = f a" unfolding sum0 by simp also have "... = A $ a $ b" unfolding f_def using False a_not_i by auto finally show ?thesis . qed qed qed lemma det_bezout_matrix: fixes A::"'a::{bezout_domain}^'cols^'rows::{finite,wellorder}" assumes ib: "is_bezout_ext bezout" and a_less_b: "a < b" and aj: "A $ a $ j \ 0" shows "det (bezout_matrix A a b j bezout) = 1" proof - let ?B = "bezout_matrix A a b j bezout" let ?a = "(A $ a $ j)" let ?b = "(A $ b $ j)" let ?z = "bezout ?a ?b" from ib have foo: "(\a b. let (p, q, u, v, gcd_a_b) = bezout a b in p * a + q * b = gcd_a_b \ gcd_a_b dvd a \ gcd_a_b dvd b \ (\d'. d' dvd a \ d' dvd b \ d' dvd gcd_a_b) \ gcd_a_b * u = - b \ gcd_a_b * v = a)" using is_bezout_ext_def [of bezout] by simp obtain p q u v d where bz: "(p, q, u, v, d) = ?z" by (cases ?z, auto) hence pib: "p * ?a + q * ?b = d \ d dvd ?a \ d dvd ?b \ (\d'. d' dvd ?a \ d' dvd ?b \ d' dvd d) \ d * u = - ?b \ d * v = ?a" using foo [of ?a ?b] by fastforce hence pa_bq_d: "p * ?a + ?b * q = d" by (simp add: mult.commute) have a_not_b: "a \ b" using a_less_b by auto have d_dvd_a: "d dvd ?a" using pib by auto have UNIV_rw: "UNIV = insert b (insert a (UNIV - {a} - {b}))" by auto show ?thesis proof (cases "p = 0") case True note p0=True have q_not_0: "q \ 0" proof (rule ccontr, simp) assume q: "q = 0" have "d = 0" using pib by (metis True q add.right_neutral mult.commute mult_zero_right) hence "A $ a $ j = 0 \ A $ b $ j = 0" by (metis aj d_dvd_a dvd_0_left_iff) thus False using aj by auto qed have d_not_0: "d \ 0" by (metis aj d_dvd_a dvd_0_left_iff) have qb_not_0: "q *(-?b) \ 0" by (metis d_not_0 mult_cancel_left1 neg_equal_0_iff_equal no_zero_divisors p0 pa_bq_d q_not_0 right_minus) have "det (interchange_rows ?B a b) = (\i\UNIV. (interchange_rows ?B a b) $ i $ i)" proof (rule det_upperdiagonal) fix i ja::'rows assume ja_i: "ja = -1" proof - define f where "f i = interchange_rows (bezout_matrix A a b j bezout) a b $ i $ i" for i have prod_rw: "prod f (insert a (UNIV - {a} - {b})) = f a * prod f (UNIV - {a} - {b})" by (rule prod.insert, simp_all) have prod1: "prod f (UNIV - {a} - {b}) = 1" by (rule prod.neutral) (simp add: f_def interchange_rows_def bezout_matrix_def Let_def) have "prod f UNIV = prod f (insert b (insert a (UNIV - {a} - {b})))" using UNIV_rw by simp also have "... = f b * prod f (insert a (UNIV - {a} - {b}))" proof (rule prod.insert, simp) show "b \ insert a (UNIV - {a} - {b})" using a_not_b by auto qed also have "... = f b * f a" unfolding prod_rw prod1 by auto also have "... = q * u" using a_not_b using bz [symmetric] unfolding f_def interchange_rows_def bezout_matrix_def Let_def by auto also have "... = -1" proof - let ?r = "q * u" have du_b: " d * u = -?b" using pib by auto hence "q * (-?b) = d * ?r" by (metis mult.left_commute) also have "... = (p * ?a + ?b * q) * ?r" unfolding pa_bq_d by auto also have "... = ?b * q * ?r" using True by auto also have "... = q * (-?b) * (-?r)" by auto finally show ?thesis using qb_not_0 unfolding mult_cancel_left1 by (metis minus_minus) qed finally show ?thesis unfolding f_def . qed finally have det_inter_1: "det (interchange_rows ?B a b) = - 1" . have "det (bezout_matrix A a b j bezout) = - 1 * det (interchange_rows ?B a b)" unfolding det_interchange_rows using a_not_b by auto thus ?thesis unfolding det_inter_1 by simp next case False define mult_b_dp where "mult_b_dp = mult_row ?B b (d * p)" define sum_ab where "sum_ab = row_add mult_b_dp b a ?b" have "det (sum_ab) = prod (\i. sum_ab $ i $ i) UNIV" proof (rule det_upperdiagonal) fix i j::'rows assume j_less_i: "j < i" have "d * p * u + ?b * p = 0" using pib by (metis eq_neg_iff_add_eq_0 mult_minus_left semiring_normalization_rules(16)) thus "sum_ab $ i $ j = 0" unfolding sum_ab_def mult_b_dp_def unfolding row_add_def unfolding mult_row_def bezout_matrix_def using a_not_b j_less_i a_less_b unfolding bz [symmetric] by auto qed also have "... = d * p" proof - define f where "f i = sum_ab $ i $ i" for i have prod_rw: "prod f (insert a (UNIV - {a} - {b})) = f a * prod f (UNIV - {a} - {b})" by (rule prod.insert, simp_all) have prod1: "prod f (UNIV - {a} - {b}) = 1" by (rule prod.neutral) (simp add: f_def sum_ab_def row_add_def mult_b_dp_def mult_row_def bezout_matrix_def Let_def) have ap_bq_d: "A $ a $ j * p + A $ b $ j * q = d" by (metis mult.commute pa_bq_d) have "prod f UNIV = prod f (insert b (insert a (UNIV - {a} - {b})))" using UNIV_rw by simp also have "... = f b * prod f (insert a (UNIV - {a} - {b}))" proof (rule prod.insert, simp) show "b \ insert a (UNIV - {a} - {b})" using a_not_b by auto qed also have "... = f b * f a" unfolding prod_rw prod1 by auto also have "... = (d * p * v + ?b * q) * p" unfolding f_def sum_ab_def row_add_def mult_b_dp_def mult_row_def bezout_matrix_def unfolding bz [symmetric] using a_not_b by auto also have "... = d * p" using pib ap_bq_d semiring_normalization_rules(16) by auto finally show ?thesis unfolding f_def . qed finally have "det (sum_ab) = d * p" . moreover have "det (sum_ab) = d * p * det ?B" unfolding sum_ab_def unfolding det_row_add'[OF not_sym[OF a_not_b]] unfolding mult_b_dp_def unfolding det_mult_row .. ultimately show ?thesis by (metis (erased, opaque_lifting) False aj d_dvd_a dvd_0_left_iff mult_cancel_left1 mult_eq_0_iff) qed qed lemma invertible_bezout_matrix: fixes A::"'a::{bezout_ring_div}^'cols^'rows::{finite,wellorder}" assumes ib: "is_bezout_ext bezout" and a_less_b: "a < b" and aj: "A $ a $ j \ 0" shows "invertible (bezout_matrix A a b j bezout)" unfolding invertible_iff_is_unit unfolding det_bezout_matrix[OF assms] by simp lemma echelon_form_upt_k_bezout_matrix: fixes A k and i::"'b::mod_type" assumes e: "echelon_form_upt_k A k" and ib: "is_bezout_ext bezout" and Aik_0: "A $ i $ from_nat k \ 0" and zero_i: "is_zero_row_upt_k i k A" and i_less_n: "i n" using i_less_n by simp have zero_n: "is_zero_row_upt_k n k A" by (metis assms(5) e echelon_form_upt_k_condition1 zero_i) have zero_i2: "is_zero_row_upt_k i (to_nat (from_nat k::'c)) A" using zero_i by (metis k ncols_def to_nat_from_nat_id) have zero_n2: "is_zero_row_upt_k n (to_nat (from_nat k::'c)) A" using zero_n by (metis k ncols_def to_nat_from_nat_id) show ?thesis unfolding echelon_form_upt_k_def proof (auto) fix ia j assume ia: "is_zero_row_upt_k ia k ?B" and ia_j: "ia < j" have ia_A: "is_zero_row_upt_k ia k A" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'c assume j_less_k: "to_nat j < k" have "A $ ia $ j = ?B $ ia $ j" proof (rule bezout_matrix_preserves_previous_columns [symmetric, OF ib i_not_n Aik_0 _ zero_i2 zero_n2]) show "j < from_nat k" using j_less_k k by (metis from_nat_mono from_nat_to_nat_id ncols_def) qed also have "... = 0" by (metis ia is_zero_row_upt_k_def j_less_k) finally show "A $ ia $ j = 0" . qed show "is_zero_row_upt_k j k ?B" proof (unfold is_zero_row_upt_k_def, clarify) fix ja::'c assume ja_less_k: "to_nat ja < k" have "?B $ j $ ja = A $ j $ ja" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_n Aik_0 _ zero_i2 zero_n2]) show "ja < from_nat k" using ja_less_k k by (metis from_nat_mono from_nat_to_nat_id ncols_def) qed also have "... = 0" by (metis e echelon_form_upt_k_condition1 ia_A ia_j is_zero_row_upt_k_def ja_less_k) finally show "?B $ j $ ja = 0" . qed next fix ia j assume ia_j: "ia < j" and not_zero_ia_B: "\ is_zero_row_upt_k ia k ?B" and not_zero_j_B: "\ is_zero_row_upt_k j k ?B" obtain na where na: "to_nat na < k" and Biana: "?B $ ia $ na \ 0" using not_zero_ia_B unfolding is_zero_row_upt_k_def by auto obtain na2 where na2: "to_nat na2 < k" and Bjna2: "?B $ j $ na2 \ 0" using not_zero_j_B unfolding is_zero_row_upt_k_def by auto have na_less_fn: "na < from_nat k" by (metis from_nat_mono from_nat_to_nat_id k na ncols_def) have "A $ ia $ na = ?B $ ia $ na" by (rule bezout_matrix_preserves_previous_columns [symmetric, OF ib i_not_n Aik_0 na_less_fn zero_i2 zero_n2]) also have "... \ 0" using Biana by simp finally have A: "A $ ia $ na \ 0" . have na_less_fn2: "na2 < from_nat k" by (metis from_nat_mono from_nat_to_nat_id k na2 ncols_def) have "A $ j $ na2 = ?B $ j $ na2" by (rule bezout_matrix_preserves_previous_columns [symmetric, OF ib i_not_n Aik_0 na_less_fn2 zero_i2 zero_n2]) also have "... \ 0" using Bjna2 by simp finally have A2: "A $ j $ na2 \ 0" . have not_zero_ia_A: "\ is_zero_row_upt_k ia k A" unfolding is_zero_row_upt_k_def using na A by auto have not_zero_j_A: "\ is_zero_row_upt_k j k A" unfolding is_zero_row_upt_k_def using na2 A2 by auto obtain na where A: "A $ ia $ na \ 0" and na_less_k: "to_nat na 0" and na2_less_k: "to_nat na2 0) = (LEAST na. A $ ia $ na \ 0)" proof (rule Least_equality) have "?B $ ia $ (LEAST na. A $ ia $ na \ 0) = A $ ia $ (LEAST na. A $ ia $ na \ 0)" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_n Aik_0 _ zero_i2 zero_n2]) show "(LEAST na. A $ ia $ na \ 0) < from_nat k" using Least_le A na_less_fn by fastforce qed also have "... \ 0" by (metis (mono_tags) A LeastI) finally show "?B $ ia $ (LEAST na. A $ ia $ na \ 0) \ 0" . fix y assume B_ia_y: "?B $ ia $ y \ 0" show "(LEAST na. A $ ia $ na \ 0) \ y" proof (cases "y 0" using B_ia_y . finally show "A $ ia $ y \ 0" . qed next case False show ?thesis using False by (metis (mono_tags) A Least_le dual_order.strict_iff_order le_less_trans na_less_fn not_le) qed qed have least_eq2: "(LEAST na. ?B $ j $ na \ 0) = (LEAST na. A $ j $ na \ 0)" proof (rule Least_equality) have "?B $ j $ (LEAST na. A $ j $ na \ 0) = A $ j $ (LEAST na. A $ j $ na \ 0)" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_n Aik_0 _ zero_i2 zero_n2]) show "(LEAST na. A $ j $ na \ 0) < from_nat k" using Least_le A2 na_less_fn2 by fastforce qed also have "... \ 0" by (metis (mono_tags) A2 LeastI) finally show "?B $ j $ (LEAST na. A $ j $ na \ 0) \ 0" . fix y assume B_ia_y: "?B $ j $ y \ 0" show "(LEAST na. A $ j $ na \ 0) \ y" proof (cases "y 0" using B_ia_y . finally show "A $ j $ y \ 0" . qed next case False show ?thesis using False by (metis (mono_tags) A2 Least_le dual_order.strict_iff_order le_less_trans na_less_fn2 not_le) qed qed show "(LEAST na. ?B $ ia $ na \ 0) < (LEAST na. ?B $ j $ na \ 0)" unfolding least_eq least_eq2 by (rule echelon_form_upt_k_condition2[OF e ia_j not_zero_ia_A not_zero_j_A]) qed qed lemma bezout_matrix_preserves_rest: assumes ib: "is_bezout_ext bezout" and a_not_n: "a\n" and i_not_n: "i\n" and a_not_i: "a\i" and Aik_0: "A $ i $ k \ 0" and zero_ikA: "is_zero_row_upt_k i (to_nat k) A" shows "(bezout_matrix A i n k bezout ** A) $ a $ b = A $ a $ b" unfolding matrix_matrix_mult_def unfolding bezout_matrix_def Let_def proof (auto simp add: a_not_n i_not_n a_not_i) have UNIV_rw: "UNIV = insert a (UNIV - {a})" by auto let ?f="(\k. (if a = k then 1 else 0) * A $ k $ b)" have sum0: "sum ?f (UNIV - {a}) = 0" by (rule sum.neutral, auto) have "sum ?f UNIV = sum ?f (insert a (UNIV - {a}))" using UNIV_rw by simp also have "... = ?f a + sum ?f (UNIV - {a})" by (rule sum.insert, simp_all) also have "... = ?f a" using sum0 by auto also have "... = A $ a $ b" by simp finally show "sum ?f UNIV = A $ a $ b" . qed text\Code equations to execute the bezout matrix\ definition "bezout_matrix_row A a b j bezout x = (let (p, q, u, v, d) = bezout (A $ a $ j) (A $ b $ j) in vec_lambda (\y. if x = a \ y = a then p else if x = a \ y = b then q else if x = b \ y = a then u else if x = b \ y = b then v else if x = y then 1 else 0))" lemma bezout_matrix_row_code [code abstract]: "vec_nth (bezout_matrix_row A a b j bezout x) = (let (p, q, u, v, d) = bezout (A $ a $ j) (A $ b $ j) in (\y. if x = a \ y = a then p else if x = a \ y = b then q else if x = b \ y = a then u else if x = b \ y = b then v else if x = y then 1 else 0))" unfolding bezout_matrix_row_def by (cases "bezout (A $ a $ j) (A $ b $ j)") auto lemma [code abstract]: "vec_nth (bezout_matrix A a b j bezout) = bezout_matrix_row A a b j bezout" unfolding bezout_matrix_def unfolding bezout_matrix_row_def[abs_def] Let_def by (cases "bezout (A $ a $ j) (A $ b $ j)") auto subsubsection\Properties of the bezout iterate function\ lemma bezout_iterate_not_zero: assumes Aik_0: "A $ i $ from_nat k \ 0" and n: "n n" and ib: "is_bezout_ext bezout" shows "bezout_iterate A n i (from_nat k) bezout $ i $ from_nat k \ 0" using Aik_0 n a proof (induct n arbitrary: A) case 0 have "to_nat i = 0" by (metis "0.prems"(3) le_0_eq) hence i0: "i=0" by (metis to_nat_eq_0) show ?case using "0.prems"(1) unfolding i0 by auto next case (Suc n) show ?case proof (cases "Suc n = to_nat i") case True show ?thesis unfolding bezout_iterate.simps using True Suc.prems(1) by simp next case False let ?B="(bezout_matrix A i (from_nat (Suc n)) (from_nat k) bezout ** A)" have i_le_n: "to_nat i < Suc n" using Suc.prems(3) False by auto have "bezout_iterate A (Suc n) i (from_nat k) bezout $ i $ from_nat k = bezout_iterate ?B n i (from_nat k) bezout $ i $ from_nat k" unfolding bezout_iterate.simps using i_le_n by auto also have "... \ 0" proof (rule Suc.hyps, rule bezout_matrix_not_zero[OF ib]) show "i \ from_nat (Suc n)" by (metis False Suc.prems(2) nrows_def to_nat_from_nat_id) show "A $ i $ from_nat k \ 0" by (rule Suc.prems(1)) show "n < nrows ?B" by (metis Suc.prems(2) Suc_lessD nrows_def) show "to_nat i \ n" using i_le_n by auto qed finally show ?thesis . qed qed lemma bezout_iterate_preserves: fixes A k and i::"'b::mod_type" assumes e: "echelon_form_upt_k A k" and ib: "is_bezout_ext bezout" and Aik_0: "A $ i $ from_nat k \ 0" and n: "n n" and k: "k from_nat (Suc n)" by (metis False Suc.prems(2) nrows_def to_nat_from_nat_id) let ?B="(bezout_matrix A i (from_nat (Suc n)) (from_nat k) bezout ** A)" have i_le_n: "to_nat i < Suc n" by (metis False Suc.prems(3) le_imp_less_or_eq) have "bezout_iterate A (Suc n) i (from_nat k) bezout $ a $ b = bezout_iterate ?B n i (from_nat k) bezout $ a $ b" unfolding bezout_iterate.simps using i_le_n by auto also have "... = ?B $ a $ b" proof (rule Suc.hyps) show "?B $ i $ from_nat k \ 0" by (metis False Suc.prems(1) Suc.prems(2) bezout_matrix_not_zero ib nrows_def to_nat_from_nat_id) show "n < nrows ?B" by (metis Suc.prems(2) Suc_lessD nrows_def) show "k < ncols ?B" by (metis Suc.prems(4) ncols_def) show "to_nat i \ n" using i_le_n by auto show "is_zero_row_upt_k i k ?B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'c assume j_k: "to_nat j < k" have j_k2: "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id j_k k ncols_def) have "?B $ i $ j = A $ i $ j" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_fn Suc.prems(1) j_k2], unfold to_nat_from_nat_id[OF Suc.prems(4)[unfolded ncols_def]]) show "is_zero_row_upt_k i k A" by (rule Suc.prems(5)) show "is_zero_row_upt_k (from_nat (Suc n)) k A" using echelon_form_upt_k_condition1[OF Suc.prems(6) Suc.prems(5)] by (metis Suc.prems(2) from_nat_mono from_nat_to_nat_id i_le_n nrows_def) qed also have "... = 0" by (metis Suc.prems(5) is_zero_row_upt_k_def j_k) finally show "?B $ i $ j = 0" . qed show "echelon_form_upt_k ?B k" proof (rule echelon_form_upt_k_bezout_matrix) show "echelon_form_upt_k A k" by (metis Suc.prems(6)) show "is_bezout_ext bezout" by (rule ib) show "A $ i $ from_nat k \ 0" by (rule Suc.prems(1)) show "is_zero_row_upt_k i k A" by (rule Suc.prems(5)) have "(from_nat (to_nat i)::'b)\from_nat (Suc n)" by (rule from_nat_mono'[OF Suc.prems(3) Suc.prems(2)[unfolded nrows_def]]) thus "i < from_nat (Suc n)" using i_not_fn by auto show "k < ncols A" by (rule Suc.prems(4)) qed qed also have "... = A $ a $ b" proof (rule bezout_matrix_preserves_previous_columns[OF ib]) show "i \ from_nat (Suc n)" by (metis False Suc.prems(2) nrows_def to_nat_from_nat_id) show "A $ i $ from_nat k \ 0" by (rule Suc.prems(1)) show "b < from_nat k" by (rule assms(5)) show "is_zero_row_upt_k i (to_nat (from_nat k::'c)) A" unfolding to_nat_from_nat_id[OF Suc.prems(4)[unfolded ncols_def]] by (rule Suc.prems(5)) show "is_zero_row_upt_k (from_nat (Suc n)) (to_nat (from_nat k::'c)) A" unfolding to_nat_from_nat_id[OF Suc.prems(4)[unfolded ncols_def]] by (metis Suc.prems(2) Suc.prems(5) Suc.prems(6) add_to_nat_def echelon_form_upt_k_condition1 from_nat_mono i_le_n monoid_add_class.add.right_neutral nrows_def to_nat_0) qed finally show ?thesis . qed qed lemma bezout_iterate_preserves_below_n: assumes e: "echelon_form_upt_k A k" and ib: "is_bezout_ext bezout" and Aik_0: "A $ i $ from_nat k \ 0" and n: "n n" and zero_upt_k_i: "is_zero_row_upt_k i k A" shows "bezout_iterate A n i (from_nat k) bezout $ a $ b = A $ a $ b" using Aik_0 n i_le_n k zero_upt_k_i e n_less_a proof (induct n arbitrary: A) case 0 show ?case unfolding bezout_iterate.simps .. next case (Suc n) show ?case proof (cases "Suc n = to_nat i") case True show ?thesis unfolding bezout_iterate.simps using True by simp next case False have i_not_fn:" i \ from_nat (Suc n)" by (metis False Suc.prems(2) nrows_def to_nat_from_nat_id) let ?B="(bezout_matrix A i (from_nat (Suc n)) (from_nat k) bezout ** A)" have i_le_n: "to_nat i < Suc n" by (metis False Suc.prems(3) le_imp_less_or_eq) have zero_ikB: "is_zero_row_upt_k i k ?B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'b assume j_k: "to_nat j < k" have j_k2: "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id j_k k ncols_def) have "?B $ i $ j = A $ i $ j" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_fn Suc.prems(1) j_k2], unfold to_nat_from_nat_id[OF Suc.prems(4)[unfolded ncols_def]]) show "is_zero_row_upt_k i k A" by (rule Suc.prems(5)) show "is_zero_row_upt_k (from_nat (Suc n)) k A" using echelon_form_upt_k_condition1[OF Suc.prems(6) Suc.prems(5)] by (metis Suc.prems(2) from_nat_mono from_nat_to_nat_id i_le_n nrows_def) qed also have "... = 0" by (metis Suc.prems(5) is_zero_row_upt_k_def j_k) finally show "?B $ i $ j = 0" . qed have "bezout_iterate A (Suc n) i (from_nat k) bezout $ a $ b = bezout_iterate ?B n i (from_nat k) bezout $ a $ b" unfolding bezout_iterate.simps using i_le_n by auto also have "... = ?B $ a $ b" proof (rule Suc.hyps) show "?B $ i $ from_nat k \ 0" by (metis Suc.prems(1) bezout_matrix_not_zero i_not_fn ib) show "n < nrows ?B" by (metis Suc.prems(2) Suc_lessD nrows_def) show "to_nat i \ n" by (metis i_le_n less_Suc_eq_le) show "k < ncols ?B" by (metis Suc.prems(4) ncols_def) show "is_zero_row_upt_k i k ?B" by (rule zero_ikB) show "echelon_form_upt_k ?B k" proof (rule echelon_form_upt_k_bezout_matrix[OF Suc.prems(6) ib Suc.prems(1) Suc.prems(5) _ Suc.prems(4)]) show "i < from_nat (Suc n)" by (metis (no_types) Suc.prems(7) add_to_nat_def dual_order.strict_iff_order from_nat_mono i_le_n le_less_trans monoid_add_class.add.right_neutral to_nat_0 to_nat_less_card) qed show "n < to_nat a" by (metis Suc.prems(7) Suc_lessD) qed also have "... = A $ a $ b" proof (rule bezout_matrix_preserves_rest[OF ib _ _ _ Suc.prems(1)]) show "a \ from_nat (Suc n)" by (metis Suc.prems(7) add_to_nat_def from_nat_mono less_irrefl monoid_add_class.add.right_neutral to_nat_0 to_nat_less_card) show "i \ from_nat (Suc n)" by (rule i_not_fn) show "a \ i" by (metis assms(7) n_less_a not_le) show "is_zero_row_upt_k i (to_nat (from_nat k::'b)) A" by (metis Suc.prems(4) Suc.prems(5) ncols_def to_nat_from_nat_id) qed finally show ?thesis . qed qed lemma bezout_iterate_zero_column_k: fixes A::"'a::bezout_domain^'cols::{mod_type}^'rows::{mod_type}" assumes e: "echelon_form_upt_k A k" and ib: "is_bezout_ext bezout" and Aik_0: "A $ i $ from_nat k \ 0" and n: "nn" and zero_upt_k_i: "is_zero_row_upt_k i k A" shows "bezout_iterate A n i (from_nat k) bezout $ a $ from_nat k = 0" using e Aik_0 n k a_n zero_upt_k_i proof (induct n arbitrary: A) case 0 show ?case unfolding bezout_iterate.simps using "0.prems"(5) i_le_a to_nat_from_nat to_nat_le to_nat_mono by fastforce next case (Suc n) show ?case proof (cases "Suc n = to_nat i") case True show ?thesis unfolding bezout_iterate.simps using True by (metis Suc.prems(5) i_le_a leD to_nat_mono) next case False have i_not_fn:" i \ from_nat (Suc n)" by (metis False Suc.prems(3) nrows_def to_nat_from_nat_id) let ?B="(bezout_matrix A i (from_nat (Suc n)) (from_nat k) bezout ** A)" have i_le_n: "to_nat i < Suc n" by (metis Suc.prems(5) i_le_a le_less_trans not_le to_nat_mono) have zero_ikB: "is_zero_row_upt_k i k ?B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j_k: "to_nat j < k" have j_k2: "j < from_nat k" using from_nat_mono[OF j_k Suc.prems(4)[unfolded ncols_def]] unfolding from_nat_to_nat_id . have "?B $ i $ j = A $ i $ j" proof (rule bezout_matrix_preserves_previous_columns[OF ib i_not_fn Suc.prems(2) j_k2], unfold to_nat_from_nat_id[OF Suc.prems(4)[unfolded ncols_def]]) show "is_zero_row_upt_k i k A" by (rule Suc.prems(6)) show "is_zero_row_upt_k (from_nat (Suc n)) k A" using echelon_form_upt_k_condition1[OF Suc.prems(1)] by (metis (mono_tags) Suc.prems(3) Suc.prems(6) add_to_nat_def from_nat_mono i_le_n monoid_add_class.add.right_neutral nrows_def to_nat_0) qed also have "... = 0" by (metis Suc.prems(6) is_zero_row_upt_k_def j_k) finally show "?B $ i $ j = 0" . qed have "bezout_iterate A (Suc n) i (from_nat k) bezout $ a $ (from_nat k) = bezout_iterate ?B n i (from_nat k) bezout $ a $ (from_nat k)" unfolding bezout_iterate.simps using i_le_n by auto also have "... = 0" proof (cases "to_nat a = Suc n") case True have "bezout_iterate ?B n i (from_nat k) bezout $ a $ (from_nat k) = ?B $ a $ from_nat k" proof (rule bezout_iterate_preserves_below_n[OF _ ib]) show "echelon_form_upt_k ?B k" by (metis (erased, opaque_lifting) Suc.prems(1) Suc.prems(2) Suc.prems(4) Suc.prems(6) True echelon_form_upt_k_bezout_matrix from_nat_to_nat_id i_le_a ib) show "?B $ i $ from_nat k \ 0" by (metis Suc.prems(2) bezout_matrix_not_zero i_not_fn ib) show "n < nrows ?B" by (metis Suc.prems(3) Suc_lessD nrows_def) show "n < to_nat a" by (metis True lessI) show "k < ncols ?B" by (metis Suc.prems(4) ncols_def) show "to_nat i \ n" by (metis i_le_n less_Suc_eq_le) show "is_zero_row_upt_k i k ?B" by (rule zero_ikB) qed also have "... = 0" by (metis Suc.prems(2) True bezout_matrix_works2 i_not_fn ib to_nat_from_nat) finally show ?thesis . next case False show ?thesis proof (rule Suc.hyps) show "echelon_form_upt_k ?B k" proof (rule echelon_form_upt_k_bezout_matrix [OF Suc.prems(1) ib Suc.prems(2) Suc.prems(6) _ Suc.prems(4)]) show "i < from_nat (Suc n)" by (metis (mono_tags) Suc.prems(3) from_nat_mono from_nat_to_nat_id i_le_n nrows_def) qed show "?B $ i $ from_nat k \ 0" by (metis Suc.prems(2) bezout_matrix_not_zero i_not_fn ib) show "n < nrows ?B" by (metis Suc.prems(3) Suc_lessD nrows_def) show "k < ncols ?B" by (metis Suc.prems(4) ncols_def) show "to_nat a \ n" by (metis False Suc.prems(5) le_SucE) show "is_zero_row_upt_k i k ?B" by (rule zero_ikB) qed qed finally show ?thesis . qed qed subsubsection\Proving the correctness\ lemma condition1_index_le_zero_row: fixes A k defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" assumes e: "echelon_form_upt_k A k" and "is_zero_row_upt_k a (Suc k) A" shows "from_nat i\a" proof (rule ccontr) have zero_ik: "is_zero_row_upt_k a k A" by (metis assms(3) is_zero_row_upt_k_le) assume a: "\ from_nat i \ (a::'a)" hence ai: "a < from_nat i" by simp show False proof (cases "(from_nat i::'a)=0") case True thus ?thesis using ai least_mod_type[of a] unfolding True from_nat_0 by auto next case False from a have "a \ from_nat i - 1" by (intro leI) (auto dest: le_Suc) also from False have "i \ 0" by (intro notI) (simp_all add: from_nat_0) hence "i = (i - 1) + 1" by simp also have "from_nat \ = from_nat (i - 1) + 1" by (rule from_nat_suc) finally have ai2: "a \ from_nat (i - 1)" by simp have "i = to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" using i False by (metis from_nat_0) hence "i - 1 = to_nat (GREATEST n. \ is_zero_row_upt_k n k A)" by simp hence "from_nat (i - 1) = (GREATEST n. \ is_zero_row_upt_k n k A)" using from_nat_to_nat_id by auto hence "\ is_zero_row_upt_k (from_nat (i - 1)) k A" using False GreatestI_ex i by (metis from_nat_to_nat_id to_nat_0) moreover have "is_zero_row_upt_k (from_nat (i - 1)) k A" using echelon_form_upt_k_condition1[OF e zero_ik] using ai2 zero_ik by (cases "a = from_nat (i - 1)", auto) ultimately show False by contradiction qed qed lemma condition1_part1: fixes A k defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" assumes e: "echelon_form_upt_k A k" and a: "is_zero_row_upt_k a (Suc k) A" and ab: "a < b" and all_zero: "\m\from_nat i. A $ m $ from_nat k = 0" shows "is_zero_row_upt_k b (Suc k) A" proof (rule is_zero_row_upt_k_suc) have zero_ik: "is_zero_row_upt_k a k A" by (metis assms(3) is_zero_row_upt_k_le) show "is_zero_row_upt_k b k A" using echelon_form_upt_k_condition1[OF e zero_ik] using ab by auto have "from_nat i\a" using condition1_index_le_zero_row[OF e a] all_zero unfolding i by auto thus "A $ b $ from_nat k = 0" using all_zero ab by auto qed lemma condition1_part2: fixes A k defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" assumes e: "echelon_form_upt_k A k" and a: "is_zero_row_upt_k a (Suc k) A" and ab: "a < b" and i_last: "i = nrows A" and all_zero: "\m>from_nat (nrows A). A $ m $ from_nat k = 0" shows "is_zero_row_upt_k b (Suc k) A" proof - have zero_ik: "is_zero_row_upt_k a k A" by (metis assms(3) is_zero_row_upt_k_le) have i_le_a: "from_nat i\a" using condition1_index_le_zero_row[OF e a] unfolding i . have "(from_nat (nrows A)::'a) = 0" unfolding nrows_def using from_nat_CARD . thus ?thesis using ab i_last i_le_a by (metis all_zero e echelon_form_upt_k_condition1 is_zero_row_upt_k_suc le_less_trans zero_ik) qed lemma condition1_part3: fixes A k bezout defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B: "B \ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes e: "echelon_form_upt_k A k" and ib: "is_bezout_ext bezout" and a: "is_zero_row_upt_k a (Suc k) B" and "a < b" and all_zero: "\m>from_nat i. A $ m $ from_nat k = 0" and i_not_last: "i \ nrows A" and i_le_m: "from_nat i \ m" and Amk_not_0: "A $ m $ from_nat k \ 0" shows "is_zero_row_upt_k b (Suc k) A" proof (rule is_zero_row_upt_k_suc) have AB: "A = B" unfolding B echelon_form_of_column_k_def Let_def using all_zero by auto have i_le_a: "from_nat i\a" using condition1_index_le_zero_row[OF e a[unfolded AB[symmetric]]] unfolding i . show "A $ b $ from_nat k = 0" by (metis i_le_a all_zero assms(6) le_less_trans) show "is_zero_row_upt_k b k A" by (metis (poly_guards_query) AB a assms(6) e echelon_form_upt_k_condition1 is_zero_row_upt_k_le) qed lemma condition1_part4: fixes A k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B: "B\ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes e: "echelon_form_upt_k A k" assumes a: "is_zero_row_upt_k a (Suc k) A" and i_nrows: "i = nrows A" shows "is_zero_row_upt_k b (Suc k) A" proof - have eq_G: "from_nat (i - 1) = (GREATEST n. \ is_zero_row_upt_k n k A)" by (metis One_nat_def Suc_eq_plus1 i_nrows diff_Suc_Suc diff_zero from_nat_to_nat_id i nrows_not_0) hence a_le: "a\from_nat (i - 1)" by (metis One_nat_def Suc_pred i_nrows lessI not_less not_less_eq nrows_def to_nat_from_nat_id to_nat_less_card to_nat_mono zero_less_card_finite) have not_zero_G: "\ is_zero_row_upt_k (from_nat(i - 1)) k A" unfolding eq_G by (metis (mono_tags) GreatestI_ex i_nrows i nrows_not_0) hence "\ is_zero_row_upt_k a k A" by (metis a_le dual_order.strict_iff_order e echelon_form_upt_k_condition1) hence "\ is_zero_row_upt_k a (Suc k) A" by (metis is_zero_row_upt_k_le) thus ?thesis using a by contradiction qed lemma condition1_part5: fixes A::"'a::bezout_domain^'cols::{mod_type}^'rows::{mod_type}" and k bezout defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B: "B \ fst((echelon_form_of_column_k bezout) (A,i) k)" assumes ib: "is_bezout_ext bezout" and e: "echelon_form_upt_k A k" assumes zero_a_B: "is_zero_row_upt_k a (Suc k) B" and ab: "a < b" and im: "from_nat i < m" and Amk_not_0: "A $ m $ from_nat k \ 0" and not_last_row: "i \ nrows A" and k: "k 0 \ (from_nat i) \ n)) (nrows A - Suc 0) (from_nat i) (from_nat k) bezout)" proof (rule is_zero_row_upt_k_suc) let ?least="(LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)" let ?interchange="(interchange_rows A (from_nat i) ?least)" let ?bezout_iterate="(bezout_iterate ?interchange (nrows A - Suc 0) (from_nat i) (from_nat k) bezout)" have B_eq: "B = ?bezout_iterate" unfolding B echelon_form_of_column_k_def Let_def fst_conv snd_conv using im Amk_not_0 not_last_row by auto have zero_ikA: "is_zero_row_upt_k (from_nat i) k A" proof (cases "\m. is_zero_row_upt_k m k A") case True thus ?thesis by simp next case False hence i_eq: "i=to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" unfolding i by auto show ?thesis proof (rule row_greater_greatest_is_zero, simp add: i_eq from_nat_to_nat_greatest,rule Suc_le') show " (GREATEST m. \ is_zero_row_upt_k m k A) + 1 \ 0" proof - have "\x\<^sub>1. \ x\<^sub>1 < i \ \ to_nat (GREATEST R. \ is_zero_row_upt_k R k A) < x\<^sub>1" using i_eq by linarith thus "(GREATEST m. \ is_zero_row_upt_k m k A) + 1 \ 0" by (metis One_nat_def add_Suc_right neq_iff from_nat_to_nat_greatest i_eq monoid_add_class.add.right_neutral nat.distinct(1) not_last_row nrows_def to_nat_0 to_nat_from_nat_id to_nat_less_card) qed qed qed have zero_interchange: "is_zero_row_upt_k (from_nat i) k ?interchange" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j_less_k: "to_nat j < k" have i_le_least: "from_nat i\?least" by (metis (mono_tags, lifting) Amk_not_0 LeastI2_wellorder less_imp_le im) hence zero_least_kA: "is_zero_row_upt_k ?least k A" using echelon_form_upt_k_condition1[OF e zero_ikA] by (metis (poly_guards_query) dual_order.strict_iff_order zero_ikA) have "?interchange $ from_nat i $ j = A $ ?least $ j" by simp also have "... = 0" using zero_least_kA j_less_k unfolding is_zero_row_upt_k_def by simp finally show "?interchange $ from_nat i $ j = 0" . qed have zero_a_k: "is_zero_row_upt_k a k A" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j_less_k: "to_nat j < k" have "?interchange $ a $ j = ?bezout_iterate $ a $ j" proof (rule bezout_iterate_preserves[symmetric]) show "echelon_form_upt_k ?interchange k" proof (rule echelon_form_upt_k_interchange[OF e zero_ikA Amk_not_0 _ k]) show "from_nat i \ m" using im by auto qed show "is_bezout_ext bezout" using ib . show "?interchange $ (from_nat i) $ from_nat k \ 0" by (metis (mono_tags, lifting) Amk_not_0 LeastI_ex dual_order.strict_iff_order im interchange_rows_i) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id j_less_k k ncols_def) show "to_nat (from_nat i::'rows) \ nrows A - Suc 0" - by (metis Suc_eq_plus1 Suc_le_mono Suc_pred - discrete nrows_def to_nat_less_card zero_less_card_finite) + by (simp add: nrows_def le_diff_conv2 Suc_le_eq to_nat_less_card) show "k < ncols ?interchange" using k unfolding ncols_def by auto show "is_zero_row_upt_k (from_nat i) k ?interchange" using zero_interchange . qed also have "... = 0" using zero_a_B j_less_k unfolding B_eq is_zero_row_upt_k_def by auto finally have *: "?interchange $ a $ j = 0" . show "A $ a $ j = 0" proof (cases "a=from_nat i") case True show ?thesis unfolding True using zero_ikA j_less_k unfolding is_zero_row_upt_k_def by auto next case False note a_not_i=False show ?thesis proof (cases "a=?least") case True show ?thesis using zero_interchange unfolding True is_zero_row_upt_k_def using j_less_k by auto next case False note a_not_least=False have "?interchange $ a $ j = A $ a $ j" using a_not_least a_not_i by (metis (erased, lifting) interchange_rows_preserves) thus ?thesis unfolding * .. qed qed qed hence zero_b_k: "is_zero_row_upt_k b k A" by (metis ab e echelon_form_upt_k_condition1) have i_le_a: "from_nat i\a" unfolding i proof (auto simp add: from_nat_to_nat_greatest from_nat_0) show "0 \ a" by (metis least_mod_type) fix m assume m: "\ is_zero_row_upt_k m k A" have "(GREATEST n. \ is_zero_row_upt_k n k A) < a" by (metis (no_types, lifting) GreatestI_ex neq_iff e echelon_form_upt_k_condition1 m zero_a_k) thus "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 \ a" by (metis le_Suc) qed have i_not_b: "from_nat i \ b" using i_le_a ab by simp show "is_zero_row_upt_k b k ?bezout_iterate" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j_less_k: "to_nat j < k" have "?bezout_iterate $ b $ j = ?interchange $ b $ j" proof (rule bezout_iterate_preserves) show "echelon_form_upt_k ?interchange k" proof (rule echelon_form_upt_k_interchange[OF e zero_ikA Amk_not_0 _ k]) show "from_nat i \ m" using im by auto qed show "is_bezout_ext bezout" using ib . show "?interchange $ from_nat i $ from_nat k \ 0" by (metis (mono_tags, lifting) Amk_not_0 LeastI_ex dual_order.strict_iff_order im interchange_rows_i) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id j_less_k k ncols_def) show "to_nat (from_nat i::'rows) \ nrows A - Suc 0" - by (metis Suc_eq_plus1 Suc_le_mono Suc_pred - discrete nrows_def to_nat_less_card zero_less_card_finite) + by (simp add: nrows_def le_diff_conv2 Suc_le_eq to_nat_less_card) show "k < ncols ?interchange" using k unfolding ncols_def by auto show "is_zero_row_upt_k (from_nat i) k ?interchange" by (rule zero_interchange) qed also have "... = A $ b $ j" proof (cases "b=?least") case True have "?interchange $ b $ j = A $ (from_nat i) $ j" using True by auto also have "... = A $ b $ j" using zero_b_k zero_ikA j_less_k unfolding is_zero_row_upt_k_def by auto finally show ?thesis . next case False show ?thesis using False using interchange_rows_preserves[OF i_not_b] by (metis (no_types, lifting)) qed also have "... = 0" using zero_b_k j_less_k unfolding is_zero_row_upt_k_def by auto finally show "?bezout_iterate $ b $ j = 0" . qed show "?bezout_iterate $ b $ from_nat k = 0" proof (rule bezout_iterate_zero_column_k[OF _ ib]) show "echelon_form_upt_k ?interchange k" proof (rule echelon_form_upt_k_interchange[OF e zero_ikA Amk_not_0 _ k]) show "from_nat i \ m" using im by auto qed show "?interchange $ from_nat i $ from_nat k \ 0" by (metis (mono_tags, lifting) Amk_not_0 LeastI_ex dual_order.strict_iff_order im interchange_rows_i) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "from_nat i < b" by (metis ab i_le_a le_less_trans) show "k < ncols ?interchange" by (metis (full_types, lifting) k ncols_def) show "to_nat b \ nrows A - Suc 0" by (metis Suc_pred leD not_less_eq_eq nrows_def to_nat_less_card zero_less_card_finite) show "is_zero_row_upt_k (from_nat i) k ?interchange" by (rule zero_interchange) qed qed lemma condition2_part1: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" and k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B:"B \ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes e: "echelon_form_upt_k A k" and ab: "a < b" and not_zero_aB: "\ is_zero_row_upt_k a (Suc k) B" and not_zero_bB: "\ is_zero_row_upt_k b (Suc k) B" and all_zero: "\m\from_nat i. A $ m $ from_nat k = 0" shows "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" proof - have B_eq_A: "B=A" unfolding B echelon_form_of_column_k_def Let_def fst_conv snd_conv using all_zero by auto show ?thesis proof (cases "\m. is_zero_row_upt_k m k A") case True have i0: "i = 0" unfolding i using True by simp have "is_zero_row_upt_k a k B" using True unfolding B_eq_A by auto moreover have "B $ a $ from_nat k = 0" using all_zero unfolding i0 from_nat_0 by (metis B_eq_A least_mod_type) ultimately have "is_zero_row_upt_k a (Suc k) B" by (rule is_zero_row_upt_k_suc) thus ?thesis using not_zero_aB by contradiction next case False note not_all_zero=False have i2: "i = to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" unfolding i using False by auto have not_zero_aA: "\ is_zero_row_upt_k a k A" by (metis (erased, lifting) B_eq_A GreatestI_ex add_to_nat_def all_zero neq_iff e echelon_form_upt_k_condition1 i2 is_zero_row_upt_k_suc le_Suc not_all_zero not_zero_aB to_nat_1) moreover have not_zero_bA: "\ is_zero_row_upt_k b k A" by (metis (erased, lifting) B_eq_A GreatestI_ex add_to_nat_def all_zero neq_iff e echelon_form_upt_k_condition1 i2 is_zero_row_upt_k_suc le_Suc not_all_zero not_zero_bB to_nat_1) ultimately show ?thesis using echelon_form_upt_k_condition2[OF e ab] by simp qed qed lemma condition2_part2: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" and k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" assumes e: "echelon_form_upt_k A k" and ab: "a < b" and all_zero: "\m>from_nat (nrows A). A $ m $ from_nat k = 0" and i_nrows: "i = nrows A" shows "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" proof - have not_all_zero: "\ (\m. is_zero_row_upt_k m k A)" by (metis i i_nrows nrows_not_0) have "(GREATEST m. \ is_zero_row_upt_k m k A) + 1 = 0" by (metis (mono_tags, lifting) add_0_right One_nat_def Suc_le' add_Suc_right i i_nrows less_not_refl less_trans_Suc nrows_def to_nat_less_card to_nat_mono zero_less_card_finite) hence g_minus_1: "(GREATEST m. \ is_zero_row_upt_k m k A) = - 1" by (simp add: a_eq_minus_1) have "\ is_zero_row_upt_k a k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show "a \ (GREATEST m. \ is_zero_row_upt_k m k A)" by (simp add: Greatest_is_minus_1 g_minus_1) qed moreover have "\ is_zero_row_upt_k b k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show "b \ (GREATEST m. \ is_zero_row_upt_k m k A)" by (simp add: Greatest_is_minus_1 g_minus_1) qed ultimately show ?thesis using echelon_form_upt_k_condition2[OF e ab] by simp qed lemma condition2_part3: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" and k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B:"B \ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes e: "echelon_form_upt_k A k" and k: "k is_zero_row_upt_k a (Suc k) B" and not_zero_bB: "\ is_zero_row_upt_k b (Suc k) B" and all_zero: "\m>from_nat i. A $ m $ from_nat k = 0" and i_ma: "from_nat i \ ma" and A_ma_k: "A $ ma $ from_nat k \ 0" shows "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" proof - have B_eq_A: "B=A" unfolding B echelon_form_of_column_k_def Let_def fst_conv snd_conv using all_zero by simp have not_all_zero: "\ (\m. is_zero_row_upt_k m k A)" by (metis B_eq_A ab all_zero from_nat_0 i is_zero_row_upt_k_suc le_less_trans least_mod_type not_zero_bB) have i2: "i = to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" unfolding i using not_all_zero by auto have not_zero_aA: "\ is_zero_row_upt_k a k A" proof - have "\x\<^sub>1 x\<^sub>2. from_nat (to_nat (x\<^sub>1::'rows) + 1) \ x\<^sub>2 \ \ x\<^sub>1 < x\<^sub>2" by (metis (no_types) add_to_nat_def le_Suc to_nat_1) moreover { assume "\ is_zero_row_upt_k b k A" hence "\ is_zero_row_upt_k a k A" using ab e echelon_form_upt_k_condition1 by blast } ultimately show "\ is_zero_row_upt_k a k A" by (metis B_eq_A greatest_less_zero_row ab all_zero le_imp_less_or_eq e i2 is_zero_row_upt_k_suc not_all_zero not_zero_aB not_zero_bB) qed show ?thesis proof (cases "\ is_zero_row_upt_k b k A") case True thus ?thesis using not_zero_aA echelon_form_upt_k_condition2[OF e ab] by simp next case False note zero_bA=False obtain v where Aav: "A $ a $ v \ 0" and v: "v 0) \ v" by (rule Least_le, simp add: Aav) have b_ge_greatest: "b>(GREATEST n. \ is_zero_row_upt_k n k A)" using False by (simp add: greatest_less_zero_row e not_all_zero) have i_eq_b: "from_nat i = b" proof (rule ccontr, cases "from_nat i < b") case True hence Abk_0: "A $ b $ from_nat k = 0" using all_zero by auto have "is_zero_row_upt_k b (Suc k) B" proof (rule is_zero_row_upt_k_suc) show "is_zero_row_upt_k b k B" using zero_bA unfolding B_eq_A by simp show "B $ b $ from_nat k = 0" using Abk_0 unfolding B_eq_A by simp qed thus False using not_zero_bB by contradiction next case False assume i_not_b: "from_nat i \ b" hence b_less_i: "from_nat i > b" using False by simp thus False using b_ge_greatest unfolding i by (metis (no_types, lifting) False Suc_less add_to_nat_def i2 i_not_b to_nat_1) qed have Abk_not_0: "A $ b $ from_nat k \ 0" using False not_zero_bB unfolding B_eq_A is_zero_row_upt_k_def by (metis B_eq_A False is_zero_row_upt_k_suc not_zero_bB) have "(LEAST n. A $ b $ n \ 0) = from_nat k" proof (rule Least_equality) show "A $ b $ from_nat k \ 0" by (rule Abk_not_0) show "\y. A $ b $ y \ 0 \ from_nat k \ y" by (metis False is_zero_row_upt_k_def k ncols_def not_less to_nat_from_nat_id to_nat_mono) qed thus ?thesis using least_v v by auto qed qed lemma condition2_part4: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" and k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" assumes e: "echelon_form_upt_k A k" and ab: "a < b" and i_nrows: "i = nrows A" shows "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" proof - have not_all_zero: "\ (\m. is_zero_row_upt_k m k A)" by (metis i_nrows i nrows_not_0) then have "i = to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" by (simp add: i) then have "nrows A = to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1" by (simp add: i_nrows) then have "CARD('rows) = mod_type_class.to_nat (GREATEST n. \ is_zero_row_upt_k n k A) + 1" unfolding nrows_def . then have "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 = 0" using to_nat_plus_one_less_card by auto hence g: "(GREATEST n. \ is_zero_row_upt_k n k A) = -1" by (simp add: a_eq_minus_1) have "\ is_zero_row_upt_k a k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show "a \ (GREATEST m. \ is_zero_row_upt_k m k A)" by (simp add: Greatest_is_minus_1 g) qed moreover have "\ is_zero_row_upt_k b k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show "b \ (GREATEST m. \ is_zero_row_upt_k m k A)" by (simp add: Greatest_is_minus_1 g) qed ultimately show ?thesis using echelon_form_upt_k_condition2[OF e ab] by simp qed lemma condition2_part5: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" and k bezout i defines i:"i\(if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B:"B \ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes ib: "is_bezout_ext bezout" and e: "echelon_form_upt_k A k" and k: "k is_zero_row_upt_k a (Suc k) B" and not_zero_bB: "\ is_zero_row_upt_k b (Suc k) B" and i_m:"from_nat i < m" and A_mk: "A $ m $ from_nat k \ 0" and i_not_nrows: "i \ nrows A" shows "(LEAST n. B $ a $ n \ 0) < (LEAST n. B $ b $ n \ 0)" proof - have B_eq: "B = bezout_iterate (interchange_rows A (from_nat i) (LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)) (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" unfolding B echelon_form_of_column_k_def Let_def fst_conv snd_conv using i_m A_mk i_not_nrows by auto let ?least="(LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)" let ?interchange="interchange_rows A (from_nat i) ?least" let ?greatest="(GREATEST n. \ is_zero_row_upt_k n k A)" have nrows_less: "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by auto have interchange_ik_not_zero: "?interchange $ from_nat i $ from_nat k \ 0" by (metis (mono_tags, lifting) A_mk LeastI_ex dual_order.strict_iff_order i_m interchange_rows_i) have k2: "k < ncols ?interchange" using k unfolding ncols_def by simp have to_nat_b: "to_nat b \ nrows A - Suc 0" by (metis Suc_pred leD not_less_eq_eq nrows_def to_nat_less_card zero_less_card_finite) have to_nat_from_nat_i: "to_nat (from_nat i::'rows) \ nrows A - Suc 0" using i_not_nrows unfolding nrows_def by (metis Suc_pred less_Suc_eq_le to_nat_less_card zero_less_card_finite) have not_all_zero: "\ (\m. is_zero_row_upt_k m k A)" proof (rule ccontr) assume all_zero: "\\(\m. is_zero_row_upt_k m k A)" hence zero_aA: "is_zero_row_upt_k a k A" and zero_bA: "is_zero_row_upt_k b k A" by auto have echelon_interchange: "echelon_form_upt_k ?interchange k" proof (rule echelon_form_upt_k_interchange[OF e _ A_mk _ k]) show "is_zero_row_upt_k (from_nat i) k A" using all_zero by auto show "from_nat i \ m" using i_m by auto qed have zero_i_interchange: "is_zero_row_upt_k (from_nat i) k ?interchange" using all_zero unfolding is_zero_row_upt_k_def by auto have zero_bB: "is_zero_row_upt_k b k B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j: "to_nat j < k" have "B $ b $ j = ?interchange $ b $ j" proof (unfold B_eq, rule bezout_iterate_preserves [OF echelon_interchange ib interchange_ik_not_zero nrows_less _ to_nat_from_nat_i k2 zero_i_interchange]) show "j < from_nat k" using j by (metis from_nat_mono from_nat_to_nat_id k ncols_def) qed also have "... = 0" using all_zero j unfolding is_zero_row_upt_k_def interchange_rows_def by auto finally show "B $ b $ j = 0" . qed have i_not_b: "from_nat i \ b" using i all_zero ab least_mod_type by (metis leD from_nat_0) have "B $ b $ from_nat k \ 0" by (metis is_zero_row_upt_k_suc not_zero_bB zero_bB) moreover have "B $ b $ from_nat k = 0" proof (unfold B_eq, rule bezout_iterate_zero_column_k [OF echelon_interchange ib interchange_ik_not_zero nrows_less _ k2 to_nat_b zero_i_interchange]) show "from_nat i < b" by (metis all_zero antisym_conv1 from_nat_0 i i_not_b least_mod_type) qed ultimately show False by contradiction qed have i2: "i = to_nat (?greatest) + 1" unfolding i using not_all_zero by auto have g_rw: "(from_nat (to_nat ?greatest + 1)) = ?greatest + 1" by (metis add_to_nat_def to_nat_1) have zero_least_kA: "is_zero_row_upt_k ?least k A" proof (rule row_greater_greatest_is_zero) have "?greatest < from_nat i" by (metis Suc_eq_plus1 Suc_le' add_to_nat_def neq_iff from_nat_0 from_nat_mono i2 i_not_nrows not_less_eq nrows_def to_nat_1 to_nat_less_card zero_less_Suc) also have "... \ ?least" by (metis (mono_tags, lifting) A_mk LeastI_ex dual_order.strict_iff_order i_m) finally show "?greatest < ?least" . qed have zero_ik_interchange: "is_zero_row_upt_k (from_nat i) k ?interchange" by (metis (no_types, lifting) interchange_rows_i is_zero_row_upt_k_def zero_least_kA) have echelon_form_interchange: "echelon_form_upt_k ?interchange k" proof (rule echelon_form_upt_k_interchange[OF e _ A_mk _ k]) show "is_zero_row_upt_k (from_nat i) k A" by (metis (mono_tags) greatest_ge_nonzero_row' Greatest_is_minus_1 Suc_le' a_eq_minus_1 e g_rw i2 row_greater_greatest_is_zero zero_least_kA) show "from_nat i \ m" using i_m by simp qed have b_le_i: "b \ from_nat i" proof (rule ccontr) assume "\ b \ from_nat i" hence b_gr_i: "b > from_nat i" by simp have "is_zero_row_upt_k b (Suc k) B" proof (rule is_zero_row_upt_k_suc) show "B $ b $ from_nat k = 0" by (unfold B_eq, rule bezout_iterate_zero_column_k[OF echelon_form_interchange ib interchange_ik_not_zero nrows_less b_gr_i k2 to_nat_b zero_ik_interchange]) show "is_zero_row_upt_k b k B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j_k: "to_nat j < k" have "B $ b $ j = ?interchange $ b $ j" proof (unfold B_eq, rule bezout_iterate_preserves[OF echelon_form_interchange ib interchange_ik_not_zero nrows_less _ to_nat_from_nat_i k2 zero_ik_interchange]) show "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id j_k k ncols_def) qed also have "... = 0" by (metis (erased, lifting) b_gr_i echelon_form_interchange echelon_form_upt_k_condition1 is_zero_row_upt_k_def j_k zero_ik_interchange) finally show "B $ b $ j = 0" . qed qed thus False using not_zero_bB by contradiction qed hence a_less_i: "a < from_nat i" using ab by simp have not_zero_aA: "\ is_zero_row_upt_k a k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show " a \ (GREATEST m. \ is_zero_row_upt_k m k A)" using a_less_i unfolding i2 g_rw by (metis le_Suc not_le) qed have least_eq1:"(LEAST n. B $ a $ n \ 0) = (LEAST n. A $ a $ n \ 0)" proof (rule Least_equality) have "B $ a $ (LEAST n. A $ a $ n \ 0) = ?interchange $ a $ (LEAST n. A $ a $ n \ 0)" proof (unfold B_eq, rule bezout_iterate_preserves[OF echelon_form_interchange ib interchange_ik_not_zero nrows_less _ to_nat_from_nat_i k2 zero_ik_interchange]) obtain j::'cols where j: "to_nat j < k" and Aaj: "A $ a $ j \ 0" using not_zero_aA unfolding is_zero_row_upt_k_def by auto have "(LEAST n. A $ a $ n \ 0) \ j" by (rule Least_le, simp add: Aaj) also have "... < from_nat k" by (metis (full_types) from_nat_mono from_nat_to_nat_id j k ncols_def) finally show "(LEAST n. A $ a $ n \ 0) < from_nat k" . qed also have "... = A $ a $ (LEAST n. A $ a $ n \ 0)" by (metis (no_types, lifting) ab b_le_i interchange_rows_preserves leD not_zero_aA zero_least_kA) also have "... \ 0" by (metis (mono_tags) LeastI is_zero_row_def' is_zero_row_imp_is_zero_row_upt not_zero_aA) finally show "B $ a $ (LEAST n. A $ a $ n \ 0) \ 0" . fix y assume Bay:"B $ a $ y \ 0" show "(LEAST n. A $ a $ n \ 0) \ y" proof (cases "y 0" using Bay by simp thus ?thesis using Least_le by fast next case False obtain j::'cols where j: "to_nat j < k" and Aaj: "A $ a $ j \ 0" using not_zero_aA unfolding is_zero_row_upt_k_def by auto have "(LEAST n. A $ a $ n \ 0) \ j" by (rule Least_le, simp add: Aaj) also have "... < from_nat k" by (metis (full_types) from_nat_mono from_nat_to_nat_id j k ncols_def) also have "...\ y" using False by auto finally show ?thesis by simp qed qed show ?thesis proof (cases "b=from_nat i") case True have zero_bB: "is_zero_row_upt_k b k B" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume jk:"to_nat j < k" have jk2: "j < from_nat k" by (metis from_nat_mono from_nat_to_nat_id jk k ncols_def) have "B $ b $ j = ?interchange $ b $ j" by (unfold B_eq, rule bezout_iterate_preserves[OF echelon_form_interchange ib interchange_ik_not_zero nrows_less jk2 to_nat_from_nat_i k2 zero_ik_interchange]) also have "... = A $ ?least $ j" unfolding True by auto also have "... = 0" using zero_least_kA jk unfolding is_zero_row_upt_k_def by simp finally show "B $ b $ j = 0" . qed have least_eq2: "(LEAST n. B $ b $ n \ 0) = from_nat k" proof (rule Least_equality) show "B $ b $ from_nat k \ 0" unfolding B_eq True by (rule bezout_iterate_not_zero[OF interchange_ik_not_zero nrows_less to_nat_from_nat_i ib]) show "\y. B $ b $ y \ 0 \ from_nat k \ y" by (metis is_zero_row_upt_k_def le_less_linear to_nat_le zero_bB) qed obtain j::'cols where j: "to_nat j < k" and Abj: "A $ a $ j \ 0" using not_zero_aA unfolding is_zero_row_upt_k_def by auto have "(LEAST n. A $ a $ n \ 0) \ j" by (rule Least_le, simp add: Abj) also have "... < from_nat k" by (metis (full_types) from_nat_mono from_nat_to_nat_id j k ncols_def) finally show ?thesis unfolding least_eq1 least_eq2 . next case False note b_not_i=False hence b_less_i: "b < from_nat i" using b_le_i by simp have not_zero_bA: "\ is_zero_row_upt_k b k A" proof (rule greatest_ge_nonzero_row'[OF e _ not_all_zero]) show " b \ (GREATEST m. \ is_zero_row_upt_k m k A)" using b_less_i unfolding i2 g_rw by (metis le_Suc not_le) qed have least_eq2: "(LEAST n. B $ b $ n \ 0) = (LEAST n. A $ b $ n \ 0)" proof (rule Least_equality) have "B $ b $ (LEAST n. A $ b $ n \ 0) = ?interchange $ b $ (LEAST n. A $ b $ n \ 0)" proof (unfold B_eq, rule bezout_iterate_preserves[OF echelon_form_interchange ib interchange_ik_not_zero nrows_less _ to_nat_from_nat_i k2 zero_ik_interchange]) obtain j::'cols where j: "to_nat j < k" and Abj: "A $ b $ j \ 0" using not_zero_bA unfolding is_zero_row_upt_k_def by auto have "(LEAST n. A $ b $ n \ 0) \ j" by (rule Least_le, simp add: Abj) also have "... < from_nat k" by (metis (full_types) from_nat_mono from_nat_to_nat_id j k ncols_def) finally show "(LEAST n. A $ b $ n \ 0) < from_nat k" . qed also have "... = A $ b $ (LEAST n. A $ b $ n \ 0)" by (metis (mono_tags) b_not_i interchange_rows_preserves not_zero_bA zero_least_kA) also have "... \ 0" by (metis (mono_tags) LeastI is_zero_row_def' is_zero_row_imp_is_zero_row_upt not_zero_bA) finally show "B $ b $ (LEAST n. A $ b $ n \ 0) \ 0" . fix y assume Bby:"B $ b $ y \ 0" show "(LEAST n. A $ b $ n \ 0) \ y" proof (cases "y 0" using Bby by simp thus ?thesis using Least_le by fast next case False obtain j::'cols where j: "to_nat j < k" and Abj: "A $ b $ j \ 0" using not_zero_bA unfolding is_zero_row_upt_k_def by auto have "(LEAST n. A $ b $ n \ 0) \ j" by (rule Least_le, simp add: Abj) also have "... < from_nat k" by (metis (full_types) from_nat_mono from_nat_to_nat_id j k ncols_def) also have "...\ y" using False by auto finally show ?thesis by simp qed qed show ?thesis unfolding least_eq1 least_eq2 by (rule echelon_form_upt_k_condition2[OF e ab not_zero_aA not_zero_bA]) qed qed lemma echelon_echelon_form_column_k: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" and k bezout defines i:"i \ (if \m. is_zero_row_upt_k m k A then 0 else to_nat ((GREATEST n. \ is_zero_row_upt_k n k A)) + 1)" defines B: "B \ fst ((echelon_form_of_column_k bezout) (A,i) k)" assumes ib: "is_bezout_ext bezout" and e: "echelon_form_upt_k A k" and k: "k a < b \ is_zero_row_upt_k b (Suc k) B" proof (unfold B echelon_form_of_column_k_def Let_def fst_conv snd_conv, auto) assume 1: "is_zero_row_upt_k a (Suc k) A" and 2: "a < b" and 3: "\m\from_nat i. A $ m $ from_nat k = 0" show "is_zero_row_upt_k b (Suc k) A" by (rule condition1_part1[OF e 1 2 3[unfolded i]]) next assume 1: "is_zero_row_upt_k a (Suc k) A" and 2: "a < b" and 3: "i = nrows A " and 4: "\m>from_nat (nrows A). A $ m $ from_nat k = 0" show "is_zero_row_upt_k b (Suc k) A" by (rule condition1_part2[OF e 1 2 3[unfolded i] 4]) next fix m assume 1: "is_zero_row_upt_k a (Suc k) ?B2" and 2: "a < b" and 3: "\m>from_nat i. A $ m $ from_nat k = 0" and 4: "i \ nrows A" and 5: "from_nat i \ m" and 6: "A $ m $ from_nat k \ 0" show "is_zero_row_upt_k b (Suc k) A" using condition1_part3[OF e ib _ 2 _ _ _ 6] using 1 3 4 5 unfolding i echelon_form_of_column_k_def Let_def fst_conv snd_conv by auto next fix m::'c assume 1: "is_zero_row_upt_k a (Suc k) A" and 2: "i = nrows A" show "is_zero_row_upt_k b (Suc k) A" by (rule condition1_part4[OF e 1 2[unfolded i]]) next let ?B2="(fst (if \m\from_nat i. A $ m $ from_nat k = 0 then (A, i) else if \m>from_nat i. A $ m $ from_nat k = 0 then (A, i + 1) else (bezout_iterate (interchange_rows A (from_nat i) (LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)) (nrows A - 1) (from_nat i) (from_nat k) bezout, i + 1)))" fix m assume 1: "is_zero_row_upt_k a (Suc k) ?B2" and 2: "a < b" and 3: "from_nat i < m" and 4: "A $ m $ from_nat k \ 0" and 5: "i \ nrows A" show "is_zero_row_upt_k b (Suc k) (bezout_iterate (interchange_rows A (from_nat i) (LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)) (nrows A - Suc 0) (from_nat i) (from_nat k) bezout)" using condition1_part5[OF ib e _ 2 _ 4 _ k] using 1 3 5 unfolding i echelon_form_of_column_k_def Let_def fst_conv snd_conv by auto qed next fix a b assume ab: "a < b" and not_zero_aB: "\ is_zero_row_upt_k a (Suc k) B" and not_zero_bB: "\ is_zero_row_upt_k b (Suc k) B" show "(LEAST n. B $ a $ n \ 0) < (LEAST n. B $ b $ n \ 0)" proof (unfold B echelon_form_of_column_k_def Let_def fst_conv snd_conv, auto) assume all_zero: "\m\from_nat i. A $ m $ from_nat k = 0" show "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" using condition2_part1[OF e ab] not_zero_aB not_zero_bB all_zero unfolding B i by simp next assume 1: "\m>from_nat (nrows A). A $ m $ from_nat k = 0" and 2: "i = nrows A" show "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" using condition2_part2[OF e ab 1] 2 unfolding i by simp next fix ma assume 1: "\m>from_nat i. A $ m $ from_nat k = 0" and 2: "from_nat i \ ma" and 3: "A $ ma $ from_nat k \ 0" show "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" using condition2_part3[OF e k ab _ _ _ _ 3] using 1 2 not_zero_aB not_zero_bB unfolding i B by auto next assume "i = nrows A" thus "(LEAST n. A $ a $ n \ 0) < (LEAST n. A $ b $ n \ 0)" using condition2_part4[OF e ab] unfolding i by simp next let ?B2="bezout_iterate (interchange_rows A (from_nat i) (LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)) (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" fix m assume 1: "from_nat i < m" and 2: "A $ m $ from_nat k \ 0" and 3: "i \ nrows A" have B_eq: "B=?B2" unfolding B echelon_form_of_column_k_def Let_def using 1 2 3 by auto show "(LEAST n. ?B2 $ a $ n \ 0) < (LEAST n. ?B2 $ b $ n \ 0)" using condition2_part5[OF ib e k ab _ _ _ 2] 1 3 not_zero_aB not_zero_bB unfolding i[symmetric] B[symmetric] unfolding B_eq by auto qed qed lemma echelon_foldl_condition1: assumes ib: "is_bezout_ext bezout" and "A $ ma $ from_nat (Suc k) \ 0" and k: "km. \ is_zero_row_upt_k m (Suc (Suc k)) (bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ from_nat (Suc k) \ 0)) (nrows A - Suc 0) 0 (from_nat (Suc k)) bezout)" proof (rule exI[of _ 0], unfold is_zero_row_upt_k_def, auto, rule exI[of _ "from_nat (Suc k)"], rule conjI) show "to_nat (from_nat (Suc k)) < Suc (Suc k)" by (metis from_nat_mono from_nat_to_nat_id less_irrefl not_less_eq to_nat_less_card) show "bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ from_nat (Suc k) \ 0)) (nrows A - Suc 0) 0 (from_nat (Suc k)) bezout $ 0 $ from_nat (Suc k) \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib]) show "interchange_rows A 0 (LEAST n. A $ n $ from_nat (Suc k) \ 0) $ 0 $ from_nat (Suc k) \ 0" by (metis (mono_tags, lifting) LeastI_ex assms(2) interchange_rows_i) show "nrows A - Suc 0 < nrows (interchange_rows A 0 (LEAST n. A $ n $ from_nat (Suc k) \ 0))" unfolding nrows_def by simp show "to_nat 0 \ nrows A - Suc 0" unfolding to_nat_0 nrows_def by simp qed qed lemma echelon_foldl_condition2: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes n: "\ is_zero_row_upt_k ma k A" and all_zero: "\m\ (GREATEST n. \ is_zero_row_upt_k n k A)+1. A $ m $ from_nat k = 0" shows "(GREATEST n. \ is_zero_row_upt_k n k A) = (GREATEST n. \ is_zero_row_upt_k n (Suc k) A)" proof (rule Greatest_equality[symmetric]) show "\ is_zero_row_upt_k (GREATEST n. \ is_zero_row_upt_k n k A) (Suc k) A" by (metis GreatestI_ex n is_zero_row_upt_k_le) fix y assume y: "\ is_zero_row_upt_k y (Suc k) A" show "y \ (GREATEST n. \ is_zero_row_upt_k n k A)" proof (rule ccontr) assume " \ y \ (GREATEST n. \ is_zero_row_upt_k n k A)" hence y2: "y > (GREATEST n. \ is_zero_row_upt_k n k A)" by simp hence "is_zero_row_upt_k y k A" by (metis row_greater_greatest_is_zero) moreover have "A $ y $ from_nat k = 0" by (metis (no_types, lifting) all_zero le_Suc y2) ultimately have "is_zero_row_upt_k y (Suc k) A" by (rule is_zero_row_upt_k_suc) thus False using y by contradiction qed qed lemma echelon_foldl_condition3: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" and Am0: "A $ m $ from_nat k \ 0" and all_zero: "\m. is_zero_row_upt_k m k A" and e: "echelon_form_upt_k A k" and k: "k < ncols A" shows "to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0)) (nrows A - (Suc 0)) 0 (from_nat k) bezout)) = 0" proof (unfold to_nat_eq_0, rule Greatest_equality) let ?interchange="(interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0))" let ?b="(bezout_iterate ?interchange (nrows A - (Suc 0)) 0 (from_nat k) bezout)" have b0k: "?b $ 0 $ from_nat k \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib]) show "interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0) $ 0 $ from_nat k \ 0" by (metis (mono_tags, lifting) LeastI Am0 interchange_rows_i) show "nrows A - (Suc 0) < nrows (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0))" unfolding nrows_def by simp show "to_nat 0 \ nrows A - (Suc 0)" unfolding to_nat_0 nrows_def by simp qed have least_eq: "(LEAST n. A $ n $ from_nat k \ 0) = (LEAST n. A $ n $ from_nat k \ 0 \ 0 \ n)" by (metis least_mod_type) have all_zero_below: "\a>0. ?b $ a $ from_nat k = 0" proof (auto) fix a::'rows assume a: "0 < a" show "bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0)) (nrows A - Suc 0) 0 (from_nat k) bezout $ a $ from_nat k = 0" proof (rule bezout_iterate_zero_column_k[OF _ ib _ _ a]) show "echelon_form_upt_k (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0)) k" proof(unfold from_nat_0[symmetric] least_eq, rule echelon_form_upt_k_interchange[OF e _ Am0 _ k]) show "is_zero_row_upt_k (from_nat 0) k A" by (metis all_zero) show "from_nat 0 \ m" unfolding from_nat_0 by (metis least_mod_type) qed show "interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0) $ 0 $ from_nat k \ 0" by (metis (mono_tags, lifting) Am0 LeastI interchange_rows_i) show "nrows A - Suc 0 < nrows (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0))" unfolding nrows_def by simp show "k < ncols (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0))" using k unfolding ncols_def by simp show "to_nat a \ nrows A - Suc 0" by (metis (erased, opaque_lifting) One_nat_def Suc_leI Suc_le_D diff_Suc_eq_diff_pred not_le nrows_def to_nat_less_card zero_less_diff) show "is_zero_row_upt_k 0 k (interchange_rows A 0 (LEAST n. A $ n $ from_nat k \ 0))" by (metis all_zero interchange_rows_i is_zero_row_upt_k_def) qed qed show "\ is_zero_row_upt_k 0 (Suc k) ?b" by (metis b0k is_zero_row_upt_k_def k lessI ncols_def to_nat_from_nat_id) fix y assume y: "\ is_zero_row_upt_k y (Suc k) ?b" show "y \ 0" proof (rule ccontr) assume "\ y \ 0" hence y2: "y>0" by simp have "is_zero_row_upt_k y (Suc k) ?b" proof (rule is_zero_row_upt_k_suc) show "is_zero_row_upt_k y k ?b" proof (unfold is_zero_row_upt_k_def, clarify) fix j::'cols assume j: "to_nat j < k" have "?b $ y $ j = ?interchange $ y $ j" proof (rule bezout_iterate_preserves[OF _ ib]) show "echelon_form_upt_k ?interchange k" proof (unfold least_eq from_nat_0[symmetric], rule echelon_form_upt_k_interchange[OF e _ Am0 _ k]) show "is_zero_row_upt_k (from_nat 0) k A" by (metis all_zero) show "from_nat 0 \ m" by (metis from_nat_0 least_mod_type) qed show "?interchange $ 0 $ from_nat k \ 0" by (metis (mono_tags, lifting) Am0 LeastI interchange_rows_i) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "j < from_nat k" by (metis (full_types) j from_nat_mono from_nat_to_nat_id k ncols_def) show "to_nat 0 \ nrows A - Suc 0" by (metis le0 to_nat_0) show "k < ncols ?interchange" using k unfolding ncols_def by simp show "is_zero_row_upt_k 0 k ?interchange" by (metis all_zero interchange_rows_i is_zero_row_upt_k_def) qed also have "... = 0" by (metis all_zero dual_order.strict_iff_order interchange_rows_j interchange_rows_preserves is_zero_row_upt_k_def j y2) finally show "?b $ y $ j = 0" . qed show "?b $ y $ from_nat k = 0" using all_zero_below using y2 by auto qed thus False using y by contradiction qed qed lemma echelon_foldl_condition4: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes all_zero: "\m>(GREATEST n. \ is_zero_row_upt_k n k A)+1. A $ m $ from_nat k = 0" and greatest_nrows: "Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n k A)) \ nrows A" and le_mb: "(GREATEST n. \ is_zero_row_upt_k n k A)+1 \ mb" and A_mb_k: "A $ mb $ from_nat k \ 0" shows "Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n k A)) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) A)" proof - let ?greatest = "(GREATEST n. \ is_zero_row_upt_k n k A)" have mb_eq: "mb=(GREATEST n. \ is_zero_row_upt_k n k A) + 1" by (metis all_zero le_mb A_mb_k le_less ) have "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 = (GREATEST n. \ is_zero_row_upt_k n (Suc k) A)" proof (rule Greatest_equality[symmetric]) show "\ is_zero_row_upt_k (?greatest + 1) (Suc k) A" by (metis (no_types, lifting) A_mb_k is_zero_row_upt_k_def less_Suc_eq less_trans mb_eq not_less_eq to_nat_from_nat_id to_nat_less_card) fix y assume y: "\ is_zero_row_upt_k y (Suc k) A" show "y \ ?greatest + 1" proof (rule ccontr) assume "\ y \ (GREATEST n. \ is_zero_row_upt_k n k A) + 1" hence y_greatest: "y > ?greatest + 1" by simp have "is_zero_row_upt_k y (Suc k) A" proof (rule is_zero_row_upt_k_suc) show "is_zero_row_upt_k y k A" proof (rule row_greater_greatest_is_zero) show "?greatest < y" using y_greatest greatest_nrows unfolding nrows_def by (metis Suc_eq_plus1 dual_order.strict_implies_order le_Suc' suc_not_zero to_nat_plus_one_less_card') qed show "A $ y $ from_nat k = 0" using all_zero y_greatest unfolding from_nat_to_nat_greatest by auto qed thus False using y by contradiction qed qed thus ?thesis by (metis (mono_tags, lifting) Suc_eq_plus1 Suc_lessI add_to_nat_def greatest_nrows nrows_def to_nat_1 to_nat_from_nat_id to_nat_less_card) qed lemma echelon_foldl_condition5: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes mb: "\ is_zero_row_upt_k mb k A" and nrows: "Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n k A)) = nrows A" shows "nrows A = Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) A))" by (metis (no_types, lifting) GreatestI Suc_lessI Suc_less_eq mb nrows from_nat_mono from_nat_to_nat_id is_zero_row_upt_k_le not_greater_Greatest nrows_def to_nat_less_card) lemma echelon_foldl_condition6: fixes A::"'a::{bezout_ring}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" and g_mc: "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 \ mc" and A_mc_k: "A $ mc $ from_nat k \ 0" shows "\m. \ is_zero_row_upt_k m (Suc k) (bezout_iterate (interchange_rows A ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) (LEAST n. A $ n $ from_nat k \ 0 \ (GREATEST n. \ is_zero_row_upt_k n k A) + 1 \ n)) (nrows A - Suc 0) ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) (from_nat k) bezout)" proof - let ?greatest="(GREATEST n. \ is_zero_row_upt_k n k A)" let ?interchange="interchange_rows A (?greatest + 1) (LEAST n. A $ n $ from_nat k \ 0 \ ?greatest + 1 \ n)" let ?B="(bezout_iterate ?interchange (nrows A - Suc 0) (?greatest + 1) (from_nat k) bezout)" have "?B $ (?greatest + 1) $ from_nat k \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib]) show "?interchange $ (?greatest + 1) $ from_nat k \ 0" by (metis (mono_tags, lifting) LeastI_ex g_mc A_mc_k interchange_rows_i) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "to_nat (?greatest + 1) \ nrows A - Suc 0" by (metis Suc_pred less_Suc_eq_le nrows_def to_nat_less_card zero_less_card_finite) qed thus ?thesis by (metis (no_types, lifting) from_nat_mono from_nat_to_nat_id is_zero_row_upt_k_def less_irrefl not_less_eq to_nat_less_card) qed lemma echelon_foldl_condition7: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" and e: "echelon_form_upt_k A k" and k: "k is_zero_row_upt_k mb k A" and not_nrows: "Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n k A)) \ nrows A" and g_mc: "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 \ mc" and A_mc_k: "A $ mc $ from_nat k \ 0" shows "Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n k A)) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (bezout_iterate (interchange_rows A ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) (LEAST n. A $ n $ from_nat k \ 0 \ (GREATEST n. \ is_zero_row_upt_k n k A) + 1 \ n)) (nrows A - Suc 0) ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) (from_nat k) bezout))" proof - let ?greatest="(GREATEST n. \ is_zero_row_upt_k n k A)" let ?interchange="interchange_rows A (?greatest + 1) (LEAST n. A $ n $ from_nat k \ 0 \ ?greatest + 1 \ n)" let ?B="(bezout_iterate ?interchange (nrows A - Suc 0) (?greatest + 1) (from_nat k) bezout)" have g_rw: "(GREATEST n. \ is_zero_row_upt_k n k A) + 1 = from_nat (to_nat ((GREATEST n. \ is_zero_row_upt_k n k A) + 1))" unfolding from_nat_to_nat_id .. have B_gk:"?B $ (?greatest + 1) $ from_nat k \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib]) show "?interchange $ ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) $ from_nat k \ 0" by (metis (mono_tags, lifting) LeastI_ex g_mc A_mc_k interchange_rows_i) show "nrows A - Suc 0 < nrows (?interchange)" unfolding nrows_def by simp show "to_nat (?greatest + 1) \ nrows A - Suc 0" by (metis Suc_pred less_Suc_eq_le nrows_def to_nat_less_card zero_less_card_finite) qed have "(GREATEST n. \ is_zero_row_upt_k n (Suc k) ?B) = ?greatest + 1" proof (rule Greatest_equality) show "\ is_zero_row_upt_k (?greatest + 1) (Suc k) ?B" by (metis (no_types, lifting) B_gk from_nat_mono from_nat_to_nat_id is_zero_row_upt_k_def less_irrefl not_less_eq to_nat_less_card) fix y assume y: "\ is_zero_row_upt_k y (Suc k) ?B" show "y \ ?greatest + 1" proof (rule ccontr) assume " \ y \ (GREATEST n. \ is_zero_row_upt_k n k A) + 1" hence y_gr: " y > (GREATEST n. \ is_zero_row_upt_k n k A) + 1" by simp hence y_gr2: "y > (GREATEST n. \ is_zero_row_upt_k n k A)" by (metis (erased, lifting) Suc_eq_plus1 leI le_Suc' less_irrefl less_trans not_nrows nrows_def suc_not_zero to_nat_plus_one_less_card') have echelon_interchange: "echelon_form_upt_k ?interchange k" proof (subst (1 2) from_nat_to_nat_id [of "(GREATEST n. \ is_zero_row_upt_k n k A) + 1", symmetric], rule echelon_form_upt_k_interchange[OF e _ A_mc_k _ k]) show "is_zero_row_upt_k (from_nat (to_nat ((GREATEST n. \ is_zero_row_upt_k n k A) + 1))) k A" by (metis Suc_eq_plus1 Suc_le' g_rw not_nrows nrows_def row_greater_greatest_is_zero suc_not_zero) show "from_nat (to_nat ((GREATEST n. \ is_zero_row_upt_k n k A) + 1)) \ mc" by (metis g_mc g_rw) qed have i: "?interchange $ (?greatest + 1) $ from_nat k \ 0" by (metis (mono_tags, lifting) A_mc_k LeastI_ex g_mc interchange_rows_i) have zero_greatest: "is_zero_row_upt_k (?greatest + 1) k A" by (metis Suc_eq_plus1 Suc_le' not_nrows nrows_def row_greater_greatest_is_zero suc_not_zero) { fix j::'cols assume "to_nat j < k" have "?greatest < ?greatest + 1" by (metis greatest_less_zero_row e mb zero_greatest) also have "... \(LEAST n. A $ n $ from_nat k \ 0 \ (?greatest + 1) \ n)" by (metis (mono_tags, lifting) A_mc_k LeastI_ex g_mc) finally have least_less: "?greatest < (LEAST n. A $ n $ from_nat k \ 0 \ (?greatest + 1) \ n)" . have "is_zero_row_upt_k (LEAST n. A $ n $ from_nat k \ 0 \ (?greatest + 1) \ n) k A" by (rule row_greater_greatest_is_zero[OF least_less]) } hence zero_g1: "is_zero_row_upt_k (?greatest + 1) k ?interchange" unfolding is_zero_row_upt_k_def by auto hence zero_y: "is_zero_row_upt_k y k ?interchange" by (metis (erased, lifting) echelon_form_upt_k_condition1' echelon_interchange y_gr) have "is_zero_row_upt_k y (Suc k) ?B" proof (rule is_zero_row_upt_k_suc) show "?B $ y $ from_nat k = 0" proof (rule bezout_iterate_zero_column_k[OF echelon_interchange ib i _ y_gr _ _ zero_g1]) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "k < ncols ?interchange" using k unfolding ncols_def by simp show "to_nat y \ nrows A - Suc 0" by (metis One_nat_def Suc_eq_plus1 Suc_leI nrows_def le_diff_conv2 to_nat_less_card zero_less_card_finite) qed show "is_zero_row_upt_k y k ?B" proof (subst is_zero_row_upt_k_def, clarify) fix j::'cols assume j: "to_nat j < k" have "?B $ y $ j = ?interchange $ y $ j" proof (rule bezout_iterate_preserves[OF echelon_interchange ib i _ _ _ _ zero_g1]) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "j < from_nat k" using j by (metis (poly_guards_query) from_nat_mono from_nat_to_nat_id k ncols_def) show "to_nat ((GREATEST n. \ is_zero_row_upt_k n k A) + 1) \ nrows A - Suc 0" by (metis Suc_pred less_Suc_eq_le nrows_def to_nat_less_card zero_less_card_finite) show "k < ncols ?interchange" using k unfolding ncols_def . qed also have "... = 0" using zero_y unfolding is_zero_row_upt_k_def using j by simp finally show "?B $ y $ j = 0" . qed qed thus False using y by contradiction qed qed thus ?thesis by (metis (erased, lifting) Suc_eq_plus1 add_to_nat_def not_nrows nrows_def suc_not_zero to_nat_1 to_nat_from_nat_id to_nat_plus_one_less_card') qed lemma fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes k: "km. is_zero_row_upt_k m (Suc k) (fst (foldl (echelon_form_of_column_k bezout) (A, 0) [0.. is_zero_row_upt_k n (Suc k) (fst (foldl (echelon_form_of_column_k bezout) (A, 0) [0..m. is_zero_row_upt_k m 0 A then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n 0 A) + 1) = 0" unfolding is_zero_row_upt_k_def by auto show "echelon_form_upt_k (echelon_form_of_upt_k A 0 bezout) (Suc 0)" unfolding echelon_form_of_upt_k_def by (auto, subst i_rw[symmetric], rule echelon_echelon_form_column_k[OF ib echelon_form_upt_k_0], simp add: ncols_def) have rw_upt: "[0..m. is_zero_row_upt_k m (Suc 0) (fst (foldl (echelon_form_of_column_k bezout) (A, 0) [0.. is_zero_row_upt_k n (Suc 0) (fst (foldl (echelon_form_of_column_k bezout) (A, 0) [0.. 0" and all_zero: "\m. bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ m $ 0 = 0" have "bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ 0 $ 0 = bezout_iterate ?interchange (nrows A - Suc 0) 0 (from_nat 0) bezout $ 0 $ from_nat 0" using from_nat_0 by metis also have "... \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib], simp_all add: nrows_def) show "A $ (LEAST n. A $ n $ 0 \ 0) $ from_nat 0 \ 0" by (metis (mono_tags) LeastI \A $ m $ 0 \ 0\ from_nat_0) show "to_nat 0 \ CARD('rows) - Suc 0" by (metis le0 to_nat_0) qed finally have " bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ 0 $ 0 \ 0" . thus "A $ mb $ 0 = 0" using all_zero by auto next fix m assume Am0: "A $ m $ 0 \ 0" and all_zero: "\m>0. A $ m $ 0 = 0" thus "(GREATEST n. A $ n $ 0 \ 0) = 0" by (metis (mono_tags, lifting) GreatestI neq_iff not_less0 to_nat_0 to_nat_mono) next fix m ma mb assume "A $ m $ 0 \ 0" and "bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ 0 \ 0)) (nrows A - Suc 0) 0 0 bezout $ ma $ 0 \ 0" have "bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ 0 $ 0 = bezout_iterate ?interchange (nrows A - Suc 0) 0 (from_nat 0) bezout $ 0 $ from_nat 0" using from_nat_0 by metis also have "... \ 0" proof (rule bezout_iterate_not_zero[OF _ _ _ ib], simp_all add: nrows_def) show "A $ (LEAST n. A $ n $ 0 \ 0) $ from_nat 0 \ 0" by (metis (mono_tags) LeastI \A $ m $ 0 \ 0\ from_nat_0) show "to_nat 0 \ CARD('rows) - Suc 0" by (metis le0 to_nat_0) qed finally have 1: "bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ 0 $ 0 \ 0" . have 2: "\m>0. bezout_iterate ?interchange (nrows A - Suc 0) 0 0 bezout $ m $ 0 = 0" proof (auto) fix b::'rows assume b: "0 0" by (metis (mono_tags, lifting) LeastI_ex \A $ m $ 0 \ 0\ from_nat_0 interchange_rows_i) show "nrows A - Suc 0 < nrows (?interchange)" unfolding nrows_def by simp show "0 < b" using b . show "0 < ncols (?interchange)" unfolding ncols_def by auto show "to_nat b \ nrows A - Suc 0" - by (metis Suc_eq_plus1 discrete less_one add.left_neutral not_le - nrows_def nrows_not_0 le_diff_conv2 to_nat_less_card) + by (simp add: nrows_def le_diff_conv2 Suc_le_eq to_nat_less_card) show "is_zero_row_upt_k 0 0 (?interchange)" by (metis is_zero_row_utp_0) qed finally show "bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ 0 \ 0)) (nrows A - Suc 0) 0 0 bezout $ b $ 0 = 0" . qed show "(GREATEST n. bezout_iterate (interchange_rows A 0 (LEAST n. A $ n $ 0 \ 0)) (nrows A - Suc 0) 0 0 bezout $ n $ 0 \ 0) = 0" apply (rule Greatest_equality, simp add: 1) using 2 by force qed next fix k let ?fold="(foldl (echelon_form_of_column_k bezout)(A, 0) [0.. echelon_form_upt_k (echelon_form_of_upt_k A k bezout) (Suc k))" and "(k < ncols A \ foldl (echelon_form_of_column_k bezout) (A, 0) [0..m. is_zero_row_upt_k m (Suc k) (fst ?fold2) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (fst ?fold2)) + 1))" and Suc_k: "Suc k < ncols A" hence hyp_foldl: "foldl (echelon_form_of_column_k bezout) (A, 0) [0..m. is_zero_row_upt_k m (Suc k) (fst ?fold2) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (fst ?fold2)) + 1)" and hyp_echelon: "echelon_form_upt_k (echelon_form_of_upt_k A k bezout) (Suc k)" by auto have rw: "[0..m. is_zero_row_upt_k m (Suc k) (echelon_form_of_upt_k A k bezout) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (echelon_form_of_upt_k A k bezout)) + 1)" unfolding echelon_form_of_upt_k_def using hyp_foldl by fast show "echelon_form_upt_k (echelon_form_of_upt_k A (Suc k) bezout) (Suc (Suc k))" unfolding echelon_form_of_upt_k_def unfolding rw unfolding foldl_append unfolding foldl.simps unfolding rw2 proof (rule echelon_echelon_form_column_k[OF ib hyp_echelon]) show "Suc k < ncols (echelon_form_of_upt_k A k bezout)" using Suc_k unfolding ncols_def . qed show "foldl (echelon_form_of_column_k bezout) (A, 0) [0..m. is_zero_row_upt_k m (Suc (Suc k)) (fst ?fold) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (fst ?fold)) + 1)" proof (rule prod_eqI, metis fst_conv) define A' where "A' = fst ?fold2" let ?greatest="(GREATEST n. \ is_zero_row_upt_k n (Suc k) A')" have k: "k < ncols A'" using Suc_k unfolding ncols_def by auto have k2: "Suc k < ncols A'" using Suc_k unfolding ncols_def by auto have fst_snd_foldl: "snd ?fold2 = snd (fst ?fold2, if \m. is_zero_row_upt_k m (Suc k) (fst ?fold2) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc k) (fst ?fold2)) + 1)" using hyp_foldl by simp have ncols_eq: "ncols A = ncols A'" unfolding A'_def ncols_def .. have rref_A': "echelon_form_upt_k A' (Suc k)" using hyp_echelon unfolding A'_def echelon_form_of_upt_k_def . show "snd ?fold = snd (fst ?fold, if \m. is_zero_row_upt_k m (Suc (Suc k)) (fst ?fold) then 0 else to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (fst ?fold)) + 1)" using [[unfold_abs_def = false]] unfolding fst_conv snd_conv unfolding rw unfolding foldl_append unfolding foldl.simps unfolding echelon_form_of_column_k_def Let_def split_beta fst_snd_foldl unfolding A'_def[symmetric] proof (auto simp add: least_mod_type from_nat_0 from_nat_to_nat_greatest) fix m assume "A' $ m $ from_nat (Suc k) \ 0" thus "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" unfolding is_zero_row_upt_k_def by (metis add_to_nat_def from_nat_mono less_irrefl monoid_add_class.add.right_neutral not_less_eq to_nat_0 to_nat_less_card)+ next fix m assume "\ is_zero_row_upt_k m (Suc k) A'" thus "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) A'" by (metis is_zero_row_upt_k_le)+ next fix m assume "\ma. is_zero_row_upt_k ma (Suc k) A'" and "\mb. A' $ mb $ from_nat (Suc k) = 0" thus "is_zero_row_upt_k m (Suc (Suc k)) A'" by (metis is_zero_row_upt_k_suc) next fix ma assume "\m>0. A' $ m $ from_nat (Suc k) = 0" and "\m. is_zero_row_upt_k m (Suc k) A'" and "\ is_zero_row_upt_k ma (Suc (Suc k)) A'" thus "to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A') = 0" and "to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A') = 0" by (metis (erased, lifting) GreatestI_ex le_less is_zero_row_upt_k_suc least_mod_type to_nat_0)+ next fix m assume "\m>0. A' $ m $ from_nat (Suc k) = 0" and "\ is_zero_row_upt_k m (Suc k) A'" and "\m\?greatest+1. A' $ m $ from_nat (Suc k) = 0" thus "?greatest = (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A')" by (metis (mono_tags, lifting) echelon_form_upt_k_condition1 from_nat_0 is_zero_row_upt_k_le is_zero_row_upt_k_suc less_nat_zero_code neq_iff rref_A' to_nat_le) next fix m ma assume "\m>?greatest+1. A' $ m $ from_nat (Suc k) = 0" and "\m>0. A' $ m $ from_nat (Suc k) = 0" and"Suc (to_nat ?greatest) \ nrows A'" and "?greatest + 1 \ ma" and "A' $ ma $ from_nat (Suc k) \ 0" thus "Suc (to_nat ?greatest) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A')" by (metis (mono_tags) Suc_eq_plus1 less_linear leD least_mod_type nrows_def suc_not_zero) next fix m ma assume "\m>?greatest + 1. A' $ m $ from_nat (Suc k) = 0" and "\m>0. A' $ m $ from_nat (Suc k) = 0" and "\ is_zero_row_upt_k m (Suc k) A'" and "Suc (to_nat ?greatest) = nrows A'" and "\ is_zero_row_upt_k ma (Suc (Suc k)) A'" thus "nrows A' = Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A'))" by (metis echelon_foldl_condition5) next fix ma assume 1: "A' $ ma $ from_nat (Suc k) \ 0" show "\m. \ is_zero_row_upt_k m (Suc (Suc k)) (bezout_iterate (interchange_rows A' 0 (LEAST n. A' $ n $ from_nat (Suc k) \ 0)) (nrows A' - Suc 0) 0 (from_nat (Suc k)) bezout)" and "\m. \ is_zero_row_upt_k m (Suc (Suc k)) (bezout_iterate (interchange_rows A' 0 (LEAST n. A' $ n $ from_nat (Suc k) \ 0)) (nrows A' - Suc 0) 0 (from_nat (Suc k)) bezout)" by (rule echelon_foldl_condition1[OF ib 1 k])+ next fix m ma mb assume 1: "\ is_zero_row_upt_k ma (Suc k) A'" and 2:"\m\?greatest + 1. A' $ m $ from_nat (Suc k) = 0" show "?greatest = (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A')" by (rule echelon_foldl_condition2[OF 1 2]) next fix m assume 1: "A' $ m $ from_nat (Suc k) \ 0" and 2: "\m. is_zero_row_upt_k m (Suc k) A'" show "to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (bezout_iterate (interchange_rows A' 0 (LEAST n. A' $ n $ from_nat (Suc k) \ 0)) (nrows A' - Suc 0) 0 (from_nat (Suc k)) bezout)) = 0" and "to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (bezout_iterate (interchange_rows A' 0 (LEAST n. A' $ n $ from_nat (Suc k) \ 0)) (nrows A' - Suc 0) 0 (from_nat (Suc k)) bezout)) = 0" by (rule echelon_foldl_condition3[OF ib 1 2 rref_A'], metis ncols_def Suc_k)+ next fix m assume "\m>?greatest + 1. A' $ m $ from_nat (Suc k) = 0" and "0 < m" and "A' $ m $ from_nat (Suc k) \ 0" and "Suc (to_nat ?greatest) = nrows A'" thus "nrows A' = Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A'))" by (metis (mono_tags) Suc_eq_plus1 Suc_le' from_nat_suc from_nat_to_nat_id not_less_eq nrows_def to_nat_less_card to_nat_mono) next fix mb assume 1: "\m>?greatest+1. A' $ m $ from_nat (Suc k) = 0" and 2: "Suc (to_nat ?greatest) \ nrows A'" and 3: "?greatest+1 \ mb" and 4: "A' $ mb $ from_nat (Suc k) \ 0" show "Suc (to_nat ?greatest) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A')" by (rule echelon_foldl_condition4[OF 1 2 3 4]) next fix m assume "?greatest + 1 < m" and "A' $ m $ from_nat (Suc k) \ 0" and "\m>0. A' $ m $ from_nat (Suc k) = 0 " thus "nrows A' = Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A'))" by (metis le_less_trans least_mod_type) next fix m assume "?greatest + 1 < m" and "A' $ m $ from_nat (Suc k) \ 0" and "\m>0. A' $ m $ from_nat (Suc k) = 0" thus "\m. \ is_zero_row_upt_k m (Suc (Suc k)) (bezout_iterate (interchange_rows A' (?greatest + 1) (LEAST n. A' $ n $ from_nat (Suc k) \ 0 \ ?greatest + 1 \ n)) (nrows A' - Suc 0) (?greatest + 1) (from_nat (Suc k)) bezout)" by (metis le_less_trans least_mod_type) next fix mb assume "\ is_zero_row_upt_k mb (Suc k) A'" and "Suc (to_nat ?greatest) = nrows A'" thus " nrows A' = Suc (to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) A'))" by (rule echelon_foldl_condition5) next fix m assume "(GREATEST n. \ is_zero_row_upt_k n (Suc k) A') + 1 < m" and "A' $ m $ from_nat (Suc k) \ 0" and "\m>0. A' $ m $ from_nat (Suc k) = 0" thus "Suc (to_nat ?greatest) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (bezout_iterate (interchange_rows A' (?greatest + 1) (LEAST n. A' $ n $ from_nat (Suc k) \ 0 \ ?greatest + 1 \ n)) (nrows A' - Suc 0) (?greatest + 1) (from_nat (Suc k)) bezout))" by (metis le_less_trans least_mod_type) next fix mc assume "?greatest + 1 \ mc" and "A' $ mc $ from_nat (Suc k) \ 0" thus "\m. \ is_zero_row_upt_k m (Suc (Suc k)) (bezout_iterate (interchange_rows A' (?greatest + 1) (LEAST n. A' $ n $ from_nat (Suc k) \ 0 \ ?greatest + 1 \ n)) (nrows A' - Suc 0) (?greatest + 1) (from_nat (Suc k)) bezout)" using echelon_foldl_condition6[OF ib] by blast next fix mb mc assume 1: "\ is_zero_row_upt_k mb (Suc k) A'" and 2: "Suc (to_nat ?greatest) \ nrows A'" and 3: "?greatest + 1 \ mc" and 4:"A' $ mc $ from_nat (Suc k) \ 0" show " Suc (to_nat ?greatest) = to_nat (GREATEST n. \ is_zero_row_upt_k n (Suc (Suc k)) (bezout_iterate (interchange_rows A' (?greatest + 1) (LEAST n. A' $ n $ from_nat (Suc k) \ 0 \ ?greatest + 1 \ n)) (nrows A' - Suc 0) (?greatest + 1) (from_nat (Suc k)) bezout))" by (rule echelon_foldl_condition7[OF ib rref_A' k2 1 2 3 4 ]) qed qed qed subsubsection\Proving the existence of invertible matrices which do the transformations\ lemma bezout_iterate_invertible: fixes A::"'a::{bezout_domain}^'cols^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" assumes "nn" and "A $ i $ j \ 0" shows "\P. invertible P \ P**A = bezout_iterate A n i j bezout" using assms proof (induct n arbitrary: A) case 0 show ?case unfolding bezout_iterate.simps by (simp add: exI[of _ "mat 1"] matrix_mul_lid invertible_def) next case (Suc n) show ?case proof (cases "Suc n = to_nat i") case True show ?thesis unfolding bezout_iterate.simps using True Suc.prems(1) by (simp add: exI[of _ "mat 1"] matrix_mul_lid invertible_def) next case False have i_le_n: "to_nat i < Suc n" using Suc.prems(3) False by auto let ?B="(bezout_matrix A i (from_nat (Suc n)) j bezout ** A)" have b: "bezout_iterate A (Suc n) i j bezout = bezout_iterate ?B n i j bezout" unfolding bezout_iterate.simps using i_le_n by auto have "\P. invertible P \ P**?B = bezout_iterate ?B n i j bezout" proof (rule Suc.hyps[OF ib _]) show "n < nrows ?B" using Suc.prems (2) unfolding nrows_def by simp show "to_nat i \ n" using i_le_n by auto show "?B $ i $ j \ 0" by (metis False Suc.prems(2) Suc.prems(4) bezout_matrix_not_zero ib nrows_def to_nat_from_nat_id) qed from this obtain P where inv_P: "invertible P" and P: "P**?B = bezout_iterate ?B n i j bezout" by blast show ?thesis proof (rule exI[of _ "P ** bezout_matrix A i (from_nat (Suc n)) j bezout"], rule conjI, rule invertible_mult) show "P ** bezout_matrix A i (from_nat (Suc n)) j bezout ** A = bezout_iterate A (Suc n) i j bezout" using P unfolding b by (metis matrix_mul_assoc) have "det (bezout_matrix A i (from_nat (Suc n)) j bezout) = 1" proof (rule det_bezout_matrix[OF ib]) show "i < from_nat (Suc n)" using i_le_n from_nat_mono[of "to_nat i" "Suc n"] Suc.prems(2) unfolding nrows_def by (metis from_nat_to_nat_id) show "A $ i $ j \ 0" by (rule Suc.prems(4)) qed thus "invertible (bezout_matrix A i (mod_type_class.from_nat (Suc n)) j bezout)" unfolding invertible_iff_is_unit by simp show "invertible P" using inv_P . qed qed qed lemma echelon_form_of_column_k_invertible: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" shows "\P. invertible P \ P ** A = fst ((echelon_form_of_column_k bezout) (A,i) k)" proof - have "\P. invertible P \ P ** A = A" by (simp add: exI[of _ "mat 1"] matrix_mul_lid invertible_def) thus ?thesis proof (unfold echelon_form_of_column_k_def Let_def, auto) fix P m ma let ?least = "(LEAST n. A $ n $ from_nat k \ 0 \ from_nat i \ n)" let ?interchange ="(interchange_rows A (from_nat i) ?least)" assume i: "i \ nrows A" and i2: "mod_type_class.from_nat i \ ma" and ma: "A $ ma $ mod_type_class.from_nat k \ 0" have "\P. invertible P \ P ** ?interchange = bezout_iterate ?interchange (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" proof (rule bezout_iterate_invertible[OF ib]) show "nrows A - Suc 0 < nrows ?interchange" unfolding nrows_def by simp show "to_nat (from_nat i::'rows) \ nrows A - Suc 0" by (metis Suc_leI Suc_le_mono Suc_pred nrows_def to_nat_less_card zero_less_card_finite) show "?interchange $ from_nat i $ from_nat k \ 0" by (metis (mono_tags, lifting) LeastI_ex i2 ma interchange_rows_i) qed from this obtain P where inv_P: "invertible P" and P: "P ** ?interchange = bezout_iterate ?interchange (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" by blast show "\P. invertible P \ P ** A = bezout_iterate ?interchange (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" proof (rule exI[of _ "P ** interchange_rows (mat 1) (from_nat i) ?least"], rule conjI, rule invertible_mult) show "P ** interchange_rows (mat 1) (from_nat i) ?least ** A = bezout_iterate ?interchange (nrows A - Suc 0) (from_nat i) (from_nat k) bezout" using P by (metis (no_types, lifting) interchange_rows_mat_1 matrix_mul_assoc) show "invertible P" by (rule inv_P) show "invertible (interchange_rows (mat 1) (from_nat i) ?least)" by (simp add: invertible_interchange_rows) qed qed qed lemma echelon_form_of_upt_k_invertible: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" shows "\P. invertible P \ P**A = (echelon_form_of_upt_k A k bezout)" proof (induct k) case 0 show ?case unfolding echelon_form_of_upt_k_def by (simp add: echelon_form_of_column_k_invertible[OF ib]) next case (Suc k) have set_rw: "[0..Final results\ lemma echelon_form_echelon_form_of: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext bezout" shows "echelon_form (echelon_form_of A bezout)" proof - have n: "ncols A - 1 < ncols A" unfolding ncols_def by auto show ?thesis unfolding echelon_form_def echelon_form_of_def using echelon_echelon_form_of_upt_k[OF n ib] unfolding ncols_def by simp qed lemma echelon_form_of_invertible: fixes A::"'a::{bezout_domain}^'cols::{mod_type}^'rows::{mod_type}" assumes ib: "is_bezout_ext (bezout)" shows "\P. invertible P \ P ** A = (echelon_form_of A bezout) \ echelon_form (echelon_form_of A bezout)" using echelon_form_of_upt_k_invertible[OF ib] echelon_form_echelon_form_of[OF ib] unfolding echelon_form_of_def by fast text\Executable version\ corollary echelon_form_echelon_form_of_euclidean: fixes A::"'a::{euclidean_ring_gcd}^'cols::{mod_type}^'rows::{mod_type}" shows "echelon_form (echelon_form_of_euclidean A)" using echelon_form_echelon_form_of is_bezout_ext_euclid_ext2 unfolding echelon_form_of_euclidean_def by auto corollary echelon_form_of_euclidean_invertible: fixes A::"'a::{euclidean_ring_gcd}^'cols::{mod_type}^'rows::{mod_type}" shows "\P. invertible P \ P**A = (echelon_form_of A euclid_ext2) \ echelon_form (echelon_form_of A euclid_ext2)" using echelon_form_of_invertible[OF is_bezout_ext_euclid_ext2] . subsection\More efficient code equations\ definition "echelon_form_of_column_k_efficient bezout A' k = (let (A, i) = A'; from_nat_k = from_nat k; from_nat_i = from_nat i; all_zero_below_i = (\m>from_nat_i. A $ m $ from_nat_k = 0) in if (i = nrows A) \ (A $ from_nat_i $ from_nat_k = 0) \ all_zero_below_i then (A, i) else if all_zero_below_i then (A, i + 1) else let n = (LEAST n. A $ n $ from_nat_k \ 0 \ from_nat_i \ n); interchange_A = interchange_rows A (from_nat_i) n in (bezout_iterate (interchange_A) (nrows A - 1) (from_nat_i) (from_nat_k) bezout, i + 1))" lemma echelon_form_of_column_k_efficient[code]: "(echelon_form_of_column_k bezout) (A,i) k = (echelon_form_of_column_k_efficient bezout) (A,i) k" unfolding echelon_form_of_column_k_def echelon_form_of_column_k_efficient_def unfolding Let_def by force end (*********** Possible future work: ***********) (* - Other kind of Echelon Forms (minimal Echelon, Howell...) and more applications: ranks, system of equations... - Hermite Normal Form and its uniqueness. HNF is unique over PID (not in PIR) with respect to a given complete set of associates and a given complete set of residues. In general terms, with the integers we use the positive integers, in F[x] we use monic polynomials... - Smith Normal Form and its uniqueness. *) diff --git a/thys/Gabow_SCC/Gabow_Skeleton.thy b/thys/Gabow_SCC/Gabow_Skeleton.thy --- a/thys/Gabow_SCC/Gabow_Skeleton.thy +++ b/thys/Gabow_SCC/Gabow_Skeleton.thy @@ -1,2655 +1,2652 @@ section \Skeleton for Gabow's SCC Algorithm \label{sec:skel}\ theory Gabow_Skeleton imports CAVA_Automata.Digraph begin (* TODO: convenience locale, consider merging this with invariants *) locale fr_graph = graph G for G :: "('v, 'more) graph_rec_scheme" + assumes finite_reachableE_V0[simp, intro!]: "finite (E\<^sup>* `` V0)" text \ In this theory, we formalize a skeleton of Gabow's SCC algorithm. The skeleton serves as a starting point to develop concrete algorithms, like enumerating the SCCs or checking emptiness of a generalized Büchi automaton. \ section \Statistics Setup\ text \ We define some dummy-constants that are included into the generated code, and may be mapped to side-effecting ML-code that records statistics and debug information about the execution. In the skeleton algorithm, we count the number of visited nodes, and include a timing for the whole algorithm. \ definition stat_newnode :: "unit => unit" \ \Invoked if new node is visited\ where [code]: "stat_newnode \ \_. ()" definition stat_start :: "unit => unit" \ \Invoked once if algorithm starts\ where [code]: "stat_start \ \_. ()" definition stat_stop :: "unit => unit" \ \Invoked once if algorithm stops\ where [code]: "stat_stop \ \_. ()" lemma [autoref_rules]: "(stat_newnode,stat_newnode) \ unit_rel \ unit_rel" "(stat_start,stat_start) \ unit_rel \ unit_rel" "(stat_stop,stat_stop) \ unit_rel \ unit_rel" by auto abbreviation "stat_newnode_nres \ RETURN (stat_newnode ())" abbreviation "stat_start_nres \ RETURN (stat_start ())" abbreviation "stat_stop_nres \ RETURN (stat_stop ())" lemma discard_stat_refine[refine]: "m1\m2 \ stat_newnode_nres \ m1 \ m2" "m1\m2 \ stat_start_nres \ m1 \ m2" "m1\m2 \ stat_stop_nres \ m1 \ m2" by simp_all section \Abstract Algorithm\ text \ In this section, we formalize an abstract version of a path-based SCC algorithm. Later, this algorithm will be refined to use Gabow's data structure. \ subsection \Preliminaries\ definition path_seg :: "'a set list \ nat \ nat \ 'a set" \ \Set of nodes in a segment of the path\ where "path_seg p i j \ \{p!k|k. i\k \ ki \ path_seg p i j = {}" "path_seg p i (Suc i) = p!i" unfolding path_seg_def apply auto [] apply (auto simp: le_less_Suc_eq) [] done lemma path_seg_drop: "\(set (drop i p)) = path_seg p i (length p)" unfolding path_seg_def by (fastforce simp: in_set_drop_conv_nth Bex_def) lemma path_seg_butlast: "p\[] \ path_seg p 0 (length p - Suc 0) = \(set (butlast p))" apply (cases p rule: rev_cases, simp) apply (fastforce simp: path_seg_def nth_append in_set_conv_nth) done definition idx_of :: "'a set list \ 'a \ nat" \ \Index of path segment that contains a node\ where "idx_of p v \ THE i. i v\p!i" lemma idx_of_props: assumes p_disjoint_sym: "\i j v. i j v\p!i \ v\p!j \ i=j" assumes ON_STACK: "v\\(set p)" shows "idx_of p v < length p" and "v \ p ! idx_of p v" proof - from ON_STACK obtain i where "i p ! i" by (auto simp add: in_set_conv_nth) moreover hence "\jp ! j \ i=j" using p_disjoint_sym by auto ultimately show "idx_of p v < length p" and "v \ p ! idx_of p v" unfolding idx_of_def by (metis (lifting) theI')+ qed lemma idx_of_uniq: assumes p_disjoint_sym: "\i j v. i j v\p!i \ v\p!j \ i=j" assumes A: "ip!i" shows "idx_of p v = i" proof - from A p_disjoint_sym have "\jp ! j \ i=j" by auto with A show ?thesis unfolding idx_of_def by (metis (lifting) the_equality) qed subsection \Invariants\ text \The state of the inner loop consists of the path \p\ of collapsed nodes, the set \D\ of finished (done) nodes, and the set \pE\ of pending edges.\ type_synonym 'v abs_state = "'v set list \ 'v set \ ('v\'v) set" context fr_graph begin definition touched :: "'v set list \ 'v set \ 'v set" \ \Touched: Nodes that are done or on path\ where "touched p D \ D \ \(set p)" definition vE :: "'v set list \ 'v set \ ('v \ 'v) set \ ('v \ 'v) set" \ \Visited edges: No longer pending edges from touched nodes\ where "vE p D pE \ (E \ (touched p D \ UNIV)) - pE" lemma vE_ss_E: "vE p D pE \ E" \ \Visited edges are edges\ unfolding vE_def by auto end locale outer_invar_loc \ \Invariant of the outer loop\ = fr_graph G for G :: "('v,'more) graph_rec_scheme" + fixes it :: "'v set" \ \Remaining nodes to iterate over\ fixes D :: "'v set" \ \Finished nodes\ assumes it_initial: "it\V0" \ \Only start nodes to iterate over\ assumes it_done: "V0 - it \ D" \ \Nodes already iterated over are visited\ assumes D_reachable: "D\E\<^sup>*``V0" \ \Done nodes are reachable\ assumes D_closed: "E``D \ D" \ \Done is closed under transitions\ begin lemma locale_this: "outer_invar_loc G it D" by unfold_locales definition (in fr_graph) "outer_invar \ \it D. outer_invar_loc G it D" lemma outer_invar_this[simp, intro!]: "outer_invar it D" unfolding outer_invar_def apply simp by unfold_locales end locale invar_loc \ \Invariant of the inner loop\ = fr_graph G for G :: "('v, 'more) graph_rec_scheme" + fixes v0 :: "'v" fixes D0 :: "'v set" fixes p :: "'v set list" fixes D :: "'v set" fixes pE :: "('v\'v) set" assumes v0_initial[simp, intro!]: "v0\V0" assumes D_incr: "D0 \ D" assumes pE_E_from_p: "pE \ E \ (\(set p)) \ UNIV" \ \Pending edges are edges from path\ assumes E_from_p_touched: "E \ (\(set p) \ UNIV) \ pE \ UNIV \ touched p D" \ \Edges from path are pending or touched\ assumes D_reachable: "D\E\<^sup>*``V0" \ \Done nodes are reachable\ assumes p_connected: "Suc i p!i \ p!Suc i \ (E-pE) \ {}" \ \CNodes on path are connected by non-pending edges\ assumes p_disjoint: "\i \ p!i \ p!j = {}" \ \CNodes on path are disjoint\ assumes p_sc: "U\set p \ U\U \ (vE p D pE \ U\U)\<^sup>*" \ \Nodes in CNodes are mutually reachable by visited edges\ assumes root_v0: "p\[] \ v0\hd p" \ \Root CNode contains start node\ assumes p_empty_v0: "p=[] \ v0\D" \ \Start node is done if path empty\ assumes D_closed: "E``D \ D" \ \Done is closed under transitions\ (*assumes D_vis: "E\D\D \ vE" -- "All edges from done nodes are visited"*) assumes vE_no_back: "\i \ vE p D pE \ p!j \ p!i = {}" \ \Visited edges do not go back on path\ assumes p_not_D: "\(set p) \ D = {}" \ \Path does not contain done nodes\ begin abbreviation ltouched where "ltouched \ touched p D" abbreviation lvE where "lvE \ vE p D pE" lemma locale_this: "invar_loc G v0 D0 p D pE" by unfold_locales definition (in fr_graph) "invar \ \v0 D0 (p,D,pE). invar_loc G v0 D0 p D pE" lemma invar_this[simp, intro!]: "invar v0 D0 (p,D,pE)" unfolding invar_def apply simp by unfold_locales lemma finite_reachableE_v0[simp, intro!]: "finite (E\<^sup>*``{v0})" apply (rule finite_subset[OF _ finite_reachableE_V0]) using v0_initial by auto lemma D_vis: "E\D\UNIV \ lvE" \ \All edges from done nodes are visited\ unfolding vE_def touched_def using pE_E_from_p p_not_D by blast lemma vE_touched: "lvE \ ltouched \ ltouched" \ \Visited edges only between touched nodes\ using E_from_p_touched D_closed unfolding vE_def touched_def by blast lemma lvE_ss_E: "lvE \ E" \ \Visited edges are edges\ unfolding vE_def by auto lemma path_touched: "\(set p) \ ltouched" by (auto simp: touched_def) lemma D_touched: "D \ ltouched" by (auto simp: touched_def) lemma pE_by_vE: "pE = (E \ \(set p) \ UNIV) - lvE" \ \Pending edges are edges from path not yet visited\ unfolding vE_def touched_def using pE_E_from_p by auto lemma pick_pending: "p\[] \ pE \ last p \ UNIV = (E-lvE) \ last p \ UNIV" \ \Pending edges from end of path are non-visited edges from end of path\ apply (subst pE_by_vE) by auto lemma p_connected': assumes A: "Suc i p!Suc i \ lvE \ {}" proof - from A p_not_D have "p!i \ set p" "p!Suc i \ set p" by auto with p_connected[OF A] show ?thesis unfolding vE_def touched_def by blast qed end subsubsection \Termination\ context fr_graph begin text \The termination argument is based on unprocessed edges: Reachable edges from untouched nodes and pending edges.\ definition "unproc_edges v0 p D pE \ (E \ (E\<^sup>*``{v0} - (D \ \(set p))) \ UNIV) \ pE" text \ In each iteration of the loop, either the number of unprocessed edges decreases, or the path length decreases.\ definition "abs_wf_rel v0 \ inv_image (finite_psubset <*lex*> measure length) (\(p,D,pE). (unproc_edges v0 p D pE, p))" lemma abs_wf_rel_wf[simp, intro!]: "wf (abs_wf_rel v0)" unfolding abs_wf_rel_def by auto end subsection \Abstract Skeleton Algorithm\ context fr_graph begin definition (in fr_graph) initial :: "'v \ 'v set \ 'v abs_state" where "initial v0 D \ ([{v0}], D, (E \ {v0}\UNIV))" definition (in -) collapse_aux :: "'a set list \ nat \ 'a set list" where "collapse_aux p i \ take i p @ [\(set (drop i p))]" definition (in -) collapse :: "'a \ 'a abs_state \ 'a abs_state" where "collapse v PDPE \ let (p,D,pE)=PDPE; i=idx_of p v; p = collapse_aux p i in (p,D,pE)" definition (in -) select_edge :: "'a abs_state \ ('a option \ 'a abs_state) nres" where "select_edge PDPE \ do { let (p,D,pE) = PDPE; e \ SELECT (\e. e \ pE \ last p \ UNIV); case e of None \ RETURN (None,(p,D,pE)) | Some (u,v) \ RETURN (Some v, (p,D,pE - {(u,v)})) }" definition (in fr_graph) push :: "'v \ 'v abs_state \ 'v abs_state" where "push v PDPE \ let (p,D,pE) = PDPE; p = p@[{v}]; pE = pE \ (E\{v}\UNIV) in (p,D,pE)" definition (in -) pop :: "'v abs_state \ 'v abs_state" where "pop PDPE \ let (p,D,pE) = PDPE; (p,V) = (butlast p, last p); D = V \ D in (p,D,pE)" text \The following lemmas match the definitions presented in the paper:\ lemma "select_edge (p,D,pE) \ do { e \ SELECT (\e. e \ pE \ last p \ UNIV); case e of None \ RETURN (None,(p,D,pE)) | Some (u,v) \ RETURN (Some v, (p,D,pE - {(u,v)})) }" unfolding select_edge_def by simp lemma "collapse v (p,D,pE) \ let i=idx_of p v in (take i p @ [\(set (drop i p))],D,pE)" unfolding collapse_def collapse_aux_def by simp lemma "push v (p, D, pE) \ (p @ [{v}], D, pE \ E \ {v} \ UNIV)" unfolding push_def by simp lemma "pop (p, D, pE) \ (butlast p, last p \ D, pE)" unfolding pop_def by auto thm pop_def[unfolded Let_def, no_vars] thm select_edge_def[unfolded Let_def] definition skeleton :: "'v set nres" \ \Abstract Skeleton Algorithm\ where "skeleton \ do { let D = {}; r \ FOREACHi outer_invar V0 (\v0 D0. do { if v0\D0 then do { let s = initial v0 D0; (p,D,pE) \ WHILEIT (invar v0 D0) (\(p,D,pE). p \ []) (\(p,D,pE). do { \ \Select edge from end of path\ (vo,(p,D,pE)) \ select_edge (p,D,pE); ASSERT (p\[]); case vo of Some v \ do { \ \Found outgoing edge to node \v\\ if v \ \(set p) then do { \ \Back edge: Collapse path\ RETURN (collapse v (p,D,pE)) } else if v\D then do { \ \Edge to new node. Append to path\ RETURN (push v (p,D,pE)) } else do { \ \Edge to done node. Skip\ RETURN (p,D,pE) } } | None \ do { ASSERT (pE \ last p \ UNIV = {}); \ \No more outgoing edges from current node on path\ RETURN (pop (p,D,pE)) } }) s; ASSERT (p=[] \ pE={}); RETURN D } else RETURN D0 }) D; RETURN r }" end subsection \Invariant Preservation\ context fr_graph begin lemma set_collapse_aux[simp]: "\(set (collapse_aux p i)) = \(set p)" apply (subst (2) append_take_drop_id[of _ p,symmetric]) apply (simp del: append_take_drop_id) unfolding collapse_aux_def by auto lemma touched_collapse[simp]: "touched (collapse_aux p i) D = touched p D" unfolding touched_def by simp lemma vE_collapse_aux[simp]: "vE (collapse_aux p i) D pE = vE p D pE" unfolding vE_def by simp lemma touched_push[simp]: "touched (p @ [V]) D = touched p D \ V" unfolding touched_def by auto end subsubsection \Corollaries of the invariant\ text \In this section, we prove some more corollaries of the invariant, which are helpful to show invariant preservation\ context invar_loc begin lemma cnode_connectedI: "\ip!i; v\p!i\ \ (u,v)\(lvE \ p!i\p!i)\<^sup>*" using p_sc[of "p!i"] by (auto simp: in_set_conv_nth) lemma cnode_connectedI': "\ip!i; v\p!i\ \ (u,v)\(lvE)\<^sup>*" by (metis inf.cobounded1 rtrancl_mono_mp cnode_connectedI) lemma p_no_empty: "{} \ set p" proof assume "{}\set p" then obtain i where IDX: "i p!i\{}" using p_no_empty by (metis nth_mem) lemma p_disjoint_sym: "\ip!i; v\p!j\ \ i=j" by (metis disjoint_iff_not_equal linorder_neqE_nat p_disjoint) lemma pi_ss_path_seg_eq[simp]: assumes A: "ilength p" shows "p!i\path_seg p l u \ l\i \ ipath_seg p l u" from A obtain x where "x\p!i" by (blast dest: p_no_empty_idx) with B obtain i' where C: "x\p!i'" "l\i'" "i'i _ \x\p!i\ \x\p!i'\] \i' \u\length p\ have "i=i'" by simp with C show "l\i \ ilength p" "l2length p" shows "path_seg p l1 u1 \ path_seg p l2 u2 \ l2\l1 \ u1\u2" proof assume S: "path_seg p l1 u1 \ path_seg p l2 u2" have "p!l1 \ path_seg p l1 u1" using A by simp also note S finally have 1: "l2\l1" using A by simp have "p!(u1 - 1) \ path_seg p l1 u1" using A by simp also note S finally have 2: "u1\u2" using A by auto from 1 2 show "l2\l1 \ u1\u2" .. next assume "l2\l1 \ u1\u2" thus "path_seg p l1 u1 \ path_seg p l2 u2" using A apply (clarsimp simp: path_seg_def) [] apply (metis dual_order.strict_trans1 dual_order.trans) done qed lemma pathI: assumes "x\p!i" "y\p!j" assumes "i\j" "j path_seg p i (Suc j)" shows "(x,y)\(lvE \ seg\seg)\<^sup>*" \ \We can obtain a path between cnodes on path\ using assms(3,1,2,4) unfolding seg_def proof (induction arbitrary: y rule: dec_induct) case base thus ?case by (auto intro!: cnode_connectedI) next case (step j) let ?seg = "path_seg p i (Suc j)" let ?seg' = "path_seg p i (Suc (Suc j))" have SSS: "?seg \ ?seg'" apply (subst path_seg_ss_eq) using step.hyps step.prems by auto from p_connected'[OF \Suc j < length p\] obtain u v where UV: "(u,v)\lvE" "u\p!j" "v\p!Suc j" by auto have ISS: "p!j \ ?seg'" "p!Suc j \ ?seg'" using step.hyps step.prems by simp_all from p_no_empty_idx[of j] \Suc j < length p\ obtain x' where "x'\p!j" by auto with step.IH[of x'] \x\p!i\ \Suc j < length p\ have t: "(x,x')\(lvE\?seg\?seg)\<^sup>*" by auto have "(x,x')\(lvE\?seg'\?seg')\<^sup>*" using SSS by (auto intro: rtrancl_mono_mp[OF _ t]) also from cnode_connectedI[OF _ \x'\p!j\ \u\p!j\] \Suc j < length p\ have t: "(x', u) \ (lvE \ p ! j \ p ! j)\<^sup>*" by auto have "(x', u) \ (lvE\?seg'\?seg')\<^sup>*" using ISS by (auto intro: rtrancl_mono_mp[OF _ t]) also have "(u,v)\lvE\?seg'\?seg'" using UV ISS by auto also from cnode_connectedI[OF \Suc j < length p\ \v\p!Suc j\ \y\p!Suc j\] have t: "(v, y) \ (lvE \ p ! Suc j \ p ! Suc j)\<^sup>*" by auto have "(v, y) \ (lvE\?seg'\?seg')\<^sup>*" using ISS by (auto intro: rtrancl_mono_mp[OF _ t]) finally show "(x,y)\(lvE\?seg'\?seg')\<^sup>*" . qed lemma p_reachable: "\(set p) \ E\<^sup>*``{v0}" \ \Nodes on path are reachable\ proof fix v assume A: "v\\(set p)" then obtain i where "ip!i" by (metis UnionE in_set_conv_nth) moreover from A root_v0 have "v0\p!0" by (cases p) auto ultimately have t: "(v0,v)\(lvE \ path_seg p 0 (Suc i) \ path_seg p 0 (Suc i))\<^sup>*" by (auto intro: pathI) from lvE_ss_E have "(v0,v)\E\<^sup>*" by (auto intro: rtrancl_mono_mp[OF _ t]) thus "v\E\<^sup>*``{v0}" by auto qed lemma touched_reachable: "ltouched \ E\<^sup>*``V0" \ \Touched nodes are reachable\ unfolding touched_def using p_reachable D_reachable by blast lemma vE_reachable: "lvE \ E\<^sup>*``V0 \ E\<^sup>*``V0" apply (rule order_trans[OF vE_touched]) using touched_reachable by blast lemma pE_reachable: "pE \ E\<^sup>*``{v0} \ E\<^sup>*``{v0}" proof safe fix u v assume E: "(u,v)\pE" with pE_E_from_p p_reachable have "(v0,u)\E\<^sup>*" "(u,v)\E" by blast+ thus "(v0,u)\E\<^sup>*" "(v0,v)\E\<^sup>*" by auto qed lemma D_closed_vE_rtrancl: "lvE\<^sup>*``D \ D" by (metis D_closed Image_closed_trancl eq_iff reachable_mono lvE_ss_E) lemma D_closed_path: "\path E u q w; u\D\ \ set q \ D" proof - assume a1: "path E u q w" assume "u \ D" hence f1: "{u} \ D" using bot.extremum by force have "set q \ E\<^sup>* `` {u}" using a1 by (metis insert_subset path_nodes_reachable) thus "set q \ D" using f1 by (metis D_closed rtrancl_reachable_induct subset_trans) qed lemma D_closed_path_vE: "\path lvE u q w; u\D\ \ set q \ D" by (metis D_closed_path path_mono lvE_ss_E) lemma path_in_lastnode: assumes P: "path lvE u q v" assumes [simp]: "p\[]" assumes ND: "u\last p" "v\last p" shows "set q \ last p" \ \A path from the last Cnode to the last Cnode remains in the last Cnode\ (* TODO: This can be generalized in two directions: either 1) The path end anywhere. Due to vE_touched we can infer that it ends in last cnode or 2) We may use any cnode, not only the last one *) using P ND proof (induction) case (path_prepend u v l w) from \(u,v)\lvE\ vE_touched have "v\ltouched" by auto hence "v\\(set p)" unfolding touched_def proof assume "v\D" moreover from \path lvE v l w\ have "(v,w)\lvE\<^sup>*" by (rule path_is_rtrancl) ultimately have "w\D" using D_closed_vE_rtrancl by auto with \w\last p\ p_not_D have False by (metis IntI Misc.last_in_set Sup_inf_eq_bot_iff assms(2) bex_empty path_prepend.hyps(2)) thus ?thesis .. qed then obtain i where "ip!i" by (metis UnionE in_set_conv_nth) have "i=length p - 1" proof (rule ccontr) assume "i\length p - 1" with \i have "i < length p - 1" by simp with vE_no_back[of i "length p - 1"] \i have "lvE \ last p \ p!i = {}" by (simp add: last_conv_nth) with \(u,v)\lvE\ \u\last p\ \v\p!i\ show False by auto qed with \v\p!i\ have "v\last p" by (simp add: last_conv_nth) with path_prepend.IH \w\last p\ \u\last p\ show ?case by auto qed simp lemma loop_in_lastnode: assumes P: "path lvE u q u" assumes [simp]: "p\[]" assumes ND: "set q \ last p \ {}" shows "u\last p" and "set q \ last p" \ \A loop that touches the last node is completely inside the last node\ proof - from ND obtain v where "v\set q" "v\last p" by auto then obtain q1 q2 where [simp]: "q=q1@v#q2" by (auto simp: in_set_conv_decomp) from P have "path lvE v (v#q2@q1) v" by (auto simp: path_conc_conv path_cons_conv) from path_in_lastnode[OF this \p\[]\ \v\last p\ \v\last p\] show "set q \ last p" by simp from P show "u\last p" apply (cases q, simp) apply simp using \set q \ last p\ apply (auto simp: path_cons_conv) done qed lemma no_D_p_edges: "E \ D \ \(set p) = {}" using D_closed p_not_D by auto lemma idx_of_props: assumes ON_STACK: "v\\(set p)" shows "idx_of p v < length p" and "v \ p ! idx_of p v" using idx_of_props[OF _ assms] p_disjoint_sym by blast+ end subsubsection \Auxiliary Lemmas Regarding the Operations\ lemma (in fr_graph) vE_initial[simp]: "vE [{v0}] {} (E \ {v0} \ UNIV) = {}" unfolding vE_def touched_def by auto context invar_loc begin lemma vE_push: "\ (u,v)\pE; u\last p; v\\(set p); v\D \ \ vE (p @ [{v}]) D ((pE - {(u,v)}) \ E\{v}\UNIV) = insert (u,v) lvE" unfolding vE_def touched_def using pE_E_from_p by auto lemma vE_remove[simp]: "\p\[]; (u,v)\pE\ \ vE p D (pE - {(u,v)}) = insert (u,v) lvE" unfolding vE_def touched_def using pE_E_from_p by blast lemma vE_pop[simp]: "p\[] \ vE (butlast p) (last p \ D) pE = lvE" unfolding vE_def touched_def by (cases p rule: rev_cases) auto lemma pE_fin: "p=[] \ pE={}" using pE_by_vE by auto lemma (in invar_loc) lastp_un_D_closed: assumes NE: "p \ []" assumes NO': "pE \ (last p \ UNIV) = {}" shows "E``(last p \ D) \ (last p \ D)" \ \On pop, the popped CNode and D are closed under transitions\ proof (intro subsetI, elim ImageE) from NO' have NO: "(E - lvE) \ (last p \ UNIV) = {}" by (simp add: pick_pending[OF NE]) let ?i = "length p - 1" from NE have [simp]: "last p = p!?i" by (metis last_conv_nth) fix u v assume E: "(u,v)\E" assume UI: "u\last p \ D" hence "u\p!?i \ D" by simp { assume "u\last p" "v\last p" moreover from E NO \u\last p\ have "(u,v)\lvE" by auto ultimately have "v\D \ v\\(set p)" using vE_touched unfolding touched_def by auto moreover { assume "v\\(set p)" then obtain j where V: "jp!j" by (metis UnionE in_set_conv_nth) with \v\last p\ have "jj _] \(u,v)\lvE\ V \u\last p\ have False by auto } ultimately have "v\D" by blast } with E UI D_closed show "v\last p \ D" by auto qed end subsubsection \Preservation of Invariant by Operations\ context fr_graph begin lemma (in outer_invar_loc) invar_initial_aux: assumes "v0\it - D" shows "invar v0 D (initial v0 D)" unfolding invar_def initial_def apply simp apply unfold_locales apply simp_all using assms it_initial apply auto [] using D_reachable it_initial assms apply auto [] using D_closed apply auto [] using assms apply auto [] done lemma invar_initial: "\outer_invar it D0; v0\it; v0\D0\ \ invar v0 D0 (initial v0 D0)" unfolding outer_invar_def apply (drule outer_invar_loc.invar_initial_aux) by auto lemma outer_invar_initial[simp, intro!]: "outer_invar V0 {}" unfolding outer_invar_def apply unfold_locales by auto lemma invar_pop: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes NO': "pE \ (last p \ UNIV) = {}" shows "invar v0 D0 (pop (p,D,pE))" unfolding invar_def pop_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp have [simp]: "set p = insert (last p) (set (butlast p))" using NE by (cases p rule: rev_cases) auto from p_disjoint have lp_dj_blp: "last p \ \(set (butlast p)) = {}" apply (cases p rule: rev_cases) apply simp apply (fastforce simp: in_set_conv_nth nth_append) done { fix i assume A: "Suc i < length (butlast p)" hence A': "Suc i < length p" by auto from nth_butlast[of i p] A have [simp]: "butlast p ! i = p ! i" by auto from nth_butlast[of "Suc i" p] A have [simp]: "butlast p ! Suc i = p ! Suc i" by auto from p_connected[OF A'] have "butlast p ! i \ butlast p ! Suc i \ (E - pE) \ {}" by simp } note AUX_p_connected = this (*have [simp]: "(E \ (last p \ D \ \set (butlast p)) \ UNIV - pE) = vE" unfolding vE_def touched_def by auto*) show "invar_loc G v0 D0 (butlast p) (last p \ D) pE" apply unfold_locales unfolding vE_pop[OF NE] apply simp using D_incr apply auto [] using pE_E_from_p NO' apply auto [] using E_from_p_touched apply (auto simp: touched_def) [] using D_reachable p_reachable NE apply auto [] apply (rule AUX_p_connected, assumption+) [] using p_disjoint apply (simp add: nth_butlast) using p_sc apply simp using root_v0 apply (cases p rule: rev_cases) apply auto [2] using root_v0 p_empty_v0 apply (cases p rule: rev_cases) apply auto [2] apply (rule lastp_un_D_closed, insert NO', auto) [] using vE_no_back apply (auto simp: nth_butlast) [] using p_not_D lp_dj_blp apply auto [] done qed thm invar_pop[of v_0 D_0, no_vars] lemma invar_collapse: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and "u\last p" assumes BACK: "v\\(set p)" defines "i \ idx_of p v" defines "p' \ collapse_aux p i" shows "invar v0 D0 (collapse v (p,D,pE - {(u,v)}))" unfolding invar_def collapse_def apply simp unfolding i_def[symmetric] p'_def[symmetric] proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp let ?thesis="invar_loc G v0 D0 p' D (pE - {(u,v)})" have SETP'[simp]: "\(set p') = \(set p)" unfolding p'_def by simp have IL: "i < length p" and VMEM: "v\p!i" using idx_of_props[OF BACK] unfolding i_def by auto have [simp]: "length p' = Suc i" unfolding p'_def collapse_aux_def using IL by auto have P'_IDX_SS: "\j p'!j" unfolding p'_def collapse_aux_def using IL by (auto simp add: nth_append path_seg_drop) from \u\last p\ have "u\p!(length p - 1)" by (auto simp: last_conv_nth) have defs_fold: "vE p' D (pE - {(u,v)}) = insert (u,v) lvE" "touched p' D = ltouched" by (simp_all add: p'_def E) { fix j assume A: "Suc j < length p'" hence "Suc j < length p" using IL by simp from p_connected[OF this] have "p!j \ p!Suc j \ (E-pE) \ {}" . moreover from P'_IDX_SS A have "p!j\p'!j" and "p!Suc j \ p'!Suc j" by auto ultimately have "p' ! j \ p' ! Suc j \ (E - (pE - {(u, v)})) \ {}" by blast } note AUX_p_connected = this have P_IDX_EQ[simp]: "\j. j < i \ p'!j = p!j" unfolding p'_def collapse_aux_def using IL by (auto simp: nth_append) have P'_LAST[simp]: "p'!i = path_seg p i (length p)" (is "_ = ?last_cnode") unfolding p'_def collapse_aux_def using IL by (auto simp: nth_append path_seg_drop) { fix j k assume A: "j < k" "k < length p'" have "p' ! j \ p' ! k = {}" proof (safe, simp) fix v assume "v\p'!j" and "v\p'!k" with A have "v\p!j" by simp show False proof (cases) assume "k=i" with \v\p'!k\ obtain k' where "v\p!k'" "i\k'" "k' p ! k' = {}" using A by (auto intro!: p_disjoint) with \v\p!j\ \v\p!k'\ show False by auto next assume "k\i" with A have "kj this] also have "p!j = p'!j" using \j \k by simp also have "p!k = p'!k" using \k by simp finally show False using \v\p'!j\ \v\p'!k\ by auto qed qed } note AUX_p_disjoint = this { fix U assume A: "U\set p'" then obtain j where "j U \ (insert (u, v) lvE \ U \ U)\<^sup>*" proof cases assume [simp]: "j=i" show ?thesis proof (clarsimp) fix x y assume "x\path_seg p i (length p)" "y\path_seg p i (length p)" then obtain ix iy where IX: "x\p!ix" "i\ix" "ixp!iy" "i\iy" "iy ?last_cnode" by (subst path_seg_ss_eq) auto from IY have SS2: "path_seg p i (Suc iy) \ ?last_cnode" by (subst path_seg_ss_eq) auto let ?rE = "\R. (lvE \ R\R)" let ?E = "(insert (u,v) lvE \ ?last_cnode \ ?last_cnode)" from pathI[OF \x\p!ix\ \u\p!(length p - 1)\] have "(x,u)\(?rE (path_seg p ix (Suc (length p - 1))))\<^sup>*" using IX by auto hence "(x,u)\?E\<^sup>*" apply (rule rtrancl_mono_mp[rotated]) using SS1 by auto also have "(u,v)\?E" using \i apply (clarsimp) apply (intro conjI) apply (rule rev_subsetD[OF \u\p!(length p - 1)\]) apply (simp) apply (rule rev_subsetD[OF VMEM]) apply (simp) done also from pathI[OF \v\p!i\ \y\p!iy\] have "(v,y)\(?rE (path_seg p i (Suc iy)))\<^sup>*" using IY by auto hence "(v,y)\?E\<^sup>*" apply (rule rtrancl_mono_mp[rotated]) using SS2 by auto finally show "(x,y)\?E\<^sup>*" . qed next assume "j\i" with \j have [simp]: "ji have "p!j\set p" by (metis Suc_lessD in_set_conv_nth less_trans_Suc) thus ?thesis using p_sc[of U] \p!j\set p\ apply (clarsimp) apply (subgoal_tac "(a,b)\(lvE \ p ! j \ p ! j)\<^sup>*") apply (erule rtrancl_mono_mp[rotated]) apply auto done qed } note AUX_p_sc = this { fix j k assume A: "j p' ! k \ p' ! j = {}" proof - have "{(u,v)} \ p' ! k \ p' ! j = {}" apply auto by (metis IL P_IDX_EQ Suc_lessD VMEM \j < i\ less_irrefl_nat less_trans_Suc p_disjoint_sym) moreover have "lvE \ p' ! k \ p' ! j = {}" proof (cases "ki by auto next case False with A have [simp]: "k=i" by simp show ?thesis proof (rule disjointI, clarsimp simp: \j) fix x y assume B: "(x,y)\lvE" "x\path_seg p i (length p)" "y\p!j" then obtain ix where "x\p!ix" "i\ix" "ix[]" assumes E: "(u,v)\pE" and UIL: "u\last p" assumes VNE: "v\\(set p)" "v\D" shows "invar v0 D0 (push v (p,D,pE - {(u,v)}))" unfolding invar_def push_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp let ?thesis = "invar_loc G v0 D0 (p @ [{v}]) D (pE - {(u, v)} \ E \ {v} \ UNIV)" note defs_fold = vE_push[OF E UIL VNE] touched_push { fix i assume SILL: "Suc i < length (p @ [{v}])" have "(p @ [{v}]) ! i \ (p @ [{v}]) ! Suc i \ (E - (pE - {(u, v)} \ E \ {v} \ UNIV)) \ {}" proof (cases "i = length p - 1") case True thus ?thesis using SILL E pE_E_from_p UIL VNE by (simp add: nth_append last_conv_nth) fast next case False with SILL have SILL': "Suc i < length p" by simp with SILL' VNE have X1: "v\p!i" "v\p!Suc i" by auto from p_connected[OF SILL'] obtain a b where "a\p!i" "b\p!Suc i" "(a,b)\E" "(a,b)\pE" by auto with X1 have "a\v" "b\v" by auto with \(a,b)\E\ \(a,b)\pE\ have "(a,b)\(E - (pE - {(u, v)} \ E \ {v} \ UNIV))" by auto with \a\p!i\ \b\p!Suc i\ show ?thesis using SILL' by (simp add: nth_append; blast) qed } note AUX_p_connected = this { fix U assume A: "U \ set (p @ [{v}])" have "U \ U \ (insert (u, v) lvE \ U \ U)\<^sup>*" proof cases assume "U\set p" with p_sc have "U\U \ (lvE \ U\U)\<^sup>*" . thus ?thesis by (metis (lifting, no_types) Int_insert_left_if0 Int_insert_left_if1 in_mono insert_subset rtrancl_mono_mp subsetI) next assume "U\set p" with A have "U={v}" by simp thus ?thesis by auto qed } note AUX_p_sc = this { fix i j assume A: "i < j" "j < length (p @ [{v}])" have "insert (u, v) lvE \ (p @ [{v}]) ! j \ (p @ [{v}]) ! i = {}" proof (cases "j=length p") case False with A have "ji this VNE show ?thesis by (auto simp add: nth_append) next from p_not_D A have PDDJ: "p!i \ D = {}" by (auto simp: Sup_inf_eq_bot_iff) case True thus ?thesis using A apply (simp add: nth_append) apply (rule conjI) using UIL A p_disjoint_sym apply (metis Misc.last_in_set NE UnionI VNE(1)) using vE_touched VNE PDDJ apply (auto simp: touched_def) [] done qed } note AUX_vE_no_back = this show ?thesis apply unfold_locales unfolding defs_fold apply simp using D_incr apply auto [] using pE_E_from_p apply auto [] using E_from_p_touched VNE apply (auto simp: touched_def) [] apply (rule D_reachable) apply (rule AUX_p_connected, assumption+) [] using p_disjoint \v\\(set p)\ apply (auto simp: nth_append) [] apply (rule AUX_p_sc, assumption+) [] using root_v0 apply simp apply simp apply (rule D_closed) apply (rule AUX_vE_no_back, assumption+) [] using p_not_D VNE apply auto [] done qed lemma invar_skip: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" and UIL: "u\last p" assumes VNP: "v\\(set p)" and VD: "v\D" shows "invar v0 D0 (p,D,pE - {(u, v)})" unfolding invar_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp let ?thesis = "invar_loc G v0 D0 p D (pE - {(u, v)})" note defs_fold = vE_remove[OF NE E] show ?thesis apply unfold_locales unfolding defs_fold apply simp using D_incr apply auto [] using pE_E_from_p apply auto [] using E_from_p_touched VD apply (auto simp: touched_def) [] apply (rule D_reachable) using p_connected apply auto [] apply (rule p_disjoint, assumption+) [] apply (drule p_sc) apply (erule order_trans) apply (rule rtrancl_mono) apply blast [] apply (rule root_v0, assumption+) [] apply (rule p_empty_v0, assumption+) [] apply (rule D_closed) using vE_no_back VD p_not_D apply clarsimp apply (metis Suc_lessD UnionI VNP less_trans_Suc nth_mem) apply (rule p_not_D) done qed lemma fin_D_is_reachable: \ \When inner loop terminates, all nodes reachable from start node are finished\ assumes INV: "invar v0 D0 ([], D, pE)" shows "D \ E\<^sup>*``{v0}" proof - from INV interpret invar_loc G v0 D0 "[]" D pE unfolding invar_def by auto from p_empty_v0 rtrancl_reachable_induct[OF order_refl D_closed] D_reachable show ?thesis by auto qed lemma fin_reachable_path: \ \When inner loop terminates, nodes reachable from start node are reachable over visited edges\ assumes INV: "invar v0 D0 ([], D, pE)" assumes UR: "u\E\<^sup>*``{v0}" shows "path (vE [] D pE) u q v \ path E u q v" proof - from INV interpret invar_loc G v0 D0 "[]" D pE unfolding invar_def by auto show ?thesis proof assume "path lvE u q v" thus "path E u q v" using path_mono[OF lvE_ss_E] by blast next assume "path E u q v" thus "path lvE u q v" using UR proof induction case (path_prepend u v p w) with fin_D_is_reachable[OF INV] have "u\D" by auto with D_closed \(u,v)\E\ have "v\D" by auto from path_prepend.prems path_prepend.hyps have "v\E\<^sup>*``{v0}" by auto with path_prepend.IH fin_D_is_reachable[OF INV] have "path lvE v p w" by simp moreover from \u\D\ \v\D\ \(u,v)\E\ D_vis have "(u,v)\lvE" by auto ultimately show ?case by (auto simp: path_cons_conv) qed simp qed qed lemma invar_outer_newnode: assumes A: "v0\D0" "v0\it" assumes OINV: "outer_invar it D0" assumes INV: "invar v0 D0 ([],D',pE)" shows "outer_invar (it-{v0}) D'" proof - from OINV interpret outer_invar_loc G it D0 unfolding outer_invar_def . from INV interpret inv: invar_loc G v0 D0 "[]" D' pE unfolding invar_def by simp from fin_D_is_reachable[OF INV] have [simp]: "v0\D'" by auto show ?thesis unfolding outer_invar_def apply unfold_locales using it_initial apply auto [] using it_done inv.D_incr apply auto [] using inv.D_reachable apply assumption using inv.D_closed apply assumption done qed lemma invar_outer_Dnode: assumes A: "v0\D0" "v0\it" assumes OINV: "outer_invar it D0" shows "outer_invar (it-{v0}) D0" proof - from OINV interpret outer_invar_loc G it D0 unfolding outer_invar_def . show ?thesis unfolding outer_invar_def apply unfold_locales using it_initial apply auto [] using it_done A apply auto [] using D_reachable apply assumption using D_closed apply assumption done qed lemma pE_fin': "invar x \ ([], D, pE) \ pE={}" unfolding invar_def by (simp add: invar_loc.pE_fin) end subsubsection \Termination\ context invar_loc begin lemma unproc_finite[simp, intro!]: "finite (unproc_edges v0 p D pE)" \ \The set of unprocessed edges is finite\ proof - have "unproc_edges v0 p D pE \ E\<^sup>*``{v0} \ E\<^sup>*``{v0}" unfolding unproc_edges_def using pE_reachable by auto thus ?thesis by (rule finite_subset) simp qed lemma unproc_decreasing: \ \As effect of selecting a pending edge, the set of unprocessed edges decreases\ assumes [simp]: "p\[]" and A: "(u,v)\pE" "u\last p" shows "unproc_edges v0 p D (pE-{(u,v)}) \ unproc_edges v0 p D pE" using A unfolding unproc_edges_def by fastforce end context fr_graph begin lemma abs_wf_pop: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes NO: "pE \ last aba \ UNIV = {}" shows "(pop (p,D,pE), (p, D, pE)) \ abs_wf_rel v0" unfolding pop_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp let ?thesis = "((butlast p, last p \ D, pE), p, D, pE) \ abs_wf_rel v0" have "unproc_edges v0 (butlast p) (last p \ D) pE = unproc_edges v0 p D pE" unfolding unproc_edges_def apply (cases p rule: rev_cases, simp) apply auto done thus ?thesis by (auto simp: abs_wf_rel_def) qed lemma abs_wf_collapse: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" "u\last p" shows "(collapse v (p,D,pE-{(u,v)}), (p, D, pE))\ abs_wf_rel v0" unfolding collapse_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp define i where "i = idx_of p v" let ?thesis = "((collapse_aux p i, D, pE-{(u,v)}), (p, D, pE)) \ abs_wf_rel v0" have "unproc_edges v0 (collapse_aux p i) D (pE-{(u,v)}) = unproc_edges v0 p D (pE-{(u,v)})" unfolding unproc_edges_def by (auto) also note unproc_decreasing[OF NE E] finally show ?thesis by (auto simp: abs_wf_rel_def) qed lemma abs_wf_push: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" "u\last p" and A: "v\D" "v\\(set p)" shows "(push v (p,D,pE-{(u,v)}), (p, D, pE)) \ abs_wf_rel v0" unfolding push_def apply simp proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp let ?thesis = "((p@[{v}], D, pE-{(u,v)} \ E\{v}\UNIV), (p, D, pE)) \ abs_wf_rel v0" have "unproc_edges v0 (p@[{v}]) D (pE-{(u,v)} \ E\{v}\UNIV) = unproc_edges v0 p D (pE-{(u,v)})" unfolding unproc_edges_def using E A pE_reachable by auto also note unproc_decreasing[OF NE E] finally show ?thesis by (auto simp: abs_wf_rel_def) qed lemma abs_wf_skip: assumes INV: "invar v0 D0 (p,D,pE)" assumes NE[simp]: "p\[]" assumes E: "(u,v)\pE" "u\last p" shows "((p, D, pE-{(u,v)}), (p, D, pE)) \ abs_wf_rel v0" proof - from INV interpret invar_loc G v0 D0 p D pE unfolding invar_def by simp from unproc_decreasing[OF NE E] show ?thesis by (auto simp: abs_wf_rel_def) qed end subsubsection \Main Correctness Theorem\ context fr_graph begin lemmas invar_preserve = invar_initial invar_pop invar_push invar_skip invar_collapse abs_wf_pop abs_wf_collapse abs_wf_push abs_wf_skip outer_invar_initial invar_outer_newnode invar_outer_Dnode text \The main correctness theorem for the dummy-algorithm just states that it satisfies the invariant when finished, and the path is empty. \ theorem skeleton_spec: "skeleton \ SPEC (\D. outer_invar {} D)" proof - note [simp del] = Union_iff note [[goals_limit = 4]] show ?thesis unfolding skeleton_def select_edge_def select_def apply (refine_vcg WHILEIT_rule[OF abs_wf_rel_wf]) apply (vc_solve solve: invar_preserve simp: pE_fin' finite_V0) apply auto done qed text \Short proof, as presented in the paper\ context notes [refine] = refine_vcg begin theorem "skeleton \ SPEC (\D. outer_invar {} D)" unfolding skeleton_def select_edge_def select_def by (refine_vcg WHILEIT_rule[OF abs_wf_rel_wf]) (auto intro: invar_preserve simp: pE_fin' finite_V0) end end subsection "Consequences of Invariant when Finished" context fr_graph begin lemma fin_outer_D_is_reachable: \ \When outer loop terminates, exactly the reachable nodes are finished\ assumes INV: "outer_invar {} D" shows "D = E\<^sup>*``V0" proof - from INV interpret outer_invar_loc G "{}" D unfolding outer_invar_def by auto from it_done rtrancl_reachable_induct[OF order_refl D_closed] D_reachable show ?thesis by auto qed end section \Refinement to Gabow's Data Structure\text_raw\\label{sec:algo-ds}\ text \ The implementation due to Gabow \<^cite>\"Gabow00"\ represents a path as a stack \S\ of single nodes, and a stack \B\ that contains the boundaries of the collapsed segments. Moreover, a map \I\ maps nodes to their stack indices. As we use a tail-recursive formulation, we use another stack \P :: (nat \ 'v set) list\ to represent the pending edges. The entries in \P\ are sorted by ascending first component, and \P\ only contains entries with non-empty second component. An entry \(i,l)\ means that the edges from the node at \S[i]\ to the nodes stored in \l\ are pending. \ subsection \Preliminaries\ primrec find_max_nat :: "nat \ (nat\bool) \ nat" \ \Find the maximum number below an upper bound for which a predicate holds\ where "find_max_nat 0 _ = 0" | "find_max_nat (Suc n) P = (if (P n) then n else find_max_nat n P)" lemma find_max_nat_correct: "\P 0; 0 \ find_max_nat u P = Max {i. i P i}" apply (induction u) apply auto apply (rule Max_eqI[THEN sym]) apply auto [3] apply (case_tac u) apply simp apply clarsimp by (metis less_SucI less_antisym) lemma find_max_nat_param[param]: assumes "(n,n')\nat_rel" assumes "\j j'. \(j,j')\nat_rel; j' \ (P j,P' j')\bool_rel" shows "(find_max_nat n P,find_max_nat n' P') \ nat_rel" using assms by (induction n arbitrary: n') auto context begin interpretation autoref_syn . lemma find_max_nat_autoref[autoref_rules]: assumes "(n,n')\nat_rel" assumes "\j j'. \(j,j')\nat_rel; j' \ (P j,P'$j')\bool_rel" shows "(find_max_nat n P, (OP find_max_nat ::: nat_rel \ (nat_rel\bool_rel) \ nat_rel) $n'$P' ) \ nat_rel" using find_max_nat_param[OF assms] by simp end subsection \Gabow's Datastructure\ subsubsection \Definition and Invariant\ datatype node_state = STACK nat | DONE type_synonym 'v oGS = "'v \ node_state" definition oGS_\ :: "'v oGS \ 'v set" where "oGS_\ I \ {v. I v = Some DONE}" locale oGS_invar = fixes I :: "'v oGS" assumes I_no_stack: "I v \ Some (STACK j)" type_synonym 'a GS = "'a list \ nat list \ ('a \ node_state) \ (nat \ 'a set) list" locale GS = fixes SBIP :: "'a GS" begin definition "S \ (\(S,B,I,P). S) SBIP" definition "B \ (\(S,B,I,P). B) SBIP" definition "I \ (\(S,B,I,P). I) SBIP" definition "P \ (\(S,B,I,P). P) SBIP" definition seg_start :: "nat \ nat" \ \Start index of segment, inclusive\ where "seg_start i \ B!i" definition seg_end :: "nat \ nat" \ \End index of segment, exclusive\ where "seg_end i \ if i+1 = length B then length S else B!(i+1)" definition seg :: "nat \ 'a set" \ \Collapsed set at index\ where "seg i \ {S!j | j. seg_start i \ j \ j < seg_end i }" definition "p_\ \ map seg [0.. \Collapsed path\ definition "D_\ \ {v. I v = Some DONE}" \ \Done nodes\ definition "pE_\ \ { (u,v) . \j I. (j,I)\set P \ u = S!j \ v\I }" \ \Pending edges\ definition "\ \ (p_\,D_\,pE_\)" \ \Abstract state\ end lemma GS_sel_simps[simp]: "GS.S (S,B,I,P) = S" "GS.B (S,B,I,P) = B" "GS.I (S,B,I,P) = I" "GS.P (S,B,I,P) = P" unfolding GS.S_def GS.B_def GS.I_def GS.P_def by auto context GS begin lemma seg_start_indep[simp]: "GS.seg_start (S',B,I',P') = seg_start" unfolding GS.seg_start_def[abs_def] by (auto) lemma seg_end_indep[simp]: "GS.seg_end (S,B,I',P') = seg_end" unfolding GS.seg_end_def[abs_def] by auto lemma seg_indep[simp]: "GS.seg (S,B,I',P') = seg" unfolding GS.seg_def[abs_def] by auto lemma p_\_indep[simp]: "GS.p_\ (S,B,I',P') = p_\" unfolding GS.p_\_def by auto lemma D_\_indep[simp]: "GS.D_\ (S',B',I,P') = D_\" unfolding GS.D_\_def by auto lemma pE_\_indep[simp]: "GS.pE_\ (S,B',I',P) = pE_\" unfolding GS.pE_\_def by auto definition find_seg \ \Abs-path index for stack index\ where "find_seg j \ Max {i. i B!i\j}" definition S_idx_of \ \Stack index for node\ where "S_idx_of v \ case I v of Some (STACK i) \ i" end locale GS_invar = GS + assumes B_in_bound: "set B \ {0..[] \ B\[] \ B!0=0" assumes S_distinct: "distinct S" assumes I_consistent: "(I v = Some (STACK j)) \ (j v = S!j)" assumes P_sorted: "sorted (map fst P)" assumes P_distinct: "distinct (map fst P)" assumes P_bound: "set P \ {0..Collect ((\) {})" begin lemma locale_this: "GS_invar SBIP" by unfold_locales end definition "oGS_rel \ br oGS_\ oGS_invar" lemma oGS_rel_sv[intro!,simp,relator_props]: "single_valued oGS_rel" unfolding oGS_rel_def by auto definition "GS_rel \ br GS.\ GS_invar" lemma GS_rel_sv[intro!,simp,relator_props]: "single_valued GS_rel" unfolding GS_rel_def by auto context GS_invar begin lemma empty_eq: "S=[] \ B=[]" using B_in_bound B0 by auto lemma B_in_bound': "i B!i < length S" using B_in_bound nth_mem by fastforce lemma seg_start_bound: assumes A: "i length S" proof (cases "i+1=length B") case True thus ?thesis by (simp add: seg_end_def) next case False with A have "i+1 seg_start i < seg_end i" unfolding seg_start_def seg_end_def using B_in_bound' distinct_sorted_mono[OF B_sorted B_distinct] by auto lemma seg_end_less_start: "\i \ seg_end i \ seg_start j" unfolding seg_start_def seg_end_def by (auto simp: distinct_sorted_mono_iff[OF B_distinct B_sorted]) lemma find_seg_bounds: assumes A: "j j" and "j < seg_end (find_seg j)" and "find_seg j < length B" proof - let ?M = "{i. i B!i\j}" from A have [simp]: "B\[]" using empty_eq by (cases S) auto have NE: "?M\{}" using A B0 by (cases B) auto have F: "finite ?M" by auto from Max_in[OF F NE] have LEN: "find_seg j < length B" and LB: "B!find_seg j \ j" unfolding find_seg_def by auto thus "find_seg j < length B" by - from LB show LB': "seg_start (find_seg j) \ j" unfolding seg_start_def by simp moreover show UB': "j < seg_end (find_seg j)" unfolding seg_end_def proof (split if_split, intro impI conjI) show "j length B" with LEN have P1: "find_seg j + 1 < length B" by simp show "j < B ! (find_seg j + 1)" proof (rule ccontr, simp only: linorder_not_less) assume P2: "B ! (find_seg j + 1) \ j" with P1 Max_ge[OF F, of "find_seg j + 1", folded find_seg_def] show False by simp qed qed qed lemma find_seg_correct: assumes A: "j seg (find_seg j)" and "find_seg j < length B" using find_seg_bounds[OF A] unfolding seg_def by auto lemma set_p_\_is_set_S: "\(set p_\) = set S" apply rule unfolding p_\_def seg_def[abs_def] using seg_end_bound apply fastforce [] apply (auto simp: in_set_conv_nth) using find_seg_bounds apply (fastforce simp: in_set_conv_nth) done lemma S_idx_uniq: "\i \ S!i=S!j \ i=j" using S_distinct by (simp add: nth_eq_iff_index_eq) lemma S_idx_of_correct: assumes A: "v\\(set p_\)" shows "S_idx_of v < length S" and "S!S_idx_of v = v" proof - from A have "v\set S" by (simp add: set_p_\_is_set_S) then obtain j where G1: "j_disjoint_sym: shows "\i j v. i \ j \ v\p_\!i \ v\p_\!j \ i=j" proof (intro allI impI, elim conjE) fix i j v assume A: "i < length p_\" "j < length p_\" "v \ p_\ ! i" "v \ p_\ ! j" from A have LI: "i_def) from A have B1: "seg_start j < seg_end i" and B2: "seg_start i < seg_end j" unfolding p_\_def seg_def[abs_def] apply clarsimp_all apply (subst (asm) S_idx_uniq) apply (metis dual_order.strict_trans1 seg_end_bound) apply (metis dual_order.strict_trans1 seg_end_bound) apply simp apply (subst (asm) S_idx_uniq) apply (metis dual_order.strict_trans1 seg_end_bound) apply (metis dual_order.strict_trans1 seg_end_bound) apply simp done from B1 have B1: "(B!j < B!Suc i \ Suc i < length B) \ i=length B - 1" using LI unfolding seg_start_def seg_end_def by (auto split: if_split_asm) from B2 have B2: "(B!i < B!Suc j \ Suc j < length B) \ j=length B - 1" using LJ unfolding seg_start_def seg_end_def by (auto split: if_split_asm) from B1 have B1: "j i=length B - 1" using LI LJ distinct_sorted_strict_mono_iff[OF B_distinct B_sorted] by auto from B2 have B2: "i j=length B - 1" using LI LJ distinct_sorted_strict_mono_iff[OF B_distinct B_sorted] by auto from B1 B2 show "i=j" using LI LJ by auto qed end subsection \Refinement of the Operations\ definition GS_initial_impl :: "'a oGS \ 'a \ 'a set \ 'a GS" where "GS_initial_impl I v0 succs \ ( [v0], [0], I(v0\(STACK 0)), if succs={} then [] else [(0,succs)])" context GS begin definition "push_impl v succs \ let _ = stat_newnode (); j = length S; S = S@[v]; B = B@[j]; I = I(v \ STACK j); P = if succs={} then P else P@[(j,succs)] in (S,B,I,P)" definition mark_as_done where "\l u I. mark_as_done l u I \ do { (_,I)\WHILET (\(l,I). l(l,I). do { ASSERT (l DONE))}) (l,I); RETURN I }" definition mark_as_done_abs where "\l u I. mark_as_done_abs l u I \ (\v. if v\{S!j | j. l\j \ jllength S\ \ mark_as_done l u I \ SPEC (\r. r = mark_as_done_abs l u I)" unfolding mark_as_done_def mark_as_done_abs_def apply (refine_rcg WHILET_rule[where I="\(l',I'). I' = (\v. if v\{S!j | j. l\j \ j l\l' \ l'\u" and R="measure (\(l',_). u-l')" ] refine_vcg) apply (auto intro!: ext simp: less_Suc_eq) done definition "pop_impl \ do { let lsi = length B - 1; ASSERT (lsi mark_as_done (seg_start lsi) (seg_end lsi) I; ASSERT (B\[]); let S = take (last B) S; ASSERT (B\[]); let B = butlast B; RETURN (S,B,I,P) }" definition "sel_rem_last \ if P=[] then RETURN (None,(S,B,I,P)) else do { let (j,succs) = last P; ASSERT (length B - 1 < length B); if j \ seg_start (length B - 1) then do { ASSERT (succs\{}); v \ SPEC (\x. x\succs); let succs = succs - {v}; ASSERT (P\[] \ length P - 1 < length P); let P = (if succs={} then butlast P else P[length P - 1 := (j,succs)]); RETURN (Some v,(S,B,I,P)) } else RETURN (None,(S,B,I,P)) }" definition "find_seg_impl j \ find_max_nat (length B) (\i. B!i\j)" lemma (in GS_invar) find_seg_impl: "j find_seg_impl j = find_seg j" unfolding find_seg_impl_def thm find_max_nat_correct apply (subst find_max_nat_correct) apply (simp add: B0) apply (simp add: B0) apply (simp add: find_seg_def) done definition "idx_of_impl v \ do { ASSERT (\i. I v = Some (STACK i)); let j = S_idx_of v; ASSERT (j do { i\idx_of_impl v; ASSERT (i+1 \ length B); let B = take (i+1) B; RETURN (S,B,I,P) }" end lemma (in -) GS_initial_correct: assumes REL: "(I,D)\oGS_rel" assumes A: "v0\D" shows "GS.\ (GS_initial_impl I v0 succs) = ([{v0}],D,{v0}\succs)" (is ?G1) and "GS_invar (GS_initial_impl I v0 succs)" (is ?G2) proof - from REL have [simp]: "D = oGS_\ I" and I: "oGS_invar I" by (simp_all add: oGS_rel_def br_def) from I have [simp]: "\j v. I v \ Some (STACK j)" by (simp add: oGS_invar_def) show ?G1 unfolding GS.\_def GS_initial_impl_def apply (simp split del: if_split) apply (intro conjI) unfolding GS.p_\_def GS.seg_def[abs_def] GS.seg_start_def GS.seg_end_def apply (auto) [] using A unfolding GS.D_\_def apply (auto simp: oGS_\_def) [] unfolding GS.pE_\_def apply auto [] done show ?G2 unfolding GS_initial_impl_def apply unfold_locales apply auto done qed context GS_invar begin lemma push_correct: assumes A: "v\\(set p_\)" and B: "v\D_\" shows "GS.\ (push_impl v succs) = (p_\@[{v}],D_\,pE_\ \ {v}\succs)" (is ?G1) and "GS_invar (push_impl v succs)" (is ?G2) proof - note [simp] = Let_def have A1: "GS.D_\ (push_impl v succs) = D_\" using B by (auto simp: push_impl_def GS.D_\_def) have iexI: "\a b j P. \a!j = b!j; P j\ \ \j'. a!j = b!j' \ P j'" by blast have A2: "GS.p_\ (push_impl v succs) = p_\ @ [{v}]" unfolding push_impl_def GS.p_\_def GS.seg_def[abs_def] GS.seg_start_def GS.seg_end_def apply (clarsimp split del: if_split) apply clarsimp apply safe apply (((rule iexI)?, (auto simp: nth_append nat_in_between_eq dest: order.strict_trans[OF _ B_in_bound'] )) [] ) + done have iexI2: "\j I Q. \(j,I)\set P; (j,I)\set P \ Q j\ \ \j. Q j" by blast have A3: "GS.pE_\ (push_impl v succs) = pE_\ \ {v} \ succs" unfolding push_impl_def GS.pE_\_def using P_bound apply (force simp: nth_append elim!: iexI2) done show ?G1 unfolding GS.\_def by (simp add: A1 A2 A3) show ?G2 apply unfold_locales unfolding push_impl_def apply simp_all using B_in_bound B_sorted B_distinct apply (auto simp: sorted_append) [3] using B_in_bound B0 apply (cases S) apply (auto simp: nth_append) [2] using S_distinct A apply (simp add: set_p_\_is_set_S) using A I_consistent apply (auto simp: nth_append set_p_\_is_set_S split: if_split_asm) [] using P_sorted P_distinct P_bound apply (auto simp: sorted_append) [3] done qed lemma no_last_out_P_aux: assumes NE: "p_\\[]" and NS: "pE_\ \ last p_\ \ UNIV = {}" shows "set P \ {0.. UNIV" proof - { fix j I assume jII: "(j,I)\set P" and JL: "last B\j" with P_bound have JU: "j{}" by auto with JL JU have "S!j \ last p_\" using NE unfolding p_\_def apply (auto simp: last_map seg_def seg_start_def seg_end_def last_conv_nth) done moreover from jII have "{S!j} \ I \ pE_\" unfolding pE_\_def by auto moreover note INE NS ultimately have False by blast } thus ?thesis by fastforce qed lemma pop_correct: assumes NE: "p_\\[]" and NS: "pE_\ \ last p_\ \ UNIV = {}" shows "pop_impl \ \GS_rel (SPEC (\r. r=(butlast p_\, D_\ \ last p_\, pE_\)))" proof - have iexI: "\a b j P. \a!j = b!j; P j\ \ \j'. a!j = b!j' \ P j'" by blast have [simp]: "\n. n - Suc 0 \ n \ n\0" by auto from NE have BNE: "B\[]" unfolding p_\_def by auto { fix i j assume B: "j last B" by (simp add: last_conv_nth) finally have "j < last B" . hence "take (last B) S ! j = S ! j" and "take (B!(length B - Suc 0)) S !j = S!j" by (simp_all add: last_conv_nth BNE) } note AUX1=this { fix v j have "(mark_as_done_abs (seg_start (length B - Suc 0)) (seg_end (length B - Suc 0)) I v = Some (STACK j)) \ (j < length S \ j < last B \ v = take (last B) S ! j)" apply (simp add: mark_as_done_abs_def) apply safe [] using I_consistent apply (clarsimp_all simp: seg_start_def seg_end_def last_conv_nth BNE simp: S_idx_uniq) apply (force) apply (subst nth_take) apply force apply force done } note AUX2 = this define ci where "ci = ( take (last B) S, butlast B, mark_as_done_abs (seg_start (length B - Suc 0)) (seg_end (length B - Suc 0)) I, P)" have ABS: "GS.\ ci = (butlast p_\, D_\ \ last p_\, pE_\)" apply (simp add: GS.\_def ci_def) apply (intro conjI) apply (auto simp del: map_butlast simp add: map_butlast[symmetric] butlast_upt simp add: GS.p_\_def GS.seg_def[abs_def] GS.seg_start_def GS.seg_end_def simp: nth_butlast last_conv_nth nth_take AUX1 cong: if_cong intro!: iexI dest: order.strict_trans[OF _ B_in_bound'] ) [] apply (auto simp: GS.D_\_def p_\_def last_map BNE seg_def mark_as_done_abs_def) [] using AUX1 no_last_out_P_aux[OF NE NS] apply (auto simp: GS.pE_\_def mark_as_done_abs_def elim!: bex2I) [] done have INV: "GS_invar ci" apply unfold_locales apply (simp_all add: ci_def) using B_in_bound B_sorted B_distinct apply (cases B rule: rev_cases, simp) apply (auto simp: sorted_append order.strict_iff_order) [] using B_sorted BNE apply (auto simp: sorted_butlast) [] using B_distinct BNE apply (auto simp: distinct_butlast) [] using B0 apply (cases B rule: rev_cases, simp add: BNE) apply (auto simp: nth_append split: if_split_asm) [] using S_distinct apply (auto) [] apply (rule AUX2) using P_sorted P_distinct apply (auto) [2] using P_bound no_last_out_P_aux[OF NE NS] apply (auto simp: in_set_conv_decomp) done show ?thesis unfolding pop_impl_def apply (refine_rcg SPEC_refine refine_vcg order_trans[OF mark_as_done_aux]) apply (simp_all add: BNE seg_start_less_end seg_end_bound) apply (fold ci_def) unfolding GS_rel_def apply (rule brI) apply (simp_all add: ABS INV) done qed lemma sel_rem_last_correct: assumes NE: "p_\\[]" shows "sel_rem_last \ \(Id \\<^sub>r GS_rel) (select_edge (p_\,D_\,pE_\))" proof - { fix l i a b b' have "\i \ map fst (l[i:=(a,b')]) = map fst l" by (induct l arbitrary: i) (auto split: nat.split) } note map_fst_upd_snd_eq = this from NE have BNE[simp]: "B\[]" unfolding p_\_def by simp have INVAR: "sel_rem_last \ SPEC (GS_invar o snd)" unfolding sel_rem_last_def apply (refine_rcg refine_vcg) using locale_this apply (cases SBIP) apply simp apply simp using P_bound apply (cases P rule: rev_cases, auto) [] apply simp apply simp apply (intro impI conjI) apply (unfold_locales, simp_all) [] using B_in_bound B_sorted B_distinct B0 S_distinct I_consistent apply auto [6] using P_sorted P_distinct apply (auto simp: map_butlast sorted_butlast distinct_butlast) [2] using P_bound apply (auto dest: in_set_butlastD) [] apply (unfold_locales, simp_all) [] using B_in_bound B_sorted B_distinct B0 S_distinct I_consistent apply auto [6] using P_sorted P_distinct apply (auto simp: last_conv_nth map_fst_upd_snd_eq) [2] using P_bound apply (cases P rule: rev_cases, simp) apply (auto) [] using locale_this apply (cases SBIP) apply simp done { assume NS: "pE_\\last p_\\UNIV = {}" hence "sel_rem_last \ SPEC (\r. case r of (None,SBIP') \ SBIP'=SBIP | _ \ False)" unfolding sel_rem_last_def apply (refine_rcg refine_vcg) apply (cases SBIP) apply simp apply simp using P_bound apply (cases P rule: rev_cases, auto) [] apply simp using no_last_out_P_aux[OF NE NS] apply (auto simp: seg_start_def last_conv_nth) [] apply (cases SBIP) apply simp done } note SPEC_E = this { assume NON_EMPTY: "pE_\\last p_\\UNIV \ {}" then obtain j succs P' where EFMT: "P = P'@[(j,succs)]" unfolding pE_\_def by (cases P rule: rev_cases) auto with P_bound have J_UPPER: "j{}" by auto have J_LOWER: "seg_start (length B - Suc 0) \ j" proof (rule ccontr) assume "\(seg_start (length B - Suc 0) \ j)" hence "j < seg_start (length B - 1)" by simp with P_sorted EFMT have P_bound': "set P \ {0.. UNIV" by (auto simp: sorted_append) hence "pE_\ \ last p_\\UNIV = {}" by (auto simp: p_\_def last_conv_nth seg_def pE_\_def S_idx_uniq seg_end_def) thus False using NON_EMPTY by simp qed from J_UPPER J_LOWER have SJL: "S!j\last p_\" unfolding p_\_def seg_def[abs_def] seg_end_def by (auto simp: last_map) from EFMT have SSS: "{S!j}\succs\pE_\" unfolding pE_\_def by auto { fix v assume "v\succs" with SJL SSS have G: "(S!j,v)\pE_\ \ last p_\\UNIV" by auto { fix j' succs' assume "S ! j' = S ! j" "(j', succs') \ set P'" with J_UPPER P_bound S_idx_uniq EFMT have "j'=j" by auto with P_distinct \(j', succs') \ set P'\ EFMT have False by auto } note AUX3=this have G1: "GS.pE_\ (S,B,I,P' @ [(j, succs - {v})]) = pE_\ - {(S!j, v)}" unfolding GS.pE_\_def using AUX3 by (auto simp: EFMT) { assume "succs\{v}" hence "GS.pE_\ (S,B,I,P' @ [(j, succs - {v})]) = GS.pE_\ (S,B,I,P')" unfolding GS.pE_\_def by auto with G1 have "GS.pE_\ (S,B,I,P') = pE_\ - {(S!j, v)}" by simp } note G2 = this note G G1 G2 } note AUX3 = this have "sel_rem_last \ SPEC (\r. case r of (Some v,SBIP') \ \u. (u,v)\(pE_\\last p_\\UNIV) \ GS.\ SBIP' = (p_\,D_\,pE_\-{(u,v)}) | _ \ False)" unfolding sel_rem_last_def apply (refine_rcg refine_vcg) using SNE apply (vc_solve simp: J_LOWER EFMT) apply (frule AUX3(1)) apply safe apply (drule (1) AUX3(3)) apply (auto simp: EFMT GS.\_def) [] apply (drule AUX3(2)) apply (auto simp: GS.\_def) [] done } note SPEC_NE=this have SPEC: "sel_rem_last \ SPEC (\r. case r of (None, SBIP') \ SBIP' = SBIP \ pE_\ \ last p_\ \ UNIV = {} \ GS_invar SBIP | (Some v, SBIP') \ \u. (u, v) \ pE_\ \ last p_\ \ UNIV \ GS.\ SBIP' = (p_\, D_\, pE_\ - {(u, v)}) \ GS_invar SBIP' )" using INVAR apply (cases "pE_\ \ last p_\ \ UNIV = {}") apply (frule SPEC_E) apply (auto split: option.splits simp: pw_le_iff; blast; fail) apply (frule SPEC_NE) apply (auto split: option.splits simp: pw_le_iff; blast; fail) done have X1: "(\y. (y=None \ \ y) \ (\a b. y=Some (a,b) \ \ y a b)) \ (\ None \ (\a b. \ (Some (a,b)) a b))" for \ \ by auto show ?thesis apply (rule order_trans[OF SPEC]) unfolding select_edge_def select_def apply (simp add: pw_le_iff refine_pw_simps prod_rel_sv del: SELECT_pw split: option.splits prod.splits) apply (fastforce simp: br_def GS_rel_def GS.\_def) done qed lemma find_seg_idx_of_correct: assumes A: "v\\(set p_\)" shows "(find_seg (S_idx_of v)) = idx_of p_\ v" proof - note S_idx_of_correct[OF A] idx_of_props[OF p_\_disjoint_sym A] from find_seg_correct[OF \S_idx_of v < length S\] have "find_seg (S_idx_of v) < length p_\" and "S!S_idx_of v \ p_\!find_seg (S_idx_of v)" unfolding p_\_def by auto from idx_of_uniq[OF p_\_disjoint_sym this] \S ! S_idx_of v = v\ show ?thesis by auto qed lemma idx_of_correct: assumes A: "v\\(set p_\)" shows "idx_of_impl v \ SPEC (\x. x=idx_of p_\ v \ x_is_set_S) apply (erule S_idx_of_correct) apply (simp add: find_seg_impl find_seg_idx_of_correct) by (metis find_seg_correct(2) find_seg_impl) lemma collapse_correct: assumes A: "v\\(set p_\)" shows "collapse_impl v \\GS_rel (SPEC (\r. r=collapse v \))" proof - { fix i assume "i" hence ILEN: "i_def) let ?SBIP' = "(S, take (Suc i) B, I, P)" { have [simp]: "GS.seg_start ?SBIP' i = seg_start i" by (simp add: GS.seg_start_def) have [simp]: "GS.seg_end ?SBIP' i = seg_end (length B - 1)" using ILEN by (simp add: GS.seg_end_def min_absorb2) { fix j assume B: "seg_start i \ j" "j < seg_end (length B - Suc 0)" hence "ji have "(length B - Suc 0) < length B" by auto from seg_end_bound[OF this] have "seg_end (length B - Suc 0) \ length S" . finally show ?thesis . qed have "i \ find_seg j \ find_seg j < length B \ seg_start (find_seg j) \ j \ j < seg_end (find_seg j)" proof (intro conjI) show "i\find_seg j" by (metis le_trans not_less B(1) find_seg_bounds(2) seg_end_less_start ILEN \j < length S\) qed (simp_all add: find_seg_bounds[OF \j]) } note AUX1 = this { fix Q and j::nat assume "Q j" hence "\i. S!j = S!i \ Q i" by blast } note AUX_ex_conj_SeqSI = this have "GS.seg ?SBIP' i = \ (seg ` {i.. (S, take (Suc i) B, I, P) = collapse_aux p_\ i" unfolding GS.p_\_def collapse_aux_def apply (simp add: min_absorb2 drop_map) apply (rule conjI) apply (auto simp: GS.seg_def[abs_def] GS.seg_start_def GS.seg_end_def take_map) [] apply (simp add: AUX2) done } note AUX1 = this from A obtain i where [simp]: "I v = Some (STACK i)" using I_consistent set_p_\_is_set_S by (auto simp: in_set_conv_nth) { have "(collapse_aux p_\ (idx_of p_\ v), D_\, pE_\) = GS.\ (S, take (Suc (idx_of p_\ v)) B, I, P)" unfolding GS.\_def using idx_of_props[OF p_\_disjoint_sym A] by (simp add: AUX1) } note ABS=this { have "GS_invar (S, take (Suc (idx_of p_\ v)) B, I, P)" apply unfold_locales apply simp_all using B_in_bound B_sorted B_distinct apply (auto simp: sorted_wrt_take dest: in_set_takeD) [3] using B0 S_distinct apply auto [2] using I_consistent apply simp using P_sorted P_distinct P_bound apply auto [3] done } note INV=this show ?thesis unfolding collapse_impl_def apply (refine_rcg SPEC_refine refine_vcg order_trans[OF idx_of_correct]) - - apply fact - apply (metis discrete) - - apply (simp add: collapse_def \_def find_seg_impl) - unfolding GS_rel_def + apply fact + apply simp + apply (simp add: collapse_def \_def find_seg_impl GS_rel_def) apply (rule brI) apply (rule ABS) apply (rule INV) done qed end text \Technical adjustment for avoiding case-splits for definitions extracted from GS-locale\ lemma opt_GSdef: "f \ g \ f s \ case s of (S,B,I,P) \ g (S,B,I,P)" by auto lemma ext_def: "f\g \ f x \ g x" by auto context fr_graph begin definition "push_impl v s \ GS.push_impl s v (E``{v})" lemmas push_impl_def_opt = push_impl_def[abs_def, THEN ext_def, THEN opt_GSdef, unfolded GS.push_impl_def GS_sel_simps] text \Definition for presentation\ lemma "push_impl v (S,B,I,P) \ (S@[v], B@[length S], I(v\STACK (length S)), if E``{v}={} then P else P@[(length S,E``{v})])" unfolding push_impl_def GS.push_impl_def GS.P_def GS.S_def by (auto simp: Let_def) lemma GS_\_split: "GS.\ s = (p,D,pE) \ (p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s)" "(p,D,pE) = GS.\ s \ (p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s)" by (auto simp add: GS.\_def) lemma push_refine: assumes A: "(s,(p,D,pE))\GS_rel" "(v,v')\Id" assumes B: "v\\(set p)" "v\D" shows "(push_impl v s, push v' (p,D,pE))\GS_rel" proof - from A have [simp]: "p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s" "v'=v" and INV: "GS_invar s" by (auto simp add: GS_rel_def br_def GS_\_split) from INV B show ?thesis by (auto simp: GS_rel_def br_def GS_invar.push_correct push_impl_def push_def) qed definition "pop_impl s \ GS.pop_impl s" lemmas pop_impl_def_opt = pop_impl_def[abs_def, THEN opt_GSdef, unfolded GS.pop_impl_def GS.mark_as_done_def GS.seg_start_def GS.seg_end_def GS_sel_simps] lemma pop_refine: assumes A: "(s,(p,D,pE))\GS_rel" assumes B: "p \ []" "pE \ last p \ UNIV = {}" shows "pop_impl s \ \GS_rel (RETURN (pop (p,D,pE)))" proof - from A have [simp]: "p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s" and INV: "GS_invar s" by (auto simp add: GS_rel_def br_def GS_\_split) show ?thesis unfolding pop_impl_def[abs_def] pop_def apply (rule order_trans[OF GS_invar.pop_correct]) using INV B apply (simp_all add: Un_commute RETURN_def) done qed thm pop_refine[no_vars] definition "collapse_impl v s \ GS.collapse_impl s v" lemmas collapse_impl_def_opt = collapse_impl_def[abs_def, THEN ext_def, THEN opt_GSdef, unfolded GS.collapse_impl_def GS_sel_simps] lemma collapse_refine: assumes A: "(s,(p,D,pE))\GS_rel" "(v,v')\Id" assumes B: "v'\\(set p)" shows "collapse_impl v s \\GS_rel (RETURN (collapse v' (p,D,pE)))" proof - from A have [simp]: "p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s" "v'=v" and INV: "GS_invar s" by (auto simp add: GS_rel_def br_def GS_\_split) show ?thesis unfolding collapse_impl_def[abs_def] apply (rule order_trans[OF GS_invar.collapse_correct]) using INV B by (simp_all add: GS.\_def RETURN_def) qed definition "select_edge_impl s \ GS.sel_rem_last s" lemmas select_edge_impl_def_opt = select_edge_impl_def[abs_def, THEN opt_GSdef, unfolded GS.sel_rem_last_def GS.seg_start_def GS_sel_simps] lemma select_edge_refine: assumes A: "(s,(p,D,pE))\GS_rel" assumes NE: "p \ []" shows "select_edge_impl s \ \(Id \\<^sub>r GS_rel) (select_edge (p,D,pE))" proof - from A have [simp]: "p=GS.p_\ s \ D=GS.D_\ s \ pE=GS.pE_\ s" and INV: "GS_invar s" by (auto simp add: GS_rel_def br_def GS_\_split) from INV NE show ?thesis unfolding select_edge_impl_def using GS_invar.sel_rem_last_correct[OF INV] NE by (simp) qed definition "initial_impl v0 I \ GS_initial_impl I v0 (E``{v0})" lemma initial_refine: "\v0\D0; (I,D0)\oGS_rel; (v0i,v0)\Id\ \ (initial_impl v0i I,initial v0 D0)\GS_rel" unfolding initial_impl_def GS_rel_def br_def apply (simp_all add: GS_initial_correct) apply (auto simp: initial_def) done definition "path_is_empty_impl s \ GS.S s = []" lemma path_is_empty_refine: "GS_invar s \ path_is_empty_impl s \ GS.p_\ s=[]" unfolding path_is_empty_impl_def GS.p_\_def GS_invar.empty_eq by auto definition (in GS) "is_on_stack_impl v \ case I v of Some (STACK _) \ True | _ \ False" lemma (in GS_invar) is_on_stack_impl_correct: shows "is_on_stack_impl v \ v\\(set p_\)" unfolding is_on_stack_impl_def using I_consistent[of v] apply (force simp: set_p_\_is_set_S in_set_conv_nth split: option.split node_state.split) done definition "is_on_stack_impl v s \ GS.is_on_stack_impl s v" lemmas is_on_stack_impl_def_opt = is_on_stack_impl_def[abs_def, THEN ext_def, THEN opt_GSdef, unfolded GS.is_on_stack_impl_def GS_sel_simps] lemma is_on_stack_refine: "\ GS_invar s \ \ is_on_stack_impl v s \ v\\(set (GS.p_\ s))" unfolding is_on_stack_impl_def GS_rel_def br_def by (simp add: GS_invar.is_on_stack_impl_correct) definition (in GS) "is_done_impl v \ case I v of Some DONE \ True | _ \ False" lemma (in GS_invar) is_done_impl_correct: shows "is_done_impl v \ v\D_\" unfolding is_done_impl_def D_\_def apply (auto split: option.split node_state.split) done definition "is_done_oimpl v I \ case I v of Some DONE \ True | _ \ False" definition "is_done_impl v s \ GS.is_done_impl s v" lemma is_done_orefine: "\ oGS_invar s \ \ is_done_oimpl v s \ v\oGS_\ s" unfolding is_done_oimpl_def oGS_rel_def br_def by (auto simp: oGS_invar_def oGS_\_def split: option.splits node_state.split) lemma is_done_refine: "\ GS_invar s \ \ is_done_impl v s \ v\GS.D_\ s" unfolding is_done_impl_def GS_rel_def br_def by (simp add: GS_invar.is_done_impl_correct) lemma oinitial_refine: "(Map.empty, {}) \ oGS_rel" by (auto simp: oGS_rel_def br_def oGS_\_def oGS_invar_def) end subsection \Refined Skeleton Algorithm\ context fr_graph begin lemma I_to_outer: assumes "((S, B, I, P), ([], D, {})) \ GS_rel" shows "(I,D)\oGS_rel" using assms unfolding GS_rel_def oGS_rel_def br_def oGS_\_def GS.\_def GS.D_\_def GS_invar_def oGS_invar_def apply (auto simp: GS.p_\_def) done definition skeleton_impl :: "'v oGS nres" where "skeleton_impl \ do { stat_start_nres; let I=Map.empty; r \ FOREACHi (\it I. outer_invar it (oGS_\ I)) V0 (\v0 I0. do { if \is_done_oimpl v0 I0 then do { let s = initial_impl v0 I0; (S,B,I,P)\WHILEIT (invar v0 (oGS_\ I0) o GS.\) (\s. \path_is_empty_impl s) (\s. do { \ \Select edge from end of path\ (vo,s) \ select_edge_impl s; case vo of Some v \ do { if is_on_stack_impl v s then do { collapse_impl v s } else if \is_done_impl v s then do { \ \Edge to new node. Append to path\ RETURN (push_impl v s) } else do { \ \Edge to done node. Skip\ RETURN s } } | None \ do { \ \No more outgoing edges from current node on path\ pop_impl s } }) s; RETURN I } else RETURN I0 }) I; stat_stop_nres; RETURN r }" subsubsection \Correctness Theorem\ lemma "skeleton_impl \ \oGS_rel skeleton" using [[goals_limit = 1]] unfolding skeleton_impl_def skeleton_def apply (refine_rcg bind_refine' select_edge_refine push_refine pop_refine collapse_refine initial_refine oinitial_refine inj_on_id ) using [[goals_limit = 5]] apply refine_dref_type apply (vc_solve (nopre) solve: asm_rl I_to_outer simp: GS_rel_def br_def GS.\_def oGS_rel_def oGS_\_def is_on_stack_refine path_is_empty_refine is_done_refine is_done_orefine ) done lemmas skeleton_refines = select_edge_refine push_refine pop_refine collapse_refine initial_refine oinitial_refine lemmas skeleton_refine_simps = GS_rel_def br_def GS.\_def oGS_rel_def oGS_\_def is_on_stack_refine path_is_empty_refine is_done_refine is_done_orefine text \Short proof, for presentation\ context notes [[goals_limit = 1]] notes [refine] = inj_on_id bind_refine' begin lemma "skeleton_impl \ \oGS_rel skeleton" unfolding skeleton_impl_def skeleton_def by (refine_rcg skeleton_refines, refine_dref_type) (vc_solve (nopre) solve: asm_rl I_to_outer simp: skeleton_refine_simps) end end end diff --git a/thys/Gaussian_Integers/Gaussian_Integers.thy b/thys/Gaussian_Integers/Gaussian_Integers.thy --- a/thys/Gaussian_Integers/Gaussian_Integers.thy +++ b/thys/Gaussian_Integers/Gaussian_Integers.thy @@ -1,2376 +1,2363 @@ (* File: Gaussian_Integers.thy Author: Manuel Eberl, TU München *) section \Gaussian Integers\ theory Gaussian_Integers imports "HOL-Computational_Algebra.Computational_Algebra" "HOL-Number_Theory.Number_Theory" begin subsection \Auxiliary material\ lemma coprime_iff_prime_factors_disjoint: fixes x y :: "'a :: factorial_semiring" assumes "x \ 0" "y \ 0" shows "coprime x y \ prime_factors x \ prime_factors y = {}" proof assume "coprime x y" have False if "p \ prime_factors x" "p \ prime_factors y" for p proof - from that assms have "p dvd x" "p dvd y" by (auto simp: prime_factors_dvd) with \coprime x y\ have "p dvd 1" using coprime_common_divisor by auto with that assms show False by (auto simp: prime_factors_dvd) qed thus "prime_factors x \ prime_factors y = {}" by auto next assume disjoint: "prime_factors x \ prime_factors y = {}" show "coprime x y" proof (rule coprimeI) fix d assume d: "d dvd x" "d dvd y" show "is_unit d" proof (rule ccontr) assume "\is_unit d" moreover from this and d assms have "d \ 0" by auto ultimately obtain p where p: "prime p" "p dvd d" using prime_divisor_exists by auto with d and assms have "p \ prime_factors x \ prime_factors y" by (auto simp: prime_factors_dvd) with disjoint show False by auto qed qed qed lemma product_dvd_irreducibleD: fixes a b x :: "'a :: algebraic_semidom" assumes "irreducible x" assumes "a * b dvd x" shows "a dvd 1 \ b dvd 1" proof - from assms obtain c where "x = a * b * c" by auto hence "x = a * (b * c)" by (simp add: mult_ac) from irreducibleD[OF assms(1) this] show "a dvd 1 \ b dvd 1" by (auto simp: is_unit_mult_iff) qed lemma prime_elem_mult_dvdI: assumes "prime_elem p" "p dvd c" "b dvd c" "\p dvd b" shows "p * b dvd c" proof - from assms(3) obtain a where c: "c = a * b" using mult.commute by blast with assms(2) have "p dvd a * b" by simp with assms have "p dvd a" by (subst (asm) prime_elem_dvd_mult_iff) auto with c show ?thesis by (auto intro: mult_dvd_mono) qed lemma prime_elem_power_mult_dvdI: fixes p :: "'a :: algebraic_semidom" assumes "prime_elem p" "p ^ n dvd c" "b dvd c" "\p dvd b" shows "p ^ n * b dvd c" proof (cases "n = 0") case False from assms(3) obtain a where c: "c = a * b" using mult.commute by blast with assms(2) have "p ^ n dvd b * a" by (simp add: mult_ac) hence "p ^ n dvd a" by (rule prime_power_dvd_multD[OF assms(1)]) (use assms False in auto) with c show ?thesis by (auto intro: mult_dvd_mono) qed (use assms in auto) lemma prime_mod_4_cases: fixes p :: nat assumes "prime p" shows "p = 2 \ [p = 1] (mod 4) \ [p = 3] (mod 4)" proof (cases "p = 2") case False with prime_gt_1_nat[of p] assms have "p > 2" by auto have "\4 dvd p" using assms product_dvd_irreducibleD[of p 2 2] by (auto simp: prime_elem_iff_irreducible simp flip: prime_elem_nat_iff) hence "p mod 4 \ 0" by (auto simp: mod_eq_0_iff_dvd) moreover have "p mod 4 \ 2" proof assume "p mod 4 = 2" hence "p mod 4 mod 2 = 0" by (simp add: cong_def) thus False using \prime p\ \p > 2\ prime_odd_nat[of p] by (auto simp: mod_mod_cancel) qed moreover have "p mod 4 \ {0,1,2,3}" by auto ultimately show ?thesis by (auto simp: cong_def) qed auto lemma of_nat_prod_mset: "of_nat (prod_mset A) = prod_mset (image_mset of_nat A)" by (induction A) auto lemma multiplicity_0_left [simp]: "multiplicity 0 x = 0" by (cases "x = 0") (auto simp: not_dvd_imp_multiplicity_0) lemma is_unit_power [intro]: "is_unit x \ is_unit (x ^ n)" by (subst is_unit_power_iff) auto lemma (in factorial_semiring) pow_divides_pow_iff: assumes "n > 0" shows "a ^ n dvd b ^ n \ a dvd b" proof (cases "b = 0") case False show ?thesis proof assume dvd: "a ^ n dvd b ^ n" with \b \ 0\ have "a \ 0" using \n > 0\ by (auto simp: power_0_left) show "a dvd b" proof (rule multiplicity_le_imp_dvd) fix p :: 'a assume p: "prime p" from dvd \b \ 0\ have "multiplicity p (a ^ n) \ multiplicity p (b ^ n)" by (intro dvd_imp_multiplicity_le) auto thus "multiplicity p a \ multiplicity p b" using p \a \ 0\ \b \ 0\ \n > 0\ by (simp add: prime_elem_multiplicity_power_distrib) qed fact+ qed (auto intro: dvd_power_same) qed (use assms in \auto simp: power_0_left\) lemma multiplicity_power_power: fixes p :: "'a :: {factorial_semiring, algebraic_semidom}" assumes "n > 0" shows "multiplicity (p ^ n) (x ^ n) = multiplicity p x" proof (cases "x = 0 \ p = 0 \ is_unit p") case True thus ?thesis using \n > 0\ by (auto simp: power_0_left is_unit_power_iff multiplicity_unit_left) next case False show ?thesis proof (intro antisym multiplicity_geI) have "(p ^ multiplicity p x) ^ n dvd x ^ n" by (intro dvd_power_same) (simp add: multiplicity_dvd) thus "(p ^ n) ^ multiplicity p x dvd x ^ n" by (simp add: mult_ac flip: power_mult) next have "(p ^ n) ^ multiplicity (p ^ n) (x ^ n) dvd x ^ n" by (simp add: multiplicity_dvd) hence "(p ^ multiplicity (p ^ n) (x ^ n)) ^ n dvd x ^ n" by (simp add: mult_ac flip: power_mult) thus "p ^ multiplicity (p ^ n) (x ^ n) dvd x" by (subst (asm) pow_divides_pow_iff) (use assms in auto) qed (use False \n > 0\ in \auto simp: is_unit_power_iff\) qed lemma even_square_cong_4_int: - fixes x :: int - assumes "even x" - shows "[x ^ 2 = 0] (mod 4)" + \[x\<^sup>2 = 0] (mod 4)\ if \even x\ for x :: int proof - - from assms have "even \x\" - by simp - hence [simp]: "\x\ mod 2 = 0" - by presburger - have "(\x\ ^ 2) mod 4 = ((\x\ mod 4) ^ 2) mod 4" - by (simp add: power_mod) - also from assms have "\x\ mod 4 = 0 \ \x\ mod 4 = 2" - using mod_double_modulus[of 2 "\x\"] by simp - hence "((\x\ mod 4) ^ 2) mod 4 = 0" - by auto - finally show ?thesis by (simp add: cong_def) + from that obtain y where \x = 2 * y\ .. + then show ?thesis by (simp add: cong_def) qed -lemma even_square_cong_4_nat: "even (x::nat) \ [x ^ 2 = 0] (mod 4)" - using even_square_cong_4_int[of "int x"] by (auto simp flip: cong_int_iff) +lemma even_square_cong_4_nat: + \[x\<^sup>2 = 0] (mod 4)\ if \even x\ for x :: nat + using that even_square_cong_4_int [of \int x\] by (simp flip: cong_int_iff) lemma odd_square_cong_4_int: - fixes x :: int - assumes "odd x" - shows "[x ^ 2 = 1] (mod 4)" + \[x\<^sup>2 = 1] (mod 4)\ if \odd x\ for x :: int proof - - from assms have "odd \x\" + from that obtain y where \x = 2 * y + 1\ .. + then have \x\<^sup>2 = 4 * (y\<^sup>2 + y) + 1\ + by (simp add: power2_eq_square algebra_simps) + also have \\ mod 4 = ((4 * (y\<^sup>2 + y)) mod 4 + 1 mod 4) mod 4\ + by (simp only: mod_simps) + also have \\ = 1 mod 4\ by simp - hence [simp]: "\x\ mod 2 = 1" - by presburger - have "(\x\ ^ 2) mod 4 = ((\x\ mod 4) ^ 2) mod 4" - by (simp add: power_mod) - also from assms have "\x\ mod 4 = 1 \ \x\ mod 4 = 3" - using mod_double_modulus[of 2 "\x\"] by simp - hence "((\x\ mod 4) ^ 2) mod 4 = 1" - by auto - finally show ?thesis by (simp add: cong_def) + finally show ?thesis + by (simp only: cong_def) qed -lemma odd_square_cong_4_nat: "odd (x::nat) \ [x ^ 2 = 1] (mod 4)" - using odd_square_cong_4_int[of "int x"] by (auto simp flip: cong_int_iff) +lemma odd_square_cong_4_nat: + \[x\<^sup>2 = 1] (mod 4)\ if \odd x\ for x :: nat + using that odd_square_cong_4_int [of \int x\] by (simp flip: cong_int_iff) text \ Gaussian integers will require a notion of an element being a power up to a unit, so we introduce this here. This should go in the library eventually. \ definition is_nth_power_upto_unit where "is_nth_power_upto_unit n x \ (\u. is_unit u \ is_nth_power n (u * x))" lemma is_nth_power_upto_unit_base: "is_nth_power n x \ is_nth_power_upto_unit n x" by (auto simp: is_nth_power_upto_unit_def intro: exI[of _ 1]) lemma is_nth_power_upto_unitI: assumes "normalize (x ^ n) = normalize y" shows "is_nth_power_upto_unit n y" proof - from associatedE1[OF assms] obtain u where "is_unit u" "u * y = x ^ n" by metis thus ?thesis by (auto simp: is_nth_power_upto_unit_def intro!: exI[of _ u]) qed lemma is_nth_power_upto_unit_conv_multiplicity: fixes x :: "'a :: factorial_semiring" assumes "n > 0" shows "is_nth_power_upto_unit n x \ (\p. prime p \ n dvd multiplicity p x)" proof (cases "x = 0") case False show ?thesis proof safe fix p :: 'a assume p: "prime p" assume "is_nth_power_upto_unit n x" then obtain u y where uy: "is_unit u" "u * x = y ^ n" by (auto simp: is_nth_power_upto_unit_def elim!: is_nth_powerE) from p uy assms False have [simp]: "y \ 0" by (auto simp: power_0_left) have "multiplicity p (u * x) = multiplicity p (y ^ n)" by (subst uy(2) [symmetric]) simp also have "multiplicity p (u * x) = multiplicity p x" by (simp add: multiplicity_times_unit_right uy(1)) finally show "n dvd multiplicity p x" using False and p and uy and assms by (auto simp: prime_elem_multiplicity_power_distrib) next assume *: "\p. prime p \ n dvd multiplicity p x" have "multiplicity p ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n) = multiplicity p x" if "prime p" for p proof - from that and * have "n dvd multiplicity p x" by blast have "multiplicity p x = 0" if "p \ prime_factors x" using that and \prime p\ by (simp add: prime_factors_multiplicity) with that and * and assms show ?thesis unfolding prod_power_distrib power_mult [symmetric] by (subst multiplicity_prod_prime_powers) (auto simp: in_prime_factors_imp_prime elim: dvdE) qed with assms False have "normalize ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n) = normalize x" by (intro multiplicity_eq_imp_eq) (auto simp: multiplicity_prod_prime_powers) thus "is_nth_power_upto_unit n x" by (auto intro: is_nth_power_upto_unitI) qed qed (use assms in \auto simp: is_nth_power_upto_unit_def\) lemma is_nth_power_upto_unit_0_left [simp, intro]: "is_nth_power_upto_unit 0 x \ is_unit x" proof assume "is_unit x" thus "is_nth_power_upto_unit 0 x" unfolding is_nth_power_upto_unit_def by (intro exI[of _ "1 div x"]) auto next assume "is_nth_power_upto_unit 0 x" then obtain u where "is_unit u" "u * x = 1" by (auto simp: is_nth_power_upto_unit_def) thus "is_unit x" by (metis dvd_triv_right) qed lemma is_nth_power_upto_unit_unit [simp, intro]: assumes "is_unit x" shows "is_nth_power_upto_unit n x" using assms by (auto simp: is_nth_power_upto_unit_def intro!: exI[of _ "1 div x"]) lemma is_nth_power_upto_unit_1_left [simp, intro]: "is_nth_power_upto_unit 1 x" by (auto simp: is_nth_power_upto_unit_def intro: exI[of _ 1]) lemma is_nth_power_upto_unit_mult_coprimeD1: fixes x y :: "'a :: factorial_semiring" assumes "coprime x y" "is_nth_power_upto_unit n (x * y)" shows "is_nth_power_upto_unit n x" proof - consider "n = 0" | "x = 0" "n > 0" | "x \ 0" "y = 0" "n > 0" | "n > 0" "x \ 0" "y \ 0" by force thus ?thesis proof cases assume [simp]: "n = 0" from assms have "is_unit (x * y)" by auto hence "is_unit x" using is_unit_mult_iff by blast thus ?thesis using assms by auto next assume "x = 0" "n > 0" thus ?thesis by (auto simp: is_nth_power_upto_unit_def) next assume *: "x \ 0" "y = 0" "n > 0" with assms show ?thesis by auto next assume *: "n > 0" and [simp]: "x \ 0" "y \ 0" show ?thesis proof (subst is_nth_power_upto_unit_conv_multiplicity[OF \n > 0\]; safe) fix p :: 'a assume p: "prime p" show "n dvd multiplicity p x" proof (cases "p dvd x") case False thus ?thesis by (simp add: not_dvd_imp_multiplicity_0) next case True have "n dvd multiplicity p (x * y)" using assms(2) \n > 0\ p by (auto simp: is_nth_power_upto_unit_conv_multiplicity) also have "\ = multiplicity p x + multiplicity p y" using p by (subst prime_elem_multiplicity_mult_distrib) auto also have "\p dvd y" using \coprime x y\ \p dvd x\ p not_prime_unit coprime_common_divisor by blast hence "multiplicity p y = 0" by (rule not_dvd_imp_multiplicity_0) finally show ?thesis by simp qed qed qed qed lemma is_nth_power_upto_unit_mult_coprimeD2: fixes x y :: "'a :: factorial_semiring" assumes "coprime x y" "is_nth_power_upto_unit n (x * y)" shows "is_nth_power_upto_unit n y" using assms is_nth_power_upto_unit_mult_coprimeD1[of y x] by (simp_all add: mult_ac coprime_commute) subsection \Definition\ text \ Gaussian integers are the ring $\mathbb{Z}[i]$ which is formed either by formally adjoining an element $i$ with $i^2 = -1$ to $\mathbb{Z}$ or by taking all the complex numbers with integer real and imaginary part. We define them simply by giving an appropriate ring structure to $\mathbb{Z}^2$, with the first component representing the real part and the second component the imaginary part: \ codatatype gauss_int = Gauss_Int (ReZ: int) (ImZ: int) text \ The following is the imaginary unit $i$ in the Gaussian integers, which we will denote as \\\<^sub>\\: \ primcorec gauss_i where "ReZ gauss_i = 0" | "ImZ gauss_i = 1" lemma gauss_int_eq_iff: "x = y \ ReZ x = ReZ y \ ImZ x = ImZ y" by (cases x; cases y) auto (*<*) bundle gauss_int_notation begin notation gauss_i ("\\<^sub>\") end bundle no_gauss_int_notation begin no_notation (output) gauss_i ("\\<^sub>\") end bundle gauss_int_output_notation begin notation (output) gauss_i ("\") end unbundle gauss_int_notation (*>*) text \ Next, we define the canonical injective homomorphism from the Gaussian integers into the complex numbers: \ primcorec gauss2complex where "Re (gauss2complex z) = of_int (ReZ z)" | "Im (gauss2complex z) = of_int (ImZ z)" declare [[coercion gauss2complex]] lemma gauss2complex_eq_iff [simp]: "gauss2complex z = gauss2complex u \ z = u" by (simp add: complex_eq_iff gauss_int_eq_iff) text \ Gaussian integers also have conjugates, just like complex numbers: \ primcorec gauss_cnj where "ReZ (gauss_cnj z) = ReZ z" | "ImZ (gauss_cnj z) = -ImZ z" text \ In the remainder of this section, we prove that Gaussian integers are a commutative ring of characteristic 0 and several other trivial algebraic properties. \ instantiation gauss_int :: comm_ring_1 begin primcorec zero_gauss_int where "ReZ zero_gauss_int = 0" | "ImZ zero_gauss_int = 0" primcorec one_gauss_int where "ReZ one_gauss_int = 1" | "ImZ one_gauss_int = 0" primcorec uminus_gauss_int where "ReZ (uminus_gauss_int x) = -ReZ x" | "ImZ (uminus_gauss_int x) = -ImZ x" primcorec plus_gauss_int where "ReZ (plus_gauss_int x y) = ReZ x + ReZ y" | "ImZ (plus_gauss_int x y) = ImZ x + ImZ y" primcorec minus_gauss_int where "ReZ (minus_gauss_int x y) = ReZ x - ReZ y" | "ImZ (minus_gauss_int x y) = ImZ x - ImZ y" primcorec times_gauss_int where "ReZ (times_gauss_int x y) = ReZ x * ReZ y - ImZ x * ImZ y" | "ImZ (times_gauss_int x y) = ReZ x * ImZ y + ImZ x * ReZ y" instance by intro_classes (auto simp: gauss_int_eq_iff algebra_simps) end lemma gauss_i_times_i [simp]: "\\<^sub>\ * \\<^sub>\ = (-1 :: gauss_int)" and gauss_cnj_i [simp]: "gauss_cnj \\<^sub>\ = -\\<^sub>\" by (simp_all add: gauss_int_eq_iff) lemma gauss_cnj_eq_0_iff [simp]: "gauss_cnj z = 0 \ z = 0" by (auto simp: gauss_int_eq_iff) lemma gauss_cnj_eq_self: "Im z = 0 \ gauss_cnj z = z" and gauss_cnj_eq_minus_self: "Re z = 0 \ gauss_cnj z = -z" by (auto simp: gauss_int_eq_iff) lemma ReZ_of_nat [simp]: "ReZ (of_nat n) = of_nat n" and ImZ_of_nat [simp]: "ImZ (of_nat n) = 0" by (induction n; simp)+ lemma ReZ_of_int [simp]: "ReZ (of_int n) = n" and ImZ_of_int [simp]: "ImZ (of_int n) = 0" by (induction n; simp)+ lemma ReZ_numeral [simp]: "ReZ (numeral n) = numeral n" and ImZ_numeral [simp]: "ImZ (numeral n) = 0" by (subst of_nat_numeral [symmetric], subst ReZ_of_nat ImZ_of_nat, simp)+ lemma gauss2complex_0 [simp]: "gauss2complex 0 = 0" and gauss2complex_1 [simp]: "gauss2complex 1 = 1" and gauss2complex_i [simp]: "gauss2complex \\<^sub>\ = \" and gauss2complex_add [simp]: "gauss2complex (x + y) = gauss2complex x + gauss2complex y" and gauss2complex_diff [simp]: "gauss2complex (x - y) = gauss2complex x - gauss2complex y" and gauss2complex_mult [simp]: "gauss2complex (x * y) = gauss2complex x * gauss2complex y" and gauss2complex_uminus [simp]: "gauss2complex (-x) = -gauss2complex x" and gauss2complex_cnj [simp]: "gauss2complex (gauss_cnj x) = cnj (gauss2complex x)" by (simp_all add: complex_eq_iff) lemma gauss2complex_of_nat [simp]: "gauss2complex (of_nat n) = of_nat n" by (simp add: complex_eq_iff) lemma gauss2complex_eq_0_iff [simp]: "gauss2complex x = 0 \ x = 0" and gauss2complex_eq_1_iff [simp]: "gauss2complex x = 1 \ x = 1" and zero_eq_gauss2complex_iff [simp]: "0 = gauss2complex x \ x = 0" and one_eq_gauss2complex_iff [simp]: "1 = gauss2complex x \ x = 1" by (simp_all add: complex_eq_iff gauss_int_eq_iff) lemma gauss_i_times_gauss_i_times [simp]: "\\<^sub>\ * (\\<^sub>\ * x) = (-x :: gauss_int)" by (subst mult.assoc [symmetric], subst gauss_i_times_i) auto lemma gauss_i_neq_0 [simp]: "\\<^sub>\ \ 0" "0 \ \\<^sub>\" and gauss_i_neq_1 [simp]: "\\<^sub>\ \ 1" "1 \ \\<^sub>\" and gauss_i_neq_of_nat [simp]: "\\<^sub>\ \ of_nat n" "of_nat n \ \\<^sub>\" and gauss_i_neq_of_int [simp]: "\\<^sub>\ \ of_int n" "of_int n \ \\<^sub>\" and gauss_i_neq_numeral [simp]: "\\<^sub>\ \ numeral m" "numeral m \ \\<^sub>\" by (auto simp: gauss_int_eq_iff) lemma gauss_cnj_0 [simp]: "gauss_cnj 0 = 0" and gauss_cnj_1 [simp]: "gauss_cnj 1 = 1" and gauss_cnj_cnj [simp]: "gauss_cnj (gauss_cnj z) = z" and gauss_cnj_uminus [simp]: "gauss_cnj (-a) = -gauss_cnj a" and gauss_cnj_add [simp]: "gauss_cnj (a + b) = gauss_cnj a + gauss_cnj b" and gauss_cnj_diff [simp]: "gauss_cnj (a - b) = gauss_cnj a - gauss_cnj b" and gauss_cnj_mult [simp]: "gauss_cnj (a * b) = gauss_cnj a * gauss_cnj b" and gauss_cnj_of_nat [simp]: "gauss_cnj (of_nat n1) = of_nat n1" and gauss_cnj_of_int [simp]: "gauss_cnj (of_int n2) = of_int n2" and gauss_cnj_numeral [simp]: "gauss_cnj (numeral n3) = numeral n3" by (simp_all add: gauss_int_eq_iff) lemma gauss_cnj_power [simp]: "gauss_cnj (a ^ n) = gauss_cnj a ^ n" by (induction n) auto lemma gauss_cnj_sum [simp]: "gauss_cnj (sum f A) = (\x\A. gauss_cnj (f x))" by (induction A rule: infinite_finite_induct) auto lemma gauss_cnj_prod [simp]: "gauss_cnj (prod f A) = (\x\A. gauss_cnj (f x))" by (induction A rule: infinite_finite_induct) auto lemma of_nat_dvd_of_nat: assumes "a dvd b" shows "of_nat a dvd (of_nat b :: 'a :: comm_semiring_1)" using assms by auto lemma of_int_dvd_imp_dvd_gauss_cnj: fixes z :: gauss_int assumes "of_int n dvd z" shows "of_int n dvd gauss_cnj z" proof - from assms obtain u where "z = of_int n * u" by blast hence "gauss_cnj z = of_int n * gauss_cnj u" by simp thus ?thesis by auto qed lemma of_nat_dvd_imp_dvd_gauss_cnj: fixes z :: gauss_int assumes "of_nat n dvd z" shows "of_nat n dvd gauss_cnj z" using of_int_dvd_imp_dvd_gauss_cnj[of "int n"] assms by simp lemma of_int_dvd_of_int_gauss_int_iff: "(of_int m :: gauss_int) dvd of_int n \ m dvd n" proof assume "of_int m dvd (of_int n :: gauss_int)" then obtain a :: gauss_int where "of_int n = of_int m * a" by blast thus "m dvd n" by (auto simp: gauss_int_eq_iff) qed auto lemma of_nat_dvd_of_nat_gauss_int_iff: "(of_nat m :: gauss_int) dvd of_nat n \ m dvd n" using of_int_dvd_of_int_gauss_int_iff[of "int m" "int n"] by simp lemma gauss_cnj_dvd: assumes "a dvd b" shows "gauss_cnj a dvd gauss_cnj b" proof - from assms obtain c where "b = a * c" by blast hence "gauss_cnj b = gauss_cnj a * gauss_cnj c" by simp thus ?thesis by auto qed lemma gauss_cnj_dvd_iff: "gauss_cnj a dvd gauss_cnj b \ a dvd b" using gauss_cnj_dvd[of a b] gauss_cnj_dvd[of "gauss_cnj a" "gauss_cnj b"] by auto lemma gauss_cnj_dvd_left_iff: "gauss_cnj a dvd b \ a dvd gauss_cnj b" by (subst gauss_cnj_dvd_iff [symmetric]) auto lemma gauss_cnj_dvd_right_iff: "a dvd gauss_cnj b \ gauss_cnj a dvd b" by (rule gauss_cnj_dvd_left_iff [symmetric]) instance gauss_int :: idom proof fix z u :: gauss_int assume "z \ 0" "u \ 0" hence "gauss2complex z * gauss2complex u \ 0" by simp also have "gauss2complex z * gauss2complex u = gauss2complex (z * u)" by simp finally show "z * u \ 0" unfolding gauss2complex_eq_0_iff . qed instance gauss_int :: ring_char_0 by intro_classes (auto intro!: injI simp: gauss_int_eq_iff) subsection \Pretty-printing\ text \ The following lemma collection provides better pretty-printing of Gaussian integers so that e.g.\ evaluation with the `value' command produces nicer results. \ lemma gauss_int_code_post [code_post]: "Gauss_Int 0 0 = 0" "Gauss_Int 0 1 = \\<^sub>\" "Gauss_Int 0 (-1) = -\\<^sub>\" "Gauss_Int 1 0 = 1" "Gauss_Int 1 1 = 1 + \\<^sub>\" "Gauss_Int 1 (-1) = 1 - \\<^sub>\" "Gauss_Int (-1) 0 = -1" "Gauss_Int (-1) 1 = -1 + \\<^sub>\" "Gauss_Int (-1) (-1) = -1 - \\<^sub>\" "Gauss_Int (numeral b) 0 = numeral b" "Gauss_Int (-numeral b) 0 = -numeral b" "Gauss_Int (numeral b) 1 = numeral b + \\<^sub>\" "Gauss_Int (-numeral b) 1 = -numeral b + \\<^sub>\" "Gauss_Int (numeral b) (-1) = numeral b - \\<^sub>\" "Gauss_Int (-numeral b) (-1) = -numeral b - \\<^sub>\" "Gauss_Int 0 (numeral b) = numeral b * \\<^sub>\" "Gauss_Int 0 (-numeral b) = -numeral b * \\<^sub>\" "Gauss_Int 1 (numeral b) = 1 + numeral b * \\<^sub>\" "Gauss_Int 1 (-numeral b) = 1 - numeral b * \\<^sub>\" "Gauss_Int (-1) (numeral b) = -1 + numeral b * \\<^sub>\" "Gauss_Int (-1) (-numeral b) = -1 - numeral b * \\<^sub>\" "Gauss_Int (numeral a) (numeral b) = numeral a + numeral b * \\<^sub>\" "Gauss_Int (numeral a) (-numeral b) = numeral a - numeral b * \\<^sub>\" "Gauss_Int (-numeral a) (numeral b) = -numeral a + numeral b * \\<^sub>\" "Gauss_Int (-numeral a) (-numeral b) = -numeral a - numeral b * \\<^sub>\" by (simp_all add: gauss_int_eq_iff) value "\\<^sub>\ ^ 3" value "2 * (3 + \\<^sub>\)" value "(2 + \\<^sub>\) * (2 - \\<^sub>\)" subsection \Norm\ text \ The square of the complex norm (or complex modulus) on the Gaussian integers gives us a norm that always returns a natural number. We will later show that this is also a Euclidean norm (in the sense of a Euclidean ring). \ definition gauss_int_norm :: "gauss_int \ nat" where "gauss_int_norm z = nat (ReZ z ^ 2 + ImZ z ^ 2)" lemma gauss_int_norm_0 [simp]: "gauss_int_norm 0 = 0" and gauss_int_norm_1 [simp]: "gauss_int_norm 1 = 1" and gauss_int_norm_i [simp]: "gauss_int_norm \\<^sub>\ = 1" and gauss_int_norm_cnj [simp]: "gauss_int_norm (gauss_cnj z) = gauss_int_norm z" and gauss_int_norm_of_nat [simp]: "gauss_int_norm (of_nat n) = n ^ 2" and gauss_int_norm_of_int [simp]: "gauss_int_norm (of_int m) = nat (m ^ 2)" and gauss_int_norm_of_numeral [simp]: "gauss_int_norm (numeral n') = numeral (Num.sqr n')" by (simp_all add: gauss_int_norm_def nat_power_eq) lemma gauss_int_norm_uminus [simp]: "gauss_int_norm (-z) = gauss_int_norm z" by (simp add: gauss_int_norm_def) lemma gauss_int_norm_eq_0_iff [simp]: "gauss_int_norm z = 0 \ z = 0" proof assume "gauss_int_norm z = 0" hence "ReZ z ^ 2 + ImZ z ^ 2 \ 0" by (simp add: gauss_int_norm_def) moreover have "ReZ z ^ 2 + ImZ z ^ 2 \ 0" by simp ultimately have "ReZ z ^ 2 + ImZ z ^ 2 = 0" by linarith thus "z = 0" by (auto simp: gauss_int_eq_iff) qed auto lemma gauss_int_norm_pos_iff [simp]: "gauss_int_norm z > 0 \ z \ 0" using gauss_int_norm_eq_0_iff[of z] by (auto intro: Nat.gr0I) lemma real_gauss_int_norm: "real (gauss_int_norm z) = norm (gauss2complex z) ^ 2" by (auto simp: cmod_def gauss_int_norm_def) lemma gauss_int_norm_mult: "gauss_int_norm (z * u) = gauss_int_norm z * gauss_int_norm u" proof - have "real (gauss_int_norm (z * u)) = real (gauss_int_norm z * gauss_int_norm u)" unfolding of_nat_mult by (simp add: real_gauss_int_norm norm_power norm_mult power_mult_distrib) thus ?thesis by (subst (asm) of_nat_eq_iff) qed lemma self_mult_gauss_cnj: "z * gauss_cnj z = of_nat (gauss_int_norm z)" by (simp add: gauss_int_norm_def gauss_int_eq_iff algebra_simps power2_eq_square) lemma gauss_cnj_mult_self: "gauss_cnj z * z = of_nat (gauss_int_norm z)" by (subst mult.commute, rule self_mult_gauss_cnj) lemma self_plus_gauss_cnj: "z + gauss_cnj z = of_int (2 * ReZ z)" and self_minus_gauss_cnj: "z - gauss_cnj z = of_int (2 * ImZ z) * \\<^sub>\" by (auto simp: gauss_int_eq_iff) lemma gauss_int_norm_dvd_mono: assumes "a dvd b" shows "gauss_int_norm a dvd gauss_int_norm b" proof - from assms obtain c where "b = a * c" by blast hence "gauss_int_norm b = gauss_int_norm (a * c)" by metis thus ?thesis by (simp add: gauss_int_norm_mult) qed text \ A Gaussian integer is a unit iff its norm is 1, and this is the case precisely for the four elements \\1\ and \\\\: \ lemma is_unit_gauss_int_iff: "x dvd 1 \ x \ {1, -1, \\<^sub>\, -\\<^sub>\ :: gauss_int}" and is_unit_gauss_int_iff': "x dvd 1 \ gauss_int_norm x = 1" proof - have "x dvd 1" if "x \ {1, -1, \\<^sub>\, -\\<^sub>\}" proof - from that have *: "x * gauss_cnj x = 1" by (auto simp: gauss_int_norm_def) show "x dvd 1" by (subst * [symmetric]) simp qed moreover have "gauss_int_norm x = 1" if "x dvd 1" using gauss_int_norm_dvd_mono[OF that] by simp moreover have "x \ {1, -1, \\<^sub>\, -\\<^sub>\}" if "gauss_int_norm x = 1" proof - from that have *: "(ReZ x)\<^sup>2 + (ImZ x)\<^sup>2 = 1" by (auto simp: gauss_int_norm_def nat_eq_iff) hence "ReZ x ^ 2 \ 1" and "ImZ x ^ 2 \ 1" using zero_le_power2[of "ImZ x"] zero_le_power2[of "ReZ x"] by linarith+ hence "\ReZ x\ \ 1" and "\ImZ x\ \ 1" by (auto simp: abs_square_le_1) hence "ReZ x \ {-1, 0, 1}" and "ImZ x \ {-1, 0, 1}" by auto thus "x \ {1, -1, \\<^sub>\, -\\<^sub>\ :: gauss_int}" using * by (auto simp: gauss_int_eq_iff) qed ultimately show "x dvd 1 \ x \ {1, -1, \\<^sub>\, -\\<^sub>\ :: gauss_int}" and "x dvd 1 \ gauss_int_norm x = 1" by blast+ qed lemma is_unit_gauss_i [simp, intro]: "(gauss_i :: gauss_int) dvd 1" by (simp add: is_unit_gauss_int_iff) lemma gauss_int_norm_eq_Suc_0_iff: "gauss_int_norm x = Suc 0 \ x dvd 1" by (simp add: is_unit_gauss_int_iff') lemma is_unit_gauss_cnj [intro]: "z dvd 1 \ gauss_cnj z dvd 1" by (simp add: is_unit_gauss_int_iff') lemma is_unit_gauss_cnj_iff [simp]: "gauss_cnj z dvd 1 \ z dvd 1" by (simp add: is_unit_gauss_int_iff') subsection \Division and normalisation\ text \ We define a rounding operation that takes a complex number and returns a Gaussian integer by rounding the real and imaginary parts separately: \ primcorec round_complex :: "complex \ gauss_int" where "ReZ (round_complex z) = round (Re z)" | "ImZ (round_complex z) = round (Im z)" text \ The distance between a rounded complex number and the original one is no more than $\frac{1}{2}\sqrt{2}$: \ lemma norm_round_complex_le: "norm (z - gauss2complex (round_complex z)) ^ 2 \ 1 / 2" proof - have "(Re z - ReZ (round_complex z)) ^ 2 \ (1 / 2) ^ 2" using of_int_round_abs_le[of "Re z"] by (subst abs_le_square_iff [symmetric]) (auto simp: abs_minus_commute) moreover have "(Im z - ImZ (round_complex z)) ^ 2 \ (1 / 2) ^ 2" using of_int_round_abs_le[of "Im z"] by (subst abs_le_square_iff [symmetric]) (auto simp: abs_minus_commute) ultimately have "(Re z - ReZ (round_complex z)) ^ 2 + (Im z - ImZ (round_complex z)) ^ 2 \ (1 / 2) ^ 2 + (1 / 2) ^ 2" by (rule add_mono) thus "norm (z - gauss2complex (round_complex z)) ^ 2 \ 1 / 2" by (simp add: cmod_def power2_eq_square) qed lemma dist_round_complex_le: "dist z (gauss2complex (round_complex z)) \ sqrt 2 / 2" proof - have "dist z (gauss2complex (round_complex z)) ^ 2 = norm (z - gauss2complex (round_complex z)) ^ 2" by (simp add: dist_norm) also have "\ \ 1 / 2" by (rule norm_round_complex_le) also have "\ = (sqrt 2 / 2) ^ 2" by (simp add: power2_eq_square) finally show ?thesis by (rule power2_le_imp_le) auto qed text \ We can now define division on Gaussian integers simply by performing the division in the complex numbers and rounding the result. This also gives us a remainder operation defined accordingly for which the norm of the remainder is always smaller than the norm of the divisor. We can also define a normalisation operation that returns a canonical representative for each association class. Since the four units of the Gaussian integers are \\1\ and \\\\, each association class (other than \0\) has four representatives, one in each quadrant. We simply define the on in the upper-right quadrant (i.e.\ the one with non-negative imaginary part and positive real part) as the canonical one. Thus, the Gaussian integers form a Euclidean ring. This gives us many things, most importantly the existence of GCDs and LCMs and unique factorisation. \ instantiation gauss_int :: algebraic_semidom begin definition divide_gauss_int :: "gauss_int \ gauss_int \ gauss_int" where "divide_gauss_int a b = round_complex (gauss2complex a / gauss2complex b)" instance proof fix a :: gauss_int show "a div 0 = 0" by (auto simp: gauss_int_eq_iff divide_gauss_int_def) next fix a b :: gauss_int assume "b \ 0" thus "a * b div b = a" by (auto simp: gauss_int_eq_iff divide_gauss_int_def) qed end instantiation gauss_int :: semidom_divide_unit_factor begin definition unit_factor_gauss_int :: "gauss_int \ gauss_int" where "unit_factor_gauss_int z = (if z = 0 then 0 else if ImZ z \ 0 \ ReZ z > 0 then 1 else if ReZ z \ 0 \ ImZ z > 0 then \\<^sub>\ else if ImZ z \ 0 \ ReZ z < 0 then -1 else -\\<^sub>\)" instance proof show "unit_factor (0 :: gauss_int) = 0" by (simp add: unit_factor_gauss_int_def) next fix z :: gauss_int assume "is_unit z" thus "unit_factor z = z" by (subst (asm) is_unit_gauss_int_iff) (auto simp: unit_factor_gauss_int_def) next fix z :: gauss_int assume z: "z \ 0" thus "is_unit (unit_factor z)" by (subst is_unit_gauss_int_iff) (auto simp: unit_factor_gauss_int_def) next fix z u :: gauss_int assume "is_unit z" hence "z \ {1, -1, \\<^sub>\, -\\<^sub>\}" by (subst (asm) is_unit_gauss_int_iff) thus "unit_factor (z * u) = z * unit_factor u" by (safe; auto simp: unit_factor_gauss_int_def gauss_int_eq_iff[of u 0]) qed end instantiation gauss_int :: normalization_semidom begin definition normalize_gauss_int :: "gauss_int \ gauss_int" where "normalize_gauss_int z = (if z = 0 then 0 else if ImZ z \ 0 \ ReZ z > 0 then z else if ReZ z \ 0 \ ImZ z > 0 then -\\<^sub>\ * z else if ImZ z \ 0 \ ReZ z < 0 then -z else \\<^sub>\ * z)" instance proof show "normalize (0 :: gauss_int) = 0" by (simp add: normalize_gauss_int_def) next fix z :: gauss_int show "unit_factor z * normalize z = z" by (auto simp: normalize_gauss_int_def unit_factor_gauss_int_def algebra_simps) qed end lemma normalize_gauss_int_of_nat [simp]: "normalize (of_nat n :: gauss_int) = of_nat n" and normalize_gauss_int_of_int [simp]: "normalize (of_int m :: gauss_int) = of_int \m\" and normalize_gauss_int_of_numeral [simp]: "normalize (numeral n' :: gauss_int) = numeral n'" by (auto simp: normalize_gauss_int_def) lemma normalize_gauss_i [simp]: "normalize \\<^sub>\ = 1" by (simp add: normalize_gauss_int_def) lemma gauss_int_norm_normalize [simp]: "gauss_int_norm (normalize x) = gauss_int_norm x" by (simp add: normalize_gauss_int_def gauss_int_norm_mult) lemma normalized_gauss_int: assumes "normalize z = z" shows "ReZ z \ 0" "ImZ z \ 0" using assms by (cases "ReZ z" "0 :: int" rule: linorder_cases; cases "ImZ z" "0 :: int" rule: linorder_cases; simp add: normalize_gauss_int_def gauss_int_eq_iff)+ lemma normalized_gauss_int': assumes "normalize z = z" "z \ 0" shows "ReZ z > 0" "ImZ z \ 0" using assms by (cases "ReZ z" "0 :: int" rule: linorder_cases; cases "ImZ z" "0 :: int" rule: linorder_cases; simp add: normalize_gauss_int_def gauss_int_eq_iff)+ lemma normalized_gauss_int_iff: "normalize z = z \ z = 0 \ ReZ z > 0 \ ImZ z \ 0" by (cases "ReZ z" "0 :: int" rule: linorder_cases; cases "ImZ z" "0 :: int" rule: linorder_cases; simp add: normalize_gauss_int_def gauss_int_eq_iff)+ instantiation gauss_int :: idom_modulo begin definition modulo_gauss_int :: "gauss_int \ gauss_int \ gauss_int" where "modulo_gauss_int a b = a - a div b * b" instance proof fix a b :: gauss_int show "a div b * b + a mod b = a" by (simp add: modulo_gauss_int_def) qed end lemma gauss_int_norm_mod_less_aux: assumes [simp]: "b \ 0" shows "2 * gauss_int_norm (a mod b) \ gauss_int_norm b" proof - define a' b' where "a' = gauss2complex a" and "b' = gauss2complex b" have [simp]: "b' \ 0" by (simp add: b'_def) have "gauss_int_norm (a mod b) = norm (gauss2complex (a - round_complex (a' / b') * b)) ^ 2" unfolding modulo_gauss_int_def by (subst real_gauss_int_norm [symmetric]) (auto simp add: divide_gauss_int_def a'_def b'_def) also have "gauss2complex (a - round_complex (a' / b') * b) = a' - gauss2complex (round_complex (a' / b')) * b'" by (simp add: a'_def b'_def) also have "\ = (a' / b' - gauss2complex (round_complex (a' / b'))) * b'" by (simp add: field_simps) also have "norm \ ^ 2 = norm (a' / b' - gauss2complex (round_complex (a' / b'))) ^ 2 * norm b' ^ 2" by (simp add: norm_mult power_mult_distrib) also have "\ \ 1 / 2 * norm b' ^ 2" by (intro mult_right_mono norm_round_complex_le) auto also have "norm b' ^ 2 = gauss_int_norm b" by (simp add: b'_def real_gauss_int_norm) finally show ?thesis by linarith qed lemma gauss_int_norm_mod_less: assumes [simp]: "b \ 0" shows "gauss_int_norm (a mod b) < gauss_int_norm b" proof - have "gauss_int_norm b > 0" by simp thus "gauss_int_norm (a mod b) < gauss_int_norm b" using gauss_int_norm_mod_less_aux[OF assms, of a] by presburger qed lemma gauss_int_norm_dvd_imp_le: assumes "b \ 0" shows "gauss_int_norm a \ gauss_int_norm (a * b)" proof (cases "a = 0") case False thus ?thesis using assms by (intro dvd_imp_le gauss_int_norm_dvd_mono) auto qed auto instantiation gauss_int :: euclidean_ring begin definition euclidean_size_gauss_int :: "gauss_int \ nat" where [simp]: "euclidean_size_gauss_int = gauss_int_norm" instance proof show "euclidean_size (0 :: gauss_int) = 0" by simp next fix a b :: gauss_int assume [simp]: "b \ 0" show "euclidean_size (a mod b) < euclidean_size b" using gauss_int_norm_mod_less[of b a] by simp show "euclidean_size a \ euclidean_size (a * b)" by (simp add: gauss_int_norm_dvd_imp_le) qed end instance gauss_int :: normalization_euclidean_semiring .. instantiation gauss_int :: euclidean_ring_gcd begin definition gcd_gauss_int :: "gauss_int \ gauss_int \ gauss_int" where "gcd_gauss_int \ normalization_euclidean_semiring_class.gcd" definition lcm_gauss_int :: "gauss_int \ gauss_int \ gauss_int" where "lcm_gauss_int \ normalization_euclidean_semiring_class.lcm" definition Gcd_gauss_int :: "gauss_int set \ gauss_int" where "Gcd_gauss_int \ normalization_euclidean_semiring_class.Gcd" definition Lcm_gauss_int :: "gauss_int set \ gauss_int" where "Lcm_gauss_int \ normalization_euclidean_semiring_class.Lcm" instance by intro_classes (simp_all add: gcd_gauss_int_def lcm_gauss_int_def Gcd_gauss_int_def Lcm_gauss_int_def) end lemma multiplicity_gauss_cnj: "multiplicity (gauss_cnj a) (gauss_cnj b) = multiplicity a b" unfolding multiplicity_def gauss_cnj_power [symmetric] gauss_cnj_dvd_iff .. lemma multiplicity_gauss_int_of_nat: "multiplicity (of_nat a) (of_nat b :: gauss_int) = multiplicity a b" unfolding multiplicity_def of_nat_power [symmetric] of_nat_dvd_of_nat_gauss_int_iff .. lemma gauss_int_dvd_same_norm_imp_associated: assumes "z1 dvd z2" "gauss_int_norm z1 = gauss_int_norm z2" shows "normalize z1 = normalize z2" proof (cases "z1 = 0") case [simp]: False from assms(1) obtain u where u: "z2 = z1 * u" by blast from assms have "gauss_int_norm u = 1" by (auto simp: gauss_int_norm_mult u) hence "is_unit u" by (simp add: is_unit_gauss_int_iff') with u show ?thesis by simp qed (use assms in auto) lemma gcd_of_int_gauss_int: "gcd (of_int a :: gauss_int) (of_int b) = of_int (gcd a b)" proof (induction "nat \b\" arbitrary: a b rule: less_induct) case (less b a) show ?case proof (cases "b = 0") case False have "of_int (gcd a b) = (of_int (gcd b (a mod b)) :: gauss_int)" by (subst gcd_red_int) auto also have "\ = gcd (of_int b) (of_int (a mod b))" using False by (intro less [symmetric]) (auto intro!: abs_mod_less) also have "a mod b = (a - a div b * b)" by (simp add: minus_div_mult_eq_mod) also have "of_int \ = of_int (-(a div b)) * of_int b + (of_int a :: gauss_int)" by (simp add: algebra_simps) also have "gcd (of_int b) \ = gcd (of_int b) (of_int a)" by (rule gcd_add_mult) finally show ?thesis by (simp add: gcd.commute) qed auto qed lemma coprime_of_int_gauss_int: "coprime (of_int a :: gauss_int) (of_int b) = coprime a b" unfolding coprime_iff_gcd_eq_1 gcd_of_int_gauss_int by auto lemma gcd_of_nat_gauss_int: "gcd (of_nat a :: gauss_int) (of_nat b) = of_nat (gcd a b)" using gcd_of_int_gauss_int[of "int a" "int b"] by simp lemma coprime_of_nat_gauss_int: "coprime (of_nat a :: gauss_int) (of_nat b) = coprime a b" unfolding coprime_iff_gcd_eq_1 gcd_of_nat_gauss_int by auto lemma gauss_cnj_dvd_self_iff: "gauss_cnj z dvd z \ ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" proof assume "gauss_cnj z dvd z" hence "normalize (gauss_cnj z) = normalize z" by (rule gauss_int_dvd_same_norm_imp_associated) auto then obtain u :: gauss_int where "is_unit u" and u: "gauss_cnj z = u * z" using associatedE1 by blast hence "u \ {1, -1, \\<^sub>\, -\\<^sub>\}" by (simp add: is_unit_gauss_int_iff) thus "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" proof (elim insertE emptyE) assume [simp]: "u = \\<^sub>\" have "ReZ z = ReZ (gauss_cnj z)" by simp also have "gauss_cnj z = \\<^sub>\ * z" using u by simp also have "ReZ \ = -ImZ z" by simp finally show "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" by auto next assume [simp]: "u = -\\<^sub>\" have "ReZ z = ReZ (gauss_cnj z)" by simp also have "gauss_cnj z = -\\<^sub>\ * z" using u by simp also have "ReZ \ = ImZ z" by simp finally show "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" by auto next assume [simp]: "u = 1" have "ImZ z = -ImZ (gauss_cnj z)" by simp also have "gauss_cnj z = z" using u by simp finally show "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" by auto next assume [simp]: "u = -1" have "ReZ z = ReZ (gauss_cnj z)" by simp also have "gauss_cnj z = -z" using u by simp also have "ReZ \ = -ReZ z" by simp finally show "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" by auto qed next assume "ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" thus "gauss_cnj z dvd z" proof safe assume "\ReZ z\ = \ImZ z\" then obtain u :: int where "is_unit u" and u: "ImZ z = u * ReZ z" using associatedE2[of "ReZ z" "ImZ z"] by auto from \is_unit u\ have "u \ {1, -1}" by auto hence "z = gauss_cnj z * (of_int u * \\<^sub>\)" using u by (auto simp: gauss_int_eq_iff) thus ?thesis by (metis dvd_triv_left) qed (auto simp: gauss_cnj_eq_self gauss_cnj_eq_minus_self) qed lemma self_dvd_gauss_cnj_iff: "z dvd gauss_cnj z \ ReZ z = 0 \ ImZ z = 0 \ \ReZ z\ = \ImZ z\" using gauss_cnj_dvd_self_iff[of z] by (subst (asm) gauss_cnj_dvd_left_iff) auto subsection \Prime elements\ text \ Next, we analyse what the prime elements of the Gaussian integers are. First, note that according to the conventions of Isabelle's computational algebra library, a prime element is called a prime iff it is also normalised, i.e.\ in our case it lies in the upper right quadrant. As a first fact, we can show that a Gaussian integer whose norm is \\\-prime must be $\mathbb{Z}[i]$-prime: \ lemma prime_gauss_int_norm_imp_prime_elem: assumes "prime (gauss_int_norm q)" shows "prime_elem q" proof - have "irreducible q" proof (rule irreducibleI) fix a b assume "q = a * b" hence "gauss_int_norm q = gauss_int_norm a * gauss_int_norm b" by (simp_all add: gauss_int_norm_mult) thus "is_unit a \ is_unit b" using assms by (auto dest!: prime_product simp: gauss_int_norm_eq_Suc_0_iff) qed (use assms in \auto simp: is_unit_gauss_int_iff'\) thus "prime_elem q" using irreducible_imp_prime_elem_gcd by blast qed text \ Also, a conjugate is a prime element iff the original element is a prime element: \ lemma prime_elem_gauss_cnj [intro]: "prime_elem z \ prime_elem (gauss_cnj z)" by (auto simp: prime_elem_def gauss_cnj_dvd_left_iff) lemma prime_elem_gauss_cnj_iff [simp]: "prime_elem (gauss_cnj z) \ prime_elem z" using prime_elem_gauss_cnj[of z] prime_elem_gauss_cnj[of "gauss_cnj z"] by auto subsubsection \The factorisation of 2\ text \ 2 factors as $-i (1 + i)^2$ in the Gaussian integers, where $-i$ is a unit and $1 + i$ is prime. \ lemma gauss_int_2_eq: "2 = -\\<^sub>\ * (1 + \\<^sub>\) ^ 2" by (simp add: gauss_int_eq_iff power2_eq_square) lemma prime_elem_one_plus_i_gauss_int: "prime_elem (1 + \\<^sub>\)" by (rule prime_gauss_int_norm_imp_prime_elem) (auto simp: gauss_int_norm_def) lemma prime_one_plus_i_gauss_int: "prime (1 + \\<^sub>\)" by (simp add: prime_def prime_elem_one_plus_i_gauss_int gauss_int_eq_iff normalize_gauss_int_def) lemma prime_factorization_2_gauss_int: "prime_factorization (2 :: gauss_int) = {#1 + \\<^sub>\, 1 + \\<^sub>\#}" proof - have "prime_factorization (2 :: gauss_int) = (prime_factorization (prod_mset {#1 + gauss_i, 1 + gauss_i#}))" by (subst prime_factorization_unique) (auto simp: gauss_int_eq_iff normalize_gauss_int_def) also have "prime_factorization (prod_mset {#1 + gauss_i, 1 + gauss_i#}) = {#1 + gauss_i, 1 + gauss_i#}" using prime_one_plus_i_gauss_int by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed subsubsection \Inert primes\ text \ Any \\\-prime congruent 3 modulo 4 is also a Gaussian prime. These primes are called \<^emph>\inert\, because they do not decompose when moving from \\\ to $\mathbb{Z}[i]$. \ lemma gauss_int_norm_not_3_mod_4: "[gauss_int_norm z \ 3] (mod 4)" proof - have A: "ReZ z mod 4 \ {0..3}" "ImZ z mod 4 \ {0..3}" by auto have B: "{0..3} = {0, 1, 2, 3 :: int}" by auto have "[ReZ z ^ 2 + ImZ z ^ 2 = (ReZ z mod 4) ^ 2 + (ImZ z mod 4) ^ 2] (mod 4)" by (intro cong_add cong_pow) (auto simp: cong_def) moreover have "((ReZ z mod 4) ^ 2 + (ImZ z mod 4) ^ 2) mod 4 \ 3 mod 4" using A unfolding B by auto ultimately have "[ReZ z ^ 2 + ImZ z ^ 2 \ 3] (mod 4)" unfolding cong_def by metis hence "[int (nat (ReZ z ^ 2 + ImZ z ^ 2)) \ int 3] (mod (int 4))" by simp thus ?thesis unfolding gauss_int_norm_def by (subst (asm) cong_int_iff) qed lemma prime_elem_gauss_int_of_nat: fixes n :: nat assumes prime: "prime n" and "[n = 3] (mod 4)" shows "prime_elem (of_nat n :: gauss_int)" proof (intro irreducible_imp_prime_elem irreducibleI) from assms show "of_nat n \ (0 :: gauss_int)" by (auto simp: gauss_int_eq_iff) next show "\is_unit (of_nat n :: gauss_int)" using assms by (subst is_unit_gauss_int_iff) (auto simp: gauss_int_eq_iff) next fix a b :: gauss_int assume *: "of_nat n = a * b" hence "gauss_int_norm (a * b) = gauss_int_norm (of_nat n)" by metis hence *: "gauss_int_norm a * gauss_int_norm b = n ^ 2" by (simp add: gauss_int_norm_mult power2_eq_square flip: nat_mult_distrib) from prime_power_mult_nat[OF prime this] obtain i j :: nat where ij: "gauss_int_norm a = n ^ i" "gauss_int_norm b = n ^ j" by blast have "i + j = 2" proof - have "n ^ (i + j) = n ^ 2" using ij * by (simp add: power_add) from prime_power_inj[OF prime this] show ?thesis by simp qed hence "i = 0 \ j = 2 \ i = 1 \ j = 1 \ i = 2 \ j = 0" by auto thus "is_unit a \ is_unit b" proof (elim disjE) assume "i = 1 \ j = 1" with ij have "gauss_int_norm a = n" by auto hence "[gauss_int_norm a = n] (mod 4)" by simp also have "[n = 3] (mod 4)" by fact finally have "[gauss_int_norm a = 3] (mod 4)" . moreover have "[gauss_int_norm a \ 3] (mod 4)" by (rule gauss_int_norm_not_3_mod_4) ultimately show ?thesis by contradiction qed (use ij in \auto simp: is_unit_gauss_int_iff'\) qed theorem prime_gauss_int_of_nat: fixes n :: nat assumes prime: "prime n" and "[n = 3] (mod 4)" shows "prime (of_nat n :: gauss_int)" using prime_elem_gauss_int_of_nat[OF assms] unfolding prime_def by simp subsubsection \Non-inert primes\ text \ Any \\\-prime congruent 1 modulo 4 factors into two conjugate Gaussian primes. \ lemma minimal_QuadRes_neg1: assumes "QuadRes n (-1)" "n > 1" "odd n" obtains x :: nat where "x \ (n - 1) div 2" and "[x ^ 2 + 1 = 0] (mod n)" proof - from \QuadRes n (-1)\ obtain x where "[x ^ 2 = (-1)] (mod (int n))" by (auto simp: QuadRes_def) hence "[x ^ 2 + 1 = -1 + 1] (mod (int n))" by (intro cong_add) auto also have "x ^ 2 + 1 = int (nat \x\ ^ 2 + 1)" by simp finally have "[int (nat \x\ ^ 2 + 1) = int 0] (mod (int n))" by simp hence "[nat \x\ ^ 2 + 1 = 0] (mod n)" by (subst (asm) cong_int_iff) define x' where "x' = (if nat \x\ mod n \ (n - 1) div 2 then nat \x\ mod n else n - (nat \x\ mod n))" have x'_quadres: "[x' ^ 2 + 1 = 0] (mod n)" proof (cases "nat \x\ mod n \ (n - 1) div 2") case True hence "[x' ^ 2 + 1 = (nat \x\ mod n) ^ 2 + 1] (mod n)" by (simp add: x'_def) also have "[(nat \x\ mod n) ^ 2 + 1 = nat \x\ ^ 2 + 1] (mod n)" by (intro cong_add cong_pow) (auto simp: cong_def) also have "[nat \x\ ^ 2 + 1 = 0] (mod n)" by fact finally show ?thesis . next case False hence "[int (x' ^ 2 + 1) = (int n - int (nat \x\ mod n)) ^ 2 + 1] (mod int n)" using \n > 1\ by (simp add: x'_def of_nat_diff add_ac) also have "[(int n - int (nat \x\ mod n)) ^ 2 + 1 = (0 - int (nat \x\ mod n)) ^ 2 + 1] (mod int n)" by (intro cong_add cong_pow) (auto simp: cong_def) also have "[(0 - int (nat \x\ mod n)) ^ 2 + 1 = int ((nat \x\ mod n) ^ 2 + 1)] (mod (int n))" by (simp add: add_ac) finally have "[x' ^ 2 + 1 = (nat \x\ mod n)\<^sup>2 + 1] (mod n)" by (subst (asm) cong_int_iff) also have "[(nat \x\ mod n)\<^sup>2 + 1 = nat \x\ ^ 2 + 1] (mod n)" by (intro cong_add cong_pow) (auto simp: cong_def) also have "[nat \x\ ^ 2 + 1 = 0] (mod n)" by fact finally show ?thesis . qed moreover have x'_le: "x' \ (n - 1) div 2" using \odd n\ by (auto elim!: oddE simp: x'_def) ultimately show ?thesis by (intro that[of x']) qed text \ Let \p\ be some prime number that is congruent 1 modulo 4. \ locale noninert_gauss_int_prime = fixes p :: nat assumes prime_p: "prime p" and cong_1_p: "[p = 1] (mod 4)" begin lemma p_gt_2: "p > 2" and odd_p: "odd p" proof - from prime_p and cong_1_p have "p > 1" "p \ 2" by (auto simp: prime_gt_Suc_0_nat cong_def) thus "p > 2" by auto with prime_p show "odd p" using primes_dvd_imp_eq two_is_prime_nat by blast qed text \ -1 is a quadratic residue modulo \p\, so there exists some \x\ such that $x^2 + 1$ is divisible by \p\. Moreover, we can choose \x\ such that it is positive and no greater than $\frac{1}{2}(p-1)$: \ lemma minimal_QuadRes_neg1: obtains x where "x > 0" "x \ (p - 1) div 2" "[x ^ 2 + 1 = 0] (mod p)" proof - have "[Legendre (-1) (int p) = (- 1) ^ ((p - 1) div 2)] (mod (int p))" using prime_p p_gt_2 by (intro euler_criterion) auto also have "[p - 1 = 1 - 1] (mod 4)" using p_gt_2 by (intro cong_diff_nat cong_refl) (use cong_1_p in auto) hence "2 * 2 dvd p - 1" by (simp add: cong_0_iff) hence "even ((p - 1) div 2)" using dvd_mult_imp_div by blast hence "(-1) ^ ((p - 1) div 2) = (1 :: int)" by simp finally have "Legendre (-1) (int p) mod p = 1" using p_gt_2 by (auto simp: cong_def) hence "Legendre (-1) (int p) = 1" using p_gt_2 by (auto simp: Legendre_def cong_def zmod_minus1 split: if_splits) hence "QuadRes p (-1)" by (simp add: Legendre_def split: if_splits) from minimal_QuadRes_neg1[OF this] p_gt_2 odd_p obtain x where x: "x \ (p - 1) div 2" "[x ^ 2 + 1 = 0] (mod p)" by auto have "x > 0" using x p_gt_2 by (auto intro!: Nat.gr0I simp: cong_def) from x and this show ?thesis by (intro that[of x]) auto qed text \ We can show from this that \p\ is not prime as a Gaussian integer. \ lemma not_prime: "\prime_elem (of_nat p :: gauss_int)" proof assume prime: "prime_elem (of_nat p :: gauss_int)" obtain x where x: "x > 0" "x \ (p - 1) div 2" "[x ^ 2 + 1 = 0] (mod p)" using minimal_QuadRes_neg1 . have "of_nat p dvd (of_nat (x ^ 2 + 1) :: gauss_int)" using x by (intro of_nat_dvd_of_nat) (auto simp: cong_0_iff) also have eq: "of_nat (x ^ 2 + 1) = ((of_nat x + \\<^sub>\) * (of_nat x - \\<^sub>\) :: gauss_int)" using \x > 0\ by (simp add: algebra_simps gauss_int_eq_iff power2_eq_square of_nat_diff) finally have "of_nat p dvd ((of_nat x + \\<^sub>\) * (of_nat x - \\<^sub>\) :: gauss_int)" . from prime and this have "of_nat p dvd (of_nat x + \\<^sub>\ :: gauss_int) \ of_nat p dvd (of_nat x - \\<^sub>\ :: gauss_int)" by (rule prime_elem_dvd_multD) hence dvd: "of_nat p dvd (of_nat x + \\<^sub>\ :: gauss_int)" "of_nat p dvd (of_nat x - \\<^sub>\ :: gauss_int)" by (auto dest: of_nat_dvd_imp_dvd_gauss_cnj) have "of_nat (p ^ 2) = (of_nat p * of_nat p :: gauss_int)" by (simp add: power2_eq_square) also from dvd have "\ dvd ((of_nat x + \\<^sub>\) * (of_nat x - \\<^sub>\))" by (intro mult_dvd_mono) also have "\ = of_nat (x ^ 2 + 1)" by (rule eq [symmetric]) finally have "p ^ 2 dvd (x ^ 2 + 1)" by (subst (asm) of_nat_dvd_of_nat_gauss_int_iff) hence "p ^ 2 \ x ^ 2 + 1" by (intro dvd_imp_le) auto moreover have "p ^ 2 > x ^ 2 + 1" proof - have "x ^ 2 + 1 \ ((p - 1) div 2) ^ 2 + 1" using x by (intro add_mono power_mono) auto also have "\ \ (p - 1) ^ 2 + 1" by auto also have "(p - 1) * (p - 1) < (p - 1) * (p + 1)" using p_gt_2 by (intro mult_strict_left_mono) auto hence "(p - 1) ^ 2 + 1 < p ^ 2" by (simp add: algebra_simps power2_eq_square) finally show ?thesis . qed ultimately show False by linarith qed text \ Any prime factor of \p\ in the Gaussian integers must have norm \p\. \ lemma norm_prime_divisor: fixes q :: gauss_int assumes q: "prime_elem q" "q dvd of_nat p" shows "gauss_int_norm q = p" proof - from assms obtain r where r: "of_nat p = q * r" by auto have "p ^ 2 = gauss_int_norm (of_nat p)" by simp also have "\ = gauss_int_norm q * gauss_int_norm r" by (auto simp: r gauss_int_norm_mult) finally have *: "gauss_int_norm q * gauss_int_norm r = p ^ 2" by simp hence "\i j. gauss_int_norm q = p ^ i \ gauss_int_norm r = p ^ j" using prime_p by (intro prime_power_mult_nat) then obtain i j where ij: "gauss_int_norm q = p ^ i" "gauss_int_norm r = p ^ j" by blast have ij_eq_2: "i + j = 2" proof - from * have "p ^ (i + j) = p ^ 2" by (simp add: power_add ij) thus ?thesis using p_gt_2 by (subst (asm) power_inject_exp) auto qed hence "i = 0 \ j = 2 \ i = 1 \ j = 1 \ i = 2 \ j = 0" by auto hence "i = 1" proof (elim disjE) assume "i = 2 \ j = 0" hence "is_unit r" using ij by (simp add: gauss_int_norm_eq_Suc_0_iff) hence "prime_elem (of_nat p :: gauss_int)" using \prime_elem q\ by (simp add: prime_elem_mult_unit_left r mult.commute[of _ r]) with not_prime show "i = 1" by contradiction qed (use q ij in \auto simp: gauss_int_norm_eq_Suc_0_iff\) thus ?thesis using ij by simp qed text \ We now show two lemmas that characterise the two prime factors of \p\ in the Gaussian integers: they are two conjugates $x\pm iy$ for positive integers \x\ and \y\ such that $x^2 + y^2 = p$. \ lemma prime_divisor_exists: obtains q where "prime q" "prime_elem (gauss_cnj q)" "ReZ q > 0" "ImZ q > 0" "of_nat p = q * gauss_cnj q" "gauss_int_norm q = p" proof - have "\q::gauss_int. q dvd of_nat p \ prime q" by (rule prime_divisor_exists) (use prime_p in \auto simp: is_unit_gauss_int_iff'\) then obtain q :: gauss_int where q: "prime q" "q dvd of_nat p" by blast from \prime q\ have [simp]: "q \ 0" by auto have "normalize q = q" using q by simp hence q_signs: "ReZ q > 0" "ImZ q \ 0" by (subst (asm) normalized_gauss_int_iff; simp)+ from q have "gauss_int_norm q = p" using norm_prime_divisor[of q] by simp moreover from this have "gauss_int_norm (gauss_cnj q) = p" by simp hence "prime_elem (gauss_cnj q)" using prime_p by (intro prime_gauss_int_norm_imp_prime_elem) auto moreover have "of_nat p = q * gauss_cnj q" using \gauss_int_norm q = p\ by (simp add: self_mult_gauss_cnj) moreover have "ImZ q \ 0" proof assume [simp]: "ImZ q = 0" define m where "m = nat (ReZ q)" have [simp]: "q = of_nat m" using q_signs by (auto simp: gauss_int_eq_iff m_def) with q have "m dvd p" by (simp add: of_nat_dvd_of_nat_gauss_int_iff) with prime_p have "m = 1 \ m = p" using prime_nat_iff by blast with q show False using not_prime by auto qed with q_signs have "ImZ q > 0" by simp ultimately show ?thesis using q q_signs by (intro that[of q]) qed theorem prime_factorization: obtains q1 q2 where "prime q1" "prime q2" "prime_factorization (of_nat p) = {#q1, q2#}" "gauss_int_norm q1 = p" "gauss_int_norm q2 = p" "q2 = \\<^sub>\ * gauss_cnj q1" "ReZ q1 > 0" "ImZ q1 > 0" "ReZ q1 > 0" "ImZ q2 > 0" proof - obtain q where q: "prime q" "prime_elem (gauss_cnj q)" "ReZ q > 0" "ImZ q > 0" "of_nat p = q * gauss_cnj q" "gauss_int_norm q = p" using prime_divisor_exists by metis from \prime q\ have [simp]: "q \ 0" by auto define q' where "q' = normalize (gauss_cnj q)" have "prime_factorization (of_nat p) = prime_factorization (prod_mset {#q, q'#})" by (subst prime_factorization_unique) (auto simp: q q'_def) also have "\ = {#q, q'#}" using q by (subst prime_factorization_prod_mset_primes) (auto simp: q'_def) finally have "prime_factorization (of_nat p) = {#q, q'#}" . moreover have "q' = \\<^sub>\ * gauss_cnj q" using q by (auto simp: normalize_gauss_int_def q'_def) moreover have "prime q'" using q by (auto simp: q'_def) ultimately show ?thesis using q by (intro that[of q q']) (auto simp: q'_def gauss_int_norm_mult) qed end text \ In particular, a consequence of this is that any prime congruent 1 modulo 4 can be written as a sum of squares of positive integers. \ lemma prime_cong_1_mod_4_gauss_int_norm_exists: fixes p :: nat assumes "prime p" "[p = 1] (mod 4)" shows "\z. gauss_int_norm z = p \ ReZ z > 0 \ ImZ z > 0" proof - from assms interpret noninert_gauss_int_prime p by unfold_locales from prime_divisor_exists obtain q where q: "prime q" "of_nat p = q * gauss_cnj q" "ReZ q > 0" "ImZ q > 0" "gauss_int_norm q = p" by metis have "p = gauss_int_norm q" using q by simp thus ?thesis using q by blast qed subsubsection \Full classification of Gaussian primes\ text \ Any prime in the ring of Gaussian integers is of the form \<^item> \1 + \\<^sub>\\ \<^item> \p\ where \p \ \\ is prime in \\\ and congruent 1 modulo 4 \<^item> $x + iy$ where $x,y$ are positive integers and $x^2 + y^2$ is a prime congruent 3 modulo 4 or an associated element of one of these. \ theorem gauss_int_prime_classification: fixes x :: gauss_int assumes "prime x" obtains (one_plus_i) "x = 1 + \\<^sub>\" | (cong_3_mod_4) p where "x = of_nat p" "prime p" "[p = 3] (mod 4)" | (cong_1_mod_4) "prime (gauss_int_norm x)" "[gauss_int_norm x = 1] (mod 4)" "ReZ x > 0" "ImZ x > 0" "ReZ x \ ImZ x" proof - define N where "N = gauss_int_norm x" have "x dvd x * gauss_cnj x" by simp also have "\ = of_nat (gauss_int_norm x)" by (simp add: self_mult_gauss_cnj) finally have "x \ prime_factors (of_nat N)" using assms by (auto simp: in_prime_factors_iff N_def) also have "N = prod_mset (prime_factorization N)" using assms unfolding N_def by (subst prod_mset_prime_factorization_nat) auto also have "(of_nat \ :: gauss_int) = prod_mset (image_mset of_nat (prime_factorization N))" by (subst of_nat_prod_mset) auto also have "prime_factors \ = (\p\prime_factors N. prime_factors (of_nat p))" by (subst prime_factorization_prod_mset) auto finally obtain p where p: "p \ prime_factors N" "x \ prime_factors (of_nat p)" by auto have "prime p" using p by auto hence "\(2 * 2) dvd p" using product_dvd_irreducibleD[of p 2 2] by (auto simp flip: prime_elem_iff_irreducible) hence "[p \ 0] (mod 4)" using p by (auto simp: cong_0_iff in_prime_factors_iff) hence "p mod 4 \ {1,2,3}" by (auto simp: cong_def) thus ?thesis proof (elim singletonE insertE) assume "p mod 4 = 2" hence "p mod 4 mod 2 = 0" by simp hence "p mod 2 = 0" by (simp add: mod_mod_cancel) with \prime p\ have [simp]: "p = 2" using prime_prime_factor two_is_prime_nat by blast have "prime_factors (of_nat p) = {1 + \\<^sub>\ :: gauss_int}" by (simp add: prime_factorization_2_gauss_int) with p show ?thesis using that(1) by auto next assume *: "p mod 4 = 3" hence "prime_factors (of_nat p) = {of_nat p :: gauss_int}" using prime_gauss_int_of_nat[of p] \prime p\ by (subst prime_factorization_prime) (auto simp: cong_def) with p show ?thesis using that(2)[of p] * by (auto simp: cong_def) next assume *: "p mod 4 = 1" then interpret noninert_gauss_int_prime p by unfold_locales (use \prime p\ in \auto simp: cong_def\) obtain q1 q2 :: gauss_int where q12: "prime q1" "prime q2" "prime_factorization (of_nat p) = {#q1, q2#}" "gauss_int_norm q1 = p" "gauss_int_norm q2 = p" "q2 = \\<^sub>\ * gauss_cnj q1" "ReZ q1 > 0" "ImZ q1 > 0" "ReZ q1 > 0" "ImZ q2 > 0" using prime_factorization by metis from p q12 have "x = q1 \ x = q2" by auto with q12 have **: "gauss_int_norm x = p" "ReZ x > 0" "ImZ x > 0" by auto have "ReZ x \ ImZ x" proof assume "ReZ x = ImZ x" hence "even (gauss_int_norm x)" by (auto simp: gauss_int_norm_def nat_mult_distrib) hence "even p" using \gauss_int_norm x = p\ by simp with \p mod 4 = 1\ show False by presburger qed thus ?thesis using that(3) \prime p\ * ** by (simp add: cong_def) qed qed lemma prime_gauss_int_norm_squareD: fixes z :: gauss_int assumes "prime z" "gauss_int_norm z = p ^ 2" shows "prime p \ z = of_nat p" using assms(1) proof (cases rule: gauss_int_prime_classification) case one_plus_i have "prime (2 :: nat)" by simp also from one_plus_i have "2 = p ^ 2" using assms(2) by (auto simp: gauss_int_norm_def) finally show ?thesis by (simp add: prime_power_iff) next case (cong_3_mod_4 p) thus ?thesis using assms by auto next case cong_1_mod_4 with assms show ?thesis by (auto simp: prime_power_iff) qed lemma gauss_int_norm_eq_prime_squareD: assumes "prime p" and "[p = 3] (mod 4)" and "gauss_int_norm z = p ^ 2" shows "normalize z = of_nat p" and "prime_elem z" proof - have "\q::gauss_int. q dvd z \ prime q" by (rule prime_divisor_exists) (use assms in \auto simp: is_unit_gauss_int_iff'\) then obtain q :: gauss_int where q: "q dvd z" "prime q" by blast have "gauss_int_norm q dvd gauss_int_norm z" by (rule gauss_int_norm_dvd_mono) fact also have "\ = p ^ 2" by fact finally obtain i where i: "i \ 2" "gauss_int_norm q = p ^ i" by (subst (asm) divides_primepow_nat) (use assms q in auto) from i assms q have "i \ 0" by (auto intro!: Nat.gr0I simp: gauss_int_norm_eq_Suc_0_iff) moreover from i assms q have "i \ 1" using gauss_int_norm_not_3_mod_4[of q] by auto ultimately have "i = 2" using i by auto with i have "gauss_int_norm q = p ^ 2" by auto hence [simp]: "q = of_nat p" using prime_gauss_int_norm_squareD[of q p] q by auto have "normalize (of_nat p) = normalize z" using q assms by (intro gauss_int_dvd_same_norm_imp_associated) auto thus *: "normalize z = of_nat p" by simp have "prime (normalize z)" using prime_gauss_int_of_nat[of p] assms by (subst *) auto thus "prime_elem z" by simp qed text \ The following can be used as a primality test for Gaussian integers. It effectively reduces checking the primality of a Gaussian integer to checking the primality of an integer. A Gaussian integer is prime if either its norm is either \\\-prime or the square of a \\\-prime that is congruent 3 modulo 4. \ lemma prime_elem_gauss_int_iff: fixes z :: gauss_int defines "n \ gauss_int_norm z" shows "prime_elem z \ prime n \ (\p. n = p ^ 2 \ prime p \ [p = 3] (mod 4))" proof assume "prime n \ (\p. n = p ^ 2 \ prime p \ [p = 3] (mod 4))" thus "prime_elem z" by (auto intro: gauss_int_norm_eq_prime_squareD(2) prime_gauss_int_norm_imp_prime_elem simp: n_def) next assume "prime_elem z" hence "prime (normalize z)" by simp thus "prime n \ (\p. n = p ^ 2 \ prime p \ [p = 3] (mod 4))" proof (cases rule: gauss_int_prime_classification) case one_plus_i have "n = gauss_int_norm (normalize z)" by (simp add: n_def) also have "normalize z = 1 + \\<^sub>\" by fact also have "gauss_int_norm \ = 2" by (simp add: gauss_int_norm_def) finally show ?thesis by simp next case (cong_3_mod_4 p) have "n = gauss_int_norm (normalize z)" by (simp add: n_def) also have "normalize z = of_nat p" by fact also have "gauss_int_norm \ = p ^ 2" by simp finally show ?thesis using cong_3_mod_4 by simp next case cong_1_mod_4 thus ?thesis by (simp add: n_def) qed qed subsubsection \Multiplicities of primes\ text \ In this section, we will show some results connecting the multiplicity of a Gaussian prime \p\ in a Gaussian integer \z\ to the \\\-multiplicity of the norm of \p\ in the norm of \z\. \ text \ The multiplicity of the Gaussian prime \<^term>\1 + \\<^sub>\\ in an integer \c\ is simply twice the \\\-multiplicity of 2 in \c\: \ lemma multiplicity_prime_1_plus_i_aux: "multiplicity (1 + \\<^sub>\) (of_nat c) = 2 * multiplicity 2 c" proof (cases "c = 0") case [simp]: False have "2 * multiplicity 2 c = multiplicity 2 (c ^ 2)" by (simp add: prime_elem_multiplicity_power_distrib) also have "multiplicity 2 (c ^ 2) = multiplicity (of_nat 2) (of_nat c ^ 2 :: gauss_int)" by (simp flip: multiplicity_gauss_int_of_nat) also have "of_nat 2 = (-\\<^sub>\) * (1 + \\<^sub>\) ^ 2" by (simp add: algebra_simps power2_eq_square) also have "multiplicity \ (of_nat c ^ 2) = multiplicity ((1 + \\<^sub>\) ^ 2) (of_nat c ^ 2)" by (subst multiplicity_times_unit_left) auto also have "\ = multiplicity (1 + \\<^sub>\) (of_nat c)" by (subst multiplicity_power_power) auto finally show ?thesis .. qed auto text \ Tha multiplicity of an inert Gaussian prime $q\in\mathbb{Z}$ in a Gaussian integer \z\ is precisely half the \\\-multiplicity of \q\ in the norm of \z\. \ lemma multiplicity_prime_cong_3_mod_4: assumes "prime (of_nat q :: gauss_int)" shows "multiplicity q (gauss_int_norm z) = 2 * multiplicity (of_nat q) z" proof (cases "z = 0") case [simp]: False have "multiplicity q (gauss_int_norm z) = multiplicity (of_nat q) (of_nat (gauss_int_norm z) :: gauss_int)" by (simp add: multiplicity_gauss_int_of_nat) also have "\ = multiplicity (of_nat q) (z * gauss_cnj z)" by (simp add: self_mult_gauss_cnj) also have "\ = multiplicity (of_nat q) z + multiplicity (gauss_cnj (of_nat q)) (gauss_cnj z)" using assms by (subst prime_elem_multiplicity_mult_distrib) auto also have "multiplicity (gauss_cnj (of_nat q)) (gauss_cnj z) = multiplicity (of_nat q) z" by (subst multiplicity_gauss_cnj) auto also have "\ + \ = 2 * \" by simp finally show ?thesis . qed auto text \ For Gaussian primes \p\ whose norm is congruent 1 modulo 4, the $\mathbb{Z}[i]$-multiplicity of \p\ in an integer \c\ is just the \\\-multiplicity of their norm in \c\. \ lemma multiplicity_prime_cong_1_mod_4_aux: fixes p :: gauss_int assumes "prime_elem p" "ReZ p > 0" "ImZ p > 0" "ImZ p \ ReZ p" shows "multiplicity p (of_nat c) = multiplicity (gauss_int_norm p) c" proof (cases "c = 0") case [simp]: False show ?thesis proof (intro antisym multiplicity_geI) define k where "k = multiplicity p (of_nat c)" have "p ^ k dvd of_nat c" by (simp add: multiplicity_dvd k_def) moreover have "gauss_cnj p ^ k dvd of_nat c" using multiplicity_dvd[of "gauss_cnj p" "of_nat c"] multiplicity_gauss_cnj[of p "of_nat c"] by (simp add: k_def) moreover have "\p dvd gauss_cnj p" using assms by (subst self_dvd_gauss_cnj_iff) auto hence "\p dvd gauss_cnj p ^ k" using assms prime_elem_dvd_power by blast ultimately have "p ^ k * gauss_cnj p ^ k dvd of_nat c" using assms by (intro prime_elem_power_mult_dvdI) auto also have "p ^ k * gauss_cnj p ^ k = of_nat (gauss_int_norm p ^ k)" by (simp flip: self_mult_gauss_cnj add: power_mult_distrib) finally show "gauss_int_norm p ^ k dvd c" by (subst (asm) of_nat_dvd_of_nat_gauss_int_iff) next define k where "k = multiplicity (gauss_int_norm p) c" have "p ^ k dvd (p * gauss_cnj p) ^ k" by (intro dvd_power_same) auto also have "\ = of_nat (gauss_int_norm p ^ k)" by (simp add: self_mult_gauss_cnj) also have "\ dvd of_nat c" unfolding of_nat_dvd_of_nat_gauss_int_iff by (auto simp: k_def multiplicity_dvd) finally show "p ^ k dvd of_nat c" . qed (use assms in \auto simp: gauss_int_norm_eq_Suc_0_iff\) qed auto text \ The multiplicity of a Gaussian prime with norm congruent 1 modulo 4 in some Gaussian integer \z\ and the multiplicity of its conjugate in \z\ sum to the the \\\-multiplicity of their norm in the norm of \z\: \ lemma multiplicity_prime_cong_1_mod_4: fixes p :: gauss_int assumes "prime_elem p" "ReZ p > 0" "ImZ p > 0" "ImZ p \ ReZ p" shows "multiplicity (gauss_int_norm p) (gauss_int_norm z) = multiplicity p z + multiplicity (gauss_cnj p) z" proof (cases "z = 0") case [simp]: False have "multiplicity (gauss_int_norm p) (gauss_int_norm z) = multiplicity p (of_nat (gauss_int_norm z))" using assms by (subst multiplicity_prime_cong_1_mod_4_aux) auto also have "\ = multiplicity p (z * gauss_cnj z)" by (simp add: self_mult_gauss_cnj) also have "\ = multiplicity p z + multiplicity p (gauss_cnj z)" using assms by (subst prime_elem_multiplicity_mult_distrib) auto also have "multiplicity p (gauss_cnj z) = multiplicity (gauss_cnj p) z" by (subst multiplicity_gauss_cnj [symmetric]) auto finally show ?thesis . qed auto text \ The multiplicity of the Gaussian prime \<^term>\1 + \\<^sub>\\ in a Gaussian integer \z\ is precisely the \\\-multiplicity of 2 in the norm of \z\: \ lemma multiplicity_prime_1_plus_i: "multiplicity (1 + \\<^sub>\) z = multiplicity 2 (gauss_int_norm z)" proof (cases "z = 0") case [simp]: False note [simp] = prime_elem_one_plus_i_gauss_int have "2 * multiplicity 2 (gauss_int_norm z) = multiplicity (1 + \\<^sub>\) (of_nat (gauss_int_norm z))" by (rule multiplicity_prime_1_plus_i_aux [symmetric]) also have "\ = multiplicity (1 + \\<^sub>\) (z * gauss_cnj z)" by (simp add: self_mult_gauss_cnj) also have "\ = multiplicity (1 + \\<^sub>\) z + multiplicity (gauss_cnj (1 - \\<^sub>\)) (gauss_cnj z)" by (subst prime_elem_multiplicity_mult_distrib) auto also have "multiplicity (gauss_cnj (1 - \\<^sub>\)) (gauss_cnj z) = multiplicity (1 - \\<^sub>\) z" by (subst multiplicity_gauss_cnj) auto also have "1 - \\<^sub>\ = (-\\<^sub>\) * (1 + \\<^sub>\)" by (simp add: algebra_simps) also have "multiplicity \ z = multiplicity (1 + \\<^sub>\) z" by (subst multiplicity_times_unit_left) auto also have "\ + \ = 2 * \" by simp finally show ?thesis by simp qed auto subsection \Coprimality of an element and its conjugate\ text \ Using the classification of the primes, we now show that if the real and imaginary parts of a Gaussian integer are coprime and its norm is odd, then it is coprime to its own conjugate. \ lemma coprime_self_gauss_cnj: assumes "coprime (ReZ z) (ImZ z)" and "odd (gauss_int_norm z)" shows "coprime z (gauss_cnj z)" proof (rule coprimeI) fix d assume "d dvd z" "d dvd gauss_cnj z" have *: False if "p \ prime_factors z" "p \ prime_factors (gauss_cnj z)" for p proof - from that have p: "prime p" "p dvd z" "p dvd gauss_cnj z" by auto define p' where "p' = gauss_cnj p" define d where "d = gauss_int_norm p" have of_nat_d_eq: "of_nat d = p * p'" by (simp add: p'_def self_mult_gauss_cnj d_def) have "prime_elem p" "prime_elem p'" "p dvd z" "p' dvd z" "p dvd gauss_cnj z" "p' dvd gauss_cnj z" using that by (auto simp: in_prime_factors_iff p'_def gauss_cnj_dvd_left_iff) have "prime p" using that by auto then obtain q where q: "prime q" "of_nat q dvd z" proof (cases rule: gauss_int_prime_classification) case one_plus_i hence "2 = gauss_int_norm p" by (auto simp: gauss_int_norm_def) also have "gauss_int_norm p dvd gauss_int_norm z" using p by (intro gauss_int_norm_dvd_mono) auto finally have "even (gauss_int_norm z)" . with \odd (gauss_int_norm z)\ show ?thesis by contradiction next case (cong_3_mod_4 q) thus ?thesis using that[of q] p by simp next case cong_1_mod_4 hence "\p dvd p'" unfolding p'_def by (subst self_dvd_gauss_cnj_iff) auto hence "p * p' dvd z" using p by (intro prime_elem_mult_dvdI) (auto simp: p'_def gauss_cnj_dvd_left_iff) also have "p * p' = of_nat (gauss_int_norm p)" by (simp add: p'_def self_mult_gauss_cnj) finally show ?thesis using that[of "gauss_int_norm p"] cong_1_mod_4 by simp qed have "of_nat q dvd gcd (2 * of_int (ReZ z)) (2 * \\<^sub>\ * of_int (ImZ z))" proof (rule gcd_greatest) have "of_nat q dvd (z + gauss_cnj z)" using q by (auto simp: gauss_cnj_dvd_right_iff) also have "\ = 2 * of_int (ReZ z)" by (simp add: self_plus_gauss_cnj) finally show "of_nat q dvd (2 * of_int (ReZ z) :: gauss_int)" . next have "of_nat q dvd (z - gauss_cnj z)" using q by (auto simp: gauss_cnj_dvd_right_iff) also have "\ = 2 * \\<^sub>\ * of_int (ImZ z)" by (simp add: self_minus_gauss_cnj) finally show "of_nat q dvd (2 * \\<^sub>\ * of_int (ImZ z))" . qed also have "\ = 2" proof - have "odd (ReZ z) \ odd (ImZ z)" using assms by (auto simp: gauss_int_norm_def even_nat_iff) thus ?thesis proof assume "odd (ReZ z)" hence "coprime (of_int (ReZ z)) (of_int 2 :: gauss_int)" unfolding coprime_of_int_gauss_int coprime_right_2_iff_odd . thus ?thesis using assms by (subst gcd_mult_left_right_cancel) (auto simp: coprime_of_int_gauss_int coprime_commute is_unit_left_imp_coprime is_unit_right_imp_coprime gcd_proj1_if_dvd gcd_proj2_if_dvd) next assume "odd (ImZ z)" hence "coprime (of_int (ImZ z)) (of_int 2 :: gauss_int)" unfolding coprime_of_int_gauss_int coprime_right_2_iff_odd . hence "gcd (2 * of_int (ReZ z)) (2 * \\<^sub>\ * of_int (ImZ z)) = gcd (2 * of_int (ReZ z)) (2 * \\<^sub>\)" using assms by (subst gcd_mult_right_right_cancel) (auto simp: coprime_of_int_gauss_int coprime_commute is_unit_left_imp_coprime is_unit_right_imp_coprime) also have "\ = normalize (2 * gcd (of_int (ReZ z)) \\<^sub>\)" by (subst gcd_mult_left) auto also have "gcd (of_int (ReZ z)) \\<^sub>\ = 1" by (subst coprime_iff_gcd_eq_1 [symmetric], rule is_unit_right_imp_coprime) auto finally show ?thesis by simp qed qed finally have "of_nat q dvd (of_nat 2 :: gauss_int)" by simp hence "q dvd 2" by (simp only: of_nat_dvd_of_nat_gauss_int_iff) with \prime q\ have "q = 2" using primes_dvd_imp_eq two_is_prime_nat by blast with q have "2 dvd z" by auto have "2 dvd gauss_int_norm 2" by simp also have "\ dvd gauss_int_norm z" using \2 dvd z\ by (intro gauss_int_norm_dvd_mono) finally show False using \odd (gauss_int_norm z)\ by contradiction qed fix d :: gauss_int assume d: "d dvd z" "d dvd gauss_cnj z" show "is_unit d" proof (rule ccontr) assume "\is_unit d" moreover from d assms have "d \ 0" by auto ultimately obtain p where p: "prime p" "p dvd d" using prime_divisorE by blast with d have "p \ prime_factors z" "p \ prime_factors (gauss_cnj z)" using assms by (auto simp: in_prime_factors_iff) with *[of p] show False by blast qed qed subsection \Square decompositions of prime numbers congruent 1 mod 4\ lemma prime_1_mod_4_sum_of_squares_unique_aux: fixes p x y :: nat assumes "prime p" "[p = 1] (mod 4)" "x ^ 2 + y ^ 2 = p" shows "x > 0 \ y > 0 \ x \ y" proof safe from assms show "x > 0" "y > 0" by (auto intro!: Nat.gr0I simp: prime_power_iff) next assume "x = y" with assms have "p = 2 * x ^ 2" by simp with \prime p\ have "p = 2" by (auto dest: prime_product) with \[p = 1] (mod 4)\ show False by (simp add: cong_def) qed text \ Any prime number congruent 1 modulo 4 can be written \<^emph>\uniquely\ as a sum of two squares $x^2 + y^2$ (up to commutativity of the addition). Additionally, we have shown above that \x\ and \y\ are both positive and \x \ y\. \ lemma prime_1_mod_4_sum_of_squares_unique: fixes p :: nat assumes "prime p" "[p = 1] (mod 4)" shows "\!(x,y). x \ y \ x ^ 2 + y ^ 2 = p" proof (rule ex_ex1I) obtain z where z: "gauss_int_norm z = p" using prime_cong_1_mod_4_gauss_int_norm_exists[OF assms] by blast show "\z. case z of (x,y) \ x \ y \ x ^ 2 + y ^ 2 = p" proof (cases "\ReZ z\ \ \ImZ z\") case True with z show ?thesis by (intro exI[of _ "(nat \ReZ z\, nat \ImZ z\)"]) (auto simp: gauss_int_norm_def nat_add_distrib simp flip: nat_power_eq) next case False with z show ?thesis by (intro exI[of _ "(nat \ImZ z\, nat \ReZ z\)"]) (auto simp: gauss_int_norm_def nat_add_distrib simp flip: nat_power_eq) qed next fix z1 z2 assume z1: "case z1 of (x, y) \ x \ y \ x\<^sup>2 + y\<^sup>2 = p" assume z2: "case z2 of (x, y) \ x \ y \ x\<^sup>2 + y\<^sup>2 = p" define z1' :: gauss_int where "z1' = of_nat (fst z1) + \\<^sub>\ * of_nat (snd z1)" define z2' :: gauss_int where "z2' = of_nat (fst z2) + \\<^sub>\ * of_nat (snd z2)" from assms interpret noninert_gauss_int_prime p by unfold_locales auto have norm_z1': "gauss_int_norm z1' = p" using z1 by (simp add: z1'_def gauss_int_norm_def case_prod_unfold nat_add_distrib nat_power_eq) have norm_z2': "gauss_int_norm z2' = p" using z2 by (simp add: z2'_def gauss_int_norm_def case_prod_unfold nat_add_distrib nat_power_eq) have sgns: "fst z1 > 0" "snd z1 > 0" "fst z2 > 0" "snd z2 > 0" "fst z1 \ snd z1" "fst z2 \ snd z2" using prime_1_mod_4_sum_of_squares_unique_aux[OF assms, of "fst z1" "snd z1"] z1 prime_1_mod_4_sum_of_squares_unique_aux[OF assms, of "fst z2" "snd z2"] z2 by auto have [simp]: "normalize z1' = z1'" "normalize z2' = z2'" using sgns by (subst normalized_gauss_int_iff; simp add: z1'_def z2'_def)+ have "prime z1'" "prime z2'" using norm_z1' norm_z2' assms unfolding prime_def by (auto simp: prime_gauss_int_norm_imp_prime_elem) have "of_nat p = z1' * gauss_cnj z1'" by (simp add: self_mult_gauss_cnj norm_z1') hence "z1' dvd of_nat p" by simp also have "of_nat p = z2' * gauss_cnj z2'" by (simp add: self_mult_gauss_cnj norm_z2') finally have "z1' dvd z2' \ z1' dvd gauss_cnj z2'" using assms by (subst (asm) prime_elem_dvd_mult_iff) (simp add: norm_z1' prime_gauss_int_norm_imp_prime_elem) thus "z1 = z2" proof assume "z1' dvd z2'" with \prime z1'\ \prime z2'\ have "z1' = z2'" by (simp add: primes_dvd_imp_eq) thus ?thesis by (simp add: z1'_def z2'_def gauss_int_eq_iff prod_eq_iff) next assume dvd: "z1' dvd gauss_cnj z2'" have "normalize (\\<^sub>\ * gauss_cnj z2') = \\<^sub>\ * gauss_cnj z2'" using sgns by (subst normalized_gauss_int_iff) (auto simp: z2'_def) moreover have "prime_elem (\\<^sub>\ * gauss_cnj z2')" by (rule prime_gauss_int_norm_imp_prime_elem) (simp add: gauss_int_norm_mult norm_z2' \prime p\) ultimately have "prime (\\<^sub>\ * gauss_cnj z2')" by (simp add: prime_def) moreover from dvd have "z1' dvd \\<^sub>\ * gauss_cnj z2'" by simp ultimately have "z1' = \\<^sub>\ * gauss_cnj z2'" using \prime z1'\ by (simp add: primes_dvd_imp_eq) hence False using z1 z2 sgns by (auto simp: gauss_int_eq_iff z1'_def z2'_def) thus ?thesis .. qed qed lemma two_sum_of_squares_nat_iff: "(x :: nat) ^ 2 + y ^ 2 = 2 \ x = 1 \ y = 1" proof assume eq: "x ^ 2 + y ^ 2 = 2" have square_neq_2: "n ^ 2 \ 2" for n :: nat proof assume *: "n ^ 2 = 2" have "prime (2 :: nat)" by simp thus False by (subst (asm) * [symmetric]) (auto simp: prime_power_iff) qed from eq have "x ^ 2 < 2 ^ 2" "y ^ 2 < 2 ^ 2" by simp_all hence "x < 2" "y < 2" using power2_less_imp_less[of x 2] power2_less_imp_less[of y 2] by auto moreover have "x > 0" "y > 0" using eq square_neq_2[of x] square_neq_2[of y] by (auto intro!: Nat.gr0I) ultimately show "x = 1 \ y = 1" by auto qed auto lemma prime_sum_of_squares_unique: fixes p :: nat assumes "prime p" "p = 2 \ [p = 1] (mod 4)" shows "\!(x,y). x \ y \ x ^ 2 + y ^ 2 = p" using assms(2) proof assume [simp]: "p = 2" have **: "(\(x,y). x \ y \ x ^ 2 + y ^ 2 = p) = (\z. z = (1,1 :: nat))" using two_sum_of_squares_nat_iff by (auto simp: fun_eq_iff) thus ?thesis by (subst **) auto qed (use prime_1_mod_4_sum_of_squares_unique[of p] assms in auto) text \ We now give a simple and inefficient algorithm to compute the canonical decomposition $x ^ 2 + y ^ 2$ with $x\leq y$. \ definition prime_square_sum_nat_decomp :: "nat \ nat \ nat" where "prime_square_sum_nat_decomp p = (if prime p \ (p = 2 \ [p = 1] (mod 4)) then THE (x,y). x \ y \ x ^ 2 + y ^ 2 = p else (0, 0))" lemma prime_square_sum_nat_decomp_eqI: assumes "prime p" "x ^ 2 + y ^ 2 = p" "x \ y" shows "prime_square_sum_nat_decomp p = (x, y)" proof - have "[gauss_int_norm (of_nat x + \\<^sub>\ * of_nat y) \ 3] (mod 4)" by (rule gauss_int_norm_not_3_mod_4) also have "gauss_int_norm (of_nat x + \\<^sub>\ * of_nat y) = p" using assms by (auto simp: gauss_int_norm_def nat_add_distrib nat_power_eq) finally have "[p \ 3] (mod 4)" . with prime_mod_4_cases[of p] assms have *: "p = 2 \ [p = 1] (mod 4)" by auto have "prime_square_sum_nat_decomp p = (THE (x,y). x \ y \ x ^ 2 + y ^ 2 = p)" using * \prime p\ by (simp add: prime_square_sum_nat_decomp_def) also have "\ = (x, y)" proof (rule the1_equality) show "\!(x,y). x \ y \ x ^ 2 + y ^ 2 = p" using \prime p\ * by (rule prime_sum_of_squares_unique) qed (use assms in auto) finally show ?thesis . qed lemma prime_square_sum_nat_decomp_correct: assumes "prime p" "p = 2 \ [p = 1] (mod 4)" defines "z \ prime_square_sum_nat_decomp p" shows "fst z ^ 2 + snd z ^ 2 = p" "fst z \ snd z" proof - define z' where "z' = (THE (x,y). x \ y \ x ^ 2 + y ^ 2 = p)" have "z = z'" unfolding z_def z'_def using assms by (simp add: prime_square_sum_nat_decomp_def) also have"\!(x,y). x \ y \ x ^ 2 + y ^ 2 = p" using assms by (intro prime_sum_of_squares_unique) hence "case z' of (x, y) \ x \ y \ x ^ 2 + y ^ 2 = p" unfolding z'_def by (rule theI') finally show "fst z ^ 2 + snd z ^ 2 = p" "fst z \ snd z" by auto qed lemma sum_of_squares_nat_bound: fixes x y n :: nat assumes "x ^ 2 + y ^ 2 = n" shows "x \ n" proof (cases "x = 0") case False hence "x * 1 \ x ^ 2" unfolding power2_eq_square by (intro mult_mono) auto also have "\ \ x ^ 2 + y ^ 2" by simp also have "\ = n" by fact finally show ?thesis by simp qed auto lemma sum_of_squares_nat_bound': fixes x y n :: nat assumes "x ^ 2 + y ^ 2 = n" shows "y \ n" using sum_of_squares_nat_bound[of y x] assms by (simp add: add.commute) lemma is_singleton_conv_Ex1: "is_singleton A \ (\!x. x \ A)" proof assume "is_singleton A" thus "\!x. x \ A" by (auto elim!: is_singletonE) next assume "\!x. x \ A" thus "is_singleton A" by (metis equals0D is_singletonI') qed lemma the_elemI: assumes "is_singleton A" shows "the_elem A \ A" using assms by (elim is_singletonE) auto lemma prime_square_sum_nat_decomp_code_aux: assumes "prime p" "p = 2 \ [p = 1] (mod 4)" defines "z \ the_elem (Set.filter (\(x,y). x ^ 2 + y ^ 2 = p) (SIGMA x:{0..p}. {x..p}))" shows "prime_square_sum_nat_decomp p = z" proof - let ?A = "Set.filter (\(x,y). x ^ 2 + y ^ 2 = p) (SIGMA x:{0..p}. {x..p})" have eq: "?A = {(x,y). x \ y \ x ^ 2 + y ^ 2 = p}" using sum_of_squares_nat_bound [of _ _ p] sum_of_squares_nat_bound' [of _ _ p] by auto have z: "z \ Set.filter (\(x,y). x ^ 2 + y ^ 2 = p) (SIGMA x:{0..p}. {x..p})" unfolding z_def eq using prime_sum_of_squares_unique[OF assms(1,2)] by (intro the_elemI) (simp add: is_singleton_conv_Ex1) have "prime_square_sum_nat_decomp p = (fst z, snd z)" using z by (intro prime_square_sum_nat_decomp_eqI[OF assms(1)]) auto also have "\ = z" by simp finally show ?thesis . qed lemma prime_square_sum_nat_decomp_code [code]: "prime_square_sum_nat_decomp p = (if prime p \ (p = 2 \ [p = 1] (mod 4)) then the_elem (Set.filter (\(x,y). x ^ 2 + y ^ 2 = p) (SIGMA x:{0..p}. {x..p})) else (0, 0))" using prime_square_sum_nat_decomp_code_aux[of p] by (auto simp: prime_square_sum_nat_decomp_def) subsection \Executable factorisation of Gaussian integers\ text \ Lastly, we use all of the above to give an executable (albeit not very efficient) factorisation algorithm for Gaussian integers based on factorisation of regular integers. Note that we will only compute the set of prime factors without multiplicity, but given that, it would be fairly easy to determine the multiplicity as well. First, we need the following function that computes the Gaussian integer factors of a \\\-prime \p\: \ definition factor_gauss_int_prime_nat :: "nat \ gauss_int list" where "factor_gauss_int_prime_nat p = (if p = 2 then [1 + \\<^sub>\] else if [p = 3] (mod 4) then [of_nat p] else case prime_square_sum_nat_decomp p of (x, y) \ [of_nat x + \\<^sub>\ * of_nat y, of_nat y + \\<^sub>\ * of_nat x])" lemma factor_gauss_int_prime_nat_correct: assumes "prime p" shows "set (factor_gauss_int_prime_nat p) = prime_factors (of_nat p)" using prime_mod_4_cases[OF assms] proof (elim disjE) assume "p = 2" thus ?thesis by (auto simp: prime_factorization_2_gauss_int factor_gauss_int_prime_nat_def) next assume *: "[p = 3] (mod 4)" with assms have "prime (of_nat p :: gauss_int)" by (intro prime_gauss_int_of_nat) thus ?thesis using assms * by (auto simp: prime_factorization_prime factor_gauss_int_prime_nat_def cong_def) next assume *: "[p = 1] (mod 4)" then interpret noninert_gauss_int_prime p using \prime p\ by unfold_locales define z where "z = prime_square_sum_nat_decomp p" define x y where "x = fst z" and "y = snd z" have xy: "x ^ 2 + y ^ 2 = p" "x \ y" using prime_square_sum_nat_decomp_correct[of p] * assms by (auto simp: x_def y_def z_def) from xy have xy_signs: "x > 0" "y > 0" using prime_1_mod_4_sum_of_squares_unique_aux[of p x y] assms * by auto have norms: "gauss_int_norm (of_nat x + \\<^sub>\ * of_nat y) = p" "gauss_int_norm (of_nat y + \\<^sub>\ * of_nat x) = p" using xy by (auto simp: gauss_int_norm_def nat_add_distrib nat_power_eq) have prime: "prime (of_nat x + \\<^sub>\ * of_nat y)" "prime (of_nat y + \\<^sub>\ * of_nat x)" using norms xy_signs \prime p\ unfolding prime_def normalized_gauss_int_iff by (auto intro!: prime_gauss_int_norm_imp_prime_elem) have "normalize ((of_nat x + \\<^sub>\ * of_nat y) * (of_nat y + \\<^sub>\ * of_nat x)) = of_nat p" proof - have "(of_nat x + \\<^sub>\ * of_nat y) * (of_nat y + \\<^sub>\ * of_nat x) = (\\<^sub>\ * of_nat p :: gauss_int)" by (subst xy(1) [symmetric]) (auto simp: gauss_int_eq_iff power2_eq_square) also have "normalize \ = of_nat p" by simp finally show ?thesis . qed hence "prime_factorization (of_nat p) = prime_factorization (prod_mset {#of_nat x + \\<^sub>\ * of_nat y, of_nat y + \\<^sub>\ * of_nat x#})" using assms xy by (subst prime_factorization_unique) (auto simp: gauss_int_eq_iff) also have "\ = {#of_nat x + \\<^sub>\ * of_nat y, of_nat y + \\<^sub>\ * of_nat x#}" using prime by (subst prime_factorization_prod_mset_primes) auto finally have "prime_factors (of_nat p) = {of_nat x + \\<^sub>\ * of_nat y, of_nat y + \\<^sub>\ * of_nat x}" by simp also have "\ = set (factor_gauss_int_prime_nat p)" using * unfolding factor_gauss_int_prime_nat_def case_prod_unfold by (auto simp: cong_def x_def y_def z_def) finally show ?thesis .. qed text \ Next, we lift this to compute the prime factorisation of any integer in the Gaussian integers: \ definition prime_factors_gauss_int_of_nat :: "nat \ gauss_int set" where "prime_factors_gauss_int_of_nat n = (if n = 0 then {} else (\p\prime_factors n. set (factor_gauss_int_prime_nat p)))" lemma prime_factors_gauss_int_of_nat_correct: "prime_factors_gauss_int_of_nat n = prime_factors (of_nat n)" proof (cases "n = 0") case False from False have [simp]: "n > 0" by auto have "prime_factors (of_nat n :: gauss_int) = prime_factors (of_nat (prod_mset (prime_factorization n)))" by (subst prod_mset_prime_factorization_nat [symmetric]) auto also have "\ = prime_factors (prod_mset (image_mset of_nat (prime_factorization n)))" by (subst of_nat_prod_mset) auto also have "\ = (\p\prime_factors n. prime_factors (of_nat p))" by (subst prime_factorization_prod_mset) auto also have "\ = (\p\prime_factors n. set (factor_gauss_int_prime_nat p))" by (intro SUP_cong refl factor_gauss_int_prime_nat_correct [symmetric]) auto finally show ?thesis by (simp add: prime_factors_gauss_int_of_nat_def) qed (auto simp: prime_factors_gauss_int_of_nat_def) text \ We can now use this to factor any Gaussian integer by computing a factorisation of its norm and removing all the prime divisors that do not actually divide it. \ definition prime_factors_gauss_int :: "gauss_int \ gauss_int set" where "prime_factors_gauss_int z = (if z = 0 then {} else Set.filter (\p. p dvd z) (prime_factors_gauss_int_of_nat (gauss_int_norm z)))" lemma prime_factors_gauss_int_correct [code_unfold]: "prime_factors z = prime_factors_gauss_int z" proof (cases "z = 0") case [simp]: False define n where "n = gauss_int_norm z" from False have [simp]: "n > 0" by (auto simp: n_def) have "prime_factors_gauss_int z = Set.filter (\p. p dvd z) (prime_factors (of_nat n))" by (simp add: prime_factors_gauss_int_of_nat_correct prime_factors_gauss_int_def n_def) also have "of_nat n = z * gauss_cnj z" by (simp add: n_def self_mult_gauss_cnj) also have "prime_factors \ = prime_factors z \ prime_factors (gauss_cnj z)" by (subst prime_factors_product) auto also have "Set.filter (\p. p dvd z) \ = prime_factors z" by (auto simp: in_prime_factors_iff) finally show ?thesis by simp qed (auto simp: prime_factors_gauss_int_def) (*<*) unbundle no_gauss_int_notation (*>*) end \ No newline at end of file diff --git a/thys/Lucas_Theorem/Lucas_Theorem.thy b/thys/Lucas_Theorem/Lucas_Theorem.thy --- a/thys/Lucas_Theorem/Lucas_Theorem.thy +++ b/thys/Lucas_Theorem/Lucas_Theorem.thy @@ -1,373 +1,372 @@ (* Title: Lucas_Theorem.thy Author: Chelsea Edmonds, University of Cambridge *) theory Lucas_Theorem imports Main "HOL-Computational_Algebra.Computational_Algebra" begin notation fps_nth (infixl "$" 75) section \Extensions on Formal Power Series (FPS) Library\ text \This section presents a few extensions on the Formal Power Series (FPS) library, described in \<^cite>\"Chaieb2011"\ \ subsection \FPS Equivalence Relation \ text \ This proof requires reasoning around the equivalence of coefficients mod some prime number. This section defines an equivalence relation on FPS using the pattern described by Paulson in \<^cite>\"paulsonDefiningFunctionsEquivalence2006"\, as well as some basic lemmas for reasoning around how the equivalence holds after common operations are applied \ definition "fpsmodrel p \ { (f, g). \ n. (f $ n) mod p = (g $ n) mod p }" lemma fpsrel_iff [simp]: "(f, g) \ fpsmodrel p \ (\n. (f $ n) mod p = (g $ n) mod p)" by (simp add: fpsmodrel_def) lemma fps_equiv: "equiv UNIV (fpsmodrel p)" proof (rule equivI) show "refl (fpsmodrel p)" by (simp add: refl_on_def fpsmodrel_def) show "sym (fpsmodrel p)" by (simp add: sym_def fpsmodrel_def) show "trans (fpsmodrel p)" by (intro transI) (simp add: fpsmodrel_def) qed text \ Equivalence relation over multiplication \ lemma fps_mult_equiv_coeff: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" assumes "(f, g) \ fpsmodrel p" shows "(f*h)$n mod p = (g*h)$n mod p" proof - have "((f*h) $ n) mod p =(\i=0..n. (f$i mod p * h$(n - i) mod p) mod p) mod p" using mod_sum_eq mod_mult_left_eq by (simp add: fps_mult_nth mod_sum_eq mod_mult_left_eq) also have "... = (\i=0..n. (g$i mod p * h$(n - i) mod p) mod p) mod p" using assms by auto also have "... = ((g*h) $ n) mod p" by (simp add: mod_mult_left_eq mod_sum_eq fps_mult_nth) thus ?thesis by (simp add: calculation) qed lemma fps_mult_equiv: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" assumes "(f, g) \ fpsmodrel p" shows "(f*h, g*h) \ fpsmodrel p" using fpsmodrel_def fps_mult_equiv_coeff assms by blast text \ Equivalence relation over power operator \ lemma fps_power_equiv: fixes f g :: "('a :: {euclidean_ring_cancel}) fps" fixes x :: nat assumes "(f, g) \ fpsmodrel p" shows "(f^x, g^x) \ fpsmodrel p" using assms proof (induct x) case 0 thus ?case by (simp add: fpsmodrel_def) next case (Suc x) then have hyp: " \n. f^x $ n mod p = g ^x $ n mod p" using fpsrel_iff by blast thus ?case proof - have fact: "\n h. (g * h) $ n mod p = (f * h) $ n mod p" by (metis assms fps_mult_equiv_coeff) have "\n h. (g ^ x * h) $ n mod p = (f ^ x * h) $ n mod p" by (simp add: fps_mult_equiv_coeff hyp) then have "\n h. (h * g ^ x) $ n mod p = (h * f ^ x) $ n mod p" by (simp add: mult.commute) thus ?thesis using fact by force qed qed subsection \Binomial Coefficients \ text \The @{term "fps_binomial"} definition in the formal power series uses the @{term "n gchoose k"} operator. It's defined as being of type @{typ "'a :: field_char_0 fps"}, however the equivalence relation requires a type @{typ 'a} that supports the modulo operator. The proof of the binomial theorem based on FPS coefficients below uses the choose operator and does not put bounds on the type of @{term "fps_X"}.\ lemma binomial_coeffs_induct: fixes n k :: nat shows "(1 + fps_X)^n $ k = of_nat(n choose k)" proof (induct n arbitrary: k) case 0 thus ?case by (metis binomial_eq_0_iff binomial_n_0 fps_nth_of_nat not_gr_zero of_nat_0 of_nat_1 power_0) next case h: (Suc n) fix k have start: "(1 + fps_X)^(n + 1) = (1 + fps_X) * (1 + fps_X)^n" by auto show ?case using One_nat_def Suc_eq_plus1 Suc_pred add.commute binomial_Suc_Suc binomial_n_0 fps_mult_fps_X_plus_1_nth h.hyps neq0_conv start by (smt of_nat_add) qed subsection \Freshman's Dream Lemma on FPS \ text \ The Freshman's dream lemma modulo a prime number $p$ is a well known proof that $(1 + x^p) \equiv (1 + x)^p \mod p$\ text \ First prove that $\binom{p^n}{k} \equiv 0 \mod p$ for $k \ge 1$ and $k < p^n$. The eventual proof only ended up requiring this with $n = 1$\ lemma pn_choose_k_modp_0: fixes n k::nat assumes "prime p" "k \ 1 \ k \ p^n - 1" "n > 0" shows "(p^n choose k) mod p = 0" proof - have inequality: "k \ p^n" using assms (2) by arith have choose_take_1: "((p^n - 1) choose ( k - 1))= fact (p^n - 1) div (fact (k - 1) * fact (p^n - k))" using binomial_altdef_nat diff_le_mono inequality assms(2) by auto have "k * (p^n choose k) = k * ((fact (p^n)) div (fact k * fact((p^n) - k)))" using assms binomial_fact'[OF inequality] by auto also have "... = k * fact (p^n) div (fact k * fact((p^n) - k))" using binomial_fact_lemma div_mult_self_is_m fact_gt_zero inequality mult.assoc mult.commute nat_0_less_mult_iff by smt also have "... = k * fact (p^n) div (k * fact (k - 1) * fact((p^n) - k))" by (metis assms(2) fact_nonzero fact_num_eq_if le0 le_antisym of_nat_id) also have "... = fact (p^n) div (fact (k - 1) * fact((p^n) - k))" using assms by auto also have "... = ((p^n) * fact (p^n - 1)) div (fact (k - 1) * fact((p^n) - k))" by (metis assms(2) fact_nonzero fact_num_eq_if inequality le0 le_antisym of_nat_id) also have "... = (p^n) * (fact (p^n - 1) div (fact (k - 1) * fact((p^n) - k)))" by (metis assms(2) calculation choose_take_1 neq0_conv not_one_le_zero times_binomial_minus1_eq) finally have equality: "k * (p^n choose k) = p^n * ((p^n - 1) choose (k - 1))" using assms(2) times_binomial_minus1_eq by auto then have dvd_result: "p^n dvd (k * (p^n choose k))" by simp have "\ (p^n dvd k)" using assms (2) binomial_n_0 diff_diff_cancel nat_dvd_not_less neq0_conv by auto then have "p dvd (p^n choose k)" using mult.commute prime_imp_prime_elem prime_power_dvd_multD assms dvd_result by metis thus "?thesis" by simp qed text \ Applying the above lemma to the coefficients of $(1 + X)^p$, it is easy to show that all coefficients other than the $0$th and $p$th will be $0$ \ lemma fps_middle_coeffs: assumes "prime p" "n \ 0 \ n \ p" shows "((1 + fps_X :: int fps) ^p) $ n mod p = 0 mod p" proof - let ?f = "(1 + fps_X :: int fps)^p" - have "\ n. n > 0 \ n < p \ (p choose n) mod p = 0" using pn_choose_k_modp_0 - by (metis (no_types, lifting) add_le_imp_le_diff assms(1) diff_diff_cancel diff_is_0_eq' - discrete le_add_diff_inverse le_numeral_extra(4) power_one_right zero_le_one zero_less_one) + have "\ n. n > 0 \ n < p \ (p choose n) mod p = 0" + using pn_choose_k_modp_0 [of p _ 1] \prime p\ by auto then have middle_0: "\ n. n > 0 \ n < p \ (?f $ n) mod p = 0" using binomial_coeffs_induct by (metis of_nat_0 zmod_int) have "\ n. n > p \ ?f $ n mod p = 0" using binomial_eq_0_iff binomial_coeffs_induct mod_0 by (metis of_nat_eq_0_iff) thus ?thesis using middle_0 assms(2) nat_neq_iff by auto qed text \It follows that $(1+ X)^p$ is equivalent to $(1 + X^p)$ under our equivalence relation, as required to prove the freshmans dream lemma. \ lemma fps_freshmans_dream: assumes "prime p" shows "(((1 + fps_X :: int fps ) ^p), (1 + (fps_X)^(p))) \ fpsmodrel p" proof - let ?f = "(1 + fps_X :: int fps)^p" let ?g = "(1 + (fps_X :: int fps)^p)" have all_f_coeffs: "\ n. n \ 0 \ n \ p \ ?f $ n mod p = 0 mod p" using fps_middle_coeffs assms by blast have "?g $ 0 = 1" using assms by auto then have "?g $ 0 mod p = 1 mod p" using int_ops(2) zmod_int assms by presburger then have "?g $ p mod p = 1 mod p" using assms by auto then have "\ n . ?f $ n mod p = ?g $ n mod p" using all_f_coeffs by (simp add: binomial_coeffs_induct) thus ?thesis using fpsrel_iff by blast qed section \Lucas's Theorem Proof\ text \A formalisation of Lucas's theorem based on a generating function proof using the existing formal power series (FPS) Isabelle library\ subsection \Reasoning about Coefficients Helpers\ text \A generating function proof of Lucas's theorem relies on direct comparison between coefficients of FPS which requires a number of helper lemmas to prove formally. In particular it compares the coefficients of $(1 + X)^n \mod p$ to $(1 + X^p)^N * (1 + X) ^rn \mod p$, where $N = n / p$, and $rn = n \mod p$. This section proves that the $k$th coefficient of $(1 + X^p)^N * (1 + X) ^rn = (N choose K) * (rn choose rk)$\ text \Applying the @{term "fps_compose"} operator enables reasoning about the coefficients of $(1 + X^p)^n$ using the existing binomial theorem proof with $X^p$ instead of $X$.\ lemma fps_binomial_p_compose: assumes "p \ 0" shows "(1 + (fps_X:: ('a :: {idom} fps))^p)^n = ((1 + fps_X)^n) oo (fps_X^p)" proof - have "(1::'a fps) + fps_X ^ p = 1 + fps_X oo fps_X ^ p" by (simp add: assms fps_compose_add_distrib) then show ?thesis by (simp add: assms fps_compose_power) qed text \ Next the proof determines the value of the $k$th coefficient of $(1 + X^p)^N$. \ lemma fps_X_pow_binomial_coeffs: assumes "prime p" shows "(1 + (fps_X ::int fps)^p)^N $k = (if p dvd k then (N choose (k div p)) else 0)" proof - let ?fx = "(fps_X :: int fps)" have "(1 + ?fx^p)^N $ k = (((1 + ?fx)^N) oo (?fx^p)) $k" by (metis assms fps_binomial_p_compose not_prime_0) also have "... = (\i=0..k.((1 + ?fx)^N)$i * ((?fx^p)^i$k))" by (simp add: fps_compose_nth) finally have coeffs: "(1 + ?fx^p)^N $ k = (\i=0..k. (N choose i) * ((?fx^(p*i))$k))" using binomial_coeffs_induct sum.cong by (metis (no_types, lifting) power_mult) thus ?thesis proof (cases "p dvd k") case False \ \$p$ does not divide $k$ implies the $k$th term has a coefficient of 0\ have "\ i. \(p dvd k) \ (?fx^(p*i)) $ k = 0" by auto thus ?thesis using coeffs by (simp add: False) next case True \ \$p$ divides $k$ implies the $k$th term has a non-zero coefficient\ have contained: "k div p \ {0.. k}" by simp have "\ i. i \ k div p \ (?fx^(p*i)) $ k = 0" using assms by auto then have notdivpis0: "\ i \ ({0 .. k} - {k div p}). (?fx^(p*i)) $ k = 0" by simp have "(1 + ?fx^p)^N $ k = (N choose (k div p)) * (?fx^(p * (k div p))) $ k + (\i\({0..k} -{k div p}). (N choose i) * ((?fx^(p*i))$k))" using contained coeffs sum.remove by (metis (no_types, lifting) finite_atLeastAtMost) thus ?thesis using notdivpis0 True by simp qed qed text \ The final helper lemma proves the $k$th coefficient is equivalent to $\binom{?N}{?K}*\binom{?rn}{?rk}$ as required.\ lemma fps_div_rep_coeffs: assumes "prime p" shows "((1 + (fps_X::int fps)^p)^(n div p) * (1 + fps_X)^(n mod p)) $ k = ((n div p) choose (k div p)) * ((n mod p) choose (k mod p))" (is "((1 + (fps_X::int fps)^p)^?N * (1 + fps_X)^?rn) $ k = (?N choose ?K) * (?rn choose ?rk)") proof - \ \Initial facts with results around representation and 0 valued terms\ let ?fx = "fps_X :: int fps" have krep: "k - ?rk = ?K*p" by (simp add: minus_mod_eq_mult_div) have rk_in_range: "?rk \ {0..k}" by simp have "\ i \ p. (?rn choose i) = 0" using binomial_eq_0_iff by (metis assms(1) leD le_less_trans linorder_cases mod_le_divisor mod_less_divisor prime_gt_0_nat) then have ptok0: "\ i \ {p..k}. ((?rn choose i) * (1 + ?fx^p)^?N $ (k - i)) = 0" by simp then have notrkis0: "\i \ {0.. k}. i \ ?rk \ (?rn choose i) * (1 + ?fx^p)^?N $ (k - i) = 0" proof (cases "k < p") case True \ \When $k < p$, it presents a side case with regards to range of reasoning\ then have k_value: "k = ?rk" by simp then have "\ i < k. \ (p dvd (k - i))" using True by (metis diff_diff_cancel diff_is_0_eq dvd_imp_mod_0 less_imp_diff_less less_irrefl_nat mod_less) then show ?thesis using fps_X_pow_binomial_coeffs assms(1) k_value by simp next case False then have "\ i < p. i \ ?rk \ \(p dvd (k - i))" using mod_nat_eqI by auto then have "\ i \ {0.. ?rk \ (1 + ?fx^p)^?N $ (k - i) = 0" using assms fps_X_pow_binomial_coeffs by simp then show ?thesis using ptok0 by auto qed \ \Main body of the proof, using helper facts above\ have "((1 + fps_X^p)^?N * (1 + fps_X)^?rn) $ k = (((1 + fps_X)^?rn) * (1 + fps_X^p)^?N) $ k" by (metis (no_types, opaque_lifting) distrib_left distrib_right fps_mult_fps_X_commute fps_one_mult(1) fps_one_mult(2) power_commuting_commutes) also have "... = (\i=0..k.(of_nat(?rn choose i)) * ((1 + (fps_X)^p)^?N $ (k - i)))" by (simp add: fps_mult_nth binomial_coeffs_induct) also have "... = ((?rn choose ?rk) * (1 + ?fx^p)^?N $ (k - ?rk)) + (\i\({0..k} - {?rk}). (?rn choose i) * (1 + ?fx^p)^?N $ (k - i))" using rk_in_range sum.remove by (metis (no_types, lifting) finite_atLeastAtMost) finally have "((1 + ?fx^p)^?N * (1 + ?fx)^?rn) $ k = ((?rn choose ?rk) * (1 + ?fx^p)^?N $ (k - ?rk))" using notrkis0 by simp thus ?thesis using fps_X_pow_binomial_coeffs assms krep by auto qed (* Lucas theorem proof *) subsection \Lucas Theorem Proof\ text \ The proof of Lucas's theorem combines a generating function approach, based off \<^cite>\"Fine"\ with induction. For formalisation purposes, it was easier to first prove a well known corollary of the main theorem (also often presented as an alternative statement for Lucas's theorem), which can itself be used to backwards prove the the original statement by induction. This approach was adapted from P. Cameron's lecture notes on combinatorics \<^cite>\"petercameronNotesCombinatorics2007"\ \ subsubsection \ Proof of the Corollary \ text \ This step makes use of the coefficient equivalence arguments proved in the previous sections \ corollary lucas_corollary: fixes n k :: nat assumes "prime p" shows "(n choose k) mod p = (((n div p) choose (k div p)) * ((n mod p) choose (k mod p))) mod p" (is "(n choose k) mod p = ((?N choose ?K) * (?rn choose ?rk)) mod p") proof - let ?fx = "fps_X :: int fps" have n_rep: "n = ?N * p + ?rn" by simp have k_rep: "k =?K * p + ?rk" by simp have rhs_coeffs: "((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn)) $ k = (?N choose ?K) * (?rn choose ?rk)" using assms fps_div_rep_coeffs k_rep n_rep by blast \ \Application of coefficient reasoning\ have "((((1 + ?fx)^p)^(?N) * (1 + ?fx)^(?rn)), ((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn))) \ fpsmodrel p" using fps_freshmans_dream assms fps_mult_equiv fps_power_equiv by blast \ \Application of equivalence facts and freshmans dream lemma\ then have modrel2: "((1 + ?fx)^n, ((1 + ?fx^p)^(?N) * (1 + ?fx)^(?rn))) \ fpsmodrel p" by (metis (mono_tags, opaque_lifting) mult_div_mod_eq power_add power_mult) thus ?thesis using fpsrel_iff binomial_coeffs_induct rhs_coeffs by (metis of_nat_eq_iff zmod_int) qed subsubsection \ Proof of the Theorem \ text \The theorem statement requires a formalised way of referring to the base $p$ representation of a number. We use a definition that specifies the $i$th digit of the base $p$ representation. This definition is originally from the Hilbert's 10th Problem Formalisation project \<^cite>\"bayerDPRMTheoremIsabelle2019"\ which this work contributes to.\ definition nth_digit_general :: "nat \ nat \ nat \ nat" where "nth_digit_general num i base = (num div (base ^ i)) mod base" text \Applying induction on $d$, where $d$ is the highest power required in either $n$ or $k$'s base $p$ representation, @{thm lucas_corollary} can be used to prove the original theorem.\ theorem lucas_theorem: fixes n k d::nat assumes "n < p ^ (Suc d)" assumes "k < p ^ (Suc d)" assumes "prime p" shows "(n choose k) mod p = (\i\d. ((nth_digit_general n i p) choose (nth_digit_general k i p))) mod p" using assms proof (induct d arbitrary: n k) case 0 thus ?case using nth_digit_general_def assms by simp next case (Suc d) \ \Representation Variables\ let ?N = "n div p" let ?K = "k div p" let ?nr = "n mod p" let ?kr = "k mod p" \ \Required assumption facts\ have Mlessthan: "?N < p ^ (Suc d)" using less_mult_imp_div_less power_Suc2 assms(3) prime_ge_2_nat Suc.prems(1) by metis have Nlessthan: "?K < p ^ (Suc d)" using less_mult_imp_div_less power_Suc2 prime_ge_2_nat Suc.prems(2) assms(3) by metis have shift_bounds_fact: "(\i=(Suc 0)..(Suc (d )). ((nth_digit_general n i p) choose (nth_digit_general k i p))) = (\i=0..(d). (nth_digit_general n (Suc i) p) choose (nth_digit_general k (Suc i) p))" using prod.shift_bounds_cl_Suc_ivl by blast \ \Product manipulation helper fact\ have "(n choose k ) mod p = ((?N choose ?K) * (?nr choose ?kr)) mod p" using lucas_corollary assms(3) by blast \ \Application of corollary\ also have "...= ((\i\d. ((nth_digit_general ?N i p) choose (nth_digit_general ?K i p))) * (?nr choose ?kr)) mod p" using Mlessthan Nlessthan Suc.hyps mod_mult_cong assms(3) by blast \ \Using Inductive Hypothesis\ \ \Product manipulation steps\ also have "... = ((\i=0..(d). (nth_digit_general n (Suc i) p) choose (nth_digit_general k (Suc i) p)) * (?nr choose ?kr)) mod p" using atMost_atLeast0 nth_digit_general_def div_mult2_eq by auto also have "... = ((\i=1..(d+1). (nth_digit_general n i p) choose (nth_digit_general k i p)) * ((nth_digit_general n 0 p) choose (nth_digit_general k 0 p))) mod p" using nth_digit_general_def shift_bounds_fact by simp finally have "(n choose k ) mod p = ((\i=0..(d+1). (nth_digit_general n i p) choose (nth_digit_general k i p))) mod p" using One_nat_def atMost_atLeast0 mult.commute prod.atLeast1_atMost_eq prod.atMost_shift by (smt Suc_eq_plus1 shift_bounds_fact) thus ?case using Suc_eq_plus1 atMost_atLeast0 by presburger qed end \ No newline at end of file diff --git a/thys/MFODL_Monitor_Optimized/Monitor.thy b/thys/MFODL_Monitor_Optimized/Monitor.thy --- a/thys/MFODL_Monitor_Optimized/Monitor.thy +++ b/thys/MFODL_Monitor_Optimized/Monitor.thy @@ -1,5600 +1,5600 @@ (*<*) theory Monitor imports Formula Optimized_Join "MFOTL_Monitor.Abstract_Monitor" "HOL-Library.While_Combinator" "HOL-Library.Mapping" "Deriving.Derive" "Generic_Join.Generic_Join_Correctness" begin (*>*) section \Generic monitoring algorithm\ text \The algorithm defined here abstracts over the implementation of the temporal operators.\ subsection \Monitorable formulas\ definition "mmonitorable \ \ safe_formula \ \ Formula.future_bounded \" definition "mmonitorable_regex b g r \ safe_regex b g r \ Regex.pred_regex Formula.future_bounded r" definition is_simple_eq :: "Formula.trm \ Formula.trm \ bool" where "is_simple_eq t1 t2 = (Formula.is_Const t1 \ (Formula.is_Const t2 \ Formula.is_Var t2) \ Formula.is_Var t1 \ Formula.is_Const t2)" fun mmonitorable_exec :: "Formula.formula \ bool" where "mmonitorable_exec (Formula.Eq t1 t2) = is_simple_eq t1 t2" | "mmonitorable_exec (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var y))) = (x = y)" | "mmonitorable_exec (Formula.Pred e ts) = list_all (\t. Formula.is_Var t \ Formula.is_Const t) ts" | "mmonitorable_exec (Formula.Let p \ \) = ({0..} \ Formula.fv \ \ mmonitorable_exec \ \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Neg \) = (fv \ = {} \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Or \ \) = (fv \ = fv \ \ mmonitorable_exec \ \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.And \ \) = (mmonitorable_exec \ \ (safe_assignment (fv \) \ \ mmonitorable_exec \ \ fv \ \ fv \ \ (is_constraint \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))))" | "mmonitorable_exec (Formula.Ands l) = (let (pos, neg) = partition mmonitorable_exec l in pos \ [] \ list_all mmonitorable_exec (map remove_neg neg) \ \(set (map fv neg)) \ \(set (map fv pos)))" | "mmonitorable_exec (Formula.Exists \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Agg y \ b f \) = (mmonitorable_exec \ \ y + b \ Formula.fv \ \ {0.. Formula.fv \ \ Formula.fv_trm f \ Formula.fv \)" | "mmonitorable_exec (Formula.Prev I \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Next I \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Since \ I \) = (Formula.fv \ \ Formula.fv \ \ (mmonitorable_exec \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False)) \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Until \ I \) = (Formula.fv \ \ Formula.fv \ \ right I \ \ \ (mmonitorable_exec \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False)) \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.MatchP I r) = Regex.safe_regex Formula.fv (\g \. mmonitorable_exec \ \ (g = Lax \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))) Past Strict r" | "mmonitorable_exec (Formula.MatchF I r) = (Regex.safe_regex Formula.fv (\g \. mmonitorable_exec \ \ (g = Lax \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))) Futu Strict r \ right I \ \)" | "mmonitorable_exec _ = False" lemma cases_Neg_iff: "(case \ of formula.Neg \ \ P \ | _ \ False) \ (\\. \ = formula.Neg \ \ P \)" by (cases \) auto lemma safe_formula_mmonitorable_exec: "safe_formula \ \ Formula.future_bounded \ \ mmonitorable_exec \" proof (induct \ rule: safe_formula.induct) case (8 \ \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (9 \ \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (10 l) from "10.prems"(2) have bounded: "Formula.future_bounded \" if "\ \ set l" for \ using that by (auto simp: list.pred_set) obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp have "set poss \ set posm" proof (rule subsetI) fix x assume "x \ set poss" then have "x \ set l" "safe_formula x" using posnegs by simp_all then have "mmonitorable_exec x" using "10.hyps"(1) bounded by blast then show "x \ set posm" using \x \ set poss\ posnegm posnegs by simp qed then have "set negm \ set negs" using posnegm posnegs by auto obtain "poss \ []" "list_all safe_formula (map remove_neg negs)" "(\x\set negs. fv x) \ (\x\set poss. fv x)" using "10.prems"(1) posnegs by simp then have "posm \ []" using \set poss \ set posm\ by auto moreover have "list_all mmonitorable_exec (map remove_neg negm)" proof - let ?l = "map remove_neg negm" have "\x. x \ set ?l \ mmonitorable_exec x" proof - fix x assume "x \ set ?l" then obtain y where "y \ set negm" "x = remove_neg y" by auto then have "y \ set negs" using \set negm \ set negs\ by blast then have "safe_formula x" unfolding \x = remove_neg y\ using \list_all safe_formula (map remove_neg negs)\ by (simp add: list_all_def) show "mmonitorable_exec x" proof (cases "\z. y = Formula.Neg z") case True then obtain z where "y = Formula.Neg z" by blast then show ?thesis using "10.hyps"(2)[OF posnegs refl] \x = remove_neg y\ \y \ set negs\ posnegs bounded \safe_formula x\ by fastforce next case False then have "remove_neg y = y" by (cases y) simp_all then have "y = x" unfolding \x = remove_neg y\ by simp show ?thesis using "10.hyps"(1) \y \ set negs\ posnegs \safe_formula x\ unfolding \y = x\ by auto qed qed then show ?thesis by (simp add: list_all_iff) qed moreover have "(\x\set negm. fv x) \ (\x\set posm. fv x)" using \\ (fv ` set negs) \ \ (fv ` set poss)\ \set poss \ set posm\ \set negm \ set negs\ by fastforce ultimately show ?case using posnegm by simp next case (15 \ I \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (16 \ I \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (17 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set) next case (18 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set) qed (auto simp add: is_simple_eq_def list.pred_set) lemma safe_assignment_future_bounded: "safe_assignment X \ \ Formula.future_bounded \" unfolding safe_assignment_def by (auto split: formula.splits) lemma is_constraint_future_bounded: "is_constraint \ \ Formula.future_bounded \" by (induction rule: is_constraint.induct) simp_all lemma mmonitorable_exec_mmonitorable: "mmonitorable_exec \ \ mmonitorable \" proof (induct \ rule: mmonitorable_exec.induct) case (7 \ \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (auto simp: cases_Neg_iff intro: safe_assignment_future_bounded is_constraint_future_bounded) next case (8 l) obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp have pos_monitorable: "list_all mmonitorable_exec posm" using posnegm by (simp add: list_all_iff) have neg_monitorable: "list_all mmonitorable_exec (map remove_neg negm)" using "8.prems" posnegm by (simp add: list_all_iff) have "set posm \ set poss" using "8.hyps"(1) posnegs posnegm unfolding mmonitorable_def by auto then have "set negs \ set negm" using posnegs posnegm by auto have "poss \ []" using "8.prems" posnegm \set posm \ set poss\ by auto moreover have "list_all safe_formula (map remove_neg negs)" using neg_monitorable "8.hyps"(2)[OF posnegm refl] \set negs \ set negm\ unfolding mmonitorable_def by (auto simp: list_all_iff) moreover have "\ (fv ` set negm) \ \ (fv ` set posm)" using "8.prems" posnegm by simp then have "\ (fv ` set negs) \ \ (fv ` set poss)" using \set posm \ set poss\ \set negs \ set negm\ by fastforce ultimately have "safe_formula (Formula.Ands l)" using posnegs by simp moreover have "Formula.future_bounded (Formula.Ands l)" proof - have "list_all Formula.future_bounded posm" using pos_monitorable "8.hyps"(1) posnegm unfolding mmonitorable_def by (auto elim!: list.pred_mono_strong) moreover have fnegm: "list_all Formula.future_bounded (map remove_neg negm)" using neg_monitorable "8.hyps"(2) posnegm unfolding mmonitorable_def by (auto elim!: list.pred_mono_strong) then have "list_all Formula.future_bounded negm" proof - have "\x. x \ set negm \ Formula.future_bounded x" proof - fix x assume "x \ set negm" have "Formula.future_bounded (remove_neg x)" using fnegm \x \ set negm\ by (simp add: list_all_iff) then show "Formula.future_bounded x" by (cases x) auto qed then show ?thesis by (simp add: list_all_iff) qed ultimately have "list_all Formula.future_bounded l" using posnegm by (auto simp: list_all_iff) then show ?thesis by (auto simp: list_all_iff) qed ultimately show ?case unfolding mmonitorable_def .. next case (13 \ I \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (fastforce simp: cases_Neg_iff) next case (14 \ I \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (fastforce simp: one_enat_def cases_Neg_iff) next case (15 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated] simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set) next case (16 I r) then show ?case by (auto 0 3 elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated] simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set) qed (auto simp add: mmonitorable_def mmonitorable_regex_def is_simple_eq_def one_enat_def list.pred_set) lemma monitorable_formula_code[code]: "mmonitorable \ = mmonitorable_exec \" using mmonitorable_exec_mmonitorable safe_formula_mmonitorable_exec mmonitorable_def by blast subsection \Handling regular expressions\ datatype mregex = MSkip nat | MTestPos nat | MTestNeg nat | MPlus mregex mregex | MTimes mregex mregex | MStar mregex primrec ok where "ok _ (MSkip n) = True" | "ok m (MTestPos n) = (n < m)" | "ok m (MTestNeg n) = (n < m)" | "ok m (MPlus r s) = (ok m r \ ok m s)" | "ok m (MTimes r s) = (ok m r \ ok m s)" | "ok m (MStar r) = ok m r" primrec from_mregex where "from_mregex (MSkip n) _ = Regex.Skip n" | "from_mregex (MTestPos n) \s = Regex.Test (\s ! n)" | "from_mregex (MTestNeg n) \s = (if safe_formula (Formula.Neg (\s ! n)) then Regex.Test (Formula.Neg (Formula.Neg (Formula.Neg (\s ! n)))) else Regex.Test (Formula.Neg (\s ! n)))" | "from_mregex (MPlus r s) \s = Regex.Plus (from_mregex r \s) (from_mregex s \s)" | "from_mregex (MTimes r s) \s = Regex.Times (from_mregex r \s) (from_mregex s \s)" | "from_mregex (MStar r) \s = Regex.Star (from_mregex r \s)" primrec to_mregex_exec where "to_mregex_exec (Regex.Skip n) xs = (MSkip n, xs)" | "to_mregex_exec (Regex.Test \) xs = (if safe_formula \ then (MTestPos (length xs), xs @ [\]) else case \ of Formula.Neg \' \ (MTestNeg (length xs), xs @ [\']) | _ \ (MSkip 0, xs))" | "to_mregex_exec (Regex.Plus r s) xs = (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys in (MPlus mr ms, zs))" | "to_mregex_exec (Regex.Times r s) xs = (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys in (MTimes mr ms, zs))" | "to_mregex_exec (Regex.Star r) xs = (let (mr, ys) = to_mregex_exec r xs in (MStar mr, ys))" primrec shift where "shift (MSkip n) k = MSkip n" | "shift (MTestPos i) k = MTestPos (i + k)" | "shift (MTestNeg i) k = MTestNeg (i + k)" | "shift (MPlus r s) k = MPlus (shift r k) (shift s k)" | "shift (MTimes r s) k = MTimes (shift r k) (shift s k)" | "shift (MStar r) k = MStar (shift r k)" primrec to_mregex where "to_mregex (Regex.Skip n) = (MSkip n, [])" | "to_mregex (Regex.Test \) = (if safe_formula \ then (MTestPos 0, [\]) else case \ of Formula.Neg \' \ (MTestNeg 0, [\']) | _ \ (MSkip 0, []))" | "to_mregex (Regex.Plus r s) = (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s in (MPlus mr (shift ms (length ys)), ys @ zs))" | "to_mregex (Regex.Times r s) = (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s in (MTimes mr (shift ms (length ys)), ys @ zs))" | "to_mregex (Regex.Star r) = (let (mr, ys) = to_mregex r in (MStar mr, ys))" lemma shift_0: "shift r 0 = r" by (induct r) auto lemma shift_shift: "shift (shift r k) j = shift r (k + j)" by (induct r) auto lemma to_mregex_to_mregex_exec: "case to_mregex r of (mr, \s) \ to_mregex_exec r xs = (shift mr (length xs), xs @ \s)" by (induct r arbitrary: xs) (auto simp: shift_shift ac_simps split: formula.splits prod.splits) lemma to_mregex_to_mregex_exec_Nil[code]: "to_mregex r = to_mregex_exec r []" using to_mregex_to_mregex_exec[where xs="[]" and r = r] by (auto simp: shift_0) lemma ok_mono: "ok m mr \ m \ n \ ok n mr" by (induct mr) auto lemma from_mregex_cong: "ok m mr \ (\i < m. xs ! i = ys ! i) \ from_mregex mr xs = from_mregex mr ys" by (induct mr) auto lemma not_Neg_cases: "(\\. \ \ Formula.Neg \) \ (case \ of formula.Neg \ \ f \ | _ \ x) = x" by (cases \) auto lemma to_mregex_exec_ok: "to_mregex_exec r xs = (mr, ys) \ \zs. ys = xs @ zs \ set zs = atms r \ ok (length ys) mr" proof (induct r arbitrary: xs mr ys) case (Skip x) then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case (Test x) show ?case proof (cases "\x'. x = Formula.Neg x'") case True with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case False with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono) qed next case (Plus r1 r2) then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono) next case (Times r1 r2) then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono) next case (Star r) then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) qed lemma ok_shift: "ok (i + m) (Monitor.shift r i) \ ok m r" by (induct r) auto lemma to_mregex_ok: "to_mregex r = (mr, ys) \ set ys = atms r \ ok (length ys) mr" proof (induct r arbitrary: mr ys) case (Skip x) then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits) next case (Test x) show ?case proof (cases "\x'. x = Formula.Neg x'") case True with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case False with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono) qed next case (Plus r1 r2) then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits) next case (Times r1 r2) then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits) next case (Star r) then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits) qed lemma from_mregex_shift: "from_mregex (shift r (length xs)) (xs @ ys) = from_mregex r ys" by (induct r) (auto simp: nth_append) lemma from_mregex_to_mregex: "safe_regex m g r \ case_prod from_mregex (to_mregex r) = r" by (induct r rule: safe_regex.induct) (auto dest: to_mregex_ok intro!: from_mregex_cong simp: nth_append from_mregex_shift cases_Neg_iff split: prod.splits modality.splits) lemma from_mregex_eq: "safe_regex m g r \ to_mregex r = (mr, \s) \ from_mregex mr \s = r" using from_mregex_to_mregex[of m g r] by auto lemma from_mregex_to_mregex_exec: "safe_regex m g r \ case_prod from_mregex (to_mregex_exec r xs) = r" proof (induct r arbitrary: xs rule: safe_regex_induct) case (Plus b g r s) from Plus(3) Plus(1)[of xs] Plus(2)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (TimesF g r s) from TimesF(3) TimesF(2)[of xs] TimesF(1)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (TimesP g r s) from TimesP(3) TimesP(1)[of xs] TimesP(2)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (Star b g r) from Star(2) Star(1)[of xs] show ?case by (auto split: prod.splits) qed (auto split: prod.splits simp: cases_Neg_iff) derive linorder mregex subsubsection \LPD\ definition saturate where "saturate f = while (\S. f S \ S) f" lemma saturate_code[code]: "saturate f S = (let S' = f S in if S' = S then S else saturate f S')" unfolding saturate_def Let_def by (subst while_unfold) auto definition "MTimesL r S = MTimes r ` S" definition "MTimesR R s = (\r. MTimes r s) ` R" primrec LPD where "LPD (MSkip n) = (case n of 0 \ {} | Suc m \ {MSkip m})" | "LPD (MTestPos \) = {}" | "LPD (MTestNeg \) = {}" | "LPD (MPlus r s) = (LPD r \ LPD s)" | "LPD (MTimes r s) = MTimesR (LPD r) s \ LPD s" | "LPD (MStar r) = MTimesR (LPD r) (MStar r)" primrec LPDi where "LPDi 0 r = {r}" | "LPDi (Suc i) r = (\s \ LPD r. LPDi i s)" lemma LPDi_Suc_alt: "LPDi (Suc i) r = (\s \ LPDi i r. LPD s)" by (induct i arbitrary: r) fastforce+ definition "LPDs r = (\i. LPDi i r)" lemma LPDs_refl: "r \ LPDs r" by (auto simp: LPDs_def intro: exI[of _ 0]) lemma LPDs_trans: "r \ LPD s \ s \ LPDs t \ r \ LPDs t" by (force simp: LPDs_def LPDi_Suc_alt simp del: LPDi.simps intro: exI[of _ "Suc _"]) lemma LPDi_Test: "LPDi i (MSkip 0) \ {MSkip 0}" "LPDi i (MTestPos \) \ {MTestPos \}" "LPDi i (MTestNeg \) \ {MTestNeg \}" by (induct i) auto lemma LPDs_Test: "LPDs (MSkip 0) \ {MSkip 0}" "LPDs (MTestPos \) \ {MTestPos \}" "LPDs (MTestNeg \) \ {MTestNeg \}" unfolding LPDs_def using LPDi_Test by blast+ lemma LPDi_MSkip: "LPDi i (MSkip n) \ MSkip ` {i. i \ n}" by (induct i arbitrary: n) (auto dest: set_mp[OF LPDi_Test(1)] simp: le_Suc_eq split: nat.splits) lemma LPDs_MSkip: "LPDs (MSkip n) \ MSkip ` {i. i \ n}" unfolding LPDs_def using LPDi_MSkip by auto lemma LPDi_Plus: "LPDi i (MPlus r s) \ {MPlus r s} \ LPDi i r \ LPDi i s" by (induct i arbitrary: r s) auto lemma LPDs_Plus: "LPDs (MPlus r s) \ {MPlus r s} \ LPDs r \ LPDs s" unfolding LPDs_def using LPDi_Plus[of _ r s] by auto lemma LPDi_Times: "LPDi i (MTimes r s) \ {MTimes r s} \ MTimesR (\j \ i. LPDi j r) s \ (\j \ i. LPDi j s)" proof (induct i arbitrary: r s) case (Suc i) show ?case by (fastforce simp: MTimesR_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc]) qed simp lemma LPDs_Times: "LPDs (MTimes r s) \ {MTimes r s} \ MTimesR (LPDs r) s \ LPDs s" unfolding LPDs_def using LPDi_Times[of _ r s] by (force simp: MTimesR_def) lemma LPDi_Star: "j \ i \ LPDi j (MStar r) \ {MStar r} \ MTimesR (\j \ i. LPDi j r) (MStar r)" proof (induct i arbitrary: j r) case (Suc i) from Suc(2) show ?case by (auto 0 5 simp: MTimesR_def image_iff le_Suc_eq dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF LPDi_Times]) qed simp lemma LPDs_Star: "LPDs (MStar r) \ {MStar r} \ MTimesR (LPDs r) (MStar r)" unfolding LPDs_def using LPDi_Star[OF order_refl, of _ r] by (force simp: MTimesR_def) lemma finite_LPDs: "finite (LPDs r)" proof (induct r) case (MSkip n) then show ?case by (intro finite_subset[OF LPDs_MSkip]) simp next case (MTestPos \) then show ?case by (intro finite_subset[OF LPDs_Test(2)]) simp next case (MTestNeg \) then show ?case by (intro finite_subset[OF LPDs_Test(3)]) simp next case (MPlus r s) then show ?case by (intro finite_subset[OF LPDs_Plus]) simp next case (MTimes r s) then show ?case by (intro finite_subset[OF LPDs_Times]) (simp add: MTimesR_def) next case (MStar r) then show ?case by (intro finite_subset[OF LPDs_Star]) (simp add: MTimesR_def) qed context begin private abbreviation (input) "addLPD r \ \S. insert r S \ Set.bind (insert r S) LPD" private lemma mono_addLPD: "mono (addLPD r)" unfolding mono_def Set.bind_def by auto private lemma LPDs_aux1: "lfp (addLPD r) \ LPDs r" by (rule lfp_induct[OF mono_addLPD], auto intro: LPDs_refl LPDs_trans simp: Set.bind_def) private lemma LPDs_aux2: "LPDi i r \ lfp (addLPD r)" proof (induct i) case 0 then show ?case by (subst lfp_unfold[OF mono_addLPD]) auto next case (Suc i) then show ?case by (subst lfp_unfold[OF mono_addLPD]) (auto simp: LPDi_Suc_alt simp del: LPDi.simps) qed lemma LPDs_alt: "LPDs r = lfp (addLPD r)" using LPDs_aux1 LPDs_aux2 by (force simp: LPDs_def) lemma LPDs_code[code]: "LPDs r = saturate (addLPD r) {}" unfolding LPDs_alt saturate_def by (rule lfp_while[OF mono_addLPD _ finite_LPDs, of r]) (auto simp: LPDs_refl LPDs_trans Set.bind_def) end subsubsection \RPD\ primrec RPD where "RPD (MSkip n) = (case n of 0 \ {} | Suc m \ {MSkip m})" | "RPD (MTestPos \) = {}" | "RPD (MTestNeg \) = {}" | "RPD (MPlus r s) = (RPD r \ RPD s)" | "RPD (MTimes r s) = MTimesL r (RPD s) \ RPD r" | "RPD (MStar r) = MTimesL (MStar r) (RPD r)" primrec RPDi where "RPDi 0 r = {r}" | "RPDi (Suc i) r = (\s \ RPD r. RPDi i s)" lemma RPDi_Suc_alt: "RPDi (Suc i) r = (\s \ RPDi i r. RPD s)" by (induct i arbitrary: r) fastforce+ definition "RPDs r = (\i. RPDi i r)" lemma RPDs_refl: "r \ RPDs r" by (auto simp: RPDs_def intro: exI[of _ 0]) lemma RPDs_trans: "r \ RPD s \ s \ RPDs t \ r \ RPDs t" by (force simp: RPDs_def RPDi_Suc_alt simp del: RPDi.simps intro: exI[of _ "Suc _"]) lemma RPDi_Test: "RPDi i (MSkip 0) \ {MSkip 0}" "RPDi i (MTestPos \) \ {MTestPos \}" "RPDi i (MTestNeg \) \ {MTestNeg \}" by (induct i) auto lemma RPDs_Test: "RPDs (MSkip 0) \ {MSkip 0}" "RPDs (MTestPos \) \ {MTestPos \}" "RPDs (MTestNeg \) \ {MTestNeg \}" unfolding RPDs_def using RPDi_Test by blast+ lemma RPDi_MSkip: "RPDi i (MSkip n) \ MSkip ` {i. i \ n}" by (induct i arbitrary: n) (auto dest: set_mp[OF RPDi_Test(1)] simp: le_Suc_eq split: nat.splits) lemma RPDs_MSkip: "RPDs (MSkip n) \ MSkip ` {i. i \ n}" unfolding RPDs_def using RPDi_MSkip by auto lemma RPDi_Plus: "RPDi i (MPlus r s) \ {MPlus r s} \ RPDi i r \ RPDi i s" by (induct i arbitrary: r s) auto lemma RPDi_Suc_RPD_Plus: "RPDi (Suc i) r \ RPDs (MPlus r s)" "RPDi (Suc i) s \ RPDs (MPlus r s)" unfolding RPDs_def by (force intro!: exI[of _ "Suc i"])+ lemma RPDs_Plus: "RPDs (MPlus r s) \ {MPlus r s} \ RPDs r \ RPDs s" unfolding RPDs_def using RPDi_Plus[of _ r s] by auto lemma RPDi_Times: "RPDi i (MTimes r s) \ {MTimes r s} \ MTimesL r (\j \ i. RPDi j s) \ (\j \ i. RPDi j r)" proof (induct i arbitrary: r s) case (Suc i) show ?case by (fastforce simp: MTimesL_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc]) qed simp lemma RPDs_Times: "RPDs (MTimes r s) \ {MTimes r s} \ MTimesL r (RPDs s) \ RPDs r" unfolding RPDs_def using RPDi_Times[of _ r s] by (force simp: MTimesL_def) lemma RPDi_Star: "j \ i \ RPDi j (MStar r) \ {MStar r} \ MTimesL (MStar r) (\j \ i. RPDi j r)" proof (induct i arbitrary: j r) case (Suc i) from Suc(2) show ?case by (auto 0 5 simp: MTimesL_def image_iff le_Suc_eq dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF RPDi_Times]) qed simp lemma RPDs_Star: "RPDs (MStar r) \ {MStar r} \ MTimesL (MStar r) (RPDs r)" unfolding RPDs_def using RPDi_Star[OF order_refl, of _ r] by (force simp: MTimesL_def) lemma finite_RPDs: "finite (RPDs r)" proof (induct r) case MSkip then show ?case by (intro finite_subset[OF RPDs_MSkip]) simp next case (MTestPos \) then show ?case by (intro finite_subset[OF RPDs_Test(2)]) simp next case (MTestNeg \) then show ?case by (intro finite_subset[OF RPDs_Test(3)]) simp next case (MPlus r s) then show ?case by (intro finite_subset[OF RPDs_Plus]) simp next case (MTimes r s) then show ?case by (intro finite_subset[OF RPDs_Times]) (simp add: MTimesL_def) next case (MStar r) then show ?case by (intro finite_subset[OF RPDs_Star]) (simp add: MTimesL_def) qed context begin private abbreviation (input) "addRPD r \ \S. insert r S \ Set.bind (insert r S) RPD" private lemma mono_addRPD: "mono (addRPD r)" unfolding mono_def Set.bind_def by auto private lemma RPDs_aux1: "lfp (addRPD r) \ RPDs r" by (rule lfp_induct[OF mono_addRPD], auto intro: RPDs_refl RPDs_trans simp: Set.bind_def) private lemma RPDs_aux2: "RPDi i r \ lfp (addRPD r)" proof (induct i) case 0 then show ?case by (subst lfp_unfold[OF mono_addRPD]) auto next case (Suc i) then show ?case by (subst lfp_unfold[OF mono_addRPD]) (auto simp: RPDi_Suc_alt simp del: RPDi.simps) qed lemma RPDs_alt: "RPDs r = lfp (addRPD r)" using RPDs_aux1 RPDs_aux2 by (force simp: RPDs_def) lemma RPDs_code[code]: "RPDs r = saturate (addRPD r) {}" unfolding RPDs_alt saturate_def by (rule lfp_while[OF mono_addRPD _ finite_RPDs, of r]) (auto simp: RPDs_refl RPDs_trans Set.bind_def) end subsection \The executable monitor\ type_synonym ts = nat type_synonym 'a mbuf2 = "'a table list \ 'a table list" type_synonym 'a mbufn = "'a table list list" type_synonym 'a msaux = "(ts \ 'a table) list" type_synonym 'a muaux = "(ts \ 'a table \ 'a table) list" type_synonym 'a mr\aux = "(ts \ (mregex, 'a table) mapping) list" type_synonym 'a ml\aux = "(ts \ 'a table list \ 'a table) list" datatype mconstraint = MEq | MLess | MLessEq record args = args_ivl :: \ args_n :: nat args_L :: "nat set" args_R :: "nat set" args_pos :: bool datatype (dead 'msaux, dead 'muaux) mformula = MRel "event_data table" | MPred Formula.name "Formula.trm list" | MLet Formula.name nat "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" | MAnd "nat set" "('msaux, 'muaux) mformula" bool "nat set" "('msaux, 'muaux) mformula" "event_data mbuf2" | MAndAssign "('msaux, 'muaux) mformula" "nat \ Formula.trm" | MAndRel "('msaux, 'muaux) mformula" "Formula.trm \ bool \ mconstraint \ Formula.trm" | MAnds "nat set list" "nat set list" "('msaux, 'muaux) mformula list" "event_data mbufn" | MOr "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" | MNeg "('msaux, 'muaux) mformula" | MExists "('msaux, 'muaux) mformula" | MAgg bool nat Formula.agg_op nat "Formula.trm" "('msaux, 'muaux) mformula" | MPrev \ "('msaux, 'muaux) mformula" bool "event_data table list" "ts list" | MNext \ "('msaux, 'muaux) mformula" bool "ts list" | MSince args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'msaux" | MUntil args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'muaux" | MMatchP \ "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data mr\aux" | MMatchF \ "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data ml\aux" record ('msaux, 'muaux) mstate = mstate_i :: nat mstate_m :: "('msaux, 'muaux) mformula" mstate_n :: nat fun eq_rel :: "nat \ Formula.trm \ Formula.trm \ event_data table" where "eq_rel n (Formula.Const x) (Formula.Const y) = (if x = y then unit_table n else empty_table)" | "eq_rel n (Formula.Var x) (Formula.Const y) = singleton_table n x y" | "eq_rel n (Formula.Const x) (Formula.Var y) = singleton_table n y x" | "eq_rel n _ _ = undefined" lemma regex_atms_size: "x \ regex.atms r \ size x < regex.size_regex size r" by (induct r) auto lemma atms_size: assumes "x \ atms r" shows "size x < Regex.size_regex size r" proof - { fix y assume "y \ regex.atms r" "case y of formula.Neg z \ x = z | _ \ False" then have "size x < Regex.size_regex size r" by (cases y rule: formula.exhaust) (auto dest: regex_atms_size) } with assms show ?thesis unfolding atms_def by (auto split: formula.splits dest: regex_atms_size) qed definition init_args :: "\ \ nat \ nat set \ nat set \ bool \ args" where "init_args I n L R pos = \args_ivl = I, args_n = n, args_L = L, args_R = R, args_pos = pos\" locale msaux = fixes valid_msaux :: "args \ ts \ 'msaux \ event_data msaux \ bool" and init_msaux :: "args \ 'msaux" and add_new_ts_msaux :: "args \ ts \ 'msaux \ 'msaux" and join_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and add_new_table_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and result_msaux :: "args \ 'msaux \ event_data table" assumes valid_init_msaux: "L \ R \ valid_msaux (init_args I n L R pos) 0 (init_msaux (init_args I n L R pos)) []" assumes valid_add_new_ts_msaux: "valid_msaux args cur aux auxlist \ nt \ cur \ valid_msaux args nt (add_new_ts_msaux args nt aux) (filter (\(t, rel). enat (nt - t) \ right (args_ivl args)) auxlist)" assumes valid_join_msaux: "valid_msaux args cur aux auxlist \ table (args_n args) (args_L args) rel1 \ valid_msaux args cur (join_msaux args rel1 aux) (map (\(t, rel). (t, join rel (args_pos args) rel1)) auxlist)" assumes valid_add_new_table_msaux: "valid_msaux args cur aux auxlist \ table (args_n args) (args_R args) rel2 \ valid_msaux args cur (add_new_table_msaux args rel2 aux) (case auxlist of [] => [(cur, rel2)] | ((t, y) # ts) \ if t = cur then (t, y \ rel2) # ts else (cur, rel2) # auxlist)" and valid_result_msaux: "valid_msaux args cur aux auxlist \ result_msaux args aux = foldr (\) [rel. (t, rel) \ auxlist, left (args_ivl args) \ cur - t] {}" fun check_before :: "\ \ ts \ (ts \ 'a \ 'b) \ bool" where "check_before I dt (t, a, b) \ enat t + right I < enat dt" fun proj_thd :: "('a \ 'b \ 'c) \ 'c" where "proj_thd (t, a1, a2) = a2" definition update_until :: "args \ event_data table \ event_data table \ ts \ event_data muaux \ event_data muaux" where "update_until args rel1 rel2 nt aux = (map (\x. case x of (t, a1, a2) \ (t, if (args_pos args) then join a1 True rel1 else a1 \ rel1, if mem (nt - t) (args_ivl args) then a2 \ join rel2 (args_pos args) a1 else a2)) aux) @ [(nt, rel1, if left (args_ivl args) = 0 then rel2 else empty_table)]" lemma map_proj_thd_update_until: "map proj_thd (takeWhile (check_before (args_ivl args) nt) auxlist) = map proj_thd (takeWhile (check_before (args_ivl args) nt) (update_until args rel1 rel2 nt auxlist))" proof (induction auxlist) case Nil then show ?case by (simp add: update_until_def) next case (Cons a auxlist) then show ?case by (cases "right (args_ivl args)") (auto simp add: update_until_def split: if_splits prod.splits) qed fun eval_until :: "\ \ ts \ event_data muaux \ event_data table list \ event_data muaux" where "eval_until I nt [] = ([], [])" | "eval_until I nt ((t, a1, a2) # aux) = (if t + right I < nt then (let (xs, aux) = eval_until I nt aux in (a2 # xs, aux)) else ([], (t, a1, a2) # aux))" lemma eval_until_length: "eval_until I nt auxlist = (res, auxlist') \ length auxlist = length res + length auxlist'" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: if_splits prod.splits) lemma eval_until_res: "eval_until I nt auxlist = (res, auxlist') \ res = map proj_thd (takeWhile (check_before I nt) auxlist)" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: prod.splits) lemma eval_until_auxlist': "eval_until I nt auxlist = (res, auxlist') \ auxlist' = drop (length res) auxlist" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: if_splits prod.splits) locale muaux = fixes valid_muaux :: "args \ ts \ 'muaux \ event_data muaux \ bool" and init_muaux :: "args \ 'muaux" and add_new_muaux :: "args \ event_data table \ event_data table \ ts \ 'muaux \ 'muaux" and length_muaux :: "args \ 'muaux \ nat" and eval_muaux :: "args \ ts \ 'muaux \ event_data table list \ 'muaux" assumes valid_init_muaux: "L \ R \ valid_muaux (init_args I n L R pos) 0 (init_muaux (init_args I n L R pos)) []" assumes valid_add_new_muaux: "valid_muaux args cur aux auxlist \ table (args_n args) (args_L args) rel1 \ table (args_n args) (args_R args) rel2 \ nt \ cur \ valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux) (update_until args rel1 rel2 nt auxlist)" assumes valid_length_muaux: "valid_muaux args cur aux auxlist \ length_muaux args aux = length auxlist" assumes valid_eval_muaux: "valid_muaux args cur aux auxlist \ nt \ cur \ eval_muaux args nt aux = (res, aux') \ eval_until (args_ivl args) nt auxlist = (res', auxlist') \ res = res' \ valid_muaux args cur aux' auxlist'" locale maux = msaux valid_msaux init_msaux add_new_ts_msaux join_msaux add_new_table_msaux result_msaux + muaux valid_muaux init_muaux add_new_muaux length_muaux eval_muaux for valid_msaux :: "args \ ts \ 'msaux \ event_data msaux \ bool" and init_msaux :: "args \ 'msaux" and add_new_ts_msaux :: "args \ ts \ 'msaux \ 'msaux" and join_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and add_new_table_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and result_msaux :: "args \ 'msaux \ event_data table" and valid_muaux :: "args \ ts \ 'muaux \ event_data muaux \ bool" and init_muaux :: "args \ 'muaux" and add_new_muaux :: "args \ event_data table \ event_data table \ ts \ 'muaux \ 'muaux" and length_muaux :: "args \ 'muaux \ nat" and eval_muaux :: "args \ nat \ 'muaux \ event_data table list \ 'muaux" fun split_assignment :: "nat set \ Formula.formula \ nat \ Formula.trm" where "split_assignment X (Formula.Eq t1 t2) = (case (t1, t2) of (Formula.Var x, Formula.Var y) \ if x \ X then (y, t1) else (x, t2) | (Formula.Var x, _) \ (x, t2) | (_, Formula.Var y) \ (y, t1))" | "split_assignment _ _ = undefined" fun split_constraint :: "Formula.formula \ Formula.trm \ bool \ mconstraint \ Formula.trm" where "split_constraint (Formula.Eq t1 t2) = (t1, True, MEq, t2)" | "split_constraint (Formula.Less t1 t2) = (t1, True, MLess, t2)" | "split_constraint (Formula.LessEq t1 t2) = (t1, True, MLessEq, t2)" | "split_constraint (Formula.Neg (Formula.Eq t1 t2)) = (t1, False, MEq, t2)" | "split_constraint (Formula.Neg (Formula.Less t1 t2)) = (t1, False, MLess, t2)" | "split_constraint (Formula.Neg (Formula.LessEq t1 t2)) = (t1, False, MLessEq, t2)" | "split_constraint _ = undefined" function (in maux) (sequential) minit0 :: "nat \ Formula.formula \ ('msaux, 'muaux) mformula" where "minit0 n (Formula.Neg \) = (if fv \ = {} then MNeg (minit0 n \) else MRel empty_table)" | "minit0 n (Formula.Eq t1 t2) = MRel (eq_rel n t1 t2)" | "minit0 n (Formula.Pred e ts) = MPred e ts" | "minit0 n (Formula.Let p \ \) = MLet p (Formula.nfv \) (minit0 (Formula.nfv \) \) (minit0 n \)" | "minit0 n (Formula.Or \ \) = MOr (minit0 n \) (minit0 n \) ([], [])" | "minit0 n (Formula.And \ \) = (if safe_assignment (fv \) \ then MAndAssign (minit0 n \) (split_assignment (fv \) \) else if safe_formula \ then MAnd (fv \) (minit0 n \) True (fv \) (minit0 n \) ([], []) else if is_constraint \ then MAndRel (minit0 n \) (split_constraint \) else (case \ of Formula.Neg \ \ MAnd (fv \) (minit0 n \) False (fv \) (minit0 n \) ([], [])))" | "minit0 n (Formula.Ands l) = (let (pos, neg) = partition safe_formula l in let mpos = map (minit0 n) pos in let mneg = map (minit0 n) (map remove_neg neg) in let vpos = map fv pos in let vneg = map fv neg in MAnds vpos vneg (mpos @ mneg) (replicate (length l) []))" | "minit0 n (Formula.Exists \) = MExists (minit0 (Suc n) \)" | "minit0 n (Formula.Agg y \ b f \) = MAgg (fv \ \ {0.. b f (minit0 (b + n) \)" | "minit0 n (Formula.Prev I \) = MPrev I (minit0 n \) True [] []" | "minit0 n (Formula.Next I \) = MNext I (minit0 n \) True []" | "minit0 n (Formula.Since \ I \) = (if safe_formula \ then MSince (init_args I n (Formula.fv \) (Formula.fv \) True) (minit0 n \) (minit0 n \) ([], []) [] (init_msaux (init_args I n (Formula.fv \) (Formula.fv \) True)) else (case \ of Formula.Neg \ \ MSince (init_args I n (Formula.fv \) (Formula.fv \) False) (minit0 n \) (minit0 n \) ([], []) [] (init_msaux (init_args I n (Formula.fv \) (Formula.fv \) False)) | _ \ undefined))" | "minit0 n (Formula.Until \ I \) = (if safe_formula \ then MUntil (init_args I n (Formula.fv \) (Formula.fv \) True) (minit0 n \) (minit0 n \) ([], []) [] (init_muaux (init_args I n (Formula.fv \) (Formula.fv \) True)) else (case \ of Formula.Neg \ \ MUntil (init_args I n (Formula.fv \) (Formula.fv \) False) (minit0 n \) (minit0 n \) ([], []) [] (init_muaux (init_args I n (Formula.fv \) (Formula.fv \) False)) | _ \ undefined))" | "minit0 n (Formula.MatchP I r) = (let (mr, \s) = to_mregex r in MMatchP I mr (sorted_list_of_set (RPDs mr)) (map (minit0 n) \s) (replicate (length \s) []) [] [])" | "minit0 n (Formula.MatchF I r) = (let (mr, \s) = to_mregex r in MMatchF I mr (sorted_list_of_set (LPDs mr)) (map (minit0 n) \s) (replicate (length \s) []) [] [])" | "minit0 n _ = undefined" by pat_completeness auto termination (in maux) by (relation "measure (\(_, \). size \)") (auto simp: less_Suc_eq_le size_list_estimation' size_remove_neg dest!: to_mregex_ok[OF sym] atms_size) definition (in maux) minit :: "Formula.formula \ ('msaux, 'muaux) mstate" where "minit \ = (let n = Formula.nfv \ in \mstate_i = 0, mstate_m = minit0 n \, mstate_n = n\)" definition (in maux) minit_safe where "minit_safe \ = (if mmonitorable_exec \ then minit \ else undefined)" fun mprev_next :: "\ \ event_data table list \ ts list \ event_data table list \ event_data table list \ ts list" where "mprev_next I [] ts = ([], [], ts)" | "mprev_next I xs [] = ([], xs, [])" | "mprev_next I xs [t] = ([], xs, [t])" | "mprev_next I (x # xs) (t # t' # ts) = (let (ys, zs) = mprev_next I xs (t' # ts) in ((if mem (t' - t) I then x else empty_table) # ys, zs))" fun mbuf2_add :: "event_data table list \ event_data table list \ event_data mbuf2 \ event_data mbuf2" where "mbuf2_add xs' ys' (xs, ys) = (xs @ xs', ys @ ys')" fun mbuf2_take :: "(event_data table \ event_data table \ 'b) \ event_data mbuf2 \ 'b list \ event_data mbuf2" where "mbuf2_take f (x # xs, y # ys) = (let (zs, buf) = mbuf2_take f (xs, ys) in (f x y # zs, buf))" | "mbuf2_take f (xs, ys) = ([], (xs, ys))" fun mbuf2t_take :: "(event_data table \ event_data table \ ts \ 'b \ 'b) \ 'b \ event_data mbuf2 \ ts list \ 'b \ event_data mbuf2 \ ts list" where "mbuf2t_take f z (x # xs, y # ys) (t # ts) = mbuf2t_take f (f x y t z) (xs, ys) ts" | "mbuf2t_take f z (xs, ys) ts = (z, (xs, ys), ts)" lemma size_list_length_diff1: "xs \ [] \ [] \ set xs \ size_list (\xs. length xs - Suc 0) xs < size_list length xs" proof (induct xs) case (Cons x xs) then show ?case by (cases xs) auto qed simp fun mbufn_add :: "event_data table list list \ event_data mbufn \ event_data mbufn" where "mbufn_add xs' xs = List.map2 (@) xs xs'" function mbufn_take :: "(event_data table list \ 'b \ 'b) \ 'b \ event_data mbufn \ 'b \ event_data mbufn" where "mbufn_take f z buf = (if buf = [] \ [] \ set buf then (z, buf) else mbufn_take f (f (map hd buf) z) (map tl buf))" by pat_completeness auto termination by (relation "measure (\(_, _, buf). size_list length buf)") (auto simp: comp_def Suc_le_eq size_list_length_diff1) fun mbufnt_take :: "(event_data table list \ ts \ 'b \ 'b) \ 'b \ event_data mbufn \ ts list \ 'b \ event_data mbufn \ ts list" where "mbufnt_take f z buf ts = (if [] \ set buf \ ts = [] then (z, buf, ts) else mbufnt_take f (f (map hd buf) (hd ts) z) (map tl buf) (tl ts))" fun match :: "Formula.trm list \ event_data list \ (nat \ event_data) option" where "match [] [] = Some Map.empty" | "match (Formula.Const x # ts) (y # ys) = (if x = y then match ts ys else None)" | "match (Formula.Var x # ts) (y # ys) = (case match ts ys of None \ None | Some f \ (case f x of None \ Some (f(x \ y)) | Some z \ if y = z then Some f else None))" | "match _ _ = None" fun meval_trm :: "Formula.trm \ event_data tuple \ event_data" where "meval_trm (Formula.Var x) v = the (v ! x)" | "meval_trm (Formula.Const x) v = x" | "meval_trm (Formula.Plus x y) v = meval_trm x v + meval_trm y v" | "meval_trm (Formula.Minus x y) v = meval_trm x v - meval_trm y v" | "meval_trm (Formula.UMinus x) v = - meval_trm x v" | "meval_trm (Formula.Mult x y) v = meval_trm x v * meval_trm y v" | "meval_trm (Formula.Div x y) v = meval_trm x v div meval_trm y v" | "meval_trm (Formula.Mod x y) v = meval_trm x v mod meval_trm y v" | "meval_trm (Formula.F2i x) v = EInt (integer_of_event_data (meval_trm x v))" | "meval_trm (Formula.I2f x) v = EFloat (double_of_event_data (meval_trm x v))" definition eval_agg :: "nat \ bool \ nat \ Formula.agg_op \ nat \ Formula.trm \ event_data table \ event_data table" where "eval_agg n g0 y \ b f rel = (if g0 \ rel = empty_table then singleton_table n y (eval_agg_op \ {}) else (\k. let group = Set.filter (\x. drop b x = k) rel; M = (\y. (y, ecard (Set.filter (\x. meval_trm f x = y) group))) ` meval_trm f ` group in k[y:=Some (eval_agg_op \ M)]) ` (drop b) ` rel)" definition (in maux) update_since :: "args \ event_data table \ event_data table \ ts \ 'msaux \ event_data table \ 'msaux" where "update_since args rel1 rel2 nt aux = (let aux0 = join_msaux args rel1 (add_new_ts_msaux args nt aux); aux' = add_new_table_msaux args rel2 aux0 in (result_msaux args aux', aux'))" definition "lookup = Mapping.lookup_default empty_table" fun \_lax where "\_lax guard \s (MSkip n) = (if n = 0 then guard else empty_table)" | "\_lax guard \s (MTestPos i) = join guard True (\s ! i)" | "\_lax guard \s (MTestNeg i) = join guard False (\s ! i)" | "\_lax guard \s (MPlus r s) = \_lax guard \s r \ \_lax guard \s s" | "\_lax guard \s (MTimes r s) = join (\_lax guard \s r) True (\_lax guard \s s)" | "\_lax guard \s (MStar r) = guard" fun r\_strict where "r\_strict n \s (MSkip m) = (if m = 0 then unit_table n else empty_table)" | "r\_strict n \s (MTestPos i) = \s ! i" | "r\_strict n \s (MTestNeg i) = (if \s ! i = empty_table then unit_table n else empty_table)" | "r\_strict n \s (MPlus r s) = r\_strict n \s r \ r\_strict n \s s" | "r\_strict n \s (MTimes r s) = \_lax (r\_strict n \s r) \s s" | "r\_strict n \s (MStar r) = unit_table n" fun l\_strict where "l\_strict n \s (MSkip m) = (if m = 0 then unit_table n else empty_table)" | "l\_strict n \s (MTestPos i) = \s ! i" | "l\_strict n \s (MTestNeg i) = (if \s ! i = empty_table then unit_table n else empty_table)" | "l\_strict n \s (MPlus r s) = l\_strict n \s r \ l\_strict n \s s" | "l\_strict n \s (MTimes r s) = \_lax (l\_strict n \s s) \s r" | "l\_strict n \s (MStar r) = unit_table n" fun r\ :: "(mregex \ mregex) \ (mregex, 'a table) mapping \ 'a table list \ mregex \ 'a table" where "r\ \ X \s (MSkip n) = (case n of 0 \ empty_table | Suc m \ lookup X (\ (MSkip m)))" | "r\ \ X \s (MTestPos i) = empty_table" | "r\ \ X \s (MTestNeg i) = empty_table" | "r\ \ X \s (MPlus r s) = r\ \ X \s r \ r\ \ X \s s" | "r\ \ X \s (MTimes r s) = r\ (\t. \ (MTimes r t)) X \s s \ \_lax (r\ \ X \s r) \s s" | "r\ \ X \s (MStar r) = r\ (\t. \ (MTimes (MStar r) t)) X \s r" fun l\ :: "(mregex \ mregex) \ (mregex, 'a table) mapping \ 'a table list \ mregex \ 'a table" where "l\ \ X \s (MSkip n) = (case n of 0 \ empty_table | Suc m \ lookup X (\ (MSkip m)))" | "l\ \ X \s (MTestPos i) = empty_table" | "l\ \ X \s (MTestNeg i) = empty_table" | "l\ \ X \s (MPlus r s) = l\ \ X \s r \ l\ \ X \s s" | "l\ \ X \s (MTimes r s) = l\ (\t. \ (MTimes t s)) X \s r \ \_lax (l\ \ X \s s) \s r" | "l\ \ X \s (MStar r) = l\ (\t. \ (MTimes t (MStar r))) X \s r" lift_definition mrtabulate :: "mregex list \ (mregex \ 'b table) \ (mregex, 'b table) mapping" is "\ks f. (map_of (List.map_filter (\k. let fk = f k in if fk = empty_table then None else Some (k, fk)) ks))" . lemma lookup_tabulate: "distinct xs \ lookup (mrtabulate xs f) x = (if x \ set xs then f x else empty_table)" unfolding lookup_default_def lookup_def by transfer (auto simp: Let_def map_filter_def map_of_eq_None_iff o_def image_image dest!: map_of_SomeD split: if_splits option.splits) definition update_matchP :: "nat \ \ \ mregex \ mregex list \ event_data table list \ ts \ event_data mr\aux \ event_data table \ event_data mr\aux" where "update_matchP n I mr mrs rels nt aux = (let aux = (case [(t, mrtabulate mrs (\mr. r\ id rel rels mr \ (if t = nt then r\_strict n rels mr else {}))). (t, rel) \ aux, enat (nt - t) \ right I] of [] \ [(nt, mrtabulate mrs (r\_strict n rels))] | x # aux' \ (if fst x = nt then x # aux' else (nt, mrtabulate mrs (r\_strict n rels)) # x # aux')) in (foldr (\) [lookup rel mr. (t, rel) \ aux, left I \ nt - t] {}, aux))" definition update_matchF_base where "update_matchF_base n I mr mrs rels nt = (let X = mrtabulate mrs (l\_strict n rels) in ([(nt, rels, if left I = 0 then lookup X mr else empty_table)], X))" definition update_matchF_step where "update_matchF_step I mr mrs nt = (\(t, rels', rel) (aux', X). (let Y = mrtabulate mrs (l\ id X rels') in ((t, rels', if mem (nt - t) I then rel \ lookup Y mr else rel) # aux', Y)))" definition update_matchF :: "nat \ \ \ mregex \ mregex list \ event_data table list \ ts \ event_data ml\aux \ event_data ml\aux" where "update_matchF n I mr mrs rels nt aux = fst (foldr (update_matchF_step I mr mrs nt) aux (update_matchF_base n I mr mrs rels nt))" fun eval_matchF :: "\ \ mregex \ ts \ event_data ml\aux \ event_data table list \ event_data ml\aux" where "eval_matchF I mr nt [] = ([], [])" | "eval_matchF I mr nt ((t, rels, rel) # aux) = (if t + right I < nt then (let (xs, aux) = eval_matchF I mr nt aux in (rel # xs, aux)) else ([], (t, rels, rel) # aux))" primrec map_split where "map_split f [] = ([], [])" | "map_split f (x # xs) = (let (y, z) = f x; (ys, zs) = map_split f xs in (y # ys, z # zs))" fun eval_assignment :: "nat \ Formula.trm \ event_data tuple \ event_data tuple" where "eval_assignment (x, t) y = (y[x:=Some (meval_trm t y)])" fun eval_constraint0 :: "mconstraint \ event_data \ event_data \ bool" where "eval_constraint0 MEq x y = (x = y)" | "eval_constraint0 MLess x y = (x < y)" | "eval_constraint0 MLessEq x y = (x \ y)" fun eval_constraint :: "Formula.trm \ bool \ mconstraint \ Formula.trm \ event_data tuple \ bool" where "eval_constraint (t1, p, c, t2) x = (eval_constraint0 c (meval_trm t1 x) (meval_trm t2 x) = p)" primrec (in maux) meval :: "nat \ ts \ Formula.database \ ('msaux, 'muaux) mformula \ event_data table list \ ('msaux, 'muaux) mformula" where "meval n t db (MRel rel) = ([rel], MRel rel)" | "meval n t db (MPred e ts) = (map (\X. (\f. Table.tabulate f 0 n) ` Option.these (match ts ` X)) (case Mapping.lookup db e of None \ [{}] | Some xs \ xs), MPred e ts)" | "meval n t db (MLet p m \ \) = (let (xs, \) = meval m t db \; (ys, \) = meval n t (Mapping.update p (map (image (map the)) xs) db) \ in (ys, MLet p m \ \))" | "meval n t db (MAnd A_\ \ pos A_\ \ buf) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (zs, buf) = mbuf2_take (\r1 r2. bin_join n A_\ r1 pos A_\ r2) (mbuf2_add xs ys buf) in (zs, MAnd A_\ \ pos A_\ \ buf))" | "meval n t db (MAndAssign \ conf) = (let (xs, \) = meval n t db \ in (map (\r. eval_assignment conf ` r) xs, MAndAssign \ conf))" | "meval n t db (MAndRel \ conf) = (let (xs, \) = meval n t db \ in (map (Set.filter (eval_constraint conf)) xs, MAndRel \ conf))" | "meval n t db (MAnds A_pos A_neg L buf) = (let R = map (meval n t db) L in let buf = mbufn_add (map fst R) buf in let (zs, buf) = mbufn_take (\xs zs. zs @ [mmulti_join n A_pos A_neg xs]) [] buf in (zs, MAnds A_pos A_neg (map snd R) buf))" | "meval n t db (MOr \ \ buf) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (zs, buf) = mbuf2_take (\r1 r2. r1 \ r2) (mbuf2_add xs ys buf) in (zs, MOr \ \ buf))" | "meval n t db (MNeg \) = (let (xs, \) = meval n t db \ in (map (\r. (if r = empty_table then unit_table n else empty_table)) xs, MNeg \))" | "meval n t db (MExists \) = (let (xs, \) = meval (Suc n) t db \ in (map (\r. tl ` r) xs, MExists \))" | "meval n t db (MAgg g0 y \ b f \) = (let (xs, \) = meval (b + n) t db \ in (map (eval_agg n g0 y \ b f) xs, MAgg g0 y \ b f \))" | "meval n t db (MPrev I \ first buf nts) = (let (xs, \) = meval n t db \; (zs, buf, nts) = mprev_next I (buf @ xs) (nts @ [t]) in (if first then empty_table # zs else zs, MPrev I \ False buf nts))" | "meval n t db (MNext I \ first nts) = (let (xs, \) = meval n t db \; (xs, first) = (case (xs, first) of (_ # xs, True) \ (xs, False) | a \ a); (zs, _, nts) = mprev_next I xs (nts @ [t]) in (zs, MNext I \ first nts))" | "meval n t db (MSince args \ \ buf nts aux) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; ((zs, aux), buf, nts) = mbuf2t_take (\r1 r2 t (zs, aux). let (z, aux) = update_since args r1 r2 t aux in (zs @ [z], aux)) ([], aux) (mbuf2_add xs ys buf) (nts @ [t]) in (zs, MSince args \ \ buf nts aux))" | "meval n t db (MUntil args \ \ buf nts aux) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (aux, buf, nts) = mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [t]); (zs, aux) = eval_muaux args (case nts of [] \ t | nt # _ \ nt) aux in (zs, MUntil args \ \ buf nts aux))" | "meval n t db (MMatchP I mr mrs \s buf nts aux) = (let (xss, \s) = map_split id (map (meval n t db) \s); ((zs, aux), buf, nts) = mbufnt_take (\rels t (zs, aux). let (z, aux) = update_matchP n I mr mrs rels t aux in (zs @ [z], aux)) ([], aux) (mbufn_add xss buf) (nts @ [t]) in (zs, MMatchP I mr mrs \s buf nts aux))" | "meval n t db (MMatchF I mr mrs \s buf nts aux) = (let (xss, \s) = map_split id (map (meval n t db) \s); (aux, buf, nts) = mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [t]); (zs, aux) = eval_matchF I mr (case nts of [] \ t | nt # _ \ nt) aux in (zs, MMatchF I mr mrs \s buf nts aux))" definition (in maux) mstep :: "Formula.database \ ts \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list \ ('msaux, 'muaux) mstate" where "mstep tdb st = (let (xs, m) = meval (mstate_n st) (snd tdb) (fst tdb) (mstate_m st) in (List.enumerate (mstate_i st) xs, \mstate_i = mstate_i st + length xs, mstate_m = m, mstate_n = mstate_n st\))" subsection \Verdict delay\ context fixes \ :: Formula.trace begin fun progress :: "(Formula.name \ nat) \ Formula.formula \ nat \ nat" where "progress P (Formula.Pred e ts) j = (case P e of None \ j | Some k \ k)" | "progress P (Formula.Let p \ \) j = progress (P(p \ progress P \ j)) \ j" | "progress P (Formula.Eq t1 t2) j = j" | "progress P (Formula.Less t1 t2) j = j" | "progress P (Formula.LessEq t1 t2) j = j" | "progress P (Formula.Neg \) j = progress P \ j" | "progress P (Formula.Or \ \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.And \ \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.Ands l) j = (if l = [] then j else Min (set (map (\\. progress P \ j) l)))" | "progress P (Formula.Exists \) j = progress P \ j" | "progress P (Formula.Agg y \ b f \) j = progress P \ j" | "progress P (Formula.Prev I \) j = (if j = 0 then 0 else min (Suc (progress P \ j)) j)" | "progress P (Formula.Next I \) j = progress P \ j - 1" | "progress P (Formula.Since \ I \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.Until \ I \) j = Inf {i. \k. k < j \ k \ min (progress P \ j) (progress P \ j) \ \ \ i + right I \ \ \ k}" | "progress P (Formula.MatchP I r) j = min_regex_default (progress P) r j" | "progress P (Formula.MatchF I r) j = Inf {i. \k. k < j \ k \ min_regex_default (progress P) r j \ \ \ i + right I \ \ \ k}" definition "progress_regex P = min_regex_default (progress P)" declare progress.simps[simp del] lemmas progress_simps[simp] = progress.simps[folded progress_regex_def[THEN fun_cong, THEN fun_cong]] end definition "pred_mapping Q = pred_fun (\_. True) (pred_option Q)" definition "rel_mapping Q = rel_fun (=) (rel_option Q)" lemma pred_mapping_alt: "pred_mapping Q P = (\p \ dom P. Q (the (P p)))" unfolding pred_mapping_def pred_fun_def option.pred_set dom_def by (force split: option.splits) lemma rel_mapping_alt: "rel_mapping Q P P' = (dom P = dom P' \ (\p \ dom P. Q (the (P p)) (the (P' p))))" unfolding rel_mapping_def rel_fun_def rel_option_iff dom_def by (force split: option.splits) lemma rel_mapping_map_upd[simp]: "Q x y \ rel_mapping Q P P' \ rel_mapping Q (P(p \ x)) (P'(p \ y))" by (auto simp: rel_mapping_alt) lemma pred_mapping_map_upd[simp]: "Q x \ pred_mapping Q P \ pred_mapping Q (P(p \ x))" by (auto simp: pred_mapping_alt) lemma pred_mapping_empty[simp]: "pred_mapping Q Map.empty" by (auto simp: pred_mapping_alt) lemma pred_mapping_mono: "pred_mapping Q P \ Q \ R \ pred_mapping R P" by (auto simp: pred_mapping_alt) lemma pred_mapping_mono_strong: "pred_mapping Q P \ (\p. p \ dom P \ Q (the (P p)) \ R (the (P p))) \ pred_mapping R P" by (auto simp: pred_mapping_alt) lemma progress_mono_gen: "j \ j' \ rel_mapping (\) P P' \ progress \ P \ j \ progress \ P' \ j'" proof (induction \ arbitrary: P P') case (Pred e ts) then show ?case by (force simp: rel_mapping_alt dom_def split: option.splits) next case (Ands l) then show ?case by (auto simp: image_iff intro!: Min.coboundedI[THEN order_trans]) next case (Until \ I \) from Until(1,2)[of P P'] Until.prems show ?case by (cases "right I") (auto dest: trans_le_add1[OF \_mono] intro!: cInf_superset_mono) next case (MatchF I r) from MatchF(1)[of _ P P'] MatchF.prems show ?case by (cases "right I"; cases "regex.atms r = {}") (auto 0 3 simp: Min_le_iff progress_regex_def dest: trans_le_add1[OF \_mono] intro!: cInf_superset_mono elim!: less_le_trans order_trans) qed (force simp: Min_le_iff progress_regex_def split: option.splits)+ lemma rel_mapping_reflp: "reflp Q \ rel_mapping Q P P" by (auto simp: rel_mapping_alt reflp_def) lemmas progress_mono = progress_mono_gen[OF _ rel_mapping_reflp[unfolded reflp_def], simplified] lemma progress_le_gen: "pred_mapping (\x. x \ j) P \ progress \ P \ j \ j" proof (induction \ arbitrary: P) case (Pred e ts) then show ?case by (auto simp: pred_mapping_alt dom_def split: option.splits) next case (Ands l) then show ?case by (auto simp: image_iff intro!: Min.coboundedI[where a="progress \ P (hd l) j", THEN order_trans]) next case (Until \ I \) then show ?case by (cases "right I") (auto intro: trans_le_add1[OF \_mono] intro!: cInf_lower) next case (MatchF I r) then show ?case by (cases "right I") (auto intro: trans_le_add1[OF \_mono] intro!: cInf_lower) qed (force simp: Min_le_iff progress_regex_def split: option.splits)+ lemma progress_le: "progress \ Map.empty \ j \ j" using progress_le_gen[of _ Map.empty] by auto lemma progress_0_gen[simp]: "pred_mapping (\x. x = 0) P \ progress \ P \ 0 = 0" using progress_le_gen[of 0 P] by auto lemma progress_0[simp]: "progress \ Map.empty \ 0 = 0" by (auto simp: pred_mapping_alt) definition max_mapping :: "('b \ 'a option) \ ('b \ 'a option) \ 'b \ ('a :: linorder) option" where "max_mapping P P' x = (case (P x, P' x) of (None, None) \ None | (Some x, None) \ None | (None, Some x) \ None | (Some x, Some y) \ Some (max x y))" definition Max_mapping :: "('b \ 'a option) set \ 'b \ ('a :: linorder) option" where "Max_mapping Ps x = (if (\P \ Ps. P x \ None) then Some (Max ((\P. the (P x)) ` Ps)) else None)" lemma dom_max_mapping[simp]: "dom (max_mapping P1 P2) = dom P1 \ dom P2" unfolding max_mapping_def by (auto split: option.splits) lemma dom_Max_mapping[simp]: "dom (Max_mapping X) = (\P \ X. dom P)" unfolding Max_mapping_def by (auto split: if_splits) lemma Max_mapping_coboundedI: assumes "finite X" "\Q \ X. dom Q = dom P" "P \ X" shows "rel_mapping (\) P (Max_mapping X)" unfolding rel_mapping_alt proof (intro conjI ballI) from assms(3) have "X \ {}" by auto then show "dom P = dom (Max_mapping X)" using assms(2) by auto next fix p assume "p \ dom P" with assms show "the (P p) \ the (Max_mapping X p)" by (force simp add: Max_mapping_def intro!: Max.coboundedI imageI) qed lemma rel_mapping_trans: "P OO Q \ R \ rel_mapping P P1 P2 \ rel_mapping Q P2 P3 \ rel_mapping R P1 P3" by (force simp: rel_mapping_alt dom_def set_eq_iff) abbreviation range_mapping :: "nat \ nat \ ('b \ nat option) \ bool" where "range_mapping i j P \ pred_mapping (\x. i \ x \ x \ j) P" lemma range_mapping_relax: "range_mapping i j P \ i' \ i \ j' \ j \ range_mapping i' j' P" by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits) lemma range_mapping_max_mapping[simp]: "range_mapping i j1 P1 \ range_mapping i j2 P2 \ range_mapping i (max j1 j2) (max_mapping P1 P2)" by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits) lemma range_mapping_Max_mapping[simp]: "finite X \ X \ {} \ \x\X. range_mapping i (j x) (P x) \ range_mapping i (Max (j ` X)) (Max_mapping (P ` X))" by (force simp: pred_mapping_alt Max_mapping_def dom_def image_iff intro!: Max_ge_iff[THEN iffD2] split: if_splits) lemma pred_mapping_le: "pred_mapping ((\) i) P1 \ rel_mapping (\) P1 P2 \ pred_mapping ((\) (i :: nat)) P2" by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff) lemma pred_mapping_le': "pred_mapping ((\) j) P1 \ i \ j \ rel_mapping (\) P1 P2 \ pred_mapping ((\) (i :: nat)) P2" by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff) lemma max_mapping_cobounded1: "dom P1 \ dom P2 \ rel_mapping (\) P1 (max_mapping P1 P2)" unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits) lemma max_mapping_cobounded2: "dom P2 \ dom P1 \ rel_mapping (\) P2 (max_mapping P1 P2)" unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits) lemma max_mapping_fun_upd2[simp]: "(max_mapping P1 (P2(p := y)))(p \ x) = (max_mapping P1 P2)(p \ x)" by (auto simp: max_mapping_def) lemma rel_mapping_max_mapping_fun_upd: "dom P2 \ dom P1 \ p \ dom P2 \ the (P2 p) \ y \ rel_mapping (\) P2 ((max_mapping P1 P2)(p \ y))" by (auto simp: rel_mapping_alt max_mapping_def split: option.splits) lemma progress_ge_gen: "Formula.future_bounded \ \ \P j. dom P = S \ range_mapping i j P \ i \ progress \ P \ j" proof (induction \ arbitrary: i S) case (Pred e ts) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: option.splits if_splits simp: rel_mapping_alt pred_mapping_alt dom_def) next case (Let p \ \) from Let.prems obtain P2 j2 where P2: "dom P2 = insert p S" "range_mapping i j2 P2" "i \ progress \ P2 \ j2" by (atomize_elim, intro Let(2)) (force simp: pred_mapping_alt rel_mapping_alt dom_def)+ from Let.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (the (P2 p)) j1 P1" "the (P2 p) \ progress \ P1 \ j1" by (atomize_elim, intro Let(1)) auto let ?P12 = "max_mapping P1 P2" from P1 P2 have le1: "progress \ P1 \ j1 \ progress \ (?P12(p := P1 p)) \ (max j1 j2)" by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def) from P1 P2 have le2: "progress \ P2 \ j2 \ progress \ (?P12(p \ progress \ P1 \ j1)) \ (max j1 j2)" by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def) show ?case unfolding progress.simps proof (intro exI[of _ "?P12(p := P1 p)"] exI[of _ "max j1 j2"] conjI) show "dom (?P12(p := P1 p)) = S" using P1 P2 by (auto simp: dom_def max_mapping_def) next show "range_mapping i (max j1 j2) (?P12(p := P1 p))" using P1 P2 by (force simp add: pred_mapping_alt dom_def max_mapping_def split: option.splits) next have "i \ progress \ P2 \ j2" by fact also have "... \ progress \ (?P12(p \ progress \ P1 \ j1)) \ (max j1 j2)" using le2 by blast also have "... \ progress \ ((?P12(p := P1 p))(p\progress \ (?P12(p := P1 p)) \ (max j1 j2))) \ (max j1 j2)" by (auto intro!: progress_mono_gen simp: le1 rel_mapping_alt) finally show "i \ ..." . qed next case (Eq _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (Less _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (LessEq _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (Or \1 \2) from Or(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using Or(1)[of S i] by auto moreover from Or(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using Or(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.Or \1 \2) (max j1 j2)" by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto next case (And \1 \2) from And(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using And(1)[of S i] by auto moreover from And(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using And(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.And \1 \2) (max j1 j2)" by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto next case (Ands l) show ?case proof (cases "l = []") case True then show ?thesis by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case False then obtain \ where "\ \ set l" by (cases l) auto from Ands.prems have "\\\set l. Formula.future_bounded \" by (simp add: list.pred_set) { fix \ assume "\ \ set l" with Ands.prems obtain P j where "dom P = S" "range_mapping i j P" "i \ progress \ P \ j" by (atomize_elim, intro Ands(1)[of \ S i]) (auto simp: list.pred_set) then have "\Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)" (is "\Pj. ?P Pj") by (intro exI[of _ "(P, j)"]) auto } then have "\\\set l. \Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)" (is "\\\set l. \Pj. ?P Pj \") by blast with Ands(1) Ands.prems False have "\Pjf. \\\set l. ?P (Pjf \) \" by (auto simp: Ball_def intro: choice) then obtain Pjf where Pjf: "\\\set l. ?P (Pjf \) \" .. define Pf where "Pf = fst o Pjf" define jf where "jf = snd o Pjf" have *: "dom (Pf \) = S" "range_mapping i (jf \) (Pf \)" "i \ progress \ (Pf \) \ (jf \)" if "\ \ set l" for \ using Pjf[THEN bspec, OF that] unfolding Pf_def jf_def by auto with False show ?thesis unfolding progress.simps eq_False[THEN iffD2, OF False] if_False by ((subst Min_ge_iff; simp add: False), intro exI[where x="MAX \\set l. jf \"] exI[where x="Max_mapping (Pf ` set l)"] conjI ballI order.trans[OF *(3) progress_mono_gen] Max_mapping_coboundedI) (auto simp: False *[OF \\ \ set l\] \\ \ set l\) qed next case (Exists \) then show ?case by simp next case (Prev I \) then obtain P j where "dom P = S" "range_mapping i j P" "i \ progress \ P \ j" by (atomize_elim, intro Prev(1)) (auto simp: pred_mapping_alt dom_def) with Prev(2) have "dom P = S \ range_mapping i (max i j) P \ i \ progress \ P (formula.Prev I \) (max i j)" by (auto simp: le_Suc_eq max_def pred_mapping_alt split: if_splits elim: order.trans[OF _ progress_mono]) then show ?case by blast next case (Next I \) then obtain P j where "dom P = S" "range_mapping (Suc i) j P" "Suc i \ progress \ P \ j" by (atomize_elim, intro Next(1)) (auto simp: pred_mapping_alt dom_def) then show ?case by (intro exI[of _ P] exI[of _ j]) (auto 0 3 simp: pred_mapping_alt dom_def) next case (Since \1 I \2) from Since(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using Since(1)[of S i] by auto moreover from Since(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using Since(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.Since \1 I \2) (max j1 j2)" by (auto elim!: order.trans[OF _ progress_mono_gen] simp: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) (auto elim!: pred_mapping_le intro: max_mapping_cobounded1) next case (Until \1 I \2) from Until.prems obtain b where [simp]: "right I = enat b" by (cases "right I") (auto) obtain i' where "i < i'" and "\ \ i + b + 1 \ \ \ i'" using ex_le_\[where x="\ \ i + b + 1"] by (auto simp add: less_eq_Suc_le) then have 1: "\ \ i + b < \ \ i'" by simp from Until.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (Suc i') j1 P1" "Suc i' \ progress \ P1 \1 j1" by (atomize_elim, intro Until(1)) (auto simp: pred_mapping_alt dom_def) from Until.prems obtain P2 j2 where P2: "dom P2 = S" "range_mapping (Suc i') j2 P2" "Suc i' \ progress \ P2 \2 j2" by (atomize_elim, intro Until(2)) (auto simp: pred_mapping_alt dom_def) let ?P12 = "max_mapping P1 P2" have "i \ progress \ ?P12 (Formula.Until \1 I \2) (max j1 j2)" unfolding progress.simps proof (intro cInf_greatest, goal_cases nonempty greatest) case nonempty then show ?case by (auto simp: trans_le_add1[OF \_mono] intro!: exI[of _ "max j1 j2"]) next case (greatest x) from P1(2,3) have "i' < j1" by (auto simp: less_eq_Suc_le intro!: progress_le_gen elim!: order.trans pred_mapping_mono_strong) then have "i' < max j1 j2" by simp have "progress \ P1 \1 j1 \ progress \ ?P12 \1 (max j1 j2)" using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded1) moreover have "progress \ P2 \2 j2 \ progress \ ?P12 \2 (max j1 j2)" using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded2) ultimately have "i' \ min (progress \ ?P12 \1 (max j1 j2)) (progress \ ?P12 \2 (max j1 j2))" using P1(3) P2(3) by simp with greatest \i' < max j1 j2\ have "\ \ i' \ \ \ x + b" by simp with 1 have "\ \ i < \ \ x" by simp then show ?case by (auto dest!: less_\D) qed with P1 P2 \i < i'\ show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) (auto simp: range_mapping_relax) next case (MatchP I r) then show ?case proof (cases "regex.atms r = {}") case True with MatchP.prems show ?thesis unfolding progress.simps by (intro exI[of _ "\e. if e \ S then Some i else None"] exI[of _ i]) (auto split: if_splits simp: pred_mapping_alt regex.pred_set) next case False define pick where "pick = (\\. SOME Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj))" let ?pickP = "fst o pick" let ?pickj = "snd o pick" from MatchP have pick: "\ \ regex.atms r \ dom (?pickP \) = S \ range_mapping i (?pickj \) (?pickP \) \ i \ progress \ (?pickP \) \ (?pickj \)" for \ unfolding pick_def o_def future_bounded.simps regex.pred_set by (intro someI_ex[where P = "\Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)"]) auto with False show ?thesis unfolding progress.simps by (intro exI[of _ "Max_mapping (?pickP ` regex.atms r)"] exI[of _ "Max (?pickj ` regex.atms r)"]) (auto simp: Max_mapping_coboundedI order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen]) qed next case (MatchF I r) from MatchF.prems obtain b where [simp]: "right I = enat b" by auto obtain i' where i': "i < i'" "\ \ i + b + 1 \ \ \ i'" using ex_le_\[where x="\ \ i + b + 1"] by (auto simp add: less_eq_Suc_le) then have 1: "\ \ i + b < \ \ i'" by simp have ix: "i \ x" if "\ \ i' \ b + \ \ x" "b + \ \ i < \ \ i'" for x using less_\D[of \ i] that less_le_trans by fastforce show ?case proof (cases "regex.atms r = {}") case True with MatchF.prems i' ix show ?thesis unfolding progress.simps by (intro exI[of _ "\e. if e \ S then Some (Suc i') else None"] exI[of _ "Suc i'"]) (auto split: if_splits simp: pred_mapping_alt regex.pred_set add.commute less_Suc_eq intro!: cInf_greatest dest!: spec[of _ i'] less_imp_le[THEN \_mono, of _ i' \]) next case False then obtain \ where \: "\ \ regex.atms r" by auto define pick where "pick = (\\. SOME Pj. dom (fst Pj) = S \ range_mapping (Suc i') (snd Pj) (fst Pj) \ Suc i' \ progress \ (fst Pj) \ (snd Pj))" define pickP where "pickP = fst o pick" define pickj where "pickj = snd o pick" from MatchF have pick: "\ \ regex.atms r \ dom (pickP \) = S \ range_mapping (Suc i') (pickj \) (pickP \) \ Suc i' \ progress \ (pickP \) \ (pickj \)" for \ unfolding pick_def o_def future_bounded.simps regex.pred_set pickj_def pickP_def by (intro someI_ex[where P = "\Pj. dom (fst Pj) = S \ range_mapping (Suc i') (snd Pj) (fst Pj) \ Suc i' \ progress \ (fst Pj) \ (snd Pj)"]) auto let ?P = "Max_mapping (pickP ` regex.atms r)" let ?j = "Max (pickj ` regex.atms r)" from pick[OF \] False \ have "Suc i' \ ?j" by (intro order_trans[OF pick[THEN conjunct2, THEN conjunct2], OF \] order_trans[OF progress_le_gen]) (auto simp: Max_ge_iff dest: range_mapping_relax[of _ _ _ 0, OF _ _ order_refl, simplified]) moreover note i' 1 ix moreover from MatchF.prems have "Regex.pred_regex Formula.future_bounded r" by auto ultimately show ?thesis using \_mono[of _ ?j \] less_\D[of \ i] pick False by (intro exI[of _ "?j"] exI[of _ "?P"]) (auto 0 3 intro!: cInf_greatest order_trans[OF le_SucI[OF order_refl] order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen]] range_mapping_Max_mapping[OF _ _ ballI[OF range_mapping_relax[of "Suc i'" _ _ i, OF _ _ order_refl]]] simp: ac_simps Suc_le_eq trans_le_add2 Max_mapping_coboundedI progress_regex_def dest: spec[of _ "i'"] spec[of _ ?j]) qed qed (auto split: option.splits) lemma progress_ge: "Formula.future_bounded \ \ \j. i \ progress \ Map.empty \ j" using progress_ge_gen[of \ "{}" i \] by auto lemma cInf_restrict_nat: fixes x :: nat assumes "x \ A" shows "Inf A = Inf {y \ A. y \ x}" using assms by (auto intro!: antisym intro: cInf_greatest cInf_lower Inf_nat_def1) lemma progress_time_conv: assumes "\i \ i = \ \' i" shows "progress \ P \ j = progress \' P \ j" using assms proof (induction \ arbitrary: P) case (Until \1 I \2) have *: "i \ j - 1 \ i < j" if "j \ 0" for i using that by auto with Until show ?case proof (cases "right I") case (enat b) then show ?thesis proof (cases "j") case (Suc n) with enat * Until show ?thesis using \_mono[THEN trans_le_add1] by (auto 8 0 intro!: box_equals[OF arg_cong[where f=Inf] cInf_restrict_nat[symmetric, where x=n] cInf_restrict_nat[symmetric, where x=n]]) qed simp qed simp next case (MatchF I r) have *: "i \ j - 1 \ i < j" if "j \ 0" for i using that by auto with MatchF show ?case using \_mono[THEN trans_le_add1] by (cases "right I"; cases j) ((auto 6 0 simp: progress_le_gen progress_regex_def intro!: box_equals[OF arg_cong[where f=Inf] cInf_restrict_nat[symmetric, where x="j-1"] cInf_restrict_nat[symmetric, where x="j-1"]]) [])+ qed (auto simp: progress_regex_def) lemma Inf_UNIV_nat: "(Inf UNIV :: nat) = 0" by (simp add: cInf_eq_minimum) lemma progress_prefix_conv: assumes "prefix_of \ \" and "prefix_of \ \'" shows "progress \ P \ (plen \) = progress \' P \ (plen \)" using assms by (auto intro: progress_time_conv \_prefix_conv) lemma bounded_rtranclp_mono: fixes n :: "'x :: linorder" assumes "\i j. R i j \ j < n \ S i j" "\i j. R i j \ i \ j" shows "rtranclp R i j \ j < n \ rtranclp S i j" proof (induct rule: rtranclp_induct) case (step y z) then show ?case using assms(1,2)[of y z] by (auto elim!: rtrancl_into_rtrancl[to_pred, rotated]) qed auto lemma sat_prefix_conv_gen: assumes "prefix_of \ \" and "prefix_of \ \'" shows "i < progress \ P \ (plen \) \ dom V = dom V' \ dom P = dom V \ pred_mapping (\x. x \ plen \) P \ (\p i \. p \ dom V \ i < the (P p) \ the (V p) i = the (V' p) i) \ Formula.sat \ V v i \ \ Formula.sat \' V' v i \" proof (induction \ arbitrary: P V V' v i) case (Pred e ts) from Pred.prems(1,4) have "i < plen \" by (blast intro!: order.strict_trans2 progress_le_gen) show ?case proof (cases "V e") case None then have "V' e = None" using \dom V = dom V'\ by auto with None \_prefix_conv[OF assms(1,2) \i < plen \\] show ?thesis by simp next case (Some a) obtain a' where "V' e = Some a'" using Some \dom V = dom V'\ by auto then have "i < the (P e)" using Pred.prems(1-3) by (auto split: option.splits) then have "the (V e) i = the (V' e) i" using Some by (intro Pred.prems(5)) (simp_all add: domI) with Some \V' e = Some a'\ show ?thesis by simp qed next case (Let p \ \) let ?V = "\V \. (V(p \ \i. {v. length v = Formula.nfv \ \ Formula.sat \ V v i \}))" show ?case unfolding sat.simps proof (rule Let.IH(2)) from Let.prems show "i < progress \ (P(p \ progress \ P \ (plen \))) \ (plen \)" by simp from Let.prems show "dom (?V V \) = dom (?V V' \')" by simp from Let.prems show "dom (P(p \ progress \ P \ (plen \))) = dom (?V V \)" by simp from Let.prems show "pred_mapping (\x. x \ plen \) (P(p \ progress \ P \ (plen \)))" by (auto intro!: pred_mapping_map_upd elim!: progress_le_gen) next fix p' i \' assume 1: "p' \ dom (?V V \)" and 2: "i < the ((P(p \ progress \ P \ (plen \))) p')" show "the (?V V \ p') i = the (?V V' \' p') i" proof (cases "p' = p") case True with Let 2 show ?thesis by auto next case False with 1 2 show ?thesis by (auto intro!: Let.prems(5)) qed qed next case (Eq t1 t2) show ?case by simp next case (Neg \) then show ?case by simp next case (Or \1 \2) then show ?case by auto next case (Ands l) from Ands.prems have "\\\set l. i < progress \ P \ (plen \)" by (cases l) simp_all with Ands show ?case unfolding sat_Ands by blast next case (Exists \) then show ?case by simp next case (Prev I \) with \_prefix_conv[OF assms(1,2)] show ?case by (cases i) (auto split: if_splits) next case (Next I \) then have "Suc i < plen \" by (auto intro: order.strict_trans2[OF _ progress_le_gen[of _ P \ \]]) with Next.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro conj_cong Next) auto next case (Since \1 I \2) then have "i < plen \" by (auto elim!: order.strict_trans2[OF _ progress_le_gen]) with Since.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro conj_cong ex_cong ball_cong Since) auto next case (Until \1 I \2) from Until.prems obtain b where right[simp]: "right I = enat b" by (cases "right I") (auto simp add: Inf_UNIV_nat) from Until.prems obtain j where "\ \ i + b + 1 \ \ \ j" "j \ progress \ P \1 (plen \)" "j \ progress \ P \2 (plen \)" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"] dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "k < progress \ P \1 (plen \)" and 2: "k < progress \ P \2 (plen \)" if "\ \ k \ \ \ i + b" for k using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \])+ have 3: "k < plen \" if "\ \ k \ \ \ i + b" for k using 1[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen]) from Until.prems have "i < progress \' P (Formula.Until \1 I \2) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast then obtain j where "\ \' i + b + 1 \ \ \' j" "j \ progress \' P \1 (plen \)" "j \ progress \' P \2 (plen \)" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"] dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 11: "k < progress \ P \1 (plen \)" and 21: "k < progress \ P \2 (plen \)" if "\ \' k \ \ \' i + b" for k unfolding progress_prefix_conv[OF assms(1,2)] using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \'])+ have 31: "k < plen \" if "\ \' k \ \ \' i + b" for k using 11[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen]) show ?case unfolding sat.simps proof ((intro ex_cong iffI; elim conjE), goal_cases LR RL) case (LR j) with Until(1)[OF 1] Until(2)[OF 2] \_prefix_conv[OF assms(1,2) 3] Until.prems show ?case by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF \_mono, rotated]) next case (RL j) with Until(1)[OF 11] Until(2)[OF 21] \_prefix_conv[OF assms(1,2) 31] Until.prems show ?case by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF \_mono, rotated]) qed next case (MatchP I r) then have "i < plen \" by (force simp: pred_mapping_alt elim!: order.strict_trans2[OF _ progress_le_gen]) with MatchP.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro ex_cong conj_cong match_cong_strong MatchP) (auto simp: progress_regex_def split: if_splits) next case (MatchF I r) from MatchF.prems obtain b where right[simp]: "right I = enat b" by (cases "right I") (auto simp add: Inf_UNIV_nat) show ?case proof (cases "regex.atms r = {}") case True from MatchF.prems(1) obtain j where "\ \ i + b + 1 \ \ \ j" "j \ plen \" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "k < plen \" if "\ \ k \ \ \ i + b" for k - by (meson \_mono discrete not_le order.strict_trans2 that) + using that le_less_trans [of \\ \ k\ _ \\ \ j\] less_\D [of \ k j] by simp from MatchF.prems have "i < progress \' P (Formula.MatchF I r) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast then obtain j where "\ \' i + b + 1 \ \ \' j" "j \ plen \" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 2: "k < plen \" if "\ \' k \ \ \' i + b" for k - by (meson \_mono discrete not_le order.strict_trans2 that) + using that le_less_trans [of \\ \' k\ _ \\ \' j\] less_\D [of \' k j] by simp from MatchF.prems(1,4) True show ?thesis unfolding sat.simps conj_commute[of "left I \ _" "_ \ _"] proof (intro ex_cong conj_cong match_cong_strong, goal_cases left right sat) case (left j) then show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 1, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF \_mono, rotated])) next case (right j) then show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF \_mono, rotated])) qed auto next case False from MatchF.prems(1) False obtain j where "\ \ i + b + 1 \ \ \ j" "(\x\regex.atms r. j \ progress \ P x (plen \))" by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "\ \ regex.atms r \ k < progress \ P \ (plen \)" if "\ \ k \ \ \ i + b" for k \ using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \]) then have 2: "k < plen \" if "\ \ k \ \ \ i + b" "regex.atms r \ {}" for k using that by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of \], of k]) from MatchF.prems have "i < progress \' P (Formula.MatchF I r) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast with False obtain j where "\ \' i + b + 1 \ \ \' j" "(\x\regex.atms r. j \ progress \' P x (plen \))" by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 11: "\ \ regex.atms r \ k < progress \ P \ (plen \)" if "\ \' k \ \ \' i + b" for k \ using that using progress_prefix_conv[OF assms(1,2)] by (auto 0 3 elim!: order.strict_trans2[rotated] intro: less_\D[of \']) have 21: "k < plen \" if "\ \' k \ \ \' i + b" for k using 11[OF that(1)] False by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of \], of k]) show ?thesis unfolding sat.simps conj_commute[of "left I \ _" "_ \ _"] proof ((intro ex_cong conj_cong match_cong_strong MatchF(1)[OF _ _ MatchF(3-6)]; assumption?), goal_cases right left progress) case (right j) with False show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF \_mono, rotated])) next case (left j) with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b] by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 21, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF \_mono, rotated])) next case (progress j k z) with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b] by (elim 1[rotated]) (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim!: order.trans[OF \_mono, rotated]) qed qed qed auto lemma sat_prefix_conv: assumes "prefix_of \ \" and "prefix_of \ \'" shows "i < progress \ Map.empty \ (plen \) \ Formula.sat \ Map.empty v i \ \ Formula.sat \' Map.empty v i \" by (erule sat_prefix_conv_gen[OF assms]) auto lemma progress_remove_neg[simp]: "progress \ P (remove_neg \) j = progress \ P \ j" by (cases \) simp_all lemma safe_progress_get_and: "safe_formula \ \ Min ((\\. progress \ P \ j) ` set (get_and_list \)) = progress \ P \ j" by (induction \ rule: get_and_list.induct) auto lemma progress_convert_multiway: "safe_formula \ \ progress \ P (convert_multiway \) j = progress \ P \ j" proof (induction \ arbitrary: P rule: safe_formula_induct) case (And_safe \ \) let ?c = "convert_multiway (Formula.And \ \)" let ?c\ = "convert_multiway \" let ?c\ = "convert_multiway \" have c_eq: "?c = Formula.Ands (get_and_list ?c\ @ get_and_list ?c\)" using And_safe by simp from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) moreover from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) ultimately show ?case unfolding c_eq using And_safe.IH by (auto simp: get_and_nonempty Min.union safe_progress_get_and) next case (And_Not \ \) let ?c = "convert_multiway (Formula.And \ (Formula.Neg \))" let ?c\ = "convert_multiway \" let ?c\ = "convert_multiway \" have c_eq: "?c = Formula.Ands (Formula.Neg ?c\ # get_and_list ?c\)" using And_Not by simp from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) moreover from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) ultimately show ?case unfolding c_eq using And_Not.IH by (auto simp: get_and_nonempty Min.union safe_progress_get_and) next case (MatchP I r) from MatchP show ?case unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image by (intro if_cong arg_cong[of _ _ Min] image_cong) (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula) next case (MatchF I r) from MatchF show ?case unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image by (intro if_cong arg_cong[of _ _ Min] arg_cong[of _ _ Inf] arg_cong[of _ _ "(\) _"] image_cong Collect_cong all_cong1 imp_cong conj_cong image_cong) (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula) qed auto subsection \Specification\ definition pprogress :: "Formula.formula \ Formula.prefix \ nat" where "pprogress \ \ = (THE n. \\. prefix_of \ \ \ progress \ Map.empty \ (plen \) = n)" lemma pprogress_eq: "prefix_of \ \ \ pprogress \ \ = progress \ Map.empty \ (plen \)" unfolding pprogress_def using progress_prefix_conv by blast locale future_bounded_mfodl = fixes \ :: Formula.formula assumes future_bounded: "Formula.future_bounded \" sublocale future_bounded_mfodl \ sliceable_timed_progress "Formula.nfv \" "Formula.fv \" "relevant_events \" "\\ v i. Formula.sat \ Map.empty v i \" "pprogress \" proof (unfold_locales, goal_cases) case (1 x) then show ?case by (simp add: fvi_less_nfv) next case (2 v v' \ i) then show ?case by (simp cong: sat_fv_cong[rule_format]) next case (3 v S \ i) then show ?case using sat_slice_iff[symmetric] by simp next case (4 \ \') moreover obtain \ where "prefix_of \' \" using ex_prefix_of .. moreover have "prefix_of \ \" using prefix_of_antimono[OF \\ \ \'\ \prefix_of \' \\] . ultimately show ?case by (simp add: pprogress_eq plen_mono progress_mono) next case (5 \ x) obtain j where "x \ progress \ Map.empty \ j" using future_bounded progress_ge by blast then have "x \ pprogress \ (take_prefix j \)" by (simp add: pprogress_eq[of _ \]) then show ?case by force next case (6 \ \ \' i v) then have "i < progress \ Map.empty \ (plen \)" by (simp add: pprogress_eq) with 6 show ?case using sat_prefix_conv by blast next case (7 \ \') then have "plen \ = plen \'" by transfer (simp add: list_eq_iff_nth_eq) moreover obtain \ \' where "prefix_of \ \" "prefix_of \' \'" using ex_prefix_of by blast+ moreover have "\i < plen \. \ \ i = \ \' i" using 7 calculation by transfer (simp add: list_eq_iff_nth_eq) ultimately show ?case by (simp add: pprogress_eq progress_time_conv) qed locale verimon_spec = fixes \ :: Formula.formula assumes monitorable: "mmonitorable \" sublocale verimon_spec \ future_bounded_mfodl using monitorable by unfold_locales (simp add: mmonitorable_def) subsection \Correctness\ subsubsection \Invariants\ definition wf_mbuf2 :: "nat \ nat \ nat \ (nat \ event_data table \ bool) \ (nat \ event_data table \ bool) \ event_data mbuf2 \ bool" where "wf_mbuf2 i ja jb P Q buf \ i \ ja \ i \ jb \ (case buf of (xs, ys) \ list_all2 P [i.. list_all2 Q [i.. 'b \ 'c \ bool) \ 'a list \ 'b list \ 'c list \ bool" for P :: "('a \ 'b \ 'c \ bool)" where "list_all3 P [] [] []" | "P a1 a2 a3 \ list_all3 P q1 q2 q3 \ list_all3 P (a1 # q1) (a2 # q2) (a3 # q3)" lemma list_all3_list_all2D: "list_all3 P xs ys zs \ (length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs)" by (induct xs ys zs rule: list_all3.induct) auto lemma list_all2_list_all3I: "length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs \ list_all3 P xs ys zs" by (induct xs ys arbitrary: zs rule: list_induct2) (auto simp: list_all2_Cons1 intro: list_all3.intros) lemma list_all3_list_all2_eq: "list_all3 P xs ys zs \ (length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs)" using list_all2_list_all3I list_all3_list_all2D by blast lemma list_all3_mapD: "list_all3 P (map f xs) (map g ys) (map h zs) \ list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs" by (induct "map f xs" "map g ys" "map h zs" arbitrary: xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) lemma list_all3_mapI: "list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs \ list_all3 P (map f xs) (map g ys) (map h zs)" by (induct xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) lemma list_all3_map_iff: "list_all3 P (map f xs) (map g ys) (map h zs) \ list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs" using list_all3_mapD list_all3_mapI by blast lemmas list_all3_map = list_all3_map_iff[where g=id and h=id, unfolded list.map_id id_apply] list_all3_map_iff[where f=id and h=id, unfolded list.map_id id_apply] list_all3_map_iff[where f=id and g=id, unfolded list.map_id id_apply] lemma list_all3_conv_all_nth: "list_all3 P xs ys zs = (length xs = length ys \ length ys = length zs \ (\i < length xs. P (xs!i) (ys!i) (zs!i)))" by (auto simp add: list_all3_list_all2_eq list_all2_conv_all_nth) lemma list_all3_refl [intro?]: "(\x. x \ set xs \ P x x x) \ list_all3 P xs xs xs" by (simp add: list_all3_conv_all_nth) definition wf_mbufn :: "nat \ nat list \ (nat \ event_data table \ bool) list \ event_data mbufn \ bool" where "wf_mbufn i js Ps buf \ list_all3 (\P j xs. i \ j \ list_all2 P [i.. _ \ _ \ nat \ nat \ event_data list set \ Formula.formula \ Formula.formula \ event_data mbuf2 \ bool" where "wf_mbuf2' \ P V j n R \ \ buf \ wf_mbuf2 (min (progress \ P \ j) (progress \ P \ j)) (progress \ P \ j) (progress \ P \ j) (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) buf" definition wf_mbufn' :: "Formula.trace \ _ \ _ \ nat \ nat \ event_data list set \ Formula.formula Regex.regex \ event_data mbufn \ bool" where "wf_mbufn' \ P V j n R r buf \ (case to_mregex r of (mr, \s) \ wf_mbufn (progress_regex \ P r j) (map (\\. progress \ P \ j) \s) (map (\\ i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) buf)" lemma wf_mbuf2'_UNIV_alt: "wf_mbuf2' \ P V j n UNIV \ \ buf \ (case buf of (xs, ys) \ list_all2 (\i. wf_table n (Formula.fv \) (\v. Formula.sat \ V (map the v) i \)) [min (progress \ P \ j) (progress \ P \ j) ..< (progress \ P \ j)] xs \ list_all2 (\i. wf_table n (Formula.fv \) (\v. Formula.sat \ V (map the v) i \)) [min (progress \ P \ j) (progress \ P \ j) ..< (progress \ P \ j)] ys)" unfolding wf_mbuf2'_def wf_mbuf2_def by (simp add: mem_restr_UNIV[THEN eqTrueI, abs_def] split: prod.split) definition wf_ts :: "Formula.trace \ _ \ nat \ Formula.formula \ Formula.formula \ ts list \ bool" where "wf_ts \ P j \ \ ts \ list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j).. _ \ nat \ Formula.formula Regex.regex \ ts list \ bool" where "wf_ts_regex \ P j r ts \ list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. I \ \ Formula.Since (if pos then \ else Formula.Neg \) I \" definition (in msaux) wf_since_aux :: "Formula.trace \ _ \ event_data list set \ args \ Formula.formula \ Formula.formula \ 'msaux \ nat \ bool" where "wf_since_aux \ V R args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_msaux args cur aux auxlist \ cur = (if ne = 0 then 0 else \ \ (ne - 1)) \ sorted_wrt (\x y. fst x > fst y) auxlist \ (\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ qtable (args_n args) (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne-1) (Sincep (args_pos args) \ (point (\ \ (ne - 1) - t)) \)) X) \ (\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)))" definition wf_matchP_aux :: "Formula.trace \ _ \ nat \ event_data list set \ \ \ Formula.formula Regex.regex \ event_data mr\aux \ nat \ bool" where "wf_matchP_aux \ V n R I r aux ne \ sorted_wrt (\x y. fst x > fst y) aux \ (\t X. (t, X) \ set aux \ ne \ 0 \ t \ \ \ (ne-1) \ \ \ (ne-1) - t \ right I \ (\i. \ \ i = t) \ (case to_mregex r of (mr, \s) \ (\ms \ RPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne-1) (Formula.MatchP (point (\ \ (ne-1) - t)) (from_mregex ms \s))) (lookup X ms)))) \ (\t. ne \ 0 \ t \ \ \ (ne-1) \ \ \ (ne-1) - t \ right I \ (\i. \ \ i = t) \ (\X. (t, X) \ set aux))" lemma qtable_mem_restr_UNIV: "qtable n A(mem_restr UNIV) Q X = wf_table n A Q X" unfolding qtable_def by auto lemma (in msaux) wf_since_aux_UNIV_alt: "wf_since_aux \ V UNIV args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_msaux args cur aux auxlist \ cur = (if ne = 0 then 0 else \ \ (ne - 1)) \ sorted_wrt (\x y. fst x > fst y) auxlist \ (\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ wf_table (args_n args) (Formula.fv \) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep (args_pos args) \ (point (\ \ (ne - 1) - t)) \)) X) \ (\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)))" unfolding wf_since_aux_def qtable_mem_restr_UNIV .. definition wf_until_auxlist :: "Formula.trace \ _ \ nat \ event_data list set \ bool \ Formula.formula \ \ \ Formula.formula \ event_data muaux \ nat \ bool" where "wf_until_auxlist \ V n R pos \ I \ auxlist ne \ list_all2 (\x i. case x of (t, r1, r2) \ t = \ \ i \ qtable n (Formula.fv \) (mem_restr R) (\v. if pos then (\k\{i.. V (map the v) k \) else (\k\{i.. V (map the v) k \)) r1 \ qtable n (Formula.fv \) (mem_restr R) (\v. (\j. i \ j \ j < ne + length auxlist \ mem (\ \ j - \ \ i) I \ Formula.sat \ V (map the v) j \ \ (\k\{i.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \))) r2) auxlist [ne.. _ \ event_data list set \ args \ Formula.formula \ Formula.formula \ 'muaux \ nat \ bool" where "wf_until_aux \ V R args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_muaux args cur aux auxlist \ cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1)) \ wf_until_auxlist \ V (args_n args) R (args_pos args) \ (args_ivl args) \ auxlist ne)" lemma (in muaux) wf_until_aux_UNIV_alt: "wf_until_aux \ V UNIV args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_muaux args cur aux auxlist \ cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1)) \ list_all2 (\x i. case x of (t, r1, r2) \ t = \ \ i \ wf_table (args_n args) (Formula.fv \) (\v. if (args_pos args) then (\k\{i.. V (map the v) k \) else (\k\{i.. V (map the v) k \)) r1 \ wf_table (args_n args) (Formula.fv \) (\v. \j. i \ j \ j < ne + length auxlist \ mem (\ \ j - \ \ i) (args_ivl args) \ Formula.sat \ V (map the v) j \ \ (\k\{i.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \)) r2) auxlist [ne.. _ \ nat \ event_data list set \ \ \ Formula.formula Regex.regex \ event_data ml\aux \ nat \ nat \ bool" where "wf_matchF_aux \ V n R I r aux ne k \ (case to_mregex r of (mr, \s) \ list_all2 (\x i. case x of (t, rels, rel) \ t = \ \ i \ list_all2 (\\. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s rels \ qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. i \ j \ j < ne + length aux + k \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel) aux [ne.. V n R I r st i = (case st of (aux, Y) \ aux \ [] \ wf_matchF_aux \ V n R I r aux i 0 \ (case to_mregex r of (mr, \s) \ \ms \ LPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Regex.match (Formula.sat \ V (map the v)) (from_mregex ms \s) i (i + length aux - 1)) (lookup Y ms)))" definition lift_envs' :: "nat \ event_data list set \ event_data list set" where "lift_envs' b R = (\(xs,ys). xs @ ys) ` ({xs. length xs = b} \ R)" fun formula_of_constraint :: "Formula.trm \ bool \ mconstraint \ Formula.trm \ Formula.formula" where "formula_of_constraint (t1, True, MEq, t2) = Formula.Eq t1 t2" | "formula_of_constraint (t1, True, MLess, t2) = Formula.Less t1 t2" | "formula_of_constraint (t1, True, MLessEq, t2) = Formula.LessEq t1 t2" | "formula_of_constraint (t1, False, MEq, t2) = Formula.Neg (Formula.Eq t1 t2)" | "formula_of_constraint (t1, False, MLess, t2) = Formula.Neg (Formula.Less t1 t2)" | "formula_of_constraint (t1, False, MLessEq, t2) = Formula.Neg (Formula.LessEq t1 t2)" inductive (in maux) wf_mformula :: "Formula.trace \ nat \ _ \ _ \ nat \ event_data list set \ ('msaux, 'muaux) mformula \ Formula.formula \ bool" for \ j where Eq: "is_simple_eq t1 t2 \ \x\Formula.fv_trm t1. x < n \ \x\Formula.fv_trm t2. x < n \ wf_mformula \ j P V n R (MRel (eq_rel n t1 t2)) (Formula.Eq t1 t2)" | neq_Var: "x < n \ wf_mformula \ j P V n R (MRel empty_table) (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var x)))" | Pred: "\x\Formula.fv (Formula.Pred e ts). x < n \ \t\set ts. Formula.is_Var t \ Formula.is_Const t \ wf_mformula \ j P V n R (MPred e ts) (Formula.Pred e ts)" | Let: "wf_mformula \ j P V m UNIV \ \' \ wf_mformula \ j (P(p \ progress \ P \' j)) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \'})) n R \ \' \ {0.. Formula.fv \' \ b \ m \ m = Formula.nfv \' \ wf_mformula \ j P V n R (MLet p m \ \) (Formula.Let p \' \')" | And: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if pos then \ = Formula.And \' \' else \ = Formula.And \' (Formula.Neg \') \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_mformula \ j P V n R (MAnd (fv \') \ pos (fv \') \ buf) \" | AndAssign: "wf_mformula \ j P V n R \ \' \ x < n \ x \ Formula.fv \' \ Formula.fv_trm t \ Formula.fv \' \ (x, t) = conf \ \' = Formula.Eq (Formula.Var x) t \ \' = Formula.Eq t (Formula.Var x) \ wf_mformula \ j P V n R (MAndAssign \ conf) (Formula.And \' \')" | AndRel: "wf_mformula \ j P V n R \ \' \ \' = formula_of_constraint conf \ (let (t1, _, _, t2) = conf in Formula.fv_trm t1 \ Formula.fv_trm t2 \ Formula.fv \') \ wf_mformula \ j P V n R (MAndRel \ conf) (Formula.And \' \')" | Ands: "list_all2 (\\ \'. wf_mformula \ j P V n R \ \') l (l_pos @ map remove_neg l_neg) \ wf_mbufn (progress \ P (Formula.Ands l') j) (map (\\. progress \ P \ j) (l_pos @ map remove_neg l_neg)) (map (\\ i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (l_pos @ map remove_neg l_neg)) buf \ (l_pos, l_neg) = partition safe_formula l' \ l_pos \ [] \ list_all safe_formula (map remove_neg l_neg) \ A_pos = map fv l_pos \ A_neg = map fv l_neg \ \(set A_neg) \ \(set A_pos) \ wf_mformula \ j P V n R (MAnds A_pos A_neg l buf) (Formula.Ands l')" | Or: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ Formula.fv \' = Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_mformula \ j P V n R (MOr \ \ buf) (Formula.Or \' \')" | Neg: "wf_mformula \ j P V n R \ \' \ Formula.fv \' = {} \ wf_mformula \ j P V n R (MNeg \) (Formula.Neg \')" | Exists: "wf_mformula \ j P V (Suc n) (lift_envs R) \ \' \ wf_mformula \ j P V n R (MExists \) (Formula.Exists \')" | Agg: "wf_mformula \ j P V (b + n) (lift_envs' b R) \ \' \ y < n \ y + b \ Formula.fv \' \ {0.. Formula.fv \' \ Formula.fv_trm f \ Formula.fv \' \ g0 = (Formula.fv \' \ {0.. wf_mformula \ j P V n R (MAgg g0 y \ b f \) (Formula.Agg y \ b f \')" | Prev: "wf_mformula \ j P V n R \ \' \ first \ j = 0 \ list_all2 (\i. qtable n (Formula.fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [min (progress \ P \' j) (j-1).. P \' j] buf \ list_all2 (\i t. t = \ \ i) [min (progress \ P \' j) (j-1).. wf_mformula \ j P V n R (MPrev I \ first buf nts) (Formula.Prev I \')" | Next: "wf_mformula \ j P V n R \ \' \ first \ progress \ P \' j = 0 \ list_all2 (\i t. t = \ \ i) [progress \ P \' j - 1.. wf_mformula \ j P V n R (MNext I \ first nts) (Formula.Next I \')" | Since: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if args_pos args then \'' = \' else \'' = Formula.Neg \' \ safe_formula \'' = args_pos args \ args_ivl args = I \ args_n args = n \ args_L args = Formula.fv \' \ args_R args = Formula.fv \' \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_ts \ P j \' \' nts \ wf_since_aux \ V R args \' \' aux (progress \ P (Formula.Since \'' I \') j) \ wf_mformula \ j P V n R (MSince args \ \ buf nts aux) (Formula.Since \'' I \')" | Until: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if args_pos args then \'' = \' else \'' = Formula.Neg \' \ safe_formula \'' = args_pos args \ args_ivl args = I \ args_n args = n \ args_L args = Formula.fv \' \ args_R args = Formula.fv \' \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_ts \ P j \' \' nts \ wf_until_aux \ V R args \' \' aux (progress \ P (Formula.Until \'' I \') j) \ progress \ P (Formula.Until \'' I \') j + length_muaux args aux = min (progress \ P \' j) (progress \ P \' j) \ wf_mformula \ j P V n R (MUntil args \ \ buf nts aux) (Formula.Until \'' I \')" | MatchP: "(case to_mregex r of (mr', \s') \ list_all2 (wf_mformula \ j P V n R) \s \s' \ mr = mr') \ mrs = sorted_list_of_set (RPDs mr) \ safe_regex Past Strict r \ wf_mbufn' \ P V j n R r buf \ wf_ts_regex \ P j r nts \ wf_matchP_aux \ V n R I r aux (progress \ P (Formula.MatchP I r) j) \ wf_mformula \ j P V n R (MMatchP I mr mrs \s buf nts aux) (Formula.MatchP I r)" | MatchF: "(case to_mregex r of (mr', \s') \ list_all2 (wf_mformula \ j P V n R) \s \s' \ mr = mr') \ mrs = sorted_list_of_set (LPDs mr) \ safe_regex Futu Strict r \ wf_mbufn' \ P V j n R r buf \ wf_ts_regex \ P j r nts \ wf_matchF_aux \ V n R I r aux (progress \ P (Formula.MatchF I r) j) 0 \ progress \ P (Formula.MatchF I r) j + length aux = progress_regex \ P r j \ wf_mformula \ j P V n R (MMatchF I mr mrs \s buf nts aux) (Formula.MatchF I r)" definition (in maux) wf_mstate :: "Formula.formula \ Formula.prefix \ event_data list set \ ('msaux, 'muaux) mstate \ bool" where "wf_mstate \ \ R st \ mstate_n st = Formula.nfv \ \ (\\. prefix_of \ \ \ mstate_i st = progress \ Map.empty \ (plen \) \ wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \)" subsubsection \Initialisation\ lemma wf_mbuf2'_0: "pred_mapping (\x. x = 0) P \ wf_mbuf2' \ P V 0 n R \ \ ([], [])" unfolding wf_mbuf2'_def wf_mbuf2_def by simp lemma wf_mbufn'_0: "to_mregex r = (mr, \s) \ pred_mapping (\x. x = 0) P \ wf_mbufn' \ P V 0 n R r (replicate (length \s) [])" unfolding wf_mbufn'_def wf_mbufn_def map_replicate_const[symmetric] by (auto simp: list_all3_map intro: list_all3_refl simp: Min_eq_iff progress_regex_def) lemma wf_ts_0: "wf_ts \ P 0 \ \ []" unfolding wf_ts_def by simp lemma wf_ts_regex_0: "wf_ts_regex \ P 0 r []" unfolding wf_ts_regex_def by simp lemma (in msaux) wf_since_aux_Nil: "Formula.fv \' \ Formula.fv \' \ wf_since_aux \ V R (init_args I n (Formula.fv \') (Formula.fv \') b) \' \' (init_msaux (init_args I n (Formula.fv \') (Formula.fv \') b)) 0" unfolding wf_since_aux_def by (auto intro!: valid_init_msaux) lemma (in muaux) wf_until_aux_Nil: "Formula.fv \' \ Formula.fv \' \ wf_until_aux \ V R (init_args I n (Formula.fv \') (Formula.fv \') b) \' \' (init_muaux (init_args I n (Formula.fv \') (Formula.fv \') b)) 0" unfolding wf_until_aux_def wf_until_auxlist_def by (auto intro: valid_init_muaux) lemma wf_matchP_aux_Nil: "wf_matchP_aux \ V n R I r [] 0" unfolding wf_matchP_aux_def by simp lemma wf_matchF_aux_Nil: "wf_matchF_aux \ V n R I r [] 0 k" unfolding wf_matchF_aux_def by simp lemma fv_regex_alt: "safe_regex m g r \ Formula.fv_regex r = (\\ \ atms r. Formula.fv \)" unfolding fv_regex_alt atms_def by (auto 0 3 dest: safe_regex_safe_formula) lemmas to_mregex_atms = to_mregex_ok[THEN conjunct1, THEN equalityD1, THEN set_mp, rotated] lemma (in maux) wf_minit0: "safe_formula \ \ \x\Formula.fv \. x < n \ pred_mapping (\x. x = 0) P \ wf_mformula \ 0 P V n R (minit0 n \) \" proof (induction arbitrary: n R P V rule: safe_formula_induct) case (Eq_Const c d) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (Eq_Var1 c x) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (Eq_Var2 c x) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (neq_Var x y) then show ?case by (auto intro!: wf_mformula.neq_Var) next case (Pred e ts) then show ?case by (auto intro!: wf_mformula.Pred) next case (Let p \ \) with fvi_less_nfv show ?case by (auto simp: pred_mapping_alt dom_def intro!: wf_mformula.Let Let(4,5)) next case (And_assign \ \) then have 1: "\x\fv \. x < n" by simp from 1 \safe_assignment (fv \) \\ obtain x t where "x < n" "x \ fv \" "fv_trm t \ fv \" "\ = Formula.Eq (Formula.Var x) t \ \ = Formula.Eq t (Formula.Var x)" unfolding safe_assignment_def by (force split: formula.splits trm.splits) with And_assign show ?case by (auto intro!: wf_mformula.AndAssign split: trm.splits) next case (And_safe \ \) then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0) next case (And_constraint \ \) from \fv \ \ fv \\ \is_constraint \\ obtain t1 p c t2 where "(t1, p, c, t2) = split_constraint \" "formula_of_constraint (split_constraint \) = \" "fv_trm t1 \ fv_trm t2 \ fv \" by (induction rule: is_constraint.induct) auto with And_constraint show ?case by (auto 0 3 intro!: wf_mformula.AndRel) next case (And_Not \ \) then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0) next case (Ands l pos neg) note posneg = "Ands.hyps"(1) let ?wf_minit = "\x. wf_mformula \ 0 P V n R (minit0 n x)" let ?pos = "filter safe_formula l" let ?neg = "filter (Not \ safe_formula) l" have "list_all2 ?wf_minit ?pos pos" using Ands.IH(1) Ands.prems posneg by (auto simp: list_all_iff intro!: list.rel_refl_strong) moreover have "list_all2 ?wf_minit (map remove_neg ?neg) (map remove_neg neg)" using Ands.IH(2) Ands.prems posneg by (auto simp: list.rel_map list_all_iff intro!: list.rel_refl_strong) moreover have "list_all3 (\_ _ _. True) (?pos @ map remove_neg ?neg) (?pos @ map remove_neg ?neg) l" by (auto simp: list_all3_conv_all_nth comp_def sum_length_filter_compl) moreover have "l \ [] \ (MIN \\set l. (0 :: nat)) = 0" by (cases l) (auto simp: Min_eq_iff) ultimately show ?case using Ands.hyps Ands.prems(2) by (auto simp: wf_mbufn_def list_all3_map list.rel_map map_replicate_const[symmetric] subset_eq map_map[symmetric] map_append[symmetric] simp del: map_map map_append intro!: wf_mformula.Ands list_all2_appendI) next case (Neg \) then show ?case by (auto intro!: wf_mformula.Neg) next case (Or \ \) then show ?case by (auto intro!: wf_mformula.Or wf_mbuf2'_0) next case (Exists \) then show ?case by (auto simp: fvi_Suc_bound intro!: wf_mformula.Exists) next case (Agg y \ b f \) then show ?case by (auto intro!: wf_mformula.Agg Agg.IH fvi_plus_bound) next case (Prev I \) thm wf_mformula.Prev[where P=P] then show ?case by (auto intro!: wf_mformula.Prev) next case (Next I \) then show ?case by (auto intro!: wf_mformula.Next) next case (Since \ I \) then show ?case using wf_since_aux_Nil by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0) next case (Not_Since \ I \) then show ?case using wf_since_aux_Nil by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0) next case (Until \ I \) then show ?case using valid_length_muaux[OF valid_init_muaux[OF Until(1)]] wf_until_aux_Nil by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0) next case (Not_Until \ I \) then show ?case using valid_length_muaux[OF valid_init_muaux[OF Not_Until(1)]] wf_until_aux_Nil by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0) next case (MatchP I r) then show ?case by (auto simp: list.rel_map fv_regex_alt simp del: progress_simps split: prod.split intro!: wf_mformula.MatchP list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchP_aux_Nil dest!: to_mregex_atms) next case (MatchF I r) then show ?case by (auto simp: list.rel_map fv_regex_alt progress_le Min_eq_iff progress_regex_def simp del: progress_simps split: prod.split intro!: wf_mformula.MatchF list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchF_aux_Nil dest!: to_mregex_atms) qed lemma (in maux) wf_mstate_minit: "safe_formula \ \ wf_mstate \ pnil R (minit \)" unfolding wf_mstate_def minit_def Let_def by (auto intro!: wf_minit0 fvi_less_nfv) subsubsection \Evaluation\ lemma match_wf_tuple: "Some f = match ts xs \ wf_tuple n (\t\set ts. Formula.fv_trm t) (Table.tabulate f 0 n)" by (induction ts xs arbitrary: f rule: match.induct) (fastforce simp: wf_tuple_def split: if_splits option.splits)+ lemma match_fvi_trm_None: "Some f = match ts xs \ \t\set ts. x \ Formula.fv_trm t \ f x = None" by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits) lemma match_fvi_trm_Some: "Some f = match ts xs \ t \ set ts \ x \ Formula.fv_trm t \ f x \ None" by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits) lemma match_eval_trm: "\t\set ts. \i\Formula.fv_trm t. i < n \ Some f = match ts xs \ map (Formula.eval_trm (Table.tabulate (\i. the (f i)) 0 n)) ts = xs" proof (induction ts xs arbitrary: f rule: match.induct) case (3 x ts y ys) from 3(1)[symmetric] 3(2,3) show ?case by (auto 0 3 dest: match_fvi_trm_Some sym split: option.splits if_splits intro!: eval_trm_fv_cong) qed (auto split: if_splits) lemma wf_tuple_tabulate_Some: "wf_tuple n A (Table.tabulate f 0 n) \ x \ A \ x < n \ \y. f x = Some y" unfolding wf_tuple_def by auto lemma ex_match: "wf_tuple n (\t\set ts. Formula.fv_trm t) v \ \t\set ts. (\x\Formula.fv_trm t. x < n) \ (Formula.is_Var t \ Formula.is_Const t) \ \f. match ts (map (Formula.eval_trm (map the v)) ts) = Some f \ v = Table.tabulate f 0 n" proof (induction ts "map (Formula.eval_trm (map the v)) ts" arbitrary: v rule: match.induct) case (3 x ts y ys) then show ?case proof (cases "x \ (\t\set ts. Formula.fv_trm t)") case True with 3 show ?thesis by (auto simp: insert_absorb dest!: wf_tuple_tabulate_Some meta_spec[of _ v]) next case False with 3(3,4) have *: "map (Formula.eval_trm (map the v)) ts = map (Formula.eval_trm (map the (v[x := None]))) ts" by (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong) from False 3(2-4) obtain f where "match ts (map (Formula.eval_trm (map the v)) ts) = Some f" "v[x := None] = Table.tabulate f 0 n" unfolding * by (atomize_elim, intro 3(1)[of "v[x := None]"]) (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong) moreover from False this have "f x = None" "length v = n" by (auto dest: match_fvi_trm_None[OF sym] arg_cong[of _ _ length]) ultimately show ?thesis using 3(3) by (auto simp: list_eq_iff_nth_eq wf_tuple_def) qed qed (auto simp: wf_tuple_def intro: nth_equalityI) lemma eq_rel_eval_trm: "v \ eq_rel n t1 t2 \ is_simple_eq t1 t2 \ \x\Formula.fv_trm t1 \ Formula.fv_trm t2. x < n \ Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2" by (cases t1; cases t2) (simp_all add: is_simple_eq_def singleton_table_def split: if_splits) lemma in_eq_rel: "wf_tuple n (Formula.fv_trm t1 \ Formula.fv_trm t2) v \ is_simple_eq t1 t2 \ Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2 \ v \ eq_rel n t1 t2" by (cases t1; cases t2) (auto simp: is_simple_eq_def singleton_table_def wf_tuple_def unit_table_def intro!: nth_equalityI split: if_splits) lemma table_eq_rel: "is_simple_eq t1 t2 \ table n (Formula.fv_trm t1 \ Formula.fv_trm t2) (eq_rel n t1 t2)" by (cases t1; cases t2; simp add: is_simple_eq_def) lemma wf_tuple_Suc_fviD: "wf_tuple (Suc n) (Formula.fvi b \) v \ wf_tuple n (Formula.fvi (Suc b) \) (tl v)" unfolding wf_tuple_def by (simp add: fvi_Suc nth_tl) lemma table_fvi_tl: "table (Suc n) (Formula.fvi b \) X \ table n (Formula.fvi (Suc b) \) (tl ` X)" unfolding table_def by (auto intro: wf_tuple_Suc_fviD) lemma wf_tuple_Suc_fvi_SomeI: "0 \ Formula.fvi b \ \ wf_tuple n (Formula.fvi (Suc b) \) v \ wf_tuple (Suc n) (Formula.fvi b \) (Some x # v)" unfolding wf_tuple_def by (auto simp: fvi_Suc less_Suc_eq_0_disj) lemma wf_tuple_Suc_fvi_NoneI: "0 \ Formula.fvi b \ \ wf_tuple n (Formula.fvi (Suc b) \) v \ wf_tuple (Suc n) (Formula.fvi b \) (None # v)" unfolding wf_tuple_def by (auto simp: fvi_Suc less_Suc_eq_0_disj) lemma qtable_project_fv: "qtable (Suc n) (fv \) (mem_restr (lift_envs R)) P X \ qtable n (Formula.fvi (Suc 0) \) (mem_restr R) (\v. \x. P ((if 0 \ fv \ then Some x else None) # v)) (tl ` X)" using neq0_conv by (fastforce simp: image_iff Bex_def fvi_Suc elim!: qtable_cong dest!: qtable_project) lemma mem_restr_lift_envs'_append[simp]: "length xs = b \ mem_restr (lift_envs' b R) (xs @ ys) = mem_restr R ys" unfolding mem_restr_def lift_envs'_def by (auto simp: list_all2_append list.rel_map intro!: exI[where x="map the xs"] list.rel_refl) lemma nth_list_update_alt: "xs[i := x] ! j = (if i < length xs \ i = j then x else xs ! j)" by auto lemma wf_tuple_upd_None: "wf_tuple n A xs \ A - {i} = B \ wf_tuple n B (xs[i:=None])" unfolding wf_tuple_def by (auto simp: nth_list_update_alt) lemma mem_restr_upd_None: "mem_restr R xs \ mem_restr R (xs[i:=None])" unfolding mem_restr_def by (auto simp: list_all2_conv_all_nth nth_list_update_alt) lemma mem_restr_dropI: "mem_restr (lift_envs' b R) xs \ mem_restr R (drop b xs)" unfolding mem_restr_def lift_envs'_def by (auto simp: append_eq_conv_conj list_all2_append2) lemma mem_restr_dropD: assumes "b \ length xs" and "mem_restr R (drop b xs)" shows "mem_restr (lift_envs' b R) xs" proof - let ?R = "\a b. a \ None \ a = Some b" from assms(2) obtain v where "v \ R" and "list_all2 ?R (drop b xs) v" unfolding mem_restr_def .. show ?thesis unfolding mem_restr_def proof have "list_all2 ?R (take b xs) (map the (take b xs))" by (auto simp: list.rel_map intro!: list.rel_refl) moreover note \list_all2 ?R (drop b xs) v\ ultimately have "list_all2 ?R (take b xs @ drop b xs) (map the (take b xs) @ v)" by (rule list_all2_appendI) then show "list_all2 ?R xs (map the (take b xs) @ v)" by simp show "map the (take b xs) @ v \ lift_envs' b R" unfolding lift_envs'_def using assms(1) \v \ R\ by auto qed qed lemma wf_tuple_append: "wf_tuple a {x \ A. x < a} xs \ wf_tuple b {x - a | x. x \ A \ x \ a} ys \ wf_tuple (a + b) A (xs @ ys)" unfolding wf_tuple_def by (auto simp: nth_append eq_diff_iff) lemma wf_tuple_map_Some: "length xs = n \ {0.. A \ wf_tuple n A (map Some xs)" unfolding wf_tuple_def by auto lemma wf_tuple_drop: "wf_tuple (b + n) A xs \ {x - b | x. x \ A \ x \ b} = B \ wf_tuple n B (drop b xs)" unfolding wf_tuple_def by force lemma ecard_image: "inj_on f A \ ecard (f ` A) = ecard A" unfolding ecard_def by (auto simp: card_image dest: finite_imageD) lemma meval_trm_eval_trm: "wf_tuple n A x \ fv_trm t \ A \ \i\A. i < n \ meval_trm t x = Formula.eval_trm (map the x) t" unfolding wf_tuple_def by (induction t) simp_all lemma list_update_id: "xs ! i = z \ xs[i:=z] = xs" by (induction xs arbitrary: i) (auto split: nat.split) lemma qtable_wf_tupleD: "qtable n A P Q X \ \x\X. wf_tuple n A x" unfolding qtable_def table_def by blast lemma qtable_eval_agg: assumes inner: "qtable (b + n) (Formula.fv \) (mem_restr (lift_envs' b R)) (\v. Formula.sat \ V (map the v) i \) rel" and n: "\x\Formula.fv (Formula.Agg y \ b f \). x < n" and fresh: "y + b \ Formula.fv \" and b_fv: "{0.. Formula.fv \" and f_fv: "Formula.fv_trm f \ Formula.fv \" and g0: "g0 = (Formula.fv \ \ {0.. b f \)) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)) (eval_agg n g0 y \ b f rel)" (is "qtable _ ?fv _ ?Q ?rel'") proof - define M where "M = (\v. {(x, ecard Zs) | x Zs. Zs = {zs. length zs = b \ Formula.sat \ V (zs @ v) i \ \ Formula.eval_trm (zs @ v) f = x} \ Zs \ {}})" have f_fvi: "Formula.fvi_trm b f \ Formula.fvi b \" using f_fv by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b]) show ?thesis proof (cases "g0 \ rel = empty_table") case True then have [simp]: "Formula.fvi b \ = {}" by (auto simp: g0 fvi_iff_fv(1)[where b=b]) then have [simp]: "Formula.fvi_trm b f = {}" using f_fvi by auto show ?thesis proof (rule qtableI) show "table n ?fv ?rel'" by (simp add: eval_agg_def True) next fix v assume "wf_tuple n ?fv v" "mem_restr R v" have "\ Formula.sat \ V (zs @ map the v) i \" if [simp]: "length zs = b" for zs proof - let ?zs = "map2 (\z i. if i \ Formula.fv \ then Some z else None) zs [0.. fv \. x < b} ?zs" by (simp add: wf_tuple_def) then have "wf_tuple (b + n) (Formula.fv \) (?zs @ v[y:=None])" using \wf_tuple n ?fv v\ True by (auto simp: g0 intro!: wf_tuple_append wf_tuple_upd_None) then have "\ Formula.sat \ V (map the (?zs @ v[y:=None])) i \" using True \mem_restr R v\ by (auto simp del: map_append dest!: in_qtableI[OF inner, rotated -1] intro!: mem_restr_upd_None) also have "Formula.sat \ V (map the (?zs @ v[y:=None])) i \ \ Formula.sat \ V (zs @ map the v) i \" using True by (auto simp: g0 nth_append intro!: sat_fv_cong) finally show ?thesis . qed then have M_empty: "M (map the v) = {}" unfolding M_def by blast show "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" if "v \ eval_agg n g0 y \ b f rel" using M_empty True that n by (simp add: M_def eval_agg_def g0 singleton_table_def) have "v \ singleton_table n y (the (v ! y))" "length v = n" using \wf_tuple n ?fv v\ unfolding wf_tuple_def singleton_table_def by (auto simp add: tabulate_alt map_nth intro!: trans[OF map_cong[where g="(!) v", simplified nth_map, OF refl], symmetric]) then show "v \ eval_agg n g0 y \ b f rel" if "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" using M_empty True that n by (simp add: M_def eval_agg_def g0) qed next case non_default_case: False have union_fv: "{0.. (\x. x + b) ` Formula.fvi b \ = fv \" using b_fv by (auto simp: fvi_iff_fv(1)[where b=b] intro!: image_eqI[where b=x and x="x - b" for x]) have b_n: "\x\fv \. x < b + n" proof fix x assume "x \ fv \" show "x < b + n" proof (cases "x \ b") case True with \x \ fv \\ have "x - b \ ?fv" by (simp add: fvi_iff_fv(1)[where b=b]) then show ?thesis using n f_fvi by (auto simp: Un_absorb2) qed simp qed define M' where "M' = (\k. let group = Set.filter (\x. drop b x = k) rel; images = meval_trm f ` group in (\y. (y, ecard (Set.filter (\x. meval_trm f x = y) group))) ` images)" have M'_M: "M' (drop b x) = M (map the (drop b x))" if "x \ rel" "mem_restr (lift_envs' b R) x" for x proof - from that have wf_x: "wf_tuple (b + n) (fv \) x" by (auto elim!: in_qtableE[OF inner]) then have wf_zs_x: "wf_tuple (b + n) (fv \) (map Some zs @ drop b x)" if "length zs = b" for zs using that b_fv by (auto intro!: wf_tuple_append wf_tuple_map_Some wf_tuple_drop) have 1: "(length zs = b \ Formula.sat \ V (zs @ map the (drop b x)) i \ \ Formula.eval_trm (zs @ map the (drop b x)) f = y) \ (\a. a \ rel \ take b a = map Some zs \ drop b a = drop b x \ meval_trm f a = y)" (is "?A \ (\a. ?B a)") for y zs proof (intro iffI conjI) assume ?A then have "?B (map Some zs @ drop (length zs) x)" using in_qtableI[OF inner wf_zs_x] \mem_restr (lift_envs' b R) x\ meval_trm_eval_trm[OF wf_zs_x f_fv b_n] by (auto intro!: mem_restr_dropI) then show "\a. ?B a" .. next assume "\a. ?B a" then obtain a where "?B a" .. then have "a \ rel" and a_eq: "a = map Some zs @ drop b x" using append_take_drop_id[of b a] by auto then have "length a = b + n" using inner unfolding qtable_def table_def by (blast intro!: wf_tuple_length) then show "length zs = b" using wf_tuple_length[OF wf_x] unfolding a_eq by simp then have "mem_restr (lift_envs' b R) a" using \mem_restr _ x\ unfolding a_eq by (auto intro!: mem_restr_dropI) then show "Formula.sat \ V (zs @ map the (drop b x)) i \" using in_qtableE[OF inner \a \ rel\] by (auto simp: a_eq sat_fv_cong[THEN iffD1, rotated -1]) from \?B a\ show "Formula.eval_trm (zs @ map the (drop b x)) f = y" using meval_trm_eval_trm[OF wf_zs_x f_fv b_n, OF \length zs = b\] unfolding a_eq by simp qed have 2: "map Some (map the (take b a)) = take b a" if "a \ rel" for a using that b_fv inner[THEN qtable_wf_tupleD] unfolding table_def wf_tuple_def by (auto simp: list_eq_iff_nth_eq) have 3: "ecard {zs. \a. a \ rel \ take b a = map Some zs \ drop b a = drop b x \ P a} = ecard {a. a \ rel \ drop b a = drop b x \ P a}" (is "ecard ?A = ecard ?B") for P proof - have "ecard ?A = ecard ((\zs. map Some zs @ drop b x) ` ?A)" by (auto intro!: ecard_image[symmetric] inj_onI) also have "(\zs. map Some zs @ drop b x) ` ?A = ?B" by (subst (1 2) eq_commute) (auto simp: image_iff, metis "2" append_take_drop_id) finally show ?thesis . qed show ?thesis unfolding M_def M'_def by (auto simp: non_default_case Let_def image_def Set.filter_def 1 3, metis "2") qed have drop_lift: "mem_restr (lift_envs' b R) x" if "x \ rel" "mem_restr R ((drop b x)[y:=z])" for x z proof - have "(drop b x)[y:=None] = (drop b x)[y:=drop b x ! y]" proof - from \x \ rel\ have "drop b x ! y = None" using fresh n inner[THEN qtable_wf_tupleD] by (simp add: add.commute wf_tuple_def) then show ?thesis by simp qed then have "(drop b x)[y:=None] = drop b x" by simp moreover from \x \ rel\ have "length x = b + n" using inner[THEN qtable_wf_tupleD] by (simp add: wf_tuple_def) moreover from that(2) have "mem_restr R ((drop b x)[y:=z, y:=None])" by (rule mem_restr_upd_None) ultimately show ?thesis by (auto intro!: mem_restr_dropD) qed { fix v assume "mem_restr R v" have "v \ (\k. k[y:=Some (eval_agg_op \ (M' k))]) ` drop b ` rel \ v \ (\k. k[y:=Some (eval_agg_op \ (M (map the k)))]) ` drop b ` rel" (is "v \ ?A \ v \ ?B") proof assume "v \ ?A" then obtain v' where *: "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M' (drop b v')))]" by auto then have "M' (drop b v') = M (map the (drop b v'))" using \mem_restr R v\ by (auto intro!: M'_M drop_lift) with * show "v \ ?B" by simp next assume "v \ ?B" then obtain v' where *: "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M (map the (drop b v'))))]" by auto then have "M (map the (drop b v')) = M' (drop b v')" using \mem_restr R v\ by (auto intro!: M'_M[symmetric] drop_lift) with * show "v \ ?A" by simp qed then have "v \ eval_agg n g0 y \ b f rel \ v \ (\k. k[y:=Some (eval_agg_op \ (M (map the k)))]) ` drop b ` rel" by (simp add: non_default_case eval_agg_def M'_def Let_def) } note alt = this show ?thesis proof (rule qtableI) show "table n ?fv ?rel'" using inner[THEN qtable_wf_tupleD] n f_fvi by (auto simp: eval_agg_def non_default_case table_def wf_tuple_def Let_def nth_list_update fvi_iff_fv[where b=b] add.commute) next fix v assume "wf_tuple n ?fv v" "mem_restr R v" then have length_v: "length v = n" by (simp add: wf_tuple_def) show "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" if "v \ eval_agg n g0 y \ b f rel" proof - from that obtain v' where "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M (map the (drop b v'))))]" using alt[OF \mem_restr R v\] by blast then have length_v': "length v' = b + n" using inner[THEN qtable_wf_tupleD] by (simp add: wf_tuple_def) have "Formula.sat \ V (map the v') i \" using \v' \ rel\ \mem_restr R v\ by (auto simp: \v = _\ elim!: in_qtableE[OF inner] intro!: drop_lift \v' \ rel\) then have "Formula.sat \ V (map the (take b v') @ map the v) i \" proof (rule sat_fv_cong[THEN iffD1, rotated], intro ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < length v'" using \x \ fv \\ b_n by (simp add: length_v') ultimately show "map the v' ! x = (map the (take b v') @ map the v) ! x" by (auto simp: \v = _\ nth_append) qed then have 1: "M (map the v) \ {}" by (force simp: M_def length_v') have "y < length (drop b v')" using n by (simp add: length_v') moreover have "Formula.sat \ V (zs @ map the v) i \ \ Formula.sat \ V (zs @ map the (drop b v')) i \" if "length zs = b" for zs proof (intro sat_fv_cong ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < length v'" using \x \ fv \\ b_n by (simp add: length_v') ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x" by (auto simp: \v = _\ that nth_append) qed moreover have "Formula.eval_trm (zs @ map the v) f = Formula.eval_trm (zs @ map the (drop b v')) f" if "length zs = b" for zs proof (intro eval_trm_fv_cong ballI) fix x assume "x \ fv_trm f" then have "x \ y + b" using f_fv fresh by blast moreover have "x < length v'" using \x \ fv_trm f\ f_fv b_n by (auto simp: length_v') ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x" by (auto simp: \v = _\ that nth_append) qed ultimately have "map the v ! y = eval_agg_op \ (M (map the v))" by (simp add: M_def \v = _\ conj_commute cong: conj_cong) with 1 show ?thesis by (auto simp: M_def) qed show "v \ eval_agg n g0 y \ b f rel" if sat_Agg: "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" proof - obtain zs where "length zs = b" and "map Some zs @ v[y:=None] \ rel" proof (cases "fv \ \ {0.. empty_table" by (simp add: g0) then obtain x where "x \ rel" by auto have "(\i < n. (v[y:=None]) ! i = None)" using True \wf_tuple n ?fv v\ f_fv by (fastforce simp: wf_tuple_def fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b]) moreover have x: "(\i < n. drop b x ! i = None) \ length x = b + n" using True \x \ rel\ inner[THEN qtable_wf_tupleD] f_fv by (auto simp: wf_tuple_def) ultimately have "v[y:=None] = drop b x" unfolding list_eq_iff_nth_eq by (auto simp: length_v) with \x \ rel\ have "take b x @ v[y:=None] \ rel" by simp moreover have "map (Some \ the) (take b x) = take b x" using True \x \ rel\ inner[THEN qtable_wf_tupleD] b_fv by (subst map_cong[where g=id, OF refl]) (auto simp: wf_tuple_def in_set_conv_nth) ultimately have "map Some (map the (take b x)) @ v[y:=None] \ rel" by simp then show thesis using x[THEN conjunct2] by (fastforce intro!: that[rotated]) next case False with sat_Agg obtain zs where "length zs = b" and "Formula.sat \ V (zs @ map the v) i \" by auto then have "Formula.sat \ V (zs @ map the (v[y:=None])) i \" using fresh by (auto simp: map_update not_less nth_append elim!: sat_fv_cong[THEN iffD1, rotated] intro!: nth_list_update_neq[symmetric]) then have "map Some zs @ v[y:=None] \ rel" using b_fv f_fv fresh by (auto intro!: in_qtableI[OF inner] wf_tuple_append wf_tuple_map_Some wf_tuple_upd_None \wf_tuple n ?fv v\ mem_restr_upd_None \mem_restr R v\ simp: \length zs = b\ set_eq_iff fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b]) force+ with that \length zs = b\ show thesis by blast qed then have 1: "v[y:=None] \ drop b ` rel" by (intro image_eqI) auto have y_length: "y < length v" using n by (simp add: length_v) moreover have "Formula.sat \ V (zs @ map the (v[y:=None])) i \ \ Formula.sat \ V (zs @ map the v) i \" if "length zs = b" for zs proof (intro sat_fv_cong ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < b + length v" using \x \ fv \\ b_n by (simp add: length_v) ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x" by (auto simp: that nth_append) qed moreover have "Formula.eval_trm (zs @ map the (v[y:=None])) f = Formula.eval_trm (zs @ map the v) f" if "length zs = b" for zs proof (intro eval_trm_fv_cong ballI) fix x assume "x \ fv_trm f" then have "x \ y + b" using f_fv fresh by blast moreover have "x < b + length v" using \x \ fv_trm f\ f_fv b_n by (auto simp: length_v) ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x" by (auto simp: that nth_append) qed ultimately have "map the v ! y = eval_agg_op \ (M (map the (v[y:=None])))" using sat_Agg by (simp add: M_def cong: conj_cong) (simp cong: rev_conj_cong) then have 2: "v ! y = Some (eval_agg_op \ (M (map the (v[y:=None]))))" using \wf_tuple n ?fv v\ y_length by (auto simp add: wf_tuple_def) show ?thesis unfolding alt[OF \mem_restr R v\] by (rule image_eqI[where x="v[y:=None]"]) (use 1 2 in \auto simp: y_length list_update_id\) qed qed qed qed lemma mprev: "mprev_next I xs ts = (ys, xs', ts') \ list_all2 P [i.. list_all2 (\i t. t = \ \ i) [i.. i \ j' \ i < j \ list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then P i X else X = empty_table) [i.. list_all2 P [min j' (j-1).. list_all2 (\i t. t = \ \ i) [min j' (j-1).. list_all2 P [Suc i.. list_all2 (\i t. t = \ \ i) [i.. Suc i \ j' \ i < j \ list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then P (Suc i) X else X = empty_table) [i.. list_all2 P [Suc (min (j'-1) (j-1)).. list_all2 (\i t. t = \ \ i) [min (j'-1) (j-1).. A \ A \ set xs \ x \ foldr (\) xs {}" by (induction xs) auto lemma in_foldr_UnE: "x \ foldr (\) xs {} \ (\A. A \ set xs \ x \ A \ P) \ P" by (induction xs) auto lemma sat_the_restrict: "fv \ \ A \ Formula.sat \ V (map the (restrict A v)) i \ = Formula.sat \ V (map the v) i \" by (rule sat_fv_cong) (auto intro!: map_the_restrict) lemma eps_the_restrict: "fv_regex r \ A \ Regex.eps (Formula.sat \ V (map the (restrict A v))) i r = Regex.eps (Formula.sat \ V (map the v)) i r" by (rule eps_fv_cong) (auto intro!: map_the_restrict) lemma sorted_wrt_filter[simp]: "sorted_wrt R xs \ sorted_wrt R (filter P xs)" by (induct xs) auto lemma concat_map_filter[simp]: "concat (map f (filter P xs)) = concat (map (\x. if P x then f x else []) xs)" by (induct xs) auto lemma map_filter_alt: "map f (filter P xs) = concat (map (\x. if P x then [f x] else []) xs)" by (induct xs) auto lemma (in maux) update_since: assumes pre: "wf_since_aux \ V R args \ \ aux ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel2" and result_eq: "(rel, aux') = update_since args rel1 rel2 (\ \ ne) aux" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \" and args_R: "args_R args = Formula.fv \" and args_pos: "args_pos args = pos" shows "wf_since_aux \ V R args \ \ aux' (Suc ne)" and "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ I \)) rel" proof - let ?wf_tuple = "\v. wf_tuple n (Formula.fv \) v" note sat.simps[simp del] from pre[unfolded wf_since_aux_def] obtain cur auxlist where aux: "valid_msaux args cur aux auxlist" "sorted_wrt (\x y. fst y < fst x) auxlist" "\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right I \ (\i. \ \ i = t) \ qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - t)) \)) X" "\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right I \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)" and cur_def: "cur = (if ne = 0 then 0 else \ \ (ne - 1))" unfolding args_ivl args_n args_pos by blast from pre[unfolded wf_since_aux_def] have fv_sub: "Formula.fv \ \ Formula.fv \" by simp define aux0 where "aux0 = join_msaux args rel1 (add_new_ts_msaux args (\ \ ne) aux)" define auxlist0 where "auxlist0 = [(t, join rel pos rel1). (t, rel) \ auxlist, \ \ ne - t \ right I]" have tabL: "table (args_n args) (args_L args) rel1" using qtable1[unfolded qtable_def] unfolding args_n[symmetric] args_L[symmetric] by simp have cur_le: "cur \ \ \ ne" unfolding cur_def by auto have valid0: "valid_msaux args (\ \ ne) aux0 auxlist0" unfolding aux0_def auxlist0_def using valid_join_msaux[OF valid_add_new_ts_msaux[OF aux(1)], OF cur_le tabL] by (auto simp: args_ivl args_pos cur_def map_filter_alt split_beta cong: map_cong) from aux(2) have sorted_auxlist0: "sorted_wrt (\x y. fst x > fst y) auxlist0" unfolding auxlist0_def by (induction auxlist) (auto simp add: sorted_wrt_append) have in_auxlist0_1: "(t, X) \ set auxlist0 \ ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ qtable n (Formula.fv \) (mem_restr R) (\v. (Formula.sat \ V (map the v) (ne-1) (Sincep pos \ (point (\ \ (ne-1) - t)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \))) X" for t X unfolding auxlist0_def using fvi_subset by (auto 0 1 elim!: qtable_join[OF _ qtable1] simp: sat_the_restrict dest!: aux(3)) then have in_auxlist0_le_\: "(t, X) \ set auxlist0 \ t \ \ \ ne" for t X by (meson \_mono diff_le_self le_trans) have in_auxlist0_2: "ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ \i. \ \ i = t \ \X. (t, X) \ set auxlist0" for t proof - fix t assume "ne \ 0" "t \ \ \ (ne-1)" "\ \ ne - t \ right I" "\i. \ \ i = t" then obtain X where "(t, X) \ set auxlist" by (atomize_elim, intro aux(4)) (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono \_mono) with \\ \ ne - t \ right I\ have "(t, join X pos rel1) \ set auxlist0" unfolding auxlist0_def by (auto elim!: bexI[rotated] intro!: exI[of _ X]) then show "\X. (t, X) \ set auxlist0" by blast qed have auxlist0_Nil: "auxlist0 = [] \ ne = 0 \ ne \ 0 \ (\t. t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t))" using in_auxlist0_2 by (auto) have aux'_eq: "aux' = add_new_table_msaux args rel2 aux0" using result_eq unfolding aux0_def update_since_def Let_def by simp define auxlist' where auxlist'_eq: "auxlist' = (case auxlist0 of [] \ [(\ \ ne, rel2)] | x # auxlist' \ (if fst x = \ \ ne then (fst x, snd x \ rel2) # auxlist' else (\ \ ne, rel2) # x # auxlist'))" have tabR: "table (args_n args) (args_R args) rel2" using qtable2[unfolded qtable_def] unfolding args_n[symmetric] args_R[symmetric] by simp have valid': "valid_msaux args (\ \ ne) aux' auxlist'" unfolding aux'_eq auxlist'_eq using valid_add_new_table_msaux[OF valid0 tabR] by (auto simp: not_le split: list.splits option.splits if_splits) have sorted_auxlist': "sorted_wrt (\x y. fst x > fst y) auxlist'" unfolding auxlist'_eq using sorted_auxlist0 in_auxlist0_le_\ by (cases auxlist0) fastforce+ have in_auxlist'_1: "t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ (point (\ \ ne - t)) \)) X" if auxlist': "(t, X) \ set auxlist'" for t X proof (cases auxlist0) case Nil with auxlist' show ?thesis unfolding auxlist'_eq using qtable2 auxlist0_Nil by (auto simp: zero_enat_def[symmetric] sat_Since_rec[where i=ne] dest: spec[of _ "\ \ (ne-1)"] elim!: qtable_cong[OF _ refl]) next case (Cons a as) show ?thesis proof (cases "t = \ \ ne") case t: True show ?thesis proof (cases "fst a = \ \ ne") case True with auxlist' Cons t have "X = snd a \ rel2" unfolding auxlist'_eq using sorted_auxlist0 by (auto split: if_splits) moreover from in_auxlist0_1[of "fst a" "snd a"] Cons have "ne \ 0" "fst a \ \ \ (ne - 1)" "\ \ ne - fst a \ right I" "\i. \ \ i = fst a" "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - fst a)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \)) (snd a)" by (auto simp: True[symmetric] zero_enat_def[symmetric]) ultimately show ?thesis using qtable2 t True by (auto simp: sat_Since_rec[where i=ne] sat.simps(6) elim!: qtable_union) next case False with auxlist' Cons t have "X = rel2" unfolding auxlist'_eq using sorted_auxlist0 in_auxlist0_le_\[of "fst a" "snd a"] by (auto split: if_splits) with auxlist' Cons t False show ?thesis unfolding auxlist'_eq using qtable2 in_auxlist0_2[of "\ \ (ne-1)"] in_auxlist0_le_\[of "fst a" "snd a"] sorted_auxlist0 by (auto simp: sat_Since_rec[where i=ne] sat.simps(3) zero_enat_def[symmetric] enat_0_iff not_le elim!: qtable_cong[OF _ refl] dest!: le_\_less meta_mp) qed next case False with auxlist' Cons have "(t, X) \ set auxlist0" unfolding auxlist'_eq by (auto split: if_splits) then have "ne \ 0" "t \ \ \ (ne - 1)" "\ \ ne - t \ right I" "\i. \ \ i = t" "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - t)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \)) X" using in_auxlist0_1 by blast+ with False auxlist' Cons show ?thesis unfolding auxlist'_eq using qtable2 by (fastforce simp: sat_Since_rec[where i=ne] sat.simps(6) diff_diff_right[where i="\ \ ne" and j="\ \ _ + \ \ ne" and k="\ \ (ne - 1)", OF trans_le_add2, simplified] elim!: qtable_cong[OF _ refl] order_trans dest: le_\_less) qed qed have in_auxlist'_2: "\X. (t, X) \ set auxlist'" if "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" for t proof (cases "t = \ \ ne") case True then show ?thesis proof (cases auxlist0) case Nil with True show ?thesis unfolding auxlist'_eq by (simp add: zero_enat_def[symmetric]) next case (Cons a as) with True show ?thesis unfolding auxlist'_eq by (cases "fst a = \ \ ne") (auto simp: zero_enat_def[symmetric]) qed next case False with that have "ne \ 0" using le_\_less neq0_conv by blast moreover from False that have "t \ \ \ (ne-1)" by (metis One_nat_def Suc_leI Suc_pred \_mono diff_is_0_eq' order.antisym neq0_conv not_le) ultimately obtain X where "(t, X) \ set auxlist0" using \\ \ ne - t \ right I\ \\i. \ \ i = t\ using \_mono[of "ne - 1" "ne" \] by (atomize_elim, cases "right I") (auto intro!: in_auxlist0_2 simp del: \_mono) then show ?thesis unfolding auxlist'_eq using False \\ \ ne - t \ right I\ by (auto intro: exI[of _ X] split: list.split) qed show "wf_since_aux \ V R args \ \ aux' (Suc ne)" unfolding wf_since_aux_def args_ivl args_n args_pos by (auto simp add: fv_sub dest: in_auxlist'_1 intro: sorted_auxlist' in_auxlist'_2 intro!: exI[of _ auxlist'] valid') have "rel = result_msaux args aux'" using result_eq by (auto simp add: update_since_def Let_def) with valid' have rel_eq: "rel = foldr (\) [rel. (t, rel) \ auxlist', left I \ \ \ ne - t] {}" by (auto simp add: args_ivl valid_result_msaux intro!: arg_cong[where f = "\x. foldr (\) (concat x) {}"] split: option.splits) have rel_alt: "rel = (\(t, rel) \ set auxlist'. if left I \ \ \ ne - t then rel else empty_table)" unfolding rel_eq by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI) show "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ I \)) rel" unfolding rel_alt proof (rule qtable_Union[where Qi="\(t, X) v. left I \ \ \ ne - t \ Formula.sat \ V (map the v) ne (Sincep pos \ (point (\ \ ne - t)) \)"], goal_cases finite qtable equiv) case (equiv v) show ?case proof (rule iffI, erule sat_Since_point, goal_cases left right) case (left j) then show ?case using in_auxlist'_2[of "\ \ j", OF _ _ exI, OF _ _ refl] by auto next case right then show ?case by (auto elim!: sat_Since_pointD dest: in_auxlist'_1) qed qed (auto dest!: in_auxlist'_1 intro!: qtable_empty) qed lemma fv_regex_from_mregex: "ok (length \s) mr \ fv_regex (from_mregex mr \s) \ (\\ \ set \s. fv \)" by (induct mr) (auto simp: Bex_def in_set_conv_nth)+ lemma qtable_\_lax: assumes "ok (length \s) mr" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" and "fv_regex (from_mregex mr \s) \ A" and "qtable n A (mem_restr R) Q guard" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s) \ Q v) (\_lax guard rels mr)" using assms proof (induct mr) case (MPlus mr1 mr2) from MPlus(3-6) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) then have "fv_regex (from_mregex mr1 \s) \ A" "fv_regex (from_mregex mr2 \s) \ A" using fv_regex_from_mregex[of \s mr1] fv_regex_from_mregex[of \s mr2] by (auto simp: subset_eq) with MTimes(3-6) show ?case by (auto simp: eps_the_restrict restrict_idle intro!: qtable_join[OF MTimes(1,2)]) qed (auto split: prod.splits if_splits simp: qtable_empty_iff list_all2_conv_all_nth in_set_conv_nth restrict_idle sat_the_restrict intro: in_qtableI qtableI elim!: qtable_join[where A=A and C=A]) lemma nullary_qtable_cases: "qtable n {} P Q X \ (X = empty_table \ X = unit_table n)" by (simp add: qtable_def table_empty) lemma qtable_empty_unit_table: "qtable n {} R P empty_table \ qtable n {} R (\v. \ P v) (unit_table n)" by (auto intro: qtable_unit_table simp add: qtable_empty_iff) lemma qtable_unit_empty_table: "qtable n {} R P (unit_table n) \ qtable n {} R (\v. \ P v) empty_table" by (auto intro!: qtable_empty elim: in_qtableE simp add: wf_tuple_empty unit_table_def) lemma qtable_nonempty_empty_table: "qtable n {} R P X \ x \ X \ qtable n {} R (\v. \ P v) empty_table" by (frule nullary_qtable_cases) (auto dest: qtable_unit_empty_table) lemma qtable_r\_strict: assumes "safe_regex Past Strict (from_mregex mr \s)" "ok (length \s) mr" "A = fv_regex (from_mregex mr \s)" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s)) (r\_strict n rels mr)" using assms proof (hypsubst, induct Past Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits) next case (Test \) then show ?case by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table dest!: qtable_nonempty_empty_table split: if_splits) next case (Plus r s) then show ?case by (cases mr) (fastforce intro: qtable_union split: if_splits)+ next case (TimesP r s) then show ?case by (cases mr) (auto intro: qtable_cong[OF qtable_\_lax] split: if_splits)+ next case (Star r) then show ?case by (cases mr) (auto simp: qtable_unit_table split: if_splits) qed lemma qtable_l\_strict: assumes "safe_regex Futu Strict (from_mregex mr \s)" "ok (length \s) mr" "A = fv_regex (from_mregex mr \s)" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s)) (l\_strict n rels mr)" using assms proof (hypsubst, induct Futu Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits) next case (Test \) then show ?case by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table dest!: qtable_nonempty_empty_table split: if_splits) next case (Plus r s) then show ?case by (cases mr) (fastforce intro: qtable_union split: if_splits)+ next case (TimesF r s) then show ?case by (cases mr) (auto intro: qtable_cong[OF qtable_\_lax] split: if_splits)+ next case (Star r) then show ?case by (cases mr) (auto simp: qtable_unit_table split: if_splits) qed lemma rtranclp_False: "(\i j. False)\<^sup>*\<^sup>* = (=)" proof - have "(\i j. False)\<^sup>*\<^sup>* i j \ i = j" for i j :: 'a by (induct i j rule: rtranclp.induct) auto then show ?thesis by (auto intro: exI[of _ 0]) qed inductive ok_rctxt for \s where "ok_rctxt \s id id" | "ok_rctxt \s \ \' \ ok_rctxt \s (\t. \ (MTimes mr t)) (\t. \' (Regex.Times (from_mregex mr \s) t))" lemma ok_rctxt_swap: "ok_rctxt \s \ \' \ from_mregex (\ mr) \s = \' (from_mregex mr \s)" by (induct \ \' arbitrary: mr rule: ok_rctxt.induct) auto lemma ok_rctxt_cong: "ok_rctxt \s \ \' \ Regex.match (Formula.sat \ V v) r = Regex.match (Formula.sat \ V v) s \ Regex.match (Formula.sat \ V v) (\' r) i j = Regex.match (Formula.sat \ V v) (\' s) i j" by (induct \ \' arbitrary: r s rule: ok_rctxt.induct) simp_all lemma qtable_r\\: assumes "ok (length \s) mr" "fv_regex (from_mregex mr \s) \ A" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" and "ok_rctxt \s \ \'" and "\ms \ \ ` RPD mr. qtable n A (mem_restr R) (\v. Q (map the v) (from_mregex ms \s)) (lookup rel ms)" shows "qtable n A (mem_restr R) (\v. \s \ Regex.rpd\ \' (Formula.sat \ V (map the v)) j (from_mregex mr \s). Q (map the v) s) (r\ \ rel rels mr)" using assms proof (induct mr arbitrary: \ \') case MSkip then show ?case by (auto simp: rtranclp_False ok_rctxt_swap qtable_empty_iff elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ \ \']] split: nat.splits) next case (MPlus mr1 mr2) from MPlus(3-7) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) from MTimes(3-7) show ?case by (auto intro!: qtable_union[OF MTimes(2) qtable_\_lax[OF _ _ _ MTimes(1)]] elim!: ok_rctxt.intros(2) simp: MTimesL_def Ball_def) next case (MStar mr) from MStar(2-6) show ?case by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_rctxt.intros simp: MTimesL_def Ball_def) qed (auto simp: qtable_empty_iff) lemmas qtable_r\ = qtable_r\\[OF _ _ _ ok_rctxt.intros(1), unfolded rpd\_rpd image_id id_apply] inductive ok_lctxt for \s where "ok_lctxt \s id id" | "ok_lctxt \s \ \' \ ok_lctxt \s (\t. \ (MTimes t mr)) (\t. \' (Regex.Times t (from_mregex mr \s)))" lemma ok_lctxt_swap: "ok_lctxt \s \ \' \ from_mregex (\ mr) \s = \' (from_mregex mr \s)" by (induct \ \' arbitrary: mr rule: ok_lctxt.induct) auto lemma ok_lctxt_cong: "ok_lctxt \s \ \' \ Regex.match (Formula.sat \ V v) r = Regex.match (Formula.sat \ V v) s \ Regex.match (Formula.sat \ V v) (\' r) i j = Regex.match (Formula.sat \ V v) (\' s) i j" by (induct \ \' arbitrary: r s rule: ok_lctxt.induct) simp_all lemma qtable_l\\: assumes "ok (length \s) mr" "fv_regex (from_mregex mr \s) \ A" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" and "ok_lctxt \s \ \'" and "\ms \ \ ` LPD mr. qtable n A (mem_restr R) (\v. Q (map the v) (from_mregex ms \s)) (lookup rel ms)" shows "qtable n A (mem_restr R) (\v. \s \ Regex.lpd\ \' (Formula.sat \ V (map the v)) j (from_mregex mr \s). Q (map the v) s) (l\ \ rel rels mr)" using assms proof (induct mr arbitrary: \ \') case MSkip then show ?case by (auto simp: rtranclp_False ok_lctxt_swap qtable_empty_iff elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ \ \']] split: nat.splits) next case (MPlus mr1 mr2) from MPlus(3-7) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) from MTimes(3-7) show ?case by (auto intro!: qtable_union[OF MTimes(1) qtable_\_lax[OF _ _ _ MTimes(2)]] elim!: ok_lctxt.intros(2) simp: MTimesR_def Ball_def) next case (MStar mr) from MStar(2-6) show ?case by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_lctxt.intros simp: MTimesR_def Ball_def) qed (auto simp: qtable_empty_iff) lemmas qtable_l\ = qtable_l\\[OF _ _ _ ok_lctxt.intros(1), unfolded lpd\_lpd image_id id_apply] lemma RPD_fv_regex_le: "ms \ RPD mr \ fv_regex (from_mregex ms \s) \ fv_regex (from_mregex mr \s)" by (induct mr arbitrary: ms) (auto simp: MTimesL_def split: nat.splits)+ lemma RPD_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPD mr \ safe_regex Past g (from_mregex ms \s)" proof (induct Past g "from_mregex mr \s" arbitrary: mr ms rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test g \) then show ?case by (cases mr) auto next case (Plus g r s mrs) then show ?case proof (cases mrs) case (MPlus mr ms) with Plus(3-5) show ?thesis by (auto dest!: Plus(1,2)) qed auto next case (TimesP g r s mrs) then show ?case proof (cases mrs) case (MTimes mr ms) with TimesP(3-5) show ?thesis by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le TimesP(1,2)) qed auto next case (Star g r) then show ?case proof (cases mr) case (MStar x6) with Star(2-4) show ?thesis by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le elim!: safe_cosafe[rotated] dest!: Star(1)) qed auto qed lemma RPDi_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPDi n mr ==> safe_regex Past g (from_mregex ms \s)" by (induct n arbitrary: ms mr) (auto dest: RPD_safe) lemma RPDs_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPDs mr ==> safe_regex Past g (from_mregex ms \s)" unfolding RPDs_def by (auto dest: RPDi_safe) lemma RPD_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPD mr \ fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" proof (induct Past Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto split: nat.splits) next case (Test \) then show ?case by (cases mr) auto next case (Plus r s) then show ?case by (cases mr) auto next case (TimesP r s) then show ?case by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le split: modality.splits) next case (Star r) then show ?case by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le) qed lemma RPDi_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPDi n mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" by (induct n arbitrary: ms mr) (auto 5 0 dest: RPD_safe_fv_regex RPD_safe) lemma RPDs_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPDs mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" unfolding RPDs_def by (auto dest: RPDi_safe_fv_regex) lemma RPD_ok: "ok m mr \ ms \ RPD mr \ ok m ms" proof (induct mr arbitrary: ms) case (MPlus mr1 mr2) from MPlus(3,4) show ?case by (auto elim: MPlus(1,2)) next case (MTimes mr1 mr2) from MTimes(3,4) show ?case by (auto elim: MTimes(1,2) simp: MTimesL_def) next case (MStar mr) from MStar(2,3) show ?case by (auto elim: MStar(1) simp: MTimesL_def) qed (auto split: nat.splits) lemma RPDi_ok: "ok m mr \ ms \ RPDi n mr \ ok m ms" by (induct n arbitrary: ms mr) (auto intro: RPD_ok) lemma RPDs_ok: "ok m mr \ ms \ RPDs mr \ ok m ms" unfolding RPDs_def by (auto intro: RPDi_ok) lemma LPD_fv_regex_le: "ms \ LPD mr \ fv_regex (from_mregex ms \s) \ fv_regex (from_mregex mr \s)" by (induct mr arbitrary: ms) (auto simp: MTimesR_def split: nat.splits)+ lemma LPD_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPD mr ==> safe_regex Futu g (from_mregex ms \s)" proof (induct Futu g "from_mregex mr \s" arbitrary: mr ms rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test g \) then show ?case by (cases mr) auto next case (Plus g r s mrs) then show ?case proof (cases mrs) case (MPlus mr ms) with Plus(3-5) show ?thesis by (auto dest!: Plus(1,2)) qed auto next case (TimesF g r s mrs) then show ?case proof (cases mrs) case (MTimes mr ms) with TimesF(3-5) show ?thesis by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits dest: TimesF(1,2)) qed auto next case (Star g r) then show ?case proof (cases mr) case (MStar x6) with Star(2-4) show ?thesis by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le elim!: safe_cosafe[rotated] dest!: Star(1)) qed auto qed lemma LPDi_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPDi n mr ==> safe_regex Futu g (from_mregex ms \s)" by (induct n arbitrary: ms mr) (auto dest: LPD_safe) lemma LPDs_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPDs mr ==> safe_regex Futu g (from_mregex ms \s)" unfolding LPDs_def by (auto dest: LPDi_safe) lemma LPD_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPD mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" proof (induct Futu Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test \) then show ?case by (cases mr) auto next case (Plus r s) then show ?case by (cases mr) auto next case (TimesF r s) then show ?case by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits) next case (Star r) then show ?case by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le) qed lemma LPDi_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPDi n mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" by (induct n arbitrary: ms mr) (auto 5 0 dest: LPD_safe_fv_regex LPD_safe) lemma LPDs_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPDs mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" unfolding LPDs_def by (auto dest: LPDi_safe_fv_regex) lemma LPD_ok: "ok m mr \ ms \ LPD mr \ ok m ms" proof (induct mr arbitrary: ms) case (MPlus mr1 mr2) from MPlus(3,4) show ?case by (auto elim: MPlus(1,2)) next case (MTimes mr1 mr2) from MTimes(3,4) show ?case by (auto elim: MTimes(1,2) simp: MTimesR_def) next case (MStar mr) from MStar(2,3) show ?case by (auto elim: MStar(1) simp: MTimesR_def) qed (auto split: nat.splits) lemma LPDi_ok: "ok m mr \ ms \ LPDi n mr \ ok m ms" by (induct n arbitrary: ms mr) (auto intro: LPD_ok) lemma LPDs_ok: "ok m mr \ ms \ LPDs mr \ ok m ms" unfolding LPDs_def by (auto intro: LPDi_ok) lemma update_matchP: assumes pre: "wf_matchP_aux \ V n R I r aux ne" and safe: "safe_regex Past Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (RPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel) \s rels" and result_eq: "(rel, aux') = update_matchP n I mr mrs rels (\ \ ne) aux" shows "wf_matchP_aux \ V n R I r aux' (Suc ne)" and "qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP I r)) rel" proof - let ?wf_tuple = "\v. wf_tuple n (Formula.fv_regex r) v" let ?update = "\rel t. mrtabulate mrs (\mr. r\ id rel rels mr \ (if t = \ \ ne then r\_strict n rels mr else {}))" note sat.simps[simp del] define aux0 where "aux0 = [(t, ?update rel t). (t, rel) \ aux, enat (\ \ ne - t) \ right I]" have sorted_aux0: "sorted_wrt (\x y. fst x > fst y) aux0" using pre[unfolded wf_matchP_aux_def, THEN conjunct1] unfolding aux0_def by (induction aux) (auto simp add: sorted_wrt_append) { fix ms assume "ms \ RPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Past Strict (from_mregex ms \s)" "ok (length \s) ms" "RPD ms \ RPDs mr" using safe RPDs_safe RPDs_safe_fv_regex mr from_mregex_to_mregex RPDs_ok to_mregex_ok RPDs_trans by fastforce+ } note * = this have **: "\ \ ne - (\ \ i + \ \ ne - \ \ (ne - Suc 0)) = \ \ (ne - Suc 0) - \ \ i" for i by (metis (no_types, lifting) Nat.diff_diff_right \_mono add.commute add_diff_cancel_left diff_le_self le_add2 order_trans) have ***: "\ \ i = \ \ ne" if "\ \ ne \ \ \ i" "\ \ i \ \ \ (ne - Suc 0)" "ne > 0" for i by (metis (no_types, lifting) Suc_pred \_mono diff_le_self le_\_less le_antisym not_less_eq that) then have in_aux0_1: "(t, X) \ set aux0 \ ne \ 0 \ t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ (\ms\RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms))" for t X unfolding aux0_def using safe mr mrs by (auto simp: lookup_tabulate map_of_map_restrict restrict_map_def finite_RPDs * ** RPDs_trans diff_le_mono2 intro!: sat_MatchP_rec[of \ _ _ ne, THEN iffD2] qtable_union[OF qtable_r\[OF _ _ qtables] qtable_r\_strict[OF _ _ _ qtables], of ms "fv_regex r" "\v r. Formula.sat \ V v (ne - Suc 0) (Formula.MatchP (point 0) r)" _ ms for ms] qtable_cong[OF qtable_r\[OF _ _ qtables], of ms "fv_regex r" "\v r. Formula.sat \ V v (ne - Suc 0) (Formula.MatchP (point (\ \ (ne - Suc 0) - \ \ i)) r)" _ _ "(\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - \ \ i)) (from_mregex ms \s)))" for ms i] dest!: assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct1, rule_format] sat_MatchP_rec["of" \ _ _ ne, THEN iffD1] elim!: bspec order.trans[OF _ \_mono] bexI[rotated] split: option.splits if_splits) (* slow 7 sec *) then have in_aux0_le_\: "(t, X) \ set aux0 \ t \ \ \ ne" for t X by (meson \_mono diff_le_self le_trans) have in_aux0_2: "ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ \i. \ \ i = t \ \X. (t, X) \ set aux0" for t proof - fix t assume "ne \ 0" "t \ \ \ (ne-1)" "\ \ ne - t \ right I" "\i. \ \ i = t" then obtain X where "(t, X) \ set aux" by (atomize_elim, intro assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct2, rule_format]) (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono \_mono) with \\ \ ne - t \ right I\ have "(t, ?update X t) \ set aux0" unfolding aux0_def by (auto simp: id_def elim!: bexI[rotated] intro!: exI[of _ X]) then show "\X. (t, X) \ set aux0" by blast qed have aux0_Nil: "aux0 = [] \ ne = 0 \ ne \ 0 \ (\t. t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t))" using in_aux0_2 by (cases "ne = 0") (auto) have aux'_eq: "aux' = (case aux0 of [] \ [(\ \ ne, mrtabulate mrs (r\_strict n rels))] | x # aux' \ (if fst x = \ \ ne then x # aux' else (\ \ ne, mrtabulate mrs (r\_strict n rels)) # x # aux'))" using result_eq unfolding aux0_def update_matchP_def Let_def by simp have sorted_aux': "sorted_wrt (\x y. fst x > fst y) aux'" unfolding aux'_eq using sorted_aux0 in_aux0_le_\ by (cases aux0) (fastforce)+ have in_aux'_1: "t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ (\ms\RPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms))" if aux': "(t, X) \ set aux'" for t X proof (cases aux0) case Nil with aux' show ?thesis unfolding aux'_eq using safe mrs qtables aux0_Nil * by (auto simp: zero_enat_def[symmetric] sat_MatchP_rec[where i=ne] lookup_tabulate finite_RPDs split: option.splits intro!: qtable_cong[OF qtable_r\_strict] dest: spec[of _ "\ \ (ne-1)"]) next case (Cons a as) show ?thesis proof (cases "t = \ \ ne") case t: True show ?thesis proof (cases "fst a = \ \ ne") case True with aux' Cons t have "X = snd a" unfolding aux'_eq using sorted_aux0 by auto moreover from in_aux0_1[of "fst a" "snd a"] Cons have "ne \ 0" "fst a \ \ \ ne" "\ \ ne - fst a \ right I" "\i. \ \ i = fst a" "\ms \ RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - fst a)) (from_mregex ms \s))) (lookup (snd a) ms)" by auto ultimately show ?thesis using t True by auto next case False with aux' Cons t have "X = mrtabulate mrs (r\_strict n rels)" unfolding aux'_eq using sorted_aux0 in_aux0_le_\[of "fst a" "snd a"] by auto with aux' Cons t False show ?thesis unfolding aux'_eq using safe mrs qtables * in_aux0_2[of "\ \ (ne-1)"] in_aux0_le_\[of "fst a" "snd a"] sorted_aux0 by (auto simp: sat_MatchP_rec[where i=ne] zero_enat_def[symmetric] enat_0_iff not_le lookup_tabulate finite_RPDs split: option.splits intro!: qtable_cong[OF qtable_r\_strict] dest!: le_\_less meta_mp) qed next case False with aux' Cons have "(t, X) \ set aux0" unfolding aux'_eq by (auto split: if_splits) then have "ne \ 0" "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" "\ms \ RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms)" using in_aux0_1 by blast+ with False aux' Cons show ?thesis unfolding aux'_eq by auto qed qed have in_aux'_2: "\X. (t, X) \ set aux'" if "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" for t proof (cases "t = \ \ ne") case True then show ?thesis proof (cases aux0) case Nil with True show ?thesis unfolding aux'_eq by simp next case (Cons a as) with True show ?thesis unfolding aux'_eq using eq_fst_iff[of t a] by (cases "fst a = \ \ ne") auto qed next case False with that have "ne \ 0" using le_\_less neq0_conv by blast moreover from False that have "t \ \ \ (ne-1)" by (metis One_nat_def Suc_leI Suc_pred \_mono diff_is_0_eq' order.antisym neq0_conv not_le) ultimately obtain X where "(t, X) \ set aux0" using \\ \ ne - t \ right I\ \\i. \ \ i = t\ by atomize_elim (auto intro!: in_aux0_2) then show ?thesis unfolding aux'_eq using False by (auto intro: exI[of _ X] split: list.split) qed show "wf_matchP_aux \ V n R I r aux' (Suc ne)" unfolding wf_matchP_aux_def using mr by (auto dest: in_aux'_1 intro: sorted_aux' in_aux'_2) have rel_eq: "rel = foldr (\) [lookup rel mr. (t, rel) \ aux', left I \ \ \ ne - t] {}" unfolding aux'_eq aux0_def using result_eq by (simp add: update_matchP_def Let_def) have rel_alt: "rel = (\(t, rel) \ set aux'. if left I \ \ \ ne - t then lookup rel mr else empty_table)" unfolding rel_eq by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI) show "qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP I r)) rel" unfolding rel_alt proof (rule qtable_Union[where Qi="\(t, X) v. left I \ \ \ ne - t \ Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) r)"], goal_cases finite qtable equiv) case (equiv v) show ?case proof (rule iffI, erule sat_MatchP_point, goal_cases left right) case (left j) then show ?case using in_aux'_2[of "\ \ j", OF _ _ exI, OF _ _ refl] by auto next case right then show ?case by (auto elim!: sat_MatchP_pointD dest: in_aux'_1) qed qed (auto dest!: in_aux'_1 intro!: qtable_empty dest!: bspec[OF _ RPDs_refl] simp: from_mregex_eq[OF safe mr]) qed lemma length_update_until: "length (update_until args rel1 rel2 nt aux) = Suc (length aux)" unfolding update_until_def by simp lemma wf_update_until_auxlist: assumes pre: "wf_until_auxlist \ V n R pos \ I \ auxlist ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length auxlist) \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length auxlist) \) rel2" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_pos: "args_pos args = pos" shows "wf_until_auxlist \ V n R pos \ I \ (update_until args rel1 rel2 (\ \ (ne + length auxlist)) auxlist) ne" unfolding wf_until_auxlist_def length_update_until unfolding update_until_def list.rel_map add_Suc_right upt.simps eqTrueI[OF le_add1] if_True proof (rule list_all2_appendI, unfold list.rel_map, goal_cases old new) case old show ?case proof (rule list.rel_mono_strong[OF assms(1)[unfolded wf_until_auxlist_def]]; safe, goal_cases mono1 mono2) case (mono1 i X Y v) then show ?case by (fastforce simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq elim!: qtable_join[OF _ qtable1] qtable_union[OF _ qtable1]) next case (mono2 i X Y v) then show ?case using fvi_subset by (auto 0 3 simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq split: if_splits elim!: qtable_union[OF _ qtable_join_fixed[OF qtable2]] elim: qtable_cong[OF _ refl] intro: exI[of _ "ne + length auxlist"]) (* slow 8 sec*) qed next case new then show ?case by (auto intro!: qtable_empty qtable1 qtable2[THEN qtable_cong] exI[of _ "ne + length auxlist"] simp: args_ivl args_n args_pos less_Suc_eq zero_enat_def[symmetric]) qed lemma (in muaux) wf_update_until: assumes pre: "wf_until_aux \ V R args \ \ aux ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length_muaux args aux) \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length_muaux args aux) \) rel2" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \" and args_R: "args_R args = Formula.fv \" and args_pos: "args_pos args = pos" shows "wf_until_aux \ V R args \ \ (add_new_muaux args rel1 rel2 (\ \ (ne + length_muaux args aux)) aux) ne \ length_muaux args (add_new_muaux args rel1 rel2 (\ \ (ne + length_muaux args aux)) aux) = Suc (length_muaux args aux)" proof - from pre obtain cur auxlist where valid_aux: "valid_muaux args cur aux auxlist" and cur: "cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1))" and pre_list: "wf_until_auxlist \ V n R pos \ I \ auxlist ne" unfolding wf_until_aux_def args_ivl args_n args_pos by auto have length_aux: "length_muaux args aux = length auxlist" using valid_length_muaux[OF valid_aux] . define nt where "nt \ \ \ (ne + length_muaux args aux)" have nt_mono: "cur \ nt" unfolding cur nt_def length_aux by simp define auxlist' where "auxlist' \ update_until args rel1 rel2 (\ \ (ne + length auxlist)) auxlist" have length_auxlist': "length auxlist' = Suc (length auxlist)" unfolding auxlist'_def by (auto simp add: length_update_until) have tab1: "table (args_n args) (args_L args) rel1" using qtable1 unfolding args_n[symmetric] args_L[symmetric] by (auto simp add: qtable_def) have tab2: "table (args_n args) (args_R args) rel2" using qtable2 unfolding args_n[symmetric] args_R[symmetric] by (auto simp add: qtable_def) have fv_sub: "fv \ \ fv \" using pre unfolding wf_until_aux_def by auto moreover have valid_add_new_auxlist: "valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux) auxlist'" using valid_add_new_muaux[OF valid_aux tab1 tab2 nt_mono] unfolding auxlist'_def nt_def length_aux . moreover have "length_muaux args (add_new_muaux args rel1 rel2 nt aux) = Suc (length_muaux args aux)" using valid_length_muaux[OF valid_add_new_auxlist] unfolding length_auxlist' length_aux[symmetric] . moreover have "wf_until_auxlist \ V n R pos \ I \ auxlist' ne" using wf_update_until_auxlist[OF pre_list qtable1[unfolded length_aux] qtable2[unfolded length_aux] fv_sub args_ivl args_n args_pos] unfolding auxlist'_def . moreover have "\ \ (ne + length auxlist) = (if ne + length auxlist' = 0 then 0 else \ \ (ne + length auxlist' - 1))" unfolding cur length_auxlist' by auto ultimately show ?thesis unfolding wf_until_aux_def nt_def length_aux args_ivl args_n args_pos by fast qed lemma length_update_matchF_base: "length (fst (update_matchF_base I mr mrs nt entry st)) = Suc 0" by (auto simp: Let_def update_matchF_base_def) lemma length_update_matchF_step: "length (fst (update_matchF_step I mr mrs nt entry st)) = Suc (length (fst st))" by (auto simp: Let_def update_matchF_step_def split: prod.splits) lemma length_foldr_update_matchF_step: "length (fst (foldr (update_matchF_step I mr mrs nt) aux base)) = length aux + length (fst base)" by (induct aux arbitrary: base) (auto simp: Let_def length_update_matchF_step) lemma length_update_matchF: "length (update_matchF n I mr mrs rels nt aux) = Suc (length aux)" unfolding update_matchF_def update_matchF_base_def length_foldr_update_matchF_step by (auto simp: Let_def) lemma wf_update_matchF_base_invar: assumes safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" shows "wf_matchF_invar \ V n R I r (update_matchF_base n I mr mrs rels (\ \ j)) j" proof - have from_mregex: "from_mregex mr \s = r" using safe mr using from_mregex_eq by blast { fix ms assume "ms \ LPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Futu Strict (from_mregex ms \s)" "ok (length \s) ms" "LPD ms \ LPDs mr" using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans by fastforce+ } note * = this show ?thesis unfolding wf_matchF_invar_def wf_matchF_aux_def update_matchF_base_def mr prod.case Let_def mrs using safe by (auto simp: * from_mregex qtables qtable_empty_iff zero_enat_def[symmetric] lookup_tabulate finite_LPDs eps_match less_Suc_eq LPDs_refl intro!: qtable_cong[OF qtable_l\_strict[where \s=\s]] intro: qtables exI[of _ j] split: option.splits) qed lemma Un_empty_table[simp]: "rel \ empty_table = rel" "empty_table \ rel = rel" unfolding empty_table_def by auto lemma wf_matchF_invar_step: assumes wf: "wf_matchF_invar \ V n R I r st (Suc i)" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" and rel: "qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. i \ j \ j < i + length (fst st) \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel" and entry: "entry = (\ \ i, rels, rel)" and nt: "nt = \ \ (i + length (fst st))" shows "wf_matchF_invar \ V n R I r (update_matchF_step I mr mrs nt entry st) i" proof - have from_mregex: "from_mregex mr \s = r" using safe mr using from_mregex_eq by blast { fix ms assume "ms \ LPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Futu Strict (from_mregex ms \s)" "ok (length \s) ms" "LPD ms \ LPDs mr" using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans by fastforce+ } note * = this { fix aux X ms assume "st = (aux, X)" "ms \ LPDs mr" with wf mr have "qtable n (fv_regex r) (mem_restr R) (\v. Regex.match (Formula.sat \ V (map the v)) (from_mregex ms \s) i (i + length aux)) (l\ (\x. x) X rels ms)" by (intro qtable_cong[OF qtable_l\[where \s=\s and A="fv_regex r" and Q="\v r. Regex.match (Formula.sat \ V v) r (Suc i) (i + length aux)", OF _ _ qtables]]) (auto simp: wf_matchF_invar_def * LPDs_trans lpd_match[of i] elim!: bspec) } note l\ = this have "lookup (mrtabulate mrs f) ms = f ms" if "ms \ LPDs mr" for ms and f :: "mregex \ 'a table" using that mrs by (fastforce simp: lookup_tabulate finite_LPDs split: option.splits)+ then show ?thesis using wf mr mrs entry nt LPDs_trans by (auto 0 3 simp: Let_def wf_matchF_invar_def update_matchF_step_def wf_matchF_aux_def mr * LPDs_refl list_all2_Cons1 append_eq_Cons_conv upt_eq_Cons_conv Suc_le_eq qtables lookup_tabulate finite_LPDs id_def l\ from_mregex less_Suc_eq intro!: qtable_union[OF rel l\] qtable_cong[OF rel] intro: exI[of _ "i + length _"] split: if_splits prod.splits) qed lemma wf_update_matchF_invar: assumes pre: "wf_matchF_aux \ V n R I r aux ne (length (fst st) - 1)" and wf: "wf_matchF_invar \ V n R I r st (ne + length aux)" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and j: "j = ne + length aux + length (fst st) - 1" shows "wf_matchF_invar \ V n R I r (foldr (update_matchF_step I mr mrs (\ \ j)) aux st) ne" using pre wf unfolding j proof (induct aux arbitrary: ne) case (Cons entry aux) from Cons(1)[of "Suc ne"] Cons(2,3) show ?case unfolding foldr.simps o_apply by (intro wf_matchF_invar_step[where rels = "fst (snd entry)" and rel = "snd (snd entry)"]) (auto simp: safe mr mrs wf_matchF_aux_def wf_matchF_invar_def list_all2_Cons1 append_eq_Cons_conv Suc_le_eq upt_eq_Cons_conv length_foldr_update_matchF_step add.assoc split: if_splits) qed simp lemma wf_update_matchF: assumes pre: "wf_matchF_aux \ V n R I r aux ne 0" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and nt: "nt = \ \ (ne + length aux)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length aux) \) rel) \s rels" shows "wf_matchF_aux \ V n R I r (update_matchF n I mr mrs rels nt aux) ne 0" unfolding update_matchF_def using wf_update_matchF_base_invar[OF safe mr mrs qtables, of I] unfolding nt by (intro wf_update_matchF_invar[OF _ _ safe mr mrs, unfolded wf_matchF_invar_def split_beta, THEN conjunct2, THEN conjunct1]) (auto simp: length_update_matchF_base wf_matchF_invar_def update_matchF_base_def Let_def pre) lemma wf_until_aux_Cons: "wf_until_auxlist \ V n R pos \ I \ (a # aux) ne \ wf_until_auxlist \ V n R pos \ I \ aux (Suc ne)" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong) lemma wf_matchF_aux_Cons: "wf_matchF_aux \ V n R I r (entry # aux) ne i \ wf_matchF_aux \ V n R I r aux (Suc ne) i" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong split: prod.splits) lemma wf_until_aux_Cons1: "wf_until_auxlist \ V n R pos \ I \ ((t, a1, a2) # aux) ne \ t = \ \ ne" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc) lemma wf_matchF_aux_Cons1: "wf_matchF_aux \ V n R I r ((t, rels, rel) # aux) ne i \ t = \ \ ne" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits) lemma wf_until_aux_Cons3: "wf_until_auxlist \ V n R pos \ I \ ((t, a1, a2) # aux) ne \ qtable n (Formula.fv \) (mem_restr R) (\v. (\j. ne \ j \ j < Suc (ne + length aux) \ mem (\ \ j - \ \ ne) I \ Formula.sat \ V (map the v) j \ \ (\k\{ne.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \))) a2" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc) lemma wf_matchF_aux_Cons3: "wf_matchF_aux \ V n R I r ((t, rels, rel) # aux) ne i \ qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. ne \ j \ j < Suc (ne + length aux + i) \ mem (\ \ j - \ \ ne) I \ Regex.match (Formula.sat \ V (map the v)) r ne j)) rel" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits) lemma upt_append: "a \ b \ b \ c \ [a.. ja'" "jb \ jb'" shows "wf_mbuf2 i ja' jb' P Q (mbuf2_add xs ys buf)" using assms unfolding wf_mbuf2_def by (auto 0 3 simp: list_all2_append2 upt_append dest: list_all2_lengthD intro: exI[where x="[i..j j'. [j..) js js'" shows "wf_mbufn i js' Ps (mbufn_add xss buf)" unfolding wf_mbufn_def list_all3_conv_all_nth proof safe show "length Ps = length js'" "length js' = length (mbufn_add xss buf)" using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth by auto next fix k assume k: "k < length Ps" then show "i \ js' ! k" using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth by (auto 0 4 dest: spec[of _ i]) from k have " [i..i z. \x y. P i x \ Q i y \ z = f x y) [i.. xs = [] \ ys = []" "list_all3 P xs [] zs \ xs = [] \ zs = []" "list_all3 P [] ys zs \ ys = [] \ zs = []" unfolding list_all3_conv_all_nth by auto lemma list_all3_Cons: "list_all3 P xs ys (z # zs) \ (\x xs' y ys'. xs = x # xs' \ ys = y # ys' \ P x y z \ list_all3 P xs' ys' zs)" "list_all3 P xs (y # ys) zs \ (\x xs' z zs'. xs = x # xs' \ zs = z # zs' \ P x y z \ list_all3 P xs' ys zs')" "list_all3 P (x # xs) ys zs \ (\y ys' z zs'. ys = y # ys' \ zs = z # zs' \ P x y z \ list_all3 P xs ys' zs')" unfolding list_all3_conv_all_nth by (auto simp: length_Suc_conv Suc_length_conv nth_Cons split: nat.splits) lemma list_all3_mono_strong: "list_all3 P xs ys zs \ (\x y z. x \ set xs \ y \ set ys \ z \ set zs \ P x y z \ Q x y z) \ list_all3 Q xs ys zs" by (induct xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) definition Mini where "Mini i js = (if js = [] then i else Min (set js))" lemma wf_mbufn_in_set_Mini: assumes "wf_mbufn i js Ps buf" shows "[] \ set buf \ Mini i js = i" unfolding in_set_conv_nth proof (elim exE conjE) fix k have "i \ j" if "j \ set js" for j using that assms unfolding wf_mbufn_def list_all3_conv_all_nth in_set_conv_nth by auto moreover assume "k < length buf" "buf ! k = []" ultimately show ?thesis using assms unfolding Mini_def wf_mbufn_def list_all3_conv_all_nth by (auto 0 4 dest!: spec[of _ k] intro: Min_eqI simp: in_set_conv_nth) qed lemma wf_mbufn_notin_set: assumes "wf_mbufn i js Ps buf" shows "[] \ set buf \ j \ set js \ i < j" using assms unfolding wf_mbufn_def list_all3_conv_all_nth by (cases "i \ set js") (auto intro: le_neq_implies_less simp: in_set_conv_nth) lemma wf_mbufn_map_tl: "wf_mbufn i js Ps buf \ [] \ set buf \ wf_mbufn (Suc i) js Ps (map tl buf)" by (auto simp: wf_mbufn_def list_all3_map Suc_le_eq dest: rel_funD[OF tl_transfer] elim!: list_all3_mono_strong le_neq_implies_less) lemma list_all3_list_all2I: "list_all3 (\x y z. Q x z) xs ys zs \ list_all2 Q xs zs" by (induct xs ys zs rule: list_all3.induct) auto lemma mbuf2t_take_eqD: assumes "mbuf2t_take f z buf nts = (z', buf', nts')" and "wf_mbuf2 i ja jb P Q buf" and "list_all2 R [i.. j" "jb \ j" shows "wf_mbuf2 (min ja jb) ja jb P Q buf'" and "list_all2 R [min ja jb.. set buf") case True from rec.prems(2) have "\j\set js. i \ j" by (auto simp: in_set_conv_nth list_all3_conv_all_nth) moreover from True rec.prems(2) have "i \ set js" by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) ultimately have "Mini i js = i" unfolding Mini_def by (auto intro!: antisym[OF Min.coboundedI Min.boundedI]) with rec.prems nonempty True show ?thesis by simp next case False from nonempty rec.prems(2) have "Mini i js = Mini (Suc i) js" unfolding Mini_def by auto show ?thesis unfolding \Mini i js = Mini (Suc i) js\ proof (rule rec.IH) show "\ (buf = [] \ [] \ set buf)" using nonempty False by simp show "list_all3 (\P j xs. Suc i \ j \ list_all2 P [Suc i..k. k \ set js \ k \ j" and "k = Mini (i + length nts) js" shows "wf_mbufn k js Ps buf'" and "list_all2 R [k.. []" using that by (fastforce simp flip: length_greater_0_conv) with 2 show ?case using wf_mbufn_in_set_Mini[OF 2(2)] wf_mbufn_notin_set[OF 2(2)] by (subst (asm) mbufnt_take.simps) (force simp: Mini_def wf_mbufn_def dest!: IH(2)[rotated, OF _ wf_mbufn_map_tl[OF 2(2)] *] split: if_splits prod.splits) qed lemma mbuf2t_take_induct[consumes 5, case_names base step]: assumes "mbuf2t_take f z buf nts = (z', buf', nts')" and "wf_mbuf2 i ja jb P Q buf" and "list_all2 R [i.. j" "jb \ j" and "U i z" and "\k x y t z. i \ k \ Suc k \ ja \ Suc k \ jb \ P k x \ Q k y \ R k t \ U k z \ U (Suc k) (f x y t z)" shows "U (min ja jb) z'" using assms unfolding wf_mbuf2_def by (induction f z buf nts arbitrary: i z' buf' nts' rule: mbuf2t_take.induct) (auto simp add: list_all2_Cons2 upt_eq_Cons_conv less_eq_Suc_le min_absorb1 min_absorb2 elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated] split: prod.split) lemma list_all2_hdD: assumes "list_all2 P [i.. []" shows "P i (hd xs)" "i < j" using assms unfolding list_all2_conv_all_nth by (auto simp: hd_conv_nth intro: zero_less_diff[THEN iffD1] dest!: spec[of _ 0]) lemma mbufn_take_induct[consumes 3, case_names base step]: assumes "mbufn_take f z buf = (z', buf')" and "wf_mbufn i js Ps buf" and "U i z" and "\k xs z. i \ k \ Suc k \ Mini i js \ list_all2 (\P x. P k x) Ps xs \ U k z \ U (Suc k) (f xs z)" shows "U (Mini i js) z'" using assms unfolding wf_mbufn_def proof (induction f z buf arbitrary: i z' buf' rule: mbufn_take.induct) case rec: (1 f z buf) show ?case proof (cases "buf = []") case True with rec.prems show ?thesis unfolding Mini_def by simp next case nonempty: False show ?thesis proof (cases "[] \ set buf") case True from rec.prems(2) have "\j\set js. i \ j" by (auto simp: in_set_conv_nth list_all3_conv_all_nth) moreover from True rec.prems(2) have "i \ set js" by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) ultimately have "Mini i js = i" unfolding Mini_def by (auto intro!: antisym[OF Min.coboundedI Min.boundedI]) with rec.prems nonempty True show ?thesis by simp next case False with nonempty rec.prems(2) have more: "Suc i \ Mini i js" using diff_is_0_eq not_le unfolding Mini_def by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) then have "Mini i js = Mini (Suc i) js" unfolding Mini_def by auto show ?thesis unfolding \Mini i js = Mini (Suc i) js\ proof (rule rec.IH) show "\ (buf = [] \ [] \ set buf)" using nonempty False by simp show "mbufn_take f (f (map hd buf) z) (map tl buf) = (z', buf')" using nonempty False rec.prems by simp show "list_all3 (\P j xs. Suc i \ j \ list_all2 P [Suc i..k xs z. Suc i \ k \ Suc k \ Mini (Suc i) js \ list_all2 (\P. P k) Ps xs \ U k z \ U (Suc k) (f xs z)" by (rule rec.prems(4)) (auto simp: Mini_def) qed qed qed qed lemma mbufnt_take_induct[consumes 5, case_names base step]: assumes "mbufnt_take f z buf nts = (z', buf', nts')" and "wf_mbufn i js Ps buf" and "list_all2 R [i..k. k \ set js \ k \ j" and "U i z" and "\k xs t z. i \ k \ Suc k \ Mini j js \ list_all2 (\P x. P k x) Ps xs \ R k t \ U k z \ U (Suc k) (f xs t z)" shows "U (Mini (i + length nts) js) z'" using assms proof (induction f z buf nts arbitrary: i z' buf' nts' rule: mbufnt_take.induct) case (1 f z buf nts) then have *: "list_all2 R [Suc i.. []" "nts = []" using that unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)] by (fastforce simp: Mini_def list_all3_Cons neq_Nil_conv) with 1(2-7) list_all2_hdD[OF 1(4)] show ?case unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)] wf_mbufn_notin_set[OF 1(3)] by (subst (asm) mbufnt_take.simps) (auto simp add: Mini_def list.rel_map Suc_le_eq elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated] list_all3_mono_strong[THEN list_all3_list_all2I[of _ _ js]] list_all2_hdD dest!: 1(1)[rotated, OF _ wf_mbufn_map_tl[OF 1(3)] * _ 1(7)] split: prod.split if_splits) qed lemma mbuf2_take_add': assumes eq: "mbuf2_take f (mbuf2_add xs ys buf) = (zs, buf')" and pre: "wf_mbuf2' \ P V j n R \ \ buf" and rm: "rel_mapping (\) P P'" and xs: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" shows "wf_mbuf2' \ P' V j' n R \ \ buf'" and "list_all2 (\i Z. \X Y. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) Y \ Z = f X Y) [min (progress \ P \ j) (progress \ P \ j).. P' \ j') (progress \ P' \ j')] zs" using pre rm unfolding wf_mbuf2'_def by (force intro!: mbuf2_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm])+ lemma mbuf2t_take_add': assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and pre_buf: "wf_mbuf2' \ P V j n R \ \ buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j)..i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" shows "wf_mbuf2' \ P' V j' n R \ \ buf'" and "wf_ts \ P' j' \ \ nts'" using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def wf_ts_def by (auto intro!: mbuf2t_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm] progress_le_gen) lemma ok_0_atms: "ok 0 mr \ regex.atms (from_mregex mr []) = {}" by (induct mr) auto lemma ok_0_progress: "ok 0 mr \ progress_regex \ P (from_mregex mr []) j = j" by (drule ok_0_atms) (auto simp: progress_regex_def) lemma atms_empty_atms: "safe_regex m g r \ atms r = {} \ regex.atms r = {}" by (induct r rule: safe_regex_induct) (auto split: if_splits simp: cases_Neg_iff) lemma atms_empty_progress: "safe_regex m g r \ atms r = {} \ progress_regex \ P r j = j" by (drule atms_empty_atms) (auto simp: progress_regex_def) lemma to_mregex_empty_progress: "safe_regex m g r \ to_mregex r = (mr, []) \ progress_regex \ P r j = j" using from_mregex_eq ok_0_progress to_mregex_ok atms_empty_atms by fastforce lemma progress_regex_le: "pred_mapping (\x. x \ j) P \ progress_regex \ P r j \ j" by (auto intro!: progress_le_gen simp: Min_le_iff progress_regex_def) lemma Neg_acyclic: "formula.Neg x = x \ P" by (induct x) auto lemma case_Neg_in_iff: "x \ (case y of formula.Neg \' \ {\'} | _ \ {}) \ y = formula.Neg x" by (cases y) auto lemma atms_nonempty_progress: "safe_regex m g r \ atms r \ {} \ (\\. progress \ P \ j) ` atms r = (\\. progress \ P \ j) ` regex.atms r" by (frule atms_empty_atms; simp) (auto 0 3 simp: atms_def image_iff case_Neg_in_iff elim!: disjE_Not2 dest: safe_regex_safe_formula) lemma to_mregex_nonempty_progress: "safe_regex m g r \ to_mregex r = (mr, \s) \ \s \ [] \ progress_regex \ P r j = (MIN \\set \s. progress \ P \ j)" using atms_nonempty_progress to_mregex_ok unfolding progress_regex_def by fastforce lemma to_mregex_progress: "safe_regex m g r \ to_mregex r = (mr, \s) \ progress_regex \ P r j = (if \s = [] then j else (MIN \\set \s. progress \ P \ j))" using to_mregex_nonempty_progress to_mregex_empty_progress unfolding progress_regex_def by auto lemma mbufnt_take_add': assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and safe: "safe_regex m g r" and mr: "to_mregex r = (mr, \s)" and pre_buf: "wf_mbufn' \ P V j n R r buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j..\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) (map2 upt (map (\\. progress \ P \ j) \s) (map (\\. progress \ P' \ j') \s)) xss" and "j \ j'" shows "wf_mbufn' \ P' V j' n R r buf'" and "wf_ts_regex \ P' j' r nts'" using pre_buf pre_nts bounded rm mr safe xss \j \ j'\ unfolding wf_mbufn'_def wf_ts_regex_def using atms_empty_progress[of m g r] to_mregex_ok[OF mr] by (auto 0 3 simp: list.rel_map to_mregex_empty_progress to_mregex_nonempty_progress Mini_def intro: progress_mono_gen[OF \j \ j'\ rm] list.rel_refl_strong progress_le_gen dest: list_all2_lengthD elim!: mbufnt_take_eqD[OF eq wf_mbufn_add]) lemma mbuf2t_take_add_induct'[consumes 6, case_names base step]: assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and pre_buf: "wf_mbuf2' \ P V j n R \ \ buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j)..i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" and base: "U (min (progress \ P \ j) (progress \ P \ j)) z" and step: "\k X Y z. min (progress \ P \ j) (progress \ P \ j) \ k \ Suc k \ progress \ P' \ j' \ Suc k \ progress \ P' \ j' \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \) X \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \) Y \ U k z \ U (Suc k) (f X Y (\ \ k) z)" shows "U (min (progress \ P' \ j') (progress \ P' \ j')) z'" using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def by (blast intro!: mbuf2t_take_induct[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm] progress_le_gen base step) lemma mbufnt_take_add_induct'[consumes 6, case_names base step]: assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and safe: "safe_regex m g r" and mr: "to_mregex r = (mr, \s)" and pre_buf: "wf_mbufn' \ P V j n R r buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j..\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) (map2 upt (map (\\. progress \ P \ j) \s) (map (\\. progress \ P' \ j') \s)) xss" and "j \ j'" and base: "U (progress_regex \ P r j) z" and step: "\k Xs z. progress_regex \ P r j \ k \ Suc k \ progress_regex \ P' r j' \ list_all2 (\\. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \)) \s Xs \ U k z \ U (Suc k) (f Xs (\ \ k) z)" shows "U (progress_regex \ P' r j') z'" using pre_buf pre_nts bounded rm \j \ j'\ to_mregex_progress[OF safe mr, of \ P' j'] to_mregex_empty_progress[OF safe, of mr \ P j] base unfolding wf_mbufn'_def mr prod.case by (fastforce dest!: mbufnt_take_induct[OF eq wf_mbufn_add[OF _ xss] pre_nts, of U] simp: list.rel_map le_imp_diff_is_add ac_simps Mini_def intro: progress_mono_gen[OF \j \ j'\ rm] list.rel_refl_strong progress_le_gen intro!: base step dest: list_all2_lengthD split: if_splits) lemma progress_Until_le: "progress \ P (Formula.Until \ I \) j \ min (progress \ P \ j) (progress \ P \ j)" by (cases "right I") (auto simp: trans_le_add1 intro!: cInf_lower) lemma progress_MatchF_le: "progress \ P (Formula.MatchF I r) j \ progress_regex \ P r j" by (cases "right I") (auto simp: trans_le_add1 progress_regex_def intro!: cInf_lower) lemma list_all2_upt_Cons: "P a x \ list_all2 P [Suc a.. Suc a \ b \ list_all2 P [a.. list_all2 P [b.. a \ b \ b \ c \ list_all2 P [a..x. R x x) xs ys" by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth) lemma map_split_map: "map_split f (map g xs) = map_split (f o g) xs" by (induct xs) auto lemma map_split_alt: "map_split f xs = (map (fst o f) xs, map (snd o f) xs)" by (induct xs) (auto split: prod.splits) lemma fv_formula_of_constraint: "fv (formula_of_constraint (t1, p, c, t2)) = fv_trm t1 \ fv_trm t2" by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct) simp_all lemma (in maux) wf_mformula_wf_set: "wf_mformula \ j P V n R \ \' \ wf_set n (Formula.fv \')" unfolding wf_set_def proof (induction rule: wf_mformula.induct) case (AndRel P V n R \ \' \' conf) then show ?case by (auto simp: fv_formula_of_constraint dest!: subsetD) next case (Ands P V n R l l_pos l_neg l' buf A_pos A_neg) from Ands.IH have "\\'\set (l_pos @ map remove_neg l_neg). \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of "_ @ _"] del: set_append) then have "\\'\set (l_pos @ l_neg). \x\fv \'. x < n" by (auto dest: bspec[where x="remove_neg _"]) then show ?case using Ands.hyps(2) by auto next case (Agg P V b n R \ \' y f g0 \) then have "Formula.fvi_trm b f \ Formula.fvi b \'" by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b]) with Agg show ?case by (auto 0 3 simp: Un_absorb2 fvi_iff_fv[where b=b]) next case (MatchP r P V n R \s mr mrs buf nts I aux) then obtain \s' where conv: "to_mregex r = (mr, \s')" by blast with MatchP have "\\'\set \s'. \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of \s']) with conv show ?case by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF \safe_regex _ _ r\]) next case (MatchF r P V n R \s mr mrs buf nts I aux) then obtain \s' where conv: "to_mregex r = (mr, \s')" by blast with MatchF have "\\'\set \s'. \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of \s']) with conv show ?case by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF \safe_regex _ _ r\]) qed (auto simp: fvi_Suc split: if_splits) lemma qtable_mmulti_join: assumes pos: "list_all3 (\A Qi X. qtable n A P Qi X \ wf_set n A) A_pos Q_pos L_pos" and neg: "list_all3 (\A Qi X. qtable n A P Qi X \ wf_set n A) A_neg Q_neg L_neg" and C_eq: "C = \(set A_pos)" and L_eq: "L = L_pos @ L_neg" and "A_pos \ []" and fv_subset: "\(set A_neg) \ \(set A_pos)" and restrict_pos: "\x. wf_tuple n C x \ P x \ list_all (\A. P (restrict A x)) A_pos" and restrict_neg: "\x. wf_tuple n C x \ P x \ list_all (\A. P (restrict A x)) A_neg" and Qs: "\x. wf_tuple n C x \ P x \ Q x \ list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos \ list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" shows "qtable n C P Q (mmulti_join n A_pos A_neg L)" proof (rule qtableI) from pos have 1: "list_all2 (\A X. table n A X \ wf_set n A) A_pos L_pos" by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def) moreover from neg have "list_all2 (\A X. table n A X \ wf_set n A) A_neg L_neg" by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def) ultimately have L: "list_all2 (\A X. table n A X \ wf_set n A) (A_pos @ A_neg) (L_pos @ L_neg)" by (rule list_all2_appendI) note in_join_iff = mmulti_join_correct[OF \A_pos \ []\ L] from 1 have take_eq: "take (length A_pos) (L_pos @ L_neg) = L_pos" by (auto dest!: list_all2_lengthD) from 1 have drop_eq: "drop (length A_pos) (L_pos @ L_neg) = L_neg" by (auto dest!: list_all2_lengthD) note mmulti_join.simps[simp del] show "table n C (mmulti_join n A_pos A_neg L)" unfolding C_eq L_eq table_def by (simp add: in_join_iff) show "Q x" if "x \ mmulti_join n A_pos A_neg L" "wf_tuple n C x" "P x" for x using that(2,3) proof (rule Qs[THEN iffD2, OF _ _ conjI]) have pos': "list_all2 (\A. (\) (restrict A x)) A_pos L_pos" and neg': "list_all2 (\A. (\) (restrict A x)) A_neg L_neg" using that(1) unfolding L_eq in_join_iff take_eq drop_eq by simp_all show "list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos" using pos pos' restrict_pos that(2,3) by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def) have fv_subset': "\i. i < length A_neg \ A_neg ! i \ C" using fv_subset unfolding C_eq by (auto simp: Sup_le_iff) show "list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" using neg neg' restrict_neg that(2,3) by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def wf_tuple_restrict_simple[OF _ fv_subset']) qed show "x \ mmulti_join n A_pos A_neg L" if "wf_tuple n C x" "P x" "Q x" for x unfolding L_eq in_join_iff take_eq drop_eq proof (intro conjI) from that have pos': "list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos" and neg': "list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" using Qs[THEN iffD1] by auto show "wf_tuple n (\A\set A_pos. A) x" using \wf_tuple n C x\ unfolding C_eq by simp show "list_all2 (\A. (\) (restrict A x)) A_pos L_pos" using pos pos' restrict_pos that(1,2) by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def C_eq wf_tuple_restrict_simple[OF _ Sup_upper]) show "list_all2 (\A. (\) (restrict A x)) A_neg L_neg" using neg neg' restrict_neg that(1,2) by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def) qed qed lemma nth_filter: "i < length (filter P xs) \ (\i'. i' < length xs \ P (xs ! i') \ Q (xs ! i')) \ Q (filter P xs ! i)" by (metis (lifting) in_set_conv_nth set_filter mem_Collect_eq) lemma nth_partition: "i < length xs \ (\i'. i' < length (filter P xs) \ Q (filter P xs ! i')) \ (\i'. i' < length (filter (Not \ P) xs) \ Q (filter (Not \ P) xs ! i')) \ Q (xs ! i)" by (metis (no_types, lifting) in_set_conv_nth set_filter mem_Collect_eq comp_apply) lemma qtable_bin_join: assumes "qtable n A P Q1 X" "qtable n B P Q2 Y" "\ b \ B \ A" "C = A \ B" "\x. wf_tuple n C x \ P x \ P (restrict A x) \ P (restrict B x)" "\x. b \ wf_tuple n C x \ P x \ Q x \ Q1 (restrict A x) \ Q2 (restrict B x)" "\x. \ b \ wf_tuple n C x \ P x \ Q x \ Q1 (restrict A x) \ \ Q2 (restrict B x)" shows "qtable n C P Q (bin_join n A X b B Y)" using qtable_join[OF assms] bin_join_table[of n A X B Y b] assms(1,2) by (auto simp add: qtable_def) lemma restrict_update: "y \ A \ y < length x \ restrict A (x[y:=z]) = restrict A x" unfolding restrict_def by (auto simp add: nth_list_update) lemma qtable_assign: assumes "qtable n A P Q X" "y < n" "insert y A = A'" "y \ A" "\x'. wf_tuple n A' x' \ P x' \ P (restrict A x')" "\x. wf_tuple n A x \ P x \ Q x \ Q' (x[y:=Some (f x)])" "\x'. wf_tuple n A' x' \ P x' \ Q' x' \ Q (restrict A x') \ x' ! y = Some (f (restrict A x'))" shows "qtable n A' P Q' ((\x. x[y:=Some (f x)]) ` X)" (is "qtable _ _ _ _ ?Y") proof (rule qtableI) from assms(1) have "table n A X" unfolding qtable_def by simp then show "table n A' ?Y" unfolding table_def wf_tuple_def using assms(2,3) by (auto simp: nth_list_update) next fix x' assume "x' \ ?Y" "wf_tuple n A' x'" "P x'" then obtain x where "x \ X" and x'_eq: "x' = x[y:=Some (f x)]" by blast then have "wf_tuple n A x" using assms(1) unfolding qtable_def table_def by blast then have "y < length x" using assms(2) by (simp add: wf_tuple_def) with \wf_tuple n A x\ have "restrict A x' = x" unfolding x'_eq by (simp add: restrict_update[OF assms(4)] restrict_idle) with \wf_tuple n A' x'\ \P x'\ have "P x" using assms(5) by blast with \wf_tuple n A x\ \x \ X\ have "Q x" using assms(1) by (elim in_qtableE) with \wf_tuple n A x\ \P x\ show "Q' x'" unfolding x'_eq by (rule assms(6)) next fix x' assume "wf_tuple n A' x'" "P x'" "Q' x'" then have "wf_tuple n A (restrict A x')" using assms(3) by (auto intro!: wf_tuple_restrict_simple) moreover have "P (restrict A x')" using \wf_tuple n A' x'\ \P x'\ by (rule assms(5)) moreover have "Q (restrict A x')" and y: "x' ! y = Some (f (restrict A x'))" using \wf_tuple n A' x'\ \P x'\ \Q' x'\ by (auto dest!: assms(7)) ultimately have "restrict A x' \ X" by (intro in_qtableI[OF assms(1)]) moreover have "x' = (restrict A x')[y:=Some (f (restrict A x'))]" using y assms(2,3) \wf_tuple n A (restrict A x')\ \wf_tuple n A' x'\ by (auto simp: list_eq_iff_nth_eq wf_tuple_def nth_list_update nth_restrict) ultimately show "x' \ ?Y" by simp qed lemma sat_the_update: "y \ fv \ \ Formula.sat \ V (map the (x[y:=z])) i \ = Formula.sat \ V (map the x) i \" by (rule sat_fv_cong) (metis map_update nth_list_update_neq) lemma progress_constraint: "progress \ P (formula_of_constraint c) j = j" by (induction rule: formula_of_constraint.induct) simp_all lemma qtable_filter: assumes "qtable n A P Q X" "\x. wf_tuple n A x \ P x \ Q x \ R x \ Q' x" shows "qtable n A P Q' (Set.filter R X)" (is "qtable _ _ _ _ ?Y") proof (rule qtableI) from assms(1) have "table n A X" unfolding qtable_def by simp then show "table n A ?Y" unfolding table_def wf_tuple_def by simp next fix x assume "x \ ?Y" "wf_tuple n A x" "P x" with assms show "Q' x" by (auto elim!: in_qtableE) next fix x assume "wf_tuple n A x" "P x" "Q' x" with assms show "x \ Set.filter R X" by (auto intro!: in_qtableI) qed lemma eval_constraint_sat_eq: "wf_tuple n A x \ fv_trm t1 \ A \ fv_trm t2 \ A \ \i\A. i < n \ eval_constraint (t1, p, c, t2) x = Formula.sat \ V (map the x) i (formula_of_constraint (t1, p, c, t2))" by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct) (simp_all add: meval_trm_eval_trm) declare progress_le_gen[simp] definition "wf_envs \ j P P' V db = (dom V = dom P \ Mapping.keys db = dom P \ {p. p \ fst ` \ \ j} \ rel_mapping (\) P P' \ pred_mapping (\i. i \ j) P \ pred_mapping (\i. i \ Suc j) P' \ (\p \ Mapping.keys db - dom P. the (Mapping.lookup db p) = [{ts. (p, ts) \ \ \ j}]) \ (\p \ dom P. list_all2 (\i X. X = the (V p) i) [the (P p).. event_data list) set \ Formula.database" is "\X p. (if p \ fst ` X then Some [{ts. (p, ts) \ X}] else None)" . lemma wf_envs_mk_db: "wf_envs \ j Map.empty Map.empty Map.empty (mk_db (\ \ j))" unfolding wf_envs_def mk_db_def by transfer (force split: if_splits simp: image_iff rel_mapping_alt) lemma wf_envs_update: assumes wf_envs: "wf_envs \ j P P' V db" and m_eq: "m = Formula.nfv \" and in_fv: "{0 ..< m} \ fv \" and xs: "list_all2 (\i. qtable m (Formula.fv \) (mem_restr UNIV) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ (Suc j)] xs" shows "wf_envs \ j (P(p \ progress \ P \ j)) (P'(p \ progress \ P' \ (Suc j))) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \})) (Mapping.update p (map (image (map the)) xs) db)" unfolding wf_envs_def proof (intro conjI ballI, goal_cases) case 3 from assms show ?case by (auto simp: wf_envs_def pred_mapping_alt progress_le progress_mono_gen intro!: rel_mapping_map_upd) next case (6 p') with assms show ?case by (cases "p' \ dom P") (auto simp: wf_envs_def lookup_update') next case (7 p') from xs in_fv have "list_all2 (\x y. map the ` y = {v. length v = m \ Formula.sat \ V v x \}) [progress \ P \ j.. P' \ (Suc j)] xs" by (elim list.rel_mono_strong) (auto 0 3 simp: wf_tuple_def nth_append elim!: in_qtableE in_qtableI intro!: image_eqI[where x="map Some _"]) moreover have "list_all2 (\i X. X = the (V p') i) [the (P p').. p'" proof - from that 7 have "p' \ dom P" by simp with wf_envs show ?thesis by (simp add: wf_envs_def) qed ultimately show ?case by (simp add: list.rel_map image_iff lookup_update') qed (use assms in \auto simp: wf_envs_def\) lemma wf_envs_P_simps[simp]: "wf_envs \ j P P' V db \ pred_mapping (\i. i \ j) P" "wf_envs \ j P P' V db \ pred_mapping (\i. i \ Suc j) P'" "wf_envs \ j P P' V db \ rel_mapping (\) P P'" unfolding wf_envs_def by auto lemma wf_envs_progress_le[simp]: "wf_envs \ j P P' V db \ progress \ P \ j \ j" "wf_envs \ j P P' V db \ progress \ P' \ (Suc j) \ Suc j" unfolding wf_envs_def by auto lemma wf_envs_progress_regex_le[simp]: "wf_envs \ j P P' V db \ progress_regex \ P r j \ j" "wf_envs \ j P P' V db \ progress_regex \ P' r (Suc j) \ Suc j" unfolding wf_envs_def by (auto simp: progress_regex_le) lemma wf_envs_progress_mono[simp]: "wf_envs \ j P P' V db \ a \ b \ progress \ P \ a \ progress \ P' \ b" unfolding wf_envs_def by (auto simp: progress_mono_gen) lemma qtable_wf_tuple_cong: "qtable n A P Q X \ A = B \ (\v. wf_tuple n A v \ P v \ Q v = Q' v) \ qtable n B P Q' X" unfolding qtable_def table_def by blast lemma (in maux) meval: assumes "wf_mformula \ j P V n R \ \'" "wf_envs \ j P P' V db" shows "case meval n (\ \ j) db \ of (xs, \\<^sub>n) \ wf_mformula \ (Suc j) P' V n R \\<^sub>n \' \ list_all2 (\i. qtable n (Formula.fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [progress \ P \' j.. P' \' (Suc j)] xs" using assms proof (induction \ arbitrary: db P P' V n R \') case (MRel rel) then show ?case by (cases rule: wf_mformula.cases) (auto simp add: ball_Un intro: wf_mformula.intros table_eq_rel eq_rel_eval_trm in_eq_rel qtable_empty qtable_unit_table intro!: qtableI) next case (MPred e ts) then show ?case proof (cases "e \ dom P") case True with MPred(2) have "e \ Mapping.keys db" "e \ dom P'" "e \ dom V" "list_all2 (\i X. X = the (V e) i) [the (P e).. dom P'" "e \ dom V" unfolding wf_envs_def rel_mapping_alt by auto moreover from False MPred(2) have *: "e \ fst ` \ \ j \ e \ Mapping.keys db" unfolding wf_envs_def by auto from False MPred(2) have "e \ Mapping.keys db \ Mapping.lookup db e = Some [{ts. (e, ts) \ \ \ j}]" unfolding wf_envs_def keys_dom_lookup by (metis Diff_iff domD option.sel) with * have "(case Mapping.lookup db e of None \ [{}] | Some xs \ xs) = [{ts. (e, ts) \ \ \ j}]" by (cases "e \ fst ` \ \ j") (auto simp: image_iff keys_dom_lookup split: option.splits) ultimately show ?thesis by (cases rule: wf_mformula.cases) (fastforce intro!: wf_mformula.Pred qtableI bexI[where P="\x. _ = tabulate x 0 n", OF refl] elim!: list.rel_mono_strong bexI[rotated] dest: ex_match simp: list.rel_map table_def match_wf_tuple in_these_eq match_eval_trm image_iff list.map_comp) qed next case (MLet p m \1 \2) from MLet.prems(1) obtain \1' \2' where Let: "\' = Formula.Let p \1' \2'" and 1: "wf_mformula \ j P V m UNIV \1 \1'" and fv: "m = Formula.nfv \1'" "{0.. fv \1'" and 2: "wf_mformula \ j (P(p \ progress \ P \1' j)) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \1'})) n R \2 \2'" by (cases rule: wf_mformula.cases) auto obtain xs \1_new where e1: "meval m (\ \ j) db \1 = (xs, \1_new)" and wf1: "wf_mformula \ (Suc j) P' V m UNIV \1_new \1'" and res1: "list_all2 (\i. qtable m (fv \1') (mem_restr UNIV) (\v. Formula.sat \ V (map the v) i \1')) [progress \ P \1' j.. P' \1' (Suc j)] xs" using MLet(1)[OF 1(1) MLet.prems(2)] by (auto simp: eqTrueI[OF mem_restr_UNIV, abs_def]) from MLet(2)[OF 2 wf_envs_update[OF MLet.prems(2) fv res1]] wf1 e1 fv show ?case unfolding Let by (auto simp: fun_upd_def intro!: wf_mformula.Let) next case (MAnd A_\ \ pos A_\ \ buf) from MAnd.prems show ?case by (cases rule: wf_mformula.cases) (auto simp: sat_the_restrict simp del: bin_join.simps dest!: MAnd.IH split: if_splits prod.splits intro!: wf_mformula.And qtable_bin_join elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)]) next case (MAndAssign \ conf) from MAndAssign.prems obtain \'' x t \'' where wf_envs: "wf_envs \ j P P' V db" and \'_eq: "\' = formula.And \'' \''" and wf_\: "wf_mformula \ j P V n R \ \''" and "x < n" and "x \ fv \''" and fv_t_subset: "fv_trm t \ fv \''" and conf: "(x, t) = conf" and \''_eqs: "\'' = formula.Eq (trm.Var x) t \ \'' = formula.Eq t (trm.Var x)" by (cases rule: wf_mformula.cases) from wf_\ wf_envs obtain xs \\<^sub>n where meval_eq: "meval n (\ \ j) db \ = (xs, \\<^sub>n)" and wf_\\<^sub>n: "wf_mformula \ (Suc j) P' V n R \\<^sub>n \''" and xs: "list_all2 (\i. qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i \'')) [progress \ P \'' j.. P' \'' (Suc j)] xs" by (auto dest!: MAndAssign.IH) have progress_eqs: "progress \ P \' j = progress \ P \'' j" "progress \ P' \' (Suc j) = progress \ P' \'' (Suc j)" using \''_eqs wf_envs_progress_le[OF wf_envs] by (auto simp: \'_eq) show ?case proof (simp add: meval_eq, intro conjI) show "wf_mformula \ (Suc j) P' V n R (MAndAssign \\<^sub>n conf) \'" unfolding \'_eq by (rule wf_mformula.AndAssign) fact+ next show "list_all2 (\i. qtable n (fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [progress \ P \' j.. P' \' (Suc j)] (map ((`) (eval_assignment conf)) xs)" unfolding list.rel_map progress_eqs conf[symmetric] eval_assignment.simps using xs proof (rule list.rel_mono_strong) fix i X assume qtable: "qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i \'') X" then show "qtable n (fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \') ((\y. y[x := Some (meval_trm t y)]) ` X)" proof (rule qtable_assign) show "x < n" by fact show "insert x (fv \'') = fv \'" using \''_eqs fv_t_subset by (auto simp: \'_eq) show "x \ fv \''" by fact next fix v assume wf_v: "wf_tuple n (fv \') v" and "mem_restr R v" then show "mem_restr R (restrict (fv \'') v)" by simp assume sat: "Formula.sat \ V (map the v) i \'" then have A: "Formula.sat \ V (map the (restrict (fv \'') v)) i \''" (is ?A) by (simp add: \'_eq sat_the_restrict) have "map the v ! x = Formula.eval_trm (map the v) t" using sat \''_eqs by (auto simp: \'_eq) also have "... = Formula.eval_trm (map the (restrict (fv \'') v)) t" using fv_t_subset by (auto simp: map_the_restrict intro!: eval_trm_fv_cong) finally have "map the v ! x = meval_trm t (restrict (fv \'') v)" using meval_trm_eval_trm[of n "fv \''" "restrict (fv \'') v" t] fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_\] by (fastforce simp: \'_eq intro!: wf_tuple_restrict) then have B: "v ! x = Some (meval_trm t (restrict (fv \'') v))" (is ?B) using \''_eqs wf_v \x < n\ by (auto simp: wf_tuple_def \'_eq) from A B show "?A \ ?B" .. next fix v assume wf_v: "wf_tuple n (fv \'') v" and "mem_restr R v" and sat: "Formula.sat \ V (map the v) i \''" let ?v = "v[x := Some (meval_trm t v)]" from sat have A: "Formula.sat \ V (map the ?v) i \''" using \x \ fv \''\ by (simp add: sat_the_update) have "y \ fv_trm t \ x \ y" for y using fv_t_subset \x \ fv \''\ by auto then have B: "Formula.sat \ V (map the ?v) i \''" using \''_eqs meval_trm_eval_trm[of n "fv \''" v t] \x < n\ fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_\] by (auto simp: wf_tuple_def map_update intro!: eval_trm_fv_cong) from A B show "Formula.sat \ V (map the ?v) i \'" by (simp add: \'_eq) qed qed qed next case (MAndRel \ conf) from MAndRel.prems show ?case by (cases rule: wf_mformula.cases) (auto simp: progress_constraint progress_le list.rel_map fv_formula_of_constraint Un_absorb2 wf_mformula_wf_set[unfolded wf_set_def] split: prod.splits dest!: MAndRel.IH[where db=db and P=P and P'=P'] eval_constraint_sat_eq[THEN iffD2] intro!: wf_mformula.AndRel elim!: list.rel_mono_strong qtable_filter eval_constraint_sat_eq[THEN iffD1]) next case (MAnds A_pos A_neg l buf) note mbufn_take.simps[simp del] mbufn_add.simps[simp del] mmulti_join.simps[simp del] from MAnds.prems obtain pos neg l' where wf_l: "list_all2 (wf_mformula \ j P V n R) l (pos @ map remove_neg neg)" and wf_buf: "wf_mbufn (progress \ P (formula.Ands l') j) (map (\\. progress \ P \ j) (pos @ map remove_neg neg)) (map (\\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (pos @ map remove_neg neg)) buf" and posneg: "(pos, neg) = partition safe_formula l'" and "pos \ []" and safe_neg: "list_all safe_formula (map remove_neg neg)" and A_eq: "A_pos = map fv pos" "A_neg = map fv neg" and fv_subset: "\ (set A_neg) \ \ (set A_pos)" and "\' = Formula.Ands l'" by (cases rule: wf_mformula.cases) simp have progress_eq: "progress \ P' (formula.Ands l') (Suc j) = Mini (progress \ P (formula.Ands l') j) (map (\\. progress \ P' \ (Suc j)) (pos @ map remove_neg neg))" using \pos \ []\ posneg by (auto simp: Mini_def image_Un[symmetric] Collect_disj_eq[symmetric] intro!: arg_cong[where f=Min]) have join_ok: "qtable n (\ (fv ` set l')) (mem_restr R) (\v. list_all (Formula.sat \ V (map the v) k) l') (mmulti_join n A_pos A_neg L)" if args_ok: "list_all2 (\x. qtable n (fv x) (mem_restr R) (\v. Formula.sat \ V (map the v) k x)) (pos @ map remove_neg neg) L" for k L proof (rule qtable_mmulti_join) let ?ok = "\A Qi X. qtable n A (mem_restr R) Qi X \ wf_set n A" let ?L_pos = "take (length A_pos) L" let ?L_neg = "drop (length A_pos) L" have 1: "length pos \ length L" using args_ok by (auto dest!: list_all2_lengthD) show "list_all3 ?ok A_pos (map (\\ v. Formula.sat \ V (map the v) k \) pos) ?L_pos" using args_ok wf_l unfolding A_eq by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth nth_append split: if_splits intro!: wf_mformula_wf_set[of \ j P V n R] dest: order.strict_trans2[OF _ 1]) from args_ok have prems_neg: "list_all2 (\\. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k (remove_neg \))) neg ?L_neg" by (auto simp: A_eq list_all2_append1 list.rel_map) show "list_all3 ?ok A_neg (map (\\ v. Formula.sat \ V (map the v) k (remove_neg \)) neg) ?L_neg" using prems_neg wf_l unfolding A_eq by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length nth_append less_diff_conv split: if_splits intro!: wf_mformula_wf_set[of \ j P V n R _ "remove_neg _", simplified]) show "\(fv ` set l') = \(set A_pos)" using fv_subset posneg unfolding A_eq by auto show "L = take (length A_pos) L @ drop (length A_pos) L" by simp show "A_pos \ []" using \pos \ []\ A_eq by simp fix x :: "event_data tuple" assume "wf_tuple n (\ (fv ` set l')) x" and "mem_restr R x" then show "list_all (\A. mem_restr R (restrict A x)) A_pos" and "list_all (\A. mem_restr R (restrict A x)) A_neg" by (simp_all add: list.pred_set) have "list_all Formula.is_Neg neg" using posneg safe_neg by (auto 0 3 simp add: list.pred_map elim!: list.pred_mono_strong intro: formula.exhaust[of \ "Formula.is_Neg \" for \]) then have "list_all (\\. Formula.sat \ V (map the v) i (remove_neg \) \ \ Formula.sat \ V (map the v) i \) neg" for v i by (fastforce simp: Formula.is_Neg_def elim!: list.pred_mono_strong) then show "list_all (Formula.sat \ V (map the x) k) l' = (list_all2 (\A Qi. Qi (restrict A x)) A_pos (map (\\ v. Formula.sat \ V (map the v) k \) pos) \ list_all2 (\A Qi. \ Qi (restrict A x)) A_neg (map (\\ v. Formula.sat \ V (map the v) k (remove_neg \)) neg))" using posneg by (auto simp add: A_eq list_all2_conv_all_nth list_all_length sat_the_restrict elim: nth_filter nth_partition[where P=safe_formula and Q="Formula.sat _ _ _ _"]) qed fact from MAnds.prems(2) show ?case unfolding \\' = Formula.Ands l'\ by (auto 0 3 simp add: list.rel_map progress_eq map2_map_map list_all3_map list_all3_list_all2_conv list.pred_map simp del: set_append map_append progress_simps split: prod.splits intro!: wf_mformula.Ands[OF _ _ posneg \pos \ []\ safe_neg A_eq fv_subset] list.rel_mono_strong[OF wf_l] wf_mbufn_add[OF wf_buf] list.rel_flip[THEN iffD1, OF list.rel_mono_strong, OF wf_l] list.rel_refl join_ok[unfolded list.pred_set] dest!: MAnds.IH[OF _ _ MAnds.prems(2), rotated] elim!: wf_mbufn_take list_all2_appendI elim: mbufn_take_induct[OF _ wf_mbufn_add[OF wf_buf]]) next case (MOr \ \ buf) from MOr.prems show ?case by (cases rule: wf_mformula.cases) (auto dest!: MOr.IH split: if_splits prod.splits intro!: wf_mformula.Or qtable_union elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)]) next case (MNeg \) have *: "qtable n {} (mem_restr R) (\v. P v) X \ \ qtable n {} (mem_restr R) (\v. \ P v) empty_table \ x \ X \ False" for P x X using nullary_qtable_cases qtable_unit_empty_table by fastforce from MNeg.prems show ?case by (cases rule: wf_mformula.cases) (auto 0 4 intro!: wf_mformula.Neg dest!: MNeg.IH simp add: list.rel_map dest: nullary_qtable_cases qtable_unit_empty_table intro!: qtable_empty_unit_table elim!: list.rel_mono_strong elim: *) next case (MExists \) from MExists.prems show ?case by (cases rule: wf_mformula.cases) (force simp: list.rel_map fvi_Suc sat_fv_cong nth_Cons' intro!: wf_mformula.Exists dest!: MExists.IH qtable_project_fv elim!: list.rel_mono_strong table_fvi_tl qtable_cong sat_fv_cong[THEN iffD1, rotated -1] split: if_splits)+ next case (MAgg g0 y \ b f \) from MAgg.prems show ?case using wf_mformula_wf_set[OF MAgg.prems(1), unfolded wf_set_def] by (cases rule: wf_mformula.cases) (auto 0 3 simp add: list.rel_map simp del: sat.simps fvi.simps split: prod.split intro!: wf_mformula.Agg qtable_eval_agg dest!: MAgg.IH elim!: list.rel_mono_strong) next case (MPrev I \ first buf nts) from MPrev.prems show ?case proof (cases rule: wf_mformula.cases) case (Prev \) let ?xs = "fst (meval n (\ \ j) db \)" let ?\ = "snd (meval n (\ \ j) db \)" let ?ls = "fst (mprev_next I (buf @ ?xs) (nts @ [\ \ j]))" let ?rs = "fst (snd (mprev_next I (buf @ ?xs) (nts @ [\ \ j])))" let ?ts = "snd (snd (mprev_next I (buf @ ?xs) (nts @ [\ \ j])))" let ?P = "\i X. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X" let ?min = "min (progress \ P' \ (Suc j)) (Suc j - 1)" from Prev MPrev.IH[OF _ MPrev.prems(2), of n R \] have IH: "wf_mformula \ (Suc j) P' V n R ?\ \" and "list_all2 ?P [progress \ P \ j.. P' \ (Suc j)] ?xs" by auto with Prev(4,5) MPrev.prems(2) have "list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then ?P i X else X = empty_table) [min (progress \ P \ j) (j - 1).. list_all2 ?P [?min.. P' \ (Suc j)] ?rs \ list_all2 (\i t. t = \ \ i) [?min.. P \ j)) j = Suc (min (progress \ P \ j) (j-1))" if "j > 0" using that by auto ultimately show ?thesis using Prev(1,3) MPrev.prems(2) IH by (auto simp: map_Suc_upt[symmetric] upt_Suc[of 0] list.rel_map qtable_empty_iff simp del: upt_Suc elim!: wf_mformula.Prev list.rel_mono_strong split: prod.split if_split_asm) qed next case (MNext I \ first nts) from MNext.prems show ?case proof (cases rule: wf_mformula.cases) case (Next \) have min[simp]: "min (progress \ P \ j - Suc 0) (j - Suc 0) = progress \ P \ j - Suc 0" "min (progress \ P' \ (Suc j) - Suc 0) j = progress \ P' \ (Suc j) - Suc 0" using wf_envs_progress_le[OF MNext.prems(2), of \] by auto let ?xs = "fst (meval n (\ \ j) db \)" let ?ys = "case (?xs, first) of (_ # xs, True) \ xs | _ \ ?xs" let ?\ = "snd (meval n (\ \ j) db \)" let ?ls = "fst (mprev_next I ?ys (nts @ [\ \ j]))" let ?rs = "fst (snd (mprev_next I ?ys (nts @ [\ \ j])))" let ?ts = "snd (snd (mprev_next I ?ys (nts @ [\ \ j])))" let ?P = "\i X. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X" let ?min = "min (progress \ P' \ (Suc j) - 1) (Suc j - 1)" from Next MNext.IH[OF _ MNext.prems(2), of n R \] have IH: "wf_mformula \ (Suc j) P' V n R ?\ \" "list_all2 ?P [progress \ P \ j.. P' \ (Suc j)] ?xs" by auto with Next have "list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then ?P (Suc i) X else X = empty_table) [progress \ P \ j - 1.. list_all2 ?P [Suc ?min.. P' \ (Suc j)] ?rs \ list_all2 (\i t. t = \ \ i) [?min.. P \ j < progress \ P' \ (Suc j)" using that wf_envs_progress_le[OF MNext.prems(2), of \] by (intro mnext) (auto simp: list_all2_Cons2 upt_eq_Cons_conv intro!: list_all2_upt_append list_all2_appendI split: list.splits) then show ?thesis using wf_envs_progress_le[OF MNext.prems(2), of \] wf_envs_progress_mono[OF MNext.prems(2), of j "Suc j" \, simplified] Next IH by (cases "progress \ P' \ (Suc j) > progress \ P \ j") (auto 0 3 simp: qtable_empty_iff le_Suc_eq le_diff_conv elim!: wf_mformula.Next list.rel_mono_strong list_all2_appendI split: prod.split list.splits if_split_asm) (* slow 5 sec*) qed next case (MSince args \ \ buf nts aux) note sat.simps[simp del] from MSince.prems obtain \'' \''' \'' I where Since_eq: "\' = Formula.Since \''' I \''" and pos: "if args_pos args then \''' = \'' else \''' = Formula.Neg \''" and pos_eq: "safe_formula \''' = args_pos args" and \: "wf_mformula \ j P V n R \ \''" and \: "wf_mformula \ j P V n R \ \''" and fvi_subset: "Formula.fv \'' \ Formula.fv \''" and buf: "wf_mbuf2' \ P V j n R \'' \'' buf" and nts: "wf_ts \ P j \'' \'' nts" and aux: "wf_since_aux \ V R args \'' \'' aux (progress \ P (Formula.Since \''' I \'') j)" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \''" and args_R: "args_R args = Formula.fv \''" by (cases rule: wf_mformula.cases) (auto) have \''': "Formula.fv \''' = Formula.fv \''" "progress \ P \''' j = progress \ P \'' j" "progress \ P' \''' j = progress \ P' \'' j" for j using pos by (simp_all split: if_splits) from MSince.prems(2) have nts_snoc: "list_all2 (\i t. t = \ \ i) [min (progress \ P \'' j) (progress \ P \'' j).. \ j])" using nts unfolding wf_ts_def by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI) have update: "wf_since_aux \ V R args \'' \'' (snd (zs, aux')) (progress \ P' (Formula.Since \''' I \'') (Suc j)) \ list_all2 (\i. qtable n (Formula.fv \''' \ Formula.fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Since \''' I \''))) [progress \ P (Formula.Since \''' I \'') j.. P' (Formula.Since \''' I \'') (Suc j)] (fst (zs, aux'))" if eval_\: "fst (meval n (\ \ j) db \) = xs" and eval_\: "fst (meval n (\ \ j) db \) = ys" and eq: "mbuf2t_take (\r1 r2 t (zs, aux). case update_since args r1 r2 t aux of (z, x) \ (zs @ [z], x)) ([], aux) (mbuf2_add xs ys buf) (nts @ [\ \ j]) = ((zs, aux'), buf', nts')" for xs ys zs aux' buf' nts' unfolding progress_simps \''' proof (rule mbuf2t_take_add_induct'[where j=j and j'="Suc j" and z'="(zs, aux')", OF eq wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc], goal_cases xs ys _ base step) case xs then show ?case using MSince.IH(1)[OF \ MSince.prems(2)] eval_\ by auto next case ys then show ?case using MSince.IH(2)[OF \ MSince.prems(2)] eval_\ by auto next case base then show ?case using aux by (simp add: \''') next case (step k X Y z) then show ?case using fvi_subset pos by (auto 0 3 simp: args_ivl args_n args_L args_R Un_absorb1 elim!: update_since(1) list_all2_appendI dest!: update_since(2) split: prod.split if_splits) qed simp with MSince.IH(1)[OF \ MSince.prems(2)] MSince.IH(2)[OF \ MSince.prems(2)] show ?case by (auto 0 3 simp: Since_eq split: prod.split intro: wf_mformula.Since[OF _ _ pos pos_eq args_ivl args_n args_L args_R fvi_subset] elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc] mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc]) next case (MUntil args \ \ buf nts aux) note sat.simps[simp del] progress_simps[simp del] from MUntil.prems obtain \'' \''' \'' I where Until_eq: "\' = Formula.Until \''' I \''" and pos: "if args_pos args then \''' = \'' else \''' = Formula.Neg \''" and pos_eq: "safe_formula \''' = args_pos args" and \: "wf_mformula \ j P V n R \ \''" and \: "wf_mformula \ j P V n R \ \''" and fvi_subset: "Formula.fv \'' \ Formula.fv \''" and buf: "wf_mbuf2' \ P V j n R \'' \'' buf" and nts: "wf_ts \ P j \'' \'' nts" and aux: "wf_until_aux \ V R args \'' \'' aux (progress \ P (Formula.Until \''' I \'') j)" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \''" and args_R: "args_R args = Formula.fv \''" and length_aux: "progress \ P (Formula.Until \''' I \'') j + length_muaux args aux = min (progress \ P \'' j) (progress \ P \'' j)" by (cases rule: wf_mformula.cases) (auto) define pos where args_pos: "pos = args_pos args" have \''': "progress \ P \''' j = progress \ P \'' j" "progress \ P' \''' j = progress \ P' \'' j" for j using pos by (simp_all add: progress.simps split: if_splits) from MUntil.prems(2) have nts_snoc: "list_all2 (\i t. t = \ \ i) [min (progress \ P \'' j) (progress \ P \'' j).. \ j])" using nts unfolding wf_ts_def by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI) { fix xs ys zs aux' aux'' buf' nts' assume eval_\: "fst (meval n (\ \ j) db \) = xs" and eval_\: "fst (meval n (\ \ j) db \) = ys" and eq1: "mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [\ \ j]) = (aux', buf', nts')" and eq2: "eval_muaux args (case nts' of [] \ \ \ j | nt # _ \ nt) aux' = (zs, aux'')" define ne where "ne \ progress \ P (Formula.Until \''' I \'') j" have update1: "wf_until_aux \ V R args \'' \'' aux' (progress \ P (Formula.Until \''' I \'') j) \ ne + length_muaux args aux' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j))" using MUntil.IH(1)[OF \ MUntil.prems(2)] eval_\ MUntil.IH(2)[OF \ MUntil.prems(2)] eval_\ nts_snoc nts_snoc length_aux aux fvi_subset unfolding \''' by (elim mbuf2t_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MUntil.prems(2)] buf]) (auto simp: args_n args_L args_R ne_def wf_update_until) then obtain cur auxlist' where valid_aux': "valid_muaux args cur aux' auxlist'" and cur: "cur = (if ne + length auxlist' = 0 then 0 else \ \ (ne + length auxlist' - 1))" and wf_auxlist': "wf_until_auxlist \ V n R pos \'' I \'' auxlist' (progress \ P (Formula.Until \''' I \'') j)" unfolding wf_until_aux_def ne_def args_ivl args_n args_pos by auto have length_aux': "length_muaux args aux' = length auxlist'" using valid_length_muaux[OF valid_aux'] . have nts': "wf_ts \ P' (Suc j) \'' \'' nts'" using MUntil.IH(1)[OF \ MUntil.prems(2)] eval_\ MUntil.IH(2)[OF \ MUntil.prems(2)] MUntil.prems(2) eval_\ nts_snoc unfolding wf_ts_def by (intro mbuf2t_take_eqD(2)[OF eq1]) (auto intro: wf_mbuf2_add buf[unfolded wf_mbuf2'_def]) define zs'' where "zs'' = fst (eval_until I (case nts' of [] \ \ \ j | nt # x \ nt) auxlist')" define auxlist'' where "auxlist'' = snd (eval_until I (case nts' of [] \ \ \ j | nt # x \ nt) auxlist')" have current_w_le: "cur \ (case nts' of [] \ \ \ j | nt # x \ nt)" proof (cases nts') case Nil have p_le: "min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) - 1 \ j" using wf_envs_progress_le[OF MUntil.prems(2)] by (auto simp: min_def le_diff_conv) then show ?thesis unfolding cur conjunct2[OF update1, unfolded length_aux'] using Nil by auto next case (Cons nt x) have progress_\''': "progress \ P' \'' (Suc j) = progress \ P' \''' (Suc j)" using pos by (auto simp add: progress.simps split: if_splits) have "nt = \ \ (min (progress \ P' \'' (Suc j)) (progress \ P' \'' (Suc j)))" using nts'[unfolded wf_ts_def Cons] unfolding list_all2_Cons2 upt_eq_Cons_conv by auto then show ?thesis unfolding cur conjunct2[OF update1, unfolded length_aux'] Cons progress_\''' by (auto split: if_splits list.splits intro!: \_mono) qed have valid_aux'': "valid_muaux args cur aux'' auxlist''" using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist''] by (auto simp add: args_ivl zs''_def auxlist''_def) have length_aux'': "length_muaux args aux'' = length auxlist''" using valid_length_muaux[OF valid_aux''] . have eq2': "eval_until I (case nts' of [] \ \ \ j | nt # _ \ nt) auxlist' = (zs, auxlist'')" using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist''] by (auto simp add: args_ivl zs''_def auxlist''_def) have length_aux'_aux'': "length_muaux args aux' = length zs + length_muaux args aux''" using eval_until_length[OF eq2'] unfolding length_aux' length_aux'' . have "i \ progress \ P' (Formula.Until \''' I \'') (Suc j) \ wf_until_auxlist \ V n R pos \'' I \'' auxlist' i \ i + length auxlist' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ wf_until_auxlist \ V n R pos \'' I \'' auxlist'' (progress \ P' (Formula.Until \''' I \'') (Suc j)) \ i + length zs = progress \ P' (Formula.Until \''' I \'') (Suc j) \ i + length zs + length auxlist'' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ list_all2 (\i. qtable n (Formula.fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Until (if pos then \'' else Formula.Neg \'') I \''))) [i.. V n R pos \'' I \'' aux' (Suc i)" by (rule wf_until_aux_Cons) from Cons.prems(2) have 1: "t = \ \ i" unfolding \a = (t, a1, a2)\ by (rule wf_until_aux_Cons1) from Cons.prems(2) have 3: "qtable n (Formula.fv \'') (mem_restr R) (\v. (\j\i. j < Suc (i + length aux') \ mem (\ \ j - \ \ i) I \ Formula.sat \ V (map the v) j \'' \ (\k\{i.. V (map the v) k \'' else \ Formula.sat \ V (map the v) k \''))) a2" unfolding \a = (t, a1, a2)\ by (rule wf_until_aux_Cons3) from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j))" by simp have "i \ progress \ P' (Formula.Until \''' I \'') (Suc j)" if "enat (case nts' of [] \ \ \ j | nt # x \ nt) \ enat t + right I" using that nts' unfolding wf_ts_def progress.simps by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv \''' intro!: cInf_lower \_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits) moreover have "Suc i \ progress \ P' (Formula.Until \''' I \'') (Suc j)" if "enat t + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" proof - from that obtain m where m: "right I = enat m" by (cases "right I") auto have \_min: "\ \ (min j k) = min (\ \ j) (\ \ k)" for k by (simp add: min_of_mono monoI) have le_progress_iff[simp]: "(Suc j) \ progress \ P' \ (Suc j) \ progress \ P' \ (Suc j) = (Suc j)" for \ using wf_envs_progress_le[OF MUntil.prems(2), of \] by auto have min_Suc[simp]: "min j (Suc j) = j" by auto let ?X = "{i. \k. k < Suc j \ k \min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ enat (\ \ k) \ enat (\ \ i) + right I}" let ?min = "min j (min (progress \ P' \'' (Suc j)) (progress \ P' \'' (Suc j)))" have "\ \ ?min \ \ \ j" by (rule \_mono) auto from m have "?X \ {}" by (auto dest!: \_mono[of _ "progress \ P' \'' (Suc j)" \] simp: not_le not_less \''' intro!: exI[of _ "progress \ P' \'' (Suc j)"]) from m show ?thesis using that nts' unfolding wf_ts_def progress.simps by (intro cInf_greatest[OF \?X \ {}\]) (auto simp: 1 \''' not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq simp del: upt_Suc split: list.splits if_splits dest!: spec[of _ "?min"] less_le_trans[of "\ \ i + m" "\ \ _" "\ \ _ + m"] less_\D) qed moreover have *: "k < progress \ P' \ (Suc j)" if "enat (\ \ i) + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" "enat (\ \ k - \ \ i) \ right I" "\ = \'' \ \ = \''" for k \ proof - from that(1,2) obtain m where "right I = enat m" "\ \ i + m < (case nts' of [] \ \ \ j | nt # x \ nt)" "\ \ k - \ \ i \ m" by (cases "right I") auto with that(3) nts' progress_le[of \ \'' "Suc j"] progress_le[of \ \'' "Suc j"] show ?thesis unfolding wf_ts_def le_diff_conv by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "\ \ k"] less_\D) qed ultimately show ?case using Cons.prems Suc_i_aux'[simplified] unfolding \a = (t, a1, a2)\ by (auto simp: \''' 1 sat.simps upt_conv_Cons dest!: Cons.IH[OF _ aux' Suc_i_aux'] simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl]) qed thm this note wf_aux'' = this[OF wf_envs_progress_mono[OF MUntil.prems(2) le_SucI[OF order_refl]] wf_auxlist' conjunct2[OF update1, unfolded ne_def length_aux']] have "progress \ P (formula.Until \''' I \'') j + length auxlist' = progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist''" using wf_aux'' valid_aux'' length_aux'_aux'' by (auto simp add: ne_def length_aux' length_aux'') then have "cur = (if progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist'' = 0 then 0 else \ \ (progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist'' - 1))" unfolding cur ne_def by auto then have "wf_until_aux \ V R args \'' \'' aux'' (progress \ P' (formula.Until \''' I \'') (Suc j)) \ progress \ P (formula.Until \''' I \'') j + length zs = progress \ P' (formula.Until \''' I \'') (Suc j) \ progress \ P (formula.Until \''' I \'') j + length zs + length_muaux args aux'' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ list_all2 (\i. qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (formula.Until (if pos then \'' else formula.Neg \'') I \''))) [progress \ P (formula.Until \''' I \'') j.. P (formula.Until \''' I \'') j + length zs] zs" using wf_aux'' valid_aux'' fvi_subset unfolding wf_until_aux_def length_aux'' args_ivl args_n args_pos by (auto simp only: length_aux'') } note update = this from MUntil.IH(1)[OF \ MUntil.prems(2)] MUntil.IH(2)[OF \ MUntil.prems(2)] pos pos_eq fvi_subset show ?case by (auto 0 4 simp: args_ivl args_n args_pos Until_eq \''' progress.simps(6) split: prod.split if_splits dest!: update[OF refl refl, rotated] intro!: wf_mformula.Until[OF _ _ _ _ args_ivl args_n args_L args_R fvi_subset] elim!: list.rel_mono_strong qtable_cong elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc] mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc]) next case (MMatchP I mr mrs \s buf nts aux) note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del] from MMatchP.prems obtain r \s where eq: "\' = Formula.MatchP I r" and safe: "safe_regex Past Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (RPDs mr)" and \s: "list_all2 (wf_mformula \ j P V n R) \s \s" and buf: "wf_mbufn' \ P V j n R r buf" and nts: "wf_ts_regex \ P j r nts" and aux: "wf_matchP_aux \ V n R I r aux (progress \ P (Formula.MatchP I r) j)" by (cases rule: wf_mformula.cases) (auto) have nts_snoc: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. \ j])" using nts unfolding wf_ts_regex_def by (auto simp add: wf_envs_progress_regex_le[OF MMatchP.prems(2)] intro: list_all2_appendI) have update: "wf_matchP_aux \ V n R I r (snd (zs, aux')) (progress \ P' (Formula.MatchP I r) (Suc j)) \ list_all2 (\i. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.MatchP I r))) [progress \ P (Formula.MatchP I r) j.. P' (Formula.MatchP I r) (Suc j)] (fst (zs, aux'))" if eval: "map (fst o meval n (\ \ j) db) \s = xss" and eq: "mbufnt_take (\rels t (zs, aux). case update_matchP n I mr mrs rels t aux of (z, x) \ (zs @ [z], x)) ([], aux) (mbufn_add xss buf) (nts @ [\ \ j]) = ((zs, aux'), buf', nts')" for xss zs aux' buf' nts' unfolding progress_simps proof (rule mbufnt_take_add_induct'[where j'="Suc j" and z'="(zs, aux')", OF eq wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc], goal_cases xss _ base step) case xss then show ?case using eval \s by (auto simp: list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] dest!: MMatchP.IH(1)[OF _ _ MMatchP.prems(2)] elim!: list.rel_mono_strong split: prod.splits) next case base then show ?case using aux by auto next case (step k Xs z) then show ?case by (auto simp: Un_absorb1 mrs safe mr elim!: update_matchP(1) list_all2_appendI dest!: update_matchP(2) split: prod.split) qed simp then show ?case using \s by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ \s \s] list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc] mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc] dest!: MMatchP.IH[OF _ _ MMatchP.prems(2)] split: prod.splits) next case (MMatchF I mr mrs \s buf nts aux) note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del] progress_simps[simp del] from MMatchF.prems obtain r \s where eq: "\' = Formula.MatchF I r" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and \s: "list_all2 (wf_mformula \ j P V n R) \s \s" and buf: "wf_mbufn' \ P V j n R r buf" and nts: "wf_ts_regex \ P j r nts" and aux: "wf_matchF_aux \ V n R I r aux (progress \ P (Formula.MatchF I r) j) 0" and length_aux: "progress \ P (Formula.MatchF I r) j + length aux = progress_regex \ P r j" by (cases rule: wf_mformula.cases) (auto) have nts_snoc: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. \ j])" using nts unfolding wf_ts_regex_def by (auto simp add: wf_envs_progress_regex_le[OF MMatchF.prems(2)] intro: list_all2_appendI) { fix xss zs aux' aux'' buf' nts' assume eval: "map (fst o meval n (\ \ j) db) \s = xss" and eq1: "mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [\ \ j]) = (aux', buf', nts')" and eq2: "eval_matchF I mr (case nts' of [] \ \ \ j | nt # _ \ nt) aux' = (zs, aux'')" have update1: "wf_matchF_aux \ V n R I r aux' (progress \ P (Formula.MatchF I r) j) 0 \ progress \ P (Formula.MatchF I r) j + length aux' = progress_regex \ P' r (Suc j)" using eval nts_snoc nts_snoc length_aux aux \s by (elim mbufnt_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf]) (auto simp: length_update_matchF list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)] elim: wf_update_matchF[OF _ safe mr mrs] elim!: list.rel_mono_strong) from MMatchF.prems(2) have nts': "wf_ts_regex \ P' (Suc j) r nts'" using eval eval nts_snoc \s unfolding wf_ts_regex_def by (intro mbufnt_take_eqD(2)[OF eq1 wf_mbufn_add[where js'="map (\\. progress \ P' \ (Suc j)) \s", OF buf[unfolded wf_mbufn'_def mr prod.case]]]) (auto simp: to_mregex_progress[OF safe mr] Mini_def list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] list_all2_Cons1 elim!: list.rel_mono_strong intro!: list.rel_refl_strong dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)]) have "i \ progress \ P' (Formula.MatchF I r) (Suc j) \ wf_matchF_aux \ V n R I r aux' i 0 \ i + length aux' = progress_regex \ P' r (Suc j) \ wf_matchF_aux \ V n R I r aux'' (progress \ P' (Formula.MatchF I r) (Suc j)) 0 \ i + length zs = progress \ P' (Formula.MatchF I r) (Suc j) \ i + length zs + length aux'' = progress_regex \ P' r (Suc j) \ list_all2 (\i. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.MatchF I r))) [i.. V n R I r aux' (Suc i) 0" by (rule wf_matchF_aux_Cons) from Cons.prems(2) have 1: "t = \ \ i" unfolding \a = (t, rels, rel)\ by (rule wf_matchF_aux_Cons1) from Cons.prems(2) have 3: "qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j\i. j < Suc (i + length aux') \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel" unfolding \a = (t, rels, rel)\ using wf_matchF_aux_Cons3 by force from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' = progress_regex \ P' r (Suc j)" by simp have "i \ progress \ P' (Formula.MatchF I r) (Suc j)" if "enat (case nts' of [] \ \ \ j | nt # x \ nt) \ enat t + right I" using that nts' unfolding wf_ts_regex_def progress_simps by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv intro!: cInf_lower \_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits) moreover have "Suc i \ progress \ P' (Formula.MatchF I r) (Suc j)" if "enat t + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" proof - from that obtain m where m: "right I = enat m" by (cases "right I") auto have \_min: "\ \ (min j k) = min (\ \ j) (\ \ k)" for k by (simp add: min_of_mono monoI) have le_progress_iff[simp]: "Suc j \ progress \ P' \ (Suc j) \ progress \ P' \ (Suc j) = (Suc j)" for \ using wf_envs_progress_le[OF MMatchF.prems(2), of \] by auto have min_Suc[simp]: "min j (Suc j) = j" by auto let ?X = "{i. \k. k < Suc j \ k \ progress_regex \ P' r (Suc j) \ enat (\ \ k) \ enat (\ \ i) + right I}" let ?min = "min j (progress_regex \ P' r (Suc j))" have "\ \ ?min \ \ \ j" by (rule \_mono) auto from m have "?X \ {}" by (auto dest!: less_\D add_lessD1 simp: not_le not_less) from m show ?thesis using that nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r] unfolding wf_ts_regex_def progress_simps by (intro cInf_greatest[OF \?X \ {}\]) (auto simp: 1 not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq simp del: upt_Suc split: list.splits if_splits dest!: spec[of _ "?min"] less_le_trans[of "\ \ i + m" "\ \ _" "\ \ _ + m"] less_\D) qed moreover have *: "k < progress_regex \ P' r (Suc j)" if "enat (\ \ i) + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" "enat (\ \ k - \ \ i) \ right I" for k proof - from that(1,2) obtain m where "right I = enat m" "\ \ i + m < (case nts' of [] \ \ \ j | nt # x \ nt)" "\ \ k - \ \ i \ m" by (cases "right I") auto with nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r] show ?thesis unfolding wf_ts_regex_def le_diff_conv by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "\ \ k"] less_\D) qed ultimately show ?case using Cons.prems Suc_i_aux'[simplified] unfolding \a = (t, rels, rel)\ by (auto simp: 1 sat.simps upt_conv_Cons dest!: Cons.IH[OF _ aux' Suc_i_aux'] simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl]) qed note this[OF progress_mono_gen[OF le_SucI, OF order.refl] conjunct1[OF update1] conjunct2[OF update1]] } note update = this[OF refl, rotated] with MMatchF.prems(2) show ?case using \s by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ \s \s] list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc] mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc] dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)] update split: prod.splits) qed subsubsection \Monitor step\ lemma (in maux) wf_mstate_mstep: "wf_mstate \ \ R st \ last_ts \ \ snd tdb \ wf_mstate \ (psnoc \ tdb) R (snd (mstep (map_prod mk_db id tdb) st))" unfolding wf_mstate_def mstep_def Let_def by (fastforce simp add: progress_mono le_imp_diff_is_add split: prod.splits elim!: prefix_of_psnocE dest: meval[OF _ wf_envs_mk_db] list_all2_lengthD) definition "flatten_verdicts Vs = (\ (set (map (\(i, X). (\v. (i, v)) ` X) Vs)))" lemma flatten_verdicts_append[simp]: "flatten_verdicts (Vs @ Us) = flatten_verdicts Vs \ flatten_verdicts Us" by (induct Vs) (auto simp: flatten_verdicts_def) lemma (in maux) mstep_output_iff: assumes "wf_mstate \ \ R st" "last_ts \ \ snd tdb" "prefix_of (psnoc \ tdb) \" "mem_restr R v" shows "(i, v) \ flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) \ progress \ Map.empty \ (plen \) \ i \ i < progress \ Map.empty \ (Suc (plen \)) \ wf_tuple (Formula.nfv \) (Formula.fv \) v \ Formula.sat \ Map.empty (map the v) i \" proof - from prefix_of_psnocE[OF assms(3,2)] have "prefix_of \ \" "\ \ (plen \) = fst tdb" "\ \ (plen \) = snd tdb" by auto moreover from assms(1) \prefix_of \ \\ have "mstate_n st = Formula.nfv \" "mstate_i st = progress \ Map.empty \ (plen \)" "wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \" unfolding wf_mstate_def by blast+ moreover from meval[OF \wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \\ wf_envs_mk_db] obtain Vs st' where "meval (mstate_n st) (\ \ (plen \)) (mk_db (\ \ (plen \))) (mstate_m st) = (Vs, st')" "wf_mformula \ (Suc (plen \)) Map.empty Map.empty (mstate_n st) R st' \" "list_all2 (\i. qtable (mstate_n st) (fv \) (mem_restr R) (\v. Formula.sat \ Map.empty (map the v) i \)) [progress \ Map.empty \ (plen \).. Map.empty \ (Suc (plen \))] Vs" by blast moreover from this assms(4) have "qtable (mstate_n st) (fv \) (mem_restr R) (\v. Formula.sat \ Map.empty (map the v) i \) (Vs ! (i - progress \ Map.empty \ (plen \)))" if "progress \ Map.empty \ (plen \) \ i" "i < progress \ Map.empty \ (Suc (plen \))" using that by (auto simp: list_all2_conv_all_nth dest!: spec[of _ "(i - progress \ Map.empty \ (plen \))"]) ultimately show ?thesis using assms(4) unfolding mstep_def Let_def flatten_verdicts_def by (auto simp: in_set_enumerate_eq list_all2_conv_all_nth progress_mono le_imp_diff_is_add elim!: in_qtableE in_qtableI intro!: bexI[of _ "(i, Vs ! (i - progress \ Map.empty \ (plen \)))"]) qed subsubsection \Monitor function\ locale verimon = verimon_spec + maux lemma (in verimon) mstep_mverdicts: assumes wf: "wf_mstate \ \ R st" and le[simp]: "last_ts \ \ snd tdb" and restrict: "mem_restr R v" shows "(i, v) \ flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) \ (i, v) \ M (psnoc \ tdb) - M \" proof - obtain \ where p2: "prefix_of (psnoc \ tdb) \" using ex_prefix_of by blast with le have p1: "prefix_of \ \" by (blast elim!: prefix_of_psnocE) show ?thesis unfolding M_def by (auto 0 3 simp: p2 progress_prefix_conv[OF _ p1] sat_prefix_conv[OF _ p1] not_less pprogress_eq[OF p1] pprogress_eq[OF p2] dest: mstep_output_iff[OF wf le p2 restrict, THEN iffD1] spec[of _ \] mstep_output_iff[OF wf le _ restrict, THEN iffD1] progress_sat_cong[OF p1] intro: mstep_output_iff[OF wf le p2 restrict, THEN iffD2] p1) qed context maux begin primrec msteps0 where "msteps0 [] st = ([], st)" | "msteps0 (tdb # \) st = (let (V', st') = mstep (map_prod mk_db id tdb) st; (V'', st'') = msteps0 \ st' in (V' @ V'', st''))" primrec msteps0_stateless where "msteps0_stateless [] st = []" | "msteps0_stateless (tdb # \) st = (let (V', st') = mstep (map_prod mk_db id tdb) st in V' @ msteps0_stateless \ st')" lemma msteps0_msteps0_stateless: "fst (msteps0 w st) = msteps0_stateless w st" by (induct w arbitrary: st) (auto simp: split_beta) lift_definition msteps :: "Formula.prefix \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list \ ('msaux, 'muaux) mstate" is msteps0 . lift_definition msteps_stateless :: "Formula.prefix \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list" is msteps0_stateless . lemma msteps_msteps_stateless: "fst (msteps w st) = msteps_stateless w st" by transfer (rule msteps0_msteps0_stateless) lemma msteps0_snoc: "msteps0 (\ @ [tdb]) st = (let (V', st') = msteps0 \ st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))" by (induct \ arbitrary: st) (auto split: prod.splits) lemma msteps_psnoc: "last_ts \ \ snd tdb \ msteps (psnoc \ tdb) st = (let (V', st') = msteps \ st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))" by transfer' (auto simp: msteps0_snoc split: list.splits prod.splits if_splits) definition monitor where "monitor \ \ = msteps_stateless \ (minit_safe \)" end lemma Suc_length_conv_snoc: "(Suc n = length xs) = (\y ys. xs = ys @ [y] \ length ys = n)" by (cases xs rule: rev_cases) auto lemma (in verimon) wf_mstate_msteps: "wf_mstate \ \ R st \ mem_restr R v \ \ \ \' \ X = msteps (pdrop (plen \) \') st \ wf_mstate \ \' R (snd X) \ ((i, v) \ flatten_verdicts (fst X)) = ((i, v) \ M \' - M \)" proof (induct "plen \' - plen \" arbitrary: X st \ \') case 0 from 0(1,4,5) have "\ = \'" "X = ([], st)" by (transfer; auto)+ with 0(2) show ?case unfolding flatten_verdicts_def by simp next case (Suc x) from Suc(2,5) obtain \'' tdb where "x = plen \'' - plen \" "\ \ \''" "\' = psnoc \'' tdb" "pdrop (plen \) (psnoc \'' tdb) = psnoc (pdrop (plen \) \'') tdb" "last_ts (pdrop (plen \) \'') \ snd tdb" "last_ts \'' \ snd tdb" "\'' \ psnoc \'' tdb" proof (atomize_elim, transfer, elim exE, goal_cases prefix) case (prefix _ _ \' _ \_tdb) then show ?case proof (cases \_tdb rule: rev_cases) case (snoc \ tdb) with prefix show ?thesis by (intro bexI[of _ "\' @ \"] exI[of _ tdb]) (force simp: sorted_append append_eq_Cons_conv split: list.splits if_splits)+ qed simp qed with Suc(1)[OF this(1) Suc.prems(1,2) this(2) refl] Suc.prems show ?case unfolding msteps_msteps_stateless[symmetric] by (auto simp: msteps_psnoc split_beta mstep_mverdicts dest: mono_monitor[THEN set_mp, rotated] intro!: wf_mstate_mstep) qed lemma (in verimon) wf_mstate_msteps_stateless: assumes "wf_mstate \ \ R st" "mem_restr R v" "\ \ \'" shows "(i, v) \ flatten_verdicts (msteps_stateless (pdrop (plen \) \') st) \ (i, v) \ M \' - M \" using wf_mstate_msteps[OF assms refl] unfolding msteps_msteps_stateless by simp lemma (in verimon) wf_mstate_msteps_stateless_UNIV: "wf_mstate \ \ UNIV st \ \ \ \' \ flatten_verdicts (msteps_stateless (pdrop (plen \) \') st) = M \' - M \" by (auto dest: wf_mstate_msteps_stateless[OF _ mem_restr_UNIV]) lemma (in verimon) mverdicts_Nil: "M pnil = {}" by (simp add: M_def pprogress_eq) context maux begin lemma minit_safe_minit: "mmonitorable \ \ minit_safe \ = minit \" unfolding minit_safe_def monitorable_formula_code by simp lemma wf_mstate_minit_safe: "mmonitorable \ \ wf_mstate \ pnil R (minit_safe \)" using wf_mstate_minit minit_safe_minit mmonitorable_def by metis end lemma (in verimon) monitor_mverdicts: "flatten_verdicts (monitor \ \) = M \" unfolding monitor_def using monitorable by (subst wf_mstate_msteps_stateless_UNIV[OF wf_mstate_minit_safe, simplified]) (auto simp: mmonitorable_def mverdicts_Nil) subsection \Collected correctness results\ context verimon begin text \We summarize the main results proved above. \begin{enumerate} \item The term @{term M} describes semantically the monitor's expected behaviour: \begin{itemize} \item @{thm[source] mono_monitor}: @{thm mono_monitor[no_vars]} \item @{thm[source] sound_monitor}: @{thm sound_monitor[no_vars]} \item @{thm[source] complete_monitor}: @{thm complete_monitor[no_vars]} \item @{thm[source] sliceable_M}: @{thm sliceable_M[no_vars]} \end{itemize} \item The executable monitor's online interface @{term minit_safe} and @{term mstep} preserves the invariant @{term wf_mstate} and produces the the verdicts according to @{term M}: \begin{itemize} \item @{thm[source] wf_mstate_minit_safe}: @{thm wf_mstate_minit_safe[no_vars]} \item @{thm[source] wf_mstate_mstep}: @{thm wf_mstate_mstep[no_vars]} \item @{thm[source] mstep_mverdicts}: @{thm mstep_mverdicts[no_vars]} \end{itemize} \item The executable monitor's offline interface @{term monitor} implements @{term M}: \begin{itemize} \item @{thm[source] monitor_mverdicts}: @{thm monitor_mverdicts[no_vars]} \end{itemize} \end{enumerate} \ end (*<*) end (*>*) diff --git a/thys/Padic_Field/Padic_Field_Powers.thy b/thys/Padic_Field/Padic_Field_Powers.thy --- a/thys/Padic_Field/Padic_Field_Powers.thy +++ b/thys/Padic_Field/Padic_Field_Powers.thy @@ -1,11429 +1,11420 @@ theory Padic_Field_Powers imports Ring_Powers Padic_Field_Polynomials Generated_Boolean_Algebra Padic_Field_Topology begin text\This theory is intended to develop the necessary background on subsets of powers of a $p$-adic field to prove Macintyre's quantifier elimination theorem. In particular, we define semi-algebraic subsets of $\mathbb{Q}_p^n$, semi-algebraic functions $\mathbb{Q}_p^n \to \mathbb{Q}_p$, and semi- algebraic mappings $\mathbb{Q}_p^n \to \mathbb{Q}_p^m$ for arbitrary $n, m \in \mathbb{N}$. In addition we prove that many common sets and functions are semi-algebraic. We are closely following the paper \<^cite>\"denef1986"\ by Denef, where an algebraic proof of Mactinyre's theorem is developed.\ (**************************************************************************************************) (**************************************************************************************************) section\Cartesian Powers of $p$-adic Fields\ (**************************************************************************************************) (**************************************************************************************************) lemma list_tl: "tl (t#x) = x" using List.list.sel(3) by auto lemma list_hd: "hd (t#x) = t" unfolding List.list.sel(1) by auto sublocale padic_fields < cring_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding cring_coord_rings_axioms_def cring_coord_rings_def using Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.is_UP_cring) by auto sublocale padic_fields < Qp: domain_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding domain_coord_rings_def cring_coord_rings_axioms_def cring_coord_rings_def using Qp.domain_axioms Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.UP_cring_axioms) by auto context padic_fields begin no_notation Zp.to_fun (infixl\\\ 70) no_notation ideal_prod (infixl "\\" 80) notation evimage (infixr "\\" 90) and euminus_set ("_ \<^sup>c\" 70) type_synonym padic_tuple = "padic_number list" type_synonym padic_function = "padic_number \ padic_number" type_synonym padic_nary_function = "padic_tuple \ padic_number" type_synonym padic_function_tuple = "padic_nary_function list" type_synonym padic_nary_function_poly = "nat \ padic_nary_function" (**************************************************************************************************) (**************************************************************************************************) subsection\Polynomials over $\mathbb{Q}_p$ and Polynomial Maps\ (**************************************************************************************************) (**************************************************************************************************) lemma last_closed': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "t \ carrier Q\<^sub>p" using assms last_closed[of n "x@[t]" Q\<^sub>p] by (metis (full_types) cartesian_power_car_memE gr0I last_snoc length_append_singleton less_not_refl zero_less_Suc) lemma segment_in_car': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" proof- have 0: "length x = n" by (metis Suc_inject assms cartesian_power_car_memE length_append_singleton) have "set x \ set (x@[t])" by (metis rotate1.simps(2) set_rotate1 set_subset_Cons) then have 1: "set x \ carrier Q\<^sub>p" using assms cartesian_power_car_memE''[of "x@[t]" Q\<^sub>p "Suc n"] by blast show ?thesis using 0 1 assms cartesian_power_car_memI[of x n Q\<^sub>p] by blast qed lemma Qp_zero: "Q\<^sub>p\<^bsup>0\<^esup> = nil_ring" unfolding cartesian_power_def by simp lemma Qp_zero_carrier: "carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" by (simp add: Qp_zero) text\Abbreviation for constant polynomials\ abbreviation(input) Qp_to_IP where "Qp_to_IP k \ Qp.indexed_const k" lemma Qp_to_IP_car: assumes "k \ carrier Q\<^sub>p" shows "Qp_to_IP k \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Qp.indexed_const_closed by blast lemma(in cring_coord_rings) smult_closed: assumes "a \ carrier R" assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q \ carrier (R[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Pring_smult_closed by (simp add: R.Pring_smult_closed) lemma Qp_poly_smult_cfs: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) m = a \ (P m)" using assms unfolding coord_ring_def using Qp.Pring_smult_cfs by blast lemma Qp_smult_r_distr: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q) = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q)" using assms unfolding coord_ring_def using Qp.Pring_smult_r_distr by blast lemma Qp_smult_l_distr: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \ b) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (b \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P)" using assms unfolding coord_ring_def using Qp.Pring_smult_l_distr by blast abbreviation(input) Qp_funs where "Qp_funs n \ Fun\<^bsub>n\<^esub> Q\<^sub>p" (**************************************************************************************************) (**************************************************************************************************) subsection\Evaluation of Polynomials in $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) Qp_ev where "Qp_ev P q \ (eval_at_point Q\<^sub>p q P)" lemma Qp_ev_one: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \" unfolding coord_ring_def by (metis Qp.Pring_one eval_at_point_const Qp.one_closed assms) lemma Qp_ev_zero: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \"unfolding coord_ring_def by (metis Qp.Pring_zero eval_at_point_const Qp.zero_closed assms) lemma Qp_eval_pvar_pow: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "k < n" assumes "(m::nat) \ 0" shows "Qp_ev ((pvar Q\<^sub>p k)[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> m) a = ((a!k)[^]m)" by (metis eval_at_point_nat_pow eval_pvar assms(1) assms(2) pvar_closed) text\composition of polynomials over $\mathbb{Q}_p$\ definition Qp_poly_comp where "Qp_poly_comp m fs = poly_compose (length fs) m fs" text\lemmas about polynomial maps\ lemma Qp_is_poly_tupleI: assumes "\i. i < length fs\ fs!i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_poly_tuple m fs" unfolding is_poly_tuple_def using assms using cartesian_power_car_memE'' cartesian_power_car_memI' by blast lemma Qp_is_poly_tuple_append: assumes "is_poly_tuple m fs" assumes "is_poly_tuple m gs" shows "is_poly_tuple m (fs@gs)" proof(rule Qp_is_poly_tupleI) show "\i. i < length (fs @ gs) \ (fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" proof- fix i assume A: "i < length (fs @ gs)" show "(fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" apply(cases "i < length fs") using assms is_poly_tupleE[of m fs i] nth_append[of fs gs i] apply presburger proof- assume B: "\ i < length fs" then have C: "length fs \ i \ i < length (fs @ gs)" using A not_le by blast then have "i - length fs < length gs" using length_append[of fs gs] by linarith then show ?thesis using A assms is_poly_tupleE[of m gs "i - length fs"] nth_append[of fs gs i] B by presburger qed qed qed lemma Qp_poly_mapE: assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j \ carrier Q\<^sub>p" using assms poly_map_closed cartesian_power_car_memE' by blast lemma Qp_poly_mapE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = length fs" unfolding poly_map_def using Qp.cring_axioms poly_tuple_evalE' by (metis assms restrict_def) lemma Qp_poly_mapE'': assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "n \ 0" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j = (Qp_ev (fs!j) as)" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) nth_map restrict_apply') lemma poly_map_apply: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n fs as = poly_tuple_eval fs as" unfolding poly_map_def restrict_def by (simp add: assms) lemma poly_map_pullbackI: assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ poly_map n fs \\<^bsub>n\<^esub> S" using assms poly_map_apply by blast lemma poly_map_pullbackI': assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ ((poly_map n fs) -` S)" by (simp add: assms(3)) text\lemmas about polynomial composition\ lemma poly_compose_ring_hom: assumes "is_poly_tuple m fs" assumes "length fs = n" shows "(ring_hom_ring (Q\<^sub>p[\\<^bsub>n\<^esub>]) (Q\<^sub>p[\\<^bsub>m\<^esub>]) (Qp_poly_comp m fs))" unfolding Qp_poly_comp_def by (simp add: assms(1) assms(2) poly_compose_ring_hom) lemma poly_compose_closed: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(Qp_poly_comp m fs f) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" using Qp.cring_axioms assms unfolding Qp_poly_comp_def using poly_compose_closed by blast lemma poly_compose_add: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_add unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_mult: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_mult unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_const: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "a \ carrier Q\<^sub>p" shows "Qp_poly_comp m fs (Qp_to_IP a) = Qp_to_IP a" using Qp.cring_axioms assms poly_compose_const unfolding is_poly_tuple_def Qp_poly_comp_def by metis lemma Qp_poly_comp_eval: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" shows "Qp_ev (Qp_poly_comp m fs f) as = Qp_ev f (poly_map m fs as)" proof- have "(restrict (poly_tuple_eval fs) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) as) = poly_tuple_eval fs as" unfolding restrict_def by (simp add: assms) thus ?thesis using assms Qp.cring_axioms poly_compose_eval unfolding Qp_poly_comp_def poly_map_def by presburger qed (**************************************************************************************************) (**************************************************************************************************) subsection\Mapping Univariate Polynomials to Multivariable Polynomials in One Variable\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) to_Qp_x where "to_Qp_x \ (IP_to_UP (0::nat) :: (nat multiset \ padic_number) \ nat \ padic_number)" abbreviation(input) from_Qp_x where "from_Qp_x \ UP_to_IP Q\<^sub>p (0::nat)" lemma from_Qp_x_closed: assumes "q \ carrier Q\<^sub>p_x" shows "from_Qp_x q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" using assms UP_to_IP_closed unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_closed: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "to_Qp_x q \ carrier Q\<^sub>p_x" using assms Qp.IP_to_UP_closed[of q "0::nat"] Qp.cring_axioms unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_from_Qp_x: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "from_Qp_x (to_Qp_x q) = q" using assms UP_to_IP_inv[of q "0::nat"] Qp.Pring_car unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_to_Qp_x: assumes "q \ carrier Q\<^sub>p_x" shows "to_Qp_x (from_Qp_x q) = q" by (meson UPQ.IP_to_UP_inv assms) text\ring hom properties of these maps\ lemma to_Qp_x_ring_hom: "to_Qp_x \ ring_hom (Q\<^sub>p[\\<^bsub>1\<^esub>]) Q\<^sub>p_x" using IP_to_UP_ring_hom[of "0::nat"] ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_ring_hom: "from_Qp_x \ ring_hom Q\<^sub>p_x (Q\<^sub>p[\\<^bsub>1\<^esub>])" using UP_to_IP_ring_hom ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_add: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis (mono_tags, lifting) assms(1) assms(2) from_Qp_x_ring_hom ring_hom_add) lemma from_Qp_x_mult: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis assms(1) assms(2) from_Qp_x_ring_hom ring_hom_mult) text\equivalence of evaluation maps\ lemma Qp_poly_Qp_x_eval: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" shows "Qp_ev P a = (to_Qp_x P)\(Qp.to_R a)" proof- have 0: "(IP_to_UP 0 P) \ (a ! 0) = ((IP_to_UP 0 P) \ (if 0 < length a then a ! 0 else \))" using assms by (metis (mono_tags, lifting) cartesian_power_car_memE gr0I zero_neq_one) have 1: "closed_fun Q\<^sub>p (\n. if n < length a then a ! n else \)" proof fix n show "(if n < length a then a ! n else \) \ carrier Q\<^sub>p" apply(cases "n < length a") using assms apply (metis cartesian_power_car_memE cartesian_power_car_memE') by (meson Qp.cring_axioms cring.cring_simprules(2)) qed have 2: " P \ Pring_set Q\<^sub>p {0::nat}" using assms unfolding coord_ring_def by (metis Qp.Pring_car UPQ.UP_to_IP_closed assms(1) to_Qp_x_closed to_Qp_x_from_Qp_x) have 3: "total_eval Q\<^sub>p (\i. if i < length a then a ! i else \) P = IP_to_UP 0 P \ (if 0 < length a then a ! 0 else \)" using 1 2 assms IP_to_UP_poly_eval[of P "0::nat" "(\i. if i < length a then a ! i else \)" ] UPQ.to_fun_def by presburger then show ?thesis using 0 unfolding eval_at_point_def by blast qed lemma Qp_x_Qp_poly_eval: assumes "P \ carrier Q\<^sub>p_x" assumes "a \ carrier Q\<^sub>p" shows "P \ a = Qp_ev (from_Qp_x P) (to_R1 a)" proof- have "Qp_ev (from_Qp_x P) (to_R1 a) = (to_Qp_x (from_Qp_x P))\(Qp.to_R (Qp.to_R1 a))" using Qp_poly_Qp_x_eval assms(1) assms(2) from_Qp_x_closed Qp.to_R1_closed by blast then show ?thesis using assms by (metis UPQ.IP_to_UP_inv Qp.to_R_to_R1) qed (**************************************************************************************************) (**************************************************************************************************) subsection\$n^{th}$-Power Sets over $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) definition P_set where "P_set (n::nat) = {a \ nonzero Q\<^sub>p. (\y \ carrier Q\<^sub>p . (y[^] n) = a)}" lemma P_set_carrier: "P_set n \ carrier Q\<^sub>p" unfolding P_set_def nonzero_def by blast lemma P_set_memI: assumes "a \ carrier Q\<^sub>p" assumes "a \ \" assumes "b \ carrier Q\<^sub>p" assumes "b[^](n::nat) = a" shows "a \ P_set n" unfolding P_set_def using assms by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) lemma P_set_nonzero: "P_set n \ nonzero Q\<^sub>p" unfolding P_set_def by blast lemma P_set_nonzero': assumes "a \ P_set n" shows "a \ nonzero Q\<^sub>p" "a \ carrier Q\<^sub>p" using assms P_set_nonzero P_set_carrier apply blast using assms P_set_carrier by blast lemma P_set_one: assumes "n \ 0" shows "\ \ P_set (n::nat)" proof- have 0: "\[^]n = \" using Qp.nat_pow_one by blast have 1: "\ \ carrier Q\<^sub>p" by blast then show ?thesis using one_nonzero unfolding P_set_def using 0 by blast qed lemma zeroth_P_set: "P_set 0 = {\}" proof show "P_set 0 \ {\}" unfolding P_set_def proof fix x assume "x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^](0::nat)) = a}" then have "\y\carrier Q\<^sub>p. (y[^](0::nat)) = x" by blast then obtain a where a_def: "a \ carrier Q\<^sub>p \ (a[^](0::nat)) = x" by blast then show "x \ {\}" using Qp.nat_pow_0 by blast qed show "{\} \ P_set 0" using P_set_memI[of \ \ 0] Qp.nat_pow_one Qp.one_closed local.one_neq_zero by blast qed lemma P_set_mult_closed: assumes "n \ 0" assumes "a \ P_set n" assumes "b \ P_set n" shows "a \ b \ P_set n" proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ (a0 [^] n = a)" using assms(2) unfolding P_set_def by blast obtain b0 where b0_def: "b0 \ carrier Q\<^sub>p \ (b0 [^] n = b)" using assms(3) unfolding P_set_def by blast have "(a0 \ b0) [^] n = a0 [^] n \ b0 [^] n" using a0_def b0_def Qp.nat_pow_distrib by blast then have 0: "a \ b = (a0 \ b0) [^] n" using a0_def b0_def by blast have 1: "a0 \ b0 \ carrier Q\<^sub>p" by (meson Qp.cring_axioms a0_def b0_def cring.cring_simprules(5)) have 2: "a \ b \ nonzero Q\<^sub>p" using assms nonzero_is_submonoid unfolding P_set_def by (metis (no_types, lifting) "0" "1" Qp.integral Qp_nat_pow_nonzero a0_def b0_def mem_Collect_eq not_nonzero_Qp) then show ?thesis using 0 1 assms unfolding P_set_def by blast qed lemma P_set_inv_closed: assumes "a \ P_set n" shows "inv a \ P_set n" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis Qp.inv_one singletonD) next case False then show ?thesis proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ a0[^]n = a" using assms P_set_def[of n] by blast have "a0 \ nonzero Q\<^sub>p" apply(rule ccontr) proof- assume "a0 \ nonzero Q\<^sub>p " then have "a0 = \" using a0_def by (meson not_nonzero_Qp) then show False using a0_def assms by (metis (mono_tags, lifting) False P_set_def Qp.cring_axioms \a0 \ nonzero Q\<^sub>p\ cring_def mem_Collect_eq neq0_conv ring.pow_zero) qed then have "(inv a0)[^]n = inv a" using a0_def \a0 \ carrier Q\<^sub>p \ (a0[^]n) = a\ \a0 \ nonzero Q\<^sub>p\ Units_nonzero monoid.nat_pow_of_inv[of Q\<^sub>p a n] Qp.nat_pow_of_inv Units_eq_nonzero by presburger then show ?thesis by (metis P_set_memI Qp.nat_pow_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero \a0 \ nonzero Q\<^sub>p\ a0_def inv_in_frac(1) inv_in_frac(2)) qed qed lemma P_set_val: assumes "a \ P_set (n::nat)" shows "(ord a) mod n = 0" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis mod_by_0 of_nat_0 ord_one singletonD) next case False then show ?thesis proof- obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^] n) = a" using assms P_set_def by blast have an: "a \ nonzero Q\<^sub>p" using P_set_def assms by blast have bn: "b \ nonzero Q\<^sub>p" proof(rule ccontr) assume "b \ nonzero Q\<^sub>p" then have "b = \\<^bsub> Q\<^sub>p\<^esub>" using b_def not_nonzero_Qp by metis then have "(b[^] n) = \" using False Qp.cring_axioms cring_def ring.pow_zero by blast then show False using b_def an Qp.not_nonzero_memI by blast qed then have "ord a = n * (ord b)" using b_def an nonzero_nat_pow_ord by blast then show ?thesis by (metis mod_mult_self1_is_0) qed qed lemma P_set_pow: assumes "n > 0" assumes "s \ P_set n" shows "s[^]k \ P_set (n*k)" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ y[^]n = s" using assms unfolding P_set_def by blast then have 0: "y \ nonzero Q\<^sub>p" using assms P_set_nonzero'(1) Qp_nonzero_nat_pow by blast have 1: "y[^](n*k) = s[^] k" using 0 y_def assms Qp.nat_pow_pow by blast hence 2: "s[^]k \ nonzero Q\<^sub>p" using 0 by (metis Qp_nat_pow_nonzero) thus ?thesis unfolding P_set_def using 1 y_def by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) text\ In this section we introduce the notion of a $p$-adic semialgebraic set. Intuitively, these are the subsets of $\mathbb{Q}_p^n$ which are definable by first order quantifier-free formulas in the standard first-order language of rings, with an additional relation symbol included for the relation $\text{ val}(x) \leq \text{ val}(y)$, interpreted according to the definiton of the $p$-adic valuation on $\mathbb{Q}_p$. In fact, by Macintyre's quantifier elimination theorem for the first-order theory of $\mathbb{Q}_p$ in this language, one can equivalently remove the ``quantifier-free" clause from the latter definition. The definition we give here is also equivalent, and due to Denef in \<^cite>\"denef1986"\. The given definition here is desirable mainly for its utility in producing a proof of Macintyre's theorem, which is our overarching goal. \ (********************************************************************************************) (********************************************************************************************) subsubsection\Defining Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition basic_semialg_set where "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). \y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)}" lemma basic_semialg_set_zero_set: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "Qp_ev P q = \" assumes "n \ 0" shows "q \ basic_semialg_set (m::nat) (n::nat) P" proof- have "\ = (\[^]n)" using assms(4) Qp.nat_pow_zero by blast then show ?thesis unfolding basic_semialg_set_def using assms Qp.cring_axioms cring.cring_simprules(2) by blast qed lemma basic_semialg_set_def': assumes "n \ 0" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ (P_set n)}" proof show "basic_semialg_set m n P \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" proof fix x assume A: "x \ basic_semialg_set m n P" show "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" apply(cases "Qp_ev P x = \") using A basic_semialg_set_def apply blast unfolding basic_semialg_set_def P_set_def proof assume A0: "Qp_ev P x \ \" have A1: " \y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n)" using A basic_semialg_set_def by blast have A2: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A basic_semialg_set_def by blast show " x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (Qp_ev P x = \ \ Qp_ev P x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a})" by (metis (mono_tags, lifting) A1 A2 Qp.nonzero_memI assms(2) eval_at_point_closed mem_Collect_eq) qed qed show "{q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} \ basic_semialg_set m n P" proof fix x assume A: " x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" then have A':"x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by blast show "x \ basic_semialg_set m n P" using A A' apply(cases "Qp_ev P x = \") using assms basic_semialg_set_zero_set[of P m x n] apply blast proof- assume B: "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} " assume B': "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assume B'': "Qp_ev P x \ \ " show "x \ basic_semialg_set m n P" unfolding basic_semialg_set_def P_set_def proof have "\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n) " using A nonzero_def [of Q\<^sub>p] unfolding P_set_def proof - assume "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a}}" then have "Qp_ev P x \ nonzero Q\<^sub>p \ (\r. r \ carrier Q\<^sub>p \ (r[^]n) = Qp_ev P x)" using B'' by blast then show ?thesis by blast qed then show "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n))" using B' by blast qed qed qed qed lemma basic_semialg_set_memI: assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "y \ carrier Q\<^sub>p" assumes "Qp_ev P q = (y[^]n)" shows "q \ basic_semialg_set m n P" using assms(1) assms(2) assms(3) basic_semialg_set_def by blast lemma basic_semialg_set_memE: assumes "q \ basic_semialg_set m n P" shows "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" "\y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)" using assms basic_semialg_set_def apply blast using assms basic_semialg_set_def by blast definition is_basic_semialg :: "nat \ ((nat \ int) \ (nat \ int)) set list set \ bool" where "is_basic_semialg m S \ (\ (n::nat) \ 0. (\ P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]). S = basic_semialg_set m n P))" abbreviation(input) basic_semialgs where "basic_semialgs m \ {S. (is_basic_semialg m S)}" definition semialg_sets where "semialg_sets n = gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" lemma carrier_is_semialg: "(carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ semialg_sets n " unfolding semialg_sets_def using gen_boolean_algebra.universe by blast lemma empty_set_is_semialg: " {} \ semialg_sets n" using carrier_is_semialg[of n] unfolding semialg_sets_def using gen_boolean_algebra.complement by (metis Diff_cancel) lemma semialg_intersect: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms(1) assms(2) gen_boolean_algebra_intersect semialg_sets_def by blast lemma semialg_union: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms gen_boolean_algebra.union semialg_sets_def by blast lemma semialg_complement: assumes "A \ semialg_sets n" shows "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) \ semialg_sets n " using assms gen_boolean_algebra.complement semialg_sets_def by blast lemma semialg_zero: assumes "A \ semialg_sets 0" shows "A = {[]} \ A = {}" using assms unfolding semialg_sets_def cartesian_power_def proof- assume A0: " A \ gen_boolean_algebra (carrier (RDirProd_list (R_list 0 Q\<^sub>p))) (basic_semialgs 0)" show " A = {[]} \ A = {}" proof- have "A \ {[]} \ A = {}" proof assume A1: "A \ {[]}" show "A = {}" proof- have "(R_list 0 Q\<^sub>p) = []" by simp then have "(carrier (RDirProd_list (R_list 0 Q\<^sub>p))) = {[]}" using RDirProd_list_nil by simp then show ?thesis using A0 A1 by (metis gen_boolean_algebra_subset subset_singletonD) qed qed then show ?thesis by linarith qed qed lemma basic_semialg_is_semialg: assumes "is_basic_semialg n A" shows "A \ semialg_sets n" by (metis (no_types, lifting) assms gen_boolean_algebra.simps inf_absorb1 is_basic_semialg_def mem_Collect_eq basic_semialg_set_def semialg_sets_def subsetI) lemma basic_semialg_is_semialg': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "A \ semialg_sets n" using assms basic_semialg_is_semialg is_basic_semialg_def by blast definition is_semialgebraic where "is_semialgebraic n S = (S \ semialg_sets n)" lemma is_semialgebraicE: assumes "is_semialgebraic n S" shows "S \ semialg_sets n" using assms is_semialgebraic_def by blast lemma is_semialgebraic_closed: assumes "is_semialgebraic n S" shows "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using is_semialgebraicE[of n S] unfolding semialg_sets_def using assms gen_boolean_algebra_subset is_semialgebraicE semialg_sets_def by blast lemma is_semialgebraicI: assumes "S \ semialg_sets n" shows "is_semialgebraic n S" by (simp add: assms is_semialgebraic_def) lemma basic_semialg_is_semialgebraic: assumes "is_basic_semialg n A" shows "is_semialgebraic n A" using assms basic_semialg_is_semialg is_semialgebraicI by blast lemma basic_semialg_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "is_semialgebraic n A" using assms(1) assms(2) assms(3) basic_semialg_is_semialg' is_semialgebraicI by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Algebraic Sets over $p$-adic Fields\ (********************************************************************************************) (********************************************************************************************) lemma p_times_square_not_square: assumes "a \ nonzero Q\<^sub>p" shows "\ \ (a [^] (2::nat)) \ P_set (2::nat)" proof assume A: "\ \ (a[^](2::nat)) \ P_set (2::nat)" then have "\ \ (a[^](2::nat)) \ nonzero Q\<^sub>p" unfolding P_set_def by blast then obtain b where b_def: "b \ carrier Q\<^sub>p \ b[^](2::nat) = \ \ (a[^](2::nat))" using A P_set_def by blast have "b \ nonzero Q\<^sub>p" apply(rule ccontr) using b_def assms by (metis A P_set_nonzero'(1) Qp.nat_pow_zero not_nonzero_Qp zero_neq_numeral) then have LHS: "ord (b[^](2::nat)) = 2* (ord b)" using nonzero_nat_pow_ord by presburger have "ord( \ \ (a[^](2::nat))) = 1 + 2* ord a" using assms nonzero_nat_pow_ord Qp_nat_pow_nonzero ord_mult ord_p p_nonzero by presburger then show False using b_def LHS by presburger qed lemma p_times_square_not_square': assumes "a \ carrier Q\<^sub>p" shows "\ \ (a [^] (2::nat)) = \ \ a = \" by (metis Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero assms p_nonzero) lemma zero_set_semialg_set: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev q a = \ \( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev q a) [^] (2::nat)) = y[^](2::nat)) " proof show "Qp_ev q a = \ \ \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" proof- assume "Qp_ev q a = \" then have "\ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" by (metis Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null zero_neq_numeral) then have "\ \ carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" using Qp.cring_axioms cring.cring_simprules(2) by blast then show "\y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" by blast qed show " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat)) \ Qp_ev q a = \" proof- assume A: " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat))" then obtain b where b_def: "b\carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (b[^](2::nat))" by blast show "Qp_ev q a = \" proof(rule ccontr) assume " Qp_ev q a \ \" then have " Qp_ev q a \ nonzero Q\<^sub>p" using assms eval_at_point_closed[of a n q] nonzero_def proof - have "Qp_ev q a \ carrier Q\<^sub>p" using \\a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>); q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ \ Qp_ev q a \ carrier Q\<^sub>p\ \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ by fastforce then have "Qp_ev q a \ {r \ carrier Q\<^sub>p. r \ \}" using \Qp_ev q a \ \\ by force then show ?thesis by (metis nonzero_def ) qed then have "\ \ (Qp_ev q a[^](2::nat)) \ nonzero Q\<^sub>p" by (metis Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero not_nonzero_Qp p_nonzero p_times_square_not_square') then have "\ \ (Qp_ev q a[^](2::nat)) \ P_set (2::nat)" using b_def unfolding P_set_def by blast then show False using \Qp_ev q a \ nonzero Q\<^sub>p\ p_times_square_not_square by blast qed qed qed lemma alg_as_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" shows "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" proof have 00: "\x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ Qp_ev q x = \ \ (Qp_ev P x) [^] (2::nat)" using assms eval_at_point_smult MP.nat_pow_closed Qp.int_inc_closed eval_at_point_nat_pow by presburger show "V\<^bsub>Q\<^sub>p\<^esub> n P \ basic_semialg_set n 2 q" proof fix x assume A: "x \ V\<^bsub>Q\<^sub>p\<^esub> n P " show "x \ basic_semialg_set n (2::nat) q " proof- have P: "Qp_ev P x = \" using A zero_setE(2) by blast have "Qp_ev q x = \" proof- have "Qp_ev q x = \ \ (Qp_ev (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) x)" using assms eval_at_point_smult[of x n "(P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" \] basic_semialg_set_def by (meson A MP.nat_pow_closed Qp.int_inc_closed zero_setE(1)) then show ?thesis by (metis A P Qp.int_inc_closed Qp.integral_iff Qp.nat_pow_zero Qp.zero_closed assms(1) eval_at_point_nat_pow neq0_conv zero_less_numeral zero_setE(1)) qed then have 0: "Qp_ev q x = \ \ Qp_ev q x \ P_set (2::nat)" by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A zero_setE(1) by blast then show ?thesis using 0 basic_semialg_set_def' by (metis (no_types, opaque_lifting) Qp.nat_pow_zero Qp.zero_closed \eval_at_point Q\<^sub>p x q = \\ basic_semialg_set_memI zero_neq_numeral) qed qed show "basic_semialg_set n 2 q \ V\<^bsub>Q\<^sub>p\<^esub> n P" proof fix x assume A: "x \ basic_semialg_set n 2 q" have 0: "\ Qp_ev q x \ P_set 2" proof assume "Qp_ev q x \ P_set 2" then have 0: "Qp_ev q x \ nonzero Q\<^sub>p \ (\y \ carrier Q\<^sub>p . (y[^] (2::nat)) = Qp_ev q x)" using P_set_def by blast have "( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev P x) [^] (2::nat)) = y[^](2::nat))" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ (y[^] (2::nat)) = Qp_ev q x" using 0 by blast have "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_memE(1) by blast then have "Qp_ev q x = \ \ ((Qp_ev P x) [^] (2::nat))" using assms eval_at_point_scalar_mult 00 by blast then have "(y[^] (2::nat)) = \ \ ((Qp_ev P x) [^] (2::nat))" using y_def by blast then show ?thesis using y_def by blast qed then have "Qp_ev P x = \" by (metis (no_types, lifting) A assms(1) basic_semialg_set_def mem_Collect_eq zero_set_semialg_set) then have "Qp_ev q x = \" using assms eval_at_point_smult by (metis "00" A Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null basic_semialg_set_memE(1) zero_neq_numeral) then show False using 0 Qp.not_nonzero_memI by blast qed show " x \ V\<^bsub>Q\<^sub>p\<^esub> n P" apply(rule zero_setI) using A basic_semialg_set_memE(1) apply blast using A 0 00[of x] by (metis assms(1) basic_semialg_set_memE(1) basic_semialg_set_memE(2) zero_set_semialg_set) qed qed lemma is_zero_set_imp_basic_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_basic_semialg n S" unfolding is_basic_semialg_def proof- obtain q where q_def: "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" by blast have 0: "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" using alg_as_semialg[of P n q] q_def assms(1) by linarith have "(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms(1) by blast then have "\ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms q_def Qp.int_inc_closed local.smult_closed by blast then have 1: "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" by (metis q_def ) then show "\m. m \ 0 \ (\P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). S = basic_semialg_set n m P)" using 0 assms by (metis zero_neq_numeral) qed lemma is_zero_set_imp_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_semialgebraic n S" using assms(1) assms(2) basic_semialg_is_semialg is_semialgebraicI is_zero_set_imp_basic_semialg by blast text\Algebraic sets are semialgebraic\ lemma is_algebraic_imp_is_semialg: assumes "is_algebraic Q\<^sub>p n S" shows "is_semialgebraic n S" proof(rule is_semialgebraicI) obtain ps where ps_def: "finite ps \ ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ S = affine_alg_set Q\<^sub>p n ps" using is_algebraicE by (metis assms) have "ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n ps \ semialg_sets n" apply(rule finite.induct[of ps]) apply (simp add: ps_def) using affine_alg_set_empty[of n] apply (simp add: carrier_is_semialg) proof fix A a assume IH: "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n A \ semialg_sets n" assume P: "insert a A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" have "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using P by blast then show "affine_alg_set Q\<^sub>p n (insert a A) \ semialg_sets n" using IH P semialg_intersect[of "affine_alg_set Q\<^sub>p n A" n "affine_alg_set Q\<^sub>p n {a}" ] is_zero_set_imp_semialg affine_alg_set_insert[of n a A] by (metis Int_commute affine_alg_set_singleton insert_subset is_semialgebraicE) qed then show "S \ semialg_sets n" using ps_def by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\Basic Lemmas about the Semialgebraic Predicate\ (********************************************************************************************) (********************************************************************************************) text\Finite and cofinite sets are semialgebraic\ lemma finite_is_semialg: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "finite F" shows "is_semialgebraic n F" using Qp.finite_sets_are_algebraic is_algebraic_imp_is_semialg[of n F] assms(1) assms(2) by blast definition is_cofinite where "is_cofinite n F = finite (ring_pow_comp Q\<^sub>p n F)" lemma is_cofiniteE: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "finite (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)" using assms(2) is_cofinite_def by (simp add: ring_pow_comp_def) lemma complement_is_semialg: assumes "is_semialgebraic n F" shows "is_semialgebraic n ((carrier (Q\<^sub>p\<^bsup>n\<^esup>)) - F)" using assms is_semialgebraic_def semialg_complement by blast lemma cofinite_is_semialgebraic: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "is_semialgebraic n F" using assms ring_pow_comp_inv[of F Q\<^sub>p n] complement_is_semialg[of n "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] finite_is_semialg[of "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] is_cofiniteE[of F] by (simp add: ring_pow_comp_def) lemma diff_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A - B)" apply(rule is_semialgebraicI) using assms unfolding semialg_sets_def using gen_boolean_algebra_diff is_semialgebraicE semialg_sets_def by blast lemma intersection_is_semialg: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_intersect by blast lemma union_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_union by blast lemma carrier_is_semialgebraic: "is_semialgebraic n (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using carrier_is_semialg by (simp add: carrier_is_semialg is_semialgebraic_def) lemma empty_is_semialgebraic: "is_semialgebraic n {}" by (simp add: empty_set_is_semialg is_semialgebraic_def) (********************************************************************************************) (********************************************************************************************) subsubsection\One-Dimensional Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition one_var_semialg where "one_var_semialg S = ((to_R1 ` S) \ (semialg_sets 1))" definition univ_basic_semialg_set where "univ_basic_semialg_set (m::nat) P = {a \ carrier Q\<^sub>p. (\y \ carrier Q\<^sub>p. (P \ a = (y[^]m)))}" text\Equivalence of univ\_basic\_semialg\_sets and semialgebraic subsets of $\mathbb{Q}^1$ \ lemma univ_basic_semialg_set_to_semialg_set: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "to_R1 ` (univ_basic_semialg_set m P) = basic_semialg_set 1 m (from_Qp_x P)" proof show "(\a. [a]) ` univ_basic_semialg_set m P \ basic_semialg_set 1 m (from_Qp_x P)" proof fix x assume A: "x \ (\a. [a]) ` univ_basic_semialg_set m P" then obtain b y where by_def:"b \ carrier Q\<^sub>p \ y \ carrier Q\<^sub>p \ (P \ b) = (y[^]m) \ x = [b]" unfolding univ_basic_semialg_set_def by blast then have "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A Qp.to_R1_closed[of b] unfolding univ_basic_semialg_set_def by blast then show "x \ basic_semialg_set 1 m (from_Qp_x P)" using by_def Qp_x_Qp_poly_eval assms unfolding basic_semialg_set_def by blast qed show "basic_semialg_set 1 m (from_Qp_x P) \ (\a. [a]) ` univ_basic_semialg_set m P" proof fix x assume A: "x \ basic_semialg_set 1 m (from_Qp_x P)" then obtain b where b_def: "b \ carrier Q\<^sub>p \ x = [b]" unfolding basic_semialg_set_def by (metis (mono_tags, lifting) mem_Collect_eq Qp.to_R1_to_R Qp.to_R_pow_closed) obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev (from_Qp_x P) [b] = (y[^]m))" using A b_def unfolding basic_semialg_set_def by blast have " P \ b = (y[^]m)" using assms y_def b_def Qp_x_Qp_poly_eval by blast then show " x \ (\a. [a]) ` univ_basic_semialg_set m P" using y_def b_def unfolding basic_semialg_set_def univ_basic_semialg_set_def by blast qed qed definition is_univ_semialgebraic where "is_univ_semialgebraic S = (S \ carrier Q\<^sub>p \ is_semialgebraic 1 (to_R1 ` S))" lemma is_univ_semialgebraicE: assumes "is_univ_semialgebraic S" shows "is_semialgebraic 1 (to_R1 ` S)" using assms is_univ_semialgebraic_def by blast lemma is_univ_semialgebraicI: assumes "is_semialgebraic 1 (to_R1 ` S)" shows "is_univ_semialgebraic S" proof- have "S \ carrier Q\<^sub>p" proof fix x assume "x \ S" then have "(to_R1 x) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms by (smt Collect_mono_iff gen_boolean_algebra_subset image_def is_semialgebraicE mem_Collect_eq semialg_sets_def Qp.to_R1_carrier) then show "x \ carrier Q\<^sub>p" using assms by (metis nth_Cons_0 Qp.to_R_pow_closed) qed then show ?thesis using assms unfolding is_univ_semialgebraic_def by blast qed lemma univ_basic_semialg_set_is_univ_semialgebraic: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "is_univ_semialgebraic (univ_basic_semialg_set m P)" using assms by (metis (mono_tags, lifting) basic_semialg_is_semialgebraic' from_Qp_x_closed is_univ_semialgebraic_def mem_Collect_eq subsetI univ_basic_semialg_set_def univ_basic_semialg_set_to_semialg_set) lemma intersection_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms intersection_is_semialg[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis le_infI1 Qp.to_R1_intersection) lemma union_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms union_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis Un_subset_iff image_Un) lemma diff_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A - B)" using assms diff_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (smt Diff_subset subset_trans Qp.to_R1_diff) lemma finite_is_univ_semialgebraic: assumes "A \ carrier Q\<^sub>p" assumes "finite A" shows "is_univ_semialgebraic A" using assms finite_is_semialg[of "((\a. [a]) ` A)" ] to_R1_finite[of A] unfolding is_univ_semialgebraic_def by (metis Qp.to_R1_carrier Qp.to_R1_subset) (********************************************************************************************) (********************************************************************************************) subsubsection\Defining the $p$-adic Valuation Semialgebraically\ (********************************************************************************************) (********************************************************************************************) lemma Qp_square_root_criterion0: assumes "p \ 2" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "val a \ val b" assumes "a \ \" assumes "b \ \" assumes "val a \ 0" shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" proof- have 0: "(to_Zp a) \ carrier Z\<^sub>p" using assms(2) to_Zp_closed by blast have 1: "(to_Zp b) \ carrier Z\<^sub>p" using assms(3) to_Zp_closed by blast have 2: "a \ \\<^sub>p" using val_ring_val_criterion assms(2) assms(5) assms(7) by blast have 3: "b \ \\<^sub>p" using assms val_ring_val_criterion[of b] dual_order.trans by blast have 4: "val_Zp (to_Zp b) = val b" using 3 Zp_def \_def padic_fields.to_Zp_val padic_fields_axioms by blast have 5: "val_Zp (to_Zp a) = val a" using Q\<^sub>p_def Zp_def assms(2) assms(7) padic_fields.Qp_val_ringI padic_fields.to_Zp_val padic_fields_axioms by blast have "\y \ carrier Z\<^sub>p. (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" using 0 1 2 4 5 assms Zp_square_root_criterion[of "(to_Zp a)" "(to_Zp b)"] by (metis "3" to_Zp_inc to_Zp_zero zero_in_val_ring) then obtain y where y_def: "y \ carrier Z\<^sub>p \ (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" by blast have 6: "a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \ \b[^](2::nat) = ((\ y) [^] (2::nat))" proof- have 0: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = ((\ y) [^] (2::nat))" using Qp_nonzero_nat_pow nat_pow_closed inc_pow nat_inc_zero inc_is_hom \_def y_def ring_hom_nat_pow[of Z\<^sub>p Q\<^sub>p \ y 2] Q\<^sub>p_def Qp.ring_axioms Zp.ring_axioms by blast have 1: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = \ ((to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" using y_def by presburger have 2: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = \ ((to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ ( \

\\<^bsub>Z\<^sub>p\<^esub>(to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" using "1" Zp.m_closed Zp_int_inc_closed assms(2) assms(3) inc_of_sum pow_closed to_Zp_closed by presburger hence 3: "\ (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat)) = (\ (to_Zp a))[^](2::nat) \ (\ \

) \ \ ((to_Zp b)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat))" using Qp_nonzero_nat_pow nat_pow_closed inc_pow nat_inc_zero inc_is_hom \_def y_def ring_hom_nat_pow[of Z\<^sub>p Q\<^sub>p \ _ 2] Q\<^sub>p_def Qp.ring_axioms Zp.ring_axioms Zp_int_inc_closed assms(2) assms(3) inc_of_prod pow_closed to_Zp_closed by metis then show ?thesis using "0" "4" val_ring_ord_criterion assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) inc_pow not_nonzero_Zp ord_of_nonzero(1) p_inc to_Zp_closed to_Zp_inc by (metis to_Zp_zero val_pos val_ringI zero_in_val_ring) qed have "(\ y) \ carrier Q\<^sub>p" using frac_closed local.inc_def y_def inc_closed by blast then show ?thesis using 6 by blast qed lemma eint_minus_ineq': assumes "(a::eint) \ b" shows "a -b \ 0" by (metis assms eint_minus_ineq eint_ord_simps(3) idiff_infinity idiff_self order_trans top.extremum_unique top_eint_def) lemma Qp_square_root_criterion: assumes "p \ 2" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "ord b \ ord a" assumes "a \ \" assumes "b \ \" shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" proof- have "\k::int. k \ min (ord a) (ord b) \ k mod 2 = 0" proof- let ?k = "if (min (ord a) (ord b)) mod 2 = 0 then min (ord a) (ord b) else (min (ord a) (ord b)) - 1" have "?k \ min (ord a) (ord b) \ ?k mod 2 = 0" apply(cases "(min (ord a) (ord b)) mod 2 = 0 ") apply presburger by presburger then show ?thesis by meson qed then obtain k where k_def: "k \ min (ord a) (ord b) \ k mod 2 = 0" by meson obtain a0 where a0_def: "a0 = (\[^](-k)) \ a" by blast obtain b0 where b0_def: "b0 = (\[^](-k)) \ b" by blast have 0: "a0 \ nonzero Q\<^sub>p" using Qp.cring_axioms Qp.field_axioms Ring.integral a0_def assms(2) assms(5) cring_simprules(5) not_nonzero_Qp p_intpow_closed(1) p_nonzero by (metis Qp_int_pow_nonzero cring.cring_simprules(5)) have 1: "val a0 = val a - k" using a0_def assms(2) assms(5) val_mult p_nonzero p_intpow_closed(1) by (metis Qp.m_comm Qp_int_pow_nonzero p_intpow_inv'' val_fract val_p_int_pow) have 11: "val b0 = val b - k" using assms(3) assms(6) b0_def val_mult p_nonzero p_intpow_closed(1) by (metis Qp.m_lcomm Qp.one_closed Qp.r_one Qp_int_pow_nonzero p_intpow_inv'' val_fract val_p_int_pow) have A: "val a \ k" using k_def val_ord assms by (smt eint_ord_simps(1) not_nonzero_Qp) have B: "val b \ k" using k_def val_ord assms by (smt eint_ord_simps(1) not_nonzero_Qp) then have 2: "val a0 \ 0" using A 1 assms k_def eint_minus_ineq eint_ord_code(5) local.eint_minus_ineq' by presburger have 3: "val a0 \ val b0" using 1 11 assms by (metis eint.distinct(2) eint_minus_ineq eint_ord_simps(1) val_def) have 4: "a0 \ \" using a0_def "0" Qp.nonzero_memE(2) by blast have 5: "b0 \ \" using b0_def by (metis "4" Qp.integral_iff a0_def assms(2) assms(3) assms(6) p_intpow_closed(1)) have "\y \ carrier Q\<^sub>p. a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat) = (y [^] (2::nat))" using Qp_square_root_criterion0[of a0 b0] assms 2 3 4 5 b0_def a0_def Qp.m_closed p_intpow_closed(1) by metis then obtain y where y_def: " y \ carrier Q\<^sub>p \ a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat) = (y [^] (2::nat))" by blast then have 6: " (\[^] (2 * k)) \ (a0[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b0[^](2::nat)) = (\[^] (2 * k)) \ (y [^] (2::nat))" by presburger then have 8: "((\[^] (2 * k)) \ (a0[^](2::nat))) \\<^bsub>Q\<^sub>p\<^esub>((\[^] (2 * k)) \ (\\b0[^](2::nat))) = (\[^] (2 * k)) \ (y [^] (2::nat))" using 6 Qp.r_distr[of "(a0[^](2::nat))" " (\\b0[^](2::nat))" "(\[^] (2 * k))"] by (metis Qp.add.int_pow_closed Qp.m_closed Qp.nat_pow_closed Qp.one_closed a0_def assms(2) assms(3) b0_def p_inc p_intpow_closed(1) y_def) have 9: "(\[^](int 2*k)) = (\[^]k)[^](2::nat)" using Qp_int_nat_pow_pow[of \ k 2] by (metis mult_of_nat_commute p_nonzero) then have "((\[^]k)[^](2::nat) \ (a0[^](2::nat))) \\<^bsub>Q\<^sub>p\<^esub> (\[^]k)[^](2::nat) \ (\\b0[^](2::nat)) = (\[^]k)[^](2::nat) \ (y [^] (2::nat))" by (metis "8" int_eq_iff_numeral) then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\[^]k)[^](2::nat)) \ (\\b0[^](2::nat)) = ((\[^]k)[^](2::nat)) \ (y [^] (2::nat))" by (metis Qp.cring_axioms a0_def assms(2) comm_monoid.nat_pow_distrib cring.cring_simprules(5) cring_def p_intpow_closed(1)) then have 10: "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\[^]k)[^](2::nat)) \ (\\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" using comm_monoid.nat_pow_distrib y_def by (metis Qp.comm_monoid_axioms p_intpow_closed(1)) then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((((\[^]k)[^](2::nat)) \ \)\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" using 10 monoid.m_assoc[of Q\<^sub>p "((\[^]k)[^](2::nat))" \ " b0[^](2::nat)"] by (metis Qp.int_inc_closed Qp.m_assoc Qp.m_closed Qp.nat_pow_closed assms(3) b0_def p_intpow_closed(1)) then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>((\ \ ((\[^]k)[^](2::nat)) )\b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" by (metis Qp.group_commutes_pow Qp.int_inc_closed Qp.m_comm p_intpow_closed(1)) then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ (((\[^]k)[^](2::nat)) \b0[^](2::nat)) = ((\[^]k) \ y) [^] (2::nat)" by (metis "10" Qp.int_inc_closed Qp.m_closed Qp.m_lcomm Qp.nat_pow_closed assms(3) b0_def p_intpow_closed(1)) then have "((\[^]k) \ a0)[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ ((\[^]k) \b0)[^](2::nat) = ((\[^]k) \ y) [^] (2::nat)" by (metis Qp.m_closed Qp.nat_pow_distrib assms(3) b0_def p_intpow_closed(1)) then have "a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub>\ \ b[^](2::nat) = ((\[^]k) \ y) [^] (2::nat)" by (metis Qp.l_one Qp.m_assoc a0_def assms(2) assms(3) b0_def p_intpow_closed(1) p_intpow_inv) then show ?thesis by (meson Qp.cring_axioms cring.cring_simprules(5) p_intpow_closed(1) y_def) qed lemma Qp_val_ring_alt_def0: assumes "a \ nonzero Q\<^sub>p" assumes "ord a \ 0" shows "\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" proof- have "\y \ carrier Z\<^sub>p. \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub> (3::nat))\\<^bsub>Z\<^sub>p\<^esub> ((to_Zp a) [^]\<^bsub>Z\<^sub>p\<^esub> (4::nat)) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" using padic_integers.Zp_semialg_eq[of p "to_Zp a"] prime assms to_Zp_def by (metis (no_types, lifting) Qp.nonzero_closed Qp.not_nonzero_memI Zp_def val_ring_ord_criterion not_nonzero_Zp padic_integers_axioms to_Zp_closed to_Zp_inc to_Zp_zero zero_in_val_ring) then obtain y where y_def: "y \ carrier Z\<^sub>p \ \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> (\

[^]\<^bsub>Z\<^sub>p\<^esub> (3::nat))\\<^bsub>Z\<^sub>p\<^esub> ((to_Zp a) [^]\<^bsub>Z\<^sub>p\<^esub> (4::nat)) = (y [^]\<^bsub>Z\<^sub>p\<^esub> (2::nat))" by blast then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((\ y)[^](2::nat))" using Group.nat_pow_0 Group.nat_pow_Suc nonzero_def val_ring_ord_criterion assms inc_of_nonzero inc_of_prod inc_of_sum inc_pow m_closed nat_inc_closed nat_pow_closed not_nonzero_Zp numeral_2_eq_2 p_natpow_inc to_Zp_closed to_Zp_inc by (smt Qp.nonzero_closed Qp.nonzero_memE(2) Zp.monom_term_car p_pow_nonzero(1) pow_closed to_Zp_zero zero_in_val_ring) then have "(\ y) \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((\ y)[^](2::nat))" using y_def inc_closed by blast then show ?thesis by blast qed text\Defining the valuation semialgebraically for odd primes\ lemma P_set_ord_semialg_odd_p: assumes "p \ 2" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" proof(cases "a = \") case True show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" proof show "val b \ val a \ \y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" proof- assume A: "val b \ val a" then have "val b \ \" by (metis True local.val_zero) then have "b = \" using assms(3) local.val_zero val_ineq by presburger then have "(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (\[^](2::nat))" using True by (metis Qp.int_inc_zero Qp.int_nat_pow_rep Qp.nonzero_closed Qp.r_null Qp.r_zero assms(3) p_nonzero zero_power2) then show ?thesis using \b = \\ assms(3) by blast qed show "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat)) \ val b \ val a" proof- assume "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" then obtain y where y_def: "y \ carrier Q\<^sub>p \(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" by blast then have 0: "\ \ (b[^](2::nat)) = (y[^](2::nat))" by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.int_inc_closed Qp.nat_pow_closed Qp.not_nonzero_memI Qp_nonzero_nat_pow True assms(2) assms(3) local.monom_term_car not_nonzero_Qp zero_less_numeral) have "b = \" apply(rule ccontr) using 0 assms y_def p_times_square_not_square[of b] unfolding P_set_def by (metis (no_types, opaque_lifting) P_set_memI Qp.nat_pow_closed True \b \ nonzero Q\<^sub>p \ \ \ b [^] 2 \ P_set 2\ not_nonzero_Qp p_times_square_not_square') then show ?thesis using eint_ord_code(3) local.val_zero by presburger qed qed next case False then show ?thesis proof(cases "b = \") case True then have "(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (a[^](2::nat))" by (metis Qp.add.l_cancel_one' Qp.int_inc_zero Qp.int_nat_pow_rep Qp.nat_pow_closed Qp.nonzero_closed Qp.r_null assms(2) assms(3) p_nonzero zero_power2) then have 0: "(\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" using assms(2) by blast have 1: "val a \ val b" using True assms local.val_zero eint_ord_code(3) by presburger show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" using 0 1 by blast next case F: False show "val a \ val b \ (\y \ carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\ \ (b[^](2::nat))) = (y[^](2::nat)))" proof show "val b \ val a \ \y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" proof- assume "val b \ val a " then have "ord b \ ord a" using F False by (metis eint_ord_simps(1) val_def) then show "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" using assms Qp_square_root_criterion[of a b] False F by blast qed show "\y\carrier Q\<^sub>p.(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat)) \ val b \ val a" proof- assume "\y\carrier Q\<^sub>p. (a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" then obtain y where y_def: "y \ carrier Q\<^sub>p \(a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat)) = (y[^](2::nat))" by blast have 0: "ord (a[^](2::nat)) = 2* ord a" by (metis (mono_tags, opaque_lifting) False Suc_1 assms(2) int_eq_iff_numeral nat_numeral nonzero_nat_pow_ord not_nonzero_Qp) have 1: "ord (\ \ (b[^](2::nat))) = 1 + 2* ord b" proof- have 0: "ord (\ \ (b[^](2::nat))) = ord \ + ord (b[^](2::nat))" using F Qp_nat_pow_nonzero assms(3) not_nonzero_Qp ord_mult p_nonzero by metis have 1: "ord (b[^](2::nat)) = 2* ord b" using F assms by (metis (mono_tags, opaque_lifting) Suc_1 int_eq_iff_numeral nat_numeral nonzero_nat_pow_ord not_nonzero_Qp) then show ?thesis using "0" ord_p by linarith qed show "val b \ val a" proof(rule ccontr) assume "\ val b \ val a" then have "val b \ val a \ val a \ val b" by (metis linear) then have "ord a > ord b" using F False assms by (metis \\ val a \ val b\ eint_ord_simps(1) le_less not_less_iff_gr_or_eq val_def) then have "ord (a[^](2::nat)) > ord (\ \ (b[^](2::nat)))" using 0 1 by linarith then have "ord ((a[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^](2::nat))) = ord (\ \ (b[^](2::nat)))" by (meson F False Qp.int_inc_closed Qp_nat_pow_nonzero assms(2) assms(3) local.monom_term_car not_nonzero_Qp ord_ultrametric_noteq p_times_square_not_square') then have A0: "ord (y[^](2::nat)) = 1 + 2* ord b" by (metis "1" \y \ carrier Q\<^sub>p \ (a[^]2) \\<^bsub>Q\<^sub>p\<^esub> \ \ (b[^]2) = (y[^]2)\) have A1: "(y[^](2::nat)) \ nonzero Q\<^sub>p" using y_def 0 1 by (smt F False Qp.nonzero_closed Qp_nat_pow_nonzero assms(2) assms(3) diff_ord_nonzero local.monom_term_car not_nonzero_Qp p_nonzero p_times_square_not_square') have A2: "y \ nonzero Q\<^sub>p" using A1 Qp_nonzero_nat_pow pos2 y_def by blast have A3: "ord (y[^](2::nat)) = 2* ord y" using A2 nonzero_nat_pow_ord by presburger then show False using A0 by presburger qed qed qed qed qed text\Defining the valuation ring semialgebraically for all primes\ lemma Qp_val_ring_alt_def: assumes "a \ carrier Q\<^sub>p" shows "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" proof(cases "a = \") case True then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = \" by (metis Qp.add.l_cancel_one' Qp.integral_iff Qp.nat_pow_closed Qp.not_nonzero_memI Qp.one_closed Qp_nonzero_nat_pow assms not_nonzero_Qp p_natpow_closed(1) zero_less_numeral) then have "\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (\[^](2::nat))" using Qp.nat_pow_one by blast then show ?thesis using True zero_in_val_ring by blast next case False show "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" proof show "a \ \\<^sub>p \ (\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" using assms Qp_val_ring_alt_def0[of a] False by (meson not_nonzero_Qp ord_nonneg) show "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))) \ a \ \\<^sub>p" proof- assume "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" then obtain y where y_def: "y \ carrier Q\<^sub>p \\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" by blast then have "(\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using Qp.ring_simprules by (smt Qp.nat_pow_closed assms p_natpow_closed(1)) then have "ord ((\[^](3::nat))\ (a[^](4::nat))) = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" by presburger then have "3 + ord (a[^](4::nat)) = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" by (metis False Qp_nat_pow_nonzero assms not_nonzero_Qp of_nat_numeral ord_mult ord_p_pow_nat p_nonzero) then have 0: "3 + 4* ord a = ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \)" using assms False nonzero_nat_pow_ord[of a "(4::nat)"] by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) have "ord a \ 0" proof(rule ccontr) assume "\ 0 \ ord a" then have 00: "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) < 0" using 0 by linarith have yn: "y \ nonzero Q\<^sub>p" apply(rule ccontr) using y_def 0 by (metis "00" Qp.not_eq_diff_nonzero Qp.one_closed Qp.one_nonzero Qp.pow_zero \\ [^] 3 \ a [^] 4 = y [^] 2 \ \\ diff_ord_nonzero less_numeral_extra(3) local.one_neq_zero not_nonzero_Qp ord_one zero_less_numeral) then have "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) = ord (y[^](2::nat))" using y_def ord_ultrametric_noteq''[of "(y[^](2::nat))" "\" ] by (metis "00" False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_pow_nonzero Qp.not_eq_diff_nonzero Qp.one_nonzero Qp.r_right_minus_eq \\ [^] 3 \ a [^] 4 = y [^] 2 \ \\ assms ord_one ord_ultrametric_noteq p_nonzero) then have "ord ((y[^](2::nat)) \\<^bsub>Q\<^sub>p\<^esub> \) = 2* ord y" using y_def Qp_nat_pow_nonzero Qp_nonzero_nat_pow nonzero_nat_pow_ord[of y "(2::nat)"] yn by linarith then have "3 + (4* ord a) = 2* ord y" using "00" "0" by linarith then show False by presburger qed then show "a \ \\<^sub>p" using False val_ring_ord_criterion assms by blast qed qed qed lemma Qp_val_alt_def: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "val b \ val a \ (\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)))" proof show "val a \ val b \ \y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" proof- assume A: "val a \ val b" show "\y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" proof(cases "b = \") case True then have "a = \" using A assms(1) val_ineq by blast then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (\[^](2::nat))" by (metis Qp.nat_pow_zero Qp.r_null Qp.r_zero True assms(2) p_natpow_closed(1) zero_neq_numeral) then show ?thesis using True A assms(2) by blast next case False assume B: "b \ \" show "\y\carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat)) \ (a[^](4::nat)) = (y[^](2::nat))" proof(cases "a = \") case True then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (b[^](4::nat))" using Qp.cring_axioms Qp.nat_pow_closed assms(2) cring_def p_natpow_closed(1) ring.pow_zero zero_less_numeral by (metis Qp.add.l_cancel_one' Qp.integral_iff assms(1)) then have "(b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((b[^](2::nat))[^] (2::nat))" by (metis Qp_nat_pow_pow assms(2) mult_2_right numeral_Bit0) then have "(b[^](2::nat)) \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = ((b[^](2::nat))[^] (2::nat))" using Qp.nat_pow_closed assms(2) by blast then show ?thesis by blast next case False have F0: "b \ nonzero Q\<^sub>p" using B assms(2) not_nonzero_Qp by metis have F1: "a \ nonzero Q\<^sub>p" using False assms(1) not_nonzero_Qp by metis then have "(a \

b) \ nonzero Q\<^sub>p" using B by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) inv_in_frac(3)) then have "val a \ val b" using F0 F1 A by blast then have "val (a \
b) \ 0" using F0 F1 val_fract assms(1) local.eint_minus_ineq' by presburger obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat)) = (y[^](2::nat))" using Qp_val_ring_alt_def0 by (meson B False Qp.integral Qp.nonzero_closed \(a \
b) \ nonzero Q\<^sub>p\ \0 \ val (a \
b)\ assms(1) assms(2) inv_in_frac(1) inv_in_frac(2) ord_nonneg val_ringI) then have "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = (b[^](4::nat)) \ (y[^](2::nat))" by presburger then have F2: "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis Qp.nat_pow_pow assms(2) mult_2_right numeral_Bit0) have F3: "((b[^](4::nat)) \ \) \\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat))[^] (2::nat)) \ (y[^](2::nat))" proof- have 0: "(\[^](3::nat)) \ (a \
b[^](4::nat)) \ carrier Q\<^sub>p " proof- have "(a \
b[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed Qp_nat_pow_nonzero by presburger then show ?thesis by (meson Qp.cring_axioms cring.cring_simprules(5) p_natpow_closed(1)) qed have 1: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then show ?thesis using 0 F2 ring.ring_simprules(23)[of Q\<^sub>p "\" "(\[^](3::nat))\ ((a \
b)[^](4::nat))" "(b[^](4::nat))"] Qp.cring_axioms Qp.nonzero_mult_closed Qp.ring_axioms Qp_nat_pow_nonzero \(a \
b) \ nonzero Q\<^sub>p\ p_nonzero by blast qed have F4: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then have "((b[^](4::nat)) \ \) = (b[^](4::nat))" using Qp.r_one by blast then have F5: "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F3 by presburger have "((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat)))" proof- have 0: "(b[^](4::nat)) \ carrier Q\<^sub>p" using F4 by blast have 1: "(\[^](3::nat)) \ carrier Q\<^sub>p" by blast have 2: "((a \
b)[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed by blast show ?thesis using 0 1 2 monoid.m_assoc[of Q\<^sub>p] comm_monoid.m_comm[of Q\<^sub>p] using Qp.m_lcomm by presburger qed then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F5 by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b \(a \
b))[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F0 Qp.nat_pow_distrib assms(1) assms(2) fract_closed by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis F0 assms(1) local.fract_cancel_right) then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" using Qp.nat_pow_closed Qp.nat_pow_distrib assms(2) y_def by blast then have "((b[^](2::nat))\ y) \ carrier Q\<^sub>p \ (b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" by (meson Qp.cring_axioms Qp.nat_pow_closed assms(2) cring.cring_simprules(5) y_def) then show ?thesis by blast qed qed qed show "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \ val a \ val b" proof- assume A: "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" show "val a \ val b" proof(cases "a = \") case True then show ?thesis using eint_ord_code(3) local.val_zero by presburger next case False have "b \ \" proof(rule ccontr) assume "\ b \ \" then have "\y \ carrier Q\<^sub>p. (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.nat_pow_closed Qp.nonzero_memE(2) Qp_nonzero_nat_pow assms(1) assms(2) local.monom_term_car not_nonzero_Qp p_natpow_closed(1) zero_less_numeral) then obtain y where y_def: "y \ carrier Q\<^sub>p \ (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" by blast have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" proof- have 00: "(\[^](3::nat)) \ nonzero Q\<^sub>p" using Qp_nat_pow_nonzero p_nonzero by blast have 01: "(a[^](4::nat)) \ nonzero Q\<^sub>p" using False Qp_nat_pow_nonzero assms(1) not_nonzero_Qp Qp.nonzero_memI by presburger then show ?thesis using ord_mult[of "\[^](3::nat)" "a[^](4::nat)"] by (metis (no_types, lifting) "00" False assms(1) nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral ord_p_pow_nat) qed have 1: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 2* (ord y)" proof- have "y \ \" proof(rule ccontr) assume " \ y \ \" then have "(\[^](3::nat))\ (a[^](4::nat)) = \" using y_def Qp.cring_axioms cring_def pos2 ring.pow_zero by blast then show False by (metis False Qp.integral Qp.nat_pow_closed Qp.nonzero_pow_nonzero Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) p_natpow_closed(1) p_nonzero) qed then show ?thesis using y_def by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) qed then show False using 0 by presburger qed then have F0: "b \ nonzero Q\<^sub>p" using assms(2) not_nonzero_Qp by metis have F1: "a \ nonzero Q\<^sub>p" using False assms(1) not_nonzero_Qp by metis obtain y where y_def: "y \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by blast show ?thesis proof(rule ccontr) assume " \ val a \ val b " then have F2: "ord a < ord b" using F0 F1 assms by (metis False \b \ \\ eint_ord_simps(1) leI val_def) have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using F0 ord_mult F1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_p_pow_nat p_natpow_closed(2) by presburger have 1: " ord (b[^](4::nat)) = 4* ord b" using F0 nonzero_nat_pow_ord by presburger have 2: "(4 * (ord b)) > 4 * (ord a)" using F2 by linarith have 3: "(4 * (ord b)) \ 3 + 4* ord a" proof(rule ccontr) assume "\ (4 * (ord b)) \ 3 + 4* ord a" then have "(4 * (ord b)) > 3 + 4* ord a" by linarith then have 30: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using "0" "1" F0 F1 Qp_nat_pow_nonzero Qp.nat_pow_closed assms(1) monom_term_car not_nonzero_Qp ord_ultrametric_noteq p_natpow_closed(1) p_nonzero by (metis Qp.integral) have "y \ nonzero Q\<^sub>p" proof(rule ccontr) assume A: "y \ nonzero Q\<^sub>p" then have "y = \" using y_def Qp.nonzero_memI by blast then have "b [^] 4 \ \ [^] 3 \ a [^] 4 = \" by (smt "0" "1" A F0 False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_mult_closed Qp.nonzero_pow_nonzero Qp.pow_zero assms(1) diff_ord_nonzero not_nonzero_Qp p_nonzero pos2 y_def) then show False by (smt "0" "1" A F0 F1 Qp.integral Qp.nat_pow_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero assms(1) diff_ord_nonzero not_nonzero_Qp p_natpow_closed(1) p_nonzero y_def) qed then have 31: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 2* ord y" using nonzero_nat_pow_ord y_def by presburger then show False using 30 by presburger qed show False using 2 3 by presburger qed qed qed qed text\The polynomial in two variables which semialgebraically defines the valuation relation\ definition Qp_val_poly where "Qp_val_poly = (pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" lemma Qp_val_poly_closed: "Qp_val_poly \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" proof- have "(pvar Q\<^sub>p 1) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed one_less_numeral_iff semiring_norm(76) by blast then have 0: "(pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using ring.Pring_is_ring[of Q\<^sub>p "{0::nat..2-1}"] monoid.nat_pow_closed[of "coord_ring Q\<^sub>p 2"] Qp.cring_axioms cring.axioms(1) ring.Pring_is_monoid by blast have 1: "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed pos2 by blast have 2: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using 1 local.smult_closed p_natpow_closed(1) by blast then show ?thesis unfolding Qp_val_poly_def using 0 by blast qed lemma Qp_val_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "Qp_ev Qp_val_poly [a, b] = (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))" proof- have 0: "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" proof(rule cartesian_power_car_memI) show "length [a, b] = 2" by simp have "set [a,b] = {a,b}" by auto then show "set [a, b] \ carrier Q\<^sub>p" using assms by (simp add: \a \ carrier Q\<^sub>p\ \b \ carrier Q\<^sub>p\) qed obtain f where f_def: "f = ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))" by blast obtain g where g_def: "g = (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" by blast have 1: "Qp_val_poly = f \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> g" unfolding Qp_val_poly_def using f_def g_def by blast have 1: "Qp_ev (pvar Q\<^sub>p (0::nat)) [a,b] = a" using eval_pvar by (metis \[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)\ nth_Cons_0 pos2) have 2: "Qp_ev (pvar Q\<^sub>p (1::nat)) [a,b] = b" using eval_pvar by (metis (no_types, lifting) "0" One_nat_def add_diff_cancel_right' assms(2) cartesian_power_car_memE gr_zeroI less_numeral_extra(1) less_numeral_extra(4) list.size(4) nth_Cons_pos Qp.to_R1_closed Qp.to_R_to_R1 zero_less_diff) have 3: "Qp_ev ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (b[^](4::nat))" by (metis "0" "2" eval_at_point_nat_pow local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 4: "Qp_ev ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (a[^](4::nat))" using "0" "1" eval_at_point_nat_pow local.pvar_closed pos2 by presburger then have 5: "Qp_ev (poly_scalar_mult Q\<^sub>p (\[^](3::nat)) ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))) [a,b] = (\[^](3::nat))\ (a[^](4::nat))" using eval_at_point_smult[of "[a,b]" 2 "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)" "\[^](3::nat)" ] 2 by (metis "0" MP.nat_pow_closed eval_at_point_scalar_mult local.pvar_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis proof- have 00: "[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" by (simp add: "0") have 01: " pvar Q\<^sub>p 1 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 02: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (pvar Q\<^sub>p 0 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat)) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed local.smult_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis unfolding Qp_val_poly_def using 00 01 02 by (metis (no_types, lifting) "3" "4" MP.nat_pow_closed eval_at_point_add eval_at_point_smult local.pvar_closed p_natpow_closed(1) zero_less_numeral) qed qed lemma Qp_2I: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule cartesian_power_car_memI) using assms apply (simp add: assms(1) assms(2)) using assms by (simp add: assms(1) assms(2)) lemma pair_id: assumes "length as = 2" shows "as = [as!0, as!1]" using assms by (smt One_nat_def diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons' nth_equalityI numeral_2_eq_2) lemma Qp_val_semialg: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "val b \ val a \ [a,b] \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val a \ val b \ [a, b] \ basic_semialg_set 2 2 Qp_val_poly" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def by (metis (mono_tags, lifting) assms(1) assms(2) mem_Collect_eq) show "[a, b] \ basic_semialg_set 2 2 Qp_val_poly \ val a \ val b" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def using assms(1) assms(2) by blast qed definition val_relation_set where "val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as!1) \ val (as!0)}" lemma val_relation_setE: assumes "as \ val_relation_set" shows "as!0 \ carrier Q\<^sub>p \ as!1 \ carrier Q\<^sub>p \ as = [as!0,as!1] \ val (as!1) \ val (as!0)" using assms unfolding val_relation_set_def by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff pair_id pos2 semiring_norm(76)) lemma val_relation_setI: assumes "as!0 \ carrier Q\<^sub>p" assumes "as!1 \ carrier Q\<^sub>p" assumes "length as = 2" assumes "val (as!1) \ val(as!0)" shows "as \ val_relation_set" unfolding val_relation_set_def using assms Qp_2I[of "as!0" "as!1"] by (metis (no_types, lifting) mem_Collect_eq pair_id) lemma val_relation_semialg: "val_relation_set = basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val_relation_set \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof fix as assume A: "as \ val_relation_set" have 0: "length as = 2" unfolding val_relation_set_def by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq val_relation_set_def) have 1: "as = [as ! 0, as ! 1]" by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq pair_id val_relation_set_def) show "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" using A 1 val_relation_setE[of as] Qp_val_semialg[of "as!0" "as!1"] by presburger qed show "basic_semialg_set 2 (2::nat) Qp_val_poly \ val_relation_set" proof fix as assume "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" then show "as \ val_relation_set" using val_relation_setI[of as] by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff Qp_val_semialg basic_semialg_set_def val_relation_set_def padic_fields_axioms pair_id pos2 semiring_norm(76)) qed qed lemma val_relation_is_semialgebraic: "is_semialgebraic 2 val_relation_set" proof - have "{rs \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (rs ! 0) \ val (rs ! 1)} = basic_semialg_set (Suc 1) (Suc 1) Qp_val_poly" using Suc_1 val_relation_semialg val_relation_set_def by presburger then show ?thesis by (metis (no_types) Qp_val_poly_closed Suc_1 basic_semialg_is_semialgebraic' val_relation_set_def zero_neq_numeral) qed lemma Qp_val_ring_is_semialg: obtains P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" proof- obtain P where P_def: "P = (\[^](3::nat)) \\<^bsub>Q\<^sub>p_x \<^esub>(X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat) \\<^bsub>Q\<^sub>p_x\<^esub> \\<^bsub>Q\<^sub>p_x\<^esub>" by blast have 0: "P \ carrier Q\<^sub>p_x" proof- have 0: "(X_poly Q\<^sub>p) \ carrier Q\<^sub>p_x" using UPQ.X_closed by blast then show ?thesis using P_def UPQ.P.nat_pow_closed p_natpow_closed(1) by blast qed have 1: "\\<^sub>p = univ_basic_semialg_set 2 P" proof show "\\<^sub>p \ univ_basic_semialg_set 2 P" proof fix x assume A: "x \ \\<^sub>p" show "x \ univ_basic_semialg_set 2 P" proof- have x_car: "x \ carrier Q\<^sub>p" using A val_ring_memE by blast then have "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat)))" using A Qp_val_ring_alt_def[of x] by blast then obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat))" by blast have "y \ carrier Q\<^sub>p \ P \ x = (y[^](2::nat))" proof- have "P \ x = \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat))" proof- have "((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat)) \ carrier Q\<^sub>p_x" using UPQ.monom_closed p_natpow_closed(1) by blast then have "P \ x = (((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using P_def x_car UPQ.to_fun_plus by blast then have 0: "P \ x = (\[^](3::nat)) \(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_smult p_natpow_closed(1) x_car by presburger have "(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) = (x[^](4::nat))" using UPQ.to_fun_X_pow x_car by blast then have "P \ x = (\[^](3::nat)) \(x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using "0" UPQ.to_fun_one x_car by presburger then show ?thesis using y_def Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1) x_car by presburger qed then show ?thesis using y_def by blast qed then show ?thesis unfolding univ_basic_semialg_set_def using x_car by blast qed qed show "univ_basic_semialg_set 2 P \ \\<^sub>p" proof fix x assume A: "x \ univ_basic_semialg_set (2::nat) P" then obtain y where y_def: "y \ carrier Q\<^sub>p \ (P \ x) = (y[^](2::nat))" unfolding univ_basic_semialg_set_def by blast have x_car: "x \ carrier Q\<^sub>p" using A by (metis (no_types, lifting) mem_Collect_eq univ_basic_semialg_set_def) have 0: "(P \ x) = (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using P_def x_car UPQ.UP_one_closed UPQ.monom_closed UPQ.monom_rep_X_pow UPQ.to_fun_monom UPQ.to_fun_one UPQ.to_fun_plus p_natpow_closed(1) by presburger have 1: "y \ carrier Q\<^sub>p \ (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ = (y[^](2::nat))" using "0" y_def by blast then show "x \ \\<^sub>p" using x_car Qp_val_ring_alt_def[of x] y_def by (metis Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1)) qed qed show ?thesis using 0 1 that by blast qed lemma Qp_val_ring_is_univ_semialgebraic: "is_univ_semialgebraic \\<^sub>p" proof- obtain P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" using Qp_val_ring_is_semialg by blast then show ?thesis by (metis univ_basic_semialg_set_is_univ_semialgebraic zero_neq_numeral) qed lemma Qp_val_ring_is_semialgebraic: "is_semialgebraic 1 (to_R1` \\<^sub>p)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Inverse Images of Semialgebraic Sets by Polynomial Maps\ (********************************************************************************************) (********************************************************************************************) lemma basic_semialg_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S = basic_semialg_set k m f" assumes "m \0" shows "poly_map n fs \\<^bsub>n\<^esub> S = basic_semialg_set n m (Qp_poly_comp n fs f)" proof show "poly_map n fs \\<^bsub>n\<^esub> S \ basic_semialg_set n m (Qp_poly_comp n fs f)" proof fix x assume A: "x \ poly_map n fs \\<^bsub>n\<^esub> S" then have 0: "poly_map n fs x \ S" proof - have "\n f. {rs. rs \ S} \ {rs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). \r. r \ carrier Q\<^sub>p \ Qp_ev f rs = (r[^](n::nat))}" by (metis (no_types) Collect_mem_eq \S = basic_semialg_set k m f\ basic_semialg_set_def eq_iff) then show ?thesis using A by blast qed have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A assms by (meson evimage_eq) have "\y \ (carrier Q\<^sub>p). Qp_ev f (poly_map n fs x) = (y[^]m)" using A 0 assms basic_semialg_set_def by blast then have "\y \ (carrier Q\<^sub>p). Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using 1 assms Qp_poly_comp_eval by blast then show "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" using "1" basic_semialg_set_def by blast qed show "basic_semialg_set n m (Qp_poly_comp n fs f) \ poly_map n fs \\<^bsub>n\<^esub> S" proof fix x assume A: "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_def by blast have 1: "(poly_map n fs x) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using "0" poly_map_closed assms(2) assms(3) by blast show "x \ poly_map n fs \\<^bsub>n\<^esub> S" proof- have "\y \ carrier Q\<^sub>p. Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using A basic_semialg_set_def by blast then have 2: "\y \ carrier Q\<^sub>p. Qp_ev f (poly_map n fs x) = (y[^]m)" using assms Qp_poly_comp_eval by (metis (no_types, lifting) A basic_semialg_set_def mem_Collect_eq) have 3: "poly_map n fs x \ S" using assms 0 1 basic_semialg_set_def[of k m f] "2" by blast show ?thesis using "0" "3" by blast qed qed qed lemma basic_semialg_pullback': assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "A \ basic_semialgs k" shows "poly_map n fs \\<^bsub>n\<^esub> A \ (basic_semialgs n)" proof- obtain f m where fm_def: "m \0 \f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>]) \ A = basic_semialg_set k m f" using assms by (metis is_basic_semialg_def mem_Collect_eq) then have "poly_map n fs \\<^bsub>n\<^esub> A = basic_semialg_set n m (Qp_poly_comp n fs f)" using assms basic_semialg_pullback[of f k n fs A m] by linarith then show ?thesis unfolding is_basic_semialg_def by (metis (mono_tags, lifting) assms(1) assms(2) fm_def mem_Collect_eq poly_compose_closed) qed lemma semialg_pullback: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "poly_map n fs \\<^bsub>n\<^esub> S \ semialg_sets n" unfolding semialg_sets_def apply(rule gen_boolean_algebra.induct[of S "(carrier (Q\<^sub>p\<^bsup>k\<^esup>))" "basic_semialgs k"]) using assms semialg_sets_def apply blast apply (metis assms(1) assms(2) carrier_is_semialgebraic evimageI2 extensional_vimage_closed is_semialgebraicE poly_map_closed semialg_sets_def subsetI subset_antisym) apply (metis Int_absorb2 assms(1) assms(2) basic_semialg_is_semialg basic_semialg_is_semialgebraic basic_semialg_pullback' is_semialgebraic_closed mem_Collect_eq semialg_sets_def) apply (metis evimage_Un semialg_sets_def semialg_union) by (metis assms(1) assms(2) carrier_is_semialgebraic diff_is_semialgebraic evimage_Diff extensional_vimage_closed is_semialgebraicE is_semialgebraicI poly_map_closed poly_map_pullbackI semialg_sets_def subsetI subset_antisym) lemma pullback_is_semialg: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "is_semialgebraic n (poly_map n fs \\<^bsub>n\<^esub> S)" using assms(1) assms(2) assms(3) is_semialgebraicI padic_fields_axioms semialg_pullback by blast text\Equality and inequality sets for a pair of polynomials\ definition val_ineq_set where "val_ineq_set n f g = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" lemma poly_map_length : assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = m" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) length_map restrict_apply') lemma val_ineq_set_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g = poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set " proof show "val_ineq_set n f g \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" proof fix x assume "x \ val_ineq_set n f g" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ val (Qp_ev f x) \ val (Qp_ev g x)" by (metis (mono_tags, lifting) mem_Collect_eq val_ineq_set_def) have 1: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" unfolding poly_map_def poly_tuple_eval_def using 0 by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') have 2: "poly_map n [g,f] x \ val_relation_set" apply(rule val_relation_setI) using 1 0 assms apply (metis eval_at_point_closed nth_Cons_0) using 1 0 assms apply (metis One_nat_def eval_at_point_closed diff_Suc_1 less_numeral_extra(1) nth_Cons_pos Qp.to_R_to_R1) using poly_map_length assms 0 apply (metis "1" Qp_2I cartesian_power_car_memE eval_at_point_closed) by (metis "0" "1" One_nat_def nth_Cons_0 nth_Cons_Suc) have 3: "is_poly_tuple n [g, f]" using assms by (smt One_nat_def diff_Suc_1 Qp_is_poly_tupleI length_Suc_conv less_SucE less_one list.size(3) nth_Cons') then show "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" using 0 1 2 by blast qed show "poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set \ val_ineq_set n f g" proof fix x have 0: "is_poly_tuple n [g, f]" using Qp_is_poly_tupleI assms by (metis (no_types, lifting) diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons') assume A: "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ poly_map n [g,f] x \ val_relation_set" using 0 by (meson evimageD extensional_vimage_closed subsetD) have 2: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" by (metis "1" Qp_poly_mapE' length_0_conv poly_map_cons) show "x \ val_ineq_set n f g" using 0 1 2 unfolding val_ineq_set_def val_relation_set_def by (metis (no_types, lifting) "1" list.inject mem_Collect_eq nth_Cons_0 poly_map_apply val_relation_setE) qed qed lemma val_ineq_set_is_semialg: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g \ semialg_sets n" proof- have 0: "val_relation_set \ semialg_sets 2" using val_relation_semialg basic_semialg_is_semialg' by (metis Qp_val_poly_closed zero_neq_numeral) show ?thesis using val_ineq_set_pullback semialg_pullback[of n "[g,f]" 2 "val_relation_set" ] by (metis (no_types, lifting) "0" assms(1) assms(2) diff_Suc_1 Qp_is_poly_tupleI length_Cons less_Suc0 less_SucE list.size(3) nth_Cons_0 nth_Cons_pos numeral_2_eq_2 zero_neq_numeral) qed lemma val_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n (val_ineq_set n f g)" using assms(1) assms(2) is_semialgebraicI val_ineq_set_is_semialg by blast lemma val_ineq_setI: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ (val_ineq_set n f g)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "val (Qp_ev f x) \ val (Qp_ev g x)" using assms unfolding val_ineq_set_def apply blast using assms unfolding val_ineq_set_def by blast lemma val_ineq_setE: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "val (Qp_ev f x) \ val (Qp_ev g x)" shows "x \ (val_ineq_set n f g)" using assms unfolding val_ineq_set_def by blast lemma val_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast lemma val_eq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" proof- have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 1: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 2: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} \ {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule equalityI, rule subsetI , rule IntI) unfolding mem_Collect_eq using le_less apply blast apply (metis order_refl) apply(rule subsetI, erule IntE) unfolding mem_Collect_eq by (meson less_le_trans not_less_iff_gr_or_eq) show ?thesis unfolding 2 apply(rule intersection_is_semialg) using 0 apply blast using 1 by blast qed lemma equalityI'': assumes "\x. A x \ B x" assumes "\x. B x \ A x" shows "{x. A x} = {x. B x}" using assms by blast lemma val_strict_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" proof- have 0: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} - {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule equalityI', rule DiffI) unfolding le_less mem_Collect_eq apply blast unfolding mem_Collect_eq using neq_iff apply blast apply(erule DiffE) unfolding mem_Collect_eq by blast show ?thesis unfolding 0 apply(rule diff_is_semialgebraic) using assms val_ineq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def apply blast using assms val_eq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def by blast qed lemma constant_poly_val_exists: shows "\g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" proof- obtain a where a_def: "a \ carrier Q\<^sub>p \ val a = c" by (meson Qp.minus_closed Qp.nonzero_closed dist_nonempty' p_nonzero) obtain g where g_def: "g = coord_const a" by blast show ?thesis using a_def g_def Qp_to_IP_car by (metis (no_types, opaque_lifting) Qp_to_IP_car a_def eval_at_point_const g_def le_less subset_iff) qed lemma val_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" apply(rule val_ineq_set_is_semialgebraic') using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_ineq_set_is_semialgebraic''': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule val_ineq_set_is_semialgebraic') using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_eq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule val_eq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" apply(rule equalityI'') using g_def apply fastforce using g_def by metis show ?thesis using 0 unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" apply(rule equalityI'') using assms g_def apply fastforce using assms g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma(in cring) R1_memE: assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "x = [(hd x)]" using assms cartesian_power_car_memE by (metis diff_is_0_eq' hd_conv_nth le_eq_less_or_eq length_0_conv length_tl list.exhaust list.sel(3) normalize.cases nth_Cons_0 zero_neq_one) lemma(in cring) R1_memE': assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "hd x \ carrier R" using R1_memE assms cartesian_power_car_memE[of x R 1] cartesian_power_car_memE'[of x R 1 0] by (metis hd_conv_nth less_numeral_extra(1) list.size(3) zero_neq_one) lemma univ_val_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x \ c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c}" apply(rule val_ineq_set_is_semialgebraic'') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c} = to_R1 ` {x \ carrier Q\<^sub>p. val x \ c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a \ c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_strict_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x < c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) < c} = to_R1 ` {x \ carrier Q\<^sub>p. val x < c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a < c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_eq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x = c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c}" apply(rule val_eq_set_is_semialgebraic') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c} = to_R1 ` {x \ carrier Q\<^sub>p. val x = c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" using A 0 unfolding mem_Collect_eq 1 by blast qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a = c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\One Dimensional $p$-adic Balls are Semialgebraic\ (********************************************************************************************) (********************************************************************************************) lemma coord_ring_one_def: "Pring Q\<^sub>p {(0::nat)} = (Q\<^sub>p[\\<^bsub>1\<^esub>])" proof- have "{(0::nat)} = {..<1}" by auto thus ?thesis unfolding coord_ring_def by auto qed lemma times_p_pow_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]n \ a" shows "val b = val a + n" using val_mult[of "\[^]n" a] assms unfolding assms(2) val_p_int_pow by (metis add.commute p_intpow_closed(1)) lemma times_p_pow_neg_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]-n \ a" shows "val b = val a - n" by (metis Qp.m_comm Qp_int_pow_nonzero assms(1) assms(2) p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) lemma eint_minus_int_pos: assumes "a - eint n \ 0" shows "a \ n" using assms apply(induction a) apply (metis diff_ge_0_iff_ge eint_ord_simps(1) idiff_eint_eint zero_eint_def) by simp text\\p\-adic balls as pullbacks of polynomial maps\ lemma balls_as_pullbacks: assumes "c \ carrier Q\<^sub>p" shows "\P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]). to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p)" proof- obtain P0 where P0_def: "P0 = (to_poly (\[^](-n))) \\<^bsub>Q\<^sub>p_x\<^esub>((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c)" by blast have 0: "P0 \ carrier Q\<^sub>p_x" proof- have P0: "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have P1: "(to_poly (\[^](-n))) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast then show ?thesis using P0_def P0 P1 by blast qed have 1: "\x. x \ carrier Q\<^sub>p \ P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" proof- fix x assume A: "x \ carrier Q\<^sub>p" have P0: "(to_poly (\[^](-n))) \ x = (\[^](-n))" using A UPQ.to_fun_to_poly p_intpow_closed(1) by blast have P1: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (x \\<^bsub>Q\<^sub>p\<^esub> c)" by (metis A UPQ.to_fun_X_minus X_poly_minus_def assms) have P2: "to_poly (\[^](-n)) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast have P3: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = to_poly (\[^]- n) \ x \ (((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x)" using A P0_def P0 P1 P2 P3 to_fun_mult[of "to_poly (\[^](-n))" "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c" x] UPQ.to_fun_mult by blast then have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c) " by (metis P0 P1) then show "P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" using P0_def by metis qed have 2: " (\a. [a]) ` B\<^bsub>n\<^esub>[c] = poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof show "(\a. [a]) ` B\<^bsub>n\<^esub>[c] \ poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof fix x assume A: "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" then obtain a where a_def: "x = [a] \ a \ B\<^bsub>n\<^esub>[c]" by blast have P0: "P0 \ a \ \\<^sub>p" proof- have "B\<^bsub>n\<^esub>[c] \ carrier Q\<^sub>p" using c_ball_in_Qp by blast hence a_closed: "a \ carrier Q\<^sub>p" using a_def by blast have P0: "P0 \ a = (\[^](-n)) \ (a \ c)" using 1 a_def c_ballE(1) by blast then have P1: "val (P0 \ a) = val (\[^](-n)) + val (a \ c)" using val_mult[of "\[^]-n" "a \ c"] a_closed assms Qp.minus_closed p_intpow_closed(1) by presburger then have P2: "val (P0 \ a) = val (a \\<^bsub>Q\<^sub>p\<^esub> c) - n" by (metis P0 Qp.m_comm Qp.minus_closed Qp_int_pow_nonzero assms local.a_closed p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) have P3: "val (a \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using a_def c_ballE(2) by blast then have "val (P0 \ a) \ -n + n" using P2 by (metis add.commute diff_conv_add_uminus diff_self local.eint_minus_ineq' zero_eint_def) then have P4: "val (P0 \ a) \ 0" by (metis add.commute add.right_inverse zero_eint_def) have P5: "P0 \ a \ carrier Q\<^sub>p" using "0" UPQ.to_fun_closed local.a_closed by blast then show ?thesis using P4 using val_ring_val_criterion by blast qed have "poly_map 1 [from_Qp_x P0] x = [Qp_ev (from_Qp_x P0) [a]] " using a_def poly_map_def[of 1 "[from_Qp_x P0]"] poly_tuple_eval_def[of ] by (metis Qp_poly_mapE' c_ballE(1) length_0_conv poly_map_cons Qp.to_R1_closed) then have "poly_map 1 [from_Qp_x P0] x = [P0 \ a] " using Qp_x_Qp_poly_eval[of P0 a] by (metis "0" a_def c_ballE(1)) then have P1: "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using P0 by blast have P2: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A c_ballE(1) Qp.to_R1_closed by blast have P3: "is_poly_tuple 1 [from_Qp_x P0]" apply(rule Qp_is_poly_tupleI) by (metis "0" Qp_is_poly_tupleI from_Qp_x_closed gr_implies_not0 is_poly_tupleE is_poly_tuple_Cons list.size(3) zero_neq_one) show "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p" using P3 P2 P1 unfolding evimage_def poly_map_def by blast qed have 20: "is_poly_tuple 1 [from_Qp_x P0]" using 0 UP_to_IP_closed[of P0 "0::nat"] unfolding is_poly_tuple_def by (metis (no_types, lifting) empty_set from_Qp_x_closed list.simps(15) singletonD subset_code(1)) show "poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" proof fix x assume A: "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" have A0: "(\a. [a]) ` \\<^sub>p \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def Qp.to_R1_car_subset Qp_val_ring_is_semialgebraic is_semialgebraic_closed by presburger have "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using A0 A 20 by blast then obtain a where a_def: "a \ \\<^sub>p \ (poly_map 1 [from_Qp_x P0] x) = [a]" by blast have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A by (meson evimage_eq) then obtain y where y_def: "x = [y] \ y \ carrier Q\<^sub>p" using A by (metis Qp.to_R1_to_R Qp.to_R_pow_closed) have "(poly_map 1 [from_Qp_x P0] x) = [(Qp_ev (from_Qp_x P0) [y])]" unfolding poly_map_def poly_tuple_eval_def using x_closed by (smt "20" One_nat_def length_Suc_conv list.size(3) nth_Cons_0 nth_map poly_tuple_eval_closed poly_tuple_eval_def restrict_apply' Qp.to_R1_to_R y_def zero_less_Suc) then have "(poly_map 1 [from_Qp_x P0] x) = [P0 \ y]" by (metis "0" Qp_x_Qp_poly_eval y_def) then have "[a] = [P0 \ y]" using a_def by presburger then have A1: "a = (\[^](-n)) \ (y \\<^bsub>Q\<^sub>p\<^esub> c)" using 1[of y] y_def by blast have "y \ B\<^bsub>n\<^esub>[c]" proof- have B0: "val a = val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n" using A1 y_def Qp.minus_closed assms times_p_pow_neg_val by blast have B1: "val a \ 0" using a_def val_ring_memE by blast then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n \ 0" using B0 by metis then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using eint_minus_int_pos by blast then show "y \ B\<^bsub>n\<^esub>[c]" using c_ballI y_def by blast qed then show "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" using y_def by blast qed qed then show ?thesis by (meson "0" from_Qp_x_closed) qed lemma ball_is_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1` B\<^bsub>n\<^esub>[c])" proof- obtain P where P_def: "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]) \ to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p) " using assms balls_as_pullbacks[of c n] by meson have "is_poly_tuple 1 [P]" using P_def unfolding is_poly_tuple_def by (metis (no_types, opaque_lifting) list.inject list.set_cases neq_Nil_conv subset_code(1)) then show ?thesis using assms P_def pullback_is_semialg[of 1 "[P]" 1 "((\a. [a]) ` \\<^sub>p) "] by (metis (mono_tags, lifting) One_nat_def Qp_val_ring_is_semialgebraic is_semialgebraic_def length_Suc_conv list.distinct(1) list.size(3)) qed lemma ball_is_univ_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (B\<^bsub>n\<^esub>[c])" using assms ball_is_semialgebraic c_ball_in_Qp is_univ_semialgebraic_def by presburger abbreviation Qp_to_R1_set where "Qp_to_R1_set S \ to_R1 ` S" (********************************************************************************************) (********************************************************************************************) subsubsection\Finite Unions and Intersections of Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition are_semialgebraic where "are_semialgebraic n Xs = (\ x. x \ Xs \ is_semialgebraic n x)" lemma are_semialgebraicI: assumes "\x. x \ Xs \ is_semialgebraic n x " shows "are_semialgebraic n Xs" using are_semialgebraic_def assms by blast lemma are_semialgebraicE: assumes "are_semialgebraic n Xs" assumes "x \ Xs" shows "is_semialgebraic n x" using are_semialgebraic_def assms(1) assms(2) by blast definition are_univ_semialgebraic where "are_univ_semialgebraic Xs = (\ x. x \ Xs \ is_univ_semialgebraic x)" lemma are_univ_semialgebraicI: assumes "\x. x \ Xs \ is_univ_semialgebraic x " shows "are_univ_semialgebraic Xs" using are_univ_semialgebraic_def assms by blast lemma are_univ_semialgebraicE: assumes "are_univ_semialgebraic Xs" assumes "x \ Xs" shows "is_univ_semialgebraic x" using are_univ_semialgebraic_def assms(1) assms(2) by blast lemma are_univ_semialgebraic_semialgebraic: assumes "are_univ_semialgebraic Xs" shows "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" apply(rule are_semialgebraicI) using are_univ_semialgebraicE assms image_iff is_univ_semialgebraicE by (metis (no_types, lifting)) lemma to_R1_set_union: "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" using image_iff by blast lemma to_R1_inter: assumes "Xs \ {}" shows "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" proof show "to_R1 ` (\ Xs) \ \ (Qp_to_R1_set ` Xs)" by blast show "\ (Qp_to_R1_set ` Xs) \ to_R1 ` (\ Xs)" proof fix x assume A: "x \ \ (Qp_to_R1_set ` Xs)" then have 0: "\S. S \ Xs \ x \ (Qp_to_R1_set S)" by blast obtain S where "S \ Xs \ x \ (Qp_to_R1_set S)" using assms 0 by blast then obtain b where b_def: "b \ S \ x = [b]" by blast have "b \ (\ Xs)" using "0" b_def by blast then show "x \ to_R1 ` (\ Xs)" using b_def by blast qed qed lemma finite_union_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply (simp add: empty_is_semialgebraic) by (metis Sup_insert insert_subset is_semialgebraicI union_is_semialgebraic) lemma finite_union_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n " shows "is_semialgebraic n (\ Xs)" using assms(1) assms(2) finite_union_is_semialgebraic by blast lemma(in padic_fields) finite_union_is_semialgebraic'': assumes "finite S" assumes "\x. x \ S \ is_semialgebraic m (F x)" shows "is_semialgebraic m (\ x \ S. F x)" using assms finite_union_is_semialgebraic[of "F`S" m] unfolding is_semialgebraic_def by blast lemma finite_union_is_univ_semialgebraic': assumes "finite Xs" assumes "are_univ_semialgebraic Xs" shows "is_univ_semialgebraic (\ Xs)" proof- have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms finite_union_is_semialgebraic'[of "((`) (\a. [a]) ` Xs)"] to_R1_set_union[of Xs] by (metis (no_types, lifting) are_semialgebraicE are_univ_semialgebraic_semialgebraic finite_imageI is_semialgebraicE subsetI) then show ?thesis using is_univ_semialgebraicI by blast qed lemma finite_intersection_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ Xs \{} \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply auto[1] proof fix A::"((nat \ int) \ (nat \ int)) set list set set" fix a assume 0: "finite A" assume 1: "A \ semialg_sets n \ A \ {} \ is_semialgebraic n (\ A) " assume 2: "insert a A \ semialg_sets n \ insert a A \ {}" show "is_semialgebraic n (\ (insert a A))" proof(cases "A = {}") case True then have "insert a A = {a}" by simp then show ?thesis by (metis "2" cInf_singleton insert_subset is_semialgebraicI) next case False then have "A \ semialg_sets n \ A \ {}" using "2" by blast then have "is_semialgebraic n (\ A) " using "1" by linarith then show ?thesis using 0 1 2 intersection_is_semialg by (metis Inf_insert insert_subset is_semialgebraicI) qed qed lemma finite_intersection_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (simp add: assms(1) assms(2) finite_intersection_is_semialgebraic) lemma finite_intersection_is_semialgebraic'': assumes "finite Xs" assumes "are_semialgebraic n Xs \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (meson are_semialgebraicE assms(1) assms(2) finite_intersection_is_semialgebraic' is_semialgebraicE subsetI) lemma finite_intersection_is_univ_semialgebraic: assumes "finite Xs" assumes "are_univ_semialgebraic Xs" assumes "Xs \ {}" shows "is_univ_semialgebraic (\ Xs)" proof- have "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" using are_univ_semialgebraic_semialgebraic assms(2) by blast then have "is_semialgebraic 1 (\ (Qp_to_R1_set ` Xs))" using assms finite_intersection_is_semialgebraic''[of "Qp_to_R1_set ` Xs" 1] by blast then have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms to_R1_inter[of Xs] by presburger then show ?thesis using is_univ_semialgebraicI by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Cartesian Products of Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) lemma Qp_times_basic_semialg_right: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" proof show "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ basic_semialg_set (n + m) k a" proof fix x assume "x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" then obtain as bs where as_bs_def: "as \ (basic_semialg_set n k a) \ bs \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ x = as@bs" using cartesian_product_memE[of x "basic_semialg_set n k a" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p n] basic_semialg_set_def by (metis (no_types, lifting) append_take_drop_id basic_semialg_set_memE(1) subsetI) have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using as_bs_def basic_semialg_set_memE(1) cartesian_product_closed' by blast have 1: "(Qp_ev a x = Qp_ev a as)" using as_bs_def poly_eval_cartesian_prod[of as n bs m a] assms basic_semialg_set_memE(1) by blast obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev a as = (y[^]k))" using as_bs_def using basic_semialg_set_memE(2)[of as n k a] by blast show "x \ basic_semialg_set (n + m) k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "0") apply (simp add: y_def) using "1" y_def by blast qed show "basic_semialg_set (n + m) k a \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" proof fix x assume A: "x \ basic_semialg_set (n + m) k a" have A0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A basic_semialg_set_memE(1) by blast have A1: "set x \ carrier Q\<^sub>p" using A0 by (metis (no_types, lifting) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth subsetI) have A2: "length x = n + m" using A0 cartesian_power_car_memE by blast obtain as where as_def: "as = take n x" by blast obtain bs where bs_def: "bs = drop n x" by blast have 0: "x = as@bs" using A as_def bs_def by (metis append_take_drop_id) have 1: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) using as_def A2 apply (simp add: A2 min.absorb2) by (metis (no_types, lifting) A1 as_def dual_order.trans set_take_subset) have 2: "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule cartesian_power_car_memI) using bs_def A2 apply (simp add: A2) by (metis A1 bs_def order.trans set_drop_subset) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)" using basic_semialg_set_memE A by meson have 3: "as \ basic_semialg_set n k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using \y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)\ apply blast using y_def A 1 0 2 assms(1) poly_eval_cartesian_prod by blast show " x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" using 3 2 "0" by (metis (mono_tags, lifting) as_def basic_semialg_set_memE(1) bs_def cartesian_product_memI subsetI) qed qed lemma Qp_times_basic_semialg_right_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: " a \ carrier (Q\<^sub>p[\\<^bsub>n+m\<^esub>])" using assms poly_ring_car_mono'(2) by blast have 2: "is_semialgebraic (n + m) (basic_semialg_set (n + m) k a)" using assms basic_semialg_is_semialgebraic'[of a "n+m" k "basic_semialg_set (n + m) k a"] "1" by blast show ?thesis using 0 2 by metis qed lemma Qp_times_basic_semialg_right_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_right_is_semialgebraic[of k P] using assms(1) by blast qed lemma cartesian_product_memE': assumes "x \ cartesian_product A B" obtains a b where "a \ A \ b \ B \ x = a@b" using assms unfolding cartesian_product_def by blast lemma Qp_times_basic_semialg_left: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) = basic_semialg_set (n+m) k (shift_vars n m a)" proof show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) \ basic_semialg_set (n + m) k (shift_vars n m a)" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" then obtain as bs where as_bs_def: "as \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ bs \ (basic_semialg_set n k a) \ x = as@bs " using cartesian_product_memE' by blast have 0: "Qp_ev (shift_vars n m a) x = Qp_ev a bs" using A as_bs_def assms shift_vars_eval[of a n as m bs ] by (metis (no_types, lifting) basic_semialg_set_memE(1)) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a bs = (y[^]k)" using as_bs_def basic_semialg_set_memE(2) by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A as_bs_def by (metis (no_types, lifting) add.commute basic_semialg_set_memE(1) cartesian_product_closed') show "x \ basic_semialg_set (n + m) k (shift_vars n m a)" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using y_def apply blast using "0" y_def by blast qed show "basic_semialg_set (n + m) k (shift_vars n m a) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) " proof fix x assume A: "x \ basic_semialg_set (n + m) k (shift_vars n m a)" then obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev (shift_vars n m a) x = (y[^]k)" using assms basic_semialg_set_memE[of x "n + m" k "shift_vars n m a"] shift_vars_closed[of a m] Qp.cring_axioms by blast have "x \ carrier (Q\<^sub>p\<^bsup>m+n\<^esup>)" using A basic_semialg_set_memE(1) by (metis add.commute) then have "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using cartesian_product_carrier by blast then obtain as bs where as_bs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson cartesian_product_memE') have "bs \ (basic_semialg_set n k a)" apply(rule basic_semialg_set_memI[of _ _ y]) using as_bs_def apply blast apply (simp add: y_def) using y_def shift_vars_eval[of a n as m bs ] as_bs_def assms(1) by metis then show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" using as_bs_def unfolding cartesian_product_def by blast qed qed lemma Qp_times_basic_semialg_left_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a))" using basic_semialg_is_semialgebraic'[of a "n+m" k] Qp_times_basic_semialg_left by (metis assms(1) assms(2) basic_semialg_is_semialgebraic is_basic_semialg_def neq0_conv shift_vars_closed) lemma Qp_times_basic_semialg_left_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) A)" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic[of k P] using assms(1) by blast qed lemma product_of_basic_semialgs_is_semialg: assumes "k > 0" assumes "l > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: "cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b) = basic_semialg_set (m + n) l (shift_vars m n b)" using Qp_times_basic_semialg_left assms by blast have 2: "(cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b)) = cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b)" proof- have 0: "basic_semialg_set n k a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using basic_semialg_set_memE(1) by blast have 1: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by simp have 2: "carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by simp have 3: "basic_semialg_set m l b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using basic_semialg_set_memE(1) by blast show ?thesis using 0 1 2 3 cartesian_product_intersection[of "(basic_semialg_set n k a)" Q\<^sub>p n "(carrier (Q\<^sub>p\<^bsup>m\<^esup>))" m "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))" "(basic_semialg_set m l b)"] by (smt Collect_cong inf.absorb_iff1 inf.absorb_iff2) qed then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic Qp_times_basic_semialg_right_is_semialgebraic assms by (metis (no_types, lifting) add.commute intersection_is_semialg) qed lemma product_of_basic_semialgs_is_semialg': assumes "A \ (basic_semialgs n)" assumes "B \ (basic_semialgs m)" shows "is_semialgebraic (n + m) (cartesian_product A B)" proof- obtain k a where ka_def: "k > 0 \ a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ A = (basic_semialg_set n k a)" using assms by (metis is_basic_semialg_def mem_Collect_eq neq0_conv) obtain l b where lb_def: "l > 0 \ b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]) \ B = (basic_semialg_set m l b)" by (metis assms(2) gr_zeroI is_basic_semialg_def mem_Collect_eq) show ?thesis using ka_def lb_def assms product_of_basic_semialgs_is_semialg by blast qed lemma car_times_semialg_is_semialg: assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms is_semialgebraic_def semialg_sets_def apply blast apply (simp add: carrier_is_semialgebraic cartesian_product_carrier) proof- show "\A. A \ basic_semialgs m \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- fix A assume A: "A \ basic_semialgs m " then have " (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = A" by (metis basic_semialg_set_memE(1) inf_absorb1 is_basic_semialg_def mem_Collect_eq subsetI) then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" using add.commute[of n m] assms A Qp_times_basic_semialg_left_is_semialgebraic' by (simp add: \n + m = m + n\) qed show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- fix A C assume A: " A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" " is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" then have B: "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" using union_is_semialgebraic by blast show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- have 0: "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(1) gen_boolean_algebra_subset by blast have 1: " C \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(3) gen_boolean_algebra_subset by blast then show ?thesis using 0 A B using cartesian_product_binary_union_right[of A Q\<^sub>p m C "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))"] unfolding is_semialgebraic_def semialg_sets_def by presburger qed qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" then have "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using gen_boolean_algebra_subset by blast then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" using A cartesian_product_car_complement_right[of A Q\<^sub>p m n] unfolding is_semialgebraic_def semialg_sets_def by (metis (mono_tags, lifting) semialg_complement semialg_sets_def) qed qed lemma basic_semialg_times_semialg_is_semialg: assumes "A \ basic_semialgs n" assumes "is_semialgebraic m B" shows " is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms(2) is_semialgebraic_def semialg_sets_def apply blast using Qp_times_basic_semialg_right_is_semialgebraic' assms(1) apply blast apply (metis assms(1) basic_semialg_is_semialgebraic inf.absorb1 is_semialgebraic_closed mem_Collect_eq product_of_basic_semialgs_is_semialg') apply (metis (no_types, lifting) cartesian_product_binary_union_right is_semialgebraicI is_semialgebraic_closed semialg_sets_def union_is_semialgebraic) proof- show "\Aa. Aa \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product A Aa) \ is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - Aa))" proof- fix B assume A: "B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B))" using A assms cartesian_product_complement_right[of B Q\<^sub>p m A n] add.commute[of n m] proof - have f1: "\n B. \ is_semialgebraic n B \ B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson is_semialgebraic_closed) have "is_basic_semialg n A" using \A \ {S. is_basic_semialg n S}\ by blast then have f2: "is_semialgebraic n A" using padic_fields.basic_semialg_is_semialgebraic padic_fields_axioms by blast have "B \ semialg_sets m" using \B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {S. is_basic_semialg m S}\ semialg_sets_def by blast then have "is_semialgebraic m B" by (meson padic_fields.is_semialgebraicI padic_fields_axioms) then show ?thesis using f2 f1 by (metis (no_types) Qp_times_basic_semialg_right_is_semialgebraic' \A \ {S. is_basic_semialg n S}\ \\B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>); A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \ cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) - cartesian_product A B = cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B)\ \is_semialgebraic (n + m) (cartesian_product A B)\ diff_is_semialgebraic) qed qed qed text\Semialgebraic sets are closed under cartesian products\ lemma cartesian_product_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of A "carrier (Q\<^sub>p\<^bsup>n\<^esup>)""basic_semialgs n"]) using assms is_semialgebraicE semialg_sets_def apply blast using assms car_times_semialg_is_semialg apply blast using assms basic_semialg_times_semialg_is_semialg apply (simp add: Int_absorb2 basic_semialg_is_semialgebraic is_semialgebraic_closed) proof- show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product C B) \ is_semialgebraic (n + m) (cartesian_product (A \ C) B)" proof- fix A C assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product C B)" show "is_semialgebraic (n + m) (cartesian_product (A \ C) B)" using A cartesian_product_binary_union_left[of A Q\<^sub>p n C B] by (metis (no_types, lifting) gen_boolean_algebra_subset union_is_semialgebraic) qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" using assms A cartesian_product_complement_left[of A Q\<^sub>p n B m] unfolding is_semialgebraic_def semialg_sets_def by (metis car_times_semialg_is_semialg diff_is_semialgebraic is_semialgebraicE is_semialgebraicI is_semialgebraic_closed semialg_sets_def) qed qed (**************************************************************************************************) (**************************************************************************************************) subsection\$N^{th}$ Power Residues\ (**************************************************************************************************) (**************************************************************************************************) definition nth_root_poly where "nth_root_poly (n::nat) a = ((X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> n) \\<^bsub>Q\<^sub>p_x\<^esub> (to_poly a)" lemma nth_root_poly_closed: assumes "a \ carrier Q\<^sub>p" shows "nth_root_poly n a \ carrier Q\<^sub>p_x" using assms unfolding nth_root_poly_def by (meson UPQ.P.minus_closed UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_poly_closed) lemma nth_root_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "(nth_root_poly n a) \ b = (b[^]n) \\<^bsub>Q\<^sub>p\<^esub> a" using assms unfolding nth_root_poly_def using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_X_pow UPQ.to_fun_diff UPQ.to_fun_to_poly UPQ.to_poly_closed by presburger text\Hensel's lemma gives us this criterion for the existence of \n\-th roots\ lemma nth_root_poly_root: assumes "(n::nat) > 1" assumes "a \ \\<^sub>p" assumes "a \ \" assumes "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > 2* val ([n]\\)" shows "(\ b \ \\<^sub>p. ((b[^]n) = a))" proof- obtain \ where alpha_def: "\ \ carrier Z\<^sub>p \ \ \ = a" using assms(2) by blast have 0: "\ \ carrier Z\<^sub>p" by (simp add: alpha_def) have 1: "\ \ \\<^bsub>Z\<^sub>p\<^esub>" using assms alpha_def inc_of_one by blast obtain N where N_def: "N = [n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" by blast have N_closed: "N \ carrier Z\<^sub>p" using N_def Zp_nat_mult_closed by blast have 2: "\ ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = ([n]\ \)" proof(induction n) case 0 have 00: "[(0::nat)] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = \\<^bsub>Z\<^sub>p\<^esub>" using Zp_nat_inc_zero by blast have 01: "[(0::nat)] \\<^bsub>Q\<^sub>p\<^esub> \ = \" using Qp.nat_inc_zero by blast then show ?case using 00 inc_of_nat by blast next case (Suc n) then show ?case using inc_of_nat by blast qed have 3: "val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \) = val (\ \\<^bsub>Q\<^sub>p\<^esub> a)" using alpha_def Zp.one_closed ring_hom_one[of \ Z\<^sub>p Q\<^sub>p] inc_is_hom Zp.ring_hom_minus[of Q\<^sub>p \ "\\<^bsub>Z\<^sub>p\<^esub>" \ ] Qp.ring_axioms unfolding \_def by (metis Q\<^sub>p_def Zp.minus_closed Zp_def padic_fields.val_of_inc padic_fields_axioms) have 4: "([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) \ nonzero Z\<^sub>p" proof- have 40: "int n \ 0" using of_nat_0_le_iff by blast have "nat (int n) = n" using nat_int by blast hence "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = [int n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" using 40 unfolding add_pow_def int_pow_def nat_pow_def proof - have "(if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) 0 else rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n) = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n" by (meson of_nat_less_0_iff) then show "rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n = (let f = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) in if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> f (nat (- int n)) else f (nat (int n)))" using \nat (int n) = n\ by presburger qed thus ?thesis using Zp_char_0[of n] Zp.not_nonzero_memE Zp_char_0' assms(1) gr_implies_not_zero by blast qed then have 5: "val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = val ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))" using 2 ord_of_inc by (metis N_closed N_def val_of_inc) then have 6: "(val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \)) > 2*(val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>))" using assms 3 by presburger have "\ b \ carrier Z\<^sub>p. (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" using Zp_nth_root_lemma[of \ n] assms "0" "1" "6" by blast then obtain b where b_def: "b \ carrier Z\<^sub>p \ (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" by blast then have "\ (b [^]\<^bsub>Z\<^sub>p\<^esub>n) = a" using alpha_def by blast then have "(\ b) [^] n = a" by (metis Qp.nat_inc_zero Q\<^sub>p_def Qp.nat_pow_zero Zp.nat_pow_0 Zp.nat_pow_zero Zp_nat_inc_zero \_def alpha_def assms(3) b_def frac_inc_of_nat inc_of_one inc_pow not_nonzero_Qp) then show ?thesis using b_def by blast qed text\All points sufficiently close to 1 have nth roots\ lemma eint_nat_times_2: "2*(n::nat) = 2*eint n" using times_eint_simps(1) by (metis mult.commute mult_2_right of_nat_add) lemma P_set_of_one: "P_set 1 = nonzero Q\<^sub>p" apply(rule equalityI) apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq apply blast apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq using Qp.nat_pow_eone by blast lemma nth_power_fact: assumes "(n::nat) \ 1" shows "\ (m::nat) > 0. \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof(cases "n = 1") case True have "\ u \ carrier Q\<^sub>p. ac 1 u = 1 \ val u = 0 \ u \ P_set n" unfolding True P_set_of_one by (metis iless_Suc_eq padic_fields.val_ring_memE padic_fields.zero_in_val_ring padic_fields_axioms val_nonzero zero_eint_def) then show ?thesis by blast next case F: False obtain m where m_def: "m = 1 + nat ( 2*(ord ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))))" by blast then have m_pos: "m > 0" by linarith have "\ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof fix u assume A: "u \ carrier Q\<^sub>p" show " ac m u = 1 \ val u = 0 \ u \ P_set n" proof assume B: "ac m u = 1 \ val u = 0" then have 0: "val u = val \" by (smt A ac_def not_nonzero_Qp val_one val_ord zero_eint_def) have 1: "ac m u = ac m \" by (metis B Qp.one_nonzero ac_p ac_p_int_pow_factor angular_component_factors_x angular_component_p inc_of_one m_pos p_nonzero) have 2: "u \ nonzero Q\<^sub>p" proof- have "ac m \ = 0" by (meson ac_def) then have "u \ \" by (metis B zero_neq_one) then show ?thesis using A not_nonzero_Qp Qp.nonzero_memI by presburger qed then have 3: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using m_pos 0 1 ac_ord_prop[of "\" u "0::int" m] by (metis B Qp.one_nonzero add.right_neutral eint.inject val_ord zero_eint_def) show "u \ P_set n" proof(cases "u = \") case True then show ?thesis by (metis P_set_one insert_iff zeroth_P_set) next case False have F0: "u \ \\<^sub>p" apply(rule val_ring_memI, rule A) unfolding 0 val_one by auto have F1: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using False 3 by blast have "ord (\ \ u) \ m" by (metis A F1 False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) val_ord) hence F2: "ord (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using m_def F1 A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) int_nat_eq of_nat_1 of_nat_add val_ord[of "\ \ u"] eint_nat_times_2 by linarith have "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(val ([n]\ \))" proof- have 0: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using F2 val_ord[of "\ \ u"] A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(2) by presburger have "n > 0" using assms by linarith hence "eint (ord ([n] \ \)) = val ([n] \ \)" using val_ord_nat_inc[of n] by blast hence "2*ord ([n]\ \) = 2*val ([n]\ \)" by (metis inc_of_nat times_eint_simps(1)) thus ?thesis using 0 val_ord[of "\ \ u"] assms by presburger qed then have "(\ b \ \\<^sub>p. ((b[^]n) = u))" using m_def False nth_root_poly_root[of n u] F0 assms F by linarith then have "(\ b \ carrier Q\<^sub>p. ((b[^]n) = u))" using val_ring_memE by blast then show "u \ P_set n" using P_set_def[of n] 2 by blast qed qed qed then show ?thesis using m_pos by blast qed definition pow_res where "pow_res (n::nat) x = {a. a \ carrier Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. (a = x \ (y[^]n)))}" lemma nonzero_pow_res: assumes "x \ nonzero Q\<^sub>p" shows "pow_res (n::nat) x \ nonzero Q\<^sub>p" proof fix a assume "a \ pow_res n x" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (a = x \ (y[^]n))" using pow_res_def by blast then show "a \ nonzero Q\<^sub>p" using assms Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero by blast qed lemma pow_res_of_zero: shows "pow_res n \ = {\}" unfolding pow_res_def apply(rule equalityI) apply(rule subsetI) unfolding mem_Collect_eq apply (metis Qp.integral_iff Qp.nat_pow_closed Qp.nonzero_closed Qp.zero_closed insertCI) apply(rule subsetI) unfolding mem_Collect_eq by (metis Qp.nat_pow_one Qp.one_nonzero Qp.r_one Qp.zero_closed equals0D insertE) lemma equal_pow_resI: assumes "x \ carrier Q\<^sub>p" assumes "y \ pow_res n x" shows "pow_res n x = pow_res n y" proof have y_closed: "y \ carrier Q\<^sub>p" using assms unfolding pow_res_def by blast obtain c where c_def: "c \ nonzero Q\<^sub>p \ y = x \ (c[^]n)" using assms pow_res_def by blast have "((inv c)[^]n) = inv (c[^]n)" using c_def Qp.field_axioms Qp.nat_pow_of_inv Units_eq_nonzero by blast then have "y \ ((inv c)[^]n) = x" using y_closed c_def assms Qp.inv_cancelL(2) Qp.nonzero_closed Qp_nat_pow_nonzero Units_eq_nonzero by presburger then have P0: "(inv c) \ nonzero Q\<^sub>p \ x =y \ ((inv c)[^]n) " using c_def nonzero_inverse_Qp by blast show "pow_res n x \ pow_res n y" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = x \ (b[^]n)" using A pow_res_def by blast then have 0: "b \ nonzero Q\<^sub>p \ a = y \ ((inv c)[^]n) \ (b[^]n)" using \y \ inv c [^] n = x\ by blast have "b \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" proof- have "(inv c)[^]n \ (b[^]n) = ((inv c) \ b)[^]n" using c_def b_def assms P0 Qp.nat_pow_distrib Qp.nonzero_closed by presburger then have " y \ (((inv c)[^]n) \ (b[^]n)) = y \ (((inv c) \ b)[^]n)" by presburger then show ?thesis using y_closed 0 P0 Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) by presburger qed then have "((inv c) \ b) \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" by (metis P0 Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed not_nonzero_Qp) then show "a \ pow_res n y" using pow_res_def \a \ carrier Q\<^sub>p\ by blast qed show "pow_res n y \ pow_res n x" proof fix a assume A: "a \ pow_res n y" then have 0: "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = y \ (b[^]n)" using A pow_res_def by blast then have "a = (x \ (c[^]n)) \ (b[^]n)" using c_def by blast then have "a = x \ ((c[^]n) \ (b[^]n))" by (meson Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) b_def c_def) then have "a = x \ ((c \ b)[^]n)" using Qp.nat_pow_distrib Qp.nonzero_closed b_def c_def by presburger then have "(c \ b) \ nonzero Q\<^sub>p \ a = x \ ((c \ b)[^]n)" by (metis Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed b_def c_def not_nonzero_Qp) then show "a \ pow_res n x" using pow_res_def 0 by blast qed qed lemma zeroth_pow_res: assumes "x \ carrier Q\<^sub>p" shows "pow_res 0 x = {x}" apply(rule equalityI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms apply (metis Qp.nat_pow_0 Qp.r_one singletonI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms by (metis Qp.nat_pow_0 Qp.one_nonzero Qp.r_one equals0D insertE) lemma Zp_car_zero_res: assumes "x \ carrier Z\<^sub>p" shows "x 0 = 0" using assms unfolding Zp_def using Zp_def Zp_defs(3) padic_set_zero_res prime by blast lemma zeroth_ac: assumes "x \ carrier Q\<^sub>p" shows "ac 0 x = 0" apply(cases "x = \ ") unfolding ac_def apply presburger using assms angular_component_closed[of x] Zp_car_zero_res unfolding nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_imp_nonzero: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "x \ nonzero Q\<^sub>p" using assms unfolding ac_def nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_val_ord: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "val x = ord x" using nonzero_ac_imp_nonzero assms val_ord by blast lemma pow_res_equal_ord: assumes "n > 0" shows "\m > 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof- obtain m where m_def_0: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n)" using assms nth_power_fact[of n] by (metis less_imp_le_nat less_one linorder_neqE_nat nat_le_linear zero_less_iff_neq_zero) then have m_def: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ ord u = 0 \ u \ P_set n)" by (smt nonzero_ac_val_ord zero_eint_def) have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y " then have 0: "ac m (x \
y) = 1" using ac_inv[of y m] ac_mult by (metis ac_inv'''(1) ac_mult' m_def nonzero_inverse_Qp) have 1: "ord (x \
y) = 0" using A ord_fract by presburger have 2: "(x \
y) \ nonzero Q\<^sub>p" using A by (metis Qp.nonzero_closed Qp.nonzero_mult_closed local.fract_cancel_right nonzero_inverse_Qp not_nonzero_Qp zero_fract) have 3: "(x \
y) \ P_set n" using m_def 0 1 2 nonzero_def by (smt Qp.nonzero_closed) then obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^]n) = (x \
y)" using P_set_def by blast then have "(x \
y) \ y = (b[^]n) \ y" by presburger then have "x = (b[^]n) \ y" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_left) then have "x = y \(b[^]n)" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_right) then have "x \ pow_res n y" unfolding pow_res_def using A b_def by (metis (mono_tags, lifting) "2" Qp.nat_pow_0 Qp.nonzero_closed Qp_nonzero_nat_pow mem_Collect_eq not_gr_zero) then show "pow_res n x = pow_res n y" using A equal_pow_resI[of x y n] unfolding nonzero_def by (metis (mono_tags, lifting) A Qp.nonzero_closed equal_pow_resI) qed qed qed then show ?thesis using m_def by blast qed lemma pow_res_equal: assumes "n > 0" shows "\m> 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = (ord y mod n) \ pow_res n x = pow_res n y" proof- obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms pow_res_equal_ord by meson have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n" show "pow_res n x = pow_res n y" proof- have A0: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p" using A by blast have A1: "ac m x = ac m y" using A by blast have A2: "ord x = ord y mod int n" using A by blast obtain k where k_def: "k = ord x" by blast obtain l where l_def: "ord y = ord x + (l:: int)*(int n)" using assms A2 by (metis A k_def mod_eqE mod_mod_trivial mult_of_nat_commute) have m_def': "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" using m_def by blast have 0: "ord (y \ (\[^](- l*n))) = ord x" proof- have 0: "ord (y \ (\[^](- l*n))) = ord y + (ord (\[^](- l*n)))" using ord_mult p_nonzero A0 Qp_int_pow_nonzero by blast have 1: "ord (\[^](- l*n)) = - l*n" using ord_p_pow_int[of "-l*n"] by blast then have "ord (y \ (\[^](- l*n))) = ord y - l*n" using 0 by linarith then show ?thesis using k_def l_def by linarith qed have 1: "ac m (y \ (\[^](- l*n))) = ac m y" using assms ac_p_int_pow_factor_right[of ] m_def A Qp.nonzero_closed by presburger have 2: "y \ (\[^](- l*n)) \ nonzero Q\<^sub>p" using A0 Qp_int_pow_nonzero[of \ "- l*n"] Qp.cring_axioms nonzero_def cring.cring_simprules(5) fract_cancel_left not_nonzero_Qp p_intpow_inv'' p_nonzero zero_fract Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_memI Qp.nonzero_mult_closed minus_mult_commute mult_minus_right p_intpow_closed(1) p_intpow_closed(2) by presburger then have 3: "pow_res n (y \ (\[^](- l*n))) = pow_res n x" using 2 A0 m_def'[of "y \ (\[^](- l*n))" x] "0" "1" A1 by linarith have 4: "(y \ (\[^](- l*n))) = (y \ ((\[^]- l)[^]n))" using Qp_int_nat_pow_pow[of \ "-l" n] p_nonzero by presburger have "y \ (\[^](- l*n))\ pow_res n y " using "2" "4" Qp_int_pow_nonzero nonzero_def p_nonzero unfolding pow_res_def nonzero_def proof - assume a1: "\x n. x \ {a \ carrier Q\<^sub>p. a \ \} \ x [^] (n::int) \ {a \ carrier Q\<^sub>p. a \ \}" assume a2: "\ \ {a \ carrier Q\<^sub>p. a \ \}" assume a3: "y \ \ [^] (- l * int n) \ {a \ carrier Q\<^sub>p. a \ \}" have f4: "\ [^] (- 1 * l) \ {r \ carrier Q\<^sub>p. r \ \}" using a2 a1 by presburger have f5: "- l = - 1 * l" by linarith then have f6: "y \ \ [^] (- 1 * l * int n) = y \ (\ [^] (- 1 * l)) [^] n" using \y \ \ [^] (- l * int n) = y \ (\ [^] - l) [^] n\ by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. r \ \}" using f5 a3 by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f4 by blast then have "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f6 f5 by presburger then show "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra\{r \ carrier Q\<^sub>p. r \ \}. r = y \ ra [^] n}" by meson qed then have "pow_res n (y \ (\[^](- l*n))) = pow_res n y" using equal_pow_resI[of "(y \ (\[^](- l*n)))" y n] "2" A0 assms Qp.nonzero_mult_closed p_intpow_closed(2) by (metis (mono_tags, opaque_lifting) "3" A Qp.nonzero_closed equal_pow_resI) then show ?thesis using 3 by blast qed qed qed qed then show ?thesis using m_def by blast qed definition pow_res_classes where "pow_res_classes n = {S. \x \ nonzero Q\<^sub>p. S = pow_res n x }" lemma pow_res_semialg_def: assumes "x \ nonzero Q\<^sub>p" assumes "n \ 1" shows "\P \ carrier Q\<^sub>p_x. pow_res n x = (univ_basic_semialg_set n P) - {\}" proof- have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (inv x) \ a = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))" unfolding pow_res_def by blast then obtain y where y_def: "y \ nonzero Q\<^sub>p \a = x \ (y[^]n)" by blast then have "y \ nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" proof - show ?thesis by (metis (no_types, opaque_lifting) Qp.m_assoc Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ y [^] n)\ assms(1) local.fract_cancel_right nonzero_inverse_Qp y_def) qed then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" using assms \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))\ by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)} \ pow_res n x" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" show "a \ pow_res n x" proof- have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using A by blast then obtain y where y_def: "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ a = x \(y[^]n)" by (metis Qp.l_one Qp.m_assoc Qp.nonzero_closed Qp.not_nonzero_memI \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = y [^] n)\ assms(1) field_inv(2) inv_in_frac(1)) then show ?thesis by (metis (mono_tags, lifting) \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))\ mem_Collect_eq pow_res_def) qed qed qed obtain P where P_def: "P = up_ring.monom Q\<^sub>p_x (inv x) 1" by blast have P_closed: "P \ carrier Q\<^sub>p_x" using P_def Qp.nonzero_closed Qp.nonzero_memE(2) UPQ.is_UP_monomE(1) UPQ.is_UP_monomI assms(1) inv_in_frac(1) by presburger have P_eval: "\a. a \ carrier Q\<^sub>p \ (P \ a) = (inv x) \ a" using P_def to_fun_monom[of ] by (metis Qp.nat_pow_eone Qp.nonzero_closed Qp.not_nonzero_memI assms(1) inv_in_frac(1)) have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using 0 by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using P_eval by (metis (mono_tags, lifting) mem_Collect_eq) qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ pow_res n x" proof fix a assume "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" using P_eval \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then show "a \ pow_res n x" using 0 by blast qed qed have 1: "univ_basic_semialg_set n P - {\} = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "univ_basic_semialg_set n P - {\} \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume A: "a \ univ_basic_semialg_set n P - {\}" then have A0: "a \ carrier Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" unfolding univ_basic_semialg_set_def by blast then have A0': "a \ nonzero Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" using A by (metis DiffD2 not_nonzero_Qp singletonI) then obtain y where y_def: "y\carrier Q\<^sub>p \ P \ a = (y[^]n)" by blast have A1: "(P \ a) \ \" using P_eval A0' Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) assms(1) inv_in_frac(1) inv_in_frac(2) by presburger have A2: "y \ nonzero Q\<^sub>p" proof- have A20: "(y[^]n) \\" using A1 y_def by blast have "y \ \" apply(rule ccontr) using A20 assms by (metis Qp.nat_pow_eone Qp.semiring_axioms Qp.zero_closed le_zero_eq semiring.nat_pow_zero) then show ?thesis using y_def A1 not_nonzero_Qp Qp.not_nonzero_memE by blast qed then have "y \ nonzero Q\<^sub>p \ P \ a = (y[^]n)" using y_def by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using A0 nonzero_def by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ univ_basic_semialg_set n P - {\}" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y \\ \ y\ carrier Q\<^sub>p \ P \ a = (y[^]n)" by (metis (mono_tags, opaque_lifting) Qp.nonzero_closed Qp.not_nonzero_memI) then have "a \\" using P_eval by (metis Qp.m_comm Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.zero_closed assms(1) inv_in_frac(1) zero_fract) then show "a \ univ_basic_semialg_set n P - {\}" unfolding univ_basic_semialg_set_def using A \y \ \ \ y \ carrier Q\<^sub>p \ P \ a = (y[^]n)\ by blast qed qed show ?thesis using 0 1 by (metis (no_types, lifting) P_closed) qed lemma pow_res_is_univ_semialgebraic: assumes "x \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (pow_res n x)" proof(cases "n = 0") case True have T0: "pow_res n x = {x}" unfolding True using assms by (simp add: assms zeroth_pow_res) have "[x] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms Qp.to_R1_closed by blast hence "is_semialgebraic 1 {[x]}" using is_algebraic_imp_is_semialg singleton_is_algebraic by blast thus ?thesis unfolding T0 using assms by (simp add: \x \ carrier Q\<^sub>p\ finite_is_univ_semialgebraic) next case False show ?thesis proof(cases "x = \") case True then show ?thesis using finite_is_univ_semialgebraic False pow_res_of_zero by (metis Qp.zero_closed empty_subsetI finite.emptyI finite.insertI insert_subset) next case F: False then show ?thesis using False pow_res_semialg_def[of x n] diff_is_univ_semialgebraic[of _ "{\}"] finite_is_univ_semialgebraic[of "{\}"] by (metis Qp.zero_closed assms empty_subsetI finite.emptyI finite.insertI insert_subset less_one less_or_eq_imp_le linorder_neqE_nat not_nonzero_Qp univ_basic_semialg_set_is_univ_semialgebraic) qed qed lemma pow_res_is_semialg: assumes "x \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1 ` (pow_res n x))" using assms pow_res_is_univ_semialgebraic is_univ_semialgebraicE by blast lemma pow_res_refl: assumes "x \ carrier Q\<^sub>p" shows "x \ pow_res n x" proof- have "x = x \ (\ [^]n)" using assms Qp.nat_pow_one Qp.r_one by presburger thus ?thesis using assms unfolding pow_res_def mem_Collect_eq using Qp.one_nonzero by blast qed lemma equal_pow_resE: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "n > 0" assumes "pow_res n a = pow_res n b" shows "\ s \ P_set n. a = b \ s" proof- have "a \ pow_res n b" using assms pow_res_refl by blast then obtain y where y_def: " y \ nonzero Q\<^sub>p \ a = b \ y[^]n" unfolding pow_res_def by blast thus ?thesis unfolding P_set_def using Qp.nonzero_closed Qp_nat_pow_nonzero by blast qed lemma pow_res_one: assumes "x \ nonzero Q\<^sub>p" shows "pow_res 1 x = nonzero Q\<^sub>p" proof show "pow_res 1 x \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of x 1] by blast show "nonzero Q\<^sub>p \ pow_res 1 x" proof fix y assume A: "y \ nonzero Q\<^sub>p" then have 0: "\ \ nonzero Q\<^sub>p \ y = x \ ((inv x)\ y)[^](1::nat)" using assms Qp.m_comm Qp.nat_pow_eone Qp.nonzero_closed Qp.nonzero_mult_closed Qp.one_nonzero local.fract_cancel_right nonzero_inverse_Qp by presburger have 1: "(inv x)\ y \ nonzero Q\<^sub>p" using A assms by (metis Qp.Units_m_closed Units_eq_nonzero nonzero_inverse_Qp) then show "y \ pow_res 1 x" unfolding pow_res_def using 0 1 A Qp.nonzero_closed by blast qed qed lemma pow_res_zero: assumes "n > 0" shows "pow_res n \ = {\}" proof show "pow_res n \ \ {\}" unfolding pow_res_def using Qp.l_null Qp.nat_pow_closed Qp.nonzero_closed by blast show "{\} \ pow_res n \" using assms unfolding pow_res_def using Qp.l_null Qp.one_closed Qp.one_nonzero empty_subsetI insert_subset by blast qed lemma equal_pow_resI': assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ P_set n" assumes "a = b \ c" assumes "n > 0" shows "pow_res n a = pow_res n b" proof- obtain y where y_def: "c = y[^]n \ y \ carrier Q\<^sub>p" using assms unfolding P_set_def by blast have c_nonzero: "c \ nonzero Q\<^sub>p" using P_set_nonzero'(1) assms(3) by blast have y_nonzero: "y \ nonzero Q\<^sub>p" using y_def c_nonzero Qp_nonzero_nat_pow assms(5) by blast have 0: "a \ pow_res n b" using assms y_nonzero y_def unfolding pow_res_def by blast show ?thesis apply(cases "b = \") using pow_res_zero Qp.l_null Qp.nonzero_closed assms(4) c_nonzero apply presburger by (metis "0" assms(1) assms(2) assms(5) not_nonzero_Qp equal_pow_resI) qed lemma equal_pow_resI'': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "a \ inv b \ P_set n" shows "pow_res n a = pow_res n b" using assms equal_pow_resI'[of a b "a \ inv b" n] Qp.nonzero_closed local.fract_cancel_right by blast lemma equal_pow_resI''': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "c \ nonzero Q\<^sub>p" assumes "pow_res n (c \ a) = pow_res n (c \ b)" shows "pow_res n a = pow_res n b" proof- have 0: "c \a \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) assms(4)) have 1: "c \b \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(3) assms(4)) have "c\a \ pow_res n (c\b)" proof(cases "n = 1") case True then show ?thesis using assms 0 1 pow_res_one[of "c\b"] by blast next case False then have "n \ 2" using assms(1) by linarith then show ?thesis using assms 0 1 pow_res_refl[of "c\a" n] unfolding nonzero_def by blast qed then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (c \ a) = (c \ b)\y[^]n" using assms unfolding pow_res_def by blast then have "a = b\y[^]n" using assms by (metis Qp.m_assoc Qp.m_lcancel Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero) then show ?thesis by (metis P_set_memI Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) assms(3) equal_pow_resI' y_def) qed lemma equal_pow_resI'''': assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "a = b \ u" assumes "u \ P_set n" shows "pow_res n a = pow_res n b" proof(cases "a = \") case True then have "b = \" using assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.not_nonzero_memI mem_Collect_eq) then show ?thesis using pow_res_zero using True by blast next case False then have 0: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast have 1: "b \ nonzero Q\<^sub>p" using 0 assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral_iff Qp.nonzero_closed mem_Collect_eq not_nonzero_Qp) have 2: "a \ (inv b)\ P_set n" using assms 0 1 by (metis P_set_nonzero'(2) Qp.inv_cancelR(1) Qp.m_comm Qp.nonzero_memE(2) Units_eq_nonzero inv_in_frac(1)) then show ?thesis using equal_pow_resI'' by (meson "0" "1" assms(1) equal_pow_resI) qed lemma Zp_Units_ord_zero: assumes "a \ Units Z\<^sub>p" shows "ord_Zp a = 0" proof- have "inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p" apply(rule Zp.nonzero_memI, rule Zp.Units_inv_closed, rule assms) using assms Zp.Units_inverse in_Units_imp_not_zero by blast then have "ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub> a) = ord_Zp a + ord_Zp (inv \<^bsub>Z\<^sub>p\<^esub> a)" using assms ord_Zp_mult Zp.Units_nonzero zero_not_one by (metis Zp.zero_not_one) then show ?thesis by (smt Zp.Units_closed Zp.Units_r_inv Zp.integral_iff Zp.nonzero_closed \inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p\ assms ord_Zp_one ord_pos) qed lemma pow_res_nth_pow: assumes "a \ nonzero Q\<^sub>p" assumes "n > 0" shows "pow_res n (a[^]n) = pow_res n \" proof show "pow_res n (a [^] n) \ pow_res n \" proof fix x assume A: "x \ pow_res n (a [^] n)" then show "x \ pow_res n \" by (metis P_set_memI Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.one_closed assms(1) assms(2) equal_pow_resI') qed show "pow_res n \ \ pow_res n (a [^] n)" proof fix x assume A: "x \ pow_res n \" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ x = \\y[^]n" unfolding pow_res_def by blast then have 0: "x = y[^]n" using Qp.l_one Qp.nonzero_closed by blast have "y[^]n = a[^]n \ (inv a \ y)[^]n" proof- have "a[^]n \ (inv a \ y)[^]n = a[^]n \ (inv a)[^]n \ y[^]n" using Qp.Units_inv_closed Qp.m_assoc Qp.nat_pow_closed Qp.nat_pow_distrib Qp.nonzero_closed Units_eq_nonzero assms(1) y_def by presburger then show ?thesis by (metis Qp.Units_inv_inv Qp.inv_cancelR(1) Qp.nat_pow_distrib Qp.nonzero_closed Qp.nonzero_mult_closed Units_eq_nonzero assms(1) nonzero_inverse_Qp y_def) qed then show "x \ pow_res n (a [^] n)" using y_def A assms unfolding pow_res_def mem_Collect_eq by (metis "0" Qp.integral Qp.m_closed Qp.nonzero_closed Qp.not_nonzero_memI inv_in_frac(1) inv_in_frac(2) not_nonzero_Qp) qed qed lemma pow_res_of_p_pow: assumes "n > 0" shows "pow_res n (\[^]((l::int)*n)) = pow_res n \" proof- have 0: "\[^]((l::int)*n) = (\[^]l)[^]n" using Qp_p_int_nat_pow_pow by blast have "\[^]((l::int)*n) \ P_set n" using P_set_memI[of _ "\[^]l"] by (metis "0" Qp.not_nonzero_memI Qp_int_pow_nonzero p_intpow_closed(1) p_nonzero) thus ?thesis using "0" assms p_intpow_closed(2) pow_res_nth_pow by presburger qed lemma pow_res_nonzero: assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n b" shows "b \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of a n] pow_res_zero[of n] by (metis insert_subset not_nonzero_Qp) lemma pow_res_mult: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "d \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n c" assumes "pow_res n b = pow_res n d" shows "pow_res n (a \ b) = pow_res n (c \ d)" proof(cases "a \ nonzero Q\<^sub>p") case True then have "c \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where alpha_def: "\ \ nonzero Q\<^sub>p \ a = c \ \[^]n" using assms True pow_res_refl[of a n] unfolding assms unfolding pow_res_def by blast show ?thesis proof(cases "b \ nonzero Q\<^sub>p") case T: True then have "d \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where beta_def: "\ \ nonzero Q\<^sub>p \ b = d \ \[^]n" using T pow_res_refl[of b n] unfolding assms unfolding pow_res_def using assms by blast then have "a \ b = (c \ d) \ (\[^]n \ \[^] n)" using Qp.m_assoc Qp.m_lcomm Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero alpha_def assms(3) assms(4) assms(5) by presburger then have 0: "a \ b = (c \ d) \ ((\ \ \)[^] n)" by (metis Qp.nat_pow_distrib Qp.nonzero_closed alpha_def beta_def) show ?thesis apply(intro equal_pow_resI'[of _ _ "(\ \ \)[^] n"] Qp.ring_simprules assms P_set_memI[of _ "\ \ \"] Qp.nat_pow_closed nonzero_memE 0 Qp_nat_pow_nonzero ) using alpha_def beta_def apply auto apply(intro nonzero_memI Qp.nonzero_mult_closed) using alpha_def beta_def nonzero_memE apply auto by (meson Qp.integral_iff) next case False then have "b = \" by (meson assms not_nonzero_Qp) then have "d = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \b = \\ assms by presburger qed next case False then have "a = \" by (meson assms not_nonzero_Qp) then have "c = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \a = \\ assms Qp.l_null by presburger qed lemma pow_res_p_pow_factor: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" shows "pow_res n a = pow_res n (\[^]((l::int)*n) \ a)" proof(cases "a = \") case True then show ?thesis using Qp.r_null p_intpow_closed(1) by presburger next case False then show ?thesis using assms pow_res_of_p_pow using Qp.m_comm Qp.one_closed Qp.r_one p_intpow_closed(1) pow_res_mult by presburger qed lemma pow_res_classes_finite: assumes "n \ 1" shows "finite (pow_res_classes n)" proof(cases "n = 1") case True have "pow_res_classes n = {(nonzero Q\<^sub>p)}" using True pow_res_one unfolding pow_res_classes_def using Qp.one_nonzero by blast then show ?thesis by auto next case False then have n_bound: "n \ 2" using assms by linarith obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms False pow_res_equal_ord n_bound by (metis gr_zeroI le_numeral_extra(2)) obtain f where f_def: "f = (\ \ \. (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)))" by blast have 0: "\x. x \ nonzero Q\<^sub>p \ pow_res n x = f (ac m x) (ord x)" proof- fix x assume A: "x \ nonzero Q\<^sub>p" obtain \ where eta_def: "\ = ac m x" by blast obtain \ where nu_def: "\ = ord x" by blast have "\y \pow_res n x. ac m y = ac m x \ ord y = ord x" using pow_res_refl A assms neq0_conv Qp.nonzero_closed by blast hence "pow_res n x \ (pow_res_classes n) \ (\ y \ (pow_res n x). ac m y = \ \ ord y = \)" unfolding nu_def eta_def using assms unfolding pow_res_classes_def using A by blast then have 0: "(\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" by blast have "f \ \ = (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" using f_def by blast then have 1: "f \ \ \ (pow_res_classes n) \ ((\ y \ (f \ \). ac m y = \ \ ord y = \))" using 0 someI_ex[of "\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)"] unfolding f_def by blast then obtain y where y_def: "y \ (f \ \) \ ac m y = ac m x \ ord y = ord x" unfolding nu_def eta_def by blast obtain a where a_def: "a \ nonzero Q\<^sub>p \ (f \ \) = pow_res n a" using 1 unfolding pow_res_classes_def by blast then have 2: "y \ pow_res n a" using y_def by blast have 3: "y \ nonzero Q\<^sub>p" using y_def nonzero_pow_res[of a n] a_def by blast then have 4: "pow_res n y = pow_res n a" using 3 y_def a_def equal_pow_resI[of y a n] n_bound Qp.nonzero_closed by (metis equal_pow_resI) have 5: "pow_res n y = f \ \" using 4 a_def by blast then show "pow_res n x = f (ac m x) (ord x)" unfolding eta_def nu_def using "3" A m_def y_def by blast qed obtain N where N_def: "N > 0 \ (\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n)" using n_bound nth_power_fact assms by blast have 1: "\x. x \ nonzero Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. ord y \ 0 \ ord y < n \ pow_res n x = pow_res n y)" proof- fix x assume x_def: "x \ nonzero Q\<^sub>p" then obtain k where k_def: "k = ord x mod n" by blast then obtain l where l_def: "ord x = (int n)*l + k" using cancel_div_mod_rules(2)[of n "ord x"0] unfolding k_def by (metis group_add_class.add.right_cancel) have "x = (\[^](ord x)) \ \ (angular_component x)" using x_def angular_component_factors_x by blast then have "x = (\[^](n*l + k)) \ \ (angular_component x)" unfolding l_def by blast hence "x = \[^](int n*l) \ (\[^] k) \ \ (angular_component x)" by (metis p_intpow_add) hence 0: "x = (\[^]l)[^]n \ (\[^] k) \ \ (angular_component x)" using p_pow_factor[of n l k] \x = \ [^] (int n * l + k) \ \ (angular_component x)\ by presburger have 1: "\ (angular_component x) \ carrier Q\<^sub>p" using x_def angular_component_closed inc_closed by blast hence 2: "x = (\[^]l)[^]n \ ((\[^] k) \ \ (angular_component x))" using 0 by (metis Qp.m_assoc Qp.nat_pow_closed p_intpow_closed(1)) obtain a where a_def: "a = (\[^] k) \ \ (angular_component x)" by blast have 30: "angular_component x \ Units Z\<^sub>p" using angular_component_unit x_def by blast then have 3: "\ (angular_component x) \ Units Q\<^sub>p" by (metis Units_eq_nonzero Zp.Units_closed in_Units_imp_not_zero inc_of_nonzero not_nonzero_Qp) have 4: "\ (angular_component x) \ nonzero Q\<^sub>p" using 3 Units_nonzero_Qp by blast have a_nonzero: "a \ nonzero Q\<^sub>p" unfolding a_def 4 by (meson "3" Qp.UnitsI(1) Qp.Units_m_closed Units_nonzero_Qp p_intpow_closed(1) p_intpow_inv) have 5: "x = a \(\[^]l)[^]n" using 2 a_nonzero unfolding a_def using Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed p_intpow_closed(1) by presburger hence "x \ pow_res n a" unfolding pow_res_def using Qp.nonzero_closed Qp_int_pow_nonzero p_nonzero x_def by blast hence 6:"pow_res n a = pow_res n x" using x_def a_def equal_pow_resI[of x a n] a_nonzero n_bound Qp.nonzero_closed equal_pow_resI by blast have 7: "ord (\ (angular_component x)) = 0" proof- have "ord_Zp (angular_component x) = 0" using 30 Zp_Units_ord_zero by blast then have "val_Zp (angular_component x) = 0" using "30" unit_imp_val_Zp0 by blast then have "val (\ (angular_component x)) = 0" by (metis angular_component_closed val_of_inc x_def) then show ?thesis using angular_component_closed x_def by (metis "30" Zp.Units_closed \ord_Zp (angular_component x) = 0\ in_Units_imp_not_zero not_nonzero_Qp ord_of_inc) qed have 8: "ord a = k" unfolding a_def using 3 4 7 ord_mult[of "\ [^] k" "\ (angular_component x)"] ord_p_pow_int[of k] p_pow_nonzero using Qp_int_pow_nonzero p_nonzero by presburger have 9: "k < n" unfolding k_def using assms by auto from 6 8 9 assms have \0 \ ord a\ \ord a < int n\ \pow_res n x = pow_res n a\ by (auto simp add: k_def) with a_nonzero show "\y\nonzero Q\<^sub>p. 0 \ ord y \ ord y < int n \ pow_res n x = pow_res n y" by auto qed have 2: "\x. x \ (pow_res_classes n) \ \ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" proof- fix a assume A: "a \ (pow_res_classes n)" then obtain x where x_def: "x \ nonzero Q\<^sub>p \ a = pow_res n x" unfolding pow_res_classes_def by blast then obtain x' where x'_def: "x' \ nonzero Q\<^sub>p \ ord x' \ 0 \ ord x' < n \ pow_res n x' = a" using 1[of x] unfolding x_def by blast hence 20: "f (ac m x') (ord x') = a" using 0 by blast have 21: "ac m x' \ Units (Zp_res_ring m)" using x'_def ac_units m_def by presburger then have 22: "ac m x' \ Units (Zp_res_ring m) \ (ord x') \ ({0.. a = f (ac m x') (ord x')" using x'_def 20 atLeastLessThan_iff by blast then show "\ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. a = f \ \" by blast qed obtain F where F_def: "F = (\ps. f (fst ps) (snd ps))" by blast have 3: "\x. x \ (pow_res_classes n) \ \ ps \ Units (Zp_res_ring m) \ {0.. pow_res_classes n" obtain \ \ where eta_nu_def: " \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" using 2 A by blast then have "F (\, \) = x" unfolding F_def by (metis fst_conv snd_conv) then show " \ ps \ Units (Zp_res_ring m) \ {0.. F ` (Units (Zp_res_ring m) \ {0.. {0.. pow_res_classes n" shows "is_univ_semialgebraic S" proof- obtain x where x_def: "x\nonzero Q\<^sub>p \ S = pow_res n x" using assms unfolding pow_res_classes_def by blast then show ?thesis using pow_res_is_univ_semialgebraic using Qp.nonzero_closed by blast qed lemma pow_res_classes_semialg: assumes "S \ pow_res_classes n" shows "is_semialgebraic 1 (to_R1` S)" using pow_res_classes_univ_semialg assms(1) is_univ_semialgebraicE by blast definition nth_pow_wits where "nth_pow_wits n = (\ S. (SOME x. x \ (S \ \\<^sub>p)))` (pow_res_classes n)" lemma nth_pow_wits_finite: assumes "n > 0" shows "finite (nth_pow_wits n)" proof- have "n \ 1" by (simp add: assms leI) thus ?thesis unfolding nth_pow_wits_def using assms pow_res_classes_finite[of n] by blast qed lemma nth_pow_wits_covers: assumes "n > 0" assumes "x \ nonzero Q\<^sub>p" shows "\y \ (nth_pow_wits n). y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" proof- have PP: "(pow_res n x) \ pow_res_classes n" unfolding pow_res_classes_def using assms by blast obtain k where k_def: "val x = eint k" using assms val_ord by blast obtain N::int where N_def: "N = (if k < 0 then -k else k)" by blast then have N_nonneg: "N \ 0" unfolding N_def by presburger have 0: "int n \ 1" using assms by linarith have "N*(int n) + k \ 0" proof(cases "k<0") case True then have "N = -k" unfolding N_def by presburger then have "N*n + k = k*(1- int n)" using distrib_left[of k 1 "-int n"] mult_cancel_left2 mult_minus_left by (metis add.inverse_inverse diff_minus_eq_add minus_mult_minus neg_equal_iff_equal uminus_add_conv_diff) then show ?thesis using 0 True zero_less_mult_iff[of k "1 - int n"] proof - have "0 \ N * (int n - 1)" by (meson "0" N_nonneg diff_ge_0_iff_ge zero_le_mult_iff) then show ?thesis by (metis (no_types) \N = - k\ add.commute distrib_left minus_add_cancel mult_minus1_right uminus_add_conv_diff) qed next case False then have "N = k" unfolding N_def by presburger then show ?thesis using 0 False by (metis N_nonneg add_increasing2 mult_nonneg_nonneg of_nat_0_le_iff) qed then have 1: "ord (\[^](N*n)\x) \ 0" using ord_mult k_def val_ord assms by (metis Qp_int_pow_nonzero eint.inject ord_p_pow_int p_nonzero) have 2: "\[^](N*n)\x \ pow_res n x" proof- have "\[^](N*n) = (\[^]N)[^]n" using Qp_p_int_nat_pow_pow by blast then have "\[^]N \ nonzero Q\<^sub>p \ \[^](N*n)\x = x \ (\[^]N)[^]n" by (metis Qp.m_comm Qp.nonzero_closed Qp_int_pow_nonzero assms(2) p_nonzero) then show ?thesis unfolding pow_res_def by (metis (mono_tags, lifting) Qp.m_closed Qp.nonzero_closed assms(2) mem_Collect_eq p_intpow_closed(1)) qed have 3: "\[^](N*n)\x \ \\<^sub>p" using 1 assms by (metis Q\<^sub>p_def Qp.nonzero_mult_closed Qp_int_pow_nonzero Z\<^sub>p_def val_ring_ord_criterion \_def p_nonzero padic_fields.zero_in_val_ring padic_fields_axioms) have 4: "x \ pow_res n (\[^](N*n)\x)" using 2 equal_pow_resI[of x "\[^](N*n)\x" n] pow_res_refl[of "\[^](N*n)\x" n] assms Qp.nonzero_mult_closed p_intpow_closed(2) pow_res_refl Qp.nonzero_closed by metis have 5: "\[^](N*n)\x \ (pow_res n x \ \\<^sub>p)" using 2 3 by blast have 6: "(SOME z. z \ (pow_res n x) \ \\<^sub>p) \ pow_res n x \ \\<^sub>p" using 5 by (meson someI) obtain y where y_def: "y = (SOME z. z \ (pow_res n x) \ \\<^sub>p)" by blast then have A: "y \ pow_res n x" using "6" by blast then have "pow_res n x = pow_res n y" using equal_pow_resI[of x y n] assms y_def Qp.nonzero_closed nonzero_pow_res by blast then have 7: "x \ pow_res n y" using pow_res_refl[of x n] assms unfolding nonzero_def by blast have 8: "y \ nonzero Q\<^sub>p " using y_def PP 6 A nonzero_pow_res[of x n] assms by blast have 9: "y \ \\<^sub>p" using y_def "6" by blast have "y\(\S. SOME x. x \ S \ \\<^sub>p) ` pow_res_classes n \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" using y_def PP 6 7 8 9 A nonzero_pow_res[of x n] assms by blast then show ?thesis unfolding nth_pow_wits_def by blast qed lemma nth_pow_wits_closed: assumes "n > 0" assumes "x \ nth_pow_wits n" shows "x \ carrier Q\<^sub>p" "x \ \\<^sub>p" "x \ nonzero Q\<^sub>p" "\ y \ pow_res_classes n. y = pow_res n x" proof- obtain c where c_def: "c \ pow_res_classes n \ x = (SOME x. x \ (c \ \\<^sub>p))" by (metis (no_types, lifting) assms(2) image_iff nth_pow_wits_def) then obtain y where y_def: "y \ nonzero Q\<^sub>p \ c = pow_res n y" unfolding pow_res_classes_def by blast then obtain a where a_def: "a \ (nth_pow_wits n) \ a \ nonzero Q\<^sub>p \ a \ \\<^sub>p \ y \ pow_res n a" using nth_pow_wits_covers[of n y] assms(1) by blast have 00: "pow_res n a = c" using equal_pow_resI[of a y n] y_def assms a_def unfolding nonzero_def by blast then have P :"a \ c \ \\<^sub>p" using pow_res_refl[of a n] assms a_def unfolding 00 nonzero_def by blast then show 0: "x \ \\<^sub>p" using c_def by (metis Collect_mem_eq Int_Collect tfl_some) then show "x \ carrier Q\<^sub>p" using val_ring_memE by blast have 1: "c \ nonzero Q\<^sub>p" using c_def nonzero_pow_res[of y n] unfolding pow_res_classes_def using assms(1) y_def by blast have "(SOME x. x \ (c \ \\<^sub>p)) \ (c \ \\<^sub>p)" using P tfl_some by (smt Int_def someI_ex) then have 2: "x \ c" using c_def by blast thus "x \ nonzero Q\<^sub>p" using 1 by blast show "\ y \ pow_res_classes n. y = pow_res n x" using 00 2 c_def P a_def equal_pow_resI[of a x n] 0 val_ring_memE assms(1) by blast qed lemma finite_extensional_funcset: assumes "finite A" assumes "finite (B::'b set)" shows "finite (A \\<^sub>E B)" using finite_PiE[of A "\_. B"] assms by blast lemma nth_pow_wits_exists: assumes "m > 0" assumes "c \ pow_res_classes m" shows "\x. x \ c \ \\<^sub>p" proof- obtain x where x_def: "x \ nonzero Q\<^sub>p \ pow_res m x = c" using assms unfolding pow_res_classes_def by blast obtain y where y_def: "y \ (nth_pow_wits m) \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res m y" using nth_pow_wits_covers assms x_def by blast have 0: "pow_res m x = pow_res m y" using x_def y_def equal_pow_resI Qp.nonzero_closed assms(1) by blast then have 1: "y \ pow_res m x" using pow_res_refl[of y m ] y_def assms unfolding nonzero_def by blast thus ?thesis using x_def y_def assms by blast qed lemma pow_res_classes_mem_eq: assumes "m > 0" assumes "a \ pow_res_classes m" assumes "x \ a" shows "a = pow_res m x" proof- obtain y where y_def: "y \ nonzero Q\<^sub>p \ a = pow_res m y" using assms unfolding pow_res_classes_def by blast then show ?thesis using assms equal_pow_resI[of y x m] by (meson Qp.nonzero_closed nonzero_pow_res equal_pow_resI subsetD) qed lemma nth_pow_wits_neq_pow_res: assumes "m > 0" assumes "x \ nth_pow_wits m" assumes "y \ nth_pow_wits m" assumes "x \ y" shows "pow_res m x \ pow_res m y" proof- obtain a where a_def: "a \ pow_res_classes m \ x = (\ S. (SOME x. x \ (S \ \\<^sub>p))) a" using assms unfolding nth_pow_wits_def by blast obtain b where b_def: "b \ pow_res_classes m \ y = (\ S. (SOME x. x \ (S \ \\<^sub>p))) b" using assms unfolding nth_pow_wits_def by blast have a_neq_b: "a \ b" using assms a_def b_def by blast have 0: "x \ a \ \\<^sub>p" using a_def nth_pow_wits_exists[of m a] assms by (meson someI_ex) have 1: "y \ b \ \\<^sub>p" using b_def nth_pow_wits_exists[of m b] assms by (meson someI_ex) have 2: "pow_res m x = a" using a_def pow_res_classes_mem_eq[of m a x] assms 0 by blast have 3: "pow_res m y = b" using b_def pow_res_classes_mem_eq[of m b y] assms 1 by blast show ?thesis by (simp add: "2" "3" a_neq_b) qed lemma nth_pow_wits_disjoint_pow_res: assumes "m > 0" assumes "x \ nth_pow_wits m" assumes "y \ nth_pow_wits m" assumes "x \ y" shows "pow_res m x \ pow_res m y = {}" using assms nth_pow_wits_neq_pow_res disjoint_iff_not_equal by (metis (no_types, opaque_lifting) nth_pow_wits_closed(4) pow_res_classes_mem_eq) lemma nth_power_fact': assumes "0 < (n::nat)" shows "\m>0. \u\carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" using nth_power_fact[of n] assms by (metis less_one less_or_eq_imp_le linorder_neqE_nat neq0_conv) lemma equal_pow_res_criterion: assumes "N > 0" assumes "n > 0" assumes "\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "a = b \ (\ \ c)" assumes "val c \ N" shows "pow_res n a = pow_res n b" proof(cases "b = \") case True then have "a = \" using assms Qp.add.m_closed Qp.l_null Qp.one_closed by presburger then show ?thesis using True by blast next case False then have F0: "a \
b = \ \ c" by (metis Qp.Units_one_closed Qp.add.m_closed Qp.inv_cancelR(2) Qp.one_closed Qp.unit_factor assms(4) assms(5) assms(6) assms(7) field_inv(2) inv_in_frac(1)) have "0 < eint N" using assms by (metis eint_ord_simps(2) of_nat_0_less_iff zero_eint_def) hence F1: "val \ < val c" using assms less_le_trans[of 0 N "val c"] unfolding val_one by blast hence F2: " val \ = val (\ \ c)" using assms val_one one_nonzero Qp.add.m_comm Qp.one_closed val_ultrametric_noteq by metis have "val \ + eint (int N) \ val (\ \ (\ \ c))" proof- have "val (\ \ (\ \ c)) = val c" using Qp.add.inv_closed Qp.minus_eq Qp.minus_sum Qp.one_closed Qp.r_neg2 assms(6) val_minus by presburger thus ?thesis unfolding val_one using assms F1 by (metis add.left_neutral) qed hence F3: "ac N \ = ac N (\ \ c)" using F2 F1 assms ac_val[of \ "\ \ c" N] by (metis Qp.add.m_closed Qp.one_closed val_nonzero) have F4: "\ \ c \ P_set n" using assms F1 F2 F3 val_one ac_one by (metis Qp.add.m_closed Qp.one_closed Qp.one_nonzero ac_inv'' ac_inv'''(1) ac_one') then show ?thesis using assms(2) assms(4) assms(5) assms(7) equal_pow_resI' by blast qed lemma pow_res_nat_pow: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n b" shows "pow_res n (a[^](k::nat)) = pow_res n (b[^]k)" apply(induction k) using assms apply (metis Group.nat_pow_0) using assms pow_res_mult by (smt Qp.nat_pow_Suc2 Qp.nat_pow_closed) lemma pow_res_mult': assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "d \ carrier Q\<^sub>p" assumes "e \ carrier Q\<^sub>p" assumes "f \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n d" assumes "pow_res n b = pow_res n e" assumes "pow_res n c = pow_res n f" shows "pow_res n (a \ b \ c) = pow_res n (d \ e \ f)" proof- have "pow_res n (a \ b) = pow_res n (d \ e)" using pow_res_mult assms by meson then show ?thesis using pow_res_mult assms by (meson Qp.m_closed) qed lemma pow_res_disjoint: assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "a \ pow_res n \" shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" using assms unfolding pow_res_def using Qp.l_one Qp.nonzero_closed by blast lemma pow_res_disjoint': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "pow_res n a \ pow_res n \" shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" using assms pow_res_disjoint pow_res_refl by (metis pow_res_nth_pow) lemma pow_res_one_imp_nth_pow: assumes "n > 0" assumes "a \ pow_res n \" shows "\y \ nonzero Q\<^sub>p. a = y[^]n" using assms unfolding pow_res_def using Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed by blast lemma pow_res_eq: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ pow_res n a" shows "pow_res n b = pow_res n a" proof(cases "a = \") case True then show ?thesis using assms by (metis pow_res_zero singletonD) next case False then have a_nonzero: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast show ?thesis proof(cases "n = 1") case True then show ?thesis using a_nonzero assms using pow_res_one Q\<^sub>p_def Zp_def padic_fields_axioms by blast next case False then have "n \ 2" using assms(1) by linarith then show ?thesis using False a_nonzero assms Qp.nonzero_closed nonzero_pow_res equal_pow_resI by blast qed qed lemma pow_res_classes_n_eq_one: "pow_res_classes 1 = {nonzero Q\<^sub>p}" unfolding pow_res_classes_def using pow_res_one Qp.one_nonzero by blast lemma nth_pow_wits_closed': assumes "n > 0" assumes "x \ nth_pow_wits n" shows "x \ \\<^sub>p \ x \ nonzero Q\<^sub>p" using nth_pow_wits_closed assms by blast (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Sets Defined by Congruences\ (**************************************************************************************************) (**************************************************************************************************) (********************************************************************************************) (********************************************************************************************) subsubsection\$p$-adic ord Congruence Sets\ (********************************************************************************************) (********************************************************************************************) lemma carrier_is_univ_semialgebraic: "is_univ_semialgebraic (carrier Q\<^sub>p)" apply(rule is_univ_semialgebraicI) using Qp.to_R1_carrier carrier_is_semialgebraic by presburger lemma nonzero_is_univ_semialgebraic: "is_univ_semialgebraic (nonzero Q\<^sub>p)" proof- have "nonzero Q\<^sub>p = carrier Q\<^sub>p - {\}" unfolding nonzero_def by blast then show ?thesis using diff_is_univ_semialgebraic[of "carrier Q\<^sub>p" "{\}"] by (metis Diff_empty Diff_insert0 carrier_is_univ_semialgebraic empty_subsetI finite.emptyI finite.insertI finite_is_univ_semialgebraic insert_subset) qed definition ord_congruence_set where "ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x mod n = a}" lemma ord_congruence_set_nonzero: "ord_congruence_set n a \ nonzero Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq ord_congruence_set_def subsetI) lemma ord_congruence_set_closed: "ord_congruence_set n a \ carrier Q\<^sub>p" using nonzero_def ord_congruence_set_nonzero unfolding nonzero_def by (meson Qp.nonzero_closed ord_congruence_set_nonzero subset_iff) lemma ord_congruence_set_memE: assumes "x \ ord_congruence_set n a" shows "x \ nonzero Q\<^sub>p" "ord x mod n = a" using assms ord_congruence_set_nonzero apply blast by (metis (mono_tags, lifting) assms mem_Collect_eq ord_congruence_set_def) lemma ord_congruence_set_memI: assumes "x \ nonzero Q\<^sub>p" assumes "ord x mod n = a" shows "x \ ord_congruence_set n a" using assms by (metis (mono_tags, lifting) mem_Collect_eq ord_congruence_set_def) text\ We want to prove that ord\_congruence\_set is a finite union of semialgebraic sets, hence is also semialgebraic. \ lemma pow_res_ord_cong: assumes "x \ carrier Q\<^sub>p" assumes "x \ ord_congruence_set n a" shows "pow_res n x \ ord_congruence_set n a" proof fix y assume A: "y \ pow_res n x" show "y \ ord_congruence_set (int n) a" proof- obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]n)" using A pow_res_def[of n x] by blast have 0: "x \ nonzero Q\<^sub>p" using assms(2) ord_congruence_set_memE(1) by blast have 1: "y \ nonzero Q\<^sub>p" using A by (metis "0" Qp.integral Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero a_def not_nonzero_Qp) have 2: "ord y = ord x + n* ord a" using a_def 0 1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_mult by presburger show ?thesis apply(rule ord_congruence_set_memI) using assms ord_congruence_set_memE 2 1 apply blast using "2" assms(2) ord_congruence_set_memE(2) by presburger qed qed lemma pow_res_classes_are_univ_semialgebraic: shows "are_univ_semialgebraic (pow_res_classes n)" apply(rule are_univ_semialgebraicI) using pow_res_classes_univ_semialg by blast lemma ord_congruence_set_univ_semialg: assumes "n \ 0" shows "is_univ_semialgebraic (ord_congruence_set n a)" proof(cases "n = 0") case True have T0: "ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x = a}" unfolding ord_congruence_set_def True by presburger have T1: "{x \ nonzero Q\<^sub>p. ord x = a} = {x \ nonzero Q\<^sub>p. val x = a}" apply(rule equalityI'') using val_ord apply blast using val_ord by (metis eint.inject) have T2: "{x \ nonzero Q\<^sub>p. val x = a} = {x \ carrier Q\<^sub>p. val x = a}" apply(rule equalityI'') using Qp.nonzero_closed apply blast by (metis iless_Suc_eq val_nonzero val_val_ring_prod zero_in_val_ring) show ?thesis unfolding T0 T1 T2 using univ_val_eq_set_is_univ_semialgebraic by blast next case False obtain F where F_def: "F = {S \ (pow_res_classes (nat n)). S \(ord_congruence_set n a) }" by blast have 0: "F \ pow_res_classes (nat n)" using F_def by blast have 1: "finite F" using 0 False nat_mono[of 1 n] nat_numeral[] pow_res_classes_finite[of "nat n"] rev_finite_subset by (smt assms nat_one_as_int) have 2: "are_univ_semialgebraic F" apply(rule are_univ_semialgebraicI) using 0 pow_res_classes_are_univ_semialgebraic by (metis (mono_tags) are_univ_semialgebraicE are_univ_semialgebraic_def assms nat_mono nat_numeral subset_iff) have 3: "\ F = (ord_congruence_set n a)" proof show "\ F \ ord_congruence_set n a" using F_def by blast show "ord_congruence_set n a \ \ F" proof fix x assume A: "x \ ord_congruence_set n a" have x_nonzero: "x \ nonzero Q\<^sub>p" using A ord_congruence_set_memE(1) by blast have 0: "pow_res (nat n) x \ F" using A pow_res_classes_def F_def by (smt nonzero_def assms mem_Collect_eq nat_0_le ord_congruence_set_memE(1) pow_res_ord_cong) have 1: "x \ pow_res (nat n) x" using False x_nonzero assms pow_res_refl[of x "nat n"] using Qp.nonzero_closed by blast show "x \ \ F" using 0 1 by blast qed qed then show ?thesis using "1" "2" finite_union_is_univ_semialgebraic' by fastforce qed lemma ord_congruence_set_is_semialg: assumes "n \ 0" shows "is_semialgebraic 1 (Qp_to_R1_set (ord_congruence_set n a))" using assms is_univ_semialgebraicE ord_congruence_set_univ_semialg by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Congruence Sets for the order of the Evaluation of a Polynomial\ (********************************************************************************************) (********************************************************************************************) lemma poly_map_singleton: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n [f] x = [(Qp_ev f x)]" unfolding poly_map_def poly_tuple_eval_def using assms by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') definition poly_cong_set where "poly_cong_set n f m a = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). (Qp_ev f x) \ \ \ (ord (Qp_ev f x) mod m = a)}" lemma poly_cong_set_as_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "poly_cong_set n f m a = poly_map n [f] \\<^bsub>n\<^esub>(Qp_to_R1_set (ord_congruence_set m a))" proof show "poly_cong_set n f m a \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" proof fix x assume A: "x \ poly_cong_set n f m a" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis (no_types, lifting) mem_Collect_eq poly_cong_set_def) have 1: "(Qp_ev f x) \ \ " by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) have 2: "(ord (Qp_ev f x) mod m = a)" by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) have 3: "(Qp_ev f x) \ (ord_congruence_set m a)" using "0" "1" "2" eval_at_point_closed assms not_nonzero_Qp ord_congruence_set_memI by metis show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" proof- have 00: "poly_map n [f] x = [(Qp_ev f x)]" using "0" assms poly_map_singleton by blast have 01: "[eval_at_point Q\<^sub>p x f] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using "0" assms eval_at_point_closed Qp.to_R1_closed by blast hence 02: "poly_map n [f] x \ (\a. [a]) ` ord_congruence_set m a" using 3 "00" by blast then show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" using 0 unfolding evimage_def by blast qed qed show "poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a \ poly_cong_set n f m a" proof fix x assume A: "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` (ord_congruence_set m a))" have 0: "((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using ord_congruence_set_closed Qp.to_R1_carrier by blast have "is_poly_tuple n [f]" using assms unfolding is_poly_tuple_def by (simp add: assms) then have 1:"poly_map n [f] \\<^bsub>n\<^esub>((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using 0 A assms One_nat_def by (metis extensional_vimage_closed) then have 2: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A unfolding evimage_def by blast then have 3: "poly_map n [f] x \ ((\a. [a]) ` ord_congruence_set m a)" using A assms 0 One_nat_def by blast have "poly_map n [f] x = [(Qp_ev f x)]" using "2" assms poly_map_singleton by blast then have "Qp_ev f x \ ord_congruence_set m a" using 3 by (metis (mono_tags, lifting) image_iff list.inject) then show "x \ poly_cong_set n f m a" unfolding poly_cong_set_def by (metis (mono_tags, lifting) "2" Qp.nonzero_memE(2) mem_Collect_eq ord_congruence_set_memE(1) ord_congruence_set_memE(2)) qed qed lemma singleton_poly_tuple: "is_poly_tuple n [f] \ f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" unfolding is_poly_tuple_def by (metis (no_types, lifting) list.distinct(1) list.set_cases list.set_intros(1) set_ConsD subset_code(1)) lemma poly_cong_set_is_semialgebraic: assumes "m \ 0" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n (poly_cong_set n f m a)" proof- have 0: "(\a. [a]) ` ord_congruence_set m a \ semialg_sets 1" using assms ord_congruence_set_is_semialg[of m a] unfolding is_semialgebraic_def by blast have 1: "length [f] = 1" by simp hence " poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a \ semialg_sets n" using 0 singleton_poly_tuple[of n f] zero_neq_one assms pullback_is_semialg[of n "[f]" 1 "(\a. [a]) ` ord_congruence_set m a"] unfolding is_semialgebraic_def by blast thus ?thesis using assms poly_cong_set_as_pullback[of f n m a] unfolding is_semialgebraic_def by presburger qed (********************************************************************************************) (********************************************************************************************) subsubsection\Congruence Sets for Angular Components\ (********************************************************************************************) (********************************************************************************************) text\If a set is a union of \n\-th power residues, then it is semialgebraic.\ lemma pow_res_union_imp_semialg: assumes "n \ 1" assumes "S \ nonzero Q\<^sub>p" assumes "\x. x \ S \ pow_res n x \ S" shows "is_univ_semialgebraic S" proof- obtain F where F_def: "F = {T. T \ pow_res_classes n \ T \ S}" by blast have 0: "F \ pow_res_classes n" using F_def by blast have 1: "finite F" using 0 pow_res_classes_finite[of n] assms(1) finite_subset by auto have 2: "are_univ_semialgebraic F" using 0 by (meson are_univ_semialgebraicE are_univ_semialgebraicI assms(1) pow_res_classes_are_univ_semialgebraic padic_fields_axioms subsetD) have 3: "S = \ F" proof show "S \ \ F" proof fix x assume A: "x \ S" then have "pow_res n x \ S" using assms(3) by blast then have "pow_res n x \ F" using A assms(2) F_def pow_res_classes_def by (smt mem_Collect_eq subsetD) then have "pow_res n x \ \ F" by blast then show "x \ \ F" using A assms(1) assms(2) pow_res_refl[of x n] unfolding nonzero_def by blast qed show "\ F \ S" using F_def by blast qed show ?thesis using 1 2 3 finite_union_is_univ_semialgebraic' by blast qed definition ac_cong_set1 where "ac_cong_set1 n y = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = ac n y}" lemma ac_cong_set1_is_univ_semialg: assumes "n > 0" assumes "b \ nonzero Q\<^sub>p" assumes "b \ \\<^sub>p" shows "is_univ_semialgebraic (ac_cong_set1 n b)" proof(cases "n = 1 \ p = 2") case True have "(ac_cong_set1 n b) = nonzero Q\<^sub>p" proof have 0: "Units (Zp_res_ring n) = {1}" proof show "Units (Zp_res_ring n) \ {1}" proof fix x assume A: "x \ Units (Zp_res_ring n)" have 0: "carrier (Zp_res_ring n) = {0..(int 2) - 1}" using True by (metis assms(1) int_ops(3) p_residues power_one_right residues.res_carrier_eq) have 1: "carrier (Zp_res_ring n) = {0..(1::int)}" proof- have "int 2 - 1 = (1::int)" by linarith then show ?thesis using 0 by presburger qed have 15: "{0..(1::int)} = {0, (1::int)}" using atLeastAtMostPlus1_int_conv [of 0 "0::int"] by (smt atLeastAtMost_singleton insert_commute) have 2: "carrier (Zp_res_ring n) = {0,(1::int)}" using "1" "15" by blast have 3: "0 \ Units (Zp_res_ring n)" using True zero_not_in_residue_units by blast have "x \ carrier (Zp_res_ring n)" using A unfolding Units_def by blast then have "x = 1" using A 2 3 by (metis "1" atLeastAtMost_iff atLeastatMost_empty atLeastatMost_empty_iff2 linorder_neqE_linordered_idom mod_by_1 mod_pos_pos_trivial ) then show "x \ {1}" by simp qed show "{1} \ Units (Zp_res_ring n)" by (meson assms(1) empty_subsetI insert_subset residue_1_unit(1)) qed show "ac_cong_set1 n b \ nonzero Q\<^sub>p" by (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) show "nonzero Q\<^sub>p \ ac_cong_set1 n b" proof fix x assume A: "x \ nonzero Q\<^sub>p" then have P0: "ac n x = 1" using 0 ac_units assms(1) by blast have P1: "ac n b = 1" using assms 0 ac_units assms(1) by blast then have "ac n x = ac n b" using P0 by metis then show " x \ ac_cong_set1 n b" unfolding ac_cong_set1_def using A proof - have "x \ {r \ carrier Q\<^sub>p. r \ \}" by (metis (no_types) \x \ nonzero Q\<^sub>p\ nonzero_def ) then show "x \ {r \ carrier Q\<^sub>p. r \ \ \ ac n r = ac n b}" using \ac n x = ac n b\ by force qed qed qed then show "is_univ_semialgebraic (ac_cong_set1 n b)" by (simp add: nonzero_is_univ_semialgebraic) next case F: False have F0: "2 \ card (Units (Zp_res_ring n))" proof(cases "n = 1") case True then have "field (Zp_res_ring n)" using p_res_ring_1_field by blast then have F00: "Units (Zp_res_ring n) = carrier (Zp_res_ring n) - {\\<^bsub>Zp_res_ring n\<^esub>}" using field.field_Units by blast have F01: "\\<^bsub>Zp_res_ring n\<^esub> \ carrier (Zp_res_ring n)" using assms(1) cring.cring_simprules(2) padic_integers.R_cring padic_integers_axioms by blast have F02: "card (carrier (Zp_res_ring n)) = p \ finite (carrier (Zp_res_ring n))" by (smt F01 True nat_eq_iff2 p_res_ring_zero p_residue_ring_car_memE(1) power_one_right residue_ring_card) have F03: "\\<^bsub>residue_ring (p ^ n)\<^esub> \ carrier (residue_ring (p ^ n)) " using F01 by blast have F04: "int (card (carrier (residue_ring (p ^ n)))) \ int (card {\\<^bsub>residue_ring (p ^ n)\<^esub>}) " by (smt F02 F03 nat_int of_nat_0_le_iff of_nat_1 of_nat_power p_res_ring_0 p_res_ring_zero p_residue_ring_car_memE(1) power_increasing power_one_right residue_ring_card) have "card (carrier (residue_ring (p ^ n))) - 1 = p - 1" using F02 prime by (metis Totient.of_nat_eq_1_iff True less_imp_le_nat less_one nat_int nat_less_eq_zless of_nat_1 of_nat_diff of_nat_zero_less_power_iff p_residues pos_int_cases power_0 power_one_right residue_ring_card residues.m_gt_one zero_le_one) hence F05: "card (carrier (residue_ring (p ^ n)) - {\\<^bsub>residue_ring (p ^ n)\<^esub>}) = p - 1" using F02 F03 F04 card_Diff_singleton_if[of "(carrier (Zp_res_ring n))" "\\<^bsub>residue_ring (p^n)\<^esub>"] True int_ops(6)[of "card (carrier (residue_ring (p ^ n)))" "card {\\<^bsub>residue_ring (p ^ n)\<^esub>}"] p_res_ring_zero p_residue_ring_car_memE(1) by (metis) hence F06: "card (Units (Zp_res_ring n)) = p -1" using True F02 F01 F00 by (metis p_res_ring_zero) have F04: "p - 1 \2 " using F prime by (meson True linorder_cases not_less prime_ge_2_int zle_diff1_eq) then show ?thesis using F03 F06 by linarith next case False then show ?thesis by (metis assms(1) less_imp_le_nat mod2_gr_0 mod_less nat_le_linear nat_neq_iff residue_units_card_geq_2) qed show ?thesis apply(rule pow_res_union_imp_semialg[of "card (Units (Zp_res_ring n))"]) using F0 assms apply linarith apply (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) proof- fix x assume AA: "x \ ac_cong_set1 n b" show "pow_res (card (Units (Zp_res_ring n))) x \ ac_cong_set1 n b" proof fix y assume A: "y \ pow_res (card (Units (Zp_res_ring n))) x" show "y \ ac_cong_set1 n b" proof- obtain k where k_def: "k = card (Units (Zp_res_ring n))" by blast have "k \2" using assms k_def F F0 by blast then obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]k)" using k_def A pow_res_def[of k x] by blast have 0: "x \ nonzero Q\<^sub>p" using AA ac_cong_set1_def by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) have 1: "y \ nonzero Q\<^sub>p" by (metis "0" Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero \\thesis. (\a. a \ nonzero Q\<^sub>p \ y = x \ a [^] k \ thesis) \ thesis\) have "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)" using a_def 0 1 Qp_nat_pow_nonzero ac_mult' by blast then have 2: "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" proof- have "ac n (a[^]k) = ac n a [^]\<^bsub>Zp_res_ring n\<^esub> k" using a_def assms(1) ac_nat_pow'[of a n k] by linarith then show ?thesis using \ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)\ by presburger qed then have "ac n y = ac n x" proof- have "(ac n a) \ Units (Zp_res_ring n)" by (metis (mono_tags, opaque_lifting) a_def ac_units assms(1) ) then have "(ac n a)^k mod (p^n) = 1" using k_def a_def ac_nat_pow ac_nat_pow' assms(1) residue_units_nilpotent using neq0_conv by presburger then have 00: "(ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k = 1" by (metis a_def ac_nat_pow ac_nat_pow' mod_by_1 power_0 zero_neq_one) have "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x \\<^bsub>Zp_res_ring n\<^esub> \\<^bsub>Zp_res_ring n\<^esub>" using 00 assms(1) p_res_ring_one by presburger hence "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x" by (metis "0" Qp.nonzero_closed Qp.one_nonzero Qp.r_one ac_mult' ac_one' assms(1)) then show ?thesis using 2 "0" 00 by linarith qed then show ?thesis using "1" AA nonzero_def ac_cong_set1_def[of n b] mem_Collect_eq by smt qed qed qed qed definition ac_cong_set where "ac_cong_set n k = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" lemma ac_cong_set_is_univ_semialg: assumes "n >0 " assumes "k \ Units (Zp_res_ring n)" shows "is_univ_semialgebraic (ac_cong_set n k)" proof- have "k \ carrier (Zp_res_ring n)" using assms(2) Units_def[of "Zp_res_ring n"] by blast then have k_n: "([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) n = k" using assms by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) obtain b where b_def: "b = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" by blast have 0: "k mod p \ 0" using assms residue_UnitsE[of n k] by (metis le_eq_less_or_eq le_refl less_one nat_le_linear p_residues power_0 power_one_right residues.mod_in_res_units residues_def zero_less_one zero_neq_one zero_not_in_residue_units zero_power) then have "val_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = 0" using val_Zp_p_int_unit by blast then have 1: "val b = 0" by (metis Zp_int_inc_closed b_def val_of_inc) have 2: "b \ \\<^sub>p" using b_def Zp_int_mult_closed by blast have "ord_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0" using 0 ord_Zp_p_int_unit by blast have "ac_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using "0" Zp_int_inc_closed ac_Zp_of_Unit ord_Zp_p_int_unit \val_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0\ by blast then have "(angular_component b) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using b_def 1 2 angular_component_ord_zero[of b] by (metis Qp.int_inc_zero Qp.one_closed val_ring_memE Zp.int_inc_zero Zp.one_closed Zp.one_nonzero Zp_int_inc_closed angular_component_of_inclusion inc_closed inc_of_int inc_of_one inc_to_Zp local.val_zero not_nonzero_Qp val_ineq val_one zero_in_val_ring) then have "ac n b = k" using ac_def[of n b] k_n by (metis Qp_char_0_int Zp_defs(1) ac_def b_def inc_of_int inc_of_one) then have 3: "(ac_cong_set n k) = (ac_cong_set1 n b)" unfolding ac_cong_set_def ac_cong_set1_def by meson have 4: "b \ nonzero Q\<^sub>p" using 1 2 val_nonzero by (metis Qp.one_closed val_ring_memE Zp_def \_def local.one_neq_zero not_nonzero_Qp padic_fields.val_ring_memE padic_fields_axioms val_ineq val_one) then show ?thesis using 1 2 3 assms ac_cong_set1_is_univ_semialg[of n b] val_nonzero[of b 1] by presburger qed definition val_ring_constant_ac_set where "val_ring_constant_ac_set n k = {a \ \\<^sub>p. val a = 0 \ ac n a = k}" lemma val_nonzero': assumes "a \ carrier Q\<^sub>p" assumes "val a = eint k" shows "a \ nonzero Q\<^sub>p" using val_nonzero[of a "k + 1"] by (metis Suc_ile_eq assms(1) assms(2) eint_ord_code(3) val_nonzero) lemma val_ord': assumes "a \ carrier Q\<^sub>p" assumes "a \\" shows "val a = ord a" by (meson assms(1) assms(2) not_nonzero_Qp val_ord) lemma val_ring_constant_ac_set_is_univ_semialgebraic: assumes "n > 0" assumes "k \ 0" shows "is_univ_semialgebraic (val_ring_constant_ac_set n k)" proof(cases "val_ring_constant_ac_set n k = {}") case True then show ?thesis by (metis equals0D order_refl pow_res_union_imp_semialg subsetI) next case False then obtain b where b_def: "b \ val_ring_constant_ac_set n k" by blast have 0: "val_ring_constant_ac_set n k = q_ball n k 0 \" proof show "val_ring_constant_ac_set n k \ q_ball n k 0 \" proof fix x assume A: "x \ val_ring_constant_ac_set n k" then show "x \ q_ball n k 0 \" proof- have 0: "x \ \\<^sub>p \ val x = 0 \ ac n x = k" using A unfolding val_ring_constant_ac_set_def by blast then have x_car: "x \ carrier Q\<^sub>p" using val_ring_memE by blast then have 00: "x = x \ \" using Qp.ring_simprules by metis then have 1: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" using 0 by presburger have 2: "val (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" using 0 00 by metis have 3: "x \ nonzero Q\<^sub>p" proof(rule ccontr) assume " x \ nonzero Q\<^sub>p " then have "x = \" using Qp.nonzero_memI x_car by blast then show False using 0 val_zero by (metis ac_def assms(2)) qed have 4: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" proof(rule ccontr) assume "ord (x \ \) \ 0" then have "val (x \ \) \ 0" by (metis "00" "3" Qp.one_closed equal_val_imp_equal_ord(1) ord_one val_one) then show False using "2" by blast qed show ?thesis using 0 1 4 unfolding q_ball_def using x_car by blast qed qed show "q_ball n k 0 \ \ val_ring_constant_ac_set n k" proof fix x assume A: "x \ q_ball n k 0 \" then have 0: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" using q_ballE'(1) by blast have 1: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" using q_ball_def A by blast have 2: "x \ carrier Q\<^sub>p" using A q_ball_def by blast have 3: "ord x = 0" using 2 1 ring.ring_simprules[of Q\<^sub>p] by (metis Qp.ring_axioms) have 4: "ac n x = k" using 0 2 1 cring.axioms(1)[of Q\<^sub>p] ring.ring_simprules[of Q\<^sub>p] by (metis Qp.ring_axioms) have 5: "x \ \\<^sub>p" using Qp_val_ringI[of x] 2 3 val_ord val_nonzero' by (metis Qp.integral_iff val_ring_memE Zp.nonzero_closed angular_component_closed angular_component_ord_zero image_eqI local.numer_denom_facts(1) local.numer_denom_facts(2) local.numer_denom_facts(4) not_nonzero_Qp) have 6: "x \ \" using 4 assms ac_def[of n x] by meson have 7: "val x = 0" using 6 3 2 assms val_ord' zero_eint_def by presburger show " x \ val_ring_constant_ac_set n k" unfolding val_ring_constant_ac_set_def using 7 6 5 4 by blast qed qed obtain b where b_def: "b \ q_ball n k (0::int) \" using "0" b_def by blast have 1: "b \ carrier Q\<^sub>p \ ac n b = k" using b_def unfolding q_ball_def by (metis (mono_tags, lifting) "0" b_def mem_Collect_eq val_ring_constant_ac_set_def) then have 2: "b \ nonzero Q\<^sub>p" using 1 assms by (metis ac_def not_nonzero_Qp) have "q_ball n k 0 \ = B\<^bsub>0 + int n\<^esub>[b]" using 1 b_def nonzero_def [of Q\<^sub>p] assms 0 2 c_ball_q_ball[of b n k "\" b 0] by (meson Qp.cring_axioms cring.cring_simprules(2)) then have "is_univ_semialgebraic (q_ball n k (0::int) \) " using 1 ball_is_univ_semialgebraic[of b "0 + int n"] by metis then show ?thesis using 0 by presburger qed definition val_ring_constant_ac_sets where "val_ring_constant_ac_sets n = val_ring_constant_ac_set n ` (Units (Zp_res_ring n))" lemma val_ring_constant_ac_sets_are_univ_semialgebraic: assumes "n > 0" shows "are_univ_semialgebraic (val_ring_constant_ac_sets n)" proof(rule are_univ_semialgebraicI) have 0: "\ coprime 0 p" using coprime_0_right_iff[of p] coprime_commute[of p 0] coprime_int_iff[of "nat p" 0] nat_dvd_1_iff_1 prime_gt_1_nat zdvd1_eq by (metis not_prime_unit prime) have "(0::int) \(Units (Zp_res_ring n))" apply(rule ccontr) using 0 assms residues.cring[of "p ^ n"] unfolding residues_def by (smt less_one not_gr_zero power_le_imp_le_exp power_less_imp_less_exp residue_UnitsE) fix x assume A: "x \ val_ring_constant_ac_sets n" then obtain k where k_def: "x = val_ring_constant_ac_set n k \ k \ Units (Zp_res_ring n)" by (metis image_iff val_ring_constant_ac_sets_def) then show "is_univ_semialgebraic x" using assms by (metis \0 \ Units (Zp_res_ring n)\ val_ring_constant_ac_set_is_univ_semialgebraic) qed definition ac_cong_set3 where "ac_cong_set3 n = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = ac n b) \ as = [a, b] }" definition ac_cong_set2 where "ac_cong_set2 n k = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k \ as = [a, b] }" lemma ac_cong_set2_cartesian_product: assumes "k \ Units (Zp_res_ring n)" assumes "n > 0" shows "ac_cong_set2 n k = cartesian_product (to_R1` (ac_cong_set n k)) (to_R1` (val_ring_constant_ac_set n k))" proof show "ac_cong_set2 n k \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" proof fix x assume A: "x \ ac_cong_set2 n k" show "x \ (cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k))" unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 1]) apply (metis (mono_tags, lifting) mem_Collect_eq subsetI Qp.to_R1_car_subset) apply (metis (no_types, lifting) val_ring_memE mem_Collect_eq subsetI Qp.to_R1_car_subset) proof- obtain a b where ab_def: "x = [a,b] \ a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k" using A unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def by blast have 0: "take 1 x = [a]" by (simp add: ab_def) have 1: "drop 1 x = [b]" by (simp add: ab_def) have 2: "a \ {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" using ab_def nonzero_def by (smt mem_Collect_eq) have 3: "b \ {a \ \\<^sub>p. val a = 0 \ ac n a = k}" using ab_def by blast show "take 1 x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" using 0 2 by blast show "drop 1 x \ (\a. [a]) ` {a \ \\<^sub>p. val a = 0 \ ac n a = k}" using 1 3 by blast qed qed show "cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k) \ ac_cong_set2 n k" proof fix x have 0: "(\a. [a]) ` ac_cong_set n k \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms by (metis (no_types, lifting) ac_cong_set_def mem_Collect_eq subsetI Qp.to_R1_car_subset) have 1: "((\a. [a]) ` val_ring_constant_ac_set n k) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" by (smt val_ring_memE mem_Collect_eq subsetI Qp.to_R1_carrier Qp.to_R1_subset val_ring_constant_ac_set_def) assume A: "x \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" then have "length x = 2" using 0 1 A cartesian_product_closed[of "((\a. [a]) ` ac_cong_set n k)" Q\<^sub>p 1 "((\a. [a]) ` val_ring_constant_ac_set n k)" 1] by (metis (no_types, lifting) cartesian_power_car_memE one_add_one subset_iff) then obtain a b where ab_def: "take 1 x = [a] \ drop 1 x = [b]" by (metis One_nat_def add_diff_cancel_left' drop0 drop_Cons_numeral numerals(1) pair_id plus_1_eq_Suc take0 take_Cons_numeral) have 2: " a \ (ac_cong_set n k) \ b \ (val_ring_constant_ac_set n k)" proof- have P0: "take 1 x \ (\a. [a]) ` ac_cong_set n k" using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] by blast have P1: "drop 1 x \ (\a. [a]) ` val_ring_constant_ac_set n k" using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] by blast have P2: "[a] \ (\a. [a]) ` ac_cong_set n k" using P0 ab_def by metis have P3: "[b] \ (\a. [a]) ` val_ring_constant_ac_set n k" using P1 ab_def by metis show ?thesis using P2 P3 by blast qed have 3: "a \ nonzero Q\<^sub>p" using 2 assms nonzero_def [of Q\<^sub>p] ac_cong_set_def[of n k] by blast have 4: "x = [a,b]" by (metis (no_types, lifting) \length x = 2\ ab_def less_numeral_extra(1) nth_Cons_0 nth_take nth_via_drop pair_id) then have "\a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ ac n a = k \ ac n b = k \ x = [a, b]" using 2 3 ab_def unfolding val_ring_constant_ac_set_def ac_cong_set_def by blast then show "x \ ac_cong_set2 n k" unfolding ac_cong_set2_def val_ring_constant_ac_set_def ac_cong_set_def by blast qed qed lemma ac_cong_set2_is_semialg: assumes "k \ Units (Zp_res_ring n)" assumes "n > 0" shows "is_semialgebraic 2 (ac_cong_set2 n k)" using ac_cong_set_is_univ_semialg ac_cong_set2_cartesian_product[of k n] cartesian_product_is_semialgebraic[of 1 "((\a. [a]) ` ac_cong_set n k)" 1 " ((\a. [a]) ` val_ring_constant_ac_set n k)"] by (metis assms(1) assms(2) is_univ_semialgebraicE less_one less_or_eq_imp_le nat_neq_iff one_add_one val_ring_constant_ac_set_is_univ_semialgebraic zero_not_in_residue_units) lemma ac_cong_set3_as_union: assumes "n > 0" shows "ac_cong_set3 n = \ (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" proof show "ac_cong_set3 n \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" proof fix x assume A: "x \ ac_cong_set3 n" then have 0: "x \ (ac_cong_set2 n (ac n (x!0)))" unfolding ac_cong_set2_def ac_cong_set3_def by (smt mem_Collect_eq nth_Cons_0) have 1: "(ac n (x!0)) \ Units (Zp_res_ring n)" using A unfolding ac_cong_set3_def by (smt ac_units assms mem_Collect_eq nth_Cons_0) then show "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" using 0 by blast qed show "\ (ac_cong_set2 n ` Units (Zp_res_ring n)) \ ac_cong_set3 n" proof fix x assume A: "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" obtain k where k_def: "x \ (ac_cong_set2 n k) \ k \ (Units (Zp_res_ring n))" using A by blast have 0: "k mod p \ 0" using k_def One_nat_def Suc_le_eq assms less_numeral_extra(1) power_one_right residues.m_gt_one residues.mod_in_res_units by (metis p_residues residue_UnitsE zero_not_in_residue_units) obtain b where b_def: "b = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" by blast have "k \0" using 0 mod_0 by blast then have 1: "b \ nonzero Z\<^sub>p" using 0 b_def int_unit by (metis Zp.Units_nonzero Zp.zero_not_one) have 10: "ord_Zp b = 0" using 0 1 using b_def ord_Zp_p_int_unit by blast have 2: "\ b \ nonzero Q\<^sub>p" using k_def using "1" inc_of_nonzero by blast have 3: "angular_component (\ b) = ac_Zp b" using "1" angular_component_of_inclusion by blast have 4: "ac_Zp b = b" using 1 10 by (metis "3" Zp.r_one ac_Zp_factors' angular_component_closed inc_of_nonzero int_pow_0 mult_comm ord_Zp_def) have 5: "ac_Zp b n = k" proof- have "k \ carrier (Zp_res_ring n)" using k_def unfolding Units_def by blast then show ?thesis using b_def k_def 4 Zp_int_inc_res mod_pos_pos_trivial by (metis p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) qed then have "ac n (\ b) = k" using 10 1 2 3 4 unfolding ac_def using Qp.not_nonzero_memI by metis then show "x \ ac_cong_set3 n" unfolding ac_cong_set3_def using k_def unfolding ac_cong_set2_def by (smt mem_Collect_eq) qed qed lemma ac_cong_set3_is_semialgebraic: assumes "n > 0" shows "is_semialgebraic 2 (ac_cong_set3 n)" proof- have 0: "finite (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" using assms residues.finite_Units[of "p^n"] unfolding residues_def using p_residues residues.finite_Units by blast have 1: "are_semialgebraic 2 (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" apply(rule are_semialgebraicI) using ac_cong_set2_is_semialg assms by blast show ?thesis using 0 1 ac_cong_set3_as_union by (metis (no_types, lifting) are_semialgebraicE assms finite_union_is_semialgebraic' is_semialgebraicE subsetI) qed (**************************************************************************************************) (**************************************************************************************************) subsection\Permutations of indices of semialgebraic sets\ (**************************************************************************************************) (**************************************************************************************************) lemma fun_inv_permute: assumes "\ permutes {.. permutes {.. \ (fun_inv \) = id" "(fun_inv \) \ \ = id" using assms unfolding fun_inv_def using permutes_inv apply blast using assms permutes_inv_o(1) apply blast using assms permutes_inv_o(2) by blast lemma poly_tuple_pullback_eq_poly_map_vimage: assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "S \ carrir (Q\<^sub>p\<^bsup>m\<^esup>)" shows "poly_map n fs \\<^bsub>n\<^esub> S = poly_tuple_pullback n S fs" unfolding poly_map_def poly_tuple_pullback_def evimage_def restrict_def using assms by (smt vimage_inter_cong) lemma permutation_is_semialgebraic: assumes "is_semialgebraic n S" assumes "\ permutes {.. ` S)" proof- have "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms gen_boolean_algebra_subset is_semialgebraic_def semialg_sets_def by blast then have "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" using Qp.cring_axioms assms pullback_by_permutation_of_poly_list'[of \ n S] unfolding poly_map_def by blast then have 0: "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" using poly_tuple_pullback_def by blast have 1: "(fun_inv \) permutes {..) (pvar_list Q\<^sub>p n))"] permutation_of_poly_list_is_poly_list[of n "(pvar_list Q\<^sub>p n)" "fun_inv \"] pvar_list_is_poly_tuple[of n] assms poly_tuple_pullback_eq_poly_map_vimage by (metis "0" \S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ is_semialgebraic_def length_permute_list pvar_list_length) qed lemma permute_list_closed: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) using assms cartesian_power_car_memE length_permute_list apply blast using assms cartesian_power_car_memE'' permute_list_set by blast lemma permute_list_closed': assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) apply (metis assms(2) cartesian_power_car_memE length_permute_list) using assms cartesian_power_car_memE'[of "permute_list \ a" Q\<^sub>p n] by (metis cartesian_power_car_memE in_set_conv_nth length_permute_list set_permute_list subsetI) lemma permute_list_compose_inv: assumes "\ permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "permute_list \ (permute_list (fun_inv \) a) = a" "permute_list (fun_inv \) (permute_list \ a) = a" using assms apply (metis cartesian_power_car_memE fun_inv_permute(3) permute_list_compose permute_list_id) using assms by (metis cartesian_power_car_memE fun_inv_permute(2) fun_inv_permute(1) permute_list_compose permute_list_id) lemma permutation_is_semialgebraic_imp_is_semialgebraic: assumes "is_semialgebraic n (permute_list \ ` S)" assumes "\ permutes {..) ` (permute_list \ ` S) = S" proof- have 0: "(permute_list \ ` S) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms unfolding is_semialgebraic_def semialg_sets_def using gen_boolean_algebra_subset by blast have 1: "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" proof fix x assume "x \ S" then show "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using 0 assms by (meson image_subset_iff permute_list_closed') qed show ?thesis proof show "permute_list (fun_inv \) ` permute_list \ ` S \ S" using 0 assms permute_list_compose_inv[of \] "1" image_iff image_subset_iff subsetD by smt show "S \ permute_list (fun_inv \) ` permute_list \ ` S" using 0 assms permute_list_compose_inv[of \] by (smt "1" image_iff subset_eq) qed qed then show ?thesis using permutation_is_semialgebraic by (metis assms(1) assms(2) fun_inv_permute(1)) qed lemma split_cartesian_product_is_semialgebraic: assumes "i \ n" assumes "is_semialgebraic n A" assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (split_cartesian_product n m i A B)" using assms cartesian_product_is_semialgebraic scp_permutes[of i n m] permutation_is_semialgebraic[of "n + m" "cartesian_product A B" "(scp_permutation n m i)"] unfolding split_cartesian_product_def by blast definition reverse_val_relation_set where "reverse_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)}" lemma Qp_2_car_memE: assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" shows "x = [x!0, x!1]" proof- have "length x = 2" using assms cartesian_power_car_memE by blast then show ?thesis using pair_id by blast qed definition flip where "flip = (\i::nat. (if i = 0 then 1 else (if i = 1 then 0 else i)))" lemma flip_permutes: "flip permutes {0,1}" unfolding permutes_def flip_def by (smt mem_simps(1)) lemma flip_eval: "flip 0 = 1" "flip 1 = 0" unfolding flip_def by auto lemma flip_x: assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" shows "permute_list flip x = [x!1, x!0]" proof- have 0: "x = [x!0, x!1]" using assms Qp_2_car_memE by blast have 1: "length (permute_list flip x) = length [x!1, x!0]" using 0 unfolding permute_list_def by (metis length_Cons length_map map_nth) have 2: "\i. i < 2 \ permute_list flip x ! i = [x!1, x!0] ! i" proof- fix i::nat assume A: "i < 2" show "permute_list flip x ! i = [x!1, x!0] ! i" using 0 unfolding permute_list_def by (smt flip_eval(1) flip_eval(2) length_Cons length_greater_0_conv list.simps(8) map_upt_Suc numeral_nat(7) upt_rec) qed have "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" proof- have 0: "length x = 2" using assms cartesian_power_car_memE by blast show "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" using 2 unfolding 0 by blast qed thus ?thesis using 1 by (metis length_permute_list nth_equalityI) qed lemma permute_with_flip_closed: assumes "x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" shows "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" apply(rule permute_list_closed) using assms apply blast proof- have "{0::nat, 1} = {..<2::nat}" by auto thus "flip permutes {..<2}" using flip_permutes by auto qed lemma reverse_val_relation_set_semialg: "is_semialgebraic 2 reverse_val_relation_set" proof- have 1: "reverse_val_relation_set = permute_list flip ` val_relation_set" apply(rule equalityI') proof- show " \x. x \ reverse_val_relation_set \ x \ permute_list flip ` val_relation_set" proof- fix x assume A: "x \ reverse_val_relation_set" have 0: "permute_list flip x = [x ! 1, x ! 0]" using flip_x[of x] A unfolding reverse_val_relation_set_def by blast have 1: "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule permute_with_flip_closed) using A unfolding reverse_val_relation_set_def by blast have 2: "permute_list flip x \ val_relation_set" using 1 A unfolding 0 reverse_val_relation_set_def val_relation_set_def mem_Collect_eq by (metis Qp_2_car_memE list_hd list_tl) show "x \ permute_list flip ` val_relation_set" using flip_x[of x] A unfolding reverse_val_relation_set_def val_relation_set_def mem_Collect_eq by (metis (no_types, lifting) "1" "2" Qp_2_car_memE flip_x image_eqI list_tl nth_Cons_0 val_relation_set_def) qed show "\x. x \ permute_list flip ` val_relation_set \ x \ reverse_val_relation_set" proof- fix x assume a: " x \ permute_list flip ` val_relation_set" then obtain y where y_def: "y \ val_relation_set \x = permute_list flip y" by blast have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using y_def basic_semialg_set_memE(1) val_relation_semialg by blast have y_length: " length y = 2" using y_def basic_semialg_set_memE val_relation_semialg by (metis cartesian_power_car_memE) obtain a b where ab_def: "y = [a,b]" using y_length pair_id by blast have 0: "a = y!0" using ab_def by (metis nth_Cons_0) have 1: "b = y!1" using ab_def by (metis cancel_comm_monoid_add_class.diff_cancel eq_numeral_extra(2) nth_Cons') have a_closed: "a \ carrier Q\<^sub>p" using 0 y_closed unfolding 0 by (meson cartesian_power_car_memE' rel_simps(75) zero_order(5)) have b_closed: "b \ carrier Q\<^sub>p" proof- have "1 < (2::nat)" by linarith thus ?thesis using y_closed unfolding 1 by (meson cartesian_power_car_memE') qed have 2: "x = [b, a]" using flip_x[of y] y_def y_closed unfolding ab_def unfolding 0 1 using \y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ permute_list flip y = [y ! 1, y ! 0]\ y_closed y_def by presburger have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using y_def unfolding val_relation_set_def using permute_with_flip_closed[of y] by blast show " x \ reverse_val_relation_set" using x_closed y_def unfolding val_relation_set_def reverse_val_relation_set_def mem_Collect_eq 2 0 1 by (metis Qp_2_car_memE list_hd list_tl) qed qed show ?thesis unfolding 1 apply(rule permutation_is_semialgebraic) using val_relation_is_semialgebraic apply blast using flip_permutes by (metis Suc_1 insert_commute lessThan_0 lessThan_Suc numeral_nat(7)) qed definition strict_val_relation_set where "strict_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) < val (as ! 1)}" definition val_diag where "val_diag = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)}" lemma val_diag_semialg: "is_semialgebraic 2 val_diag" proof- have "val_diag = val_relation_set \reverse_val_relation_set" apply(rule equalityI') apply(rule IntI) unfolding val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq apply simp apply simp apply(erule IntE) unfolding mem_Collect_eq using basic_trans_rules(24) by blast then show ?thesis using intersection_is_semialg by (simp add: reverse_val_relation_set_semialg val_relation_is_semialgebraic) qed lemma strict_val_relation_set_is_semialg: "is_semialgebraic 2 strict_val_relation_set" proof- have 0: "strict_val_relation_set = reverse_val_relation_set - val_diag" apply(rule equalityI') apply(rule DiffI) unfolding strict_val_relation_set_def val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq using order_le_less apply blast proof show "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1) \ x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) = val (x ! 1) \ False" using order_less_le by blast show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)} - {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)} \ x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1)" apply(erule DiffE) unfolding mem_Collect_eq using order_le_less by blast qed show ?thesis unfolding 0 apply(rule diff_is_semialgebraic ) using reverse_val_relation_set_semialg apply blast using val_diag_semialg by blast qed lemma singleton_length: "length [a] = 1" by auto lemma take_closed': assumes "m > 0" assumes "x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" shows "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule take_closed[of m "m+l"]) apply simp using assms by blast lemma triple_val_ineq_set_semialg: shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}" proof- have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule equalityI') show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!0, x!1]" by blast have a_length: "length a = 2" proof- have "a = x!0 #[x!1]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] by presburger qed obtain b where b_def: "b = [x!2]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule cartesian_power_car_memI') apply (simp add: a_length) unfolding 0 using A unfolding mem_Collect_eq using cartesian_power_car_memE' by fastforce show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] by simp have 2: "x = a@b" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of a b] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (a @ b) ! i" apply(cases "i = 0") apply (metis a_def append.simps(2) nth_Cons_0) apply(cases "(i:: nat) = 1") apply (simp add: a_def) proof- assume a: "i \0" "i \ 1" then have "i = 2" using A1 by presburger thus ?thesis by (metis a_length b_def nth_append_length) qed qed have 3: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) apply blast using 3 a_closed apply blast proof- have "drop 2 x = b" unfolding 2 unfolding 3 using 0 by simp then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using b_closed by blast qed qed show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast have "(0::nat)< 2" by presburger hence 0: "x!0 = a!0" unfolding ab_def using a_length by (metis append.simps(2) nth_Cons_0 pair_id) have "(1::nat)< 2" by presburger hence 1: "x!1 = a!1" unfolding ab_def using a_length by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) qed have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) using reverse_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) qed have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}= {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" by blast show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast qed lemma triple_val_ineq_set_semialg': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}" proof- have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule equalityI') show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!0, x!1]" by blast have a_length: "length a = 2" proof- have "a = x!0 #[x!1]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] by presburger qed obtain b where b_def: "b = [x!2]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def 0 A unfolding mem_Collect_eq by (meson Qp_2I cartesian_power_car_memE' rel_simps(49) rel_simps(51) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] by simp have 2: "x = a@b" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of a b] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (a @ b) ! i" apply(cases "i = 0") apply (metis a_def append.simps(2) nth_Cons_0) apply(cases "(i:: nat) = 1") apply (simp add: a_def) proof- assume a: "i \0" "i \ 1" then have "i = 2" using A1 by presburger thus ?thesis by (metis a_length b_def nth_append_length) qed qed have 3: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) apply blast using 3 a_closed apply blast proof- have "drop 2 x = b" unfolding 2 unfolding 3 using 0 by simp then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using b_closed by blast qed qed show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast have "(0::nat)< 2" by presburger hence 0: "x!0 = a!0" unfolding ab_def using a_length by (metis append.simps(2) nth_Cons_0 pair_id) have "(1::nat)< 2" by presburger hence 1: "x!1 = a!1" unfolding ab_def using a_length by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) qed have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ strict_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) using strict_val_relation_set_def apply blast using take_closed[of 1 3 x Q\<^sub>p] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding strict_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding strict_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) qed have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}= {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" by blast show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast qed lemma triple_val_ineq_set_semialg'': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ strict_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) using strict_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq using one_le_numeral apply blast using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding strict_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding strict_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) qed lemma triple_val_ineq_set_semialg''': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) using reverse_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) qed (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Functions\ (**************************************************************************************************) (**************************************************************************************************) text\ The most natural way to define a semialgebraic function $f: \mathbb{Q}_p^n \to \mathbb{Q}_p$ is a function whose graph is a semialgebraic subset of $\mathbb{Q}_p^{n+1}$. However, the definition given here is slightly different, and devised by Denef in \<^cite>\"denef1986"\ in order to prove Macintyre's theorem. As Denef notes, we can use Macintyre's theorem to deduce that the given definition perfectly aligns with the intuitive one. \ (********************************************************************************************) (********************************************************************************************) subsubsection\Defining Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Apply a function f to the tuple consisting of the first n indices, leaving the remaining indices unchanged\ definition partial_image where "partial_image m f xs = (f (take m xs))#(drop m xs)" definition partial_pullback where "partial_pullback m f l S = (partial_image m f) \\<^bsub>m+l\<^esub> S " lemma partial_pullback_memE: assumes "as \ partial_pullback m f l S" shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "partial_image m f as \ S" using assms apply (metis evimage_eq partial_pullback_def) using assms unfolding partial_pullback_def by blast lemma partial_pullback_closed: "partial_pullback m f l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using partial_pullback_memE(1) by blast lemma partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" assumes "(f (take m as))#(drop m as) \ S" shows "as \ partial_pullback m f k S" using assms unfolding partial_pullback_def partial_image_def evimage_def by blast lemma partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" shows "partial_image n f x = (f as)#bs" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(3) cartesian_power_car_memE) have 1: "drop n x = bs" by (metis "0" append_take_drop_id assms(3) same_append_eq) show ?thesis using 0 1 unfolding partial_image_def by blast qed lemma partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" assumes "x \ partial_pullback n f k S" shows "(f as)#bs \ S" using partial_pullback_memE[of x n f k S] partial_image_def[of n f x] by (metis assms(1) assms(2) assms(3) assms(4) partial_image_eq) text\Partial pullbacks have the same algebraic properties as pullbacks\ lemma partial_pullback_intersect: "partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" unfolding partial_pullback_def by simp lemma partial_pullback_union: "partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" unfolding partial_pullback_def by simp lemma cartesian_power_drop: assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "drop n x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" apply(rule cartesian_power_car_memI) using assms cartesian_power_car_memE apply (metis add_diff_cancel_left' length_drop) using assms cartesian_power_car_memE'' by (metis order.trans set_drop_subset) lemma partial_pullback_complement: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (partial_pullback m f l S) " apply(rule equalityI) using partial_pullback_def[of m f l "(carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S)"] partial_pullback_def[of m f l S] apply (smt Diff_iff evimage_Diff partial_pullback_memE(1) subsetI) proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - partial_pullback m f l S" show " x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) " apply(rule partial_pullback_memI) using A apply blast proof have 00: "Suc l = l + 1" by auto have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" by (meson A DiffD1 cartesian_power_drop) have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A by (meson DiffD1 le_add1 take_closed) have "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>l+1\<^esup>) " using assms 0 1 00 cartesian_power_cons[of "drop m x" Q\<^sub>p l "f (take m x)"] by blast thus "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) " using 00 by metis show "f (take m x) # drop m x \ S" using A unfolding partial_pullback_def partial_image_def by blast qed qed lemma partial_pullback_carrier: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" apply(rule equalityI) using partial_pullback_memE(1) apply blast proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" show "x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>))" apply(rule partial_pullback_memI) using A cartesian_power_drop[of x m l] assms apply blast proof- have "f (take m x) \ carrier Q\<^sub>p" using A assms take_closed[of m "m+l" x Q\<^sub>p] by (meson Pi_mem le_add1) then show "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)" using cartesian_power_drop[of x m l] by (metis A add.commute cartesian_power_cons plus_1_eq_Suc) qed qed text\Definition 1.4 from Denef\ definition is_semialg_function where "is_semialg_function m f = ((f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p) \ (\l \ 0. \S \ semialg_sets (1 + l). is_semialgebraic (m + l) (partial_pullback m f l S)))" lemma is_semialg_function_closed: assumes "is_semialg_function m f" shows "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using is_semialg_function_def assms by blast lemma is_semialg_functionE: assumes "is_semialg_function m f" assumes "is_semialgebraic (1 + k) S" shows " is_semialgebraic (m + k) (partial_pullback m f k S)" using is_semialg_function_def assms by (meson is_semialgebraicE le0) lemma is_semialg_functionI: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" shows "is_semialg_function m f" using assms unfolding is_semialg_function_def by blast text\Semialgebraicity for functions can be verified on basic semialgebraic sets \ lemma is_semialg_functionI': assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" shows "is_semialg_function m f" apply(rule is_semialg_functionI) using assms(1) apply blast proof- show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" proof- fix k S assume A: "S \ semialg_sets (1 + k)" show "is_semialgebraic (m + k) (partial_pullback m f k S)" apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" "basic_semialgs (1 + k)"]) using A unfolding semialg_sets_def apply blast using partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger apply (metis assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg partial_pullback_carrier partial_pullback_intersect plus_1_eq_Suc) using partial_pullback_union union_is_semialgebraic apply presburger using assms(1) complement_is_semialg partial_pullback_complement plus_1_eq_Suc by presburger qed qed text\Graphs of semialgebraic functions are semialgebraic\ abbreviation graph where "graph \ fun_graph Q\<^sub>p" lemma graph_memE: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "x \ graph m f" shows "f (take m x) = x!m" "x = (take m x)@[f (take m x)]" "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" proof- obtain a where a_def: "a\carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ x = a @ [f a]" using assms unfolding fun_graph_def by blast then have 0: "a = take m x" by (metis append_eq_conv_conj cartesian_power_car_memE) then show "f (take m x) = x!m" by (metis a_def cartesian_power_car_memE nth_append_length) show "x = (take m x)@[f (take m x)]" using "0" a_def by blast show "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using "0" a_def by blast qed lemma graph_memI: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "f (take m x) = x!m" assumes "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" shows "x \ graph m f" proof- have 0: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule take_closed[of _ "m + 1"]) apply simp using assms(3) by blast have "x = (take m x)@[x!m]" by (metis \take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)\ add.commute assms(3) cartesian_power_car_memE length_append_singleton lessI nth_equalityI nth_take plus_1_eq_Suc take_Suc_conv_app_nth) then have "x = (take m x)@[f (take m x)]" using assms(2) by presburger then show ?thesis using assms 0 unfolding fun_graph_def by blast qed lemma graph_mem_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "x \ graph m f" shows "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" proof(rule cartesian_power_car_memI') show "length x = m + 1" using assms graph_memE[of f m x] by (smt Groups.add_ac(2) cartesian_power_car_memE fun_graph_def length_append_singleton mem_Collect_eq plus_1_eq_Suc) show "\i. i < m + 1 \ x ! i \ carrier Q\<^sub>p" proof- fix i assume A: "i < m + 1" then show "x ! i \ carrier Q\<^sub>p" proof(cases "i = m") case True then show ?thesis using graph_memE[of f m x] by (metis PiE assms(1) assms(2)) next case False then show ?thesis using graph_memE[of f m x] by (metis \i < m + 1\ add.commute assms(1) assms(2) cartesian_power_car_memE' less_SucE nth_take plus_1_eq_Suc) qed qed qed lemma graph_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "graph m f \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" using assms graph_mem_closed by blast text\The \m\-dimensional diagonal set is semialgebraic\ notation diagonal ("\ ") lemma diag_is_algebraic: shows "is_algebraic Q\<^sub>p (n + n) (\ n)" using Qp.cring_axioms diagonal_is_algebraic by blast lemma diag_is_semialgebraic: shows "is_semialgebraic (n + n) (\ n)" using diag_is_algebraic is_algebraic_imp_is_semialg by blast text\Transposition permutations\ definition transpose where "transpose i j = (Fun.swap i j id)" lemma transpose_permutes: assumes "i< n" assumes "j < n" shows "transpose i j permutes {..x. x \ {.. Fun.swap i j id x = x" using assms by (auto simp: Transposition.transpose_def) show "\y. \!x. Fun.swap i j id x = y" proof fix y show "\!x. Fun.swap i j id x = y" using swap_id_eq[of i j y] by (metis eq_id_iff swap_apply(1) swap_apply(2) swap_id_eq swap_self) qed qed lemma transpose_alt_def: "transpose a b x = (if x = a then b else if x = b then a else x)" using swap_id_eq by (simp add: transpose_def) definition last_to_first where "last_to_first n = (\i. if i = (n-1) then 0 else if i < n-1 then i + 1 else i)" definition first_to_last where "first_to_last n = fun_inv (last_to_first n)" lemma last_to_first_permutes: assumes "(n::nat) > 0" shows "last_to_first n permutes {..x. x \ {.. last_to_first n x = x" proof fix x show " x \ {.. last_to_first n x = x" proof assume A: "x \ {.. x < n" by blast then have "x \ n" by linarith then show "last_to_first n x = x" unfolding last_to_first_def using assms by auto qed qed show "\y. \!x. last_to_first n x = y" proof fix y show "\!x. last_to_first n x = y" proof(cases "y = 0") case True then have 0: "last_to_first n (n-1) = y" using last_to_first_def by (simp add: last_to_first_def) have 1: "\x. last_to_first n x = y \ x = n-1" unfolding last_to_first_def using True by (metis add_gr_0 less_numeral_extra(1) not_gr_zero) show ?thesis using 0 1 by blast next case False then show ?thesis proof(cases "y < n") case True then have 0: "last_to_first n (y-1) = y" using False True unfolding last_to_first_def using add.commute by auto have 1: "\x. last_to_first n x = y \ x =(y-1)" unfolding last_to_first_def using True False by auto show ?thesis using 0 1 by blast next case F: False then have 0: "y \ n" using not_less by blast then have 1: "last_to_first n y = y" by (simp add: \\x. x \ {.. last_to_first n x = x\) have 2: "\x. last_to_first n x = y \ x =y" using 0 unfolding last_to_first_def using False by presburger then show ?thesis using 1 2 by blast qed qed qed qed definition graph_swap where "graph_swap n f = permute_list ((first_to_last (n+1))) ` (graph n f)" lemma last_to_first_eq: assumes "length as = n" shows "permute_list (last_to_first (n+1)) (a#as) = (as@[a])" proof- have 0: "\i. i < (n+1) \ permute_list (last_to_first (n + 1)) (a # as) ! i = (as@[a]) ! i" proof- fix i assume A: "i < n+1" show "permute_list (last_to_first (n + 1)) (a # as) ! i = (as @ [a]) ! i" proof(cases "i = n") case True have 0: "(as @ [a]) ! i = a" by (metis True assms nth_append_length) have 1: "length (a#as) = n + 1" by (simp add: assms) have 2: "i < length (a # as)" using "1" A by linarith have 3: "last_to_first (n + 1) permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "a \ carrier Q\<^sub>p" shows "permute_list (first_to_last (n+1)) (as@[a]) = (a#as)" proof- have "length as = n" using assms(1) cartesian_power_car_memE by blast then show ?thesis using last_to_first_eq last_to_first_permutes[of n] permute_list_compose_inv(2)[of "(last_to_first (n + 1))" n "a # as"] unfolding first_to_last_def by (metis add_gr_0 assms(1) assms(2) cartesian_power_append last_to_first_permutes less_one permute_list_closed' permute_list_compose_inv(2)) qed lemma graph_swapI: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "(f as)#as \ graph_swap n f" proof- have 0: "as@[f as] \ graph n f" using assms using graph_memI[of f n] fun_graph_def by blast have 1: "f as \ carrier Q\<^sub>p" using assms by blast then show ?thesis using assms 0 first_to_last_eq[of as "n" "f as"] unfolding graph_swap_def by (metis image_eqI) qed lemma graph_swapE: assumes "x \ graph_swap n f" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "hd x = f (tl x)" proof- obtain y where y_def: "y \ graph n f \ x = permute_list (first_to_last (n+1)) y" using assms graph_swap_def by (smt image_def mem_Collect_eq) then have "take n y \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(2) graph_memE(3) by blast then show "hd x = f (tl x)" by (metis (no_types, lifting) add.commute assms(2) cartesian_power_car_memE' first_to_last_eq graph_memE(1) graph_memE(2) graph_mem_closed lessI list.sel(1) list.sel(3) plus_1_eq_Suc y_def) qed text\Semialgebraic functions have semialgebraic graphs\ lemma graph_as_partial_pullback: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback n f 1 (\ 1) = graph n f" proof show "partial_pullback n f 1 (\ 1) \ graph n f" proof fix x assume A: "x \ partial_pullback n f 1 (\ 1)" then have 0: "f (take n x) # drop n x \ \ 1" by (metis local.partial_image_def partial_pullback_memE(2)) then have 1: "length (f (take n x) # drop n x) = 2" using diagonal_def by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq one_add_one) then obtain b where b_def: "[b] = drop n x" by (metis list.inject pair_id) then have "[f (take n x), b] \ \ 1" using "0" by presburger then have "b = f (take n x)" using 0 by (smt One_nat_def Qp.cring_axioms diagonal_def drop0 drop_Suc_Cons list.inject mem_Collect_eq take_Suc_Cons) then have "x = (take n x)@[f (take n x)]" by (metis append_take_drop_id b_def) then show "x \ graph n f" using graph_memI[of f n x] by (metis (no_types, lifting) A \b = f (take n x)\ assms b_def nth_via_drop partial_pullback_memE(1)) qed show "graph n f \ partial_pullback n f 1 (\ 1)" proof fix x assume A: "x \ graph n f " then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" using assms graph_mem_closed by blast have "x = (take n x) @ [f (take n x)]" using A graph_memE(2)[of f n x] assms by blast then have "partial_image n f x = [f (take n x), f (take n x)]" by (metis append_take_drop_id local.partial_image_def same_append_eq) then have "partial_image n f x \ \ 1" using assms 0 diagonal_def[of 1] Qp.cring_axioms diagonalI[of "partial_image n f x"] by (metis (no_types, lifting) A append_Cons append_eq_conv_conj cartesian_power_car_memE cartesian_power_car_memE' graph_memE(1) less_add_one self_append_conv2 Qp.to_R1_closed) then show "x \ partial_pullback n f 1 (\ 1)" unfolding partial_pullback_def using 0 by blast qed qed lemma semialg_graph: assumes "is_semialg_function n f" shows "is_semialgebraic (n + 1) (graph n f)" using assms graph_as_partial_pullback[of f n] unfolding is_semialg_function_def by (metis diag_is_semialgebraic is_semialgebraicE less_imp_le_nat less_numeral_extra(1)) text\Functions induced by polynomials are semialgebraic\ definition var_list_segment where "var_list_segment i j = map (\i. pvar Q\<^sub>p i) [i..< j]" lemma var_list_segment_length: assumes "i \ j" shows "length (var_list_segment i j) = j - i" using assms var_list_segment_def by fastforce lemma var_list_segment_entry: assumes "k < j - i" assumes "i \ j" shows "var_list_segment i j ! k = pvar Q\<^sub>p (i + k)" using assms var_list_segment_length unfolding var_list_segment_def using nth_map_upt by blast lemma var_list_segment_is_poly_tuple: assumes "i \j" assumes "j \ n" shows "is_poly_tuple n (var_list_segment i j)" apply(rule Qp_is_poly_tupleI) using assms var_list_segment_entry var_list_segment_length Qp.cring_axioms pvar_closed[of _ n] by (metis (no_types, opaque_lifting) add.commute add_lessD1 diff_add_inverse le_Suc_ex less_diff_conv) lemma map_by_var_list_segment: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j \ n" assumes "i \ j" shows "poly_map n (var_list_segment i j) as = list_segment i j as" apply(rule nth_equalityI ) unfolding poly_map_def var_list_segment_def list_segment_def restrict_def poly_tuple_eval_def apply (metis (full_types) assms(1) length_map) using assms eval_pvar[of _ n as] Qp.cring_axioms length_map add.commute length_upt less_diff_conv less_imp_add_positive nth_map nth_upt trans_less_add2 by (smt le_add_diff_inverse2) lemma map_by_var_list_segment_to_length: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i \ n" shows "poly_map n (var_list_segment i n) as = drop i as" apply(rule nth_equalityI ) apply (metis Qp_poly_mapE' assms(1) assms(2) cartesian_power_car_memE length_drop var_list_segment_length) using assms map_by_var_list_segment[of as n n i] list_segment_drop[of i as] cartesian_power_car_memE[of as Q\<^sub>p n] map_nth[of ] nth_drop nth_map[of _ "[i..p)" ] nth_map[of _ "map (pvar Q\<^sub>p) [i..p as"] unfolding poly_map_def poly_tuple_eval_def var_list_segment_def restrict_def list_segment_def by (smt add.commute add_eq_self_zero drop_map drop_upt le_Suc_ex le_refl) lemma map_tail_by_var_list_segment: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "a \ carrier Q\<^sub>p" assumes "i < n" shows "poly_map (n+1) (var_list_segment 1 (n+1)) (a#as) = as" proof- have 0: "(a#as) \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" using assms by (meson cartesian_power_cons) have 1: "length as = n" using assms cartesian_power_car_memE by blast have 2: "drop 1 (a # as) = as" using 0 1 using list_segment_drop[of 1 "a#as"] by (metis One_nat_def drop0 drop_Suc_Cons ) have "1 \n + 1" by auto then show ?thesis using 0 2 map_by_var_list_segment_to_length[of "a#as" "n+1" 1] by presburger qed lemma Qp_poly_tuple_Cons: assumes "is_poly_tuple n fs" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" assumes "k \n" shows "is_poly_tuple n (f#fs)" using is_poly_tuple_Cons[of n fs f] poly_ring_car_mono[of k n] assms by blast lemma poly_map_Cons: assumes "is_poly_tuple n fs" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n (f#fs) a = (Qp_ev f a)#poly_map n fs a" using assms poly_map_cons by blast lemma poly_map_append': assumes "is_poly_tuple n fs" assumes "is_poly_tuple n gs" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n (fs@gs) a = poly_map n fs a @ poly_map n gs a" using assms(3) poly_map_append by blast lemma partial_pullback_by_poly: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" shows "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n+k) S (f# (var_list_segment n (n+k)))" proof show "partial_pullback n (Qp_ev f) k S \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" proof fix x assume A: " x \ partial_pullback n (Qp_ev f) k S" then obtain as bs where as_bs_def: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ x = as @ bs" using partial_pullback_memE(1)[of x n "(Qp_ev f)" k S] cartesian_power_decomp by metis then have 0: "(Qp_ev f as#bs) \ S" using A partial_pullback_memE' by blast have 1: "Qp_ev f as = Qp_ev f (as@bs)" using assms as_bs_def poly_eval_cartesian_prod[of as n bs k f] Qp.cring_axioms [of ] by metis then have 2: "((Qp_ev f x) #bs) \ S" using "0" as_bs_def by presburger have 3: "bs = list_segment n (n+k) x" using as_bs_def list_segment_drop[of n x] by (metis (no_types, lifting) add_cancel_right_right add_diff_cancel_left' append_eq_append_conv append_take_drop_id cartesian_power_car_memE length_0_conv length_append length_map length_upt linorder_neqE_nat list_segment_def not_add_less1) have 4: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" using Qp_poly_tuple_Cons var_list_segment_is_poly_tuple by (metis add.commute assms(1) dual_order.refl le_add2) have 5: "f \ carrier (Q\<^sub>p [\\<^bsub>n + k\<^esub>])" using poly_ring_car_mono[of n "n + k"] assms le_add1 by blast have 6: "is_poly_tuple (n + k) (var_list_segment n (n + k))" by (simp add: var_list_segment_is_poly_tuple) have 7: "x \ carrier (Q\<^sub>p\<^bsup>n + k\<^esup>)" using as_bs_def cartesian_power_concat(1) by blast hence 8: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#poly_map (n+k) (var_list_segment n (n + k)) x" using 5 6 7 A poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 4 unfolding partial_pullback_def evimage_def by blast hence 6: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#bs" using 3 "7" le_add1 le_refl map_by_var_list_segment by presburger show " x \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" unfolding poly_tuple_pullback_def using 6 by (metis "2" "7" IntI poly_map_apply vimage_eq) qed show "poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k)) \ partial_pullback n (Qp_ev f) k S" proof fix x assume A: "x \ poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" have 0: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" using Qp_poly_tuple_Cons assms(1) le_add1 var_list_segment_is_poly_tuple by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using A unfolding poly_tuple_pullback_def by blast have 2: "poly_map (n+k) (f # var_list_segment n (n + k)) x \ S" using 1 assms A unfolding poly_map_def poly_tuple_pullback_def restrict_def by (metis (no_types, opaque_lifting) Int_commute add.commute evimage_def evimage_eq) have 3: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#(drop n x)" using poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 1 assms(1) map_by_var_list_segment_to_length le_add1 poly_map_cons by presburger have 4: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f (take n x))#(drop n x)" using assms 1 3 eval_at_points_higher_pow[of f n "n + k" "x"] le_add1 by (metis nat_le_iff_add) show "x \ partial_pullback n (Qp_ev f) k S" apply(rule partial_pullback_memI) using 1 apply blast using 2 3 4 by metis qed qed lemma poly_is_semialg: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialg_function n (Qp_ev f)" proof(rule is_semialg_functionI) show "Qp_ev f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" using assms by (meson Pi_I eval_at_point_closed) show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" proof- fix k::nat fix S assume A: "S \ semialg_sets (1 + k)" have 0: "is_poly_tuple (n + k) (f # var_list_segment n (n + k))" by (metis add.commute assms le_add2 order_refl Qp_poly_tuple_Cons var_list_segment_is_poly_tuple) have 1: "length (f # var_list_segment n (n + k)) = k + 1" by (metis add.commute add_diff_cancel_left' le_add1 length_Cons plus_1_eq_Suc var_list_segment_length) have 2: "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" using A assms partial_pullback_by_poly[of f n S k] unfolding semialg_sets_def using gen_boolean_algebra_subset by blast then show "is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" using add.commute[of 1 k] 0 1 assms(1) pullback_is_semialg[of "n+k" "(f # var_list_segment n (n + k))" "k+1" S] by (metis A is_semialgebraicI is_semialgebraic_closed poly_tuple_pullback_eq_poly_map_vimage) qed qed text\Families of polynomials defined by semialgebraic coefficient functions\ lemma semialg_function_on_carrier: assumes "is_semialg_function n f" assumes "restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" shows "is_semialg_function n g" proof(rule is_semialg_functionI) have 0: "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" using assms(1) is_semialg_function_closed by blast show "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" then show " g x \ carrier Q\<^sub>p" using assms(2) 0 by (metis (no_types, lifting) PiE restrict_Pi_cancel) qed show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n g k S)" proof- fix k S assume A: "S \ semialg_sets (1 + k)" have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" using A assms(1) is_semialg_functionE is_semialgebraicI by blast have 2: "(partial_pullback n f k S) = (partial_pullback n g k S)" unfolding partial_pullback_def partial_image_def evimage_def proof show "(\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" proof fix x assume "x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) " have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms by (meson \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ inf_le2 le_add1 subset_iff take_closed) then have "f (take n x) = g (take n x)" using assms unfolding restrict_def by meson then show " x \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using assms \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ by blast qed show "(\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" proof fix x assume A: "x \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms by (meson A inf_le2 le_add1 subset_iff take_closed) then have "f (take n x) = g (take n x)" using assms unfolding restrict_def by meson then show "x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using A by blast qed qed then show "is_semialgebraic (n + k) (partial_pullback n g k S)" using 1 by auto qed qed lemma semialg_function_on_carrier': assumes "is_semialg_function n f" assumes "\a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = g a" shows "is_semialg_function n g" using assms semialg_function_on_carrier unfolding restrict_def by (meson restrict_ext semialg_function_on_carrier) lemma constant_function_is_semialg: assumes "n > 0" assumes "x \ carrier Q\<^sub>p" assumes "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = x" shows "is_semialg_function n f" proof(rule semialg_function_on_carrier[of _ "Qp_ev (Qp_to_IP x)"]) show "is_semialg_function n (Qp_ev (Qp_to_IP x))" using assms poly_is_semialg[of "(Qp_to_IP x)"] Qp_to_IP_car by blast have 0: "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = Qp_ev (Qp_to_IP x) a" using eval_at_point_const assms by blast then show "restrict (Qp_ev (Qp_to_IP x)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" by (metis (no_types, lifting) restrict_ext) qed lemma cartesian_product_singleton_factor_projection_is_semialg: assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A {b})" shows "is_semialgebraic m A" proof- obtain f where f_def: "f = map (pvar Q\<^sub>p) [0..p" "[0.. ((nat \ int) \ (nat \ int)) set) list) = map (\i::nat. Qp.indexed_const (b ! i)) [(0::nat)..i. i \ set [0::nat..< n] \ b!i \ carrier Q\<^sub>p" using assms(2) cartesian_power_car_memE'[of b Q\<^sub>p n] by blast hence 1: "\i. i \ set [0::nat..< n] \ Qp.indexed_const (b ! i) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" using assms Qp_to_IP_car by blast show ?thesis unfolding is_poly_tuple_def g_def apply(rule subsetI) using set_map[of "\i. Qp.indexed_const (b ! i)" "[0..x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ poly_tuple_eval (f@g) x = x@b" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" have 30: "poly_tuple_eval f x = x" proof- have 300: "length (poly_tuple_eval f x) = length x" unfolding poly_tuple_eval_def using cartesian_power_car_memE by (metis "4" A length_map) have "\i. i < length x \ poly_tuple_eval f x ! i = x ! i" unfolding f_def poly_tuple_eval_def using nth_map by (metis "4" A add_cancel_right_left cartesian_power_car_memE eval_pvar f_def length_map nth_upt) thus ?thesis using 300 by (metis nth_equalityI) qed have 31: "poly_tuple_eval g x = b" proof- have 310: "length (poly_tuple_eval g x) = length b" unfolding poly_tuple_eval_def g_def using cartesian_power_car_memE by (metis assms(2) length_map map_nth) have 311: "length b = n" using assms cartesian_power_car_memE by blast hence "\i. i < n \ poly_tuple_eval g x ! i = b ! i" proof- fix i assume "i < n" thus "poly_tuple_eval g x ! i = b ! i" unfolding g_def poly_tuple_eval_def using eval_at_point_const[of "b!i" x m] 310 nth_map by (metis "311" A assms(2) cartesian_power_car_memE' length_map map_nth) qed thus ?thesis using 311 310 nth_equalityI by (metis list_eq_iff_nth_eq) qed have 32: "poly_tuple_eval (f @ g) x = poly_map m (f@g) x" unfolding poly_map_def restrict_def using A by (simp add: A) have 33: "poly_tuple_eval f x = poly_map m f x" unfolding poly_map_def restrict_def using A by (simp add: A) have 34: "poly_tuple_eval g x = poly_map m g x" unfolding poly_map_def restrict_def using A by (simp add: A) show "poly_tuple_eval (f @ g) x = x @ b" using assms 1 2 30 31 poly_map_append[of x m f g] A unfolding 32 33 34 by (simp add: A \b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\) qed have 4: "A = (poly_tuple_eval (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b}))" proof show "A \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" proof(rule subsetI) fix x assume A: "x \ A" then have 0: "poly_tuple_eval (f@g) x = x@b" using 3 assms by blast then show " x \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" using A cartesian_product_memE by (smt Un_upper1 assms(1) assms(2) cartesian_product_memI' evimageI2 in_mono insert_is_Un mk_disjoint_insert singletonI) qed show "poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" proof(rule subsetI) fix x assume A: "x \ (poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b})" then have "poly_tuple_eval (f @ g) x \ cartesian_product A {b}" by blast then have "x@b \ cartesian_product A {b}" using A 3 by (metis evimage_eq) then show "x \ A" using A by (metis append_same_eq cartesian_product_memE' singletonD) qed qed have 5: "A = poly_map m (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b})" proof show "A \ poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" unfolding poly_map_def evimage_def restrict_def using 4 by (smt IntI assms(1) evimageD in_mono subsetI vimageI) show "poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" unfolding poly_map_def evimage_def restrict_def using 4 by (smt Int_iff evimageI2 subsetI vimage_eq) qed have 6: "length (f @ g) = m + n" unfolding f_def g_def by (metis index_list_length length_append length_map map_nth) show ?thesis using 2 5 6 assms pullback_is_semialg[of m "f@g" "m+n" "cartesian_product A {b}"] by (metis is_semialgebraicE zero_eq_add_iff_both_eq_0) qed lemma cartesian_product_factor_projection_is_semialg: assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "B \ {}" assumes "is_semialgebraic (m+n) (cartesian_product A B)" shows "is_semialgebraic m A" proof- obtain b where b_def: "b \ B" using assms by blast have "is_semialgebraic n {b}" using assms b_def is_algebraic_imp_is_semialg singleton_is_algebraic by blast hence 0: "is_semialgebraic (m+n) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b})" using car_times_semialg_is_semialg assms(4) by blast have "(cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b}) \ (cartesian_product A B) = (cartesian_product A {b})" using assms b_def cartesian_product_intersection[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m "{b}" n A B] by (metis (no_types, lifting) Int_absorb1 Int_empty_left Int_insert_left_if1 \is_semialgebraic n {b}\ is_semialgebraic_closed set_eq_subset) hence "is_semialgebraic (m+n) (cartesian_product A {b})" using assms 0 intersection_is_semialg by metis thus ?thesis using assms cartesian_product_singleton_factor_projection_is_semialg by (meson \is_semialgebraic n {b}\ insert_subset is_semialgebraic_closed) qed lemma partial_pullback_cartesian_product: assumes "\ \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" shows "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) = partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) " proof show "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" proof fix x assume A: "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain y t where yt_def: "x = y@[t] \ y \ partial_pullback m \ 0 S \ t \ carrier Q\<^sub>p" by (metis cartesian_product_memE' Qp.to_R1_to_R Qp.to_R_pow_closed) then have "[\ y] \ S" using partial_pullback_memE unfolding partial_image_def by (metis (no_types, lifting) add.right_neutral append.right_neutral cartesian_power_drop le_zero_eq take_closed partial_pullback_memE' take_eq_Nil) then have 0: "[\ y]@[t] \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using cartesian_product_memI' yt_def by (metis assms(2) carrier_is_semialgebraic is_semialgebraic_closed Qp.to_R1_closed) have 1: " x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" using A yt_def by (metis add.right_neutral cartesian_power_append partial_pullback_memE(1)) show "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" apply(rule partial_pullback_memI) using "1" apply blast using yt_def 0 by (smt Cons_eq_appendI add.right_neutral local.partial_image_def partial_image_eq partial_pullback_memE(1) self_append_conv2 Qp.to_R1_closed) qed show "partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule subsetI) fix x assume A: "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" using assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] by blast have 1: "\ (take m x) # drop m x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using A assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] unfolding partial_image_def by blast have 2: "\ (take m (take m x)) # drop m (take m x) = [\ (take m x)]" using 0 1 by (metis add.commute add.right_neutral append.right_neutral append_take_drop_id take0 take_drop) show "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p m _ 1]) apply (metis add_cancel_right_right partial_pullback_closed) apply blast apply(rule partial_pullback_memI[of _ m 0 \ S]) using 0 apply (metis Nat.add_0_right le_iff_add take_closed) using 2 apply (metis (no_types, lifting) "1" add.commute add.right_neutral assms(2) cartesian_product_memE(1) list.inject plus_1_eq_Suc take_Suc_Cons take_drop) using 0 cartesian_power_drop by blast qed qed lemma cartesian_product_swap: assumes "A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A B)" shows "is_semialgebraic (m+n) (cartesian_product B A)" proof- obtain f where f_def: "f = (\i. (if i < m then n + i else (if i < m+n then i - m else i)))" by blast have 0: "\i. i \ {.. f i \ {n..i. i \ {m.. f i \ {..i. i \ {.. f i \ {..x. x \ {.. f x = x" unfolding f_def by simp show "\y. \!x. f x = y" proof fix y show "\!x. f x = y" proof(cases "y < n") case True have T0: "f (y+m) = y" unfolding f_def using True by simp have "\i. f i = y \ i \ {m..i. f i = y \ i = y+m" using T0 unfolding f_def by auto thus ?thesis using T0 by blast next case False show ?thesis proof(cases "y \ {n..i. f i = y \ i \ {..i. f i = y \ i = y- n" using f_def by force then show ?thesis using T0 by blast next case F: False then show ?thesis using 0 1 2 unfolding f_def using False add_diff_inverse_nat lessThan_iff by auto qed qed qed qed have "permute_list f ` (cartesian_product A B) = (cartesian_product B A)" proof show "permute_list f ` cartesian_product A B \ cartesian_product B A" proof fix x assume A: " x \ permute_list f ` cartesian_product A B" then obtain a b where ab_def: "a \ A \b \ B \ x = permute_list f (a@b)" by (metis (mono_tags, lifting) cartesian_product_memE' image_iff) have 0: "x = permute_list f (a@b)" using ab_def by blast have 1: "length a = n" using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast have 2: "length b = m" using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast have 3: "length x = m + n" using 1 2 0 f_permutes by simp have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" unfolding 0 using permute_list_nth by (metis "0" "3" f_permutes length_permute_list trans_less_add1) hence 5: "\i. i < m \ x ! i = b!i" unfolding f_def using 1 2 by (metis "4" f_def nth_append_length_plus) have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" unfolding 0 using f_def permute_list_nth f_permutes by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 ordered_cancel_comm_monoid_diff_class.diff_add) have 7: "x = b@a" proof(rule nth_equalityI) show "length x = length (b @ a)" using 1 2 3 by simp show "\i. i < length x \ x ! i = (b @ a) ! i" unfolding 3 using 1 2 4 5 by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) qed show "x \ cartesian_product B A" unfolding 7 using ab_def unfolding cartesian_product_def by blast qed show "cartesian_product B A \ permute_list f ` cartesian_product A B" proof fix y assume A: "y \ cartesian_product B A" then obtain b a where ab_def: "b \ B \ a \ A \ y = b@a" using cartesian_product_memE' by blast obtain x where 0: "x = permute_list f (a@b)" by blast have 1: "length a = n" using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast have 2: "length b = m" using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast have 3: "length x = m + n" using 1 2 0 f_permutes by simp have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" unfolding 0 using permute_list_nth by (metis "0" "3" f_permutes length_permute_list trans_less_add1) hence 5: "\i. i < m \ x ! i = b!i" unfolding f_def using 1 2 by (metis "4" f_def nth_append_length_plus) have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" unfolding 0 using f_def permute_list_nth f_permutes by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 ordered_cancel_comm_monoid_diff_class.diff_add) have 7: "x = b@a" proof(rule nth_equalityI) show "length x = length (b @ a)" using 1 2 3 by simp show "\i. i < length x \ x ! i = (b @ a) ! i" unfolding 3 using 1 2 4 5 by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) qed show "y \ permute_list f ` cartesian_product A B" using ab_def 7 cartesian_product_memI'[of _ Q\<^sub>p] unfolding 0 by (metis assms(1) assms(2) image_eqI) qed qed thus ?thesis using assms f_permutes permutation_is_semialgebraic by metis qed lemma Qp_zero_subset_is_semialg: assumes "S \ carrier (Q\<^sub>p\<^bsup>0\<^esup>)" shows "is_semialgebraic 0 S" proof(cases "S = {}") case True then show ?thesis by (simp add: empty_is_semialgebraic) next case False then have "S = carrier (Q\<^sub>p\<^bsup>0\<^esup>)" using assms unfolding Qp_zero_carrier by blast then show ?thesis by (simp add: carrier_is_semialgebraic) qed lemma cartesian_product_empty_list: "cartesian_product A {[]} = A" "cartesian_product {[]} A = A" proof show "cartesian_product A {[]} \ A" apply(rule subsetI) unfolding cartesian_product_def by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) show "A \ cartesian_product A {[]}" apply(rule subsetI) unfolding cartesian_product_def by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) show "cartesian_product {[]} A = A" proof show "cartesian_product {[]} A \ A" apply(rule subsetI) unfolding cartesian_product_def by (smt append_self_conv2 bex_empty insert_compr mem_Collect_eq) show "A \ cartesian_product {[]} A" apply(rule subsetI) unfolding cartesian_product_def by blast qed qed lemma cartesian_product_singleton_factor_projection_is_semialg': assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A {b})" shows "is_semialgebraic m A" proof(cases "n > 0") case True show ?thesis proof(cases "m > 0") case T: True then show ?thesis using assms True cartesian_product_singleton_factor_projection_is_semialg by blast next case False then show ?thesis using Qp_zero_subset_is_semialg assms by blast qed next case False then have F0: "b = []" using assms Qp_zero_carrier by blast have "cartesian_product A {b} = A" unfolding F0 by (simp add: cartesian_product_empty_list(1)) then show ?thesis using assms False by (metis add.right_neutral gr0I) qed (**************************************************************************************************) (**************************************************************************************************) subsection \More on graphs of functions\ (**************************************************************************************************) (**************************************************************************************************) text\This section lays the groundwork for showing that semialgebraic functions are closed under various algebraic operations\ text\The take and drop functions on lists are polynomial maps\ lemma function_restriction: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ S" assumes "n \ k" shows "(g \ (take n)) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ S" proof fix x assume "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" then have "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(2) take_closed by blast then show "(g \ take n) x \ S" using assms comp_apply by (metis Pi_iff comp_def) qed lemma partial_pullback_restriction: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" shows "partial_pullback k (g \ take n) m S = split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" proof(rule equalityI) show "partial_pullback k (g \ take n) m S \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" proof fix x assume A: "x \ partial_pullback k (g \ take n) m S" obtain as bs where asbs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using partial_pullback_memE[of x k "g \ take n" m S] A cartesian_power_decomp[of x Q\<^sub>p k m] by metis have 0: "((g \ (take n)) as)#bs \ S" using asbs_def partial_pullback_memE'[of as k bs m x] A by blast have 1: "(g (take n as))#bs \ S" using 0 by (metis comp_apply) have 2: "take n as @ bs \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" by (meson asbs_def assms(2) cartesian_power_concat(1) less_imp_le_nat take_closed) have 3: "(take n as)@bs \ (partial_pullback n g m S)" using 1 2 partial_pullback_memI[of "(take n as)@bs" n m g S] by (metis (mono_tags, opaque_lifting) asbs_def assms(2) local.partial_image_def nat_less_le partial_image_eq subsetD subset_refl take_closed) have 4: "drop n as \ (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" using asbs_def assms(2) drop_closed by blast show " x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" using split_cartesian_product_memI[of "take n as" bs "partial_pullback n g m S" "drop n as" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p "n + m" "k - n" n ] 4 by (metis (no_types, lifting) "3" append.assoc append_take_drop_id asbs_def assms(2) cartesian_power_car_memE less_imp_le_nat partial_pullback_memE(1) subsetI take_closed) qed show "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ partial_pullback k (g \ take n) m S" proof fix x assume A: "x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" show "x \ partial_pullback k (g \ take n) m S" proof(rule partial_pullback_memI) have 0: "(partial_pullback n g m S) \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using partial_pullback_closed by blast then have "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ carrier (Q\<^sub>p\<^bsup>n + m + (k - n)\<^esup>)" using assms A split_cartesian_product_closed[of "partial_pullback n g m S" Q\<^sub>p "n + m" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" "k - n" n] using le_add1 by blast then show P: "x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)" by (smt A Nat.add_diff_assoc2 add.commute add_diff_cancel_left' assms(2) le_add1 less_imp_le_nat subsetD) have "take n x @ drop (n + (k - n)) x \ partial_pullback n g m S" using 0 A split_cartesian_product_memE[of x "n + m" "k - n" n "partial_pullback n g m S" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p] le_add1 by blast have 1: "g (take n x) # drop k x \ S" using partial_pullback_memE by (metis (no_types, lifting) \take n x @ drop (n + (k - n)) x \ partial_pullback n g m S\ \x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)\ add.assoc assms(2) cartesian_power_drop le_add1 le_add_diff_inverse less_imp_le_nat partial_pullback_memE' take_closed) have 2: "g (take n x) = (g \ take n) (take k x)" using assms P comp_apply[of g "take n" "take k x"] by (metis add.commute append_same_eq append_take_drop_id less_imp_add_positive take_add take_drop) then show "(g \ take n) (take k x) # drop k x \ S" using "1" by presburger qed qed qed lemma comp_take_is_semialg: assumes "is_semialg_function n g" assumes "n < k" assumes "0 < n" shows "is_semialg_function k (g \ (take n))" proof(rule is_semialg_functionI) show "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" using assms function_restriction[of g n "carrier Q\<^sub>p" k] dual_order.strict_implies_order is_semialg_function_closed by blast show "\ka S. S \ semialg_sets (1 + ka) \ is_semialgebraic (k + ka) (partial_pullback k (g \ take n) ka S)" proof- fix l S assume A: "S \ semialg_sets (1 + l)" have 0: "is_semialgebraic (n + l) (partial_pullback n g l S) " using assms A is_semialg_functionE is_semialgebraicI by blast have "is_semialgebraic (n + l + (k - n)) (split_cartesian_product (n + l) (k - n) n (partial_pullback n g l S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)))" using A 0 split_cartesian_product_is_semialgebraic[of _ _ "partial_pullback n g l S" _ "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)"] add_gr_0 assms(2) assms(3) carrier_is_semialgebraic le_add1 zero_less_diff by presburger then show "is_semialgebraic (k + l) (partial_pullback k (g \ take n) l S)" using partial_pullback_restriction[of g n k l S] by (metis (no_types, lifting) add.assoc add.commute assms(1) assms(2) is_semialg_function_closed le_add_diff_inverse less_imp_le_nat) qed qed text\Restriction of a graph to a semialgebraic domain\ lemma graph_formula: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "graph n g = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). g (take n as) = as!n}" using assms graph_memI fun_graph_def[of Q\<^sub>p n g] by (smt Collect_cong Suc_eq_plus1 graph_memE(1) graph_mem_closed mem_Collect_eq) definition restricted_graph where "restricted_graph n g S = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). take n as \ S \ g (take n as) = as!n }" lemma restricted_graph_closed: "restricted_graph n g S \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" by (metis (no_types, lifting) mem_Collect_eq restricted_graph_def subsetI) lemma restricted_graph_memE: assumes "a \ restricted_graph n g S" shows "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" "take n a \ S" "g (take n a) = a!n" using assms using restricted_graph_closed apply blast apply (metis (no_types, lifting) assms mem_Collect_eq restricted_graph_def) using assms unfolding restricted_graph_def by blast lemma restricted_graph_mem_formula: assumes "a \ restricted_graph n g S" shows "a = (take n a)@[g (take n a)]" proof- have "length a = Suc n" using assms by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq restricted_graph_def) then have "a = (take n a)@[a!n]" by (metis append_eq_append_conv_if hd_drop_conv_nth lessI take_hd_drop) then show ?thesis by (metis assms restricted_graph_memE(3)) qed lemma restricted_graph_memI: assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" assumes "take n a \ S" assumes "g (take n a) = a!n" shows "a \ restricted_graph n g S" using assms restricted_graph_def by blast lemma restricted_graph_memI': assumes "a \ S" assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "(a@[g a]) \ restricted_graph n g S" proof- have "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(1) assms(3) by blast then have "g a \ carrier Q\<^sub>p" using assms by blast then have 0: "a @ [g a] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" using assms by (metis (no_types, lifting) add.commute cartesian_power_append plus_1_eq_Suc subsetD) have 1: "take n (a @ [g a]) \ S" using assms by (metis (no_types, lifting) append_eq_conv_conj cartesian_power_car_memE subsetD) show ?thesis using assms restricted_graph_memI[of "a@[g a]" n S g] by (metis "0" \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ append_eq_conv_conj cartesian_power_car_memE nth_append_length) qed lemma restricted_graph_subset: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S \ graph n g" proof fix x assume A: "x \ restricted_graph n g S" show "x \ graph n g" apply(rule graph_memI) using assms(1) apply blast using A restricted_graph_memE(3) apply blast by (metis A add.commute plus_1_eq_Suc restricted_graph_memE(1)) qed lemma restricted_graph_subset': assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof fix a assume A: "a \ restricted_graph n g S" then have "a = (take n a)@[g (take n a)]" using restricted_graph_mem_formula by blast then show "a \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using cartesian_product_memI' A unfolding restricted_graph_def by (metis (mono_tags, lifting) assms(2) last_closed' mem_Collect_eq subsetI Qp.to_R1_closed) qed lemma restricted_graph_intersection: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" proof show "restricted_graph n g S \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using assms restricted_graph_subset restricted_graph_subset' by (meson Int_subset_iff) show "graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ restricted_graph n g S" proof fix x assume A: " x \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" show "x \ restricted_graph n g S" apply(rule restricted_graph_memI) using A graph_memE[of g n x] apply (metis (no_types, lifting) Int_iff add.commute assms(1) graph_mem_closed plus_1_eq_Suc) using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] using assms(2) apply blast using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] using assms(1) by blast qed qed lemma restricted_graph_is_semialgebraic: assumes "is_semialg_function n g" assumes "is_semialgebraic n S" shows "is_semialgebraic (n+1) (restricted_graph n g S)" proof- have 0: "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" using assms is_semialg_function_closed is_semialgebraic_closed restricted_graph_intersection by presburger have 1: "is_semialgebraic (n + 1) (graph n g)" using assms semialg_graph by blast have 2: "is_semialgebraic (n + 1) (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" using cartesian_product_is_semialgebraic[of n S 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] assms carrier_is_semialgebraic less_one by presburger then show ?thesis using 0 1 2 intersection_is_semialg[of "n+1" "graph n g" "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] by presburger qed lemma take_closed: assumes "n \ k" assumes "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" shows "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms take_closed by blast lemma take_compose_closed: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" shows "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" then have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms less_imp_le_nat take_closed by blast then have "g (take n x) \ carrier Q\<^sub>p" using assms(1) by blast then show "(g \ take n) x \ carrier Q\<^sub>p" using comp_apply[of g "take n" x] by presburger qed lemma take_graph_formula: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" assumes "0 < n" shows "graph k (g \ (take n)) = {as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>). g (take n as) = as!k}" proof- have "\as. as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>) \ (g \ take n) (take k as) = g (take n as) " using assms comp_apply take_take[of n k] proof - fix as :: "((nat \ int) \ (nat \ int)) set list" show "(g \ take n) (take k as) = g (take n as)" by (metis (no_types) \n < k\ comp_eq_dest_lhs min.strict_order_iff take_take) qed then show ?thesis using take_compose_closed[of g n k] assms comp_apply[of g "take n"] graph_formula[of "g \ (take n)" k] by (smt Collect_cong Suc_eq_plus1) qed lemma graph_memI': assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" assumes "take n a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "g (take n a) = a!n" shows "a \ graph n g" using assms fun_graph_def[of Q\<^sub>p n g] by (smt cartesian_power_car_memE eq_imp_le lessI mem_Collect_eq take_Suc_conv_app_nth take_all) lemma graph_memI'': assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "(a@[g a]) \ graph n g " using assms fun_graph_def by blast lemma graph_as_restricted_graph: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "graph n f = restricted_graph n f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" apply(rule equalityI) apply (metis Suc_eq_plus1 assms graph_memE(1) graph_memE(3) graph_mem_closed restricted_graph_memI subsetI) by (simp add: assms restricted_graph_subset) definition double_graph where "double_graph n f g = {as \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>). f (take n as) = as!n \ g (take n as) = as!(n + 1)}" lemma double_graph_rep: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "double_graph n f g = restricted_graph (n + 1) (g \ take n) (graph n f)" proof show "double_graph n f g \ restricted_graph (n + 1) (g \ take n) (graph n f)" proof fix x assume A: "x \ double_graph n f g" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>) \ f (take n x) = x!n \ g (take n x) = x!(n + 1)" using double_graph_def by blast have 1: "take (n+1) x \ graph n f" apply(rule graph_memI) using assms(2) apply blast apply (metis "0" append_eq_conv_conj cartesian_power_car_memE le_add1 length_take less_add_same_cancel1 less_numeral_extra(1) min.absorb2 nth_take take_add) by (metis (no_types, opaque_lifting) "0" Suc_eq_plus1 Suc_n_not_le_n add_cancel_right_right dual_order.antisym le_iff_add not_less_eq_eq one_add_one plus_1_eq_Suc take_closed) show " x \ restricted_graph (n + 1) (g \ take n) (graph n f)" apply(rule restricted_graph_memI) apply (metis "0" One_nat_def add_Suc_right numeral_2_eq_2) using "1" apply blast using 0 take_take[of n "n + 1" x] comp_apply by (metis le_add1 min.absorb1) qed show "restricted_graph (n + 1) (g \ take n) (graph n f) \ double_graph n f g" proof fix x assume A: "x \ restricted_graph (n + 1) (g \ take n) (graph n f)" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc (n + 1)\<^esup>) \ take (n + 1) x \ graph n f \ (g \ take n) (take (n + 1) x) = x ! (n + 1)" using restricted_graph_memE[of x "n+1" "(g \ take n)" "graph n f" ] by blast then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" using 0 by (metis Suc_1 add_Suc_right) have 2: " f (take n x) = x ! n" using 0 take_take[of n "n + 1" x] graph_memE[of f n "take (n + 1) x"] by (metis assms(2) le_add1 less_add_same_cancel1 less_numeral_extra(1) min.absorb1 nth_take) have 3: "g (take n x) = x ! (n + 1)" using 0 comp_apply take_take[of n "n + 1" x] by (metis le_add1 min.absorb1) then show "x \ double_graph n f g" unfolding double_graph_def using 1 2 3 by blast qed qed lemma double_graph_is_semialg: assumes "n > 0" assumes "is_semialg_function n f" assumes "is_semialg_function n g" shows "is_semialgebraic (n+2) (double_graph n f g)" using double_graph_rep[of g n f] assms restricted_graph_is_semialgebraic[of n "g \ take n" "graph n f"] by (metis (no_types, lifting) Suc_eq_plus1 add_Suc_right is_semialg_function_closed less_add_same_cancel1 less_numeral_extra(1) one_add_one restricted_graph_is_semialgebraic comp_take_is_semialg semialg_graph) definition add_vars :: "nat \ nat \ padic_tuple \ padic_number" where "add_vars i j as = as!i \\<^bsub>Q\<^sub>p\<^esub> as!j" lemma add_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" assumes "j < n" shows "add_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" unfolding add_vars_def using assms eval_at_point_add[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] eval_pvar by (metis pvar_closed) lemma add_vars_is_semialg: assumes "i < n" assumes "j < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (add_vars i j)" proof- have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] by blast then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] by blast then show ?thesis using assms add_vars_rep semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "add_vars i j" ] by (metis (no_types, lifting) restrict_ext) qed definition mult_vars :: "nat \ nat \ padic_tuple \ padic_number" where "mult_vars i j as = as!i \ as!j" lemma mult_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" assumes "j < n" shows "mult_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" unfolding mult_vars_def using assms eval_at_point_mult[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] eval_pvar[of i n as] eval_pvar[of j n as ] by (metis pvar_closed) lemma mult_vars_is_semialg: assumes "i < n" assumes "j < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (mult_vars i j)" proof- have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] by blast then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] by blast then show ?thesis using assms mult_vars_rep semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "mult_vars i j" ] by (metis (no_types, lifting) restrict_ext) qed definition minus_vars :: "nat \ padic_tuple \ padic_number" where "minus_vars i as = \\<^bsub>Q\<^sub>p\<^esub> as!i" lemma minus_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" shows "minus_vars i as = Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)) as" unfolding minus_vars_def using assms eval_pvar[of i n as] eval_at_point_a_inv[of as n "pvar Q\<^sub>p i"] by (metis pvar_closed) lemma minus_vars_is_semialg: assumes "i < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (minus_vars i)" proof- have 0: "pvar Q\<^sub>p i \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] Qp.cring_axioms by presburger have "is_semialg_function n (Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)))" apply(rule poly_is_semialg ) using "0" by blast then show ?thesis using assms minus_vars_rep[of a i n] semialg_function_on_carrier[of n _ "minus_vars i" ] by (metis (no_types, lifting) minus_vars_rep restrict_ext) qed definition extended_graph where "extended_graph n f g h = {as \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>). f (take n as) = as!n \ g (take n as) = as! (n + 1) \ h [(f (take n as)),(g (take n as))] = as! (n + 2) }" lemma extended_graph_rep: "extended_graph n f g h = restricted_graph (n + 2) (h \ (drop n)) (double_graph n f g)" proof show "extended_graph n f g h \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" proof fix x assume "x \ extended_graph n f g h" then have A: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \f (take n x) = x!n \ g (take n x) = x! (n + 1) \ h [(f (take n x)),(g (take n x))] = x! (n + 2)" unfolding extended_graph_def by blast then have 0: "take (n + 2) x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" proof - have "Suc (Suc n) \ n + numeral (num.One + num.Bit0 num.One)" by simp then show ?thesis by (metis (no_types) \x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \ f (take n x) = x ! n \ g (take n x) = x ! (n + 1) \ h [f (take n x), g (take n x)] = x ! (n + 2)\ add_2_eq_Suc' add_One_commute semiring_norm(5) take_closed) qed have 1: "f (take n (take (n + 2) x)) = (take (n + 2) x) ! n" using A by (metis Suc_1 add.commute append_same_eq append_take_drop_id less_add_same_cancel1 nth_take take_add take_drop zero_less_Suc) have 2: " g (take n (take (n + 2) x)) = (take (n + 2) x) ! (n + 1)" using A by (smt add.assoc add.commute append_same_eq append_take_drop_id less_add_same_cancel1 less_numeral_extra(1) nth_take one_add_one take_add take_drop) then have 3: "take (n + 2) x \ double_graph n f g" unfolding double_graph_def using 0 1 2 by blast have 4: "drop n (take (n + 2) x) = [(f (take n x)),(g (take n x))]" proof- have 40: "take (n + 2) x ! (n + 1) = x! (n + 1)" by (metis add.commute add_2_eq_Suc' lessI nth_take plus_1_eq_Suc) have 41: "take (n + 2) x ! n = x! n" by (metis Suc_1 less_SucI less_add_same_cancel1 less_numeral_extra(1) nth_take) have 42: "take (n + 2) x ! (n + 1) = g (take n x)" using 40 A by blast have 43: "take (n + 2) x ! n = f (take n x)" using 41 A by blast show ?thesis using A 42 43 by (metis "0" add_cancel_right_right cartesian_power_car_memE cartesian_power_drop le_add_same_cancel1 nth_drop pair_id zero_le_numeral) qed then have 5: "(h \ drop n) (take (n + 2) x) = x ! (n + 2)" using 3 A by (metis add_2_eq_Suc' comp_eq_dest_lhs) show "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" using restricted_graph_def A 3 5 by (metis (no_types, lifting) One_nat_def Suc_1 add_Suc_right numeral_3_eq_3 restricted_graph_memI) qed show "restricted_graph (n + 2) (h \ drop n) (double_graph n f g) \ extended_graph n f g h" proof fix x assume A: "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" then have 0: "take (n+2) x \ double_graph n f g" using restricted_graph_memE(2) by blast have 1: "(h \ drop n) (take (n+2) x) = x! (n+2) " by (meson A restricted_graph_memE(3) padic_fields_axioms) have 2: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>)" using A by (metis (no_types, opaque_lifting) Suc3_eq_add_3 add.commute add_2_eq_Suc' restricted_graph_closed subsetD) have 3: "length x = n + 3" using "2" cartesian_power_car_memE by blast have 4: "drop n (take (n+2) x) = [x!n, x!(n+1)]" proof- have "length (take (n+2) x) = n+2" by (simp add: "3") then have 40:"length (drop n (take (n+2) x)) = 2" by (metis add_2_eq_Suc' add_diff_cancel_left' length_drop) have 41: "(drop n (take (n+2) x))!0 = x!n" using 3 by (metis Nat.add_0_right \length (take (n + 2) x) = n + 2\ add_gr_0 le_add1 less_add_same_cancel1 less_numeral_extra(1) nth_drop nth_take one_add_one) have 42: "(drop n (take (n+2) x))!1 = x!(n+1)" using 3 nth_take nth_drop A by (metis add.commute le_add1 less_add_same_cancel1 less_numeral_extra(1) one_add_one take_drop) show ?thesis using 40 41 42 by (metis pair_id) qed have "(take n x) = take n (take (n+2) x)" using take_take 3 by (metis le_add1 min.absorb1) then have 5: "f (take n x) = x ! n" using 0 double_graph_def[of n f g] 3 by (smt Suc_1 less_add_same_cancel1 mem_Collect_eq nth_take zero_less_Suc) have 6: "g (take n x) = x ! (n + 1) " using 0 double_graph_def[of n f g] 3 take_take[of n "n+2" x] by (smt Suc_1 \take n x = take n (take (n + 2) x)\ add_Suc_right lessI mem_Collect_eq nth_take) have 7: " h [f (take n x), g (take n x)] = x ! (n + 2)" using 4 A comp_apply by (metis "1" "5" "6") show " x \ extended_graph n f g h" unfolding extended_graph_def using 2 5 6 7 A by blast qed qed lemma function_tuple_eval_closed: assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "function_tuple_eval Q\<^sub>p n fs x \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" using function_tuple_eval_closed[of Q\<^sub>p n fs x] assms by blast definition k_graph where "k_graph n fs = {x \ carrier (Q\<^sub>p\<^bsup>n + length fs\<^esup>). x = (take n x)@ (function_tuple_eval Q\<^sub>p n fs (take n x)) }" lemma k_graph_memI: assumes "is_function_tuple Q\<^sub>p n fs" assumes "x = as@function_tuple_eval Q\<^sub>p n fs as" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "x \ k_graph n fs" proof- have "take n x = as" using assms by (metis append_eq_conv_conj cartesian_power_car_memE) then show ?thesis unfolding k_graph_def using assms by (smt append_eq_conv_conj cartesian_power_car_memE cartesian_power_car_memI'' length_append local.function_tuple_eval_closed mem_Collect_eq) qed text\composing a function with a function tuple\ lemma Qp_function_tuple_comp_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "length fs = n" assumes "is_function_tuple Q\<^sub>p m fs" shows "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms function_tuple_comp_closed by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Tuples of Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Predicate for a tuple of semialgebraic functions\ definition is_semialg_function_tuple where "is_semialg_function_tuple n fs = (\ f \ set fs. is_semialg_function n f)" lemma is_semialg_function_tupleI: assumes "\ f. f \ set fs \ is_semialg_function n f" shows "is_semialg_function_tuple n fs" using assms is_semialg_function_tuple_def by blast lemma is_semialg_function_tupleE: assumes "is_semialg_function_tuple n fs" assumes "i < length fs" shows "is_semialg_function n (fs ! i)" by (meson assms(1) assms(2) in_set_conv_nth is_semialg_function_tuple_def padic_fields_axioms) lemma is_semialg_function_tupleE': assumes "is_semialg_function_tuple n fs" assumes "f \ set fs" shows "is_semialg_function n f" using assms(1) assms(2) is_semialg_function_tuple_def by blast lemma semialg_function_tuple_is_function_tuple: assumes "is_semialg_function_tuple n fs" shows "is_function_tuple Q\<^sub>p n fs" apply(rule is_function_tupleI) using assms is_semialg_function_closed is_semialg_function_tupleE' by blast lemma const_poly_function_tuple_comp_is_semialg: assumes "n > 0" assumes "is_semialg_function_tuple n fs" assumes "a \ carrier Q\<^sub>p" shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (Qp_to_IP a))" apply(rule semialg_function_on_carrier[of n "Qp_ev (Qp_to_IP a)"]) using poly_is_semialg[of "(Qp_to_IP a)"] using assms(1) assms(3) Qp_to_IP_car apply blast using poly_function_tuple_comp_eq[of n fs "(Qp_to_IP a)"] assms unfolding restrict_def by (metis (no_types, opaque_lifting) eval_at_point_const poly_function_tuple_comp_constant semialg_function_tuple_is_function_tuple) lemma pvar_poly_function_tuple_comp_is_semialg: assumes "n > 0" assumes "is_semialg_function_tuple n fs" assumes "i < length fs" shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (pvar Q\<^sub>p i))" apply(rule semialg_function_on_carrier[of n "fs!i"]) using assms(2) assms(3) is_semialg_function_tupleE apply blast by (metis assms(2) assms(3) poly_function_tuple_comp_pvar restrict_ext semialg_function_tuple_is_function_tuple) text\Polynomial functions with semialgebraic coefficients\ definition point_to_univ_poly :: "nat \ padic_tuple \ padic_univ_poly" where "point_to_univ_poly n a = ring_cfs_to_univ_poly n a" definition tuple_partial_image where "tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs (take m x))@(drop m x)" lemma tuple_partial_image_closed: assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "tuple_partial_image n fs x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using assms unfolding tuple_partial_image_def by (meson cartesian_power_concat(1) cartesian_power_drop function_tuple_eval_closed le_add1 take_closed) lemma tuple_partial_image_indices: assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" assumes "i < length fs" shows "(tuple_partial_image n fs x) ! i = (fs!i) (take n x)" proof- have 0: "(function_tuple_eval Q\<^sub>p n fs (take n x)) ! i = (fs!i) (take n x)" using assms unfolding function_tuple_eval_def by (meson nth_map) have 1: "length (function_tuple_eval Q\<^sub>p n fs (take n x)) > i" by (metis assms(4) function_tuple_eval_def length_map) show ?thesis using 0 1 assms unfolding tuple_partial_image_def by (metis nth_append) qed lemma tuple_partial_image_indices': assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" assumes "i < l" shows "(tuple_partial_image n fs x) ! (length fs + i) = x!(n + i)" using assms unfolding tuple_partial_image_def by (metis (no_types, lifting) cartesian_power_car_memE function_tuple_eval_closed le_add1 nth_append_length_plus nth_drop take_closed) definition tuple_partial_pullback where "tuple_partial_pullback n fs l S = ((tuple_partial_image n fs)-`S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" lemma tuple_partial_pullback_memE: assumes "as \ tuple_partial_pullback m fs l S" shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "tuple_partial_image m fs as \ S" using assms apply (metis (no_types, opaque_lifting) Int_iff add.commute tuple_partial_pullback_def) using assms unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_closed: "tuple_partial_pullback m fs l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using tuple_partial_pullback_memE by blast lemma tuple_partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" assumes "is_function_tuple Q\<^sub>p m fs" assumes "((function_tuple_eval Q\<^sub>p m fs) (take m as))@(drop m as) \ S" shows "as \ tuple_partial_pullback m fs k S" using assms unfolding tuple_partial_pullback_def tuple_partial_image_def by blast lemma tuple_partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" shows "tuple_partial_image n fs x = ((function_tuple_eval Q\<^sub>p n fs) as)@bs" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(3) cartesian_power_car_memE) have 1: "drop n x = bs" by (metis "0" append_take_drop_id assms(3) same_append_eq) show ?thesis using assms 0 1 unfolding tuple_partial_image_def by presburger qed lemma tuple_partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" assumes "x \ tuple_partial_pullback n fs k S" shows "(function_tuple_eval Q\<^sub>p n fs as)@bs \ S" using tuple_partial_pullback_memE[of x n fs k S] tuple_partial_image_def[of n fs x] by (metis assms(1) assms(2) assms(3) assms(4) tuple_partial_image_eq) text\tuple partial pullbacks have the same algebraic properties as pullbacks\ lemma tuple_partial_pullback_intersect: "tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_union: "tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_complement: assumes "is_function_tuple Q\<^sub>p m fs" shows "tuple_partial_pullback m fs l ((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (tuple_partial_pullback m fs l S) " apply(rule equalityI) using tuple_partial_pullback_def[of m fs l "((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S)"] tuple_partial_pullback_def[of m fs l S] apply blast proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - tuple_partial_pullback m fs l S" show " x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>) - S) " apply(rule tuple_partial_pullback_memI) using A apply blast using assms apply blast proof have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" by (meson A DiffD1 cartesian_power_drop) have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A by (meson DiffD1 le_add1 take_closed) show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using 0 1 assms using cartesian_power_concat(1) function_tuple_eval_closed by blast show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ S" using A unfolding tuple_partial_pullback_def tuple_partial_image_def by blast qed qed lemma tuple_partial_pullback_carrier: assumes "is_function_tuple Q\<^sub>p m fs" shows "tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" apply(rule equalityI) using tuple_partial_pullback_memE(1) apply blast proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" show "x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>))" apply(rule tuple_partial_pullback_memI) using A cartesian_power_drop[of x m l] take_closed assms apply blast using assms apply blast proof- have "function_tuple_eval Q\<^sub>p m fs (take m x) \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" using A take_closed assms function_tuple_eval_closed le_add1 by blast then show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using cartesian_power_drop[of x m l] A cartesian_power_concat(1) by blast qed qed definition is_semialg_map_tuple where "is_semialg_map_tuple m fs = (is_function_tuple Q\<^sub>p m fs \ (\l \ 0. \S \ semialg_sets ((length fs) + l). is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)))" lemma is_semialg_map_tuple_closed: assumes "is_semialg_map_tuple m fs" shows "is_function_tuple Q\<^sub>p m fs" using is_semialg_map_tuple_def assms by blast lemma is_semialg_map_tupleE: assumes "is_semialg_map_tuple m fs" assumes "is_semialgebraic ((length fs) + l) S" shows " is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)" using is_semialg_map_tuple_def[of m fs] assms is_semialgebraicE[of "((length fs) + l)" S] by blast lemma is_semialg_map_tupleI: assumes "is_function_tuple Q\<^sub>p m fs" assumes "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" shows "is_semialg_map_tuple m fs" using assms unfolding is_semialg_map_tuple_def by blast text\Semialgebraicity for maps can be verified on basic semialgebraic sets\ lemma is_semialg_map_tupleI': assumes "is_function_tuple Q\<^sub>p m fs" assumes "\k S. S \ basic_semialgs ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" shows "is_semialg_map_tuple m fs" apply(rule is_semialg_map_tupleI) using assms(1) apply blast proof- show "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- fix k S assume A: "S \ semialg_sets ((length fs) + k)" show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>length fs + k\<^esup>)" "basic_semialgs ((length fs) + k)"]) using A unfolding semialg_sets_def apply blast using tuple_partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger using assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg tuple_partial_pullback_carrier tuple_partial_pullback_intersect apply presburger using tuple_partial_pullback_union union_is_semialgebraic apply presburger using assms(1) complement_is_semialg tuple_partial_pullback_complement plus_1_eq_Suc by presburger qed qed text\ The goal of this section is to show that tuples of semialgebraic functions are semialgebraic maps. \ text\The function $(x_0, x, y) \mapsto (x_0, f(x), y)$\ definition twisted_partial_image where "twisted_partial_image n m f xs = (take n xs)@ partial_image m f (drop n xs)" text\The set ${(x_0, x, y) \mid (x_0, f(x), y) \in S}$\ text\Convention: a function which produces a subset of (Qp (i + j +k)) will receive the 3 arity parameters in sequence, at the very beginning of the function\ definition twisted_partial_pullback where "twisted_partial_pullback n m l f S = ((twisted_partial_image n m f)-`S) \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" lemma twisted_partial_pullback_memE: assumes "as \ twisted_partial_pullback n m l f S" shows "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" "twisted_partial_image n m f as \ S" using assms apply (metis (no_types, opaque_lifting) Int_iff add.commute twisted_partial_pullback_def subset_iff) using assms unfolding twisted_partial_pullback_def by blast lemma twisted_partial_pullback_closed: "twisted_partial_pullback n m l f S \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" using twisted_partial_pullback_memE(1) by blast lemma twisted_partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" assumes "(take n as)@((f (take m (drop n as)))#(drop (n + m) as)) \ S" shows "as \ twisted_partial_pullback n m l f S" using assms unfolding twisted_partial_pullback_def twisted_partial_image_def by (metis (no_types, lifting) IntI add.commute drop_drop local.partial_image_def vimageI) lemma twisted_partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" assumes "x = as @ bs @ cs" shows "twisted_partial_image n m f x = as@((f bs)#cs)" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(4) cartesian_power_car_memE) have 1: "twisted_partial_image n m f x = as@(partial_image m f (bs@cs))" using 0 assms twisted_partial_image_def by (metis append_eq_conv_conj cartesian_power_car_memE) have 2: "(partial_image m f (bs@cs)) = (f bs)#cs" using partial_image_eq[of bs m cs l "bs@cs" f] assms by blast show ?thesis using assms 0 1 2 by (metis ) qed lemma twisted_partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" assumes "x = as @ bs @ cs" assumes "x \ twisted_partial_pullback n m l f S" shows "as@((f bs)#cs) \ S" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) twisted_partial_image_eq twisted_partial_pullback_memE(2)) text\partial pullbacks have the same algebraic properties as pullbacks\ text\permutation which moves the entry at index \i\ to 0\ definition twisting_permutation where "twisting_permutation (i::nat) = (\j. if j < i then j + 1 else (if j = i then 0 else j))" lemma twisting_permutation_permutes: assumes "i < n" shows "twisting_permutation i permutes {..x. x > i \ twisting_permutation i x = x" unfolding twisting_permutation_def by auto have 1: "(\x. x \ {.. twisting_permutation i x = x)" using 0 assms by auto have 2: "(\y. \!x. twisting_permutation i x = y)" proof fix y show " \!x. twisting_permutation i x = y" proof(cases "y = 0") case True show "\!x. twisting_permutation i x = y" by (metis Suc_eq_plus1 True add_eq_0_iff_both_eq_0 less_nat_zero_code nat_neq_iff twisting_permutation_def zero_neq_one) next case False show ?thesis proof(cases "y \i") case True show ?thesis proof show "twisting_permutation i (y - 1) = y" using True by (metis False add.commute add_diff_inverse_nat diff_less gr_zeroI le_eq_less_or_eq less_imp_diff_less less_one twisting_permutation_def) show "\x. twisting_permutation i x = y \ x = y - 1" using True False twisting_permutation_def by force qed next - case F: False - then show ?thesis using False - unfolding twisting_permutation_def - by (metis add_leD1 add_leD2 add_le_same_cancel2 discrete le_numeral_extra(3) - less_imp_not_less ) + case False + then show ?thesis + by (auto simp add: twisting_permutation_def) qed qed qed show ?thesis using 1 2 by (simp add: permutes_def) qed lemma twisting_permutation_action: assumes "length as = i" shows "permute_list (twisting_permutation i) (b#(as@bs)) = as@(b#bs)" proof- have 0: "length (permute_list (twisting_permutation i) (b#(as@bs))) = length (as@(b#bs))" by (metis add.assoc length_append length_permute_list list.size(4)) have "\j. j < length (as@(b#bs)) \ (permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" proof- fix j assume A: "j < length (as@(b#bs))" show "(permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" proof(cases "j < i") case True then have T0: "twisting_permutation i j = j + 1" using twisting_permutation_def by auto then have T1: "(b # as @ bs) ! twisting_permutation i j = as!j" using assms by (simp add: assms True nth_append) have T2: "(permute_list (twisting_permutation i) (b # as @ bs)) ! j = as!j" proof- have "twisting_permutation i permutes {..length as = i\) then have "permute_list (fun_inv TI) (as@(b#bs)) = permute_list ((fun_inv TI) \ TI) (b#(as@bs))" using 0 1 by (metis TI_def fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv length_permute_list permute_list_compose) then show ?thesis by (metis "0" Nil_is_append_conv TI_def fun_inv_permute(3) length_greater_0_conv list.distinct(1) permute_list_id) qed lemma twisting_semialg: assumes "is_semialgebraic n S" assumes "n > i" shows "is_semialgebraic n ((permute_list ((twisting_permutation i)) ` S))" proof- obtain TI where TI_def: "TI = twisting_permutation i" by blast have 0: "TI permutes {..<(n::nat)}" using assms TI_def twisting_permutation_permutes[of i n] by blast have "(TI) permutes {.. i" shows "is_semialgebraic n ((permute_list (fun_inv (twisting_permutation i)) ` S))" proof- obtain TI where TI_def: "TI = twisting_permutation i" by blast have 0: "TI permutes {..<(n::nat)}" using assms TI_def twisting_permutation_permutes[of i n] by blast have "(fun_inv TI) permutes {..Defining a permutation that does: $(x0, x1, y) \mapsto (x_1, x0, y)$\ definition tp_1 where "tp_1 i j = (\ n. (if n n \ n < i + j then n - i else n)))" lemma permutes_I: assumes "\x. x \ S \ f x = x" assumes "\y. y \ S \ \!x \ S. f x = y" assumes "\x. x \ S \ f x \ S" shows "f permutes S" proof- have 0 : "(\x. x \ S \ f x = x) " using assms(1) by blast have 1: "(\y. \!x. f x = y)" proof fix y show "\!x. f x = y" apply(cases "y \ S") apply (metis "0" assms(2)) proof- assume "y \ S" then show "\!x. f x = y" by (metis assms(1) assms(3)) qed qed show ?thesis using assms 1 unfolding permutes_def by blast qed lemma tp_1_permutes: "(tp_1 (i::nat) j) permutes {..< i + j}" proof(rule permutes_I) show "\x. x \ {.. tp_1 i j x = x" proof- fix x assume A: "x \ {..y. y \ {.. \!x. x \ {.. tp_1 i j x = y" proof- fix y assume A: "y \ {..!x. x \ {.. tp_1 i j x = y" proof(cases "y < j") case True then have 0:"tp_1 i j (y + i) = y" by (simp add: tp_1_def) have 1: "\x. x \ y + i \ tp_1 i j x \ y" proof- fix x assume A: " x \ y + i" show "tp_1 i j x \ y" apply(cases "x < j") apply (metis A True add.commute le_add_diff_inverse le_eq_less_or_eq nat_neq_iff not_add_less1 tp_1_def trans_less_add2) by (metis A True add.commute le_add_diff_inverse less_not_refl tp_1_def trans_less_add1) qed show ?thesis using 0 1 by (metis A \\x. x \ {.. tp_1 i j x = x\) next case False then have "y - j < i" using A by auto then have "tp_1 i j (y - j) = y" using False tp_1_def by (simp add: tp_1_def) then show ?thesis by (smt A False \\x. x \ {.. tp_1 i j x = x\ add.commute add_diff_inverse_nat add_left_imp_eq less_diff_conv2 not_less tp_1_def padic_fields_axioms) qed qed show "\x. x \ {.. tp_1 i j x \ {.. {.. carrier (Q\<^sub>p\<^bsup>i\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>j\<^esup>)" assumes "c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "permute_list (tp_1 i j) (b@a@c)= a@b@c" proof- have 0:"length (permute_list (tp_1 i j) (b@a@c))= length (a@b@c)" by (metis add.commute append.assoc length_append length_permute_list) have "\x. x < length (a@b@c) \ (permute_list (tp_1 i j) (b@a@c)) ! x= (a@b@c) ! x" proof- fix x assume A: "x < length (a@b@c)" have B: "length (a @ b @ c) = i + j + length c" using add.assoc assms(1) assms(2) cartesian_power_car_memE length_append by metis have C: "tp_1 i j permutes {..length (permute_list (tp_1 i j) (b @ a @ c)) = length (a @ b @ c)\ length_permute_list nth_append permute_list_nth) next case False show ?thesis proof(cases "x < i + j") case True then have T0: "(tp_1 i j x) = x - i" by (meson False not_less tp_1_def) have "x - i < length b" using E False True by linarith then have T1: "permute_list (tp_1 i j) (b@ a @ c) ! x = b!(x-i)" using nth_append by (metis A C T0 \length (permute_list (tp_1 i j) (b @ a @ c)) = length (a @ b @ c)\ length_permute_list permute_list_nth) then show ?thesis by (metis D False \x - i < length b\ nth_append) next case False then have "(tp_1 i j x) = x" by (meson tp_1_def trans_less_add1) then show ?thesis by (smt A C D E False add.commute add_diff_inverse_nat append.assoc length_append nth_append_length_plus permute_list_nth) qed qed qed then show ?thesis using 0 by (metis list_eq_iff_nth_eq) qed definition tw where "tw i j = permute_list (tp_1 j i)" lemma tw_is_semialg: assumes "n > 0" assumes "is_semialgebraic n S" assumes "n \ i + j" shows "is_semialgebraic n ((tw i j)`S)" unfolding tw_def using assms tp_1_permutes'[of j i "n - (j + i)"] permutation_is_semialgebraic[of n S] by (metis add.commute le_add_diff_inverse) lemma twisted_partial_pullback_factored: assumes "f \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n+1+ l\<^esup>)" assumes "Y = partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" shows "twisted_partial_pullback n m l f S = (tw m n) ` Y" proof show "twisted_partial_pullback n m l f S \ tw m n ` Y" proof fix x assume A: "x \ twisted_partial_pullback n m l f S" then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" using twisted_partial_pullback_memE(1) by blast obtain a where a_def: "a = take n x" by blast obtain b where b_def: "b = take m (drop n x)" by blast obtain c where c_def: "c = (drop (n + m) x)" by blast have x_eq:"x = a@(b@c)" by (metis a_def append.assoc append_take_drop_id b_def c_def take_add) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis (no_types, lifting) a_def dual_order.trans le_add1 take_closed x_closed) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" proof- have "drop n x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" by (metis (no_types, lifting) add.assoc cartesian_power_drop x_closed) then show ?thesis using b_def le_add1 take_closed by blast qed have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using c_def cartesian_power_drop x_closed by blast have B: "a@((f b)#c) \ S" using A twisted_partial_pullback_memE' by (smt a_closed a_def add.commute append_take_drop_id b_closed b_def c_closed c_def drop_drop) have "permute_list (fun_inv (twisting_permutation n)) (a@((f b)#c)) = (f b)#(a@c)" using assms twisting_permutation_action'[of a n "f b" c] a_closed cartesian_power_car_memE by blast then have C: "(f b)#(a@c) \ (permute_list (fun_inv (twisting_permutation n)) ` S)" by (metis B image_eqI) have C: "b@(a@c) \ partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" proof(rule partial_pullback_memI) show "b @ a @ c \ carrier (Q\<^sub>p\<^bsup>m + (n + l)\<^esup>)" using a_closed b_closed c_closed cartesian_power_concat(1) by blast have 0: "(take m (b @ a @ c)) = b" by (metis append.right_neutral b_closed cartesian_power_car_memE diff_is_0_eq diff_self_eq_0 take0 take_all take_append) have 1: "drop m (b @ a @ c) = a@c" by (metis "0" append_take_drop_id same_append_eq) show "f (take m (b @ a @ c)) # drop m (b @ a @ c) \ permute_list (fun_inv (twisting_permutation n)) ` S" using 0 1 C by presburger qed have D: "tw m n (b@(a@c)) = a@(b@c)" using assms tw_def a_closed b_closed c_closed by (metis tp_1_permutation_action x_eq) then show "x \ tw m n ` Y" using x_eq C assms by (metis image_eqI) qed show "tw m n ` Y \ twisted_partial_pullback n m l f S" proof fix x assume A: "x \ tw m n ` Y" then obtain y where y_def: "x = tw m n y \ y \ Y" by blast obtain as where as_def: "as \ (permute_list (fun_inv (twisting_permutation n)) ` S) \ as = partial_image m f y" using partial_pullback_memE by (metis assms(3) y_def) obtain s where s_def: "s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as" using as_def by blast obtain b where b_def: "b = take m y" by blast obtain a where a_def: "a = take n (drop m y)" by blast obtain c where c_def: "c = drop (n + m) y" by blast have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>m + n + l\<^esup>)" by (metis add.assoc assms(3) partial_pullback_memE(1) y_def) then have y_eq: "y = b@a@c" using a_def b_def c_def by (metis append_take_drop_id drop_drop) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis a_def add.commute cartesian_power_drop le_add1 take_closed take_drop y_closed) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using add_leD2 b_def le_add1 take_closed y_closed by (meson trans_le_add1) have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using c_def cartesian_power_drop y_closed by (metis add.commute) have ac_closed: "a@c \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using a_closed c_closed cartesian_power_concat(1) by blast then have C: " local.partial_image m f y = f b # a @ c" using b_closed y_eq partial_image_eq[of b m "a@c" "n + l" y f] by blast then have as_eq: "as = (f b)#(a@c)" using as_def by force have B: "(tw m n) y = a@b@c" using y_eq tw_def[of n m] tp_1_permutation_action by (smt a_closed b_closed c_closed tw_def) then have "x = a@(b@c)" by (simp add: y_def) then have "twisted_partial_image n m f x = a@((f b)# c)" using a_closed b_closed c_closed twisted_partial_image_eq by blast then have D: "permute_list (twisting_permutation n) as = twisted_partial_image n m f x" using as_eq twisting_permutation_action[of a n "f b" c ] by (metis a_closed cartesian_power_car_memE) have "permute_list (twisting_permutation n) as \ S" proof- have S: "length s > n" using s_def assms cartesian_power_car_memE le_add1 le_neq_implies_less le_trans less_add_same_cancel1 less_one not_add_less1 by (metis (no_types, lifting) subset_iff) have "permute_list (twisting_permutation n) as = permute_list (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)" using fun_inv_def s_def by blast then have "permute_list (twisting_permutation n) as = permute_list ((twisting_permutation n) \ (fun_inv (twisting_permutation n))) s" using fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv length_permute_list twisting_permutation_permutes[of n "length s"] permute_list_compose[of "fun_inv (twisting_permutation n)" s "twisting_permutation n"] by (metis S permute_list_compose) then have "permute_list (twisting_permutation n) as = permute_list (id) s" by (metis S \permute_list (twisting_permutation n) as = permute_list (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)\ fun_inv_permute(3) length_greater_0_conv length_permute_list permute_list_compose twisting_permutation_permutes) then have "permute_list (twisting_permutation n) as = s" by simp then show ?thesis using s_def by (simp add: \s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as\) qed then show "x \ twisted_partial_pullback n m l f S" unfolding twisted_partial_pullback_def using D by (smt \x = a @ b @ c\ a_closed append.assoc append_eq_conv_conj b_closed c_closed cartesian_power_car_memE cartesian_power_concat(1) length_append list.inject local.partial_image_def twisted_partial_image_def twisted_partial_pullback_def twisted_partial_pullback_memI) qed qed lemma twisted_partial_pullback_is_semialgebraic: assumes "is_semialg_function m f" assumes "is_semialgebraic (n + 1 + l) S" shows "is_semialgebraic (n + m + l)(twisted_partial_pullback n m l f S)" proof- have "(fun_inv (twisting_permutation n)) permutes {.. carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" apply(rule cartesian_power_car_memI) apply (smt ab_semigroup_add_class.add_ac(1) add.commute append_take_drop_id assms augment_def cartesian_power_car_memE cartesian_power_drop length_append) using assms cartesian_power_car_memE'' unfolding augment_def by (metis (no_types, opaque_lifting) append_take_drop_id cartesian_power_concat(2) nat_le_iff_add take_closed) lemma tuple_partial_image_factor: assumes "is_function_tuple Q\<^sub>p m fs" assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "length fs = n" assumes "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" shows "tuple_partial_image m (fs@[f]) x = twisted_partial_image n m f (tuple_partial_image m fs (augment m x))" proof- obtain a where a_def: "a = take m x" by blast obtain b where b_def: "b = drop m x" by blast have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using a_def assms(4) le_add1 take_closed by (meson dual_order.trans) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using assms(4) b_def cartesian_power_drop by (metis (no_types, lifting)) have A: "(augment m x) = a @ (a @ b)" using a_def augment_def b_def by blast have 0: "tuple_partial_image m fs (augment m x) = ((function_tuple_eval Q\<^sub>p m fs) a) @ a @ b" using A a_closed b_closed tuple_partial_image_eq[of a m "a@b" "m + l" "augment m x" fs] cartesian_power_concat(1) by blast have 1: "tuple_partial_image m (fs@[f]) x = ((function_tuple_eval Q\<^sub>p m fs a) @ [f a])@b" using 0 tuple_partial_image_eq[of a m b l x "fs@[f]"] unfolding function_tuple_eval_def by (metis (no_types, lifting) a_closed a_def append_take_drop_id b_closed b_def list.simps(8) list.simps(9) map_append) have 2: "tuple_partial_image m (fs@[f]) x = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" using 1 by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv2 same_append_eq) have 3: "tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs a) @ b" using a_def b_def 2 tuple_partial_image_eq[of a m b l x fs ] assms tuple_partial_image_def by blast have 4: "twisted_partial_image n m f (tuple_partial_image m fs (augment m x)) = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" using twisted_partial_image_eq[of _ n _ m _ l] 0 assms(1) assms(3) b_closed local.a_closed local.function_tuple_eval_closed by blast show ?thesis using 2 4 by presburger qed definition diagonalize where "diagonalize n m S = S \ cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" lemma diagaonlize_is_semiaglebraic: assumes "is_semialgebraic (n + n + m) S" shows "is_semialgebraic (n + n + m) (diagonalize n m S)" proof(cases "m = 0") case True then have 0: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) = {[]}" unfolding cartesian_power_def by simp have 1: "\ n \ carrier (Q\<^sub>p\<^bsup>n+n\<^esup>)" using Qp.cring_axioms assms diagonalE(2) by blast then have "cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = \ n" using 0 cartesian_product_empty_right[of "\ n" Q\<^sub>p "n + n" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] by linarith then have "diagonalize n m S = S \ (\ n)" using diagonalize_def by presburger then show ?thesis using intersection_is_semialg True assms diag_is_semialgebraic by auto next case False have "is_semialgebraic (n + n + m) (cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" using carrier_is_semialgebraic[of m] cartesian_product_is_semialgebraic[of "n + n" "\ n" m "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] diag_is_semialgebraic[of n] False by blast then show ?thesis using intersection_is_semialg assms(1) diagonalize_def by presburger qed lemma list_segment_take: assumes "length a \n" shows "list_segment 0 n a = take n a" proof- have 0: "length (list_segment 0 n a) = length (take n a)" using assms unfolding list_segment_def by (metis (no_types, lifting) Groups.add_ac(2) add_diff_cancel_left' append_take_drop_id le_Suc_ex length_append length_drop length_map map_nth) have "\i. i < n \ list_segment 0 n a !i = take n a ! i" unfolding list_segment_def using assms by (metis add.left_neutral diff_zero nth_map_upt nth_take) then show ?thesis using 0 by (metis assms diff_zero le0 list_segment_length nth_equalityI) qed lemma augment_inverse_is_semialgebraic: assumes "is_semialgebraic (n+n+l) S" shows "is_semialgebraic (n+l) ((augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>))" proof- obtain Ps where Ps_def: "Ps = (var_list_segment 0 n)" by blast obtain Qs where Qs_def: "Qs = (var_list_segment n (n+l))" by blast obtain Fs where Fs_def: "Fs = Ps@Ps@Qs" by blast have 0: "is_poly_tuple (n+l) Ps" by (simp add: Ps_def var_list_segment_is_poly_tuple) have 1: "is_poly_tuple (n+l) Qs" by (simp add: Qs_def var_list_segment_is_poly_tuple) have 2: "is_poly_tuple (n+l) (Ps@Qs)" using Qp_is_poly_tuple_append[of "n+l" Ps Qs] by (metis (no_types, opaque_lifting) "0" "1" add.commute) have "is_poly_tuple (n+l) Fs" using 0 2 Qp_is_poly_tuple_append[of "n + l" Ps "Ps@Qs"] Fs_def assms by blast have 3: "\x. x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ augment n x = poly_map (n + l) Fs x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" have 30: "poly_map (n+l) Ps x = take n x" using Ps_def map_by_var_list_segment[of x "n + l" n 0] list_segment_take[of n x] cartesian_power_car_memE[of x Q\<^sub>p "n+l"] by (simp add: A) have 31: "poly_map (n + l) Qs x = drop n x" using Qs_def map_by_var_list_segment_to_length[of x "n + l" n] A le_add1 by blast have 32: "poly_map (n + l) (Ps@Qs) x = take n x @ drop n x" using poly_map_append[of x "n+l" Ps Qs ] by (metis "30" "31" A append_take_drop_id) show "augment n x = poly_map (n + l) Fs x" using 30 32 poly_map_append by (metis A Fs_def poly_map_append augment_def) qed have 4: "(augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) = poly_tuple_pullback (n + l) S Fs" proof show "augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ poly_tuple_pullback (n + l) S Fs" proof fix x assume A: "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" then have 40: "augment n x \ S" by blast have 41: "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" using 40 assms unfolding augment_def using is_semialgebraic_closed by blast have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" proof- have "take n x @ x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" using augment_def A by (metis "41" append_take_drop_id) then have 0: "drop n (take n x @ x) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" by (metis (no_types, lifting) add.assoc cartesian_power_drop) have "drop n (take n x @ x) = x" proof- have "length x \ n" using A by (metis IntD2 cartesian_power_car_memE le_add1) then have "length (take n x) = n" by (metis add_right_cancel append_take_drop_id le_add_diff_inverse length_append length_drop) then show ?thesis by (metis append_eq_conv_conj) qed then show ?thesis using 0 by presburger qed then show "x \ poly_tuple_pullback (n + l) S Fs" using 41 3 unfolding poly_tuple_pullback_def by (metis (no_types, opaque_lifting) "40" add.commute cartesian_power_car_memE evimageI evimage_def poly_map_apply) qed show "poly_tuple_pullback (n + l) S Fs \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" proof fix x assume A: "x \ poly_tuple_pullback (n + l) S Fs" have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using A unfolding poly_tuple_pullback_def by blast then show "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using 3 by (metis (no_types, lifting) A poly_map_apply poly_tuple_pullback_def vimage_inter_cong) qed qed then show ?thesis using assms pullback_is_semialg[of "n + l" Fs] using poly_tuple_pullback_eq_poly_map_vimage unfolding restrict_def evimage_def Fs_def by (smt "4" Ex_list_of_length Fs_def Ps_def Qs_def \is_poly_tuple (n + l) Fs\ add.commute add_diff_cancel_left' append_assoc diff_zero is_semialgebraic_closed le_add2 length_append not_add_less1 not_gr_zero padic_fields.is_semialgebraicE padic_fields_axioms var_list_segment_length zero_le) qed lemma tuple_partial_pullback_is_semialg_map_tuple_induct: assumes "is_semialg_map_tuple m fs" assumes "is_semialg_function m f" assumes "length fs = n" shows "is_semialg_map_tuple m (fs@[f])" proof(rule is_semialg_map_tupleI) have 0: "is_function_tuple Q\<^sub>p m fs" using assms is_semialg_map_tuple_def by blast show "is_function_tuple Q\<^sub>p m (fs @ [f])" proof(rule is_function_tupleI) have A0: "set (fs @ [f]) = insert f (set fs)" by simp have A1: "set fs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using 0 is_function_tuple_def by blast then show "set (fs @ [f]) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms 0 by (metis (no_types, lifting) A0 is_semialg_function_closed list.simps(15) set_ConsD subset_code(1)) qed show "\k S. S \ semialg_sets (length (fs @ [f]) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m (fs @ [f]) k S)" proof- fix l S assume A: "S \ semialg_sets (length (fs @ [f]) + l)" then have B: "S \ semialg_sets (n + l + 1)" using assms by (metis (no_types, lifting) add.commute add_Suc_right add_diff_cancel_left' append_Nil2 diff_Suc_1 length_Suc_conv length_append) show "is_semialgebraic (m + l) (tuple_partial_pullback m (fs @ [f]) l S)" proof- obtain S0 where S0_def: "S0 = tuple_partial_pullback m fs (l+1) S" by blast have 0: "is_semialgebraic (m + l + 1) S0" using B assms is_semialg_map_tupleE[of m fs "l + 1" S] by (metis S0_def add.assoc is_semialgebraicI) obtain S1 where S1_def: "S1 = twisted_partial_pullback m m l f S0" by blast then have "is_semialgebraic (m + m + l) S1" using S1_def assms(1) 0 twisted_partial_pullback_is_semialgebraic[of m f m l S0] by (simp add: assms(2)) then have L: "is_semialgebraic (m + m + l) (diagonalize m l S1)" using assms diagaonlize_is_semiaglebraic by blast have 1: "(tuple_partial_pullback m (fs @ [f]) l S) = (augment m -` (diagonalize m l S1)) \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof show "tuple_partial_pullback m (fs @ [f]) l S \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof fix x assume P0: "x \ tuple_partial_pullback m (fs @ [f]) l S " show "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof show "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using tuple_partial_pullback_closed P0 by blast show "x \ augment m -` diagonalize m l S1" proof- obtain a where a_def: "a = take m x" by blast then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ le_add1 take_closed by blast obtain b where b_def: "b = drop m x" by blast then have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ cartesian_power_drop by blast have x_eq: "x = a@b" using a_def b_def by (metis append_take_drop_id) have X0: "a @ a @ b = augment m x" by (metis a_def augment_def b_def) have "a @ a @ b \ diagonalize m l S1" proof- have "length (a@a) = m + m" using a_closed cartesian_power_car_memE length_append by blast then have "take (m + m) (a @ a @ b) = a@a" by (metis append.assoc append_eq_conv_conj) then have X00: "take (m + m) (a @ a @ b) \ \ m" using diagonalI[of "a@a"] a_def a_closed by (metis append_eq_conv_conj cartesian_power_car_memE) then have X01: "a @ a @ b \ cartesian_product (\ m) (carrier (Q\<^sub>p\<^bsup>l\<^esup>))" using a_closed b_closed cartesian_product_memI[of "\ m" Q\<^sub>p "m+m" "carrier (Q\<^sub>p\<^bsup>l\<^esup>)" l "a @ a @ b"] unfolding diagonal_def by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed cartesian_power_drop mem_Collect_eq subsetI) have X02: "twisted_partial_image m m f (a @ a @ b) = a @ ((f a)# b)" using twisted_partial_image_eq[of a m a m b l _ f] a_closed b_closed by blast have "a @ a @ b \ S1" proof- have "twisted_partial_image m m f (a @ a @ b) \ S0" proof- have X020:"tuple_partial_image m fs (a @ ((f a)# b)) = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" using tuple_partial_image_eq[of a m "(f a)# b" "l + 1" _ fs] by (metis (no_types, lifting) a_closed append_Cons append_eq_conv_conj cartesian_power_car_memE self_append_conv2 tuple_partial_image_def) have X021: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b \ S" proof- have X0210: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b = (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" unfolding function_tuple_eval_def by (metis (mono_tags, lifting) append.assoc list.simps(8) list.simps(9) map_append) have X0211: "(function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b = tuple_partial_image m (fs @ [f]) x" using x_eq tuple_partial_image_eq[of a m b l x "fs@[f]"] by (simp add: a_closed b_closed) have "tuple_partial_image m (fs @ [f]) x \ S" using P0 tuple_partial_pullback_memE(2) by blast then show ?thesis using X0211 X0210 by presburger qed have X022: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" proof- have X0220: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = tuple_partial_image m fs (a @ ((f a)# b))" using X02 by presburger have X0221: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)# b)" using tuple_partial_image_eq by (metis X02 X020 append_Cons self_append_conv2) then show ?thesis unfolding function_tuple_eval_def by (metis X02 X020 X0221 append_same_eq) qed have X023: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) \ S" using X02 X020 X021 by presburger have "twisted_partial_image m m f (a @ a @ b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" proof- have "a @ ((f a)# b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" apply(rule cartesian_power_car_memI) apply (metis a_closed add.commute b_closed cartesian_power_car_memE length_Cons length_append plus_1_eq_Suc) proof- have "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms(2) is_semialg_function_closed by blast then have "f a \ carrier Q\<^sub>p" using a_closed assms by blast then show "set (a @ f a # b) \ carrier Q\<^sub>p" using assms a_closed b_closed by (meson cartesian_power_car_memE'' cartesian_power_concat(1) cartesian_power_cons) qed then show ?thesis using X02 by presburger qed then show ?thesis using S0_def X023 tuple_partial_pullback_def[of m fs "l+1" S ] by blast qed then show ?thesis using X02 S1_def twisted_partial_pullback_def by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed drop_drop local.partial_image_def twisted_partial_image_def twisted_partial_pullback_memI) qed then show ?thesis using X01 diagonalize_def[of m l S1] by blast qed then show ?thesis by (metis X0 vimageI) qed qed qed show "augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) \ tuple_partial_pullback m (fs @ [f]) l S" proof fix x assume A: "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" then have X0: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" by blast obtain a where a_def: "a = take m x" by blast then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using X0 le_add1 take_closed by blast obtain b where b_def: "b = drop m x" by blast then have a_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using X0 cartesian_power_drop by blast have X1: "augment m x = a@a@b" using a_def augment_def b_def by blast have X2: "a@a@b \ diagonalize m l S1" using A X1 by (metis Int_iff vimage_eq) have X3: "a@a@b \ S1" using X2 diagonalize_def by blast have X4: "twisted_partial_image m m f (a@a@b) \ S0" using X3 S1_def twisted_partial_pullback_memE(2) by blast have X5: "a@((f a)#b) \ S0" using X4 twisted_partial_image_eq[of a m a m b l _ f] by (metis X0 a_closed a_def le_add1 take_closed) have X6: "tuple_partial_image m fs (a@((f a)#b)) \ S" using S0_def X5 tuple_partial_pullback_memE(2) by blast have X7: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) \ S" using X6 using tuple_partial_image_eq by (metis X0 a_def append_eq_conv_conj cartesian_power_car_memE le_add1 take_closed tuple_partial_image_def) have X8: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) = tuple_partial_image m (fs @ [f]) x" proof- have X80: "tuple_partial_image m (fs @ [f]) x = (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" using tuple_partial_image_def a_def b_def by blast then show ?thesis unfolding function_tuple_eval_def by (metis (no_types, lifting) append_Cons append_eq_append_conv2 list.simps(9) map_append self_append_conv2) qed show "x \ tuple_partial_pullback m (fs @ [f]) l S" using X8 X7 tuple_partial_pullback_def by (metis X0 \is_function_tuple Q\<^sub>p m (fs @ [f])\ tuple_partial_image_def tuple_partial_pullback_memI) qed qed then show ?thesis using augment_inverse_is_semialgebraic by (simp add: L) qed qed qed lemma singleton_tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" assumes "length fs = 1" shows "is_semialg_map_tuple m fs" proof(rule is_semialg_map_tupleI) show "is_function_tuple Q\<^sub>p m fs" by (simp add: assms(1) semialg_function_tuple_is_function_tuple) show "\k S. S \ semialg_sets (length fs + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- fix k S assume A: "S \ semialg_sets (length fs + k)" show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- obtain f where f_def: "fs = [f]" using assms by (metis One_nat_def length_0_conv length_Suc_conv) have 0: "is_semialg_function m f" using f_def assms is_semialg_function_tupleE'[of m fs f] by simp have 1: "\x. tuple_partial_image m fs x = partial_image m f x" unfolding function_tuple_eval_def tuple_partial_image_def partial_image_def by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv_if f_def list.simps(8) list.simps(9)) have 2: "tuple_partial_pullback m fs k S = partial_pullback m f k S" proof show "tuple_partial_pullback m fs k S \ partial_pullback m f k S" using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def by (metis (no_types, lifting) set_eq_subset vimage_inter_cong) show "partial_pullback m f k S \ tuple_partial_pullback m fs k S" using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def by blast qed then show ?thesis by (metis "0" A assms(2) is_semialg_functionE is_semialgebraicI) qed qed qed lemma empty_tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" assumes "length fs = 0" shows "is_semialg_map_tuple m fs" apply(rule is_semialg_map_tupleI) using assms(1) semialg_function_tuple_is_function_tuple apply blast proof- fix k S assume A: "S \ semialg_sets (length fs + k)" then have 0: "is_semialgebraic k S" by (metis add.left_neutral assms(2) is_semialgebraicI) have 1: "tuple_partial_pullback m fs k S = cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" proof have 1: "\x. function_tuple_eval Q\<^sub>p m fs (take m x) = []" using assms unfolding function_tuple_eval_def by blast show "tuple_partial_pullback m fs k S \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" apply(rule subsetI) apply(rule cartesian_product_memI[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k]) apply blast using 0 is_semialgebraic_closed apply blast using 0 assms unfolding 1 tuple_partial_pullback_def tuple_partial_image_def apply (meson IntD2 le_add1 take_closed) by (metis append_Nil evimageD evimage_def) have 2: "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" using is_semialgebraic_closed[of k S] 0 assms cartesian_product_closed[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k] by blast show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ tuple_partial_pullback m fs k S" apply(rule subsetI) apply(rule tuple_partial_pullback_memI) using 2 apply blast using assms semialg_function_tuple_is_function_tuple apply blast unfolding 1 by (metis carrier_is_semialgebraic cartesian_product_memE(2) is_semialgebraic_closed self_append_conv2) qed show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" unfolding 1 using "0" car_times_semialg_is_semialg by blast qed lemma tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" shows "is_semialg_map_tuple m fs" proof- have "\n fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" proof- fix n show " \ fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" apply(induction n) using singleton_tuple_partial_pullback_is_semialg_map_tuple empty_tuple_partial_pullback_is_semialg_map_tuple apply blast proof- fix n fs assume IH: "(\fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs)" assume A: "is_semialg_function_tuple m fs \ length fs = Suc n" then obtain gs f where gs_f_def: "fs = gs@[f]" by (metis length_Suc_conv list.discI rev_exhaust) have gs_length: "length gs = n" using gs_f_def by (metis A length_append_singleton nat.inject) have 0: "set gs \ set fs" by (simp add: gs_f_def subsetI) have 1: "is_semialg_function_tuple m gs" apply(rule is_semialg_function_tupleI) using 0 A gs_f_def is_semialg_function_tupleE'[of m fs] by blast then have 2: "is_semialg_map_tuple m gs" using IH gs_length by blast have 3: "is_semialg_function m f" using gs_f_def A by (metis gs_length is_semialg_function_tupleE lessI nth_append_length) then show "is_semialg_map_tuple m fs" using assms 2 gs_f_def tuple_partial_pullback_is_semialg_map_tuple_induct by blast qed qed then show ?thesis using assms by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\Semialgebraic Functions are Closed under Composition with Semialgebraic Tuples\ (********************************************************************************************) (********************************************************************************************) lemma function_tuple_comp_partial_pullback: assumes "is_semialg_function_tuple m fs" assumes "length fs = n" assumes "is_semialg_function n f" assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" shows "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" proof- have 0: "\x. partial_image m (function_tuple_comp Q\<^sub>p fs f) x = partial_image n f (tuple_partial_image m fs x)" unfolding partial_image_def function_tuple_comp_def tuple_partial_image_def using comp_apply[of f "function_tuple_eval Q\<^sub>p 0 fs"] unfolding function_tuple_eval_def proof - fix x :: "((nat \ int) \ (nat \ int)) set list" assume a1: "\x. (f \ (\x. map (\f. f x) fs)) x = f (map (\f. f x) fs)" have f2: "\f rs. drop n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = rs" by (simp add: assms(2)) have "\f rs. take n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = map f fs" by (simp add: assms(2)) then show "(f \ (\rs. map (\f. f rs) fs)) (take m x) # drop m x = f (take n (map (\f. f (take m x)) fs @ drop m x)) # drop n (map (\f. f (take m x)) fs @ drop m x)" using f2 a1 by presburger qed show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" proof show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S \ tuple_partial_pullback m fs k (partial_pullback n f k S)" proof fix x assume A: "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" then have 1: "partial_image m (function_tuple_comp Q\<^sub>p fs f) x \ S" using partial_pullback_memE(2) by blast have 2: " partial_image n f (tuple_partial_image m fs x) \ S" using 0 1 by presburger have 3: "x \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" using A assms by (metis partial_pullback_memE(1)) have 4: "tuple_partial_image m fs x \ partial_pullback n f k S" apply(rule partial_pullback_memI) apply (metis "0" "3" add_cancel_left_left assms(1) assms(2) cartesian_power_drop drop0 list.inject local.partial_image_def not_gr_zero semialg_function_tuple_is_function_tuple tuple_partial_image_closed) by (metis "2" local.partial_image_def) show " x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" apply(rule tuple_partial_pullback_memI) apply (simp add: "3") using assms(1) semialg_function_tuple_is_function_tuple apply blast by (metis "4" tuple_partial_image_def) qed show " tuple_partial_pullback m fs k (partial_pullback n f k S) \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" proof fix x assume A: "x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" show "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S " proof- have "partial_image n f (tuple_partial_image m fs x) \ S" using A partial_pullback_memE(2) tuple_partial_pullback_memE(2) by blast show ?thesis apply(rule partial_pullback_memI) apply (meson A subset_eq tuple_partial_pullback_closed) by (metis "0" \local.partial_image n f (tuple_partial_image m fs x) \ S\ local.partial_image_def) qed qed qed qed lemma semialg_function_tuple_comp: assumes "is_semialg_function_tuple m fs" assumes "length fs = n" assumes "is_semialg_function n f" shows "is_semialg_function m (function_tuple_comp Q\<^sub>p fs f)" proof(rule is_semialg_functionI) show "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using function_tuple_comp_closed[of f Q\<^sub>p n fs] assms(1) assms(2) assms(3) is_semialg_function_closed semialg_function_tuple_is_function_tuple by blast show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" proof- fix k S assume A0: "S \ semialg_sets (1 + k)" show "is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" proof- have 0: "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" using function_tuple_comp_partial_pullback[of m fs n f S k] assms \S \ semialg_sets (1 + k)\ is_semialgebraicI is_semialgebraic_closed by blast have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" using assms A0 is_semialg_functionE is_semialgebraicI by blast have 2: "is_semialgebraic (m + k) (tuple_partial_pullback m fs k (partial_pullback n f k S))" using 1 0 assms tuple_partial_pullback_is_semialg_map_tuple[of m fs] is_semialg_map_tupleE[of m fs k "partial_pullback n f k S"] by blast then show ?thesis using 0 by simp qed qed qed (********************************************************************************************) (********************************************************************************************) subsubsection\Algebraic Operations on Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Defining the set of extensional semialgebraic functions\ definition Qp_add_fun where "Qp_add_fun xs = xs!0 \\<^bsub>Q\<^sub>p\<^esub> xs!1" definition Qp_mult_fun where "Qp_mult_fun xs = xs!0 \ xs!1" text\Inversion function on first coordinates of Qp tuples. Arbitrarily redefined at 0 to map to 0\ definition Qp_invert where "Qp_invert xs = (if ((xs!0) = \) then \ else (inv (xs!0)))" text\Addition is semialgebraic\ lemma addition_is_semialg: "is_semialg_function 2 Qp_add_fun" proof- have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = (Qp_ev (pvar Q\<^sub>p 0) x) \\<^bsub>Q\<^sub>p\<^esub> (Qp_ev (pvar Q\<^sub>p 1) x)" by (metis A One_nat_def eval_at_point_add pvar_closed less_Suc_eq numeral_2_eq_2) then show " Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" by (metis A Qp_add_fun_def add_vars_def add_vars_rep one_less_numeral_iff pos2 semiring_norm(76)) qed then have 1: "restrict Qp_add_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" by (meson restrict_ext) have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] by (meson MP.add.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) then show ?thesis using 1 semialg_function_on_carrier[of 2 "Qp_add_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] semialg_function_on_carrier by presburger qed text\Multiplication is semialgebraic:\ lemma multiplication_is_semialg: "is_semialg_function 2 Qp_mult_fun" proof- have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = (Qp_ev (pvar Q\<^sub>p 0) x) \ (Qp_ev (pvar Q\<^sub>p 1) x)" by (metis A One_nat_def eval_at_point_mult pvar_closed less_Suc_eq numeral_2_eq_2) then show " Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" by (metis A Qp_mult_fun_def mult_vars_def mult_vars_rep one_less_numeral_iff pos2 semiring_norm(76)) qed then have 1: "restrict Qp_mult_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" by (meson restrict_ext) have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] by (meson MP.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) thus ?thesis using 1 semialg_function_on_carrier[of 2 "Qp_mult_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] semialg_function_on_carrier by presburger qed text\Inversion is semialgebraic:\ lemma(in field) field_nat_pow_inv: assumes "a \ carrier R" assumes "a \ \" shows "inv (a [^] (n::nat)) = (inv a) [^] (n :: nat)" apply(induction n) using inv_one local.nat_pow_0 apply presburger using assms nat_pow_of_inv by (metis Units_one_closed field_inv(2) field_inv(3) unit_factor) lemma Qp_invert_basic_semialg: assumes "is_basic_semialg (1 + k) S" shows "is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" proof- obtain P n where P_n_def: "(n::nat) \ 0 \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>]) \ S = basic_semialg_set (1+k) n P \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>])" using assms is_basic_semialg_def by meson obtain d::nat where d_def: "d = deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P)" by auto obtain l where l_def: "l = ((- d) mod n)" by blast have 10: "n > 0" using P_n_def by blast have 11: "l \ 0" using 10 by (simp add: l_def) - have 1: "(l + d) mod n = 0" - by (metis add_eq_0_iff equation_minus_iff l_def mod_0 mod_add_cong mod_minus_eq zmod_int) - then obtain m::int where m_def: "(l + int d) = m * int n" - using d_def l_def - by (metis mult_of_nat_commute zmod_eq_0D) + have 1: "int n dvd l + int d" + by (simp add: l_def ac_simps mod_0_imp_dvd mod_add_right_eq) + then obtain m::int where m_def: "l + int d = int n * m" .. with 10 have \m = (l + d) div n\ by simp with 10 11 have 2: "m \ 0" by (simp add: div_int_pos_iff) - obtain N where N_def: "N = m*n" + obtain N where N_def: "N = m * n" by blast from 11 have 3: "N \ d" - by (simp add: N_def flip: m_def) + using m_def by (simp add: N_def ac_simps) have 4: "deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P) \ nat N" using d_def N_def 3 by linarith have 5: " P \ carrier (coord_ring Q\<^sub>p (Suc k))" by (metis P_n_def plus_1_eq_Suc) have 6: " \q\carrier (coord_ring Q\<^sub>p (Suc k)). \x\carrier Q\<^sub>p - {\}. \a\carrier (Q\<^sub>p\<^bsup>k\<^esup>). Qp_ev q (insert_at_index a x 0) = (x[^]nat N) \ Qp_ev P (insert_at_index a (inv x) 0)" using 3 4 d_def to_univ_poly_one_over_poly''[of 0 k P "nat N"] "5" Qp.field_axioms by blast obtain q where q_def: "q \ carrier (coord_ring Q\<^sub>p (Suc k)) \ ( \ x \ carrier Q\<^sub>p - {\}. ( \ a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). eval_at_point Q\<^sub>p (insert_at_index a x 0) q = (x[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index a (inv x) 0) P)))" using 6 by blast obtain T where T_def: "T = basic_semialg_set (1+k) n q" by auto have "is_basic_semialg (1 + k) T" proof- have "q \ carrier ( Q\<^sub>p[\\<^bsub>Suc k\<^esub>])" using q_def by presburger then show ?thesis using T_def is_basic_semialg_def by (metis P_n_def plus_1_eq_Suc) qed then have T_semialg: "is_semialgebraic (1+k) T" using T_def basic_semialg_is_semialg[of "1+k" T] is_semialgebraicI by blast obtain Nz where Nz_def: "Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 \ \}" by blast have Nz_semialg: "is_semialgebraic (1+k) Nz" proof- obtain Nzc where Nzc_def: "Nzc = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" by blast have 0: "Nzc = zero_set Q\<^sub>p (Suc k) (pvar Q\<^sub>p 0)" unfolding zero_set_def using Nzc_def by (metis (no_types, lifting) Collect_cong eval_pvar zero_less_Suc) have 1: "is_algebraic Q\<^sub>p (1+k) Nzc" using 0 pvar_closed[of ] by (metis is_algebraicI' plus_1_eq_Suc zero_less_Suc) then have 2: "is_semialgebraic (1+k) Nzc" using is_algebraic_imp_is_semialg by blast have 3: "Nz = carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>) - Nzc" using Nz_def Nzc_def by blast then show ?thesis using 2 by (simp add: complement_is_semialg) qed have 7: "(partial_pullback 1 Qp_invert k S) \ Nz = T \ Nz" proof show "partial_pullback 1 Qp_invert k S \ Nz \ T \ Nz" proof fix c assume A: "c \ partial_pullback 1 Qp_invert k S \ Nz" show "c \ T \ Nz" proof- have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using A partial_pullback_closed[of 1 Qp_invert k S] by blast obtain x a where xa_def: "c = (x#a)" using c_closed by (metis Suc_eq_plus1 add.commute cartesian_power_car_memE length_Suc_conv) have x_closed: "x \ carrier Q\<^sub>p" using xa_def c_closed by (metis (no_types, lifting) append_Cons cartesian_power_decomp list.inject Qp.to_R1_to_R Qp.to_R_pow_closed) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using xa_def c_closed by (metis One_nat_def cartesian_power_drop drop0 drop_Suc_Cons) have 0: "c \ Nz" using A by blast then have "x \ \" using Nz_def xa_def by (metis (mono_tags, lifting) mem_Collect_eq nth_Cons_0) have 1: "Qp_invert [x] = inv x" unfolding Qp_invert_def by (metis \x \ \\ nth_Cons_0) have 2: "partial_image 1 Qp_invert c \ S" using A partial_pullback_memE[of c 1 "Qp_invert" k S] by blast have 3: "inv x # a \ S" proof- have 30: "[x] = take 1 c" by (simp add: xa_def) have 31: "a = drop 1 c" by (simp add: xa_def) show ?thesis using 1 30 31 partial_image_def[of 1 "Qp_invert" c] xa_def "2" by metis qed obtain y where y_def: "y \ carrier Q\<^sub>p \ eval_at_point Q\<^sub>p (inv x # a) P = y [^] n" using 3 P_n_def basic_semialg_set_memE(2) by blast then have 4: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P = x [^] (nat N) \ y [^] n" by presburger have 5: "x [^] (nat N) \ y [^] n = ((x[^]m)\y)[^]n" proof- have 50: "x [^] (N) \ y [^] n = x [^] (m*n) \ y [^] n" using N_def by blast have 51: "x [^] (m*n) = (x[^]m)[^]n" using Qp_int_nat_pow_pow \x \ \\ not_nonzero_Qp x_closed by metis have 52: "x [^] (m*n)\ y [^] n = ((x[^]m) \ y) [^] n" proof- have 0: "x [^] (m*n)\ y [^] n= (x[^]m)[^]n \ (y[^] n)" using "51" by presburger have 1: "(x[^]m)[^]n \ (y[^] n) = ((x[^]m) \ y) [^] n" apply(induction n) using Qp.nat_pow_0 Qp.one_closed Qp.r_one apply presburger using x_closed y_def by (metis Qp.nat_pow_distrib Qp.nonzero_closed Qp_int_pow_nonzero \x \ \\ not_nonzero_Qp) then show ?thesis using "0" by blast qed have 53: "x [^] N = x [^] (nat N)" - proof- - from 11 m_def N_def [symmetric] have "N \ 0" - by simp - then show ?thesis - by (metis pow_nat) - qed + using 11 m_def N_def by (simp add: ac_simps) then show ?thesis using 50 52 by presburger qed have 6: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P = ((x[^]m)\y)[^]n" using "4" "5" by blast have 7: "eval_at_point Q\<^sub>p c q = ((x[^]m)\y)[^]n" proof- have 70: "(insert_at_index a (inv x) 0) = inv x # a" using insert_at_index.simps by (metis (no_types, lifting) append_eq_append_conv2 append_same_eq append_take_drop_id drop0 same_append_eq) have 71: "(insert_at_index a x) 0 = x # a" by simp then show ?thesis using 6 q_def by (metis "70" DiffI \x \ \\ a_closed empty_iff insert_iff x_closed xa_def) qed have 8: "(x[^]m)\y \ carrier Q\<^sub>p" proof- have 80: "x[^]m \ carrier Q\<^sub>p" using \x \ \\ x_closed Qp_int_pow_nonzero[of x m] unfolding nonzero_def by blast then show ?thesis using y_def by blast qed then have "c \ T" using T_def basic_semialg_set_def "7" c_closed by blast then show ?thesis by (simp add: \c \ T\ "0") qed qed show "T \ Nz \ partial_pullback 1 Qp_invert k S \ Nz" proof fix x assume A: "x \ T \ Nz" show " x \ partial_pullback 1 Qp_invert k S \ Nz " proof- have " x \ partial_pullback 1 Qp_invert k S" proof(rule partial_pullback_memI) show x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using T_def A by (meson IntD1 basic_semialg_set_memE(1)) show "Qp_invert (take 1 x) # drop 1 x \ S" proof- have 00: "x!0 \ \" using A Nz_def by blast then have 0: "Qp_invert (take 1 x) # drop 1 x = inv (x!0) # drop 1 x" unfolding Qp_invert_def by (smt One_nat_def lessI nth_take) have "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ cartesian_power_drop by blast obtain a where a_def: "a = (x!0)" by blast have a_closed: "a \ carrier Q\<^sub>p" using 00 a_def A Nz_def cartesian_power_car_memE'[of x Q\<^sub>p "Suc k" 0] inv_in_frac(1) by blast have a_nz: "a \ \" using a_def Nz_def A by blast obtain b where b_def: "b = drop 1 x" by blast have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using b_def A Nz_def \drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)\ by blast have abx: "x = a#b" using a_def b_def x_closed by (metis (no_types, lifting) One_nat_def append_Cons append_Nil append_eq_conv_conj cartesian_power_car_memE cartesian_power_decomp lessI nth_take Qp.to_R1_to_R) have 1: "Qp_invert (take 1 x) # drop 1 x = (inv a)#b" using "0" a_def b_def by blast have 22: "eval_at_point Q\<^sub>p (insert_at_index b a 0) q = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" using q_def a_closed a_nz b_closed by blast obtain c where c_def: "c \ carrier Q\<^sub>p \ Qp_ev q x = (c[^]n)" using A T_def unfolding basic_semialg_set_def by blast obtain c' where c'_def: "c' = (inv a)[^]m \ c" by blast have c'_closed: "c' \ carrier Q\<^sub>p" using c_def a_def a_closed a_nz Qp_int_pow_nonzero nonzero_def c'_def inv_in_frac(3) Qp.m_closed Qp.nonzero_closed by presburger have 3: "(eval_at_point Q\<^sub>p ((inv a) # b) P) = (c'[^]n)" proof- have 30: "x = insert_at_index b a 0" using abx by simp have 31: "(c[^]n) = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" using 22 30 c_def by blast have 32: "insert_at_index b (inv a) 0 = (inv a) # b" using insert_at_index.simps by (metis drop0 self_append_conv2 take0) have 33: "(c[^]n) = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using "31" "32" by presburger have 34: "(inv a) # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" apply(rule cartesian_power_car_memI'') apply (metis b_closed cartesian_power_car_memE length_Suc_conv plus_1_eq_Suc) using a_closed a_nz b_closed apply (metis One_nat_def inv_in_frac(1) take0 take_Suc_Cons Qp.to_R1_closed) by (metis abx b_closed b_def drop_Cons' not_Cons_self2) have 35: "(eval_at_point Q\<^sub>p ((inv a) # b) P) \ carrier Q\<^sub>p" using 34 P_n_def eval_at_point_closed by blast have "inv(a[^] (nat N)) \ (c[^]n) = inv(a[^] (nat N)) \ ((a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P))" using 31 "33" by presburger then have 6: "inv(a[^] (nat N)) \ (c[^]n) = inv(a[^] (nat N)) \ (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 35 monoid.m_assoc[of Q\<^sub>p] Qp.monoid_axioms Qp.nat_pow_closed Qp.nonzero_pow_nonzero a_nz inv_in_frac(1) local.a_closed by presburger have 37:"inv(a[^] (nat N)) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" proof- have "inv(a[^] (nat N)) \ (a[^] (nat N)) = \ " using a_closed a_nz Qp.nat_pow_closed Qp.nonzero_pow_nonzero field_inv(1) by blast then have "inv(a[^] (nat N)) \ (c[^]n) = \ \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 6 by presburger then show ?thesis using 35 Qp.l_one by blast qed have 38:"(inv a)[^] (nat N) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 37 group.nat_pow_inv[of Q\<^sub>p a "nat N"] a_closed Qp.field_axioms field.field_nat_pow_inv[of Q\<^sub>p] by (metis a_nz) have 39:"((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 2 38 monoid.nat_pow_pow[of Q\<^sub>p "inv a" ] N_def by (smt "3" Qp_int_nat_pow_pow a_closed a_nz inv_in_frac(3) of_nat_0_le_iff pow_nat) have 310:"((((inv a)[^]m) \ c)[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" proof- have AA: "(inv a)[^]m \ carrier Q\<^sub>p" using Qp_int_pow_nonzero nonzero_def a_closed a_nz inv_in_frac(3) Qp.nonzero_closed by presburger have "((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = ((((inv a)[^]m) \ c)[^]n)" using Qp.nat_pow_distrib[of "(inv a)[^]m" c n] a_closed a_def c_def AA by blast then show ?thesis using "39" by blast qed then show ?thesis using c'_def by blast qed have 4: "inv a # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" by (metis a_closed a_nz add.commute b_closed cartesian_power_cons inv_in_frac(1)) then have 5: "((inv a) # b) \ S" using 3 P_n_def c'_closed basic_semialg_set_memI[of "(inv a) # b" "1 + k" c' P n] by blast have 6: "Qp_invert (take 1 x) # drop 1 x = inv a # b" using a_def b_def unfolding Qp_invert_def using "1" Qp_invert_def by blast show ?thesis using 5 6 by presburger qed qed then show ?thesis using A by blast qed qed qed have 8: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) \ Nz)" using "7" Nz_semialg T_semialg intersection_is_semialg by auto have 9: "(partial_pullback 1 Qp_invert k S) - Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S" proof show "partial_pullback 1 Qp_invert k S - Nz \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" proof fix x assume A: " x \ partial_pullback 1 Qp_invert k S - Nz" have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" using A by (metis DiffD1 partial_pullback_memE(1) plus_1_eq_Suc) have 1: "take 1 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" by (metis "0" le_add1 plus_1_eq_Suc take_closed) have 2: "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using "0" cartesian_power_drop plus_1_eq_Suc by presburger have 3: " x = take 1 x @ drop 1 x " using 0 by (metis append_take_drop_id) have 4: "Qp_invert (take 1 x) # drop 1 x \ S" using A partial_pullback_memE'[of "take 1 x" 1 "drop 1 x" k x Qp_invert S] 1 2 3 by blast have 5: "x!0 = \" using A 0 Nz_def by blast have 6: "Qp_invert (take 1 x) # drop 1 x = x" proof- have "(take 1 x) =[x!0]" using 0 by (metis "1" "3" append_Cons nth_Cons_0 Qp.to_R1_to_R) then have "Qp_invert (take 1 x) = \" unfolding Qp_invert_def using 5 by (metis less_one nth_take) then show ?thesis using 0 5 by (metis "3" Cons_eq_append_conv \take 1 x = [x ! 0]\ self_append_conv2) qed have "x \ S" using 6 4 by presburger then show "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" using Nz_def A 0 by blast qed show "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S \ partial_pullback 1 Qp_invert k S - Nz" proof fix x assume A: "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" have A0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" using A by blast have A1: "x!0 = \" using A by blast have A2: "x \ S" using A by blast show " x \ partial_pullback 1 Qp_invert k S - Nz" proof show "x \ Nz" using Nz_def A1 by blast show " x \ partial_pullback 1 Qp_invert k S" proof(rule partial_pullback_memI) show "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using A0 by (simp add: A0) show "Qp_invert (take 1 x) # drop 1 x \ S" proof- have "Qp_invert (take 1 x) = \" unfolding Qp_invert_def using A0 A1 by (metis less_numeral_extra(1) nth_take) then have "Qp_invert (take 1 x) # drop 1 x = x" using A0 A1 A2 by (metis (no_types, lifting) Cons_eq_append_conv Qp_invert_def \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ append_take_drop_id inv_in_frac(2) le_add_same_cancel1 self_append_conv2 take_closed Qp.to_R1_to_R Qp.to_R_pow_closed zero_le) then show ?thesis using A2 by presburger qed qed qed qed qed have 10: "is_semialgebraic (1+k) {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" proof- have "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} = V\<^bsub>Q\<^sub>p\<^esub> (Suc k) (pvar Q\<^sub>p 0)" unfolding zero_set_def using eval_pvar[of 0 "Suc k"] Qp.cring_axioms by blast then show ?thesis using is_zero_set_imp_basic_semialg pvar_closed[of 0 "Suc k"] Qp.cring_axioms is_zero_set_imp_semialg plus_1_eq_Suc zero_less_Suc by presburger qed have 11: "is_semialgebraic (1+k) ({xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S)" using 10 assms basic_semialg_is_semialgebraic intersection_is_semialg by blast have 12: "(partial_pullback 1 Qp_invert k S) = ((partial_pullback 1 Qp_invert k S) \ Nz) \ ((partial_pullback 1 Qp_invert k S) - Nz)" by blast have 13: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) - Nz)" using 11 9 by metis show ?thesis using 8 12 13 by (metis "7" Int_Diff_Un Int_commute plus_1_eq_Suc union_is_semialgebraic) qed lemma Qp_invert_is_semialg: "is_semialg_function 1 Qp_invert" proof(rule is_semialg_functionI') show 0: "Qp_invert \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" then obtain a where a_def: "x = [a]" by (metis Qp.to_R1_to_R) have a_closed: "a \ carrier Q\<^sub>p" using a_def A cartesian_power_concat(1) last_closed' by blast show " Qp_invert x \ carrier Q\<^sub>p" apply(cases "a = \") unfolding Qp_invert_def using a_def a_closed apply (meson Qp.to_R_to_R1) by (metis a_closed a_def inv_in_frac(1) Qp.to_R_to_R1) qed show "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" using Qp_invert_basic_semialg by blast qed lemma Taylor_deg_1_expansion'': assumes "f \ carrier Q\<^sub>p_x" assumes "\n. f n \ \\<^sub>p" assumes "a \ \\<^sub>p " assumes "b \ \\<^sub>p" shows "\c c' c''. c = to_fun f a \ c' = deriv f a \ c \ \\<^sub>p \ c' \ \\<^sub>p \c'' \ \\<^sub>p \ to_fun f (b) = c \ c' \ (b \ a) \ (c'' \ (b \ a)[^](2::nat))" proof- obtain S where S_def: "S = (Q\<^sub>p \ carrier := \\<^sub>p \)" by blast have 1: "f \ carrier (UP S)" unfolding S_def using val_ring_subring UPQ.poly_cfs_subring[of \\<^sub>p f] assms by blast have 2: " f \ carrier (UP (Q\<^sub>p\carrier := \\<^sub>p\))" using val_ring_subring 1 assms poly_cfs_subring[of \\<^sub>p] by blast have 3: "\c\\\<^sub>p. f \ b = f \ a \ UPQ.deriv f a \ (b \ a) \ c \ (b \ a) [^] (2::nat)" using UP_subring_taylor_appr'[of \\<^sub>p f b a] UP_subring_taylor_appr[of \\<^sub>p f b a] val_ring_subring 1 2 assms by blast then show ?thesis using UP_subring_taylor_appr[of \\<^sub>p f b a] assms UP_subring_deriv_closed[of \\<^sub>p f a] UP_subring_eval_closed[of \\<^sub>p f a] 2 val_ring_subring by blast qed end (**************************************************************************************************) (**************************************************************************************************) subsection\Sets Defined by Residues of Valuation Ring Elements\ (**************************************************************************************************) (**************************************************************************************************) sublocale padic_fields < Res: cring "Zp_res_ring (Suc n)" using p_residues residues.cring by blast context padic_fields begin definition Qp_res where "Qp_res x n = to_Zp x n " lemma Qp_res_closed: assumes "x \ \\<^sub>p" shows "Qp_res x n \ carrier (Zp_res_ring n)" unfolding Qp_res_def using assms val_ring_memE residue_closed to_Zp_closed by blast lemma Qp_res_add: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_sum to_Zp_add by presburger lemma Qp_res_mult: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_prod to_Zp_mult by presburger lemma Qp_res_diff: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_diff to_Zp_minus by (meson val_ring_res) lemma Qp_res_zero: shows "Qp_res \ n = 0" unfolding Qp_res_def to_Zp_zero using residue_of_zero(2) by blast lemma Qp_res_one: assumes "n > 0" shows "Qp_res \ n = (1::int)" using assms unfolding Qp_res_def to_Zp_one using residue_of_one(2) by blast lemma Qp_res_nat_inc: shows "Qp_res ([(n::nat)]\\) n = n mod p^n" unfolding Qp_res_def unfolding to_Zp_nat_inc using Zp_nat_inc_res by blast lemma Qp_res_int_inc: shows "Qp_res ([(k::int)]\\) n = k mod p^n" unfolding Qp_res_def unfolding to_Zp_int_inc using Zp_int_inc_res by blast lemma Qp_poly_res_monom: assumes "a \ \\<^sub>p" assumes "x \ \\<^sub>p" assumes "Qp_res a n = 0" assumes "k > 0" shows "Qp_res (up_ring.monom (UP Q\<^sub>p) a k \ x) n = 0" proof- have 0: "up_ring.monom (UP Q\<^sub>p) a k \ x = a \ x [^] k" apply(rule UPQ.to_fun_monom[of a x k]) using assms val_ring_memE apply blast using assms val_ring_memE by blast have 1: "x[^]k \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast show ?thesis unfolding 0 using Qp_res_mult[of a "x[^]k" n] assms using "1" residue_times_zero_r by presburger qed lemma Qp_poly_res_zero: assumes "q \ carrier (UP Q\<^sub>p)" assumes "\i. q i \ \\<^sub>p" assumes "\i. Qp_res (q i) n = 0" assumes "x \ \\<^sub>p" shows "Qp_res (q \ x) n = 0" proof- have "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0" proof(rule UPQ.poly_induct[of q], rule assms, rule ) fix p assume A: "p \ carrier (UP Q\<^sub>p)" " deg Q\<^sub>p p = 0" " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" have 0: "p \ x = p 0" using assms by (metis A(1) A(2) val_ring_memE UPQ.ltrm_deg_0 UPQ.to_fun_ctrm) show "Qp_res (p \ x) n = 0" unfolding 0 using A by blast next fix p assume A0: "(\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0)" "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p" show "(\i. p i \ \\<^sub>p \ Qp_res (p i) n = 0) \ Qp_res (p \ x) n = 0" proof assume A1: " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" obtain k where k_def: "k = deg Q\<^sub>p p" by blast obtain q where q_def: "q = UPQ.trunc p" by blast have q_closed: "q \ carrier (UP Q\<^sub>p)" unfolding q_def using A0(2) UPQ.trunc_closed by blast have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p p" unfolding q_def using A0(2) A0(3) UPQ.trunc_degree by blast have 9: "\i. i < deg Q\<^sub>p p \ q i = p i" unfolding q_def using A0(2) UPQ.trunc_cfs by blast have 90: "\i. \ i < deg Q\<^sub>p p \ q i = \" unfolding q_def proof - fix i :: nat assume "\ i < deg Q\<^sub>p p" then have "deg Q\<^sub>p q < i" using q_deg by linarith then show "Cring_Poly.truncate Q\<^sub>p p i = \" using UPQ.deg_gtE q_closed q_def by blast qed have 10: "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0)" proof fix i show "q i \ \\<^sub>p \ Qp_res (q i) n = 0" apply(cases "i < deg Q\<^sub>p p") using A1 9[of i] apply presburger unfolding q_def using Qp_res_zero 90 by (metis q_def zero_in_val_ring) qed have 11: "Qp_res (q \ x) n = 0" using 10 A1 A0 q_closed q_deg by blast have 12: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k" unfolding k_def q_def using A0(2) UPQ.trunc_simps(1) by blast have 13: "p \ x = q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x" proof- have 0: " (q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k) \ x = q \ x \ up_ring.monom (UP Q\<^sub>p) (p k) k \ x" apply(rule UPQ.to_fun_plus) using A0(2) UPQ.ltrm_closed k_def apply blast unfolding q_def apply(rule UPQ.trunc_closed, rule A0) using assms val_ring_memE by blast show ?thesis using 0 12 by metis qed have 14: "(up_ring.monom (UP Q\<^sub>p) (p k) k) \ x \ \\<^sub>p" apply(rule val_ring_poly_eval) using A0(2) UPQ.ltrm_closed k_def apply blast using UPQ.cfs_monom[of "p k" k ] A1 zero_in_val_ring using A0(2) UPQ.ltrm_cfs k_def apply presburger using assms(4) by blast have 15: "Qp_res ((up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" apply(rule Qp_poly_res_monom) using A1 apply blast using assms apply blast using A1 apply blast unfolding k_def using A0 by blast have 16: "Qp_res (q \ x) n = 0" using A0 10 11 by blast have 17: "q \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule q_closed) using 10 apply blast by(rule assms) have 18: "Qp_res (q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" using Qp_res_add[of "q \ x" "up_ring.monom (UP Q\<^sub>p) (p k) k \ x" n] 14 17 unfolding 15 16 by (metis "10" Qp_res_add UPQ.cfs_add UPQ.coeff_of_sum_diff_degree0 q_closed q_deg) show "Qp_res (p \ x) n = 0" using 13 18 by metis qed qed thus ?thesis using assms by blast qed lemma Qp_poly_res_eval_0: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "\i. g i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" shows "Qp_res (f \ x) n = Qp_res (g \ x) n" proof- obtain F where F_def: "F = f \\<^bsub>UP Q\<^sub>p\<^esub>g" by blast have F_closed: "F \ carrier (UP Q\<^sub>p)" unfolding F_def using assms by blast have F_cfs: "\i. F i = (f i) \ (g i)" unfolding F_def using assms UPQ.cfs_minus by blast have F_cfs_res: "\i. Qp_res (F i) n = Qp_res (f i) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g i) n" unfolding F_cfs apply(rule Qp_res_diff) using assms apply blast using assms by blast have 0: "\i. Qp_res (f i) n = Qp_res (g i) n" using assms by blast have F_cfs_res': "\i. Qp_res (F i) n = 0" unfolding F_cfs_res 0 by (metis diff_self mod_0 residue_minus) have 1: "\i. F i \ \\<^sub>p" unfolding F_cfs using assms using val_ring_minus_closed by blast have 2: "Qp_res (F \ x) n = 0" by(rule Qp_poly_res_zero, rule F_closed, rule 1, rule F_cfs_res', rule assms) have 3: "F \ x = f \ x \ g \ x" unfolding F_def using assms by (meson assms UPQ.to_fun_diff val_ring_memE) have 4: "Qp_res (F \ x) n = Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n" unfolding 3 apply(rule Qp_res_diff, rule val_ring_poly_eval, rule assms) using assms apply blast using assms apply blast apply(rule val_ring_poly_eval, rule assms) using assms apply blast by(rule assms) have 5: "f \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule assms) using assms apply blast using assms by blast have 6: "g \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule assms) using assms apply blast by(rule assms) show "Qp_res (f \ x) n = Qp_res (g \ x) n" using 5 6 2 Qp_res_closed[of "f \ x" n] Qp_res_closed[of "g \ x" n] unfolding 4 proof - assume "Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n = 0" then show ?thesis by (metis (no_types) Qp_res_def 5 6 res_diff_zero_fact(1) residue_of_diff to_Zp_closed val_ring_memE) qed qed lemma Qp_poly_res_eval_1: assumes "f \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "Qp_res x n = Qp_res y n" shows "Qp_res (f \ x) n = Qp_res (f \ y) n" proof- have "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" apply(rule UPQ.poly_induct[of f], rule assms) proof fix f assume A: "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f = 0" "\i. f i \ \\<^sub>p" show "Qp_res (f \ x) n = Qp_res (f \ y) n" proof- obtain a where a_def: "a \ carrier Q\<^sub>p \ f = to_polynomial Q\<^sub>p a" using assms by (metis A(1) A(2) UPQ.lcf_closed UPQ.to_poly_inverse) have a_eq: "f = to_polynomial Q\<^sub>p a" using a_def by blast have 0: "f \ x = a" using a_def assms unfolding a_eq by (meson UPQ.to_fun_to_poly val_ring_memE) have 1: "f \ y = a" using a_def assms unfolding a_eq by (meson UPQ.to_fun_to_poly val_ring_memE) show " Qp_res (f \ x) n = Qp_res (f \ y) n" unfolding 0 1 by blast qed next fix f assume A: " (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p f \ (\i. q i \ \\<^sub>p) \ Qp_res (q \ x) n = Qp_res (q \ y) n)" "f \ carrier (UP Q\<^sub>p)" " 0 < deg Q\<^sub>p f" show "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" proof assume A1: "\i. f i \ \\<^sub>p" obtain q where q_def: "q = UPQ.trunc f" by blast have q_closed: "q \ carrier (UP Q\<^sub>p)" using q_def A UPQ.trunc_closed by presburger have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p f" using q_def A UPQ.trunc_degree by blast have q_cfs: "\i. q i \ \\<^sub>p" proof fix i show "q i \ \\<^sub>p" apply(cases "i < deg Q\<^sub>p f") unfolding q_def using A A1 UPQ.trunc_cfs apply presburger using q_deg q_closed proof - assume "\ i < deg Q\<^sub>p f" then have "deg Q\<^sub>p f \ i" by (meson diff_is_0_eq neq0_conv zero_less_diff) then show "Cring_Poly.truncate Q\<^sub>p f i \ \\<^sub>p" by (metis (no_types) UPQ.deg_eqI diff_is_0_eq' le_trans nat_le_linear neq0_conv q_closed q_def q_deg zero_in_val_ring zero_less_diff) qed qed hence 0: "Qp_res (q \ x) n = Qp_res (q \ y) n" using A q_closed q_deg by blast have 1: "Qp_res (UPQ.ltrm f \ x) n = Qp_res (UPQ.ltrm f \ y) n" proof- have 10: "UPQ.ltrm f \ x = (f (deg Q\<^sub>p f)) \ x[^](deg Q\<^sub>p f)" using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger have 11: "UPQ.ltrm f \ y = (f (deg Q\<^sub>p f)) \ y[^](deg Q\<^sub>p f)" using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger obtain d where d_def: "d = deg Q\<^sub>p f" by blast have 12: "Qp_res (x[^]d) n = Qp_res (y[^]d) n" apply(induction d) using Qp.nat_pow_0 apply presburger using Qp_res_mult assms Qp.nat_pow_Suc val_ring_nat_pow_closed by presburger hence 13: "Qp_res (x [^] deg Q\<^sub>p f) n = Qp_res (y [^] deg Q\<^sub>p f) n" unfolding d_def by blast have 14: "x [^] deg Q\<^sub>p f \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast have 15: "y [^] deg Q\<^sub>p f \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast have 16: "Qp_res (f (deg Q\<^sub>p f) \ x [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (x [^] deg Q\<^sub>p f) n" apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " x[^](deg Q\<^sub>p f)" n]) using A1 apply blast by(rule 14) have 17: "Qp_res (f (deg Q\<^sub>p f) \ y [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (y [^] deg Q\<^sub>p f) n" apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " y[^](deg Q\<^sub>p f)" n]) using A1 apply blast by(rule 15) show ?thesis unfolding 10 11 16 17 13 by blast qed have f_decomp: "f = q \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.ltrm f" using A unfolding q_def using UPQ.trunc_simps(1) by blast have 2: "f \ x = q \ x \ (UPQ.ltrm f \ x)" using A f_decomp q_closed q_cfs by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(2)) have 3: "f \ y = q \ y \ (UPQ.ltrm f \ y)" using A f_decomp q_closed q_cfs by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(3)) show 4: " Qp_res (f \ x) n = Qp_res (f \ y) n " unfolding 2 3 using assms q_cfs Qp_res_add 0 1 by (metis (no_types, opaque_lifting) "2" "3" A(2) A1 Qp_res_def poly_eval_cong) qed qed thus ?thesis using assms by blast qed lemma Qp_poly_res_eval_2: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "\i. g i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" assumes "Qp_res x n = Qp_res y n" shows "Qp_res (f \ x) n = Qp_res (g \ y) n" proof- have 0: "Qp_res (f \ x) n = Qp_res (g \ x) n" using Qp_poly_res_eval_0 assms by blast have 1: "Qp_res (g \ x) n = Qp_res (g \ y) n" using Qp_poly_res_eval_1 assms by blast show ?thesis unfolding 0 1 by blast qed definition poly_res_class where "poly_res_class n d f = {q \ carrier (UP Q\<^sub>p). deg Q\<^sub>p q \ d \ (\i. q i \ \\<^sub>p \ Qp_res (f i) n = Qp_res (q i) n) }" lemma poly_res_class_closed: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "deg Q\<^sub>p f \ d" assumes "deg Q\<^sub>p g \ d" assumes "g \ poly_res_class n d f" shows "poly_res_class n d f = poly_res_class n d g" unfolding poly_res_class_def apply(rule equalityI) apply(rule subsetI) unfolding mem_Collect_eq apply(rule conjI, blast, rule conjI, blast) using assms unfolding poly_res_class_def mem_Collect_eq apply presburger apply(rule subsetI) unfolding mem_Collect_eq apply(rule conjI, blast, rule conjI, blast) using assms unfolding poly_res_class_def mem_Collect_eq by presburger lemma poly_res_class_memE: assumes "f \ poly_res_class n d g" shows "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f \ d" "f i \ \\<^sub>p" "Qp_res (g i) n = Qp_res (f i) n" using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq by blast definition val_ring_polys where "val_ring_polys = {f \ carrier (UP Q\<^sub>p). (\i. f i \ \\<^sub>p)} " lemma val_ring_polys_closed: "val_ring_polys \ carrier (UP Q\<^sub>p)" unfolding val_ring_polys_def by blast lemma val_ring_polys_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "\i. f i \ \\<^sub>p" shows "f \ val_ring_polys" using assms unfolding val_ring_polys_def by blast lemma val_ring_polys_memE: assumes "f \ val_ring_polys" shows "f \ carrier (UP Q\<^sub>p)" "f i \ \\<^sub>p" using assms unfolding val_ring_polys_def apply blast using assms unfolding val_ring_polys_def by blast definition val_ring_polys_grad where "val_ring_polys_grad d = {f \ val_ring_polys. deg Q\<^sub>p f \ d}" lemma val_ring_polys_grad_closed: "val_ring_polys_grad d \ val_ring_polys" unfolding val_ring_polys_grad_def by blast lemma val_ring_polys_grad_closed': "val_ring_polys_grad d \ carrier (UP Q\<^sub>p)" unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma val_ring_polys_grad_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "\i. f i \ \\<^sub>p" assumes "deg Q\<^sub>p f \ d" shows "f \ val_ring_polys_grad d" using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma val_ring_polys_grad_memE: assumes "f \ val_ring_polys_grad d" shows "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f \ d" "f i \ \\<^sub>p" using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma poly_res_classes_in_val_ring_polys_grad: assumes "f \ val_ring_polys_grad d" shows "poly_res_class n d f \ val_ring_polys_grad d" apply(rule subsetI, rule val_ring_polys_grad_memI) apply(rule poly_res_class_memE[of _ n d f], blast) apply(rule poly_res_class_memE[of _ n d f], blast) by(rule poly_res_class_memE[of _ n d f], blast) lemma poly_res_class_disjoint: assumes "f \ val_ring_polys_grad d" assumes "f \ poly_res_class n d g" shows "poly_res_class n d f \ poly_res_class n d g = {}" apply(rule equalityI) apply(rule subsetI) using assms unfolding poly_res_class_def mem_Collect_eq Int_iff apply (metis val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3)) by blast lemma poly_res_class_refl: assumes "f \ val_ring_polys_grad d" shows "f \ poly_res_class n d f" unfolding poly_res_class_def mem_Collect_eq using assms val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3) by blast lemma poly_res_class_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "deg Q\<^sub>p f \ d" assumes "\i. f i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" shows "f \ poly_res_class n d g" unfolding poly_res_class_def mem_Collect_eq using assms by metis definition poly_res_classes where "poly_res_classes n d = poly_res_class n d ` val_ring_polys_grad d" lemma poly_res_classes_disjoint: assumes "A \ poly_res_classes n d" assumes "B \ poly_res_classes n d" assumes "g \ A - B" shows "A \ B = {}" proof- obtain a where a_def: "a \ val_ring_polys_grad d \ A = poly_res_class n d a" using assms unfolding poly_res_classes_def by blast obtain b where b_def: "b \ val_ring_polys_grad d \ B = poly_res_class n d b" using assms unfolding poly_res_classes_def by blast have 0: "\f. f \ A \ B \ False" proof- fix f assume A: "f \ A \ B" have 1: "\i. Qp_res (g i) n \ Qp_res (f i) n" proof(rule ccontr) assume B: "\i. Qp_res (g i) n \ Qp_res (f i) n" then have 2: "\i. Qp_res (g i) n = Qp_res (f i) n" by blast have 3: "g \ poly_res_class n d a" using a_def assms by blast have 4: "\i. Qp_res (b i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using assms A b_def by blast have 5: "\i. Qp_res (a i) n = Qp_res (g i) n" apply(rule poly_res_class_memE[of _ n d]) using assms A a_def by blast have 6: "g \ poly_res_class n d b" apply(rule poly_res_class_memI, rule poly_res_class_memE[of _ n d a], rule 3, rule poly_res_class_memE[of _ n d a], rule 3, rule poly_res_class_memE[of _ n d a], rule 3) unfolding 2 4 by blast show False using 6 b_def assms by blast qed then obtain i where i_def: "Qp_res (g i) n \ Qp_res (f i) n" by blast have 2: "\i. Qp_res (a i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using A a_def by blast have 3: "\i. Qp_res (b i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using A b_def by blast have 4: "\i. Qp_res (a i) n = Qp_res (g i) n" apply(rule poly_res_class_memE[of _ n d]) using assms a_def by blast show False using i_def 2 unfolding 4 2 by blast qed show ?thesis using 0 by blast qed definition int_fun_to_poly where "int_fun_to_poly (f::nat \ int) i = [(f i)]\\" lemma int_fun_to_poly_closed: assumes "\i. i > d \ f i = 0" shows "int_fun_to_poly f \ carrier (UP Q\<^sub>p)" apply(rule UPQ.UP_car_memI[of d]) using assms unfolding int_fun_to_poly_def using Qp.int_inc_zero apply presburger by(rule Qp.int_inc_closed) lemma int_fun_to_poly_deg: assumes "\i. i > d \ f i = 0" shows "deg Q\<^sub>p (int_fun_to_poly f) \ d" apply(rule UPQ.deg_leqI, rule int_fun_to_poly_closed, rule assms, blast) unfolding int_fun_to_poly_def using assms using Qp.int_inc_zero by presburger lemma Qp_res_mod_triv: assumes "a \ \\<^sub>p" shows "Qp_res a n mod p ^ n = Qp_res a n" using assms Qp_res_closed[of a n] by (meson mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) lemma int_fun_to_poly_is_class_wit: assumes "f \ poly_res_class n d g" shows "(int_fun_to_poly (\i::nat. Qp_res (f i) n)) \ poly_res_class n d g" proof(rule poly_res_class_memI[of ], rule int_fun_to_poly_closed[of d]) show 0: "\i. d < i \ Qp_res (f i) n = 0" proof- fix i assume A: "d < i" hence 0: "deg Q\<^sub>p f < i" using A assms poly_res_class_memE(2)[of f n d g] by linarith have 1: "f i = \" using 0 assms poly_res_class_memE[of f n d g] using UPQ.UP_car_memE(2) by blast show "Qp_res (f i) n = 0" unfolding 1 Qp_res_zero by blast qed show "deg Q\<^sub>p (int_fun_to_poly (\i. Qp_res (f i) n)) \ d" by(rule int_fun_to_poly_deg, rule 0, blast) show "\i. int_fun_to_poly (\i. Qp_res (f i) n) i \ \\<^sub>p" unfolding int_fun_to_poly_def using Qp.int_mult_closed Qp_val_ringI val_of_int_inc by blast show "\i. Qp_res (int_fun_to_poly (\i. Qp_res (f i) n) i) n = Qp_res (g i) n" unfolding int_fun_to_poly_def Qp_res_int_inc using Qp_res_mod_triv assms poly_res_class_memE(4) Qp_res_closed UPQ.cfs_closed by (metis poly_res_class_memE(3)) qed lemma finite_support_funs_finite: "finite (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" proof- have 0: "finite (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n))" apply(rule finite_PiE, blast) using residue_ring_card[of n] by blast obtain g where g_def: "g = (\f. (\i::nat. if i \ {..d} then f i else (0::int)))" by blast have 1: "g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) = (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" proof(rule equalityI, rule subsetI) fix x assume A: "x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" obtain f where f_def: "f \ (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) \ x = g f" using A by blast have x_eq: "x = g f" using f_def by blast show "x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" proof(rule, rule) fix i assume A: "i \ {..d}" show "x i \ carrier (Zp_res_ring n)" proof(cases "i \ {..d}") case True then have T0: "f i \ carrier (Zp_res_ring n)" using f_def by blast have "x i = f i" unfolding x_eq g_def using True by metis thus ?thesis using T0 by metis next case False then have F0: "x i = 0" unfolding x_eq g_def by metis show ?thesis unfolding F0 by (metis residue_mult_closed residue_times_zero_r) qed next show "x \ {f. \i>d. f i = 0}" proof(rule, rule, rule) fix i assume A: "d < i" then have 0: "i \ {..d}" by simp thus "x i = 0" unfolding x_eq g_def by metis qed qed next show "({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" proof(rule subsetI) fix x assume A: " x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" show " x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" proof- obtain h where h_def: "h = restrict x {..d}" by blast have 0: "\i. i \ {..d} \ h i = x i" unfolding h_def restrict_apply by metis have 1: "\i. i \ {..d} \ h i = undefined" unfolding h_def restrict_apply by metis have 2: "\i. i \ {..d} \ h i \ carrier (Zp_res_ring n)" using A 0 unfolding 0 by blast have 3: "h \ {..d} \\<^sub>E carrier (residue_ring (p ^ n))" by(rule, rule 2, blast, rule 1, blast) have 4: "\i. i \ {..d} \ x i = 0" using A unfolding Int_iff mem_Collect_eq by (metis atMost_iff eq_imp_le le_simps(1) linorder_neqE_nat) have 5: "x = g h" proof fix i show "x i = g h i" unfolding g_def apply(cases "i \ {..d}") using 0 apply metis unfolding 4 by metis qed show ?thesis unfolding 5 using 3 by blast qed qed qed have 2: "finite (g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)))" using 0 by blast show ?thesis using 2 unfolding 1 by blast qed lemma poly_res_classes_finite: "finite (poly_res_classes n d)" proof- have 0: "poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}) = poly_res_classes n d" proof(rule equalityI, rule subsetI) fix x assume A: " x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" then obtain f where f_def: "f \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} \ x = poly_res_class n d (int_fun_to_poly f)" by blast have x_eq: "x = poly_res_class n d (int_fun_to_poly f)" using f_def by blast show "x \ poly_res_classes n d" proof- have 0: "int_fun_to_poly f \ val_ring_polys_grad d" apply(rule val_ring_polys_grad_memI, rule int_fun_to_poly_closed[of d]) using f_def apply blast using int_fun_to_poly_def apply (metis Qp.int_inc_closed padic_fields.int_fun_to_poly_def padic_fields.val_of_int_inc padic_fields_axioms val_ring_memI) apply(rule int_fun_to_poly_deg) using f_def by blast show ?thesis unfolding poly_res_classes_def x_eq using 0 by blast qed next show "poly_res_classes n d \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" proof(rule subsetI) fix x assume A: " x \ poly_res_classes n d" show "x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" proof- obtain f where f_def: "f \ val_ring_polys_grad d \ x = poly_res_class n d f" using A unfolding poly_res_classes_def by blast have x_eq: "x = poly_res_class n d f" using f_def by blast obtain h where h_def: "h = (\i::nat. Qp_res (f i) n)" by blast have 0: "\i. i > d \ f i = \" proof- fix i assume A: "i > d" have "i > deg Q\<^sub>p f" apply(rule le_less_trans[of _ d]) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast by(rule A) then show "f i = \" using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq using UPQ.deg_leE by blast qed have 1: "\i. i > d \ h i = 0" unfolding h_def 0 Qp_res_zero by blast have 2: "x = poly_res_class n d (int_fun_to_poly h)" unfolding x_eq apply(rule poly_res_class_closed) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast apply(rule int_fun_to_poly_closed[of d], rule 1, blast) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast apply(rule int_fun_to_poly_deg, rule 1, blast) unfolding h_def apply(rule int_fun_to_poly_is_class_wit, rule poly_res_class_refl) using f_def by blast have 3: "h \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" apply(rule , rule ) unfolding h_def apply(rule Qp_res_closed, rule val_ring_polys_grad_memE[of _ d]) using f_def apply blast unfolding mem_Collect_eq apply(rule, rule) unfolding 0 Qp_res_zero by blast show ?thesis unfolding 2 using 3 by blast qed qed qed have 1: "finite (poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}))" using finite_support_funs_finite by blast show ?thesis using 1 unfolding 0 by blast qed lemma Qp_res_eq_zeroI: assumes "a \ \\<^sub>p" assumes "val a \ n" shows "Qp_res a n = 0" proof- have 0: "val_Zp (to_Zp a) \ n" using assms to_Zp_val by presburger have 1: "to_Zp a n = 0" apply(rule zero_below_val_Zp, rule to_Zp_closed) using val_ring_closed assms apply blast by(rule 0) thus ?thesis unfolding Qp_res_def by blast qed lemma Qp_res_eqI: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "Qp_res (a \ b) n = 0" shows "Qp_res a n = Qp_res b n" using assms by (metis Qp_res_def val_ring_memE res_diff_zero_fact(1) to_Zp_closed to_Zp_minus) lemma Qp_res_eqI': assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "val (a \ b) \ n" shows "Qp_res a n = Qp_res b n" apply(rule Qp_res_eqI, rule assms, rule assms, rule Qp_res_eq_zeroI) using assms Q\<^sub>p_def Zp_def \_def padic_fields.val_ring_minus_closed padic_fields_axioms apply blast by(rule assms) lemma Qp_res_eqE: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "Qp_res a n = Qp_res b n" shows "val (a \ b) \ n" proof- have 0: "val (a \ b) = val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b)" using assms by (metis to_Zp_minus to_Zp_val val_ring_minus_closed) have 1: "(to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) n = 0" using assms unfolding Qp_res_def by (meson val_ring_memE res_diff_zero_fact'' to_Zp_closed) have 2: "val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) \ n" apply(cases "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>") proof - assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>" have "\n. eint (int n) \ val_Zp \\<^bsub>Z\<^sub>p\<^esub>" by (metis (no_types) Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) then show ?thesis using a1 by presburger next assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ \\<^bsub>Z\<^sub>p\<^esub>" have 00: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ carrier Z\<^sub>p" using assms by (meson val_ring_memE Zp.cring_simprules(4) to_Zp_closed) show ?thesis using 1 a1 ord_Zp_geq[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b" n] 00 val_ord_Zp[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b"] eint_ord_code by metis qed thus ?thesis unfolding 0 by blast qed lemma notin_closed: "(\ ((c::eint) \ x \ x \ d)) = (x < c \ d < x)" by auto lemma Qp_res_neqI: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "val (a \ b) < n" shows "Qp_res a n \ Qp_res b n" apply(rule ccontr) using Qp_res_eqE[of a b n] assms using notin_closed by blast lemma Qp_res_equal: assumes "a \ \\<^sub>p" assumes "l = Qp_res a n" shows "Qp_res a n = Qp_res ([l]\\) n " unfolding Qp_res_int_inc assms using assms Qp_res_mod_triv by presburger definition Qp_res_class where "Qp_res_class n b = {a \ \\<^sub>p. Qp_res a n = Qp_res b n}" definition Qp_res_classes where "Qp_res_classes n = Qp_res_class n ` \\<^sub>p" lemma val_ring_int_inc_closed: "[(k::int)]\\ \ \\<^sub>p" proof- have 0: "[(k::int)]\\ = \ ([(k::int)]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using inc_of_int by blast thus ?thesis by blast qed lemma val_ring_nat_inc_closed: "[(k::nat)]\\ \ \\<^sub>p" proof- have 0: "[k]\\ = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using inc_of_nat by blast thus ?thesis by blast qed lemma Qp_res_classes_wits: "Qp_res_classes n = (\l::int. Qp_res_class n ([l]\\)) ` (carrier (Zp_res_ring n))" proof- obtain F where F_def: "F = (\l::int. Qp_res_class n ([l]\\))" by blast have 0: "Qp_res_classes n = F ` (carrier (Zp_res_ring n))" proof(rule equalityI, rule subsetI) fix x assume A: "x \ Qp_res_classes n" then obtain a where a_def: "a \ \\<^sub>p \ x = Qp_res_class n a" unfolding Qp_res_classes_def by blast have 1: "Qp_res a n = Qp_res ([(Qp_res a n)]\\) n " apply(rule Qp_res_equal) using a_def apply blast by blast have 2: "Qp_res_class n a = Qp_res_class n ([(Qp_res a n)]\\)" unfolding Qp_res_class_def using 1 by metis have 3: "x = Qp_res_class n ([(Qp_res a n)]\\)" using a_def unfolding 2 by blast have 4: "a \ \\<^sub>p" using a_def by blast show " x \ F ` carrier (Zp_res_ring n)" unfolding F_def 3 using Qp_res_closed[of a n] 4 by blast next show "F ` carrier (residue_ring (p ^ n)) \ Qp_res_classes n" proof(rule subsetI) fix x assume A: "x \ F ` (carrier (Zp_res_ring n))" then obtain l where l_def: "l \ carrier (Zp_res_ring n) \ x = F l" using A by blast have 0: "x = F l" using l_def by blast show "x \ Qp_res_classes n" unfolding 0 F_def Qp_res_classes_def using val_ring_int_inc_closed by blast qed qed then show ?thesis unfolding F_def by blast qed lemma Qp_res_classes_finite: "finite (Qp_res_classes n)" by (metis Qp_res_classes_wits finite_atLeastLessThan_int finite_imageI p_res_ring_car) definition Qp_cong_set where "Qp_cong_set \ a = {x \ \\<^sub>p. to_Zp x \ = a \}" lemma Qp_cong_set_as_ball: assumes "a \ carrier Z\<^sub>p" assumes "a \ = 0" shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[\]" proof- have 0: "\ a \ carrier Q\<^sub>p" using assms inc_closed[of a] by blast show ?thesis proof show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\]" proof fix x assume A: "x \ Qp_cong_set \ a" show "x \ B\<^bsub>\ \<^esub>[\]" proof(rule c_ballI) show t0: "x \ carrier Q\<^sub>p" using A unfolding Qp_cong_set_def using val_ring_memE by blast show "eint (int \) \ val (x \ \)" proof- have t1: "to_Zp x \ = 0" using A unfolding Qp_cong_set_def by (metis (mono_tags, lifting) assms(2) mem_Collect_eq) have t2: "val_Zp (to_Zp x) \ \" apply(cases "to_Zp x = \\<^bsub>Z\<^sub>p\<^esub>") apply (metis Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) using ord_Zp_geq[of "to_Zp x" \] A unfolding Qp_cong_set_def by (metis (no_types, lifting) val_ring_memE eint_ord_simps(1) t1 to_Zp_closed to_Zp_def val_ord_Zp) then show ?thesis using A unfolding Qp_cong_set_def mem_Collect_eq using val_ring_memE by (metis Qp_res_eqE Qp_res_eq_zeroI Qp_res_zero to_Zp_val zero_in_val_ring) qed qed qed show "B\<^bsub>int \\<^esub>[\] \ Qp_cong_set \ a" proof fix x assume A: "x \ B\<^bsub>int \\<^esub>[\]" then have 0: "val x \ \" using assms c_ballE[of x \ \] by (smt Qp.minus_closed Qp.r_right_minus_eq Qp_diff_diff) have 1: "to_Zp x \ carrier Z\<^sub>p" using A 0 assms c_ballE(1) to_Zp_closed by blast have 2: "x \ \\<^sub>p" using 0 A val_ringI c_ballE by (smt Q\<^sub>p_def Zp_def \_def eint_ord_simps(1) of_nat_0 of_nat_le_0_iff val_ring_ord_criterion padic_fields_axioms val_ord' zero_in_val_ring) then have "val_Zp (to_Zp x) \ \" using 0 1 A assms c_ballE[of x \ \] to_Zp_val by presburger then have "to_Zp x \ = 0" using 1 zero_below_val_Zp by blast then show " x \ Qp_cong_set \ a" unfolding Qp_cong_set_def using assms(2) 2 by (metis (mono_tags, lifting) mem_Collect_eq) qed qed qed lemma Qp_cong_set_as_ball': assumes "a \ carrier Z\<^sub>p" assumes "val_Zp a < eint (int \)" shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[(\ a)]" proof show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\ a]" proof fix x assume A: "x \ Qp_cong_set \ a" then have 0: "to_Zp x \ = a \" unfolding Qp_cong_set_def by blast have 1: "x \ \\<^sub>p" using A unfolding Qp_cong_set_def by blast have 2: "to_Zp x \ carrier Z\<^sub>p" using 1 val_ring_memE to_Zp_closed by blast have 3: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" using 0 assms 2 val_Zp_dist_def val_Zp_dist_res_eq2 by presburger have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) > val_Zp a" using 3 assms(2) less_le_trans[of "val_Zp a" "eint (int \)" "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a)" ] by blast then have 5: "val_Zp (to_Zp x) = val_Zp a" using assms 2 equal_val_Zp by blast have 7: "val (x \ (\ a)) \ \" using 3 5 1 by (metis "2" Zp.minus_closed assms(1) inc_of_diff to_Zp_inc val_of_inc) then show "x \ B\<^bsub>int \\<^esub>[\ a]" using c_ballI[of x \ "\ a"] 1 assms val_ring_memE by blast qed show "B\<^bsub>int \\<^esub>[\ a] \ Qp_cong_set \ a" proof fix x assume A: "x \ B\<^bsub>int \\<^esub>[\ a]" then have 0: "val (x \ \ a) \ \" using c_ballE by blast have 1: "val (\ a) = val_Zp a" using assms Zp_def \_def padic_fields.val_of_inc padic_fields_axioms by metis then have 2: "val (x \ \ a) > val (\ a)" using 0 assms less_le_trans[of "val (\ a)" "eint (int \)" "val (x \ \ a)"] by metis have "\ a \ carrier Q\<^sub>p" using assms(1) inc_closed by blast then have B: "val x = val (\ a)" using 2 A assms c_ballE(1)[of x \ "\ a"] by (metis ultrametric_equal_eq) have 3: "val_Zp (to_Zp x) = val_Zp a" by (metis "1" A \val x = val (\ a)\ assms(1) c_ballE(1) to_Zp_val val_pos val_ringI) have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" using 0 A 3 by (metis B Zp.minus_closed assms(1) c_ballE(1) inc_of_diff to_Zp_closed to_Zp_inc val_of_inc val_pos val_ring_val_criterion) then have 5: "to_Zp x \ = a \" by (meson A Zp.minus_closed assms(1) c_ballE(1) res_diff_zero_fact(1) to_Zp_closed zero_below_val_Zp) have 6: "x \ \\<^sub>p" proof- have "val x \ 0" using B assms 1 val_pos by presburger then show ?thesis using A c_ballE(1) val_ringI by blast qed then show "x \ Qp_cong_set \ a" unfolding Qp_cong_set_def using "5" by blast qed qed lemma Qp_cong_set_is_univ_semialgebraic: assumes "a \ carrier Z\<^sub>p" shows "is_univ_semialgebraic (Qp_cong_set \ a)" proof(cases "a \ = 0") case True then show ?thesis using ball_is_univ_semialgebraic[of \ \] Qp.zero_closed Qp_cong_set_as_ball assms by metis next case False then have "\ \ 0" using assms residues_closed[of a 0] by (meson p_res_ring_0') then obtain n where n_def: "Suc n = \" by (metis lessI less_Suc_eq_0_disj) then have "val_Zp a < eint (int \)" using below_val_Zp_zero[of a n] by (smt False assms eint_ile eint_ord_simps(1) eint_ord_simps(2) zero_below_val_Zp) then show ?thesis using ball_is_univ_semialgebraic[of "\ a" \] Qp.zero_closed Qp_cong_set_as_ball'[of a \] assms inc_closed by presburger qed lemma constant_res_set_semialg: assumes "l \ carrier (Zp_res_ring n)" shows "is_univ_semialgebraic {x \ \\<^sub>p. Qp_res x n = l}" proof- have 0: "{x \ \\<^sub>p. Qp_res x n = l} = Qp_cong_set n ([l]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" apply(rule equalityI') unfolding mem_Collect_eq Qp_cong_set_def Qp_res_def apply (metis val_ring_memE Zp_int_inc_rep nat_le_linear p_residue_padic_int to_Zp_closed) using assms by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) show ?thesis unfolding 0 apply(rule Qp_cong_set_is_univ_semialgebraic) by(rule Zp.int_inc_closed) qed end end diff --git a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy --- a/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy +++ b/thys/Planarity_Certificates/Planarity/Executable_Permutations.thy @@ -1,1033 +1,1034 @@ section \Permutations as Products of Disjoint Cycles\ theory Executable_Permutations imports "HOL-Combinatorics.Permutations" Graph_Theory.Auxiliary List_Aux begin subsection \Cyclic Permutations\ definition list_succ :: "'a list \ 'a \ 'a" where "list_succ xs x = (if x \ set xs then xs ! ((index xs x + 1) mod length xs) else x)" text \ We demonstrate the functions on the following simple lemmas @{lemma "list_succ [1 :: int, 2, 3] 1 = 2" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 2 = 3" by code_simp} @{lemma "list_succ [1 :: int, 2, 3] 3 = 1" by code_simp} \ lemma list_succ_altdef: "list_succ xs x = (let n = index xs x in if n + 1 = length xs then xs ! 0 else if n + 1 < length xs then xs ! (n + 1) else x)" using index_le_size[of xs x] unfolding list_succ_def index_less_size_conv[symmetric] by (auto simp: Let_def) lemma list_succ_Nil: "list_succ [] = id" by (simp add: list_succ_def fun_eq_iff) lemma list_succ_singleton: "list_succ [x] = list_succ []" by (simp add: fun_eq_iff list_succ_def) lemma list_succ_short: assumes "length xs < 2" shows "list_succ xs = id" using assms by (cases xs) (rename_tac [2] y ys, case_tac [2] ys, auto simp: list_succ_Nil list_succ_singleton) lemma list_succ_simps: "index xs x + 1 = length xs \ list_succ xs x = xs ! 0" "index xs x + 1 < length xs \ list_succ xs x = xs ! (index xs x + 1)" "length xs \ index xs x \ list_succ xs x = x" by (auto simp: list_succ_altdef) lemma list_succ_not_in: assumes "x \ set xs" shows "list_succ xs x = x" using assms by (auto simp: list_succ_def) lemma list_succ_list_succ_rev: assumes "distinct xs" shows "list_succ (rev xs) (list_succ xs x) = x" proof - { assume "index xs x + 1 < length xs" moreover then have "length xs - Suc (Suc (length xs - Suc (Suc (index xs x)))) = index xs x" by linarith ultimately have ?thesis using assms by (simp add: list_succ_def index_rev index_nth_id rev_nth) } moreover { assume A: "index xs x + 1 = length xs" moreover from A have "xs \ []" by auto moreover with A have "last xs = xs ! index xs x" by (cases "length xs") (auto simp: last_conv_nth) ultimately have ?thesis using assms by (auto simp add: list_succ_def rev_nth index_rev index_nth_id last_conv_nth) } moreover { assume A: "index xs x \ length xs" then have "x \ set xs" by (metis index_less less_irrefl) then have ?thesis by (auto simp: list_succ_def) } - ultimately show ?thesis by (metis discrete le_less not_less) + ultimately show ?thesis + by linarith qed lemma inj_list_succ: "distinct xs \ inj (list_succ xs)" by (metis injI list_succ_list_succ_rev) lemma inv_list_succ_eq: "distinct xs \ inv (list_succ xs) = list_succ (rev xs)" by (metis distinct_rev inj_imp_inv_eq inj_list_succ list_succ_list_succ_rev) lemma bij_list_succ: "distinct xs \ bij (list_succ xs)" by (metis bij_def inj_list_succ distinct_rev list_succ_list_succ_rev surj_def) lemma list_succ_permutes: assumes "distinct xs" shows "list_succ xs permutes set xs" using assms by (auto simp: permutes_conv_has_dom bij_list_succ has_dom_def list_succ_def) lemma permutation_list_succ: assumes "distinct xs" shows "permutation (list_succ xs)" using list_succ_permutes[OF assms] by (auto simp: permutation_permutes) lemma list_succ_nth: assumes "distinct xs" "n < length xs" shows "list_succ xs (xs ! n) = xs ! (Suc n mod length xs)" using assms by (auto simp: list_succ_def index_nth_id) lemma list_succ_last[simp]: assumes "distinct xs" "xs \ []" shows "list_succ xs (last xs) = hd xs" using assms by (auto simp: list_succ_def hd_conv_nth) lemma list_succ_rotate1[simp]: assumes "distinct xs" shows "list_succ (rotate1 xs) = list_succ xs" proof (rule ext) fix y show "list_succ (rotate1 xs) y = list_succ xs y" using assms proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases "x = y") case True then have "index (xs @ [y]) y = length xs" using \distinct (x # xs)\ by (simp add: index_append) with True show ?thesis by (cases "xs=[]") (auto simp: list_succ_def nth_append) next case False then show ?thesis apply (cases "index xs y + 1 < length xs") apply (auto simp:list_succ_def index_append nth_append) by (metis Suc_lessI index_less_size_conv mod_self nth_Cons_0 nth_append nth_append_length) qed qed qed lemma list_succ_rotate[simp]: assumes "distinct xs" shows "list_succ (rotate n xs) = list_succ xs" using assms by (induct n) auto lemma list_succ_in_conv: "list_succ xs x \ set xs \ x \ set xs" by (auto simp: list_succ_def not_nil_if_in_set ) lemma list_succ_in_conv1: assumes "A \ set xs = {}" shows "list_succ xs x \ A \ x \ A" by (metis assms disjoint_iff_not_equal list_succ_in_conv list_succ_not_in) lemma list_succ_commute: assumes "set xs \ set ys = {}" shows "list_succ xs (list_succ ys x) = list_succ ys (list_succ xs x)" proof - have "\x. x \ set xs \ list_succ ys x = x" "\x. x \ set ys \ list_succ xs x = x" using assms by (blast intro: list_succ_not_in)+ then show ?thesis by (cases "x \ set xs \ set ys") (auto simp: list_succ_in_conv list_succ_not_in) qed subsection \Arbitrary Permutations\ fun lists_succ :: "'a list list \ 'a \ 'a" where "lists_succ [] x = x" | "lists_succ (xs # xss) x = list_succ xs (lists_succ xss x)" definition distincts :: "'a list list \ bool" where "distincts xss \ distinct xss \ (\xs \ set xss. distinct xs \ xs \ []) \ (\xs \ set xss. \ys \ set xss. xs \ ys \ set xs \ set ys = {})" lemma distincts_distinct: "distincts xss \ distinct xss" by (auto simp: distincts_def) lemma distincts_Nil[simp]: "distincts []" by (simp add: distincts_def) lemma distincts_single: "distincts [xs] \ distinct xs \ xs \ []" by (auto simp add: distincts_def) lemma distincts_Cons: "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (set xs \ (\ys \ set xss. set ys)) = {}" (is "?L \ ?R") proof assume ?L then show ?R by (auto simp: distincts_def) next assume ?R then have "distinct (xs # xss)" apply (auto simp: disjoint_iff_not_equal distincts_distinct) apply (metis length_greater_0_conv nth_mem) done moreover from \?R\ have "\xs \ set (xs # xss). distinct xs \ xs \ []" by (auto simp: distincts_def) moreover from \?R\ have "\xs' \ set (xs # xss). \ys \ set (xs # xss). xs' \ ys \ set xs' \ set ys = {}" by (simp add: distincts_def) blast ultimately show ?L unfolding distincts_def by (intro conjI) qed lemma distincts_Cons': "distincts (xs # xss) \ xs \ [] \ distinct xs \ distincts xss \ (\ys \ set xss. set xs \ set ys = {})" (is "?L \ ?R") unfolding distincts_Cons by blast lemma distincts_rev: "distincts (map rev xss) \ distincts xss" by (simp add: distincts_def distinct_map) lemma length_distincts: assumes "distincts xss" shows "length xss = card (set ` set xss)" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons xs xss) then have "set xs \ set ` set xss" using equals0I[of "set xs"] by (auto simp: distincts_Cons disjoint_iff_not_equal ) with Cons show ?case by (auto simp add: distincts_Cons) qed lemma distincts_remove1: "distincts xss \ distincts (remove1 xs xss)" by (auto simp: distincts_def) lemma distinct_Cons_remove1: "x \ set xs \ distinct (x # remove1 x xs) = distinct xs" by (induct xs) auto lemma set_Cons_remove1: "x \ set xs \ set (x # remove1 x xs) = set xs" by (induct xs) auto lemma distincts_Cons_remove1: "xs \ set xss \ distincts (xs # remove1 xs xss) = distincts xss" by (simp only: distinct_Cons_remove1 set_Cons_remove1 distincts_def) lemma distincts_inj_on_set: assumes "distincts xss" shows "inj_on set (set xss)" by (rule inj_onI) (metis assms distincts_def inf.idem set_empty) lemma distincts_distinct_set: assumes "distincts xss" shows "distinct (map set xss)" using assms by (auto simp: distinct_map distincts_distinct distincts_inj_on_set) lemma distincts_distinct_nth: assumes "distincts xss" "n < length xss" shows "distinct (xss ! n)" using assms by (auto simp: distincts_def) lemma lists_succ_not_in: assumes "x \ (\xs\set xss. set xs)" shows "lists_succ xss x = x" using assms by (induct xss) (auto simp: list_succ_not_in) lemma lists_succ_in_conv: "lists_succ xss x \ (\xs\set xss. set xs) \ x \ (\xs\set xss. set xs)" by (induct xss) (auto simp: list_succ_in_conv lists_succ_not_in list_succ_not_in) lemma lists_succ_in_conv1: assumes "A \ (\xs\set xss. set xs) = {}" shows "lists_succ xss x \ A \ x \ A" by (metis Int_iff assms emptyE lists_succ_in_conv lists_succ_not_in) lemma lists_succ_Cons_pf: "lists_succ (xs # xss) = list_succ xs o lists_succ xss" by auto lemma lists_succ_Nil_pf: "lists_succ [] = id" by (simp add: fun_eq_iff) lemmas lists_succ_simps_pf = lists_succ_Cons_pf lists_succ_Nil_pf lemma lists_succ_permutes: assumes "distincts xss" shows "lists_succ xss permutes (\xs \ set xss. set xs)" using assms proof (induction xss) case Nil then show ?case by auto next case (Cons xs xss) have "list_succ xs permutes (set xs)" using Cons by (intro list_succ_permutes) (simp add: distincts_def in_set_member) moreover have "lists_succ xss permutes (\ys \ set xss. set ys)" using Cons by (auto simp: Cons distincts_def) ultimately show "lists_succ (xs # xss) permutes (\ys \ set (xs # xss). set ys)" using Cons by (auto simp: lists_succ_Cons_pf intro: permutes_compose permutes_subset) qed lemma bij_lists_succ: "distincts xss \ bij (lists_succ xss)" by (induct xss) (auto simp: lists_succ_simps_pf bij_comp bij_list_succ distincts_Cons) lemma lists_succ_snoc: "lists_succ (xss @ [xs]) = lists_succ xss o list_succ xs" by (induct xss) auto lemma inv_lists_succ_eq: assumes "distincts xss" shows "inv (lists_succ xss) = lists_succ (rev (map rev xss))" proof - have *: "\f g. inv (\b. f (g b)) = inv (f o g)" by (simp add: o_def) have **: "lists_succ [] = id" by auto show ?thesis using assms by (induct xss) (auto simp: * ** lists_succ_snoc lists_succ_Cons_pf o_inv_distrib inv_list_succ_eq distincts_Cons bij_list_succ bij_lists_succ) qed lemma lists_succ_remove1: assumes "distincts xss" "xs \ set xss" shows "lists_succ (xs # remove1 xs xss) = lists_succ xss" using assms proof (induct xss) case Nil then show ?case by simp next case (Cons ys xss) show ?case proof cases assume "xs = ys" then show ?case by simp next assume "xs \ ys" with Cons.prems have inter: "set xs \ set ys = {}" and "xs \ set xss" by (auto simp: distincts_Cons) have dists: "distincts (xs # remove1 xs xss)" "distincts (xs # ys # remove1 xs xss)" using \distincts (ys # xss)\ \xs \ set xss\ by (auto simp: distincts_def) have "list_succ xs \ (list_succ ys \ lists_succ (remove1 xs xss)) = list_succ ys \ (list_succ xs \ lists_succ (remove1 xs xss))" using inter unfolding fun_eq_iff comp_def by (subst list_succ_commute) auto also have "\ = list_succ ys o (lists_succ (xs # remove1 xs xss))" using dists by (simp add: lists_succ_Cons_pf distincts_Cons) also have "\ = list_succ ys o lists_succ xss" using \xs \ set xss\ \distincts (ys # xss)\ by (simp add: distincts_Cons Cons.hyps) finally show "lists_succ (xs # remove1 xs (ys # xss)) = lists_succ (ys # xss)" using Cons dists by (auto simp: lists_succ_Cons_pf distincts_Cons) qed qed lemma lists_succ_no_order: assumes "distincts xss" "distincts yss" "set xss = set yss" shows "lists_succ xss = lists_succ yss" using assms proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) have "xs \ set xss" "xs \ set yss" using Cons.prems by (auto dest: distincts_distinct) have "lists_succ xss = lists_succ (remove1 xs yss)" using Cons.prems \xs \ _\ by (intro Cons.hyps) (auto simp add: distincts_Cons distincts_remove1 distincts_distinct) then have "lists_succ (xs # xss) = lists_succ (xs # remove1 xs yss)" using Cons.prems \xs \ _\ by (simp add: lists_succ_Cons_pf distincts_Cons_remove1) then show ?case using Cons.prems \xs \ _\ by (simp add: lists_succ_remove1) qed section \List Orbits\ text \Computes the orbit of @{term x} under @{term f}\ definition orbit_list :: "('a \ 'a) \ 'a \ 'a list" where "orbit_list f x \ iterate 0 (funpow_dist1 f x x) f x" partial_function (tailrec) orbit_list_impl :: "('a \ 'a) \ 'a \ 'a list \ 'a \ 'a list" where "orbit_list_impl f s acc x = (let x' = f x in if x' = s then rev (x # acc) else orbit_list_impl f s (x # acc) x')" context notes [simp] = length_fold_remove1_le begin text \Computes the list of orbits\ fun orbits_list :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list f [] = []" | "orbits_list f (x # xs) = orbit_list f x # orbits_list f (fold remove1 (orbit_list f x) xs)" fun orbits_list_impl :: "('a \ 'a) \ 'a list \ 'a list list" where "orbits_list_impl f [] = []" | "orbits_list_impl f (x # xs) = (let fc = orbit_list_impl f x [] x in fc # orbits_list_impl f (fold remove1 fc xs))" declare orbit_list_impl.simps[code] end abbreviation sset :: "'a list list \ 'a set set" where "sset xss \ set ` set xss" lemma iterate_funpow_step: assumes "f x \ y" "y \ orbit f x" shows "iterate 0 (funpow_dist1 f x y) f x = x # iterate 0 (funpow_dist1 f (f x) y) f (f x)" proof - from assms have A: "y \ orbit f (f x)" by (simp add: orbit_step) have "iterate 0 (funpow_dist1 f x y) f x = x # iterate 1 (funpow_dist1 f x y) f x" (is "_ = _ # ?it") unfolding iterate_def by (simp add: upt_rec) also have "?it = map (\n. (f ^^ n) x) (map Suc [0.. = map (\n. (f ^^ n) (f x)) [0.. = iterate 0 (funpow_dist1 f (f x) y) f (f x)" unfolding iterate_def unfolding iterate_def by (simp add: funpow_dist_step[OF assms(1) A]) finally show ?thesis . qed lemma orbit_list_impl_conv: assumes "y \ orbit f x" shows "orbit_list_impl f y acc x = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using assms proof (induct n\"funpow_dist1 f x y" arbitrary: x acc) case (Suc x) show ?case proof cases assume "f x = y" then show ?thesis by (subst orbit_list_impl.simps) (simp add: Let_def iterate_def funpow_dist_0) next assume not_y :"f x \ y" have y_in_succ: "y \ orbit f (f x)" by (intro orbit_step Suc.prems not_y) have "orbit_list_impl f y acc x = orbit_list_impl f y (x # acc) (f x)" using not_y by (subst orbit_list_impl.simps) simp also have "\ = rev (x # acc) @ iterate 0 (funpow_dist1 f (f x) y) f (f x)" (is "_ = ?rev @ ?it") by (intro Suc funpow_dist_step not_y y_in_succ) also have "\ = rev acc @ iterate 0 (funpow_dist1 f x y) f x" using not_y Suc.prems by (simp add: iterate_funpow_step) finally show ?thesis . qed qed lemma orbit_list_conv_impl: assumes "x \ orbit f x" shows "orbit_list f x = orbit_list_impl f x [] x" unfolding orbit_list_impl_conv[OF assms] orbit_list_def by simp lemma set_orbit_list: assumes "x \ orbit f x" shows "set (orbit_list f x) = orbit f x" by (simp add: orbit_list_def orbit_conv_funpow_dist1[OF assms] set_iterate) lemma set_orbit_list': assumes "permutation f" shows "set (orbit_list f x) = orbit f x" using assms by (simp add: permutation_self_in_orbit set_orbit_list) lemma distinct_orbit_list: assumes "x \ orbit f x" shows "distinct (orbit_list f x)" by (simp del: upt_Suc add: orbit_list_def iterate_def distinct_map inj_on_funpow_dist1[OF assms]) lemma distinct_orbit_list': assumes "permutation f" shows "distinct (orbit_list f x)" using assms by (simp add: permutation_self_in_orbit distinct_orbit_list) lemma orbits_list_conv_impl: assumes "permutation f" shows "orbits_list f xs = orbits_list_impl f xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case using assms by (cases xs) (auto simp: assms less less_Suc_eq_le length_fold_remove1_le orbit_list_conv_impl permutation_self_in_orbit Let_def) qed lemma orbit_list_not_nil[simp]: "orbit_list f x \ []" by (simp add: orbit_list_def) lemma sset_orbits_list: assumes "permutation f" shows "sset (orbits_list f xs) = (orbit f) ` set xs" proof (induct "length xs" arbitrary: xs rule: less_induct) case less show ?case proof (cases xs) case Nil then show ?thesis by simp next case (Cons x' xs') let ?xs'' = "fold remove1 (orbit_list f x') xs'" have A: "sset (orbits_list f ?xs'') = orbit f ` set ?xs''" using Cons by (simp add: less_Suc_eq_le length_fold_remove1_le less.hyps) have B: "set (orbit_list f x') = orbit f x'" by (rule set_orbit_list) (simp add: permutation_self_in_orbit assms) have "orbit f ` set (fold remove1 (orbit_list f x') xs') \ orbit f ` set xs'" using set_fold_remove1[of _ xs'] by auto moreover have "orbit f ` set xs' - {orbit f x'} \ (orbit f ` set (fold remove1 (orbit_list f x') xs'))" (is "?L \ ?R") proof fix A assume "A \ ?L" then obtain y where "A = orbit f y" "y \ set xs'" by auto have "A \ orbit f x'" using \A \ ?L\ by auto from \A = _\ \A \ _\ have "y \ orbit f x'" by (meson assms cyclic_on_orbit orbit_cyclic_eq3 permutation_permutes) with \y \ _\ have "y \ set (fold remove1 (orbit_list f x') xs')" by (auto simp: set_fold_remove1' set_orbit_list permutation_self_in_orbit assms) then show "A \ ?R" using \A = _\ by auto qed ultimately show ?thesis by (auto simp: A B Cons) qed qed subsection \Relation to @{term cyclic_on}\ lemma list_succ_orbit_list: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "list_succ (orbit_list f s) = f" proof - have "distinct (orbit_list f s)" "\x. x \ set (orbit_list f s) \ x = f x" using assms by (simp_all add: distinct_orbit_list set_orbit_list) moreover have "\i. i < length (orbit_list f s) \ orbit_list f s ! (Suc i mod length (orbit_list f s)) = f (orbit_list f s ! i)" using funpow_dist1_prop[OF \s \ orbit f s\] by (auto simp: orbit_list_def funpow_mod_eq) ultimately show ?thesis by (auto simp: list_succ_def fun_eq_iff) qed lemma list_succ_funpow_conv: assumes A: "distinct xs" "x \ set xs" shows "(list_succ xs ^^ n) x = xs ! ((index xs x + n) mod length xs)" proof - have "xs \ []" using assms by auto then show ?thesis by (induct n) (auto simp: hd_conv_nth A index_nth_id list_succ_def mod_simps) qed lemma orbit_list_succ: assumes "distinct xs" "x \ set xs" shows "orbit (list_succ xs) x = set xs" proof (intro set_eqI iffI) fix y assume "y \ orbit (list_succ xs) x" then show "y \ set xs" by induct (auto simp: list_succ_in_conv \x \ set xs\) next fix y assume "y \ set xs" moreover { fix i j have "i < length xs \ j < length xs \ \n. xs ! j = xs ! ((i + n) mod length xs)" using assms by (auto simp: exI[where x="j + (length xs - i)"]) } ultimately show "y \ orbit (list_succ xs) x" using assms by (auto simp: orbit_altdef_permutation permutation_list_succ list_succ_funpow_conv index_nth_id in_set_conv_nth) qed lemma cyclic_on_list_succ: assumes "distinct xs" "xs \ []" shows "cyclic_on (list_succ xs) (set xs)" using assms last_in_set by (auto simp: cyclic_on_def orbit_list_succ) lemma obtain_orbit_list_func: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" obtains xs where "f = list_succ xs" "set xs = orbit f s" "distinct xs" "hd xs = s" proof - { from assms have "f = list_succ (orbit_list f s)" by (simp add: list_succ_orbit_list) moreover have "set (orbit_list f s) = orbit f s" "distinct (orbit_list f s)" by (auto simp: set_orbit_list distinct_orbit_list assms) moreover have "hd (orbit_list f s) = s" by (simp add: orbit_list_def iterate_def hd_map del: upt_Suc) ultimately have "\xs. f = list_succ xs \ set xs = orbit f s \ distinct xs \ hd xs = s" by blast } then show ?thesis by (metis that) qed lemma cyclic_on_obtain_list_succ: assumes "cyclic_on f S" "\x. x \ S \ f x = x" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" proof - from assms obtain s where s: "s \ orbit f s" "\x. x \ orbit f s \ f x = x" "S = orbit f s" by (auto simp: cyclic_on_def) then show ?thesis by (metis that obtain_orbit_list_func) qed lemma cyclic_on_obtain_list_succ': assumes "cyclic_on f S" "f permutes S" obtains xs where "f = list_succ xs" "set xs = S" "distinct xs" using assms unfolding permutes_def by (metis cyclic_on_obtain_list_succ) lemma list_succ_unique: assumes "s \ orbit f s" "\x. x \ orbit f s \ f x = x" shows "\!xs. f = list_succ xs \ distinct xs \ hd xs = s \ set xs = orbit f s" proof - from assms obtain xs where xs: "f = list_succ xs" "distinct xs" "hd xs = s" "set xs = orbit f s" by (rule obtain_orbit_list_func) moreover { fix zs assume A: "f = list_succ zs" "distinct zs" "hd zs = s" "set zs = orbit f s" then have "zs \ []" using \s \ orbit f s\ by auto from \distinct xs\ \distinct zs\ \set xs = orbit f s\ \set zs = orbit f s\ have len: "length xs = length zs" by (metis distinct_card) { fix n assume "n < length xs" then have "zs ! n = xs ! n" proof (induct n) case 0 with A xs \zs \ []\ show ?case by (simp add: hd_conv_nth nth_rotate_conv_nth) next case (Suc n) then have "list_succ zs (zs ! n) = list_succ xs (xs! n)" using \f = list_succ xs\ \f = list_succ zs\ by simp with \Suc n < _\ show ?case by (simp add:list_succ_nth len \distinct xs\ \distinct zs\) qed } then have "zs = xs" by (metis len nth_equalityI) } ultimately show ?thesis by metis qed lemma distincts_orbits_list: assumes "distinct as" "permutation f" shows "distincts (orbits_list f as)" using assms(1) proof (induct "length as" arbitrary: as rule: less_induct) case less show ?case proof (cases as) case Nil then show ?thesis by simp next case (Cons a as') let ?as' = "fold remove1 (orbit_list f a) as'" from Cons less.prems have A: "distincts (orbits_list f (fold remove1 (orbit_list f a) as'))" by (intro less) (auto simp: distinct_fold_remove1 length_fold_remove1_le less_Suc_eq_le) have B: "set (orbit_list f a) \ \(sset (orbits_list f (fold remove1 (orbit_list f a) as'))) = {}" proof - have "orbit f a \ set (fold remove1 (orbit_list f a) as') = {}" using assms less.prems Cons by (simp add: set_fold_remove1_distinct set_orbit_list') then have "orbit f a \ \ (orbit f ` set (fold remove1 (orbit_list f a) as')) = {}" by auto (metis assms(2) cyclic_on_orbit disjoint_iff_not_equal permutation_self_in_orbit[OF assms(2)] orbit_cyclic_eq3 permutation_permutes) then show ?thesis using assms by (auto simp: set_orbit_list' sset_orbits_list disjoint_iff_not_equal) qed show ?thesis using A B assms by (auto simp: distincts_Cons Cons distinct_orbit_list') qed qed lemma cyclic_on_lists_succ': assumes "distincts xss" shows "A \ sset xss \ cyclic_on (lists_succ xss) A" using assms proof (induction xss arbitrary: A) case Nil then show ?case by auto next case (Cons xs xss A) then have inter: "set xs \ (\ys\set xss. set ys) = {}" by (auto simp: distincts_Cons) note pcp[OF _ _ inter] = permutes_comp_preserves_cyclic1 permutes_comp_preserves_cyclic2 from Cons show "cyclic_on (lists_succ (xs # xss)) A" by (cases "A = set xs") (auto intro: pcp simp: cyclic_on_list_succ list_succ_permutes lists_succ_permutes lists_succ_Cons_pf distincts_Cons) qed lemma cyclic_on_lists_succ: assumes "distincts xss" shows "\xs. xs \ set xss \ cyclic_on (lists_succ xss) (set xs)" using assms by (auto intro: cyclic_on_lists_succ') lemma permutes_as_lists_succ: assumes "distincts xss" assumes ls_eq: "\xs. xs \ set xss \ list_succ xs = perm_restrict f (set xs)" assumes "f permutes (\(sset xss))" shows "f = lists_succ xss" using assms proof (induct xss arbitrary: f) case Nil then show ?case by simp next case (Cons xs xss) let ?sets = "\xss. \ys \ set xss. set ys" have xs: "distinct xs" "xs \ []" using Cons by (auto simp: distincts_Cons) have f_xs: "perm_restrict f (set xs) = list_succ xs" using Cons by simp have co_xs: "cyclic_on (perm_restrict f (set xs)) (set xs)" unfolding f_xs using xs by (rule cyclic_on_list_succ) have perm_xs: "perm_restrict f (set xs) permutes set xs" unfolding f_xs using \distinct xs\ by (rule list_succ_permutes) have perm_xss: "perm_restrict f (?sets xss) permutes (?sets xss)" proof - have "perm_restrict f (?sets (xs # xss) - set xs) permutes (?sets (xs # xss) - set xs)" using Cons co_xs by (intro perm_restrict_diff_cyclic) (auto simp: cyclic_on_perm_restrict) also have "?sets (xs # xss) - set xs = ?sets xss" using Cons by (auto simp: distincts_Cons) finally show ?thesis . qed have f_xss: "perm_restrict f (?sets xss) = lists_succ xss" proof - have *: "\xs. xs \ set xss \ ((\x\set xss. set x) \ set xs) = set xs" by blast with perm_xss Cons.prems show ?thesis by (intro Cons.hyps) (auto simp: distincts_Cons perm_restrict_perm_restrict *) qed from Cons.prems show "f = lists_succ (xs # xss)" by (simp add: lists_succ_Cons_pf distincts_Cons f_xss[symmetric] perm_restrict_union perm_xs perm_xss) qed lemma cyclic_on_obtain_lists_succ: assumes permutes: "f permutes S" and S: "S = \(sset css)" and dists: "distincts css" and cyclic: "\cs. cs \ set css \ cyclic_on f (set cs)" obtains xss where "f = lists_succ xss" "distincts xss" "map set xss = map set css" "map hd xss = map hd css" proof - let ?fc = "\cs. perm_restrict f (set cs)" define some_list where "some_list cs = (SOME xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs)" for cs { fix cs assume "cs \ set css" then have "cyclic_on (?fc cs) (set cs)" "\x. x \ set cs \ ?fc cs x = x" "hd cs \ set cs" using cyclic dists by (auto simp add: cyclic_on_perm_restrict perm_restrict_def distincts_def) then have "hd cs \ orbit (?fc cs) (hd cs)" "\x. x \ orbit (?fc cs) (hd cs) \ ?fc cs x = x" "hd cs \ set cs" "set cs = orbit (?fc cs) (hd cs)" by (auto simp: cyclic_on_alldef) then have "\xs. ?fc cs = list_succ xs \ set xs = set cs \ distinct xs \ hd xs = hd cs" by (metis obtain_orbit_list_func) then have "?fc cs = list_succ (some_list cs) \ set (some_list cs) = set cs \ distinct (some_list cs) \ hd (some_list cs) = hd cs" unfolding some_list_def by (rule someI_ex) then have "?fc cs = list_succ (some_list cs)" "set (some_list cs) = set cs" "distinct (some_list cs)" "hd (some_list cs) = hd cs" by auto } note sl_cs = this have "\cs. cs \ set css \ cs \ []" using dists by (auto simp: distincts_def) then have some_list_ne: "\cs. cs \ set css \ some_list cs \ []" by (metis set_empty sl_cs(2)) have set: "map set (map some_list css) = map set css" "map hd (map some_list css) = map hd css" using sl_cs(2,4) by (auto simp add: map_idI) have distincts: "distincts (map some_list css)" proof - have c_dist: "\xs ys. \xs\set css; ys\set css; xs \ ys\ \ set xs \ set ys = {}" using dists by (auto simp: distincts_def) have "distinct (map some_list css)" proof - have "inj_on some_list (set css)" using sl_cs(2) c_dist by (intro inj_onI) (metis inf.idem set_empty) with \distincts css\ show ?thesis by (auto simp: distincts_distinct distinct_map) qed moreover have "\xs\set (map some_list css). distinct xs \ xs \ []" using sl_cs(3) some_list_ne by auto moreover from c_dist have "(\xs\set (map some_list css). \ys\set (map some_list css). xs \ ys \ set xs \ set ys = {})" using sl_cs(2) by auto ultimately show ?thesis by (simp add: distincts_def) qed have f: "f = lists_succ (map some_list css)" using distincts proof (rule permutes_as_lists_succ) fix xs assume "xs \ set (map some_list css)" then show "list_succ xs = perm_restrict f (set xs)" using sl_cs(1) sl_cs(2) by auto next have "S = (\xs\set (map some_list css). set xs)" using S sl_cs(2) by auto with permutes show "f permutes \(sset (map some_list css))" by simp qed from f distincts set show ?thesis .. qed subsection \Permutations of a List\ lemma length_remove1_less: assumes "x \ set xs" shows "length (remove1 x xs) < length xs" proof - from assms have "0 < length xs" by auto with assms show ?thesis by (auto simp: length_remove1) qed context notes [simp] = length_remove1_less begin fun permutations :: "'a list \ 'a list list" where permutations_Nil: "permutations [] = [[]]" | permutations_Cons: "permutations xs = [y # ys. y <- xs, ys <- permutations (remove1 y xs)]" end declare permutations_Cons[simp del] text \ The function above returns all permutations of a list. The function below computes only those which yield distinct cyclic permutation functions (cf. @{term list_succ}). \ fun cyc_permutations :: "'a list \ 'a list list" where "cyc_permutations [] = [[]]" | "cyc_permutations (x # xs) = map (Cons x) (permutations xs)" lemma nil_in_permutations[simp]: "[] \ set (permutations xs) \ xs = []" by (induct xs) (auto simp: permutations_Cons) lemma permutations_not_nil: assumes "xs \ []" shows "permutations xs = concat (map (\x. map ((#) x) (permutations (remove1 x xs))) xs)" using assms by (cases xs) (auto simp: permutations_Cons) lemma set_permutations_step: assumes "xs \ []" shows "set (permutations xs) = (\x \ set xs. Cons x ` set (permutations (remove1 x xs)))" using assms by (cases xs) (auto simp: permutations_Cons) lemma in_set_permutations: assumes "distinct xs" shows "ys \ set (permutations xs) \ distinct ys \ set xs = set ys" (is "?L xs ys \ ?R xs ys") using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by auto next case (Suc n) then have "xs \ []" by auto show ?case proof assume "?L xs ys" then obtain y ys' where "ys = y # ys'" "y \ set xs" "ys' \ set (permutations (remove1 (hd ys) xs))" using \xs \ []\ by (auto simp: permutations_not_nil) moreover then have "?R (remove1 y xs) ys'" using Suc.prems Suc.hyps(2) by (intro Suc.hyps(1)[THEN iffD1]) (auto simp: length_remove1) ultimately show "?R xs ys" using Suc by auto next assume "?R xs ys" with \xs \ []\ obtain y ys' where "ys = y # ys'" "y \ set xs" by (cases ys) auto moreover then have "ys' \ set (permutations (remove1 y xs))" using Suc \?R xs ys\ by (intro Suc.hyps(1)[THEN iffD2]) (auto simp: length_remove1) ultimately show "?L xs ys" using \xs \ []\ by (auto simp: permutations_not_nil) qed qed lemma in_set_cyc_permutations: assumes "distinct xs" shows "ys \ set (cyc_permutations xs) \ distinct ys \ set xs = set ys \ hd ys = hd xs" (is "?L xs ys \ ?R xs ys") proof (cases xs) case (Cons x xs) with assms show ?thesis by (cases ys) (auto simp: in_set_permutations intro!: imageI) qed auto lemma in_set_cyc_permutations_obtain: assumes "distinct xs" "distinct ys" "set xs = set ys" obtains n where "rotate n ys \ set (cyc_permutations xs)" proof (cases xs) case Nil with assms have "rotate 0 ys \ set (cyc_permutations xs)" by auto then show ?thesis .. next case (Cons x xs') let ?ys' = "rotate (index ys x) ys" have "ys \ []" "x \ set ys" using Cons assms by auto then have "distinct ?ys' \ set xs = set ?ys' \ hd ?ys' = hd xs" using assms Cons by (auto simp add: hd_rotate_conv_nth) with \distinct xs\ have "?ys' \ set (cyc_permutations xs)" by (rule in_set_cyc_permutations[THEN iffD2]) then show ?thesis .. qed lemma list_succ_set_cyc_permutations: assumes "distinct xs" "xs \ []" shows "list_succ ` set (cyc_permutations xs) = {f. f permutes set xs \ cyclic_on f (set xs)}" (is "?L = ?R") proof (intro set_eqI iffI) fix f assume "f \ ?L" moreover have "\ys. set xs = set ys \ xs \ [] \ ys \ []" by auto ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutations list_succ_permutes cyclic_on_list_succ) next fix f assume "f \ ?R" then obtain ys where ys: "list_succ ys = f" "distinct ys" "set ys = set xs" by (auto elim: cyclic_on_obtain_list_succ') moreover with \distinct xs\ obtain n where "rotate n ys \ set (cyc_permutations xs)" by (auto elim: in_set_cyc_permutations_obtain) then have "list_succ (rotate n ys) \ ?L" by simp ultimately show "f \ ?L" by simp qed subsection \Enumerating Permutations from List Orbits\ definition cyc_permutationss :: "'a list list \ 'a list list list" where "cyc_permutationss = product_lists o map cyc_permutations" lemma cyc_permutationss_Nil[simp]: "cyc_permutationss [] = [[]]" by (auto simp: cyc_permutationss_def) lemma in_set_cyc_permutationss: assumes "distincts xss" shows "yss \ set (cyc_permutationss xss) \ distincts yss \ map set xss = map set yss \ map hd xss = map hd yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" "map hd xss = map hd yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map cyc_permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_cyc_permutations) } note Y = this show "?thesis" unfolding cyc_permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma lists_succ_set_cyc_permutationss: assumes "distincts xss" shows "lists_succ ` set (cyc_permutationss xss) = {f. f permutes \(sset xss) \ (\c \ sset xss. cyclic_on f c)}" (is "?L = ?R") using assms proof (intro set_eqI iffI) fix f assume "f \ ?L" then obtain yss where "yss \ set (cyc_permutationss xss)" "f = lists_succ yss" by (rule imageE) moreover from \yss \ _\ assms have "set (map set xss) = set (map set yss)" by (auto simp: in_set_cyc_permutationss) then have "sset xss = sset yss" by simp ultimately show "f \ ?R" using assms by (auto simp: in_set_cyc_permutationss cyclic_on_lists_succ') (metis lists_succ_permutes) next fix f assume "f \ ?R" then have "f permutes \(sset xss)" "\cs. cs \ set xss \ cyclic_on f (set cs)" by auto from this(1) refl assms this(2) obtain yss where "f = lists_succ yss" "distincts yss" "map set yss = map set xss" "map hd yss = map hd xss" by (rule cyclic_on_obtain_lists_succ) with assms show "f \ ?L" by (auto intro!: imageI simp: in_set_cyc_permutationss) qed subsection \Lists of Permutations\ definition permutationss :: "'a list list \ 'a list list list" where "permutationss = product_lists o map permutations" lemma permutationss_Nil[simp]: "permutationss [] = [[]]" by (auto simp: permutationss_def) lemma permutationss_Cons: "permutationss (xs # xss) = concat (map (\ys. map (Cons ys) (permutationss xss)) (permutations xs))" by (auto simp: permutationss_def) lemma in_set_permutationss: assumes "distincts xss" shows "yss \ set (permutationss xss) \ distincts yss \ map set xss = map set yss" proof - { assume A: "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" then have "length yss = length xss" by (auto simp: list_all2_lengthD) then have "\(sset xss) = \(sset yss)" "distincts yss" "map set xss = map set yss" using A assms by (induct yss xss rule: list_induct2) (auto simp: distincts_Cons in_set_permutations) } note X = this { assume A: "distincts yss" "map set xss = map set yss" then have "length yss = length xss" by (auto dest: map_eq_imp_length_eq) then have "list_all2 (\x ys. x \ set ys) yss (map permutations xss)" using A assms by (induct yss xss rule: list_induct2) (auto simp: in_set_permutations distincts_Cons) } note Y = this show "?thesis" unfolding permutationss_def by (auto simp: product_lists_set intro: X Y) qed lemma set_permutationss: assumes "distincts xss" shows "set (permutationss xss) = {yss. distincts yss \ map set xss = map set yss}" using in_set_permutationss[OF assms] by blast lemma permutationss_complete: assumes "distincts xss" "distincts yss" "xss \ []" and "set ` set xss = set ` set yss" shows "set yss \ set ` set (permutationss xss)" proof - have "length xss = length yss" using assms by (simp add: length_distincts) from \sset xss = _\ have "\yss'. set yss' = set yss \ map set yss' = map set xss" using assms(1-2) proof (induct xss arbitrary: yss) case Nil then show ?case by simp next case (Cons xs xss) from \sset (xs # xss) = sset yss\ obtain ys where ys: "ys \ set yss" "set ys = set xs" by auto (metis imageE insertI1) with \distincts yss\ have "set ys \ sset (remove1 ys yss)" by (fastforce simp: distincts_def) moreover from \distincts (xs # xss)\ have "set xs \ sset xss" by (fastforce simp: distincts_def) ultimately have "sset xss = sset (remove1 ys yss)" using \distincts yss\ \sset (xs # xss) = sset yss\ apply (auto simp: distincts_distinct \set ys = set xs\[symmetric]) apply (smt Diff_insert_absorb \ys \ set yss\ image_insert insert_Diff rev_image_eqI) by (metis \ys \ set yss\ image_eqI insert_Diff insert_iff) then obtain yss' where "set yss' = set (remove1 ys yss) \ map set yss' = map set xss" using Cons by atomize_elim (auto simp: distincts_Cons distincts_remove1) then have "set (ys # yss') = set yss \ map set (ys # yss') = map set (xs # xss)" using ys set_remove1_eq \distincts yss\ by (auto simp: distincts_distinct) then show ?case .. qed then obtain yss' where "set yss' = set yss" "map set yss' = map set xss" by blast then have "distincts yss'" using \distincts xss\ \distincts yss\ unfolding distincts_def by simp_all (metis \length xss = length yss\ card_distinct distinct_card length_map) then have "set yss' \ set ` set (permutationss xss)" using \distincts xss\ \map set yss' = _\ by (auto simp: set_permutationss) then show ?thesis using \set yss' = _\ by auto qed lemma permutations_complete: (* could generalize with multi-sets *) assumes "distinct xs" "distinct ys" "set xs = set ys" shows "ys \ set (permutations xs)" using assms proof (induct "length xs" arbitrary: xs ys) case 0 then show ?case by simp next case (Suc n) from Suc.hyps have "xs \ []" by auto then obtain y ys' where [simp]: "ys = y # ys'" "y \ set xs" using Suc.prems by (cases ys) auto have "ys' \ set (permutations (remove1 y xs))" using Suc.prems \Suc n = _\ by (intro Suc.hyps) (simp_all add: length_remove1 ) then show ?case using \xs \ []\ by (auto simp: set_permutations_step) qed end diff --git a/thys/Three_Squares/Three_Squares.thy b/thys/Three_Squares/Three_Squares.thy --- a/thys/Three_Squares/Three_Squares.thy +++ b/thys/Three_Squares/Three_Squares.thy @@ -1,1436 +1,1436 @@ (* Title: Three_Squares/Three_Squares.thy Author: Anton Danilkin Author: Loïc Chevalier *) section \Legendre's three squares theorem and its consequences\ theory Three_Squares imports "Dirichlet_L.Dirichlet_Theorem" Residues_Properties Quadratic_Forms begin subsection \Legendre's three squares theorem\ definition quadratic_residue_alt :: "int \ int \ bool" where "quadratic_residue_alt m a = (\x y. x\<^sup>2 - a = y * m)" lemma quadratic_residue_alt_equiv: "quadratic_residue_alt = QuadRes" unfolding quadratic_residue_alt_def QuadRes_def by (metis cong_iff_dvd_diff cong_modulus_mult dvdE dvd_refl mult.commute) lemma sq_nat_abs: "(nat \v\)\<^sup>2 = nat (v\<^sup>2)" by (simp add: nat_power_eq[symmetric]) text \Lemma 1.7 from @{cite nathanson1996}.\ lemma three_squares_using_quadratic_residue: fixes n d' :: nat assumes "n \ 2" assumes "d' > 0" assumes "QuadRes (d' * n - 1) (- d')" shows "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" proof - define a where "a \ d' * n - 1" from assms(3) obtain x y where "x\<^sup>2 + int d' = y * int a" unfolding a_def quadratic_residue_alt_equiv[symmetric] quadratic_residue_alt_def by auto hence Hxy: "x\<^sup>2 + d' = y * a" by auto have "y \ 1" using assms Hxy unfolding a_def by (smt (verit) bot_nat_0.not_eq_extremum mult_le_0_iff of_nat_0_le_iff of_nat_le_0_iff power2_eq_square zero_le_square) moreover from Hxy have Hxy\<^sub>2: "d' = y * a - x\<^sup>2" by (simp add: algebra_simps) define M where "M \ mat3 y x 1 x a 0 1 0 n" moreover have Msym: "mat_sym M" unfolding mat3_sym_def M_def mat3.sel by simp moreover have Mdet_eq_1: "mat_det M = 1" proof - have "mat_det M = (y * a - x\<^sup>2) * n - a" unfolding mat3_det_def M_def mat3.sel power2_eq_square by (simp add: algebra_simps) also have "... = int d' * n - a" unfolding Hxy\<^sub>2 by simp also have "... = 1" unfolding a_def using assms int_ops by force finally show ?thesis . qed moreover have "mat_det (mat2 y x x a) > 0" using Hxy\<^sub>2 assms unfolding mat2_det_def power2_eq_square by simp ultimately have "qf_positive_definite M" by (auto simp add: lemma_1_3(4)) hence "M ~ 1" using Msym and Mdet_eq_1 by (simp add: qf3_det_one_equiv_canonical) moreover have "M $$ vec3 0 0 1 = n" unfolding M_def qf3_app_def mat3_app_def vec3_dot_def mat3.sel vec3.sel by (simp add: algebra_simps) hence "n \ range (($$) M)" by (metis rangeI) ultimately have "n \ range (($$) (1 :: mat3))" using qf3_equiv_preserves_range by simp then obtain u :: vec3 where "1 $$ u = n" by auto hence " = n" unfolding qf3_app_def mat3_app_def one_mat3_def by simp hence "\x\<^sub>1 x\<^sub>2 x\<^sub>3. int n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" unfolding vec3_dot_def power2_eq_square by metis hence "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = (nat \x\<^sub>1\)\<^sup>2 + (nat \x\<^sub>2\)\<^sup>2 + (nat \x\<^sub>3\)\<^sup>2" unfolding sq_nat_abs apply (simp add: nat_add_distrib[symmetric]) apply (metis nat_int) done thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by blast qed lemma prime_linear_combination: fixes a m :: nat assumes "m > 1" assumes "coprime a m" obtains j :: nat where "prime (a + m * j) \ j \ 0" proof - assume 0: "\j. prime (a + m * j) \ j \ 0 \ thesis" have 1: "infinite {p. prime p \ [p = a] (mod m)}" using assms by (rule Dirichlet_Theorem.residues_nat.Dirichlet[unfolded residues_nat_def]) have 2: "finite {j. prime (nat (int a + int m * j)) \ j \ 0}" apply (rule finite_subset[of _ "{- (int a) div (int m)..0}"]) subgoal apply (rule subsetI) subgoal for j proof - assume 1: "j \ {j. prime (nat (int a + int m * j)) \ j \ 0}" have "int a + int m * j \ 0" using 1 prime_ge_0_int by force hence "int m * j \ - int a" by auto hence "j \ (- int a) div int m" using assms 1 by (smt (verit) unique_euclidean_semiring_class.cong_def coprime_crossproduct_nat coprime_iff_invertible_int coprime_int_iff int_distrib(3) int_ops(2) int_ops(7) mem_Collect_eq mult_cancel_right1 zdiv_mono1 nonzero_mult_div_cancel_left of_nat_0_eq_iff of_nat_le_0_iff prime_ge_2_int prime_nat_iff_prime) thus "j \ {- (int a) div (int m)..0}" using 1 by auto qed done subgoal by blast done have "{p. prime p \ [p = a] (mod m)} = {p. prime p \ (\j. int p = int a + int m * j)}" unfolding cong_sym_eq[of _ a] unfolding cong_int_iff[symmetric] cong_iff_lin .. also have "... = {p. prime p \ (\j. p = nat (int a + int m * j))}" by (metis (opaque_lifting) nat_int nat_eq_iff prime_ge_0_int prime_nat_iff_prime) also have "... = (\j. nat (int a + int m * j)) ` {j. prime (nat (int a + int m * j))}" by blast finally have "infinite ((\j. nat (int a + int m * j)) ` {j. prime (nat (int a + int m * j))})" using 1 by metis hence "infinite {j. prime (nat (int a + int m * j))}" using finite_imageI by blast hence "infinite ({j. prime (nat (int a + int m * j))} - {j. prime (nat (int a + int m * j)) \ j \ 0})" using 2 Diff_infinite_finite by blast hence "infinite {j. prime (nat (int a + int m * j)) \ j > 0}" by (rule back_subst[of infinite]; auto) hence "infinite (int ` {j. prime (nat (int a + int m * j)) \ j \ (0 :: nat)})" apply (rule back_subst[of infinite]) unfolding image_def using zero_less_imp_eq_int apply auto done hence "infinite {j. prime (nat (int a + int m * j)) \ (j :: nat) \ 0}" using finite_imageI by blast hence "infinite {j. prime (a + m * j) \ j \ 0}" apply (rule back_subst[of infinite]) apply (auto simp add: int_ops nat_plus_as_int) done thus thesis using 0 not_finite_existsD by blast qed text \Lemma 1.8 from @{cite nathanson1996}.\ lemma three_squares_using_mod_four: fixes n :: nat assumes "n mod 4 = 2" shows "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" proof - have "n > 1" using assms by auto have "coprime (n - 1) (4 * n)" by (smt (verit, del_insts) Suc_pred assms bot_nat_0.not_eq_extremum coprime_commute coprime_diff_one_right_nat coprime_mod_right_iff coprime_mult_left_iff diff_Suc_1 mod_Suc mod_mult_self1_is_0 mult_0_right numeral_2_eq_2 zero_neq_numeral) then obtain j where H_j: "prime ((n - 1) + (4 * n) * j) \ j \ 0" using prime_linear_combination[of "4 * n" "n - 1"] \n > 1\ by auto have "j > 0" using H_j by blast define d' where "d' \ 4 * j + 1" define p where "p \ d' * n - 1" have "prime p" unfolding p_def d'_def using conjunct1[OF H_j] apply (rule back_subst[of prime]) using \n > 1\ apply (simp add: algebra_simps nat_minus_as_int nat_plus_as_int) done have "p mod 4 = 1" unfolding p_def apply (subst mod_diff_eq_nat) subgoal unfolding d'_def using \n > 1\ \j > 0\ by simp subgoal apply (subst mod_mult_eq[symmetric]) unfolding assms d'_def apply simp done done have "d' * n mod 4 = 2" using assms p_def d'_def \p mod 4 = 1\ by (metis mod_mult_cong mod_mult_self4 nat_mult_1) hence "d' mod 4 = 1" using assms by (simp add: d'_def) have "QuadRes p (- d')" proof - have d'_expansion: "d' = (\q\prime_factors d'. q ^ multiplicity q d')" using prime_factorization_nat unfolding d'_def by auto have "odd d'" unfolding d'_def by simp hence d'_prime_factors_odd: "q \ prime_factors d' \ odd q" for q by fastforce have d'_prime_factors_gt_2: "q \ prime_factors d' \ q > 2" for q using d'_prime_factors_odd by (metis even_numeral in_prime_factors_imp_prime order_less_le prime_ge_2_nat) have "[p = - 1] (mod d')" unfolding p_def cong_iff_dvd_diff apply simp using \n > 1\ apply (smt (verit) Suc_as_int Suc_pred add_gr_0 d'_def dvd_nat_abs_iff dvd_triv_left less_numeral_extra(1) mult_pos_pos of_nat_less_0_iff order_le_less_trans zero_less_one_class.zero_le_one) done hence d'_prime_factors_2_p_mod: "q \ prime_factors d' \ [p = - 1] (mod q)" for q by (rule cong_dvd_modulus; auto) have "d' mod 4 = (\q\prime_factors d'. q ^ multiplicity q d') mod 4" using d'_expansion by argo also have "... = (\q\prime_factors d'. (q mod 4) ^ multiplicity q d') mod 4" apply (subst mod_prod_eq[symmetric]) apply (subst power_mod[symmetric]) apply (subst mod_prod_eq) apply blast done also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (q mod 4) ^ multiplicity q d') mod 4" apply (rule arg_cong[of _ _ "\x. x mod 4"]) apply (rule prod.mono_neutral_right) subgoal by blast subgoal by blast subgoal unfolding unique_euclidean_semiring_class.cong_def apply (rule ballI) using d'_prime_factors_odd apply simp apply (metis One_nat_def dvd_0_right even_even_mod_4_iff even_numeral mod_exhaust_less_4) done done also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. ((int q) mod 4) ^ multiplicity q d') mod 4" by (simp add: int_ops) also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. ((- 1) mod 4) ^ multiplicity q d') mod 4" apply (rule arg_cong[of _ _ "\x. x mod 4"]) apply (rule prod.cong[OF refl]) unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int apply simp apply (metis nat_int of_nat_mod of_nat_numeral) done also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (- 1) ^ multiplicity q d') mod 4" apply (subst mod_prod_eq[symmetric]) apply (subst power_mod) apply (subst mod_prod_eq) apply blast done finally have "[\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (- 1) ^ multiplicity q d' = 1 :: int] (mod 4)" using \d' mod 4 = 1\ by (simp add: unique_euclidean_semiring_class.cong_def) hence prod_prime_factors_minus_one: "(\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (- 1) ^ multiplicity q d') = (1 :: int)" unfolding power_sum[symmetric] unfolding minus_one_power_iff unique_euclidean_semiring_class.cong_def by presburger have "p > 2" using \prime p\ \p mod 4 = 1\ nat_less_le prime_ge_2_nat by force have d'_prime_factors_Legendre: "q \ prime_factors d' \ Legendre p q = Legendre q p" for q proof - assume "q \ prime_factors d'" have "prime q" using \q \ prime_factors d'\ by blast have "q > 2" using d'_prime_factors_gt_2 \q \ prime_factors d'\ by blast show "Legendre p q = Legendre q p" using \prime p\ \p > 2\ \p mod 4 = 1\ \prime q\ \q > 2\ Legendre_equal[of p q] unfolding unique_euclidean_semiring_class.cong_def using zmod_int[of p 4] by auto qed have "Legendre (- d') p = Legendre (- 1) p * Legendre d' p" using \prime p\ Legendre_mult[of p "- 1" d'] by auto also have "... = Legendre d' p" using \prime p\ \p > 2\ \p mod 4 = 1\ Legendre_minus_one[of p] unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int by auto also have "... = (\q\prime_factors d'. Legendre (q ^ multiplicity q d') p)" apply (subst d'_expansion) using \prime p\ \p > 2\ Legendre_prod[of p] apply auto done also have "... = (\q\prime_factors d'. (Legendre q p) ^ multiplicity q d')" using \prime p\ \p > 2\ Legendre_power by auto also have "... = (\q\prime_factors d'. (Legendre p q) ^ multiplicity q d')" using d'_prime_factors_Legendre by auto also have "... = (\q\prime_factors d'. (Legendre (- 1) q) ^ multiplicity q d')" apply (rule prod.cong[OF refl]) using d'_prime_factors_2_p_mod apply (metis Legendre_cong) done also have "... = (\q\prime_factors d'. (if [q = 1] (mod 4) then 1 else - 1) ^ multiplicity q d')" apply (rule prod.cong[OF refl]) subgoal for q using Legendre_minus_one_alt[of q] d'_prime_factors_gt_2[of q] by (smt (verit) cong_int_iff in_prime_factors_iff int_eq_iff_numeral less_imp_of_nat_less numeral_Bit0 numeral_One of_nat_1 prime_nat_int_transfer) done also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (if [q = 1] (mod 4) then 1 else - 1) ^ multiplicity q d')" apply (rule prod.mono_neutral_right) subgoal by blast subgoal by blast subgoal unfolding unique_euclidean_semiring_class.cong_def apply (rule ballI) using d'_prime_factors_odd apply simp apply (metis One_nat_def dvd_0_right even_even_mod_4_iff even_numeral mod_exhaust_less_4) done done also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 4)}. (- 1) ^ multiplicity q d')" by (rule prod.cong[OF refl]; simp add: unique_euclidean_semiring_class.cong_def) also have "... = 1" using prod_prime_factors_minus_one . finally show "QuadRes p (- d')" unfolding Legendre_def by (metis one_neq_neg_one one_neq_zero) qed thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using \n > 1\ three_squares_using_quadratic_residue[of n d'] unfolding d'_def p_def by auto qed lemma three_mod_eight_power_iff: fixes n :: nat shows "(3 :: int) ^ n mod 8 = (if even n then 1 else 3)" proof (induction n) case 0 thus ?case by auto next case (Suc n) thus ?case apply (cases "even n") subgoal using mod_mult_left_eq[of 3 8 "3 ^ n"] apply simp apply presburger done subgoal using mod_mult_left_eq[of 3 8 "3 * 3 ^ n"] apply simp apply presburger done done qed text \Lemma 1.9 from @{cite nathanson1996}.\ lemma three_squares_using_mod_eight: fixes n :: nat assumes "n mod 8 \ {1, 3, 5}" shows "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" proof cases assume "n = 1" hence "n = 1\<^sup>2 + 0\<^sup>2 + 0\<^sup>2" unfolding power2_eq_square by auto thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by blast next assume "n \ 1" hence "n > 1" using assms by auto have H_n: "(n mod 8 = 1 \ P) \ (n mod 8 = 3 \ P) \ (n mod 8 = 5 \ P) \ P" for P using assms by auto define c :: nat where "c \ if n mod 8 = 3 then 1 else 3" have "c * n \ 1" unfolding c_def using \n > 1\ by auto obtain k where H_k: "2 * k = c * n - 1" using H_n unfolding c_def by (smt (verit, ccfv_threshold) dvd_mod even_mult_iff even_numeral odd_numeral odd_one odd_two_times_div_two_nat) have k_mod_4: "k mod 4 = (if n mod 8 = 5 then 3 else 1)" (is "k mod 4 = ?v") proof - have "c * n mod 8 = (if n mod 8 = 5 then 7 else 3)" using H_n proof cases assume "n mod 8 = 1" have "3 * n mod 8 = 3" using \n mod 8 = 1\ mod_mult_right_eq[of 3 n 8] by auto thus ?thesis unfolding c_def using \n mod 8 = 1\ by auto next assume "n mod 8 = 3" thus ?thesis unfolding c_def by auto next assume "n mod 8 = 5" have "3 * n mod 8 = 7" using \n mod 8 = 5\ mod_mult_right_eq[of 3 n 8] by auto thus ?thesis unfolding c_def using \n mod 8 = 5\ by auto qed hence "2 * k mod 8 = (if n mod 8 = 5 then 6 else 2)" unfolding H_k using \c * n \ 1\ mod_diff_eq_nat by simp hence "2 * (k mod 4) = 2 * ?v" by (simp add: mult_mod_right) thus ?thesis by simp qed have "coprime k (4 * n)" using k_mod_4 H_k \c * n \ 1\ by (metis One_nat_def coprime_Suc_left_nat coprime_commute coprime_diff_one_right_nat coprime_mod_left_iff coprime_mult_right_iff mult_2_right numeral_2_eq_2 numeral_3_eq_3 numeral_Bit0 order_less_le_trans zero_less_one zero_neq_numeral) then obtain j where H_j: "prime (k + (4 * n) * j) \ j \ 0" using prime_linear_combination[of k "n - 1"] \n > 1\ by (metis One_nat_def Suc_leI bot_nat_0.not_eq_extremum mult_is_0 nat_1_eq_mult_iff nat_less_le prime_linear_combination zero_neq_numeral) have "j > 0" using H_j by blast define p where "p \ k + (4 * n) * j" have "prime p" unfolding p_def using conjunct1[OF H_j] apply (rule back_subst[of prime]) apply (simp add: int_ops nat_plus_as_int) done have "[p = k] (mod 4 * n)" unfolding p_def unique_euclidean_semiring_class.cong_def by auto have "p > 2" using \prime p\ \[p = k] (mod 4 * n)\ \coprime k (4 * n)\ by (metis cong_dvd_iff cong_dvd_modulus_nat coprime_common_divisor_nat dvd_mult2 even_numeral le_neq_implies_less odd_one prime_ge_2_nat) define d' where "d' \ 8 * j + c" have "d' > 1" unfolding d'_def using \j > 0\ by simp have H_2_p: "2 * p = d' * n - 1" unfolding p_def d'_def using \c * n \ 1\ H_k by (smt (verit, del_insts) Nat.add_diff_assoc add.commute add_mult_distrib mult.commute mult_2 numeral_Bit0) have "QuadRes p (- d')" proof - have d'_expansion: "d' = (\q\prime_factors d'. q ^ multiplicity q d')" using \j > 0\ prime_factorization_nat unfolding d'_def by auto have "odd d'" unfolding c_def d'_def by simp hence d'_prime_factors_odd: "q \ prime_factors d' \ odd q" for q by fastforce have d'_prime_factors_gt_2: "q \ prime_factors d' \ q > 2" for q using d'_prime_factors_odd by (metis even_numeral in_prime_factors_imp_prime order_less_le prime_ge_2_nat) have "[2 * p = - 1] (mod d')" using \n > 1\ \d' > 1\ unfolding H_2_p cong_iff_dvd_diff by (simp add: int_ops less_1_mult order_less_imp_not_less) hence d'_prime_factors_2_p_mod: "q \ prime_factors d' \ [2 * p = - 1] (mod q)" for q by (rule cong_dvd_modulus; auto) have "q \ prime_factors d' \ coprime (2 * int p) q" for q using d'_prime_factors_2_p_mod by (metis cong_imp_coprime cong_sym coprime_1_left coprime_minus_left_iff mult_2 of_nat_add) hence d'_prime_factors_coprime: "q \ prime_factors d' \ coprime (int p) q" for q using d'_expansion by auto have Legendre_using_quadratic_reciprocity: "Legendre (- d') p = (\q\prime_factors d'. (Legendre p q) ^ multiplicity q d')" proof cases assume "n mod 8 \ {1, 3}" have "k mod 4 = 1" using \n mod 8 \ {1, 3}\ k_mod_4 by auto hence "p mod 4 = 1" using \[p = k] (mod 4 * n)\ by (metis unique_euclidean_semiring_class.cong_def cong_modulus_mult_nat) hence "[int p = 1] (mod 4)" by (metis cong_mod_left cong_refl int_ops(2) int_ops(3) of_nat_mod) have d'_prime_factors_Legendre: "q \ prime_factors d' \ Legendre p q = Legendre q p" for q proof - assume "q \ prime_factors d'" have "prime q" using \q \ prime_factors d'\ by blast have "q > 2" using d'_prime_factors_gt_2 \q \ prime_factors d'\ by blast show "Legendre p q = Legendre q p" using \prime p\ \p > 2\ \[int p = 1] (mod 4)\ \prime q\ \q > 2\ Legendre_equal[of p q] by auto qed have "Legendre (- d') p = Legendre (- 1) p * Legendre d' p" using \prime p\ Legendre_mult[of p "- 1" d'] by auto also have "... = Legendre d' p" using \prime p\ \p > 2\ \[int p = 1] (mod 4)\ Legendre_minus_one by auto also have "... = (\q\prime_factors d'. Legendre (q ^ multiplicity q d') p)" apply (subst d'_expansion) using \prime p\ \p > 2\ Legendre_prod[of p] apply auto done also have "... = (\q\prime_factors d'. (Legendre q p) ^ multiplicity q d')" using \prime p\ \p > 2\ Legendre_power by auto also have "... = (\q\prime_factors d'. (Legendre p q) ^ multiplicity q d')" using d'_prime_factors_Legendre by auto finally show ?thesis . next assume "n mod 8 \ {1, 3}" hence "n mod 8 = 5" using assms by auto have "[p = 3] (mod 4)" using \n mod 8 = 5\ k_mod_4 \[p = k] (mod 4 * n)\ by (metis cong_mod_right cong_modulus_mult_nat) have "[d' = 3] (mod 8)" using \n mod 8 = 5\ unfolding d'_def c_def cong_iff_dvd_diff by (simp add: unique_euclidean_semiring_class.cong_def) have "[d' = - 1] (mod 4)" using \[d' = 3] (mod 8)\ cong_modulus_mult[of d' 3 4 2] unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int by auto have d'_prime_factors_cases: "q \ prime_factors d' \ multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" for q proof - assume "q \ prime_factors d'" consider "[q = 0] (mod 4)" | "[q = 1] (mod 4)" | "[q = 2] (mod 4)" | "[q = 3] (mod 4)" unfolding unique_euclidean_semiring_class.cong_def by (simp; linarith) thus "multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" proof cases assume "[q = 0] (mod 4)" hence False using d'_prime_factors_odd \q \ prime_factors d'\ by (meson cong_0_iff dvd_trans even_numeral) thus "multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" by blast next assume "[q = 1] (mod 4)" thus "multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" by blast next assume "[q = 2] (mod 4)" hence "q = 2" using d'_prime_factors_odd \q \ prime_factors d'\ by (metis unique_euclidean_semiring_class.cong_def dvd_mod_iff even_numeral) thus "multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" by (simp add: \odd d'\ not_dvd_imp_multiplicity_0) next assume "[q = 3] (mod 4)" thus "multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)" by blast qed qed have "d' = (\q\{q\prime_factors d'. True}. q ^ multiplicity q d')" using d'_expansion by auto also have "... = (\q\{q\prime_factors d'. multiplicity q d' = 0 \ [q = 1] (mod 4) \ [q = 3] (mod 4)}. q ^ multiplicity q d')" using d'_prime_factors_cases by meson also have "... = (\q\{q\prime_factors d'. multiplicity q d' = 0} \ {q\prime_factors d'. [q = 1] (mod 4) \ [q = 3] (mod 4)}. q ^ multiplicity q d')" by (rule prod.cong; blast) also have "... = (\q\{q\prime_factors d'. [q = 1] (mod 4) \ [q = 3] (mod 4)}. q ^ multiplicity q d')" by (rule prod.mono_neutral_left[symmetric]; auto) also have "... = (\q\{q\prime_factors d'. [q = 1] (mod 4)} \ {q\prime_factors d'. [q = 3] (mod 4)}. q ^ multiplicity q d')" by (rule prod.cong; blast) also have "... = (\q\{q\prime_factors d'. [q = 1] (mod 4)}. q ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. q ^ multiplicity q d')" by (rule prod.union_disjoint; auto simp add: unique_euclidean_semiring_class.cong_def) finally have d'_expansion_mod_4: "d' = (\q\{q\prime_factors d'. [q = 1] (mod 4)}. q ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. q ^ multiplicity q d')" . have "int d' mod 4 = int ((\q\{q\prime_factors d'. [q = 1] (mod 4)}. q ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. q ^ multiplicity q d') mod 4)" using d'_expansion_mod_4 by presburger also have "... = ((\q\{q\prime_factors d'. [q = 1] (mod 4)}. ((q mod 4) ^ multiplicity q d') mod 4) mod 4) * ((\q\{q\prime_factors d'. [q = 3] (mod 4)}. ((q mod 4) ^ multiplicity q d') mod 4) mod 4) mod 4" unfolding mod_mult_eq mod_prod_eq power_mod .. also have "... = int (((\q\{q\prime_factors d'. [q = 1] (mod 4)}. (1 ^ multiplicity q d') mod 4) mod 4) * ((\q\{q\prime_factors d'. [q = 3] (mod 4)}. (3 ^ multiplicity q d') mod 4) mod 4) mod 4)" unfolding unique_euclidean_semiring_class.cong_def by auto also have "... = ((\q\{q\prime_factors d'. [q = 3] (mod 4)}. (((int 3) mod 4) ^ multiplicity q d') mod 4) mod 4) mod 4" by (simp add: int_ops) also have "... = ((\q\{q\prime_factors d'. [q = 3] (mod 4)}. (((- 1) mod 4) ^ multiplicity q d') mod 4) mod 4) mod 4" by auto also have "... = (\q\{q\prime_factors d'. [q = 3] (mod 4)}. ((- 1) ^ multiplicity q d')) mod 4" unfolding power_mod mod_prod_eq mod_mod_trivial .. finally have "[d' = \q\{q\prime_factors d'. [q = 3] (mod 4)}. ((- 1) ^ multiplicity q d')] (mod 4)" unfolding unique_euclidean_semiring_class.cong_def . hence "[\q\{q\prime_factors d'. [q = 3] (mod 4)}. ((- 1) ^ multiplicity q d') = - 1 :: int] (mod 4)" using \[d' = - 1] (mod 4)\ unfolding unique_euclidean_semiring_class.cong_def by argo hence prod_d'_prime_factors_q_3_mod_4_minus_one: "(\q\{q\prime_factors d'. [q = 3] (mod 4)}. ((- 1) ^ multiplicity q d')) = (- 1 :: int)" unfolding power_sum[symmetric] unfolding minus_one_power_iff unique_euclidean_semiring_class.cong_def by auto have d'_prime_factors_q_1_mod_4_Legendre: "q \ prime_factors d' \ [q = 1] (mod 4) \ Legendre p q = Legendre q p" for q proof - assume "q \ prime_factors d'" assume "[q = 1] (mod 4)" have "prime q" using \q \ prime_factors d'\ by blast have "q > 2" using d'_prime_factors_gt_2 \q \ prime_factors d'\ by blast show "Legendre p q = Legendre q p" using \prime p\ \p > 2\ \[p = 3] (mod 4)\ \[q = 1] (mod 4)\ \prime q\ \q > 2\ Legendre_equal[of p q] unfolding unique_euclidean_semiring_class.cong_def using zmod_int[of q 4] by auto qed have d'_prime_factors_q_3_mod_4_Legendre: "q \ prime_factors d' \ [q = 3] (mod 4) \ Legendre p q = - Legendre q p" for q proof - assume "q \ prime_factors d'" assume "[q = 3] (mod 4)" have "prime q" using \q \ prime_factors d'\ by blast have "q > 2" using d'_prime_factors_gt_2 \q \ prime_factors d'\ by blast have "p \ q" using d'_prime_factors_coprime[of q] \q \ prime_factors d'\ \prime p\ by auto show "Legendre p q = - Legendre q p" using \prime p\ \p > 2\ \[p = 3] (mod 4)\ \[q = 3] (mod 4)\ \prime q\ \q > 2\ \p \ q\ Legendre_opposite[of p q] unfolding unique_euclidean_semiring_class.cong_def using zmod_int[of p 4] zmod_int[of q 4] by fastforce qed have d'_prime_factors_q_0_2_mod_4: "q \ prime_factors d' \ ([q = 0] (mod 4) \ [q = 2] (mod 4)) \ Legendre p q = 1" for q unfolding unique_euclidean_semiring_class.cong_def using d'_prime_factors_odd mod_mod_cancel[of 2 4 q] by fastforce have "Legendre (- d') p = Legendre (- 1) p * Legendre d' p" using \prime p\ Legendre_mult[of p "- 1" d'] by auto also have "... = - Legendre d' p" using \prime p\ \p > 2\ \[p = 3] (mod 4)\ Legendre_minus_one[of p] unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int by (auto simp add: cong_0_iff Legendre_def) also have "... = - (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre q p) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (Legendre q p) ^ multiplicity q d')" apply (subst d'_expansion_mod_4) using \prime p\ \p > 2\ Legendre_mult[of p] Legendre_prod[of p "\q. q ^ multiplicity q d'"] Legendre_power[of p] apply simp done also have "... = - (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (- Legendre p q) ^ multiplicity q d')" using d'_prime_factors_q_1_mod_4_Legendre d'_prime_factors_q_3_mod_4_Legendre by auto also have "... = - (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (Legendre p q * (- 1)) ^ multiplicity q d')" by auto also have "... = - (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (- 1) ^ multiplicity q d')" unfolding power_mult_distrib prod.distrib by auto also have "... = (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (Legendre p q) ^ multiplicity q d')" unfolding prod_d'_prime_factors_q_3_mod_4_minus_one by auto also have "... = (\q\{q\prime_factors d'. [q = 0] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 2] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 3] (mod 4)}. (Legendre p q) ^ multiplicity q d')" using d'_prime_factors_q_0_2_mod_4 by auto also have "... = (\q\{q\prime_factors d'. [q = 0] (mod 4)} \ {q\prime_factors d'. [q = 1] (mod 4)}. (Legendre p q) ^ multiplicity q d') * (\q\{q\prime_factors d'. [q = 2] (mod 4)} \ {q\prime_factors d'. [q = 3] (mod 4)}. (Legendre p q) ^ multiplicity q d')" using prod.union_disjoint[of "{q\prime_factors d'. [q = 0] (mod 4)}" "{q\prime_factors d'. [q = 1] (mod 4)}" "\q. (Legendre p (int q)) ^ multiplicity q d'"] prod.union_disjoint[of "{q\prime_factors d'. [q = 2] (mod 4)}" "{q\prime_factors d'. [q = 3] (mod 4)}" "\q. (Legendre p (int q)) ^ multiplicity q d'"] by (fastforce simp add: unique_euclidean_semiring_class.cong_def) also have "... = (\q\({q\prime_factors d'. [q = 0] (mod 4)} \ {q\prime_factors d'. [q = 1] (mod 4)}) \ ({q\prime_factors d'. [q = 2] (mod 4)} \ {q\prime_factors d'. [q = 3] (mod 4)}). (Legendre p q) ^ multiplicity q d')" by (rule prod.union_disjoint[symmetric]; auto simp add: unique_euclidean_semiring_class.cong_def) also have "... = (\q\{q\prime_factors d'. [q = 0] (mod 4) \ [q = 1] (mod 4) \ [q = 2] (mod 4) \ [q = 3] (mod 4)}. (Legendre p q) ^ multiplicity q d')" by (rule prod.cong; auto) also have "... = (\q\prime_factors d'. (Legendre p q) ^ multiplicity q d')" by (rule prod.cong; auto simp add: unique_euclidean_semiring_class.cong_def) finally show ?thesis . qed have "q \ prime_factors d' \ Legendre 4 q = 1" for q proof - assume "q \ prime_factors d'" have "q dvd 4 \ q \ 4" by (simp add: dvd_imp_le) hence "q dvd 4 \ q \ {0, 1, 2, 3, 4}" by auto hence "q dvd 4 \ q \ {1, 2, 4}" by auto hence "\ q dvd 4" using \q \ prime_factors d'\ d'_prime_factors_odd[of q] by (metis empty_iff even_numeral in_prime_factors_imp_prime insert_iff not_prime_1) hence "\ int q dvd 4" by presburger thus "Legendre 4 q = 1" unfolding Legendre_def QuadRes_def cong_0_iff power2_eq_square by (metis cong_refl mult_2 numeral_Bit0) qed hence "Legendre (- d') p = (\q\prime_factors d'. (Legendre (2 * 2) q * Legendre p q) ^ multiplicity q d')" using Legendre_using_quadratic_reciprocity by auto also have "... = (\q\prime_factors d'. (Legendre 2 q * Legendre (2 * p) q) ^ multiplicity q d')" apply (rule prod.cong[OF refl]) using d'_prime_factors_gt_2 Legendre_mult in_prime_factors_imp_prime by (metis int_ops(7) of_nat_numeral prime_nat_int_transfer mult.assoc) also have "... = (\q\prime_factors d'. (Legendre 2 q * Legendre (- 1) q) ^ multiplicity q d')" apply (rule prod.cong[OF refl]) using d'_prime_factors_2_p_mod Legendre_cong unfolding unique_euclidean_semiring_class.cong_def apply metis done also have "... = (\q\prime_factors d'. ((if [q = 1] (mod 8) \ [q = 7] (mod 8) then 1 else - 1) * (if [q = 1] (mod 4) then 1 else - 1)) ^ multiplicity q d')" apply (rule prod.cong[OF refl]) subgoal for q apply (rule arg_cong2[of _ _ _ _ "\a b. (a * b) ^ multiplicity q d'"]) subgoal using Legendre_two_alt[of q] d'_prime_factors_gt_2[of q] unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int by force subgoal using Legendre_minus_one_alt[of q] d'_prime_factors_gt_2[of q] unfolding unique_euclidean_semiring_class.cong_def nat_mod_as_int by force done done also have "... = (\q\prime_factors d'. ((if [q = 5] (mod 8) \ [q = 7] (mod 8) then - 1 else 1)) ^ multiplicity q d')" apply (rule prod.cong) subgoal by blast subgoal for q apply (rule arg_cong[of _ _ "\a. a ^ multiplicity q d'"]) unfolding unique_euclidean_semiring_class.cong_def apply (simp; presburger) done done also have "... = (\q\prime_factors d'. (if [q = 5] (mod 8) \ [q = 7] (mod 8) then (- 1) ^ multiplicity q d' else 1))" by (rule prod.cong; auto) also have "... = (\q\{q\prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. (- 1) ^ multiplicity q d')" using prod.inter_filter[symmetric] by fast also have "... = (- 1) ^ (\q\{q\prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. multiplicity q d')" by (simp add: power_sum) finally have Legendre_using_sum: "Legendre (- d') p = (- 1) ^ (\q\{q\prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. multiplicity q d')" . have "[\q\{q\prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. multiplicity q d' = 0] (mod 2)" proof - have "d' = (\q\{q \ prime_factors d'. [q = 1] (mod 8) \ [q = 3] (mod 8) \ [q = 5] (mod 8) \ [q = 7] (mod 8)}. q ^ multiplicity q d')" apply (subst d'_expansion) apply (rule prod.cong) subgoal apply (rule Set.set_eqI) subgoal for q apply (rule iffI) subgoal using d'_prime_factors_odd[of q] unfolding unique_euclidean_semiring_class.cong_def apply simp apply presburger done subgoal by blast done done subgoal by blast done also have "... = (\q\({q \ prime_factors d'. [q = 1] (mod 8)} \ {q \ prime_factors d'. [q = 3] (mod 8)}) \ ({q \ prime_factors d'. [q = 5] (mod 8)} \ {q \ prime_factors d'. [q = 7] (mod 8)}). q ^ multiplicity q d')" by (rule prod.cong; auto) also have "... = (\q\({q \ prime_factors d'. [q = 1] (mod 8)} \ {q \ prime_factors d'. [q = 3] (mod 8)}). q ^ multiplicity q d') * (\q\({q \ prime_factors d'. [q = 5] (mod 8)} \ {q \ prime_factors d'. [q = 7] (mod 8)}). q ^ multiplicity q d')" by (rule prod.union_disjoint; force simp add: unique_euclidean_semiring_class.cong_def) also have "... = (\q\{q \ prime_factors d'. [q = 1] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 3] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. q ^ multiplicity q d')" using prod.union_disjoint[of "{q\prime_factors d'. [q = 1] (mod 8)}" "{q\prime_factors d'. [q = 3] (mod 8)}" "\q. q ^ multiplicity q d'"] prod.union_disjoint[of "{q\prime_factors d'. [q = 5] (mod 8)}" "{q\prime_factors d'. [q = 7] (mod 8)}" "\q. q ^ multiplicity q d'"] by (force simp add: unique_euclidean_semiring_class.cong_def) finally have "int (d' mod 8) = (\q\{q \ prime_factors d'. [q = 1] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 3] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. q ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. q ^ multiplicity q d') mod 8" by auto also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. q ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. q ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. q ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. q ^ multiplicity q d') mod 8) mod 8" by (metis (no_types, lifting) mod_mult_eq mod_mod_trivial) also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. (q ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. (q ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (q ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (q ^ multiplicity q d') mod 8) mod 8) mod 8" unfolding mod_prod_eq .. also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. ((q mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. ((q mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. ((q mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. ((q mod 8) ^ multiplicity q d') mod 8) mod 8) mod 8" unfolding power_mod .. also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. (((int q) mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. (((int q) mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (((int q) mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (((int q) mod 8) ^ multiplicity q d') mod 8) mod 8) mod 8" by (simp add: int_ops) also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. ((1 mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. ((3 mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (((- 3) mod 8) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (((- 1) mod 8) ^ multiplicity q d') mod 8) mod 8) mod 8" unfolding cong_int_iff[symmetric] unique_euclidean_semiring_class.cong_def by auto also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. (1 ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. (3 ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. ((- 3) ^ multiplicity q d') mod 8) mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. ((- 1) ^ multiplicity q d') mod 8) mod 8) mod 8" unfolding power_mod .. also have "... = ((\q\{q \ prime_factors d'. [q = 1] (mod 8)}. 1 ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. 3 ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (- 3) ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d') mod 8) mod 8" unfolding mod_prod_eq .. also have "... = (\q\{q \ prime_factors d'. [q = 1] (mod 8)}. 1 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 3] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (- 3) ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d') mod 8" by (metis (no_types, lifting) mod_mult_eq mod_mod_trivial) also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (- 3) ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d') mod 8" by auto also have "... = (\q\{q \ prime_factors d'. [q = 3] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. 3 ^ multiplicity q d' * (- 1) ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d') mod 8" unfolding power_mult_distrib[symmetric] by auto also have "... = ((\q\{q \ prime_factors d'. [q = 3] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)}. 3 ^ multiplicity q d')) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8)}. (- 1) ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d')) mod 8" unfolding prod.distrib by (simp add: algebra_simps) also have "... = ((\q\{q \ prime_factors d'. [q = 3] (mod 8)} \ {q \ prime_factors d'. [q = 5] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8)} \ {q \ prime_factors d'. [q = 7] (mod 8)}. (- 1) ^ multiplicity q d')) mod 8" apply (subst prod.union_disjoint[of "{q\prime_factors d'. [q = 5] (mod 8)}" "{q\prime_factors d'. [q = 7] (mod 8)}" "\q. (- 1) ^ multiplicity q d'"] ) apply ((force simp add: unique_euclidean_semiring_class.cong_def)+)[3] apply (subst prod.union_disjoint[of "{q\prime_factors d'. [q = 3] (mod 8)}" "{q\prime_factors d'. [q = 5] (mod 8)}" "\q. 3 ^ multiplicity q d'"] ) apply ((force simp add: unique_euclidean_semiring_class.cong_def)+)[3] apply blast done also have "... = ((\q\{q \ prime_factors d'. [q = 3] (mod 8) \ [q = 5] (mod 8)}. 3 ^ multiplicity q d') * (\q\{q \ prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. (- 1) ^ multiplicity q d')) mod 8" by (rule arg_cong2[of _ _ _ _ "\A B. ((\q\A. _ q) * (\q\B. _ q)) mod 8"]; auto) also have "... = (((\q\{q \ prime_factors d'. [q = 3] (mod 8) \ [q = 5] (mod 8)}. 3 ^ multiplicity q d') mod 8) * ((\q\{q \ prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. (- 1) ^ multiplicity q d') mod 8)) mod 8" unfolding mod_mult_eq .. also have "... = ((3 ^ (\q\{q \ prime_factors d'. [q = 3] (mod 8) \ [q = 5] (mod 8)}. multiplicity q d') mod 8) * ((- 1) ^ (\q\{q \ prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. multiplicity q d') mod 8)) mod 8" unfolding power_sum .. also have "... = int (case ((\q\{q \ prime_factors d'. [q = 3] (mod 8) \ [q = 5] (mod 8)}. multiplicity q d') mod 2, (\q\{q \ prime_factors d'. [q = 5] (mod 8) \ [q = 7] (mod 8)}. multiplicity q d') mod 2) of (0 , 0 ) \ 1 | (0 , Suc 0) \ 7 | (Suc 0, 0 ) \ 3 | (Suc 0, Suc 0) \ 5)" (is "_ = int ?d'_mod_8") unfolding three_mod_eight_power_iff minus_one_power_iff by (simp; simp add: odd_iff_mod_2_eq_one) finally have d'_mod_8: "d' mod 8 = ?d'_mod_8"by linarith have "[d' = 1] (mod 8) \ [d' = 3] (mod 8)" unfolding d'_def c_def unique_euclidean_semiring_class.cong_def using assms by auto hence "?d'_mod_8 = 1 \ ?d'_mod_8 = 3" unfolding unique_euclidean_semiring_class.cong_def d'_mod_8 by auto thus ?thesis unfolding unique_euclidean_semiring_class.cong_def by (smt (z3) Collect_cong One_nat_def cong_exp_iff_simps(11) even_mod_2_iff even_numeral nat.case(2) numeral_eq_iff numerals(1) old.nat.simps(4) parity_cases prod.simps(2) semiring_norm(84)) qed hence "Legendre (- d') p = 1" using Legendre_using_sum unfolding minus_one_power_iff cong_0_iff by argo thus "QuadRes p (- d')" unfolding Legendre_def by (metis one_neq_neg_one one_neq_zero) qed from \QuadRes p (- d')\ obtain x\<^sub>0 y where "x\<^sub>0\<^sup>2 - (- d') = y * (int p)" unfolding quadratic_residue_alt_equiv[symmetric] quadratic_residue_alt_def by auto hence "(int p) dvd (x\<^sub>0\<^sup>2 - - d')" by simp define x :: int where "x \ if odd x\<^sub>0 then x\<^sub>0 else (x\<^sub>0 + p)" have "even (4 * int n * j)" by simp moreover have "odd k" using \coprime k (4 * n)\ by auto ultimately have "odd (int p)" unfolding p_def by simp have "odd x" unfolding x_def using \odd (int p)\ by auto have "QuadRes (2 * p) (- d')" unfolding quadratic_residue_alt_equiv[symmetric] quadratic_residue_alt_def proof - have "2 dvd (x\<^sup>2 - - d')" unfolding d'_def c_def using \odd x\ by auto moreover from \(int p) dvd (x\<^sub>0\<^sup>2 - - d')\ have "(int p) dvd (x\<^sup>2 - - d')" unfolding x_def power2_eq_square apply (simp add: algebra_simps) unfolding add.assoc[symmetric, of d' "x\<^sub>0 * x\<^sub>0"] apply auto done moreover have "coprime 2 (int p)" using \odd (int p)\ by auto ultimately have "(int (2 * p)) dvd (x\<^sup>2 - - d')" by (simp add: divides_mult) hence "(x\<^sup>2 - - d') mod (int (2 * p)) = 0" by simp - hence "\y. x\<^sup>2 - - d' = int (2 * p) * y" by (simp add: zmod_eq_0D) + hence "\y. x\<^sup>2 - - d' = int (2 * p) * y" by auto hence "\y. x\<^sup>2 - - d' = y * int (2 * p)" by (simp add: algebra_simps) thus "\x y. x\<^sup>2 - - d' = y * int (2 * p)" by (rule exI[where ?x=x]) qed have "n \ 2" using \1 < n\ by auto moreover have "0 < nat d'" unfolding d'_def using \j > 0\ by simp moreover have "QuadRes (int (nat d' * n - 1)) (- d')" using \d' > 1\ H_2_p \QuadRes (2 * p) (- d')\ by (simp add: int_ops) ultimately show "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using three_squares_using_quadratic_residue[of n "nat d'"] by auto qed lemma power_two_mod_eight: fixes n :: nat shows "n\<^sup>2 mod 8 \ {0, 1, 4}" proof - have 0: "n\<^sup>2 mod 8 = (n mod 8)\<^sup>2 mod 8" unfolding power2_eq_square by (simp add: mod_mult_eq) have "n mod 8 \ {0, 1, 2, 3, 4, 5, 6, 7}" by auto hence "(n mod 8)\<^sup>2 mod 8 \ {0, 1, 4}" unfolding power2_eq_square by auto thus "n\<^sup>2 mod 8 \ {0, 1, 4}" unfolding 0 . qed lemma power_two_mod_four: fixes n :: nat shows "n\<^sup>2 mod 4 \ {0, 1}" using power_two_mod_eight[of n] mod_mod_cancel[of 4 8 "n\<^sup>2"] by auto text \Theorem 1.4 from @{cite nathanson1996}.\ theorem three_squares_iff: fixes n :: nat shows "(\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2) \ (\a k. n = 4 ^ a * (8 * k + 7))" proof assume "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" then obtain x\<^sub>1 x\<^sub>2 x\<^sub>3 where 0: "n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by blast show "\a k. n = 4 ^ a * (8 * k + 7)" proof assume "\a k. n = 4 ^ a * (8 * k + 7)" then obtain a k where 1: "n = 4 ^ a * (8 * k + 7)" by blast from 0 1 show False proof (induction a arbitrary: n x\<^sub>1 x\<^sub>2 x\<^sub>3 rule: nat.induct) fix n x\<^sub>1 x\<^sub>2 x\<^sub>3 :: nat assume 2: "n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" assume "n = 4 ^ 0 * (8 * k + 7)" hence 3: "n mod 8 = 7" unfolding 1 by auto have "(x\<^sub>1\<^sup>2 mod 8 + x\<^sub>2\<^sup>2 mod 8 + x\<^sub>3\<^sup>2 mod 8) mod 8 = 7" unfolding 2 3[symmetric] by (meson mod_add_cong mod_mod_trivial) thus False using power_two_mod_eight[of x\<^sub>1] power_two_mod_eight[of x\<^sub>2] power_two_mod_eight[of x\<^sub>3] by auto next fix a' n x\<^sub>1 x\<^sub>2 x\<^sub>3 :: nat assume 4: "\n' x\<^sub>1' x\<^sub>2' x\<^sub>3' :: nat. n' = x\<^sub>1'\<^sup>2 + x\<^sub>2'\<^sup>2 + x\<^sub>3'\<^sup>2 \ n' = 4 ^ a' * (8 * k + 7) \ False" assume 5: "n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" assume "n = 4 ^ Suc a' * (8 * k + 7)" hence "n = 4 * (4 ^ a' * (8 * k + 7))" (is "n = 4 * ?m") by auto hence 6: "4 * ?m = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" unfolding 5 by auto have "(x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2) mod 4 = 0" using 6 by presburger hence "((x\<^sub>1\<^sup>2 mod 4) + (x\<^sub>2\<^sup>2 mod 4) + (x\<^sub>3\<^sup>2 mod 4)) mod 4 = 0" by presburger hence "x\<^sub>1\<^sup>2 mod 4 = 0 \ x\<^sub>2\<^sup>2 mod 4 = 0 \ x\<^sub>3\<^sup>2 mod 4 = 0" using power_two_mod_four[of x\<^sub>1] power_two_mod_four[of x\<^sub>2] power_two_mod_four[of x\<^sub>3] by (auto; presburger) hence "even x\<^sub>1 \ even x\<^sub>2 \ even x\<^sub>3" by (metis dvd_0_right even_even_mod_4_iff even_power) hence "4 * ?m = 4 * ((x\<^sub>1 div 2)\<^sup>2 + (x\<^sub>2 div 2)\<^sup>2 + (x\<^sub>3 div 2)\<^sup>2)" unfolding 6 by fastforce hence "?m = (x\<^sub>1 div 2)\<^sup>2 + (x\<^sub>2 div 2)\<^sup>2 + (x\<^sub>3 div 2)\<^sup>2" by auto thus False using 4 by blast qed qed next assume 7: "\a k. n = 4 ^ a * (8 * k + 7)" show "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" proof cases assume "n = 0" thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by auto next assume 8: "n \ 0" have "n > 0 \ \a m. n = 4 ^ a * m \ \ 4 dvd m" proof (induction n rule: less_induct) fix n :: nat assume 9: "\n'. n' < n \ n' > 0 \ \a' m'. n' = 4 ^ a' * m' \ \ 4 dvd m'" assume 10: "n > 0" show "\a m. n = 4 ^ a * m \ \ 4 dvd m" proof cases assume 11: "4 dvd n" have "n div 4 < n" "n div 4 > 0" using 10 11 by auto then obtain a' m' where 12: "n div 4 = 4 ^ a' * m' \ \ 4 dvd m'" using 9 by blast have "n = 4 ^ (Suc a') * m' \ \ 4 dvd m'" using 11 12 by auto thus "\a m. n = 4 ^ a * m \ \ 4 dvd m" by blast next assume "\ 4 dvd n" hence "n = 4 ^ 0 * n \ \ 4 dvd n" by auto thus "\a m. n = 4 ^ a * m \ \ 4 dvd m" by blast qed qed then obtain a m where 13: "n = 4 ^ a * m" "\ 4 dvd m" using 8 by auto have 14: "m mod 8 \ 7" proof assume "m mod 8 = 7" then obtain k where "m = 8 * k + 7" by (metis div_mod_decomp mult.commute) hence "n = 4 ^ a * (8 * k + 7)" unfolding 13 by blast thus False using 7 by blast qed have "m mod 4 = 2 \ m mod 8 \ {1, 3, 5, 7}" using 13 by (simp; presburger) hence "m mod 4 = 2 \ m mod 8 \ {1, 3, 5}" using 14 by blast hence "\x\<^sub>1 x\<^sub>2 x\<^sub>3. m = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using three_squares_using_mod_four three_squares_using_mod_eight by blast hence "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = (2 ^ a * x\<^sub>1)\<^sup>2 + (2 ^ a * x\<^sub>2)\<^sup>2 + (2 ^ a * x\<^sub>3)\<^sup>2" unfolding 13 power2_eq_square unfolding mult.assoc[of "2 ^ a"] unfolding mult.commute[of "2 ^ a"] unfolding mult.assoc unfolding power_add[symmetric] unfolding mult_2[symmetric] unfolding power_mult unfolding mult.assoc[symmetric] unfolding add_mult_distrib[symmetric] unfolding mult.commute[of "4 ^ a"] by simp thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by blast qed qed text \Theorem 1.5 from @{cite nathanson1996}.\ theorem odd_three_squares_using_mod_eight: fixes n :: nat assumes "n mod 8 = 3" shows "\x\<^sub>1 x\<^sub>2 x\<^sub>3. odd x\<^sub>1 \ odd x\<^sub>2 \ odd x\<^sub>3 \ n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" proof - obtain x\<^sub>1 x\<^sub>2 x\<^sub>3 where 0: "n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using assms three_squares_using_mod_eight by blast have "(x\<^sub>1\<^sup>2 mod 8 + x\<^sub>2\<^sup>2 mod 8 + x\<^sub>3\<^sup>2 mod 8) mod 8 = 3" unfolding 0 assms[symmetric] by (meson mod_add_cong mod_mod_trivial) hence "x\<^sub>1\<^sup>2 mod 8 = 1 \ x\<^sub>2\<^sup>2 mod 8 = 1 \ x\<^sub>3\<^sup>2 mod 8 = 1" using power_two_mod_eight[of x\<^sub>1] power_two_mod_eight[of x\<^sub>2] power_two_mod_eight[of x\<^sub>3] by auto hence "odd x\<^sub>1 \ odd x\<^sub>2 \ odd x\<^sub>3" by (metis dvd_mod even_numeral even_power odd_one pos2) hence "odd x\<^sub>1 \ odd x\<^sub>2 \ odd x\<^sub>3 \ n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using 0 by blast thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3. odd x\<^sub>1 \ odd x\<^sub>2 \ odd x\<^sub>3 \ n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" by blast qed subsection \Consequences\ lemma four_decomposition: fixes n :: nat shows "\x y z. n = x\<^sup>2 + y\<^sup>2 + z\<^sup>2 + z" proof - have "(4 * n + 1) mod 8 \ {1, 3, 5}" by (simp; presburger) then obtain x y z where 0: "4 * n + 1 = x\<^sup>2 + y\<^sup>2 + z\<^sup>2" using three_squares_using_mod_eight by blast hence 1: "1 = (x\<^sup>2 + y\<^sup>2 + z\<^sup>2) mod 4" by (metis Suc_0_mod_numeral(2) Suc_eq_plus1 mod_add_left_eq mod_mult_self1_is_0) show ?thesis proof cases assume "even x" then obtain x' where H_x: "x = 2 * x'" by blast show ?thesis proof cases assume "even y" then obtain y' where H_y: "y = 2 * y'" by blast show ?thesis proof cases assume "even z" then obtain z' where H_z: "z = 2 * z'" by blast show ?thesis using 1 unfolding H_x H_y H_z by auto next assume "odd z" then obtain z' where H_z: "z = 2 * z' + 1" using oddE by blast have "n = x'\<^sup>2 + y'\<^sup>2 + z'\<^sup>2 + z'" using 0 unfolding H_x H_y H_z power2_eq_square by auto thus ?thesis by blast qed next assume "odd y" then obtain y' where H_y: "y = 2 * y' + 1" using oddE by blast show ?thesis proof cases assume "even z" then obtain z' where H_z: "z = 2 * z'" by blast have "n = x'\<^sup>2 + z'\<^sup>2 + y'\<^sup>2 + y'" using 0 unfolding H_x H_y H_z power2_eq_square by auto thus ?thesis by blast next assume "odd z" then obtain z' where H_z: "z = 2 * z' + 1" using oddE by blast show ?thesis using 1 unfolding H_x H_y H_z power2_eq_square by (metis dvd_mod even_add even_mult_iff even_numeral odd_one) qed qed next assume "odd x" then obtain x' where H_x: "x = 2 * x' + 1" using oddE by blast show ?thesis proof cases assume "even y" then obtain y' where H_y: "y = 2 * y'" by blast show ?thesis proof cases assume "even z" then obtain z' where H_z: "z = 2 * z'" by blast have "n = y'\<^sup>2 + z'\<^sup>2 + x'\<^sup>2 + x'" using 0 unfolding H_x H_y H_z power2_eq_square by auto thus ?thesis by blast next assume "odd z" then obtain z' where H_z: "z = 2 * z' + 1" using oddE by blast show ?thesis using 1 unfolding H_x H_y H_z power2_eq_square by (metis dvd_mod even_add even_mult_iff even_numeral odd_one) qed next assume "odd y" then obtain y' where H_y: "y = 2 * y' + 1" using oddE by blast show ?thesis proof cases assume "even z" then obtain z' where H_z: "z = 2 * z'" by blast show ?thesis using 1 unfolding H_x H_y H_z power2_eq_square by (metis dvd_mod even_add even_mult_iff even_numeral odd_one) next assume "odd z" then obtain z' where H_z: "z = 2 * z' + 1" using oddE by blast show ?thesis using 1 unfolding H_x H_y H_z power2_eq_square by (simp add: mod_add_eq[symmetric]) qed qed qed qed theorem four_decomposition_int: fixes n :: int shows "(\x y z. n = x\<^sup>2 + y\<^sup>2 + z\<^sup>2 + z) \ n \ 0" proof assume "\x y z. n = x\<^sup>2 + y\<^sup>2 + z\<^sup>2 + z" then obtain x y z where 0: "n = x\<^sup>2 + y\<^sup>2 + z\<^sup>2 + z" by blast show "n \ 0" unfolding 0 power2_eq_square by (smt (verit) div_pos_neg_trivial mult_le_0_iff nonzero_mult_div_cancel_right sum_squares_ge_zero) next assume "n \ 0" thus "\x y z. n = x\<^sup>2 + y\<^sup>2 + z\<^sup>2 + z" using four_decomposition[of "nat n"] by (smt (verit) int_eq_iff int_plus of_nat_power) qed theorem four_squares: fixes n :: nat shows "\x\<^sub>1 x\<^sub>2 x\<^sub>3 x\<^sub>4. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2 + x\<^sub>4\<^sup>2" proof cases assume "\a k. n = 4 ^ a * (8 * k + 7)" then obtain a k where "n = 4 ^ a * (8 * k + 7)" by blast hence 0: "n = 4 ^ a * (8 * k + 6) + (2 ^ a)\<^sup>2" by (metis add_mult_distrib left_add_mult_distrib mult.commute mult_numeral_1 numeral_Bit0 numeral_plus_numeral power2_eq_square power_mult_distrib semiring_norm(5)) have "\a' k'. 4 ^ a * (8 * k + 6) = 4 ^ a' * (8 * k' + 7)" proof assume "\a' k'. 4 ^ a * (8 * k + 6) = 4 ^ a' * (8 * k' + 7)" then obtain a' k' where 1: "4 ^ a * (8 * k + 6) = 4 ^ a' * (8 * k' + 7)" by blast show False proof (cases rule: linorder_cases[of a a']) assume "a < a'" hence 2: "a' = a + (a' - a)" "a' - a > 0" by auto have 3: "4 ^ a * (8 * k + 6) = 4 ^ a * 4 ^ (a' - a) * (8 * k' + 7)" using 1 2 by (metis power_add) have "2 = (8 * k + 6) mod 4" by presburger also have "... = (4 ^ (a' - a) * (8 * k' + 7)) mod 4" using 3 by auto also have "... = 0" using 2 by auto finally show False by auto next assume "a = a'" hence "8 * k + 6 = 8 * k' + 7" using 1 by auto thus False by presburger next assume "a > a'" hence 4: "a = a' + (a - a')" "a - a' > 0" by auto have 5: "4 ^ a' * 4 ^ (a - a') * (8 * k + 6) = 4 ^ a' * (8 * k' + 7)" using 1 4 by (metis power_add) have "0 = (4 ^ (a - a') * (8 * k + 6)) mod 4" using 4 by auto also have "... = (8 * k' + 7) mod 4" using 5 by auto also have "... = 3" by presburger finally show False by auto qed qed then obtain x\<^sub>1 x\<^sub>2 x\<^sub>3 where "4 ^ a * (8 * k + 6) = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using three_squares_iff by blast thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3 x\<^sub>4. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2 + x\<^sub>4\<^sup>2" unfolding 0 by auto next assume "\a k. n = 4 ^ a * (8 * k + 7)" hence "\x\<^sub>1 x\<^sub>2 x\<^sub>3. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2" using three_squares_iff by blast thus "\x\<^sub>1 x\<^sub>2 x\<^sub>3 x\<^sub>4. n = x\<^sub>1\<^sup>2 + x\<^sub>2\<^sup>2 + x\<^sub>3\<^sup>2 + x\<^sub>4\<^sup>2" by (metis zero_power2 add_0) qed end