diff --git a/thys/BenOr_Kozen_Reif/BKR_Algorithm.thy b/thys/BenOr_Kozen_Reif/BKR_Algorithm.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/BKR_Algorithm.thy @@ -0,0 +1,146 @@ +theory BKR_Algorithm + imports + "More_Matrix" + "Sturm_Tarski.Sturm_Tarski" +begin + +section "Setup" + +definition retrieve_polys:: "real poly list \ nat list \ real poly list" + where "retrieve_polys qss index_list = (map (nth qss) index_list)" + +definition construct_NofI:: "real poly \ real poly list \ rat" + where "construct_NofI p I = rat_of_int (changes_R_smods p ((pderiv p)*(prod_list I)))" + +definition construct_rhs_vector:: "real poly \ real poly list \ nat list list \ rat vec" + where "construct_rhs_vector p qs Is = vec_of_list (map (\ I.(construct_NofI p (retrieve_polys qs I))) Is)" + +section "Base Case" + +definition base_case_info:: "(rat mat \ (nat list list \ rat list list))" + where "base_case_info = + ((mat_of_rows_list 2 [[1,1], [1,-1]]), ([[],[0]], [[1],[-1]]))" + +(* When p, q are coprime, this will actually be an int vec, which is why taking the floor is okay *) +definition base_case_solve_for_lhs:: "real poly \ real poly \ rat vec" + where "base_case_solve_for_lhs p q = (mult_mat_vec (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -1/2]]) (construct_rhs_vector p [q] [[], [0]]))" + +thm "gauss_jordan_compute_inverse" + +primrec matr_option:: "nat \ 'a::{one, zero} mat option \ 'a mat" + where "matr_option dimen None = 1\<^sub>m dimen" + | "matr_option dimen (Some c) = c" + +(* For smooth code export, we want to use a computable notion of matrix equality *) +definition mat_equal:: "'a:: field mat \ 'a :: field mat \ bool" + where "mat_equal A B = (dim_row A = dim_row B \ dim_col A = dim_col B \ (mat_to_list A) = (mat_to_list B))" + +definition mat_inverse_var :: "'a :: field mat \ 'a mat option" where + "mat_inverse_var A = (if dim_row A = dim_col A then + let one = 1\<^sub>m (dim_row A) in + (case gauss_jordan A one of + (B, C) \ if (mat_equal B one) then Some C else None) else None)" + +(* Now solve for LHS in general. + Because mat_inverse returns an option type, we pattern match on this. + Notice that when we call this function in the algorithm, the matrix we pass will always be invertible, + given how the construction works. *) +definition solve_for_lhs:: "real poly \ real poly list \ nat list list \ rat mat \ rat vec" + where "solve_for_lhs p qs subsets matr = + mult_mat_vec (matr_option (dim_row matr) (mat_inverse_var matr)) (construct_rhs_vector p qs subsets)" + +section "Smashing" + +definition subsets_smash::"nat \ nat list list \ nat list list \ nat list list" + where "subsets_smash n s1 s2 = concat (map (\l1. map (\ l2. l1 @ (map ((+) n) l2)) s2) s1)" + +definition signs_smash::"'a list list \ 'a list list \ 'a list list" + where "signs_smash s1 s2 = concat (map (\l1. map (\ l2. l1 @ l2) s2) s1)" + +definition smash_systems:: "real poly \ real poly list \ real poly list \ nat list list \ nat list list \ + rat list list \ rat list list \ rat mat \ rat mat \ + real poly list \ (rat mat \ (nat list list \ rat list list))" + where "smash_systems p qs1 qs2 subsets1 subsets2 signs1 signs2 mat1 mat2 = + (qs1@qs2, (kronecker_product mat1 mat2, (subsets_smash (length qs1) subsets1 subsets2, signs_smash signs1 signs2)))" + +fun combine_systems:: "real poly \ (real poly list \ (rat mat \ (nat list list \ rat list list))) \ (real poly list \ (rat mat \ (nat list list \ rat list list))) + \ (real poly list \ (rat mat \ (nat list list \ rat list list)))" + where "combine_systems p (qs1, m1, sub1, sgn1) (qs2, m2, sub2, sgn2) = + (smash_systems p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2)" + +(* 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! *) +section "Reduction" +definition find_nonzeros_from_input_vec:: "rat vec \ nat list" + where "find_nonzeros_from_input_vec lhs_vec = filter (\i. lhs_vec $ i \ 0) [0..< dim_vec lhs_vec]" + +definition take_indices:: "'a list \ nat list \ 'a list" + where "take_indices subsets indices = map ((!) subsets) indices" + +definition take_cols_from_matrix:: "'a mat \ nat list \ 'a mat" + where "take_cols_from_matrix matr indices_to_keep = + mat_of_cols (dim_row matr) ((take_indices (cols matr) indices_to_keep):: 'a vec list)" + +definition take_rows_from_matrix:: "'a mat \ nat list \ 'a mat" + where "take_rows_from_matrix matr indices_to_keep = + mat_of_rows (dim_col matr) ((take_indices (rows matr) indices_to_keep):: 'a vec list)" + +fun reduce_mat_cols:: "'a mat \ rat vec \ 'a mat" + where "reduce_mat_cols A lhs_vec = take_cols_from_matrix A (find_nonzeros_from_input_vec lhs_vec)" + +(* Find which rows to drop. *) +definition rows_to_keep:: "('a::field) mat \ nat list" where + "rows_to_keep A = map snd (pivot_positions (gauss_jordan_single (A\<^sup>T)))" + +fun reduction_step:: "rat mat \ rat list list \ nat list list \ rat vec \ rat mat \ (nat list list \ rat list list)" + where "reduction_step A signs subsets lhs_vec = + (let reduce_cols_A = (reduce_mat_cols A lhs_vec); + rows_keep = rows_to_keep reduce_cols_A in + (take_rows_from_matrix reduce_cols_A rows_keep, + (take_indices subsets rows_keep, + take_indices signs (find_nonzeros_from_input_vec lhs_vec))))" + +fun reduce_system:: "real poly \ (real poly list \ (rat mat \ (nat list list \ rat list list))) \ (rat mat \ (nat list list \ rat list list))" + where "reduce_system p (qs,m,subs,signs) = + reduction_step m signs subs (solve_for_lhs p qs subs m)" + +section "Overall algorithm " + (* + Find the matrix, subsets, signs for an input p and qs. + The "rat mat" in the output is the matrix. The "nat list list" is the list of subsets. + The "rat list list" is the list of signs. + + We will want to call this when p is nonzero and when every q in qs is pairwise coprime to p. + Properties of this algorithm are proved in BKR_Proofs.thy. + *) +fun calculate_data:: "real poly \ real poly list \ (rat mat \ (nat list list \ rat list list))" + where + "calculate_data p qs = + ( let len = length qs in + if len = 0 then + (\(a,b,c).(a,b,map (drop 1) c)) (reduce_system p ([1],base_case_info)) + else if len \ 1 then reduce_system p (qs,base_case_info) + else + (let q1 = take (len div 2) qs; left = calculate_data p q1; + q2 = drop (len div 2) qs; right = calculate_data p q2; + comb = combine_systems p (q1,left) (q2,right) in + reduce_system p comb + ) + )" + +(* Extract the list of consistent sign assignments *) +definition find_consistent_signs_at_roots:: "real poly \ real poly list \ rat list list" + where [code]: + "find_consistent_signs_at_roots p qs = + ( let (M,S,\) = calculate_data p qs in \ )" + +lemma find_consistent_signs_at_roots_thm: + shows "find_consistent_signs_at_roots p qs = snd (snd (calculate_data p qs))" + by (simp add: case_prod_beta find_consistent_signs_at_roots_def) + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/BKR_Decision.thy b/thys/BenOr_Kozen_Reif/BKR_Decision.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/BKR_Decision.thy @@ -0,0 +1,2034 @@ +theory BKR_Decision + imports BKR_Algorithm + "Berlekamp_Zassenhaus.Factorize_Rat_Poly" + "Algebraic_Numbers.Real_Roots" + "BKR_Proofs" + "HOL.Deriv" +begin + +section "Algorithm" + +subsection "Parsing" + (* Formula type *) +datatype 'a fml = + And "'a fml" "'a fml" + | Or "'a fml" "'a fml" + | Gt 'a (* 'a > 0 *) + | Geq 'a (* 'a \ 0 *) + | Lt 'a (* 'a < 0 *) + | Leq 'a (* 'a \ 0 *) + | Eq 'a (* 'a = 0 *) + | Neq 'a (* 'a \ 0 *) + +(* Evaluating a formula over a lookup semantics where 'a is nat *) +primrec lookup_sem :: "nat fml \ ('a::linordered_field list) \ bool" + where + "lookup_sem (And l r) ls = (lookup_sem l ls \ lookup_sem r ls)" + | "lookup_sem (Or l r) ls = (lookup_sem l ls \ lookup_sem r ls)" + | "lookup_sem (Gt p) ls = (ls ! p > 0)" + | "lookup_sem (Geq p) ls = (ls ! p \ 0)" + | "lookup_sem (Lt p) ls = (ls ! p < 0)" + | "lookup_sem (Leq p) ls = (ls ! p \ 0)" + | "lookup_sem (Eq p) ls = (ls ! p = 0)" + | "lookup_sem (Neq p) ls = (ls ! p \ 0)" + +(* (compute) all polynomials mentioned in a formula *) +primrec poly_list :: "'a fml \ 'a list" + where + "poly_list (And l r) = poly_list l @ poly_list r" + | "poly_list (Or l r) = poly_list l @ poly_list r" + | "poly_list (Gt p) = [p]" + | "poly_list (Geq p) = [p]" + | "poly_list (Lt p) = [p]" + | "poly_list (Leq p) = [p]" + | "poly_list (Eq p) = [p]" + | "poly_list (Neq p) = [p]" + +primrec index_of_aux :: "'a list \ 'a \ nat \ nat" where + "index_of_aux [] y n = n" +| "index_of_aux (x#xs) y n = + (if x = y then n else index_of_aux xs y (n+1))" + +definition index_of :: "'a list \ 'a \ nat" where + "index_of xs y = index_of_aux xs y 0" + +definition convert :: "'a fml \ (nat fml \ 'a list)" + where + "convert fml = ( + let ps = remdups (poly_list fml) + in + (map_fml (index_of ps) fml, ps) + )" + + +subsection "Factoring" + +(* Makes sure the result of factorize_rat_poly is monic *) +definition factorize_rat_poly_monic :: "rat poly \ (rat \ (rat poly \ nat) list)" + where + "factorize_rat_poly_monic p = ( + let (c,fs) = factorize_rat_poly p ; + lcs = prod_list (map (\(f,i). (lead_coeff f) ^ Suc i) fs) ; + fs = map (\(f,i). (normalize f, i)) fs + in + (c * lcs,fs) + )" + +(* Factoring an input list of polynomials *) +definition factorize_polys :: "rat poly list \ (rat poly list \ (rat \ (nat \ nat) list) list)" + where + "factorize_polys ps = ( + let fact_ps = map factorize_rat_poly_monic ps; + factors = remdups (map fst (concat (map snd fact_ps))) ; + data = map (\(c,fs). (c, map (\(f,pow). (index_of factors f, pow) ) fs)) fact_ps + in + (factors,data) + )" + +(* After turning a polynomial into factors, + this turns a sign condition on the factors + into a sign condition for the polynomial *) +definition undo_factorize :: "rat \ (nat \ nat) list \ rat list \ rat" + where + "undo_factorize cfs signs = + squash + (case cfs of (c,fs) \ + (c * prod_list (map (\(f,pow). (signs ! f) ^ Suc pow) fs))) + " + +definition undo_factorize_polys :: "(rat \ (nat \ nat) list) list \ rat list \ rat list" + where + "undo_factorize_polys ls signs = map (\l. undo_factorize l signs) ls" + +subsection "Auxiliary Polynomial" +definition crb:: "real poly \ int" where + "crb p = ceiling (2 + max_list_non_empty (map (\ i. norm (coeff p i)) [0 ..< degree p]) + / norm (lead_coeff p))" + +(* Because we are using prod_list instead of lcm, it's important that this is called + when ps is pairwise coprime. *) +definition coprime_r :: "real poly list \ real poly" + where + "coprime_r ps = pderiv (prod_list ps) * ([:-(crb (prod_list ps)),1:]) * ([:(crb (prod_list ps)),1:])" + + +subsection "Setting Up the Procedure" + (* 0 indexed *) +definition insertAt :: "nat \ 'a \ 'a list \ 'a list" where + "insertAt n x ls = take n ls @ x # (drop n ls)" + +(* 0 indexed *) +definition removeAt :: "nat \ 'a list \ 'a list" where + "removeAt n ls = take n ls @ (drop (n+1) ls)" + +definition find_sgas_aux:: "real poly list \ rat list list" + where "find_sgas_aux in_list = + concat (map (\i. + map (\v. insertAt i 0 v) (find_consistent_signs_at_roots (in_list ! i) (removeAt i in_list)) + ) [0.. rat list list" + where + "find_sgas ps = ( + let r = coprime_r ps in + find_consistent_signs_at_roots r ps @ find_sgas_aux ps + )" + +(* Putting the sign condition preprocessing together with BKR *) +definition find_consistent_signs :: "rat poly list \ rat list list" + where + "find_consistent_signs ps = ( + let (fs,data) = factorize_polys ps; + sgas = find_sgas (map (map_poly of_rat) fs); + rsgas = map (undo_factorize_polys data) sgas + in + (if fs = [] then [(map (\x. if poly x 0 < 0 then -1 else if poly x 0 = 0 then 0 else 1) ps)] else rsgas) + )" + + +subsection "Deciding Univariate Problems" +definition decide_universal :: "rat poly fml \ bool" + where [code]: + "decide_universal fml = ( + let (fml_struct,polys) = convert fml; + conds = find_consistent_signs polys + in + list_all (lookup_sem fml_struct) conds + )" + +definition decide_existential :: "rat poly fml \ bool" + where [code]: + "decide_existential fml = ( + let (fml_struct,polys) = convert fml; + conds = find_consistent_signs polys + in + find (lookup_sem fml_struct) conds \ None + )" + +section "Proofs" +subsection "Parsing and Semantics" + (* Evaluating a formula where 'a is a real poly *) +primrec real_sem :: "real poly fml \ real \ bool" + where + "real_sem (And l r) x = (real_sem l x \ real_sem r x)" + | "real_sem (Or l r) x = (real_sem l x \ real_sem r x)" + | "real_sem (Gt p) x = (poly p x > 0)" + | "real_sem (Geq p) x = (poly p x \ 0)" + | "real_sem (Lt p) x = (poly p x < 0)" + | "real_sem (Leq p) x = (poly p x \ 0)" + | "real_sem (Eq p) x = (poly p x = 0)" + | "real_sem (Neq p) x = (poly p x \ 0)" + +(* Evaluating a formula where 'a is a rat poly *) +primrec fml_sem :: "rat poly fml \ real \ bool" + where + "fml_sem (And l r) x = (fml_sem l x \ fml_sem r x)" + | "fml_sem (Or l r) x = (fml_sem l x \ fml_sem r x)" + | "fml_sem (Gt p) x = (rpoly p x > 0)" + | "fml_sem (Geq p) x = (rpoly p x \ 0)" + | "fml_sem (Lt p) x = (rpoly p x < 0)" + | "fml_sem (Leq p) x = (rpoly p x \ 0)" + | "fml_sem (Eq p) x = (rpoly p x = 0)" + | "fml_sem (Neq p) x = (rpoly p x \ 0)" + +lemma poly_list_set_fml: + shows "set (poly_list fml) = set_fml fml" + apply (induction) by auto + +lemma convert_semantics_lem: + assumes "\p. p \ set (poly_list fml) \ + ls ! (index_of ps p) = rpoly p x" + shows "fml_sem fml x = lookup_sem (map_fml (index_of ps) fml) ls" + using assms apply (induct fml) + by auto + +lemma index_of_aux_more: + shows "index_of_aux ls p n \ n" + apply (induct ls arbitrary: n) + apply auto + using Suc_leD by blast + +lemma index_of_aux_lookup: + assumes "p \ set ls" + shows "(index_of_aux ls p n) - n < length ls" + "ls ! ((index_of_aux ls p n) - n) = p" + using assms apply (induct ls arbitrary: n) + apply auto + apply (metis Suc_diff_Suc index_of_aux_more lessI less_Suc_eq_0_disj less_le_trans) + by (metis Suc_diff_Suc index_of_aux_more lessI less_le_trans nth_Cons_Suc) + +lemma index_of_lookup: + assumes "p \ set ls" + shows "index_of ls p < length ls" + "ls ! (index_of ls p) = p" + apply (metis assms index_of_aux_lookup(1) index_of_def minus_nat.diff_0) + by (metis assms index_of_aux_lookup(2) index_of_def minus_nat.diff_0) + +lemma convert_semantics: + shows "fml_sem fml x = lookup_sem (fst (convert fml)) (map (\p. rpoly p x) (snd (convert fml)))" + unfolding convert_def Let_def apply simp + apply (intro convert_semantics_lem) + by (simp add: index_of_lookup(1) index_of_lookup(2)) + +lemma convert_closed: + shows "\i. i \ set_fml (fst (convert fml)) \ i < length (snd (convert fml))" + unfolding convert_def Let_def + apply (auto simp add: fml.set_map) + by (simp add: index_of_lookup(1) poly_list_set_fml) + +(* Rational sign vector of polynomials qs with rational coefficients at x *) +definition sign_vec::"rat poly list \ real \ rat list" + where "sign_vec qs x \ + map (squash \ (\p. rpoly p x)) qs" + +(* The set of all rational sign vectors for qs wrt the set S + When S = UNIV, then this quantifies over all reals *) +definition consistent_sign_vectors::"rat poly list \ real set \ rat list set" + where "consistent_sign_vectors qs S = (sign_vec qs) ` S" + +lemma sign_vec_semantics: + assumes "\i. i \ set_fml fml \ i < length ls" + shows "lookup_sem fml (map (\p. rpoly p x) ls) = lookup_sem fml (sign_vec ls x)" + using assms apply (induction) + by (auto simp add: sign_vec_def squash_def) + +(* The universal and existential decision procedure is easy if we know the consistent sign vectors *) +lemma universal_lookup_sem: + assumes "\i. i \ set_fml fml \ i < length qs" + assumes "set signs = consistent_sign_vectors qs UNIV" + shows "(\x::real. lookup_sem fml (map (\p. rpoly p x) qs)) \ + list_all (lookup_sem fml) signs" + using assms(2) unfolding consistent_sign_vectors_def list_all_iff + by (simp add: assms(1) sign_vec_semantics) + +lemma existential_lookup_sem: + assumes "\i. i \ set_fml fml \ i < length qs" + assumes "set signs = consistent_sign_vectors qs UNIV" + shows "(\x::real. lookup_sem fml (map (\p. rpoly p x) qs)) \ + find (lookup_sem fml) signs \ None" + using assms(2) unfolding consistent_sign_vectors_def find_None_iff + by (simp add: assms(1) sign_vec_semantics) + +subsection "Factoring Lemmas" + (*definition real_factorize_list:: "rat poly list \ real poly list" +where "real_factorize_list qs = map (map_poly of_rat) (fst(factorize_polys qs))" +*) +interpretation of_rat_poly_hom: map_poly_comm_semiring_hom of_rat.. +interpretation of_rat_poly_hom: map_poly_comm_ring_hom of_rat.. +interpretation of_rat_poly_hom: map_poly_idom_hom of_rat.. + +lemma finite_prod_map_of_rat_poly_hom: + shows "poly (real_of_rat_poly (\(a,b)\s. f a b)) y = (\(a,b)\s. poly (real_of_rat_poly (f a b)) y)" + apply (simp add: of_rat_poly_hom.hom_prod poly_prod) + by (simp add: case_prod_app prod.case_distrib) + +lemma sign_vec_index_of: + assumes "f \ set ftrs" + shows "sign_vec ftrs x ! (index_of ftrs f) = squash (rpoly f x)" + by (simp add: assms index_of_lookup(1) index_of_lookup(2) sign_vec_def) + +lemma squash_idem: + shows "squash (squash x) = squash x" + unfolding squash_def by auto + +lemma squash_mult: + shows "squash ((a::real) * b) = squash a * squash b" + unfolding squash_def apply auto + using less_not_sym mult_neg_neg apply blast + using mult_less_0_iff by blast + +lemma squash_prod_list: + shows "squash (prod_list (ls::real list)) = prod_list (map squash ls)" + apply (induction ls) + unfolding squash_def apply auto + apply (simp add: mult_less_0_iff) + by (simp add: zero_less_mult_iff) + +lemma squash_pow: + shows "squash ((x::real) ^ (y::nat)) = (squash x) ^ y" + unfolding squash_def apply auto + by (auto simp add: zero_less_power_eq) + +lemma squash_real_of_rat[simp]: + shows "squash (real_of_rat x) = squash x" + unfolding squash_def by auto + +lemma factorize_rat_poly_monic_irreducible_monic: + assumes "factorize_rat_poly_monic f = (c,fs)" + assumes "(fi,i) \ set fs" + shows "irreducible fi \ monic fi" +proof - + obtain c' fs' where cfs: "factorize_rat_poly f = (c',fs')" + by (meson surj_pair) + then have fs: "fs = map (\(f,i). (normalize f, i)) fs'" + using factorize_rat_poly_monic_def assms by auto + obtain "fi'" where "(fi',i) \ set fs'" "fi = normalize fi'" + using assms(2) unfolding fs by auto + thus ?thesis using factorize_rat_poly irreducible_normalize_iff + by (metis cfs monic_normalize not_irreducible_zero) +qed + +lemma square_free_normalize: + assumes "square_free p" + shows "square_free (normalize p)" + by (metis assms square_free_multD(3) unit_factor_mult_normalize) + +lemma coprime_normalize: + assumes "coprime a b" + shows "coprime (normalize a) b" + using assms by auto + +lemma undo_normalize: + shows "a = Polynomial.smult (unit_factor (lead_coeff a)) (normalize a)" + by (metis add.right_neutral mult_pCons_right mult_zero_right normalize_mult_unit_factor pCons_0_hom.hom_zero unit_factor_poly_def) + +lemma finite_smult_distr: + assumes "distinct fs" + shows "(\(x,y)\set fs. Polynomial.smult ((f x y)::rat) (g x y)) = + Polynomial.smult (\(x,y)\set fs. f x y) (\(x,y)\set fs. g x y)" + using assms +proof (induction fs) + case Nil + then show ?case by auto +next + case (Cons a fs) + then show ?case apply auto + using mult.commute mult_smult_right prod.case_distrib smult_smult split_cong split_conv + by (simp add: Groups.mult_ac(2) split_beta) +qed + +lemma normalize_coprime_degree: + assumes "normalize (f::rat poly) = normalize g" + assumes "coprime f g" + shows "degree f = 0" +proof - + have "f dvd g" by (simp add: assms(1) associatedD2) + then have "f dvd 1" + using assms(2) associatedD1 by auto + thus ?thesis + using Missing_Polynomial_Factorial.is_unit_field_poly by blast +qed + +lemma factorize_rat_poly_monic_square_free_factorization: + assumes res: "factorize_rat_poly_monic f = (c,fs)" + shows "square_free_factorization f (c,fs)" +proof (unfold square_free_factorization_def split, intro conjI impI allI) + obtain c' fs' where cfs: "factorize_rat_poly f = (c',fs')" + by (meson surj_pair) + then have fs: "fs = map (\(f,i). (normalize f, i)) fs'" + using factorize_rat_poly_monic_def assms by auto + have sq: "square_free_factorization f (c',fs')" + using cfs factorize_rat_poly(1) by blast + obtain lcs where lcs: "lcs = prod_list (map (\(f,i). lead_coeff f ^ Suc i) fs')" by force + have c: "c = c' * lcs" using assms unfolding factorize_rat_poly_monic_def cfs Let_def lcs by auto + show "f = 0 \ c = 0" using c cfs by auto + show "f = 0 \ fs = []" using fs cfs by auto + have dist: "distinct fs'" using sq square_free_factorizationD(5) by blast + show dist2: "distinct fs" unfolding fs + unfolding distinct_conv_nth apply auto + proof - + fix i j + assume ij: "i < length fs'" "j < length fs'" "i \ j" + assume eq: "(case fs' ! i of + (f, x) \ (normalize f, x)) = + (case fs' ! j of + (f, x) \ (normalize f, x))" + obtain f a where fa: "fs' ! i = (f,a)" by force + obtain g where g: "fs' ! j = (g,a)" "normalize f = normalize g" + using eq fa apply auto + by (metis case_prod_conv prod.collapse prod.inject) + have "f \ g" using dist ij fa g + using nth_eq_iff_index_eq by fastforce + then have "coprime f g" + using square_free_factorizationD(3)[OF sq, of f a g a] fa g ij + apply auto + using nth_mem by force + then have "degree f = 0" + by (simp add: g(2) normalize_coprime_degree) + thus False + using fa ij(1) nth_mem sq square_free_factorizationD'(3) by fastforce + qed + have ceq: "c = c' * (\(a, i)\set fs'. (lead_coeff a) ^ Suc i)" using c lcs + by (simp add: dist prod.distinct_set_conv_list) + have fseq: " (\(a, i)\set fs. a ^ Suc i) = (\(a, i)\set fs'. (normalize a) ^ Suc i)" + apply (subst prod.distinct_set_conv_list[OF dist]) + apply (subst prod.distinct_set_conv_list[OF dist2]) + unfolding fs apply (auto simp add: o_def ) + by (metis (no_types, lifting) case_prod_conv old.prod.exhaust) + + have "f = Polynomial.smult c' (\(a, i)\set fs'. a ^ Suc i)" using sq square_free_factorizationD(1) by blast + moreover have "... = Polynomial.smult c' (\(a, i)\set fs'. (Polynomial.smult ((unit_factor (lead_coeff a))) (normalize a)) ^ Suc i)" + apply (subst undo_normalize[symmetric]) by auto + moreover have "... = Polynomial.smult c' + (\(a, i)\set fs'. (Polynomial.smult ((lead_coeff a) ^ Suc i) ((normalize a) ^ Suc i)))" + apply (subst smult_power) by auto + moreover have "... = Polynomial.smult c' + (Polynomial.smult (\(a, i)\set fs'. ((lead_coeff a) ^ Suc i)) + (\(a, i)\set fs'. (normalize a) ^ Suc i))" + apply (subst finite_smult_distr) by (auto simp add: dist) + moreover have "... = Polynomial.smult (c' * (\(a, i)\set fs'. (lead_coeff a) ^ Suc i)) + (\(a, i)\set fs'. (normalize a) ^ Suc i)" + using smult_smult by blast + moreover have "... = Polynomial.smult c (\(a, i)\set fs. a ^ Suc i)" + unfolding ceq fseq by auto + ultimately show "f = Polynomial.smult c (\(a, i)\set fs. a ^ Suc i)" by auto + fix a i + assume ai: "(a,i) \ set fs" + obtain a' where a': "(a',i) \ set fs'" "a = normalize a'" using ai unfolding fs by auto + show "square_free a" using square_free_normalize a' + using sq square_free_factorizationD(2) by blast + show "0 < degree a" using degree_normalize a' + using sq square_free_factorizationD'(3) by fastforce + fix b j + assume bj: "(b,j) \ set fs" "(a,i) \ (b,j)" + obtain b' where b': "(b',j) \ set fs'" "b = normalize b'" using bj unfolding fs by auto + show "algebraic_semidom_class.coprime a b" using a' b' apply auto + using bj(2) sq square_free_factorizationD(3) by fastforce +qed + +lemma undo_factorize_correct: + assumes "factorize_rat_poly_monic p = (c,fs)" + assumes "\f p. (f,p) \ set fs \ f \ set ftrs" + shows "undo_factorize (c,map (\(f,pow). (index_of ftrs f, pow)) fs) (sign_vec ftrs x) = squash (rpoly p x)" +proof - + have p: "p = smult c (\(a, i)\ set fs. a ^ Suc i)" + using assms(1) factorize_rat_poly_monic_square_free_factorization square_free_factorizationD(1) by blast + have fs: "distinct fs" + using assms(1) factorize_rat_poly_monic_square_free_factorization square_free_factorizationD(5) by blast + have "rpoly p x = ((real_of_rat c) * rpoly (\(a, i)\ set fs. a ^ Suc i) x)" + using p by (simp add: of_rat_hom.map_poly_hom_smult) + moreover have "... = ((real_of_rat c) * rpoly (\ai\ set fs. case ai of (a,i) \ a ^ Suc i) x)" + by blast + moreover have "... = ((real_of_rat c) * (\ai\ set fs. case ai of (a,i) \ rpoly (a ^ Suc i) x))" + by (simp add: finite_prod_map_of_rat_poly_hom) + moreover have "... = ((real_of_rat c) * (\ai\ set fs. case ai of (a,i) \ (rpoly a x) ^ Suc i))" + by (metis (mono_tags, lifting) of_rat_poly_hom.hom_power poly_hom.hom_power split_cong) + moreover have "... = ((real_of_rat c) * (prod_list (map (\ai. case ai of (a,i) \ (rpoly a x) ^ Suc i) fs)))" + by (simp add: fs prod.distinct_set_conv_list) + ultimately have "rpoly p x = ((real_of_rat c) * (prod_list (map (\ai. case ai of (a,i) \ (rpoly a x) ^ Suc i) fs)))" by auto + + then have "squash (rpoly p x) = squash c * prod_list (map squash (map (\ai. case ai of (a,i) \ (rpoly a x) ^ Suc i) fs))" + by (auto simp add: squash_mult squash_prod_list o_def) + moreover have "... = squash c * prod_list (map (\ai. case ai of (a,i) \ squash ((rpoly a x) ^ Suc i)) fs)" + apply (simp add: o_def) + by (simp add: prod.case_distrib) + ultimately have rp:"squash(rpoly p x) = squash c * prod_list (map (\ai. case ai of (a,i) \ squash (rpoly a x) ^ Suc i) fs)" + using squash_pow + by presburger + have "undo_factorize + (c, map (\(f, pow).(index_of ftrs f, pow)) fs) (sign_vec ftrs x) = + squash + (c * (\xa\fs. case xa of (f, y) \ sign_vec ftrs x ! index_of ftrs f ^ Suc y))" + unfolding undo_factorize_def apply (auto simp add: o_def) + by (metis (mono_tags, lifting) case_prod_conv old.prod.exhaust) + moreover have "... = squash + (c * (\xa\fs. case xa of (f, y) \ (squash (rpoly f x)) ^ Suc y))" + using assms(2) sign_vec_index_of + map_eq_conv split_cong + apply (auto) + by (smt map_eq_conv split_cong) + ultimately show ?thesis using rp + by (metis (mono_tags, lifting) of_rat_hom.hom_mult squash_idem squash_mult squash_real_of_rat) +qed + +lemma length_sign_vec[simp]: + shows "length (sign_vec ps x) = length ps" unfolding sign_vec_def by auto + +lemma factorize_polys_has_factors: + assumes "factorize_polys ps = (ftrs,data)" + assumes "p \ set ps" + assumes "factorize_rat_poly_monic p = (c,fs)" + shows "set (map fst fs) \ set ftrs" + using assms unfolding factorize_polys_def Let_def apply auto + by (metis UN_iff fst_conv image_eqI snd_conv) + +lemma factorize_polys_undo_factorize_polys: + assumes "factorize_polys ps = (ftrs,data)" + shows "undo_factorize_polys data (sign_vec ftrs x) = sign_vec ps x" + unfolding list_eq_iff_nth_eq undo_factorize_polys_def apply auto +proof - + show leq:"length data = length ps" + using assms unfolding factorize_polys_def by (auto simp add: Let_def) + fix i + assume il:"i < length data" + obtain c fs where cfs: "factorize_rat_poly_monic (ps ! i) = (c,fs)" + by (meson surj_pair) + then have fsts:"set (map fst fs) \ set ftrs" + using assms factorize_polys_has_factors il leq nth_mem by fastforce + have *:"data ! i = (c,map (\(f,pow). (index_of ftrs f, pow)) fs)" + using assms unfolding factorize_polys_def + using cfs il by (auto simp add: Let_def cfs) + have "undo_factorize (data ! i) (sign_vec ftrs x) = squash (rpoly (ps ! i) x)" unfolding * + apply (subst undo_factorize_correct[of "ps ! i"]) + apply (auto simp add: cfs) + using fsts by auto + thus "undo_factorize (data ! i) (sign_vec ftrs x) = sign_vec ps x ! i" + using leq il sign_vec_def by auto +qed + +lemma factorize_polys_irreducible_monic: + assumes "factorize_polys ps = (fs,data)" + shows "distinct fs" "\f. f \ set fs \ irreducible f \ monic f" + using assms unfolding factorize_polys_def Let_def apply auto + using factorize_rat_poly_monic_irreducible_monic + apply (metis prod.collapse) + using factorize_rat_poly_monic_irreducible_monic + by (metis prod.collapse) + +lemma factorize_polys_square_free: + assumes "factorize_polys ps = (fs,data)" + shows "\f. f \ set fs \ square_free f" + using assms factorize_polys_irreducible_monic(2) irreducible_imp_square_free by blast + +lemma irreducible_monic_coprime: + assumes f: "monic f" "irreducible (f::rat poly)" + assumes g: "monic g" "irreducible (g::rat poly)" + assumes "f \ g" + shows "coprime f g" + by (metis (no_types, lifting) assms(5) coprime_0(2) coprime_def' f(1) f(2) g(1) g(2) irreducible_normalized_divisors normalize_dvd_iff normalize_idem normalize_monic) + +lemma factorize_polys_coprime: + assumes "factorize_polys ps = (fs,data)" + shows "\f g. f \ set fs \ g \ set fs \ f \ g \ coprime f g" + using assms factorize_polys_irreducible_monic(2) irreducible_monic_coprime by auto + +lemma coprime_rat_poly_real_poly: + assumes "coprime p (q::rat poly)" + shows "coprime (real_of_rat_poly p) ((real_of_rat_poly q)::real poly)" + by (metis assms gcd_dvd_1 of_rat_hom.map_poly_gcd of_rat_poly_hom.hom_dvd_1) + +lemma coprime_rat_poly_iff_coprimereal_poly: + shows "coprime p (q::rat poly) \ coprime (real_of_rat_poly p) ((real_of_rat_poly q)::real poly)" +proof - + have forward: "coprime p (q::rat poly) \ coprime (real_of_rat_poly p) ((real_of_rat_poly q)::real poly)" + using coprime_rat_poly_real_poly by auto + have backward: "coprime (real_of_rat_poly p) ((real_of_rat_poly q)::real poly) \ coprime p (q::rat poly)" + proof - + assume copr_real: "comm_monoid_mult_class.coprime (real_of_rat_poly p) (real_of_rat_poly q)" + have "degree (gcd p (q::rat poly)) > 0 \ False" + proof - + assume deg: "degree (gcd p (q::rat poly)) > 0" + then have "\y. y dvd p \ y dvd q \ degree y > 0" + by blast + then obtain y where yprop: "y dvd p \ y dvd q \ degree y > 0" + by auto + then have "(real_of_rat_poly y) dvd (real_of_rat_poly p) \ + (real_of_rat_poly y ) dvd (real_of_rat_poly q) \ degree y > 0" + by simp + then show "False" + using copr_real apply (auto) + by fastforce + qed + then show "comm_monoid_mult_class.coprime p (q::rat poly)" + using comm_monoid_gcd_class.gcd_dvd_1 + by (metis Missing_Polynomial_Factorial.is_unit_field_poly copr_real gcd_zero_iff' neq0_conv of_rat_poly_hom.hom_zero) + qed + show ?thesis + using forward backward by auto +qed + +lemma factorize_polys_map_distinct: + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "distinct fss" + using factorize_polys_irreducible_monic[OF assms(1)] + unfolding assms(2) + apply (simp add: distinct_conv_nth) + by (metis of_rat_eq_iff of_rat_hom.coeff_map_poly_hom poly_eqI) + +lemma factorize_polys_map_square_free: + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "\f. f \ set fss \ square_free f" + using factorize_polys_square_free[OF assms(1)] + using assms(2) field_hom_0'.square_free_map_poly of_rat_hom.field_hom_0'_axioms by auto + +lemma factorize_polys_map_coprime: + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "\f g. f \ set fss \ g \ set fss \ f \ g \ coprime f g" + using factorize_polys_coprime[OF assms(1)] coprime_rat_poly_real_poly unfolding assms(2) + by auto + +lemma coprime_prod_list: + assumes "\p. p \ set ps \ p \ 0" + assumes "coprime (prod_list ps) (q::real poly)" + shows "\p. p \ set ps \ coprime p q" +proof - + fix p + assume "p \ set ps" + then obtain r where r: "prod_list ps = r * p" + using remove1_retains_prod by blast + show "coprime p q" + apply (rule coprime_prod[of r 1]) + using assms r apply auto + by blast +qed + +(* basically copied from square_free_factorizationD' *) +lemma factorize_polys_square_free_prod_list: + assumes "factorize_polys ps = (fs,data)" + shows "square_free (prod_list fs)" +proof (rule square_freeI) + from factorize_polys_coprime[OF assms] + have coprime: "\p q. p \ set fs \ q \ set fs \ p \ q \ coprime p q" . + from factorize_polys_square_free[OF assms] + have sq: "\p. p \ set fs \ square_free p" . + thus "prod_list fs \ 0" unfolding prod_list_zero_iff + using square_free_def by blast + fix q + assume "degree q > 0" "q * q dvd prod_list fs" + from irreducible\<^sub>d_factor[OF this(1)] this(2) obtain q where + irr: "irreducible q" and dvd: "q * q dvd prod_list fs" unfolding dvd_def by auto + hence dvd': "q dvd prod_list fs" unfolding dvd_def by auto + from irreducible_dvd_prod_list[OF irr dvd'] obtain b where + mem: "b \ set fs" and dvd1: "q dvd b" by auto + from dvd1 obtain k where b: "b = q * k" unfolding dvd_def by auto + from split_list[OF mem] b obtain bs1 bs2 where bs: "fs = bs1 @ b # bs2" by auto + from irr have q0: "q \ 0" and dq: "degree q > 0" unfolding irreducible\<^sub>d_def by auto + have "square_free (q * k)" using sq b mem by auto + from this[unfolded square_free_def, THEN conjunct2, rule_format, OF dq] + have qk: "\ q dvd k" by simp + from dvd[unfolded bs b] have "q * q dvd q * (k * prod_list (bs1 @ bs2))" + by (auto simp: ac_simps) + with q0 have "q dvd k * prod_list (bs1 @ bs2)" by auto + with irr qk have "q dvd prod_list (bs1 @ bs2)" by auto + from irreducible_dvd_prod_list[OF irr this] obtain b' where + mem': "b' \ set (bs1 @ bs2)" and dvd2: "q dvd b'" by fastforce + from dvd1 dvd2 have "q dvd gcd b b'" by auto + with dq is_unit_iff_degree[OF q0] have cop: "\ coprime b b'" by force + from mem' have "b' \ set fs" unfolding bs by auto + have b': "b' = b" using coprime + using \b' \ set fs\ cop mem by blast + with mem' bs factorize_polys_irreducible_monic(1)[OF assms] show False by auto +qed + +lemma factorize_polys_map_square_free_prod_list: + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "square_free (prod_list fss)" + using factorize_polys_square_free_prod_list[OF assms(1)] unfolding assms(2) + by (simp add: of_rat_hom.square_free_map_poly) + +lemma factorize_polys_map_coprime_pderiv: + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "\f. f \ set fss \ coprime f (pderiv (prod_list fss))" +proof - + fix f + assume f: "f \ set fss" + from factorize_polys_map_square_free[OF assms] + have sq: "\p. p \ set fss \ square_free p" . + have z: "\p. p \ set fss \ p \ 0" using sq square_free_def by blast + have c: "coprime (prod_list fss) (pderiv (prod_list fss))" + apply (simp add: separable_def[symmetric] square_free_iff_separable[symmetric]) + using factorize_polys_map_square_free_prod_list[OF assms] . + from coprime_prod_list[OF z c f] + show "coprime f (pderiv (prod_list fss))" by auto +qed + +definition pairwise_coprime_list:: "rat poly list \ bool" + where "pairwise_coprime_list qs = + (\m < length qs. \ n < length qs. + m \ n \ coprime (qs ! n) (qs ! m))" + +(* Restating factorize_polys_map_coprime to match later definitions *) +lemma coprime_factorize: + fixes qs:: "rat poly list" + shows "pairwise_coprime_list (fst(factorize_polys qs))" +proof - + let ?fs = "fst(factorize_polys qs)" + have "(\m < length ?fs. \ n < length ?fs. + m \ n \ coprime (?fs ! n) (?fs ! m))" + proof clarsimp + fix m n + assume "m < length (fst (factorize_polys qs))" + assume "n < length (fst (factorize_polys qs))" + assume "m \ n" + show " algebraic_semidom_class.coprime (fst (factorize_polys qs) ! n) + (fst (factorize_polys qs) ! m)" + by (metis \m < length (fst (factorize_polys qs))\ \m \ n\ \n < length (fst (factorize_polys qs))\ coprime_iff_coprime distinct_conv_nth factorize_polys_coprime factorize_polys_def factorize_polys_irreducible_monic(1) fstI nth_mem) + qed + then show ?thesis unfolding pairwise_coprime_list_def by auto +qed + +lemma squarefree_factorization_degree: + assumes "square_free_factorization p (c,fs)" + shows "degree p = sum_list (map (\(f,c). (c+1) * degree f) fs)" +proof - + have "p = + Polynomial.smult c + (\(a, i)\set fs. a ^ Suc i)" using assms unfolding square_free_factorization_def + by blast + then have "degree p = degree (\(a, i)\set fs. a ^ Suc i)" + using assms square_free_factorizationD(4) by fastforce + also have "... = degree (prod_list (map (\(f,c). f ^ Suc c) fs))" + by (metis assms prod.distinct_set_conv_list square_free_factorizationD(5)) + also have "... = (\(a, i)\fs. degree (a ^ Suc i))" + apply (subst degree_prod_list_eq) + apply (auto simp add: o_def) + using assms degree_0 square_free_factorizationD(2) apply blast + using assms degree_0 square_free_factorizationD(2) apply blast + by (simp add: prod.case_distrib) + ultimately show ?thesis + by (smt Polynomial.degree_power_eq add.commute assms degree_0 map_eq_conv plus_1_eq_Suc split_cong square_free_factorizationD(2)) +qed + +subsection "Auxiliary Polynomial Lemmas" +definition roots_of_coprime_r:: "real poly list \ real set" + where "roots_of_coprime_r qs = {x. poly (coprime_r qs) x = 0}" + +lemma crb_lem_pos: + fixes x:: "real" + fixes p:: "real poly" + assumes x: "poly p x = 0" + assumes p: "p \ 0" + shows "x < crb p" + using cauchy_root_bound[of p x] apply (auto) + unfolding crb_def apply (auto) + using p x + by linarith + +lemma crb_lem_neg: + fixes x:: "real" + fixes p:: "real poly" + assumes x: "poly p x = 0" + assumes p: "p \ 0" + shows "x > -crb p" + using cauchy_root_bound[of p x] apply (auto) + unfolding crb_def apply (auto) + using p x by linarith + +(* Show that the product of the polynomial list is 0 at x iff there is a polynomial + in the list that is 0 at x *) +lemma prod_zero: + shows "\x . poly (prod_list (qs:: rat poly list)) x = 0 \ (\q \ set (qs). poly q x = 0)" + apply auto + using poly_prod_list_zero_iff apply blast + using poly_prod_list_zero_iff by blast + +lemma coprime_r_zero1: "poly (coprime_r qs) (crb (prod_list qs)) = 0" + by (simp add: coprime_r_def) + +lemma coprime_r_zero2: "poly (coprime_r qs) (-crb (prod_list qs)) = 0" + by (simp add: coprime_r_def) + +lemma coprime_mult: + fixes a:: "real poly" + fixes b:: "real poly" + fixes c:: "real poly" + assumes "algebraic_semidom_class.coprime a b" + assumes "algebraic_semidom_class.coprime a c" + shows "algebraic_semidom_class.coprime a (b*c)" + using assms(1) assms(2) by auto + +(* Will be needed when we call the BKR roots on coprime_r *) +lemma coprime_r_coprime_prop: + fixes ps:: "rat poly list" + assumes "factorize_polys ps = (fs,data)" + assumes "fss = map real_of_rat_poly fs" + shows "\f. f \ set fss \ coprime f (coprime_r fss)" +proof clarsimp + fix f:: "real poly" + assume f_in: "f \ set fss" + have nonz_prod: "prod_list fss \ 0" using factorize_polys_map_square_free apply (auto) + using assms(1) assms(2) square_free_def by fastforce + have nonz_f: "f \ 0" using f_in factorize_polys_map_square_free apply (auto) + using assms(1) assms(2) square_free_def by fastforce + have copr_pderiv: "algebraic_semidom_class.coprime f (pderiv (prod_list fss))" using factorize_polys_map_coprime_pderiv + apply (auto) + using f_in assms(1) assms(2) by auto + have z_iff: "\x. poly f x = 0 \ poly (prod_list fss) x = 0" + using f_in apply (auto) + using poly_prod_list_zero_iff by blast + let ?inf_p = "[:-(crb (prod_list fss)),1:]::real poly" + have copr_inf: "algebraic_semidom_class.coprime f ([:-(crb (prod_list fss)),1:])" + proof - + have zero_prop: "\x. poly ?inf_p x = 0 \ x = crb (prod_list fss)" + by auto + have "poly (prod_list fss) (crb (prod_list fss)) \ 0" + proof - + have h: "\x. poly (prod_list fss) x = 0 \ x < (crb (prod_list fss))" + using nonz_prod crb_lem_pos[where p = "prod_list fss"] + by auto + then show ?thesis by auto + qed + then have nonzero: "poly f (crb (prod_list fss)) \ 0" + using z_iff by auto + then have "\(\x. poly f x = 0 \ poly ?inf_p x = 0)" + by simp + have is_unit_gcd: "is_unit (gcd ?inf_p f)" + using prime_elem_imp_gcd_eq prime_elem_iff_irreducible linear_irreducible_field + apply (auto) using nonzero + proof - + have f1: "\x0. - (x0::real) = - 1 * x0" + by simp + have "(1::real) \ 0" + by auto + then have "is_unit (gcd (pCons (- 1 * real_of_int (crb (prod_list fss))) 1) f)" + using f1 by (metis (no_types) is_unit_gcd nonzero one_poly_eq_simps(1) poly_eq_0_iff_dvd prime_elem_imp_coprime prime_elem_linear_field_poly) + then show "degree (gcd (pCons (- real_of_int (crb (prod_list fss))) 1) f) = 0" + by simp + qed + then show ?thesis + using is_unit_gcd + by (metis gcd.commute gcd_eq_1_imp_coprime is_unit_gcd_iff) + qed + let ?ninf_p = "[:(crb (prod_list fss)),1:]::real poly" + have copr_neg_inf: "algebraic_semidom_class.coprime f ([:(crb (prod_list fss)),1:])" + proof - + have h: "\x. poly f x = 0 \ poly (prod_list fss) x = 0" + using f_in apply (auto) + using poly_prod_list_zero_iff by blast + have zero_prop: "\x. poly ?ninf_p x = 0 \ x = -crb (prod_list fss)" + by auto + have "poly (prod_list fss) (-crb (prod_list fss)) \ 0" + proof - + have h: "\x. poly (prod_list fss) x = 0 \ x > (-crb (prod_list fss))" + using nonz_prod crb_lem_neg[where p = "prod_list fss"] + by auto + then show ?thesis by auto + qed + then have nonzero: "poly f (-crb (prod_list fss)) \ 0" + using z_iff by auto + then have "\(\x. poly f x = 0 \ poly ?ninf_p x = 0)" + using zero_prop by auto + have is_unit_gcd: "is_unit (gcd ?ninf_p f)" + using prime_elem_imp_gcd_eq prime_elem_iff_irreducible linear_irreducible_field + apply (auto) using nonzero + proof - + have f1: "(1::real) \ 0" + by auto + have "\ pCons (real_of_int (crb (prod_list fss))) 1 dvd f" + using nonzero by auto + then show "degree (gcd (pCons (real_of_int (crb (prod_list fss))) 1) f) = 0" + using f1 by (metis (no_types) Missing_Polynomial_Factorial.is_unit_field_poly coprime_imp_gcd_eq_1 is_unit_gcd_iff one_poly_eq_simps(1) prime_elem_imp_coprime prime_elem_linear_field_poly) + qed + then show ?thesis + using is_unit_gcd + by (metis gcd.commute gcd_eq_1_imp_coprime is_unit_gcd_iff) + qed + show "algebraic_semidom_class.coprime f (coprime_r fss)" + using copr_pderiv coprime_mult unfolding coprime_r_def + using copr_inf copr_neg_inf by blast +qed + +lemma coprime_r_nonzero: + fixes ps:: "rat poly list" + assumes "factorize_polys ps = (fs,data)" + assumes nonempty_fs: "fs \ []" + assumes fss_is: "fss = map real_of_rat_poly fs" + shows "(coprime_r fss) \ 0" +proof - + have nonempty_fss: "fss \ []" using nonempty_fs fss_is by auto + have deg_f: "\f \ set (fs). degree f > 0" + using factorize_polys_irreducible_monic + apply (auto) + using assms(1) irreducible_degree_field by blast + then have deg_fss: "\f \ set (fss). degree f > 0" + using fss_is by simp + then have fss_nonz: "\f \ set (fss). f \ 0" + by auto + have "fss \ [] \ ((\f \ set (fss). (degree f > 0 \ f \ 0)) \ degree (prod_list fss) > 0)" + proof (induct fss) + case Nil + then show ?case + by blast + next + case (Cons a fss) + show ?case + proof clarsimp + assume z_lt: "0 < degree a" + assume anonz: "a \ 0" + assume fnonz: "\f\set fss. 0 < degree f \ f \ 0" + have h: "degree (a * prod_list fss) = degree a + degree (prod_list fss) " + using degree_mult_eq[where p = "a", where q = "prod_list fss"] anonz fnonz + by auto + then show "0 < degree (a * prod_list fss)" + using z_lt Cons.hyps by auto + qed + qed + then have "degree (prod_list fss) > 0" + using nonempty_fss deg_fss fss_nonz by auto + then have pderiv_nonzero: "pderiv (prod_list fss) \ 0" + by (simp add: pderiv_eq_0_iff) + have "(([:-(crb (prod_list fss)),1:]) * ([:(crb (prod_list fss)),1:])) \ 0" + by auto + then show ?thesis using pderiv_nonzero + unfolding coprime_r_def apply (auto) + by (metis offset_poly_eq_0_lemma right_minus_eq synthetic_div_unique_lemma) +qed + +lemma Rolle_pderiv: + fixes q:: "real poly" + fixes x1 x2:: "real" + shows "(x1 < x2 \ poly q x1 = 0 \ poly q x2 = 0) \ (\w. x1 < w \ w < x2 \ poly (pderiv q) w = 0)" + using Rolle_deriv apply (auto) + by (metis DERIV_unique Rolle continuous_at_imp_continuous_on poly_DERIV poly_differentiable poly_isCont) + +lemma coprime_r_roots_prop: + fixes qs:: "real poly list" + assumes pairwise_rel_prime: "\q1 q2. (q1 \ q2 \ (List.member qs q1) \ (List.member qs q2))\ coprime q1 q2" + shows "\x1. \x2. ((x1 < x2 \ (\q1 \ set (qs). (poly q1 x1) = 0) \ (\q2\ set(qs). (poly q2 x2) = 0)) \ (\q. x1 < q \ q < x2 \ poly (coprime_r qs) q = 0))" +proof clarsimp + fix x1:: "real" + fix x2:: "real" + fix q1:: "real poly" + fix q2:: "real poly" + assume "x1 < x2" + assume q1_in: "q1 \ set qs" + assume q1_0: "poly q1 x1 = 0" + assume q2_in: "q2 \ set qs" + assume q2_0: "poly q2 x2 = 0" + have prod_z_x1: "poly (prod_list qs) x1 = 0" using q1_in q1_0 + using poly_prod_list_zero_iff by blast + have prod_z_x2: "poly (prod_list qs) x2 = 0" using q2_in q2_0 + using poly_prod_list_zero_iff by blast + have "\w>x1. w < x2 \ poly (pderiv (prod_list qs)) w = 0" + using Rolle_pderiv[where q = "prod_list qs"] prod_z_x1 prod_z_x2 + using \x1 < x2\ by blast + then obtain w where w_def: "w > x1 \w < x2 \ poly (pderiv (prod_list qs)) w = 0" + by auto + then have "poly (coprime_r qs) w = 0" + unfolding coprime_r_def + by simp + then show "\q>x1. q < x2 \ poly (coprime_r qs) q = 0" + using w_def by blast +qed + +subsection "Setting Up the Procedure: Lemmas" +definition has_no_zeros::"rat list \ bool" + where "has_no_zeros l = (0 \ set l)" + +lemma hnz_prop: "has_no_zeros l \ \(\k < length l. l ! k = 0)" + unfolding has_no_zeros_def + by (simp add: in_set_conv_nth) + +definition cast_rat_list:: "rat poly list \ real poly list" + where "cast_rat_list qs = map real_of_rat_poly qs" + +definition consistent_sign_vectors_r::"real poly list \ real set \ rat list set" + where "consistent_sign_vectors_r qs S = (signs_at qs) ` S" + +lemma consistent_sign_vectors_consistent_sign_vectors_r: + shows"consistent_sign_vectors_r (cast_rat_list qs) S = consistent_sign_vectors qs S" + unfolding consistent_sign_vectors_r_def cast_rat_list_def consistent_sign_vectors_def + sign_vec_def signs_at_def + by auto + +(* Relies on coprime_rat_poly_real_poly *) +lemma coprime_over_reals_coprime_over_rats: + fixes qs:: "rat poly list" + assumes csa_in: "csa \ (consistent_sign_vectors qs UNIV)" + assumes p1p2: "p1\p2 \ p1 < length csa \ p2 < length csa \ csa ! p1 = 0 \ csa ! p2 = 0" + shows "\ algebraic_semidom_class.coprime (nth qs p1) (nth qs p2) " +proof - + have isx: "\(x::real). csa = (sign_vec qs x)" + using csa_in unfolding consistent_sign_vectors_def by auto + then obtain x where havex: "csa = (sign_vec qs x)" by auto + then have expolys: "poly (real_of_rat_poly (nth qs p1)) x = 0 \ poly (real_of_rat_poly (nth qs p2)) x = 0" + using havex unfolding sign_vec_def squash_def + by (smt class_field.neg_1_not_0 length_map map_map nth_map one_neq_zero p1p2) + then have "\ coprime (real_of_rat_poly (nth qs p1)) ((real_of_rat_poly (nth qs p2))::real poly)" + apply (auto) using coprime_poly_0 + by blast + then show ?thesis + using coprime_rat_poly_real_poly by auto +qed + +(* This and the following lemma are designed to show that if you have two sgas that aren't the same, + there's a 0 in between! The proof uses IVT. It hones in on the component that's changed sign + (either from 1 to {0, -1} or from -1 to {0, 1}) *) +lemma zero_above: + fixes qs:: "rat poly list" + fixes x1:: "real" + assumes hnz: "has_no_zeros (sign_vec qs x1)" + shows "(\ x2 > x1. ((sign_vec qs x1) \ (sign_vec qs x2)) \ + (\(r::real) > x1. (r \ x2 \ (\ q \ set(qs). rpoly q r = 0))))" +proof clarsimp + fix x2 + assume x1_lt: "x1 < x2" + assume diff_sign_vec: "sign_vec qs x1 \ sign_vec qs x2" + then have "\q \ set qs. squash (rpoly q x1) \ squash (rpoly q x2)" + unfolding sign_vec_def + by simp + then obtain q where q_prop: "q \ set qs \ squash (rpoly q x1) \ squash (rpoly q x2)" + by auto + then have q_in: "q \ set qs" by auto + have poss1: "squash (rpoly q x1) = -1 \ squash (rpoly q x2) = 1 \ (\r>x1. r \ x2 \ (\q\set qs. rpoly q r = 0))" + using poly_IVT_pos[of x1 x2] using x1_lt unfolding squash_def apply (auto) + using q_prop by fastforce + have poss2: "squash (rpoly q x1) = 1 \ squash (rpoly q x2) = -1 \ (\r>x1. r \ x2 \ (\q\set qs. rpoly q r = 0))" + using poly_IVT_neg[of x1 x2] using x1_lt unfolding squash_def apply (auto) + using q_prop by fastforce + have poss3: "squash (rpoly q x2) = 0 \ (\r>x1. r \ x2 \ (\q\set qs. rpoly q r = 0))" + using x1_lt unfolding squash_def apply (auto) + using q_prop by blast + have "(q \ set qs \ rpoly q x1 = 0) \ \has_no_zeros (sign_vec qs x1)" + unfolding has_no_zeros_def sign_vec_def apply auto + by (smt image_iff squash_def) + have not_poss4: "squash (rpoly q x1) \ 0" + using hnz q_in unfolding squash_def + using \q \ set qs \ rpoly q x1 = 0 \ \ has_no_zeros (sign_vec qs x1)\ by auto + then show "\r>x1. r \ x2 \ (\q\set qs. rpoly q r = 0)" + using q_prop poss1 poss2 poss3 not_poss4 + apply (auto) + apply (meson squash_def) + apply (metis squash_def) + apply (metis squash_def) by (meson squash_def) +qed + +lemma zero_below: + fixes qs:: "rat poly list" + fixes x1:: "real" + assumes hnz: "has_no_zeros (sign_vec qs x1)" + shows "\x2 < x1. ((sign_vec qs x1) \ (sign_vec qs x2)) \ + (\(r::real) < x1. (r \ x2 \ (\ q \ set(qs). rpoly q r = 0)))" +proof clarsimp + fix x2 + assume x1_gt: "x2 < x1" + assume diff_sign_vec: "sign_vec qs x1 \ sign_vec qs x2" + then have "\q \ set qs. squash (rpoly q x1) \ squash (rpoly q x2)" + unfolding sign_vec_def + by simp + then obtain q where q_prop: "q \ set qs \ squash (rpoly q x1) \ squash (rpoly q x2)" + by auto + then have q_in: "q \ set qs" by auto + have poss1: "squash (rpoly q x1) = -1 \ squash (rpoly q x2) = 1 \ (\r x2 \ (\ q \ set(qs). rpoly q r = 0)))" + using poly_IVT_neg[of x2 x1] using x1_gt unfolding squash_def apply (auto) + using q_prop by fastforce + have poss2: "squash (rpoly q x1) = 1 \ squash (rpoly q x2) = -1 \ (\r x2 \ (\ q \ set(qs). rpoly q r = 0)))" + using poly_IVT_pos[of x2 x1] using x1_gt unfolding squash_def apply (auto) + using q_prop by fastforce + have poss3: "squash (rpoly q x2) = 0 \ (\r x2 \ (\ q \ set(qs). rpoly q r = 0)))" + using x1_gt unfolding squash_def apply (auto) + using q_prop by blast + have "(q \ set qs \ rpoly q x1 = 0) \ \has_no_zeros (sign_vec qs x1)" + unfolding has_no_zeros_def sign_vec_def apply auto + using image_iff squash_def + by (smt image_iff squash_def) + have not_poss4: "squash (rpoly q x1) \ 0" + using hnz q_in unfolding squash_def + using \q \ set qs \ rpoly q x1 = 0 \ \ has_no_zeros (sign_vec qs x1)\ by auto + then show "(\r x2 \ (\ q \ set(qs). rpoly q r = 0)))" + using q_prop poss1 poss2 poss3 not_poss4 apply (auto) + apply (meson squash_def) + apply (metis squash_def) + apply (metis squash_def) + by (meson squash_def) +qed + +lemma sorted_list_lemma: + fixes l:: "real list" + fixes a b:: "real" + fixes n:: "nat" + assumes "a < b" + assumes "(n + 1) < length l" + assumes strict_sort: "strict_sorted l" + assumes lt_a: "(l ! n) < a" + assumes b_lt: "b < (l ! (n+1))" + shows "\(\(x::real). (List.member l x \ a \ x \ x \ b))" +proof - + have sorted_hyp_var: "\q1 < length l. \q2 < length l. (q1 < q2 \ + (l ! q1) < (l ! q2))" + apply (auto) + by (metis (no_types, lifting) strict_sort sorted_wrt_iff_nth_less strict_sorted_sorted_wrt) + then have sorted_hyp_var2: "\q1 < length l. \q2 < length l. ((l ! q1) < (l ! q2)) \ q1 < q2" + using linorder_neqE_nat + by (metis Groups.add_ac(2) add_mono_thms_linordered_field(5) less_irrefl) + have "(\(x::real). (List.member l x \ a \ x \ x \ b)) \ False" + proof - + assume "(\(x::real). (List.member l x \ a \ x \ x \ b))" + then obtain x where x_prop: "List.member l x \ a \ x \ x \ b" by auto + then have l_prop: "List.member l x \ (l ! n) < x \ x < (l ! (n+1))" + using lt_a b_lt by auto + have nth_l: "l ! n < x" using l_prop by auto + have np1th_l: "x < l ! (n+1)" using l_prop by auto + have "\k. k < length l \ nth l k = x" using l_prop + by (meson in_set_member index_of_lookup(1) index_of_lookup(2)) + then obtain k where k_prop: "k < length l \ nth l k = x" by auto + have n_lt: "n < k" + using nth_l sorted_hyp_var k_prop add_lessD1 assms(2) linorder_neqE_nat nat_SN.gt_trans + by (meson sorted_hyp_var2) + have k_gt: "k < n + 1" + using sorted_hyp_var np1th_l k_prop + using assms(2) sorted_hyp_var2 by blast + show False + using n_lt k_gt by auto + qed + then show ?thesis by auto +qed + +(* This lemma shows that any zero of coprime_r is either between two roots or it's smaller than the +least root (neg inf) or it's greater than the biggest root (pos inf). *) +lemma roots_of_coprime_r_location_property: + fixes qs:: "rat poly list" + fixes sga:: "rat list" + fixes zer_list + assumes pairwise_rel_prime: "pairwise_coprime_list qs" + assumes all_squarefree: "\q. q \ set qs \ rsquarefree q" + assumes x1_prop: "sga = sign_vec qs x1" + assumes hnz: "has_no_zeros sga" + assumes zer_list_prop: "zer_list = sorted_list_of_set {(x::real). (\q \ set(qs). (rpoly q x = 0))}" + shows "zer_list \ [] \ ((x1 < (zer_list ! 0)) \ (x1 > (zer_list ! (length zer_list - 1)) \ + (\ n < (length zer_list - 1). x1 > (zer_list ! n) \ x1 < (zer_list ! (n+1)))))" +proof - + let ?zer_list = "sorted_list_of_set {(x::real). (\q \ set(qs). (rpoly q x = 0))} :: real list" + show ?thesis + proof - + have "((\q. (List.member qs q) \ q \ 0) \ has_no_zeros (sign_vec qs x1)) \ \ List.member ?zer_list x1" + proof (induct qs) + case Nil + then show ?case apply (auto) + by (simp add: member_rec(2)) + next + case (Cons a qs) + then show ?case + proof clarsimp + assume imp: "((\q. List.member qs q \ q \ 0) \ + has_no_zeros (sign_vec qs x1) \ + \ List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) + x1)" + assume nonz: "\q. List.member (a # qs) q \ q \ 0" + assume hnz: " has_no_zeros (sign_vec (a # qs) x1)" + assume mem_list: "List.member + (sorted_list_of_set {x. rpoly a x = 0 \ (\q\set qs. rpoly q x = 0)}) + x1" + have "has_no_zeros (sign_vec (a # qs) x1) \ has_no_zeros (sign_vec qs x1)" + proof - + assume hnz: "has_no_zeros (sign_vec (a # qs) x1)" + have same_vec: "sign_vec (a # qs) x1 = ((if rpoly a x1 > 0 then 1 else if rpoly a x1 = 0 then 0 else -1) # sign_vec (qs) x1)" + unfolding sign_vec_def squash_def by auto + have "has_no_zeros ((if rpoly a x1 > 0 then 1 else if rpoly a x1 = 0 then 0 else -1) # sign_vec (qs) x1) + \ has_no_zeros (sign_vec (qs) x1)" + by (simp add: has_no_zeros_def) + then show "has_no_zeros (sign_vec qs x1)" using hnz same_vec by auto + qed + then have nmem: "\ List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) x1" + using hnz nonz imp apply (auto) + by (simp add: member_rec(1)) + have "\q \set qs. q \ 0" + using nonz using in_set_member apply (auto) by fastforce + then have "\q \set qs. finite {x. rpoly q x = 0}" + by (simp add: poly_roots_finite) + then have fin_set: "finite {x. \q\set qs. rpoly q x = 0}" + by auto + have not_in: "x1 \ {x. \q\set qs. rpoly q x = 0}" using fin_set nmem set_sorted_list_of_set + all_squarefree + apply (auto) + by (simp add: List.member_def \finite {x. \q\set qs. rpoly q x = 0}\) + have x1_in: "x1 \ {x. rpoly a x = 0 \ (\q\set qs. rpoly q x = 0)}" + using mem_list sorted_list_of_set + proof - + have f1: "\r R. ((r::real) \ R \ \ List.member (sorted_list_of_set R) r) \ infinite R" + by (metis in_set_member set_sorted_list_of_set) + have "finite {r. rpoly a (r::real) = 0}" + by (metis (full_types) List.finite_set member_rec(1) nonz real_roots_of_rat_poly(1)) + then show ?thesis + using f1 \finite {x. \q\set qs. rpoly q x = 0}\ mem_list by fastforce + qed + have "rpoly a x1 \ 0" using hnz + unfolding has_no_zeros_def sign_vec_def squash_def by auto + then show "False" using not_in x1_in + by auto + qed + qed + then have non_mem: "\ List.member ?zer_list x1" + using all_squarefree unfolding rsquarefree_def hnz apply (auto) + using hnz x1_prop + by (simp add: in_set_member) + have "?zer_list \ [] \ ((x1 \ (?zer_list ! 0)) \ (x1 \ (?zer_list ! (length ?zer_list - 1)))) +\ (\ n < (length ?zer_list - 1). x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1)))" + proof - + assume "?zer_list \ []" + assume x1_asm: "(x1 \ (?zer_list ! 0)) \ (x1 \ (?zer_list ! (length ?zer_list - 1)))" + have nm1: "x1 \ ?zer_list ! 0" using non_mem + using \sorted_list_of_set {x. \q\set qs. rpoly q x = 0} \ []\ in_set_member + by (metis (no_types, lifting) in_set_conv_nth length_greater_0_conv) + have nm2: "x1 \ ?zer_list ! (length ?zer_list -1)" using non_mem + by (metis (no_types, lifting) One_nat_def \sorted_list_of_set {x. \q\set qs. rpoly q x = 0} \ []\ diff_Suc_less in_set_member length_greater_0_conv nth_mem) + then have x_asm_var: "x1 > (?zer_list ! 0) \ x1 < ?zer_list ! (length ?zer_list -1)" + using x1_asm nm1 nm2 by auto + have "(\n. (n < (length ?zer_list - 1) \ x1 \ (?zer_list ! n) \ x1 \ (?zer_list ! (n+1)))) \ False" + proof - + assume assump: "(\n. (n < (length ?zer_list - 1) \ x1 \ (?zer_list ! n) \ x1 \ (?zer_list ! (n+1))))" + have zer_case: "x1 \ ?zer_list ! 0" using x_asm_var by auto + have all_n: "\ n. (n < (length ?zer_list - 1) \ x1 \ ?zer_list ! n) " + proof - + fix n + show n_lt: "(n < (length ?zer_list - 1) \ x1 \ ?zer_list ! n)" + proof (induct n) + case 0 + then show ?case using zer_case + by blast + next + case (Suc n) + then show ?case + using assump + using Suc_eq_plus1 Suc_lessD by presburger + qed + qed + have "(length ?zer_list - 2) \ length ?zer_list -1" + using diff_le_mono2 one_le_numeral by blast + have "x1 \ ?zer_list ! (length ?zer_list - 1)" + proof - + have h1: "length ?zer_list = 1 \ x1 \ ?zer_list ! (length ?zer_list - 1)" + using assump zer_case by auto + have h2: "length ?zer_list > 1 \ x1 \ ?zer_list ! (length ?zer_list - 1)" + using all_n assump apply (auto) + by (metis (mono_tags, lifting) Suc_diff_Suc lessI) + then show ?thesis using h1 h2 apply (auto) + using zer_case by blast + qed + then show False using all_n x_asm_var + by linarith + qed + then show ?thesis + using x1_asm + by (smt One_nat_def Suc_pred \sorted_list_of_set {x. \q\set qs. rpoly q x = 0} \ []\ in_set_member length_greater_0_conv less_SucI non_mem nth_mem) + qed + then have h1: "(?zer_list \ [] \ (x1 \ (?zer_list ! 0)) \ (x1 \ (?zer_list ! (length ?zer_list - 1))) \ + (\ n < (length ?zer_list - 1). x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1))))" + by blast + then show ?thesis apply (auto) + using zer_list_prop not_less + by auto + qed +qed + +(* This lemma is essentially saying that the zeros of coprime_r capture all relevant sample points. +From roots_of_coprime_r_location_property, we know that any zero of coprime_r is either between two + roots, or it's smaller than the least root (neg inf), or it's greater than the biggest root (pos inf). +Then, since the polynomials have constant signs within those intervals, the zeros of coprime_r + capture all the relevant information. +*) +lemma roots_of_coprime_r_capture_sgas_without_zeros: + fixes qs:: "rat poly list" + fixes sga:: "rat list" + assumes pairwise_rel_prime: "pairwise_coprime_list qs" + assumes all_squarefree: "\q. q \ set qs \ rsquarefree q" + assumes ex_x1: "sga = sign_vec qs x1" + assumes hnz: "has_no_zeros sga" + shows "(\w \ (roots_of_coprime_r (cast_rat_list qs)). sga = (sign_vec qs w))" +proof - + obtain x1 where x1_prop: "sga = (sign_vec qs x1)" using ex_x1 by auto + let ?zer_list = "sorted_list_of_set {(x::real). (\q \ set(qs). (rpoly q x = 0))} :: real list" + have strict_sorted_h: "strict_sorted ?zer_list" using sorted_sorted_list_of_set + strict_sorted_iff by auto + then have sorted_hyp: "\q < length ?zer_list. (q + 1 < length ?zer_list) \ + (?zer_list ! q) < (?zer_list ! (q +1))" + apply (auto) using lessI sorted_wrt_iff_nth_less strict_sorted_sorted_wrt + by (metis (no_types, lifting) length_sorted_list_of_set strict_sorted_h) + then have sorted_hyp_var: "\q1 < length ?zer_list. \q2 < length ?zer_list. (q1 < q2 \ + (?zer_list ! q1) < (?zer_list ! q2))" + by (metis (no_types, lifting) sorted_wrt_iff_nth_less strict_sorted_h strict_sorted_sorted_wrt) + then have sorted_hyp_var2: "\q1 < length ?zer_list. ((?zer_list ! q1)::real) \ (?zer_list ! (length ?zer_list - 1))" + apply (auto) + by (smt (z3) Suc_pred diff_less gr_implies_not0 less_SucE zero_less_Suc) + have nonz_q: "\q \set qs. q \ 0" + using all_squarefree unfolding rsquarefree_def using in_set_member by auto + then have "\q \set qs. finite {x. rpoly q x = 0}" + by (simp add: poly_roots_finite) + then have fin_set: "finite {x. \q\set qs. rpoly q x = 0}" + by auto + have x1_and_roots: "?zer_list \ [] \ ((x1 < (?zer_list ! 0)) \ (x1 > (?zer_list ! (length ?zer_list - 1)) \ + (\ n < (length ?zer_list - 1). x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1)))))" + using roots_of_coprime_r_location_property x1_prop assms by auto + have x2gt: "\x2>x1. sign_vec qs x1 \ sign_vec qs x2 \ (\r>x1. r \ x2 \ (\q\set qs. rpoly q r = 0))" + using hnz x1_prop zero_above[of qs x1] by auto + have x2lt: "\x2 sign_vec qs x2 \ (\r r \ (\q\set qs. rpoly q r = 0))" + using hnz x1_prop zero_below[of qs x1] by (auto) + have triv_neg_inf_h: "?zer_list = [] \ sga = (sign_vec qs (-crb (prod_list (cast_rat_list qs))))" + proof - + assume empty_zer: "(?zer_list:: real list) = []" + let ?zer_set = "{x. \q\set qs. rpoly q x = 0}:: real set" + have fin_zer: "finite ?zer_set" using fin_set by auto + have "finite ?zer_set \ (sorted_list_of_set ?zer_set = []) = (?zer_set = {})" + using fin_zer sorted_list_of_set_eq_Nil_iff[where A = "?zer_set"] by auto + then have "(sorted_list_of_set ?zer_set = []) = (?zer_set = {})" + using fin_zer by auto + then have nozers: "?zer_set = {}" + using empty_zer by auto + then have "\(\(r::real). (\(q::rat poly)\set qs. rpoly q r = 0))" + using nozers by auto + then have "\y. sign_vec qs x1 = sign_vec qs y" + proof - + fix y + have gt_prop: "x1 > y \ sign_vec qs x1 = sign_vec qs y" + using hnz x1_prop zero_below[of qs x1] apply (auto) + using \\r. \q\set qs. rpoly q r = 0\ by blast + have lt_prop: "x1 < y \ sign_vec qs x1 = sign_vec qs y" + using zero_above[of qs x1] apply (auto) + using \\r. \q\set qs. rpoly q r = 0\ x2gt by blast + show ?thesis using gt_prop lt_prop apply (auto) + apply (metis \\r. \q\set qs. rpoly q r = 0\ linorder_neqE_linordered_idom x2gt x2lt) + using x2gt x2lt apply (auto) + apply (metis \\r. \q\set qs. rpoly q r = 0\ linorder_neqE_linordered_idom) + apply (metis \\r. \q\set qs. rpoly q r = 0\ linorder_neqE_linordered_idom) + by (metis \\r. \q\set qs. rpoly q r = 0\ linorder_neqE_linordered_idom) + qed + then show ?thesis + by (simp add: x1_prop) + qed + have neg_inf_h: "?zer_list \[] \ (x1 < (?zer_list ! 0) \ sga = (sign_vec qs (-crb (prod_list (cast_rat_list qs)))))" + proof - + let ?neg_crb = "-crb (prod_list (cast_rat_list qs))" + assume len_nontriv: "?zer_list \[]" + assume x1_lt: "x1 < ?zer_list ! 0" + have r_gt: "\r. (\q\set qs. rpoly q r = 0) \ r \ (?zer_list ! 0)" + proof clarsimp + fix q::"rat poly" + fix r:: "real" + assume q_in: "q \ set qs" + assume "rpoly q r = 0" + then have "r \ {x. \q\set qs. rpoly q x = 0}" using q_in by auto + then have "List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) r" + using in_set_member set_sorted_list_of_set fin_set apply (auto) + by (smt \r \ {x. \q\set qs. rpoly q x = 0}\ fin_set in_set_member set_sorted_list_of_set) + then show "sorted_list_of_set {x. \q\set qs. rpoly q x = 0} ! 0 \ r" + using sorted_hyp_var + by (metis (no_types, lifting) gr_implies_not0 in_set_conv_nth in_set_member not_less sorted_iff_nth_mono sorted_sorted_list_of_set) + qed + have prod_zer: "\x. (\q\set qs. rpoly q x = 0) \ (poly (prod_list (cast_rat_list qs)) x) = 0" + using prod_list_zero_iff[where xs = "(cast_rat_list qs)"] unfolding cast_rat_list_def + apply (auto) + using nonz_q apply blast + by (metis (no_types, hide_lams) image_eqI list.set_map of_rat_poly_hom.prod_list_map_hom poly_prod_list_zero_iff) + have "?zer_list \[] \ List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) + (?zer_list ! 0)" + using nth_Cons_0 apply (auto) + by (meson gr0I in_set_member length_0_conv nth_mem) + then have "?zer_list \[] \ (?zer_list ! 0) + \ {x. \q\set qs. rpoly q x = 0}" + using in_set_member[where x = "(sorted_list_of_set {x. \q\set qs. rpoly q x = 0} ! 0)", + where xs = "sorted_list_of_set {x. \q\set qs. rpoly q x = 0}"] + set_sorted_list_of_set fin_set + by blast + then have "?zer_list \[] \ (\q\set qs. rpoly q (?zer_list ! 0) = 0)" + by blast + then have poly_zer: "?zer_list \[] \ (poly (prod_list (cast_rat_list qs)) (?zer_list ! 0)) = 0" + using prod_zer by auto + have "\q. List.member (cast_rat_list qs) q \q \ 0" using nonz_q + unfolding cast_rat_list_def using in_set_member imageE image_set map_poly_zero of_rat_eq_0_iff + by smt + then have "(prod_list (cast_rat_list qs)) \ 0" + using prod_list_zero_iff in_set_member apply (auto) + by fastforce + then have crb_lt: "?zer_list \[] \ ?neg_crb < ?zer_list ! 0" + using crb_lem_neg[where p = "(prod_list (cast_rat_list qs))", where x = "sorted_list_of_set {x. \q\set qs. rpoly q x = 0} ! 0"] apply (auto) + using poly_zer + by blast + have crb_gt_x1: "?zer_list \[] \ (?neg_crb > x1 \ (sga \ (sign_vec qs ?neg_crb)) \ (\r>x1. r \ ?neg_crb \ (\q\set qs. rpoly q r = 0)))" + using x2gt apply (auto) + using crb_lt r_gt x1_prop by fastforce + have crb_lt_x1: "?neg_crb < x1 \ (sga \ (sign_vec qs ?neg_crb)) \ (\r r \ (\q\set qs. rpoly q r = 0))" + using x2lt apply (auto) + using x1_lt r_gt x1_prop by fastforce + show ?thesis using len_nontriv crb_gt_x1 crb_lt_x1 x1_lt crb_lt r_gt apply (auto) + using x1_prop by blast + qed + have pos_inf_h: "?zer_list \ [] \ (x1 > (?zer_list ! (length ?zer_list - 1)) \ sga = (sign_vec qs (crb (prod_list (cast_rat_list qs)))))" + proof - + let ?pos_crb = "crb (prod_list (cast_rat_list qs))" + assume len_nontriv: "?zer_list \[]" + assume x1_lt: "x1 > ?zer_list ! (length ?zer_list - 1)" + have r_gt: "\r. (\q\set qs. rpoly q r = 0) \ r \ (?zer_list ! (length ?zer_list - 1))" + proof - + fix r:: "real" + assume q_in: "(\q\set qs. rpoly q r = 0)" + then obtain q::"rat poly" where q_prop: "q \ set qs \ rpoly q r = 0" by auto + then have "r \ {x. \q\set qs. rpoly q x = 0}" using q_in by auto + then have "List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) r" + using in_set_member set_sorted_list_of_set fin_set by smt + then have "\n < (length ?zer_list). r = ?zer_list ! n" + by (metis (no_types, lifting) in_set_conv_nth in_set_member) + then obtain n where n_prop: "n < length ?zer_list \ r = ?zer_list ! n" + by auto + then show "r \ (?zer_list ! (length ?zer_list - 1))" + proof - + have "\q1. q1 < length ?zer_list \ (?zer_list ! q1) \ (?zer_list ! (length ?zer_list - 1:: nat))" + using sorted_hyp_var2 by auto + then have "(?zer_list ! n) \ (?zer_list ! (length ?zer_list - 1))" + using n_prop by auto + then show ?thesis using n_prop by auto + qed + qed + have prod_zer: "\x. (\q\set qs. rpoly q x = 0) \ (poly (prod_list (cast_rat_list qs)) x) = 0" + using prod_list_zero_iff[where xs = "(cast_rat_list qs)"] unfolding cast_rat_list_def + apply (auto) + using nonz_q apply blast + by (metis (no_types, hide_lams) image_eqI list.set_map of_rat_poly_hom.prod_list_map_hom poly_prod_list_zero_iff) + have "?zer_list \[] \ List.member (sorted_list_of_set {x. \q\set qs. rpoly q x = 0}) + (?zer_list ! (length ?zer_list -1))" + using nth_Cons_0 apply (auto) + by (metis (no_types, lifting) diff_less in_set_conv_nth in_set_member length_greater_0_conv length_sorted_list_of_set zero_less_Suc) + then have "?zer_list \[] \ (?zer_list ! (length ?zer_list -1)) + \ {x. \q\set qs. rpoly q x = 0}" + using in_set_member[where x = "(sorted_list_of_set {x. \q\set qs. rpoly q x = 0} ! (length ?zer_list -1))", + where xs = "sorted_list_of_set {x. \q\set qs. rpoly q x = 0}"] + set_sorted_list_of_set fin_set + by blast + then have "?zer_list \[] \ (\q\set qs. rpoly q (?zer_list ! (length ?zer_list -1)) = 0)" + by blast + then have poly_zer: "?zer_list \[] \ (poly (prod_list (cast_rat_list qs)) (?zer_list ! (length ?zer_list -1))) = 0" + using prod_zer by auto + have "\q. List.member (cast_rat_list qs) q \q \ 0" using nonz_q + unfolding cast_rat_list_def using in_set_member imageE image_set map_poly_zero of_rat_eq_0_iff + by smt + then have "(prod_list (cast_rat_list qs)) \ 0" + using prod_list_zero_iff in_set_member apply (auto) + by fastforce + then have crb_lt: "?zer_list \[] \ ?pos_crb > ?zer_list ! (length ?zer_list -1)" + using crb_lem_pos[where p = "(prod_list (cast_rat_list qs))", where x = "sorted_list_of_set {x. \q\set qs. rpoly q x = 0} ! (length ?zer_list -1)"] apply (auto) + using poly_zer + by simp + have crb_gt_x1: "?zer_list \[] \ ((?pos_crb::real) > (x1::real) \ (sga \ (sign_vec (qs::rat poly list) (?pos_crb::real))) \ (\(r::real) ?pos_crb \ (\(q::rat poly)\set qs. rpoly q r = 0)))" + using x2gt apply (auto) + using crb_lt r_gt x1_prop + using x1_lt by fastforce + have crb_lt_x1: "?pos_crb < x1 \ (sga \ (sign_vec qs ?pos_crb)) \ (\r>x1. ?pos_crb \ r \ (\q\set qs. poly (real_of_rat_poly q) r = 0))" + using x2lt apply (auto) + using x1_lt r_gt x1_prop + by (meson \prod_list (cast_rat_list qs) \ 0\ crb_lem_pos not_less prod_zer) + show ?thesis using len_nontriv crb_gt_x1 crb_lt_x1 x1_lt crb_lt r_gt apply (auto) + using x1_prop + by blast + qed + have between_h: "(\ n < (length ?zer_list - 1). x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1))) \ (\w \ (roots_of_coprime_r (cast_rat_list qs)). sga = (sign_vec qs w))" + proof - + assume "(\ n < (length ?zer_list - 1). x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1)))" + then obtain n where n_prop: "n < (length ?zer_list - 1) \ x1 > (?zer_list ! n) \ x1 < (?zer_list ! (n+1))" + by auto + have "\q1 q2. (q1 \ q2 \ (List.member (cast_rat_list qs) q1) \ (List.member (cast_rat_list qs) q2))\ coprime q1 q2" + using pairwise_rel_prime coprime_rat_poly_iff_coprimereal_poly + unfolding pairwise_coprime_list_def + by (smt cast_rat_list_def imageE image_set in_set_conv_nth in_set_member) + then have all_prop: "\x1. \x2. ((x1 < x2 \ (\q1 \ set (cast_rat_list(qs)). (poly q1 x1) = 0) \ (\q2\ set((cast_rat_list(qs))). (poly q2 x2) = 0)) \ (\q. x1 < q \ q < x2 \ poly (coprime_r (cast_rat_list qs)) q = 0))" + using coprime_r_roots_prop + by auto + have exq1: "(\q1 \ set (cast_rat_list(qs)). (poly q1 (?zer_list ! n)) = 0)" + unfolding cast_rat_list_def using n_prop apply (auto) + by (smt (verit, ccfv_SIG) One_nat_def Suc_eq_plus1 Suc_lessD fin_set length_sorted_list_of_set less_diff_conv mem_Collect_eq nth_mem set_sorted_list_of_set) + have exq2: "(\q2 \ set (cast_rat_list(qs)). (poly q2 (?zer_list ! (n+1))) = 0)" + unfolding cast_rat_list_def using n_prop One_nat_def Suc_eq_plus1 fin_set less_diff_conv mem_Collect_eq nth_mem set_sorted_list_of_set + by (smt (verit, ccfv_SIG) image_eqI set_map) + have n_prop2: "(((?zer_list ! n) < (?zer_list ! (n+1)) \ (\q1 \ set (cast_rat_list(qs)). (poly q1 (?zer_list ! n)) = 0) \ (\q2\ set((cast_rat_list(qs))). (poly q2 (?zer_list ! (n+1))) = 0)))" + using exq1 exq2 sorted_hyp n_prop by auto + then have "(\q. (?zer_list ! n) < q \ q < (?zer_list ! (n+1)) \ poly (coprime_r (cast_rat_list qs)) q = 0)" + using all_prop n_prop2 by auto + then have "\w \ (roots_of_coprime_r (cast_rat_list qs)). (?zer_list ! n) < w \ w < (?zer_list ! (n+1))" + apply (auto) + using roots_of_coprime_r_def by auto + then obtain w where w_prop: "w \ (roots_of_coprime_r (cast_rat_list qs)) \ (?zer_list ! n) < w \ w < (?zer_list ! (n+1))" by auto + have n_lt1: "n < length ?zer_list" using n_prop + using add_lessD1 less_diff_conv by blast + have n_lt2: "n + 1 < length ?zer_list" using n_prop + using less_diff_conv by blast + have sorted_hyp_var3: "?zer_list ! n < ?zer_list ! (n + 1)" using sorted_hyp + n_lt1 n_lt2 by auto + then have helper: "w > x1 \ \(\(x::real). (List.member ?zer_list x \ x1 \ x \ x \ w))" + using n_prop w_prop x1_prop strict_sorted_h sorted_list_lemma[where n = "n", where l = ?zer_list, where a = "x1", where b = "w"] sorted_hyp_var3 + by auto + have no_root1: "w > x1 \ \(\r>x1. r \ w \ (\q\set qs. rpoly q r = 0))" + proof - + assume "w > x1" + then have nex: "\(\(x::real). (List.member ?zer_list x \ x1 \ x \ x \ w))" + using helper by auto + have "(\r>x1. r \ w \ (\q\set qs. rpoly q r = 0)) \ False" + proof - + assume "(\r>x1. r \ w \ (\q\set qs. rpoly q r = 0))" + then obtain r where r_prop: "r > x1 \r \ w \(\q\set qs. rpoly q r = 0)" by auto + then have "List.member ?zer_list r \x1 \ r \x1 \ w " + by (smt fin_set in_set_member mem_Collect_eq set_sorted_list_of_set) + then show ?thesis using nex r_prop + by blast + qed + then show ?thesis by auto + qed + have helper2: "w < x1 \ \(\(x::real). (List.member ?zer_list x \ w \ x \ x \ x1))" + using n_prop w_prop x1_prop strict_sorted_h sorted_list_lemma[where n = "n", where l = ?zer_list, where a = "w", where b = "x1"] sorted_hyp_var3 + by auto + have no_root2: "w < x1 \ \(\r r \ (\q\set qs. rpoly q r = 0))" + proof - + assume "w < x1" + then have nex: "\(\(x::real). (List.member ?zer_list x \ w \ x \ x \ x1))" + using helper2 by auto + have "(\r r \ (\q\set qs. rpoly q r = 0)) \ False" + proof - + assume "(\r r\ (\q\set qs. rpoly q r = 0))" + then obtain r where r_prop: "r < x1 \ w \ r \(\q\set qs. rpoly q r = 0)" by auto + then have "List.member ?zer_list r \ w \ r \ r \ x1 " + by (smt fin_set in_set_member mem_Collect_eq set_sorted_list_of_set) + then show ?thesis using nex r_prop + by blast + qed + then show ?thesis by auto + qed + then have w_gt: "w > x1 \ (sign_vec qs x1) = (sign_vec qs w)" + using no_root1 n_prop x2gt by auto + have w_lt: "w < x1 \ (sign_vec qs x1) = (sign_vec qs w)" + using no_root2 n_prop x2lt by auto + then have "sga = (sign_vec qs w)" using w_gt w_lt x1_prop by auto + then show "(\w \ (roots_of_coprime_r (cast_rat_list qs)). sga = (sign_vec qs w))" + using w_prop by auto + qed + show ?thesis using triv_neg_inf_h neg_inf_h pos_inf_h between_h x1_and_roots + by (metis (mono_tags, lifting) coprime_r_zero1 coprime_r_zero2 mem_Collect_eq roots_of_coprime_r_def) +qed + +(* This lemma heavily relies on the main BKR_Proofs result and also the lemma named + roots_of_coprime_r_capture_sgas_without_zeros *) +lemma find_csas_lemma_nozeros: + fixes qs:: "rat poly list" + assumes fs: "factorize_polys qs = (fs,data)" + assumes "fs \ []" + shows "(csa \ (consistent_sign_vectors fs UNIV) \ has_no_zeros csa) \ + csa \ set (find_consistent_signs_at_roots (coprime_r (cast_rat_list fs)) (cast_rat_list fs))" +proof - + let ?new_l = "cast_rat_list fs" + let ?copr = "coprime_r ?new_l" + have copr_nonz: "?copr \ 0" + using coprime_r_nonzero[OF assms(1-2)] unfolding cast_rat_list_def by auto + have nontriv_list: "0 < length ?new_l" + using assms cast_rat_list_def by (auto) + have pairwise_cp: "(\q. q \ set ?new_l \ + algebraic_semidom_class.coprime ?copr q)" using coprime_r_coprime_prop[OF assms(1)] + apply (auto) + by (metis cast_rat_list_def comm_monoid_mult_class.coprime_commute coprime_iff_coprime list.set_map) + have set_fsga: "set(find_consistent_signs_at_roots ?copr ?new_l) = set(characterize_consistent_signs_at_roots ?copr ?new_l)" + using find_consistent_signs_at_roots[OF copr_nonz pairwise_cp] by auto + then have csa_in_hyp: "csa \ set (find_consistent_signs_at_roots ?copr ?new_l) + \ csa \ set(characterize_consistent_signs_at_roots ?copr ?new_l)" by auto + have forward: "(csa \ (consistent_sign_vectors fs UNIV) \ has_no_zeros csa) + \ csa \ set(characterize_consistent_signs_at_roots ?copr ?new_l)" + proof - + assume csa_in: "(csa \ (consistent_sign_vectors fs UNIV) \ has_no_zeros csa)" + have fin: "finite {x. poly (coprime_r (cast_rat_list fs)) x = 0}" + using copr_nonz poly_roots_finite + by (simp add: poly_roots_finite fs) + have pcl: "pairwise_coprime_list fs" + using coprime_factorize fs + by (metis fst_conv) + have sqf: "\q. q \ set fs \ rsquarefree q" + using factorize_polys_square_free[OF assms(1)] + by (metis square_free_rsquarefree) + obtain x1 where x1:"csa = sign_vec fs x1" + using consistent_sign_vectors_def csa_in by auto + have hnz: "has_no_zeros csa" using csa_in by auto + obtain w where w_prop: "w\roots_of_coprime_r (cast_rat_list fs)" "csa = sign_vec fs w" + using roots_of_coprime_r_capture_sgas_without_zeros[OF pcl sqf x1 hnz] + by auto + have w_root: "poly (coprime_r (cast_rat_list fs)) w = 0" + using w_prop + by (simp add: roots_of_coprime_r_def) + then have "w \ {x. poly (coprime_r (cast_rat_list fs)) x = 0}" + by auto + then have w_ins: "w \ set (characterize_root_list_p (coprime_r (cast_rat_list fs)))" + using fin set_sorted_list_of_set[where A="{x. poly (coprime_r (cast_rat_list fs)) x = 0}"] + unfolding characterize_root_list_p_def + by auto + have "map ((\x. if 0 < x then 1 else if x < 0 then - 1 else 0) \ (\p. rpoly p w)) fs = + map ((\x. if 0 < x then 1 else - 1) \ (\p. rpoly p w)) fs" + proof - + have "\(\x \ set fs. rpoly x w = 0)" + proof clarsimp + fix x + assume x_in: "x \ set fs" + assume x_zer: "rpoly x w = 0" + then have "\k < length fs. nth fs k = x" + using x_in + by (simp add: in_set_conv_nth) + then obtain k where k_prop: "k < length fs \ fs ! k = x" + by auto + then have "(sign_vec fs w) ! k = 0" using x_zer apply (auto) + unfolding sign_vec_def squash_def by auto + then have "\ (has_no_zeros (sign_vec fs w))" + apply (auto) + by (simp add: hnz_prop k_prop) + then show False using hnz unfolding sign_vec_def squash_def + using \\ has_no_zeros (sign_vec fs w)\ w_prop(2) by blast + qed + then show ?thesis using hnz unfolding sign_vec_def squash_def by auto + qed + then have "map ((\x. if 0 < x then 1 else if x < 0 then - 1 else 0) \ (\p. rpoly p w)) fs = + map (\q. if 0 < poly q w then 1 else - 1) + (cast_rat_list fs)" + unfolding cast_rat_list_def by auto + then have "csa = map (\q. if 0 < poly q w then 1 else - 1) + (cast_rat_list fs)" + by (simp add: comp_def sign_vec_def squash_def w_prop(2)) + then show ?thesis unfolding characterize_consistent_signs_at_roots_def + apply (auto) unfolding signs_at_def using w_ins w_prop + using consistent_sign_vectors_consistent_sign_vectors_r consistent_sign_vectors_def consistent_sign_vectors_r_def signs_at_def by force + qed + have backward: "csa \ set(characterize_consistent_signs_at_roots ?copr ?new_l) \ + (csa \ (consistent_sign_vectors fs UNIV) \ has_no_zeros csa)" + proof - + assume csa_in: "csa \ set(characterize_consistent_signs_at_roots ?copr ?new_l)" + have csa_in_old_set: "csa \ set (characterize_consistent_signs_at_roots_copr ?copr ?new_l)" + using csa_list_copr_rel copr_nonz csa_in find_consistent_signs_at_roots_copr pairwise_cp set_fsga by auto + have "\(x::real). \ (l::real poly list). rec_list True (\h T. If (h = 0) False) + (map (\q. if 0 < poly q x then (1::rat) else (-1::rat)) + l)" + proof clarsimp + fix x::"real" + fix l::"real poly list" + show " rec_list True (\h T. If (h = 0) False) + (map (\q. if 0 < poly q x then (1::rat) else (-1::rat)) l) " + proof (induct l) + case Nil + then show ?case + by simp + next + case (Cons a l) + then show ?case by auto + qed + qed + then have "\x. rec_list True (\h T. If (h = 0) False) + (map (\q. if 0 < poly q x then (1::rat) else - 1) + (cast_rat_list fs))" + by auto + then have hnz: "has_no_zeros csa" + using csa_in_old_set + unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def + apply (auto) unfolding has_no_zeros_def + by auto + have "\w \ set(characterize_root_list_p ?copr). csa = consistent_sign_vec_copr ?new_l w" + using csa_in_old_set using characterize_consistent_signs_at_roots_copr_def by auto + then obtain w where w_prop: "w \ set (characterize_root_list_p ?copr) \ csa = consistent_sign_vec_copr ?new_l w" + by auto + then have "poly ?copr w = 0" unfolding characterize_root_list_p_def + by (simp add: copr_nonz poly_roots_finite) + then have copr_prop: "\p \ set(?new_l). poly p w \ 0" + using w_prop coprime_r_coprime_prop apply (auto) + by (meson coprime_poly_0 in_set_member pairwise_cp) + then have "consistent_sign_vec_copr ?new_l w = sign_vec fs w" + unfolding sign_vec_def squash_def consistent_sign_vec_copr_def + cast_rat_list_def by auto + then show ?thesis using hnz w_prop apply (auto) + using consistent_sign_vectors_def by blast + qed + have "(csa \ (consistent_sign_vectors fs UNIV) \ has_no_zeros csa) + \ csa \ set(characterize_consistent_signs_at_roots ?copr ?new_l)" + using forward backward by blast + then show ?thesis using csa_in_hyp by auto +qed + +lemma range_eq: + fixes a + fixes b + shows "a = b \ range a = range b" + by simp + +lemma removeAt_distinct: + shows "distinct fss \ distinct (removeAt i fss)" + unfolding removeAt_def + by (simp add: set_take_disj_set_drop_if_distinct) + +lemma consistent_signs_atw: + assumes "\p. p \ set fs \ poly p x \ 0" + shows "consistent_sign_vec_copr fs x = signs_at fs x" + unfolding consistent_sign_vec_copr_def signs_at_def squash_def o_def + by (simp add: assms) + +(* This could be an alternate (equivalent) definition *) +lemma characterize_consistent_signs_at_roots_rw: + assumes "p \ 0" + assumes copr: "\q. q \ set fs \ coprime p q" + shows "set (characterize_consistent_signs_at_roots p fs) = + consistent_sign_vectors_r fs ({x. poly p x = 0})" + by (simp add: assms(1) characterize_consistent_signs_at_roots_def consistent_sign_vectors_r_def poly_roots_finite characterize_root_list_p_def) + +lemma signs_at_insert: + assumes "i \ length qs" + shows "insertAt i (squash (poly p x)) (signs_at qs x) = + signs_at (insertAt i p qs) x" + unfolding insertAt_def signs_at_def o_def using assms take_map apply auto + apply (subst drop_map) by simp + +lemma length_removeAt: + assumes "i < length ls" + shows "length (removeAt i ls) = length ls - 1" + unfolding removeAt_def using assms by auto + +lemma insertAt_removeAt: + assumes "i < length ls" + shows "insertAt i (ls!i) (removeAt i ls) = ls" + unfolding insertAt_def removeAt_def using assms apply auto + by (simp add: Cons_nth_drop_Suc) + +lemma insertAt_nth: + assumes "i \ length ls" + shows "insertAt i n ls ! i = n" + unfolding insertAt_def using assms + by (simp add: nth_append_take) + +lemma length_signs_at[simp]: + shows "length (signs_at qs x) = length qs" + unfolding signs_at_def by auto + +lemma find_csa_at_index: + assumes "i < length fss" + assumes "\p q. p \ set fss \ q \ set fss \ p \ q \ coprime p q" + assumes "\p. p \ set fss \ p \ 0" + assumes "distinct fss" + shows " + set (map (\v. insertAt i 0 v) (find_consistent_signs_at_roots (fss ! i) (removeAt i fss))) = + {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" +proof - + from removeAt_distinct[OF assms(4)] + have neq: "fss ! i \ 0" using assms by simp + have cop: "\q. q \ set (removeAt i fss) \ coprime (fss ! i) q" using assms + unfolding removeAt_def + apply auto + using dual_order.strict_trans find_first_unique image_iff less_imp_le_nat less_not_refl nth_image nth_mem + apply (smt atLeastLessThan_iff dual_order.strict_trans find_first_unique image_iff less_imp_le_nat less_not_refl nth_image nth_mem) + by (metis Cons_nth_drop_Suc distinct.simps(2) distinct_drop in_set_dropD nth_mem) + from find_consistent_signs_at_roots[OF neq] + have "set (find_consistent_signs_at_roots (fss ! i) (removeAt i fss)) = + set (characterize_consistent_signs_at_roots (fss ! i) (removeAt i fss))" + using cop by auto + then have eq: "set (map (insertAt i 0) + (find_consistent_signs_at_roots (fss ! i) + (removeAt i fss))) = insertAt i 0 ` set (characterize_consistent_signs_at_roots (fss ! i) (removeAt i fss)) " + by simp + from characterize_consistent_signs_at_roots_rw[OF neq cop] + have eq2: "set (characterize_consistent_signs_at_roots (fss ! i) (removeAt i fss)) = consistent_sign_vectors_r (removeAt i fss) {x. poly (fss ! i) x = 0}" . + have ile: "i \ length (removeAt i fss)" + using length_removeAt[OF assms(1)] assms(1) by linarith + have 1: "\xb. poly (fss ! i) xb = 0 \ + insertAt i 0 (signs_at (removeAt i fss) xb) \ range (signs_at fss)" + proof - + fix x + assume "poly (fss ! i) x = 0" + then have "insertAt i 0 (signs_at (removeAt i fss) x) = + insertAt i (squash (poly (fss ! i) x)) (signs_at (removeAt i fss) x)" by (auto simp add: squash_def) + also have "... = signs_at (insertAt i (fss ! i) (removeAt i fss)) x" + apply (intro signs_at_insert) + using length_removeAt[OF assms(1)] + using assms(1) by linarith + also have "... = signs_at fss x" unfolding insertAt_removeAt[OF assms(1)] by auto + ultimately have *:"insertAt i 0 (signs_at (removeAt i fss) x) = signs_at fss x" by auto + thus "insertAt i 0 (signs_at (removeAt i fss) x) \ range (signs_at fss)" by auto + qed + have 2: "\xa. signs_at fss xa ! i = 0 \ + i \ length (removeAt i fss) \ + signs_at fss xa + \ insertAt i 0 ` + signs_at (removeAt i fss) ` + {x. poly (fss ! i) x = 0}" + proof - + fix x + assume "signs_at fss x ! i = 0" + then have z:"poly (fss ! i) x = 0" unfolding signs_at_def o_def squash_def + by (smt assms(1) class_field.zero_not_one nth_map zero_neq_neg_one) + then have "insertAt i 0 (signs_at (removeAt i fss) x) = + insertAt i (squash (poly (fss ! i) x)) (signs_at (removeAt i fss) x)" by (auto simp add: squash_def) + also have "... = signs_at (insertAt i (fss ! i) (removeAt i fss)) x" + apply (intro signs_at_insert) + using length_removeAt[OF assms(1)] + using assms(1) by linarith + also have "... = signs_at fss x" unfolding insertAt_removeAt[OF assms(1)] by auto + ultimately have *:"insertAt i 0 (signs_at (removeAt i fss) x) = signs_at fss x" by auto + thus "signs_at fss x \ insertAt i 0 ` signs_at (removeAt i fss) ` {x. poly (fss ! i) x = 0}" + using z + by (metis (mono_tags, lifting) mem_Collect_eq rev_image_eqI) + qed + have "insertAt i 0 ` consistent_sign_vectors_r (removeAt i fss) {x. poly (fss ! i) x = 0} = + {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" + unfolding consistent_sign_vectors_r_def + apply (auto simp add: 1) + apply (subst insertAt_nth) + using ile by (auto simp add: 2) + thus ?thesis unfolding eq eq2 . +qed + +lemma in_set_concat_map: + assumes "i < length ls" + assumes "x \ set (f (ls ! i))" + shows "x \ set (concat (map f ls))" + using assms by auto + +lemma find_csas_lemma_onezero_gen: + fixes qs:: "rat poly list" + assumes fs: "factorize_polys qs = (fs,data)" + assumes fss: "fss = cast_rat_list fs" + shows "(csa \ (consistent_sign_vectors_r fss UNIV) \ \(has_no_zeros csa)) + \ csa \ set (find_sgas_aux fss)" +proof - + have a:"(\p q. p \ set fss \ + q \ set fss \ + p \ q \ coprime p + q)" + using cast_rat_list_def factorize_polys_map_coprime fs fss by blast + have b:"(\p. p \ set fss \ p \ 0)" + using factorize_polys_map_square_free_prod_list semidom_class.prod_list_zero_iff square_free_def + using cast_rat_list_def fs fss by blast + have c:"distinct fss" + using factorize_polys_map_distinct[OF assms(1)] assms(2) unfolding cast_rat_list_def + by auto + have forwards: "csa \ (consistent_sign_vectors_r fss UNIV) \ + \ (has_no_zeros csa) + \ csa \ set (find_sgas_aux fss)" + unfolding find_sgas_aux_def + proof - + assume csa: "csa \ (consistent_sign_vectors_r fss UNIV)" + assume hnz: "\(has_no_zeros csa)" + then obtain i where i: "i < length csa" "csa ! i = 0" unfolding hnz_prop by auto + then have cin: "csa \ {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" using csa by auto + have 1:"i < length fss" using i csa unfolding consistent_sign_vectors_r_def by auto + from find_csa_at_index[OF 1 a b c] + have eq: "set (map (\v. insertAt i 0 v) (find_consistent_signs_at_roots (fss ! i) (removeAt i fss))) = + {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" by auto + show "csa \ set (concat (map (\i. map (insertAt i 0) (find_consistent_signs_at_roots (fss ! i) (removeAt i fss))) [0.. set (find_sgas_aux fss) \ + \ (has_no_zeros csa) \ csa \ (consistent_sign_vectors_r fss UNIV)" + proof - + assume *: "csa \ set (find_sgas_aux fss)" + then obtain i where i: "i < length fss" + "csa \ set (map (\v. insertAt i 0 v) (find_consistent_signs_at_roots (fss ! i) (removeAt i fss)))" + unfolding find_sgas_aux_def + by auto + from find_csa_at_index[OF i(1) a b c] + have eq: "set (map (\v. insertAt i 0 v) (find_consistent_signs_at_roots (fss ! i) (removeAt i fss))) = + {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" by auto + have *: "csa \ {v \ consistent_sign_vectors_r fss UNIV. v ! i = 0}" using i eq by auto + then have "length csa = length fss" unfolding consistent_sign_vectors_r_def by auto + thus ?thesis unfolding has_no_zeros_def using * apply (auto simp add:in_set_conv_nth) + using i(1) by auto + qed + show ?thesis + using forwards backwards by blast +qed + +lemma mem_append: "List.member (l1@l2) m \ (List.member l1 m \ List.member l2 m)" + by (simp add: List.member_def) + +lemma same_sign_cond_rationals_reals: + fixes qs:: "rat poly list" + assumes lenh: "length (fst(factorize_polys qs)) > 0" + shows "set(find_sgas (map (map_poly of_rat) (fst(factorize_polys qs)))) = consistent_sign_vectors (fst(factorize_polys qs)) UNIV" +proof - + let ?ftrs = "(fst(factorize_polys qs))" + have pairwise_rel_prime: "pairwise_coprime_list (fst(factorize_polys qs))" + using factorize_polys_coprime + by (simp add: coprime_factorize) + have all_squarefree:"\q. (List.member (fst(factorize_polys qs)) q) \ (rsquarefree q)" + using factorize_polys_square_free + by (metis in_set_member prod.collapse square_free_rsquarefree) + have allnonzero: "\q. (List.member ?ftrs q) \ q \ 0" + using all_squarefree apply (auto) + using rsquarefree_def by blast + have h1: "\ csa. (csa \ (consistent_sign_vectors ?ftrs UNIV) \ \ (has_no_zeros csa)) + \ csa \ set (find_sgas_aux (cast_rat_list ?ftrs))" + using lenh find_csas_lemma_onezero_gen pairwise_rel_prime allnonzero + by (metis consistent_sign_vectors_consistent_sign_vectors_r eq_fst_iff) + have h2: "\csa. (csa \ (consistent_sign_vectors ?ftrs UNIV) \ has_no_zeros csa) \ + List.member (find_consistent_signs_at_roots (coprime_r (cast_rat_list ?ftrs)) (cast_rat_list ?ftrs)) csa" + using lenh find_csas_lemma_nozeros pairwise_rel_prime allnonzero + by (metis in_set_member length_greater_0_conv prod.collapse) + have h3: "\ csa. List.member (find_sgas (map (map_poly of_rat) ?ftrs)) csa \ + ((List.member (find_sgas_aux (cast_rat_list ?ftrs)) csa) \ (List.member (find_consistent_signs_at_roots (coprime_r (cast_rat_list ?ftrs)) (cast_rat_list ?ftrs)) csa))" + unfolding find_sgas_def cast_rat_list_def using mem_append + by metis + have h4: "\ csa. List.member (find_sgas (map (map_poly of_rat) ?ftrs)) csa \ + ((csa \ (consistent_sign_vectors ?ftrs UNIV) \ has_no_zeros csa) \ (csa \ (consistent_sign_vectors ?ftrs UNIV) \ \ (has_no_zeros csa)))" + using h1 h2 h3 apply (auto) apply (simp add: in_set_member) by (simp add: in_set_member) + have h5: "\csa. (csa \ (consistent_sign_vectors ?ftrs UNIV) \ has_no_zeros csa) \ (csa \ (consistent_sign_vectors ?ftrs UNIV) \ \ (has_no_zeros csa)) + \ csa \ (consistent_sign_vectors ?ftrs UNIV)" + by auto + then have "\ csa. List.member (find_sgas (map (map_poly of_rat) ?ftrs)) csa \ csa \ (consistent_sign_vectors ?ftrs UNIV)" + using h4 + by blast + then show ?thesis using in_set_member apply (auto) + apply (simp add: in_set_member) + by (simp add: in_set_member) +qed + +lemma factorize_polys_undo_factorize_polys_set: + assumes "factorize_polys ps = (ftrs,data)" + assumes "sgas = find_sgas (map (map_poly of_rat) ftrs)" + assumes sgas_set: "set (sgas) = consistent_sign_vectors ftrs UNIV" + shows "set (map (undo_factorize_polys data) sgas) = consistent_sign_vectors ps UNIV" +proof - + have h: "\x. undo_factorize_polys data (sign_vec ftrs x) = sign_vec ps x" + using factorize_polys_undo_factorize_polys + by (simp add: assms(1)) + have h1: "\x. sign_vec ps x \ set (map (undo_factorize_polys data) sgas)" + using sgas_set + by (metis UNIV_I consistent_sign_vectors_def h image_eqI image_set) + then have subset_h: "consistent_sign_vectors ps UNIV \ set (map (undo_factorize_polys data) sgas)" + unfolding consistent_sign_vectors_def by auto + have supset_h: "consistent_sign_vectors ps UNIV \ set (map (undo_factorize_polys data) sgas)" + proof - + have "\ ele. ele \ set (map (undo_factorize_polys data) sgas) \ + (\n < length sgas. ele = (undo_factorize_polys data (nth sgas n)))" + using index_of_lookup(1) index_of_lookup(2) by fastforce + then have "\ ele. ele \ set (map (undo_factorize_polys data) sgas) \ + (\x. ele = (undo_factorize_polys data (sign_vec ftrs x)))" + using sgas_set apply (auto) using consistent_sign_vectors_def by auto + then show ?thesis using h + using consistent_sign_vectors_def by auto + qed + show ?thesis using subset_h supset_h + by blast +qed + +lemma main_step_aux1: + fixes qs:: "rat poly list" + assumes empty: "(fst(factorize_polys qs)) = []" + shows "set (find_consistent_signs qs) = consistent_sign_vectors qs UNIV" +proof - + have set_eq: "set (find_consistent_signs qs) = {(map (\x. if poly x 0 < 0 then -1 else if poly x 0 = 0 then 0 else 1) qs)}" + using empty unfolding find_consistent_signs_def apply (auto) + apply (metis fst_conv) + by (metis prod.collapse) + have deg_q_prop: "fst(factorize_polys qs) = [] \ (\q \set(qs). degree q = 0)" + apply (rule ccontr) + proof clarsimp + fix q + assume *:"fst(factorize_polys qs) = []" + assume q: "q \ set qs" "0 < degree q" + obtain arb where "factorize_rat_poly_monic q = (arb,[])" + using * q unfolding factorize_polys_def apply (auto simp add:Let_def) + by (metis prod.collapse) + from squarefree_factorization_degree[OF factorize_rat_poly_monic_square_free_factorization[OF this]] + have "degree q = 0" by auto + thus False using q by auto + qed + have deg_q_cons: "(\q \set(qs). degree q = 0) \ (consistent_sign_vectors qs UNIV = {(map (\x. if poly x 0 < 0 then -1 else if poly x 0 = 0 then 0 else 1) qs)})" + proof - + assume deg_z: "(\q \set(qs). degree q = 0)" + then have "\q \set(qs). \x y. poly q x = poly q y" + apply (auto) + by (meson constant_def constant_degree) + then have csv: "consistent_sign_vectors qs UNIV = {sign_vec qs 0}" + unfolding consistent_sign_vectors_def sign_vec_def apply (auto) + apply (metis deg_z degree_0_id of_rat_hom.map_poly_hom_coeff_lift poly_0_coeff_0 poly_const_conv squash_real_of_rat) + by (metis (mono_tags, lifting) class_semiring.add.one_closed comp_def image_iff list.map_cong0 of_rat_hom.poly_map_poly_0) + have "sign_vec qs 0 = (map (\x. if poly x 0 < 0 then -1 else if poly x 0 = 0 then 0 else 1) qs)" + unfolding sign_vec_def squash_def by auto + then show "consistent_sign_vectors qs UNIV = {(map (\x. if poly x 0 < 0 then -1 else if poly x 0 = 0 then 0 else 1) qs)}" + using csv by auto + qed + then show ?thesis + using deg_q_prop deg_q_cons set_eq + by (simp add: empty) +qed + +lemma main_step_aux2: + fixes qs:: "rat poly list" + assumes lenh: "length (fst(factorize_polys qs)) > 0" + shows "set (find_consistent_signs qs) = consistent_sign_vectors qs UNIV" +proof - + let ?fs = "fst(factorize_polys qs)" + let ?data = "snd(factorize_polys qs)" + let ?sgas = "find_sgas (map (map_poly of_rat) ?fs)" + have h0: "set (?sgas) = consistent_sign_vectors ?fs UNIV" + using lenh same_sign_cond_rationals_reals coprime_factorize by auto + then have h1: "set (map (undo_factorize_polys ?data) ?sgas) = consistent_sign_vectors qs UNIV" + using factorize_polys_undo_factorize_polys_set + by simp + then show ?thesis using lenh apply (auto) + apply (smt case_prod_unfold find_consistent_signs_def h1 main_step_aux1) + by (simp add: find_consistent_signs_def prod.case_eq_if) +qed + +lemma main_step: + fixes qs:: "rat poly list" + shows "set (find_consistent_signs qs) = consistent_sign_vectors qs UNIV" + using main_step_aux1 main_step_aux2 by auto + +subsection "Decision Procedure: Main Lemmas" + +lemma decide_univ_lem_helper: + assumes "(fml_struct,polys) = convert fml" + shows "(\x::real. lookup_sem fml_struct (map (\p. rpoly p x) polys)) \ + (decide_universal fml)" + using universal_lookup_sem main_step unfolding decide_universal_def apply (auto) + apply (metis assms convert_closed fst_conv snd_conv) + by (metis (full_types) assms convert_closed fst_conv snd_conv) + +lemma decide_exis_lem_helper: + assumes "(fml_struct,polys) = convert fml" + shows "(\x::real. lookup_sem fml_struct (map (\p. rpoly p x) polys)) \ + (decide_existential fml)" + using existential_lookup_sem main_step unfolding decide_existential_def apply (auto) + apply (metis assms convert_closed fst_conv snd_conv) + by (metis (full_types) assms convert_closed fst_conv snd_conv) + +theorem decision_procedure: + shows "(\x::real. fml_sem fml x) \ (decide_universal fml)" + "\x::real. fml_sem fml x \ (decide_existential fml)" + using convert_semantics_lem decide_univ_lem_helper apply (auto) + apply (simp add: convert_semantics) + apply (metis convert_def convert_semantics fst_conv snd_conv) + using convert_semantics_lem + by (metis convert_def convert_semantics decide_exis_lem_helper fst_conv snd_conv) + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/BKR_Proofs.thy b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/BKR_Proofs.thy @@ -0,0 +1,2223 @@ +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 + case (Cons a ls) + then show ?case + apply (auto simp add: nth_append) + using div_if mod_geq by auto +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 + 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) + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy b/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/Matrix_Equation_Construction.thy @@ -0,0 +1,881 @@ +theory Matrix_Equation_Construction + +imports "BKR_Algorithm" +begin + +section "Results with Sturm's Theorem" + +lemma relprime: + fixes q::"real poly" + assumes "coprime p q" + assumes "p \ 0" + assumes "q \ 0" + shows "changes_R_smods p (pderiv p) = card {x. poly p x = 0 \ poly q x > 0} + card {x. poly p x = 0 \ poly q x < 0}" +proof - + have 1: "{x. poly p x = 0 \ poly q x = 0} = {}" + using assms(1) coprime_poly_0 by auto + have 2: "changes_R_smods p (pderiv p) = int (card {x . poly p x = 0})" using sturm_R by auto + have 3: "{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} = {}" by auto + have "{x . poly p x = 0} = {x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0} \ {x. poly p x = 0 \ poly q x = 0}" by force + then have "{x . poly p x = 0} = {x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0}" using 1 by auto + then have "(card {x . poly p x = 0}) = (card ({x. poly p x = 0 \ poly q x > 0} \{x. poly p x = 0 \ poly q x < 0}))" by presburger + then have 4: "(card {x . poly p x = 0}) = card {x. poly p x = 0 \ poly q x > 0} + card {x. poly p x = 0 \ poly q x < 0}" using 3 by (simp add: card_Un_disjoint assms(2) poly_roots_finite) + show ?thesis by (simp add: "2" "4") +qed + +(* This is the same proof as card_eq_sum *) +lemma card_eq_const_sum: + fixes k:: real + assumes "finite A" + shows "k*card A = sum (\x. k) A" +proof - + have "plus \ (\_. Suc 0) = (\_. Suc)" + by (simp add: fun_eq_iff) + then have "Finite_Set.fold (plus \ (\_. Suc 0)) = Finite_Set.fold (\_. Suc)" + by (rule arg_cong) + then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" + by (blast intro: fun_cong) + then show ?thesis + by (simp add: card.eq_fold sum.eq_fold) +qed + +lemma restate_tarski: + fixes q::"real poly" + assumes "coprime p q" + assumes "p \ 0" + assumes "q \ 0" + shows "changes_R_smods p ((pderiv p) * q) = card {x. poly p x = 0 \ poly q x > 0} - int(card {x. poly p x = 0 \ poly q x < 0})" +proof - + have 3: "taq {x. poly p x=0} q \ \y\{x. poly p x=0}. sign (poly q y)" by (simp add: taq_def) + have 4: "{x. poly p x=0} = {x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} \ {x. poly p x = 0 \ poly q x = 0}" by force + then have 5: "{x. poly p x=0} = {x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}" using assms(1) coprime_poly_0 by auto + then have 6: "\y\{x. poly p x=0}. sign (poly q y) \ \y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. sign (poly q y)" by presburger + then have 12: "taq {x. poly p x=0} q \ \y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. sign (poly q y)" using 3 by linarith + have 7: "{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0} = {}" by auto + then have 8: "\y\{x. poly p x = 0 \ poly q x > 0} \ {x. poly p x = 0 \ poly q x < 0}. sign (poly q y) \ (\y\{x. poly p x = 0 \ poly q x > 0}.sign (poly q y)) + (\y\{x. poly p x = 0 \ poly q x < 0}.sign(poly q y))" by (simp add: assms(2) poly_roots_finite sum.union_disjoint) + then have 13: "taq {x. poly p x=0} q \ (\y\{x. poly p x = 0 \ poly q x > 0}.sign (poly q y)) + (\y\{x. poly p x = 0 \ poly q x < 0}.sign(poly q y))" using 12 by linarith + then have 9: "taq {x. poly p x = 0} q \ (\y\{x. poly p x = 0 \ poly q x > 0}.1) + (\y\{x. poly p x = 0 \ poly q x < 0}.(-1))" by simp + have 10: "(\y\{x. poly p x = 0 \ poly q x > 0}.1) = card {x. poly p x = 0 \ poly q x > 0}" using card_eq_sum by auto + have 11: " (\y\{x. poly p x = 0 \ poly q x < 0}.(-1)) = -1*card {x. poly p x = 0 \ poly q x < 0}" using card_eq_const_sum by simp + have 14: "taq {x. poly p x = 0} q \ card {x. poly p x = 0 \ poly q x > 0} + -1*card {x. poly p x = 0 \ poly q x < 0}" using 9 10 11 by simp + have 1: "changes_R_smods p (pderiv p * q) = taq {x. poly p x=0} q" using sturm_tarski_R by simp + then have 15: "changes_R_smods p (pderiv p * q) = card {x. poly p x = 0 \ poly q x > 0} + (-1*card {x. poly p x = 0 \ poly q x < 0})" using 14 by linarith + have 16: "(-1*card {x. poly p x = 0 \ poly q x < 0}) = - card {x. poly p x = 0 \ poly q x < 0}" by auto + then show ?thesis using 15 by linarith +qed + +lemma restate_tarski2: + fixes q::"real poly" + assumes "p \ 0" + shows "changes_R_smods p ((pderiv p) * q) = + int(card {x. poly p x = 0 \ poly q x > 0}) - + int(card {x. poly p x = 0 \ poly q x < 0})" + unfolding sturm_tarski_R[symmetric] taq_def +proof - + let ?all = "{x. poly p x=0}" + let ?lt = "{x. poly p x=0 \ poly q x < 0}" + let ?gt = "{x. poly p x=0 \ poly q x > 0}" + let ?eq = "{x. poly p x=0 \ poly q x = 0}" + have eq: "?all = ?lt \ ?gt \ ?eq" by force + from poly_roots_finite[OF assms] have fin: "finite ?all" . + show "(\x | poly p x = 0. sign (poly q x)) = int (card ?gt) - int (card ?lt)" + unfolding eq + apply (subst sum_Un) + apply (auto simp add:fin) + apply (subst sum_Un) + by (auto simp add:fin) +qed + +lemma coprime_set_prod: + fixes I:: "real poly set" + shows "finite I \ ((\ q \ I. (coprime p q)) \ (coprime p (\ I)))" +proof (induct rule: finite_induct) + case empty + then show ?case + by simp +next + case (insert x F) + then show ?case using coprime_mult_right_iff + by simp +qed + +lemma finite_nonzero_set_prod: + fixes I:: "real poly set" + shows nonzero_hyp: "finite I \ ((\ q \ I. q \ 0) \ \ I \ 0)" +proof (induct rule: finite_induct) + case empty + then show ?case + by simp +next + case (insert x F) + have h: "\ (insert x F) = x * (\ F)" + by (simp add: insert.hyps(1) insert.hyps(2)) + have h_xin: "x \ insert x F" + by simp + have hq: "(\ q \ (insert x F). q \ 0) \ x \ 0" using h_xin + by blast + show ?case using h hq + using insert.hyps(3) by auto +qed + +section "Setting up the construction: Definitions" + +definition characterize_root_list_p:: "real poly \ real list" + where "characterize_root_list_p p \ sorted_list_of_set({x. poly p x = 0}::real set)" + +(************** Renegar's N(I); towards defining the RHS of the matrix equation **************) + +lemma construct_NofI_prop: + fixes p:: "real poly" + fixes I:: "real poly list" + assumes nonzero: "p\0" + shows "construct_NofI p I = + rat_of_int (int (card {x. poly p x = 0 \ poly (prod_list I) x > 0}) - + int (card {x. poly p x = 0 \ poly (prod_list I) x < 0}))" + unfolding construct_NofI_def + using assms restate_tarski2 nonzero rsquarefree_def + by (simp add: rsquarefree_def) + +definition construct_s_vector:: "real poly \ real poly list list \ rat vec" + where "construct_s_vector p Is = vec_of_list (map (\ I.(construct_NofI p I)) Is)" + +(* Consistent sign assignments *) +definition squash::"'a::linordered_field \ rat" + where "squash x = (if x > 0 then 1 + else if x < 0 then -1 + else 0)" + +definition signs_at::"real poly list \ real \ rat list" + where "signs_at qs x \ + map (squash \ (\q. poly q x)) qs" + +definition characterize_consistent_signs_at_roots:: "real poly \ real poly list \ rat list list" + where "characterize_consistent_signs_at_roots p qs = + (remdups (map (signs_at qs) (characterize_root_list_p p)))" + +(* An alternate version designed to be used when every polynomial in qs is relatively prime to p*) +definition consistent_sign_vec_copr::"real poly list \ real \ rat list" + where "consistent_sign_vec_copr qs x \ + map (\ q. if (poly q x > 0) then (1::rat) else (-1::rat)) qs" + +definition characterize_consistent_signs_at_roots_copr:: "real poly \ real poly list \ rat list list" + where "characterize_consistent_signs_at_roots_copr p qss = + (remdups (map (consistent_sign_vec_copr qss) (characterize_root_list_p p)))" + +lemma csa_list_copr_rel: + fixes p:: "real poly" + fixes qs:: "real poly list" + assumes nonzero: "p\0" + assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" + shows "characterize_consistent_signs_at_roots p qs = characterize_consistent_signs_at_roots_copr p qs" +proof - + have "\q \ set(qs). \ x \ set (characterize_root_list_p p). poly q x \ 0" + using pairwise_rel_prime + using coprime_poly_0 in_set_member nonzero poly_roots_finite characterize_root_list_p_def by fastforce + then have h: "\q \ set(qs). \ x \ set (characterize_root_list_p p). squash (poly q x) = (if (poly q x > 0) then (1::rat) else (-1::rat))" + by (simp add: squash_def) + have "map (\r. map (\p. if 0 < poly p r then 1 else - 1) qs) (characterize_root_list_p p) = map (\r. map (squash \ (\p. poly p r)) qs) (characterize_root_list_p p)" + by (simp add: h) + thus ?thesis unfolding characterize_consistent_signs_at_roots_def characterize_consistent_signs_at_roots_copr_def + signs_at_def consistent_sign_vec_copr_def + by presburger +qed + +(************** Towards defining Renegar's polynomial function and the LHS of the matrix equation **************) + +definition list_constr:: "nat list \ nat \ bool" + where "list_constr L n \ list_all (\x. x < n) L" + +definition all_list_constr:: "nat list list \ nat \ bool" + where "all_list_constr L n \ (\x. List.member L x \ list_constr x n)" + +(* The first input is the subset; the second input is the consistent sign assignment. + We want to map over the first list and pull out all of the elements in the second list with + corresponding positions, and then multiply those together. +*) +definition z:: "nat list \ rat list \ rat" + where "z index_list sign_asg \ (prod_list (map (nth sign_asg) index_list))" + +definition mtx_row:: "rat list list \ nat list \ rat list" + where "mtx_row sign_list index_list \ (map ( (z index_list)) sign_list)" + +definition matrix_A:: "rat list list \ nat list list \ rat mat" + where "matrix_A sign_list subset_list = + (mat_of_rows_list (length sign_list) (map (\i .(mtx_row sign_list i)) subset_list))" + +definition alt_matrix_A:: "rat list list \ nat list list \ rat mat" + where "alt_matrix_A signs subsets = (mat (length subsets) (length signs) + (\(i, j). z (subsets ! i) (signs ! j)))" + +lemma alt_matrix_char: "alt_matrix_A signs subsets = matrix_A signs subsets" +proof - + have h0: "(\i j. i < length subsets \ + j < length signs \ + map (\index_list. map (z index_list) signs) subsets ! i ! j = z (subsets ! i) (signs ! j))" + proof - + fix i + fix j + assume i_lt: "i < length subsets" + assume j_lt: "j < length signs" + show "((map (\index_list. map (z index_list) signs) subsets) ! i) ! j = z (subsets ! i) (signs ! j)" + proof - + have h0: "(map (\index_list. map (z index_list) signs) subsets) ! i = map (z (subsets ! i)) signs" + using nth_map i_lt + by blast + then show ?thesis using nth_map j_lt + by simp + qed + qed + have h: " mat (length subsets) (length signs) (\(i, j). z (subsets ! i) (signs ! j)) = + mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z index_list) signs) subsets ! i ! y)" + using h0 eq_matI[where A = "mat (length subsets) (length signs) (\(i, j). z (subsets ! i) (signs ! j))", + where B = "mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z index_list) signs) subsets ! i ! y)"] + by auto + show ?thesis unfolding alt_matrix_A_def matrix_A_def mat_of_rows_list_def apply (auto) unfolding mtx_row_def + using h by blast +qed + +lemma subsets_are_rows: "\i < (length subsets). row (alt_matrix_A signs subsets) i = vec (length signs) (\j. z (subsets ! i) (signs ! j))" + unfolding row_def unfolding alt_matrix_A_def by auto + +lemma signs_are_cols: "\i < (length signs). col (alt_matrix_A signs subsets) i = vec (length subsets) (\j. z (subsets ! j) (signs ! i))" + unfolding col_def unfolding alt_matrix_A_def by auto + +(* ith entry of LHS vector is the number of (distinct) real zeros of p where the sign vector of the qs is the ith entry of signs.*) +definition construct_lhs_vector:: "real poly \ real poly list \ rat list list \ rat vec" + where "construct_lhs_vector p qs signs \ + vec_of_list (map (\w. rat_of_int (int (length (filter (\v. v = w) (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))))) signs)" + +(* Putting all of the pieces of the construction together *) +definition satisfy_equation:: "real poly \ real poly list \ nat list list \ rat list list \ bool" + where "satisfy_equation p qs subset_list sign_list = + (mult_mat_vec (matrix_A sign_list subset_list) (construct_lhs_vector p qs sign_list) = (construct_rhs_vector p qs subset_list))" + +section "Setting up the construction: Proofs" + +(* Some matrix lemmas *) +lemma row_mat_of_rows_list: + assumes "list_all (\r. length r = nc) rs" + assumes "i < length rs" + shows "row (mat_of_rows_list nc rs) i = vec_of_list (nth rs i)" + by (smt assms(1) assms(2) dim_col_mat(1) dim_vec_of_list eq_vecI index_row(2) index_vec list_all_length mat_of_rows_list_def row_mat split_conv vec_of_list_index) + + +lemma mult_mat_vec_of_list: + assumes "length ls = nc" + assumes "list_all (\r. length r = nc) rs" + shows "mat_of_rows_list nc rs *\<^sub>v vec_of_list ls = + vec_of_list (map (\r. vec_of_list r \ vec_of_list ls) rs)" + unfolding mult_mat_vec_def + using row_mat_of_rows_list assms + apply auto + by (smt dim_row_mat(1) dim_vec dim_vec_of_list eq_vecI index_map_vec(1) index_map_vec(2) index_vec list_all_length mat_of_rows_list_def row_mat_of_rows_list vec_of_list_index) + +lemma mtx_row_length: + "list_all (\r. length r = length signs) (map (mtx_row signs) ls)" + apply (induction ls) + by (auto simp add: mtx_row_def) + +thm construct_lhs_vector_def +thm poly_roots_finite + +(* Recharacterize the LHS vector *) +lemma construct_lhs_vector_clean: + assumes "p \ 0" + assumes "i < length signs" + shows "(construct_lhs_vector p qs signs) $ i = + card {x. poly p x = 0 \ ((consistent_sign_vec_copr qs x) = (nth signs i))}" +proof - + from poly_roots_finite[OF assms(1)] have "finite {x. poly p x = 0}" . + then have eq: "(Collect + ((\v. v = signs ! i) \ + consistent_sign_vec_copr qs) \ + set (sorted_list_of_set + {x. poly p x = 0})) = + {x. poly p x = 0 \ consistent_sign_vec_copr qs x = signs ! i}" + by auto + show ?thesis + unfolding construct_lhs_vector_def vec_of_list_index characterize_root_list_p_def + apply auto + apply (subst nth_map[OF assms(2)]) + apply auto + apply (subst distinct_length_filter) + using eq by auto +qed + +lemma construct_lhs_vector_cleaner: + assumes "p \ 0" + shows "(construct_lhs_vector p qs signs) = + vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ ((consistent_sign_vec_copr qs x) = s)})) signs)" + apply (rule eq_vecI) + apply (auto simp add: construct_lhs_vector_clean[OF assms] ) + apply (simp add: vec_of_list_index) + unfolding construct_lhs_vector_def + using assms construct_lhs_vector_clean construct_lhs_vector_def apply auto[1] + by simp + +(* Show that because our consistent sign vectors consist of 1 and -1's, z returns 1 or -1 + when applied to a consistent sign vector *) +lemma z_signs: + assumes "list_all (\i. i < length signs) I" + assumes "list_all (\s. s = 1 \ s = -1) signs" + shows "(z I signs = 1) \ (z I signs = -1)" using assms +proof (induction I) + case Nil + then show ?case + by (auto simp add:z_def) +next + case (Cons a I) + moreover have "signs ! a = 1 \ signs ! a = -1" + by (metis (mono_tags, lifting) add_Suc_right calculation(2) calculation(3) gr0_conv_Suc list.size(4) list_all_length nth_Cons_0) + ultimately show ?case + by (auto simp add:z_def) +qed + +lemma z_lemma: + fixes I:: "nat list" + fixes sign:: "rat list" + assumes consistent: "sign \ set (characterize_consistent_signs_at_roots_copr p qs)" + assumes welldefined: "list_constr I (length qs)" + shows "(z I sign = 1) \ (z I sign = -1)" +proof (rule z_signs) + have "length sign = length qs" using consistent + by (auto simp add: characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) + thus "list_all (\i. i < length sign) I" + using welldefined + by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) + show "list_all (\s. s = 1 \ s = - 1) sign" using consistent + apply (auto simp add: list.pred_map characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def) + using Ball_set + by force +qed + +(* Show that all consistent sign vectors on roots of polynomials are in characterize_consistent_signs_at_roots_copr *) +lemma in_set: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec_copr qs x" + assumes welldefined: "list_constr I (length qs)" + shows "sign \ set (characterize_consistent_signs_at_roots_copr p qs)" +proof - + have h1: "consistent_sign_vec_copr qs x \ + set (remdups (map (consistent_sign_vec_copr qs) (sorted_list_of_set {x. poly p x = 0})))" + using root_p apply auto apply (subst set_sorted_list_of_set) + using nonzero poly_roots_finite rsquarefree_def apply blast by auto + thus ?thesis unfolding characterize_consistent_signs_at_roots_copr_def characterize_root_list_p_def using sign_fix + by blast +qed + +(* Since all of the polynomials in qs are relatively prime to p, products of subsets of these + polynomials are also relatively prime to p *) +lemma nonzero_product: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + fixes I:: "nat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes welldefined: "list_constr I (length qs)" + shows "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (poly (prod_list (retrieve_polys qs I)) x < 0)" +proof - + have "\x. x \ set (retrieve_polys qs I) \ coprime p x" + unfolding retrieve_polys_def + by (smt in_set_conv_nth in_set_member length_map list_all_length list_constr_def nth_map pairwise_rel_prime_1 welldefined) + then have coprimeh: "coprime p (prod_list (retrieve_polys qs I))" + using prod_list_coprime_right by auto + thus ?thesis using root_p + using coprime_poly_0 linorder_neqE_linordered_idom by blast +qed + +(* The next few lemmas relate z to the signs of the product of subsets of polynomials of qs *) +lemma horiz_vector_helper_pos_ind: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec_copr qs x" + shows "(list_constr I (length qs)) \ (poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" +proof (induct I) + case Nil + then show ?case + by (simp add: retrieve_polys_def z_def) +next + case (Cons a I) + have welldef: "list_constr (a#I) (length qs) \ (list_constr I (length qs))" + unfolding list_constr_def list_all_def by auto + have set_hyp: "list_constr I (length qs) \ sign \ set (characterize_consistent_signs_at_roots_copr p qs)" + using in_set using nonzero root_p sign_fix by blast + have z_hyp: "list_constr I (length qs) \ ((z I sign = 1) \ (z I sign = -1))" + using set_hyp z_lemma[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast + have sign_hyp: "sign = map (\ q. if (poly q x > 0) then 1 else -1) qs" + using sign_fix unfolding consistent_sign_vec_copr_def by blast + have ind_hyp_1: "list_constr (a#I) (length qs) \ + ((0 < poly (prod_list (retrieve_polys qs I)) x) = (z I sign = 1))" + using welldef Cons.hyps by auto + have ind_hyp_2: "list_constr (a#I) (length qs) \ + ((0 > poly (prod_list (retrieve_polys qs I)) x) = (z I sign = -1))" + using welldef z_hyp Cons.hyps nonzero_product + using pairwise_rel_prime_1 nonzero root_p by auto + have h1: "prod_list (retrieve_polys qs (a # I)) = (nth qs a)*(prod_list (retrieve_polys qs I))" + by (simp add: retrieve_polys_def) + have h2: "(z (a # I) sign) = (nth sign a)*(z I sign)" + by (metis (mono_tags, hide_lams) list.simps(9) prod_list.Cons z_def) + have h3help: "list_constr (a#I) (length qs) \ a < length qs" unfolding list_constr_def + by simp + then have h3: "list_constr (a#I) (length qs) \ + ((nth sign a) = (if (poly (nth qs a) x > 0) then 1 else -1))" + using nth_map sign_hyp by auto + have h2: "(0 < poly ((nth qs a)*(prod_list (retrieve_polys qs I))) x) \ + ((0 < poly (nth qs a) x \ (0 < poly (prod_list (retrieve_polys qs I)) x)) \ + (0 > poly (nth qs a) x \ (0 > poly (prod_list (retrieve_polys qs I)) x)))" + by (simp add: zero_less_mult_iff) + have final_hyp_a: "list_constr (a#I) (length qs) \ (((0 < poly (nth qs a) x \ (0 < poly (prod_list (retrieve_polys qs I)) x)) + \ (0 > poly (nth qs a) x \ (0 > poly (prod_list (retrieve_polys qs I)) x))) = + ((nth sign a)*(z I sign) = 1))" + proof - + have extra_hyp_a: "list_constr (a#I) (length qs) \ (0 < poly (nth qs a) x = ((nth sign a) = 1))" using h3 + by simp + have extra_hyp_b: "list_constr (a#I) (length qs) \ (0 > poly (nth qs a) x = ((nth sign a) = -1))" + using h3 apply (auto) using coprime_poly_0 h3help in_set_member nth_mem pairwise_rel_prime_1 root_p by fastforce + have ind_hyp_1: "list_constr (a#I) (length qs) \ (((0 < poly (nth qs a) x \ (z I sign = 1)) \ + (0 > poly (nth qs a) x \ (z I sign = -1))) + = ((nth sign a)*(z I sign) = 1))" using extra_hyp_a extra_hyp_b + using zmult_eq_1_iff + by (simp add: h3) + then show ?thesis + using ind_hyp_1 ind_hyp_2 by (simp add: Cons.hyps welldef) + qed + then show ?case + using h1 z_def by (simp add: zero_less_mult_iff) +qed + +lemma horiz_vector_helper_pos: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec_copr qs x" + assumes welldefined: "list_constr I (length qs)" + shows "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" + using horiz_vector_helper_pos_ind + using pairwise_rel_prime_1 nonzero root_p sign_fix welldefined by blast + +lemma horiz_vector_helper_neg: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec_copr qs x" + assumes welldefined: "list_constr I (length qs)" + shows "(poly (prod_list (retrieve_polys qs I)) x < 0) \ (z I sign = -1)" +proof - + have set_hyp: "list_constr I (length qs) \ sign \ set (characterize_consistent_signs_at_roots_copr p qs)" + using in_set using nonzero root_p sign_fix by blast + have z_hyp: "list_constr I (length qs) \ ((z I sign = 1) \ (z I sign = -1))" + using set_hyp z_lemma[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast + have poly_hyp: "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (poly (prod_list (retrieve_polys qs I)) x < 0)" + using nonzero_product + using pairwise_rel_prime_1 nonzero root_p + using welldefined by blast + have pos_hyp: "(poly (prod_list (retrieve_polys qs I)) x > 0) \ (z I sign = 1)" using horiz_vector_helper_pos + using pairwise_rel_prime_1 nonzero root_p sign_fix welldefined by blast + show ?thesis using z_hyp poly_hyp pos_hyp apply (auto) + using welldefined by blast +qed + +(* Recharacterize the dot product *) +lemma vec_of_list_dot_rewrite: + assumes "length xs = length ys" + shows "vec_of_list xs \ vec_of_list ys = + sum_list (map2 (*) xs ys)" + using assms +proof (induction xs arbitrary:ys) + case Nil + then show ?case by auto +next + case (Cons a xs) + then show ?case apply auto + by (smt (verit, best) Suc_length_conv list.simps(9) old.prod.case scalar_prod_vCons sum_list.Cons vec_of_list_Cons zip_Cons_Cons) +qed + +lemma lhs_dot_rewrite: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + shows + "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = + sum_list (map (\s. (z I s) * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs)" +proof - + have "p \ 0" using nonzero by auto + from construct_lhs_vector_cleaner[OF this] + have rhseq: "construct_lhs_vector p qs signs = + vec_of_list + (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs)" by auto + have "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = + sum_list (map2 (*) (mtx_row signs I) (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) signs))" + unfolding rhseq + apply (intro vec_of_list_dot_rewrite) + by (auto simp add: mtx_row_def) + thus ?thesis unfolding mtx_row_def + using map2_map_map + by (auto simp add: map2_map_map) +qed + +lemma sum_list_distinct_filter: + fixes f:: "'a \ int" + assumes "distinct xs" "distinct ys" + assumes "set ys \ set xs" + assumes "\x. x \ set xs - set ys \ f x = 0" + shows "sum_list (map f xs) = sum_list (map f ys)" + by (metis List.finite_set assms(1) assms(2) assms(3) assms(4) sum.mono_neutral_cong_left sum_list_distinct_conv_sum_set) + +(* If we have a superset of the signs, we can drop to just the consistent ones *) +lemma construct_lhs_vector_drop_consistent: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" + assumes welldefined: "list_constr I (length qs)" + shows + "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = + (vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ + (construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)))" +proof - + have h0: "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_nat (card + {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ z I sgn = 0" + proof - + have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (card + {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn} \ {}" + by fastforce + then show ?thesis + proof - + { fix iis :: "rat list" + have ff1: "0 \ p" + using nonzero rsquarefree_def by blast + obtain rr :: "(real \ bool) \ real" where + ff2: "\p. p (rr p) \ Collect p = {}" + by moura + { assume "\is. is = iis \ {r. poly p r = 0 \ consistent_sign_vec_copr qs r = is} \ {}" + then have "\is. consistent_sign_vec_copr qs (rr (\r. poly p r = 0 \ consistent_sign_vec_copr qs r = is)) = iis \ {r. poly p r = 0 \ consistent_sign_vec_copr qs r = is} \ {}" + using ff2 + by (metis (mono_tags, lifting)) + then have "\r. poly p r = 0 \ consistent_sign_vec_copr qs r = iis" + using ff2 by smt + then have "iis \ consistent_sign_vec_copr qs ` set (sorted_list_of_set {r. poly p r = 0})" + using ff1 poly_roots_finite by fastforce } + then have "iis \ set signs \ iis \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ \ 0 < rat_of_int (int (card {r. poly p r = 0 \ consistent_sign_vec_copr qs r = iis}))" + by (metis (no_types) \\sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (int (card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn})) \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn} \ {}\ characterize_root_list_p_def) } + then show ?thesis + by fastforce + qed + qed + then have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ ((0 = rat_of_nat (card + {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = sgn}) \ z I sgn = 0))" + by auto + then have hyp: "\ s. s \ set signs \ s \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ (z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = 0)" + by auto + then have "(\s\ set(signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + proof - + have "set(signs) =(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p)))" + by blast + then have sum_rewrite: "(\s\ set(signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + by auto + then have sum_split: "(\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + = +(\s\ (set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) ++ (\s\ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + by (metis (no_types, lifting) List.finite_set sum.Int_Diff) + have sum_zero: "(\s\ (set(signs)-(consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = 0" + using hyp + by (simp add: hyp) + show ?thesis using sum_rewrite sum_split sum_zero by linarith + qed + then have set_eq: "set (remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p))) = set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))" + using all_info + by (simp add: characterize_consistent_signs_at_roots_copr_def subset_antisym) + have hyp1: "(\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\set (signs). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + using distinct_signs sum_list_distinct_conv_sum_set by blast + have hyp2: "(\s\remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + = (\s\ set (remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + using sum_list_distinct_conv_sum_set by blast + have set_sum_eq: "(\s\(set (signs) \ (consistent_sign_vec_copr qs ` set (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\ set (remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p))). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + using set_eq by auto + then have "(\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + using set_sum_eq hyp1 hyp2 + using \(\s\set signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set signs \ consistent_sign_vec_copr qs ` set (characterize_root_list_p p). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ by linarith + then have "consistent_sign_vec_copr qs ` set (characterize_root_list_p p) \ set signs \ + (\p qss. + characterize_consistent_signs_at_roots_copr p qss = + remdups (map (consistent_sign_vec_copr qss) (characterize_root_list_p p))) \ + (\s\signs. z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + (\s\remdups + (map (consistent_sign_vec_copr qs) + (characterize_root_list_p p)). z I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + by linarith + then show ?thesis unfolding lhs_dot_rewrite[OF nonzero] + apply (auto intro!: sum_list_distinct_filter simp add: distinct_signs characterize_consistent_signs_at_roots_copr_def) + using all_info characterize_consistent_signs_at_roots_copr_def by auto[1] +qed + +(* Both matrix_equation_helper_step and matrix_equation_main_step relate the matrix construction + to the Tarski queries, i.e. relate the product of a row of the matrix and the LHS vector to a + Tarski query on the RHS *) +lemma matrix_equation_helper_step: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" + assumes welldefined: "list_constr I (length qs)" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = + rat_of_int (card {x. poly p x = 0 \ poly (prod_list (retrieve_polys qs I)) x > 0}) - + rat_of_int (card {x. poly p x = 0 \ poly (prod_list (retrieve_polys qs I)) x < 0})" +proof - + have "finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))" by auto + let ?gt = "(set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = 1})" + let ?lt = " (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = -1})" + have eq: "set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) = ?gt \ ?lt" + apply auto + by (metis characterize_root_list_p_def horiz_vector_helper_neg horiz_vector_helper_pos_ind nonzero nonzero_product pairwise_rel_prime_1 poly_roots_finite sorted_list_of_set(1) welldefined) + (* First, drop the signs that are irrelevant *) + from construct_lhs_vector_drop_consistent[OF assms(1-4)] have + "vec_of_list (mtx_row signs I) \ construct_lhs_vector p qs signs = + vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ + construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)" . + (* Now we split the sum *) + from lhs_dot_rewrite[OF assms(1)] + moreover have "... = + (\s\characterize_consistent_signs_at_roots_copr p qs. + z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" . + moreover have "... = + (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). + z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" unfolding characterize_consistent_signs_at_roots_copr_def sum_code[symmetric] + by (auto) + ultimately have "... = + (\s\?gt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + + (\s\?lt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))" + apply (subst eq) + apply (rule sum.union_disjoint) + by auto + (* Now recharacterize lt, gt*) + have setroots: "set (characterize_root_list_p p) = {x. poly p x = 0}" unfolding characterize_root_list_p_def + using poly_roots_finite nonzero rsquarefree_def set_sorted_list_of_set by blast + have *: "\s. {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s} = + {x \{x. poly p x = 0}. consistent_sign_vec_copr qs x = s}" + by auto + have lem_e1: "\x. x \ {x. poly p x = 0} \ + card + {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. + consistent_sign_vec_copr qs x = s} = + (if 0 < poly (prod_list (retrieve_polys qs I)) x then 1 else 0)" + proof - + fix x + assume rt: "x \ {x. poly p x = 0}" + then have 1: "{s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. consistent_sign_vec_copr qs x = s} = + {s. z I s = 1 \ consistent_sign_vec_copr qs x = s}" + by auto + from horiz_vector_helper_pos[OF assms(1) assms(5) rt] + have 2: "... = {s. (0 < poly (prod_list (retrieve_polys qs I)) x) \ consistent_sign_vec_copr qs x = s}" + using welldefined by blast + have 3: "... = (if (0 < poly (prod_list (retrieve_polys qs I)) x) then {consistent_sign_vec_copr qs x} else {})" + by auto + thus "card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. consistent_sign_vec_copr qs x = s} = + (if 0 < poly (prod_list (retrieve_polys qs I)) x then 1 else 0) " using 1 2 3 by auto + qed + have e1: "(\s\consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = 1}. + card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = + (sum (\x. if (poly (prod_list (retrieve_polys qs I)) x) > 0 then 1 else 0) {x. poly p x = 0})" + unfolding * apply (rule sum_multicount_gen) + using \finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))\ setroots apply auto[1] + apply (metis List.finite_set setroots) + using lem_e1 by auto + have gtchr: "(\s\?gt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + rat_of_int (card {x. poly p x = 0 \ 0 < poly (prod_list (retrieve_polys qs I)) x})" + apply (auto simp add: setroots) + apply (subst of_nat_sum[symmetric]) + apply (subst of_nat_eq_iff) + apply (subst e1) + apply (subst card_eq_sum) + apply (rule sum.mono_neutral_cong_right) + apply (metis List.finite_set setroots) + by auto + have lem_e2: "\x. x \ {x. poly p x = 0} \ + card + {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. + consistent_sign_vec_copr qs x = s} = + (if poly (prod_list (retrieve_polys qs I)) x < 0 then 1 else 0)" + proof - + fix x + assume rt: "x \ {x. poly p x = 0}" + then have 1: "{s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. consistent_sign_vec_copr qs x = s} = + {s. z I s = -1 \ consistent_sign_vec_copr qs x = s}" + by auto + from horiz_vector_helper_neg[OF assms(1) assms(5) rt] + have 2: "... = {s. (0 > poly (prod_list (retrieve_polys qs I)) x) \ consistent_sign_vec_copr qs x = s}" + using welldefined by blast + have 3: "... = (if (0 > poly (prod_list (retrieve_polys qs I)) x) then {consistent_sign_vec_copr qs x} else {})" + by auto + thus "card {s \ consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = -1}. consistent_sign_vec_copr qs x = s} = + (if poly (prod_list (retrieve_polys qs I)) x < 0 then 1 else 0)" using 1 2 3 by auto + qed + have e2: " (\s\consistent_sign_vec_copr qs ` {x. poly p x = 0} \ {s. z I s = - 1}. + card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) = + (sum (\x. if (poly (prod_list (retrieve_polys qs I)) x) < 0 then 1 else 0) {x. poly p x = 0})" + unfolding * apply (rule sum_multicount_gen) + using \finite (set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)))\ setroots apply auto[1] + apply (metis List.finite_set setroots) + using lem_e2 by auto + have ltchr: "(\s\?lt. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = + - rat_of_int (card {x. poly p x = 0 \ 0 > poly (prod_list (retrieve_polys qs I)) x})" + apply (auto simp add: setroots sum_negf) + apply (subst of_nat_sum[symmetric]) + apply (subst of_nat_eq_iff) + apply (subst e2) + apply (subst card_eq_sum) + apply (rule sum.mono_neutral_cong_right) + apply (metis List.finite_set setroots) + by auto + show ?thesis using gtchr ltchr + using \(\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = 1}. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) + (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)) \ {s. z I s = - 1}. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \(\s\characterize_consistent_signs_at_roots_copr p qs. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s})) = (\s\set (map (consistent_sign_vec_copr qs) (characterize_root_list_p p)). z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs) = (\s\characterize_consistent_signs_at_roots_copr p qs. z I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}))\ \vec_of_list (mtx_row signs I) \ construct_lhs_vector p qs signs = vec_of_list (mtx_row (characterize_consistent_signs_at_roots_copr p qs) I) \ construct_lhs_vector p qs (characterize_consistent_signs_at_roots_copr p qs)\ + by linarith +qed + +(* A clean restatement of the helper lemma *) +lemma matrix_equation_main_step: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" + assumes welldefined: "list_constr I (length qs)" + assumes pairwise_rel_prime_1: "\q. ((List.member qs q) \ (coprime p q))" + shows "(vec_of_list (mtx_row signs I) \ (construct_lhs_vector p qs signs)) = + construct_NofI p (retrieve_polys qs I)" + unfolding construct_NofI_prop[OF nonzero] + using matrix_equation_helper_step[OF assms] + by linarith + +lemma map_vec_vec_of_list_eq_intro: + assumes "map f xs = map g ys" + shows "map_vec f (vec_of_list xs) = map_vec g (vec_of_list ys)" + by (metis assms vec_of_list_map) + +(* Shows that as long as we have a "basis" of sign assignments (see assumptions all_info, welldefined), + and some other mild assumptions on our inputs (given in nonzero, distinct_signs, pairwise_rel_prime), + the construction will be satisfied *) +theorem matrix_equation: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes subsets:: "nat list list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" + assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" + assumes welldefined: "all_list_constr (subsets) (length qs)" + shows "satisfy_equation p qs subsets signs" + unfolding satisfy_equation_def matrix_A_def + construct_lhs_vector_def construct_rhs_vector_def all_list_constr_def + apply (subst mult_mat_vec_of_list) + apply (auto simp add: mtx_row_length intro!: map_vec_vec_of_list_eq_intro) + using matrix_equation_main_step[OF assms(1-3) _ assms(4), unfolded construct_lhs_vector_def] + using all_list_constr_def in_set_member welldefined by fastforce + +(* Prettifying some theorems*) +definition roots:: "real poly \ real set" + where "roots p = {x. poly p x = 0}" + +definition sgn::"'a::linordered_field \ rat" + where "sgn x = (if x > 0 then 1 + else if x < 0 then -1 + else 0)" + +definition sgn_vec::"real poly list \ real \ rat list" + where "sgn_vec qs x \ map (sgn \ (\q. poly q x)) qs" + +definition consistent_signs_at_roots:: "real poly \ real poly list \ rat list set" + where "consistent_signs_at_roots p qs = + (sgn_vec qs) ` (roots p)" + +lemma consistent_signs_at_roots_eq: + assumes "p \ 0" + shows "consistent_signs_at_roots p qs = + set (characterize_consistent_signs_at_roots p qs)" + unfolding consistent_signs_at_roots_def characterize_consistent_signs_at_roots_def + characterize_root_list_p_def + apply auto + apply (subst set_sorted_list_of_set) + using assms poly_roots_finite apply blast + unfolding sgn_vec_def sgn_def signs_at_def squash_def o_def + using roots_def apply auto[1] + by (smt Collect_cong assms image_iff poly_roots_finite roots_def sorted_list_of_set(1)) + +abbreviation w_vec:: "real poly \ real poly list \ rat list list \ rat vec" + where "w_vec \ construct_lhs_vector" + +abbreviation v_vec:: "real poly \ real poly list \ nat list list \ rat vec" + where "v_vec \ construct_rhs_vector" + +abbreviation M_mat:: "rat list list \ nat list list \ rat mat" + where "M_mat \ matrix_A" + +theorem matrix_equation_pretty: + assumes "p\0" + assumes "\q. q \ set qs \ coprime p q" + assumes "distinct signs" + assumes "consistent_signs_at_roots p qs \ set signs" + assumes "\l i. l \ set subsets \ i \ set l \ i < length qs" + shows "M_mat signs subsets *\<^sub>v w_vec p qs signs = v_vec p qs subsets" + unfolding satisfy_equation_def[symmetric] + apply (rule matrix_equation[OF assms(1) assms(3)]) + apply (metis assms(1) assms(2) assms(4) consistent_signs_at_roots_eq csa_list_copr_rel member_def) + apply (simp add: assms(2) in_set_member) + using Ball_set all_list_constr_def assms(5) list_constr_def member_def by fastforce + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/More_Matrix.thy b/thys/BenOr_Kozen_Reif/More_Matrix.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/More_Matrix.thy @@ -0,0 +1,2014 @@ +theory More_Matrix + imports "Jordan_Normal_Form.Matrix" + "Jordan_Normal_Form.DL_Rank" + "Jordan_Normal_Form.VS_Connect" + "Jordan_Normal_Form.Gauss_Jordan_Elimination" +begin + +section "Kronecker Product" + +definition kronecker_product :: "'a :: ring mat \ 'a mat \ 'a mat" where + "kronecker_product A B = + (let ra = dim_row A; ca = dim_col A; + rb = dim_row B; cb = dim_col B + in + mat (ra*rb) (ca*cb) + (\(i,j). + A $$ (i div rb, j div cb) * + B $$ (i mod rb, j mod cb) + ))" + +lemma arith: + assumes "d < a" + assumes "c < b" + shows "b*d+c < a*(b::nat)" +proof - + have "b*d+c < b*(d+1)" + by (simp add: assms(2)) + thus ?thesis + by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right assms(1) less_le_trans mult.commute mult_le_cancel2) +qed + +lemma dim_kronecker[simp]: + "dim_row (kronecker_product A B) = dim_row A * dim_row B" + "dim_col (kronecker_product A B) = dim_col A * dim_col B" + unfolding kronecker_product_def Let_def by auto + +lemma kronecker_inverse_index: + assumes "r < dim_row A" "s < dim_col A" + assumes "v < dim_row B" "w < dim_col B" + shows "kronecker_product A B $$ (dim_row B*r+v, dim_col B*s+w) = A $$ (r,s) * B $$ (v,w)" +proof - + from arith[OF assms(1) assms(3)] + have "dim_row B*r+v < dim_row A * dim_row B" . + moreover from arith[OF assms(2) assms(4)] + have "dim_col B * s + w < dim_col A * dim_col B" . + ultimately show ?thesis + unfolding kronecker_product_def Let_def + using assms by auto +qed + +lemma kronecker_distr_left: + assumes "dim_row B = dim_row C" "dim_col B = dim_col C" + shows "kronecker_product A (B+C) = kronecker_product A B + kronecker_product A C" + unfolding kronecker_product_def Let_def + using assms apply (auto simp add: mat_eq_iff) + by (metis (no_types, lifting) distrib_left index_add_mat(1) mod_less_divisor mult_eq_0_iff neq0_conv not_less_zero) + +lemma kronecker_distr_right: + assumes "dim_row B = dim_row C" "dim_col B = dim_col C" + shows "kronecker_product (B+C) A = kronecker_product B A + kronecker_product C A" + unfolding kronecker_product_def Let_def + using assms by (auto simp add: mat_eq_iff less_mult_imp_div_less distrib_right) + +lemma index_mat_mod[simp]: "nr > 0 & nc > 0 \ mat nr nc f $$ (i mod nr,j mod nc) = f (i mod nr,j mod nc)" + by auto + +lemma kronecker_assoc: + shows "kronecker_product A (kronecker_product B C) = kronecker_product (kronecker_product A B) C" + unfolding kronecker_product_def Let_def + apply (case_tac "dim_row B * dim_row C > 0 & dim_col B * dim_col C > 0") + apply (auto simp add: mat_eq_iff less_mult_imp_div_less) + by (smt div_mult2_eq div_mult_mod_eq kronecker_inverse_index less_mult_imp_div_less linordered_semiring_strict_class.mult_pos_pos mod_less_divisor mod_mult2_eq mult.assoc mult.commute) + +lemma sum_sum_mod_div: + "(\ia = 0::nat..ja = 0..ia = 0..ia. (ia div y, ia mod y)) {0.. {0.. (\ia. (ia div y, ia mod y)) ` {0.. {0.. (\ia. (ia div y, ia mod y)) ` {0..ia. (ia div y, ia mod y)) ` {0.. {0.. {0.. {0..ia. (ia div y, ia mod y)) ` {0..ia = 0::nat..ja = 0..(x, y)\{0.. {0..ia. (ia div y, ia mod y)"]) + using 1 2 by auto +qed + +(* Kronecker product distributes over matrix multiplication *) +lemma kronecker_of_mult: + assumes "dim_col (A :: 'a :: comm_ring mat) = dim_row C" + assumes "dim_col B = dim_row D" + shows "kronecker_product A B * kronecker_product C D = kronecker_product (A * C) (B * D)" + unfolding kronecker_product_def Let_def mat_eq_iff +proof clarsimp + fix i j + assume ij: "i < dim_row A * dim_row B" "j < dim_col C * dim_col D" + have 1: "(A * C) $$ (i div dim_row B, j div dim_col D) = + row A (i div dim_row B) \ col C (j div dim_col D)" + using ij less_mult_imp_div_less by (auto intro!: index_mult_mat) + have 2: "(B * D) $$ (i mod dim_row B, j mod dim_col D) = + row B (i mod dim_row B) \ col D (j mod dim_col D)" + using ij apply (auto intro!: index_mult_mat) + using gr_implies_not0 apply fastforce + using gr_implies_not0 by fastforce + have 3: "\x. x < dim_row C * dim_row D \ + A $$ (i div dim_row B, x div dim_row D) * + B $$ (i mod dim_row B, x mod dim_row D) * + (C $$ (x div dim_row D, j div dim_col D) * + D $$ (x mod dim_row D, j mod dim_col D)) = + row A (i div dim_row B) $ (x div dim_row D) * + col C (j div dim_col D) $ (x div dim_row D) * + (row B (i mod dim_row B) $ (x mod dim_row D) * + col D (j mod dim_col D) $ (x mod dim_row D))" + proof - + fix x + assume *:"x < dim_row C * dim_row D" + have 1: "row A (i div dim_row B) $ (x div dim_row D) = A $$ (i div dim_row B, x div dim_row D)" + by (simp add: * assms(1) less_mult_imp_div_less row_def) + have 2: "row B (i mod dim_row B) $ (x mod dim_row D) = B $$ (i mod dim_row B, x mod dim_row D)" + by (metis "*" assms(2) ij(1) index_row(1) mod_less_divisor nat_0_less_mult_iff neq0_conv not_less_zero) + have 3: "col C (j div dim_col D) $ (x div dim_row D) = C $$ (x div dim_row D, j div dim_col D)" + by (simp add: "*" ij(2) less_mult_imp_div_less) + have 4: "col D (j mod dim_col D) $ (x mod dim_row D) = D $$ (x mod dim_row D, j mod dim_col D)" + by (metis "*" Euclidean_Division.div_eq_0_iff gr_implies_not0 ij(2) index_col mod_div_trivial mult_not_zero) + show "A $$ (i div dim_row B, x div dim_row D) * + B $$ (i mod dim_row B, x mod dim_row D) * + (C $$ (x div dim_row D, j div dim_col D) * + D $$ (x mod dim_row D, j mod dim_col D)) = + row A (i div dim_row B) $ (x div dim_row D) * + col C (j div dim_col D) $ (x div dim_row D) * + (row B (i mod dim_row B) $ (x mod dim_row D) * + col D (j mod dim_col D) $ (x mod dim_row D))" unfolding 1 2 3 4 + by (simp add: mult.assoc mult.left_commute) + qed + have *: "(A * C) $$ (i div dim_row B, j div dim_col D) * + (B * D) $$ (i mod dim_row B, j mod dim_col D) = + (\ia = 0..j. A $$ (i div dim_row B, j div dim_col B) * + B $$ (i mod dim_row B, j mod dim_col B)) \ + vec (dim_row C * dim_row D) + (\i. C $$ (i div dim_row D, j div dim_col D) * + D $$ (i mod dim_row D, j mod dim_col D)) = + (A * C) $$ (i div dim_row B, j div dim_col D) * + (B * D) $$ (i mod dim_row B, j mod dim_col D)" + unfolding * scalar_prod_def + by (auto simp add: assms sum_product sum_sum_mod_div intro!: sum.cong) +qed + +lemma inverts_mat_length: + assumes "square_mat A" "inverts_mat A B" "inverts_mat B A" + shows "dim_row B = dim_row A" "dim_col B = dim_col A" + apply (metis assms(1) assms(3) index_mult_mat(3) index_one_mat(3) inverts_mat_def square_mat.simps) + by (metis assms(1) assms(2) index_mult_mat(3) index_one_mat(3) inverts_mat_def square_mat.simps) + +lemma less_mult_imp_mod_less: + "m mod i < i" if "m < n * i" for m n i :: nat + using gr_implies_not_zero that by fastforce + +lemma kronecker_one: + shows "kronecker_product ((1\<^sub>m x)::'a :: ring_1 mat) (1\<^sub>m y) = 1\<^sub>m (x*y)" + unfolding kronecker_product_def Let_def + apply (auto simp add:mat_eq_iff less_mult_imp_div_less less_mult_imp_mod_less) + by (metis div_mult_mod_eq) + +lemma kronecker_invertible: + assumes "invertible_mat (A :: 'a :: comm_ring_1 mat)" "invertible_mat B" + shows "invertible_mat (kronecker_product A B)" +proof - + obtain Ai where Ai: "inverts_mat A Ai" "inverts_mat Ai A" using assms invertible_mat_def by blast + obtain Bi where Bi: "inverts_mat B Bi" "inverts_mat Bi B" using assms invertible_mat_def by blast + have "square_mat (kronecker_product A B)" + by (metis (no_types, lifting) assms(1) assms(2) dim_col_mat(1) dim_row_mat(1) invertible_mat_def kronecker_product_def square_mat.simps) + moreover have "inverts_mat (kronecker_product A B) (kronecker_product Ai Bi)" + using Ai Bi unfolding inverts_mat_def + by (metis (no_types, lifting) dim_kronecker(1) index_mult_mat(3) index_one_mat(3) kronecker_of_mult kronecker_one) + moreover have "inverts_mat (kronecker_product Ai Bi) (kronecker_product A B)" + using Ai Bi unfolding inverts_mat_def + by (metis (no_types, lifting) dim_kronecker(1) index_mult_mat(3) index_one_mat(3) kronecker_of_mult kronecker_one) + ultimately show ?thesis unfolding invertible_mat_def by blast +qed + +section "More DL Rank" + +(* conjugate matrices *) +instantiation mat :: (conjugate) conjugate +begin + +definition conjugate_mat :: "'a :: conjugate mat \ 'a mat" + where "conjugate m = mat (dim_row m) (dim_col m) (\(i,j). conjugate (m $$ (i,j)))" + +lemma dim_row_conjugate[simp]: "dim_row (conjugate m) = dim_row m" + unfolding conjugate_mat_def by auto + +lemma dim_col_conjugate[simp]: "dim_col (conjugate m) = dim_col m" + unfolding conjugate_mat_def by auto + +lemma carrier_vec_conjugate[simp]: "m \ carrier_mat nr nc \ conjugate m \ carrier_mat nr nc" + by (auto) + +lemma mat_index_conjugate[simp]: + shows "i < dim_row m \ j < dim_col m \ conjugate m $$ (i,j) = conjugate (m $$ (i,j))" + unfolding conjugate_mat_def by auto + +lemma row_conjugate[simp]: "i < dim_row m \ row (conjugate m) i = conjugate (row m i)" + by (auto) + +lemma col_conjugate[simp]: "i < dim_col m \ col (conjugate m) i = conjugate (col m i)" + by (auto) + +lemma rows_conjugate: "rows (conjugate m) = map conjugate (rows m)" + by (simp add: list_eq_iff_nth_eq) + +lemma cols_conjugate: "cols (conjugate m) = map conjugate (cols m)" + by (simp add: list_eq_iff_nth_eq) + +instance +proof + fix a b :: "'a mat" + show "conjugate (conjugate a) = a" + unfolding mat_eq_iff by auto + let ?a = "conjugate a" + let ?b = "conjugate b" + show "conjugate a = conjugate b \ a = b" + by (metis dim_col_conjugate dim_row_conjugate mat_index_conjugate conjugate_cancel_iff mat_eq_iff) +qed + +end + +abbreviation conjugate_transpose :: "'a::conjugate mat \ 'a mat" + where "conjugate_transpose A \ conjugate (A\<^sup>T)" + +notation conjugate_transpose ("(_\<^sup>H)" [1000]) + +lemma transpose_conjugate: + shows "(conjugate A)\<^sup>T = A\<^sup>H" + unfolding conjugate_mat_def + by auto + +lemma vec_module_col_helper: + fixes A:: "('a :: field) mat" + shows "(0\<^sub>v (dim_row A) \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A)))" +proof - + have "\v. (0::'a) \\<^sub>v v + v = v" + by auto + then show "0\<^sub>v (dim_row A) \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" + by (metis cols_dim module_vec_def right_zero_vec smult_carrier_vec vec_space.prod_in_span zero_carrier_vec) +qed + +lemma vec_module_col_helper2: + fixes A:: "('a :: field) mat" + shows "\a x. x \ LinearCombinations.module.span class_ring + \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, + zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ + (set (cols A)) \ + (\a b v. (a + b) \\<^sub>v v = a \\<^sub>v v + b \\<^sub>v v) \ + a \\<^sub>v x + \ LinearCombinations.module.span class_ring + \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, + zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ + (set (cols A))" +proof - + fix a :: 'a and x :: "'a vec" + assume "x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" + then show "a \\<^sub>v x \ LinearCombinations.module.span class_ring \carrier = carrier_vec (dim_row A), mult = undefined, one = undefined, zero = 0\<^sub>v (dim_row A), add = (+), smult = (\\<^sub>v)\ (set (cols A))" + by (metis (full_types) cols_dim idom_vec.smult_in_span module_vec_def) +qed + +lemma vec_module_col: "module (class_ring :: 'a :: field ring) + (module_vec TYPE('a) + (dim_row A) + \carrier := + LinearCombinations.module.span + class_ring (module_vec TYPE('a) (dim_row A)) (set (cols A))\)" +proof - + interpret abelian_group "module_vec TYPE('a) (dim_row A) + \carrier := + LinearCombinations.module.span + class_ring (module_vec TYPE('a) (dim_row A)) (set (cols A))\" + apply (unfold_locales) + apply (auto simp add:module_vec_def) + apply (metis cols_dim module_vec_def partial_object.select_convs(1) ring.simps(2) vec_vs vectorspace.span_add1) + apply (metis assoc_add_vec cols_dim module_vec_def vec_space.cV vec_vs vectorspace.span_closed) + using vec_module_col_helper[of A] apply (auto) + apply (metis cols_dim left_zero_vec module_vec_def partial_object.select_convs(1) vec_vs vectorspace.span_closed) + apply (metis cols_dim module_vec_def partial_object.select_convs(1) right_zero_vec vec_vs vectorspace.span_closed) + apply (metis cols_dim comm_add_vec module_vec_def vec_space.cV vec_vs vectorspace.span_closed) + unfolding Units_def apply auto + by (smt cols_dim module_vec_def partial_object.select_convs(1) uminus_l_inv_vec uminus_r_inv_vec vec_space.vec_neg vec_vs vectorspace.span_closed vectorspace.span_neg) + show ?thesis + apply (unfold_locales) + unfolding class_ring_simps apply auto + unfolding module_vec_simps using add_smult_distrib_vec apply auto + apply (auto simp add:module_vec_def) + using vec_module_col_helper2 + apply blast + using cols_dim module_vec_def partial_object.select_convs(1) smult_add_distrib_vec vec_vs vectorspace.span_closed + by (smt (z3)) +qed + +(* The columns of a matrix form a vectorspace *) +lemma vec_vs_col: "vectorspace (class_ring :: 'a :: field ring) + (module_vec TYPE('a) (dim_row A) + \carrier := + LinearCombinations.module.span + class_ring + (module_vec TYPE('a) + (dim_row A)) + (set (cols A))\)" + unfolding vectorspace_def + using vec_module_col class_field + by (auto simp: class_field_def) + +lemma cols_mat_mul_map: + shows "cols (A * B) = map ((*\<^sub>v) A) (cols B)" + unfolding list_eq_iff_nth_eq + by auto + +lemma cols_mat_mul: + shows "set (cols (A * B)) = (*\<^sub>v) A ` set (cols B)" + by (simp add: cols_mat_mul_map) + +lemma set_obtain_sublist: + assumes "S \ set ls" + obtains ss where "distinct ss" "S = set ss" + using assms finite_distinct_list infinite_super by blast + +lemma mul_mat_of_cols: + assumes "A \ carrier_mat nr n" + assumes "\j. j < length cs \ cs ! j \ carrier_vec n" + shows "A * (mat_of_cols n cs) = mat_of_cols nr (map ((*\<^sub>v) A) cs)" + unfolding mat_eq_iff + using assms apply auto + apply (subst mat_of_cols_index) + by auto + +lemma helper: + fixes x y z ::"'a :: {conjugatable_ring, comm_ring}" + shows "x * (y * z) = y * x * z" + by (simp add: mult.assoc mult.left_commute) + +lemma cscalar_prod_conjugate_transpose: + fixes x y ::"'a :: {conjugatable_ring, comm_ring} vec" + assumes "A \ carrier_mat nr nc" + assumes "x \ carrier_vec nr" + assumes "y \ carrier_vec nc" + shows "x \c (A *\<^sub>v y) = (A\<^sup>H *\<^sub>v x) \c y" + unfolding mult_mat_vec_def scalar_prod_def + using assms apply (auto simp add: sum_distrib_left sum_distrib_right sum_conjugate conjugate_dist_mul) + apply (subst sum.swap) + by (meson helper mult.assoc mult.left_commute sum.cong) + +lemma mat_mul_conjugate_transpose_vec_eq_0: + fixes v ::"'a :: {conjugatable_ordered_ring,semiring_no_zero_divisors,comm_ring} vec" + assumes "A \ carrier_mat nr nc" + assumes "v \ carrier_vec nr" + assumes "A *\<^sub>v (A\<^sup>H *\<^sub>v v) = 0\<^sub>v nr" + shows "A\<^sup>H *\<^sub>v v = 0\<^sub>v nc" +proof - + have "(A\<^sup>H *\<^sub>v v) \c (A\<^sup>H *\<^sub>v v) = (A *\<^sub>v (A\<^sup>H *\<^sub>v v)) \c v" + by (metis (mono_tags, lifting) Matrix.carrier_vec_conjugate assms(1) assms(2) assms(3) carrier_matD(2) conjugate_zero_vec cscalar_prod_conjugate_transpose dim_row_conjugate index_transpose_mat(2) mult_mat_vec_def scalar_prod_left_zero scalar_prod_right_zero vec_carrier) + also have "... = 0" + by (simp add: assms(2) assms(3)) + (* this step requires real entries *) + ultimately have "(A\<^sup>H *\<^sub>v v) \c (A\<^sup>H *\<^sub>v v) = 0" by auto + thus ?thesis + apply (subst conjugate_square_eq_0_vec[symmetric]) + using assms(1) carrier_dim_vec apply fastforce + by auto +qed + +lemma row_mat_of_cols: + assumes "i < nr" + shows "row (mat_of_cols nr ls) i = vec (length ls) (\j. (ls ! j) $i)" + by (smt assms dim_vec eq_vecI index_row(1) index_row(2) index_vec mat_of_cols_carrier(2) mat_of_cols_carrier(3) mat_of_cols_index) + +lemma mat_of_cols_cons_mat_vec: + fixes v ::"'a::comm_ring vec" + assumes "v \ carrier_vec (length ls)" + assumes "dim_vec a = nr" + shows + "mat_of_cols nr (a # ls) *\<^sub>v (vCons m v) = + m \\<^sub>v a + mat_of_cols nr ls *\<^sub>v v" + unfolding mult_mat_vec_def vec_eq_iff + using assms by + (auto simp add: row_mat_of_cols vec_Suc o_def mult.commute) + +lemma smult_vec_zero: + fixes v ::"'a::ring vec" + shows "0 \\<^sub>v v = 0\<^sub>v (dim_vec v)" + unfolding smult_vec_def vec_eq_iff + by (auto) + +lemma helper2: + fixes A ::"'a::comm_ring mat" + fixes v ::"'a vec" + assumes "v \ carrier_vec (length ss)" + assumes "\x. x \ set ls \ dim_vec x = nr" + shows + "mat_of_cols nr ss *\<^sub>v v = + mat_of_cols nr (ls @ ss) *\<^sub>v (0\<^sub>v (length ls) @\<^sub>v v)" + using assms(2) +proof (induction ls) + case Nil + then show ?case by auto +next + case (Cons a ls) + then show ?case apply (auto simp add:zero_vec_Suc) + apply (subst mat_of_cols_cons_mat_vec) + by (auto simp add:assms smult_vec_zero) +qed + +lemma mat_of_cols_mult_mat_vec_permute_list: + fixes v ::"'a::comm_ring list" + assumes "f permutes {..v vec_of_list (permute_list f v) = + mat_of_cols nr ss *\<^sub>v vec_of_list v" + unfolding mat_of_cols_def mult_mat_vec_def vec_eq_iff scalar_prod_def +proof clarsimp + fix i + assume "i < nr" + from sum.permute[OF assms(1)] + have "(\iaia. ss ! f ia $ i * v ! f ia) \ f) {..ia = 0..ia = 0..ia = 0..\g. sum g {.. f) {.. assms(2) comp_apply lessThan_atLeast0 sum.cong) + show "(\ia = 0..j. permute_list f ss ! j $ i) $ ia * + vec_of_list (permute_list f v) $ ia) = + (\ia = 0..j. ss ! j $ i) $ ia * vec_of_list v $ ia)" + using assms * by (auto simp add: permute_list_nth vec_of_list_index) +qed + +(* permute everything in a subset of the indices to the back *) +lemma subindex_permutation: + assumes "distinct ss" "set ss \ {..i. i \ set ss) [0..i. i \ set ss) [0..i. i \ set ss) [0..i. i \ set ss) + [0.. {..i. i \ set ss) [0.. set ls" + obtains ids where "distinct ids" "set ids \ {.. {..ids. distinct ids \ + set ids \ {.. + ss = map ((!) ls) ids \ thesis) \ + thesis" using 1 2 3 by blast +qed + +lemma helper3: + fixes A ::"'a::comm_ring mat" + assumes A: "A \ carrier_mat nr nc" + assumes ss:"distinct ss" "set ss \ set (cols A)" + assumes "v \ carrier_vec (length ss)" + obtains c where "mat_of_cols nr ss *\<^sub>v v = A *\<^sub>v c" "dim_vec c = nc" +proof - + from distinct_list_subset_nths[OF ss] + obtain ids where ids: "distinct ids" "set ids \ {..x. x \ set ?ls \ dim_vec x = nr" + using A by auto + let ?cs1 = "(list_of_vec (0\<^sub>v (length ?ls) @\<^sub>v v))" + from helper2[OF assms(4) ] + have "mat_of_cols nr ss *\<^sub>v v = mat_of_cols nr (?ls @ ss) *\<^sub>v vec_of_list (?cs1)" + using * + by (metis vec_list) + also have "... = mat_of_cols nr (permute_list f (?ls @ ss)) *\<^sub>v vec_of_list (permute_list f ?cs1)" + apply (auto intro!: mat_of_cols_mult_mat_vec_permute_list[symmetric]) + apply (metis cols_length f(1) f(2) length_append length_map length_permute_list) + using assms(4) by auto + also have "... = A *\<^sub>v vec_of_list (permute_list f ?cs1)" using f(2) assms by auto + ultimately show + "(\c. mat_of_cols nr ss *\<^sub>v v = A *\<^sub>v c \ dim_vec c = nc \ thesis) \ thesis" + by (metis A assms(4) carrier_matD(2) carrier_vecD cols_length dim_vec_of_list f(2) index_append_vec(2) index_zero_vec(2) length_append length_list_of_vec length_permute_list) +qed + +lemma mat_mul_conjugate_transpose_sub_vec_eq_0: + fixes A ::"'a :: {conjugatable_ordered_ring,semiring_no_zero_divisors,comm_ring} mat" + assumes "A \ carrier_mat nr nc" + assumes "distinct ss" "set ss \ set (cols (A\<^sup>H))" + assumes "v \ carrier_vec (length ss)" + assumes "A *\<^sub>v (mat_of_cols nc ss *\<^sub>v v) = 0\<^sub>v nr" + shows "(mat_of_cols nc ss *\<^sub>v v) = 0\<^sub>v nc" +proof - + have "A\<^sup>H \ carrier_mat nc nr" using assms(1) by auto + from helper3[OF this assms(2-4)] + obtain c where c: "mat_of_cols nc ss *\<^sub>v v = A\<^sup>H *\<^sub>v c" "dim_vec c = nr" by blast + have 1: "c \ carrier_vec nr" + using c carrier_vec_dim_vec by blast + have 2: "A *\<^sub>v (A\<^sup>H *\<^sub>v c) = 0\<^sub>v nr" using c assms(5) by auto + from mat_mul_conjugate_transpose_vec_eq_0[OF assms(1) 1 2] + have "A\<^sup>H *\<^sub>v c = 0\<^sub>v nc" . + thus ?thesis unfolding c(1)[symmetric] . +qed + +lemma Units_invertible: + fixes A:: "'a::semiring_1 mat" + assumes "A \ Units (ring_mat TYPE('a) n b)" + shows "invertible_mat A" + using assms unfolding Units_def invertible_mat_def + apply (auto simp add: ring_mat_def) + using inverts_mat_def by blast + +lemma invertible_Units: + fixes A:: "'a::semiring_1 mat" + assumes "invertible_mat A" + shows "A \ Units (ring_mat TYPE('a) (dim_row A) b)" + using assms unfolding Units_def invertible_mat_def + apply (auto simp add: ring_mat_def) + by (metis assms carrier_mat_triv invertible_mat_def inverts_mat_def inverts_mat_length(1) inverts_mat_length(2)) + +lemma invertible_det: + fixes A:: "'a::field mat" + assumes "A \ carrier_mat n n" + shows "invertible_mat A \ det A \ 0" + apply auto + using invertible_Units unit_imp_det_non_zero apply fastforce + using assms by (auto intro!: Units_invertible det_non_zero_imp_unit) + +context vec_space begin + +lemma find_indices_distinct: + assumes "distinct ss" + assumes "i < length ss" + shows "find_indices (ss ! i) ss = [i]" +proof - + have "set (find_indices (ss ! i) ss) = {i}" + using assms apply auto by (simp add: assms(1) assms(2) nth_eq_iff_index_eq) + thus ?thesis + by (metis distinct.simps(2) distinct_find_indices empty_iff empty_set insert_iff list.exhaust list.simps(15)) +qed + +lemma lin_indpt_lin_comb_list: + assumes "distinct ss" + assumes "lin_indpt (set ss)" + assumes "set ss \ carrier_vec n" + assumes "lincomb_list f ss = 0\<^sub>v n" + assumes "i < length ss" + shows "f i = 0" +proof - + from lincomb_list_as_lincomb[OF assms(3)] + have "lincomb_list f ss = lincomb (mk_coeff ss f) (set ss)" . + also have "... = lincomb (\v. sum f (set (find_indices v ss))) (set ss)" + unfolding mk_coeff_def + apply (subst R.sumlist_map_as_finsum) + by (auto simp add: distinct_find_indices) + ultimately have "lincomb_list f ss = lincomb (\v. sum f (set (find_indices v ss))) (set ss)" by auto + then have *:"lincomb (\v. sum f (set (find_indices v ss))) (set ss) = 0\<^sub>v n" + using assms(4) by auto + have "finite (set ss)" by simp + from not_lindepD[OF assms(2) this _ _ *] + have "(\v. sum f (set (find_indices v ss))) \ set ss \ {0}" + by auto + from funcset_mem[OF this] + have "sum f (set (find_indices (nth ss i) ss)) \ {0}" + using assms(5) by auto + thus ?thesis unfolding find_indices_distinct[OF assms(1) assms(5)] + by auto +qed + +(* Note: in this locale dim_row A = n, e.g.: +lemma foo: + assumes "dim_row A = n" + shows "rank A = vec_space.rank (dim_row A) A" + by (simp add: assms) *) + +lemma span_mat_mul_subset: + assumes "A \ carrier_mat n d" + assumes "B \ carrier_mat d nc" + shows "span (set (cols (A * B))) \ span (set (cols A))" +proof - + have *: "\v. \ca. lincomb_list v (cols (A * B)) = + lincomb_list ca (cols A)" + proof - + fix v + have "lincomb_list v (cols (A * B)) = (A * B) *\<^sub>v vec nc v" + apply (subst lincomb_list_as_mat_mult) + apply (metis assms(1) carrier_dim_vec carrier_matD(1) cols_dim index_mult_mat(2) subset_code(1)) + by (metis assms(1) assms(2) carrier_matD(1) carrier_matD(2) cols_length index_mult_mat(2) index_mult_mat(3) mat_of_cols_cols) + also have "... = A *\<^sub>v (B *\<^sub>v vec nc v)" + using assms(1) assms(2) by auto + also have "... = lincomb_list (\i. (B *\<^sub>v vec nc v) $ i) (cols A)" + apply (subst lincomb_list_as_mat_mult) + using assms(1) carrier_dim_vec cols_dim apply blast + by (metis assms(1) assms(2) carrier_matD(1) carrier_matD(2) cols_length dim_mult_mat_vec dim_vec eq_vecI index_vec mat_of_cols_cols) + ultimately have "lincomb_list v (cols (A * B)) = + lincomb_list (\i. (B *\<^sub>v vec nc v) $ i) (cols A)" by auto + thus "\ca. lincomb_list v (cols (A * B)) = lincomb_list ca (cols A)" by auto + qed + show ?thesis + apply (subst span_list_as_span[symmetric]) + apply (metis assms(1) carrier_matD(1) cols_dim index_mult_mat(2)) + apply (subst span_list_as_span[symmetric]) + using assms(1) cols_dim apply blast + by (auto simp add:span_list_def *) +qed + +lemma rank_mat_mul_right: + assumes "A \ carrier_mat n d" + assumes "B \ carrier_mat d nc" + shows "rank (A * B) \ rank A" +proof - + have "subspace class_ring (local.span (set (cols (A*B)))) + (vs (local.span (set (cols A))))" + unfolding subspace_def + by (metis assms(1) assms(2) carrier_matD(1) cols_dim index_mult_mat(2) nested_submodules span_is_submodule vec_space.span_mat_mul_subset vec_vs_col) + from vectorspace.subspace_dim[OF _ this] + have "vectorspace.dim class_ring + (vs (local.span (set (cols A))) + \carrier := local.span (set (cols (A * B)))\) \ + vectorspace.dim class_ring + (vs (local.span (set (cols A))))" + apply auto + by (metis (no_types) assms(1) carrier_matD(1) fin_dim_span_cols index_mult_mat(2) mat_of_cols_carrier(1) mat_of_cols_cols vec_vs_col) + thus ?thesis unfolding rank_def + by auto +qed + +lemma sumlist_drop: + assumes "\v. v \ set ls \ dim_vec v = n" + shows "sumlist ls = sumlist (filter (\v. v \ 0\<^sub>v n) ls)" + using assms +proof (induction ls) + case Nil + then show ?case by auto +next + case (Cons a ls) + then show ?case using dim_sumlist by auto +qed + +lemma lincomb_list_alt: + shows "lincomb_list c s = + sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) [0..v. v \ set s \ dim_vec v = n" + assumes "\i. i \ set ls \ i < length s" + shows " + sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) ls) ls) = + sumlist (map2 (\i j. i \\<^sub>v s ! j) (map (\i. c i) (filter (\i. c i \ 0) ls)) (filter (\i. c i \ 0) ls))" + using assms(2) +proof (induction ls) + case Nil + then show ?case by auto +next + case (Cons a s) + then show ?case + apply auto + apply (subst smult_l_null) + apply (simp add: assms(1) carrier_vecI) + apply (subst left_zero_vec) + apply (subst sumlist_carrier) + apply auto + by (metis (no_types, lifting) assms(1) carrier_dim_vec mem_Collect_eq nth_mem set_filter set_zip_rightD) +qed + +lemma two_set: + assumes "distinct ls" + assumes "set ls = set [a,b]" + assumes "a \ b" + shows "ls = [a,b] \ ls = [b,a]" + apply (cases ls) + using assms(2) apply auto[1] +proof - + fix x xs + assume ls:"ls = x # xs" + obtain y ys where xs:"xs = y # ys" + by (metis (no_types) \ls = x # xs\ assms(2) assms(3) list.set_cases list.set_intros(1) list.set_intros(2) set_ConsD) + have 1:"x = a \ x = b" + using \ls = x # xs\ assms(2) by auto + have 2:"y = a \ y = b" + using \ls = x # xs\ \xs = y # ys\ assms(2) by auto + have 3:"ys = []" + by (metis (no_types) \ls = x # xs\ \xs = y # ys\ assms(1) assms(2) distinct.simps(2) distinct_length_2_or_more in_set_member member_rec(2) neq_Nil_conv set_ConsD) + show "ls = [a, b] \ ls = [b, a]" using ls xs 1 2 3 assms + by auto +qed + +lemma filter_disj_inds: + assumes "i < length ls" "j < length ls" "i \ j" + shows "filter (\ia. ia \ j \ ia = i) [0.. + filter (\ia. ia \ j \ ia = i) [0..ia. ia = i \ ia = j) [0..ia. ia = i \ ia = j) [0..v. v \ set ls \ dim_vec v = n" + assumes + "\c. lincomb_list c ls = 0\<^sub>v n \ (\i. i < (length ls) \ c i = 0)" + shows "distinct ls" + unfolding distinct_conv_nth +proof clarsimp + fix i j + assume ij: "i < length ls" "j < length ls" "i \ j" + assume lsij: "ls ! i = ls ! j" + have "lincomb_list (\k. if k = i then 1 else if k = j then -1 else 0) ls = + (ls ! i) - (ls ! j)" + unfolding lincomb_list_alt + apply (subst lincomb_list_alt2[OF assms(1)]) + apply auto + using filter_disj_inds[OF ij] + apply auto + using ij(3) apply force + using assms(1) ij(2) apply auto[1] + using ij(3) apply blast + using assms(1) ij(2) by auto + also have "... = 0\<^sub>v n" unfolding lsij + apply (rule minus_cancel_vec) + using \j < length ls\ assms(1) + using carrier_vec_dim_vec nth_mem by blast + ultimately have "lincomb_list (\k. if k = i then 1 else if k = j then -1 else 0) ls = 0\<^sub>v n" by auto + from assms(2)[OF this] + show False + using \i < length ls\ by auto +qed + +end + +locale conjugatable_vec_space = vec_space f_ty n for + f_ty::"'a::conjugatable_ordered_field itself" + and n +begin + +lemma transpose_rank_mul_conjugate_transpose: + fixes A :: "'a mat" + assumes "A \ carrier_mat n nc" + shows "vec_space.rank nc A\<^sup>H \ rank (A * A\<^sup>H)" +proof - + have 1: "A\<^sup>H \ carrier_mat nc n" using assms by auto + have 2: "A * A\<^sup>H \ carrier_mat n n" using assms by auto + (* S is a maximal linearly independent set of rows A (or cols A\<^sup>T) *) + let ?P = "(\T. T \ set (cols A\<^sup>H) \ module.lin_indpt class_ring (module_vec TYPE('a) nc) T)" + have *:"\A. ?P A \ finite A \ card A \ n" + proof clarsimp + fix S + assume S: "S \ set (cols A\<^sup>H)" + have "card S \ card (set (cols A\<^sup>H))" using S + using card_mono by blast + also have "... \ length (cols A\<^sup>H)" using card_length by blast + also have "... \ n" using assms by auto + ultimately show "finite S \ card S \ n" + by (meson List.finite_set S dual_order.trans finite_subset) + qed + have **:"?P {}" + apply (subst module.lin_dep_def) + by (auto simp add: vec_module) + from maximal_exists[OF *] + obtain S where S: "maximal S ?P" using ** + by (metis (no_types, lifting)) + (* Some properties of S *) + from vec_space.rank_card_indpt[OF 1 S] + have rankeq: "vec_space.rank nc A\<^sup>H = card S" . + + have s_hyp: "S \ set (cols A\<^sup>H)" + using S unfolding maximal_def by simp + have modhyp: "module.lin_indpt class_ring (module_vec TYPE('a) nc) S" + using S unfolding maximal_def by simp + +(* switch to a list representation *) + obtain ss where ss: "set ss = S" "distinct ss" + by (metis (mono_tags) S maximal_def set_obtain_sublist) + have ss2: "set (map ((*\<^sub>v) A) ss) = (*\<^sub>v) A ` S" + by (simp add: ss(1)) + have rw_hyp: "cols (mat_of_cols n (map ((*\<^sub>v) A) ss)) = cols (A * mat_of_cols nc ss)" + unfolding cols_def apply (auto) + using mat_vec_as_mat_mat_mult[of A n nc] + by (metis (no_types, lifting) "1" assms carrier_matD(1) cols_dim mul_mat_of_cols nth_mem s_hyp ss(1) subset_code(1)) + then have rw: "mat_of_cols n (map ((*\<^sub>v) A) ss) = A * mat_of_cols nc ss" + by (metis assms carrier_matD(1) index_mult_mat(2) mat_of_cols_carrier(2) mat_of_cols_cols) + have indpt: "\c. lincomb_list c (map ((*\<^sub>v) A) ss) = 0\<^sub>v n \ + \i. (i < (length ss) \ c i = 0)" + proof clarsimp + fix c i + assume *: "lincomb_list c (map ((*\<^sub>v) A) ss) = 0\<^sub>v n" + assume i: "i < length ss" + have "\w\set (map ((*\<^sub>v) A) ss). dim_vec w = n" + using assms by auto + from lincomb_list_as_mat_mult[OF this] + have "A * mat_of_cols nc ss *\<^sub>v vec (length ss) c = 0\<^sub>v n" + using * rw by auto + then have hq: "A *\<^sub>v (mat_of_cols nc ss *\<^sub>v vec (length ss) c) = 0\<^sub>v n" + by (metis assms assoc_mult_mat_vec mat_of_cols_carrier(1) vec_carrier) + + then have eq1: "(mat_of_cols nc ss *\<^sub>v vec (length ss) c) = 0\<^sub>v nc" + apply (intro mat_mul_conjugate_transpose_sub_vec_eq_0) + using assms ss s_hyp by auto + +(* Rewrite the inner vector back to a lincomb_list *) + have dim_hyp2: "\w\set ss. dim_vec w = nc" + using ss(1) s_hyp + by (metis "1" carrier_matD(1) carrier_vecD cols_dim subsetD) + from vec_module.lincomb_list_as_mat_mult[OF this, symmetric] + have "mat_of_cols nc ss *\<^sub>v vec (length ss) c = module.lincomb_list (module_vec TYPE('a) nc) c ss" . + then have "module.lincomb_list (module_vec TYPE('a) nc) c ss = 0\<^sub>v nc" using eq1 by auto + from vec_space.lin_indpt_lin_comb_list[OF ss(2) _ _ this i] + show "c i = 0" using modhyp ss s_hyp + using "1" cols_dim by blast + qed + have distinct: "distinct (map ((*\<^sub>v) A) ss)" + by (metis (no_types, lifting) assms carrier_matD(1) dim_mult_mat_vec imageE indpt length_map lincomb_list_indpt_distinct ss2) + then have 3: "card S = card ((*\<^sub>v) A ` S)" + by (metis ss distinct_card image_set length_map) + then have 4: "(*\<^sub>v) A ` S \ set (cols (A * A\<^sup>H))" + using cols_mat_mul \S \ set (cols A\<^sup>H)\ by blast + have 5: "lin_indpt ((*\<^sub>v) A ` S)" + proof clarsimp + assume ld:"lin_dep ((*\<^sub>v) A ` S)" + have *: "finite ((*\<^sub>v) A ` S)" + by (metis List.finite_set ss2) + have **: "(*\<^sub>v) A ` S \ carrier_vec n" + using "2" "4" cols_dim by blast + from finite_lin_dep[OF * ld **] + obtain a v where + a: "lincomb a ((*\<^sub>v) A ` S) = 0\<^sub>v n" and + v: "v \ (*\<^sub>v) A ` S" "a v \ 0" by blast + obtain i where i:"v = map ((*\<^sub>v) A) ss ! i" "i < length ss" + using v unfolding ss2[symmetric] + using find_first_le nth_find_first by force + from ss2[symmetric] + have "set (map ((*\<^sub>v) A) ss)\ carrier_vec n" using ** ss2 by auto + from lincomb_as_lincomb_list_distinct[OF this distinct] have + "lincomb_list + (\i. a (map ((*\<^sub>v) A) ss ! i)) (map ((*\<^sub>v) A) ss) = 0\<^sub>v n" + using a ss2 by auto + from indpt[OF this] + show False using v i by simp + qed + from rank_ge_card_indpt[OF 2 4 5] + have "card ((*\<^sub>v) A ` S) \ rank (A * A\<^sup>H)" . + thus ?thesis using rankeq 3 by linarith +qed + +lemma conjugate_transpose_rank_le: + assumes "A \ carrier_mat n nc" + shows "vec_space.rank nc (A\<^sup>H) \ rank A" + by (metis assms carrier_matD(2) carrier_mat_triv dim_row_conjugate dual_order.trans index_transpose_mat(2) rank_mat_mul_right transpose_rank_mul_conjugate_transpose) + +lemma conjugate_finsum: + assumes f: "f : U \ carrier_vec n" + shows "conjugate (finsum V f U) = finsum V (conjugate \ f) U" + using f +proof (induct U rule: infinite_finite_induct) + case (infinite A) + then show ?case by auto +next + case empty + then show ?case by auto +next + case (insert u U) + hence f: "f : U \ carrier_vec n" "f u : carrier_vec n" by auto + then have cf: "conjugate \ f : U \ carrier_vec n" + "(conjugate \ f) u : carrier_vec n" + apply (simp add: Pi_iff) + by (simp add: f(2)) + then show ?case + unfolding finsum_insert[OF insert(1) insert(2) f] + unfolding finsum_insert[OF insert(1) insert(2) cf ] + apply (subst conjugate_add_vec[of _ n]) + using f(2) apply blast + using M.finsum_closed f(1) apply blast + by (simp add: comp_def f(1) insert.hyps(3)) +qed + +lemma rank_conjugate_le: + assumes A:"A \ carrier_mat n d" + shows "rank (conjugate (A)) \ rank A" +proof - + (* S is a maximal linearly independent set of (conjugate A) *) + let ?P = "(\T. T \ set (cols (conjugate A)) \ lin_indpt T)" + have *:"\A. ?P A \ finite A \ card A \ d" + by (metis List.finite_set assms card_length card_mono carrier_matD(2) cols_length dim_col_conjugate dual_order.trans rev_finite_subset) + have **:"?P {}" + by (simp add: finite_lin_indpt2) + from maximal_exists[OF *] + obtain S where S: "maximal S ?P" using ** + by (metis (no_types, lifting)) + have s_hyp: "S \ set (cols (conjugate A))" "lin_indpt S" + using S unfolding maximal_def + apply blast + by (metis (no_types, lifting) S maximal_def) + from rank_card_indpt[OF _ S, of d] + have rankeq: "rank (conjugate A) = card S" using assms by auto + have 1:"conjugate ` S \ set (cols A)" + using S apply auto + by (metis (no_types, lifting) cols_conjugate conjugate_id image_eqI in_mono list.set_map s_hyp(1)) + have 2: "lin_indpt (conjugate ` S)" + apply (rule ccontr) + apply (auto simp add: lin_dep_def) + proof - + fix T c v + assume T: "T \ conjugate ` S" "finite T" and + lc:"lincomb c T = 0\<^sub>v n" and "v \ T" "c v \ 0" + let ?T = "conjugate ` T" + let ?c = "conjugate \ c \ conjugate" + have 1: "finite ?T" using T by auto + have 2: "?T \ S" using T by auto + have 3: "?c \ ?T \ UNIV" by auto + have "lincomb ?c ?T = (\\<^bsub>V\<^esub>x\T. conjugate (c x) \\<^sub>v conjugate x)" + unfolding lincomb_def + apply (subst finsum_reindex) + apply auto + apply (metis "2" carrier_vec_conjugate assms carrier_matD(1) cols_dim image_eqI s_hyp(1) subsetD) + by (meson conjugate_cancel_iff inj_onI) + also have "... = (\\<^bsub>V\<^esub>x\T. conjugate (c x \\<^sub>v x)) " + by (simp add: conjugate_smult_vec) + also have "... = conjugate (\\<^bsub>V\<^esub>x\T. (c x \\<^sub>v x))" + apply(subst conjugate_finsum[of "\x.(c x \\<^sub>v x)" T]) + apply (auto simp add:o_def) + by (smt Matrix.carrier_vec_conjugate Pi_I' T(1) assms carrier_matD(1) cols_dim dim_row_conjugate imageE s_hyp(1) smult_carrier_vec subset_eq) + also have "... = conjugate (lincomb c T)" + using lincomb_def by presburger + ultimately have "lincomb ?c ?T = conjugate (lincomb c T)" by auto + then have 4:"lincomb ?c ?T = 0\<^sub>v n" using lc by auto + from not_lindepD[OF s_hyp(2) 1 2 3 4] + have "conjugate \ c \ conjugate \ conjugate ` T \ {0}" . + then have "c v = 0" + by (simp add: Pi_iff \v \ T\) + thus False using \c v \ 0\ by auto + qed + from rank_ge_card_indpt[OF A 1 2] + have 3:"card (conjugate ` S) \ rank A" . + have 4: "card (conjugate ` S) = card S" + apply (auto intro!: card_image) + by (meson conjugate_cancel_iff inj_onI) + show ?thesis using rankeq 3 4 by auto +qed + +lemma rank_conjugate: + assumes "A \ carrier_mat n d" + shows "rank (conjugate A) = rank A" + using rank_conjugate_le + by (metis carrier_vec_conjugate assms conjugate_id dual_order.antisym) + +end (* exit the context *) + +lemma conjugate_transpose_rank: + fixes A::"'a::{conjugatable_ordered_field} mat" + shows "vec_space.rank (dim_row A) A = vec_space.rank (dim_col A) (A\<^sup>H)" + using conjugatable_vec_space.conjugate_transpose_rank_le + by (metis (no_types, lifting) Matrix.transpose_transpose carrier_matI conjugate_id dim_col_conjugate dual_order.antisym index_transpose_mat(2) transpose_conjugate) + +lemma transpose_rank: + fixes A::"'a::{conjugatable_ordered_field} mat" + shows "vec_space.rank (dim_row A) A = vec_space.rank (dim_col A) (A\<^sup>T)" + by (metis carrier_mat_triv conjugatable_vec_space.rank_conjugate conjugate_transpose_rank index_transpose_mat(2)) + +lemma rank_mat_mul_left: + fixes A::"'a::{conjugatable_ordered_field} mat" + assumes "A \ carrier_mat n d" + assumes "B \ carrier_mat d nc" + shows "vec_space.rank n (A * B) \ vec_space.rank d B" + by (metis (no_types, lifting) Matrix.transpose_transpose assms(1) assms(2) carrier_matD(1) carrier_matD(2) carrier_mat_triv conjugatable_vec_space.rank_conjugate conjugate_transpose_rank index_mult_mat(3) index_transpose_mat(3) transpose_mult vec_space.rank_mat_mul_right) + +section "Results on Invertibility" + +(* Extract specific columns of a matrix *) +definition take_cols :: "'a mat \ nat list \ 'a mat" + where "take_cols A inds = mat_of_cols (dim_row A) (map ((!) (cols A)) (filter ((>) (dim_col A)) inds))" + +definition take_cols_var :: "'a mat \ nat list \ 'a mat" + where "take_cols_var A inds = mat_of_cols (dim_row A) (map ((!) (cols A)) (inds))" + +definition take_rows :: "'a mat \ nat list \ 'a mat" + where "take_rows A inds = mat_of_rows (dim_col A) (map ((!) (rows A)) (filter ((>) (dim_row A)) inds))" + +lemma cong1: + "x = y \ mat_of_cols n x = mat_of_cols n y" + by auto + +lemma nth_filter: + assumes "j < length (filter P ls)" + shows "P ((filter P ls) ! j)" + by (simp add: assms list_ball_nth) + +lemma take_cols_mat_mul: + assumes "A \ carrier_mat nr n" + assumes "B \ carrier_mat n nc" + shows "A * take_cols B inds = take_cols (A * B) inds" +proof - + have "\j. j < length (map ((!) (cols B)) (filter ((>) nc) inds)) \ + (map ((!) (cols B)) (filter ((>) nc) inds)) ! j \ carrier_vec n" + using assms apply auto + apply (subst cols_nth) + using nth_filter by auto + from mul_mat_of_cols[OF assms(1) this] + have "A * take_cols B inds = mat_of_cols nr (map (\x. A *\<^sub>v cols B ! x) (filter ((>) (dim_col B)) inds))" + unfolding take_cols_def using assms by (auto simp add: o_def) + also have "... = take_cols (A * B) inds" + unfolding take_cols_def using assms apply (auto intro!: cong1) + by (simp add: mult_mat_vec_def) + ultimately show ?thesis by auto +qed + +lemma take_cols_carrier_mat: + assumes "A \ carrier_mat nr nc" + obtains n where "take_cols A inds \ carrier_mat nr n" + unfolding take_cols_def + using assms + by fastforce + +lemma take_cols_carrier_mat_strict: + assumes "A \ carrier_mat nr nc" + assumes "\i. i \ set inds \ i < nc" + shows "take_cols A inds \ carrier_mat nr (length inds)" + unfolding take_cols_def + using assms by auto + +lemma gauss_jordan_take_cols: + assumes "gauss_jordan A (take_cols A inds) = (C,D)" + shows "D = take_cols C inds" +proof - + obtain nr nc where A: "A \ carrier_mat nr nc" by auto + from take_cols_carrier_mat[OF this] + obtain n where B: "take_cols A inds \ carrier_mat nr n" by auto + from gauss_jordan_transform[OF A B assms, of undefined] + obtain P where PP:"P\Units (ring_mat TYPE('a) nr undefined)" and + CD: "C = P * A" "D = P * take_cols A inds" by blast + have P: "P \ carrier_mat nr nr" + by (metis (no_types, lifting) Units_def PP mem_Collect_eq partial_object.select_convs(1) ring_mat_def) + from take_cols_mat_mul[OF P A] + have "P * take_cols A inds = take_cols (P * A) inds" . + thus ?thesis using CD by blast +qed + +lemma dim_col_take_cols: + assumes "\j. j \ set inds \ j < dim_col A" + shows "dim_col (take_cols A inds) = length inds" + unfolding take_cols_def + using assms by auto + +lemma dim_col_take_rows[simp]: + shows "dim_col (take_rows A inds) = dim_col A" + unfolding take_rows_def by auto + +lemma cols_take_cols_subset: + shows "set (cols (take_cols A inds)) \ set (cols A)" + unfolding take_cols_def + apply (subst cols_mat_of_cols) + apply auto + using in_set_conv_nth by fastforce + +lemma dim_row_take_cols[simp]: + shows "dim_row (take_cols A ls) = dim_row A" + by (simp add: take_cols_def) + +lemma dim_row_append_rows[simp]: + shows "dim_row (A @\<^sub>r B) = dim_row A + dim_row B" + by (simp add: append_rows_def) + +lemma rows_inj: + assumes "dim_col A = dim_col B" + assumes "rows A = rows B" + shows "A = B" + unfolding mat_eq_iff + apply auto + apply (metis assms(2) length_rows) + using assms(1) apply blast + by (metis assms(1) assms(2) mat_of_rows_rows) + +lemma append_rows_index: + assumes "dim_col A = dim_col B" + assumes "i < dim_row A + dim_row B" + assumes "j < dim_col A" + shows "(A @\<^sub>r B) $$ (i,j) = (if i < dim_row A then A $$ (i,j) else B $$ (i-dim_row A,j))" + unfolding append_rows_def + apply (subst index_mat_four_block) + using assms by auto + +lemma row_append_rows: + assumes "dim_col A = dim_col B" + assumes "i < dim_row A + dim_row B" + shows "row (A @\<^sub>r B) i = (if i < dim_row A then row A i else row B (i-dim_row A))" + unfolding vec_eq_iff + using assms by (auto simp add: append_rows_def) + +lemma append_rows_mat_mul: + assumes "dim_col A = dim_col B" + shows "(A @\<^sub>r B) * C = A * C @\<^sub>r B * C" + unfolding mat_eq_iff + apply auto + apply (simp add: append_rows_def) + apply (subst index_mult_mat) + apply auto + apply (simp add: append_rows_def) + apply (subst append_rows_index) + apply auto + apply (simp add: append_rows_def) + apply (metis add.right_neutral append_rows_def assms index_mat_four_block(3) index_mult_mat(1) index_mult_mat(3) index_zero_mat(3) row_append_rows trans_less_add1) + by (metis add_cancel_right_right add_diff_inverse_nat append_rows_def assms index_mat_four_block(3) index_mult_mat(1) index_mult_mat(3) index_zero_mat(3) nat_add_left_cancel_less row_append_rows) + +lemma cardlt: + shows "card {i. i < (n::nat)} \ n" + by simp + +lemma row_echelon_form_zero_rows: + assumes row_ech: "row_echelon_form A" + assumes dim_asm: "dim_col A \ dim_row A" + shows "take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) = A" +proof - + have ex_pivot_fun: "\ f. pivot_fun A f (dim_col A)" using row_ech unfolding row_echelon_form_def by auto + have len_help: "length (pivot_positions A) = card {i. i < dim_row A \ row A i \ 0\<^sub>v (dim_col A)}" + using ex_pivot_fun pivot_positions[where A = "A",where nr = "dim_row A", where nc = "dim_col A"] + by auto + then have len_help2: "length (pivot_positions A) \ dim_row A" + by (metis (no_types, lifting) card_mono cardlt finite_Collect_less_nat le_trans mem_Collect_eq subsetI) + have fileq: "filter (\y. y < dim_row A) [0..< length (pivot_positions A)] = [0..n. card {i. i < n \ row A i \ 0\<^sub>v (dim_col A)} \ n" + proof clarsimp + fix n + have h: "\x. x \ {i. i < n \ row A i \ 0\<^sub>v (dim_col A)} \ x\{.. row A i \ 0\<^sub>v (dim_col A)} \ {.. row A i \ 0\<^sub>v (dim_col A)}::nat) \ (card {.. row A i \ 0\<^sub>v (dim_col A)}::nat) \ (n::nat)" using h2 card_lessThan[of n] + by auto + qed + then have pivot_len: "length (pivot_positions A) \ dim_row A " using len_help + by simp + have alt_char: "mat_of_rows (dim_col A) + (map ((!) (rows A)) (filter (\y. y < dim_col A) [0..i j. i < dim_row A \ + j < dim_col A \ + i < dim_row (take_rows A [0.. + take_rows A [0.. dim_row A" using pivot_len by auto + have h1: "take_rows A [0..i < dim_row A\ j_lt) + qed + let ?nc = "dim_col A" + let ?nr = "dim_row A" + have h2: "\i j. i < dim_row A \ + j < dim_col A \ + \ i < dim_row (take_rows A [0.. + 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) $$ + (i - dim_row (take_rows A [0.. i < dim_row (take_rows A [0..f. pivot_fun A f (dim_col A) \ f i = ?nc" + proof - + have half1: "\f. pivot_fun A f (dim_col A)" using assms unfolding row_echelon_form_def + by blast + have half2: "\f. pivot_fun A f (dim_col A) \ f i = ?nc " + proof clarsimp + fix f + assume is_piv: "pivot_fun A f (dim_col A)" + have len_pp: "length (pivot_positions A) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" using is_piv pivot_positions[of A ?nr ?nc f] + by auto + have "\i. (i < ?nr \ row A i \ 0\<^sub>v ?nc) \ (i < ?nr \ f i \ ?nc)" + using is_piv pivot_fun_zero_row_iff[of A f ?nc ?nr] + by blast + then have len_pp_var: "length (pivot_positions A) = card {i. i < ?nr \ f i \ ?nc}" + using len_pp by auto + have allj_hyp: "\j < ?nr. f j = ?nc \ ((Suc j) < ?nr \ f (Suc j) = ?nc)" + using is_piv unfolding pivot_fun_def + using lt_i + by (metis le_antisym le_less) + have if_then_bad: "f i \ ?nc \ (\j. j \ i \ f j \ ?nc)" + proof clarsimp + fix j + assume not_i: "f i \ ?nc" + assume j_leq: "j \ i" + assume bad_asm: "f j = ?nc" + have "\k. k \ j \ k < ?nr \ f k = ?nc" + proof - + fix k :: nat + assume a1: "j \ k" + assume a2: "k < dim_row A" + have f3: "\n. \ n < dim_row A \ f n \ f j \ \ Suc n < dim_row A \ f (Suc n) = f j" + using allj_hyp bad_asm by presburger + obtain nn :: "nat \ nat \ (nat \ bool) \ nat" where + f4: "\n na p nb nc. (\ n \ na \ Suc n \ Suc na) \ (\ p nb \ \ nc \ nb \ \ p (nn nc nb p) \ p nc) \ (\ p nb \ \ nc \ nb \ p nc \ p (Suc (nn nc nb p)))" + using inc_induct order_refl by moura + then have f5: "\p. \ p k \ p j \ p (Suc (nn j k p))" + using a1 by presburger + have f6: "\p. \ p k \ \ p (nn j k p) \ p j" + using f4 a1 by meson + { assume "nn j k (\n. n < dim_row A \ f n \ dim_col A) < dim_row A \ f (nn j k (\n. n < dim_row A \ f n \ dim_col A)) \ dim_col A" + moreover + { assume "(nn j k (\n. n < dim_row A \ f n \ dim_col A) < dim_row A \ f (nn j k (\n. n < dim_row A \ f n \ dim_col A)) \ dim_col A) \ (\ j < dim_row A \ f j = dim_col A)" + then have "\ k < dim_row A \ f k = dim_col A" + using f6 + by (metis (mono_tags, lifting)) } + ultimately have "(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A) \ \ k < dim_row A \ f k = dim_col A" + using bad_asm + by blast } + moreover + { assume "(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A)" + then have "\ k < dim_row A \ f k = dim_col A" + using f5 + proof - + have "\ (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) \ dim_col A) \ \ (j < dim_row A \ f j \ dim_col A)" + using \(\ j < dim_row A \ f j = dim_col A) \ (\ Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A)) < dim_row A \ f (Suc (nn j k (\n. n < dim_row A \ f n \ dim_col A))) = dim_col A)\ by linarith + then have "\ (k < dim_row A \ f k \ dim_col A)" + by (metis (mono_tags, lifting) a2 bad_asm f5 le_less) + then show ?thesis + by meson + qed } + ultimately show "f k = dim_col A" + using f3 a2 by (metis (lifting) Suc_lessD bad_asm) + qed + then show "False" using lt_i not_i + using j_leq by blast + qed + have "f i \ ?nc \ ({0.. {y. y < ?nr \ f y \ dim_col A})" + proof - + have h1: "f i \ dim_col A \ (\j\i. j < ?nr \ f j \ dim_col A)" + using if_then_bad lt_i by auto + then show ?thesis by auto + qed + then have gteq: "f i \ ?nc \ (card {i. i < ?nr \ f i \ dim_col A} \ (i+1))" + using card_lessThan[of ?ip] card_mono[where B = "{i. i < ?nr \ f i \ dim_col A} ", where A = "{0.. length (pivot_positions A)" using not_lt clear by auto + then show "f i = ?nc" using gteq len_pp_var by auto + qed + show ?thesis using half1 half2 + by blast + qed + then have h1a: "row A i = 0\<^sub>v (dim_col A)" + using pivot_fun_zero_row_iff[of A _ ?nc ?nr] + using lt_i by blast + then have h1: "A $$ (i, j) = 0" + using index_row(1) lt_i lt_j by fastforce + have h2a: "i - dim_row (take_rows A [0..m (dim_row A - length (pivot_positions A)) (dim_col A) $$ + (i - dim_row (take_rows A [0..m (dim_row A - length (pivot_positions A)) (dim_col A) $$ + (i - dim_row (take_rows A [0..i j. i < dim_row A \ + j < dim_col A \ + i < dim_row (take_rows A [0.. dim_row A" +proof - + have 1: "A \ carrier_mat (dim_row A) (dim_col A)" by auto + obtain f where 2: "pivot_fun A f (dim_col A)" + using assms row_echelon_form_def by blast + from pivot_positions(4)[OF 1 2] have + "length (pivot_positions A) = card {i. i < dim_row A \ row A i \ 0\<^sub>v (dim_col A)}" . + also have "... \ card {i. i < dim_row A}" + apply (rule card_mono) + by auto + ultimately show ?thesis by auto +qed + +lemma rref_pivot_positions: + assumes "row_echelon_form R" + assumes R: "R \ carrier_mat nr nc" + shows "\i j. (i,j) \ set (pivot_positions R) \ i < nr \ j < nc" +proof - + obtain f where f: "pivot_fun R f nc" + using assms(1) assms(2) row_echelon_form_def by blast + have *: "\i. i < nr \ f i \ nc" using f + using R pivot_funD(1) by blast + from pivot_positions[OF R f] + have "set (pivot_positions R) = {(i, f i) |i. i < nr \ f i \ nc}" by auto + then have **: "set (pivot_positions R) = {(i, f i) |i. i < nr \ f i < nc}" + using * + by fastforce + fix i j + assume "(i, j) \ set (pivot_positions R)" + thus "i < nr \ j < nc" using ** + by simp +qed + +lemma pivot_fun_monoton: + assumes pf: "pivot_fun A f (dim_col A)" + assumes dr: "dim_row A = nr" + shows "\ i. i < nr \ (\ k. ((k < nr \ i < k) \ f i \ f k))" +proof - + fix i + assume "i < nr" + show "(\ k. ((k < nr \ i < k) \ f i \ f k))" + proof - + fix k + show "((k < nr \ i < k) \ f i \ f k)" + proof (induct k) + case 0 + then show ?case + by blast + next + case (Suc k) + then show ?case + by (smt dr le_less_trans less_Suc_eq less_imp_le_nat pf pivot_funD(1) pivot_funD(3)) + qed + qed +qed + +lemma pivot_positions_contains: + assumes row_ech: "row_echelon_form A" + assumes dim_h: "dim_col A \ dim_row A" + assumes "pivot_fun A f (dim_col A)" + shows "\i < (length (pivot_positions A)). (i, f i) \ set (pivot_positions A)" +proof - + let ?nr = "dim_row A" + let ?nc = "dim_col A" + let ?pp = "pivot_positions A" + have i_nr: "\i < (length ?pp). i < ?nr" using rref_pivot_positions assms + using length_pivot_positions_dim_row less_le_trans by blast + have i_nc: "\i < (length ?pp). f i < ?nc" + proof clarsimp + fix i + assume i_lt: "i < length ?pp" + have fis_nc: "f i = ?nc \ (\ k > i. k < ?nr \ f k = ?nc)" + proof - + assume is_nc: "f i = ?nc" + show "(\ k > i. k < ?nr \ f k = ?nc)" + proof clarsimp + fix k + assume k_gt: "k > i" + assume k_lt: "k < ?nr" + have fk_lt: "f k \ ?nc" using pivot_funD(1)[of A ?nr f ?nc k] k_lt apply (auto) + using \pivot_fun A f (dim_col A)\ by blast + show "f k = ?nc" + using fk_lt is_nc k_gt k_lt assms pivot_fun_monoton[of A f ?nr i k] + using \pivot_fun A f (dim_col A)\ by auto + qed + qed + have ncimp: "f i = ?nc \ (\ k \i. k \ { i. i < ?nr \ row A i \ 0\<^sub>v ?nc})" + proof - + assume nchyp: "f i = ?nc" + show "(\ k \i. k \ { i. i < ?nr \ row A i \ 0\<^sub>v ?nc})" + proof clarsimp + fix k + assume i_lt: "i \ k" + assume k_lt: "k < dim_row A" + show "row A k = 0\<^sub>v (dim_col A) " + using i_lt k_lt fis_nc + using pivot_fun_zero_row_iff[of A f ?nc ?nr] + using \pivot_fun A f (dim_col A)\ le_neq_implies_less nchyp by blast + qed + qed + then have "f i = ?nc \ card { i. i < ?nr \ row A i \ 0\<^sub>v ?nc} \ i" + proof - + assume nchyp: "f i = ?nc" + have h: "{ i. i < ?nr \ row A i \ 0\<^sub>v ?nc} \ {0.. row A i \ 0\<^sub>v ?nc} \ i" + using card_lessThan + using subset_eq_atLeast0_lessThan_card by blast + qed + then show "f i < ?nc" using i_lt pivot_positions(4)[of A ?nr ?nc f] + apply (auto) + by (metis \pivot_fun A f (dim_col A)\ i_nr le_neq_implies_less not_less pivot_funD(1)) + qed + then show ?thesis + using pivot_positions(1) + by (smt \pivot_fun A f (dim_col A)\ carrier_matI i_nr less_not_refl mem_Collect_eq) +qed + +lemma pivot_positions_form_helper_1: + shows "(a, b) \ set (pivot_positions_main_gen z A nr nc i j) \ i \ a" +proof (induct i j rule: pivot_positions_main_gen.induct[of nr nc A z]) + case (1 i j) + then show ?case using pivot_positions_main_gen.simps[of z A nr nc i j] + apply (auto) + by (smt Suc_leD le_refl old.prod.inject set_ConsD) +qed + +lemma pivot_positions_form_helper_2: + shows "strict_sorted (map fst (pivot_positions_main_gen z A nr nc i j))" +proof (induct i j rule: pivot_positions_main_gen.induct[of nr nc A z]) + case (1 i j) + then show ?case using pivot_positions_main_gen.simps[of z A nr nc i j] + apply (auto) using pivot_positions_form_helper_1 + by (simp add: pivot_positions_form_helper_1 Suc_le_lessD) +qed + +lemma sorted_pivot_positions: + shows "strict_sorted (map fst (pivot_positions A))" + using pivot_positions_form_helper_2 + by (simp add: pivot_positions_form_helper_2 pivot_positions_gen_def) + +lemma pivot_positions_form: + assumes row_ech: "row_echelon_form A" + assumes dim_h: "dim_col A \ dim_row A" + shows "\ i < (length (pivot_positions A)). fst ((pivot_positions A) ! i) = i" +proof clarsimp + let ?nr = "dim_row A" + let ?nc = "dim_col A" + let ?pp = "pivot_positions A :: (nat \ nat) list" + fix i + assume i_lt: "i < length (pivot_positions A)" + have "\f. pivot_fun A f ?nc" using row_ech unfolding row_echelon_form_def + by blast + then obtain f where pf:"pivot_fun A f ?nc" + by blast + have all_f_in: "\i < (length ?pp). (i, f i) \ set ?pp" + using pivot_positions_contains pf + assms + by blast + have sorted_hyp: "\ (p::nat) (q::nat). p < (length ?pp) \ q < (length ?pp) \ p < q \ (fst (?pp ! p) < fst (?pp ! q))" + proof - + fix p::nat + fix q::nat + assume p_lt: "p < q" + assume p_welldef: "p < (length ?pp)" + assume q_welldef: "q < (length ?pp)" + show "fst (?pp ! p) < fst (?pp ! q)" + using sorted_pivot_positions p_lt p_welldef q_welldef apply (auto) + by (smt find_first_unique length_map nat_less_le nth_map p_welldef sorted_nth_mono sorted_pivot_positions strict_sorted_iff) + qed + have h: "i < (length ?pp) \ fst (pivot_positions A ! i) = i" + proof (induct i) + case 0 + have "\j. fst (pivot_positions A ! j) = 0" + by (metis all_f_in fst_conv i_lt in_set_conv_nth length_greater_0_conv list.size(3) not_less0) + then obtain j where jth:" fst (pivot_positions A ! j) = 0" + by blast + have "j \ 0 \ (fst (pivot_positions A ! 0) > 0 \ j \ 0)" + using sorted_hyp apply (auto) + by (metis all_f_in fst_conv i_lt in_set_conv_nth length_greater_0_conv list.size(3) neq0_conv not_less0) + then show ?case + using jth neq0_conv by blast + next + case (Suc i) + have ind_h: "i < length (pivot_positions A) \ fst (pivot_positions A ! i) = i" + using Suc.hyps by blast + have thesis_h: "(Suc i) < length (pivot_positions A) \ fst (pivot_positions A ! (Suc i)) = (Suc i)" + proof - + assume suc_i_lt: "(Suc i) < length (pivot_positions A)" + have fst_i_is: "fst (pivot_positions A ! i) = i" using suc_i_lt ind_h + using Suc_lessD by blast + have "(\j < (length ?pp). fst (pivot_positions A ! j) = (Suc i))" + by (metis suc_i_lt all_f_in fst_conv in_set_conv_nth) + then obtain j where jth: "j < (length ?pp) \ fst (pivot_positions A ! j) = (Suc i)" + by blast + have "j > i" + using sorted_hyp apply (auto) + by (metis Suc_lessD \fst (pivot_positions A ! i) = i\ jth less_not_refl linorder_neqE_nat n_not_Suc_n suc_i_lt) + have "j > (Suc i) \ False" + proof - + assume j_gt: "j > (Suc i)" + then have h1: "fst (pivot_positions A ! (Suc i)) > i" + using fst_i_is sorted_pivot_positions + using sorted_hyp suc_i_lt by force + have "fst (pivot_positions A ! j) > fst (pivot_positions A ! (Suc i))" + using jth j_gt sorted_hyp apply (auto) + by fastforce + then have h2: "fst (pivot_positions A ! (Suc i)) < (Suc i)" + using jth + by simp + show "False" using h1 h2 + using not_less_eq by blast + qed + show "fst (pivot_positions A ! (Suc i)) = (Suc i)" + using Suc_lessI \Suc i < j \ False\ \i < j\ jth by blast + qed + then show ?case + by blast + qed + then show "fst (pivot_positions A ! i) = i" + using i_lt by auto +qed + +lemma take_cols_pivot_eq: + assumes row_ech: "row_echelon_form A" + assumes dim_h: "dim_col A \ dim_row A" + shows "take_cols A (map snd (pivot_positions A)) = + 1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))" +proof - + let ?nr = "dim_row A" + let ?nc = "dim_col A" + have h1: " dim_col + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) = (length (pivot_positions A))" + by (simp add: append_rows_def) + have len_pivot: "length (pivot_positions A) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" + using row_ech pivot_positions(4) row_echelon_form_def by blast + have pp_leq_nc: "\f. pivot_fun A f ?nc \ (\i < ?nr. f i \ ?nc)" unfolding pivot_fun_def + by meson + have pivot_set: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ f i \ ?nc}" + using row_ech row_echelon_form_def pivot_positions(1) + by (smt (verit) Collect_cong carrier_matI) + then have pivot_set_alt: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" + using pivot_positions pivot_fun_zero_row_iff Collect_cong carrier_mat_triv + by (smt (verit, best)) + have "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. f i \ ?nc \ i < ?nr \ f i \ ?nc}" + using pivot_set pp_leq_nc by auto + then have pivot_set_var: "\f. pivot_fun A f ?nc \ set (pivot_positions A) = {(i, f i) | i. i < ?nr \ f i < ?nc}" + by auto + have "length (map snd (pivot_positions A)) = card (set (map snd (pivot_positions A)))" + using row_ech row_echelon_form_def pivot_positions(3) distinct_card[where xs = "map snd (pivot_positions A)"] + by (metis carrier_mat_triv) + then have "length (map snd (pivot_positions A)) = card (set (pivot_positions A))" + by (metis card_distinct distinct_card distinct_map length_map) + then have "length (map snd (pivot_positions A)) = card {i. i < ?nr \ row A i \ 0\<^sub>v ?nc}" + using pivot_set_alt + by (simp add: len_pivot) + then have length_asm: "length (map snd (pivot_positions A)) = length (pivot_positions A)" + using len_pivot by linarith + then have "\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A" + proof clarsimp + fix a + assume a_in: "List.member (map snd (pivot_positions A)) a" + have "\v \ set (pivot_positions A). a = snd v" + using a_in in_set_member[where xs = "(pivot_positions A)"] apply (auto) + by (metis in_set_impl_in_set_zip2 in_set_member length_map snd_conv zip_map_fst_snd) + then show "a < dim_col A" + using pivot_set_var in_set_member by auto + qed + then have h2b: "(filter (\y. y < dim_col A) (map snd (pivot_positions A))) = (map snd (pivot_positions A))" + by (meson filter_True in_set_member) + then have h2a: "length (map ((!) (cols A)) (filter (\y. y < dim_col A) (map snd (pivot_positions A)))) = length (pivot_positions A)" + using length_asm + by (simp add: h2b) + then have h2: "length (pivot_positions A) \ dim_row A \ + dim_col (take_cols A (map snd (pivot_positions A))) = (length (pivot_positions A))" + unfolding take_cols_def using mat_of_cols_carrier by auto + have h_len: "length (pivot_positions A) \ dim_row A \ + dim_col (take_cols A (map snd (pivot_positions A))) = + dim_col + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A)))" + using h1 h2 + by (simp add: h1 assms length_pivot_positions_dim_row) + have h2: "\i j. length (pivot_positions A) \ dim_row A \ + i < dim_row A \ + j < dim_col + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) \ + take_cols A (map snd (pivot_positions A)) $$ (i, j) = + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ + (i, j)" + proof - + fix i + fix j + let ?pp = "(pivot_positions A)" + assume len_lt: "length (pivot_positions A) \ dim_row A" + assume i_lt: " i < dim_row A" + assume j_lt: "j < dim_col + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A)))" + let ?w = "((map snd (pivot_positions A)) ! j)" + have breaking_it_down: "mat_of_cols (dim_row A) + (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) + = ((cols A) ! ?w) $ i" + apply (auto) + by (metis comp_apply h1 i_lt j_lt length_map mat_of_cols_index nth_map) + have h1a: "i < (length ?pp) \ (mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) + = (1\<^sub>m (length (pivot_positions A))) $$ (i, j))" + proof - + (* need to, using row_ech, rely heavily on pivot_fun_def, that num_cols \ num_rows, and row_echelon form*) + assume "i < (length ?pp)" + have "\f. pivot_fun A f ?nc" using row_ech unfolding row_echelon_form_def + by blast + then obtain f where "pivot_fun A f ?nc" + by blast + have j_nc: "j < (length ?pp)" using j_lt + by (simp add: h1) + then have j_lt_nr: "j < ?nr" using dim_h + using len_lt by linarith + then have is_this_true: "(pivot_positions A) ! j = (j, f j)" + using pivot_positions_form pivot_positions(1)[of A ?nr ?nc f] + proof - + have "pivot_positions A ! j \ set (pivot_positions A)" + using j_nc nth_mem by blast + then have "\n. pivot_positions A ! j = (n, f n) \ n < dim_row A \ f n \ dim_col A" + using \\A \ carrier_mat (dim_row A) (dim_col A); pivot_fun A f (dim_col A)\ \ set (pivot_positions A) = {(i, f i) |i. i < dim_row A \ f i \ dim_col A}\ \pivot_fun A f (dim_col A)\ by blast + then show ?thesis + by (metis (no_types) \\A. \row_echelon_form A; dim_row A \ dim_col A\ \ \i dim_h fst_conv j_nc row_ech) + qed + then have w_is: "?w = f j" + by (metis h1 j_lt nth_map snd_conv) + have h0: "i = j \ ((cols A) ! ?w) $ i = 1" using w_is pivot_funD(4)[of A ?nr f ?nc i] + by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ \i < length (pivot_positions A)\ \pivot_fun A f (dim_col A)\ cols_length i_lt in_set_member length_asm mat_of_cols_cols mat_of_cols_index nth_mem) + have h1: "i \ j \ ((cols A) ! ?w) $ i = 0" using w_is pivot_funD(5) + by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ \pivot_fun A f (dim_col A)\ cols_length h1 i_lt in_set_member j_lt len_lt length_asm less_le_trans mat_of_cols_cols mat_of_cols_index nth_mem) + show "(mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) + = (1\<^sub>m (length (pivot_positions A))) $$ (i, j))" using h0 h1 breaking_it_down + by (metis \i < length (pivot_positions A)\ h2 h_len index_one_mat(1) j_lt len_lt) + qed + have h1b: "i \ (length ?pp) \ (mat_of_cols (dim_row A) (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = 0)" + proof - + assume i_gt: "i \ (length ?pp)" + have h0a: "((cols A) ! ((map snd (pivot_positions A)) ! j)) $ i = (row A i) $ ?w" + by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ cols_length h1 i_lt in_set_member index_row(1) j_lt length_asm mat_of_cols_cols mat_of_cols_index nth_mem) + have h0b: + "take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A) = A" + using assms row_echelon_form_zero_rows[of A] + by blast + then have h0c: "(row A i) = 0\<^sub>v (dim_col A)" using i_gt + using add_diff_cancel_right' add_less_cancel_left diff_is_0_eq' dim_col_take_rows dim_row_append_rows i_lt index_zero_mat(2) index_zero_mat(3) le_add_diff_inverse len_lt less_not_refl3 row_append_rows row_zero zero_less_diff + by (smt add_diff_cancel_right' add_less_cancel_left diff_is_0_eq' dim_col_take_rows dim_row_append_rows i_lt index_zero_mat(2) index_zero_mat(3) le_add_diff_inverse len_lt less_not_refl3 row_append_rows row_zero zero_less_diff) + then show ?thesis using h0a breaking_it_down apply (auto) + by (metis \\a. List.member (map snd (pivot_positions A)) a \ a < dim_col A\ h1 in_set_member index_zero_vec(1) j_lt length_asm nth_mem) + qed + have h1: " mat_of_cols (dim_row A) + (map ((!) (cols A)) (map snd (pivot_positions A))) $$ (i, j) = + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ + (i, j) " using h1a h1b + apply (auto) + by (smt add_diff_inverse_nat add_less_cancel_left append_rows_index h1 i_lt index_one_mat(2) index_one_mat(3) index_zero_mat(1) index_zero_mat(2) index_zero_mat(3) j_lt len_lt not_less) + then show "take_cols A (map snd (pivot_positions A)) $$ (i, j) = + (1\<^sub>m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))) $$ + (i, j)" + unfolding take_cols_def + by (simp add: h2b) + qed + show ?thesis + unfolding mat_eq_iff + using length_pivot_positions_dim_row[OF assms(1)] h_len h2 by auto +qed + +lemma rref_right_mul: + assumes "row_echelon_form A" + assumes "dim_col A \ dim_row A" + shows + "take_cols A (map snd (pivot_positions A)) * take_rows A [0..m (length (pivot_positions A)) @\<^sub>r + 0\<^sub>m (dim_row A - length (pivot_positions A)) (length (pivot_positions A))" . + have 2: "take_cols A (map snd (pivot_positions A)) * take_rows A [0..r 0\<^sub>m (dim_row A - length (pivot_positions A)) (dim_col A)" + unfolding 1 + apply (auto simp add: append_rows_mat_mul) + by (smt add_diff_cancel_right' assms diff_diff_cancel dim_col_take_rows dim_row_append_rows index_zero_mat(2) left_mult_one_mat' left_mult_zero_mat' length_pivot_positions_dim_row row_echelon_form_zero_rows) + from row_echelon_form_zero_rows[OF assms] have + "... = A" . + thus ?thesis + by (simp add: "2") +qed + +context conjugatable_vec_space begin + +lemma lin_indpt_id: + shows "lin_indpt (set (cols (1\<^sub>m n)::'a vec list))" +proof - + have *: "set (cols (1\<^sub>m n)) = set (rows (1\<^sub>m n))" + by (metis cols_transpose transpose_one) + have "det (1\<^sub>m n) \ 0" using det_one by auto + from det_not_0_imp_lin_indpt_rows[OF _ this] + have "lin_indpt (set (rows (1\<^sub>m n)))" + using one_carrier_mat by blast + thus ?thesis + by (simp add: *) +qed + +lemma lin_indpt_take_cols_id: + shows "lin_indpt (set (cols (take_cols (1\<^sub>m n) inds)))" +proof - + have subset_h: "set (cols (take_cols (1\<^sub>m n) inds)) \ set (cols (1\<^sub>m n)::'a vec list)" + using cols_take_cols_subset by blast + then show ?thesis using lin_indpt_id subset_li_is_li by auto +qed + +lemma cols_id_unit_vecs: + shows "cols (1\<^sub>m d) = unit_vecs d" + unfolding unit_vecs_def list_eq_iff_nth_eq + by auto + +lemma distinct_cols_id: + shows "distinct (cols (1\<^sub>m d)::'a vec list)" + by (simp add: conjugatable_vec_space.cols_id_unit_vecs vec_space.unit_vecs_distinct) + +lemma distinct_map_nth: + assumes "distinct ls" + assumes "distinct inds" + assumes "\j. j \ set inds \ j < length ls" + shows "distinct (map ((!) ls) inds)" + by (simp add: assms(1) assms(2) assms(3) distinct_map inj_on_nth) + +lemma distinct_take_cols_id: + assumes "distinct inds" + shows "distinct (cols (take_cols (1\<^sub>m n) inds) :: 'a vec list)" + unfolding take_cols_def + apply (subst cols_mat_of_cols) + apply (auto intro!: distinct_map_nth simp add: distinct_cols_id) + using assms distinct_filter by blast + +lemma rank_take_cols: + assumes "distinct inds" + shows "rank (take_cols (1\<^sub>m n) inds) = length (filter ((>) n) inds)" + apply (subst lin_indpt_full_rank[of _ "length (filter ((>) n) inds)"]) + apply (auto simp add: lin_indpt_take_cols_id) + apply (metis (full_types) index_one_mat(2) index_one_mat(3) length_map mat_of_cols_carrier(1) take_cols_def) + by (simp add: assms distinct_take_cols_id) + +lemma rank_mul_left_invertible_mat: + fixes A::"'a mat" + assumes "invertible_mat A" + assumes "A \ carrier_mat n n" + assumes "B \ carrier_mat n nc" + shows "rank (A * B) = rank B" +proof - + obtain C where C: "inverts_mat A C" "inverts_mat C A" + using assms invertible_mat_def by blast + from C have ceq: "C * A = 1\<^sub>m n" + by (metis assms(2) carrier_matD(2) index_mult_mat(3) index_one_mat(3) inverts_mat_def) + then have *:"B = C*A*B" + using assms(3) by auto + from rank_mat_mul_left[OF assms(2-3)] + have **: "rank (A*B) \ rank B" . + have 1: "C \ carrier_mat n n" using C ceq + by (metis assms(2) carrier_matD(1) carrier_matI index_mult_mat(3) index_one_mat(3) inverts_mat_def) + have 2: "A * B \ carrier_mat n nc" using assms by auto + have "rank B = rank (C* A * B)" using * by auto + also have "... \ rank (A*B)" using rank_mat_mul_left[OF 1 2] + using "1" assms(2) assms(3) by auto + ultimately show ?thesis using ** by auto +qed + +lemma invertible_take_cols_rank: + fixes A::"'a mat" + assumes "invertible_mat A" + assumes "A \ carrier_mat n n" + assumes "distinct inds" + shows "rank (take_cols A inds) = length (filter ((>) n) inds)" +proof - + have " A = A * 1\<^sub>m n" using assms(2) by auto + then have "take_cols A inds = A * take_cols (1\<^sub>m n) inds" + by (metis assms(2) one_carrier_mat take_cols_mat_mul) + then have "rank (take_cols A inds) = rank (take_cols (1\<^sub>m n) inds)" + by (metis assms(1) assms(2) conjugatable_vec_space.rank_mul_left_invertible_mat one_carrier_mat take_cols_carrier_mat) + thus ?thesis + by (simp add: assms(3) conjugatable_vec_space.rank_take_cols) +qed + +lemma rank_take_cols_leq: + assumes R:"R \ carrier_mat n nc" + shows "rank (take_cols R ls) \ rank R" +proof - + from take_cols_mat_mul[OF R] + have "take_cols R ls = R * take_cols (1\<^sub>m nc) ls" + by (metis assms one_carrier_mat right_mult_one_mat) + thus ?thesis + by (metis assms one_carrier_mat take_cols_carrier_mat vec_space.rank_mat_mul_right) +qed + +lemma rank_take_cols_geq: + assumes R:"R \ carrier_mat n nc" + assumes t:"take_cols R ls \ carrier_mat n r" + assumes B:"B \ carrier_mat r nc" + assumes "R = (take_cols R ls) * B" + shows "rank (take_cols R ls) \ rank R" + by (metis B assms(4) t vec_space.rank_mat_mul_right) + +lemma rref_drop_pivots: + assumes row_ech: "row_echelon_form R" + assumes dims: "R \ carrier_mat n nc" + assumes order: "nc \ n" + shows "rank (take_cols R (map snd (pivot_positions R))) = rank R" +proof - + let ?B = "take_rows R [0..r. take_cols R (map snd (pivot_positions R)) \ carrier_mat n r \ ?B \ carrier_mat r nc" + proof - + have h1: + "take_cols R (map snd (pivot_positions R)) \ carrier_mat n (length (pivot_positions R))" + using assms + by (metis in_set_impl_in_set_zip2 length_map rref_pivot_positions take_cols_carrier_mat_strict zip_map_fst_snd) + have "\ f. pivot_fun R f nc" using row_ech unfolding row_echelon_form_def using dims + by blast + then have "length (pivot_positions R) = card {i. i < n \ row R i \ 0\<^sub>v nc}" + using pivot_positions[of R n nc] + using dims by auto + then have "nc \ length (pivot_positions R)" using order + using carrier_matD(1) dims dual_order.trans length_pivot_positions_dim_row row_ech by blast + then have "dim_col R \ length (pivot_positions R)" using dims by auto + then have h2: "?B \ carrier_mat (length (pivot_positions R)) nc" unfolding take_rows_def + using dims + by (smt atLeastLessThan_iff carrier_matD(2) filter_True le_eq_less_or_eq length_map length_pivot_positions_dim_row less_trans map_nth mat_of_cols_carrier(1) row_ech set_upt transpose_carrier_mat transpose_mat_of_rows) + show ?thesis using h1 h2 + by blast + qed + (* prove the other two dimensionality assumptions *) + have "rank R \ rank (take_cols R (map snd (pivot_positions R)))" + using dims ex_r rank_take_cols_geq[where R = "R", where B = "?B", where ls = "(map snd (pivot_positions R))", where nc = "nc"] + using equa by blast + thus ?thesis + using assms(2) conjugatable_vec_space.rank_take_cols_leq le_antisym by blast +qed + +lemma gjs_and_take_cols_var: + fixes A::"'a mat" + assumes A:"A \ carrier_mat n nc" + assumes order: "nc \ n" + shows "(take_cols A (map snd (pivot_positions (gauss_jordan_single A)))) = + (take_cols_var A (map snd (pivot_positions (gauss_jordan_single A))))" +proof - + let ?gjs = "(gauss_jordan_single A)" + have "\x. List.member (map snd (pivot_positions (gauss_jordan_single A))) x \ x \ dim_col A" + using rref_pivot_positions gauss_jordan_single(3) carrier_matD(2) gauss_jordan_single(2) in_set_impl_in_set_zip2 in_set_member length_map less_irrefl less_trans not_le_imp_less zip_map_fst_snd + by (smt A carrier_matD(2) gauss_jordan_single(2) in_set_impl_in_set_zip2 in_set_member length_map less_irrefl less_trans not_le_imp_less zip_map_fst_snd) + then have "(filter (\y. y < dim_col A) (map snd (pivot_positions (gauss_jordan_single A)))) = + (map snd (pivot_positions (gauss_jordan_single A)))" + by (metis (no_types, lifting) A carrier_matD(2) filter_True gauss_jordan_single(2) gauss_jordan_single(3) in_set_impl_in_set_zip2 length_map rref_pivot_positions zip_map_fst_snd) + then show ?thesis unfolding take_cols_def take_cols_var_def + by simp +qed + +lemma gauss_jordan_single_rank: + fixes A::"'a mat" + assumes A:"A \ carrier_mat n nc" + assumes order: "nc \ n" + shows "rank (take_cols A (map snd (pivot_positions (gauss_jordan_single A)))) = rank A" +proof - + let ?R = "gauss_jordan_single A" + obtain P where P:"P\Units (ring_mat TYPE('a) n undefined)" and + i: "?R = P * A" using gauss_jordan_transform[OF A] + using A assms det_mult det_non_zero_imp_unit det_one gauss_jordan_single(4) mult_not_zero one_neq_zero + by (smt A assms det_mult det_non_zero_imp_unit det_one gauss_jordan_single(4) mult_not_zero one_neq_zero) + have pcarrier: "P \ carrier_mat n n" using P unfolding Units_def + by (auto simp add: ring_mat_def) + have "invertible_mat P" using P unfolding invertible_mat_def Units_def inverts_mat_def + apply auto + apply (simp add: ring_mat_simps(5)) + by (metis index_mult_mat(2) index_one_mat(2) ring_mat_simps(1) ring_mat_simps(3)) + then + obtain Pi where Pi: "invertible_mat Pi" "Pi * P = 1\<^sub>m n" + proof - + assume a1: "\Pi. \invertible_mat Pi; Pi * P = 1\<^sub>m n\ \ thesis" + have "dim_row P = n" + by (metis (no_types) A assms(1) carrier_matD(1) gauss_jordan_single(2) i index_mult_mat(2)) + then show ?thesis + using a1 by (metis (no_types) \invertible_mat P\ index_mult_mat(3) index_one_mat(3) invertible_mat_def inverts_mat_def square_mat.simps) + qed + then have pi_carrier:"Pi \ carrier_mat n n" + by (metis carrier_mat_triv index_mult_mat(2) index_one_mat(2) invertible_mat_def square_mat.simps) + have R1:"row_echelon_form ?R" + using assms(2) gauss_jordan_single(3) by blast + have R2: "?R \ carrier_mat n nc" + using A assms(2) gauss_jordan_single(2) by auto + have Rcm: "take_cols ?R (map snd (pivot_positions ?R)) + \ carrier_mat n (length (map snd (pivot_positions ?R)))" + apply (rule take_cols_carrier_mat_strict[OF R2]) + using rref_pivot_positions[OF R1 R2] by auto + have "Pi * ?R = A" using i Pi + by (smt A \invertible_mat P\ assoc_mult_mat carrier_mat_triv index_mult_mat(2) index_mult_mat(3) index_one_mat(3) invertible_mat_def left_mult_one_mat square_mat.simps) + then have "rank (take_cols A (map snd (pivot_positions ?R))) = rank (take_cols (Pi * ?R) (map snd (pivot_positions ?R)))" + by auto + also have "... = rank ( Pi * take_cols ?R (map snd (pivot_positions ?R)))" + by (metis A gauss_jordan_single(2) pi_carrier take_cols_mat_mul) + also have "... = rank (take_cols ?R (map snd (pivot_positions ?R)))" + by (intro rank_mul_left_invertible_mat[OF Pi(1) pi_carrier Rcm]) + also have "... = rank ?R" + using assms(2) conjugatable_vec_space.rref_drop_pivots gauss_jordan_single(3) + using R1 R2 by blast + ultimately show ?thesis + using A \P \ carrier_mat n n\ \invertible_mat P\ conjugatable_vec_space.rank_mul_left_invertible_mat i + by auto +qed + +lemma lin_indpt_subset_cols: + fixes A:: "'a mat" + fixes B:: "'a vec set" + assumes "A \ carrier_mat n n" + assumes inv: "invertible_mat A" + assumes "B \ set (cols A)" + shows "lin_indpt B" +proof - + have "det A \ 0" + using assms(1) inv invertible_det by blast + then have "lin_indpt (set (rows A\<^sup>T))" + using assms(1) idom_vec.lin_dep_cols_imp_det_0 by auto + thus ?thesis using subset_li_is_li assms(3) + by auto +qed + +lemma rank_invertible_subset_cols: + fixes A:: "'a mat" + fixes B:: "'a vec list" + assumes inv: "invertible_mat A" + assumes A_square: "A \ carrier_mat n n" + assumes set_sub: "set (B) \ set (cols A)" + assumes dist_B: "distinct B" + shows "rank (mat_of_cols n B) = length B" +proof - + let ?B_mat = "(mat_of_cols n B)" + have h1: "lin_indpt (set(B))" + using assms lin_indpt_subset_cols[of A] by auto + have "set B \ carrier_vec n" + using set_sub A_square cols_dim[of A] by auto + then have cols_B: "cols (mat_of_cols n B) = B" using cols_mat_of_cols by auto + then have "maximal (set B) (\T. T \ set (B) \ lin_indpt T)" using h1 + by (simp add: maximal_def subset_antisym) + then have h2: "maximal (set B) (\T. T \ set (cols (mat_of_cols n B)) \ lin_indpt T)" + using cols_B by auto + have h3: "rank (mat_of_cols n B) = card (set B)" + using h1 h2 rank_card_indpt[of ?B_mat] + using mat_of_cols_carrier(1) by blast + then show ?thesis using assms distinct_card by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/ROOT b/thys/BenOr_Kozen_Reif/ROOT new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/ROOT @@ -0,0 +1,11 @@ +chapter AFP + +session BenOr_Kozen_Reif (AFP) = Algebraic_Numbers + + options [timeout = 600] + sessions + Sturm_Tarski + theories + BKR_Decision + Renegar_Decision + document_files + "root.tex" diff --git a/thys/BenOr_Kozen_Reif/Renegar_Algorithm.thy b/thys/BenOr_Kozen_Reif/Renegar_Algorithm.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/Renegar_Algorithm.thy @@ -0,0 +1,108 @@ +theory Renegar_Algorithm + imports BKR_Algorithm +begin + +(* There is significant overlap between Renegar's algorithm and BKR's. + However, the RHS vector is constructed differently in Renegar. The base case is also different. + In general, the _R's on definition and lemma names in this file are to emphasize that we are + working with Renegar style. +*) + +definition construct_NofI_R:: "real poly \ real poly list \ real poly list \ rat" + where "construct_NofI_R p I1 I2 = ( + let new_p = sum_list (map (\x. x^2) (p # I1)) in + rat_of_int (changes_R_smods new_p ((pderiv new_p)*(prod_list I2))))" + +(* Renegar's RHS vector will have type (nat list * nat list) list. + Note the change from BKR, where the RHS vector had type nat list list*) +definition construct_rhs_vector_R:: "real poly \ real poly list \ (nat list * nat list) list \ rat vec" + where "construct_rhs_vector_R p qs Is = + vec_of_list (map (\(I1,I2). + (construct_NofI_R p (retrieve_polys qs I1) (retrieve_polys qs I2))) Is)" + +section "Base Case" + +(* Renegar's matrix is 3x3 instead of 2x2 *) +definition base_case_info_R:: "(rat mat \ ((nat list * nat list) list \ rat list list))" + where "base_case_info_R = + ((mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]),([([], []),([0], []),([], [0])], [[1],[0],[-1]]))" + +(* When p, q are coprime, this will actually be an int vec, which is why taking the floor is okay *) +definition base_case_solve_for_lhs:: "real poly \ real poly \ rat vec" + where "base_case_solve_for_lhs p q = (mult_mat_vec (mat_of_rows_list 3 [[1/2, -1/2, 1/2], [0, 1, 0], [1/2, -1/2, -1/2]]) (construct_rhs_vector_R p [q] [([], []),([0], []),([], [0])]))" + +(* Solve for LHS in general: mat_inverse returns an option type, and we pattern match on this. + Notice that when we call this function in the algorithm, the matrix we pass will always be invertible, + given how the construction works. *) +definition solve_for_lhs_R:: "real poly \ real poly list \ (nat list * nat list) list \ rat mat \ rat vec" + where "solve_for_lhs_R p qs subsets matr = + mult_mat_vec (matr_option (dim_row matr) (mat_inverse_var matr)) (construct_rhs_vector_R p qs subsets)" + +section "Smashing" + +definition subsets_smash_R::"nat \ (nat list*nat list) list \ (nat list*nat list) list \ (nat list*nat list) list" + where "subsets_smash_R n s1 s2 = concat (map (\l1. map (\ l2. (((fst l1) @ (map ((+) n) (fst l2))), (snd l1) @ (map ((+) n) (snd l2)))) s2) s1)" + +definition smash_systems_R:: "real poly \ real poly list \ real poly list \ (nat list * nat list) list \ (nat list * nat list) list \ + rat list list \ rat list list \ rat mat \ rat mat \ + real poly list \ (rat mat \ ((nat list * nat list) list \ rat list list))" + where "smash_systems_R p qs1 qs2 subsets1 subsets2 signs1 signs2 mat1 mat2 = + (qs1@qs2, (kronecker_product mat1 mat2, (subsets_smash_R (length qs1) subsets1 subsets2, signs_smash signs1 signs2)))" + +fun combine_systems_R:: "real poly \ (real poly list \ (rat mat \ ((nat list * nat list) list \ rat list list))) \ (real poly list \ (rat mat \ ((nat list * nat list) list \ rat list list))) + \ (real poly list \ (rat mat \ ((nat list * nat list) list \ rat list list)))" + where "combine_systems_R p (qs1, m1, sub1, sgn1) (qs2, m2, sub2, sgn2) = + (smash_systems_R p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2)" + +(* 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! *) +section "Reduction" + +fun reduction_step_R:: "rat mat \ rat list list \ (nat list*nat list) list \ rat vec \ rat mat \ ((nat list*nat list) list \ rat list list)" + where "reduction_step_R A signs subsets lhs_vec = + (let reduce_cols_A = (reduce_mat_cols A lhs_vec); + rows_keep = rows_to_keep reduce_cols_A in + (take_rows_from_matrix reduce_cols_A rows_keep, + (take_indices subsets rows_keep, + take_indices signs (find_nonzeros_from_input_vec lhs_vec))))" + +fun reduce_system_R:: "real poly \ (real poly list \ (rat mat \ ((nat list*nat list) list \ rat list list))) \ (rat mat \ ((nat list*nat list) list \ rat list list))" + where "reduce_system_R p (qs,m,subs,signs) = + reduction_step_R m signs subs (solve_for_lhs_R p qs subs m)" + +section "Overall algorithm " +(* Find matrix, subsets, signs. + The "rat mat" in the output is the matrix. The "(nat list*nat list) list" is the list of subsets. + The "rat list list" is the list of signs. +*) +fun calculate_data_R:: "real poly \ real poly list \ (rat mat \ ((nat list*nat list) list \ rat list list))" + where + "calculate_data_R p qs = + ( let len = length qs in + if len = 0 then + (\(a,b,c).(a,b,map (drop 1) c)) (reduce_system_R p ([1],base_case_info_R)) + else if len \ 1 then reduce_system_R p (qs,base_case_info_R) + else + (let q1 = take (len div 2) qs; left = calculate_data_R p q1; + q2 = drop (len div 2) qs; right = calculate_data_R p q2; + comb = combine_systems_R p (q1,left) (q2,right) in + reduce_system_R p comb + ) + )" + +(* Extract the list of consistent sign assignments *) +definition find_consistent_signs_at_roots_R:: "real poly \ real poly list \ rat list list" + where [code]: + "find_consistent_signs_at_roots_R p qs = + ( let (M,S,\) = calculate_data_R p qs in \ )" + +lemma find_consistent_signs_at_roots_thm_R: + shows "find_consistent_signs_at_roots_R p qs = snd (snd (calculate_data_R p qs))" + by (simp add: case_prod_beta find_consistent_signs_at_roots_R_def) + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/Renegar_Decision.thy b/thys/BenOr_Kozen_Reif/Renegar_Decision.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/Renegar_Decision.thy @@ -0,0 +1,809 @@ +theory Renegar_Decision + imports "Renegar_Proofs" + "BKR_Decision" +begin + +(* Note that there is significant overlap between Renegar and BKR in general, so there is some + similarity between this file and BKR_Decision.thy. However, there are also notable differences + as Renegar and BKR use different auxiliary polynomials in their decision procedures. + + In general, the _R's on definition and lemma names in this file are to emphasize that we are + working with Renegar style. + +*) + +section "Algorithm" + +(* The set of all rational sign vectors for qs wrt the set S + When S = UNIV, then this quantifies over all reals *) +definition consistent_sign_vectors_R::"real poly list \ real set \ rat list set" + where "consistent_sign_vectors_R qs S = (consistent_sign_vec qs) ` S" + +primrec prod_list_var:: "real poly list \ real poly" + where "prod_list_var [] = 1" + | "prod_list_var (h#T) = (if h = 0 then (prod_list_var T) else (h* prod_list_var T))" + +primrec check_all_const_deg:: "real poly list \ bool" + where "check_all_const_deg [] = True" + | "check_all_const_deg (h#T) = (if degree h = 0 then (check_all_const_deg T) else False)" + +definition poly_f :: "real poly list \ real poly" + where + "poly_f ps = + (if (check_all_const_deg ps = True) then [:0, 1:] else + (pderiv (prod_list_var ps)) * (prod_list_var ps)* ([:-(crb (prod_list_var ps)),1:]) * ([:(crb (prod_list_var ps)),1:]))" + +definition find_consistent_signs_R :: "real poly list \ rat list list" + where + "find_consistent_signs_R ps = find_consistent_signs_at_roots_R (poly_f ps) ps" + +definition decide_universal_R :: "real poly fml \ bool" + where [code]: + "decide_universal_R fml = ( + let (fml_struct,polys) = convert fml; + conds = find_consistent_signs_R polys + in + list_all (lookup_sem fml_struct) conds + )" + +definition decide_existential_R :: "real poly fml \ bool" + where [code]: + "decide_existential_R fml = ( + let (fml_struct,polys) = convert fml; + conds = find_consistent_signs_R polys + in + find (lookup_sem fml_struct) conds \ None + )" + +subsection "Proofs" +definition roots_of_poly_f:: "real poly list \ real set" + where "roots_of_poly_f qs = {x. poly (poly_f qs) x = 0}" + +lemma prod_list_var_nonzero: + shows "prod_list_var qs \ 0" +proof (induct qs) + case Nil + then show ?case by auto +next + case (Cons a qs) + then show ?case by auto +qed + +lemma q_dvd_prod_list_var_prop: + assumes "q \ set qs" + assumes "q \ 0" + shows "q dvd prod_list_var qs" using assms +proof (induct qs) + case Nil + then show ?case by auto +next + case (Cons a qs) + then have eo: "q = a \q \ set qs" by auto + have c1: "q = a \ q dvd prod_list_var (a#qs)" + proof - + assume "q = a" + then have "prod_list_var (a#qs) = q*(prod_list_var qs)" using Cons.prems + unfolding prod_list_var_def by auto + then show ?thesis using prod_list_var_nonzero[of qs] by auto + qed + have c2: "q \ set qs \ q dvd prod_list_var qs" + using Cons.prems Cons.hyps unfolding prod_list_var_def by auto + show ?case using eo c1 c2 by auto +qed + + +lemma check_all_const_deg_prop: + shows "check_all_const_deg l = True \ (\p \ set(l). degree p = 0)" +proof (induct l) + case Nil + then show ?case by auto +next + case (Cons a l) + then show ?case by auto +qed + +(* lemma prod_zero shows that the product of the polynomial list is 0 at x iff there is a polynomial + in the list that is 0 at x *) +lemma poly_f_nonzero: + fixes qs :: "real poly list" + shows "(poly_f qs) \ 0" +proof - + have eo: "(\p \ set qs. degree p = 0) \ (\p \ set qs. degree p > 0)" + by auto + have c1: "(\p \ set qs. degree p = 0) \ (poly_f qs) \ 0" + unfolding poly_f_def using check_all_const_deg_prop by auto + have c2: "(\p \ set qs. degree p > 0) \ (poly_f qs) \ 0" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume deg_q: "0 < degree q" + assume contrad: "poly_f qs = 0" + have nonconst: "check_all_const_deg qs = False" using deg_q check_all_const_deg_prop + q_in by auto + have h1: "prod_list_var qs \ 0" using prod_list_var_nonzero by auto + then have "degree (prod_list_var qs) > 0" using q_in deg_q h1 + proof (induct qs) + case Nil + then show ?case by auto + next + case (Cons a qs) + have q_nonz: "q \ 0" using Cons.prems by auto + have q_ins: "q \ set (a # qs) " using Cons.prems by auto + then have "q = a \ q \ set qs" by auto + then have eo: " q = a \ List.member qs q" using in_set_member[of q qs] + by auto + have degq: "degree q > 0" using Cons.prems by auto + have h2: "(prod_list (a # qs)) = a* (prod_list qs)" + by auto + have isa: "q = a \ 0 < degree (prod_list_var (a # qs))" + using h2 degree_mult_eq_0[where p = "q", where q = "prod_list_var qs"] + Cons.prems by auto + have inl: "List.member qs q \ 0 < degree (prod_list_var (a # qs))" + proof - + have nonzprod: "prod_list_var (a # qs) \ 0" using prod_list_var_nonzero by auto + have "q dvd prod_list_var (a # qs)" + using q_dvd_prod_list_var_prop[where q = "q", where qs = "(a#qs)"] q_nonz q_ins + by auto + then show ?thesis using divides_degree[where p = "q", where q = "prod_list_var (a # qs)"] nonzprod degq + by auto + qed + then show ?case using eo isa by auto + qed + then have h2: "pderiv (prod_list_var qs) \ 0" using pderiv_eq_0_iff[where p = "prod_list_var qs"] + by auto + then have "pderiv (prod_list_var qs) * prod_list_var qs \ 0" + using prod_list_var_nonzero h2 by auto + then show "False" using contrad nonconst unfolding poly_f_def deg_q + by (smt (z3) mult_eq_0_iff pCons_eq_0_iff) + qed + show ?thesis using eo c1 c2 by auto +qed + +lemma poly_f_roots_prop_1: + fixes qs:: "real poly list" + assumes non_const: "check_all_const_deg qs = False" + shows "\x1. \x2. ((x1 < x2 \ (\q1 \ set (qs). q1 \ 0 \ (poly q1 x1) = 0) \ (\q2\ set(qs). q2 \ 0 \ (poly q2 x2) = 0)) \ (\q. x1 < q \ q < x2 \ poly (poly_f qs) q = 0))" +proof clarsimp + fix x1:: "real" + fix x2:: "real" + fix q1:: "real poly" + fix q2:: "real poly" + assume "x1 < x2" + assume q1_in: "q1 \ set qs" + assume q1_0: "poly q1 x1 = 0" + assume q1_nonz: "q1 \ 0" + assume q2_in: "q2 \ set qs" + assume q2_0: "poly q2 x2 = 0" + assume q2_nonz: "q2 \ 0" + have prod_z_x1: "poly (prod_list_var qs) x1 = 0" using q1_in q1_0 + using q1_nonz q_dvd_prod_list_var_prop[of q1 qs] by auto + have prod_z_x2: "poly (prod_list_var qs) x2 = 0" using q2_in q2_0 + using q2_nonz q_dvd_prod_list_var_prop[of q2 qs] by auto + have "\w>x1. w < x2 \ poly (pderiv (prod_list_var qs)) w = 0" + using Rolle_pderiv[where q = "prod_list_var qs"] prod_z_x1 prod_z_x2 + using \x1 < x2\ by blast + then obtain w where w_def: "w > x1 \w < x2 \ poly (pderiv (prod_list_var qs)) w = 0" + by auto + then have "poly (poly_f qs) w = 0" + unfolding poly_f_def using non_const + by simp + then show "\q>x1. q < x2 \ poly (poly_f qs) q = 0" + using w_def by blast +qed + +lemma main_step_aux1_R: + fixes qs:: "real poly list" + assumes non_const: "check_all_const_deg qs = True" + shows "set (find_consistent_signs_R qs) = consistent_sign_vectors_R qs UNIV" +proof - + have poly_f_is: "poly_f qs = [:0, 1:]" unfolding poly_f_def using assms + + by auto + have same: "set (find_consistent_signs_at_roots_R [:0, 1:] qs) = + set (characterize_consistent_signs_at_roots [:0, 1:] qs)" using find_consistent_signs_at_roots_R[of "[:0, 1:]" qs] + by auto + have rech: "(sorted_list_of_set {x. poly ([:0, 1:]::real poly) x = 0}) = [0]" by auto + have alldeg0: "(\p \ set qs. degree p = 0)" using non_const check_all_const_deg_prop + by auto + then have allconst: "\p \ set qs. (\(k::real). p = [:k:])" + apply (auto) + by (meson degree_eq_zeroE) + then have allconstvar: "\p \ set qs. \(x::real). \(y::real). poly p x = poly p y" + by fastforce + have e1: "set (remdups (map (signs_at qs) [0])) \ + consistent_sign_vectors_R qs UNIV" + unfolding signs_at_def squash_def consistent_sign_vectors_R_def consistent_sign_vec_def apply (simp) + by (smt (verit, best) class_ring.ring_simprules(2) comp_def image_iff length_map map_nth_eq_conv) + have e2: "consistent_sign_vectors_R qs UNIV \ set (remdups (map (signs_at qs) [0])) " + unfolding signs_at_def squash_def consistent_sign_vectors_R_def consistent_sign_vec_def apply (simp) + using allconstvar + by (smt (verit, best) comp_apply image_iff insert_iff map_eq_conv subsetI) + have "set (remdups (map (signs_at qs) [0])) = + consistent_sign_vectors_R qs UNIV" + using e1 e2 by auto + then have "set (characterize_consistent_signs_at_roots [:0, 1:] qs) = consistent_sign_vectors_R qs UNIV" + unfolding characterize_consistent_signs_at_roots_def characterize_root_list_p_def + using rech by auto + then show ?thesis using same poly_f_is unfolding find_consistent_signs_R_def + by auto +qed + +lemma sorted_list_lemma_var: + fixes l:: "real list" + fixes x:: "real" + assumes "length l > 1" + assumes strict_sort: "strict_sorted l" + assumes x_not_in: "\ (List.member l x)" + assumes lt_a: "x > (l ! 0)" + assumes b_lt: "x < (l ! (length l - 1))" + shows "(\n. n < length l - 1 \ x > l ! n \ x < l !(n+1))" using assms +proof (induct l) + case Nil + then show ?case by auto +next + case (Cons a l) + have len_gteq: "length l \ 1" using Cons.prems(1) + by (metis One_nat_def Suc_eq_plus1 list.size(4) not_le not_less_eq) + have len_one: "length l = 1 \ (\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1))" + proof - + assume len_is: "length l = 1" + then have "x > (a#l) ! 0 \ x < (a#l) !1 " using Cons.prems(4) Cons.prems(5) + by auto + then show "(\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1))" + using len_is by auto + qed + have len_gt: "length l > 1 \ (\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1))" + proof - + assume len_gt_one: "length l > 1" + have eo: "x \ l ! 0" using Cons.prems(3) apply (auto) + by (metis One_nat_def Suc_lessD in_set_member len_gt_one member_rec(1) nth_mem) + have c1: "x < l ! 0 \ (\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1)) " + proof - + assume xlt: "x < l !0" + then have "x < (a#l) ! 1 " + by simp + then show ?thesis using Cons.prems(4) len_gt_one apply (auto) + using Cons.prems(4) Suc_lessD by blast + qed + have c2: "x > l ! 0 \ (\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1)) " + proof - + assume asm: "x > l ! 0" + have xlt_1: " x < l ! (length l - 1)" + using Cons.prems(5) + by (metis Cons.prems(1) One_nat_def add_diff_cancel_right' list.size(4) nth_Cons_pos zero_less_diff) + have ssl: "strict_sorted l " using Cons.prems(2) + using strict_sorted.simps(2) by blast + have " \ List.member l x" using Cons.prems(3) + by (meson member_rec(1)) + then have " \n x < l ! (n + 1)" + using asm xlt_1 len_gt_one ssl Cons.hyps + by auto + then show ?thesis + by (metis One_nat_def Suc_eq_plus1 diff_Suc_1 less_diff_conv list.size(4) nth_Cons_Suc) + qed + show "(\n. n < length (a#l) - 1 \ x > (a#l) ! n \ x < (a#l) !(n+1))" + using eo c1 c2 + by (meson linorder_neqE_linordered_idom) + qed + then show ?case + using len_gteq len_one len_gt + apply (auto) + by (metis One_nat_def less_numeral_extra(1) linorder_neqE_nat not_less nth_Cons_0) +qed + +(* We want to show that our auxiliary polynomial has roots in all relevant intervals: + so it captures all of the zeros, and also it captures all of the points in between! *) +lemma all_sample_points_prop: + assumes is_not_const: "check_all_const_deg qs = False" + assumes s_is: "S = (characterize_root_list_p (pderiv (prod_list_var qs) * (prod_list_var qs) * ([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])))"(* properties about S*) + shows "consistent_sign_vectors_R qs UNIV = consistent_sign_vectors_R qs (set S)" +proof - + let ?zer_list = "sorted_list_of_set {(x::real). (\q \ set(qs). (q \ 0 \ poly q x = 0))} :: real list" + have strict_sorted_h: "strict_sorted ?zer_list" using sorted_sorted_list_of_set + strict_sorted_iff by auto + have poly_f_is: "poly_f qs = (pderiv (prod_list_var qs) * prod_list_var qs)* ([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])" + unfolding poly_f_def using is_not_const by auto + then have set_S_char: "set S = ({x. poly (poly_f qs) x = 0}::real set)" + using poly_roots_finite[of "poly_f qs"] set_sorted_list_of_set poly_f_nonzero[of qs] + using s_is unfolding characterize_root_list_p_def by auto + have difficult_direction: "consistent_sign_vectors_R qs UNIV \ consistent_sign_vectors_R qs (set S)" + proof clarsimp + fix x + assume "x \ consistent_sign_vectors_R qs UNIV " + then have "\y. x = (consistent_sign_vec qs y)" unfolding consistent_sign_vectors_R_def by auto + then obtain y where y_prop: "x = consistent_sign_vec qs y" by auto + then have "\ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + proof - + have c1: "(\q \ (set qs). q \ 0 \ poly q y = 0) \ (\ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y)" + proof - + assume "(\q \ (set qs). q \ 0 \ poly q y = 0)" + then obtain q where "q \ (set qs) \ q \ 0 \ poly q y = 0" by auto + then have "poly (prod_list_var qs) y = 0" + using q_dvd_prod_list_var_prop[of q qs] by auto + then have "poly (pderiv (prod_list_var qs) * (prod_list_var qs)*([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])) y = 0" + by auto + then have "y \ (set S)" + using s_is unfolding characterize_root_list_p_def + proof - + have "y \ {r. poly (pderiv (prod_list_var qs) * (prod_list_var qs)*([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])) r = 0}" + using \poly (pderiv (prod_list_var qs) * (prod_list_var qs)*([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])) y = 0\ by force + then show ?thesis + by (metis characterize_root_list_p_def is_not_const poly_f_def poly_f_nonzero poly_roots_finite s_is set_sorted_list_of_set) + qed + then show "\ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + by auto + qed + have len_gtz_prop: "length ?zer_list > 0 \ + ((\w. w < length ?zer_list \ y = ?zer_list ! w) \ + (y < ?zer_list ! 0) \ + (y > ?zer_list ! (length ?zer_list - 1)) \ + (\k < (length ?zer_list - 1). y > ?zer_list ! k \ y < ?zer_list ! (k+1)))" + proof - + let ?c = "(\w. w < length ?zer_list \ y = ?zer_list ! w) \ + (y < ?zer_list ! 0) \ + (y > ?zer_list ! (length ?zer_list - 1)) \ + (\k < (length ?zer_list - 1). y > ?zer_list ! k \ y < ?zer_list ! (k+1))" + have lis1: "length ?zer_list = 1 \ ?c" + by auto + have h1: "\(\w. w < length ?zer_list \ y = ?zer_list ! w) \ \ (List.member ?zer_list y)" + by (metis (no_types, lifting) in_set_conv_nth in_set_member) + have h2: "(length ?zer_list > 0 \ \(\w. w < length ?zer_list \ y = ?zer_list ! w) \ \ (y < ?zer_list ! 0)) \ y > ?zer_list ! 0" + by auto + have h3: "(length ?zer_list > 1 \ \(\w. w < length ?zer_list \ y = ?zer_list ! w) \ \ (y > ?zer_list ! (length ?zer_list - 1))) \ + y < ?zer_list ! (length ?zer_list - 1)" + apply (auto) + by (smt (z3) diff_Suc_less gr_implies_not0 not_gr_zero) + have "length ?zer_list > 1 \ \(\w. w < length ?zer_list \ y = ?zer_list ! w) \ \ (y < ?zer_list ! 0) \ \ (y > ?zer_list ! (length ?zer_list - 1)) + \ (\k < (length ?zer_list - 1). y > ?zer_list ! k \ y < ?zer_list ! (k+1))" + using h1 h2 h3 strict_sorted_h sorted_list_lemma_var[of ?zer_list y] + using One_nat_def Suc_lessD by presburger + then have lgt1: "length ?zer_list > 1 \ ?c" + by auto + then show ?thesis using lis1 lgt1 + by (smt (z3) diff_is_0_eq' not_less) + qed + have neg_crb_in: "(- crb (prod_list_var qs)) \ set S" + using set_S_char poly_f_is by auto + have pos_crb_in: "(crb (prod_list_var qs)) \ set S" + using set_S_char poly_f_is by auto + have set_S_nonempty: "set S \ {}" using neg_crb_in by auto + have finset: "finite {x. \q\set qs. q \ 0 \ poly q x = 0}" + proof - + have "\q \ set qs. q\ 0 \ finite {x. poly q x = 0} " + using poly_roots_finite by auto + then show ?thesis by auto + qed + have c2: "\(\q \ (set qs). q \ 0 \ poly q y = 0) \ \ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + proof - + assume "\(\q \ (set qs). q \ 0 \ poly q y = 0)" + have c_c1: "length ?zer_list = 0 \ \ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + proof - + assume "length ?zer_list = 0" + then have "\q \ set (qs). \ (x:: real). \(y::real). squash (poly q x) = squash (poly q y)" + proof clarsimp + fix q x y + assume czer: "card {x. \q\set qs. q \ 0 \ poly q x = 0} = 0" + assume qin: "q \ set qs" + have fin_means_empty: "{x. \q\set qs. q \ 0 \ poly q x = 0} = {}" + using finset czer + by auto + have qzer: "q = 0 \ squash (poly q x) = squash (poly q y)" by auto + have qnonz: "q \ 0 \ squash (poly q x) = squash (poly q y)" + proof - + assume qnonz: "q \ 0" + then have noroots: "{x. poly q x = 0} = {}" using qin finset + using Collect_empty_eq fin_means_empty by auto + have nonzsq1: "squash (poly q x) \ 0" using fin_means_empty qnonz czer qin + unfolding squash_def by auto + then have eo: "(poly q x) > 0 \ (poly q x) < 0" unfolding squash_def + apply (auto) + by presburger + have eo1: "poly q x > 0 \ poly q y > 0" + using noroots poly_IVT_pos[of y x q] poly_IVT_neg[of x y q] + apply (auto) + by (metis linorder_neqE_linordered_idom) + have eo2: "poly q x < 0 \ poly q y < 0" + using noroots poly_IVT_pos[of x y q] poly_IVT_neg[of y x q] + apply (auto) by (metis linorder_neqE_linordered_idom) + then show "squash (poly q x) = squash (poly q y)" + using eo eo1 eo2 unfolding squash_def by auto + qed + show "squash (poly q x) = squash (poly q y)" + using qzer qnonz + by blast + qed + then have "\q \ set (qs). squash (poly q y) = squash (poly q (- crb (prod_list_var qs)))" + by auto + then show "\ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + using neg_crb_in unfolding consistent_sign_vec_def squash_def + apply (auto) + by (metis (no_types, hide_lams) antisym_conv3 class_field.neg_1_not_0 equal_neg_zero less_irrefl of_int_minus) + qed + have c_c2: "length ?zer_list > 0 \ \ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + proof - + assume lengt: "length ?zer_list > 0" + let ?t = " \ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + have sg1: "(\w. w < length ?zer_list \ y = ?zer_list ! w) \ ?t" + proof - + assume "(\w. w < length ?zer_list \ y = ?zer_list ! w)" + then obtain w where w_prop: "w < length ?zer_list \ y = ?zer_list ! w" by auto + then have " y \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using finset set_sorted_list_of_set[of "{x. \q\set qs. q \ 0 \ poly q x = 0}"] + by (smt (verit, best) nth_mem) + then have "y \ {x. poly (poly_f qs) x = 0}" using poly_f_is + using \\ (\q\set qs. q \ 0 \ poly q y = 0)\ by blast + then show ?thesis using set_S_char + by blast + qed + have sg2: "(y < ?zer_list ! 0) \ ?t" + proof - + assume ylt: "y < ?zer_list ! 0" + have ynonzat_some_qs: "\q \ (set qs). q \ 0 \ poly q y \ 0" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume qnonz: "q \ 0" + assume "poly q y = 0" + then have "y \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list y" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "y \ ?zer_list ! 0" using strict_sorted_h + using \\ (\q\set qs. q \ 0 \ poly q y = 0)\ \poly q y = 0\ q_in qnonz by blast + then show "False" using ylt + by auto + qed + let ?ncrb = "(- crb (prod_list_var qs))" + have "\x \ {x. \q\set qs. q \ 0 \ poly q x = 0}. poly (prod_list_var qs) x = 0" + using q_dvd_prod_list_var_prop + by fastforce + then have "poly (prod_list_var qs) (sorted_list_of_set {x. \q\set qs. q \ 0 \ poly q x = 0} ! 0) = 0" + using finset set_sorted_list_of_set + by (metis (no_types, lifting) lengt nth_mem) + then have ncrblt: "?ncrb < ?zer_list ! 0" using prod_list_var_nonzero crb_lem_neg[of "prod_list_var qs" "?zer_list ! 0"] + by auto + have qzerh: "\q \ (set qs). q = 0 \ squash (poly q ?ncrb) = squash (poly q y)" + by auto + have "\q \ (set qs). q \ 0 \ squash (poly q ?ncrb) = squash (poly q y)" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume qnonz: "q \ 0" + have nonzylt:"\(\x \ y. poly q x = 0)" + proof clarsimp + fix x + assume xlt: "x \ y" + assume "poly q x = 0" + then have "x \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list x" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "x \ ?zer_list ! 0" using strict_sorted_h + by (metis (no_types, lifting) gr_implies_not0 in_set_conv_nth in_set_member not_less sorted_iff_nth_mono sorted_list_of_set(2)) + then show "False" using xlt ylt + by auto + qed + have nonzncrb:"\(\x \ (real_of_int ?ncrb). poly q x = 0)" + proof clarsimp + fix x + assume xlt: "x \ - real_of_int (crb (prod_list_var qs))" + assume "poly q x = 0" + then have "x \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list x" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "x \ ?zer_list ! 0" using strict_sorted_h + by (metis (no_types, lifting) gr_implies_not0 in_set_conv_nth in_set_member not_less sorted_iff_nth_mono sorted_list_of_set(2)) + then show "False" using xlt ncrblt + by auto + qed + have c1: " (poly q ?ncrb) > 0 \ (poly q y) > 0" + proof - + assume qncrbgt: "(poly q ?ncrb) > 0" + then have eq: "?ncrb = y \ poly q y > 0 " by auto + have gt: " ?ncrb > y \ poly q y > 0" using qncrbgt qnonz poly_IVT_pos[of y ?ncrb q] poly_IVT_neg[of ?ncrb y q] nonzncrb nonzylt + apply (auto) + by (meson less_eq_real_def linorder_neqE_linordered_idom) + have lt: "?ncrb < y \ poly q y > 0" using qncrbgt + using qnonz poly_IVT_pos[of y ?ncrb q] poly_IVT_neg[of ?ncrb y q] nonzncrb nonzylt + apply (auto) + by (meson less_eq_real_def linorder_neqE_linordered_idom) + then show ?thesis using eq gt lt apply (auto) + by (meson linorder_neqE_linordered_idom) + qed + have c2: "(poly q ?ncrb) < 0 \ (poly q y) < 0" + using poly_IVT_pos[of ?ncrb y q] poly_IVT_neg[of y ?ncrb q] nonzncrb nonzylt + apply (auto) + by (metis less_eq_real_def linorder_neqE_linordered_idom) + have eo: "(poly q ?ncrb) > 0 \ (poly q ?ncrb) < 0" + using nonzncrb + by auto + then show "squash (poly q (- real_of_int (crb (prod_list_var qs)))) = squash (poly q y)" + using c1 c2 + by (smt (verit, ccfv_SIG) of_int_minus squash_def) + qed + then have "\q \ (set qs). squash (poly q ?ncrb) = squash (poly q y)" + using qzerh by auto + then have "consistent_sign_vec qs ?ncrb = consistent_sign_vec qs y" + unfolding consistent_sign_vec_def squash_def + by (smt (z3) map_eq_conv) + then show ?thesis using neg_crb_in by auto + qed + have sg3: " (y > ?zer_list ! (length ?zer_list - 1)) \ ?t" + proof - + assume ygt: "y > ?zer_list ! (length ?zer_list - 1)" + have ynonzat_some_qs: "\q \ (set qs). q \ 0 \ poly q y \ 0" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume qnonz: "q \ 0" + assume "poly q y = 0" + then have "y \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list y" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "y \ ?zer_list ! (length ?zer_list - 1)" using strict_sorted_h + using \\ (\q\set qs. q \ 0 \ poly q y = 0)\ \poly q y = 0\ q_in qnonz by blast + then show "False" using ygt + by auto + qed + let ?crb = "crb (prod_list_var qs)" + have "\x \ {x. \q\set qs. q \ 0 \ poly q x = 0}. poly (prod_list_var qs) x = 0" + using q_dvd_prod_list_var_prop + by fastforce + then have "poly (prod_list_var qs) (sorted_list_of_set {x. \q\set qs. q \ 0 \ poly q x = 0} ! 0) = 0" + using finset set_sorted_list_of_set + by (metis (no_types, lifting) lengt nth_mem) + then have crbgt: "?crb > ?zer_list ! (length ?zer_list - 1)" using prod_list_var_nonzero crb_lem_pos[of "prod_list_var qs" "?zer_list ! (length ?zer_list - 1)"] + by (metis (no_types, lifting) \\x\{x. \q\set qs. q \ 0 \ poly q x = 0}. poly (prod_list_var qs) x = 0\ diff_less finset lengt less_numeral_extra(1) nth_mem set_sorted_list_of_set) + have qzerh: "\q \ (set qs). q = 0 \ squash (poly q ?crb) = squash (poly q y)" + by auto + have "\q \ (set qs). q \ 0 \ squash (poly q ?crb) = squash (poly q y)" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume qnonz: "q \ 0" + have nonzylt:"\(\x \ y. poly q x = 0)" + proof clarsimp + fix x + assume xgt: "x \ y" + assume "poly q x = 0" + then have "x \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list x" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "x \ ?zer_list ! (length ?zer_list - 1)" using strict_sorted_h + by (metis (no_types, lifting) One_nat_def Suc_leI Suc_pred diff_Suc_less in_set_conv_nth in_set_member lengt not_less sorted_iff_nth_mono sorted_list_of_set(2)) + then show "False" using xgt ygt + by auto + qed + have nonzcrb:"\(\x \ (real_of_int ?crb). poly q x = 0)" + proof clarsimp + fix x + assume xgt: "x \ real_of_int (crb (prod_list_var qs))" + assume "poly q x = 0" + then have "x \ {x. \q\set qs. q \ 0 \ poly q x = 0}" + using q_in qnonz by auto + then have "List.member ?zer_list x" + by (smt (verit, best) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have "x \ ?zer_list ! (length ?zer_list - 1)" using strict_sorted_h + by (meson \\x\{x. \q\set qs. q \ 0 \ poly q x = 0}. poly (prod_list_var qs) x = 0\ \x \ {x. \q\set qs. q \ 0 \ poly q x = 0}\ crb_lem_pos not_less prod_list_var_nonzero xgt) + then show "False" using xgt crbgt + by auto + qed + have c1: " (poly q ?crb) > 0 \ (poly q y) > 0" + proof - + assume qcrbgt: "(poly q ?crb) > 0" + then have eq: "?crb = y \ poly q y > 0 " by auto + have gt: " ?crb > y \ poly q y > 0" using qcrbgt qnonz poly_IVT_pos[of y ?crb q] poly_IVT_neg[of ?crb y q] nonzcrb nonzylt + apply (auto) + by (meson less_eq_real_def linorder_neqE_linordered_idom) + have lt: "?crb < y \ poly q y > 0" using qcrbgt + using qnonz poly_IVT_pos[of y ?crb q] poly_IVT_neg[of ?crb y q] nonzcrb nonzylt + apply (auto) + by (meson less_eq_real_def linorder_neqE_linordered_idom) + then show ?thesis using eq gt lt apply (auto) + by (meson linorder_neqE_linordered_idom) + qed + have c2: "(poly q ?crb) < 0 \ (poly q y) < 0" + using poly_IVT_pos[of ?crb y q] poly_IVT_neg[of y ?crb q] nonzcrb nonzylt + apply (auto) + by (metis less_eq_real_def linorder_neqE_linordered_idom) + have eo: "(poly q ?crb) > 0 \ (poly q ?crb) < 0" + using nonzcrb + by auto + then show "squash (poly q (real_of_int (crb (prod_list_var qs)))) = squash (poly q y)" + using c1 c2 + by (smt (verit, ccfv_SIG) of_int_minus squash_def) + qed + then have "\q \ (set qs). squash (poly q ?crb) = squash (poly q y)" + using qzerh by auto + then have "consistent_sign_vec qs ?crb = consistent_sign_vec qs y" + unfolding consistent_sign_vec_def squash_def + by (smt (z3) map_eq_conv) + then show ?thesis using pos_crb_in by auto + qed + have sg4: " (\k < (length ?zer_list - 1). y > ?zer_list ! k \ y < ?zer_list ! (k+1)) \ ?t" + proof - + assume " (\k < (length ?zer_list - 1). y > ?zer_list ! k \ y < ?zer_list ! (k+1))" + then obtain k where k_prop: "k < (length ?zer_list - 1) \ y > ?zer_list ! k \ y < ?zer_list ! (k+1)" + by auto + have ltk: "(?zer_list ! k) < (?zer_list ! (k+1)) " + using strict_sorted_h + using k_prop by linarith + have q1e: "(\q1\set qs. q1 \ 0 \ poly q1 (?zer_list ! k) = 0)" + by (smt (z3) One_nat_def Suc_lessD add.right_neutral add_Suc_right finset k_prop less_diff_conv mem_Collect_eq nth_mem set_sorted_list_of_set) + have q2e: "(\q2\set qs. q2 \ 0 \ poly q2 (?zer_list ! (k + 1)) = 0)" + by (smt (verit, del_insts) finset k_prop less_diff_conv mem_Collect_eq nth_mem set_sorted_list_of_set) + then have "(\q>(?zer_list ! k). q < (?zer_list ! (k + 1)) \ poly (poly_f qs) q = 0)" + using poly_f_roots_prop_1[of qs] q1e q2e ltk is_not_const + by auto + then have "\s \ set S. s > ?zer_list ! k \ s < ?zer_list ! (k+1)" + using poly_f_is + by (smt (z3) k_prop mem_Collect_eq set_S_char) + then obtain s where s_prop: "s \ set S \ s > ?zer_list ! k \ s < ?zer_list ! (k+1)" by auto + have qnon: "\q \ set qs. q\ 0 \ squash (poly q s) = squash (poly q y)" + proof clarsimp + fix q + assume q_in: "q \ set qs" + assume qnonz: "q \ 0" + have sgt: "s > y \ squash (poly q s) = squash (poly q y)" + proof - + assume "s > y" + then have "\x. List.member ?zer_list x \ y \ x \ x \ s" + using sorted_list_lemma[of y s k ?zer_list] k_prop strict_sorted_h s_prop y_prop + using less_diff_conv by blast + then have nox: "\x. poly q x = 0 \ y \ x \ x \ s" using q_in qnonz + by (metis (mono_tags, lifting) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have c1: "poly q s \ 0" using s_prop q_in qnonz + by (metis (mono_tags, lifting) \y < s\ less_eq_real_def ) + have c2: "poly q s > 0 \ poly q y > 0" + using poly_IVT_pos poly_IVT_neg nox + by (meson \y < s\ less_eq_real_def linorder_neqE_linordered_idom) + have c3: "poly q s < 0 \ poly q y < 0" using poly_IVT_pos poly_IVT_neg nox + by (meson \y < s\ less_eq_real_def linorder_neqE_linordered_idom) + show ?thesis using c1 c2 c3 unfolding squash_def + by auto + qed + have slt: "s < y \ squash (poly q s) = squash (poly q y)" + proof - + assume slt: "s < y" + then have "\x. List.member ?zer_list x \ s \ x \ x \ y" + using sorted_list_lemma[of s y k ?zer_list] k_prop strict_sorted_h s_prop y_prop + using less_diff_conv by blast + then have nox: "\x. poly q x = 0 \ s \ x \ x \ y" using q_in qnonz + by (metis (mono_tags, lifting) finset in_set_member mem_Collect_eq set_sorted_list_of_set) + then have c1: "poly q s \ 0" using s_prop q_in qnonz + by (metis (mono_tags, lifting) \s < y\ less_eq_real_def ) + have c2: "poly q s > 0 \ poly q y > 0" + using poly_IVT_pos poly_IVT_neg nox + by (meson \s < y\ less_eq_real_def linorder_neqE_linordered_idom) + have c3: "poly q s < 0 \ poly q y < 0" using poly_IVT_pos poly_IVT_neg nox + by (meson \s < y\ less_eq_real_def linorder_neqE_linordered_idom) + show ?thesis using c1 c2 c3 unfolding squash_def + by auto + qed + have "s = y \ squash (poly q s) = squash (poly q y)" + by auto + then show "squash (poly q s) = squash (poly q y)" + using sgt slt + by (meson linorder_neqE_linordered_idom) + qed + have "\q \ set qs. q= 0 \ squash (poly q s) = squash (poly q y)" by auto + then have "\q \ set qs. squash (poly q s) = squash (poly q y)" + using qnon + by fastforce + then show ?thesis + using s_prop unfolding squash_def consistent_sign_vec_def apply (auto) + by (metis (no_types, hide_lams) class_field.neg_1_not_0 equal_neg_zero less_irrefl linorder_neqE_linordered_idom) + qed + show ?thesis + using lengt sg1 sg2 sg3 sg4 len_gtz_prop is_not_const + by fastforce + qed + show "\ k \ (set S). consistent_sign_vec qs k = consistent_sign_vec qs y" + using c_c1 c_c2 by auto + qed + show ?thesis + using c1 c2 by auto + qed + then show "x \ consistent_sign_vectors_R qs (set S)" + using y_prop unfolding consistent_sign_vectors_R_def + by (metis imageI) + qed + have easy_direction: "consistent_sign_vectors_R qs (set S) \ consistent_sign_vectors_R qs UNIV " + using consistent_sign_vectors_R_def by auto + then show ?thesis using difficult_direction easy_direction by auto +qed + +lemma main_step_aux2_R: + fixes qs:: "real poly list" + assumes is_not_const: "check_all_const_deg qs = False" + shows "set (find_consistent_signs_R qs) = consistent_sign_vectors_R qs UNIV" +proof - + have poly_f_is: "poly_f qs = (pderiv (prod_list_var qs)) * (prod_list_var qs)* ([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])" + using is_not_const unfolding poly_f_def by auto + let ?p = "(pderiv (prod_list_var qs)) * (prod_list_var qs)* ([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:])" + let ?S = "characterize_root_list_p (pderiv (prod_list_var qs) * (prod_list_var qs) * ([:-(crb (prod_list_var qs)),1:]) * ([:(crb (prod_list_var qs)),1:]))" + have "set (remdups + (map (signs_at qs) ?S)) + = consistent_sign_vectors_R qs (set ?S)" + unfolding signs_at_def squash_def consistent_sign_vectors_R_def consistent_sign_vec_def + by (smt (verit, best) comp_apply map_eq_conv set_map set_remdups) + then have "set (characterize_consistent_signs_at_roots ?p qs) = consistent_sign_vectors_R qs UNIV" + unfolding characterize_consistent_signs_at_roots_def using assms all_sample_points_prop[of qs] + by auto + then show ?thesis + unfolding find_consistent_signs_R_def using find_consistent_signs_at_roots_R poly_f_is poly_f_nonzero[of qs] + by auto +qed + +lemma main_step_R: + fixes qs:: "real poly list" + shows "set (find_consistent_signs_R qs) = consistent_sign_vectors_R qs UNIV" + using main_step_aux1_R main_step_aux2_R by auto + +(* The universal and existential decision procedure for real polys are easy + if we know the consistent sign vectors *) +lemma consistent_sign_vec_semantics_R: + assumes "\i. i \ set_fml fml \ i < length ls" + shows "lookup_sem fml (map (\p. poly p x) ls) = lookup_sem fml (consistent_sign_vec ls x)" + using assms apply (induction) + by (auto simp add: consistent_sign_vec_def) + +lemma universal_lookup_sem_R: + assumes "\i. i \ set_fml fml \ i < length qs" + assumes "set signs = consistent_sign_vectors_R qs UNIV" + shows "(\x::real. lookup_sem fml (map (\p. poly p x) qs)) \ + list_all (lookup_sem fml) signs" + using assms(2) unfolding consistent_sign_vectors_R_def list_all_iff + by (simp add: assms(1) consistent_sign_vec_semantics_R) + +lemma existential_lookup_sem_R: + assumes "\i. i \ set_fml fml \ i < length qs" + assumes "set signs = consistent_sign_vectors_R qs UNIV" + shows "(\x::real. lookup_sem fml (map (\p. poly p x) qs)) \ + find (lookup_sem fml) signs \ None" + using assms(2) unfolding consistent_sign_vectors_R_def find_None_iff + by (simp add: assms(1) consistent_sign_vec_semantics_R) + +lemma decide_univ_lem_helper_R: + fixes fml:: "real poly fml" + assumes "(fml_struct,polys) = convert fml" + shows "(\x::real. lookup_sem fml_struct (map (\p. poly p x) polys)) \ (decide_universal_R fml)" + using assms universal_lookup_sem_R main_step_R unfolding decide_universal_R_def apply (auto) + apply (metis assms convert_closed fst_conv snd_conv) + by (metis (full_types) assms convert_closed fst_conv snd_conv) + +lemma decide_exis_lem_helper_R: + fixes fml:: "real poly fml" + assumes "(fml_struct,polys) = convert fml" + shows "(\x::real. lookup_sem fml_struct (map (\p. poly p x) polys)) \ (decide_existential_R fml)" + using assms existential_lookup_sem_R main_step_R unfolding decide_existential_R_def apply (auto) + apply (metis assms convert_closed fst_conv snd_conv) + by (metis (full_types) assms convert_closed fst_conv snd_conv) + +lemma convert_semantics_lem_R: + assumes "\p. p \ set (poly_list fml) \ + ls ! (index_of ps p) = poly p x" + shows "real_sem fml x = lookup_sem (map_fml (index_of ps) fml) ls" + using assms apply (induct fml) + by auto + +lemma convert_semantics_R: + shows "real_sem fml x = lookup_sem (fst (convert fml)) (map (\p. poly p x) (snd (convert fml)))" + unfolding convert_def Let_def apply simp + apply (intro convert_semantics_lem_R) + by (simp add: index_of_lookup(1) index_of_lookup(2)) + +(* Main result *) +theorem decision_procedure_R: + shows "(\x::real. real_sem fml x) \ (decide_universal_R fml)" + "\x::real. real_sem fml x \ (decide_existential_R fml)" + using convert_semantics_lem_R decide_univ_lem_helper_R apply (auto) + apply (simp add: convert_semantics_R) + apply (metis convert_def convert_semantics_R fst_conv snd_conv) + using convert_semantics_lem_R + by (metis convert_def convert_semantics_R decide_exis_lem_helper_R fst_conv snd_conv) + +end diff --git a/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy b/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/Renegar_Proofs.thy @@ -0,0 +1,2776 @@ +theory Renegar_Proofs + imports "Renegar_Algorithm" + "BKR_Proofs" +begin + +(* Note that there is significant overlap between Renegar and BKR in general, so there is some + similarity between this file and BKR_Proofs.thy + The main difference is that the RHS vector in Renegar is different from the RHS vector in BKR + In BKR, all of the qs are assumed to be relatively prime to p. Renegar removes this assumption. + + In general, the _R's on definition and lemma names in this file are to emphasize that we are + working with Renegar style. +*) + +section "Tarski Queries Changed" + +lemma construct_NofI_R_relation: + fixes p:: "real poly" + fixes I1:: "real poly list" + fixes I2:: "real poly list" + shows "construct_NofI_R p I1 I2 = + construct_NofI (sum_list (map power2 (p # I1))) I2" + unfolding construct_NofI_R_def construct_NofI_def + by metis + +lemma sum_list_map_power2: + shows "sum_list (map power2 ls) \ (0::real poly)" + apply (induct ls) + by auto + +lemma sum_list_map_power2_poly: + shows "poly (sum_list (map power2 (ls::real poly list))) x \ (0::real)" + apply (induct ls) + by auto + +lemma construct_NofI_R_prop_helper: + fixes p:: "real poly" + fixes I1:: "real poly list" + fixes I2:: "real poly list" + assumes nonzero: "p\0" + shows "construct_NofI_R p I1 I2 = + rat_of_int (int (card {x. poly (sum_list (map (\x. x^2) (p # I1))) x = 0 \ poly (prod_list I2) x > 0}) - + int (card {x. poly (sum_list (map (\x. x^2) (p # I1))) x = 0 \ poly (prod_list I2) x < 0}))" +proof - + show ?thesis unfolding construct_NofI_R_relation[of p I1 I2] + apply (subst construct_NofI_prop[of _ I2]) + apply auto + using assms sum_list_map_power2 + by (metis le_add_same_cancel1 power2_less_eq_zero_iff) +qed + +lemma zer_iff: + fixes p:: "real poly" + fixes ls:: "real poly list" + shows "poly (sum_list (map (\x. x^2) ls)) x = 0 \ (\i \ set ls. poly i x = 0)" +proof (induct ls) + case Nil + then show ?case by auto +next + case (Cons a I1) + then show ?case + apply simp + apply (subst add_nonneg_eq_0_iff) + using zero_le_power2 apply blast + using sum_list_map_power2_poly apply presburger + by simp +qed + +lemma construct_NofI_prop_R: + fixes p:: "real poly" + fixes I1:: "real poly list" + fixes I2:: "real poly list" + assumes nonzero: "p\0" + shows "construct_NofI_R p I1 I2 = + rat_of_int (int (card {x. poly p x = 0 \ (\q \ set I1. poly q x = 0) \ poly (prod_list I2) x > 0}) - + int (card {x. poly p x = 0 \ (\q \ set I1. poly q x = 0) \ poly (prod_list I2) x < 0}))" + unfolding construct_NofI_R_prop_helper[OF nonzero] + using zer_iff + apply auto + by (smt (verit, del_insts) Collect_cong sum_list_map_power2_poly zero_le_power2 zero_less_power2) + +section "Matrix Equation" + +definition map_sgas:: "rat list \ rat list" + where "map_sgas l = map (\r. (1 - r^2)) l" + +definition z_R:: "(nat list*nat list) \ rat list \ rat" + where "z_R index_list sign_asg \ (prod_list (map (nth (map_sgas sign_asg)) (fst(index_list))))*(prod_list (map (nth sign_asg) (snd(index_list))))" + +definition mtx_row_R:: "rat list list \ (nat list * nat list) \ rat list" + where "mtx_row_R sign_list index_list \ (map ((z_R index_list)) sign_list)" + +definition matrix_A_R:: "rat list list \ (nat list * nat list) list \ rat mat" + where "matrix_A_R sign_list subset_list = + (mat_of_rows_list (length sign_list) (map (\i .(mtx_row_R sign_list i)) subset_list))" + +definition all_list_constr_R:: "(nat list*nat list) list \ nat \ bool" + where "all_list_constr_R L n \ (\x. List.member L x \ (list_constr (fst x) n \ list_constr (snd x) n))" + +definition alt_matrix_A_R:: "rat list list \ (nat list*nat list) list \ rat mat" + where "alt_matrix_A_R signs subsets = (mat (length subsets) (length signs) + (\(i, j). z_R (subsets ! i) (signs ! j)))" + +lemma alt_matrix_char_R: "alt_matrix_A_R signs subsets = matrix_A_R signs subsets" +proof - + have h0: "(\i j. i < length subsets \ + j < length signs \ + map (\index_list. map (z_R index_list) signs) subsets ! i ! j = z_R (subsets ! i) (signs ! j))" + proof - + fix i + fix j + assume i_lt: "i < length subsets" + assume j_lt: "j < length signs" + show "((map (\index_list. map (z_R index_list) signs) subsets) ! i) ! j = z_R (subsets ! i) (signs ! j)" + proof - + have h0: "(map (\index_list. map (z_R index_list) signs) subsets) ! i = map (z_R (subsets ! i)) signs" + using nth_map i_lt + by blast + then show ?thesis using nth_map j_lt + by simp + qed + qed + have h: " mat (length subsets) (length signs) (\(i, j). z_R (subsets ! i) (signs ! j)) = + mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z_R index_list) signs) subsets ! i ! y)" + using h0 eq_matI[where A = "mat (length subsets) (length signs) (\(i, j). z_R (subsets ! i) (signs ! j))", + where B = "mat (length subsets) (length signs) (\(i, y). map (\index_list. map (z_R index_list) signs) subsets ! i ! y)"] + by auto + show ?thesis unfolding alt_matrix_A_R_def matrix_A_R_def mat_of_rows_list_def apply (auto) unfolding mtx_row_R_def + using h by blast +qed + +lemma subsets_are_rows_R: "\i < (length subsets). row (alt_matrix_A_R signs subsets) i = vec (length signs) (\j. z_R (subsets ! i) (signs ! j))" + unfolding row_def unfolding alt_matrix_A_R_def by auto + +lemma signs_are_cols_R: "\i < (length signs). col (alt_matrix_A_R signs subsets) i = vec (length subsets) (\j. z_R (subsets ! j) (signs ! i))" + unfolding col_def unfolding alt_matrix_A_R_def by auto + +definition consistent_sign_vec::"real poly list \ real \ rat list" + where "consistent_sign_vec qs x \ + map (\ q. if (poly q x > 0) then (1::rat) else (if (poly q x = 0) then (0::rat) else (-1::rat))) qs" + +definition construct_lhs_vector_R:: "real poly \ real poly list \ rat list list \ rat vec" + where "construct_lhs_vector_R p qs signs \ + vec_of_list (map (\w. rat_of_int (int (length (filter (\v. v = w) (map (consistent_sign_vec qs) (characterize_root_list_p p)))))) signs)" + +(* The ith entry of LHS vector is the number of (distinct) real zeros of p where the sign vector of + the qs is the ith entry of signs.*) + +(* Putting all of the pieces of the construction together *) +definition satisfy_equation_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat list list \ bool" + where "satisfy_equation_R p qs subset_list sign_list = + (mult_mat_vec (matrix_A_R sign_list subset_list) (construct_lhs_vector_R p qs sign_list) = (construct_rhs_vector_R p qs subset_list))" + +(* Recharacterize the LHS vector *) +lemma construct_lhs_vector_clean_R: + assumes "p \ 0" + assumes "i < length signs" + shows "(construct_lhs_vector_R p qs signs) $ i = + card {x. poly p x = 0 \ ((consistent_sign_vec qs x) = (nth signs i))}" +proof - + from poly_roots_finite[OF assms(1)] have "finite {x. poly p x = 0}" . + then have eq: "(Collect + ((\v. v = signs ! i) \ + consistent_sign_vec qs) \ + set (sorted_list_of_set + {x. poly p x = 0})) = + {x. poly p x = 0 \ consistent_sign_vec qs x = signs ! i}" + by auto + show ?thesis + unfolding construct_lhs_vector_R_def vec_of_list_index characterize_root_list_p_def + apply auto + apply (subst nth_map[OF assms(2)]) + apply auto + apply (subst distinct_length_filter) + apply (auto) + using eq + by auto +qed + +lemma construct_lhs_vector_cleaner_R: + assumes "p \ 0" + shows "(construct_lhs_vector_R p qs signs) = + vec_of_list (map (\s. rat_of_int (card {x. poly p x = 0 \ ((consistent_sign_vec qs x) = s)})) signs)" + apply (rule eq_vecI) + apply (auto simp add: construct_lhs_vector_clean_R[OF assms] ) + apply (simp add: vec_of_list_index) + unfolding construct_lhs_vector_R_def + using assms construct_lhs_vector_clean_R construct_lhs_vector_def apply auto[1] + apply (simp add: construct_lhs_vector_R_def) + by auto + +(* Show that because our consistent sign vectors consist of 1, 0's, and -1's, z returns 1, 0, or -1 + when applied to a consistent sign vector *) +lemma z_signs_R2: + fixes I:: "nat list" + fixes signs:: "rat list" + assumes lf: "list_all (\i. i < length signs) I" + assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" + shows "(prod_list (map (nth signs) I)) = 1 \ + (prod_list (map (nth signs) I)) = 0 \ + (prod_list (map (nth signs) I)) = -1" using assms +proof (induct I) + case Nil + then show ?case by auto +next + case (Cons a I) + moreover have eo: "signs ! a = 1 \ signs ! a = 0 \ signs ! a = -1" + using assms + by (smt (verit, del_insts) calculation(2) list_all_length list_all_simps(1)) + have "prod_list (map ((!) (map_sgas signs)) (a # I)) = (1 - (signs ! a)^2)*prod_list (map ((!) (map_sgas signs)) (I))" + unfolding map_sgas_def apply (auto) + using calculation(2) by auto + then show ?case using eo + using Cons.hyps calculation(2) la by auto +qed + +lemma z_signs_R1: + fixes I:: "nat list" + fixes signs:: "rat list" + assumes lf: "list_all (\i. i < length signs) I" + assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" + shows "(prod_list (map (nth (map_sgas signs)) I)) = 1 \ +(prod_list (map (nth (map_sgas signs)) I)) = 0" using assms +proof (induct I) + case Nil + then show ?case by auto +next + case (Cons a I) + moreover have "signs ! a = 1 \ signs ! a = 0 \ signs ! a = -1" + using assms + by (smt (verit, best) calculation(2) list_all_length list_all_simps(1)) + then have eo: "1 - (signs ! a)^2 = 1 \ 1 - (signs !a)^2 = 0" + using cancel_comm_monoid_add_class.diff_cancel diff_zero by fastforce + have "prod_list (map ((!) (map_sgas signs)) (a # I)) = (1 - (signs ! a)^2)*prod_list (map ((!) (map_sgas signs)) (I))" + unfolding map_sgas_def apply (auto) + using calculation(2) by auto + then show ?case using eo + using Cons.hyps calculation(2) la by auto +qed + +lemma z_signs_R: + fixes I:: "(nat list * nat list)" + fixes signs:: "rat list" + assumes lf: "list_all (\i. i < length signs) (fst(I))" + assumes ls: "list_all (\i. i < length signs) (snd(I))" + assumes la: "list_all (\s. s = 1 \ s = 0 \ s = -1) signs" + shows "(z_R I signs = 1) \ (z_R I signs = 0) \ (z_R I signs = -1)" + using assms z_signs_R2 z_signs_R1 unfolding z_R_def apply (auto) + by (metis (no_types, lifting) mult_cancel_left1 mult_minus1_right) + +lemma z_lemma_R: + fixes I:: "nat list * nat list" + fixes sign:: "rat list" + assumes consistent: "sign \ set (characterize_consistent_signs_at_roots p qs)" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows "(z_R I sign = 1) \ (z_R I sign = 0) \ (z_R I sign = -1)" +proof (rule z_signs_R) + have same: "length sign = length qs" using consistent + using characterize_consistent_signs_at_roots_def signs_at_def by fastforce + thus "(list_all (\i. i < length sign) (fst I))" + using welldefined1 + by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_def consistent_sign_vec_copr_def) + thus "(list_all (\i. i < length sign) (snd I))" + using same welldefined2 + by (auto simp add: list_constr_def characterize_consistent_signs_at_roots_def consistent_sign_vec_copr_def) + show "list_all (\s. s = 1 \ s = 0 \ s = - 1) sign" using consistent + apply (auto simp add: list.pred_map characterize_consistent_signs_at_roots_def consistent_sign_vec_def) + using Ball_set + by (simp add: list_all_length signs_at_def squash_def) +qed + +(* Show that all consistent sign vectors on roots of polynomials are in characterize_consistent_signs_at_roots *) +lemma in_set_R: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + shows "sign \ set (characterize_consistent_signs_at_roots p qs)" +proof - + have h1: "consistent_sign_vec qs x \ + set (remdups (map (signs_at qs) (sorted_list_of_set {x. poly p x = 0})))" + unfolding consistent_sign_vec_def signs_at_def squash_def + using root_p nonzero poly_roots_finite set_sorted_list_of_set apply (auto) + by (smt (verit, ccfv_SIG) Collect_cong comp_def image_eqI map_eq_conv mem_Collect_eq poly_roots_finite set_sorted_list_of_set) + thus ?thesis unfolding characterize_consistent_signs_at_roots_def characterize_root_list_p_def using sign_fix + by blast +qed + +lemma consistent_signs_prop_R: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + shows "list_all (\s. s = 1 \ s = 0 \ s = -1) sign" + using assms unfolding consistent_sign_vec_def squash_def apply (auto) + by (smt (z3) length_map list_all_length nth_map) + +(* The next few lemmas relate z_R to the signs of the product of subsets of polynomials of qs *) +lemma horiz_vector_helper_pos_ind_R1: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes asm: "list_constr I (length qs)" + shows "(prod_list (map (nth (map_sgas sign)) I)) = 1 \ + (\p \ set (retrieve_polys qs I). poly p x = 0)" + using asm +proof (induction "I") + case Nil + then show ?case unfolding map_sgas_def apply (auto) + by (simp add: retrieve_polys_def) +next + case (Cons a xa) + then have same0: "(prod_list (map (nth (map_sgas sign)) xa)) = 1 \ + (\p \ set (retrieve_polys qs xa). poly p x = 0)" unfolding list_constr_def by auto + have welldef: "a < length qs" using Cons.prems assms unfolding list_constr_def by auto + then have h: "prod_list (map ((!) (map_sgas sign)) (a#xa)) = (1 - (sign ! a)^2)*(prod_list (map ((!) (map_sgas sign)) (xa)))" + unfolding map_sgas_def using assms apply (auto) + by (metis (no_types, lifting) consistent_sign_vec_def length_map nth_map) + have "list_all (\s. s = 1 \ s = 0 \ s = -1) sign" using sign_fix unfolding consistent_sign_vec_def squash_def + apply (auto) + by (smt (z3) length_map list_all_length nth_map) + then have eo: "(prod_list (map ((!) (map_sgas sign)) (xa))) = 0 \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1" + using z_signs_R1 assms Cons.prems consistent_sign_vec_def length_map list_all_simps(1) length_map list_all_length list_constr_def + by (smt (verit, best)) + have "(sign ! a)^2 = 1 \ (sign ! a)^2 = 0" using sign_fix welldef unfolding consistent_sign_vec_def + by auto + then have s1: "(prod_list (map (nth (map_sgas sign)) (a#xa))) = 1 \ + ((sign ! a)^2 = 0 \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1)" + using eo h + by (metis (mono_tags, hide_lams) cancel_comm_monoid_add_class.diff_cancel diff_zero mult.left_neutral mult_zero_left) + have "(sign ! a)^2 = 0 \ (poly (qs ! a) x = 0)" + using sign_fix unfolding consistent_sign_vec_def welldef apply (auto) + apply (smt (z3) class_field.neg_1_not_0 class_field.zero_not_one nth_map welldef) + by (smt (z3) nth_map welldef) + then have same1:"(prod_list (map (nth (map_sgas sign)) (a#xa))) = 1 \ + ((poly (qs ! a) x = 0) \ (prod_list (map ((!) (map_sgas sign)) (xa))) = 1)" using s1 by auto + have "set (retrieve_polys qs (a#xa)) = {(qs ! a)} \ set (retrieve_polys qs xa)" + by (metis (no_types, lifting) insert_is_Un list.simps(15) list.simps(9) retrieve_polys_def) + then have same2:"(\p \ set (retrieve_polys qs (a#xa)). poly p x = 0) = ((poly (qs ! a) x = 0) + \ (\p \ set (retrieve_polys qs (xa)). poly p x = 0))" + by auto + then show ?case using same0 same1 same2 + assms by auto +qed + +lemma csv_length_same_as_qlist: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + shows "length sign = length qs" + using assms unfolding consistent_sign_vec_def by auto + +lemma horiz_vector_helper_zer_ind_R2: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes asm: "list_constr I (length qs)" + shows "(prod_list (map (nth sign) I)) = 0 \ + (poly (prod_list (retrieve_polys qs I)) x = 0)" + using asm +proof (induction "I") + case Nil + then show ?case unfolding map_sgas_def apply (auto) unfolding retrieve_polys_def + by auto next + case (Cons a xa) + have "poly (prod_list (retrieve_polys qs (a # xa))) x = (poly (qs ! a) x)*poly (prod_list (retrieve_polys qs (xa))) x" + by (simp add: retrieve_polys_def) + then show ?case using Cons.prems + by (smt (z3) Cons.IH class_field.neg_1_not_0 class_field.zero_not_one consistent_sign_vec_def list.simps(9) list_all_simps(1) list_constr_def mult_eq_0_iff nth_map prod_list.Cons sign_fix) +qed + +lemma horiz_vector_helper_pos_ind_R2: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes asm: "list_constr I (length qs)" + shows "(prod_list (map (nth sign) I)) = 1 \ + (poly (prod_list (retrieve_polys qs I)) x > 0)" + using asm +proof (induction "I") + case Nil + then show ?case unfolding map_sgas_def apply (auto) unfolding retrieve_polys_def + by auto next + case (Cons a xa) + then have ih: "(prod_list (map ((!) sign) xa) = 1) = (0 < poly (prod_list (retrieve_polys qs xa)) x)" + unfolding list_constr_def by auto + have lensame: "length sign = length qs" using sign_fix csv_length_same_as_qlist[of p x sign qs] + nonzero root_p by auto + have "poly (prod_list (retrieve_polys qs (a # xa))) x = (poly (qs ! a) x)*poly (prod_list (retrieve_polys qs (xa))) x" + by (simp add: retrieve_polys_def) + then have iff1: "(poly (prod_list (retrieve_polys qs (a#xa))) x > 0) \ + (((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ + ((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0))" + by (metis zero_less_mult_iff) + have prodsame: "(prod_list (map (nth sign) (a#xa))) = (sign ! a)* (prod_list (map (nth sign) (xa)))" + using lensame Cons.prems unfolding list_constr_def by auto + have sagt: "sign ! a = 1 \ (poly (qs ! a) x) > 0" using assms unfolding consistent_sign_vec_def + apply (auto) + apply (smt (verit, best) Cons.prems list_all_simps(1) list_constr_def neg_equal_zero nth_map zero_neq_one) + by (smt (verit, ccfv_threshold) Cons.prems list_all_simps(1) list_constr_def nth_map) + have salt: "sign ! a = -1 \ (poly (qs ! a) x) < 0" using assms unfolding consistent_sign_vec_def + apply (auto) + apply (smt (verit, ccfv_SIG) Cons.prems less_minus_one_simps(1) less_minus_one_simps(3) list_all_simps(1) list_constr_def neg_0_less_iff_less nth_map) + by (smt (verit, best) Cons.prems list_all_simps(1) list_constr_def nth_map) + have h1: "((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ + (prod_list (map (nth sign) (a#xa))) = 1" + using prodsame sagt ih by auto + have eo: "(prod_list (map ((!) sign) xa) = 1) \ (prod_list (map ((!) sign) xa) = 0) \ + (prod_list (map ((!) sign) xa) = -1)" + using Cons.prems consistent_signs_prop_R lensame list_constr_def nonzero root_p sign_fix z_signs_R2 by auto + have d1:"(prod_list (map ((!) sign) xa) = -1) \ (0 > poly (prod_list (retrieve_polys qs xa)) x)" + proof - + assume "(prod_list (map ((!) sign) xa) = -1) " + then show "(0 > poly (prod_list (retrieve_polys qs xa)) x)" + using prodsame salt ih assms Cons.prems class_field.neg_1_not_0 equal_neg_zero horiz_vector_helper_zer_ind_R2 linorder_neqE_linordered_idom list_all_simps(1) list_constr_def + apply (auto) + apply (smt (verit, ccfv_threshold) class_field.neg_1_not_0 list.set_map list_all_length semidom_class.prod_list_zero_iff) + by (smt (verit, ccfv_threshold) class_field.neg_1_not_0 list.set_map list_all_length semidom_class.prod_list_zero_iff) + qed + have d2: "(0 > poly (prod_list (retrieve_polys qs xa)) x) \ (prod_list (map ((!) sign) xa) = -1)" + using eo assms horiz_vector_helper_zer_ind_R2[where p = "p", where x = "x", where sign = "sign", where I ="I"] + apply (auto) + using ih apply force + by (metis (full_types, lifting) Cons.prems class_field.neg_1_not_0 horiz_vector_helper_zer_ind_R2 ih imageI list.set_map list_all_simps(1) list_constr_def mem_Collect_eq neg_equal_0_iff_equal semidom_class.prod_list_zero_iff) + have "(prod_list (map ((!) sign) xa) = -1) = (0 > poly (prod_list (retrieve_polys qs xa)) x)" + using d1 d2 + by blast + then have h2: "((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0) \ + (prod_list (map (nth sign) (a#xa))) = 1" + using prodsame salt ih by auto + have h3: "(prod_list (map (nth sign) (a#xa))) = 1 \ (((poly (qs ! a) x) > 0 \ poly (prod_list (retrieve_polys qs (xa))) x > 0) \ + ((poly (qs ! a) x) < 0 \ poly (prod_list (retrieve_polys qs (xa))) x < 0))" + using prodsame salt ih assms horiz_vector_helper_zer_ind_R2[where p = "p", where x = "x", where sign = "sign", where I ="I"] + by (smt (verit, ccfv_threshold) Cons.prems \poly (prod_list (retrieve_polys qs (a # xa))) x = poly (qs ! a) x * poly (prod_list (retrieve_polys qs xa)) x\ horiz_vector_helper_zer_ind_R2 mem_Collect_eq mult_cancel_left1 mult_not_zero sagt) + then show ?case using h1 h2 h3 iff1 Cons.prems by auto +qed + +lemma horiz_vector_helper_pos_ind_R: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list * nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes asm1: "list_constr (fst I) (length qs)" + assumes asm2: "list_constr (snd I) (length qs)" + shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" +proof - + have len: "length sign = length qs" using sign_fix csv_length_same_as_qlist[of p x sign qs] + nonzero root_p by auto + have d1: "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" + using assms horiz_vector_helper_pos_ind_R1[where p = "p", where qs = "qs", where sign = "sign", where x = "x", where I = "fst I"] + horiz_vector_helper_pos_ind_R2[where p = "p", where qs = "qs", where sign = "sign", where x = "x", where I = "snd I"] unfolding z_R_def + by auto + have d2: "(z_R I sign = 1) \ ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0))" + proof - + have h1: "(z_R I sign = 1) \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0)" + proof - + have "(prod_list (map (nth (map_sgas sign)) (fst I))) = 1 \ (prod_list (map (nth (map_sgas sign)) (fst I))) = 0" + using len consistent_signs_prop_R[where p = "p", where qs = "qs", where x = "x", where sign = "sign"] z_signs_R1[where signs = "sign", where I = "fst I"] assms + unfolding list_constr_def + by auto + then show ?thesis + using z_signs_R1[where signs = "sign", where I = "fst I"] horiz_vector_helper_pos_ind_R1[where sign = "sign", where I = "fst I", where p = "p", where x = "x"] assms + apply (auto) + by (metis (mono_tags, hide_lams) \prod_list (map ((!) (map_sgas sign)) (fst I)) = 1 \ prod_list (map ((!) (map_sgas sign)) (fst I)) = 0\ mult_zero_left z_R_def) + qed + then have h2: "(z_R I sign = 1) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)" + unfolding z_R_def using assms horiz_vector_helper_pos_ind_R2[where p = "p", where x = "x", where sign = "sign", where qs = "qs", where I ="snd I"] + by (metis horiz_vector_helper_pos_ind_R1 mult.left_neutral) + show ?thesis using h1 h2 by auto + qed + show ?thesis using d1 d2 by blast +qed + +lemma horiz_vector_helper_pos_R: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x > 0)) \ (z_R I sign = 1)" + using horiz_vector_helper_pos_ind_R + using nonzero root_p sign_fix welldefined1 welldefined2 by blast + +lemma horiz_vector_helper_neg_R: + fixes p:: "real poly" + assumes nonzero: "p\0" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes sign:: "rat list" + fixes x:: "real" + assumes root_p: "x \ {x. poly p x = 0}" + assumes sign_fix: "sign = consistent_sign_vec qs x" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0)) \ (z_R I sign = -1)" +proof - + have set_hyp: "sign \ set (characterize_consistent_signs_at_roots p qs)" + using in_set_R[of p x sign qs] nonzero root_p sign_fix by blast + have z_hyp: "((z_R I sign = 1) \ (z_R I sign = 0) \ (z_R I sign = -1))" + using welldefined1 welldefined2 set_hyp z_lemma_R[where sign="sign", where I = "I", where p="p", where qs="qs"] by blast + have d1: "((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0)) \ (z_R I sign = -1)" + using horiz_vector_helper_pos_R + using nonzero root_p sign_fix welldefined1 welldefined2 + by (smt (verit, ccfv_threshold) horiz_vector_helper_pos_ind_R1 horiz_vector_helper_zer_ind_R2 mem_Collect_eq mult_eq_0_iff z_R_def z_hyp) + have d2: "(z_R I sign = -1) \ ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x < 0))" + using horiz_vector_helper_pos_ind_R1 horiz_vector_helper_zer_ind_R2 nonzero root_p sign_fix welldefined1 welldefined2 + by (smt (verit, ccfv_threshold) class_field.neg_1_not_0 consistent_sign_vec_def consistent_signs_prop_R equal_neg_zero horiz_vector_helper_pos_ind_R2 length_map list_all_length list_constr_def mem_Collect_eq mem_Collect_eq mult_cancel_left1 mult_not_zero retrieve_polys_def z_R_def z_signs_R1 zero_neq_one) + then show ?thesis using d1 d2 + by linarith +qed + +lemma lhs_dot_rewrite: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + shows + "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = + sum_list (map (\s. (z_R I s) * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs)" +proof - + have "p \ 0" using nonzero by auto + from construct_lhs_vector_cleaner[OF this] + have rhseq: "construct_lhs_vector_R p qs signs = + vec_of_list + (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs)" + using construct_lhs_vector_cleaner_R nonzero by presburger + have "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = + sum_list (map2 (*) (mtx_row_R signs I) (map (\s. rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) signs))" + unfolding rhseq + apply (intro vec_of_list_dot_rewrite) + by (auto simp add: mtx_row_R_def) + thus ?thesis unfolding mtx_row_R_def + using map2_map_map + by (auto simp add: map2_map_map) +qed + +(* If we have a superset of the signs, we can drop to just the consistent ones *) +lemma construct_lhs_vector_drop_consistent_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows + "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = + (vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ + (construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)))" +proof - + have h0: "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ + 0 < rat_of_nat (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ z_R I sgn = 0" + proof - + have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (card + {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn} \ {}" + by fastforce + then show ?thesis + proof - + { fix iis :: "rat list" + have ff1: "0 \ p" + using nonzero rsquarefree_def by blast + obtain rr :: "(real \ bool) \ real" where + ff2: "\p. p (rr p) \ Collect p = {}" + by moura + { assume "\is. is = iis \ {r. poly p r = 0 \ consistent_sign_vec qs r = is} \ {}" + then have "\is. consistent_sign_vec qs (rr (\r. poly p r = 0 \ consistent_sign_vec qs r = is)) = iis \ {r. poly p r = 0 \ consistent_sign_vec qs r = is} \ {}" + using ff2 + by (metis (mono_tags, lifting)) + then have "\r. poly p r = 0 \ consistent_sign_vec qs r = iis" + using ff2 + by smt + then have "iis \ consistent_sign_vec qs ` set (sorted_list_of_set {r. poly p r = 0})" + using ff1 poly_roots_finite + by (metis (mono_tags) imageI mem_Collect_eq set_sorted_list_of_set) } + then have "iis \ set signs \ iis \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ \ 0 < rat_of_int (int (card {r. poly p r = 0 \ consistent_sign_vec qs r = iis}))" + by (metis (no_types) \\sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ 0 < rat_of_int (int (card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn})) \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn} \ {}\ characterize_root_list_p_def) } + then show ?thesis + by fastforce + qed + qed + then have "\ sgn. sgn \ set signs \ sgn \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ ((0 = rat_of_nat (card + {xa. poly p xa = 0 \ consistent_sign_vec qs xa = sgn}) \ z_R I sgn = 0))" + by auto + then have hyp: "\ s. s \ set signs \ s \ consistent_sign_vec qs ` set (characterize_root_list_p p) \ (z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = 0)" + by auto + then have "(\s\ set(signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + proof - + have "set(signs) =(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p)))" + by blast + then have sum_rewrite: "(\s\ set(signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + by auto + then have sum_split: "(\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))) \ + (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + = +(\s\ (set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) ++ (\s\ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + by (metis (no_types, lifting) List.finite_set sum.Int_Diff) + have sum_zero: "(\s\ (set(signs)-(consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" + using hyp + by (simp add: hyp) + show ?thesis using sum_rewrite sum_split sum_zero by linarith + qed + then have set_eq: "set (remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p))) = set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))" + using all_info apply (simp add: characterize_consistent_signs_at_roots_def subset_antisym) + by (smt (z3) Int_subset_iff consistent_sign_vec_def list.set_map map_eq_conv o_apply signs_at_def squash_def subsetI subset_antisym) + have hyp1: "(\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\set (signs). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + using distinct_signs sum_list_distinct_conv_sum_set by blast + have hyp2: "(\s\remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + = (\s\ set (remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + using sum_list_distinct_conv_sum_set by blast + have set_sum_eq: "(\s\(set (signs) \ (consistent_sign_vec qs ` set (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\ set (remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p))). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + using set_eq by auto + then have "(\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + using set_sum_eq hyp1 hyp2 + using \(\s\set signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = (\s\set signs \ consistent_sign_vec qs ` set (characterize_root_list_p p). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))\ by linarith + then have "consistent_sign_vec qs ` set (characterize_root_list_p p) \ set signs \ + (\p qss. + characterize_consistent_signs_at_roots p qss = + remdups (map (consistent_sign_vec qss) (characterize_root_list_p p))) \ + (\s\signs. z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + (\s\remdups + (map (consistent_sign_vec qs) + (characterize_root_list_p p)). z_R I s * rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + by linarith + then show ?thesis unfolding lhs_dot_rewrite[OF nonzero] + apply (auto intro!: sum_list_distinct_filter simp add: distinct_signs consistent_sign_vec_def characterize_consistent_signs_at_roots_def) + using all_info consistent_sign_vec_def characterize_consistent_signs_at_roots_def + by (smt (z3) list.set_map map_eq_conv o_apply set_remdups signs_at_def squash_def) +qed + +lemma matrix_equation_helper_step_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows "(vec_of_list (mtx_row_R signs I) \ (construct_lhs_vector_R p qs signs)) = + rat_of_int (card {x. poly p x = 0 \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ poly (prod_list (retrieve_polys qs (snd I))) x > 0}) - + rat_of_int (card {x. poly p x = 0 \ (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ poly (prod_list (retrieve_polys qs (snd I))) x < 0})" +proof - + have finset: "finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))" by auto + let ?gt = "(set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = 1})" + let ?lt = " (set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = -1})" + let ?zer = "(set (map (consistent_sign_vec qs) (characterize_root_list_p p)) \ {s. z_R I s = 0})" + have eq: "set (map (consistent_sign_vec qs) (characterize_root_list_p p)) = (?gt \ ?lt) \ ?zer" + proof safe + fix x + assume h:"x \ set (map (consistent_sign_vec qs) (characterize_root_list_p p))" + "z_R I x \ 0" "z_R I x \ - 1" + then have "x \ set (characterize_consistent_signs_at_roots p qs)" + unfolding characterize_consistent_signs_at_roots_def + by (smt (verit, del_insts) characterize_consistent_signs_at_roots_def characterize_root_list_p_def imageE in_set_R nonzero poly_roots_finite set_map sorted_list_of_set(1)) + thus "z_R I x = 1" + using h welldefined1 welldefined2 z_lemma_R by blast + qed + have sumeq: "(\s\(?gt\?lt). z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + = (\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + + (\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + apply (rule sum.union_disjoint) by auto + (* First, drop the signs that are irrelevant *) + from construct_lhs_vector_drop_consistent_R[OF assms(1-5)] have + "vec_of_list (mtx_row_R signs I) \ construct_lhs_vector_R p qs signs = + vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ + construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)" . + (* Now we split the sum *) + from lhs_dot_rewrite[OF assms(1)] + moreover have "... = + (\s\characterize_consistent_signs_at_roots p qs. + z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" . + moreover have "... = + (\s\set (map (consistent_sign_vec qs) (characterize_root_list_p p)). + z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" unfolding characterize_consistent_signs_at_roots_def sum_code[symmetric] + apply (auto) + by (smt (verit, best) consistent_sign_vec_def list.set_map map_eq_conv o_apply signs_at_def squash_def sum.set_conv_list) + moreover have setc1:"... = + (\s\(?gt\?lt). z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + + (\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) " + apply (subst eq) + apply (rule sum.union_disjoint) by auto + ultimately have setc: "... = + (\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + + (\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) + + (\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))" + using sumeq by auto + have "\s \ ?zer. (z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" + by auto + then have obvzer: "(\s\?zer. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = 0" + by auto + (* Now recharacterize lt, gt*) + have setroots: "set (characterize_root_list_p p) = {x. poly p x = 0}" unfolding characterize_root_list_p_def + using poly_roots_finite nonzero rsquarefree_def set_sorted_list_of_set by blast + have *: "\s. {x. poly p x = 0 \ consistent_sign_vec qs x = s} = + {x \{x. poly p x = 0}. consistent_sign_vec qs x = s}" + by auto + have e1: "(\s\consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. + card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = + (sum (\x. if (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x) > 0 then 1 else 0) {x. poly p x = 0})" + unfolding * apply (rule sum_multicount_gen) + using \finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))\ setroots apply auto[1] + apply (metis List.finite_set setroots) + proof safe + fix x + assume rt: "poly p x = 0" + then have 1: "{s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. consistent_sign_vec qs x = s} = + {s. z_R I s = 1 \ consistent_sign_vec qs x = s}" + by auto + have 2: "... = {s. (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 < poly (prod_list (retrieve_polys qs (snd I))) x) \ consistent_sign_vec qs x = s}" + using horiz_vector_helper_pos_R assms welldefined1 welldefined2 rt by blast + have 3: "... = (if (\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 < poly (prod_list (retrieve_polys qs (snd I))) x) then {consistent_sign_vec qs x} else {})" + by auto + then have "card {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. consistent_sign_vec qs x = s} = + (if ((\p \ set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 < poly (prod_list (retrieve_polys qs (snd I))) x) + then 1 else 0)" using 1 2 3 by auto + thus " card + {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = 1}. + consistent_sign_vec qs x = s} = + (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ + 0 < poly (prod_list (retrieve_polys qs (snd I))) x + then 1 else 0)" + by auto + qed + + have gtchr: "(\s\?gt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + rat_of_int (card {x. poly p x = 0 \ (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 < poly (prod_list (retrieve_polys qs (snd I))) x})" + apply (auto simp add: setroots) + apply (subst of_nat_sum[symmetric]) + apply (subst of_nat_eq_iff) + apply (subst e1) + apply (subst card_eq_sum) + apply (rule sum.mono_neutral_cong_right) + apply (metis List.finite_set setroots) + by auto + have e2: " (\s\consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = - 1}. + card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) = + (sum (\x. if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ (poly (prod_list (retrieve_polys qs (snd I))) x) < 0 then 1 else 0) {x. poly p x = 0})" + unfolding * apply (rule sum_multicount_gen) + using \finite (set (map (consistent_sign_vec qs) (characterize_root_list_p p)))\ setroots apply auto[1] + apply (metis List.finite_set setroots) + proof safe + fix x + assume rt: "poly p x = 0" + then have 1: "{s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = -1}. consistent_sign_vec qs x = s} = + {s. z_R I s = -1 \ consistent_sign_vec qs x = s}" + by auto + have 2: "... = {s. ((\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly (prod_list (retrieve_polys qs (snd I))) x) \ consistent_sign_vec qs x = s}" + using horiz_vector_helper_neg_R assms rt welldefined1 welldefined2 by blast + have 3: "... = (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ (0 > poly (prod_list (retrieve_polys qs (snd I))) x) then {consistent_sign_vec qs x} else {})" + by auto + thus "card {s \ consistent_sign_vec qs ` {x. poly p x = 0} \ {s. z_R I s = -1}. consistent_sign_vec qs x = s} = + (if (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly + (prod_list + (retrieve_polys qs (snd I))) + x + then 1 else 0)" using 1 2 3 by auto + qed + have ltchr: "(\s\?lt. z_R I s * rat_of_int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})) = + - rat_of_int (card {x. poly p x = 0 \ (\p\set (retrieve_polys qs (fst I)). poly p x = 0) \ 0 > poly (prod_list (retrieve_polys qs (snd I))) x})" + apply (auto simp add: setroots sum_negf) + apply (subst of_nat_sum[symmetric]) + apply (subst of_nat_eq_iff) + apply (subst e2) + apply (subst card_eq_sum) + apply (rule sum.mono_neutral_cong_right) + apply (metis List.finite_set setroots) + by auto + show ?thesis using gtchr ltchr obvzer setc + using \(\s\characterize_consistent_signs_at_roots p qs. z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s}))) = (\s\set (map (consistent_sign_vec qs) (characterize_root_list_p p)). z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})))\ \vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs) = (\s\characterize_consistent_signs_at_roots p qs. z_R I s * rat_of_int (int (card {x. poly p x = 0 \ consistent_sign_vec qs x = s})))\ \vec_of_list (mtx_row_R signs I) \ construct_lhs_vector_R p qs signs = vec_of_list (mtx_row_R (characterize_consistent_signs_at_roots p qs) I) \ construct_lhs_vector_R p qs (characterize_consistent_signs_at_roots p qs)\ setc1 by linarith +qed + +lemma matrix_equation_main_step_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes I:: "nat list*nat list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" + assumes welldefined1: "list_constr (fst I) (length qs)" + assumes welldefined2: "list_constr (snd I) (length qs)" + shows "(vec_of_list (mtx_row_R signs I) \ + (construct_lhs_vector_R p qs signs)) = + construct_NofI_R p (retrieve_polys qs (fst I)) (retrieve_polys qs (snd I))" +proof - + show ?thesis + unfolding construct_NofI_prop_R[OF nonzero] + using matrix_equation_helper_step_R[OF assms] + by linarith +qed + +lemma mtx_row_length_R: + "list_all (\r. length r = length signs) (map (mtx_row_R signs) ls)" + apply (induction ls) + by (auto simp add: mtx_row_R_def) + +(* Shows that as long as we have a "basis" of sign assignments (see assumptions all_info, welldefined), + and some other mild assumptions on our inputs (given in nonzero, distinct_signs), the construction + will be satisfied *) +theorem matrix_equation_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + fixes subsets:: "(nat list*nat list) list" + fixes signs:: "rat list list" + assumes nonzero: "p\0" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" + assumes welldefined: "all_list_constr_R (subsets) (length qs)" + shows "satisfy_equation_R p qs subsets signs" + unfolding satisfy_equation_R_def matrix_A_R_def + construct_lhs_vector_R_def construct_rhs_vector_R_def all_list_constr_R_def + apply (subst mult_mat_vec_of_list) + apply (auto simp add: mtx_row_length_R intro!: map_vec_vec_of_list_eq_intro) + using matrix_equation_main_step_R[OF assms(1-3), unfolded construct_lhs_vector_R_def] + using all_list_constr_R_def in_set_member welldefined by fastforce + +(* Prettifying some theorems*) + +lemma consistent_signs_at_roots_eq: + assumes "p \ 0" + shows "consistent_signs_at_roots p qs = + set (characterize_consistent_signs_at_roots p qs)" + unfolding consistent_signs_at_roots_def characterize_consistent_signs_at_roots_def + characterize_root_list_p_def + apply auto + apply (subst set_sorted_list_of_set) + using assms poly_roots_finite apply blast + unfolding sgn_vec_def sgn_def signs_at_def squash_def o_def + using roots_def apply auto[1] + by (smt Collect_cong assms image_iff poly_roots_finite roots_def sorted_list_of_set(1)) + +abbreviation w_vec_R:: "real poly \ real poly list \ rat list list \ rat vec" + where "w_vec_R \ construct_lhs_vector_R" + +abbreviation v_vec_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat vec" + where "v_vec_R \ construct_rhs_vector_R" + +abbreviation M_mat_R:: "rat list list \ (nat list*nat list) list \ rat mat" + where "M_mat_R \ matrix_A_R" + +theorem matrix_equation_pretty: + fixes subsets:: "(nat list*nat list) list" + assumes "p\0" + assumes "distinct signs" + assumes "consistent_signs_at_roots p qs \ set signs" + assumes "\a b i. (a, b) \ set ( subsets) \ (i \ set a \ i \ set b) \ i < length qs" + shows "M_mat_R signs subsets *\<^sub>v w_vec_R p qs signs = v_vec_R p qs subsets" + unfolding satisfy_equation_R_def[symmetric] + using matrix_equation_R[of p signs qs subsets] assms + using consistent_signs_at_roots_eq unfolding all_list_constr_R_def list_constr_def apply (auto) + by (metis (no_types, lifting) Ball_set in_set_member) + +section "Base Case" +definition satisfies_properties_R:: "real poly \ real poly list \ (nat list*nat list) list \ rat list list \ rat mat \ bool" + where "satisfies_properties_R p qs subsets signs matrix = + ( all_list_constr_R subsets (length qs) \ well_def_signs (length qs) signs \ distinct signs + \ satisfy_equation_R p qs subsets signs \ invertible_mat matrix \ matrix = matrix_A_R signs subsets + \ set (characterize_consistent_signs_at_roots p qs) \ set(signs) + )" + +lemma mat_base_case_R: + shows "matrix_A_R [[1],[0],[-1]] [([], []),([0], []),([], [0])] = (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" + unfolding matrix_A_R_def mtx_row_R_def z_R_def map_sgas_def apply (auto) + by (simp add: numeral_3_eq_3) + +lemma base_case_sgas_R: + fixes q p:: "real poly" + assumes nonzero: "p \ 0" + shows "set (characterize_consistent_signs_at_roots p [q]) \ {[1],[0], [- 1]}" + unfolding characterize_consistent_signs_at_roots_def signs_at_def apply (auto) + by (meson squash_def) + +lemma base_case_sgas_alt_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + assumes len1: "length qs = 1" + assumes nonzero: "p \ 0" + shows "set (characterize_consistent_signs_at_roots p qs) \ {[1], [0], [- 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_R nonzero + by auto +qed + +lemma base_case_satisfy_equation_R: + fixes q p:: "real poly" + assumes nonzero: "p \ 0" + shows "satisfy_equation_R p [q] [([], []),([0], []),([], [0])] [[1],[0],[-1]] " +proof - + have h1: "set (characterize_consistent_signs_at_roots p [q]) \ {[1], [0],[- 1]}" + using base_case_sgas_R assms by auto + have h2: "all_list_constr_R [([], []),([0], []),([], [0])] (Suc 0)" unfolding all_list_constr_R_def + by (simp add: list_constr_def member_def) + then show ?thesis using matrix_equation_R[where p = "p", where qs = "[q]", where signs = "[[1],[0],[-1]] ", where subsets = "[([], []),([0], []),([], [0])]"] + nonzero h1 h2 by auto +qed + +lemma base_case_satisfy_equation_alt_R: + fixes p:: "real poly" + fixes qs:: "real poly list" + assumes len1: "length qs = 1" + assumes nonzero: "p \ 0" + shows "satisfy_equation_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-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_R nonzero + by auto +qed + +lemma base_case_matrix_eq: + fixes q p:: "real poly" + assumes nonzero: "p \ 0" + shows "(mult_mat_vec (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]) (construct_lhs_vector_R p [q] [[1],[0],[-1]]) = + (construct_rhs_vector_R p [q] [([], []),([0], []),([], [0])]))" + using mat_base_case_R base_case_satisfy_equation_R unfolding satisfy_equation_R_def + by (simp add: nonzero) + +lemma less_three: "(n::nat) < Suc (Suc (Suc 0)) \ n = 0 \ n = 1 \ n = 2" + by auto + +lemma inverse_mat_base_case_R: + shows "inverts_mat (mat_of_rows_list 3 [[1/2, -1/2, 1/2], [0, 1, 0], [1/2, -1/2, -1/2]]::rat mat) (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" + unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto + unfolding less_three by (auto simp add: scalar_prod_def) + +lemma inverse_mat_base_case_2_R: + shows "inverts_mat (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat) (mat_of_rows_list 3 [[1/2, -1/2, 1/2], [0, 1, 0], [1/2, -1/2, -1/2]]:: rat mat)" + unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto + unfolding less_three by (auto simp add: scalar_prod_def) + +lemma base_case_invertible_mat_R: + shows "invertible_mat (matrix_A_R [[1],[0], [- 1]] [([], []),([0], []),([], [0])])" + unfolding invertible_mat_def using inverse_mat_base_case_R inverse_mat_base_case_2_R + apply (auto) + apply (simp add: mat_base_case mat_of_rows_list_def) + using mat_base_case_R by auto + +section "Inductive Step" + +(***** Need some properties of smashing; smashing signs is just as it was in BKR *****) +subsection "Lemmas on smashing subsets " + +definition subsets_first_component_list::"(nat list*nat list) list \ nat list list" + where "subsets_first_component_list I = map (\I. (fst I)) I" + +definition subsets_second_component_list::"(nat list*nat list) list \ nat list list" + where "subsets_second_component_list I = map (\I. (snd I)) I" + +definition smash_list_list::"('a list*'a list) list \ ('a list*'a list) list \ ('a list*'a list) list" + where "smash_list_list s1 s2 = concat (map (\l1. map (\l2. (fst l1 @ fst l2, snd l1 @ snd l2)) s2) s1)" + +lemma smash_list_list_property_set: + fixes l1 l2 :: "('a list*'a list) list" + fixes a b:: "nat" + shows "\ (elem :: ('a list*'a list)). (elem \ (set (smash_list_list l1 l2)) \ + (\ (elem1:: ('a list*'a list)). \ (elem2:: ('a list*'a list)). + (elem1 \ set(l1) \ elem2 \ set(l2) \ elem = (fst elem1@ fst elem2, snd elem1 @ snd elem2))))" +proof clarsimp + fix a b + assume assum: "(a, b) \ set (smash_list_list l1 l2)" + show "\aa ba. (aa, ba) \ set l1 \ (\ab bb. (ab, bb) \ set l2 \ a = aa @ ab \ b = ba @ bb)" + using assum unfolding smash_list_list_def + apply (auto) by blast +qed + + +lemma subsets_smash_property_R: + fixes subsets1 subsets2 :: "(nat list*nat list) list" + fixes n:: "nat" + shows "\ (elem :: nat list*nat list). (List.member (subsets_smash_R n subsets1 subsets2) elem) \ + (\ (elem1::nat list*nat list). \ (elem2::nat list*nat list). + (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = ((fst elem1) @ (map ((+) n) (fst elem2)), (snd elem1) @ (map ((+) n) (snd elem2)))))" +proof - + let ?fst_component2 = "subsets_first_component_list subsets2" + let ?snd_component2 = "subsets_second_component_list subsets2" + let ?new_subsets = "map (\I. ((map ((+) n)) (fst I), (map ((+) n)) (snd I))) subsets2" + (* a slightly unorthodox use of signs_smash, but it closes the proof *) + have "subsets_smash_R n subsets1 subsets2 = smash_list_list subsets1 ?new_subsets" + unfolding subsets_smash_R_def smash_list_list_def apply (auto) + by (simp add: comp_def) + then show ?thesis using smash_list_list_property_set[of subsets1 ?new_subsets] apply (auto) + using imageE in_set_member set_map smash_list_list_property_set + by (smt (z3) prod.exhaust_sel) +qed + +(* If subsets for smaller systems are well-defined, then subsets for combined systems are, too *) +subsection "Well-defined subsets preserved when smashing" + +lemma well_def_step_R: + fixes qs1 qs2 :: "real poly list" + fixes subsets1 subsets2 :: "(nat list*nat list) list" + assumes well_def_subsets1: "all_list_constr_R (subsets1) (length qs1)" + assumes well_def_subsets2: "all_list_constr_R (subsets2) (length qs2)" + shows "all_list_constr_R ((subsets_smash_R (length qs1) subsets1 subsets2)) + (length (qs1 @ qs2))" +proof - + let ?n = "(length qs1)" + have h1: "\elem. + List.member (subsets_smash_R ?n subsets1 subsets2) elem \ + (\ (elem1::nat list*nat list). \ (elem2::nat list*nat list). + (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = ((fst elem1) @ (map ((+) ?n) (fst elem2)), (snd elem1) @ (map ((+) ?n) (snd elem2)))))" + using subsets_smash_property_R by blast + have h2: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr ((fst elem1) @ map ((+) (length qs1)) (fst elem2)) (length (qs1 @ qs2))" + proof clarsimp + fix elem1 b elem2 ba + assume e1: "(elem1, b) \ set subsets1" + assume e2: "(elem2, ba) \ 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_R_def + apply (auto) + 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_R_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 + have h3: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr ((snd elem1) @ map ((+) (length qs1)) (snd elem2)) (length (qs1 @ qs2))" + proof clarsimp + fix elem1 b elem2 ba + assume e1: "(b, elem1) \ set subsets1" + assume e2: "(ba, 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_R_def + apply (auto) + 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_R_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 + show ?thesis using h1 h2 h3 unfolding all_list_constr_def + by (metis all_list_constr_R_def fst_conv snd_conv) +qed + +subsection "Consistent Sign Assignments Preserved When Smashing" + +lemma subset_helper_R: + fixes p:: "real poly" + fixes qs1 qs2 :: "real poly list" + fixes signs1 signs2 :: "rat list list" + shows "\ x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). + \ x1 \ set (characterize_consistent_signs_at_roots p qs1). + \ x2 \ set (characterize_consistent_signs_at_roots p qs2). + x = x1@x2" +proof clarsimp + fix x + assume x_in: "x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2))" + have x_in_csv: "x \ set(map (consistent_sign_vec (qs1 @ qs2)) (characterize_root_list_p p))" + using x_in unfolding characterize_consistent_signs_at_roots_def + by (smt (z3) consistent_sign_vec_def map_eq_conv o_apply set_remdups signs_at_def squash_def) + have csv_hyp: "\r. consistent_sign_vec (qs1 @ qs2) r = (consistent_sign_vec qs1 r)@(consistent_sign_vec qs2 r)" + unfolding consistent_sign_vec_def by auto + have exr_iff: "(\r \ set (characterize_root_list_p p). x = consistent_sign_vec (qs1 @ qs2) r) \ (\r \ set (characterize_root_list_p p). x = (consistent_sign_vec qs1 r)@(consistent_sign_vec qs2 r))" + using csv_hyp by auto + have exr: "x \ set(map (consistent_sign_vec (qs1 @ qs2)) (characterize_root_list_p p)) \ (\r \ set (characterize_root_list_p p). x = consistent_sign_vec (qs1 @ qs2) r)" + by auto + have mem_list1: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec qs1 r) \ set(map (consistent_sign_vec qs1) (characterize_root_list_p p))" + by simp + have mem_list2: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec qs2 r) \ set(map (consistent_sign_vec qs2) (characterize_root_list_p p))" + by simp + then show "\x1\set (characterize_consistent_signs_at_roots p qs1). + \x2\set (characterize_consistent_signs_at_roots p qs2). x = x1 @ x2" + using x_in_csv exr mem_list1 mem_list2 characterize_consistent_signs_at_roots_def exr_iff + using imageE image_eqI map_append set_map set_remdups signs_at_def x_in + by auto +qed + +lemma subset_step_R: + fixes p:: "real poly" + fixes qs1 qs2 :: "real poly list" + fixes signs1 signs2 :: "rat list list" + assumes csa1: "set (characterize_consistent_signs_at_roots p qs1) \ set (signs1)" + assumes csa2: "set (characterize_consistent_signs_at_roots p qs2) \ set (signs2)" + shows "set (characterize_consistent_signs_at_roots p + (qs1 @ qs2)) + \ set (signs_smash signs1 signs2)" +proof - + have h0: "\ x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots p qs1). \ x2 \ set (characterize_consistent_signs_at_roots p qs2). + (x = x1@x2)" using subset_helper_R[of p qs1 qs2] + by blast + then have "\x \ set (characterize_consistent_signs_at_roots p (qs1 @ qs2)). + x \ set (signs_smash (characterize_consistent_signs_at_roots p qs1) + (characterize_consistent_signs_at_roots p qs2))" + unfolding signs_smash_def apply (auto) + by fastforce + then have "\x \ set (characterize_consistent_signs_at_roots 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_matrix_A_R[simp]: + shows "dim_row (matrix_A_R signs subsets) = length subsets" + unfolding matrix_A_R_def by auto + +lemma dim_col_matrix_A_R[simp]: + shows "dim_col (matrix_A_R signs subsets) = length signs" + unfolding matrix_A_R_def by auto + +lemma length_subsets_smash_R: + shows + "length (subsets_smash_R n subs1 subs2) = length subs1 * length subs2" + unfolding subsets_smash_R_def length_concat + by (auto simp add: o_def sum_list_triv) + +lemma z_append_R: + fixes xs:: "(nat list * nat list)" + assumes "\i. i \ set (fst xs) \ i < length as" + assumes "\i. i \ set (snd xs) \ i < length as" + shows "z_R ((fst xs) @ (map ((+) (length as)) (fst ys)), (snd xs) @ (map ((+) (length as)) (snd ys))) (as @ bs) = z_R xs as * z_R ys bs" +proof - + have 1: "map ((!) (as @ bs)) (fst xs) = map ((!) as) (fst xs)" + unfolding list_eq_iff_nth_eq + using assms by (auto simp add:nth_append) + have 2: "map ((!) (as @ bs) \ (+) (length as)) (fst ys) = map ((!) bs) (fst ys)" + unfolding list_eq_iff_nth_eq + using assms by auto + have 3: "map ((!) (as @ bs)) (snd xs) = map ((!) as) (snd xs)" + unfolding list_eq_iff_nth_eq + using assms by (auto simp add:nth_append) + have 4: "map ((!) (as @ bs) \ (+) (length as)) (snd ys) = map ((!) bs) (snd ys)" + unfolding list_eq_iff_nth_eq + using assms by auto + show ?thesis + unfolding z_R_def apply auto + unfolding 1 2 3 4 apply (auto) + by (smt (z3) assms(1) comp_apply length_map map_append map_eq_conv map_sgas_def nth_append nth_append_length_plus) +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_R: + fixes qs1 :: "real poly list" + fixes subs1 subs2 :: "(nat list*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 (fst l) \ i \ set (snd l)) \ i < n1" + assumes "\j. j \ set signs1 \ length j = n1" + shows "(matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2)) = + kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R signs2 subs2)" + unfolding mat_eq_iff dim_row_matrix_A_R dim_col_matrix_A_R + length_subsets_smash_R length_signs_smash dim_kronecker +proof safe + fix i j + assume i: "i < length subs1 * length subs2" + assume 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_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2) $$ (i, j) = + z_R (subsets_smash_R n1 subs1 subs2 ! i) (signs_smash signs1 signs2 ! j)" + unfolding mat_of_rows_list_def matrix_A_R_def mtx_row_R_def + using i j by (auto simp add: length_signs_smash length_subsets_smash_R) + have 2: " ... = z_R ((fst (subs1 ! (i div length subs2)) @ map ((+) n1) (fst(subs2 ! (i mod length subs2)))), + (snd (subs1 ! (i div length subs2)) @ map ((+) n1) (snd (subs2 ! (i mod length subs2))))) + (signs1 ! (j div length signs2) @ signs2 ! (j mod length signs2))" + unfolding signs_smash_def subsets_smash_R_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_R (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * + z_R (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" + unfolding n1 + apply (subst z_append_R) + apply (auto simp add: n1[symmetric]) + using assms(1) ld(1) nth_mem + apply blast + using assms(1) ld(1) nth_mem by blast + have 4: "kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R signs2 subs2) $$ (i,j) = + z_R (subs1 ! (i div length subs2)) + (signs1 ! (j div length signs2)) * + z_R (subs2 ! (i mod length subs2)) + (signs2 ! (j mod length signs2))" + unfolding kronecker_product_def matrix_A_R_def mat_of_rows_list_def mtx_row_R_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_R (signs_smash signs1 signs2) (subsets_smash_R n1 subs1 subs2) $$ (i, j) = + kronecker_product (matrix_A_R signs1 subs1) (matrix_A_R 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_R: + fixes p:: "real poly" + fixes qs1 qs2 :: "real poly list" + fixes subsets1 subsets2 :: "(nat list*nat list) list" + fixes signs1 signs2 :: "rat list list" + assumes nonzero: "p \ 0" + assumes nontriv1: "length qs1 > 0" + assumes nontriv2: "length qs2 > 0" + 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 p qs1) \ set(signs1)" + assumes all_info2: "set (characterize_consistent_signs_at_roots p qs2) \ set(signs2)" + assumes welldefined_subsets1: "all_list_constr_R (subsets1) (length qs1)" + assumes welldefined_subsets2: "all_list_constr_R (subsets2) (length qs2)" + assumes invertibleMtx1: "invertible_mat (matrix_A_R signs1 subsets1)" + assumes invertibleMtx2: "invertible_mat (matrix_A_R signs2 subsets2)" + shows "satisfy_equation_R p (qs1@qs2) (subsets_smash_R (length qs1) subsets1 subsets2) (signs_smash signs1 signs2) + \ invertible_mat (matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R (length qs1) subsets1 subsets2))" +proof - + have h1: "invertible_mat (matrix_A_R (signs_smash signs1 signs2) (subsets_smash_R (length qs1) subsets1 subsets2))" + using matrix_construction_is_kronecker_product_R + kronecker_invertible invertibleMtx1 invertibleMtx2 + welldefined_subsets1 welldefined_subsets2 unfolding all_list_constr_R_def list_constr_def + using Ball_set in_set_member well_def_signs_def welldefined_signs1 in_set_conv_nth list_all_length + apply (auto) + by (smt (z3) Ball_set kronecker_invertible member_def) + have h2a: "distinct (signs_smash signs1 signs2)" + using distinct_signs1 distinct_signs2 welldefined_signs1 welldefined_signs2 nontriv1 nontriv2 + distinct_step by auto + have h2c: "all_list_constr_R ((subsets_smash_R (length qs1) subsets1 subsets2)) (length (qs1@qs2))" + using well_def_step_R + welldefined_subsets1 welldefined_subsets2 + by blast + have h2d: "set (characterize_consistent_signs_at_roots p (qs1@qs2)) \ set(signs_smash signs1 signs2)" + using subset_step_R all_info1 all_info2 + by simp + have h2: "satisfy_equation_R p (qs1@qs2) (subsets_smash_R (length qs1) subsets1 subsets2) (signs_smash signs1 signs2)" + using matrix_equation_R[where p="p", where qs="qs1@qs2", where subsets = "subsets_smash_R (length qs1) subsets1 subsets2", + where signs = "signs_smash signs1 signs2"] + h2a 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_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ rat mat" + where "get_matrix_R data = fst(data)" + +definition get_subsets_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ (nat list*nat list) list" + where "get_subsets_R data = fst(snd(data))" + +definition get_signs_R:: "(rat mat \ ((nat list*nat list) list \ rat list list)) \ rat list list" + where "get_signs_R data = snd(snd(data))" + +definition reduction_signs_R:: "real poly \ real poly list \ rat list list \ (nat list*nat list) list \ rat mat \ rat list list" + where "reduction_signs_R p qs signs subsets matr = + (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets matr)))" + +definition reduction_subsets_R:: "real poly \ real poly list \ rat list list \ (nat list*nat list) list \ rat mat \ (nat list*nat list) list" + where "reduction_subsets_R p qs signs subsets matr = + (take_indices subsets (rows_to_keep (reduce_mat_cols matr (solve_for_lhs_R p qs subsets matr))))" + +(* Some basic lemmas *) +lemma reduction_signs_is_get_signs_R: "reduction_signs_R p qs signs subsets m = get_signs_R (reduce_system_R p (qs, (m, (subsets, signs))))" + unfolding reduction_signs_R_def get_signs_R_def apply (simp) + using reduction_step_R.elims snd_conv + by metis + +lemma reduction_subsets_is_get_subsets_R: "reduction_subsets_R p qs signs subsets m = get_subsets_R (reduce_system_R p (qs, (m, (subsets, signs))))" + unfolding reduction_subsets_R_def get_subsets_R_def + using reduce_system.simps reduction_step.elims fst_conv snd_conv + by (metis reduce_system_R.simps reduction_step_R.simps) + +subsection "Showing sign conditions preserved when reducing" + +lemma take_indices_lem_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes arb_list :: "('a 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 "\k carrier_mat (length subsets) (length signs)" + unfolding matrix_A_R_def carrier_mat_def by auto + +lemma size_of_lhs_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes signs :: "rat list list" + shows "dim_vec (construct_lhs_vector_R p qs signs) = length signs" + unfolding construct_lhs_vector_R_def + by simp + +lemma size_of_rhs_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + shows "dim_vec (construct_rhs_vector_R p qs subsets) = length subsets" + unfolding construct_rhs_vector_R_def + by simp + +lemma same_size_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "length subsets = length signs" + using invertible_mat unfolding invertible_mat_def + using size_of_mat_R[of signs subsets] size_of_lhs_R[of p qs signs] size_of_rhs_R[of p qs subsets] + by simp + +lemma construct_lhs_matches_solve_for_lhs_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" +proof - + have matrix_equation_hyp: "(mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs) = (construct_rhs_vector_R p qs subsets))" + using match unfolding satisfy_equation_R_def by blast + then have eqn_hyp: " ((matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs)) = + mult_mat_vec (matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) (construct_rhs_vector_R p qs subsets))" + using invertible_mat + by (simp add: matrix_equation_hyp) + have match_hyp: "length subsets = length signs" using same_size_R invertible_mat by auto + then have dim_hyp1: "matrix_A_R signs subsets \ carrier_mat (length signs) (length signs)" + using size_of_mat + by auto + then have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R 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_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs))) + = (((matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs))" + using mult_assoc dim_hyp1 dim_hyp2 size_of_lhs_R by auto + have cancel_helper: "(((matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs)) + = (1\<^sub>m (length signs)) *\<^sub>v (construct_lhs_vector_R p qs signs)" + using invertible_means_mult_id[where A= "matrix_A_R signs subsets"] dim_hyp1 + by (simp add: invertible_mat match_hyp) + then have cancel_hyp: "(((matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) * (matrix_A_R signs subsets)) *\<^sub>v (construct_lhs_vector_R p qs signs)) + = (construct_lhs_vector_R p qs signs)" + apply (auto) + by (metis carrier_vec_dim_vec one_mult_mat_vec size_of_lhs_R) + then have mult_hyp: "((matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A_R signs subsets) (construct_lhs_vector_R p qs signs))) + = (construct_lhs_vector_R p qs signs)" + using mult_assoc_hyp cancel_hyp + by simp + then have "(construct_lhs_vector_R p qs signs) = (mult_mat_vec (matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets))) (construct_rhs_vector_R p qs subsets)) " + using eqn_hyp + by simp + then show ?thesis + unfolding solve_for_lhs_R_def + by (simp add: mat_inverse_same) +qed + +(* Then show that dropping columns doesn't affect the consistent sign assignments *) +lemma reduction_doesnt_break_things_signs_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" +proof - + have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets)) \ carrier_mat (length signs) (length signs)" + using invertible_mat dim_invertible + using same_size_R by fastforce + have "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" + using construct_lhs_matches_solve_for_lhs_R assms by auto + then have "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) = + vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))" + using construct_lhs_vector_cleaner_R 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_R p qs subsets (matrix_A_R signs subsets))). + (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ + (nth signs n) \ set (characterize_consistent_signs_at_roots p qs))" + proof - + have h0: "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). + (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ + rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = (nth signs n)}) = 0)" + by (metis (mono_tags, lifting) \construct_lhs_vector_R p qs signs = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)\ construct_lhs_vector_clean_R nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs_R) + have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) = 0 \ + (\ x. poly p x = 0 \ consistent_sign_vec qs x = w))" + proof clarsimp + fix x + assume card_asm: "card {xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec qs x} = 0" + assume zero_asm: "poly p x = 0" + have card_hyp: "{xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec qs x} = {}" + using card_eq_0_iff + using card_asm nonzero poly_roots_finite + by (metis (full_types) finite_Collect_conjI) + have "x \ {xa. poly p xa = 0 \ consistent_sign_vec qs xa = consistent_sign_vec 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 qs x = w}) = 0 \ + (\List.member (characterize_consistent_signs_at_roots p qs) w))" + proof clarsimp + fix w + assume card_asm: "card {x. poly p x = 0 \ consistent_sign_vec qs x = w} = 0" + assume mem_asm: "List.member (characterize_consistent_signs_at_roots p qs) w" + have h0: "\ x. poly p x = 0 \ consistent_sign_vec qs x = w" using h1 card_asm + by (simp add: h1) + have h1: "\ x. poly p x = 0 \ consistent_sign_vec qs x = w" using mem_asm + unfolding characterize_consistent_signs_at_roots_def characterize_root_list_p_def + proof - + have "w \ Collect (List.member (remdups (map (consistent_sign_vec qs) (sorted_list_of_set {r. poly p r = 0}))))" + using characterize_consistent_signs_at_roots_def mem_asm characterize_root_list_p_def + by (smt (verit, ccfv_SIG) consistent_sign_vec_def h0 imageE in_set_member list.set_map map_cong mem_Collect_eq nonzero o_apply poly_roots_finite set_remdups set_sorted_list_of_set signs_at_def squash_def) + then have f1: "w \ set (map (consistent_sign_vec qs) (sorted_list_of_set {r. poly p r = 0}))" + by (metis (no_types) in_set_member mem_Collect_eq set_remdups) + have "finite {r. poly p r = 0}" + using nonzero poly_roots_finite by blast + then show ?thesis + using f1 by auto + qed + show "False" using h0 h1 by auto + qed + then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) = 0 \ + w \ set (characterize_consistent_signs_at_roots p qs)" + by (simp add: in_set_member) + show ?thesis using h0 h3 + by blast + qed + then have "set (characterize_consistent_signs_at_roots p qs) \ set (take_indices signs + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" + using all_info + reduction_signs_set_helper_lemma[where A = "set (characterize_consistent_signs_at_roots p qs)", where B = "signs", + where C = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))"] + using dim_hyp2 solve_for_lhs_R_def by (simp add: mat_inverse_same) + then show ?thesis unfolding reduction_signs_R_def by auto +qed + +lemma reduction_deletes_bad_sign_conds_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "set (characterize_consistent_signs_at_roots p qs) = set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" +proof - + have dim_hyp2: "matr_option (dim_row (matrix_A_R signs subsets)) + (mat_inverse (matrix_A_R signs subsets)) \ carrier_mat (length signs) (length signs)" + using invertible_mat dim_invertible + using same_size_R by fastforce + have supset: "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" + proof - + have "(construct_lhs_vector_R p qs signs) = solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" + using construct_lhs_matches_solve_for_lhs_R assms by auto + then have "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) = + vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))" + using construct_lhs_vector_cleaner_R 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_R p qs subsets (matrix_A_R signs subsets))). + (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n \ 0) \ + (nth signs n) \ set (characterize_consistent_signs_at_roots p qs))" + proof - + have h0: "\ n < (dim_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))). + (((solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) $ n = 0) \ + rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = (nth signs n)}) = 0)" + by (simp add: \solve_for_lhs_R p qs subsets (M_mat_R signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec qs x = s}) signs))\ vec_of_list_index) + have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) \ 0 \ + (\ x. poly p x = 0 \ consistent_sign_vec qs x = w))" + proof clarsimp + fix w + assume card_asm: "0 < card {x. poly p x = 0 \ consistent_sign_vec qs x = w}" + show "\x. poly p x = 0 \ consistent_sign_vec 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 qs x = w}) \ 0 \ + (List.member (characterize_consistent_signs_at_roots p qs) w))" + proof clarsimp + fix w + assume card_asm: " 0 < card {x. poly p x = 0 \ consistent_sign_vec qs x = w}" + have h0: "\x. poly p x = 0 \ consistent_sign_vec qs x = w" + using card_asm + by (simp add: h1) + then show "List.member (characterize_consistent_signs_at_roots p qs) w" + unfolding characterize_consistent_signs_at_roots_def + using in_set_member nonzero poly_roots_finite characterize_root_list_p_def + by (smt (verit) characterize_consistent_signs_at_roots_def in_set_R mem_Collect_eq) + qed + then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec qs x = w}) \ 0 \ + w \ set (characterize_consistent_signs_at_roots p qs)" + by (simp add: in_set_member) + show ?thesis using h0 h3 + by (metis (no_types, lifting) \solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec 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_R p qs subsets (matrix_A_R signs subsets)))) \ + set (characterize_consistent_signs_at_roots p qs)" + using all_info + reduction_signs_set_helper_lemma2[where A = "set (characterize_consistent_signs_at_roots p qs)", where B = "signs", + where C = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))"] + using distinct_signs dim_hyp2 solve_for_lhs_R_def + by (simp add: mat_inverse_same) + then show ?thesis unfolding reduction_signs_R_def by auto + qed + have subset: "set (characterize_consistent_signs_at_roots p qs) \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" + using reduction_doesnt_break_things_signs_R[of p qs signs subsets] assms + by blast + then show ?thesis using supset subset by auto +qed + + +theorem reduce_system_sign_conditions_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "set (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = set (characterize_consistent_signs_at_roots p qs)" + unfolding get_signs_R_def + using reduction_deletes_bad_sign_conds_R[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 snd_conv + apply (metis reduction_signs_R_def) + by (metis all_info distinct_signs invertible_mat match nonzero reduction_signs_R_def snd_conv welldefined_signs1) + +subsection "Showing matrix equation preserved when reducing" + +lemma reduce_system_matrix_equation_preserved_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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_R (subsets) (length qs)" + assumes distinct_signs: "distinct signs" + assumes all_info: "set (characterize_consistent_signs_at_roots p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes invertible_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "satisfy_equation_R p qs (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" +proof - + have poly_type_hyp: "p \ 0" using nonzero by auto + have distinct_signs_hyp: "distinct (snd (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))))" + proof - + let ?sym = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets)))" + assume "j < length + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + assume neq_hyp: "i \ j" + assume "signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets + (matrix_A_R signs subsets)) ! i) = + signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets + (matrix_A_R signs subsets)) ! j)" + have h1: "find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets + (matrix_A_R signs subsets)) ! i \ find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets + (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets)))\ \j < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets) $ n \ 0) [0.. set [0..i < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ construct_lhs_matches_solve_for_lhs_R find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs_R) + have "filter (\n. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ n \ 0) [0.. set [0..j < length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ construct_lhs_matches_solve_for_lhs_R find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs_R) + then show ?thesis + using f2 by (metis \signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets))))" + using distinct_conv_nth by blast + then show ?thesis + using get_signs_R_def reduction_signs_R_def reduction_signs_is_get_signs_R by auto + qed + have all_info_hyp: "set (characterize_consistent_signs_at_roots p qs) \ set(snd (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))))" + using reduce_system_sign_conditions_R assms unfolding get_signs_R_def by auto + have welldefined_hyp: "all_list_constr_R (fst (snd (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))) (length qs)" + using welldefined_subsets rows_to_keep_lem + unfolding all_list_constr_R_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 + apply (smt (z3) mat_of_cols_carrier(2) rows_to_keep_lem) + by (smt (z3) mat_of_cols_carrier(2) nth_mem rows_to_keep_lem) + then show ?thesis using poly_type_hyp distinct_signs_hyp all_info_hyp welldefined_hyp + using matrix_equation_R unfolding get_subsets_R_def get_signs_R_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_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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_R (take_indices signs S) subsets = take_cols_from_matrix (alt_matrix_A_R signs subsets) S" +proof - + have h0a: "dim_col (take_cols_from_matrix (alt_matrix_A_R 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_R (take_indices signs S) subsets ) i = +col (take_cols_from_matrix (alt_matrix_A_R 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_R signs subsets))) S)" using asm + apply (auto) unfolding take_indices_def by auto + have h0: " vec (length subsets) (\j. z_R (subsets ! j) (map ((!) signs) S ! i)) = + vec (length subsets) (\j. z_R (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_R signs subsets))) S) ! i \ carrier_vec (dim_row (alt_matrix_A_R signs subsets))" + proof - + have "dim_col (alt_matrix_A_R signs subsets) = length (signs)" + by (simp add: alt_matrix_A_R_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_R signs subsets))) S) ! i = nth (cols (alt_matrix_A_R signs subsets)) (S ! i)" + using i_lt by auto + have "nth (cols (alt_matrix_A_R signs subsets)) (S ! i) \ carrier_vec (dim_row (alt_matrix_A_R signs subsets))" + using col_dim unfolding cols_def using nth_map well_d + by (simp add: \dim_col (alt_matrix_A_R signs subsets) = length signs\) + then show ?thesis using map_eq by auto + qed + have h1: "col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i = + col (mat_of_cols (dim_row (alt_matrix_A_R signs subsets)) (map ((!) (cols (alt_matrix_A_R 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_R signs subsets)) (map ((!) (cols (alt_matrix_A_R signs subsets))) S)) i + = nth (map ((!) (cols (alt_matrix_A_R signs subsets))) S) i " + using dim i_lt asm col_mat_of_cols[where j = "i", where n = "(dim_row (alt_matrix_A_R signs subsets))", + where vs = "(map ((!) (cols (alt_matrix_A_R signs subsets))) S)"] + by blast + have h3: "col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i = (col (alt_matrix_A_R signs subsets) (S !i))" + using h1 h2 apply (auto) + by (metis alt_matrix_char_R asm cols_nth dim_col_mat(1) in_set_member length_map mat_of_rows_list_def matrix_A_R_def nth_map nth_mem take_indices_def well_def_h) + have "vec (length subsets) (\j. z_R (subsets ! j) (signs ! (S ! i))) = (col (alt_matrix_A_R signs subsets) (S !i))" + by (metis asm in_set_member length_map nth_mem signs_are_cols_R take_indices_def well_def_h) + then have "vec (length subsets) (\j. z_R (subsets ! j) (take_indices signs S ! i)) = + col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i " + using h0 h3 + by (simp add: take_indices_def) + then show "col (alt_matrix_A_R (take_indices signs S) subsets) i = + col (take_cols_from_matrix (alt_matrix_A_R signs subsets) S) i " + using asm signs_are_cols_R[where signs = "(take_indices signs S)", where subsets = "subsets"] + by auto + qed + then show ?thesis unfolding alt_matrix_A_R_def take_cols_from_matrix_def apply (auto) + using h0a mat_col_eqI + by (metis alt_matrix_A_R_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_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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_R (take_indices signs S) subsets = take_cols_from_matrix (matrix_A_R signs subsets) S" + using reduce_system_matrix_signs_helper_aux_R alt_matrix_char_R assms by auto + +lemma reduce_system_matrix_subsets_helper_aux_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list* 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_R signs (take_indices subsets S) = take_rows_from_matrix (alt_matrix_A_R signs subsets) S" +proof - + have h0a: "dim_row (take_rows_from_matrix (alt_matrix_A_R 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_R signs (take_indices subsets S) ) i = +row (take_rows_from_matrix (alt_matrix_A_R 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_R signs subsets))) S)" using asm + apply (auto) unfolding take_indices_def by auto + have h0: "row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) i = + row (mat_of_rows (dim_col (alt_matrix_A_R signs subsets)) (map ((!) (rows (alt_matrix_A_R signs subsets))) S)) i " + unfolding take_rows_from_matrix_def take_indices_def by auto + have dim: "(map ((!) (rows (alt_matrix_A_R signs subsets))) S) ! i \ carrier_vec (dim_col (alt_matrix_A_R signs subsets))" + proof - + have "dim_col (alt_matrix_A_R signs subsets) = length (signs)" + by (simp add: alt_matrix_A_R_def) + then have lenh: "dim_col (alt_matrix_A_R 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_R signs subsets))) S) ! i = nth (rows (alt_matrix_A_R signs subsets)) (S ! i)" + using i_lt by auto + have "nth (rows (alt_matrix_A_R signs subsets)) (S ! i) \ carrier_vec (dim_col (alt_matrix_A_R signs subsets))" + using col_dim unfolding rows_def using nth_map well_d + using lenh + by (simp add: alt_matrix_A_R_def) + then show ?thesis using map_eq unfolding alt_matrix_A_R_def by auto + qed + have h1: " row (mat_of_rows (dim_col (alt_matrix_A_R signs subsets)) (map ((!) (rows (alt_matrix_A_R signs subsets))) S)) i + = (row (alt_matrix_A_R signs subsets) (S ! i)) " + using dim i_lt mat_of_rows_row[where i = "i", where n = "(dim_col (alt_matrix_A_R signs subsets))", + where vs = "(map ((!) (rows (alt_matrix_A_R signs subsets))) S)"] + using alt_matrix_char_R in_set_member nth_mem well_def_h + by fastforce + have "row (alt_matrix_A_R signs (take_indices subsets S) ) i = (row (alt_matrix_A_R signs subsets) (S ! i))" + using subsets_are_rows_R + 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_R (subsets ! i) (signs ! j))\ take_indices_def well_def_h) + qed + then show "(row (alt_matrix_A_R signs (take_indices subsets S) ) i = + row (take_rows_from_matrix (alt_matrix_A_R signs subsets) S) i)" + using h0 h1 unfolding take_indices_def by auto + qed + then show ?thesis unfolding alt_matrix_A_R_def take_rows_from_matrix_def apply (auto) + using eq_rowI + by (metis alt_matrix_A_R_def dim_col_mat(1) dim_row_mat(1) h0 h0a mat_of_rows_def take_rows_from_matrix_def) +qed + + +lemma reduce_system_matrix_subsets_helper_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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_R signs (take_indices subsets S) = take_rows_from_matrix (matrix_A_R signs subsets) S" + using assms reduce_system_matrix_subsets_helper_aux_R alt_matrix_char_R + by auto + +lemma reduce_system_matrix_match_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + assumes inv: "invertible_mat (matrix_A_R signs subsets)" + shows "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = + (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" +proof - + let ?A = "(matrix_A_R signs subsets)" + let ?lhs_vec = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))" + let ?red_mtx = "(take_rows_from_matrix (reduce_mat_cols (matrix_A_R signs subsets) ?lhs_vec) (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) ?lhs_vec)))" + have h1: "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + = matrix_A_R (reduction_signs_R p qs signs subsets (matrix_A_R signs subsets)) (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets))" + using reduction_signs_is_get_signs_R reduction_subsets_is_get_subsets_R by auto + have h1_var: "matrix_A_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + = matrix_A_R (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_R_def reduction_subsets_R_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_R p qs signs) = ?lhs_vec" + using assms construct_lhs_matches_solve_for_lhs_R + by simp + have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" + using h30 size_of_lhs_R[of p qs signs] + unfolding find_nonzeros_from_input_vec_def apply (auto) + using in_set_member by force + have h3b_a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (subsets)" + using assms h30 size_of_lhs_R same_size_R 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_R p qs subsets (matrix_A_R signs subsets) $ i \ 0) + [0.. length subsets" using length_filter_le[where P = "(\i. solve_for_lhs_R p qs subsets (matrix_A_R signs subsets) $ i \ 0)", + where xs = "[0..i. solve_for_lhs_R p qs subsets (matrix_A_R 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_R signs subsets) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) < length signs \ matrix_A_R (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) subsets = take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + using nonzero + using h3a reduce_system_matrix_signs_helper_R by auto + then have "matrix_A_R (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) subsets = take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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 dim_row_matrix_A_R rows_to_keep_def rows_to_keep_lem) + qed + have h3: "matrix_A_R (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_R reduce_system_matrix_signs_helper_R + assms + by auto + show ?thesis using h1 h2 h3 + by (metis fst_conv get_matrix_R_def h1_var reduce_system_R.simps reduction_step_R.simps) +qed + +subsection "Showing invertibility preserved when reducing" + +(* gauss_jordan_single_rank is critically helpful in this subsection *) +thm conjugatable_vec_space.gauss_jordan_single_rank + +thm vec_space.full_rank_lin_indpt + +(* 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_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes len_eq: "length subsets = length signs" + assumes inv: "invertible_mat (matrix_A_R 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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + shows "(\j. j \ set (find_nonzeros_from_input_vec + (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ + j < length (cols (matrix_A_R signs subsets)))" +proof - + fix i + fix j + assume j_in: " j \ set (find_nonzeros_from_input_vec + (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) " + let ?og_mat = "(matrix_A_R signs subsets)" + let ?lhs = "(solve_for_lhs_R 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_R 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_R p qs subsets (matrix_A_R signs subsets)) = (length signs)" + using size_of_lhs_R construct_lhs_matches_solve_for_lhs_R 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_R signs subsets))" + using mat_size by auto +qed + + +lemma take_cols_subsets_og_cols_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes len_eq: "length subsets = length signs" + assumes inv: "invertible_mat (matrix_A_R 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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + shows "set (take_indices (cols (matrix_A_R signs subsets)) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) + \ set (cols (matrix_A_R signs subsets))" +proof - + have well_def: "(\j. j \ set (find_nonzeros_from_input_vec + (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) \ + j < length (cols (matrix_A_R signs subsets)))" + using assms well_def_find_zeros_from_lhs_vec_R by auto + have "\x. x \ set (take_indices (cols (matrix_A_R signs subsets)) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) + \ x \ set (cols (matrix_A_R signs subsets))" + proof clarsimp + fix x + let ?og_list = "(cols (matrix_A_R signs subsets))" + let ?ind_list = "(find_nonzeros_from_input_vec + (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + assume x_in: "x \ set (take_indices ?og_list ?ind_list)" + show "x \ set (cols (matrix_A_R 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_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes len_eq: "length subsets = length signs" + assumes inv: "invertible_mat (matrix_A_R 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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + shows "vec_space.rank (length signs) (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))) = + (length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" +proof - + let ?og_mat = "(matrix_A_R signs subsets)" + let ?lhs = "(solve_for_lhs_R 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_R 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_R 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_R p qs subsets (matrix_A_R signs subsets))) \ + j < length (cols (matrix_A_R signs subsets)))" + using assms well_def_find_zeros_from_lhs_vec_R by auto + have dist1: "distinct + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + unfolding find_nonzeros_from_input_vec_def by auto + have clear: "set (take_indices (cols (matrix_A_R signs subsets)) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) + \ set (cols (matrix_A_R signs subsets))" + using assms take_cols_subsets_og_cols_R by auto + then have "distinct (take_indices (cols (matrix_A_R signs subsets)) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))" + unfolding take_indices_def + using dist1 dist_cols well_def conjugatable_vec_space.distinct_map_nth[where ls = "cols (matrix_A_R signs subsets)", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R signs subsets", where B = "(take_indices (cols (matrix_A_R signs subsets)) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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 reduction_doesnt_break_things_invertibility_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes len_eq: "length subsets = length signs" + assumes inv: "invertible_mat (matrix_A_R 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 p qs) \ set(signs)" + assumes match: "satisfy_equation_R p qs subsets signs" + shows "invertible_mat (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" +proof - + let ?og_mat = "(matrix_A_R signs subsets)" + let ?lhs = "(solve_for_lhs_R 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_R 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_R diff_zero find_nonzeros_from_input_vec_def inv length_filter_le length_upt match size_of_lhs_R) + 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_R 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_R 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) + by (smt (verit, ccfv_threshold) carrier_matD(2) index_transpose_mat(2) mat_of_cols_carrier(2) reduce_mat_cols.simps rows_to_keep_def step1_mat_size take_cols_from_matrix_def transpose_rank) + 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_R (get_signs_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) + (get_subsets_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs))))) = + (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" + using reduce_system_matrix_match_R[of p qs signs subsets] assms + by blast + have f2: "length (get_subsets_R (take_rows_from_matrix (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) (rows_to_keep (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))), map ((!) subsets) (rows_to_keep (mat_of_cols (dim_row (matrix_A_R signs subsets)) (map ((!) (cols (matrix_A_R signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))), map ((!) signs) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + + by (metis (no_types) \dim_col (mat_of_cols (dim_row (matrix_A_R signs subsets)) (take_indices (cols (matrix_A_R signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ \length (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) = dim_col (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ length_map reduce_mat_cols.simps reduce_system_R.simps reduction_step_R.simps reduction_subsets_R_def reduction_subsets_is_get_subsets_R 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_R p ps nss m)) = get_signs_R (reduce_system_R p (ps, m, nss, rss))" + by (metis (no_types) reduction_signs_R_def reduction_signs_is_get_signs_R take_indices_def) + have square_final_mat: "square_mat (get_matrix_R (reduce_system_R p (qs, ((matrix_A_R signs subsets), (subsets, signs)))))" + using mat_match assms size_of_mat_R same_size_R + apply (auto) using f2 f3 + by (metis (no_types, lifting) Matrix.transpose_transpose fst_conv get_matrix_R_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_R dim_row_matrix_A_R fst_conv get_matrix_R_def mat_match rank_match reduce_system_R.simps reduction_step_R.simps size_of_mat_R square_mat.elims(2)) + then show ?thesis apply (simp) + by (metis fst_conv get_matrix_R_def) +qed + +subsection "Well def signs preserved when reducing" + +lemma reduction_doesnt_break_length_signs_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes length_init : "\ x\ set(signs). length x = length qs" + assumes sat_eq: "satisfy_equation_R p qs subsets signs" + assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "\x \ set(reduction_signs_R p qs signs subsets (matrix_A_R signs subsets)). + length x = length qs" + unfolding reduction_signs_def using take_indices_lem + by (smt (verit) atLeastLessThan_iff construct_lhs_matches_solve_for_lhs_R filter_is_subset find_nonzeros_from_input_vec_def in_set_conv_nth in_set_member inv_mat length_init reduction_signs_R_def sat_eq set_upt size_of_lhs_R subset_code(1)) + +subsection "Distinct signs preserved when reducing" + +lemma reduction_signs_are_distinct_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes sat_eq: "satisfy_equation_R p qs subsets signs" + assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" + assumes distinct_init: "distinct signs" + shows "distinct (reduction_signs_R p qs signs subsets (matrix_A_R signs subsets))" +proof - + have solve_construct: "construct_lhs_vector_R p qs signs = + solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" + using construct_lhs_matches_solve_for_lhs_R assms + by simp + have h1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R 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_R p qs subsets (matrix_A_R signs subsets))) \ + j < length signs)" + proof - + fix j + assume "j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))" + show "j < length signs" using solve_construct size_of_lhs_R + by (metis \j \ set (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))\ atLeastLessThan_iff filter_is_subset find_nonzeros_from_input_vec_def set_upt subset_iff) + qed + then show ?thesis unfolding reduction_signs_R_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_R p qs subsets (matrix_A_R signs subsets)))"] + by blast +qed + +subsection "Well def subsets preserved when reducing" + +lemma reduction_doesnt_break_subsets_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list* nat list) list" + fixes signs :: "rat list list" + assumes nonzero: "p \ 0" + assumes length_init : "all_list_constr_R subsets (length qs)" + assumes sat_eq: "satisfy_equation_R p qs subsets signs" + assumes inv_mat: "invertible_mat (matrix_A_R signs subsets)" + shows "all_list_constr_R (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets)) (length qs)" + unfolding all_list_constr_R_def +proof clarsimp + fix a b + assume in_red_subsets: "List.member (reduction_subsets_R p qs signs subsets (matrix_A_R signs subsets)) (a, b) " + have solve_construct: "construct_lhs_vector_R p qs signs = + solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)" + using construct_lhs_matches_solve_for_lhs_R assms + by simp + have rows_to_keep_hyp: "\y. y \ set (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) \ + y < length subsets" + proof clarsimp + fix y + assume in_set: "y \ set (rows_to_keep + (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))))" + have in_set_2: "y \ set (rows_to_keep + (take_cols_from_matrix (matrix_A_R signs subsets) (find_nonzeros_from_input_vec (construct_lhs_vector_R p qs signs))))" + using in_set solve_construct by simp + let ?lhs_vec = "(solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))" + have h30: "(construct_lhs_vector_R p qs signs) = ?lhs_vec" + using assms construct_lhs_matches_solve_for_lhs_R + by simp + have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" + using h30 size_of_lhs_R unfolding find_nonzeros_from_input_vec_def + using atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq apply (auto) + by (smt (verit, best) atLeastLessThan_iff in_set_member mem_Collect_eq set_filter set_upt) + have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols (matrix_A_R signs subsets) (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets)))) x \ x < length (subsets)" + proof clarsimp + fix x + assume x_mem: "List.member (rows_to_keep + (take_cols_from_matrix (matrix_A_R signs subsets) + (find_nonzeros_from_input_vec (solve_for_lhs_R p qs subsets (matrix_A_R signs subsets))))) x" + show "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_R rows_to_keep_def rows_to_keep_lem + apply (auto) + by (smt (verit, best) List.member_def dim_row_matrix_A_R 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 a (length qs) \ list_constr b (length qs)" using in_red_subsets unfolding reduction_subsets_def + using take_indices_lem_R[of _ subsets] rows_to_keep_hyp + using all_list_constr_R_def in_set_conv_nth in_set_member length_init + by (metis fst_conv reduction_subsets_R_def snd_conv) +qed + +section "Overall Lemmas" +lemma combining_to_smash_R: "combine_systems_R p (qs1, m1, (sub1, sgn1)) (qs2, m2, (sub2, sgn2)) + = smash_systems_R p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2" + by simp + +lemma getter_functions_R: "calculate_data_R p qs = (get_matrix_R (calculate_data_R p qs), (get_subsets_R (calculate_data_R p qs), get_signs_R (calculate_data_R p qs))) " + unfolding get_matrix_R_def get_subsets_R_def get_signs_R_def by auto + +subsection "Key properties preserved" + +subsubsection "Properties preserved when combining and reducing systems" +lemma combining_sys_satisfies_properties_helper_R: + fixes p:: "real poly" + fixes qs1 :: "real poly list" + fixes qs2 :: "real poly list" + fixes subsets1 subsets2 :: "(nat list * nat list) list" + fixes signs1 signs2 :: "rat list list" + fixes matrix1 matrix2:: "rat mat" + assumes nonzero: "p \ 0" + assumes nontriv1: "length qs1 > 0" + assumes nontriv2: "length qs2 > 0" + assumes satisfies_properties_sys1: "satisfies_properties_R p qs1 subsets1 signs1 matrix1" + assumes satisfies_properties_sys2: "satisfies_properties_R p qs2 subsets2 signs2 matrix2" + shows "satisfies_properties_R p (qs1@qs2) (get_subsets_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) + (get_signs_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) + (get_matrix_R (snd ((combine_systems_R p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2)))))))" +proof - + let ?subsets = "(get_subsets_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) + (qs2, matrix2, subsets2, signs2))))" + let ?signs = "(get_signs_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" + let ?matrix = "(get_matrix_R (snd (combine_systems_R p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" + have h1: "all_list_constr_R ?subsets (length (qs1 @ qs2))" + using well_def_step_R[of subsets1 qs1 subsets2 qs2] assms + by (simp add: nontriv2 get_subsets_R_def satisfies_properties_R_def smash_systems_R_def) + have h2: "well_def_signs (length (qs1 @ qs2)) ?signs" + using well_def_signs_step[of qs1 qs2 signs1 signs2] + using get_signs_R_def nontriv1 nontriv2 satisfies_properties_R_def satisfies_properties_sys1 satisfies_properties_sys2 smash_systems_R_def by auto + have h3: "distinct ?signs" + using distinct_step[of _ signs1 _ signs2] assms + using combine_systems.simps get_signs_R_def satisfies_properties_R_def smash_systems_R_def snd_conv by auto + have h4: "satisfy_equation_R p (qs1 @ qs2) ?subsets ?signs" + using assms inductive_step_R[of p qs1 qs2 signs1 signs2 subsets1 subsets2] + using get_signs_R_def get_subsets_R_def satisfies_properties_R_def smash_systems_R_def + by auto + have h5: " invertible_mat ?matrix" + using assms inductive_step_R[of p qs1 qs2 signs1 signs2 subsets1 subsets2] + by (metis combining_to_smash_R fst_conv get_matrix_R_def kronecker_invertible satisfies_properties_R_def smash_systems_R_def snd_conv) + have h6: "?matrix = matrix_A_R ?signs ?subsets" + unfolding get_matrix_R_def combine_systems_R.simps smash_systems_R_def get_signs_R_def get_subsets_R_def + apply simp + apply (subst matrix_construction_is_kronecker_product_R[of subsets1 _ signs1 signs2 subsets2]) + apply (metis Ball_set all_list_constr_R_def in_set_member list_constr_def satisfies_properties_R_def satisfies_properties_sys1) + using satisfies_properties_R_def satisfies_properties_sys1 well_def_signs_def apply blast + using satisfies_properties_R_def satisfies_properties_sys1 satisfies_properties_sys2 by auto + have h7: "set (characterize_consistent_signs_at_roots p (qs1 @ qs2)) + \ set (?signs)" + using subset_step_R[of p qs1 signs1 qs2 signs2] assms + by (simp add: nonzero get_signs_R_def satisfies_properties_R_def smash_systems_R_def) + then show ?thesis unfolding satisfies_properties_R_def using h1 h2 h3 h4 h5 h6 h7 by blast +qed + +lemma combining_sys_satisfies_properties_R: + fixes p:: "real poly" + fixes qs1 :: "real poly list" + fixes qs2 :: "real poly list" + assumes nonzero: "p \ 0" + assumes nontriv1: "length qs1 > 0" + assumes nontriv2: "length qs2 > 0" + assumes satisfies_properties_sys1: "satisfies_properties_R p qs1 (get_subsets_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs1))" + assumes satisfies_properties_sys2: "satisfies_properties_R p qs2 (get_subsets_R (calculate_data_R p qs2)) (get_signs_R (calculate_data_R p qs2)) (get_matrix_R (calculate_data_R p qs2))" + shows "satisfies_properties_R p (qs1@qs2) (get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) + (get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) + (get_matrix_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" + using combining_sys_satisfies_properties_helper_R[of p qs1 qs2] + by (metis getter_functions_R nontriv1 nontriv2 nonzero satisfies_properties_sys1 satisfies_properties_sys2) + +lemma reducing_sys_satisfies_properties_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + fixes matrix:: "rat mat" + assumes nonzero: "p \ 0" + assumes nontriv: "length qs > 0" + assumes satisfies_properties_sys: "satisfies_properties_R p qs subsets signs matrix" + shows "satisfies_properties_R p qs (get_subsets_R (reduce_system_R p (qs,matrix,subsets,signs))) + (get_signs_R (reduce_system_R p (qs,matrix,subsets,signs))) + (get_matrix_R (reduce_system_R p (qs,matrix,subsets,signs)))" +proof - + have h1: " all_list_constr_R (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs))) (length qs)" + using reduction_doesnt_break_subsets_R assms reduction_subsets_is_get_subsets_R satisfies_properties_R_def satisfies_properties_sys by auto + have h2: "well_def_signs (length qs) (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduction_doesnt_break_length_signs_R[of signs qs p subsets] assms reduction_signs_is_get_signs_R satisfies_properties_R_def well_def_signs_def by auto + have h3: "distinct (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduction_signs_are_distinct_R[of p qs subsets signs] assms reduction_signs_is_get_signs_R satisfies_properties_R_def by auto + have h4: "satisfy_equation_R p qs (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs))) + (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduce_system_matrix_equation_preserved_R[of p qs signs subsets] assms satisfies_properties_R_def by auto + have h5: "invertible_mat (get_matrix_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduction_doesnt_break_things_invertibility_R assms same_size_R satisfies_properties_R_def by auto + have h6: "get_matrix_R (reduce_system_R p (qs, matrix, subsets, signs)) = + matrix_A_R (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs))) + (get_subsets_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduce_system_matrix_match_R[of p qs signs subsets] assms satisfies_properties_R_def by auto + have h7: "set (characterize_consistent_signs_at_roots p qs) \ set (get_signs_R (reduce_system_R p (qs, matrix, subsets, signs)))" + using reduction_doesnt_break_things_signs_R[of p qs signs subsets] assms reduction_signs_is_get_signs_R satisfies_properties_R_def by auto + then show ?thesis unfolding satisfies_properties_R_def using h1 h2 h3 h4 h5 h6 h7 + by blast +qed + +subsubsection "For length 1 qs" + +lemma length_1_calculate_data_satisfies_properties_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes nonzero: "p \ 0" + assumes len1: "length qs = 1" + shows "satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" +proof - + have h1: "all_list_constr_R [([], []),([0], []),([], [0])] (length qs)" + using len1 unfolding all_list_constr_R_def list_constr_def apply (auto) + apply (smt (verit, best) Ball_set in_set_member member_rec(1) member_rec(2) prod.inject) + by (smt (verit, ccfv_threshold) Ball_set in_set_member member_rec(1) member_rec(2) prod.inject) + 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],[0],[-1]]::rat list list)" + unfolding distinct_def using in_set_member by auto + have h4: "satisfy_equation_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-1]]" + using assms base_case_satisfy_equation_alt_R[of qs p] by auto + have h6: "(mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat) = (matrix_A_R ([[1],[0],[-1]]) ([([], []),([0], []),([], [0])]) :: rat mat)" + using mat_base_case_R by auto + then have h5: "invertible_mat (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" + using base_case_invertible_mat_R + by simp + have h7: "set (characterize_consistent_signs_at_roots p qs) \ set ([[1],[0],[-1]])" + using assms base_case_sgas_alt_R[of qs p] + by simp + have base_case_hyp: "satisfies_properties_R p qs [([], []),([0], []),([], [0])] [[1],[0],[-1]] (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]]::rat mat)" + using h1 h2 h3 h4 h5 h6 h7 + using satisfies_properties_R_def apply (auto) + by (simp add: well_def_signs_def) + then have key_hyp: "satisfies_properties_R p qs (get_subsets_R (reduce_system_R p (qs,base_case_info_R))) (get_signs_R (reduce_system_R p (qs,base_case_info_R))) (get_matrix_R (reduce_system_R p (qs,base_case_info_R)))" + using reducing_sys_satisfies_properties_R + by (metis base_case_info_R_def len1 nonzero nonzero zero_less_one_class.zero_less_one) + show ?thesis + by (simp add: key_hyp len1) +qed + +subsubsection "For arbitrary qs" + +lemma calculate_data_satisfies_properties_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + shows "(p \ 0 \ (length qs > 0)) + \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" +proof (induction "length qs" arbitrary: qs rule: less_induct) + case less + have len1_h: "length qs = 1 \ ( p \ 0 \ (length qs > 0)) \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" + using length_1_calculate_data_satisfies_properties_R + by blast + let ?len = "length qs" + let ?q1 = "take (?len div 2) qs" + let ?left = "calculate_data_R p ?q1" + let ?q2 = "drop (?len div 2) qs" + let ?right = "calculate_data_R p ?q2" + let ?comb = "combine_systems_R p (?q1,?left) (?q2,?right)" + let ?red = "reduce_system_R 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 q1_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ satisfies_properties_R p ?q1 (get_subsets_R (calculate_data_R p ?q1)) (get_signs_R (calculate_data_R p ?q1)) (get_matrix_R (calculate_data_R p ?q1))" + using less.hyps[of ?q1] h_q1_len + by (metis div_le_dividend div_less_dividend length_take min.absorb2 one_less_numeral_iff semiring_norm(76)) + 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)) \ satisfies_properties_R p ?q2 (get_subsets_R (calculate_data_R p ?q2)) (get_signs_R (calculate_data_R p ?q2)) (get_matrix_R (calculate_data_R p ?q2))" + using less.hyps[of ?q2] h_q2_len + 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)) \ (satisfies_properties_R p (qs) (get_subsets_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))) + (get_signs_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))) + (get_matrix_R (snd ((combine_systems_R p (?q1,calculate_data_R p ?q1) (?q2,calculate_data_R p ?q2))))))" + using q1_sat_props q2_sat_props combining_sys_satisfies_properties_R + using h_q1_len h_q2_len put_tog + by metis + then have comb_sat: "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ + (satisfies_properties_R p (qs) (get_subsets_R (snd ?comb)) (get_signs_R (snd ?comb)) (get_matrix_R (snd ?comb)))" + by blast + have red_char: "?red = (reduce_system_R p (qs,(get_matrix_R (snd ?comb)),(get_subsets_R (snd ?comb)),(get_signs_R (snd ?comb))))" + using getter_functions + by (smt (z3) combine_systems_R.simps find_consistent_signs_at_roots_R_def find_consistent_signs_at_roots_thm_R fst_conv get_matrix_R_def get_signs_R_def get_subsets_R_def prod.collapse put_tog smash_systems_R_def) + then have "length qs > 1 \ (p \ 0 \ (length qs > 0)) \ (satisfies_properties_R p qs (get_subsets_R ?red) (get_signs_R ?red) (get_matrix_R ?red))" + using reducing_sys_satisfies_properties_R comb_sat + by presburger + then have len_gt1: "length qs > 1 \ (p \ 0 \ (length qs > 0) ) \ satisfies_properties_R p qs (get_subsets_R (calculate_data_R p qs)) (get_signs_R (calculate_data_R p qs)) (get_matrix_R (calculate_data_R p qs))" + apply (auto) + by (smt (z3) div_le_dividend min.absorb2) + 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_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes nonzero: "p \ 0" + assumes len1: "length qs = 1" + shows "set (find_consistent_signs_at_roots_R p qs) = set (characterize_consistent_signs_at_roots p qs)" +proof - + let ?signs = "[[1],[0],[-1]]::rat list list" + let ?subsets = "[([], []),([0], []),([], [0])]::(nat list*nat list) list" + let ?mat = "(mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" + have mat_help: "matrix_A_R ?signs ?subsets = (mat_of_rows_list 3 [[1,1,1], [0,1,0], [1,0,-1]])" + using mat_base_case_R 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 p qs) \ set(?signs)" + using assms base_case_sgas_R by auto + have match: "satisfy_equation_R p qs ?subsets ?signs" + using ex_q base_case_satisfy_equation_R nonzero + by auto + have invertible_mat: "invertible_mat (matrix_A_R ?signs ?subsets)" + using inverse_mat_base_case_R inverse_mat_base_case_2_R unfolding invertible_mat_def using mat_base_case_R + by auto + have h: "set (get_signs_R (reduce_system_R p (qs, ((matrix_A_R ?signs ?subsets), (?subsets, ?signs))))) = + set (characterize_consistent_signs_at_roots p qs)" + using nonzero nonzero well_def_signs distinct_signs all_info match invertible_mat + reduce_system_sign_conditions_R[where p = "p", where qs = "qs", where signs = ?signs, where subsets = ?subsets] + by blast + then have "set (snd (snd (reduce_system_R p (qs, (?mat, (?subsets, ?signs)))))) = + set (characterize_consistent_signs_at_roots p qs)" + unfolding get_signs_R_def using mat_help by auto + then have "set (snd (snd (reduce_system_R p (qs, base_case_info_R)))) = set (characterize_consistent_signs_at_roots p qs)" + unfolding base_case_info_R_def + by auto + then show ?thesis using len1 + by (simp add: find_consistent_signs_at_roots_thm_R) +qed + +lemma smaller_sys_are_good_R: + fixes p:: "real poly" + fixes qs1 :: "real poly list" + fixes qs2 :: "real poly list" + fixes subsets :: "(nat list*nat list) list" + fixes signs :: "rat list list" + assumes nonzero: "p \ 0" + assumes nontriv1: "length qs1 > 0" + assumes nontriv2: "length qs2 > 0" + assumes "set(find_consistent_signs_at_roots_R p qs1) = set(characterize_consistent_signs_at_roots p qs1)" + assumes "set(find_consistent_signs_at_roots_R p qs2) = set(characterize_consistent_signs_at_roots p qs2)" + shows "set(snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) + = set(characterize_consistent_signs_at_roots p (qs1@qs2))" +proof - + let ?signs = "(get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) " + let ?subsets = "(get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) " + have h0: "satisfies_properties_R p (qs1@qs2) ?subsets ?signs + (get_matrix_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" + using calculate_data_satisfies_properties_R combining_sys_satisfies_properties_R + using nontriv1 nontriv2 nonzero nonzero + by simp + then have h1: "set(characterize_consistent_signs_at_roots p (qs1@qs2)) \ set ?signs" + unfolding satisfies_properties_R_def + by linarith + have h2: "well_def_signs (length (qs1@qs2)) ?signs" + using calculate_data_satisfies_properties_R + using h0 satisfies_properties_R_def by blast + have h3: "distinct ?signs" + using calculate_data_satisfies_properties_R + using h0 satisfies_properties_R_def by blast + have h4: "satisfy_equation_R p (qs1@qs2) ?subsets ?signs" + using calculate_data_satisfies_properties_R + using h0 satisfies_properties_R_def by blast + have h5: "invertible_mat (matrix_A_R ?signs ?subsets) " + using calculate_data_satisfies_properties_R + using h0 satisfies_properties_R_def + by auto + have h: "set (take_indices ?signs + (find_nonzeros_from_input_vec (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R ?signs ?subsets)))) + = set(characterize_consistent_signs_at_roots p (qs1@qs2))" + using h1 h2 h3 h4 h5 reduction_deletes_bad_sign_conds_R + using nonzero reduction_signs_R_def by auto + then have h: "set (characterize_consistent_signs_at_roots p (qs1@qs2)) = + set (reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets ))" + unfolding reduction_signs_R_def get_signs_R_def + by blast + have help_h: "reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets) + = (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R?signs ?subsets))))" + unfolding reduction_signs_R_def by auto + have clear_signs: "(signs_smash (get_signs_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs2))) = (get_signs_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" + using combining_to_smash get_signs_R_def getter_functions_R smash_systems_R_def snd_conv + proof - + have "combine_systems_R p (qs1, calculate_data_R p qs1) (qs2, calculate_data_R p qs2) = (qs1 @ qs2, kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2)), subsets_smash_R (length qs1) (get_subsets_R (calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2)), signs_smash (snd (snd (calculate_data_R p qs1))) (snd (snd (calculate_data_R p qs2))))" + by (metis (no_types) combine_systems_R.simps get_signs_R_def getter_functions_R smash_systems_R_def) + then show ?thesis + by (simp add: get_signs_R_def) + qed + have clear_subsets: "(subsets_smash_R (length qs1) (get_subsets_R(calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2))) = (get_subsets_R (snd ((combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))))" + using Pair_inject combining_to_smash get_subsets_R_def prod.collapse smash_systems_R_def + by (smt (z3) combine_systems_R.simps) + have "well_def_signs (length qs1) (get_signs_R (calculate_data_R p qs1))" + using calculate_data_satisfies_properties_R + using nontriv1 nonzero nonzero satisfies_properties_R_def + by auto + then have well_def_signs1: "(\j. j \ set (get_signs_R (calculate_data_R p qs1)) \ length j = (length qs1))" + using well_def_signs_def by blast + have "all_list_constr_R (get_subsets_R(calculate_data_R p qs1)) (length qs1) " + using calculate_data_satisfies_properties_R + using nontriv1 nonzero nonzero satisfies_properties_R_def + by auto + then have well_def_subsets1: "(\l i. l \ set (get_subsets_R(calculate_data_R p qs1)) \ (i \ set (fst l) \ i < (length qs1)) \ (i \ set (snd l) \ i < (length qs1)))" + unfolding all_list_constr_R_def list_constr_def + using in_set_member + by (metis in_set_conv_nth list_all_length) + have extra_matrix_same: "matrix_A_R (signs_smash (get_signs_R (calculate_data_R p qs1)) (get_signs_R (calculate_data_R p qs2))) + (subsets_smash_R (length qs1) (get_subsets_R(calculate_data_R p qs1)) (get_subsets_R (calculate_data_R p qs2))) + = kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2))" + using well_def_signs1 well_def_subsets1 + using matrix_construction_is_kronecker_product_R + using calculate_data_satisfies_properties_R nontriv1 nontriv2 nonzero nonzero satisfies_properties_R_def + by fastforce + then have matrix_same: "matrix_A_R ?signs ?subsets = kronecker_product (get_matrix_R (calculate_data_R p qs1)) (get_matrix_R (calculate_data_R p qs2))" + using clear_signs clear_subsets + by simp + have comb_sys_h: "snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2)))) = + snd(snd(reduce_system_R p (qs1@qs2, (matrix_A_R ?signs ?subsets, (?subsets, ?signs)))))" + unfolding get_signs_R_def get_subsets_R_def using matrix_same + by (metis (full_types) clear_signs clear_subsets combine_systems_R.simps get_signs_R_def get_subsets_R_def getter_functions_R smash_systems_R_def) + then have extra_h: " snd(snd(reduce_system_R p (qs1@qs2, (matrix_A_R ?signs ?subsets, (?subsets, ?signs))))) = + snd(snd(reduction_step_R (matrix_A_R ?signs ?subsets) ?signs ?subsets (solve_for_lhs_R p (qs1@qs2) ?subsets (matrix_A_R ?signs ?subsets)))) " + by simp + then have same_h: "set(snd(snd(reduce_system_R p (combine_systems_R p (qs1,calculate_data_R p qs1) (qs2,calculate_data_R p qs2))))) + = set (reduction_signs_R p (qs1@qs2) ?signs ?subsets (matrix_A_R ?signs ?subsets ))" + using comb_sys_h unfolding reduction_signs_R_def + by (metis get_signs_R_def help_h reduction_signs_is_get_signs_R) + then show ?thesis using h + by blast +qed + +lemma find_consistent_signs_at_roots_1_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + shows "(p \ 0 \ length qs > 0) \ + set(find_consistent_signs_at_roots_R p qs) = set(characterize_consistent_signs_at_roots 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 \ [] \ + set (find_consistent_signs_at_roots_R p qsa) = + set (characterize_consistent_signs_at_roots p qsa))" + assume nonzero: "p \ 0" + assume nontriv: "qs \ []" + let ?len = "length qs" + let ?q1 = "take ((?len) div 2) qs" + let ?left = "calculate_data_R p ?q1" + let ?q2 = "drop ((?len) div 2) qs" + let ?right = "calculate_data_R 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 nontriv_q2: "length qs>1 \length ?q2 > 0" + by auto + have qs_more_q2: "length qs>1 \ length qs > length ?q2" + by auto + have key_h: "set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) + else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) + (reduce_system_R p)))) = + set (characterize_consistent_signs_at_roots p qs)" + proof - + have h_len1 : "?len = 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) + else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) + (reduce_system_R p)))) = + set (characterize_consistent_signs_at_roots p qs)" + using find_consistent_signs_at_roots_len1_R[of p qs] nonzero nontriv + by (simp add: find_consistent_signs_at_roots_thm_R) + have h_len_gt1 : "?len > 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system_R p (qs, base_case_info_R) + else Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) + (reduce_system_R p)))) = + set (characterize_consistent_signs_at_roots p qs)" + proof - + have h_imp_a: "?len > 1 \ set (snd (snd (reduce_system_R p (combine_systems_R p (?q1, ?left) (?q2, ?right))))) = + set (characterize_consistent_signs_at_roots p qs)" + proof - + have h1: "?len > 1 \ set(snd(snd(?left))) = set (characterize_consistent_signs_at_roots p ?q1)" + using nontriv_q1 ind_hyp[of ?q1] qs_more_q1 + by (metis find_consistent_signs_at_roots_thm_R less_numeral_extra(3) list.size(3)) + have h2: "?len > 1 \ set(snd(snd(?right))) = set (characterize_consistent_signs_at_roots p ?q2)" + using nontriv_q2 ind_hyp[of ?q2] qs_more_q2 + by (metis (full_types) find_consistent_signs_at_roots_thm_R list.size(3) not_less_zero) + show ?thesis using nonzero nontriv_q1 nontriv_q2 h1 h2 smaller_sys_are_good_R + by (metis append_take_drop_id find_consistent_signs_at_roots_thm_R) + qed + then have h_imp: "?len > 1 \ set (snd (snd (Let (combine_systems_R p (?q1, ?left) (?q2, ?right)) + (reduce_system_R p)))) = + set (characterize_consistent_signs_at_roots 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_R p qs) = set (characterize_consistent_signs_at_roots p qs)" + using One_nat_def calculate_data.simps find_consistent_signs_at_roots_thm length_0_conv nontriv + by (smt (z3) calculate_data_R.simps find_consistent_signs_at_roots_thm_R) + qed +qed + +lemma find_consistent_signs_at_roots_0_R: + fixes p:: "real poly" + assumes "p \ 0" + shows "set(find_consistent_signs_at_roots_R p []) = + set(characterize_consistent_signs_at_roots p [])" +proof - + obtain a b c where abc: "reduce_system_R p ([1], base_case_info_R) = (a,b,c)" + using prod_cases3 by blast + have "find_consistent_signs_at_roots_R p [1] = c" using abc + by (simp add: find_consistent_signs_at_roots_thm_R) + have *: "set (find_consistent_signs_at_roots_R p [1]) = set (characterize_consistent_signs_at_roots p [1])" + apply (subst find_consistent_signs_at_roots_1_R) + using assms by auto + have "set(characterize_consistent_signs_at_roots p []) = drop 1 ` set(characterize_consistent_signs_at_roots p [1])" + unfolding characterize_consistent_signs_at_roots_def consistent_sign_vec_def signs_at_def squash_def apply simp + using drop0 drop_Suc_Cons image_cong image_image + proof - + have "(\r. []) ` set (characterize_root_list_p p) = (\r. drop (Suc 0) [1::rat]) ` set (characterize_root_list_p p)" + by force + then show "(\r. []) ` set (characterize_root_list_p p) = drop (Suc 0) ` (\r. [1::rat]) ` set (characterize_root_list_p p)" + by blast + qed + thus ?thesis using abc * + apply (auto) apply (simp add: find_consistent_signs_at_roots_thm_R) + by (simp add: find_consistent_signs_at_roots_thm_R) +qed + +lemma find_consistent_signs_at_roots_R: + fixes p:: "real poly" + fixes qs :: "real poly list" + assumes "p \ 0" + shows "set(find_consistent_signs_at_roots_R p qs) = set(characterize_consistent_signs_at_roots p qs)" + by (metis assms find_consistent_signs_at_roots_0_R find_consistent_signs_at_roots_1_R length_greater_0_conv) + +end \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/document/root.bib b/thys/BenOr_Kozen_Reif/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/document/root.bib @@ -0,0 +1,32 @@ +@article{DBLP:journals/jcss/Ben-OrKR86, + author = {Michael Ben{-}Or and + Dexter Kozen and + John H. Reif}, + title = {The Complexity of Elementary Algebra and Geometry}, + journal = {J. Comput. Syst. Sci.}, + volume = {32}, + number = {2}, + pages = {251--264}, + year = {1986}, + skipurl = {https://doi.org/10.1016/0022-0000(86)90029-2}, + doi = {10.1016/0022-0000(86)90029-2}, + timestamp = {Sat, 20 May 2017 00:25:56 +0200}, + biburl = {https://dblp.org/rec/journals/jcss/Ben-OrKR86.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} + +@article{DBLP:journals/jsc/Renegar92b, + author = {James Renegar}, + title = {On the Computational Complexity and Geometry of the First-Order Theory + of the Reals, Part {III:} {Q}uantifier Elimination}, + journal = {J. Symb. Comput.}, + volume = {13}, + number = {3}, + pages = {329--352}, + year = {1992}, + skipurl = {https://doi.org/10.1016/S0747-7171(10)80005-7}, + doi = {10.1016/S0747-7171(10)80005-7}, + timestamp = {Wed, 17 May 2017 14:25:48 +0200}, + biburl = {https://dblp.org/rec/journals/jsc/Renegar92b.bib}, + bibsource = {dblp computer science bibliography, https://dblp.org} +} \ No newline at end of file diff --git a/thys/BenOr_Kozen_Reif/document/root.tex b/thys/BenOr_Kozen_Reif/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/BenOr_Kozen_Reif/document/root.tex @@ -0,0 +1,72 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{The BKR Decision Procedure for Univariate Real Arithmetic} +\author{Katherine Cordwell, Yong Kiam Tan, and Andr\'e Platzer} +\maketitle + +\begin{abstract} +We formalize the univariate case of Ben-Or, Kozen, and Reif's decision procedure for first-order real arithmetic \cite{DBLP:journals/jcss/Ben-OrKR86} (the BKR algorithm). We also formalize the univariate case of Renegar's variation \cite{DBLP:journals/jsc/Renegar92b} of the BKR algorithm. The two formalizations differ mathematically in minor ways (that have significant impact on the multivariate case), but are quite similar in proof structure. Both rely on sign-determination (finding the set of consistent sign assignments for a set of polynomials). The method used for sign-determination is similar to Tarski's original quantifier elimination algorithm (it stores key information in a matrix equation), but with a reduction step to keep complexity low. +\end{abstract} + +\section*{Remark} Note that theories BKR\_Decision and Renegar\_Decision inherit oracles $\mathtt{holds\_by\_evaluation}$ and $\mathtt{cancel\_type\_definition}$ from Berlekamp\_Zassenhaus. + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +\section*{Acknowledgements} +This material is based upon work supported by the National Science Foundation Graduate Research Fellowship under Grants Nos. DGE1252522 and DGE1745016. +Any opinions, findings, and conclusions or recommendations expressed in this material are those of the authors and do not necessarily reflect the views of the National Science Foundation. +This research was also sponsored by the National Science Foundation under Grant No. CNS-1739629, the AFOSR under grant number FA9550-16-1-0288, and A*STAR, Singapore. +The views and conclusions contained in this document are those of the authors and should not be interpreted as representing the official policies, either expressed or implied, of any sponsoring institution, the U.S. government or any other entity. + + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,595 +1,596 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations BTree Banach_Steinhaus Bell_Numbers_Spivey +BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties CSP_RefTK DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DOM_Components DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaGeoCoq Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knuth_Bendix_Order Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mereology Mersenne_Primes MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_arithmetic_LLL_and_HNF_algorithms Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PCF PLM POPLmark-deBruijn PSemigroupsConvolution Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Shadow_DOM Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL