diff --git a/thys/BenOr_Kozen_Reif/More_Matrix.thy b/thys/BenOr_Kozen_Reif/More_Matrix.thy --- a/thys/BenOr_Kozen_Reif/More_Matrix.thy +++ b/thys/BenOr_Kozen_Reif/More_Matrix.thy @@ -1,2012 +1,2013 @@ 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..a * y + b \ {0.. add.commute add.right_neutral div_less div_mult_self1 less_zeroE mod_eq_self_iff_div_eq_0 mod_mult_self1) qed have 22:"(\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) + by (metis "*" bot_nat_0.not_eq_extremum ij(2) index_col mod_less_divisor mult_zero_right not_less_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 by (auto intro!: cong1) 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 "sorted_wrt (<) (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] by (auto simp: pivot_positions_form_helper_1 Suc_le_lessD) qed lemma sorted_pivot_positions: shows "sorted_wrt (<) (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/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy --- a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy @@ -1,3145 +1,3145 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) section \The Berlekamp Algorithm\ theory Berlekamp_Type_Based imports Jordan_Normal_Form.Matrix_Kernel Jordan_Normal_Form.Gauss_Jordan_Elimination Jordan_Normal_Form.Missing_VectorSpace Polynomial_Factorization.Square_Free_Factorization Polynomial_Factorization.Missing_Multiset Finite_Field Chinese_Remainder_Poly Poly_Mod_Finite_Field "HOL-Computational_Algebra.Field_as_Ring" begin hide_const (open) up_ring.coeff up_ring.monom Modules.module subspace Modules.module_hom subsection \Auxiliary lemmas\ context fixes g :: "'b \ 'a :: comm_monoid_mult" begin lemma prod_list_map_filter: "prod_list (map g (filter f xs)) * prod_list (map g (filter (\ x. \ f x) xs)) = prod_list (map g xs)" by (induct xs, auto simp: ac_simps) lemma prod_list_map_partition: assumes "List.partition f xs = (ys, zs)" shows "prod_list (map g xs) = prod_list (map g ys) * prod_list (map g zs)" using assms by (subst prod_list_map_filter[symmetric, of _ f], auto simp: o_def) end lemma coprime_id_is_unit: fixes a::"'b::semiring_gcd" shows "coprime a a \ is_unit a" using dvd_unit_imp_unit by auto lemma dim_vec_of_list[simp]: "dim_vec (vec_of_list x) = length x" by (transfer, auto) lemma length_list_of_vec[simp]: "length (list_of_vec A) = dim_vec A" by (transfer', auto) lemma list_of_vec_vec_of_list[simp]: "list_of_vec (vec_of_list a) = a" proof - { fix aa :: "'a list" have "map (\n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0..n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0.. list_of_vec) {v. dim_vec v = n}" proof (rule comp_inj_on) show "inj_on list_of_vec {v. dim_vec v = n}" by (auto simp add: inj_on_def, transfer, auto simp add: mk_vec_def) show "inj_on Poly (list_of_vec ` {v. dim_vec v = n})" proof (auto simp add: inj_on_def) fix x y::"'c vec" assume "n = dim_vec x" and dim_xy: "dim_vec y = dim_vec x" and Poly_eq: "Poly (list_of_vec x) = Poly (list_of_vec y)" note [simp del] = nth_list_of_vec show "list_of_vec x = list_of_vec y" proof (rule nth_equalityI, auto simp: dim_xy) have length_eq: "length (list_of_vec x ) = length (list_of_vec y)" using dim_xy by (transfer, auto) fix i assume "i < dim_vec x" thus "list_of_vec x ! i = list_of_vec y ! i" using Poly_eq unfolding poly_eq_iff coeff_Poly_eq using dim_xy unfolding nth_default_def by (auto, presburger) qed qed qed corollary inj_Poly_list_of_vec: "inj_on (Poly \ list_of_vec) (carrier_vec n)" using inj_Poly_list_of_vec' unfolding carrier_vec_def . lemma list_of_vec_rw_map: "list_of_vec m = map (\n. m $ n) [0.. []" shows "degree (Poly xs) < length xs" using xs by (induct xs, auto intro: Poly.simps(1)) lemma vec_of_list_list_of_vec[simp]: "vec_of_list (list_of_vec a) = a" by (transfer, auto simp add: mk_vec_def) lemma row_mat_of_rows_list: assumes b: "b < length A" and nc: "\i. i < length A \ length (A ! i) = nc" shows "(row (mat_of_rows_list nc A) b) = vec_of_list (A ! b)" proof (auto simp add: vec_eq_iff) show "dim_col (mat_of_rows_list nc A) = length (A ! b)" unfolding mat_of_rows_list_def using b nc by auto fix i assume i: "i < length (A ! b)" show "row (mat_of_rows_list nc A) b $ i = vec_of_list (A ! b) $ i" using i b nc unfolding mat_of_rows_list_def row_def by (transfer, auto simp add: mk_vec_def mk_mat_def) qed lemma degree_Poly_list_of_vec: assumes n: "x \ carrier_vec n" and n0: "n > 0" shows "degree (Poly (list_of_vec x)) < n" proof - have x_dim: "dim_vec x = n" using n by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] n0 n x_dim) have "degree (Poly (list_of_vec x)) < length (list_of_vec x)" by (rule degree_Poly'[OF l]) also have "... = n" using x_dim by auto finally show ?thesis . qed lemma list_of_vec_nth: assumes i: "i < dim_vec x" shows "list_of_vec x ! i = x $ i" using i by (transfer, auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth': assumes i: "i < dim_vec x" shows "coeff (Poly (list_of_vec x)) i = x $ i" using i by (auto simp add: list_of_vec_nth nth_default_def) lemma list_of_vec_row_nth: assumes x: "x < dim_col A" shows "list_of_vec (row A i) ! x = A $$ (i, x)" using x unfolding row_def by (transfer', auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth: assumes x: "x < dim_col A" shows "coeff (Poly (list_of_vec (row A i))) x = A $$ (i, x)" proof - have "coeff (Poly (list_of_vec (row A i))) x = nth_default 0 (list_of_vec (row A i)) x" unfolding coeff_Poly_eq by simp also have "... = A $$ (i, x)" using x list_of_vec_row_nth unfolding nth_default_def by (auto simp del: nth_list_of_vec) finally show ?thesis . qed lemma inj_on_list_of_vec: "inj_on list_of_vec (carrier_vec n)" unfolding inj_on_def unfolding list_of_vec_rw_map by auto lemma vec_of_list_carrier[simp]: "vec_of_list x \ carrier_vec (length x)" unfolding carrier_vec_def by simp lemma card_carrier_vec: "card (carrier_vec n:: 'b::finite vec set) = CARD('b) ^ n" proof - let ?A = "UNIV::'b set" let ?B = "{xs. set xs \ ?A \ length xs = n}" let ?C = "(carrier_vec n:: 'b::finite vec set)" have "card ?C = card ?B" proof - have "bij_betw (list_of_vec) ?C ?B" proof (unfold bij_betw_def, auto) show "inj_on list_of_vec (carrier_vec n)" by (rule inj_on_list_of_vec) fix x::"'b list" assume n: "n = length x" thus "x \ list_of_vec ` carrier_vec (length x)" unfolding image_def by auto (rule bexI[of _ "vec_of_list x"], auto) qed thus ?thesis using bij_betw_same_card by blast qed also have "... = card ?A ^ n" by (rule card_lists_length_eq, simp) finally show ?thesis . qed lemma finite_carrier_vec[simp]: "finite (carrier_vec n:: 'b::finite vec set)" by (rule card_ge_0_finite, unfold card_carrier_vec, auto) lemma row_echelon_form_dim0_row: assumes "A \ carrier_mat 0 n" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_dim0_col: assumes "A \ carrier_mat n 0" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_one_dim0[simp]: "row_echelon_form (1\<^sub>m 0)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma Poly_list_of_vec_0[simp]: "Poly (list_of_vec (0\<^sub>v 0)) = [:0:]" by (simp add: poly_eq_iff nth_default_def) lemma monic_normalize: assumes "(p :: 'b :: {field,euclidean_ring_gcd} poly) \ 0" shows "monic (normalize p)" by (simp add: assms normalize_poly_old_def) lemma exists_factorization_prod_list: fixes P::"'b::field poly list" assumes "degree (prod_list P) > 0" and "\ u. u \ set P \ degree u > 0 \ monic u" and "square_free (prod_list P)" shows "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" using assms proof (induct P) case Nil thus ?case by auto next case (Cons x P) have sf_P: "square_free (prod_list P)" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons mult.commute square_free_factor) have deg_x: "degree x > 0" using Cons.prems by auto have distinct_P: "distinct P" by (meson Cons.prems(2) Cons.prems(3) distinct.simps(2) square_free_prod_list_distinct) have "\A. finite A \ x = \A \ A \ {q. irreducible q \ monic q}" proof (rule monic_square_free_irreducible_factorization) show "monic x" by (simp add: Cons.prems(2)) show "square_free x" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons square_free_factor) qed from this obtain A where fin_A: "finite A" and xA: "x = \A" and A: "A \ {q. irreducible\<^sub>d q \ monic q}" by auto obtain A' where s: "set A' = A" and length_A': "length A' = card A" using \finite A\ distinct_card finite_distinct_list by force have A_not_empty: "A \ {}" using xA deg_x by auto have x_prod_list_A': "x = prod_list A'" proof - have "x = \A" using xA by simp also have "... = prod id A" by simp also have "... = prod id (set A')" unfolding s by simp also have "... = prod_list (map id A')" by (rule prod.distinct_set_conv_list, simp add: card_distinct length_A' s) also have "... = prod_list A'" by auto finally show ?thesis . qed show ?case proof (cases "P = []") case True show ?thesis proof (rule exI[of _ "A'"], auto simp add: True) show "prod_list A' = x" using x_prod_list_A' by simp show "Suc 0 \ length A'" using A_not_empty using s length_A' by (simp add: Suc_leI card_gt_0_iff fin_A) show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed next case False have hyp: "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule Cons.hyps[OF _ _ sf_P]) have set_P: "set P \ {}" using False by auto have "prod_list P = prod_list (map id P)" by simp also have "... = prod id (set P)" using prod.distinct_set_conv_list[OF distinct_P, of id] by simp also have "... = \(set P)" by simp finally have "prod_list P = \(set P)" . hence "degree (prod_list P) = degree (\(set P))" by simp also have "... = degree (prod id (set P))" by simp also have "... = (\i\(set P). degree (id i))" proof (rule degree_prod_eq_sum_degree) show "\i\set P. id i \ 0" using Cons.prems(2) by force qed also have "... > 0" by (metis Cons.prems(2) List.finite_set set_P gr0I id_apply insert_iff list.set(2) sum_pos) finally show "degree (prod_list P) > 0" by simp show "\u. u \ set P \ degree u > 0 \ monic u" using Cons.prems by auto qed from this obtain Q where QP: "prod_list Q = prod_list P" and length_PQ: "length P \ length Q" and monic_irr_Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast show ?thesis proof (rule exI[of _ "A' @ Q"], auto simp add: monic_irr_Q) show "prod_list A' * prod_list Q = x * prod_list P" unfolding QP x_prod_list_A' by auto have "length A' \ 0" using A_not_empty using s length_A' by auto thus "Suc (length P) \ length A' + length Q" using QP length_PQ by linarith show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed qed qed lemma normalize_eq_imp_smult: fixes p :: "'b :: {euclidean_ring_gcd} poly" assumes n: "normalize p = normalize q" shows "\ c. c \ 0 \ q = smult c p" proof(cases "p = 0") case True with n show ?thesis by (auto intro:exI[of _ 1]) next case p0: False have degree_eq: "degree p = degree q" using n degree_normalize by metis hence q0: "q \ 0" using p0 n by auto have p_dvd_q: "p dvd q" using n by (simp add: associatedD1) from p_dvd_q obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with q0 have "k \ 0" by auto then have "degree k = 0" using degree_eq degree_mult_eq p0 q by fastforce then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed lemma prod_list_normalize: fixes P :: "'b :: {idom_divide,normalization_semidom_multiplicative} poly list" shows "normalize (prod_list P) = prod_list (map normalize P)" proof (induct P) case Nil show ?case by auto next case (Cons p P) have "normalize (prod_list (p # P)) = normalize p * normalize (prod_list P)" using normalize_mult by auto also have "... = normalize p * prod_list (map normalize P)" using Cons.hyps by auto also have "... = prod_list (normalize p # (map normalize P))" by auto also have "... = prod_list (map normalize (p # P))" by auto finally show ?case . qed lemma prod_list_dvd_prod_list_subset: fixes A::"'b::comm_monoid_mult list" assumes dA: "distinct A" and dB: "distinct B" (*Maybe this condition could be avoided*) and s: "set A \ set B" shows "prod_list A dvd prod_list B" proof - have "prod_list A = prod_list (map id A)" by auto also have "... = prod id (set A)" by (rule prod.distinct_set_conv_list[symmetric, OF dA]) also have "... dvd prod id (set B)" by (rule prod_dvd_prod_subset[OF _ s], auto) also have "... = prod_list (map id B)" by (rule prod.distinct_set_conv_list[OF dB]) also have "... = prod_list B" by simp finally show ?thesis . qed end lemma gcd_monic_constant: "gcd f g \ {1, f}" if "monic f" and "degree g = 0" for f g :: "'a :: {field_gcd} poly" proof (cases "g = 0") case True moreover from \monic f\ have "normalize f = f" by (rule normalize_monic) ultimately show ?thesis by simp next case False with \degree g = 0\ have "is_unit g" by simp then have "Rings.coprime f g" by (rule is_unit_right_imp_coprime) then show ?thesis by simp qed lemma distinct_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "distinct (find_base_vectors A)" proof - note non_pivot_base = non_pivot_base[OF ref A] let ?pp = "set (pivot_positions A)" from A have dim: "dim_row A = nr" "dim_col A = nc" by auto { fix j j' assume j: "j < nc" "j \ snd ` ?pp" and j': "j' < nc" "j' \ snd ` ?pp" and neq: "j' \ j" from non_pivot_base(2)[OF j] non_pivot_base(4)[OF j' j neq] have "non_pivot_base A (pivot_positions A) j \ non_pivot_base A (pivot_positions A) j'" by auto } hence inj: "inj_on (non_pivot_base A (pivot_positions A)) (set [j\[0.. snd ` ?pp])" unfolding inj_on_def by auto thus ?thesis unfolding find_base_vectors_def Let_def unfolding distinct_map dim by auto qed lemma length_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "length (find_base_vectors A) = card (set (find_base_vectors A))" using distinct_card[OF distinct_find_base_vectors[OF ref A]] by auto subsection \Previous Results\ definition power_poly_f_mod :: "'a::field poly \ 'a poly \ nat \ 'a poly" where "power_poly_f_mod modulus = (\a n. a ^ n mod modulus)" lemma power_poly_f_mod_binary: "power_poly_f_mod m a n = (if n = 0 then 1 mod m - else let (d, r) = Euclidean_Division.divmod_nat n 2; + else let (d, r) = Euclidean_Rings.divmod_nat n 2; rec = power_poly_f_mod m ((a * a) mod m) d in if r = 0 then rec else (rec * a) mod m)" for m a :: "'a :: {field_gcd} poly" proof - note d = power_poly_f_mod_def show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding d by simp next case False - obtain q r where div: "Euclidean_Division.divmod_nat n 2 = (q,r)" by force - hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding Euclidean_Division.divmod_nat_def by auto + obtain q r where div: "Euclidean_Rings.divmod_nat n 2 = (q,r)" by force + hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding Euclidean_Rings.divmod_nat_def by auto have id: "a ^ (2 * q) = (a * a) ^ q" by (simp add: power_mult_distrib semiring_normalization_rules) show ?thesis proof (cases "r = 0") case True show ?thesis using power_mod [of "a * a" m q] - by (auto simp add: Euclidean_Division.divmod_nat_def Let_def True n d div id) + by (auto simp add: Euclidean_Rings.divmod_nat_def Let_def True n d div id) next case False with r have r: "r = 1" by simp show ?thesis by (auto simp add: d r div Let_def mod_simps) (simp add: n r mod_simps ac_simps power_mult_distrib power_mult power2_eq_square) qed qed qed fun power_polys where "power_polys mul_p u curr_p (Suc i) = curr_p # power_polys mul_p u ((curr_p * mul_p) mod u) i" | "power_polys mul_p u curr_p 0 = []" context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma fermat_theorem_mod_ring [simp]: fixes a::"'a mod_ring" shows "a ^ CARD('a) = a" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis proof transfer fix a assume "a \ {0.. 0" then have a: "1 \ a" "a < int CARD('a)" by simp_all then have [simp]: "a mod int CARD('a) = a" by simp from a have "\ int CARD('a) dvd a" by (auto simp add: zdvd_not_zless) then have "\ CARD('a) dvd nat \a\" by simp with a have "\ CARD('a) dvd nat a" by simp with prime_card have "[nat a ^ (CARD('a) - 1) = 1] (mod CARD('a))" by (rule fermat_theorem) with a have "int (nat a ^ (CARD('a) - 1) mod CARD('a)) = 1" by (simp add: cong_def) with a have "a ^ (CARD('a) - 1) mod CARD('a) = 1" by (simp add: of_nat_mod) then have "a * (a ^ (CARD('a) - 1) mod int CARD('a)) = a" by simp then have "(a * (a ^ (CARD('a) - 1) mod int CARD('a))) mod int CARD('a) = a mod int CARD('a)" by (simp only:) then show "a ^ CARD('a) mod int CARD('a) = a" by (simp add: mod_simps semiring_normalization_rules(27)) qed qed lemma mod_eq_dvd_iff_poly: "((x::'a mod_ring poly) mod n = y mod n) = (n dvd x - y)" proof assume H: "x mod n = y mod n" hence "x mod n - y mod n = 0" by simp hence "(x mod n - y mod n) mod n = 0" by simp hence "(x - y) mod n = 0" by (simp add: mod_diff_eq) thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) next assume H: "n dvd x - y" then obtain k where k: "x-y = n*k" unfolding dvd_def by blast hence "x = n*k + y" using diff_eq_eq by blast hence "x mod n = (n*k + y) mod n" by simp thus "x mod n = y mod n" by (simp add: mod_add_left_eq) qed lemma cong_gcd_eq_poly: "gcd a m = gcd b m" if "[(a::'a mod_ring poly) = b] (mod m)" using that by (simp add: cong_def) (metis gcd_mod_left mod_by_0) lemma coprime_h_c_poly: fixes h::"'a mod_ring poly" assumes "c1 \ c2" shows "coprime (h - [:c1:]) (h - [:c2:])" proof (intro coprimeI) fix d assume "d dvd h - [:c1:]" and "d dvd h - [:c2:]" hence "h mod d = [:c1:] mod d" and "h mod d = [:c2:] mod d" using mod_eq_dvd_iff_poly by simp+ hence "[:c1:] mod d = [:c2:] mod d" by simp hence "d dvd [:c2 - c1:]" by (metis (no_types) mod_eq_dvd_iff_poly diff_pCons right_minus_eq) thus "is_unit d" by (metis (no_types) assms dvd_trans is_unit_monom_0 monom_0 right_minus_eq) qed lemma coprime_h_c_poly2: fixes h::"'a mod_ring poly" assumes "coprime (h - [:c1:]) (h - [:c2:])" and "\ is_unit (h - [:c1:])" shows "c1 \ c2" using assms coprime_id_is_unit by blast lemma degree_minus_eq_right: fixes p::"'b::ab_group_add poly" shows "degree q < degree p \ degree (p - q) = degree p" using degree_add_eq_left[of "-q" p] degree_minus by auto lemma coprime_prod: fixes A::"'a mod_ring set" and g::"'a mod_ring \ 'a mod_ring poly" assumes "\x\A. coprime (g a) (g x)" shows "coprime (g a) (prod (\x. g x) A)" proof - have f: "finite A" by simp show ?thesis using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = (g x) * (\c\A. g c)" by (simp add: insert.hyps(2)) with insert.prems show ?case by (auto simp: insert.hyps(3) prod_coprime_right) qed auto qed lemma coprime_prod2: fixes A::"'b::semiring_gcd set" assumes "\x\A. coprime (a) (x)" and f: "finite A" shows "coprime (a) (prod (\x. x) A)" using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. c) = (x) * (\c\A. c)" by (simp add: insert.hyps) with insert.prems show ?case by (simp add: insert.hyps prod_coprime_right) qed auto lemma divides_prod: fixes g::"'a mod_ring \ 'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2 \ coprime (g c1) (g c2)" assumes "\c\ A. g c dvd f" shows "(\c\A. g c) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = g x * (\c\ A. g c)" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "g x dvd f" using insert.prems by auto show "prod g A dvd f" using insert.hyps(3) insert.prems by auto from insert show "Rings.coprime (g x) (prod g A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto qed (* Proof of equation 9 in the book by Knuth x^p - x = (x-0)(x-1)...(x-(p-1)) (mod p) *) lemma poly_monom_identity_mod_p: "monom (1::'a mod_ring) (CARD('a)) - monom 1 1 = prod (\x. [:0,1:] - [:x:]) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof - let ?f="(\x::'a mod_ring. [:0,1:] - [:x:])" have "?rhs dvd ?lhs" proof (rule divides_prod) { fix a::"'a mod_ring" have "poly ?lhs a = 0" by (simp add: poly_monom) hence "([:0,1:] - [:a:]) dvd ?lhs" using poly_eq_0_iff_dvd by fastforce } thus "\x\UNIV::'a mod_ring set. [:0, 1:] - [:x:] dvd monom 1 CARD('a) - monom 1 1" by fast show "\c1 c2. c1 \ UNIV \ c2 \ UNIV \ c1 \ (c2 :: 'a mod_ring) \ coprime ([:0, 1:] - [:c1:]) ([:0, 1:] - [:c2:])" by (auto dest!: coprime_h_c_poly[of _ _ "[:0,1:]"]) qed from this obtain g where g: "?lhs = ?rhs * g" using dvdE by blast have degree_lhs_card: "degree ?lhs = CARD('a)" proof - have "degree (monom (1::'a mod_ring) 1) = 1" by (simp add: degree_monom_eq) moreover have d_c: "degree (monom (1::'a mod_ring) CARD('a)) = CARD('a)" by (simp add: degree_monom_eq) ultimately have "degree (monom (1::'a mod_ring) 1) < degree (monom (1::'a mod_ring) CARD('a))" using prime_card unfolding prime_nat_iff by auto hence "degree ?lhs = degree (monom (1::'a mod_ring) CARD('a))" by (rule degree_minus_eq_right) thus ?thesis unfolding d_c . qed have degree_rhs_card: "degree ?rhs = CARD('a)" proof - have "degree (prod ?f UNIV) = sum (degree \ ?f) UNIV \ coeff (prod ?f UNIV) (sum (degree \ ?f) UNIV) = 1" by (rule degree_prod_sum_monic, auto) moreover have "sum (degree \ ?f) UNIV = CARD('a)" by auto ultimately show ?thesis by presburger qed have monic_lhs: "monic ?lhs" using degree_lhs_card by auto have monic_rhs: "monic ?rhs" by (rule monic_prod, simp) have degree_eq: "degree ?rhs = degree ?lhs" unfolding degree_lhs_card degree_rhs_card .. have g_not_0: "g \ 0" using g monic_lhs by auto have degree_g0: "degree g = 0" proof - have "degree (?rhs * g) = degree ?rhs + degree g" by (rule degree_monic_mult[OF monic_rhs g_not_0]) thus ?thesis using degree_eq g by simp qed have monic_g: "monic g" using monic_factor g monic_lhs monic_rhs by auto have "g = 1" using monic_degree_0[OF monic_g] degree_g0 by simp thus ?thesis using g by auto qed (* Proof of equation 10 in the book by Knuth v(x)^p - v(x) = (v(x)-0)(v(x)-1)...(v(x)-(p-1)) (mod p) *) lemma poly_identity_mod_p: "v^(CARD('a)) - v = prod (\x. v - [:x:]) (UNIV::'a mod_ring set)" proof - have id: "monom 1 1 \\<^sub>p v = v" "[:0, 1:] \\<^sub>p v = v" unfolding pcompose_def apply (auto) by (simp add: fold_coeffs_def) have id2: "monom 1 (CARD('a)) \\<^sub>p v = v ^ (CARD('a))" by (metis id(1) pcompose_hom.hom_power x_pow_n) show ?thesis using arg_cong[OF poly_monom_identity_mod_p, of "\ f. f \\<^sub>p v"] unfolding pcompose_hom.hom_minus pcompose_hom.hom_prod id pcompose_const id2 . qed lemma coprime_gcd: fixes h::"'a mod_ring poly" assumes "Rings.coprime (h-[:c1:]) (h-[:c2:])" shows "Rings.coprime (gcd f(h-[:c1:])) (gcd f (h-[:c2:]))" using assms coprime_divisors by blast lemma divides_prod_gcd: fixes h::"'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2\ coprime (h-[:c1:]) (h-[:c2:])" shows "(\c\A. gcd f (h - [:c:])) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "gcd f (h - [:x:]) dvd f" by simp show "(\c\A. gcd f (h - [:c:])) dvd f" using insert.hyps(3) insert.prems by auto show "Rings.coprime (gcd f (h - [:x:])) (\c\A. gcd f (h - [:c:]))" by (rule prod_coprime_right) (metis Berlekamp_Type_Based.coprime_h_c_poly coprime_gcd coprime_iff_coprime insert.hyps(2)) qed finally show ?case . qed auto qed lemma monic_prod_gcd: assumes f: "finite A" and f0: "(f :: 'b :: {field_gcd} poly) \ 0" shows "monic (\c\A. gcd f (h - [:c:]))" using f proof (induct A) case (insert x A) have rw: "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps) show ?case proof (unfold rw, rule monic_mult) show "monic (gcd f (h - [:x:]))" using poly_gcd_monic[of f] f0 using insert.prems insert_iff by blast show "monic (\c\A. gcd f (h - [:c:]))" using insert.hyps(3) insert.prems by blast qed qed auto lemma coprime_not_unit_not_dvd: fixes a::"'b::semiring_gcd" assumes "a dvd b" and "coprime b c" and "\ is_unit a" shows "\ a dvd c" using assms coprime_divisors coprime_id_is_unit by fastforce lemma divides_prod2: fixes A::"'b::semiring_gcd set" assumes f: "finite A" and "\a\A. a dvd c" and "\a1 a2. a1 \ A \ a2 \ A \ a1 \ a2 \ coprime a1 a2" shows "\A dvd c" using assms proof (induct A) case (insert x A) have "\(insert x A) = x * \A" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... dvd c" proof (rule divides_mult) show "x dvd c" by (simp add: insert.prems) show "\A dvd c" using insert by auto from insert show "Rings.coprime x (\A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto lemma coprime_polynomial_factorization: fixes a1 :: "'b :: {field_gcd} poly" assumes irr: "as \ {q. irreducible q \ monic q}" and "finite as" and a1: "a1 \ as" and a2: "a2 \ as" and a1_not_a2: "a1 \ a2" shows "coprime a1 a2" proof (rule ccontr) assume not_coprime: "\ coprime a1 a2" let ?b= "gcd a1 a2" have b_dvd_a1: "?b dvd a1" and b_dvd_a2: "?b dvd a2" by simp+ have irr_a1: "irreducible a1" using a1 irr by blast have irr_a2: "irreducible a2" using a2 irr by blast have a2_not0: "a2 \ 0" using a2 irr by auto have degree_a1: "degree a1 \ 0" using irr_a1 by auto have degree_a2: "degree a2 \ 0" using irr_a2 by auto have not_a2_dvd_a1: "\ a2 dvd a1" proof (rule ccontr, simp) assume a2_dvd_a1: "a2 dvd a1" from this obtain k where k: "a1 = a2 * k" unfolding dvd_def by auto have k_not0: "k \ 0" using degree_a1 k by auto show False proof (cases "degree a2 = degree a1") case False have "degree a2 < degree a1" using False a2_dvd_a1 degree_a1 divides_degree by fastforce hence "\ irreducible a1" using degree_a2 a2_dvd_a1 degree_a2 by (metis degree_a1 irreducible\<^sub>dD(2) irreducible\<^sub>d_multD irreducible_connect_field k neq0_conv) thus False using irr_a1 by contradiction next case True have "degree a1 = degree a2 + degree k" unfolding k using degree_mult_eq[OF a2_not0 k_not0] by simp hence "degree k = 0" using True by simp hence "k = 1" using monic_factor a1 a2 irr k monic_degree_0 by auto hence "a1 = a2" using k by simp thus False using a1_not_a2 by contradiction qed qed have b_not0: "?b \ 0" by (simp add: a2_not0) have degree_b: "degree ?b > 0" using not_coprime[simplified] b_not0 is_unit_gcd is_unit_iff_degree by blast have "degree ?b < degree a2" by (meson b_dvd_a1 b_dvd_a2 irreducibleD' dvd_trans gcd_dvd_1 irr_a2 not_a2_dvd_a1 not_coprime) hence "\ irreducible\<^sub>d a2" using degree_a2 b_dvd_a2 degree_b by (metis degree_smult_eq irreducible\<^sub>d_dvd_smult less_not_refl3) thus False using irr_a2 by auto qed (* Proof of equation 14 in the book by Knuth *) theorem Berlekamp_gcd_step: fixes f::"'a mod_ring poly" and h::"'a mod_ring poly" assumes hq_mod_f: "[h^(CARD('a)) = h] (mod f)" and monic_f: "monic f" and sf_f: "square_free f" shows "f = prod (\c. gcd f (h - [:c:])) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof (cases "f=0") case True thus ?thesis using coeff_0 monic_f zero_neq_one by auto next case False note f_not_0 = False show ?thesis proof (rule poly_dvd_antisym) show "?rhs dvd f" using coprime_h_c_poly by (intro divides_prod_gcd, auto) have "monic ?rhs" by (rule monic_prod_gcd[OF _ f_not_0], simp) thus "coeff f (degree f) = coeff ?rhs (degree ?rhs)" using monic_f by auto next show "f dvd ?rhs" proof - let ?p = "CARD('a)" obtain P where finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" using monic_square_free_irreducible_factorization[OF monic_f sf_f] by auto have f_dvd_hqh: "f dvd (h^?p - h)" using hq_mod_f unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "f = \P" using f_desc_square_free by simp also have "... dvd ?rhs" proof (rule divides_prod2[OF finite_P]) show "\a1 a2. a1 \ P \ a2 \ P \ a1 \ a2 \ coprime a1 a2" using coprime_polynomial_factorization[OF P finite_P] by simp show "\a\P. a dvd (\c\UNIV. gcd f (h - [:c:]))" proof fix fi assume fi_P: "fi \ P" show "fi dvd ?rhs" proof (rule dvd_prod, auto) show "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible (fi)" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) have fi_dvd_hc: "\c\UNIV::'a mod_ring set. fi dvd (h-[:c:])" by (rule irreducible_dvd_prod[OF _ fi_dvd_prod_hc], simp add: irr_fi) thus "\c. fi dvd h - [:c:]" by simp qed qed qed finally show "f dvd ?rhs" . qed qed qed (******* Implementation of Berlekamp's algorithm (type-based version) *******) subsection \Definitions\ definition berlekamp_mat :: "'a mod_ring poly \ 'a mod_ring mat" where "berlekamp_mat u = (let n = degree u; mul_p = power_poly_f_mod u [:0,1:] (CARD('a)); xks = power_polys mul_p u 1 n in mat_of_rows_list n (map (\ cs. let coeffs_cs = (coeffs cs); k = n - length (coeffs cs) in (coeffs cs) @ replicate k 0) xks))" definition berlekamp_resulting_mat :: "('a mod_ring) poly \ 'a mod_ring mat" where "berlekamp_resulting_mat u = (let Q = berlekamp_mat u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then Q $$ (i,j) - 1 else Q $$ (i,j)) in (gauss_jordan_single (transpose_mat QI)))" definition berlekamp_basis :: "'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_basis u = (map (Poly o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" lemma berlekamp_basis_code[code]: "berlekamp_basis u = (map (poly_of_list o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" unfolding berlekamp_basis_def poly_of_list_def .. primrec berlekamp_factorization_main :: "nat \ 'a mod_ring poly list \ 'a mod_ring poly list \ nat \ 'a mod_ring poly list" where "berlekamp_factorization_main i divs (v # vs) n = (if v = 1 then berlekamp_factorization_main i divs vs n else if length divs = n then divs else let facts = [ w . u \ divs, s \ [0 ..< CARD('a)], w \ [gcd u (v - [:of_int s:])], w \ 1]; (lin,nonlin) = List.partition (\ q. degree q = i) facts in lin @ berlekamp_factorization_main i nonlin vs (n - length lin))" | "berlekamp_factorization_main i divs [] n = divs" definition berlekamp_monic_factorization :: "nat \ 'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_monic_factorization d f = (let vs = berlekamp_basis f; n = length vs; fs = berlekamp_factorization_main d [f] vs n in fs)" subsection \Properties\ lemma power_polys_works: fixes u::"'b::unique_euclidean_semiring" assumes i: "i < n" and c: "curr_p = curr_p mod u" (*Equivalent to degree curr_p < degree u*) shows "power_polys mult_p u curr_p n ! i = curr_p * mult_p ^ i mod u" using i c proof (induct n arbitrary: curr_p i) case 0 thus ?case by simp next case (Suc n) have p_rw: "power_polys mult_p u curr_p (Suc n) ! i = (curr_p # power_polys mult_p u (curr_p * mult_p mod u) n) ! i" by simp show ?case proof (cases "i=0") case True show ?thesis using Suc.prems unfolding p_rw True by auto next case False note i_not_0 = False show ?thesis proof (cases "i < n") case True note i_less_n = True have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (i - 1)" unfolding p_rw using nth_Cons_pos False by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (i-1) mod u" by (rule Suc.hyps) (auto simp add: i_less_n less_imp_diff_less) also have "... = curr_p * mult_p ^ i mod u" using False by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . next case False hence i_n: "i = n" using Suc.prems by auto have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (n - 1)" unfolding p_rw using nth_Cons_pos i_n i_not_0 by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (n-1) mod u" proof (rule Suc.hyps) show "n - 1 < n" using i_n i_not_0 by linarith show "curr_p * mult_p mod u = curr_p * mult_p mod u mod u" by simp qed also have "... = curr_p * mult_p ^ i mod u" using i_n [symmetric] i_not_0 by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . qed qed qed lemma length_power_polys[simp]: "length (power_polys mult_p u curr_p n) = n" by (induct n arbitrary: curr_p, auto) (* Equation 12 *) lemma Poly_berlekamp_mat: assumes k: "k < degree u" shows "Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k) mod u" proof - let ?map ="(map (\cs. coeffs cs @ replicate (degree u - length (coeffs cs)) 0) (power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)))" have "row (berlekamp_mat u) k = row (mat_of_rows_list (degree u) ?map) k" by (simp add: berlekamp_mat_def Let_def) also have "... = vec_of_list (?map ! k)" proof- { fix i assume i: "i < degree u" let ?c= "power_polys (power_poly_f_mod u [:0, 1:] CARD('a)) u 1 (degree u) ! i" let ?coeffs_c="(coeffs ?c)" have "?c = 1*([:0, 1:] ^ CARD('a) mod u)^i mod u" proof (unfold power_poly_f_mod_def, rule power_polys_works[OF i]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = [:0, 1:] ^ (CARD('a) * i) mod u" by (simp add: power_mod power_mult) finally have c_rw: "?c = [:0, 1:] ^ (CARD('a) * i) mod u" . have "length ?coeffs_c \ degree u" proof - show ?thesis proof (cases "?c = 0") case True thus ?thesis by auto next case False have "length ?coeffs_c = degree (?c) + 1" by (rule length_coeffs[OF False]) also have "... = degree ([:0, 1:] ^ (CARD('a) * i) mod u) + 1" using c_rw by simp also have "... \ degree u" by (metis One_nat_def add.right_neutral add_Suc_right c_rw calculation coeffs_def degree_0 degree_mod_less discrete gr_implies_not0 k list.size(3) one_neq_zero) finally show ?thesis . qed qed then have "length ?coeffs_c + (degree u - length ?coeffs_c) = degree u" by auto } with k show ?thesis by (intro row_mat_of_rows_list, auto) qed finally have row_rw: "row (berlekamp_mat u) k = vec_of_list (?map ! k)" . have "Poly (list_of_vec (row (berlekamp_mat u) k)) = Poly (list_of_vec (vec_of_list (?map ! k)))" unfolding row_rw .. also have "... = Poly (?map ! k)" by simp also have "... = [:0,1:]^(CARD('a) * k) mod u" proof - let ?cs = "(power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)) ! k" let ?c = "coeffs ?cs @ replicate (degree u - length (coeffs ?cs)) 0" have map_k_c: "?map ! k = ?c" by (rule nth_map, simp add: k) have "(Poly (?map ! k)) = Poly (coeffs ?cs)" unfolding map_k_c Poly_append_replicate_0 .. also have "... = ?cs" by simp also have "... = power_polys ([:0, 1:] ^ CARD('a) mod u) u 1 (degree u) ! k" by (simp add: power_poly_f_mod_def) also have "... = 1* ([:0,1:]^(CARD('a)) mod u) ^ k mod u" proof (rule power_polys_works[OF k]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = ([:0,1:]^(CARD('a)) mod u) ^ k mod u" by auto also have "... = [:0,1:]^(CARD('a) * k) mod u" by (simp add: power_mod power_mult) finally show ?thesis . qed finally show ?thesis . qed corollary Poly_berlekamp_cong_mat: assumes k: "k < degree u" shows "[Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k)] (mod u)" using Poly_berlekamp_mat[OF k] unfolding cong_def by auto lemma mat_of_rows_list_dim[simp]: "mat_of_rows_list n vs \ carrier_mat (length vs) n" "dim_row (mat_of_rows_list n vs) = length vs" "dim_col (mat_of_rows_list n vs) = n" unfolding mat_of_rows_list_def by auto lemma berlekamp_mat_closed[simp]: "berlekamp_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_mat u) = degree u" "dim_col (berlekamp_mat u) = degree u" unfolding carrier_mat_def berlekamp_mat_def Let_def by auto lemma vec_of_list_coeffs_nth: assumes i: "i \ {..degree h}" and h_not0: "h \ 0" shows "vec_of_list (coeffs h) $ i = coeff h i" proof - have "vec_of_list (map (coeff h) [0..i. f i mod z) A" using f by (induct, auto simp add: poly_mod_add_left) lemma prime_not_dvd_fact: assumes kn: "k < n" and prime_n: "prime n" shows "\ n dvd fact k" using kn proof (induct k) case 0 thus ?case using prime_n unfolding prime_nat_iff by auto next case (Suc k) show ?case proof (rule ccontr, unfold not_not) assume "n dvd fact (Suc k)" also have "... = Suc k * \{1..k}" unfolding fact_Suc unfolding fact_prod by simp finally have "n dvd Suc k * \{1..k}" . hence "n dvd Suc k \ n dvd \{1..k}" using prime_dvd_mult_eq_nat[OF prime_n] by blast moreover have "\ n dvd Suc k" by (simp add: Suc.prems(1) nat_dvd_not_less) moreover hence "\ n dvd \{1..k}" using Suc.hyps Suc.prems using Suc_lessD fact_prod[of k] by (metis of_nat_id) ultimately show False by simp qed qed lemma dvd_choose_prime: assumes kn: "k < n" and k: "k \ 0" and n: "n \ 0" and prime_n: "prime n" shows "n dvd (n choose k)" proof - have "n dvd (fact n)" by (simp add: fact_num_eq_if n) moreover have "\ n dvd (fact k * fact (n-k))" proof (rule ccontr, simp) assume "n dvd fact k * fact (n - k)" hence "n dvd fact k \ n dvd fact (n - k)" using prime_dvd_mult_eq_nat[OF prime_n] by simp moreover have "\ n dvd (fact k)" by (rule prime_not_dvd_fact[OF kn prime_n]) moreover have "\ n dvd fact (n - k)" using prime_not_dvd_fact[OF _ prime_n] kn k by simp ultimately show False by simp qed moreover have "(fact n::nat) = fact k * fact (n-k) * (n choose k)" using binomial_fact_lemma kn by auto ultimately show ?thesis using prime_n by (auto simp add: prime_dvd_mult_iff) qed lemma add_power_poly_mod_ring: fixes x :: "'a mod_ring poly" shows "(x + y) ^ CARD('a) = x ^ CARD('a) + y ^ CARD('a)" proof - let ?A="{0..CARD('a)}" let ?f="\k. of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k)" have A_rw: "?A = insert CARD('a) (insert 0 (?A - {0} - {CARD('a)}))" by fastforce have sum0: "sum ?f (?A - {0} - {CARD('a)}) = 0" proof (rule sum.neutral, rule) fix xa assume xa: "xa \ {0..CARD('a)} - {0} - {CARD('a)}" have card_dvd_choose: "CARD('a) dvd (CARD('a) choose xa)" proof (rule dvd_choose_prime) show "xa < CARD('a)" using xa by simp show "xa \ 0" using xa by simp show "CARD('a) \ 0" by simp show "prime CARD('a)" by (rule prime_card) qed hence rw0: "of_int (CARD('a) choose xa) = (0 :: 'a mod_ring)" by transfer simp have "of_nat (CARD('a) choose xa) = [:of_int (CARD('a) choose xa) :: 'a mod_ring:]" by (simp add: of_nat_poly) also have "... = [:0:]" using rw0 by simp finally show "of_nat (CARD('a) choose xa) * x ^ xa * y ^ (CARD('a) - xa) = 0" by auto qed have "(x + y)^CARD('a) = (\k = 0..CARD('a). of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k))" unfolding binomial_ring by (rule sum.cong, auto) also have "... = sum ?f (insert CARD('a) (insert 0 (?A - {0} - {CARD('a)})))" using A_rw by simp also have "... = ?f 0 + ?f CARD('a) + sum ?f (?A - {0} - {CARD('a)})" by auto also have "... = x^CARD('a) + y^CARD('a)" unfolding sum0 by auto finally show ?thesis . qed lemma power_poly_sum_mod_ring: fixes f :: "'b \ 'a mod_ring poly" assumes f: "finite A" shows "(sum f A) ^ CARD('a) = sum (\i. (f i) ^ CARD('a)) A" using f by (induct, auto simp add: add_power_poly_mod_ring) lemma poly_power_card_as_sum_of_monoms: fixes h :: "'a mod_ring poly" shows "h ^ CARD('a) = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof - have "h ^ CARD('a) = (\i\degree h. monom (coeff h i) i) ^ CARD('a)" by (simp add: poly_as_sum_of_monoms) also have "... = (\i\degree h. (monom (coeff h i) i) ^ CARD('a))" by (simp add: power_poly_sum_mod_ring) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" show "monom (coeff h x) x ^ CARD('a) = monom (coeff h x) (CARD('a) * x)" by (unfold poly_eq_iff, auto simp add: monom_power) qed finally show ?thesis . qed lemma degree_Poly_berlekamp_le: assumes i: "i < degree u" shows "degree (Poly (list_of_vec (row (berlekamp_mat u) i))) < degree u" by (metis Poly_berlekamp_mat degree_0 degree_mod_less gr_implies_not0 i linorder_neqE_nat) (* Equation 12: alternative statement. *) lemma monom_card_pow_mod_sum_berlekamp: assumes i: "i < degree u" shows "monom 1 (CARD('a) * i) mod u = (\j 0" using i by simp hence set_rw: "{..degree u - 1} = {.. degree u - 1" by auto have "monom 1 (CARD('a) * i) mod u = [:0, 1:] ^ (CARD('a) * i) mod u" using x_as_monom x_pow_n by metis also have "... = ?p" unfolding Poly_berlekamp_mat[OF i] by simp also have "... = (\i\degree u - 1. monom (coeff ?p i) i)" using degree_le2 poly_as_sum_of_monoms' by fastforce also have "... = (\ij {.. v = (\i = 0.. v = (\i = 0.. v = col A j \ v" using j row_transpose by auto also have "... = (\i = 0..iiiiiiiiiii degree u" by (metis Suc_leI assms coeffs_0_eq_Nil degree_0 length_coeffs_degree list.size(3) not_le_imp_less order.asym) thus ?thesis by simp qed lemma vec_of_list_coeffs_nth': assumes i: "i \ {..degree h}" and h_not0: "h \ 0" assumes "degree h < degree u" shows "vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) $ i = coeff h i" using assms by (transfer', auto simp add: mk_vec_def coeffs_nth length_coeffs_degree nth_append) lemma vec_of_list_coeffs_replicate_nth_0: assumes i: "i \ {.. {..{..degree h}") case True thus ?thesis using assms vec_of_list_coeffs_nth' h_not0 by simp next case False have c0: "coeff h i = 0" using False le_degree by auto thus ?thesis using assms False h_not0 by (transfer', auto simp add: mk_vec_def length_coeffs_degree nth_append c0) qed qed (* Equation 13 *) lemma equation_13: fixes u h defines H: "H \ vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0)" assumes deg_le: "degree h < degree u" (*Mandatory from equation 8*) shows "[h^CARD('a) = h] (mod u) \ (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H" (is "?lhs = ?rhs") proof - have f: "finite {..degree u}" by auto have [simp]: "dim_vec H = degree u" unfolding H using dim_vec_of_list_h deg_le by simp let ?B = "(berlekamp_mat u)" let ?f = "\i. (transpose_mat ?B *\<^sub>v H) $ i" show ?thesis proof assume rhs: ?rhs have dimv_h_dimr_B: "dim_vec H = dim_row ?B" by (metis berlekamp_mat_closed(2) berlekamp_mat_closed(3) dim_mult_mat_vec index_transpose_mat(2) rhs) have degree_h_less_dim_H: "degree h < dim_vec H" by (auto simp add: deg_le) have set_rw: "{..degree u - 1} = {.. degree u - 1" using deg_le by simp hence "h = (\j\degree u - 1. monom (coeff h j) j)" using poly_as_sum_of_monoms' by fastforce also have "... = (\jj {..j H) j)" by (rule sum.cong, auto) also have "... = (\ji = 0.. {.. H) x = monom (\i = 0..ji = 0..i = 0..ji = 0..ji = 0..ji = 0.. {0..jji = 0..i = 0..i = 0..x. x mod u"], rule sum.cong, rule) fix x assume x: "x \ {0.. {..i\degree h. monom (coeff h i) (CARD('a) * i)) mod u" proof (rule arg_cong[of _ _ "\x. x mod u"]) let ?f="\i. monom (coeff h i) (CARD('a) * i)" have ss0: "(\i = degree h + 1 ..< dim_vec H. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0..< dim_vec H} = {0..degree h} \ {degree h + 1 ..< dim_vec H}" using degree_h_less_dim_H by auto have "(\i = 0..i = 0..degree h. ?f i) + (\i = degree h + 1 ..< dim_vec H. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i = 0..degree h. ?f i)" unfolding ss0 by auto finally show "(\i = 0..i\degree h. ?f i)" by (simp add: atLeast0AtMost) qed also have "... = h^CARD('a) mod u" using poly_power_card_as_sum_of_monoms by auto finally show ?lhs unfolding cong_def using deg_le by (simp add: mod_poly_less) next assume lhs: ?lhs have deg_le': "degree h \ degree u - 1" using deg_le by auto have set_rw: "{..ii \ degree u - 1. monom (coeff h i) i)" by simp also have "... = (\i\degree h. monom (coeff h i) i)" unfolding poly_as_sum_of_monoms using poly_as_sum_of_monoms' deg_le' by auto also have "... = (\i\degree h. monom (coeff h i) i) mod u" by (simp add: deg_le mod_poly_less poly_as_sum_of_monoms) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i)) mod u" using lhs unfolding cong_def poly_as_sum_of_monoms poly_power_card_as_sum_of_monoms by auto also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i)) mod u" by (rule arg_cong[of _ _ "\x. x mod u"], rule sum.cong, simp_all add: mult_monom) also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i) mod u)" by (simp add: poly_mod_sum) also have "... = (\i\degree h. monom (coeff h i) 0 * (monom 1 (CARD('a)*i) mod u))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" have h_rw: "monom (coeff h x) 0 mod u = monom (coeff h x) 0" by (metis deg_le degree_pCons_eq_if gr_implies_not_zero linorder_neqE_nat mod_poly_less monom_0) have "monom (coeff h x) 0 * monom 1 (CARD('a) * x) = smult (coeff h x) (monom 1 (CARD('a) * x))" by (simp add: monom_0) also have "... mod u = Polynomial.smult (coeff h x) (monom 1 (CARD('a) * x) mod u)" using mod_smult_left by auto also have "... = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" by (simp add: monom_0) finally show "monom (coeff h x) 0 * monom 1 (CARD('a) * x) mod u = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" . qed also have "... = (\i\degree h. monom (coeff h i) 0 * (\j {..degree h}" have "(monom 1 (CARD('a) * x) mod u) = (\jjiji=degree h+1 ..< degree u. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0.. {degree h+1..i=0..i=0..degree h. ?f i) + (\i=degree h+1 ..< degree u. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i=0..degree h. ?f i)" using ss0 by simp finally show ?thesis by (simp add: atLeast0AtMost atLeast0LessThan) qed also have "... = (\ijijjijiijii. i < degree u \ coeff h i = (\ji. i < degree u \ H $ i = (\jjjjv H = H" proof (rule eq_vecI) fix i show "dim_vec (transpose_mat ?B *\<^sub>v H) = dim_vec (H)" by auto assume i: "i < dim_vec (H)" have "(transpose_mat ?B *\<^sub>v H) $ i = row (transpose_mat ?B) i \ H" using i by simp also have "... = (\j = 0..jv H) $ i = H $ i" . qed qed qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma exists_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\s. fi dvd (h - [:s:])" proof - let ?p = "CARD('a)" have f_dvd_hqh: "f dvd (h^?p - h)" using h unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis using irreducible_dvd_prod[OF _ fi_dvd_prod_hc] irr_fi by auto qed corollary exists_unique_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\!s. fi dvd (h - [:s:])" proof - obtain c where fi_dvd: "fi dvd (h - [:c:])" using assms exists_s_factor_dvd_h_s by blast have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" by (simp add: irr_fi irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis proof (rule ex1I[of _ c], auto simp add: fi_dvd) fix c2 assume fi_dvd_hc2: "fi dvd h - [:c2:]" have *: "fi dvd (h - [:c:]) * (h - [:c2:])" using fi_dvd by auto hence "fi dvd (h - [:c:]) \ fi dvd (h - [:c2:])" using irr_fi by auto thus "c2 = c" using coprime_h_c_poly coprime_not_unit_not_dvd fi_dvd fi_dvd_hc2 fi_not_unit by blast qed qed lemma exists_two_distint: "\a b::'a mod_ring. a \ b" by (rule exI[of _ 0], rule exI[of _ 1], auto) lemma coprime_cong_mult_factorization_poly: fixes f::"'b::{field} poly" and a b p :: "'c :: {field_gcd} poly" assumes finite_P: "finite P" and P: "P \ {q. irreducible q}" and p: "\p\P. [a=b] (mod p)" and coprime_P: "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" shows "[a = b] (mod (\a\P. a))" using finite_P P p coprime_P proof (induct P) case empty thus ?case by simp next case (insert p P) have ab_mod_pP: "[a=b] (mod (p * \P))" proof (rule coprime_cong_mult_poly) show "[a = b] (mod p)" using insert.prems by auto show "[a = b] (mod \P)" using insert.prems insert.hyps by auto from insert show "Rings.coprime p (\P)" by (auto intro: prod_coprime_right) qed thus ?case by (simp add: insert.hyps(1) insert.hyps(2)) qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma W_eq_berlekamp_mat: fixes u::"'a mod_ring poly" shows "{v. [v^CARD('a) = v] (mod u) \ degree v < degree u} = {h. let H = vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) in (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H \ degree h < degree u}" using equation_13 by (auto simp add: Let_def) lemma transpose_minus_1: assumes "dim_row Q = dim_col Q" shows "transpose_mat (Q - (1\<^sub>m (dim_row Q))) = (transpose_mat Q - (1\<^sub>m (dim_row Q)))" using assms unfolding mat_eq_iff by auto lemma system_iff: fixes v::"'b::comm_ring_1 vec" assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" proof - have t1:"transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v) \ (transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v)" by (subst minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have t2:"(transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v) \ transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v)" by (subst (asm) minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have "transpose_mat Q *\<^sub>v v - v = v - v \ transpose_mat Q *\<^sub>v v = v" proof - assume a1: "transpose_mat Q *\<^sub>v v - v = v - v" have f2: "transpose_mat Q *\<^sub>v v \ carrier_vec (dim_vec v)" by (metis dim_mult_mat_vec index_transpose_mat(2) sq_Q v carrier_vec_dim_vec) then have f3: "0\<^sub>v (dim_vec v) + transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v" by (meson left_zero_vec) have f4: "0\<^sub>v (dim_vec v) = transpose_mat Q *\<^sub>v v - v" using a1 by auto have f5: "- v \ carrier_vec (dim_vec v)" by simp then have f6: "- v + transpose_mat Q *\<^sub>v v = v - v" using f2 a1 using comm_add_vec minus_add_uminus_vec by fastforce have "v - v = - v + v" by auto then have "transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v - v + v" using f6 f4 f3 f2 by (metis (no_types, lifting) a1 assoc_add_vec comm_add_vec f5 carrier_vec_dim_vec) then show ?thesis using a1 by auto qed hence "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q *\<^sub>v v) - v = v - v)" by auto also have "... = ((transpose_mat Q *\<^sub>v v) - v = 0\<^sub>v (dim_vec v))" by auto also have "... = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using t1 t2 by auto finally show ?thesis. qed lemma system_if_mat_kernel: assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q))))" proof - have "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using assms system_iff by blast also have "... = (v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q)))))" unfolding mat_kernel_def unfolding transpose_minus_1[OF sq_Q] unfolding v by auto finally show ?thesis . qed lemma degree_u_mod_irreducible\<^sub>d_factor_0: fixes v and u::"'a mod_ring poly" defines W: "W \ {v. [v ^ CARD('a) = v] (mod u)}" assumes v: "v \ W" and finite_U: "finite U" and u_U: "u = \U" and U_irr_monic: "U \ {q. irreducible q \ monic q}" and fi_U: "fi \ U" shows "degree (v mod fi) = 0" proof - have deg_fi: "degree fi > 0" using U_irr_monic using fi_U irreducible\<^sub>dD[of fi] by auto have "fi dvd u" using u_U U_irr_monic finite_U dvd_prod_eqI fi_U by blast moreover have "u dvd (v^CARD('a) - v)" using v unfolding W cong_def by (simp add: mod_eq_dvd_iff_poly) ultimately have "fi dvd (v^CARD('a) - v)" by (rule dvd_trans) then have fi_dvd_prod_vc: "fi dvd prod (\c. v - [:c:]) (UNIV::'a mod_ring set)" by (simp add: poly_identity_mod_p) have irr_fi: "irreducible fi" using fi_U U_irr_monic by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (auto simp: poly_dvd_1) have fi_dvd_vc: "\c. fi dvd v - [:c:]" using irreducible_dvd_prod[OF _ fi_dvd_prod_vc] irr_fi by auto from this obtain a where "fi dvd v - [:a:]" by blast hence "v mod fi = [:a:] mod fi" using mod_eq_dvd_iff_poly by blast also have "... = [:a:]" by (simp add: deg_fi mod_poly_less) finally show ?thesis by simp qed (* Also polynomials over a field as a vector space in HOL-Algebra.*) definition "poly_abelian_monoid = \carrier = UNIV::'a mod_ring poly set, monoid.mult = ((*)), one = 1, zero = 0, add = (+), module.smult = smult\" interpretation vector_space_poly: vectorspace class_ring poly_abelian_monoid rewrites [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 0" and [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 1" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (+)" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (*)" and [simp]: "carrier poly_abelian_monoid = UNIV" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = smult" apply unfold_locales apply (auto simp: poly_abelian_monoid_def class_field_def smult_add_left smult_add_right Units_def) by (metis add.commute add.right_inverse) lemma subspace_Berlekamp: assumes f: "degree f \ 0" shows "subspace (class_ring :: 'a mod_ring ring) {v. [v^(CARD('a)) = v] (mod f) \ (degree v < degree f)} poly_abelian_monoid" proof - { fix v :: "'a mod_ring poly" and w :: "'a mod_ring poly" assume a1: "v ^ card (UNIV::'a set) mod f = v mod f" assume "w ^ card (UNIV::'a set) mod f = w mod f" then have "(v ^ card (UNIV::'a set) + w ^ card (UNIV::'a set)) mod f = (v + w) mod f" using a1 by (meson mod_add_cong) then have "(v + w) ^ card (UNIV::'a set) mod f = (v + w) mod f" by (simp add: add_power_poly_mod_ring) } note r=this thus ?thesis using f by (unfold_locales, auto simp: zero_power mod_smult_left smult_power cong_def degree_add_less) qed lemma berlekamp_resulting_mat_closed[simp]: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" proof - let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" let ?G="(gauss_jordan_single ?A)" have "?G \carrier_mat (degree u) (degree u)" by (rule gauss_jordan_single(2)[of ?A], auto) thus "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" unfolding berlekamp_resulting_mat_def Let_def by auto qed (*find_base vectors returns a basis:*) lemma berlekamp_resulting_mat_basis: "kernel.basis (degree u) (berlekamp_resulting_mat u) (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule find_base_vectors(3)) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" have "row_echelon_form (gauss_jordan_single ?A)" by (rule gauss_jordan_single(3)[of ?A], auto) thus "row_echelon_form (berlekamp_resulting_mat u)" unfolding berlekamp_resulting_mat_def Let_def by auto qed lemma set_berlekamp_basis_eq: "(set (berlekamp_basis u)) = (Poly \ list_of_vec)` (set (find_base_vectors (berlekamp_resulting_mat u)))" by (auto simp add: image_def o_def berlekamp_basis_def) lemma berlekamp_resulting_mat_constant: assumes deg_u: "degree u = 0" shows "berlekamp_resulting_mat u = 1\<^sub>m 0" by (unfold mat_eq_iff, auto simp add: deg_u) context fixes u::"'a::prime_card mod_ring poly" begin lemma set_berlekamp_basis_constant: assumes deg_u: "degree u = 0" shows "set (berlekamp_basis u) = {}" proof - have one_carrier: "1\<^sub>m 0 \ carrier_mat 0 0" by auto have m: "mat_kernel (1\<^sub>m 0) = {(0\<^sub>v 0) :: 'a mod_ring vec}" unfolding mat_kernel_def by auto have r: "row_echelon_form (1\<^sub>m 0 :: 'a mod_ring mat)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto have "set (find_base_vectors (1\<^sub>m 0)) \ {0\<^sub>v 0 :: 'a mod_ring vec}" using find_base_vectors(1)[OF r one_carrier] unfolding m . hence "set (find_base_vectors (1\<^sub>m 0) :: 'a mod_ring vec list) = {}" using find_base_vectors(2)[OF r one_carrier] using subset_singletonD by fastforce thus ?thesis unfolding set_berlekamp_basis_eq unfolding berlekamp_resulting_mat_constant[OF deg_u] by auto qed (*Maybe [simp]*) lemma row_echelon_form_berlekamp_resulting_mat: "row_echelon_form (berlekamp_resulting_mat u)" by (rule gauss_jordan_single(3), auto simp add: berlekamp_resulting_mat_def Let_def) lemma mat_kernel_berlekamp_resulting_mat_degree_0: assumes d: "degree u = 0" shows "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (auto simp add: mat_kernel_def mult_mat_vec_def d) lemma in_mat_kernel_berlekamp_resulting_mat: assumes x: "transpose_mat (berlekamp_mat u) *\<^sub>v x = x" and x_dim: "x \ carrier_vec (degree u)" shows "x \ mat_kernel (berlekamp_resulting_mat u)" proof - let ?QI="(mat(dim_row (berlekamp_mat u)) (dim_row (berlekamp_mat u)) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" have *: "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (degree u)) = transpose_mat ?QI" by auto have "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (dim_row (berlekamp_mat u))) *\<^sub>v x = 0\<^sub>v (dim_vec x)" using system_iff[of "berlekamp_mat u" x] x_dim x by auto hence "transpose_mat ?QI *\<^sub>v x = 0\<^sub>v (degree u)" using x_dim * by auto hence "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" unfolding berlekamp_resulting_mat_def Let_def using gauss_jordan_single(1)[of "transpose_mat ?QI" "degree u" "degree u" _ x] x_dim by auto thus ?thesis by (auto simp add: mat_kernel_def x_dim) qed private abbreviation "V \ kernel.VK (degree u) (berlekamp_resulting_mat u)" private abbreviation "W \ vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u) \ (degree v < degree u)}" interpretation V: vectorspace class_ring V proof - interpret k: kernel "(degree u)" "(degree u)" "(berlekamp_resulting_mat u)" by (unfold_locales; auto) show "vectorspace class_ring V" by intro_locales qed lemma linear_Poly_list_of_vec: shows "(Poly \ list_of_vec) \ module_hom class_ring V (vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u)})" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed qed lemma linear_Poly_list_of_vec': assumes "degree u > 0" shows "(Poly \ list_of_vec) \ module_hom R V W" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "degree (Poly (list_of_vec x)) < degree u" by (rule degree_Poly_list_of_vec, insert assms x, auto simp: mat_kernel_def) qed lemma berlekamp_basis_eq_8: assumes v: "v \ set (berlekamp_basis u)" shows "[v ^ CARD('a) = v] (mod u)" proof - { fix x assume x: "x \ set (find_base_vectors (berlekamp_resulting_mat u))" have "set (find_base_vectors (berlekamp_resulting_mat u)) \ mat_kernel (berlekamp_resulting_mat u)" proof (rule find_base_vectors(1)) show "row_echelon_form (berlekamp_resulting_mat u)" by (rule row_echelon_form_berlekamp_resulting_mat) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp qed hence "x \ mat_kernel (berlekamp_resulting_mat u)" using x by auto hence "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto } thus "[v ^ CARD('a) = v] (mod u)" using v unfolding set_berlekamp_basis_eq by auto qed lemma surj_Poly_list_of_vec: assumes deg_u: "degree u > 0" shows "(Poly \ list_of_vec)` (carrier V) = carrier W" proof (auto simp add: image_def) fix xa assume xa: "xa \ mat_kernel (berlekamp_resulting_mat u)" thus "[Poly (list_of_vec xa) ^ CARD('a) = Poly (list_of_vec xa)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto show "degree (Poly (list_of_vec xa)) < degree u" proof (rule degree_Poly_list_of_vec[OF _ deg_u]) show "xa \ carrier_vec (degree u)" using xa unfolding mat_kernel_def by simp qed next fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" show "\xa \ mat_kernel (berlekamp_resulting_mat u). x = Poly (list_of_vec xa)" proof (rule bexI[of _ "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)"]) let ?X = "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)" show "x = Poly (list_of_vec (vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)))" by auto have X: "?X \ carrier_vec (degree u)" unfolding carrier_vec_def by (auto, metis Suc_leI coeffs_0_eq_Nil deg_x degree_0 le_add_diff_inverse length_coeffs_degree linordered_semidom_class.add_diff_inverse list.size(3) order.asym) have t: "transpose_mat (berlekamp_mat u) *\<^sub>v ?X = ?X" using equation_13[OF deg_x] x by auto show "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0) \ mat_kernel (berlekamp_resulting_mat u)" by (rule in_mat_kernel_berlekamp_resulting_mat[OF t X]) qed qed lemma card_set_berlekamp_basis: "card (set (berlekamp_basis u)) = length (berlekamp_basis u)" proof - have b: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by auto have "(set (berlekamp_basis u)) = (Poly \ list_of_vec) ` set (find_base_vectors (berlekamp_resulting_mat u))" unfolding set_berlekamp_basis_eq .. also have " card ... = card (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule card_image, rule subset_inj_on[OF inj_Poly_list_of_vec]) show "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier_vec (degree u)" using find_base_vectors(1)[OF row_echelon_form_berlekamp_resulting_mat b] unfolding carrier_vec_def mat_kernel_def by auto qed also have "... = length (find_base_vectors (berlekamp_resulting_mat u))" by (rule length_find_base_vectors[symmetric, OF row_echelon_form_berlekamp_resulting_mat b]) finally show ?thesis unfolding berlekamp_basis_def by auto qed context assumes deg_u0[simp]: "degree u > 0" begin interpretation Berlekamp_subspace: vectorspace class_ring W by (rule vector_space_poly.subspace_is_vs[OF subspace_Berlekamp], simp) lemma linear_map_Poly_list_of_vec': "linear_map class_ring V W (Poly \ list_of_vec)" proof (auto simp add: linear_map_def) show "vectorspace class_ring V" by intro_locales show "vectorspace class_ring W" by (rule Berlekamp_subspace.vectorspace_axioms) show "mod_hom class_ring V W (Poly \ list_of_vec)" proof (rule mod_hom.intro, unfold mod_hom_axioms_def) show "module class_ring V" by intro_locales show "module class_ring W" using Berlekamp_subspace.vectorspace_axioms by intro_locales show "Poly \ list_of_vec \ module_hom class_ring V W" by (rule linear_Poly_list_of_vec'[OF deg_u0]) qed qed lemma berlekamp_basis_basis: "Berlekamp_subspace.basis (set (berlekamp_basis u))" proof (unfold set_berlekamp_basis_eq, rule linear_map.linear_inj_image_is_basis) show "linear_map class_ring V W (Poly \ list_of_vec)" by (rule linear_map_Poly_list_of_vec') show "inj_on (Poly \ list_of_vec) (carrier V)" proof (rule subset_inj_on[OF inj_Poly_list_of_vec]) show "carrier V \ carrier_vec (degree u)" by (auto simp add: mat_kernel_def) qed show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show b: "V.basis (set (find_base_vectors (berlekamp_resulting_mat u)))" by (rule berlekamp_resulting_mat_basis) show "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using b unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed qed lemma finsum_sum: fixes f::"'a mod_ring poly" assumes f: "finite B" and a_Pi: "a \ B \ carrier R" and V: "B \ carrier W" shows "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" using f a_Pi V proof (induct B) case empty thus ?case unfolding Berlekamp_subspace.module.M.finsum_empty by auto next case (insert x V) have hyp: "(\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) V" proof (rule insert.hyps) show "a \ V \ carrier R" using insert.prems unfolding class_field_def by auto show "V \ carrier W" using insert.prems by simp qed have "(\\<^bsub>W\<^esub>v\insert x V. a v \\<^bsub>W\<^esub> v) = (a x \\<^bsub>W\<^esub> x) \\<^bsub>W\<^esub> (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" proof (rule abelian_monoid.finsum_insert) show "abelian_monoid W" by (unfold_locales) show "finite V" by fact show "x \ V" by fact show "(\v. a v \\<^bsub>W\<^esub> v) \ V \ carrier W" proof (unfold Pi_def, rule, rule allI, rule impI) fix v assume v: "v\V" show "a v \\<^bsub>W\<^esub> v \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a v \ carrier class_ring" using insert.prems v unfolding Pi_def by (simp add: class_field_def) show "v \ carrier W" using v insert.prems by auto qed qed show "a x \\<^bsub>W\<^esub> x \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a x \ carrier class_ring" using insert.prems unfolding Pi_def by (simp add: class_field_def) show "x \ carrier W" using insert.prems by auto qed qed also have "... = (a x \\<^bsub>W\<^esub> x) + (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" by auto also have "... = (a x \\<^bsub>W\<^esub> x) + sum (\v. smult (a v) v) V" unfolding hyp by simp also have "... = (smult (a x) x) + sum (\v. smult (a v) v) V" by simp also have "... = sum (\v. smult (a v) v) (insert x V)" by (simp add: insert.hyps(1) insert.hyps(2)) finally show ?case . qed lemma exists_vector_in_Berlekamp_subspace_dvd: fixes p_i::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v. v \ {h. [h^(CARD('a)) = h] (mod u) \ degree h < degree u} \ v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1)" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by auto have irr_pj: "irreducible p_j" using pj P by auto obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m f_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m f_desc_square_free f_not_0 by auto obtain i where mi: "m i = p_i" and i: "i < n" using P_m pi by blast obtain j where mj: "m j = p_j" and j: "j < n" using P_m pj by blast have ij: "i \ j" using mi mj pi_pj by auto obtain s_i and s_j::"'a mod_ring" where si_sj: "s_i \ s_j" using exists_two_distint by blast let ?u="\x. if x = i then [:s_i:] else if x = j then [:s_j:] else [:0:]" have degree_si: "degree [:s_i:] = 0" by auto have degree_sj: "degree [:s_j:] = 0" by auto have "\!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\a\{i. i < n}. [v = ?u a] (mod m a))" proof (rule chinese_remainder_unique_poly) show "\a\{i. i < n}. \b\{i. i < n}. a \ b \ Rings.coprime (m a) (m b)" proof (rule+) fix a b assume "a \ {i. i < n}" and "b \ {i. i < n}" and "a \ b" thus "Rings.coprime (m a) (m b)" using coprime_polynomial_factorization[OF P finite_P, simplified] P_m by (metis image_eqI inj_onD inj_on_m) qed show "\i\{i. i < n}. m i \ 0" by (rule not_zero) show "0 < degree (prod m {i. i < n})" unfolding degree_prod using deg_u0 by blast qed from this obtain v where v: "\a\{i. i < n}. [v = ?u a] (mod m a)" and degree_v: "degree v < (\i\{i. i < n}. degree (m i))" by blast show ?thesis proof (rule exI[of _ v], auto) show vp_v_mod: "[v ^ CARD('a) = v] (mod u)" proof (unfold f_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P]) show "P \ {q. irreducible q}" using P by blast show "\p\P. [v ^ CARD('a) = v] (mod p)" proof (rule ballI) fix p assume p: "p \ P" hence irr_p: "irreducible\<^sub>d p" using P by auto obtain k where mk: "m k = p" and k: "k < n" using P_m p by blast have "[v = ?u k] (mod p)" using v mk k by auto moreover have "?u k mod p = ?u k" apply (rule mod_poly_less) using irreducible\<^sub>dD(1)[OF irr_p] by auto ultimately obtain s where v_mod_p: "v mod p = [:s:]" unfolding cong_def by force hence deg_v_p: "degree (v mod p) = 0" by auto have "v mod p = [:s:]" by (rule v_mod_p) also have "... = [:s:]^CARD('a)" unfolding poly_const_pow by auto also have "... = (v mod p) ^ CARD('a)" using v_mod_p by auto also have "... = (v mod p) ^ CARD('a) mod p" using calculation by auto also have "... = v^CARD('a) mod p" using power_mod by blast finally show "[v ^ CARD('a) = v] (mod p)" unfolding cong_def .. qed show "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" using P coprime_polynomial_factorization finite_P by auto qed have "[v = ?u i] (mod m i)" using v i by auto hence v_pi_si_mod: "v mod p_i = [:s_i:] mod p_i" unfolding cong_def mi by auto also have "... = [:s_i:]" apply (rule mod_poly_less) using irr_pi by auto finally have v_pi_si: "v mod p_i = [:s_i:]" . have "[v = ?u j] (mod m j)" using v j by auto hence v_pj_sj_mod: "v mod p_j = [:s_j:] mod p_j" unfolding cong_def mj using ij by auto also have "... = [:s_j:]" apply (rule mod_poly_less) using irr_pj by auto finally have v_pj_sj: "v mod p_j = [:s_j:]" . show "v mod p_i = v mod p_j \ False" using si_sj v_pi_si v_pj_sj by auto show "degree (v mod p_i) = 0" unfolding v_pi_si by simp show "degree (v mod p_j) = 0" unfolding v_pj_sj by simp show "\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" using v_pi_si_mod mod_eq_dvd_iff_poly by blast have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" using v_pj_sj_mod mod_eq_dvd_iff_poly by blast have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using vp_v_mod cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" proof (rule ccontr, simp) assume gcd_w: "gcd w (v - [:s_i:]) = w" show False apply (rule \v mod p_i = v mod p_j \ False\) by (metis irreducibleE \degree (v mod p_i) = 0\ gcd_greatest_iff gcd_w irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_si) qed show "gcd w (v - [:s_i:]) \ 1" by (metis irreducibleE gcd_greatest_iff irr_pi pi_dvd_v_si pi_dvd_w) qed show "degree v < degree u" proof - have "(\i | i < n. degree (m i)) = degree (prod m {i. i < n})" by (rule degree_prod_eq_sum_degree[symmetric, OF not_zero]) thus ?thesis using degree_v unfolding degree_prod by auto qed qed qed lemma exists_vector_in_Berlekamp_basis_dvd_aux: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j" proof (rule ccontr, auto) have V_in_carrier: "B \ carrier W" using basis_V unfolding Berlekamp_subspace.basis_def by auto assume all_eq: "\v\B. v mod p_i = v mod p_j" obtain x where x: "x \ {h. [h ^ CARD('a) = h] (mod u) \ degree h < degree u}" and x_pi_pj: "x mod p_i \ x mod p_j" and "degree (x mod p_i) = 0" and "degree (x mod p_j) = 0" "(\s. gcd w (x - [:s:]) \ w \ gcd w (x - [:s:]) \ 1)" using exists_vector_in_Berlekamp_subspace_dvd[OF _ _ _ pi pj _ _ _ _ w_dvd_f monic_w pi_dvd_w] assms by meson have x_in: "x \ carrier W" using x by auto hence "(\!a. a \ B \\<^sub>E carrier class_ring \ Berlekamp_subspace.lincomb a B = x)" using Berlekamp_subspace.basis_criterion[OF finite_V V_in_carrier] using basis_V by (simp add: class_field_def) from this obtain a where a_Pi: "a \ B \\<^sub>E carrier class_ring" and lincomb_x: "Berlekamp_subspace.lincomb a B = x" by blast have fs_ss: "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" proof (rule finsum_sum) show "finite B" by fact show "a \ B \ carrier class_ring" using a_Pi by auto show "B \ carrier W" by (rule V_in_carrier) qed have "x mod p_i = Berlekamp_subspace.lincomb a B mod p_i" using lincomb_x by simp also have " ... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_i" unfolding Berlekamp_subspace.lincomb_def .. also have "... = (sum (\v. smult (a v) v) B) mod p_i" unfolding fs_ss .. also have "... = sum (\v. smult (a v) v mod p_i) B" using finite_V poly_mod_sum by blast also have "... = sum (\v. smult (a v) (v mod p_i)) B" by (meson mod_smult_left) also have "... = sum (\v. smult (a v) (v mod p_j)) B" using all_eq by auto also have "... = sum (\v. smult (a v) v mod p_j) B" by (metis mod_smult_left) also have "... = (sum (\v. smult (a v) v) B) mod p_j" by (metis (mono_tags, lifting) finite_V poly_mod_sum sum.cong) also have "... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_j" unfolding fs_ss .. also have "... = Berlekamp_subspace.lincomb a B mod p_j" unfolding Berlekamp_subspace.lincomb_def .. also have "... = x mod p_j" using lincomb_x by simp finally have "x mod p_i = x mod p_j" . thus False using x_pi_pj by contradiction qed lemma exists_vector_in_Berlekamp_basis_dvd: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ \ coprime w (v - [:s:]))" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by fast have irr_pj: "irreducible p_j" using pj P by fast obtain v where vV: "v \ B" and v_pi_pj: "v mod p_i \ v mod p_j" using assms exists_vector_in_Berlekamp_basis_dvd_aux by blast have v: "v \ {v. [v ^ CARD('a) = v] (mod u)}" using basis_V vV unfolding Berlekamp_subspace.basis_def by auto have deg_v_pi: "degree (v mod p_i) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pi]) from this obtain s_i where v_pi_si: "v mod p_i = [:s_i:]" using degree_eq_zeroE by blast have deg_v_pj: "degree (v mod p_j) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pj]) from this obtain s_j where v_pj_sj: "v mod p_j = [:s_j:]" using degree_eq_zeroE by blast have si_sj: "s_i \ s_j" using v_pi_si v_pj_sj v_pi_pj by auto have "(\s. gcd w (v - [:s:]) \ w \ \ Rings.coprime w (v - [:s:]))" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pi_si) have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pj_sj) have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using v cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" by (metis irreducibleE deg_v_pi gcd_greatest_iff irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_pj v_pi_si) show "\ Rings.coprime w (v - [:s_i:])" using irr_pi pi_dvd_v_si pi_dvd_w by (simp add: irreducible\<^sub>dD(1) not_coprimeI) qed thus ?thesis using v_pi_pj vV deg_v_pi deg_v_pj by auto qed lemma exists_bijective_linear_map_W_vec: assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" shows "\f. linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f \ bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" proof - let ?B="carrier_vec (card P)::'a mod_ring vec set" have u_not_0: "u \ 0" using deg_u0 degree_0 by force obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence n: "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m u_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m u_desc_square_free u_not_0 by auto have deg_sum_eq: "(\i\{i. i < n}. degree (m i)) = degree u" by (metis degree_prod degree_prod_eq_sum_degree not_zero) have coprime_mi_mj:"\i\{i. i < n}. \j\{i. i < n}. i \ j \ coprime (m i) (m j)" proof (rule+) fix i j assume i: "i \ {i. i < n}" and j: "j \ {i. i < n}" and ij: "i \ j" show "coprime (m i) (m j)" proof (rule coprime_polynomial_factorization[OF P finite_P]) show "m i \ P" using i P_m by auto show "m j \ P" using j P_m by auto show "m i \ m j" using inj_on_m i ij j unfolding inj_on_def by blast qed qed let ?f = "\v. vec n (\i. coeff (v mod (m i)) 0)" interpret vec_VS: vectorspace class_ring "(module_vec TYPE('a mod_ring) n)" by (rule VS_Connect.vec_vs) interpret linear_map class_ring W "(module_vec TYPE('a mod_ring) n)" ?f by (intro_locales, unfold mod_hom_axioms_def LinearCombinations.module_hom_def, auto simp add: vec_eq_iff module_vec_def mod_smult_left poly_mod_add_left) have "linear_map class_ring W (module_vec TYPE('a mod_ring) n) ?f" by (intro_locales) moreover have inj_f: "inj_on ?f (carrier W)" proof (rule Ke0_imp_inj, auto simp add: mod_hom.ker_def) show "[0 ^ CARD('a) = 0] (mod u)" by (simp add: cong_def zero_power) show "vec n (\i. 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" by (auto simp add: module_vec_def) fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" and v: "vec n (\i. coeff (x mod m i) 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" have cong_0: "\i\{i. i < n}. [x = (\i. 0) i] (mod m i)" proof (rule, unfold cong_def) fix i assume i: "i \ {i. i < n}" have deg_x_mod_mi: "degree (x mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "x \ {v. [v ^ CARD('a) = v] (mod u)}" using x by auto show "m i \ P" using P_m i by auto qed thus "x mod m i = 0 mod m i" using v unfolding module_vec_def by (auto, metis i leading_coeff_neq_0 mem_Collect_eq index_vec index_zero_vec(1)) qed moreover have deg_x2: "degree x < (\i\{i. i < n}. degree (m i))" using deg_sum_eq deg_x by simp moreover have "\i\{i. i < n}. [0 = (\i. 0) i] (mod m i)" by (auto simp add: cong_def) moreover have "degree 0 < (\i\{i. i < n}. degree (m i))" using degree_prod deg_sum_eq deg_u0 by force moreover have "\!x. degree x < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [x = (\i. 0) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) ultimately show "x = 0" by blast qed moreover have "?f ` (carrier W) = ?B" proof (auto simp add: image_def) fix xa show "n = card P" by (auto simp add: n) next fix x::"'a mod_ring vec" assume x: "x \ carrier_vec (card P)" have " \!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) from this obtain v where deg_v: "degree v < (\i\{i. i < n}. degree (m i))" and v_x_cong: "(\i \ {i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" by auto show "\xa. [xa ^ CARD('a) = xa] (mod u) \ degree xa < degree u \ x = vec n (\i. coeff (xa mod m i) 0)" proof (rule exI[of _ v], auto) show v: "[v ^ CARD('a) = v] (mod u)" proof (unfold u_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P], auto) fix y assume y: "y \ P" thus "irreducible y" using P by blast obtain i where i: "i \ {i. i < n}" and mi: "y = m i" using P_m y by blast have "irreducible (m i)" using i P_m P by auto moreover have "[v = [:x $ i:]] (mod m i)" using v_x_cong i by auto ultimately have v_mi_eq_xi: "v mod m i = [:x $ i:]" by (auto simp: cong_def intro!: mod_poly_less) have xi_pow_xi: "[:x $ i:]^CARD('a) = [:x $ i:]" by (simp add: poly_const_pow) hence "(v mod m i)^CARD('a) = v mod m i" using v_mi_eq_xi by auto hence "(v mod m i)^CARD('a) = (v^CARD('a) mod m i)" by (metis mod_mod_trivial power_mod) thus "[v ^ CARD('a) = v] (mod y)" unfolding mi cong_def v_mi_eq_xi xi_pow_xi by simp next fix p1 p2 assume "p1 \ P" and "p2 \ P" and "p1 \ p2" then show "Rings.coprime p1 p2" using coprime_polynomial_factorization[OF P finite_P] by auto qed show "degree v < degree u" using deg_v deg_sum_eq degree_prod by presburger show "x = vec n (\i. coeff (v mod m i) 0)" proof (unfold vec_eq_iff, rule conjI) show "dim_vec x = dim_vec (vec n (\i. coeff (v mod m i) 0))" using x n by simp show "\ii. coeff (v mod m i) 0)). x $ i = vec n (\i. coeff (v mod m i) 0) $ i" proof (auto) fix i assume i: "i < n" have deg_mi: "irreducible (m i)" using i P_m P by auto have deg_v_mi: "degree (v mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "v \ {v. [v ^ CARD('a) = v] (mod u)}" using v by fast show "m i \ P" using P_m i by auto qed have "v mod m i = [:x $ i:] mod m i" using v_x_cong i unfolding cong_def by auto also have "... = [:x $ i:]" using deg_mi by (auto intro!: mod_poly_less) finally show "x $ i = coeff (v mod m i) 0" by simp qed qed qed qed ultimately show ?thesis unfolding bij_betw_def n by auto qed lemma fin_dim_kernel_berlekamp: "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using berlekamp_resulting_mat_basis[of u] unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed lemma Berlekamp_subspace_fin_dim: "Berlekamp_subspace.fin_dim" proof (rule linear_map.surj_fin_dim[OF linear_map_Poly_list_of_vec']) show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show "V.fin_dim" by (rule fin_dim_kernel_berlekamp) qed context fixes P assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" begin interpretation RV: vec_space "TYPE('a mod_ring)" "card P" . lemma Berlekamp_subspace_eq_dim_vec: "Berlekamp_subspace.dim = RV.dim" proof - obtain f where lm_f: "linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f" and bij_f: "bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" using exists_bijective_linear_map_W_vec[OF finite_P u_desc_square_free P] by blast show ?thesis proof (rule linear_map.dim_eq[OF lm_f Berlekamp_subspace_fin_dim]) show "inj_on f (carrier W)" by (rule bij_betw_imp_inj_on[OF bij_f]) show " f ` carrier W = carrier RV.V" using bij_f unfolding bij_betw_def by auto qed qed lemma Berlekamp_subspace_dim: "Berlekamp_subspace.dim = card P" using Berlekamp_subspace_eq_dim_vec RV.dim_is_n by simp corollary card_berlekamp_basis_number_factors: "card (set (berlekamp_basis u)) = card P" unfolding Berlekamp_subspace_dim[symmetric] by (rule Berlekamp_subspace.dim_basis[symmetric], auto simp add: berlekamp_basis_basis) lemma length_berlekamp_basis_numbers_factors: "length (berlekamp_basis u) = card P" using card_set_berlekamp_basis card_berlekamp_basis_number_factors by auto end end end end context assumes "SORT_CONSTRAINT('a :: prime_card)" begin context fixes f :: "'a mod_ring poly" and n assumes sf: "square_free f" and n: "n = length (berlekamp_basis f)" and monic_f: "monic f" begin lemma berlekamp_basis_length_factorization: assumes f: "f = prod_list us" and d: "\ u. u \ set us \ degree u > 0" shows "length us \ n" proof (cases "degree f = 0") case True have "us = []" proof (rule ccontr) assume "us \ []" from this obtain u where u: "u \ set us" by fastforce hence deg_u: "degree u > 0" using d by auto have "degree f = degree (prod_list us)" unfolding f .. also have "... = sum_list (map degree us)" proof (rule degree_prod_list_eq) fix p assume p: "p \ set us" show "p \ 0" using d[OF p] degree_0 by auto qed also have " ... \ degree u" by (simp add: member_le_sum_list u) finally have "degree f > 0" using deg_u by auto thus False using True by auto qed thus ?thesis by simp next case False hence f_not_0: "f \ 0" using degree_0 by fastforce obtain P where fin_P: "finite P" and f_P: "f = \P" and P: "P \ {p. irreducible p \ monic p}" using monic_square_free_irreducible_factorization[OF monic_f sf] by auto have n_card_P: "n = card P" using P False f_P fin_P length_berlekamp_basis_numbers_factors n by blast have distinct_us: "distinct us" using d f sf square_free_prod_list_distinct by blast let ?us'="(map normalize us)" have distinct_us': "distinct ?us'" proof (auto simp add: distinct_map) show "distinct us" by (rule distinct_us) show "inj_on normalize (set us)" proof (auto simp add: inj_on_def, rule ccontr) fix x y assume x: "x \ set us" and y: "y \ set us" and n: "normalize x = normalize y" and x_not_y: "x \ y" from normalize_eq_imp_smult[OF n] obtain c where c0: "c \ 0" and y_smult: "y = smult c x" by blast have sf_xy: "square_free (x*y)" proof (rule square_free_factor[OF _ sf]) have "x*y = prod_list [x,y]" by simp also have "... dvd prod_list us" by (rule prod_list_dvd_prod_list_subset, auto simp add: x y x_not_y distinct_us) also have "... = f" unfolding f .. finally show "x * y dvd f" . qed have "x * y = smult c (x*x)" using y_smult mult_smult_right by auto hence sf_smult: "square_free (smult c (x*x))" using sf_xy by auto have "x*x dvd (smult c (x*x))" by (simp add: dvd_smult) hence "\ square_free (smult c (x*x))" by (metis d square_free_def x) thus "False" using sf_smult by contradiction qed qed have length_us_us': "length us = length ?us'" by simp have f_us': "f = prod_list ?us'" proof - have "f = normalize f" using monic_f f_not_0 by (simp add: normalize_monic) also have "... = prod_list ?us'" by (unfold f, rule prod_list_normalize[of us]) finally show ?thesis . qed have "\Q. prod_list Q = prod_list ?us' \ length ?us' \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule exists_factorization_prod_list) show "degree (prod_list ?us') > 0" using False f_us' by auto show "square_free (prod_list ?us')" using f_us' sf by auto fix u assume u: "u \ set ?us'" have u_not0: "u \ 0" using d u degree_0 by fastforce have "degree u > 0" using d u by auto moreover have "monic u" using u monic_normalize[OF u_not0] by auto ultimately show "degree u > 0 \ monic u" by simp qed from this obtain Q where Q_us': "prod_list Q = prod_list ?us'" and length_us'_Q: "length ?us' \ length Q" and Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast have distinct_Q: "distinct Q" proof (rule square_free_prod_list_distinct) show "square_free (prod_list Q)" using Q_us' f_us' sf by auto show "\u. u \ set Q \ degree u > 0" using Q irreducible_degree_field by auto qed have set_Q_P: "set Q = P" proof (rule monic_factorization_uniqueness) show "\(set Q) = \P" using Q_us' by (metis distinct_Q f_P f_us' list.map_ident prod.distinct_set_conv_list) qed (insert P Q fin_P, auto) hence "length Q = card P" using distinct_Q distinct_card by fastforce have "length us = length ?us'" by (rule length_us_us') also have "... \ length Q" using length_us'_Q by auto also have "... = card (set Q)" using distinct_card[OF distinct_Q] by simp also have "... = card P" using set_Q_P by simp finally show ?thesis using n_card_P by simp qed lemma berlekamp_basis_irreducible: assumes f: "f = prod_list us" and n_us: "length us = n" and us: "\ u. u \ set us \ degree u > 0" and u: "u \ set us" shows "irreducible u" proof (fold irreducible_connect_field, intro irreducible\<^sub>dI[OF us[OF u]]) fix q r :: "'a mod_ring poly" assume dq: "degree q > 0" and qu: "degree q < degree u" and dr: "degree r > 0" and uqr: "u = q * r" with us[OF u] have q: "q \ 0" and r: "r \ 0" by auto from split_list[OF u] obtain xs ys where id: "us = xs @ u # ys" by auto let ?us = "xs @ q # r # ys" have f: "f = prod_list ?us" unfolding f id uqr by simp { fix x assume "x \ set ?us" with us[unfolded id] dr dq have "degree x > 0" by auto } from berlekamp_basis_length_factorization[OF f this] have "length ?us \ n" by simp also have "\ = length us" unfolding n_us by simp also have "\ < length ?us" unfolding id by simp finally show False by simp qed end lemma not_irreducible_factor_yields_prime_factors: assumes uf: "u dvd (f :: 'b :: {field_gcd} poly)" and fin: "finite P" and fP: "f = \P" and P: "P \ {q. irreducible q \ monic q}" and u: "degree u > 0" "\ irreducible u" shows "\ pi pj. pi \ P \ pj \ P \ pi \ pj \ pi dvd u \ pj dvd u" proof - from finite_distinct_list[OF fin] obtain ps where Pps: "P = set ps" and dist: "distinct ps" by auto have fP: "f = prod_list ps" unfolding fP Pps using dist by (simp add: prod.distinct_set_conv_list) note P = P[unfolded Pps] have "set ps \ P" unfolding Pps by auto from uf[unfolded fP] P dist this show ?thesis proof (induct ps) case Nil with u show ?case using divides_degree[of u 1] by auto next case (Cons p ps) from Cons(3) have ps: "set ps \ {q. irreducible q \ monic q}" by auto from Cons(2) have dvd: "u dvd p * prod_list ps" by simp obtain k where gcd: "u = gcd p u * k" by (meson dvd_def gcd_dvd2) from Cons(3) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of u] *(2) have "gcd p u = 1 \ gcd p u = p" by auto thus ?case proof assume "gcd p u = 1" then have "Rings.coprime p u" by (rule gcd_eq_1_imp_coprime) with dvd have "u dvd prod_list ps" using coprime_dvd_mult_right_iff coprime_imp_coprime by blast from Cons(1)[OF this ps] Cons(4-5) show ?thesis by auto next assume "gcd p u = p" with gcd have upk: "u = p * k" by auto hence p: "p dvd u" by auto from dvd[unfolded upk] *(3) have kps: "k dvd prod_list ps" by auto from dvd u * have dk: "degree k > 0" by (metis gr0I irreducible_mult_unit_right is_unit_iff_degree mult_zero_right upk) from ps kps have "\ q \ set ps. q dvd k" proof (induct ps) case Nil with dk show ?case using divides_degree[of k 1] by auto next case (Cons p ps) from Cons(3) have dvd: "k dvd p * prod_list ps" by simp obtain l where gcd: "k = gcd p k * l" by (meson dvd_def gcd_dvd2) from Cons(2) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of k] *(2) have "gcd p k = 1 \ gcd p k = p" by auto thus ?case proof assume "gcd p k = 1" with dvd have "k dvd prod_list ps" by (metis dvd_triv_left gcd_greatest_mult mult.left_neutral) from Cons(1)[OF _ this] Cons(2) show ?thesis by auto next assume "gcd p k = p" with gcd have upk: "k = p * l" by auto hence p: "p dvd k" by auto thus ?thesis by auto qed qed then obtain q where q: "q \ set ps" and dvd: "q dvd k" by auto from dvd upk have qu: "q dvd u" by auto from Cons(4) q have "p \ q" by auto thus ?thesis using q p qu Cons(5) by auto qed qed qed lemma berlekamp_factorization_main: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and vs: "vs = vs1 @ vs2" and vsf: "vs = berlekamp_basis f" and n_bb: "n = length (berlekamp_basis f)" and n: "n = length us1 + n2" and us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" and us1: "\ u. u \ set us1 \ monic u \ irreducible u" and divs: "\ d. d \ set divs \ monic d \ degree d > 0" and vs1: "\ u v i. v \ set vs1 \ u \ set us1 \ set divs \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1,u}" and f: "f = prod_list (us1 @ divs)" and deg_f: "degree f > 0" and d: "\ g. g dvd f \ degree g = d \ irreducible g" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - have mon_f: "monic f" unfolding f by (rule monic_prod_list, insert divs us1, auto) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \ P" "P \ {q. irreducible q \ monic q}" by auto hence f0: "f \ 0" by auto show ?thesis using vs n us divs f us1 vs1 proof (induct vs2 arbitrary: divs n2 us1 vs1) case (Cons v vs2) show ?case proof (cases "v = 1") case False from Cons(2) vsf have v: "v \ set (berlekamp_basis f)" by auto from berlekamp_basis_eq_8[OF this] have vf: "[v ^ CARD('a) = v] (mod f)" . let ?gcd = "\ u i. gcd u (v - [:of_int i:])" let ?gcdn = "\ u i. gcd u (v - [:of_nat i:])" let ?map = "\ u. (map (\ i. ?gcd u i) [0 ..< CARD('a)])" define udivs where "udivs \ \ u. filter (\ w. w \ 1) (?map u)" { obtain xs where xs: "[0.. u. [w. i \ [0 ..< CARD('a)], w \ [?gcd u i], w \ 1])" unfolding udivs_def xs by (intro ext, auto simp: o_def, induct xs, auto) } note udivs_def' = this define facts where "facts \ [ w . u \ divs, w \ udivs u]" { fix u assume u: "u \ set divs" then obtain bef aft where divs: "divs = bef @ u # aft" by (meson split_list) from Cons(5)[OF u] have mon_u: "monic u" by simp have uf: "u dvd f" unfolding Cons(6) divs by auto from vf uf have vu: "[v ^ CARD('a) = v] (mod u)" by (rule cong_dvd_modulus_poly) from square_free_factor[OF uf sf_f] have sf_u: "square_free u" . let ?g = "?gcd u" from mon_u have u0: "u \ 0" by auto have "u = (\c\UNIV. gcd u (v - [:c:]))" using Berlekamp_gcd_step[OF vu mon_u sf_u] . also have "\ = (\i \ {0..< int CARD('a)}. ?g i)" by (rule sym, rule prod.reindex_cong[OF to_int_mod_ring_hom.inj_f range_to_int_mod_ring[symmetric]], simp add: of_int_of_int_mod_ring) finally have u_prod: "u = (\i \ {0..< int CARD('a)}. ?g i)" . let ?S = "{0.. ?S" hence "?g i \ 1" by auto moreover have mgi: "monic (?g i)" by (rule poly_gcd_monic, insert u0, auto) ultimately have "degree (?g i) > 0" using monic_degree_0 by blast note this mgi } note gS = this have int_set: "int ` set [0.. ?S" and j: "j \ ?S" and gij: "?g i = ?g j" show "i = j" proof (rule ccontr) define S where "S = {0.. S" "j \ S" "finite S" using i j unfolding S_def by auto assume ij: "i \ j" have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = ?g i * ?g j * (\i \ S. ?g i)" unfolding id using S ij by auto also have "\ = ?g i * ?g i * (\i \ S. ?g i)" unfolding gij by simp finally have dvd: "?g i * ?g i dvd u" unfolding dvd_def by auto with sf_u[unfolded square_free_def, THEN conjunct2, rule_format, OF gS(1)[OF i]] show False by simp qed qed have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = (\i \ ?S. ?g i)" by (rule sym, rule prod.setdiff_irrelevant, auto) also have "\ = \ (set (udivs u))" unfolding udivs_def set_filter set_map by (rule sym, rule prod.reindex_cong[of ?g, OF inj _ refl], auto simp: int_set[symmetric]) finally have u_udivs: "u = \(set (udivs u))" . { fix w assume mem: "w \ set (udivs u)" then obtain i where w: "w = ?g i" and i: "i \ ?S" unfolding udivs_def set_filter set_map int_set by auto have wu: "w dvd u" by (simp add: w) let ?v = "\ j. v - [:of_nat j:]" define j where "j = nat i" from i have j: "of_int i = (of_nat j :: 'a mod_ring)" "j < CARD('a)" unfolding j_def by auto from gS[OF i, folded w] have *: "degree w > 0" "monic w" "w \ 0" by auto from w have "w dvd ?v j" using j by simp hence gcdj: "?gcdn w j = w" by (metis gcd.commute gcd_left_idem j(1) w) { fix j' assume j': "j' < CARD('a)" have "?gcdn w j' \ {1,w}" proof (rule ccontr) assume not: "?gcdn w j' \ {1,w}" with gcdj have neq: "int j' \ int j" by auto (* next step will yield contradiction to square_free u *) let ?h = "?gcdn w j'" from *(3) not have deg: "degree ?h > 0" using monic_degree_0 poly_gcd_monic by auto have hw: "?h dvd w" by auto have "?h dvd ?gcdn u j'" using wu using dvd_trans by auto also have "?gcdn u j' = ?g j'" by simp finally have hj': "?h dvd ?g j'" by auto from divides_degree[OF this] deg u0 have degj': "degree (?g j') > 0" by auto hence j'1: "?g j' \ 1" by auto with j' have mem': "?g j' \ set (udivs u)" unfolding udivs_def by auto from degj' j' have j'S: "int j' \ ?S" by auto from i j have jS: "int j \ ?S" by auto from inj_on_contraD[OF inj neq j'S jS] have neq: "w \ ?g j'" using w j by auto have cop: "\ coprime w (?g j')" using hj' hw deg by (metis coprime_not_unit_not_dvd poly_dvd_1 Nat.neq0_conv) obtain w' where w': "?g j' = w'" by auto from u_udivs sf_u have "square_free (\(set (udivs u)))" by simp from square_free_prodD[OF this finite_set mem mem'] cop neq show False by simp qed } from gS[OF i, folded w] i this have "degree w > 0" "monic w" "\ j. j < CARD('a) \ ?gcdn w j \ {1,w}" by auto } note udivs = this let ?is = "filter (\ i. ?g i \ 1) (map int [0 ..< CARD('a)])" have id: "udivs u = map ?g ?is" unfolding udivs_def filter_map o_def .. have dist: "distinct (udivs u)" unfolding id distinct_map proof (rule conjI[OF distinct_filter], unfold distinct_map) have "?S = set ?is" unfolding int_set[symmetric] by auto thus "inj_on ?g (set ?is)" using inj by auto qed (auto simp: inj_on_def) from u_udivs prod.distinct_set_conv_list[OF dist, of id] have "prod_list (udivs u) = u" by auto note udivs this dist } note udivs = this have facts: "facts = concat (map udivs divs)" unfolding facts_def by auto obtain lin nonlin where part: "List.partition (\ q. degree q = d) facts = (lin,nonlin)" by force from Cons(6) have "f = prod_list us1 * prod_list divs" by auto also have "prod_list divs = prod_list facts" unfolding facts using udivs(4) by (induct divs, auto) finally have f: "f = prod_list us1 * prod_list facts" . note facts' = facts { fix u assume u: "u \ set facts" from u[unfolded facts] obtain u' where u': "u' \ set divs" and u: "u \ set (udivs u')" by auto from u' udivs(1-2)[OF u' u] prod_list_dvd[OF u, unfolded udivs(4)[OF u']] have "degree u > 0" "monic u" "\ u' \ set divs. u dvd u'" by auto } note facts = this have not1: "(v = 1) = False" using False by auto have "us = us1 @ (if length divs = n2 then divs else let (lin, nonlin) = List.partition (\q. degree q = d) facts in lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding Cons(4) facts_def udivs_def' berlekamp_factorization_main.simps Let_def not1 if_False by (rule arg_cong[where f = "\ x. us1 @ x"], rule if_cong, simp_all) (* takes time *) hence res: "us = us1 @ (if length divs = n2 then divs else lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding part by auto show ?thesis proof (cases "length divs = n2") case False with res have us: "us = (us1 @ lin) @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin)" by auto from Cons(2) have vs: "vs = (vs1 @ [v]) @ vs2" by auto have f: "f = prod_list ((us1 @ lin) @ nonlin)" unfolding f using prod_list_partition[OF part] by simp { fix u assume "u \ set ((us1 @ lin) @ nonlin)" with part have "u \ set facts \ set us1" by auto with facts Cons(7) have "degree u > 0" by (auto simp: irreducible_degree_field) } note deg = this from berlekamp_basis_length_factorization[OF sf_f n_bb mon_f f deg, unfolded Cons(3)] have "n2 \ length lin" by auto hence n: "n = length (us1 @ lin) + (n2 - length lin)" unfolding Cons(3) by auto show ?thesis proof (rule Cons(1)[OF vs n us _ f]) fix u assume "u \ set nonlin" with part have "u \ set facts" by auto from facts[OF this] show "monic u \ degree u > 0" by auto next fix u assume u: "u \ set (us1 @ lin)" { assume *: "\ (monic u \ irreducible\<^sub>d u)" with Cons(7) u have "u \ set lin" by auto with part have uf: "u \ set facts" and deg: "degree u = d" by auto from facts[OF uf] obtain u' where "u' \ set divs" and uu': "u dvd u'" by auto from this(1) have "u' dvd f" unfolding Cons(6) using prod_list_dvd[of u'] by auto with uu' have "u dvd f" by (rule dvd_trans) from facts[OF uf] d[OF this deg] * have False by auto } thus "monic u \ irreducible u" by auto next fix w u i assume w: "w \ set (vs1 @ [v])" and u: "u \ set (us1 @ lin) \ set nonlin" and i: "i < CARD('a)" from u part have u: "u \ set us1 \ set facts" by auto show "gcd u (w - [:of_nat i:]) \ {1, u}" proof (cases "u \ set us1") case True from Cons(7)[OF this] have "monic u" "irreducible u" by auto thus ?thesis by (rule monic_irreducible_gcd) next case False with u have u: "u \ set facts" by auto show ?thesis proof (cases "w = v") case True from u[unfolded facts'] obtain u' where u: "u \ set (udivs u')" and u': "u' \ set divs" by auto from udivs(3)[OF u' u i] show ?thesis unfolding True . next case False with w have w: "w \ set vs1" by auto from u obtain u' where u': "u' \ set divs" and dvd: "u dvd u'" using facts(3)[of u] dvd_refl[of u] by blast from w have "w \ set vs1 \ w = v" by auto from facts(1-2)[OF u] have u: "monic u" by auto from Cons(8)[OF w _ i] u' have "gcd u' (w - [:of_nat i:]) \ {1, u'}" by auto with dvd u show ?thesis by (rule monic_gcd_dvd) qed qed qed next case True with res have us: "us = us1 @ divs" by auto from Cons(3) True have n: "n = length us" unfolding us by auto show ?thesis unfolding us[symmetric] proof (intro conjI ballI) show f: "f = prod_list us" unfolding us using Cons(6) by simp { fix u assume "u \ set us" hence "degree u > 0" using Cons(5) Cons(7)[unfolded irreducible\<^sub>d_def] unfolding us by (auto simp: irreducible_degree_field) } note deg = this fix u assume u: "u \ set us" thus "monic u" unfolding us using Cons(5) Cons(7) by auto show "irreducible u" by (rule berlekamp_basis_irreducible[OF sf_f n_bb mon_f f n[symmetric] deg u]) qed qed next case True (* v = 1 *) with Cons(4) have us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" by simp from Cons(2) True have vs: "vs = (vs1 @ [1]) @ vs2" by auto show ?thesis proof (rule Cons(1)[OF vs Cons(3) us Cons(5-7)], goal_cases) case (3 v u i) show ?case proof (cases "v = 1") case False with 3 Cons(8)[of v u i] show ?thesis by auto next case True hence deg: "degree (v - [: of_nat i :]) = 0" by (metis (no_types, opaque_lifting) degree_pCons_0 diff_pCons diff_zero pCons_one) from 3(2) Cons(5,7)[of u] have "monic u" by auto from gcd_monic_constant[OF this deg] show ?thesis . qed qed qed next case Nil with vsf have vs1: "vs1 = berlekamp_basis f" by auto from Nil(3) have us: "us = us1 @ divs" by auto from Nil(4,6) have md: "\ u. u \ set us \ monic u \ degree u > 0" unfolding us by (auto simp: irreducible_degree_field) from Nil(7)[unfolded vs1] us have no_further_splitting_possible: "\ u v i. v \ set (berlekamp_basis f) \ u \ set us \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1, u}" by auto from Nil(5) us have prod: "f = prod_list us" by simp show ?case proof (intro conjI ballI) fix u assume u: "u \ set us" from md[OF this] have mon_u: "monic u" and deg_u: "degree u > 0" by auto from prod u have uf: "u dvd f" by (simp add: prod_list_dvd) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \P" "P \ {q. irreducible q \ monic q}" by auto show "irreducible u" proof (rule ccontr) assume irr_u: "\ irreducible u" from not_irreducible_factor_yields_prime_factors[OF uf P deg_u this] obtain pi pj where pij: "pi \ P" "pj \ P" "pi \ pj" "pi dvd u" "pj dvd u" by blast from exists_vector_in_Berlekamp_basis_dvd[OF deg_f berlekamp_basis_basis[OF deg_f, folded vs1] finite_set P pij(1-3) mon_f sf_f irr_u uf mon_u pij(4-5), unfolded vs1] obtain v s where v: "v \ set (berlekamp_basis f)" and gcd: "gcd u (v - [:s:]) \ {1,u}" using is_unit_gcd by auto from surj_of_nat_mod_ring[of s] obtain i where i: "i < CARD('a)" and s: "s = of_nat i" by auto from no_further_splitting_possible[OF v u i] gcd[unfolded s] show False by auto qed qed (insert prod md, auto) qed qed lemma berlekamp_monic_factorization: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and us: "berlekamp_monic_factorization d f = us" and d: "\ g. g dvd f \ degree g = d \ irreducible g" and deg: "degree f > 0" and mon: "monic f" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - from us[unfolded berlekamp_monic_factorization_def Let_def] deg have us: "us = [] @ berlekamp_factorization_main d [f] (berlekamp_basis f) (length (berlekamp_basis f))" by (auto) have id: "berlekamp_basis f = [] @ berlekamp_basis f" "length (berlekamp_basis f) = length [] + length (berlekamp_basis f)" "f = prod_list ([] @ [f])" by auto show "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" by (rule berlekamp_factorization_main[OF sf_f id(1) refl refl id(2) us _ _ _ id(3)], insert mon deg d, auto) qed end end diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy @@ -1,606 +1,606 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Record Based Version\ theory Finite_Field_Factorization_Record_Based imports Finite_Field_Factorization Matrix_Record_Based Poly_Mod_Finite_Field_Record_Based "HOL-Types_To_Sets.Types_To_Sets" Jordan_Normal_Form.Matrix_IArray_Impl Jordan_Normal_Form.Gauss_Jordan_IArray_Impl Polynomial_Interpolation.Improved_Code_Equations Polynomial_Factorization.Missing_List begin hide_const(open) monom coeff text \Whereas @{thm finite_field_factorization} provides a result for a polynomials over GF(p), we now develop a theorem which speaks about integer polynomials modulo p.\ lemma (in poly_mod_prime_type) finite_field_factorization_modulo_ring: assumes g: "(g :: 'a mod_ring poly) = of_int_poly f" and sf: "square_free_m f" and fact: "finite_field_factorization g = (d,gs)" and c: "c = to_int_mod_ring d" and fs: "fs = map to_int_poly gs" shows "unique_factorization_m f (c, mset fs)" proof - have [transfer_rule]: "MP_Rel f g" unfolding g MP_Rel_def by (simp add: Mp_f_representative) have sg: "square_free g" by (transfer, rule sf) have [transfer_rule]: "M_Rel c d" unfolding M_Rel_def c by (rule M_to_int_mod_ring) have fs_gs[transfer_rule]: "list_all2 MP_Rel fs gs" unfolding fs list_all2_map1 MP_Rel_def[abs_def] Mp_to_int_poly by (simp add: list.rel_refl) have [transfer_rule]: "rel_mset MP_Rel (mset fs) (mset gs)" using fs_gs using rel_mset_def by blast have [transfer_rule]: "MF_Rel (c,mset fs) (d,mset gs)" unfolding MF_Rel_def by transfer_prover from finite_field_factorization[OF sg fact] have uf: "unique_factorization Irr_Mon g (d,mset gs)" by auto from uf[untransferred] show "unique_factorization_m f (c, mset fs)" . qed text \We now have to implement @{const finite_field_factorization}.\ context fixes p :: int and ff_ops :: "'i arith_ops_record" (* finite-fields *) begin fun power_poly_f_mod_i :: "('i list \ 'i list) \ 'i list \ nat \ 'i list" where "power_poly_f_mod_i modulus a n = (if n = 0 then modulus (one_poly_i ff_ops) - else let (d,r) = Euclidean_Division.divmod_nat n 2; + else let (d,r) = Euclidean_Rings.divmod_nat n 2; rec = power_poly_f_mod_i modulus (modulus (times_poly_i ff_ops a a)) d in if r = 0 then rec else modulus (times_poly_i ff_ops rec a))" declare power_poly_f_mod_i.simps[simp del] fun power_polys_i :: "'i list \ 'i list \ 'i list \ nat \ 'i list list" where "power_polys_i mul_p u curr_p (Suc i) = curr_p # power_polys_i mul_p u (mod_field_poly_i ff_ops (times_poly_i ff_ops curr_p mul_p) u) i" | "power_polys_i mul_p u curr_p 0 = []" lemma length_power_polys_i[simp]: "length (power_polys_i x y z n) = n" by (induct n arbitrary: x y z, auto) definition berlekamp_mat_i :: "'i list \ 'i mat" where "berlekamp_mat_i u = (let n = degree_i u; ze = arith_ops_record.zero ff_ops; on = arith_ops_record.one ff_ops; mul_p = power_poly_f_mod_i (\ v. mod_field_poly_i ff_ops v u) [ze, on] (nat p); xks = power_polys_i mul_p u [on] n in mat_of_rows_list n (map (\ cs. cs @ replicate (n - length cs) ze) xks))" definition berlekamp_resulting_mat_i :: "'i list \ 'i mat" where "berlekamp_resulting_mat_i u = (let Q = berlekamp_mat_i u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then arith_ops_record.minus ff_ops (Q $$ (i,j)) (arith_ops_record.one ff_ops) else Q $$ (i,j)) in (gauss_jordan_single_i ff_ops (transpose_mat QI)))" definition berlekamp_basis_i :: "'i list \ 'i list list" where "berlekamp_basis_i u = (map (poly_of_list_i ff_ops o list_of_vec) (find_base_vectors_i ff_ops (berlekamp_resulting_mat_i u)))" primrec berlekamp_factorization_main_i :: "'i \ 'i \ nat \ 'i list list \ 'i list list \ nat \ 'i list list" where "berlekamp_factorization_main_i ze on d divs (v # vs) n = ( if v = [on] then berlekamp_factorization_main_i ze on d divs vs n else if length divs = n then divs else let of_int = arith_ops_record.of_int ff_ops; facts = filter (\ w. w \ [on]) [ gcd_poly_i ff_ops u (minus_poly_i ff_ops v (if s = 0 then [] else [of_int (int s)])) . u \ divs, s \ [0 ..< nat p]]; (lin,nonlin) = List.partition (\ q. degree_i q = d) facts in lin @ berlekamp_factorization_main_i ze on d nonlin vs (n - length lin))" | "berlekamp_factorization_main_i ze on d divs [] n = divs" definition berlekamp_monic_factorization_i :: "nat \ 'i list \ 'i list list" where "berlekamp_monic_factorization_i d f = (let vs = berlekamp_basis_i f in berlekamp_factorization_main_i (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops) d [f] vs (length vs))" partial_function (tailrec) dist_degree_factorize_main_i :: "'i \ 'i \ nat \ 'i list \ 'i list \ nat \ (nat \ 'i list) list \ (nat \ 'i list) list" where [code]: "dist_degree_factorize_main_i ze on dv v w d res = (if v = [on] then res else if d + d > dv then (dv, v) # res else let w = power_poly_f_mod_i (\ f. mod_field_poly_i ff_ops f v) w (nat p); d = Suc d; gd = gcd_poly_i ff_ops (minus_poly_i ff_ops w [ze,on]) v in if gd = [on] then dist_degree_factorize_main_i ze on dv v w d res else let v' = div_field_poly_i ff_ops v gd in dist_degree_factorize_main_i ze on (degree_i v') v' (mod_field_poly_i ff_ops w v') d ((d,gd) # res))" definition distinct_degree_factorization_i :: "'i list \ (nat \ 'i list) list" where "distinct_degree_factorization_i f = (let ze = arith_ops_record.zero ff_ops; on = arith_ops_record.one ff_ops in if degree_i f = 1 then [(1,f)] else dist_degree_factorize_main_i ze on (degree_i f) f [ze,on] 0 [])" definition finite_field_factorization_i :: "'i list \ 'i \ 'i list list" where "finite_field_factorization_i f = (if degree_i f = 0 then (lead_coeff_i ff_ops f,[]) else let a = lead_coeff_i ff_ops f; u = smult_i ff_ops (arith_ops_record.inverse ff_ops a) f; gs = (if use_distinct_degree_factorization then distinct_degree_factorization_i u else [(1,u)]); (irr,hs) = List.partition (\ (i,f). degree_i f = i) gs in (a,map snd irr @ concat (map (\ (i,g). berlekamp_monic_factorization_i i g) hs)))" end context prime_field_gen begin lemma power_polys_i: assumes i: "i < n" and [transfer_rule]: "poly_rel f f'" "poly_rel g g'" and h: "poly_rel h h'" shows "poly_rel (power_polys_i ff_ops g f h n ! i) (power_polys g' f' h' n ! i)" using i h proof (induct n arbitrary: h h' i) case (Suc n h h' i) note * = this note [transfer_rule] = *(3) show ?case proof (cases i) case 0 with Suc show ?thesis by auto next case (Suc j) with *(2-) have "j < n" by auto note IH = *(1)[OF this] show ?thesis unfolding Suc by (simp, rule IH, transfer_prover) qed qed simp lemma power_poly_f_mod_i: assumes m: "(poly_rel ===> poly_rel) m (\ x'. x' mod m')" shows "poly_rel f f' \ poly_rel (power_poly_f_mod_i ff_ops m f n) (power_poly_f_mod m' f' n)" proof - from m have m: "\ x x'. poly_rel x x' \ poly_rel (m x) (x' mod m')" unfolding rel_fun_def by auto show "poly_rel f f' \ poly_rel (power_poly_f_mod_i ff_ops m f n) (power_poly_f_mod m' f' n)" proof (induct n arbitrary: f f' rule: less_induct) case (less n f f') note f[transfer_rule] = less(2) show ?case proof (cases "n = 0") case True show ?thesis by (simp add: True power_poly_f_mod_i.simps power_poly_f_mod_binary, rule m[OF poly_rel_one]) next case False hence n: "(n = 0) = False" by simp - obtain q r where div: "Euclidean_Division.divmod_nat n 2 = (q,r)" by force - from this[unfolded Euclidean_Division.divmod_nat_def] n have "q < n" by auto + obtain q r where div: "Euclidean_Rings.divmod_nat n 2 = (q,r)" by force + from this[unfolded Euclidean_Rings.divmod_nat_def] n have "q < n" by auto note IH = less(1)[OF this] have rec: "poly_rel (power_poly_f_mod_i ff_ops m (m (times_poly_i ff_ops f f)) q) (power_poly_f_mod m' (f' * f' mod m') q)" by (rule IH, rule m, transfer_prover) have other: "poly_rel (m (times_poly_i ff_ops (power_poly_f_mod_i ff_ops m (m (times_poly_i ff_ops f f)) q) f)) (power_poly_f_mod m' (f' * f' mod m') q * f' mod m')" by (rule m, rule poly_rel_times[unfolded rel_fun_def, rule_format, OF rec f]) show ?thesis unfolding power_poly_f_mod_i.simps[of _ _ _ n] Let_def power_poly_f_mod_binary[of _ _ n] div split n if_False using rec other by auto qed qed qed lemma berlekamp_mat_i[transfer_rule]: "(poly_rel ===> mat_rel R) (berlekamp_mat_i p ff_ops) berlekamp_mat" proof (intro rel_funI) fix f f' let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" assume f[transfer_rule]: "poly_rel f f'" have deg: "degree_i f = degree f'" by transfer_prover { fix i j assume i: "i < degree f'" and j: "j < degree f'" define cs where "cs = (\cs :: 'i list. cs @ replicate (degree f' - length cs) ?ze)" define cs' where "cs' = (\cs :: 'a mod_ring poly. coeffs cs @ replicate (degree f' - length (coeffs cs)) 0)" define poly where "poly = power_polys_i ff_ops (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) f [?on] (degree f')" define poly' where "poly' = (power_polys (power_poly_f_mod f' [:0, 1:] (nat p)) f' 1 (degree f'))" have *: "poly_rel (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) (power_poly_f_mod f' [:0, 1:] (nat p))" by (rule power_poly_f_mod_i, transfer_prover, simp add: poly_rel_def one zero) have [transfer_rule]: "poly_rel (poly ! i) (poly' ! i)" unfolding poly_def poly'_def by (rule power_polys_i[OF i f *], simp add: poly_rel_def one) have *: "list_all2 R (cs (poly ! i)) (cs' (poly' ! i))" unfolding cs_def cs'_def by transfer_prover from list_all2_nthD[OF *[unfolded poly_rel_def], of j] j have "R (cs (poly ! i) ! j) (cs' (poly' ! i) ! j)" unfolding cs_def by auto hence "R (mat_of_rows_list (degree f') (map (\cs. cs @ replicate (degree f' - length cs) ?ze) (power_polys_i ff_ops (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) f [?on] (degree f'))) $$ (i, j)) (mat_of_rows_list (degree f') (map (\cs. coeffs cs @ replicate (degree f' - length (coeffs cs)) 0) (power_polys (power_poly_f_mod f' [:0, 1:] (nat p)) f' 1 (degree f'))) $$ (i, j))" unfolding mat_of_rows_list_def length_map length_power_polys_i power_polys_works length_power_polys index_mat[OF i j] split unfolding poly_def cs_def poly'_def cs'_def using i by auto } note main = this show "mat_rel R (berlekamp_mat_i p ff_ops f) (berlekamp_mat f')" unfolding berlekamp_mat_i_def berlekamp_mat_def Let_def nat_p[symmetric] deg unfolding mat_rel_def by (intro conjI allI impI, insert main, auto) qed lemma berlekamp_resulting_mat_i[transfer_rule]: "(poly_rel ===> mat_rel R) (berlekamp_resulting_mat_i p ff_ops) berlekamp_resulting_mat" proof (intro rel_funI) fix f f' assume "poly_rel f f'" from berlekamp_mat_i[unfolded rel_fun_def, rule_format, OF this] have bmi: "mat_rel R (berlekamp_mat_i p ff_ops f) (berlekamp_mat f')" . show "mat_rel R (berlekamp_resulting_mat_i p ff_ops f) (berlekamp_resulting_mat f')" unfolding berlekamp_resulting_mat_def Let_def berlekamp_resulting_mat_i_def by (rule gauss_jordan_i[unfolded rel_fun_def, rule_format], insert bmi, auto simp: mat_rel_def one intro!: minus[unfolded rel_fun_def, rule_format]) qed lemma berlekamp_basis_i[transfer_rule]: "(poly_rel ===> list_all2 poly_rel) (berlekamp_basis_i p ff_ops) berlekamp_basis" unfolding berlekamp_basis_i_def[abs_def] berlekamp_basis_code[abs_def] o_def by transfer_prover lemma berlekamp_factorization_main_i[transfer_rule]: "((=) ===> list_all2 poly_rel ===> list_all2 poly_rel ===> (=) ===> list_all2 poly_rel) (berlekamp_factorization_main_i p ff_ops (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops)) berlekamp_factorization_main" proof (intro rel_funI, clarify, goal_cases) case (1 _ d xs xs' ys ys' _ n) let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" let ?of_int = "arith_ops_record.of_int ff_ops" from 1(2) 1(1) show ?case proof (induct ys ys' arbitrary: xs xs' n rule: list_all2_induct) case (Cons y ys y' ys' xs xs' n) note trans[transfer_rule] = Cons(1,2,4) obtain clar0 clar1 clar2 where clarify: "\ s u. gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)])) = clar0 s u" "[0..u. concat (map (\s. if gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)])) \ [?on] then [gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)]))] else []) [0..concat (map (\u. map (\s. gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)]))) [0.. [?on]]" have Facts: "Facts = facts" unfolding Facts_def facts_def clarify proof (induct xs) case (Cons x xs) show ?case by (simp add: Cons, induct clar1, auto) qed simp define facts' where "facts' = concat (map (\u. concat (map (\x. if gcd u (y' - [:of_nat x:]) \ 1 then [gcd u (y' - [:of_int (int x):])] else []) [0.. x. of_int (int x) = of_nat x" "[?on] = one_poly_i ff_ops" by (auto simp: one_poly_i_def) have facts[transfer_rule]: "list_all2 poly_rel facts facts'" unfolding facts_def facts'_def apply (rule concat_transfer[unfolded rel_fun_def, rule_format]) apply (rule list.map_transfer[unfolded rel_fun_def, rule_format, OF _ trans(3)]) apply (rule concat_transfer[unfolded rel_fun_def, rule_format]) apply (rule list_all2_map_map) proof (unfold id) fix f f' x assume [transfer_rule]: "poly_rel f f'" and x: "x \ set [0.. int x" "int x < p" by auto from of_int[OF this] have rel[transfer_rule]: "R (?of_int (int x)) (of_nat x)" by auto { assume "0 < x" with * have *: "0 < int x" "int x < p" by auto have "(of_nat x :: 'a mod_ring) = of_int (int x)" by simp also have "\ \ 0" unfolding of_int_of_int_mod_ring using * unfolding p by (transfer', auto) } with rel have [transfer_rule]: "poly_rel (if x = 0 then [] else [?of_int (int x)]) [:of_nat x:]" unfolding poly_rel_def by (auto simp add: cCons_def p) show "list_all2 poly_rel (if gcd_poly_i ff_ops f (minus_poly_i ff_ops y (if x = 0 then [] else [?of_int (int x)])) \ one_poly_i ff_ops then [gcd_poly_i ff_ops f (minus_poly_i ff_ops y (if x = 0 then [] else [?of_int (int x)]))] else []) (if gcd f' (y' - [:of_nat x:]) \ 1 then [gcd f' (y' - [:of_nat x:])] else [])" by transfer_prover qed have id1: "berlekamp_factorization_main_i p ff_ops ?ze ?on d xs (y # ys) n = ( if y = [?on] then berlekamp_factorization_main_i p ff_ops ?ze ?on d xs ys n else if length xs = n then xs else (let fac = facts; (lin, nonlin) = List.partition (\q. degree_i q = d) fac in lin @ berlekamp_factorization_main_i p ff_ops ?ze ?on d nonlin ys (n - length lin)))" unfolding berlekamp_factorization_main_i.simps Facts[symmetric] by (simp add: o_def Facts_def Let_def) have id2: "berlekamp_factorization_main d xs' (y' # ys') n = ( if y' = 1 then berlekamp_factorization_main d xs' ys' n else if length xs' = n then xs' else (let fac = facts'; (lin, nonlin) = List.partition (\q. degree q = d) fac in lin @ berlekamp_factorization_main d nonlin ys' (n - length lin)))" by (simp add: o_def facts'_def nat_p) have len: "length xs = length xs'" by transfer_prover have id3: "(y = [?on]) = (y' = 1)" by (transfer_prover_start, transfer_step+, simp add: one_poly_i_def finite_field_ops_int_def) show ?case proof (cases "y' = 1") case True hence id4: "(y' = 1) = True" by simp show ?thesis unfolding id1 id2 id3 id4 if_True by (rule Cons(3), transfer_prover) next case False hence id4: "(y' = 1) = False" by simp note id1 = id1[unfolded id3 id4 if_False] note id2 = id2[unfolded id4 if_False] show ?thesis proof (cases "length xs' = n") case True thus ?thesis unfolding id1 id2 Let_def len using trans by simp next case False hence id: "(length xs' = n) = False" by simp have id': "length [q\facts . degree_i q = d] = length [q\facts'. degree q = d]" by transfer_prover have [transfer_rule]: "list_all2 poly_rel (berlekamp_factorization_main_i p ff_ops ?ze ?on d [x\facts . degree_i x \ d] ys (n - length [q\facts . degree_i q = d])) (berlekamp_factorization_main d [x\facts' . degree x \ d] ys' (n - length [q\facts' . degree q = d]))" unfolding id' by (rule Cons(3), transfer_prover) show ?thesis unfolding id1 id2 Let_def len id if_False unfolding partition_filter_conv o_def split by transfer_prover qed qed qed simp qed lemma berlekamp_monic_factorization_i[transfer_rule]: "((=) ===> poly_rel ===> list_all2 poly_rel) (berlekamp_monic_factorization_i p ff_ops) berlekamp_monic_factorization" unfolding berlekamp_monic_factorization_i_def[abs_def] berlekamp_monic_factorization_def[abs_def] Let_def by transfer_prover lemma dist_degree_factorize_main_i: "poly_rel F f \ poly_rel G g \ list_all2 (rel_prod (=) poly_rel) Res res \ list_all2 (rel_prod (=) poly_rel) (dist_degree_factorize_main_i p ff_ops (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops) (degree_i F) F G d Res) (dist_degree_factorize_main f g d res)" proof (induct f g d res arbitrary: F G Res rule: dist_degree_factorize_main.induct) case (1 v w d res V W Res) let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" note simp = dist_degree_factorize_main.simps[of v w d] dist_degree_factorize_main_i.simps[of p ff_ops ?ze ?on "degree_i V" V W d] have v[transfer_rule]: "poly_rel V v" by (rule 1) have w[transfer_rule]: "poly_rel W w" by (rule 1) have res[transfer_rule]: "list_all2 (rel_prod (=) poly_rel) Res res" by (rule 1) have [transfer_rule]: "poly_rel [?on] 1" by (simp add: one poly_rel_def) have id1: "(V = [?on]) = (v = 1)" unfolding finite_field_ops_int_def by transfer_prover have id2: "degree_i V = degree v" by transfer_prover note simp = simp[unfolded id1 id2] note IH = 1(1,2) show ?case proof (cases "v = 1") case True with res show ?thesis unfolding id2 simp by simp next case False with id1 have "(v = 1) = False" by auto note simp = simp[unfolded this if_False] note IH = IH[OF False] show ?thesis proof (cases "degree v < d + d") case True thus ?thesis unfolding id2 simp using res v by auto next case False hence "(degree v < d + d) = False" by auto note simp = simp[unfolded this if_False] let ?P = "power_poly_f_mod_i ff_ops (\f. mod_field_poly_i ff_ops f V) W (nat p)" let ?G = "gcd_poly_i ff_ops (minus_poly_i ff_ops ?P [?ze, ?on]) V" let ?g = "gcd (w ^ CARD('a) mod v - monom 1 1) v" define G where "G = ?G" define g where "g = ?g" note simp = simp[unfolded Let_def, folded G_def g_def] note IH = IH[OF False refl refl refl] have [transfer_rule]: "poly_rel [?ze,?on] (monom 1 1)" unfolding poly_rel_def by (auto simp: coeffs_monom one zero) have id: "w ^ CARD('a) mod v = power_poly_f_mod v w (nat p)" unfolding power_poly_f_mod_def by (simp add: p) have P[transfer_rule]: "poly_rel ?P (w ^ CARD('a) mod v)" unfolding id by (rule power_poly_f_mod_i[OF _ w], transfer_prover) have g[transfer_rule]: "poly_rel G g" unfolding G_def g_def by transfer_prover have id3: "(G = [?on]) = (g = 1)" by transfer_prover note simp = simp[unfolded id3] show ?thesis proof (cases "g = 1") case True from IH(1)[OF this[unfolded g_def] v P res] True show ?thesis unfolding id2 simp by simp next case False have vg: "poly_rel (div_field_poly_i ff_ops V G) (v div g)" by transfer_prover have "poly_rel (mod_field_poly_i ff_ops ?P (div_field_poly_i ff_ops V G)) (w ^ CARD('a) mod v mod (v div g))" by transfer_prover note IH = IH(2)[OF False[unfolded g_def] refl vg[unfolded G_def g_def] this[unfolded G_def g_def], folded g_def G_def] have "list_all2 (rel_prod (=) poly_rel) ((Suc d, G) # Res) ((Suc d, g) # res)" using g res by auto note IH = IH[OF this] from False have "(g = 1) = False" by simp note simp = simp[unfolded this if_False] show ?thesis unfolding id2 simp using IH by simp qed qed qed qed lemma distinct_degree_factorization_i[transfer_rule]: "(poly_rel ===> list_all2 (rel_prod (=) poly_rel)) (distinct_degree_factorization_i p ff_ops) distinct_degree_factorization" proof fix F f assume f[transfer_rule]: "poly_rel F f" have id: "(degree_i F = 1) = (degree f = 1)" by transfer_prover note d = distinct_degree_factorization_i_def distinct_degree_factorization_def let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" show "list_all2 (rel_prod (=) poly_rel) (distinct_degree_factorization_i p ff_ops F) (distinct_degree_factorization f)" proof (cases "degree f = 1") case True with id f show ?thesis unfolding d by auto next case False from False id have "?thesis = (list_all2 (rel_prod (=) poly_rel) (dist_degree_factorize_main_i p ff_ops ?ze ?on (degree_i F) F [?ze, ?on] 0 []) (dist_degree_factorize_main f (monom 1 1) 0 []))" unfolding d Let_def by simp also have \ by (rule dist_degree_factorize_main_i[OF f], auto simp: poly_rel_def coeffs_monom one zero) finally show ?thesis . qed qed lemma finite_field_factorization_i[transfer_rule]: "(poly_rel ===> rel_prod R (list_all2 poly_rel)) (finite_field_factorization_i p ff_ops) finite_field_factorization" unfolding finite_field_factorization_i_def finite_field_factorization_def Let_def lead_coeff_i_def' by transfer_prover text \Since the implementation is sound, we can now combine it with the soundness result of the finite field factorization.\ lemma finite_field_i_sound: assumes f': "f' = of_int_poly_i ff_ops (Mp f)" and berl_i: "finite_field_factorization_i p ff_ops f' = (c',fs')" and sq: "square_free_m f" and fs: "fs = map (to_int_poly_i ff_ops) fs'" and c: "c = arith_ops_record.to_int ff_ops c'" shows "unique_factorization_m f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" proof - define f'' :: "'a mod_ring poly" where "f'' = of_int_poly (Mp f)" have rel_f[transfer_rule]: "poly_rel f' f''" by (rule poly_rel_of_int_poly[OF f'], simp add: f''_def) interpret pff: idom_ops "poly_ops ff_ops" poly_rel by (rule idom_ops_poly) obtain c'' fs'' where berl: "finite_field_factorization f'' = (c'',fs'')" by force from rel_funD[OF finite_field_factorization_i rel_f, unfolded rel_prod_conv assms(2) split berl] have rel[transfer_rule]: "R c' c''" "list_all2 poly_rel fs' fs''" by auto from to_int[OF rel(1)] have cc': "c = to_int_mod_ring c''" unfolding c by simp have c: "c \ {0 ..< p}" unfolding cc' by (metis Divides.pos_mod_bound Divides.pos_mod_sign M_to_int_mod_ring atLeastLessThan_iff gr_implies_not_zero nat_le_0 nat_p not_le poly_mod.M_def zero_less_card_finite) { fix f assume "f \ set fs'" with rel(2) obtain f' where "poly_rel f f'" unfolding list_all2_conv_all_nth set_conv_nth by auto hence "is_poly ff_ops f" using fun_cong[OF Domainp_is_poly, of f] unfolding Domainp_iff[abs_def] by auto } hence fs': "Ball (set fs') (is_poly ff_ops)" by auto define mon :: "'a mod_ring poly \ bool" where "mon = monic" have [transfer_rule]: "(poly_rel ===> (=)) (monic_i ff_ops) mon" unfolding mon_def by (rule poly_rel_monic) have len: "length fs' = length fs''" by transfer_prover have fs': "fs = map to_int_poly fs''" unfolding fs proof (rule nth_map_conv[OF len], intro allI impI) fix i assume i: "i < length fs'" obtain f g where id: "fs' ! i = f" "fs'' ! i = g" by auto from i rel(2)[unfolded list_all2_conv_all_nth[of _ fs' fs'']] id have "poly_rel f g" by auto from to_int_poly_i[OF this] have "to_int_poly_i ff_ops f = to_int_poly g" . thus "to_int_poly_i ff_ops (fs' ! i) = to_int_poly (fs'' ! i)" unfolding id . qed have f: "f'' = of_int_poly f" unfolding poly_eq_iff f''_def by (simp add: to_int_mod_ring_hom.injectivity to_int_mod_ring_of_int_M Mp_coeff) have *: "unique_factorization_m f (c, mset fs)" using finite_field_factorization_modulo_ring[OF f sq berl cc' fs'] by auto have fs': "(\fi\set fs. set (coeffs fi) \ {0.. int poly \ int \ int poly list" where "finite_field_factorization_main p f_ops f \ let (c',fs') = finite_field_factorization_i p f_ops (of_int_poly_i f_ops (poly_mod.Mp p f)) in (arith_ops_record.to_int f_ops c', map (to_int_poly_i f_ops) fs')" lemma(in prime_field_gen) finite_field_factorization_main: assumes res: "finite_field_factorization_main p ff_ops f = (c,fs)" and sq: "square_free_m f" shows "unique_factorization_m f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" proof - obtain c' fs' where res': "finite_field_factorization_i p ff_ops (of_int_poly_i ff_ops (Mp f)) = (c', fs')" by force show ?thesis by (rule finite_field_i_sound[OF refl res' sq], insert res[unfolded finite_field_factorization_main_def res'], auto) qed definition finite_field_factorization_int :: "int \ int poly \ int \ int poly list" where "finite_field_factorization_int p = ( if p \ 65535 then finite_field_factorization_main p (finite_field_ops32 (uint32_of_int p)) else if p \ 4294967295 then finite_field_factorization_main p (finite_field_ops64 (uint64_of_int p)) else finite_field_factorization_main p (finite_field_ops_integer (integer_of_int p)))" context poly_mod_prime begin lemmas finite_field_factorization_main_integer = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops_integer, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas finite_field_factorization_main_uint32 = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops32, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas finite_field_factorization_main_uint64 = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops64, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemma finite_field_factorization_int: assumes sq: "poly_mod.square_free_m p f" and result: "finite_field_factorization_int p f = (c,fs)" shows "poly_mod.unique_factorization_m p f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" using finite_field_factorization_main_integer[OF _ sq, of c fs] finite_field_factorization_main_uint32[OF _ _ sq, of c fs] finite_field_factorization_main_uint64[OF _ _ sq, of c fs] result[unfolded finite_field_factorization_int_def] by (auto split: if_splits) end end diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy @@ -1,1657 +1,1657 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Finite Fields\ text \We provide four implementations for $GF(p)$ -- the field with $p$ elements for some prime $p$ -- one by int, one by integers, one by 32-bit numbers and one 64-bit implementation. Correctness of the implementations is proven by transfer rules to the type-based version of $GF(p)$.\ theory Finite_Field_Record_Based imports Finite_Field Arithmetic_Record_Based Native_Word.Uint32 Native_Word.Uint64 "HOL-Library.Code_Target_Numeral" Native_Word.Code_Target_Int_Bit begin (* mod on standard case which can immediately be mapped to target languages without considering special cases *) definition mod_nonneg_pos :: "integer \ integer \ integer" where "x \ 0 \ y > 0 \ mod_nonneg_pos x y = (x mod y)" code_printing \ \FIXME illusion of partiality\ constant mod_nonneg_pos \ (SML) "IntInf.mod/ ( _,/ _ )" and (Eval) "IntInf.mod/ ( _,/ _ )" and (OCaml) "Z.rem" and (Haskell) "Prelude.mod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ (k '% l))" definition mod_nonneg_pos_int :: "int \ int \ int" where "mod_nonneg_pos_int x y = int_of_integer (mod_nonneg_pos (integer_of_int x) (integer_of_int y))" lemma mod_nonneg_pos_int[simp]: "x \ 0 \ y > 0 \ mod_nonneg_pos_int x y = (x mod y)" unfolding mod_nonneg_pos_int_def using mod_nonneg_pos_def by simp context fixes p :: int begin definition plus_p :: "int \ int \ int" where "plus_p x y \ let z = x + y in if z \ p then z - p else z" definition minus_p :: "int \ int \ int" where "minus_p x y \ if y \ x then x - y else x + p - y" definition uminus_p :: "int \ int" where "uminus_p x = (if x = 0 then 0 else p - x)" definition mult_p :: "int \ int \ int" where "mult_p x y = (mod_nonneg_pos_int (x * y) p)" fun power_p :: "int \ nat \ int" where "power_p x n = (if n = 0 then 1 else - let (d,r) = Euclidean_Division.divmod_nat n 2; + let (d,r) = Euclidean_Rings.divmod_nat n 2; rec = power_p (mult_p x x) d in if r = 0 then rec else mult_p rec x)" text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p :: "int \ int" where "inverse_p x = (if x = 0 then 0 else power_p x (nat (p - 2)))" definition divide_p :: "int \ int \ int" where "divide_p x y = mult_p x (inverse_p y)" definition finite_field_ops_int :: "int arith_ops_record" where "finite_field_ops_int \ Arith_Ops_Record 0 1 plus_p mult_p minus_p uminus_p divide_p inverse_p (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) (\ x . x) (\ x . x) (\ x. 0 \ x \ x < p)" end context fixes p :: uint32 begin definition plus_p32 :: "uint32 \ uint32 \ uint32" where "plus_p32 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p32 :: "uint32 \ uint32 \ uint32" where "minus_p32 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p32 :: "uint32 \ uint32" where "uminus_p32 x = (if x = 0 then 0 else p - x)" definition mult_p32 :: "uint32 \ uint32 \ uint32" where "mult_p32 x y = (x * y mod p)" lemma int_of_uint32_shift: "int_of_uint32 (drop_bit k n) = (int_of_uint32 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint32_0_iff: "int_of_uint32 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint32_0: "int_of_uint32 0 = 0" unfolding int_of_uint32_0_iff by simp lemma int_of_uint32_ge_0: "int_of_uint32 n \ 0" by (transfer, auto) lemma two_32: "2 ^ LENGTH(32) = (4294967296 :: int)" by simp lemma int_of_uint32_plus: "int_of_uint32 (x + y) = (int_of_uint32 x + int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_minus: "int_of_uint32 (x - y) = (int_of_uint32 x - int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mult: "int_of_uint32 (x * y) = (int_of_uint32 x * int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mod: "int_of_uint32 (x mod y) = (int_of_uint32 x mod int_of_uint32 y)" by (transfer, unfold uint_mod two_32, rule refl) lemma int_of_uint32_inv: "0 \ x \ x < 4294967296 \ int_of_uint32 (uint32_of_int x) = x" by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p32 :: "uint32 \ uint32 \ uint32" where "power_p32 x n = (if n = 0 then 1 else let rec = power_p32 (mult_p32 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p32 rec x)" by pat_completeness auto termination proof - { fix n :: uint32 assume "n \ 0" with int_of_uint32_ge_0[of n] int_of_uint32_0_iff[of n] have "int_of_uint32 n > 0" by auto hence "0 < int_of_uint32 n" "int_of_uint32 n div 2 < int_of_uint32 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint32 n))", auto simp: int_of_uint32_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p32 :: "uint32 \ uint32" where "inverse_p32 x = (if x = 0 then 0 else power_p32 x (p - 2))" definition divide_p32 :: "uint32 \ uint32 \ uint32" where "divide_p32 x y = mult_p32 x (inverse_p32 y)" definition finite_field_ops32 :: "uint32 arith_ops_record" where "finite_field_ops32 \ Arith_Ops_Record 0 1 plus_p32 mult_p32 minus_p32 uminus_p32 divide_p32 inverse_p32 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint32_of_int int_of_uint32 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint32_code [code_unfold]: "drop_bit 1 x = (uint32_shiftr x 1)" by (simp add: uint32_shiftr_def) (* ******************************************************************************** *) subsubsection \Transfer Relation\ locale mod_ring_locale = fixes p :: int and ty :: "'a :: nontriv itself" assumes p: "p = int CARD('a)" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp lemma p2: "p \ 2" unfolding p using nontriv[where 'a = 'a] by auto lemma p2_ident: "int (CARD('a) - 2) = p - 2" using p2 unfolding p by simp definition mod_ring_rel :: "int \ 'a mod_ring \ bool" where "mod_ring_rel x x' = (x = to_int_mod_ring x')" (* domain transfer rules *) lemma Domainp_mod_ring_rel [transfer_domain_rule]: "Domainp (mod_ring_rel) = (\ v. v \ {0 ..< p})" proof - { fix v :: int assume *: "0 \ v" "v < p" have "Domainp mod_ring_rel v" proof show "mod_ring_rel v (of_int_mod_ring v)" unfolding mod_ring_rel_def using * p by auto qed } note * = this show ?thesis by (intro ext iffI, insert range_to_int_mod_ring[where 'a = 'a] *, auto simp: mod_ring_rel_def p) qed (* left/right/bi-unique *) lemma bi_unique_mod_ring_rel [transfer_rule]: "bi_unique mod_ring_rel" "left_unique mod_ring_rel" "right_unique mod_ring_rel" unfolding mod_ring_rel_def bi_unique_def left_unique_def right_unique_def by auto (* left/right-total *) lemma right_total_mod_ring_rel [transfer_rule]: "right_total mod_ring_rel" unfolding mod_ring_rel_def right_total_def by simp (* ************************************************************************************ *) subsubsection \Transfer Rules\ (* 0 / 1 *) lemma mod_ring_0[transfer_rule]: "mod_ring_rel 0 0" unfolding mod_ring_rel_def by simp lemma mod_ring_1[transfer_rule]: "mod_ring_rel 1 1" unfolding mod_ring_rel_def by simp (* addition *) lemma plus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "plus_p p x y = ((x + y) mod p)" proof (cases "p \ x + y") case False thus ?thesis using x y unfolding plus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x + y - p" "x + y - p < p" by auto from True have id: "plus_p p x y = x + y - p" unfolding plus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_plus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (plus_p p) (+)" proof - { fix x y :: "'a mod_ring" have "plus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x + y)" by (transfer, subst plus_p_mod_def, auto, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* subtraction *) lemma minus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "minus_p p x y = ((x - y) mod p)" proof (cases "x - y < 0") case False thus ?thesis using x y unfolding minus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x - y + p" "x - y + p < p" by auto from True have id: "minus_p p x y = x - y + p" unfolding minus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_minus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (minus_p p) (-)" proof - { fix x y :: "'a mod_ring" have "minus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x - y)" by (transfer, subst minus_p_mod_def, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* unary minus *) lemma mod_ring_uminus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (uminus_p p) uminus" proof - { fix x :: "'a mod_ring" have "uminus_p p (to_int_mod_ring x) = to_int_mod_ring (uminus x)" by (transfer, auto simp: uminus_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* multiplication *) lemma mod_ring_mult[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (mult_p p) ((*))" proof - { fix x y :: "'a mod_ring" have "mult_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x * y)" by (transfer, auto simp: mult_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* equality *) lemma mod_ring_eq[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> (=)) (=) (=)" by (intro rel_funI, auto simp: mod_ring_rel_def) (* power *) lemma mod_ring_power[transfer_rule]: "(mod_ring_rel ===> (=) ===> mod_ring_rel) (power_p p) (^)" proof (intro rel_funI, clarify, unfold binary_power[symmetric], goal_cases) fix x y n assume xy: "mod_ring_rel x y" from xy show "mod_ring_rel (power_p p x n) (binary_power y n)" proof (induct y n arbitrary: x rule: binary_power.induct) case (1 x n y) note 1(2)[transfer_rule] show ?case proof (cases "n = 0") case True thus ?thesis by (simp add: mod_ring_1) next case False - obtain d r where id: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force + obtain d r where id: "Euclidean_Rings.divmod_nat n 2 = (d,r)" by force let ?int = "power_p p (mult_p p y y) d" let ?gfp = "binary_power (x * x) d" from False have id': "?thesis = (mod_ring_rel (if r = 0 then ?int else mult_p p ?int y) (if r = 0 then ?gfp else ?gfp * x))" unfolding power_p.simps[of _ _ n] binary_power.simps[of _ n] Let_def id split by simp have [transfer_rule]: "mod_ring_rel ?int ?gfp" by (rule 1(1)[OF False refl id[symmetric]], transfer_prover) show ?thesis unfolding id' by transfer_prover qed qed qed declare power_p.simps[simp del] lemma ring_finite_field_ops_int: "ring_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end locale prime_field = mod_ring_locale p ty for p and ty :: "'a :: prime_card itself" begin lemma prime: "prime p" unfolding p using prime_card[where 'a = 'a] by simp (* mod *) lemma mod_ring_mod[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) ((\ x y. if y = 0 then x else 0)) (mod)" proof - { fix x y :: "'a mod_ring" have "(if to_int_mod_ring y = 0 then to_int_mod_ring x else 0) = to_int_mod_ring (x mod y)" unfolding modulo_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* normalize *) lemma mod_ring_normalize[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) ((\ x. if x = 0 then 0 else 1)) normalize" proof - { fix x :: "'a mod_ring" have "(if to_int_mod_ring x = 0 then 0 else 1) = to_int_mod_ring (normalize x)" unfolding normalize_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* unit_factor *) lemma mod_ring_unit_factor[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (\ x. x) unit_factor" proof - { fix x :: "'a mod_ring" have "to_int_mod_ring x = to_int_mod_ring (unit_factor x)" unfolding unit_factor_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* inverse *) lemma mod_ring_inverse[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (inverse_p p) inverse" proof (intro rel_funI) fix x y assume [transfer_rule]: "mod_ring_rel x y" show "mod_ring_rel (inverse_p p x) (inverse y)" unfolding inverse_p_def inverse_mod_ring_def apply (transfer_prover_start) apply (transfer_step)+ apply (unfold p2_ident) apply (rule refl) done qed (* division *) lemma mod_ring_divide[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (divide_p p) (/)" unfolding divide_p_def[abs_def] divide_mod_ring_def[abs_def] inverse_mod_ring_def[symmetric] by transfer_prover lemma mod_ring_rel_unsafe: assumes "x < CARD('a)" shows "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" proof - have id: "of_nat x = (of_int (int x) :: 'a mod_ring)" by simp show "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" unfolding id unfolding mod_ring_rel_def proof (auto simp add: assms of_int_of_int_mod_ring) assume "0 < x" with assms have "of_int_mod_ring (int x) \ (0 :: 'a mod_ring)" by (metis (no_types) less_imp_of_nat_less less_irrefl of_nat_0_le_iff of_nat_0_less_iff to_int_mod_ring_hom.hom_zero to_int_mod_ring_of_int_mod_ring) thus "of_int_mod_ring (int x) = (0 :: 'a mod_ring) \ False" by blast qed qed lemma finite_field_ops_int: "field_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_divide mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_inverse mod_ring_mod mod_ring_unit_factor mod_ring_normalize mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end text \Once we have proven the soundness of the implementation, we do not care any longer that @{typ "'a mod_ring"} has been defined internally via lifting. Disabling the transfer-rules will hide the internal definition in further applications of transfer.\ lifting_forget mod_ring.lifting text \For soundness of the 32-bit implementation, we mainly prove that this implementation implements the int-based implementation of the mod-ring.\ context mod_ring_locale begin context fixes pp :: "uint32" assumes ppp: "p = int_of_uint32 pp" and small: "p \ 65535" begin lemmas uint32_simps = int_of_uint32_0 int_of_uint32_plus int_of_uint32_minus int_of_uint32_mult definition urel32 :: "uint32 \ int \ bool" where "urel32 x y = (y = int_of_uint32 x \ y < p)" definition mod_ring_rel32 :: "uint32 \ 'a mod_ring \ bool" where "mod_ring_rel32 x y = (\ z. urel32 x z \ mod_ring_rel z y)" lemma urel32_0: "urel32 0 0" unfolding urel32_def using p2 by (simp, transfer, simp) lemma urel32_1: "urel32 1 1" unfolding urel32_def using p2 by (simp, transfer, simp) lemma le_int_of_uint32: "(x \ y) = (int_of_uint32 x \ int_of_uint32 y)" by (transfer, simp add: word_le_def) lemma urel32_plus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (plus_p32 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" let ?p = "int_of_uint32 pp" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_minus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (minus_p32 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_uminus: assumes "urel32 x y" shows "urel32 (uminus_p32 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint32 x" from assms int_of_uint32_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel32_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint32_0_iff using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_mult: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (mult_p32 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel32_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 65536 * 65536" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 4294967296" by simp show ?thesis unfolding id using small rel unfolding mult_p32_def mult_p_def Let_def urel32_def unfolding ppp by (auto simp: uint32_simps, unfold int_of_uint32_mod int_of_uint32_mult, subst mod_pos_pos_trivial[of _ 4294967296], insert le, auto) qed lemma urel32_eq: assumes "urel32 x y" "urel32 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel32_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel32_normalize: assumes x: "urel32 x y" shows "urel32 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel32_eq[OF x urel32_0] using urel32_0 urel32_1 by auto lemma urel32_mod: assumes x: "urel32 x x'" and y: "urel32 y y'" shows "urel32 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel32_eq[OF y urel32_0] using urel32_0 x by auto lemma urel32_power: "urel32 x x' \ urel32 y (int y') \ urel32 (power_p32 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel32_eq[OF y urel32_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel32_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel32_eq[OF y urel32_0] by auto from y have \int y' = int_of_uint32 y\ \int y' < p\ by (simp_all add: urel32_def) - obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force - from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Rings.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Rings.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel32 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel32_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint32.rep_eq modulo_uint32.rep_eq uint_mod \int y' = int_of_uint32 y\) done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel32_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p32.simps[of _ _ y] dr' id if_False rem using IH urel32_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel32_inverse: assumes x: "urel32 x x'" shows "urel32 (inverse_p32 pp x) (inverse_p p x')" proof - have p: "urel32 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel32_def unfolding ppp by (simp add: int_of_uint32.rep_eq minus_uint32.rep_eq uint_sub_if') show ?thesis unfolding inverse_p32_def inverse_p_def urel32_eq[OF x urel32_0] using urel32_0 urel32_power[OF x p] by auto qed lemma mod_ring_0_32: "mod_ring_rel32 0 0" using urel32_0 mod_ring_0 unfolding mod_ring_rel32_def by blast lemma mod_ring_1_32: "mod_ring_rel32 1 1" using urel32_1 mod_ring_1 unfolding mod_ring_rel32_def by blast lemma mod_ring_uminus32: "(mod_ring_rel32 ===> mod_ring_rel32) (uminus_p32 pp) uminus" using urel32_uminus mod_ring_uminus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_plus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (plus_p32 pp) (+)" using urel32_plus mod_ring_plus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_minus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (minus_p32 pp) (-)" using urel32_minus mod_ring_minus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_mult32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (mult_p32 pp) ((*))" using urel32_mult mod_ring_mult unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_eq32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> (=)) (=) (=)" using urel32_eq mod_ring_eq unfolding mod_ring_rel32_def rel_fun_def by blast lemma urel32_inj: "urel32 x y \ urel32 x z \ y = z" using urel32_eq[of x y x z] by auto lemma urel32_inj': "urel32 x z \ urel32 y z \ x = y" using urel32_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel32: "bi_unique mod_ring_rel32" "left_unique mod_ring_rel32" "right_unique mod_ring_rel32" using bi_unique_mod_ring_rel urel32_inj' unfolding mod_ring_rel32_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel32_def) lemma right_total_mod_ring_rel32: "right_total mod_ring_rel32" unfolding mod_ring_rel32_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel32 (uint32_of_int z) z" unfolding urel32_def using small unfolding ppp by (auto simp: int_of_uint32_inv) with zy show "\ x z. urel32 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel32: "Domainp mod_ring_rel32 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel32 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel32_def proof let ?i = "int_of_uint32" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel32 x (?i x)" unfolding urel32_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel32 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel32_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops32: "ring_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, auto simp: finite_field_ops32_def bi_unique_mod_ring_rel32 right_total_mod_ring_rel32 mod_ring_plus32 mod_ring_minus32 mod_ring_uminus32 mod_ring_mult32 mod_ring_eq32 mod_ring_0_32 mod_ring_1_32 Domainp_mod_ring_rel32) end end context prime_field begin context fixes pp :: "uint32" assumes *: "p = int_of_uint32 pp" "p \ 65535" begin lemma mod_ring_normalize32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. if x = 0 then 0 else 1) normalize" using urel32_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_mod32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (\x y. if y = 0 then x else 0) (mod)" using urel32_mod[OF *] mod_ring_mod unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_inverse32: "(mod_ring_rel32 ===> mod_ring_rel32) (inverse_p32 pp) inverse" using urel32_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_divide32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (divide_p32 pp) (/)" using mod_ring_inverse32 mod_ring_mult32[OF *] unfolding divide_p32_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops32: "field_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, insert ring_finite_field_ops32[OF *], auto simp: ring_ops_def finite_field_ops32_def mod_ring_divide32 mod_ring_inverse32 mod_ring_mod32 mod_ring_normalize32) end end (* now there is 64-bit time *) context fixes p :: uint64 begin definition plus_p64 :: "uint64 \ uint64 \ uint64" where "plus_p64 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p64 :: "uint64 \ uint64 \ uint64" where "minus_p64 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p64 :: "uint64 \ uint64" where "uminus_p64 x = (if x = 0 then 0 else p - x)" definition mult_p64 :: "uint64 \ uint64 \ uint64" where "mult_p64 x y = (x * y mod p)" lemma int_of_uint64_shift: "int_of_uint64 (drop_bit k n) = (int_of_uint64 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint64_0_iff: "int_of_uint64 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint64_0: "int_of_uint64 0 = 0" unfolding int_of_uint64_0_iff by simp lemma int_of_uint64_ge_0: "int_of_uint64 n \ 0" by (transfer, auto) lemma two_64: "2 ^ LENGTH(64) = (18446744073709551616 :: int)" by simp lemma int_of_uint64_plus: "int_of_uint64 (x + y) = (int_of_uint64 x + int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_minus: "int_of_uint64 (x - y) = (int_of_uint64 x - int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mult: "int_of_uint64 (x * y) = (int_of_uint64 x * int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mod: "int_of_uint64 (x mod y) = (int_of_uint64 x mod int_of_uint64 y)" by (transfer, unfold uint_mod two_64, rule refl) lemma int_of_uint64_inv: "0 \ x \ x < 18446744073709551616 \ int_of_uint64 (uint64_of_int x) = x" by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p64 :: "uint64 \ uint64 \ uint64" where "power_p64 x n = (if n = 0 then 1 else let rec = power_p64 (mult_p64 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p64 rec x)" by pat_completeness auto termination proof - { fix n :: uint64 assume "n \ 0" with int_of_uint64_ge_0[of n] int_of_uint64_0_iff[of n] have "int_of_uint64 n > 0" by auto hence "0 < int_of_uint64 n" "int_of_uint64 n div 2 < int_of_uint64 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint64 n))", auto simp: int_of_uint64_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p64 :: "uint64 \ uint64" where "inverse_p64 x = (if x = 0 then 0 else power_p64 x (p - 2))" definition divide_p64 :: "uint64 \ uint64 \ uint64" where "divide_p64 x y = mult_p64 x (inverse_p64 y)" definition finite_field_ops64 :: "uint64 arith_ops_record" where "finite_field_ops64 \ Arith_Ops_Record 0 1 plus_p64 mult_p64 minus_p64 uminus_p64 divide_p64 inverse_p64 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint64_of_int int_of_uint64 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint64_code [code_unfold]: "drop_bit 1 x = (uint64_shiftr x 1)" by (simp add: uint64_shiftr_def) text \For soundness of the 64-bit implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "uint64" assumes ppp: "p = int_of_uint64 pp" and small: "p \ 4294967295" begin lemmas uint64_simps = int_of_uint64_0 int_of_uint64_plus int_of_uint64_minus int_of_uint64_mult definition urel64 :: "uint64 \ int \ bool" where "urel64 x y = (y = int_of_uint64 x \ y < p)" definition mod_ring_rel64 :: "uint64 \ 'a mod_ring \ bool" where "mod_ring_rel64 x y = (\ z. urel64 x z \ mod_ring_rel z y)" lemma urel64_0: "urel64 0 0" unfolding urel64_def using p2 by (simp, transfer, simp) lemma urel64_1: "urel64 1 1" unfolding urel64_def using p2 by (simp, transfer, simp) lemma le_int_of_uint64: "(x \ y) = (int_of_uint64 x \ int_of_uint64 y)" by (transfer, simp add: word_le_def) lemma urel64_plus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (plus_p64 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" let ?p = "int_of_uint64 pp" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_minus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (minus_p64 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_uminus: assumes "urel64 x y" shows "urel64 (uminus_p64 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint64 x" from assms int_of_uint64_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel64_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint64_0_iff using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_mult: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (mult_p64 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel64_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 4294967296 * 4294967296" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 18446744073709551616" by simp show ?thesis unfolding id using small rel unfolding mult_p64_def mult_p_def Let_def urel64_def unfolding ppp by (auto simp: uint64_simps, unfold int_of_uint64_mod int_of_uint64_mult, subst mod_pos_pos_trivial[of _ 18446744073709551616], insert le, auto) qed lemma urel64_eq: assumes "urel64 x y" "urel64 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel64_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel64_normalize: assumes x: "urel64 x y" shows "urel64 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel64_eq[OF x urel64_0] using urel64_0 urel64_1 by auto lemma urel64_mod: assumes x: "urel64 x x'" and y: "urel64 y y'" shows "urel64 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel64_eq[OF y urel64_0] using urel64_0 x by auto lemma urel64_power: "urel64 x x' \ urel64 y (int y') \ urel64 (power_p64 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel64_eq[OF y urel64_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel64_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel64_eq[OF y urel64_0] by auto from y have \int y' = int_of_uint64 y\ \int y' < p\ by (simp_all add: urel64_def) - obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force - from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Rings.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Rings.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel64 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel64_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint64.rep_eq modulo_uint64.rep_eq uint_mod \int y' = int_of_uint64 y\) done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel64_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p64.simps[of _ _ y] dr' id if_False rem using IH urel64_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel64_inverse: assumes x: "urel64 x x'" shows "urel64 (inverse_p64 pp x) (inverse_p p x')" proof - have p: "urel64 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel64_def unfolding ppp by (simp add: int_of_uint64.rep_eq minus_uint64.rep_eq uint_sub_if') show ?thesis unfolding inverse_p64_def inverse_p_def urel64_eq[OF x urel64_0] using urel64_0 urel64_power[OF x p] by auto qed lemma mod_ring_0_64: "mod_ring_rel64 0 0" using urel64_0 mod_ring_0 unfolding mod_ring_rel64_def by blast lemma mod_ring_1_64: "mod_ring_rel64 1 1" using urel64_1 mod_ring_1 unfolding mod_ring_rel64_def by blast lemma mod_ring_uminus64: "(mod_ring_rel64 ===> mod_ring_rel64) (uminus_p64 pp) uminus" using urel64_uminus mod_ring_uminus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_plus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (plus_p64 pp) (+)" using urel64_plus mod_ring_plus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_minus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (minus_p64 pp) (-)" using urel64_minus mod_ring_minus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_mult64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (mult_p64 pp) ((*))" using urel64_mult mod_ring_mult unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_eq64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> (=)) (=) (=)" using urel64_eq mod_ring_eq unfolding mod_ring_rel64_def rel_fun_def by blast lemma urel64_inj: "urel64 x y \ urel64 x z \ y = z" using urel64_eq[of x y x z] by auto lemma urel64_inj': "urel64 x z \ urel64 y z \ x = y" using urel64_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel64: "bi_unique mod_ring_rel64" "left_unique mod_ring_rel64" "right_unique mod_ring_rel64" using bi_unique_mod_ring_rel urel64_inj' unfolding mod_ring_rel64_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel64_def) lemma right_total_mod_ring_rel64: "right_total mod_ring_rel64" unfolding mod_ring_rel64_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel64 (uint64_of_int z) z" unfolding urel64_def using small unfolding ppp by (auto simp: int_of_uint64_inv) with zy show "\ x z. urel64 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel64: "Domainp mod_ring_rel64 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel64 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel64_def proof let ?i = "int_of_uint64" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel64 x (?i x)" unfolding urel64_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel64 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel64_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops64: "ring_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, auto simp: finite_field_ops64_def bi_unique_mod_ring_rel64 right_total_mod_ring_rel64 mod_ring_plus64 mod_ring_minus64 mod_ring_uminus64 mod_ring_mult64 mod_ring_eq64 mod_ring_0_64 mod_ring_1_64 Domainp_mod_ring_rel64) end end context prime_field begin context fixes pp :: "uint64" assumes *: "p = int_of_uint64 pp" "p \ 4294967295" begin lemma mod_ring_normalize64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. if x = 0 then 0 else 1) normalize" using urel64_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_mod64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (\x y. if y = 0 then x else 0) (mod)" using urel64_mod[OF *] mod_ring_mod unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_inverse64: "(mod_ring_rel64 ===> mod_ring_rel64) (inverse_p64 pp) inverse" using urel64_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_divide64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (divide_p64 pp) (/)" using mod_ring_inverse64 mod_ring_mult64[OF *] unfolding divide_p64_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops64: "field_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, insert ring_finite_field_ops64[OF *], auto simp: ring_ops_def finite_field_ops64_def mod_ring_divide64 mod_ring_inverse64 mod_ring_mod64 mod_ring_normalize64) end end (* and a final implementation via integer *) context fixes p :: integer begin definition plus_p_integer :: "integer \ integer \ integer" where "plus_p_integer x y \ let z = x + y in if z \ p then z - p else z" definition minus_p_integer :: "integer \ integer \ integer" where "minus_p_integer x y \ if y \ x then x - y else (x + p) - y" definition uminus_p_integer :: "integer \ integer" where "uminus_p_integer x = (if x = 0 then 0 else p - x)" definition mult_p_integer :: "integer \ integer \ integer" where "mult_p_integer x y = (x * y mod p)" lemma int_of_integer_0_iff: "int_of_integer n = 0 \ n = 0" using integer_eqI by auto lemma int_of_integer_0: "int_of_integer 0 = 0" unfolding int_of_integer_0_iff by simp lemma int_of_integer_plus: "int_of_integer (x + y) = (int_of_integer x + int_of_integer y)" by simp lemma int_of_integer_minus: "int_of_integer (x - y) = (int_of_integer x - int_of_integer y)" by simp lemma int_of_integer_mult: "int_of_integer (x * y) = (int_of_integer x * int_of_integer y)" by simp lemma int_of_integer_mod: "int_of_integer (x mod y) = (int_of_integer x mod int_of_integer y)" by simp lemma int_of_integer_inv: "int_of_integer (integer_of_int x) = x" by simp lemma int_of_integer_shift: "int_of_integer (drop_bit k n) = (int_of_integer n) div (2 ^ k)" by transfer (simp add: int_of_integer_pow shiftr_integer_conv_div_pow2) context includes bit_operations_syntax begin function power_p_integer :: "integer \ integer \ integer" where "power_p_integer x n = (if n \ 0 then 1 else let rec = power_p_integer (mult_p_integer x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p_integer rec x)" by pat_completeness auto termination proof - { fix n :: integer assume "\ (n \ 0)" hence "n > 0" by auto hence "int_of_integer n > 0" by (simp add: less_integer.rep_eq) hence "0 < int_of_integer n" "int_of_integer n div 2 < int_of_integer n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_integer n))", auto simp: * int_of_integer_shift) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p_integer :: "integer \ integer" where "inverse_p_integer x = (if x = 0 then 0 else power_p_integer x (p - 2))" definition divide_p_integer :: "integer \ integer \ integer" where "divide_p_integer x y = mult_p_integer x (inverse_p_integer y)" definition finite_field_ops_integer :: "integer arith_ops_record" where "finite_field_ops_integer \ Arith_Ops_Record 0 1 plus_p_integer mult_p_integer minus_p_integer uminus_p_integer divide_p_integer inverse_p_integer (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) integer_of_int int_of_integer (\ x. 0 \ x \ x < p)" end lemma shiftr_integer_code [code_unfold]: "drop_bit 1 x = (integer_shiftr x 1)" unfolding shiftr_integer_code using integer_of_nat_1 by auto text \For soundness of the integer implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "integer" assumes ppp: "p = int_of_integer pp" begin lemmas integer_simps = int_of_integer_0 int_of_integer_plus int_of_integer_minus int_of_integer_mult definition urel_integer :: "integer \ int \ bool" where "urel_integer x y = (y = int_of_integer x \ y \ 0 \ y < p)" definition mod_ring_rel_integer :: "integer \ 'a mod_ring \ bool" where "mod_ring_rel_integer x y = (\ z. urel_integer x z \ mod_ring_rel z y)" lemma urel_integer_0: "urel_integer 0 0" unfolding urel_integer_def using p2 by simp lemma urel_integer_1: "urel_integer 1 1" unfolding urel_integer_def using p2 by simp lemma le_int_of_integer: "(x \ y) = (int_of_integer x \ int_of_integer y)" by (rule less_eq_integer.rep_eq) lemma urel_integer_plus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (plus_p_integer pp x x') (plus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" let ?p = "int_of_integer pp" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_minus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (minus_p_integer pp x x') (minus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_uminus: assumes "urel_integer x y" shows "urel_integer (uminus_p_integer pp x) (uminus_p p y)" proof - let ?x = "int_of_integer x" from assms have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel_integer_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_integer_0_iff using rel by auto show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma pp_pos: "int_of_integer pp > 0" using ppp nontriv[where 'a = 'a] unfolding p by (simp add: less_integer.rep_eq) lemma urel_integer_mult: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (mult_p_integer pp x x') (mult_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel_integer_def by auto from rel(1,3) have xx: "0 \ ?x * ?x'" by simp show ?thesis unfolding id using rel unfolding mult_p_integer_def mult_p_def Let_def urel_integer_def unfolding ppp mod_nonneg_pos_int[OF xx pp_pos] using xx pp_pos by simp qed lemma urel_integer_eq: assumes "urel_integer x y" "urel_integer x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" unfolding urel_integer_def by auto show ?thesis unfolding id integer_eq_iff .. qed lemma urel_integer_normalize: assumes x: "urel_integer x y" shows "urel_integer (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_1 by auto lemma urel_integer_mod: assumes x: "urel_integer x x'" and y: "urel_integer y y'" shows "urel_integer (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel_integer_eq[OF y urel_integer_0] using urel_integer_0 x by auto lemma urel_integer_power: "urel_integer x x' \ urel_integer y (int y') \ urel_integer (power_p_integer pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' \ 0") case True hence y: "y = 0" "y' = 0" using urel_integer_eq[OF y urel_integer_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel_integer_1) next case False hence id: "(y \ 0) = False" "(y' = 0) = False" using False y by (auto simp add: urel_integer_def not_le) (metis of_int_integer_of of_int_of_nat_eq of_nat_0_less_iff) - obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force - from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Rings.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Rings.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have aux: "\ y'. int (y' mod 2) = int y' mod 2" by presburger have "urel_integer (y AND 1) r'" unfolding r' using y unfolding urel_integer_def unfolding ppp apply (auto simp add: and_one_eq) apply (simp add: of_nat_mod) done from urel_integer_eq[OF this urel_integer_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel_integer (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel_integer_def unfolding ppp shiftr_integer_conv_div_pow2 by auto from id have "y' \ 0" by auto note IH = 1(1)[OF this refl dr'[symmetric] urel_integer_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p_integer.simps[of _ _ y] dr' id if_False rem using IH urel_integer_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel_integer_inverse: assumes x: "urel_integer x x'" shows "urel_integer (inverse_p_integer pp x) (inverse_p p x')" proof - have p: "urel_integer (pp - 2) (int (nat (p - 2)))" using p2 unfolding urel_integer_def unfolding ppp by auto show ?thesis unfolding inverse_p_integer_def inverse_p_def urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_power[OF x p] by auto qed lemma mod_ring_0__integer: "mod_ring_rel_integer 0 0" using urel_integer_0 mod_ring_0 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_1__integer: "mod_ring_rel_integer 1 1" using urel_integer_1 mod_ring_1 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_uminus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (uminus_p_integer pp) uminus" using urel_integer_uminus mod_ring_uminus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_plus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (plus_p_integer pp) (+)" using urel_integer_plus mod_ring_plus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_minus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (minus_p_integer pp) (-)" using urel_integer_minus mod_ring_minus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_mult_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (mult_p_integer pp) ((*))" using urel_integer_mult mod_ring_mult unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_eq_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> (=)) (=) (=)" using urel_integer_eq mod_ring_eq unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma urel_integer_inj: "urel_integer x y \ urel_integer x z \ y = z" using urel_integer_eq[of x y x z] by auto lemma urel_integer_inj': "urel_integer x z \ urel_integer y z \ x = y" using urel_integer_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel_integer: "bi_unique mod_ring_rel_integer" "left_unique mod_ring_rel_integer" "right_unique mod_ring_rel_integer" using bi_unique_mod_ring_rel urel_integer_inj' unfolding mod_ring_rel_integer_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel_integer_def) lemma right_total_mod_ring_rel_integer: "right_total mod_ring_rel_integer" unfolding mod_ring_rel_integer_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel_integer (integer_of_int z) z" unfolding urel_integer_def unfolding ppp by auto with zy show "\ x z. urel_integer x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel_integer: "Domainp mod_ring_rel_integer = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel_integer x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel_integer_def proof let ?i = "int_of_integer" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel_integer x (?i x)" unfolding urel_integer_def using * unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed next assume "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" then obtain b z where xz: "urel_integer x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel_integer_def unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed qed lemma ring_finite_field_ops_integer: "ring_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, auto simp: finite_field_ops_integer_def bi_unique_mod_ring_rel_integer right_total_mod_ring_rel_integer mod_ring_plus_integer mod_ring_minus_integer mod_ring_uminus_integer mod_ring_mult_integer mod_ring_eq_integer mod_ring_0__integer mod_ring_1__integer Domainp_mod_ring_rel_integer) end end context prime_field begin context fixes pp :: "integer" assumes *: "p = int_of_integer pp" begin lemma mod_ring_normalize_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. if x = 0 then 0 else 1) normalize" using urel_integer_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_mod_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (\x y. if y = 0 then x else 0) (mod)" using urel_integer_mod[OF *] mod_ring_mod unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_inverse_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (inverse_p_integer pp) inverse" using urel_integer_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_divide_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (divide_p_integer pp) (/)" using mod_ring_inverse_integer mod_ring_mult_integer[OF *] unfolding divide_p_integer_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops_integer: "field_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, insert ring_finite_field_ops_integer[OF *], auto simp: ring_ops_def finite_field_ops_integer_def mod_ring_divide_integer mod_ring_inverse_integer mod_ring_mod_integer mod_ring_normalize_integer) end end context prime_field begin (* four implementations of modular integer arithmetic for finite fields *) thm finite_field_ops64 finite_field_ops32 finite_field_ops_integer finite_field_ops_int end context mod_ring_locale begin (* four implementations of modular integer arithmetic for finite rings *) thm ring_finite_field_ops64 ring_finite_field_ops32 ring_finite_field_ops_integer ring_finite_field_ops_int end end diff --git a/thys/CRYSTALS-Kyber/Compress.thy b/thys/CRYSTALS-Kyber/Compress.thy --- a/thys/CRYSTALS-Kyber/Compress.thy +++ b/thys/CRYSTALS-Kyber/Compress.thy @@ -1,1130 +1,1137 @@ theory Compress imports Kyber_spec Mod_Plus_Minus Abs_Qr "HOL-Analysis.Finite_Cartesian_Product" begin context kyber_spec begin section \Compress and Decompress Functions\ text \Properties of the \mod+-\ function.\ lemma two_mid_lt_q: "2 * \real_of_int q/2\ < q" using oddE[OF prime_odd_int[OF q_prime q_gt_two]] by fastforce lemma mod_plus_minus_range_q: assumes "y \ {-\q/2\..\q/2\}" shows "y mod+- q = y" using mod_plus_minus_range[OF q_gt_zero, of y] unfolding mod_plus_minus_def proof (auto) have this': "y + \real_of_int q/2\ \ {0..real_of_int q/2\) mod q = (y + \real_of_int q/2\)" using mod_rangeE[OF this'] by auto then show "(y + \real_of_int q/2\) mod q - \real_of_int q/2\ = y" by auto qed text \Compression only works for $x \in \mathbb{Z}_q$ and outputs an integer in $\{0,\dots, 2^d-1\}$ , where $d$ is a positive integer with $d < \rceil\log_2 (q)\lceil$. For compression we omit the least important bits. Decompression rescales to the modulus q.\ definition compress :: "nat \ int \ int" where "compress d x = round (real_of_int (2^d * x) / real_of_int q) mod (2^d)" definition decompress :: "nat \ int \ int" where "decompress d x = round (real_of_int q * real_of_int x / real_of_int 2^d)" lemma compress_zero: "compress d 0 = 0" unfolding compress_def by auto +lemma compress_less: + \compress d x < 2 ^ d\ + by (simp add: compress_def) + lemma decompress_zero: "decompress d 0 = 0" unfolding decompress_def by auto text \Properties of the exponent $d$.\ lemma d_lt_logq: assumes "of_nat d < \(log 2 q)::real\" shows "d< log 2 q" using assms by linarith lemma twod_lt_q: assumes "of_nat d < \(log 2 q)::real\" shows "2 powr (real d) < of_int q" using assms less_log_iff[of 2 q d] d_lt_logq q_gt_zero by auto lemma prime_half: assumes "prime (p::int)" "p > 2" shows "\p / 2\ > \p / 2\" proof - have "odd p" using prime_odd_int[OF assms] . then have "\p / 2\ > p/2" by (smt (verit, best) cos_npi_int cos_zero_iff_int le_of_int_ceiling mult.commute times_divide_eq_right) then have "\p / 2\ < p/2" by (meson floor_less_iff less_ceiling_iff) then show ?thesis using \\p / 2\ > p/2\ by auto qed lemma break_point_gt_q_div_two: assumes "of_nat d < \(log 2 q)::real\" shows "\q-(q/(2*2^d))\ > \q/2\" proof - have "1/((2::real)^d) \ (1::real)" using one_le_power[of 2 d] by simp have "\q-(q/(2*2^d))\ \ q-(q/2)* (1/(2^d))" by simp moreover have "q-(q/2)* (1/(2^d)) \ q - q/2" using \1/((2::real)^d) \ (1::real)\ by (smt (z3) calculation divide_le_eq divide_nonneg_nonneg divide_self_if mult_left_mono of_int_nonneg times_divide_eq_right q_gt_zero) ultimately have "\q-(q/(2*2^d))\ \ \q/2\ " by linarith moreover have "\q/2\ > \q/2\" using prime_half[OF q_prime q_gt_two] . ultimately show ?thesis by auto qed lemma decompress_zero_unique: assumes "decompress d s = 0" "s \ {0..2^d - 1}" "of_nat d < \(log 2 q)::real\" shows "s = 0" proof - let ?x = "real_of_int q * real_of_int s / real_of_int 2^d + 1/2" have "0 \ ?x \ ?x < 1" using assms(1) unfolding decompress_def round_def using floor_correct[of ?x] by auto then have "real_of_int q * real_of_int s / real_of_int 2^d < 1/2" by linarith moreover have "real_of_int q / real_of_int 2^d > 1" using twod_lt_q[OF assms(3)] by (simp add: powr_realpow) ultimately have "real_of_int s < 1/2" by (smt (verit, best) divide_less_eq_1_pos field_sum_of_halves pos_divide_less_eq times_divide_eq_left) then show ?thesis using assms(2) by auto qed text \Range of compress and decompress functions\ lemma range_compress: assumes "x\{0..q-1}" "of_nat d < \(log 2 q)::real\" shows "compress d x \ {0..2^d - 1}" using compress_def by auto lemma range_decompress: assumes "x\{0..2^d - 1}" "of_nat d < \(log 2 q)::real\" shows "decompress d x \ {0..q-1}" using decompress_def assms proof (auto, goal_cases) case 1 then show ?case by (smt (verit, best) divide_eq_0_iff divide_numeral_1 less_divide_eq_1_pos mult_of_int_commute nonzero_mult_div_cancel_right of_int_eq_0_iff of_int_less_1_iff powr_realpow q_gt_zero q_nonzero round_0 round_mono twod_lt_q zero_less_power) next case 2 have "real_of_int q/2^d > 1" using twod_lt_q[OF assms(2)] by (simp add: powr_realpow) then have log: "real_of_int q - real_of_int q/2^d \ q-1" by simp have "x \ 2^d-1" using assms(1) by simp then have "real_of_int x \ 2^d - 1" by (simp add: int_less_real_le) then have "real_of_int q * real_of_int x / 2^d \ real_of_int q * (2^d-1) / 2^d" by (smt (verit, best) divide_strict_right_mono mult_less_cancel_left_pos of_int_pos q_gt_zero zero_less_power) also have "\ = real_of_int q * 2^d /2^d - real_of_int q/2^d" by (simp add: diff_divide_distrib right_diff_distrib) finally have "real_of_int q * real_of_int x / 2^d \ real_of_int q - real_of_int q/2^d" by simp then have "round (real_of_int q * real_of_int x / 2^d) \ round (real_of_int q - real_of_int q/2^d)" using round_mono by blast also have "\ \ q - 1" using log by (metis round_mono round_of_int) finally show ?case by auto qed text \Compression is a function qrom $\mathbb{Z} / q\mathbb{Z}$ to $\mathbb{Z} / (2^d)\mathbb{Z}$.\ lemma compress_in_range: assumes "x\{0..\q-(q/(2*2^d))\-1}" "of_nat d < \(log 2 q)::real\" shows "round (real_of_int (2^d * x) / real_of_int q) < 2^d " proof - have "(2::int)^d \ 0" by simp have "real_of_int x < real_of_int q - real_of_int q / (2 * 2^d)" using assms(1) less_ceiling_iff by auto then have "2^d * real_of_int x / real_of_int q < 2^d * (real_of_int q - real_of_int q / (2 * 2^d)) / real_of_int q" by (simp add: divide_strict_right_mono q_gt_zero) also have "\ = 2^d * ((real_of_int q / real_of_int q) - (real_of_int q / real_of_int q) / (2 * 2^d))" by (simp add:algebra_simps diff_divide_distrib) also have "\ = 2^d * (1 - 1/(2*2^d))" using q_nonzero by simp also have "\ = 2^d - 1/2" using \2^d \ 0\ by (simp add: right_diff_distrib') finally have "2^d * real_of_int x / real_of_int q < 2^d - (1::real)/(2::real)" by auto then show ?thesis unfolding round_def using floor_less_iff by fastforce qed text \When does the modulo operation in the compression function change the output? Only when \x \ \q-(q / (2*2^d))\\. Then we can determine that the compress function maps to zero. This is why we need the \mod+-\ in the definition of Compression. Otherwise the error bound would not hold.\ lemma compress_no_mod: assumes "x\{0..\q-(q / (2*2^d))\-1}" "of_nat d < \(log 2 q)::real\" shows "compress d x = round (real_of_int (2^d * x) / real_of_int q)" unfolding compress_def using compress_in_range[OF assms] assms(1) q_gt_zero by (smt (z3) atLeastAtMost_iff divide_nonneg_nonneg mod_pos_pos_trivial mult_less_cancel_left_pos of_int_nonneg of_nat_0_less_iff right_diff_distrib' round_0 round_mono zero_less_power) lemma compress_2d: assumes "x\{\q-(q/(2*2^d))\..q-1}" "of_nat d < \(log 2 q)::real\" shows "round (real_of_int (2^d * x) / real_of_int q) = 2^d " using assms proof - have "(2::int)^d \ 0" by simp have "round (real_of_int (2^d * x) / real_of_int q) \ 2^d" proof - have "real_of_int x \ real_of_int q - real_of_int q / (2 * 2^d)" using assms(1) ceiling_le_iff by auto then have "2^d * real_of_int x / real_of_int q \ 2^d * (real_of_int q - real_of_int q / (2 * 2^d)) / real_of_int q" using q_gt_zero by (simp add: divide_right_mono) also have "2^d * (real_of_int q - real_of_int q / (2 * 2^d)) / real_of_int q = 2^d * ((real_of_int q / real_of_int q) - (real_of_int q / real_of_int q) / (2 * 2^d))" by (simp add:algebra_simps diff_divide_distrib) also have "\ = 2^d * (1 - 1/(2*2^d))" using q_nonzero by simp also have "\ = 2^d - 1/2" using \2^d \ 0\ by (simp add: right_diff_distrib') finally have "2^d * real_of_int x / real_of_int q \ 2^d - (1::real)/(2::real)" by auto then show ?thesis unfolding round_def using le_floor_iff by force qed moreover have "round (real_of_int (2^d * x) / real_of_int q) \ 2^d" proof - have "d < log 2 q" using assms(2) by linarith then have "2 powr (real d) < of_int q" using less_log_iff[of 2 q d] q_gt_zero by auto have "x < q" using assms(1) by auto then have "real_of_int x/ real_of_int q < 1" by (simp add: q_gt_zero) then have "real_of_int (2^d * x) / real_of_int q < real_of_int (2^d)" by (auto) (smt (verit, best) divide_less_eq_1_pos nonzero_mult_div_cancel_left times_divide_eq_right zero_less_power) then show ?thesis unfolding round_def by linarith qed ultimately show ?thesis by auto qed lemma compress_mod: assumes "x\{\q-(q/(2*2^d))\..q-1}" "of_nat d < \(log 2 q)::real\" shows "compress d x = 0" unfolding compress_def using compress_2d[OF assms] by simp text \Error after compression and decompression of data. To prove the error bound, we distinguish the cases where the \mod+-\ is relevant or not.\ text \First let us look at the error bound for no \mod+-\ reduction.\ lemma decompress_compress_no_mod: assumes "x\{0..\q-(q/(2*2^d))\-1}" "of_nat d < \(log 2 q)::real\" shows "abs (decompress d (compress d x) - x) \ round ( real_of_int q / real_of_int (2^(d+1)))" proof - have "abs (decompress d (compress d x) - x) = abs (real_of_int (decompress d (compress d x)) - real_of_int q / real_of_int (2^d) * (real_of_int (2^d * x) / real_of_int q))" using q_gt_zero by force also have "\ \ abs (real_of_int (decompress d (compress d x)) - real_of_int q / real_of_int (2^d) * real_of_int (compress d x)) + abs (real_of_int q / real_of_int (2^d) * real_of_int (compress d x) - real_of_int q / real_of_int (2^d) * real_of_int (2^d) / real_of_int q * x)" using abs_triangle_ineq[of "real_of_int (decompress d (compress d x)) - real_of_int q / real_of_int (2^d) * real_of_int (compress d x)" "real_of_int q / real_of_int (2^d) * real_of_int (compress d x) - real_of_int q / real_of_int (2^d) * real_of_int (2^d) / real_of_int q * real_of_int x"] by auto also have "\ \ 1/2 + abs (real_of_int q / real_of_int (2^d) * (real_of_int (compress d x) - real_of_int (2^d) / real_of_int q * real_of_int x))" proof - have part_one: "abs (real_of_int (decompress d (compress d x)) - real_of_int q / real_of_int (2^d) * real_of_int (compress d x)) \ 1/2" unfolding decompress_def using of_int_round_abs_le[of "real_of_int q * real_of_int (compress d x) / real_of_int 2^d"] by simp have "real_of_int q * real_of_int (compress d x) / 2^d - real_of_int x = real_of_int q * (real_of_int (compress d x) - 2^d * real_of_int x / real_of_int q) / 2^d" by (smt (verit, best) divide_cancel_right nonzero_mult_div_cancel_left of_int_eq_0_iff q_nonzero right_diff_distrib times_divide_eq_left zero_less_power) then have part_two: "abs (real_of_int q / real_of_int (2^d) * real_of_int (compress d x) - real_of_int q / real_of_int (2^d) * real_of_int (2^d) / real_of_int q * x) = abs (real_of_int q / real_of_int (2^d) * (real_of_int (compress d x) - real_of_int (2^d) / real_of_int q * x)) " by auto show ?thesis using part_one part_two by auto qed also have "\ = 1/2 + (real_of_int q / real_of_int (2^d)) * abs (real_of_int (compress d x) - real_of_int (2^d) / real_of_int q * real_of_int x)" by (subst abs_mult) (smt (verit, best) assms(2) less_divide_eq_1_pos of_int_add of_int_hom.hom_one of_int_power powr_realpow twod_lt_q zero_less_power) also have "\ \ 1/2 + (real_of_int q / real_of_int (2^d)) * (1/2) " using compress_no_mod[OF assms] using of_int_round_abs_le[of "real_of_int (2^d) * real_of_int x / real_of_int q"] by (smt (verit, ccfv_SIG) divide_nonneg_nonneg left_diff_distrib mult_less_cancel_left_pos of_int_mult of_int_nonneg q_gt_zero times_divide_eq_left zero_le_power) finally have "real_of_int (abs (decompress d (compress d x) - x)) \ real_of_int q / real_of_int (2*2^d) + 1/2" by simp then show ?thesis unfolding round_def using le_floor_iff by fastforce qed lemma no_mod_plus_minus: assumes "abs y \ round ( real_of_int q / real_of_int (2^(d+1)))" "d>0" shows "abs y = abs (y mod+- q)" proof - have "round (real_of_int q / real_of_int (2^(d+1))) \ \q/2\" unfolding round_def proof - have "real_of_int q/real_of_int (2^d) \ real_of_int q/2" using \d>0\ proof - have "1 / real_of_int (2^d) \ 1/2" using \d>0\ inverse_of_nat_le[of 2 "2^d"] by (simp add: self_le_power) then show ?thesis using q_gt_zero by (smt (verit, best) frac_less2 of_int_le_0_iff zero_less_power) qed moreover have "real_of_int q/2 + 1 \ real_of_int q" using q_gt_two by auto ultimately have "real_of_int q / real_of_int (2^d) + 1 \ real_of_int q" by linarith then have fact: "real_of_int q / real_of_int (2 ^ (d + 1)) + 1/2 \ real_of_int q/2" by auto then show "\real_of_int q / real_of_int (2 ^ (d + 1)) + 1/2\ \ \real_of_int q/2\" using floor_mono[OF fact] by auto qed then have "abs y \ \q/2\" using assms by auto then show ?thesis using mod_plus_minus_range[OF q_gt_zero] using mod_plus_minus_def two_mid_lt_q by force qed lemma decompress_compress_no_mod_plus_minus: assumes "x\{0..\q-(q/(2*2^d))\-1}" "of_nat d < \(log 2 q)::real\" "d>0" shows "abs ((decompress d (compress d x) - x) mod+- q) \ round ( real_of_int q / real_of_int (2^(d+1)))" proof - have "abs ((decompress d (compress d x) - x) mod+- q) = abs ((decompress d (compress d x) - x)) " using no_mod_plus_minus[OF decompress_compress_no_mod [OF assms(1) assms(2)] assms(3)] by auto then show ?thesis using decompress_compress_no_mod [OF assms(1) assms(2)] by auto qed text \Now lets look at what happens when the \mod+-\ reduction comes into action.\ lemma ceiling_int: "\of_int a + b\ = a + \b\" unfolding ceiling_def by (simp add: add.commute) lemma decompress_compress_mod: assumes "x\{\q-(q/(2*2^d))\..q-1}" "of_nat d < \(log 2 q)::real\" shows "abs ((decompress d (compress d x) - x) mod+- q) \ round ( real_of_int q / real_of_int (2^(d+1)))" proof - have "(decompress d (compress d x) - x) = - x" using compress_mod[OF assms] unfolding decompress_def by auto moreover have "-x mod+- q = -x+q" proof - have "(-x + \q/2\) + q < q" using assms(1) break_point_gt_q_div_two[OF assms(2)] by force moreover have "(-x + \q/2\) + q \ 0 " using assms(1) q_gt_zero by (smt (verit, best) atLeastAtMost_iff divide_nonneg_nonneg of_int_nonneg zero_le_floor) ultimately have "(- x + \q/2\) mod q = - x + \q/2\ + q" by (metis mod_add_self2 mod_pos_pos_trivial) then show ?thesis unfolding mod_plus_minus_def by auto qed moreover have "abs (q - x) \ round ( real_of_int q / real_of_int (2^(d+1)))" proof - have "abs (q-x) = q-x" using assms(1) by auto also have "\ \ q - \q - q/(2*2^d)\" using assms(1) by simp also have "\ = - \- q/(2*2^d)\" using ceiling_int[of q "- q/(2*2^d)"] by auto also have "\ = \q/(2*2^d)\" by (simp add: ceiling_def) also have "\ \ round (q/(2*2^d))" using floor_le_round by blast finally have "abs (q-x) \ round (q/(2^(d+1)))" by auto then show ?thesis by auto qed ultimately show ?thesis by auto qed text \Together, we can determine the general error bound on decompression of compression of the data. This error needs to be small enough not to disturb the encryption and decryption process.\ lemma decompress_compress: assumes "x\{0..(log 2 q)::real\" "d>0" shows "let x' = decompress d (compress d x) in abs ((x' - x) mod+- q) \ round ( real_of_int q / real_of_int (2^(d+1)) )" proof (cases "x<\q-(q/(2*2^d))\") case True then have range_x: "x\{0..\q-(q/(2*2^d))\-1}" using assms(1) by auto show ?thesis unfolding Let_def using decompress_compress_no_mod_plus_minus[OF range_x assms(2) assms(3)] by auto next case False then have range_x: "x\{\q-(q/(2*2^d))\..q-1}" using assms(1) by auto show ?thesis unfolding Let_def using decompress_compress_mod[OF range_x assms(2)] by auto qed text \We have now defined compression only on integers (ie \{0.., corresponding to \\_q\). We need to extend this notion to the ring \\_q[X]/(X^n+1)\. Here, a compressed polynomial is the compression on every coefficient.\ text \ How to channel through the types \begin{itemize} \item \to_qr :: 'a mod_ring poly \ 'a qr\ \item \Poly :: 'a mod_ring list \ 'a mod_ring poly\ \item \map of_int_mod_ring :: int list \ 'a mod_ring list\ \item \map compress :: int list \ int list\ \item \map to_int_mod_ring :: 'a mod_ring list \ int list\ \item \coeffs :: 'a mod_ring poly \ 'a mod_ring list\ \item \of_qr :: 'a qr \ 'a mod_ring poly\ \end{itemize} \ definition compress_poly :: "nat \ 'a qr \ 'a qr" where "compress_poly d = to_qr \ Poly \ (map of_int_mod_ring) \ (map (compress d)) \ (map to_int_mod_ring) \ coeffs \ of_qr" definition decompress_poly :: "nat \ 'a qr \ 'a qr" where "decompress_poly d = to_qr \ Poly \ (map of_int_mod_ring) \ (map (decompress d)) \ (map to_int_mod_ring) \ coeffs \ of_qr" text \Lemmas for compression error for polynomials. Lemma telescope to go qrom module level down to integer coefficients and back up again.\ lemma deg_Poly': assumes "Poly xs \ 0" shows "degree (Poly xs) \ length xs - 1" proof (induct xs) case (Cons a xs) then show ?case by simp (metis Poly.simps(1) Suc_le_eq Suc_pred le_imp_less_Suc length_greater_0_conv) qed simp lemma of_int_mod_ring_eq_0: "((of_int_mod_ring x :: 'a mod_ring) = 0) \ (x mod q = 0)" by (metis CARD_a mod_0 of_int_code(2) of_int_mod_ring.abs_eq of_int_mod_ring.rep_eq of_int_of_int_mod_ring) lemma dropWhile_mod_ring: "dropWhile ((=)0) (map of_int_mod_ring xs :: 'a mod_ring list) = map of_int_mod_ring (dropWhile (\x. x mod q = 0) xs)" proof (induct xs) case (Cons x xs) have "dropWhile ((=) 0) (map of_int_mod_ring (x # xs)) = dropWhile ((=) 0) ((of_int_mod_ring x :: 'a mod_ring) # (map of_int_mod_ring xs))" by auto also have "\ = (if 0 = (of_int_mod_ring x :: 'a mod_ring) then dropWhile ((=) 0) (map of_int_mod_ring xs) else map of_int_mod_ring (x # xs))" unfolding dropWhile.simps(2)[of "((=) 0)" "of_int_mod_ring x :: 'a mod_ring" "map of_int_mod_ring xs"] by auto also have "\ = (if x mod q = 0 then map of_int_mod_ring (dropWhile (\x. x mod q = 0) xs) else map of_int_mod_ring (x # xs))" using of_int_mod_ring_eq_0 unfolding Cons.hyps by auto also have "\ = map of_int_mod_ring (dropWhile (\x. x mod q = 0) (x # xs))" unfolding dropWhile.simps(2) by auto finally show ?case by blast qed simp lemma strip_while_mod_ring: "(strip_while ((=) 0) (map of_int_mod_ring xs :: 'a mod_ring list)) = map of_int_mod_ring (strip_while (\x. x mod q = 0) xs)" unfolding strip_while_def comp_def rev_map dropWhile_mod_ring by auto lemma of_qr_to_qr_Poly: assumes "length (xs :: int list) < Suc (nat n)" "xs \ []" shows "of_qr (to_qr (Poly (map (of_int_mod_ring :: int \ 'a mod_ring) xs))) = Poly (map (of_int_mod_ring :: int \ 'a mod_ring) xs)" (is "_ = ?Poly") proof - have deg: "degree (?Poly) < n" using deg_Poly'[of "map of_int_mod_ring xs"] assms by (smt (verit, del_insts) One_nat_def Suc_pred degree_0 length_greater_0_conv length_map less_Suc_eq_le order_less_le_trans zless_nat_eq_int_zless) then show ?thesis using of_qr_to_qr[of "?Poly"] deg_mod_qr_poly[of "?Poly"] deg_qr_n by (smt (verit, best) of_nat_less_imp_less) qed lemma telescope_stripped: assumes "length (xs :: int list) < Suc (nat n)" "strip_while (\x. x mod q = 0) xs = xs" "set xs \ {0.. 'a mod_ring) xs))))) = xs" proof (cases "xs = []") case False have ge_zero: "0\x" and lt_q:"x < int CARD ('a)" if "x\set xs" for x using assms(3) CARD_a atLeastLessThan_iff that by auto have to_int_of_int: "map (to_int_mod_ring \ (of_int_mod_ring :: int \ 'a mod_ring)) xs = xs" using to_int_mod_ring_of_int_mod_ring[OF ge_zero lt_q] by (simp add: map_idI) show ?thesis using assms(2) of_qr_to_qr_Poly[OF assms(1) False] by (auto simp add: to_int_of_int strip_while_mod_ring) qed (simp) lemma map_to_of_mod_ring: assumes "set xs \ {0.. (of_int_mod_ring :: int \ 'a mod_ring)) xs = xs" using assms by (induct xs) (simp_all add: CARD_a) lemma telescope: assumes "length (xs :: int list) < Suc (nat n)" "set xs \ {0.. 'a mod_ring) xs))))) = strip_while (\x. x mod q = 0) xs" proof (cases "xs = strip_while (\x. x mod q = 0) xs") case True then show ?thesis using telescope_stripped assms by auto next case False let ?of_int = "(map (of_int_mod_ring :: int \ 'a mod_ring) xs)" have "xs \ []" using False by auto then have "(map to_int_mod_ring) (coeffs (of_qr (to_qr (Poly ?of_int)))) = (map to_int_mod_ring) (coeffs (Poly ?of_int))" using of_qr_to_qr_Poly[OF assms(1)] by auto also have "\ = (map to_int_mod_ring) (strip_while ((=) 0) ?of_int)" by auto also have "\ = map (to_int_mod_ring \ (of_int_mod_ring :: int \ 'a mod_ring)) (strip_while (\x. x mod q = 0) xs)" using strip_while_mod_ring by auto also have "\ = strip_while (\x. x mod q = 0) xs" using assms(2) proof (induct xs rule: rev_induct) case (snoc y ys) let ?to_of_mod_ring = "to_int_mod_ring \ (of_int_mod_ring :: int \ 'a mod_ring)" have "map ?to_of_mod_ring (strip_while (\x. x mod q = 0) (ys @ [y])) = (if y mod q = 0 then map ?to_of_mod_ring (strip_while (\x. x mod q = 0) ys) else map ?to_of_mod_ring ys @ [?to_of_mod_ring y])" by (subst strip_while_snoc) auto also have "\ = (if y mod q = 0 then strip_while (\x. x mod q = 0) ys else map ?to_of_mod_ring ys @ [?to_of_mod_ring y])" using snoc by fastforce also have "\ = (if y mod q = 0 then strip_while (\x. x mod q = 0) ys else ys @ [y])" using map_to_of_mod_ring[OF snoc(2)] by auto also have "\ = strip_while (\x. x mod q = 0) (ys @ [y])" by auto finally show ?case . qed simp finally show ?thesis by auto qed lemma length_coeffs_of_qr: "length (coeffs (of_qr (x ::'a qr))) < Suc (nat n)" proof (cases "x=0") case False then have "of_qr x \ 0" by simp then show ?thesis using length_coeffs_degree[of "of_qr x"] deg_of_qr[of x] using deg_qr_n by fastforce qed (auto simp add: n_gt_zero) lemma strip_while_change: assumes "\x. P x \ S x" "\x. (\ P x) \ (\ S x)" shows "strip_while P xs = strip_while S xs" proof (induct xs rule: rev_induct) case (snoc x xs) have "P x = S x" using assms[of x] by blast then show ?case by (simp add: snoc.hyps) qed simp lemma strip_while_change_subset: assumes "set xs \ s" "\x\s. P x \ S x" "\x\s. (\ P x) \ (\ S x)" shows "strip_while P xs = strip_while S xs" using assms(1) proof (induct xs rule: rev_induct) case (snoc x xs) have "x\s" using snoc(2) by simp then have "P x \ S x" and "(\ P x) \ (\ S x)" using assms(2) assms(3) by auto then have "P x = S x" by blast then show ?case using snoc.hyps snoc.prems by auto qed simp text \Estimate for decompress compress for polynomials. Using the inequality for integers, chain it up to the level of polynomials.\ lemma decompress_compress_poly: assumes "of_nat d < \(log 2 q)::real\" "d>0" shows "let x' = decompress_poly d (compress_poly d x) in abs_infty_poly (x - x') \ round ( real_of_int q / real_of_int (2^(d+1)) )" proof - let ?x' = "decompress_poly d (compress_poly d x)" have "abs_infty_q (poly.coeff (of_qr (x - ?x')) xa) \ round (real_of_int q / real_of_int (2 ^ (d + 1)))" for xa proof - let ?telescope = "(\xs. (map to_int_mod_ring) (coeffs (of_qr (to_qr (Poly (map (of_int_mod_ring :: int \ 'a mod_ring) xs))))))" define compress_x where "compress_x = map (compress d \ to_int_mod_ring) (coeffs (of_qr x))" let ?to_Poly = "(\a. Poly (map ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d) a))" have "abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (of_qr (to_qr (?to_Poly (?telescope compress_x)))) xa ) = abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (of_qr (to_qr (?to_Poly (strip_while (\x. x = 0) compress_x)))) xa )" proof (cases "x = 0") case True then have "compress_x = []" unfolding compress_x_def by auto then show ?thesis by simp next case False then have nonempty:"compress_x \ []" unfolding compress_x_def by simp have "length compress_x < Suc (nat n)" by (auto simp add: compress_x_def length_coeffs_of_qr) moreover have "set compress_x \ {0.. {0..q - 1}" for s using to_int_mod_ring_range by auto have "compress d (to_int_mod_ring (s::'a mod_ring)) \ {0..x. x mod q = 0) compress_x" by (intro telescope[of "compress_x"]) simp moreover have "strip_while (\x. x mod q = 0) compress_x = strip_while (\x. x = 0) compress_x" proof - - have "compress d s = 0" - if "compress d s mod q = 0" for s - using that unfolding compress_def using twod_lt_q - by (smt (verit, ccfv_threshold) - Euclidean_Division.pos_mod_bound - Euclidean_Division.pos_mod_sign assms(1) - compress_def mod_pos_pos_trivial of_int_add - of_int_hom.hom_one of_int_power_less_of_int_cancel_iff - powr_realpow that zero_less_power) + have \compress d s = 0\ if \compress d s mod q = 0\ for s + proof - + from \int d < \log 2 (real_of_int q)\\ twod_lt_q [of d] + have \2 ^ d < q\ + by (simp add: powr_realpow) + with compress_less [of d s] have \compress d s < q\ + by simp + then have \compress d s = compress d s mod q\ + by (simp add: compress_def) + with that show ?thesis + by simp + qed then have right: "\s. compress d s mod q = 0 \ compress d s = 0" by simp have "\ (compress d s = 0)" if "\ (compress d s mod q = 0)" for s using twod_lt_q compress_def that by force then have left: "\s. \ (compress d s mod q = 0) \ \ (compress d s = 0)" by simp have "strip_while (\x. x mod q = 0) compress_x = strip_while (\x. x mod q = 0) (map (compress d) (map to_int_mod_ring (coeffs (of_qr x))))" (is "_ = strip_while (\x. x mod q = 0) (map (compress d) ?rest)") unfolding compress_x_def by simp also have "\ = map (compress d) (strip_while ((\y. y mod q = 0) \ compress d) (map to_int_mod_ring (coeffs (of_qr x))))" using strip_while_map[of "\y. y mod q = 0" "compress d"] by blast also have "\ = map (compress d) (strip_while ((\y. y = 0) \ compress d) (map to_int_mod_ring (coeffs (of_qr x))))" by (smt (verit, best) comp_eq_dest_lhs left right strip_while_change) also have "\ = strip_while (\x. x = 0) (map (compress d) ?rest)" using strip_while_map[of "\y. y = 0" "compress d", symmetric] by blast finally show ?thesis unfolding compress_x_def by auto qed ultimately show ?thesis by auto qed also have "\ = abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (?to_Poly (strip_while (\x. x = 0) compress_x)) xa)" proof (cases "?to_Poly (strip_while (\x. x = 0) compress_x) = 0") case False have "degree (?to_Poly (strip_while (\x. x = 0) compress_x)) \ length (map ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d) (strip_while (\x. x = 0) compress_x)) - 1" using deg_Poly'[OF False] . moreover have "length (map (of_int_mod_ring \ decompress d) (strip_while (\x. x = 0) compress_x)) \ length (coeffs (of_qr x))" unfolding compress_x_def by (metis length_map length_strip_while_le) moreover have "length (coeffs (of_qr x)) - 1 < deg_qr TYPE('a)" using deg_of_qr degree_eq_length_coeffs by metis ultimately have deg: "degree (?to_Poly (strip_while (\x. x = 0) compress_x)) < deg_qr TYPE('a)" by auto show ?thesis using of_qr_to_qr' by (simp add: of_qr_to_qr'[OF deg]) qed simp also have "\ = abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (Poly (map of_int_mod_ring (strip_while (\x. x = 0) (map (decompress d) compress_x)))) xa )" proof - have "s = 0" if "decompress d s = 0" "s \ {0..2^d - 1}" for s using decompress_zero_unique[OF that assms(1)] . then have right: "\s \ {0..2^d-1}. decompress d s = 0 \ s = 0" by simp have left: "\ s \ {0..2^d-1}. decompress d s \ 0 \ s \ 0" using decompress_zero[of d] by auto have compress_x_range: "set compress_x \ {0..2^d - 1}" unfolding compress_x_def compress_def by auto have "map (decompress d) (strip_while (\x. x = 0) compress_x) = map (decompress d) (strip_while (\x. decompress d x = 0) compress_x)" using strip_while_change_subset[OF compress_x_range right left] by auto also have "\ = strip_while (\x. x = 0) (map (decompress d) compress_x)" proof - have "(\x. x = 0) \ decompress d = (\x. decompress d x = 0)" using comp_def by blast then show ?thesis using strip_while_map[symmetric, of "decompress d" "\x. x=0" compress_x] by auto qed finally have "map (decompress d) (strip_while (\x. x = 0) compress_x) = strip_while (\x. x = 0) (map (decompress d) compress_x)" by auto then show ?thesis by (metis map_map) qed also have "\ = abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (Poly (map of_int_mod_ring (strip_while (\x. x mod q = 0) (map (decompress d) compress_x)))) xa )" proof - have range: "set (map (decompress d) compress_x) \ {0..x\{0.. x mod q = 0" by auto have left: "\x\{0.. x = 0 \ \ x mod q = 0" by auto have "strip_while (\x. x = 0) (map (decompress d) compress_x) = strip_while (\x. x mod q = 0) (map (decompress d) compress_x)" using strip_while_change_subset[OF range right left] by auto then show ?thesis by auto qed also have "\ = abs_infty_q (poly.coeff (of_qr x) xa - poly.coeff (Poly (map of_int_mod_ring (map (decompress d) compress_x))) xa )" by (metis Poly_coeffs coeffs_Poly strip_while_mod_ring) also have "\ = abs_infty_q (poly.coeff (of_qr x) xa - ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa))" using coeffs_Poly proof (cases "xa < length (coeffs (?to_Poly compress_x))") case True have "poly.coeff (?to_Poly compress_x) xa = coeffs (?to_Poly compress_x) ! xa" using nth_coeffs_coeff[OF True] by simp also have "\ = strip_while ((=) 0) (map ( (of_int_mod_ring :: int \ 'a mod_ring) \ decompress d) compress_x) ! xa" using coeffs_Poly by auto also have "\ = (map ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d) compress_x) ! xa" using True by (metis coeffs_Poly nth_strip_while) also have "\ = ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (coeffs (of_qr x) ! xa)" unfolding compress_x_def by (smt (z3) True coeffs_Poly compress_x_def length_map length_strip_while_le map_map not_less nth_map order_trans) also have "\ = ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa)" by (metis (no_types, lifting) True coeffs_Poly compress_x_def length_map length_strip_while_le not_less nth_coeffs_coeff order.trans) finally have no_coeff: "poly.coeff (?to_Poly compress_x) xa = ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa)" by auto show ?thesis unfolding compress_x_def by (metis compress_x_def map_map no_coeff) next case False then have "poly.coeff (?to_Poly compress_x) xa = 0" by (metis Poly_coeffs coeff_Poly_eq nth_default_def) moreover have "((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa) = 0" proof (cases "poly.coeff (of_qr x) xa = 0") case True then show ?thesis using compress_zero decompress_zero by auto next case False then show ?thesis proof (subst ccontr, goal_cases) case 1 then have "poly.coeff (?to_Poly compress_x) xa \ 0" by (subst coeff_Poly) (metis (no_types, lifting) comp_apply compress_x_def compress_zero decompress_zero map_map nth_default_coeffs_eq nth_default_map_eq of_int_mod_ring_hom.hom_zero to_int_mod_ring_hom.hom_zero) then show ?case using \poly.coeff (?to_Poly compress_x) xa = 0\ by auto qed auto qed ultimately show ?thesis by auto qed also have "\ = abs_infty_q ( ((of_int_mod_ring :: int \ 'a mod_ring) \ decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa) - poly.coeff (of_qr x) xa)" using abs_infty_q_minus by (metis minus_diff_eq) also have "\ = \((decompress d \ compress d \ to_int_mod_ring) (poly.coeff (of_qr x) xa) - to_int_mod_ring (poly.coeff (of_qr x) xa)) mod+- q\" unfolding abs_infty_q_def using to_int_mod_ring_of_int_mod_ring by (smt (verit, best) CARD_a comp_apply mod_plus_minus_def of_int_diff of_int_mod_ring.rep_eq of_int_mod_ring_to_int_mod_ring of_int_of_int_mod_ring) also have "\ \ round (real_of_int q / real_of_int (2 ^ (d + 1)))" proof - have range_to_int_mod_ring: "to_int_mod_ring (poly.coeff (of_qr x) xa) \ {0.. round (real_of_int q / real_of_int (2 ^ (d + 1)))" by auto then show ?thesis unfolding compress_x_def decompress_poly_def compress_poly_def by (auto simp add: o_assoc) qed moreover have finite: "finite (range (abs_infty_q \ poly.coeff (of_qr (x - ?x'))))" by (metis finite_Max image_comp image_image) ultimately show ?thesis unfolding abs_infty_poly_def using Max_le_iff[OF finite] by auto qed text \More properties of compress and decompress, used for returning message at the end.\ lemma compress_1: shows "compress 1 x \ {0,1}" unfolding compress_def by auto lemma compress_poly_1: shows "\i. poly.coeff (of_qr (compress_poly 1 x)) i \ {0,1}" proof - have "poly.coeff (of_qr (compress_poly 1 x)) i \ {0,1}" for i proof - have "set (map (compress 1) ((map to_int_mod_ring \ coeffs \ of_qr) x)) \ {0,1}" using compress_1 by auto then have "set ((map (compress 1) \ map to_int_mod_ring \ coeffs \ of_qr) x) \ {0,1}" (is "set (?compressed_1) \ _") by auto then have "set (map (of_int_mod_ring :: int \ 'a mod_ring) ?compressed_1) \ {0,1}" (is "set (?of_int_compressed_1)\_") by (smt (verit, best) imageE insert_iff of_int_mod_ring_hom.hom_zero of_int_mod_ring_to_int_mod_ring set_map singletonD subsetD subsetI to_int_mod_ring_hom.hom_one) then have "nth_default 0 (?of_int_compressed_1) i \ {0,1}" by (smt (verit, best) comp_apply compress_1 compress_zero insert_iff nth_default_map_eq of_int_mod_ring_hom.hom_zero of_int_mod_ring_to_int_mod_ring singleton_iff to_int_mod_ring_hom.hom_one) moreover have "Poly (?of_int_compressed_1) = Poly (?of_int_compressed_1) mod qr_poly" proof - have "degree (Poly (?of_int_compressed_1)) < deg_qr TYPE('a)" proof (cases "Poly ?of_int_compressed_1 \ 0") case True have "degree (Poly ?of_int_compressed_1) \ length (map (of_int_mod_ring :: int \ 'a mod_ring) ?compressed_1) - 1" using deg_Poly'[OF True] by simp also have "\ = length ((coeffs \ of_qr) x) - 1" by simp also have "\ < n" unfolding comp_def using length_coeffs_of_qr by (metis deg_qr_n deg_of_qr degree_eq_length_coeffs nat_int zless_nat_conj) finally have "degree (Poly ?of_int_compressed_1) < n" using True \int (length ((coeffs \ of_qr) x) - 1) < n\ deg_Poly' by fastforce then show ?thesis using deg_qr_n by simp next case False then show ?thesis using deg_qr_pos by auto qed then show ?thesis using deg_mod_qr_poly[of "Poly (?of_int_compressed_1)", symmetric] by auto qed ultimately show ?thesis unfolding compress_poly_def comp_def using of_qr_to_qr[of "Poly (?of_int_compressed_1)"] by auto qed then show ?thesis by auto qed lemma of_int_mod_ring_mult: "of_int_mod_ring (a*b) = of_int_mod_ring a * of_int_mod_ring b" unfolding of_int_mod_ring_def by (metis (mono_tags, lifting) Rep_mod_ring_inverse mod_mult_eq of_int_mod_ring.rep_eq of_int_mod_ring_def times_mod_ring.rep_eq) lemma decompress_1: assumes "a\{0,1}" shows "decompress 1 a = round(real_of_int q/2) * a" unfolding decompress_def using assms by auto lemma decompress_poly_1: assumes "\i. poly.coeff (of_qr x) i \ {0,1}" shows "decompress_poly 1 x = to_module (round((real_of_int q)/2)) * x" proof - have "poly.coeff (of_qr (decompress_poly 1 x)) i = poly.coeff (of_qr (to_module (round((real_of_int q)/2)) * x)) i" for i proof - have "set (map to_int_mod_ring (coeffs (of_qr x))) \ {0,1}" (is "set (?int_coeffs) \ _") proof - have "set (coeffs (of_qr x)) \ {0,1}" using assms by (meson forall_coeffs_conv insert_iff subset_code(1)) then show ?thesis by auto qed then have "map (decompress 1) (?int_coeffs) = map ((*) (round (real_of_int q/2))) (?int_coeffs)" proof (induct "?int_coeffs") case (Cons a xa) then show ?case using decompress_1 by (meson map_eq_conv subsetD) qed simp then have "poly.coeff (of_qr (decompress_poly 1 x)) i = poly.coeff (of_qr (to_qr (Poly (map of_int_mod_ring (map (\a. round(real_of_int q/2) * a) (?int_coeffs)))))) i" unfolding decompress_poly_def comp_def by presburger also have "\ = poly.coeff (of_qr (to_qr (Poly (map (\a. of_int_mod_ring ((round(real_of_int q/2)) * a)) (?int_coeffs))))) i" using map_map[of of_int_mod_ring "((*) (round (real_of_int q/2)))"] by (smt (z3) map_eq_conv o_apply) also have "\ = poly.coeff (of_qr (to_qr (Poly (map (\a. of_int_mod_ring (round(real_of_int q/2)) * of_int_mod_ring a) (?int_coeffs))))) i" by (simp add: of_int_mod_ring_mult[of "(round(real_of_int q/2))"]) also have "\ = poly.coeff (of_qr (to_qr (Poly (map (\a. of_int_mod_ring (round(real_of_int q/2)) * a) (map of_int_mod_ring (?int_coeffs)))))) i" using map_map[symmetric, of "(\a. of_int_mod_ring (round (real_of_int q/2)) * a ::'a mod_ring)" "of_int_mod_ring"] unfolding comp_def by presburger also have "\ = poly.coeff (of_qr (to_qr (Polynomial.smult (of_int_mod_ring (round(real_of_int q/2))) (Poly (map of_int_mod_ring (?int_coeffs)))))) i" using smult_Poly[symmetric, of "(of_int_mod_ring (round (real_of_int q/2)))"] by metis also have "\ = poly.coeff (of_qr ((to_module (round (real_of_int q/2)) * to_qr (Poly (map of_int_mod_ring (?int_coeffs)))))) i" unfolding to_module_def using to_qr_smult_to_module [of "of_int_mod_ring (round (real_of_int q/2))"] by metis also have "\ = poly.coeff (of_qr (to_module (round (real_of_int q/2)) * to_qr (Poly (coeffs (of_qr x)))))i" by (subst map_map[of of_int_mod_ring to_int_mod_ring], unfold comp_def)(subst of_int_mod_ring_to_int_mod_ring, auto) also have "\ = poly.coeff (of_qr (to_module (round (real_of_int q/2)) * x))i" by (subst Poly_coeffs) (subst to_qr_of_qr, simp) finally show ?thesis by auto qed then have eq: "of_qr (decompress_poly 1 x) = of_qr (to_module (round((real_of_int q)/2)) * x)" by (simp add: poly_eq_iff) show ?thesis using arg_cong[OF eq, of "to_qr"] to_qr_of_qr[of "decompress_poly 1 x"] to_qr_of_qr[of "to_module (round (real_of_int q/2)) * x"] by auto qed text \Compression and decompression for vectors.\ definition map_vector :: "('b \ 'b) \ ('b, 'n) vec \ ('b, 'n::finite) vec" where "map_vector f v = (\ i. f (vec_nth v i))" text \Compression and decompression of vectors in \\_q[X]/(X^n+1)\.\ definition compress_vec :: "nat \ ('a qr, 'k) vec \ ('a qr, 'k) vec" where "compress_vec d = map_vector (compress_poly d)" definition decompress_vec :: "nat \ ('a qr, 'k) vec \ ('a qr, 'k) vec" where "decompress_vec d = map_vector (decompress_poly d)" end end diff --git a/thys/CRYSTALS-Kyber/NTT_Scheme.thy b/thys/CRYSTALS-Kyber/NTT_Scheme.thy --- a/thys/CRYSTALS-Kyber/NTT_Scheme.thy +++ b/thys/CRYSTALS-Kyber/NTT_Scheme.thy @@ -1,1050 +1,1048 @@ theory NTT_Scheme imports Crypto_Scheme Mod_Ring_Numeral "Number_Theoretic_Transform.NTT" begin section \Number Theoretic Transform for Kyber\ lemma Poly_strip_while: "Poly (strip_while ((=) 0) x) = Poly x" by (metis Poly_coeffs coeffs_Poly) locale kyber_ntt = kyber_spec "TYPE('a :: qr_spec)" "TYPE('k::finite)" + fixes type_a :: "('a :: qr_spec) itself" and type_k :: "('k ::finite) itself" and \ :: "('a::qr_spec) mod_ring" and \ :: "'a mod_ring" and \ :: "'a mod_ring" and \inv :: "'a mod_ring" and ninv :: "'a mod_ring" and mult_factor :: int assumes omega_properties: "\^n = 1" "\ \ 1" "(\ m. \^m = 1 \ m\0 \ m \ n)" and mu_properties: "\ * \ = 1" "\ \ 1" and psi_properties: "\^2 = \" "\^n = -1" and psi_psiinv: "\ * \inv = 1" and n_ninv: "(of_int_mod_ring n) * ninv = 1" and q_split: "q = mult_factor * n + 1" begin text \Some properties of the roots $\omega$ and $\psi$ and their inverses $\mu$ and $\psi_inv$.\ lemma mu_prop: "(\ m. \^m = 1 \ m\0 \ m \ n)" by (metis mu_properties(1) mult.commute mult.right_neutral omega_properties(3) power_mult_distrib power_one) lemma mu_prop': assumes "\^m' = 1" "m'\0" shows "m' \ n" using mu_prop assms by blast lemma omega_prop': assumes "\^m' = 1" "m'\0" shows "m' \ n" using omega_properties(3) assms by blast lemma psi_props: shows "\^(2*n) = 1" "\^(n*(2*a+1)) = -1" "\\1" proof - show "\^(2*n) = 1" by (simp add: omega_properties(1) power_mult psi_properties) show "\^(n*(2*a+1)) = -1" by (metis (no_types, lifting) mult.commute mult_1 power_add power_minus1_even power_mult psi_properties(2)) show "\\1" using omega_properties(2) one_power2 psi_properties(1) by blast qed lemma psi_inv_exp: "\^i * \inv ^i = 1" using left_right_inverse_power psi_psiinv by blast lemma inv_psi_exp: "\inv^i * \ ^i = 1" by (simp add: mult.commute psi_inv_exp) lemma negative_psi: assumes "i^j * \inv ^i = \^(j-i)" proof - have j: "\^j = \^(j-i) * \^i" using assms by (metis add.commute le_add_diff_inverse nat_less_le power_add) show ?thesis unfolding j by (simp add: left_right_inverse_power psi_psiinv) qed lemma negative_psi': assumes "i\j" shows "\inv^i * \ ^j = \^(j-i)" proof - have j: "\^j = \^i * \^(j-i)" using assms by (metis le_add_diff_inverse power_add) show ?thesis unfolding j mult.assoc[symmetric] using inv_psi_exp[of i] by simp qed lemma psiinv_prop: shows "\inv^2 = \" proof - show "\inv^2 = \" by (metis (mono_tags, lifting) mu_properties(1) mult.commute mult_cancel_right mult_cancel_right2 power_mult_distrib psi_properties(1) psi_psiinv) qed lemma n_ninv': "ninv * (of_int_mod_ring n) = 1" using n_ninv by (simp add: mult.commute) text \The \map2\ function for polynomials.\ definition map2_poly :: "('a mod_ring \ 'a mod_ring \ 'a mod_ring) \ 'a mod_ring poly \ 'a mod_ring poly \ 'a mod_ring poly" where "map2_poly f p1 p2 = Poly (map2 f (map (poly.coeff p1) [0..Additional lemmas on polynomials.\ lemma Poly_map_coeff: assumes "degree f < num" shows "Poly (map (poly.coeff (f)) [0..degree f" using assms by auto then show ?thesis unfolding coeff_Poly using False by (simp add: coeff_eq_0 nth_default_beyond) qed qed lemma map_upto_n_mod: "(Poly (map f [0..auto\) qed lemma coeff_of_qr_zero: assumes "i\n" shows "poly.coeff (of_qr (f :: 'a qr)) i = 0" proof - have "degree (of_qr f) < i" using deg_of_qr deg_qr_n assms order_less_le_trans by auto then show ?thesis by (subst coeff_eq_0, auto) qed text \Definition of NTT on polynomials. In contrast to the ordinary NTT, we use a different exponent on the root of unity $\psi$.\ definition ntt_coeff_poly :: "'a qr \ nat \ 'a mod_ring" where "ntt_coeff_poly g i = (\j\{0..^(j * (2*i+1)))" definition ntt_coeffs :: "'a qr \ 'a mod_ring list" where "ntt_coeffs g = map (ntt_coeff_poly g) [0.. 'a qr" where "ntt_poly g = to_qr (Poly (ntt_coeffs g))" text \Definition of inverse NTT on polynomials. The inverse transformed is already scaled such that it is the true inverse of the NTT.\ definition inv_ntt_coeff_poly :: "'a qr \ nat \ 'a mod_ring" where "inv_ntt_coeff_poly g' i = ninv * (\j\{0..inv^(i*(2*j+1)))" definition inv_ntt_coeffs :: "'a qr \ 'a mod_ring list" where "inv_ntt_coeffs g' = map (inv_ntt_coeff_poly g') [0.. 'a qr" where "inv_ntt_poly g = to_qr (Poly (inv_ntt_coeffs g))" text \Kyber is indeed in the NTT-domain with root of unity $\omega$. Note, that our ntt on polynomials uses a slightly different exponent. The root of unity $\omega$ defines an alternative NTT in Kyber.\ text \Have $7681 = 30*256 + 1$ and $3329 = 13 * 256 + 1$.\ interpretation kyber_ntt: ntt "nat q" "nat n" "nat mult_factor" \ \ proof (unfold_locales, goal_cases) case 2 then show ?case using q_gt_two by linarith next case 3 then show ?case by (smt (verit, del_insts) int_nat_eq mult.commute nat_int_add nat_mult_distrib of_nat_1 q_gt_two q_split zadd_int_left) next case 4 then show ?case using n_gt_1 by linarith qed (use CARD_a nat_int in \auto simp add: omega_properties mu_properties\) text \Multiplication in of polynomials in $R_q$ is a negacyclic convolution (because we factored by $x^n + 1$, thus $x^n\equiv -1 \mod x^n+1$). This is the reason why we needed to adapt the exponent in the NTT.\ definition qr_mult_coeffs :: "'a qr \ 'a qr \ 'a qr" (infixl "\" 70) where "qr_mult_coeffs f g = to_qr (map2_poly (*) (of_qr f) (of_qr g))" text \The definition of the exponentiation \^\ only allows for natural exponents, thus we need to cheat a bit by introducing \conv_sign x\$\equiv (-1)^x$.\ definition conv_sign :: "int \ 'a mod_ring" where "conv_sign x = (if x mod 2 = 0 then 1 else -1)" text \The definition of the negacyclic convolution.\ definition negacycl_conv :: "'a qr \ 'a qr \ 'a qr" where "negacycl_conv f g = to_qr (Poly (map (\i. \jRepresentation of f modulo \qr_poly\.\ lemma mod_div_qr_poly: "(f :: 'a mod_ring poly) = (f mod qr_poly) + qr_poly * (f div qr_poly)" by simp text \\take_deg\ returns the first $n$ coefficients of a polynomial.\ definition take_deg :: "nat \ ('b::zero) poly \ 'b poly" where "take_deg = (\n. \f. Poly (take n (coeffs f)))" text \\drop_deg\ returns the coefficients of a polynomial strarting from the $n$-th coefficient.\ definition drop_deg :: "nat \ ('b::zero) poly \ 'b poly" where "drop_deg = (\n. \f. Poly (drop n (coeffs f)))" text \\take_deg\ and \drop_deg\ return the modulo and divisor representants.\ lemma take_deg_monom_drop_deg: assumes "degree f \ n" shows "(f :: 'a mod_ring poly) = take_deg n f + (Polynomial.monom 1 n) * drop_deg n f" proof - have "min (length (coeffs f)) n = n" using assms by (metis bot_nat_0.not_eq_extremum degree_0 le_imp_less_Suc length_coeffs_degree min.absorb1 min.absorb4) then show ?thesis unfolding take_deg_def drop_deg_def apply (subst Poly_coeffs[of f,symmetric]) apply (subst append_take_drop_id[of n "coeffs f", symmetric]) apply (subst Poly_append) by (auto) qed lemma split_mod_qr_poly: assumes "degree f \ n" shows "(f :: 'a mod_ring poly) = take_deg n f - drop_deg n f + qr_poly * drop_deg n f" proof - have "(Polynomial.monom 1 n + 1) * drop_deg n f = Polynomial.monom 1 n * drop_deg n f + drop_deg n f" by (simp add: mult_poly_add_left) then show ?thesis apply (subst take_deg_monom_drop_deg[OF assms]) apply (unfold qr_poly_def qr_poly'_eq of_int_hom.map_poly_hom_add) by auto qed text \Lemmas on the degrees of \take_deg\ and \drop_deg\.\ lemma degree_drop_n: "degree (drop_deg n f) = degree f - n" unfolding drop_deg_def by (simp add: degree_eq_length_coeffs) lemma degree_drop_2n: assumes "degree f < 2*n" shows "degree (drop_deg n f) < n" using assms unfolding degree_drop_n by auto lemma degree_take_n: "degree (take_deg n f) < n" unfolding take_deg_def by (metis coeff_Poly_eq deg_qr_n deg_qr_pos degree_0 leading_coeff_0_iff nth_default_take of_nat_eq_iff) lemma deg_mult_of_qr: "degree (of_qr (f ::'a qr) * of_qr g) < 2 * n" by (metis add_less_mono deg_of_qr deg_qr_n degree_0 degree_mult_eq mult_2 mult_eq_0_iff nat_int_comparison(1)) text \Representation of a polynomial modulo \qr_poly\ using \take_deg\ and \drop_deg\.\ lemma mod_qr_poly: assumes "degree f \ n" "degree f < 2*n" shows "(f :: 'a mod_ring poly) mod qr_poly = take_deg n f - drop_deg n f " proof - have "degree (take_deg n f - drop_deg n f) < deg_qr TYPE('a)" using degree_diff_le_max[of "take_deg n f" "drop_deg n f"] degree_drop_2n[OF assms(2)] degree_take_n by (metis deg_qr_n degree_diff_less nat_int) then have "(take_deg n f - drop_deg n f) mod qr_poly = take_deg n f - drop_deg n f" by (subst deg_mod_qr_poly, auto) then show ?thesis by (subst split_mod_qr_poly[OF assms(1)], auto) qed text \Coefficients of \take_deg\, \drop_deg\ and the modulo representant.\ lemma coeff_take_deg: assumes "i n" "degree f < 2*n" "iMore lemmas on the splitting of sums.\ lemma sum_leq_split: "(\ia\i+n. f ia) = (\iaia\{n..i+n}. f ia)" proof - have *: "{..i + n} - {.. {-b..<0}" shows "x div b = -1" using assms by (smt (verit, ccfv_SIG) atLeastLessThan_iff div_minus_minus div_pos_neg_trivial) text \A coefficient of polynomial multiplication is a coefficient of the negacyclic convolution.\ lemma coeff_conv: fixes f :: "'a qr" assumes "ij = (\ia\i. poly.coeff (of_qr f) ia * poly.coeff (of_qr g) (i - ia))" unfolding coeff_mult by auto also have "\ = (\ia\i. conv_sign ((int i - int ia) div int n) * poly.coeff (of_qr f) ia * poly.coeff (of_qr g) (nat ((int i - int ia) mod n)))" proof - have "i-ia = nat ((int i - int ia) mod n)" if "ia \ i" for ia using assms that by force moreover have "conv_sign ((int i - int ia) div int n) = 1" if "ia \ i" for ia unfolding conv_sign_def using assms that by force ultimately show ?thesis by auto qed also have "\ = (\ia {i<.. degree (of_qr g)" using le_degree[OF 1] by auto have "ia > degree (of_qr f)" proof (rule ccontr) assume "\ degree (of_qr f) < ia" then have as: "ia \ degree (of_qr f)" by auto then have ni: "nat ((int i - int ia) mod int n) + ia = n + i" using that by (smt (verit, ccfv_threshold) True deg_g greaterThanLessThan_iff int_nat_eq less_imp_of_nat_less mod_add_self1 mod_pos_pos_trivial of_nat_0_le_iff of_nat_add of_nat_mono) have "n + i \ degree (of_qr f) + degree (of_qr g)" unfolding ni[symmetric] using as deg_g by auto then show False using True by auto qed then show ?case using coeff_eq_0 by blast qed then show ?thesis by (subst sum_less_split[OF \i]) (simp add: sum.neutral) qed finally show ?thesis by blast next case False then have *: "degree (of_qr f * of_qr g) \ n" by (metis add.right_neutral add_0 deg_of_qr deg_qr_n degree_0 degree_mult_eq linorder_not_le nat_int) have "poly.coeff (of_qr f * of_qr g) (i + n) = (\ia = (\ia\{i<..i]) auto also have "\ = (\ia\{i<..{0.. {i<..{i<..i that by (smt (verit, best) mod_add_self1 mod_rangeE) then have "i+n-ia = nat ((int i - ia) mod n)" if "ia\{i<.. = - (\ia\{i<.. {-n..<0}" if "ia \{i<..{i<..{i<..ia\{i<..x\{i<..ia\{i<..ia\i. conv_sign ((int i - int ia) div n) * poly.coeff (of_qr f) ia * poly.coeff (of_qr g) (nat ((int i - ia) mod n)))" proof - have "conv_sign ((int i - int ia) div n) = 1" if "ia \i" for ia using that assms conv_sign_def by force moreover have "i-ia \{0..i" for ia using that assms by auto then have "i-ia = (nat ((int i - ia) mod n))" if "ia\i" for ia using assms that by force ultimately show ?thesis unfolding coeff_mult using assms less_imp_diff_less mod_less by auto qed have calc: "poly.coeff (of_qr f * of_qr g) i - poly.coeff (of_qr f * of_qr g) (i + n) = (\iaPolynomial multiplication in $R_q$ is the negacyclic convolution.\ lemma mult_negacycl: "f * g = negacycl_conv f g" proof - have f_times_g: "f * g = to_qr ((of_qr f) * (of_qr g) mod qr_poly)" by (metis of_qr_mult to_qr_of_qr) have conv: "poly.coeff ((of_qr f) * (of_qr g) mod qr_poly) i = (\jAdditional lemmas on \ntt_coeffs\.\ lemma length_ntt_coeffs: "length (ntt_coeffs f) \ n" unfolding ntt_coeffs_def by auto lemma degree_Poly_ntt_coeffs: "degree (Poly (ntt_coeffs f)) < n" using length_ntt_coeffs by (smt (verit) deg_Poly' degree_0 degree_take_n diff_diff_cancel diff_is_0_eq le_neq_implies_less less_nat_zero_code nat_le_linear order.strict_trans1 power_0_left power_eq_0_iff) lemma Poly_ntt_coeffs_mod_qr_poly: "Poly (ntt_coeffs f) mod qr_poly = Poly (ntt_coeffs f)" using map_upto_n_mod ntt_coeffs_def by presburger lemma nth_default_map: assumes "iiWriting the convolution sign as a conditional if statement.\ lemma conv_sign_if: assumes "x {-n..<0}" using assms by simp then have "(int x - int y) div int n mod 2 = 1" using div_minus_1 by presburger then show ?case by auto next case 2 then have "(int x - int y) div int n mod 2 = 0" using assms(1) by force then show ?case by auto qed text \The convolution theorem on coefficients.\ lemma ntt_coeff_poly_mult: assumes "lx. \ y. conv_sign ((int x - int y) div int n) * poly.coeff (of_qr f) y * poly.coeff (of_qr g) (nat ((int x - int y) mod int n)))" have "ntt_coeff_poly (f*g) l = (\j = 0..^(j*(2*l+1)))" unfolding ntt_coeff_poly_def mult_negacycl by auto also have "\ = (\j=0..i^(j*(2*l+1))))" proof (subst sum.cong[of "{0..j. poly.coeff (of_qr (negacycl_conv f g)) j * \^(j*(2*l+1)))" "(\j. (\i^(j*(2*l+1))))"], goal_cases) case (2 j) then have "j ^ (j * (2 * l + 1)) = (\na ^ (j * (2 * l + 1)))" apply (subst nth_coeffs_negacycl[OF \j]) apply (subst sum_distrib_right) by auto also have "\ = (\na ^ (j * (2 * l + 1)))" unfolding f1_def by auto finally show ?case by blast qed auto also have "\ = (\ij ^ (j * (2 * l + 1))) " by (subst atLeast0LessThan, subst sum.swap, auto) also have "\ = (\i ^ (i * (2 * l + 1)) * (\jinv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))))" proof (subst sum.cong[of "{..i. (\j ^ (j * (2 * l + 1))))" "(\i. poly.coeff (of_qr f) i * \ ^ (i * (2 * l + 1)) * (\jinv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))))"], goal_cases) case (2 i) then show ?case proof (subst sum_distrib_left, subst sum.cong[of "{..j. f1 j i * \ ^ (j * (2 * l + 1)))" "(\j. poly.coeff (of_qr f) i * \ ^ (i * (2 * l + 1)) * (poly.coeff (of_qr g) (nat ((int j - int i) mod int n)) * (if int j - int i < 0 then - 1 else 1) * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))))"], goal_cases) case (2 j) then have *: "conv_sign ((int j - int i) div int n) = (if int j - int i < 0 then - 1 else 1)" using conv_sign_if by auto have "f1 j i * \ ^ (j * (2 * l + 1)) = \ ^ (i * (2 * l + 1)) * f1 j i * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))" using psi_psiinv by (simp add: left_right_inverse_power) also have "\ = poly.coeff (of_qr f) i * \ ^ (i * (2 * l + 1)) * (poly.coeff (of_qr g) (nat ((int j - int i) mod int n)) * (if int j - int i < 0 then - 1 else 1) * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1)))" unfolding f1_def mult.assoc by (simp add: "*" mult.left_commute) finally show ?case by blast qed auto qed auto also have "\ = (\i ^ (i * (2 * l + 1)) * (\x ^ (x * (2 * l + 1))))" proof - define x' where "x' = (\j i. nat ((int j - int i) mod int n))" let ?if_inv = "(\i j. (if int j - int i < 0 then - 1 else 1) * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1)))" have rewrite: "(if int j - int i < 0 then - 1 else 1) * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1)) = \ ^ ((x' j i) * (2 * l + 1))" if "ii by (metis One_nat_def add_gr_0 lessI mult_less_mono1) have "?if_inv i j = (-1) * \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))" using True by (auto split: if_splits) also have "\ = \^((n-i+j)* (2 * l + 1))" unfolding psi_props(2)[of l, symmetric] negative_psi[OF lt] by (metis comm_semiring_class.distrib diff_mult_distrib power_add) also have "\ = \ ^ ((x' j i) * (2 * l + 1))" unfolding x'_def by (smt (verit, best) True mod_add_self2 mod_pos_pos_trivial nat_int_add nat_less_le of_nat_0_le_iff of_nat_diff that(1)) finally show ?thesis by blast next case False then have "i\j" by auto have lt: "i * (2 * l + 1) \ j * (2 * l + 1)" using \i\j\ using add_gr_0 less_one mult_less_mono1 using mult_le_cancel2 by presburger have "?if_inv i j = \inv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))" using False by (auto split: if_splits) also have "\ = \^((j-i)* (2 * l + 1))" using negative_psi'[OF lt] diff_mult_distrib by presburger also have "\ = \ ^ ((x' j i) * (2 * l + 1))" unfolding x'_def by (metis \i \ j\ less_imp_diff_less mod_pos_pos_trivial nat_int of_nat_0_le_iff of_nat_diff of_nat_less_iff that(2)) finally show ?thesis by blast qed then have "(\jinv ^ (i * (2 * l + 1)) * \ ^ (j * (2 * l + 1))) = (\x ^ (x * (2 * l + 1)))" (is "(\jjj ^ ((x' j i) * (2 * l + 1)))" using rewrite[OF that] x'_def by (smt (verit, ccfv_SIG) lessThan_iff mult.assoc sum.cong) have eq: "(\j. x' j i) ` {..{..j. x' j i) {..x. poly.coeff (of_qr g) x * \ ^ (x * (2 * l + 1)))" "(\j. poly.coeff (of_qr g) (x' j i) * \ ^ (x' j i * (2 * l + 1)))"], auto) qed then show ?thesis by force qed also have "\ = (\i ^ (i * (2 * l + 1))) * (\x' ^ (x' * (2 * l + 1)))" unfolding sum_distrib_right by auto also have "\ = ntt_coeff_poly f l * ntt_coeff_poly g l" unfolding ntt_coeff_poly_def atLeast0LessThan by auto finally show ?thesis by blast qed lemma ntt_coeffs_mult: assumes "iSteps towards the convolution theorem.\ lemma nth_default_ntt_coeff_mult: "nth_default 0 (ntt_coeffs (f * g)) i = nth_default 0 (map2 (*) (map (poly.coeff (Poly (ntt_coeffs f))) [0..{0..Convolution theorem for NTT\ lemma ntt_mult: "ntt_poly (f * g) = qr_mult_coeffs (ntt_poly f) (ntt_poly g)" proof - have "Poly (ntt_coeffs (f*g)) mod qr_poly = Poly (ntt_coeffs (f*g))" using Poly_ntt_coeffs_mod_qr_poly by force also have "\ = Poly (coeffs (map2_poly (*) (Poly (ntt_coeffs f)) (Poly (ntt_coeffs g))))" unfolding map2_poly_def coeffs_Poly Poly_strip_while using Poly_ntt_coeffs_mult by auto also have "\ = (map2_poly (*) (of_qr (to_qr (Poly (ntt_coeffs f)))) (of_qr (to_qr (Poly (ntt_coeffs g))))) mod qr_poly" unfolding of_qr_to_qr map_poly_def Poly_ntt_coeffs_mod_qr_poly by (metis Poly_coeffs Poly_ntt_coeffs_mod_qr_poly calculation) finally have "[Poly (ntt_coeffs (f * g)) = (map2_poly (*) (of_qr (to_qr (Poly (ntt_coeffs f)))) (of_qr (to_qr (Poly (ntt_coeffs g)))))] (mod qr_poly)" using cong_def by blast then have "to_qr (Poly (ntt_coeffs (f * g))) = to_qr (map2_poly (*) (of_qr (to_qr (Poly (ntt_coeffs f)))) (of_qr (to_qr (Poly (ntt_coeffs g)))))" using of_qr_to_qr by auto then show ?thesis unfolding ntt_poly_def qr_mult_coeffs_def by auto qed text \Correctness of NTT on polynomials.\ lemma inv_ntt_poly_correct: "inv_ntt_poly (ntt_poly f) = f" proof - have rew_sum: "(\j = 0..i. \j = 0.. ^ (j * (2 * i + 1))) [0..inv ^ (i * (2 * j + 1))) = (\j = 0..j' = 0.. ^ (j' * (2 * j + 1))) * \inv ^ (i * (2 * j + 1)))" (is "(\j = 0..j = 0..j = 0..j' = 0.. ^ (j' * (2 * j + 1)) * \inv ^ (i * (2 * j + 1))) = (of_int_mod_ring n) * poly.coeff (of_qr f) i" if "i ^ (j * (2 * j' + 1)) * \inv ^ (i * (2 * j' + 1)) = \ ^ j * \inv ^ i * (\ ^ (j * 2) * \inv ^ (i * 2)) ^ j'" if "j'j = 0..j' = 0.. ^ (j' * (2 * j + 1)) * \inv ^ (i * (2 * j + 1))) = (\j' = 0..^j' * \inv^i * (\j = 0.. ^ (j' * 2) * \inv ^ (i * 2))^j))" apply (subst sum_distrib_left, subst sum.swap) proof (subst sum.cong[of "{0..j. \ia = 0.. ^ (j * (2 * ia + 1)) * \inv ^ (i * (2 * ia + 1)))" "(\j. \ia = 0.. ^ j * \inv ^ i * (\ ^ (j * 2) * \inv ^ (i * 2)) ^ ia)"], goal_cases) case (2 j) then show ?case proof (subst sum.cong[of "{0..ia. poly.coeff (of_qr f) j * \ ^ (j * (2 * ia + 1)) * \inv ^ (i * (2 * ia + 1)))" "(\ia. poly.coeff (of_qr f) j * \ ^ j * \inv ^ i * (\ ^ (j * 2) * \inv ^ (i * 2)) ^ ia)"], goal_cases) case (2 j') then show ?case using rew_psi[of j' j] by simp qed auto qed auto also have "\ = (\j' = 0..^j' * \inv^i * (of_int_mod_ring n) else 0))" proof (subst sum.cong[of "{0..j'. poly.coeff (of_qr f) j' * \ ^ j' * \inv ^ i * (\j=0.. ^ (j' * 2) * \inv ^ (i * 2))^j))" "(\j'. (if j' = i then poly.coeff (of_qr f) j' * \^j' * \inv^i * (of_int_mod_ring n) else 0))"], goal_cases) case (2 j') then show ?case proof (cases "j' = i") case True then have "(\j=0.. ^ (j' * 2) * \inv ^ (i * 2))^j) = of_int_mod_ring n" unfolding True psi_inv_exp by (metis kyber_ntt.sum_rules(5) mult.right_neutral power_one sum.cong) then show ?thesis using True by auto next case False have not1: "\ ^ (j' * 2) * \inv ^ (i * 2) \ 1" proof - have "\^j' * \ ^i \ 1" proof (cases "j'^j' * \ ^i = \^(i-j')" using True by (metis (no_types, lifting) le_add_diff_inverse less_or_eq_imp_le mult.assoc mult_cancel_right2 power_add power_mult psi_inv_exp psi_properties(1) psiinv_prop) show ?thesis proof (unfold *, rule ccontr) assume "\ \ ^ (i - j') \ 1" then have 1: "\ ^ (i - j') = 1" by auto show False using mu_prop'[OF 1] \j'\i\ using True less_imp_diff_less that diff_is_0_eq leD by blast qed next case False have 2: "\^j' * \ ^i = \^(j'-i)" using False by (smt (verit) Nat.add_diff_assoc ab_semigroup_mult_class.mult_ac(1) add_diff_cancel_left' left_right_inverse_power linorder_not_less mu_properties(1) mult.commute mult_numeral_1_right numeral_One power_add) show ?thesis proof (unfold 2, rule ccontr) assume "\ \ ^ (j' - i) \ 1" then have 1: "\ ^ (j' - i) = 1" by auto have "n > j' - i" using \j' \ {0.. by auto then show False using omega_prop'[OF 1] \j'\i\ using False by (meson diff_is_0_eq leD order_le_imp_less_or_eq) qed qed then show ?thesis by (metis mult.commute power_mult psi_properties(1) psiinv_prop) qed have "(1 - \ ^ (j' * 2) * \inv ^ (i * 2)) * (\j=0.. ^ (j' * 2) * \inv ^ (i * 2))^j) = 0" proof (subst kyber_ntt.geo_sum, goal_cases) case 1 then show ?case using not1 by auto next case 2 then show ?case by (metis (no_types, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel mu_properties(1) mult.commute omega_properties(1) power_mult power_mult_distrib power_one psi_properties(1) psiinv_prop) qed then have "(\j=0.. ^ (j' * 2) * \inv ^ (i * 2))^j) = 0" using not1 by auto then show ?thesis using False by auto qed qed auto also have "\ = poly.coeff (of_qr f) i * \^i * \inv^i * (of_int_mod_ring n)" by (subst sum.delta[of "{0..i in auto) also have "\ = (of_int_mod_ring n) * poly.coeff (of_qr f) i" by (simp add: psi_inv_exp) finally show ?thesis by blast qed then have rew_coeff: "(map (\i. ninv * (\j = 0..n = 0.. ^ (n * (2 * j + 1)) * \inv ^ (i * (2 * j + 1)))) [0..i. ninv * (of_int_mod_ring (int n) * poly.coeff (of_qr f) i)) [0..j = 0..i. ninv * (\j' = 0..inv ^ (i * (2 * j' + 1)))) [0.. ^ (j * (2 * i + 1))) = (\j = 0..j' = 0..inv ^ (j * (2 * j' + 1))) * \ ^ (j * (2 * i + 1)))" (is "(\j = 0..j = 0..j = 0..n = 0..inv ^ (j * (2 * n + 1))) * \ ^ (j * (2 * i + 1))) = ninv * (of_int_mod_ring (int n) * poly.coeff (of_qr f) i)" if "iinv ^ (j' * (2 * j + 1)) * \ ^ (j' * (2 * i + 1)) = (\inv ^ (j * 2) * \ ^ (i * 2)) ^ j'" if "j'inv ^ (j' * (2 * j + 1)) * \ ^ (j' * (2 * i + 1)) = \inv ^ (j' * (2 * j)) * \ ^ (j' * (2 * i)) * \inv ^ j' * \ ^ j' " by (simp add: power_add) also have "\ = (\inv ^ (2 * j) * \ ^ (2 * i))^ j'" by (smt (verit, best) inv_psi_exp kyber_ntt.exp_rule mult.assoc mult.commute mult.right_neutral power_mult) also have "\ = (\inv ^ (j * 2) * \ ^ (i * 2)) ^ j'" by (simp add: mult.commute) finally show ?thesis by blast qed have "(\j = 0..j' = 0..inv ^ (j * (2 * j' + 1))) * \ ^ (j * (2 * i + 1))) = (\j' = 0..j = 0..inv ^ (j' * 2) * \ ^ (i * 2))^j))" apply (subst sum_distrib_left, subst sum.swap, unfold mult.assoc[symmetric]) proof (subst sum.cong[of "{0..j. \ia = 0..inv ^ (ia * (2 * j + 1)) * \ ^ (ia * (2 * i + 1)))" "(\j. \n = 0..inv ^ (j * 2) * \ ^ (i * 2)) ^ n)"], goal_cases) case (2 j) then show ?case proof (subst sum.cong[of "{0..ia. ninv * poly.coeff (of_qr f) j * \inv ^ (ia * (2 * j + 1)) * \ ^ (ia * (2 * i + 1)))" "(\ia. ninv * poly.coeff (of_qr f) j * (\inv ^ (j * 2) * \ ^ (i * 2)) ^ ia)"], goal_cases) case (2 j') then show ?case using rew_psi[of j' j] by simp qed auto qed auto also have "\ = (\j' = 0..inv^j' * \^i * (of_int_mod_ring n) else 0))" (is "(\j' = 0..j' = 0..j=0..inv ^ (j' * 2) * \ ^ (i * 2))^j) = of_int_mod_ring n" unfolding True psi_inv_exp by (metis kyber_ntt.sum_const mult.commute mult.right_neutral power_one psi_inv_exp sum.cong) then show ?thesis using True by (simp add: inv_psi_exp) next case False have not1: "\inv ^ (j' * 2) * \ ^ (i * 2) \ 1" proof - have "\^j' * \ ^i \ 1" proof (cases "j'^j' * \ ^i = \^(i-j')" using True by (smt (verit, best) add.commute kyber_ntt.omega_properties(1) le_add_diff_inverse left_right_inverse_power less_or_eq_imp_le mu_properties(1) mult.left_commute mult_cancel_right1 power_add) show ?thesis proof (unfold *, rule ccontr) assume "\ \ ^ (i - j') \ 1" then have 1: "\ ^ (i - j') = 1" by auto show False using omega_prop'[OF 1] \j'\i\ using True less_imp_diff_less that diff_is_0_eq leD by blast qed next case False have 2: "\^j' * \ ^i = \^(j'-i)" using False by (smt (verit) Nat.add_diff_assoc ab_semigroup_mult_class.mult_ac(1) add_diff_cancel_left' left_right_inverse_power linorder_not_less mu_properties(1) mult.commute mult_numeral_1_right numeral_One power_add) show ?thesis proof (unfold 2, rule ccontr) assume "\ \ ^ (j' - i) \ 1" then have 1: "\ ^ (j' - i) = 1" by auto have "n > j' - i" using \j' \ {0.. by auto then show False using mu_prop'[OF 1] \j'\i\ using False by (meson diff_is_0_eq leD order_le_imp_less_or_eq) qed qed then show ?thesis by (metis mult.commute power_mult psi_properties(1) psiinv_prop) qed have "(1 - \inv ^ (j' * 2) * \ ^ (i * 2)) * (\j=0..inv ^ (j' * 2) * \ ^ (i * 2))^j) = 0" proof (subst kyber_ntt.geo_sum, goal_cases) case 1 then show ?case using not1 by auto next case 2 then show ?case by (metis (no_types, opaque_lifting) cancel_comm_monoid_add_class.diff_cancel mu_properties(1) mult.commute omega_properties(1) power_mult power_mult_distrib power_one psi_properties(1) psiinv_prop) qed then have "(\j=0..inv ^ (j' * 2) * \ ^ (i * 2))^j) = 0" using not1 by auto then show ?thesis using False by auto qed qed auto also have "\ = ninv * poly.coeff (of_qr f) i * \inv^i * \^i * (of_int_mod_ring n)" by (subst sum.delta[of "{0..i in auto) also have "\ = ninv * ((of_int_mod_ring n) * poly.coeff (of_qr f) i)" by (simp add: psi_inv_exp mult.commute) finally show ?thesis by blast qed then have rew_coeff: "(map (\i. \j = 0..n = 0..inv ^ (j * (2 * n + 1))) * \ ^ (j * (2 * i + 1))) [0..i. ninv * (of_int_mod_ring (int n) * poly.coeff (of_qr f) i)) [0..The multiplication of two polynomials can be computed by the NTT.\ lemma convolution_thm_ntt_poly: "f*g = inv_ntt_poly (qr_mult_coeffs (ntt_poly f) (ntt_poly g))" unfolding ntt_mult[symmetric] inv_ntt_poly_correct by auto end end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy b/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy --- a/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy +++ b/thys/DPRM_Theorem/Machine_Equations/Machine_Equation_Equivalence.thy @@ -1,257 +1,256 @@ subsection \Equivalence of register machine and arithmetizing equations\ theory Machine_Equation_Equivalence imports All_Equations "../Register_Machine/MachineEquations" "../Register_Machine/MultipleToSingleSteps" begin context register_machine begin lemma conclusion_4_5: assumes is_val: "is_valid_initial ic p a" and n_def: "n \ length (snd ic)" shows "(\q. terminates ic p q) = rm_equations a" proof (rule) assume "\q. terminates ic p q" then obtain q::nat where terminates: "terminates ic p q" by auto hence "q>0" using terminates_def by auto have "\c>1. cells_bounded ic p c" using terminate_c_exists terminates is_val is_valid_initial_def by blast then obtain c where c: "cells_bounded ic p c \ c > 1" by auto define b where "b \ B c" define d where "d \ D q c b" define e where "e \ E q b" define f where "f \ F q c b" have "c>1" using c by auto have "b>1" using c b_def B_def using nat_neq_iff by fastforce define r where "r \ RLe ic p b q" define s where "s \ SKe ic p b q" define z where "z \ ZLe ic p b q" interpret equations: rm_eq_fixes p n a b c d e f q r z s by unfold_locales have "equations.mask_equations" proof - have "\l d" using lm04_15_register_masking[of "ic" "p" "c" _ "q"] r_def n_def d_def b_def c by auto moreover have "\l e" using lm04_15_zero_masking z_def n_def e_def b_def c by auto moreover have "\lR+ p 0 s - b * \R- p 0 (\k. s k && z 0)" using lm04_23_multiple_register1[of "ic" "p" "a" "c" "0" "q"] is_val c terminates `q>0` r_def s_def z_def b_def bitAND_commutes by auto moreover have "\l>0. l < n \ r l = b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)" using lm04_22_multiple_register[of "ic" "p" "a" "c" _ "q"] b_def c terminates r_def s_def z_def is_val bitAND_commutes n_def `q>0` by auto moreover have "l r l < b^q" for l proof - assume "l 2 ^ Suc c - Suc 0" using `c>1` by auto have "\t. R ic p l t < 2 ^ c" using c `lt. R ic p l t < 2 ^ Suc c - Suc 0" using c_ineq by (metis dual_order.strict_trans linorder_neqE_nat not_less) have "(\t = 0..q. b ^ t * R ic p l t) = (\t = 0..(Suc (q-1)). b ^ t * R ic p l t)" using `q>0` by auto also have "... = (\t = 0..q-1. b ^ t * R ic p l t) + b^q * R ic p l q" using Set_Interval.comm_monoid_add_class.sum.atLeast0_atMost_Suc[of _ "q-1"] `q>0` by auto also have "... = (\t = 0..q-1. b ^ t * R ic p l t)" using Rlq by auto also have "... < b ^ q" using b_def R_bound base_summation_bound[of "R ic p l" "c" "q-1"] `q>0` by (auto simp: mult.commute) finally show ?thesis using r_def RLe_def by auto qed ultimately show ?thesis unfolding equations.register_equations_def equations.register_0_def equations.register_l_def equations.register_bound_def by auto qed moreover have "equations.state_equations" proof - have "equations.state_relations_from_recursion" proof - have "\d>0. d\m \ s d = b*\S+ p d (\k. s k) + b*\S- p d (\k. s k && z (modifies (p!k))) + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))" apply (auto simp: s_def z_def) using lm04_24_multiple_step_states[of "ic" "p" "a" "c" _ "q"] b_def c terminates s_def z_def is_val bitAND_commutes m_def `q>0` e_def E_def by auto moreover have "s 0 = 1 + b*\S+ p 0 (\k. s k) + b*\S- p 0 (\k. s k && z (modifies (p!k))) + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))" using lm04_25_multiple_step_state1[of "ic" "p" "a" "c" _ "q"] b_def c terminates s_def z_def is_val bitAND_commutes m_def `q>0` e_def E_def by auto ultimately show ?thesis unfolding equations.state_relations_from_recursion_def equations.state_0_def equations.state_d_def equations.state_m_def by auto qed moreover have "equations.state_unique_equations" proof - have "k s k < b ^ q" for k using state_q_bound is_val terminates \q>0\ b_def s_def m_def c by auto moreover have "k\m \ s k \ e" for k using state_mask is_val terminates \q>0\ b_def e_def s_def c by auto ultimately show ?thesis unfolding equations.state_unique_equations_def equations.state_mask_def equations.state_bound_def by auto qed moreover have "\M\m. sum s {..M} \ e" using state_sum_mask is_val terminates \q>0\ b_def e_def s_def c `b>1` m_def by auto moreover have "s m = b^q" using halting_condition_04_27[of "ic" "p" "a" "q" "c"] m_def b_def is_val `q>0` terminates s_def by auto ultimately show ?thesis unfolding equations.state_equations_def equations.state_partial_sum_mask_def equations.state_m_def by auto qed moreover have "equations.constants_equations" unfolding equations.constants_equations_def equations.constant_b_def equations.constant_d_def equations.constant_e_def equations.constant_f_def using b_def d_def e_def f_def by auto moreover have "equations.miscellaneous_equations" proof - have tapelength: "length (snd ic) > 0" using is_val is_valid_initial_def[of "ic" "p" "a"] by auto have "R ic p 0 0 = a" using is_val is_valid_initial_def[of "ic" "p" "a"] R_def List.hd_conv_nth[of "snd ic"] by auto moreover have "R ic p 0 0 < 2^c" using c tapelength by auto ultimately have "a < 2^c" by auto thus ?thesis unfolding equations.miscellaneous_equations_def equations.c_gt_0_def equations.a_bound_def equations.q_gt_0_def using \q > 0\ \c > 1\ by auto qed ultimately show "rm_equations a" unfolding rm_equations_def all_equations_def by blast next assume "rm_equations a" then obtain q b c d e f r z s where reg: "rm_eq_fixes.register_equations p n a b q r z s" and state: "rm_eq_fixes.state_equations p b e q z s" and mask: "rm_eq_fixes.mask_equations n c d e f r z" and const: "rm_eq_fixes.constants_equations b c d e f q" and misc: "rm_eq_fixes.miscellaneous_equations a c q" unfolding rm_equations_def all_equations_def by auto have fx: "rm_eq_fixes p n" unfolding rm_eq_fixes_def using local.register_machine_axioms by auto have "q>0" using misc fx rm_eq_fixes.miscellaneous_equations_def rm_eq_fixes.q_gt_0_def by auto have "b>1" using B_def const rm_eq_fixes.constants_equations_def rm_eq_fixes.constant_b_def fx by (metis One_nat_def Zero_not_Suc less_one n_not_Suc_n nat_neq_iff nat_power_eq_Suc_0_iff numeral_2_eq_2 of_nat_0 of_nat_power_eq_of_nat_cancel_iff of_nat_zero_less_power_iff pos2) have "n>0" using is_val is_valid_initial_def[of "ic" "p" "a"] n_def by auto have "m>0" using m_def is_val is_valid_initial_def[of "ic" "p"] is_valid_def[of "ic" "p"] by auto define Seq where "Seq \ (\k t. nth_digit (s k) t b)" define Req where "Req \ (\l t. nth_digit (r l) t b)" define Zeq where "Zeq \ (\l t. nth_digit (z l) t b)" (* Quick and dirty: :\| *) have mask_old: "mask_equations n r z c d e f" and reg_old: "reg_equations p r z s b a (length (snd ic)) q" and state_old: "state_equations p s z b e q (length p - 1)" and const_old: "rm_constants q c b d e f a" subgoal using mask rm_eq_fixes.mask_equations_def rm_eq_fixes.register_mask_def fx mask_equations_def rm_eq_fixes.zero_indicator_0_or_1_def rm_eq_fixes.zero_indicator_mask_def by simp subgoal using reg state mask const misc using rm_eq_fixes.register_equations_def rm_eq_fixes.register_0_def rm_eq_fixes.register_l_def rm_eq_fixes.register_bound_def reg_equations_def n_def fx by simp subgoal using state fx state_equations_def rm_eq_fixes.state_equations_def rm_eq_fixes.state_relations_from_recursion_def rm_eq_fixes.state_0_def rm_eq_fixes.state_m_def rm_eq_fixes.state_d_def rm_eq_fixes.state_unique_equations_def rm_eq_fixes.state_mask_def rm_eq_fixes.state_bound_def rm_eq_fixes.state_partial_sum_mask_def m_def by simp subgoal unfolding rm_constants_def using const misc fx rm_eq_fixes.constants_equations_def rm_eq_fixes.miscellaneous_equations_def rm_eq_fixes.constant_b_def rm_eq_fixes.constant_d_def rm_eq_fixes.constant_e_def rm_eq_fixes.constant_f_def rm_eq_fixes.c_gt_0_def rm_eq_fixes.q_gt_0_def rm_eq_fixes.a_bound_def by simp done hence RZS_eq: "l j\m \ t\q \ R ic p l t = Req l t \ Z ic p l t = Zeq l t \ S ic p j t = Seq j t" for l j t using rzs_eq[of "m" "p" "n" "ic" "a" "r" "z"] mask_old reg_old state_old const_old m_def n_def is_val `q>0` Seq_def Req_def Zeq_def by auto have R_eq: "l t\q \ R ic p l t = Req l t" for l t using RZS_eq by auto have Z_eq: "l t\q \ Z ic p l t = Zeq l t" for l t using RZS_eq by auto have S_eq: "j\m \ t\q \ S ic p j t = Seq j t" for j t using RZS_eq[of "0"] `n>0` by auto have "ishalt (p!m)" using m_def is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] by auto have "Seq m q = 1" using state nth_digit_def Seq_def `b>1` using fx rm_eq_fixes.state_equations_def rm_eq_fixes.state_relations_from_recursion_def rm_eq_fixes.state_m_def by auto hence "S ic p m q = 1" using S_eq by auto hence "fst (steps ic p q) = m" using S_def by(cases "fst (steps ic p q) = m"; auto) hence qhalt: "ishalt (p ! (fst (steps ic p q)))" using S_def `ishalt (p!m)` by auto hence rempty: "snd (steps ic p q) ! l = 0" if "l < n" for l unfolding R_def[symmetric] using R_eq[of l q] \l < n\ apply auto using reg Req_def nth_digit_def using rm_eq_fixes.register_equations_def rm_eq_fixes.register_l_def - rm_eq_fixes.register_0_def - rm_eq_fixes.register_bound_def - by (smt Euclidean_Division.div_eq_0_iff mod_if mult_0_right mult_mod_right nat_mult_1_right - zero_less_one fx) + rm_eq_fixes.register_0_def + rm_eq_fixes.register_bound_def + by auto (simp add: fx) have state_m_0: "t S ic p m t = 0" for t proof - assume "t1 < b\ \t < q\ less_imp_le not_one_le_zero power_diff) also have "... mod b = 0" using \1 < b\ \t < q\ by simp finally have mod: "b^q div b^t mod b = 0" by auto have "s m = b^q" using state fx rm_eq_fixes.state_equations_def rm_eq_fixes.state_m_def rm_eq_fixes.state_relations_from_recursion_def by auto hence "Seq m t = 0" using Seq_def nth_digit_def mod by auto with S_eq `t < q` show ?thesis by auto qed have "\k ishalt (p!k)" using is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def[of "ic" "p"] m_def by auto moreover have "t fst (steps ic p t) < length p - 1" for t proof (rule ccontr) assume asm: "\ (t < q \ fst (steps ic p t) < length p - 1)" hence "t length p - 1" by auto moreover have "fst (steps ic p t) \ length p - 1" using p_contains[of "ic" "p" "a" "t"] is_val by auto ultimately have "fst (steps ic p t) = m" using m_def by auto hence "S ic p m t = 1" using S_def by auto thus "False" using state_m_0[of "t"] `t \ ishalt (p ! (fst (steps ic p t)))" for t using m_def by auto hence no_early_halt: "t \ ishalt (p ! (fst (steps ic p t)))" for t using state_m_0 by auto have "correct_halt ic p q" using qhalt rempty correct_halt_def n_def by auto thus "(\q. terminates ic p q)" using no_early_halt terminates_def `q>0` by auto qed end end \ No newline at end of file diff --git a/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy b/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy --- a/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy +++ b/thys/DPRM_Theorem/Register_Machine/MachineMasking.thy @@ -1,659 +1,659 @@ subsection \Masking properties\ theory MachineMasking imports RegisterMachineSimulation "../Diophantine/Binary_And" begin (* [D] 4.18 *) definition E :: "nat \ nat \ nat" where "(E q b) = (\t = 0..q. b^t)" lemma e_geom_series: assumes "b \ 2" shows "(E q b = e) \ ((b-1) * e = b^(Suc q) - 1 )" (is "?P \ ?Q") proof- have "sum ((^) (int b)) {.. nat \ nat \ nat" where "(D q c b) = (\t = 0..q. (2^c - 1) * b^t)" lemma d_geom_series: assumes "b = 2^(Suc c)" shows "(D q c b = d) \ ((b-1) * d = (2^c - 1) * (b^(Suc q) - 1))" (is "?P \ ?Q") proof- have "D q c b = (2^c - 1) * E q b" by (auto simp: E_def D_def sum_distrib_left sum_distrib_right) moreover have "b \ 2" using assms by fastforce ultimately show ?thesis by (smt e_geom_series mult.left_commute mult_cancel_left) qed (* [D] 4.21 *) definition F :: "nat \ nat \ nat \ nat" where "(F q c b) = (\t = 0..q. 2^c * b^t)" lemma f_geom_series: assumes "b = 2^(Suc c)" shows "(F q c b = f) \ ( (b-1) * f = 2^c * (b^(Suc q) - 1) )" proof- have "F q c b = 2^c * E q b" by (auto simp: E_def F_def sum_distrib_left sum_distrib_right) moreover have "b \ 2" using assms by fastforce ultimately show ?thesis by (smt e_geom_series mult.left_commute mult_cancel_left) qed (* AUX LEMMAS *) lemma aux_lt_implies_mask: assumes "a < 2^k" shows "\r\k. a \ r = 0" using nth_bit_def assms apply auto proof - (* found by e *) fix r :: nat assume a1: "a < 2 ^ k" assume a2: "k \ r" - have "a div 2 ^ k = 0" - using a1 Euclidean_Division.div_eq_0_iff by blast + from a1 have "a div 2 ^ k = 0" + by simp then have "2 = (0::nat) \ a < 2 ^ r" using a2 by (metis (no_types) div_le_mono nat_zero_less_power_iff neq0_conv not_le power_diff) then show "a div 2 ^ r mod 2 = 0" by simp qed lemma lt_implies_mask: fixes a b :: nat assumes "\k. a < 2^k \ (\r b" proof - obtain k where assms: "a < 2^k \ (\rr r \ b \ r" using nth_bit_bounded by (simp add: \a < 2 ^ k \ (\r r = 1)\) hence k2: "\r\k. a \ r = 0" using aux_lt_implies_mask assms by auto show ?thesis using masks_leq_equiv by auto (metis k1 k2 le0 not_less) qed lemma mask_conversed_shift: fixes a b k :: nat assumes asm: "a \ b" shows "a * 2^k \ b * 2^k" proof - have shift: "x \ y \ 2*x \ 2*y" for x y by (induction x; auto) have "a * 2 ^ k \ b * 2 ^ k \ 2 * (a * 2 ^ k) \ 2 * (b * 2 ^ k)" for k using shift[of "a*2^k" "b*2^k"] by auto thus ?thesis by (induction k; auto simp: asm shift algebra_simps) qed lemma base_summation_bound: fixes c q :: nat and f :: "(nat \ nat)" defines b: "b \ B c" assumes bound: "\t. f t < 2 ^ Suc c - (1::nat)" shows "(\t = 0..q. f t * b^t) < b^(Suc q)" proof (induction q) case 0 then show ?case using B_def b bound less_imp_diff_less not_less_eq by auto blast next case (Suc q) have "(\t = 0..Suc q. f t * b ^ t) = f (Suc q) * b ^ (Suc q) + (\t = 0..q. f t * b ^ t)" by (auto cong: sum.cong) also have "... < (f (Suc q) + 1) * b ^ (Suc q)" using Suc.IH by auto also have "... < b * b ^ (Suc q)" by (metis bound b less_diff_conv B_def mult_less_cancel2 zero_less_numeral zero_less_power) finally show ?case by auto qed lemma mask_conserved_sum: fixes y c q :: nat and x :: "(nat \ nat)" defines b: "b \ B c" assumes mask: "\t. x t \ y" assumes xlt: "\t. x t \ 2 ^ c - Suc 0" assumes ylt: "y \ 2 ^ c - Suc 0" shows "(\t = 0..q. x t * b^t) \ (\t = 0..q. y * b^t)" proof (induction q) case 0 then show ?case using mask by auto next case (Suc q) have xb: "\t. x t < 2^Suc c - Suc 0" using xlt by (smt Suc_pred leI le_imp_less_Suc less_SucE less_trans n_less_m_mult_n numeral_2_eq_2 power.simps(2) zero_less_numeral zero_less_power) have yb: "y < 2^c" using ylt b B_def leI order_trans by fastforce have sumxlt: "(\t = 0..q. x t * b ^ t) < b^(Suc q)" using base_summation_bound xb b B_def by auto have sumylt: "(\t = 0..q. y * b ^ t) < b^(Suc q)" using base_summation_bound yb b B_def by auto have "((\t = 0..Suc q. x t * b ^ t) \ (\t = 0..Suc q. y * b ^ t)) = (x (Suc q) * b^Suc q + (\t = 0..q. x t * b ^ t) \ y * b^Suc q + (\t = 0..q. y * b ^ t))" by (auto simp: atLeast0_lessThan_Suc add.commute) also have "... = (x (Suc q) * b^Suc q \ y * b^Suc q) \ (\t = 0..q. x t * b ^ t) \ (\t = 0..q. y * b ^ t)" using mask_linear[where ?t = "Suc c * Suc q"] sumxlt sumylt Suc.IH b B_def apply auto apply (smt mask mask_conversed_shift power_Suc power_mult power_mult_distrib) by (smt mask mask_linear power_Suc power_mult power_mult_distrib) finally show ?case using mask_linear Suc.IH B_def by (metis (no_types, lifting) b mask mask_conversed_shift power_mult) qed lemma aux_powertwo_digits: fixes k c :: nat assumes "k < c" shows "nth_bit (2^c) k = 0" proof - have h: "(2::nat) ^ c div 2 ^ k = 2 ^ (c - k)" by (simp add: assms less_imp_le power_diff) thus ?thesis by (auto simp: h nth_bit_def assms) qed lemma obtain_digit_rep: fixes x c :: nat shows "x && 2^c = (\t 2^c" by (simp add: lm0245) hence "x && 2^c \ 2^c" by (simp add: masks_leq) hence h: "x && 2^c < 2^Suc c" by (smt Suc_lessD le_neq_implies_less lessI less_trans_Suc n_less_m_mult_n numeral_2_eq_2 power_Suc zero_less_power) have "\t. (x && 2^c) \ t = (nth_bit x t) * (nth_bit (2^c) t)" using bitAND_digit_mult by auto then show ?thesis using h digit_sum_repr[of "(x && 2^c)" "Suc c"] by (auto) (simp add: mult.commute semiring_normalization_rules(19)) qed lemma nth_digit_bitAND_equiv: fixes x c :: nat shows "2^c * nth_bit x c = (x && 2^c)" proof - have d1: "nth_bit (2^c) c = 1" using nth_bit_def by auto moreover have "x && 2^c = (2::nat)^c * (x \ c) * (((2::nat)^c) \ c) + (\t t) * (((2::nat)^c) \ t))" using obtain_digit_rep by (auto cong: sum.cong) moreover have "(\t x" assumes "x < 2 ^ Suc c" shows "nth_bit x c = 1" proof - obtain b where b: "x = 2^c + b" using assms(1) le_Suc_ex by auto have bb: "b < 2^c" using assms(2) b by auto have "(2 ^ c + b) div 2 ^ c mod 2 = (1 + b div 2 ^ c) mod 2" by (auto simp: div_add_self1) also have "... = 1" by (auto simp: bb) finally show ?thesis by (simp only: nth_bit_def b) qed lemma aux_bitAND_distrib: "2 * (a && b) = (2 * a) && (2 * b)" by (induct a b rule: bitAND_nat.induct; auto) lemma bitAND_distrib: "2^c * (a && b) = (2^c * a) && (2^c * b)" proof (induction c) case 0 then show ?case by auto next case (Suc c) have "2 ^ Suc c * (a && b) = 2 * (2 ^ c * (a && b))" by auto also have "... = 2 * ((2^c * a) && (2^c * b))" using Suc.IH by auto also have "... = ((2^Suc c * a) && (2^Suc c * b))" using aux_bitAND_distrib[of "2^c * a" "2^c * b"] by (auto simp add: ab_semigroup_mult_class.mult_ac(1)) finally show ?case by auto qed lemma bitAND_linear_sum: fixes x y :: "nat \ nat" and c :: nat and q :: nat defines b: "b == 2 ^ Suc c" assumes xb: "\t. x t < 2 ^ Suc c - 1" assumes yb: "\t. y t < 2 ^ Suc c - 1" shows "(\t = 0..q. (x t && y t) * b^t) = (\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)" proof (induction q) case 0 then show ?case by (auto simp: b B_def) next case (Suc q) have "(\t = 0..Suc q. (x t && y t) * b ^ t) = (x (Suc q) && y (Suc q)) * b ^ Suc q + (\t = 0..q. (x t && y t) * b ^ t)" by (auto cong: sum.cong) moreover have h0: "(x (Suc q) && y (Suc q)) * b ^ Suc q = (x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)" using b bitAND_distrib by (auto) (smt mult.commute power_Suc power_mult) moreover have h1: "(\t = 0..q. (x t && y t) * b ^ t) = (\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)" using Suc.IH by auto ultimately have h2: "(\t = 0..Suc q. (x t && y t) * b ^ t) = ((x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)) + ((\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t))" by auto have sumxb: "(\t = 0..q. x t * b ^ t) < b ^ Suc q" using base_summation_bound xb b B_def by auto have sumyb: "(\t = 0..q. y t * b ^ t) < b ^ Suc q" using base_summation_bound yb b B_def by auto have h3: "((x (Suc q) * b^Suc q) && (y (Suc q) * b^Suc q)) + ((\t = 0..q. x t * b^t) && (\t = 0..q. y t * b^t)) = ((\t = 0..q. x t * b^t) + x (Suc q) * b^Suc q) && ((\t = 0..q. y t * b^t) + y (Suc q) * b^Suc q)" using sumxb sumyb bitAND_linear h2 h0 by (auto) (smt add.commute b power_Suc power_mult) thus ?case using h2 by (auto cong: sum.cong) qed lemma dmask_aux0: fixes x :: nat assumes "x > 0" shows "(2 ^ x - Suc 0) div 2 = 2 ^ (x - 1) - Suc 0" proof - have 0: "(2^x - Suc 0) div 2 = (2^x - 2) div 2" by (smt Suc_diff_Suc Suc_pred assms dvd_power even_Suc even_Suc_div_two nat_power_eq_Suc_0_iff neq0_conv numeral_2_eq_2 zero_less_diff zero_less_power) (* can do manual parity distinction *) moreover have divides: "(2::nat) dvd 2^x" by (simp add: assms dvd_power[of x "2::nat"]) moreover have "(2^x - 2::nat) div 2 = 2^x div 2 - 2 div 2" using div_plus_div_distrib_dvd_left[of "2" "2^x" "2"] divides by auto moreover have "... = 2 ^ (x - 1) - Suc 0" by (simp add: Suc_leI assms power_diff) ultimately have 1: "(2 ^ x - Suc 0) div 2 = 2 ^ (x - 1) - Suc 0" by (smt One_nat_def) thus ?thesis by simp qed lemma dmask_aux: fixes c :: nat shows "d < c \ (2^c - Suc 0) div 2^d = 2 ^ (c - d) - Suc 0" proof (induction d) case 0 then show ?case by (auto) next case (Suc d) have d: "d < c" using Suc.prems by auto have "(2 ^ c - Suc 0) div 2 ^ Suc d = (2 ^ c - Suc 0) div 2 ^ d div 2" by (auto) (metis mult.commute div_mult2_eq) also have "... = (2 ^ (c - d) - Suc 0) div 2" by (subst Suc.IH) (auto simp: d) also have "... = 2 ^ (c - Suc d) - Suc 0" apply (subst dmask_aux0[of "c - d"]) using d by (auto) finally show ?case by auto qed (* REGISTERS *) lemma register_cells_masked: fixes l :: register and t :: nat and ic :: configuration and p :: program assumes cells_bounded: "cells_bounded ic p c" assumes l: "l < length (snd ic)" shows "R ic p l t \ 2^c - 1" proof - have a: "R ic p l t \ 2^c - 1" using cells_bounded less_Suc_eq_le using l by fastforce have b: "r < c \ nth_bit (2^c - 1) r = 1" for r apply (auto simp: nth_bit_def) apply (subst dmask_aux) apply (auto) by (metis Suc_pred dvd_power even_Suc mod_0_imp_dvd not_mod2_eq_Suc_0_eq_0 zero_less_diff zero_less_numeral zero_less_power) show ?thesis using lt_implies_mask cells_bounded l by (auto) (metis One_nat_def b) qed lemma lm04_15_register_masking: fixes c :: nat and l :: register and ic :: configuration and p :: program and q :: nat defines "b == B c" defines "d == D q c b" assumes cells_bounded: "cells_bounded ic p c" assumes l: "l < length (snd ic)" defines "r == RLe ic p b q" shows "r l \ d" proof - have "\t. R ic p l t \ 2^c - 1" using cells_bounded l by (rule register_cells_masked) hence rmasked: "\t. R ic p l t \ 2^c - 1" by (intro allI) have rlt: "\t. R ic p l t \ 2^c - 1" using cells_bounded less_Suc_eq_le l by fastforce have rlmasked: "(\t = 0..q. R ic p l t * b^t) \ (\t = 0..q. (2^c - 1) * b^t)" using rmasked rlt b_def B_def mask_conserved_sum by (auto) thus ?thesis by (auto simp: r_def d_def D_def RLe_def mult.commute cong: sum.cong) qed (* ZERO INDICATORS *) lemma zero_cells_masked: fixes l :: register and t :: nat and ic :: configuration and p :: program assumes l: "l < length (snd ic)" shows "Z ic p l t \ 1" proof - have "nth_bit 1 0 = 1" by (auto simp: nth_bit_def) thus ?thesis apply (auto) apply (rule lt_implies_mask) by (metis (full_types) One_nat_def Suc_1 Z_bounded less_Suc_eq_le less_one power_one_right) qed lemma lm04_15_zero_masking: fixes c :: nat and l :: register and ic :: configuration and p :: program and q :: nat defines "b == B c" defines "e == E q b" assumes cells_bounded: "cells_bounded ic p c" assumes l: "l < length (snd ic)" assumes c: "c > 0" defines "z == ZLe ic p b q" shows "z l \ e" proof - have "\t. Z ic p l t \ 1" using l by (rule zero_cells_masked) hence zmasked: "\t. Z ic p l t \ 1" by (intro allI) have zlt: "\t. Z ic p l t \ 2 ^ c - 1" using cells_bounded less_Suc_eq_le by fastforce have 1: "(1::nat) \ 2 ^ c - 1" using c by (simp add: Nat.le_diff_conv2 numeral_2_eq_2 self_le_power) have rlmasked: "(\t = 0..q. Z ic p l t * b^t) \ (\t = 0..q. 1 * b^t)" using zmasked zlt 1 b_def B_def mask_conserved_sum[of "Z ic p l" 1] by (auto) thus ?thesis by (auto simp: z_def e_def E_def ZLe_def mult.commute cong: sum.cong) qed (* Relation between zero indicator and register *) lemma lm04_19_zero_register_relations: fixes c :: nat and l :: register and t :: nat and ic :: configuration and p :: program assumes cells_bounded: "cells_bounded ic p c" assumes l: "l < length (snd ic)" defines "z == Z ic p" defines "r == R ic p" shows "2^c * z l t = (r l t + 2^c - 1) && 2^c" proof - have a1: "R ic p l t \ 0 \ 2^c \ R ic p l t + 2^c - 1" by auto have a2: "R ic p l t + 2^c - 1 < 2^Suc c" using cells_bounded by (simp add: l less_imp_diff_less) have "Z ic p l t = nth_bit (R ic p l t + 2^c - 1) c" apply (cases "R ic p l t = 0") subgoal by (auto simp add: Z_def R_def nth_bit_def) subgoal using cells_bounded bitAND_single_digit a1 a2 Z_def by auto done also have "2^c * nth_bit (R ic p l t + 2^c - 1) c = ((R ic p l t + 2^c - 1) && 2^c)" using nth_digit_bitAND_equiv by auto finally show ?thesis by (auto simp: z_def r_def) qed lemma lm04_20_zero_definition: fixes c :: nat and l :: register and ic :: configuration and p :: program and q :: nat defines "b == B c" defines "f == F q c b" defines "d == D q c b" assumes cells_bounded: "cells_bounded ic p c" assumes l: "l < length (snd ic)" assumes c: "c > 0" defines "z == ZLe ic p b q" defines "r == RLe ic p b q" shows "2^c * z l = (r l + d) && f" proof - have "\t. 2^c * Z ic p l t = (R ic p l t + 2^c - 1) && 2^c" by (rule lm04_19_zero_register_relations cells_bounded l) + hence raw_sums: "(\t = 0..q. 2^c * Z ic p l t * b^t) = (\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t)" by auto have "(\t = 0..q. 2^c * Z ic p l t * b^t) = 2^c * (\t = 0..q. Z ic p l t * b^t)" by (auto simp: sum_distrib_left mult.assoc cong: sum.cong) also have "... = 2^c * z l" by (auto simp: z_def ZLe_def mult.commute) finally have lhs: "(\t = 0..q. 2^c * Z ic p l t * b^t) = 2^c * z l" by auto have "(\t = 0..q. (R ic p l t + (2^c - 1)) * b^t) = (\t = 0..q. R ic p l t * b^t + (2^c - 1) * b^t)" apply (rule sum.cong) apply (auto simp: add.commute mult.commute) subgoal for x using distrib_left[of "b^x" "R ic p l x" "2^c - 1"] by (auto simp: algebra_simps) done also have "... = (\t = 0..q. (R ic p l t * b^t)) + (\t = 0..q. (2^c - 1) * b^t)" by (rule sum.distrib) also have "... = r l + d" by (auto simp: r_def RLe_def d_def D_def mult.commute) finally have split_sums: "(\t = 0..q. (R ic p l t + (2^c - 1)) * b^t) = r l + d" by auto have a1: "(2::nat) ^ c < (2::nat) ^ Suc c - 1" using c by (induct c, auto, fastforce) have a2: "\t. R ic p l t + 2 ^ c - 1 \ 2 ^ Suc c" using cells_bounded B_def by (simp add: less_imp_diff_less l) (simp add: Suc_leD l less_imp_le_nat numeral_Bit0) have "(\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t) = (\t = 0..q. (R ic p l t + 2^c - 1) * b^t) && (\t = 0..q. 2^c * b^t)" using bitAND_linear_sum[of "\t. R ic p l t + 2^c - 1" "c" "\t. 2^c"] cells_bounded b_def B_def a1 a2 apply auto by (smt One_nat_def Suc_less_eq Suc_pred a1 add.commute add_gr_0 l mult_2 nat_add_left_cancel_less power_Suc zero_less_numeral zero_less_power) also have "... = (\t = 0..q. (R ic p l t + 2^c - 1) * b^t) && f" by (auto simp: f_def F_def) also have "... = (r l + d) && f" using split_sums by auto finally have rhs: "(\t = 0..q. ((R ic p l t + 2^c - 1) && 2^c) * b^t) = (r l + d) && f" by auto show ?thesis using raw_sums lhs rhs by auto qed lemma state_mask: fixes c :: nat and l :: register and ic :: configuration and p :: program and q :: nat and a :: nat defines "b \ B c" and "m \ length p - 1" defines "e \ E q b" assumes is_val: "is_valid_initial ic p a" and q: "q > 0" and "c > 0" assumes terminate: "terminates ic p q" shows "SKe ic p b q k \ e" proof - have "1 \ 2 ^ c - Suc 0" using \c>0\ by (metis One_nat_def Suc_leI one_less_numeral_iff one_less_power semiring_norm(76) zero_less_diff) have Smask: "S ic p k t \ 1" for t by (simp add: S_def) have Sbound: "S ic p k t \ 2 ^ c - Suc 0" for t using \1\2^c-Suc 0\ by (simp add: S_def) have rlmasked: "(\t = 0..q. S ic p k t * b^t) \ (\t = 0..q. 1 * b^t)" using b_def B_def Smask Sbound mask_conserved_sum[of "S ic p k" 1] \1 \ 2^c-Suc 0\ by auto thus ?thesis using SKe_def e_def E_def by (auto simp: mult.commute) qed lemma state_sum_mask: fixes c :: nat and l :: register and ic :: configuration and p :: program and q :: nat and a :: nat defines "b \ B c" and "m \ length p - 1" defines "e \ E q b" assumes is_val: "is_valid_initial ic p a" and q: "q > 0" and "c > 0" and "b > 1" assumes "M\m" assumes terminate: "terminates ic p q" shows "(\k\M. SKe ic p b q k) \ e" proof - have e_aux: "nth_digit e t b = (if t\q then 1 else 0)" for t unfolding e_def E_def b_def B_def using `b>1` b_def nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] by (auto simp: b_def B_def) have state_unique: "\k\m. S ic p k t = 1 \ (\j\k. S ic p j t = 0)" for t using S_def by (induction t, auto) have h1: "\t. nth_digit (\k\M. SKe ic p b q k) t b \ (if t\q then 1 else 0)" proof - { fix t have aux_bound_1: "(\k\M. S ic p k t') \ 1" for t' proof (cases "\k\M. S ic p k t' = 1") case True then obtain k where k: "k\M \ S ic p k t' = 1" by auto moreover have "\j\M. j \ k \ S ic p j t' = 0" using state_unique `M<=m` k S_def by (auto) (presburger) ultimately have "(\k\M. S ic p k t') = 1" using S_def by auto then show ?thesis by auto next case False then show ?thesis using S_bounded by (auto) (metis (no_types, lifting) S_def atMost_iff eq_imp_le le_SucI sum_nonpos) qed hence aux_bound_2: "\t'. (\k\M. S ic p k t') < 2^c" by (metis Suc_1 `c>0` le_less_trans less_Suc_eq one_less_power) have h2: "(\k\M. SKe ic p b q k) = (\t = 0..q. \k\M. b ^ t * S ic p k t)" unfolding SKe_def using sum.swap by auto hence "(\k\M. SKe ic p b q k) = (\t = 0..q. b^t * (\k\M. S ic p k t))" unfolding SKe_def by (simp add: sum_distrib_left) hence "nth_digit (\k\M. SKe ic p b q k) t b = (if t\q then (\k\M. S ic p k t) else 0)" using `c>0` aux_bound_2 h2 unfolding SKe_def using nth_digit_gen_power_series[of "\t. (\k\M. S ic p k t)" "c" "q" "t"] by (smt B_def Groups.mult_ac(2) assms(7) aux_bound_1 b_def le_less_trans sum.cong) hence "nth_digit (\k\M. SKe ic p b q k) t b \ (if t\q then 1 else 0)" using aux_bound_1 by auto } thus ?thesis by auto qed moreover have "\t>q. nth_digit (\k\M. SKe ic p b q k) t b = 0" by (metis (full_types) h1 le_0_eq not_less) ultimately have "\t. \ik\M. SKe ic p b q k) t b \ i \ nth_digit e t b \ i" using aux_lt_implies_mask linorder_neqE_nat e_aux by (smt One_nat_def le_0_eq le_SucE less_or_eq_imp_le nat_zero_less_power_iff numeral_2_eq_2 zero_less_Suc) hence "\t. \ik\M. SKe ic p b q k) \ (Suc c * t + i) \ e \ (Suc c * t + i)" using digit_gen_pow2_reduct[where ?c = "Suc c" and ?a = "(\k\M. SKe ic p b q k)"] using digit_gen_pow2_reduct[where ?c = "Suc c" and ?a = e] by (simp add: b_def B_def) moreover have "\j. \t i. i < Suc c \ j = (Suc c * t + i)" using mod_less_divisor zero_less_Suc by (metis add.commute mod_mult_div_eq) ultimately have "\j. (\k\M. SKe ic p b q k) \ j \ e \ j" by metis thus ?thesis using masks_leq_equiv by auto qed end diff --git a/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy b/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy --- a/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy +++ b/thys/DPRM_Theorem/Register_Machine/MultipleToSingleSteps.thy @@ -1,1148 +1,1144 @@ subsection \From multiple to single step relations\ theory MultipleToSingleSteps imports MachineEquations CommutationRelations "../Diophantine/Binary_And" begin text \This file contains lemmas that are needed to prove the <-- direction of conclusion4.5 in the file MachineEquationEquivalence. In particular, it is shown that single step equations follow from the multiple step relations. The key idea of Matiyasevich's proof is to code all register and state values over the time into one large number. A further central statement in this file shows that the decoding of these numbers back to the single cell contents is indeed correct.\ context fixes a :: nat and ic:: configuration and p :: program and q :: nat and r z :: "register \ nat" and s :: "state \ nat" and b c d e f :: nat and m n :: nat and Req Seq Zeq (* These variables are to store the single cell entries obtained from the rows r s and z in the protocol *) assumes m_def: "m \ length p - 1" and n_def: "n \ length (snd ic)" assumes is_val: "is_valid_initial ic p a" (* rm_equations *) assumes m_eq: "mask_equations n r z c d e f" and r_eq: "reg_equations p r z s b a n q" and s_eq: "state_equations p s z b e q m" and c_eq: "rm_constants q c b d e f a" assumes Seq_def: "Seq = (\k t. nth_digit (s k) t b)" and Req_def: "Req = (\l t. nth_digit (r l) t b)" and Zeq_def: "Zeq = (\l t. nth_digit (z l) t b)" begin text \Basic properties\ lemma n_gt0: "n>0" using n_def is_val is_valid_initial_def[of "ic" "p" "a"] is_valid_def by auto lemma f_def: "f = (\t = 0..q. 2^c * b^t)" using c_eq by (simp add: rm_constants_def F_def) lemma e_def: "e = (\t = 0..q. b^t)" using c_eq by (simp add: rm_constants_def E_def) lemma d_def: "d = (\t = 0..q. (2^c - 1) * b^t)" using c_eq by (simp add: D_def rm_constants_def) lemma b_def: "b = 2^(Suc c)" using c_eq by (simp add: B_def rm_constants_def) lemma b_gt1: "b > 1" using c_eq B_def rm_constants_def by auto lemma c_gt0: "c > 0" using rm_constants_def c_eq by auto lemma h0: "1 < (2::nat)^c" using c_gt0 one_less_numeral_iff one_less_power semiring_norm(76) by blast (* this even implies b > 2, in fact b \ 4 *) lemma rl_fst_digit_zero: assumes "l < n" shows "nth_digit (r l) t b \ c = 0" proof - have "2 ^ c - (Suc 0) < 2 ^ Suc c" using c_gt0 by (simp add: less_imp_diff_less) hence "\t. nth_digit d t b = (if t\q then 2^c - 1 else 0)" using nth_digit_gen_power_series[of "\x. 2^c - 1" "c"] d_def c_gt0 b_def by (simp add: d_def) then have d_lead_digit_zero: "\t. (nth_digit d t b) \ c = 0" by (auto simp: nth_bit_def) from m_eq have "(r l) \ d" by (simp add: mask_equations_def assms) then have "\k. (r l) \ k \ d \ k" by (auto simp: masks_leq_equiv) then have "\t. (nth_digit (r l) t b \ c) \ (nth_digit d t b \ c)" using digit_gen_pow2_reduct[where ?c = "Suc c"] by (auto simp: b_def) thus ?thesis by (auto simp: d_lead_digit_zero) qed lemma e_mask_bound: assumes "x \ e" shows "nth_digit x t b \ 1" proof - have x_bounded: "nth_digit x t' b \ nth_digit e t' b" for t' proof - have "\t'. x \ t' \ e \ t'" using assms masks_leq_equiv by auto then have k_lt_c: "\t'. \k' k' \ nth_digit e t' b \ k'" using digit_gen_pow2_reduct by (auto simp: b_def) (metis power_Suc) have "k\Suc c \ x mod (2 ^ Suc c) div 2 ^ k = 0" for k x::nat - by (metis Euclidean_Division.div_eq_0_iff One_nat_def Suc_le_lessD b_gt1 b_def less_le_trans - mod_less_divisor nat_power_less_imp_less not_less not_less0 not_less_eq_eq - numeral_2_eq_2 zero_less_Suc) + by (simp only: drop_bit_take_bit flip: take_bit_eq_mod drop_bit_eq_div) simp then have "\k\Suc c. nth_digit x y b \ k = 0" for x y using b_def nth_bit_def nth_digit_def by auto then have k_gt_c: "\t'. \k'\Suc c. nth_digit x t' b \ k' \ nth_digit e t' b \ k'" by auto from k_lt_c k_gt_c have "nth_digit x t' b \ nth_digit e t' b" for t' using bitwise_leq by (meson not_le) thus ?thesis by auto qed have "\k. Suc 0 < 2^c" using c_gt0 h0 by auto hence e_aux: "nth_digit e tt b \ 1" for tt using e_def b_def c_gt0 nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] by auto show ?thesis using e_aux[of "t"] x_bounded[of "t"] using le_trans by blast qed (* --- individual bounds --- *) lemma sk_bound: shows "\t k. k\length p - 1 \ nth_digit (s k) t b \ 1" proof - have "\k\length p - 1. s k \ e" using s_eq state_equations_def m_def by auto thus ?thesis using e_mask_bound by auto qed lemma sk_bitAND_bound: shows "\t k. k\length p - 1 \ nth_digit (s k && x k) t b \ 1" using bitAND_nth_digit_commute sk_bound bitAND_lt masks_leq by (auto simp: b_def) (meson dual_order.trans) lemma s_bound: shows "\jM\m. (\k\M. s k) \ e" using s_eq state_equations_def by auto lemma sk_sum_bound: shows "\t M. M\length p - 1 \ nth_digit (\k\M. s k) t b \ 1" using sk_sum_masked e_mask_bound m_def by auto lemma sum_sk_bound: shows "(\k\length p - 1. nth_digit (s k) t b) \ 1" proof - have "\t m. m \ length p - 1 \ nth_digit (sum s {0..m}) t b < 2 ^ c" using sk_sum_bound b_def c_gt0 h0 by (metis atLeast0AtMost le_less_trans) moreover have "\t k. k \ length p - 1 \ nth_digit (s k) t b < 2 ^ c" using sk_bound b_def c_gt0 h0 by (metis le_less_trans) ultimately have "nth_digit (\k\length p - 1. s k) t b = (\k\length p - 1. nth_digit (s k) t b)" using b_def c_gt0 using finite_sum_nth_digit_commute2[where ?M = "length p - 1"] by (simp add: atMost_atLeast0) thus ?thesis using sk_sum_bound by (metis order_refl) qed lemma bitAND_sum_lt: "(\k\length p - 1. nth_digit (s k && x k) t b) \ (\k\length p - 1. Seq k t)" proof - have "(\k\length p - 1. nth_digit (s k && x k) t b) = (\k\length p - 1. nth_digit (s k) t b && nth_digit (x k) t b)" using bitAND_nth_digit_commute b_def by auto also have "... \ (\k\length p - 1. nth_digit (s k) t b)" using bitAND_lt by (simp add: sum_mono) finally show ?thesis using Seq_def by auto qed lemma states_unique_RAW: "\k\m. Seq k t = 1 \ (\j\m. j \ k \ Seq j t = 0)" proof - { fix k assume "k\m" assume skt_1: "Seq k t = 1" have "\j\m. j \ k \ Seq j t = 0" proof - { fix j assume "j\m" assume "j \ k" let ?fct = "(\k. Seq k t)" have "Seq j t = 0" proof (rule ccontr) assume "Seq j t \ 0" then have "Seq j t + Seq k t > 1" using skt_1 by auto moreover have "sum ?fct {0..m} \ sum ?fct {j, k}" using `j\m` `k\m` sum_mono2 by (metis atLeastAtMost_iff empty_subsetI finite_atLeastAtMost insert_subset le0) ultimately have "(\k\m. Seq k t) > 1" by (simp add: \j \ k\ atLeast0AtMost) thus "False" using sum_sk_bound[where ?t = t] by (auto simp: Seq_def m_def) qed } thus ?thesis by auto qed } thus ?thesis by auto qed lemma block_sum_radd_bound: shows "\t. (\R+ p l (\k. nth_digit (s k) t b)) \ 1" proof - { fix t have "(\R+ p l (\k. Seq k t)) \ (\k\length p - 1. Seq k t)" unfolding sum_radd.simps by (simp add: atMost_atLeast0 sum_mono) hence "(\R+ p l (\k. Seq k t)) \ 1" using sum_sk_bound[of t] Seq_def using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma block_sum_rsub_bound: shows "\t. (\R- p l (\k. nth_digit (s k && z l) t b)) \ 1" proof - { fix t have "(\R- p l (\k. nth_digit (s k && z l) t b)) \ (\k\length p - 1. nth_digit (s k && z l) t b)" unfolding sum_rsub.simps by (simp add: atMost_atLeast0 sum_mono) also have "... \ (\k\length p - 1. Seq k t)" using bitAND_sum_lt[where ?x = "\k. z l"] by blast finally have "(\R- p l (\k. nth_digit (s k && z l) t b)) \ 1" using sum_sk_bound[of t] Seq_def using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma block_sum_rsub_special_bound: shows "\t. (\R- p l (\k. nth_digit (s k) t b)) \ 1" proof - { fix t have "(\R- p l (\k. nth_digit (s k) t b)) \ (\k\length p - 1. nth_digit (s k) t b)" unfolding sum_rsub.simps by (simp add: atMost_atLeast0 sum_mono) then have "(\R- p l (\k. nth_digit (s k) t b)) \ 1" using sum_sk_bound[of t] using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma block_sum_sadd_bound: shows "\t. (\S+ p j (\k. nth_digit (s k) t b)) \ 1" proof - { fix t have "(\S+ p j (\k. Seq k t)) \ (\k\length p - 1. Seq k t)" unfolding sum_sadd.simps by (simp add: atMost_atLeast0 sum_mono) hence "(\S+ p j (\k. Seq k t)) \ 1" using sum_sk_bound[of t] Seq_def using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma block_sum_ssub_bound: shows "\t. (\S- p j (\k. nth_digit (s k && z (l k)) t b)) \ 1" proof - { fix t have "(\S- p j (\k. nth_digit (s k && z (l k)) t b)) \ (\k\length p - 1. nth_digit (s k && z (l k)) t b)" unfolding sum_ssub_nzero.simps by (simp add: atMost_atLeast0 sum_mono) also have "... \ (\k\length p - 1. Seq k t)" using bitAND_sum_lt[where ?x = "\k. z (l k)"] by blast finally have "(\S- p j (\k. nth_digit (s k && z (l k)) t b)) \ 1" using sum_sk_bound[of t] Seq_def using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma block_sum_szero_bound: shows "\t. (\S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)) \ 1" proof - { fix t have "(\S0 p j (\k. nth_digit (s k && e - z (l k)) t b)) \ (\k\length p - 1. nth_digit (s k && e - z (l k)) t b)" unfolding sum_ssub_zero.simps by (simp add: atMost_atLeast0 sum_mono) also have "... \ (\k\length p - 1. Seq k t)" using bitAND_sum_lt[where ?x = "\k. e - z (l k)"] by blast finally have "(\S0 p j (\k. nth_digit (s k && e - z (l k)) t b)) \ 1" using sum_sk_bound[of t] Seq_def using dual_order.trans by blast } thus ?thesis using Seq_def by auto qed lemma sum_radd_nth_digit_commute: shows "nth_digit (\R+ p l (\k. s k)) t b = \R+ p l (\k. nth_digit (s k) t b)" proof - have a1: "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" using sk_bound h0 by (meson le_less_trans) have a2: "\t. (\R+ p l (\k. nth_digit (s k) t b)) < 2^c" using block_sum_radd_bound h0 by (meson le_less_trans) show ?thesis using a1 a2 c_gt0 b_def unfolding sum_radd.simps using sum_nth_digit_commute[where ?g = "\l k. isadd (p ! k) \ l = modifies (p ! k)"] by blast qed lemma sum_rsub_nth_digit_commute: shows "nth_digit (\R- p l (\k. s k && z l)) t b = \R- p l (\k. nth_digit (s k && z l) t b)" proof - have a1: "\t. \k\length p - 1. nth_digit (s k && z l) t b < 2^c" using sk_bitAND_bound[where ?x = "\k. z l"] h0 le_less_trans by blast have a2: "\t. (\R- p l (\k. nth_digit (s k && z l) t b)) < 2^c" using block_sum_rsub_bound h0 by (meson le_less_trans) show ?thesis using a1 a2 c_gt0 b_def unfolding sum_rsub.simps using sum_nth_digit_commute [where ?g = "\l k. issub (p ! k) \ l = modifies (p ! k)" and ?fct = "\k. s k && z l"] by blast qed lemma sum_sadd_nth_digit_commute: shows "nth_digit (\S+ p j (\k. s k)) t b = \S+ p j (\k. nth_digit (s k) t b)" proof - have a1: "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" using sk_bound h0 by (meson le_less_trans) have a2: "\t. (\S+ p j (\k. nth_digit (s k) t b)) < 2^c" using block_sum_sadd_bound h0 by (meson le_less_trans) show ?thesis using a1 a2 b_def c_gt0 unfolding sum_sadd.simps using sum_nth_digit_commute[where ?g = "\j k. isadd (p ! k) \ j = goes_to (p ! k)"] by blast qed lemma sum_ssub_nth_digit_commute: shows "nth_digit (\S- p j (\k. s k && z (l k))) t b = \S- p j (\k. nth_digit (s k && z (l k)) t b)" proof - have a1: "\t. \k\length p - 1. nth_digit (s k && z (l k)) t b < 2^c" using sk_bitAND_bound[where ?x = "\k. z (l k)"] h0 le_less_trans by blast have a2: "\t. (\S- p j (\k. nth_digit (s k && z (l k)) t b)) < 2^c" using block_sum_ssub_bound h0 by (meson le_less_trans) show ?thesis using a1 a2 b_def c_gt0 unfolding sum_ssub_nzero.simps using sum_nth_digit_commute [where ?g = "\j k. issub (p ! k) \ j = goes_to (p ! k)" and ?fct = "\k. s k && z (l k)"] by blast qed lemma sum_szero_nth_digit_commute: shows "nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b = \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" proof - have a1: "\t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b < 2^c" using sk_bitAND_bound[where ?x = "\k. e - z (l k)"] h0 le_less_trans by blast have a2: "\t. (\S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)) < 2^c" using block_sum_szero_bound h0 by (meson le_less_trans) show ?thesis using a1 a2 b_def c_gt0 unfolding sum_ssub_zero.simps using sum_nth_digit_commute [where ?g = "\j k. issub (p ! k) \ j = goes_to_alt (p ! k)" and ?fct = "\k. s k && e - z (l k)"] by blast qed lemma block_bound_impl_fst_digit_zero: assumes "nth_digit x t b \ 1" shows "(nth_digit x t b) \ c = 0" using assms apply (auto simp: nth_bit_def) by (metis (no_types, opaque_lifting) c_gt0 div_le_dividend le_0_eq le_Suc_eq mod_0 mod_Suc mod_div_trivial numeral_2_eq_2 power_eq_0_iff power_mod) lemma sum_radd_block_bound: shows "nth_digit (\R+ p l (\k. s k)) t b \ 1" using block_sum_radd_bound sum_radd_nth_digit_commute by auto lemma sum_radd_fst_digit_zero: shows "(nth_digit (\R+ p l s) t b) \ c = 0" using sum_radd_block_bound block_bound_impl_fst_digit_zero by auto lemma sum_sadd_block_bound: shows "nth_digit (\S+ p j (\k. s k)) t b \ 1" using block_sum_sadd_bound sum_sadd_nth_digit_commute by auto lemma sum_sadd_fst_digit_zero: shows "(nth_digit (\S+ p j s) t b) \ c = 0" using sum_sadd_block_bound block_bound_impl_fst_digit_zero by auto lemma sum_ssub_block_bound: shows "nth_digit (\S- p j (\k. s k && z (l k))) t b \ 1" using block_sum_ssub_bound sum_ssub_nth_digit_commute by auto lemma sum_ssub_fst_digit_zero: shows "(nth_digit (\S- p j (\k. s k && z (l k))) t b) \ c = 0" using sum_ssub_block_bound block_bound_impl_fst_digit_zero by auto lemma sum_szero_block_bound: shows "nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b \ 1" using block_sum_szero_bound sum_szero_nth_digit_commute by auto lemma sum_szero_fst_digit_zero: shows "(nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b) \ c = 0" using sum_szero_block_bound block_bound_impl_fst_digit_zero by auto lemma sum_rsub_special_block_bound: shows "nth_digit (\R- p l (\k. s k)) t b \ 1" proof - have a1: "\t k. k \ length p - 1 \ nth_digit (s k) t b < 2^c" using sk_bound h0 le_less_trans by blast have a2: "\t. \R- p l (\k. nth_digit (s k) t b) < 2^c" using block_sum_rsub_special_bound h0 le_less_trans by blast have "nth_digit (\R- p l (\k. s k)) t b = \R- p l (\k. nth_digit (s k) t b)" using a1 a2 b_def c_gt0 unfolding sum_rsub.simps using sum_nth_digit_commute[where ?g = "\l k. issub (p ! k) \ l = modifies (p ! k)"] by blast thus ?thesis using block_sum_rsub_special_bound by auto qed lemma sum_state_special_block_bound: shows "nth_digit (\S+ p j (\k. s k) + \S0 p j (\k. s k && (e - z (l k)))) t b \ 1" proof - have aux_sum_zero: "\S0 p j (\k. nth_digit (s k) t b && nth_digit (e - z (l k)) t b) \ \S0 p j (\k. nth_digit (s k) t b)" unfolding sum_ssub_zero.simps by (auto simp: bitAND_lt sum_mono) have aux_addsub_excl: "(if isadd (p ! k) then Seq k t else 0) + (if issub (p ! k) then Seq k t else 0) = (if isadd (p ! k) \ issub (p ! k) then Seq k t else 0)" for k by auto have aux_sum_add_lt: "\S+ p j (\k. Seq k t) \ (\k = 0..length p - 1. if isadd (p ! k) then Seq k t else 0)" unfolding sum_sadd.simps by (simp add: sum_mono) have aux_sum_sub_lt: "\S0 p j (\k. Seq k t) \ (\k = 0..length p - 1. if issub (p ! k) then Seq k t else 0)" unfolding sum_ssub_zero.simps by (simp add: sum_mono) have "nth_digit (\S+ p j (\k. s k) + \S0 p j (\k. s k && e - z (l k))) t b = nth_digit (\S+ p j (\k. s k)) t b + nth_digit (\S0 p j (\k. s k && e - z (l k))) t b" using sum_sadd_fst_digit_zero sum_szero_fst_digit_zero block_additivity by (auto simp: c_gt0 b_def) also have "... = \S+ p j (\k. nth_digit (s k) t b) + \S0 p j (\k. nth_digit (s k && e - z (l k)) t b)" by (simp add: sum_sadd_nth_digit_commute sum_szero_nth_digit_commute) also have "... \ \S+ p j (\k. Seq k t) + \S0 p j (\k. Seq k t)" using bitAND_nth_digit_commute aux_sum_zero unfolding Seq_def by (simp add: b_def) also have "... \ (\k = 0..length p - 1. if isadd (p ! k) then Seq k t else 0) + (\k = 0..length p - 1. if issub (p ! k) then Seq k t else 0)" using aux_sum_add_lt aux_sum_sub_lt by auto also have "... = (\k\length p - 1. if (isadd (p ! k) \ issub (p ! k)) then Seq k t else 0)" using aux_addsub_excl using sum.distrib[where ?g = "\k. if isadd (p ! k) then Seq k t else 0" and ?h = "\k. if issub (p ! k) then Seq k t else 0"] by (auto simp: aux_addsub_excl atMost_atLeast0) also have "... \ (\k\length p - 1. Seq k t)" by (smt eq_iff le0 sum_mono) finally show ?thesis using sum_sk_bound[of t] Seq_def by auto qed lemma sum_state_special_fst_digit_zero: shows "(nth_digit (\S+ p j (\k. s k) + \S0 p j (\k. s k && (e - z (modifies (p!k))))) t b) \ c = 0" using sum_state_special_block_bound block_bound_impl_fst_digit_zero by auto text \Main three redution lemmas: Zero Indicators, Registers, States\ lemma Z: assumes "l 0 then Suc 0 else 0)" proof - have cond: "2^c * (z l) = (r l + d) && f" using m_eq mask_equations_def assms by auto have d_block: "\t\q. nth_digit d t b = 2^c - 1" using d_def b_def less_imp_diff_less nth_digit_gen_power_series[of "\_. 2^c-1" "c"] c_gt0 by auto have rl_bound: "t\q \ nth_digit (r l) t b \ c= 0" for t by (simp add: assms rl_fst_digit_zero) have f_block: "\t\q. nth_digit f t b = 2^c" using f_def b_def less_imp_diff_less nth_digit_gen_power_series[of "\_. 2^c" c] c_gt0 by auto then have "\t\q.\k k = 0" by (simp add: aux_powertwo_digits) moreover have AND_gen: "\t\q.\k\c. nth_digit ((r l + d) && f) t b \ k = (nth_digit (r l + d) t b \ k) * nth_digit f t b \ k" using b_def digit_gen_pow2_reduct bitAND_digit_mult digit_gen_pow2_reduct le_imp_less_Suc by presburger ultimately have "\t\q.\k k = 0" using f_def by auto moreover have "(r l + d) && f < b ^ Suc q" using lm0245[of "r l + d" f] masks_leq[of "(r l + d) && f" f] f_def proof- have "2 < b" using b_def c_gt0 gr0_conv_Suc not_less_iff_gr_or_eq by fastforce then have "b^u + b^u < b* b^ u" for u using zero_less_power[of b u] mult_less_mono1[of 2 b "b^u"] by linarith then have "(\t\{..t\{..q. b^q" "{..t\{.. f" using lm0245[of "r l + d" f] masks_leq[of "(r l + d) && f" f] by auto ultimately show ?thesis by auto qed then have rldf0: "t>q \ nth_digit ((r l + d) && f) t b = 0" for t using nth_digit_def[of "r l + d && f" t b] div_less[of "r l + d && f" "b^t"] b_def power_increasing[of "Suc q" t b] by auto moreover have "\t>q.\k k = 0" using aux_lt_implies_mask rldf0 by fastforce ultimately have AND_zero: "\t.\k k = 0" using leI by blast have "0 k< Suc c \ nth_digit (z l) t b \ k = nth_digit ((r l + d) && f) (Suc t) b \ (k - 1)" for k using b_def nth_digit_bound digit_gen_pow2_reduct[of k "Suc c" "z l" t] aux_digit_shift[of "z l" c "t + c * t + k"] digit_gen_pow2_reduct[of "k-1" "Suc c" "z l * 2^c" "Suc t"] cond by (simp add: add.commute add.left_commute mult.commute) then have aux: "0 k< Suc c \ nth_digit (z l) t b \ k = 0" for k using AND_zero by auto have zl_formula: "nth_digit (z l) t b = nth_digit (z l) t b \ 0 " using b_def digit_sum_repr[of "nth_digit (z l) t b" "Suc c"] proof - have "nth_digit (z l) t b < 2 ^ Suc c \ nth_digit (z l) t b = (\k\{0.. k * 2 ^ k)" using b_def digit_sum_repr[of "nth_digit (z l) t b" "Suc c"] by (simp add: atLeast0LessThan) hence "nth_digit (z l) t b < 2 ^ Suc c \ nth_digit (z l) t b = nth_digit (z l) t b \ 0 + (\k\{0<.. k * 2 ^ k)" by (metis One_nat_def atLeastSucLessThan_greaterThanLessThan mult_numeral_1_right numeral_1_eq_Suc_0 power_0 sum.atLeast_Suc_lessThan zero_less_Suc) thus ?thesis using aux nth_digit_bound b_def by auto qed consider (tleq) "t\q" |(tgq) "t>q" by linarith then show ?thesis proof cases case tleq then have t_bound: "t \ q" by auto have "nth_digit ((r l + d) && f) t b \ c = (nth_digit (r l + d) t b \ c)" using f_block bitAND_single_digit AND_gen t_bound by auto moreover have "nth_digit (r l + d && f) t b < 2 ^ Suc c" using nth_digit_def b_def by simp ultimately have AND_all:"nth_digit ((r l + d) && f) t b = (nth_digit (r l + d) t b \ c) * 2^c" using AND_gen AND_zero using digit_sum_repr[of "nth_digit ((r l + d) && f) t b" "Suc c"] by auto then have "\k k = 0" using cond AND_zero by metis moreover have "nth_digit (2^c * (z l)) t b \ c = nth_digit (z l) t b \ 0" using digit_gen_pow2_reduct[of c "Suc c" " (2^c * (z l))" t] digit_gen_pow2_reduct[of 0 "Suc c" "z l" t] b_def by (simp add: aux_digit_shift mult.commute) ultimately have zl0: "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b \ 0" using digit_sum_repr[of "nth_digit (2^c * (z l)) t b" "Suc c"] nth_digit_bound b_def by auto have "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b" using zl0 zl_formula by auto then have zl_block: "nth_digit (z l) t b = nth_digit (r l + d) t b \ c" using AND_all cond by auto consider (g0) "Req l t > 0" | (e0) "Req l t = 0 " by auto then show ?thesis proof cases case e0 show ?thesis using e0 apply(auto simp add: Req_def Zeq_def) subgoal proof- assume asm: "nth_digit (r l) t b = 0" have add:"((nth_digit d t b) + (nth_digit (r l) t b)) \ c = 0" by (simp add: asm d_block nth_bit_def t_bound) from d_block t_bound have "nth_digit d (t-1) b \ c = 0" using add asm by auto then have "(nth_digit (r l + d) t b) \ c = 0" using add digit_wise_block_additivity[of "r l" "t" "c" "d" "c"] rl_bound[of "t-1"] b_def asm t_bound c_gt0 by auto then show ?thesis using zl_block by simp qed done next case g0 show ?thesis using g0 apply(auto simp add: Req_def Zeq_def) subgoal proof - assume "0 < nth_digit (r l) t b" then obtain k0 where k0_def: "nth_digit (r l) t b \ k0 = 1" using aux0_digit_wise_equiv by auto then have "k0\c" using nth_digit_bound[of "r l" t c] b_def aux_lt_implies_mask by (metis Suc_leI leI zero_neq_one) then have k0bound: "k0k k = 1" using d_block t_bound nth_bit_def[of "nth_digit d t b"] by (metis One_nat_def Suc_1 Suc_diff_Suc Suc_pred dmask_aux even_add even_power odd_iff_mod_2_eq_one one_mod_two_eq_one plus_1_eq_Suc zero_less_Suc zero_less_power) ultimately have "nth_digit d t b \ k0 = 1" by simp then have "bin_carry (nth_digit d t b) (nth_digit (r l) t b) (Suc k0) = 1" using k0_def sum_carry_formula carry_bounded less_eq_Suc_le by simp moreover have "\n. Suc k0 \ n \ n < c \ bin_carry (nth_digit d t b) (nth_digit (r l) t b) n = Suc 0 \ bin_carry (nth_digit d t b) (nth_digit (r l) t b) (Suc n) = Suc 0" subgoal for n proof- assume "n c = 1" using sum_digit_formula[of "nth_digit d t b" "nth_digit (r l) t b" c] d_block nth_bit_def t_bound assms rl_fst_digit_zero by auto from d_block t_bound have "nth_digit d (t-1) b \ c = 0" by (smt aux_lt_implies_mask diff_le_self diff_less le_eq_less_or_eq le_trans zero_less_numeral zero_less_one zero_less_power) then have "(nth_digit (r l + d) t b) \ c = 1" using add b_def t_bound block_additivity assms rl_fst_digit_zero c_gt0 d_block by (simp add: add.commute) then show ?thesis using zl_block by simp qed done qed next case tgq then have t_bound: "qk k = 0" using cond AND_zero by simp (* the next two statements are already proven above but here they rely on rl0 *) moreover have "nth_digit (2^c * (z l)) t b \ c = nth_digit (z l) t b \ 0" using digit_gen_pow2_reduct[of c "Suc c" " (2^c * (z l))" t] digit_gen_pow2_reduct[of 0 "Suc c" "z l" t] b_def by (simp add: aux_digit_shift mult.commute) ultimately have zl0: "nth_digit (2^c * (z l)) t b = 2^c * nth_digit (z l) t b \ 0" using digit_sum_repr[of "nth_digit (2^c * (z l)) t b" "Suc c"] nth_digit_bound b_def by auto have "0 k< Suc c \ nth_digit (z l) t b \ k = nth_digit ((r l + d) && f) (Suc t) b \ (k - 1)" for k using b_def nth_digit_bound digit_gen_pow2_reduct[of k "Suc c" "z l" t] aux_digit_shift[of "z l" c "t + c * t + k"] digit_gen_pow2_reduct[of "k-1" "Suc c" "z l * 2^c" "Suc t"] cond by (simp add: add.commute add.left_commute mult.commute) (* Using some simplification lemmas, the result follows *) then show ?thesis using Zeq_def Req_def cond rl0 zl0 rldf0 zl_formula t_bound by auto qed qed lemma zl_le_rl: "l z l \ r l" for l proof - assume l: "l < n" have "Zeq l t \ Req l t" for t using Z l by auto hence "nth_digit (z l) t b \ nth_digit (r l) t b" for t using Zeq_def Req_def by auto thus ?thesis using digitwise_leq b_gt1 by auto qed lemma modifies_valid: "\k\m. modifies (p!k) < n" proof - have reg_check: "program_register_check p n" using is_val by (cases ic, auto simp: is_valid_initial_def n_def is_valid_def) { fix k assume "k \ m" then have "p ! k \ set p" by (metis \k \ m\ add_eq_if diff_le_self is_val le_antisym le_trans m_def n_not_Suc_n not_less not_less0 nth_mem p_contains) then have "instruction_register_check n (p ! k)" using reg_check by (auto simp: list_all_def) then have "modifies (p!k) < n" by (cases "p ! k", auto simp: n_gt0) } thus ?thesis by auto qed lemma seq_bound: "k \ length p - 1 \ Seq k t \ 1" using sk_bound Seq_def by blast lemma skzl_bitAND_to_mult: assumes "k \ length p - 1" assumes "l < n" shows "nth_digit (z l) t b && nth_digit (s k) t b = (Zeq l t) * Seq k t" proof - have "nth_digit (z l) t b && nth_digit (s k) t b = (Zeq l t) && Seq k t" using Zeq_def Seq_def by simp also have "... = (Zeq l t) * Seq k t" using bitAND_single_bit_mult_equiv[of "(Zeq l t)" "Seq k t"] seq_bound Z assms by auto finally show ?thesis by auto qed lemma skzl_bitAND_to_mult2: assumes "k \ length p - 1" assumes "\k \ length p - 1. l k < n" shows "(1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b = (1 - Zeq (l k) t) * Seq k t" proof - have "(1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b = (1 - Zeq (l k) t) && Seq k t" using Zeq_def Seq_def by simp also have "... = (1 - Zeq (l k) t) * Seq k t" using bitAND_single_bit_mult_equiv[of "(1 - Zeq (l k) t)" "Seq k t"] seq_bound Z assms by auto finally show ?thesis by auto qed lemma state_equations_digit_commute: assumes "t < q" and "j \ m" defines "l \ \k. modifies (p!k)" shows "nth_digit (s j) (Suc t) b = (\S+ p j (\k. Seq k t)) + (\S- p j (\k. Zeq (l k) t * Seq k t)) + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" proof - define o' :: nat where "o' \ if j = 0 then 1 else 0" have o'_div: "o' div b = 0" using b_gt1 by (auto simp: o'_def) have l: "\k\length p - 1. (l k) < n" using l_def by (auto simp: m_def modifies_valid) have "\k. Suc 0 < 2^c" using c_gt0 h0 by auto hence e_aux: "\tt. nth_digit e tt b = (if tt\q then Suc 0 else 0)" using e_def b_def c_gt0 nth_digit_gen_power_series[of "\k. Suc 0" "c" "q"] by auto have zl_bounded: "k\m \ \t'. nth_digit (z (l k)) t' b \ nth_digit e t' b" for k proof - assume "k\m" from m_eq have "\l e" using mask_equations_def by auto then have "\lt'. (z l) \ t' \ e \ t'" using masks_leq_equiv by auto then have k_lt_c: "\lt'. \k' k' \ nth_digit e t' b \ k'" using digit_gen_pow2_reduct by (auto simp: b_def) (metis power_Suc) have "k\Suc c \ x mod (2 ^ Suc c) div 2 ^ k = 0" for k x::nat - by (metis Euclidean_Division.div_eq_0_iff One_nat_def Suc_le_lessD b_gt1 b_def less_le_trans - mod_less_divisor nat_power_less_imp_less not_less not_less0 not_less_eq_eq - numeral_2_eq_2 zero_less_Suc) + by (simp only: drop_bit_take_bit flip: take_bit_eq_mod drop_bit_eq_div) simp then have "\k\Suc c. nth_digit x y b \ k = 0" for x y using b_def nth_bit_def nth_digit_def by auto then have k_gt_c: "\lt'. \k'\Suc c. nth_digit (z l) t' b \ k' \ nth_digit e t' b \ k'" by auto from k_lt_c k_gt_c have "\lt'. nth_digit (z l) t' b \ nth_digit e t' b" using bitwise_leq by (meson not_le) thus ?thesis by (auto simp: modifies_valid l_def `k\m`) qed have "\t k. k\m \ nth_digit (e - z (l k)) t b = nth_digit e t b - nth_digit (z (l k)) t b" using zl_bounded block_subtractivity by (auto simp: c_gt0 b_def l_def) then have sum_szero_aux: "\t k. t k\m \ nth_digit (e - z (l k)) t b = 1-nth_digit (z (l k)) t b" using e_aux by auto have skzl_bound2: "\k\length p - 1. (l k) < n \ \t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b < 2^c" proof - assume l: "\k\length p - 1. (l k) < n" have "\t. \k\length p - 1. nth_digit (s k && (e - z (l k))) t b = nth_digit (s k) t b && nth_digit (e - z (l k)) t b" using bitAND_nth_digit_commute Zeq_def b_def by auto moreover have "\tk\length p - 1. nth_digit (s k) t b && nth_digit (e - z (l k)) t b = nth_digit (s k) t b && (1 - nth_digit (z (l k)) t b)" using sum_szero_aux by (simp add: m_def) moreover have "\t. \k\length p - 1. nth_digit (s k) t b && (1 - nth_digit (z (l k)) t b) \ nth_digit (s k) t b" using Z l using lm0245 masks_leq by (simp add: lm0244) moreover have "\t. \k\length p - 1. nth_digit (s k) t b < 2^c" using sk_bound h0 by (meson le_less_trans) ultimately show ?thesis using le_less_trans by (metis lm0244 masks_leq) qed (* START *) have "s j = o' + b*\S+ p j (\k. s k) + b*\S- p j (\k. s k && z (modifies (p!k))) + b*\S0 p j (\k. s k && (e - z (modifies (p!k))))" using s_eq state_equations_def \j\m\ by (auto simp: o'_def) then have "s j div b^Suc t mod b = ( o' + b*\S+ p j (\k. s k) + b*\S- p j (\k. s k && z (modifies (p!k))) + b*\S0 p j (\k. s k && (e - z (modifies (p!k)))) ) div b div b^t mod b" by (auto simp: algebra_simps div_mult2_eq) also have "... = (\S+ p j (\k. s k) + \S- p j (\k. s k && z (modifies (p!k))) + \S0 p j (\k. s k && (e - z (modifies (p!k)))) ) div b^t mod b" using o'_div by (auto simp: algebra_simps div_mult2_eq) (smt Nat.add_0_right add_mult_distrib2 b_gt1 div_mult_self2 gr_implies_not0) (* Interchanged order in the following with add.commute to ease next step *) also have "... = nth_digit (\S- p j (\k. s k && z (l k)) + \S+ p j (\k. s k) + \S0 p j (\k. s k && (e - z (l k)))) t b" by (auto simp: nth_digit_def l_def add.commute) (* STEP 2: Commute nth_digit to the inside of all expressions. *) also have "... = nth_digit (\S- p j (\k. s k && z (l k))) t b + nth_digit (\S+ p j (\k. s k) + \S0 p j (\k. s k && (e - z (l k)))) t b" using block_additivity sum_ssub_fst_digit_zero sum_state_special_fst_digit_zero by (auto simp: l_def c_gt0 b_def add.assoc) also have "... = nth_digit (\S+ p j (\k. s k)) t b + nth_digit (\S- p j (\k. s k && z (l k))) t b + nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b" using block_additivity sum_sadd_fst_digit_zero sum_szero_fst_digit_zero by (auto simp: l_def c_gt0 b_def) also have "... = nth_digit (\S+ p j (\k. s k)) t b + (\S- p j (\k. nth_digit (s k && z (l k)) t b)) + nth_digit (\S0 p j (\k. s k && (e - z (l k)))) t b" using sum_ssub_nth_digit_commute by auto also have "... = nth_digit (\S+ p j (\k. s k)) t b + \S- p j (\k. nth_digit (s k && z (l k)) t b) + \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" using l_def l skzl_bound2 sum_szero_nth_digit_commute by (auto) also have "... = \S+ p j (\k. nth_digit (s k) t b) + \S- p j (\k. nth_digit (s k && z (l k)) t b) + \S0 p j (\k. nth_digit (s k && (e - z (l k))) t b)" using sum_sadd_nth_digit_commute by auto also have "... = \S+ p j (\k. nth_digit (s k) t b) + \S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b) + \S0 p j (\k. (nth_digit (e - z (l k)) t b) && nth_digit (s k) t b)" using bitAND_nth_digit_commute b_def by (auto simp: bitAND_commutes) also have "... = (\S+ p j (\k. nth_digit (s k) t b)) + (\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) + (\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b))" using sum_szero_aux assms sum_ssub_zero.simps m_def \t apply (auto) using sum.cong atLeastAtMost_iff by smt (* this ATP proof ensures that the \S0 only uses the sum_szero_aux lemma with k \ m because the summation of sum_ssub_zero ranges from k = 0 .. (length p - 1) *) ultimately have "s j div (b ^ Suc t) mod b = (\S+ p j (\k. nth_digit (s k) t b)) + (\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) + (\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b))" by auto moreover have "(\S- p j (\k. nth_digit (z (l k)) t b && nth_digit (s k) t b)) = (\S- p j (\k. Zeq (l k) t * Seq k t))" using skzl_bitAND_to_mult sum_ssub_nzero.simps l by (smt atLeastAtMost_iff sum.cong) moreover have "(\S0 p j (\k. (1 - nth_digit (z (l k)) t b) && nth_digit (s k) t b)) = (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" using skzl_bitAND_to_mult2 sum_ssub_zero.simps l by (auto) (smt atLeastAtMost_iff sum.cong) ultimately have "nth_digit (s j) (Suc t) b = (\S+ p j (\k. Seq k t)) + (\S- p j (\k. Zeq (l k) t * Seq k t)) + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" using Seq_def nth_digit_def by auto thus ?thesis by auto qed lemma aux_nocarry_sk: assumes "t\q" shows "i \ j \ i\m \ j\m \ nth_digit (s i) t b * nth_digit (s j) t b = 0" proof (cases "t=q") case True have "j < m \ Seq j q = 0" for j using s_bound Seq_def nth_digit_def by auto then show ?thesis using True Seq_def apply auto by (metis le_less less_nat_zero_code) next case False hence "k\m \ nth_digit (s k) t b = 1 \ (\j\m. j \ k \ nth_digit (s j) t b = 0)" for k using states_unique_RAW[of "t"] Seq_def assms by auto thus ?thesis by (auto) (metis One_nat_def le_neq_implies_less m_def not_less_eq sk_bound) qed lemma nocarry_sk: assumes "i \ j" and "i\m" and "j\m" shows "(s i) \ k * (s j) \ k = 0" proof - have reduct: "(s i) \ k = nth_digit (s i) (k div Suc c) b \ (k mod Suc c)" for i using digit_gen_pow2_reduct[of "k mod Suc c" "Suc c" "s i" "k div Suc c"] b_def using mod_less_divisor zero_less_Suc by presburger have "k div Suc c \ q \ nth_digit (s i) (k div Suc c) b * nth_digit (s j) (k div Suc c) b = 0" using aux_nocarry_sk assms by auto moreover have "k div Suc c > q \ nth_digit (s i) (k div Suc c) b * nth_digit (s j) (k div Suc c) b = 0" using nth_digit_def s_bound apply auto using b_gt1 div_greater_zero_iff leD le_less less_trans mod_less neq0_conv power_increasing_iff by (smt assms) ultimately have "nth_digit (s i) (k div Suc c) b \ (k mod Suc c) * nth_digit (s j) (k div Suc c) b \ (k mod Suc c) = 0" using nth_bit_def by auto thus ?thesis using reduct[of "i"] reduct[of "j"] by auto qed lemma commute_sum_rsub_bitAND: "\R- p l (\k. s k && z l) = \R- p l (\k. s k) && z l" proof - show ?thesis apply (auto simp: sum_rsub.simps) using m_def nocarry_sk aux_commute_bitAND_sum_if[of "m" "s" "\k. issub (p ! k) \ l = modifies (p ! k)" "z l"] by (auto simp add: atMost_atLeast0) qed lemma sum_rsub_bound: "l \R- p l (\k. s k && z l) \ r l + \R+ p l s" proof - assume "lR- p l (\k. s k) && z l \ z l" by (auto simp: lm0245 masks_leq) also have "... \ r l" using zl_le_rl `lObtaining single step register relations from multiple step register relations \ lemma mult_to_single_reg: "c>0 \ l Req l (Suc t) = Req l t + (\R+ p l (\k. Seq k t)) - (\R- p l (\k. (Zeq l t) * Seq k t))" for l t proof - assume l: "l 0" have a_div: "a div b = 0" using c_eq rm_constants_def B_def by auto have subtract_bound: "\t'. nth_digit (\R- p l (\k. s k && z l)) t' b \ nth_digit (r l + \R+ p l (\k. s k)) t' b" proof - { fix t' have "nth_digit (z l) t' b \ nth_digit (r l) t' b" using Zeq_def Req_def Z l by auto then have "nth_digit (\R- p l (\k. s k)) t' b && nth_digit (z l) t' b \ nth_digit (r l) t' b" using sum_rsub_special_block_bound by (meson dual_order.trans lm0245 masks_leq) then have "nth_digit (\R- p l (\k. s k && z l)) t' b \ nth_digit (r l) t' b" using commute_sum_rsub_bitAND bitAND_nth_digit_commute b_def by auto then have "nth_digit (\R- p l (\k. s k && z l)) t' b \ nth_digit (r l) t' b + nth_digit (\R+ p l (\k. s k)) t' b" by auto then have "nth_digit (\R- p l (\k. s k && z l)) t' b \ nth_digit (r l + \R+ p l (\k. s k)) t' b" using block_additivity rl_fst_digit_zero sum_radd_fst_digit_zero by (auto simp: b_def c l) } then show ?thesis by auto qed define a' where "a' \ (if l = 0 then a else 0)" have a'_div: "a' div b = 0" using a_div a'_def by auto (* START PROOF CHAIN *) (* STEP 1: Remove a' using divisibility properties. *) have "r l div (b * b ^ t) mod b = (a' + b*r l + b*\R+ p l (\k. s k) - b*\R- p l (\k. s k && z l)) div (b * b ^ t) mod b" using r_eq reg_equations_def by (auto simp: a'_def l) also have "... = (a' + b * (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l))) div b div b^t mod b" by (auto simp: algebra_simps div_mult2_eq) (metis Nat.add_diff_assoc add_mult_distrib2 mult_le_mono2 sum_rsub_bound l) also have "... = ((r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) + a' div b) div b ^ t mod b" using b_gt1 by auto also have "... = (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) div b ^ t mod b" using a'_div by auto also have "... = nth_digit (r l + \R+ p l (\k. s k) - \R- p l (\k. s k && z l)) t b" using nth_digit_def by auto (* STEP 2: Commute nth_digit to the inside of all expressions. *) also have "... = nth_digit (r l + \R+ p l (\k. s k)) t b - nth_digit (\R- p l (\k. s k && z l)) t b" using block_subtractivity subtract_bound by (auto simp: c b_def) also have "... = nth_digit (r l) t b + nth_digit (\R+ p l (\k. s k)) t b - nth_digit (\R- p l (\k. s k && z l)) t b" using block_additivity rl_fst_digit_zero sum_radd_fst_digit_zero by (auto simp: l b_def c) also have "... = nth_digit (r l) t b + \R+ p l (\k. nth_digit (s k) t b) - \R- p l (\k. nth_digit (s k && z l) t b)" using sum_radd_nth_digit_commute using sum_rsub_nth_digit_commute by auto ultimately have "r l div (b * b ^ t) mod b = (nth_digit (r l) t b) + \R+ p l (\k. nth_digit (s k) t b) - \R- p l (\k. nth_digit (z l) t b && nth_digit (s k) t b)" using bitAND_nth_digit_commute b_def by (simp add: bitAND_commutes) then show ?thesis using Req_def Seq_def nth_digit_def skzl_bitAND_to_mult l by (auto simp: sum_rsub.simps) (smt atLeastAtMost_iff sum.cong) qed text \Obtaining single step state relations from multiple step state relations\ lemma mult_to_single_state: fixes t j :: nat defines "l \ \k. modifies (p!k)" shows "j\m \ t Seq j (Suc t) = (\S+ p j (\k. Seq k t)) + (\S- p j (\k. Zeq (l k) t * Seq k t)) + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" proof - assume "j \ m" assume "t < q" have "nth_digit (s j) (Suc t) b = (\S+ p j (\k. Seq k t)) + (\S- p j (\k. Zeq (l k) t * Seq k t)) + (\S0 p j (\k. (1 - Zeq (l k) t) * Seq k t))" using state_equations_digit_commute \j\m\ \t l_def by auto then show ?thesis using nth_digit_def l_def Seq_def by auto qed text \Conclusion: The central equivalence showing that the cell entries obtained from r s z indeed coincide with the correct cell values when executing the register machine. This statement is proven by induction using the single step relations for Req and Seq as well as the statement for Zeq. \ lemma rzs_eq: "l j\m \ t\q \ R ic p l t = Req l t \ Z ic p l t = Zeq l t \ S ic p j t = Seq j t" proof (induction t arbitrary: j l) have "m>0" using m_def is_val is_valid_initial_def[of "ic" "p"] is_valid_def[of "ic" "p"] by auto case 0 (* STATES *) have mod_aux0: "Suc (b * k) mod b = 1" for k using euclidean_semiring_cancel_class.mod_mult_self2[of "1" "b" "k"] b_gt1 by auto have step_state0: "s 0 = 1 + b*\S+ p 0 (\k. s k) + b*\S- p 0 (\k. s k && z (modifies (p!k))) + b*\S0 p 0 (\k. s k && (e - z (modifies (p!k))))" using s_eq state_equations_def by auto hence "Seq 0 0 = 1" using Seq_def by (auto simp: nth_digit_def mod_aux0) hence S00: "Seq 0 0 = S ic p 0 0" using S_def is_val is_valid_initial_def[of "ic"] by auto have "s m = b^q" using s_eq state_equations_def by auto hence "Seq m 0 = 0" using Seq_def nth_digit_def c_eq rm_constants_def by auto hence Sm0: "S ic p m 0 = Seq m 0" using is_val is_valid_initial_def[of "ic" "p" "a"] S_def `m>0` by auto have step_states: "\d>0. d s d = b*\S+ p d (\k. s k) + b*\S- p d (\k. s k && z (modifies (p!k))) + b*\S0 p d (\k. s k && (e - z (modifies (p!k))))" using s_eq state_equations_def by auto hence "\k>0. k Seq k 0 = 0" using Seq_def by (auto simp: nth_digit_def) hence "k>0 \ k Seq k 0 = S ic p k 0" for k using S_def is_val is_valid_initial_def[of "ic"] by auto with S00 Sm0 have Sid: "k\m \ Seq k 0 = S ic p k 0" for k by (cases "k=0"; cases "k=m"; auto) (* REGISTERS *) have "b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0)) = b * (r 0 + \R+ p 0 s) - b * \R- p 0 (\k. s k && z 0)" using Nat.diff_mult_distrib2[of "b" "r 0 + \R+ p 0 s" "\R- p 0 (\k. s k && z 0)"] by auto also have "... = b * r 0 + b* \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)" using Nat.add_mult_distrib2[of "b" "r 0" "\R+ p 0 s"] by auto ultimately have distrib: "a + b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0)) = a + b * r 0 + b * \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)" by (auto simp: algebra_simps) (metis Nat.add_diff_assoc add_mult_distrib2 mult_le_mono2 n_gt0 sum_rsub_bound) hence "Req 0 0 = (a + b * r 0 + b * \R+ p 0 s - b * \R- p 0 (\k. s k && z 0)) mod b" using Req_def nth_digit_def r_eq reg_equations_def by auto also have "... = (a + b * (r 0 + \R+ p 0 s - \R- p 0 (\k. s k && z 0))) mod b" using distrib by auto finally have "Req 0 0 = a" using c_eq rm_constants_def B_def by auto hence R00: "R ic p 0 0 = Req 0 0" using R_def is_val is_valid_initial_def[of "ic" "p" "a"] by auto have rl_transform: "l>0 \ r l = b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)" using reg_equations_def r_eq `l0 \ (b * r l + b * \R+ p l s - b * \R- p l (\k. s k && z l)) mod b = 0" using Req_def nth_digit_def reg_equations_def r_eq by auto hence "l>0 \ Req l 0 = 0" using Req_def rl_transform nth_digit_def by auto hence "l>0 \ Req l 0 = R ic p l 0" using is_val is_valid_initial_def[of "ic"] R_def by auto hence Rid: "R ic p l 0 = Req l 0" using R00 by (cases "l=0"; auto) (* ZERO INDICATORS *) hence Zid: "Z ic p l 0 = Zeq l 0" using Z Z_def 0 by auto show ?case using Sid Rid Zid `lm` by auto next case (Suc t) (* INDUCTIVE HYPOTHESIS *) have "Suc t\q" using Suc by auto then have "tm \ S ic p k t = Seq k t" for k using Suc m_def by auto have Z_IH: "\l::nat. l Z ic p l t = Zeq l t" using Suc by auto (* REGISTERS *) from S_IH have S1: "k\m \ (if isadd (p ! k) \ l = modifies (p ! k) then Seq k t else 0) = (if isadd (p ! k) \ l = modifies (p ! k) then S ic p k t else 0)" for k by auto have S2: "k \ {0..length p-1} \ (if issub (p ! k) \ l = modifies (p ! k) then Zeq l t * Seq k t else 0) = (if issub (p ! k) \ l = modifies (p ! k) then Zeq l t * S ic p k t else 0)" for k using Suc m_def by auto have "Req l (Suc t) = Req l t + (\R+ p l (\k. Seq k t)) - (\R- p l (\k. (Zeq l t) * Seq k t))" using mult_to_single_reg[of "l"] `lR+ p l (\k. S ic p k t)) - (\R- p l (\k. (Z ic p l t) * S ic p k t))" using Suc sum_radd.simps sum_rsub.simps S1 S2 m_def by auto finally have R: "Req l (Suc t) = R ic p l (Suc t)" using is_val `l Suc m" by (simp add: m_def) (* STATES *) have "s m = b ^ q" using s_eq state_equations_def by auto hence "Seq m t = 0" using Seq_def \t nth_digit_def apply auto using b_gt1 bx_aux by auto hence "S ic p m t = 0" using Suc by auto hence "fst (steps ic p t) \ m" using S_def by auto hence "fst (steps ic p t) < m" using is_val m_def by (metis less_Suc_eq less_le_trans p_contains plength) hence nohalt: "\ ishalt (p ! fst (steps ic p t))" using is_valid_def[of "ic" "p"] is_valid_initial_def[of "ic" "p" "a"] m_def is_val by auto have "jm` m_def by (metis (full_types) diff_less is_val length_greater_0_conv less_imp_diff_less less_one list.size(3) nat_less_le not_less not_less_zero p_contains) have "Seq j (Suc t) = (\S+ p j (\k. Seq k t)) + (\S- p j (\k. Zeq (modifies (p!k)) t * Seq k t)) + (\S0 p j (\k. (1 - Zeq (modifies (p!k)) t) * Seq k t))" using mult_to_single_state `j\m` `tS+ p j (\k. Seq k t)) + (\S- p j (\k. Z ic p (modifies (p!k)) t * Seq k t)) + (\S0 p j (\k. (1 - Z ic p (modifies (p!k)) t) * Seq k t))" using Z_IH modifies_valid sum_ssub_zero.simps sum_ssub_nzero.simps by (auto simp: m_def, smt atLeastAtMost_iff sum.cong) also have "... = (\S+ p j (\k. S ic p k t)) + (\S- p j (\k. Z ic p (modifies (p!k)) t * S ic p k t)) + (\S0 p j (\k. (1 - Z ic p (modifies (p!k)) t) * S ic p k t))" using S_IH sum_ssub_zero.simps sum_ssub_nzero.simps sum_sadd.simps by (auto simp: m_def, smt atLeastAtMost_iff sum.cong) finally have S: "Seq j (Suc t) = S ic p j (Suc t)" using is_val lm04_07_one_step_relation_state[of "ic" "p" "a" "j"] `jResolvable Designs\ text \Resolvable designs have further structure, and can be "resolved" into a set of resolution classes. A resolution class is a subset of blocks which exactly partitions the point set. Definitions based off the handbook \<^cite>\"colbournHandbookCombinatorialDesigns2007"\ and Stinson \<^cite>\"stinsonCombinatorialDesignsConstructions2004"\. This theory includes a proof of an alternate statement of Bose's theorem\ theory Resolvable_Designs imports BIBD begin subsection \Resolutions and Resolution Classes\ text \A resolution class is a partition of the point set using a set of blocks from the design A resolution is a group of resolution classes partitioning the block collection\ context incidence_system begin definition resolution_class :: "'a set set \ bool" where "resolution_class S \ partition_on \ S \ (\ bl \ S . bl \# \)" lemma resolution_classI [intro]: "partition_on \ S \ (\ bl . bl \ S \ bl \# \) \ resolution_class S" by (simp add: resolution_class_def) lemma resolution_classD1: "resolution_class S \ partition_on \ S" by (simp add: resolution_class_def) lemma resolution_classD2: "resolution_class S \ bl \ S \ bl \# \" by (simp add: resolution_class_def) lemma resolution_class_empty_iff: "resolution_class {} \ \ = {}" by (auto simp add: resolution_class_def partition_on_def) lemma resolution_class_complete: "\ \ {} \ \ \# \ \ resolution_class {\}" by (auto simp add: resolution_class_def partition_on_space) lemma resolution_class_union: "resolution_class S \ \S = \ " by (simp add: resolution_class_def partition_on_def) lemma (in finite_incidence_system) resolution_class_finite: "resolution_class S \ finite S" using finite_elements finite_sets by (auto simp add: resolution_class_def) lemma (in design) resolution_class_sum_card: "resolution_class S \ (\bl \ S . card bl) = \" using resolution_class_union finite_blocks by (auto simp add: resolution_class_def partition_on_def card_Union_disjoint) definition resolution:: "'a set multiset multiset \ bool" where "resolution P \ partition_on_mset \ P \ (\ S \# P . distinct_mset S \ resolution_class (set_mset S))" lemma resolutionI : "partition_on_mset \ P \ (\ S . S \#P \ distinct_mset S) \ (\ S . S\# P \ resolution_class (set_mset S)) \ resolution P" by (simp add: resolution_def) lemma (in proper_design) resolution_blocks: "distinct_mset \ \ disjoint (set_mset \) \ \(set_mset \) = \ \ resolution {#\#}" unfolding resolution_def resolution_class_def partition_on_mset_def partition_on_def using design_blocks_nempty blocks_nempty by auto end subsection \Resolvable Design Locale\ text \A resolvable design is one with a resolution P\ locale resolvable_design = design + fixes partition :: "'a set multiset multiset" ("\

") assumes resolvable: "resolution \

" begin lemma resolutionD1: "partition_on_mset \ \

" using resolvable by (simp add: resolution_def) lemma resolutionD2: "S \#\

\ distinct_mset S" using resolvable by (simp add: resolution_def) lemma resolutionD3: " S\# \

\ resolution_class (set_mset S)" using resolvable by (simp add: resolution_def) lemma resolution_class_blocks_disjoint: "S \# \

\ disjoint (set_mset S)" using resolutionD3 by (simp add: partition_on_def resolution_class_def) lemma resolution_not_empty: "\ \ {#} \ \

\ {#}" using partition_on_mset_not_empty resolutionD1 by auto lemma resolution_blocks_subset: "S \# \

\ S \# \" using partition_on_mset_subsets resolutionD1 by auto end lemma (in incidence_system) resolvable_designI [intro]: "resolution \

\ design \ \ \ resolvable_design \ \ \

" by (simp add: resolvable_design.intro resolvable_design_axioms.intro) subsection \Resolvable Block Designs\ text \An RBIBD is a resolvable BIBD - a common subclass of interest for block designs\ locale r_block_design = resolvable_design + block_design begin lemma resolution_class_blocks_constant_size: "S \# \

\ bl \# S \ card bl = \" by (metis resolutionD3 resolution_classD2 uniform_alt_def_all) lemma resolution_class_size1: assumes "S \# \

" shows "\ = \ * size S" proof - have "(\bl \# S . card bl) = (\bl \ (set_mset S) . card bl)" using resolutionD2 assms by (simp add: sum_unfold_sum_mset) then have eqv: "(\bl \# S . card bl) = \" using resolutionD3 assms resolution_class_sum_card by presburger have "(\bl \# S . card bl) = (\bl \# S . \)" using resolution_class_blocks_constant_size assms by auto thus ?thesis using eqv by auto qed lemma resolution_class_size2: assumes "S \# \

" shows "size S = \ div \" using resolution_class_size1 assms by (metis nonzero_mult_div_cancel_left not_one_le_zero k_non_zero) lemma resolvable_necessary_cond_v: "\ dvd \" proof - obtain S where s_in: "S \#\

" using resolution_not_empty design_blocks_nempty by blast then have "\ * size S = \" using resolution_class_size1 by simp thus ?thesis by (metis dvd_triv_left) qed end locale rbibd = r_block_design + bibd begin lemma resolvable_design_num_res_classes: "size \

= \" proof - have k_ne0: "\ \ 0" using k_non_zero by auto have f1: "\ = (\S \# \

. size S)" by (metis partition_on_msetD1 resolutionD1 size_big_union_sum) then have "\ = (\S \# \

. \ div \)" using resolution_class_size2 f1 by auto then have f2: "\ = (size \

) * (\ div \)" by simp then have "size \

= \ div (\ div \)" using b_non_zero by auto then have "size \

= (\ * \) div \" using f2 resolvable_necessary_cond_v by (metis div_div_div_same div_dvd_div dvd_triv_right k_ne0 nonzero_mult_div_cancel_right) thus ?thesis using necessary_condition_two by (metis nonzero_mult_div_cancel_left not_one_less_zero t_design_min_v) qed lemma resolvable_necessary_cond_b: "\ dvd \" proof - have f1: "\ = (\S \# \

. size S)" by (metis partition_on_msetD1 resolutionD1 size_big_union_sum) then have "\ = (\S \# \

. \ div \)" using resolution_class_size2 f1 by auto thus ?thesis using resolvable_design_num_res_classes by simp qed subsubsection \Bose's Inequality\ text \Boses inequality is an important theorem on RBIBD's. This is a proof of an alternate statement of the thm, which does not require a linear algebraic approach, taken directly from Stinson \<^cite>\"stinsonCombinatorialDesignsConstructions2004"\\ theorem bose_inequality_alternate: "\ \ \ + \ - 1 \ \ \ \ + \" proof - + from necessary_condition_two v_non_zero have r: \\ = \ * \ div \\ + by (metis div_mult_self1_is_m) define k b v l r where intdefs: "k \ (int \)" "b \ int \" "v = int \" "l \ int \" "r \ int \" have kdvd: "k dvd (v * (r - k))" using intdefs by (simp add: resolvable_necessary_cond_v) have necess1_alt: "l * v - l = r * (k - 1)" using necessary_condition_one intdefs by (smt (verit) diff_diff_cancel int_ops(2) int_ops(6) k_non_zero nat_mult_1_right of_nat_0_less_iff of_nat_mult right_diff_distrib' v_non_zero) then have v_eq: "v = (r * (k - 1) + l) div l" using necessary_condition_one index_not_zero intdefs by (metis diff_add_cancel nonzero_mult_div_cancel_left not_one_le_zero of_nat_mult unique_euclidean_semiring_with_nat_class.of_nat_div) have ldvd: " \ x. l dvd (x * (r * (k - 1) + l))" by (metis necess1_alt diff_add_cancel dvd_mult dvd_triv_left) have "(b \ v + r - 1) \ ((\ * r) div k \ v + r - 1)" using necessary_condition_two k_non_zero intdefs by (metis (no_types, lifting) nonzero_mult_div_cancel_right not_one_le_zero of_nat_eq_0_iff of_nat_mult) also have "... \ (((v * r) - (v * k)) div k \ r - 1)" - using k_non_zero div_mult_self3 k_non_zero necessary_condition_two intdefs - by (smt (verit, ccfv_SIG) Euclidean_Division.div_eq_0_iff b_non_zero bibd_block_number mult_is_0 of_nat_eq_0_iff) + using k_non_zero k_non_zero r intdefs + by (simp add: of_nat_div algebra_simps) + (smt (verit, ccfv_threshold) One_nat_def div_mult_self4 of_nat_1 of_nat_mono) also have f2: " ... \ ((v * ( r - k)) div k \ ( r - 1))" using int_distrib(3) by (simp add: mult.commute) also have f2: " ... \ ((v * ( r - k)) \ k * ( r - 1))" using k_non_zero kdvd intdefs by auto also have "... \ ((((r * (k - 1) + l ) div l) * (r - k)) \ k * (r - 1))" using v_eq by presburger also have "... \ ( (r - k) * ((r * (k - 1) + l ) div l) \ (k * (r - 1)))" by (simp add: mult.commute) also have " ... \ ( ((r - k) * (r * (k - 1) + l )) div l \ (k * (r - 1)))" using div_mult_swap necessary_condition_one intdefs by (metis diff_add_cancel dvd_triv_left necess1_alt) also have " ... \ (((r - k) * (r * (k - 1) + l )) \ l * (k * (r - 1)))" using ldvd[of "(r - k)"] dvd_mult_div_cancel index_not_zero mult_strict_left_mono intdefs by (smt (verit) b_non_zero bibd_block_number bot_nat_0.extremum_strict div_0 less_eq_nat.simps(1) -mult_eq_0_iff mult_left_le_imp_le mult_left_mono of_nat_0 of_nat_le_0_iff of_nat_le_iff of_nat_less_iff) + mult_eq_0_iff mult_left_le_imp_le mult_left_mono of_nat_0 of_nat_le_0_iff of_nat_le_iff of_nat_less_iff) also have 1: "... \ (((r - k) * (r * (k - 1))) + ((r - k) * l ) \ l * (k * (r - 1)))" by (simp add: distrib_left) also have "... \ (((r - k) * r * (k - 1)) \ l * k * (r - 1) - ((r - k) * l ))" using mult.assoc by linarith also have "... \ (((r - k) * r * (k - 1)) \ (l * k * r) - (l * k) - ((r * l) -(k * l )))" using distrib_right by (simp add: distrib_left right_diff_distrib' left_diff_distrib') also have "... \ (((r - k) * r * (k - 1)) \ (l * k * r) - ( l * r))" by (simp add: mult.commute) also have "... \ (((r - k) * r * (k - 1)) \ (l * (k * r)) - ( l * r))" by linarith also have "... \ (((r - k) * r * (k - 1)) \ (l * (r * k)) - ( l * r))" by (simp add: mult.commute) also have "... \ (((r - k) * r * (k - 1)) \ l * r * (k - 1))" by (simp add: mult.assoc int_distrib(4)) finally have "(b \ v + r - 1) \ (r \ k + l)" using index_lt_replication mult_right_le_imp_le r_gzero mult_cancel_right k_non_zero intdefs by (smt (z3) of_nat_0_less_iff of_nat_1 of_nat_le_iff of_nat_less_iff) then have "\ \ \ + \ - 1 \ \ \ \ + \" using k_non_zero le_add_diff_inverse of_nat_1 of_nat_le_iff intdefs by linarith thus ?thesis by simp qed end end \ No newline at end of file diff --git a/thys/Dict_Construction/Test_Lazy_Case.thy b/thys/Dict_Construction/Test_Lazy_Case.thy --- a/thys/Dict_Construction/Test_Lazy_Case.thy +++ b/thys/Dict_Construction/Test_Lazy_Case.thy @@ -1,47 +1,47 @@ subsection \Interaction with \Lazy_Case\\ theory Test_Lazy_Case imports Dict_Construction Lazy_Case.Lazy_Case Show.Show_Instances begin datatype 'a tree = Node | Fork 'a "'a tree list" lemma map_tree[code]: "map_tree f t = (case t of Node \ Node | Fork x ts \ Fork (f x) (map (map_tree f) ts))" by (induction t) auto experiment begin (* FIXME proper qualified path *) text \ Dictionary construction of @{const map_tree} requires the [@{attribute fundef_cong}] rule of @{const Test_Lazy_Case.tree.case_lazy}. \ declassify valid: map_tree thm valid lemma "Test__Lazy__Case_tree_map__tree = map_tree" by (fact valid) end definition i :: "(unit \ (bool list \ string \ nat option) list) option \ string" where "i = show" experiment begin -text \This currently requires @{theory Lazy_Case.Lazy_Case} because of @{const Euclidean_Division.divmod_nat}.\ +text \This currently requires @{theory Lazy_Case.Lazy_Case} because of @{const Euclidean_Rings.divmod_nat}.\ (* FIXME get rid of Lazy_Case dependency *) declassify valid: i thm valid lemma "Test__Lazy__Case_i = i" by (fact valid) end end \ No newline at end of file diff --git a/thys/Digit_Expansions/Bits_Digits.thy b/thys/Digit_Expansions/Bits_Digits.thy --- a/thys/Digit_Expansions/Bits_Digits.thy +++ b/thys/Digit_Expansions/Bits_Digits.thy @@ -1,719 +1,719 @@ section \Digit functions\ theory Bits_Digits imports Main begin text \We define the n-th bit of a number in base 2 representation \ definition nth_bit :: "nat \ nat \ nat" (infix "\" 100) where "nth_bit num k = (num div (2 ^ k)) mod 2" text \as well as the n-th digit of a number in an arbitrary base \ definition nth_digit :: "nat \ nat \ nat \ nat" where "nth_digit num k base = (num div (base ^ k)) mod base" text \In base 2, the two definitions coincide. \ lemma nth_digit_base2_equiv:"nth_bit a k = nth_digit a k (2::nat)" by (auto simp add:nth_bit_def nth_digit_def) lemma general_digit_base: assumes "t1 > t2" and "b>1" shows "nth_digit (a * b^t1) t2 b = 0" proof - have 1: "b^t1 div b^t2 = b^(t1-t2)" using assms apply auto by (metis Suc_lessD less_imp_le less_numeral_extra(3) power_diff) have "b^t2 dvd b^t1" using `t1 > t2` by (simp add: le_imp_power_dvd) hence "(a * b^t1) div b^t2 = a * b^(t1-t2)" using div_mult_swap[of "b^t2" "b^t1" "a"] 1 by auto thus ?thesis using nth_digit_def assms by auto qed lemma nth_bit_bounded: "nth_bit a k \ 1" by (auto simp add: nth_bit_def) lemma nth_digit_bounded: "b>1 \ nth_digit a k b \ b-1" apply (auto simp add: nth_digit_def) using less_Suc_eq_le by fastforce lemma obtain_smallest: "P (n::nat) \ \k\n. P k \ (\a(P a))" by (metis ex_least_nat_le not_less_zero zero_le) subsection \Simple properties and equivalences\ text \Reduce the @{term nth_digit} function to @{term nth_bit} if the base is a power of 2\ -lemma digit_gen_pow2_reduct: "k (nth_digit a t (2^c)) \ k = a \ (c*t+k)" +lemma digit_gen_pow2_reduct: + \(nth_digit a t (2 ^ c)) \ k = a \ (c * t + k)\ if \k < c\ proof - - assume "k k = x \ k" for x + have moddiv: "(x mod 2 ^ c) \ k = x \ k" for x proof- - assume klc: "kn = c - k\ + with \k < c\ have c_nk: \c = n + k\ + by simp + obtain a b where x_def: "x = a * 2 ^ c + b" and "b < 2 ^ c" by (meson mod_div_decomp mod_less_divisor zero_less_numeral zero_less_power) then have bk: "(x mod 2 ^ c) \ k = b \ k" by simp - have "x div 2 ^ k = a*2^(c-k) + b div 2^k" - using x_def by (smt Euclidean_Division.div_eq_0_iff add_diff_inverse_nat add_self_div_2 - div_mult_self3 klc mult.left_commute nat_le_linear not_less power_add power_eq_0_iff - semiring_normalization_rules(7)) + from \b < 2 ^ c\ have \x div 2 ^ k = a * 2 ^ (c - k) + b div 2 ^ k\ + by (simp add: x_def c_nk power_add flip: mult.commute [of \2 ^ k\] mult.left_commute [of \2 ^ k\]) then have "x \ k = (a*2^(c-k) + b div 2^k) mod 2" by (simp add: nth_bit_def) - then have "x \ k = b \ k" using nth_bit_def klc by (simp add: mod2_eq_if) + then have "x \ k = b \ k" using nth_bit_def \k < c\ by (simp add: mod2_eq_if) then show ?thesis using nth_bit_def bk by linarith qed have "a div ((2 ^ c) ^ t * 2 ^ k) = a div (2 ^ c) ^ t div 2 ^ k" using div_mult2_eq by blast moreover have "a div (2 ^ c) ^ t mod 2 ^ c div 2 ^ k mod 2 = a div (2 ^ c) ^ t div 2 ^ k mod 2" using moddiv nth_bit_def by auto ultimately show "(nth_digit a t (2^c)) \ k = a \ (c*t+k)" using nth_digit_def nth_bit_def by (auto simp: power_add power_mult) qed text \Show equivalence of numbers by equivalence of all their bits (digits)\ lemma aux_even_pow2_factor: "a > 0 \ \k b. ((a::nat) = (2^k) * b \ odd b)" proof(induction a rule: full_nat_induct) case (1 n) then show ?case proof (cases "odd n") case True then show ?thesis by (metis nat_power_eq_Suc_0_iff power_Suc power_Suc0_right power_commutes) next case False have "(\t. n = 2 * t)" using False by auto then obtain t where n_def:"n = 2 * t" .. then have "t < n" using "1.prems" by linarith then have ih:"\r s. t = 2^r * s \ odd s" using 1 n_def by simp then have "\r s. n = 2^(Suc r) * s" using n_def by auto then show ?thesis by (metis ih n_def power_commutes semiring_normalization_rules(18) semiring_normalization_rules(28)) qed qed lemma aux0_digit_wise_equiv:"a > 0 \ (\k. nth_bit a k = 1)" proof - assume a_geq_0: "a > 0" consider (odd) "a mod 2 = 1" | (even) "a mod 2 = 0" by force then show ?thesis proof(cases) case odd then show ?thesis by (metis div_by_1 nth_bit_def power_0) next case even then have bk_def:"\k b. (a = (2^k) * b \ odd b)" using aux_even_pow2_factor a_geq_0 by simp then obtain b k where bk_cond:"(a = (2^k) * b \ odd b)" by blast then have "b = a div (2^k)" by simp then have digi_b:"nth_bit b 0 = 1" using bk_def using bk_cond nth_bit_def odd_iff_mod_2_eq_one by fastforce then have "nth_bit a k = (nth_bit b 0)" using nth_bit_def bk_cond by force then have "nth_bit a k = 1" using digi_b by simp then show ?thesis by blast qed qed lemma aux1_digit_wise_equiv:"(\k.(nth_bit a k = 0)) \ a = 0" (is "?P \ ?Q") proof assume "?Q" then show "?P" by (simp add: nth_bit_def) next { assume a_neq_0:"\?Q" then have "a > 0" by blast then have "\?P" using aux0_digit_wise_equiv by (metis zero_neq_one) } thus "?P \ ?Q" by blast qed lemma aux2_digit_wise_equiv: "(\r (a mod 2^k = 0)" proof(induct k) case 0 then show ?case by (auto simp add: nth_bit_def) next case (Suc k) then show ?case by (auto simp add: nth_bit_def) (metis dvd_imp_mod_0 dvd_mult_div_cancel dvd_refl even_iff_mod_2_eq_zero lessI minus_mod_eq_div_mult minus_mod_eq_mult_div mult_dvd_mono) qed lemma digit_wise_equiv: "(a = b) \ (\k. nth_bit a k = nth_bit b k)" (is "?P \ ?Q") proof assume "?P" then show "?Q" by simp next { assume notP: "\?P" have "\(\k. nth_bit a k = nth_bit b k)" if ab: "ak.(nth_bit c k = 1)" using nth_bit_def aux1_digit_wise_equiv by (metis c_def not_less0 not_mod_2_eq_1_eq_0 that zero_less_diff) then obtain k where k1:"nth_bit c k = 1" and k2:"\r < k. (nth_bit c r \ 1)" by (auto dest: obtain_smallest) then have cr0: "\r < k. (nth_bit c r = 0)" by (simp add: nth_bit_def) from aux2_digit_wise_equiv cr0 have "c mod 2^k = 0" by auto then have "a div 2 ^ k mod 2 \ b div 2 ^ k mod 2" by auto (metis b div_plus_div_distrib_dvd_right even_add k1 nth_bit_def odd_iff_mod_2_eq_one) then show ?thesis by (auto simp add: nth_bit_def) qed from this [of a b] this [of b a] notP have "\k. nth_bit a k = nth_bit b k \ a = b" using linorder_neqE_nat by auto } then show "?Q ==> ?P" by auto qed text \Represent natural numbers in their binary expansion\ lemma aux3_digit_sum_repr: assumes "b < 2^r" shows "(a*2^r + b) \ r = (a*2^r) \ r" by (auto simp add: nth_bit_def assms) lemma aux2_digit_sum_repr: assumes "n < 2^c" "r < c" shows "(a*2^c+n) \ r = n \ r" proof - have [simp]: \a*(2::nat) ^ c div 2 ^ r = a*2 ^ (c - r)\ using assms(2) by (auto simp: less_iff_Suc_add monoid_mult_class.power_add ac_simps) show ?thesis using assms by (auto simp add: nth_bit_def div_plus_div_distrib_dvd_left le_imp_power_dvd simp flip: mod_add_left_eq) qed lemma aux1_digit_sum_repr: assumes "n < 2^c" "rk k)*2^k)) \ r = n \ r" proof- define a where "a \ (\k=0.. k)*2^k))" define d where "d \ (\k=Suc(r).. k)*2^k))" define e where "e \ (\k=Suc r.. k)*2^(k-Suc(r))))" define b where "b \ (\k=0.. k)*2^k))" have ad: "(\k=0.. k)*2^k)) = a + d" using a_def d_def assms by (metis (no_types, lifting) Suc_leI a_def assms(2) d_def sum.atLeastLessThan_concat zero_le) have "d = (\k=Suc(r).. k * 2^(k - Suc r))))" using d_def apply (simp) apply (rule sum.cong; auto simp: algebra_simps) by (metis add_Suc le_add_diff_inverse numerals(2) power_Suc power_add) hence d2r: "d = 2^Suc(r)*e" using d_def e_def sum_distrib_left[of "2^(Suc r)" "\k. (n \ k) * 2^(k - Suc r)" "{Suc r..k k)*2^k)) = a +2^Suc(r)*e" by (simp add: d2r ad lessThan_atLeast0) moreover have "(\k=0.. k)*2^k)) < 2 ^ Suc(r)" using assms proof(induct r) case 0 then show ?case proof - have "n \ 0 < Suc 1" by (metis (no_types) atLeastLessThan_empty atLeastLessThan_iff lessI linorder_not_less nth_bit_bounded) then show ?thesis by simp qed next case (Suc r) have r2: "n \ Suc(r) * 2 ^ Suc(r) \ 2 ^ Suc(r)" using nth_bit_bounded by simp have "(\k=0.. k)*2^k)) < 2 ^ Suc(r)" using Suc.hyps using Suc.prems(2) Suc_lessD assms(1) by blast then show ?case using "r2" by (smt Suc_leI Suc_le_lessD add_Suc add_le_mono mult_2 power_Suc sum.atLeast0_lessThan_Suc) qed then have "a < 2 ^Suc(r)" using a_def by blast ultimately have ar:"(a+d) \ r = a \ r" using d2r by (metis (no_types, lifting) aux2_digit_sum_repr lessI semiring_normalization_rules(24) semiring_normalization_rules(7)) (* Second part of proof *) have ab: "a =(n \ r)*2^r + b" using a_def b_def d_def by simp have "(\k=0.. k)*2^k)) <2^r" proof(induct r) case 0 then show ?case by auto next case (Suc r) have r2: "n \ r * 2 ^ r \ 2 ^ r" using nth_bit_bounded by simp have "(\k = 0.. k * 2 ^ k) < 2^r" using Suc.hyps by auto then show ?case apply simp using "r2" by linarith qed then have b: "b < 2^r" using b_def by blast then have "a \ r = ((n \ r)*2^r) \ r" using ab aux3_digit_sum_repr by simp then have "a \ r = (n \ r)" using nth_bit_def by simp then show ?thesis using ar a_def by (simp add: ad lessThan_atLeast0) qed lemma digit_sum_repr: assumes "n < 2^c" shows "n = (\k < c.((n \ k) * 2^k))" proof - have "\k. (c\k \ n<2^k)" using assms less_le_trans by fastforce then have nik: "\k.( k\ c \ (n \ k = 0))" by (auto simp add: nth_bit_def) have "(\r r)*(2::nat)^r))<(2::nat)^c" proof (induct c) case 0 then show ?case by simp next case (Suc c) then show ?case using nth_bit_bounded add_mono_thms_linordered_field[of "(n \ c)* 2 ^ c" " 2 ^ c" "(\r r * 2 ^ r)" "2^c"] by simp qed then have "\k. k\ c \ (\r r)*2^r)) \ k = 0" using less_le_trans by (auto simp add: nth_bit_def) fastforce then have "\r\c.(n \ r = (\k k)*2^k)) \ r)" by (simp add: nik) moreover have "\r r = (\k k)*2^k)) \ r)" using aux1_digit_sum_repr assms by simp ultimately have "\r.(\k r = n \ r " by (metis not_less) then show ?thesis using digit_wise_equiv by presburger qed lemma digit_sum_repr_variant: "n =(\kn \ ((\k< n.((n \ k)*2^k)) = (\k< r.(n \ k)*2^k))" proof- have "\r.(2^r > r)" using less_exp by simp then have pow2: "\r.(r>n \ 2^r>n)" using less_trans by blast then have "\r.(r > n \ (n \ r = 0))" by (auto simp add: nth_bit_def) then have "\r. r > n \ (n \ r)*2^r = 0" by auto then show ?thesis using digit_sum_repr digit_sum_repr_variant pow2 by auto qed text \Digits are preserved under shifts\ lemma digit_shift_preserves_digits: assumes "b>1" shows "nth_digit (b * y) (Suc t) b = nth_digit y t b" using nth_digit_def assms by auto lemma digit_shift_inserts_zero_least_siginificant_digit: assumes "t>0" and "b>1" shows "nth_digit (1 + b * y) t b = nth_digit (b * y) t b" using nth_digit_def assms apply auto proof - assume "0 < t" assume "Suc 0 < b" hence "Suc (b * y) mod b = 1" by (simp add: Suc_times_mod_eq) hence "b * y div b = Suc (b * y) div b" using \b>1\ by (metis (no_types) One_nat_def diff_Suc_Suc gr_implies_not0 minus_mod_eq_div_mult mod_mult_self1_is_0 nonzero_mult_div_cancel_right) then show "Suc (b * y) div b ^ t mod b = b * y div b ^ t mod b" using \t>0\ by (metis Suc_pred div_mult2_eq power_Suc) qed text \Represent natural numbers in their base-b digitwise expansion\ lemma aux3_digit_gen_sum_repr: assumes "d < b^r" and "b > 1" shows "nth_digit (a*b^r + d) r b = nth_digit (a*b^r) r b" using \b>1\ by (auto simp: nth_digit_def assms) lemma aux2_digit_gen_sum_repr: assumes "n < b^c" "r < c" shows "nth_digit (a*b^c+n) r b = nth_digit n r b" proof - have [simp]: \a*b ^ c div b ^ r = a*b ^ (c - r)\ using assms(2) by (auto simp: less_iff_Suc_add monoid_mult_class.power_add ac_simps) show ?thesis using assms by (auto simp add: nth_digit_def div_plus_div_distrib_dvd_left le_imp_power_dvd simp flip: mod_add_left_eq) qed lemma aux1_digit_gen_sum_repr: assumes "n < b^c" "r1" shows "nth_digit (\k (\k=0.. (\k=Suc(r).. (\k=Suc r.. (\k=0..k=0..k=Suc(r)..k. (nth_digit n k b) * b^(k - Suc r)" "{Suc r..kk=0..b>1\ by auto then show ?thesis by simp qed next case (Suc r) have r2: "(nth_digit n (Suc r) b) * b ^ Suc(r) \ (b-1) * b ^ Suc(r)" using nth_digit_bounded[of "b" "n" "Suc r"] \b>1\ by auto moreover have "(\k=0..b>1\ by blast ultimately have "(nth_digit n (Suc r) b) * b ^ Suc(r) + (\k=0..k=0.. (b-1) * b ^ r" using nth_digit_bounded[of "b"] \b>1\ by auto have "(\k = 0..b>1\ by simp hence "nth_digit a r b = nth_digit n r b" using nth_digit_def \b>1\ by simp then show ?thesis using ar a_def by (simp add: ad lessThan_atLeast0) qed lemma aux_gen_b_factor: "a > 0 \ b>1 \ \k c. ((a::nat) = (b^k) * c \ \(c mod b = 0))" proof(induction a rule: full_nat_induct) case (1 n) show ?case proof(cases "n mod b = 0") case True then obtain t where n_def: "n = b * t" by blast hence "t < n" using 1 by auto with 1 have ih:"\r s. t = b^r * s \ \(s mod b = 0)" by (metis Suc_leI gr0I mult_0_right n_def) hence "\r s. n = b^(Suc r) * s" using n_def by auto then show ?thesis by (metis ih mult.commute mult.left_commute n_def power_Suc) next case False then show ?thesis by (metis mult.commute power_0 power_Suc power_Suc0_right) qed qed lemma aux0_digit_wise_gen_equiv: assumes "b>1" and a_geq_0: "a > 0" shows "(\k. nth_digit a k b \ 0)" proof(cases "a mod b = 0") case True hence "\k c. a = (b^k) * c \ \(c mod b = 0)" using aux_gen_b_factor a_geq_0 assms by simp then obtain c k where ck_cond:"a = (b^k) * c \ \(c mod b = 0)" by blast hence c_cond:"c = a div (b^k)" using a_geq_0 by auto hence digi_b:"nth_digit c 0 b \ 0" using ck_cond nth_digit_def by force hence "nth_digit a k b = nth_digit c 0 b" using nth_digit_def c_cond by simp hence "nth_digit a k b \ 0" using digi_b by simp then show ?thesis by blast next case False then show ?thesis by (metis div_by_1 nth_digit_def power.simps(1)) qed lemma aux1_digit_wise_gen_equiv: assumes "b>1" shows "(\k.(nth_digit a k b = 0)) \ a = 0" (is "?P \ ?Q") proof assume "?Q" then show "?P" by (simp add: nth_digit_def) next { assume a_neq_0:"\?Q" hence "a > 0" by blast hence "\?P" using aux0_digit_wise_gen_equiv \b>1\ by auto } from this show "?P \ ?Q" by blast qed lemma aux2_digit_wise_gen_equiv: "(\r (a mod b^k = 0)" proof(induct k) case 0 then show ?case by auto next case (Suc k) then show ?case apply(auto simp add: nth_digit_def) using dvd_imp_mod_0 dvd_mult_div_cancel dvd_refl lessI minus_mod_eq_div_mult minus_mod_eq_mult_div mult_dvd_mono by (metis mod_0_imp_dvd) qed text \Two numbers are the same if and only if their digits are the same\ lemma digit_wise_gen_equiv: assumes "b>1" shows "(x = y) \ (\k. nth_digit x k b = nth_digit y k b)" (is "?P \ ?Q") proof assume "?P" then show "?Q" by simp next{ assume notP: "\?P" have"\(\k. nth_digit x k b = nth_digit y k b)" if xy: "xk.(nth_digit c k b \ 0)" using nth_digit_def \b>1\ aux0_digit_wise_gen_equiv by (metis c_def that zero_less_diff) then obtain k where k1:"nth_digit c k b \ 0" and k2:"\r < k. (nth_digit c r b = 0)" apply(auto dest: obtain_smallest) done hence cr0: "\r < k. (nth_digit c r b = 0)" by (simp add: nth_digit_def) from aux2_digit_wise_gen_equiv cr0 have "c mod b^k = 0" by auto hence "x div b ^ k mod b \ y div b ^ k mod b" using y k1 \b>1\ aux1_digit_wise_gen_equiv[of "b" "c"] nth_digit_def apply auto proof - (* this ISAR proof was found by sledgehammer *) fix ka :: nat assume a1: "b ^ k dvd c" assume a2: "x div b ^ k mod b = (x + c) div b ^ k mod b" assume a3: "y = x + c" assume a4: "\num k base. nth_digit num k base = num div base ^ k mod base" have f5: "\n na. (na::nat) + n - na = n" by simp have f6: "x div b ^ k + c div b ^ k = y div b ^ k" using a3 a1 by (simp add: add.commute div_plus_div_distrib_dvd_left) have f7: "\n. (x div b ^ k + n) mod b = (y div b ^ k + n) mod b" using a3 a2 by (metis add.commute mod_add_right_eq) have "\n na nb. ((nb::nat) mod na + n - (nb + n) mod na) mod na = 0" by (metis (no_types) add.commute minus_mod_eq_mult_div mod_add_right_eq mod_mult_self1_is_0) then show "c div b ^ ka mod b = 0" using f7 f6 f5 a4 by (metis (no_types) k1) qed then show ?thesis by (auto simp add: nth_digit_def) qed from this [of x y] this [of y x] notP have "\k. nth_digit x k b = nth_digit y k b \ x = y" apply(auto) using linorder_neqE_nat by blast}then show "?Q ==> ?P" by auto qed text \A number is equal to the sum of its digits multiplied by powers of two\ lemma digit_gen_sum_repr: assumes "n < b^c" and "b>1" shows "n = (\k < c.((nth_digit n k b) * b^k))" proof - have 1: "(c\k \ nk \ (nth_digit n k b = 0)" for k by (auto simp add: nth_digit_def) have "(\rr (b-1) * b^c" using nth_digit_bounded \b>1\ by auto thus ?thesis using assms IH by (metis (no_types, lifting) bound add_mono_thms_linordered_field(1) add_mono_thms_linordered_field(5) le_less mult_eq_if not_one_le_zero) qed done hence "k\c \ nth_digit (\rr\c.(nth_digit n r b = nth_digit (\krkr. nth_digit (\kkb>1\ by auto qed lemma digit_gen_sum_repr_variant: assumes "b>1" shows "n = (\kb>1\ apply (induct n, auto) by (simp add: less_trans_Suc) then show ?thesis using digit_gen_sum_repr \b>1\ by auto qed lemma digit_gen_sum_index_variant: assumes "b>1" shows "r>n \ (\k< n.((nth_digit n k b )*b^k)) = (\k< r.(nth_digit n k b)*b^k)" proof - assume "r>n" have "b^r > r" for r using \b>1\ by (induction r, auto simp add: less_trans_Suc) hence powb: "\r.(r>n \ b^r>n)" using less_trans by auto hence "r > n \ (nth_digit n r b = 0)" for r by (auto simp add: nth_digit_def) hence "r > n \ (nth_digit n r b) * b^r = 0" for r by auto then show ?thesis using digit_gen_sum_repr digit_gen_sum_repr_variant powb \n < r\ assms by auto qed text \@{text nth_digit} extracts coefficients from a base-b digitwise expansion\ lemma nth_digit_gen_power_series: fixes c b k q defines "b \ 2^(Suc c)" assumes bound: "\k. (f k) < b" (* < 2^c makes proof easier, but is too strong for const f *) shows "nth_digit (\k=0..q. (f k) * b^k) t b = (if t\q then (f t) else 0)" proof (induction q arbitrary: t) case 0 have "b>1" using b_def using one_less_numeral_iff power_gt1 semiring_norm(76) by blast have "f 0 < b" using bound by auto hence "t>0 \ f 0 < b^t" using \b>1\ using bound less_imp_le_nat less_le_trans by (metis self_le_power) thus ?case using nth_digit_def bound by auto next case (Suc q) thus ?case proof (cases "t \ Suc q") case True have f_le_bound: "f k \ b-1" for k using bound apply auto by (metis Suc_pred b_def less_Suc_eq_le numeral_2_eq_2 zero_less_Suc zero_less_power) have series_bound: "(\k = 0..q. f k * b ^ k) < b^(Suc q)" apply (induct q) subgoal using bound by (simp add: less_imp_le_nat) subgoal for q proof - assume asm: "(\k = 0..q. f k * b ^ k) < b ^ Suc q" have "(\k = 0..q. f k * b ^ k) + f (Suc q) * (b * b ^ q) \ (\k = 0..q. f k * b ^ k) + (b-1) * (b * b ^ q)" using f_le_bound by auto also have "... < b^(Suc q) + (b-1) * (b * b ^ q)" using asm by auto also have "... \ b * b * b^q" apply auto by (metis One_nat_def Suc_neq_Zero b_def eq_imp_le mult.assoc mult_eq_if numerals(2) power_not_zero) finally show ?thesis using asm by auto qed done hence "nth_digit ((\k = 0..q. f k * b ^ k) + f (Suc q) * (b * b ^ q)) t b = f t" using Suc nth_digit_def apply (cases "t = Suc q", auto) subgoal using Suc_n_not_le_n add.commute add.left_neutral b_def bound div_mult_self1 less_imp_le_nat less_mult_imp_div_less mod_less not_one_le_zero one_less_numeral_iff one_less_power power_Suc semiring_norm(76) zero_less_Suc by auto subgoal proof - assume "t \ Suc q" hence "t < Suc q" using True by auto hence "nth_digit (f (Suc q) * b ^ Suc q + (\k = 0..q. f k * b ^ k)) t b = nth_digit (\k = 0..q. f k * b ^ k) t b" using aux2_digit_gen_sum_repr[of "\k = 0..q. f k * b ^ k" "b" "Suc q" "t" "f (Suc q)"] series_bound by auto hence "((\k = 0..q. f k * b ^ k) + f (Suc q) * (b * b ^ q)) div b ^ t mod b = (\k = 0..q. f k * b ^ k) div b ^ t mod b" using nth_digit_def by (auto simp:add.commute) thus ?thesis using Suc[of "t"] nth_digit_def True \t \ Suc q\ by auto qed done thus ?thesis using True by auto next case False (* t > Suc q *) hence "t \ Suc (Suc q)" by auto have f_le_bound: "f k \ b-1" for k using bound apply auto by (metis Suc_pred b_def less_Suc_eq_le numeral_2_eq_2 zero_less_Suc zero_less_power) have bound: "(\k = 0..q. f k * b ^ k) < b^(Suc q)" for q apply (induct q) subgoal using bound by (simp add: less_imp_le_nat) subgoal for q proof - assume asm: "(\k = 0..q. f k * b ^ k) < b ^ Suc q" have "(\k = 0..q. f k * b ^ k) + f (Suc q) * (b * b ^ q) \ (\k = 0..q. f k * b ^ k) + (b-1) * (b * b ^ q)" using f_le_bound by auto also have "... < b^(Suc q) + (b-1) * (b * b ^ q)" using asm by auto also have "... \ b * b * b^q" apply auto by (metis One_nat_def Suc_neq_Zero b_def eq_imp_le mult.assoc mult_eq_if numerals(2) power_not_zero) finally show ?thesis using asm by auto qed done have "(\k = 0..q. f k * b ^ k) + f (Suc q) * (b * b ^ q) < (b ^ Suc (Suc q))" using bound[of "Suc q"] by auto also have "... \ b^t" using \t \ Suc (Suc q)\ apply auto by (metis b_def nat_power_less_imp_less not_le numeral_2_eq_2 power_Suc zero_less_Suc zero_less_power) finally show ?thesis using \t \ Suc (Suc q)\ bound[of "Suc q"] nth_digit_def by auto qed qed text \Equivalence condition for the @{text nth_digit} function \<^cite>\"h10lecturenotes"\ (see equation 2.29)\ lemma digit_gen_equiv: assumes "b>1" shows "d = nth_digit a k b \ (\x.\y.(a = x * b^(k+1) + d*b^k +y \ d < b \ y < b^k))" (is "?P \ ?Q") proof assume p: ?P then show ?Q proof(cases "kiii (b-1) * b^ k" using assms mult_le_mono nth_digit_bounded by auto moreover have "(\iii=Suc k..ii=Suc k..n = Suc k..i = Suc k..i=Suc k..i. (nth_digit a i b) * b^(i - Suc k)"] by simp hence "a = x * b^(k+1) + d*b^k +y" using s by auto thus ?thesis using 2 3 by blast next case False then have "a < b^k" using assms power_gt_expt[of b k] by auto moreover have "d = 0" by (simp add: calculation nth_digit_def p) ultimately show ?thesis using assms by force qed next assume ?Q then obtain x y where conds: "a = x * b^(k+1) + d*b^k +y \ d < b \ y < b^k" by auto hence "nth_digit a k b = nth_digit(x * b^(k+1) + d*b^k) k b" using aux3_digit_gen_sum_repr[of y b "k" "x*b + d"] assms by (auto simp: algebra_simps) hence "nth_digit a k b = nth_digit(d*b^k) k b" using aux2_digit_gen_sum_repr[of "d*b^k" "b" "k+1" k x] conds by auto then show ?P using conds nth_digit_def by simp qed end diff --git a/thys/Finite_Fields/Ring_Characteristic.thy b/thys/Finite_Fields/Ring_Characteristic.thy --- a/thys/Finite_Fields/Ring_Characteristic.thy +++ b/thys/Finite_Fields/Ring_Characteristic.thy @@ -1,1017 +1,1016 @@ section \Characteristic of Rings\label{sec:ring_char}\ theory Ring_Characteristic imports "Finite_Fields_Factorization_Ext" "HOL-Algebra.IntRing" "HOL-Algebra.Embedded_Algebras" begin locale finite_field = field + assumes finite_carrier: "finite (carrier R)" begin lemma finite_field_min_order: "order R > 1" proof (rule ccontr) assume a:"\(1 < order R)" have "{\\<^bsub>R\<^esub>,\\<^bsub>R\<^esub>} \ carrier R" by auto hence "card {\\<^bsub>R\<^esub>,\\<^bsub>R\<^esub>} \ card (carrier R)" using card_mono finite_carrier by blast also have "... \ 1" using a by (simp add:order_def) finally have "card {\\<^bsub>R\<^esub>,\\<^bsub>R\<^esub>} \ 1" by blast thus "False" by simp qed lemma (in finite_field) order_pow_eq_self: assumes "x \ carrier R" shows "x [^] (order R) = x" proof (cases "x = \") case True have "order R > 0" using assms(1) order_gt_0_iff_finite finite_carrier by simp then obtain n where n_def:"order R = Suc n" using lessE by blast have "x [^] (order R) = \" unfolding n_def using True by (subst nat_pow_Suc, simp) thus ?thesis using True by simp next case False have x_carr:"x \ carrier (mult_of R)" using False assms by simp have carr_non_empty: "card (carrier R) > 0" using order_gt_0_iff_finite finite_carrier unfolding order_def by simp have "x [^] (order R) = x [^]\<^bsub>mult_of R\<^esub> (order R)" by (simp add:nat_pow_mult_of) also have "... = x [^]\<^bsub>mult_of R\<^esub> (order (mult_of R)+1)" using carr_non_empty unfolding order_def by (intro arg_cong[where f="\t. x [^]\<^bsub>mult_of R\<^esub> t"]) (simp) also have "... = x" using x_carr by (simp add:mult_of.pow_order_eq_1) finally show "x [^] (order R) = x" by simp qed lemma (in finite_field) order_pow_eq_self': assumes "x \ carrier R" shows "x [^] (order R ^ d) = x" proof (induction d) case 0 then show ?case using assms by simp next case (Suc d) have "x [^] order R ^ (Suc d) = x [^] (order R ^ d * order R)" by (simp add:mult.commute) also have "... = (x [^] (order R ^ d)) [^] order R" using assms by (simp add: nat_pow_pow) also have "... = (x [^] (order R ^ d))" using order_pow_eq_self assms by simp also have "... = x" using Suc by simp finally show ?case by simp qed end lemma finite_fieldI: assumes "field R" assumes "finite (carrier R)" shows "finite_field R" using assms unfolding finite_field_def finite_field_axioms_def by auto lemma (in domain) finite_domain_units: assumes "finite (carrier R)" shows "Units R = carrier R - {\}" (is "?lhs = ?rhs") proof have "Units R \ carrier R" by (simp add:Units_def) moreover have "\ \ Units R" by (meson zero_is_prime(1) primeE) ultimately show "Units R \ carrier R - {\}" by blast next have "x \ Units R" if a: "x \ carrier R - {\}" for x proof - have x_carr: "x \ carrier R" using a by blast define f where "f = (\y. y \\<^bsub>R\<^esub> x)" have "inj_on f (carrier R)" unfolding f_def by (rule inj_onI, metis DiffD1 DiffD2 a m_rcancel insertI1) hence "card (carrier R) = card (f ` carrier R)" by (metis card_image) moreover have "f ` carrier R \ carrier R" unfolding f_def by (rule image_subsetI, simp add: ring.ring_simprules x_carr) ultimately have "f ` carrier R = carrier R" using card_subset_eq assms by metis moreover have "\\<^bsub>R\<^esub> \ carrier R" by simp ultimately have "\y \ carrier R. f y = \\<^bsub>R\<^esub>" by (metis image_iff) then obtain y where y_carrier: "y \ carrier R" and y_left_inv: "y \\<^bsub>R\<^esub> x = \\<^bsub>R\<^esub>" using f_def by blast hence y_right_inv: "x \\<^bsub>R\<^esub> y = \\<^bsub>R\<^esub>" by (metis DiffD1 a cring_simprules(14)) show "x \ Units R" using y_carrier y_left_inv y_right_inv by (metis DiffD1 a divides_one factor_def) qed thus "?rhs \ ?lhs" by auto qed text \The following theorem can be found in Lidl and Niederreiter~\<^cite>\\Theorem 1.31\ in "lidl1986"\.\ theorem finite_domains_are_fields: assumes "domain R" assumes "finite (carrier R)" shows "finite_field R" proof - interpret domain R using assms by auto have "Units R = carrier R - {\\<^bsub>R\<^esub>}" using finite_domain_units[OF assms(2)] by simp then have "field R" by (simp add: assms(1) field.intro field_axioms.intro) thus ?thesis using assms(2) finite_fieldI by auto qed definition zfact_iso :: "nat \ nat \ int set" where "zfact_iso p k = Idl\<^bsub>\\<^esub> {int p} +>\<^bsub>\\<^esub> (int k)" context fixes n :: nat assumes n_gt_0: "n > 0" begin private abbreviation I where "I \ Idl\<^bsub>\\<^esub> {int n}" private lemma ideal_I: "ideal I \" by (simp add: int.genideal_ideal) lemma int_cosetI: assumes "u mod (int n) = v mod (int n)" shows "Idl\<^bsub>\\<^esub> {int n} +>\<^bsub>\\<^esub> u = Idl\<^bsub>\\<^esub> {int n} +>\<^bsub>\\<^esub> v" proof - have "u - v \ I" by (metis Idl_subset_eq_dvd assms int_Idl_subset_ideal mod_eq_dvd_iff) thus ?thesis using ideal_I int.quotient_eq_iff_same_a_r_cos by simp qed lemma zfact_iso_inj: "inj_on (zfact_iso n) {.. {.. {..\<^bsub>\\<^esub> (int x) = I +>\<^bsub>\\<^esub> (int y)" by (simp add:zfact_iso_def) hence "int x - int y \ I" by (subst int.quotient_eq_iff_same_a_r_cos[OF ideal_I], auto) hence "int x mod int n = int y mod int n" by (meson Idl_subset_eq_dvd int_Idl_subset_ideal mod_eq_dvd_iff) thus "x = y" using a by simp qed lemma zfact_iso_ran: "zfact_iso n ` {.. carrier (ZFact (int n))" unfolding zfact_iso_def ZFact_def FactRing_simps using int.a_rcosetsI by auto moreover have "x \ zfact_iso n ` {.. carrier (ZFact (int n))" for x proof - obtain y where y_def: "x = I +>\<^bsub>\\<^esub> y" using a unfolding ZFact_def FactRing_simps by auto - obtain z where z_def: "(int z) mod (int n) = y mod (int n)" "z < n" - by (metis Euclidean_Division.pos_mod_sign mod_mod_trivial - nonneg_int_cases of_nat_0_less_iff of_nat_mod n_gt_0 - unique_euclidean_semiring_numeral_class.pos_mod_bound) + define z where \z = nat (y mod int n)\ + with n_gt_0 have z_def: \int z mod int n = y mod int n\ \z < n\ + by (simp_all add: z_def nat_less_iff) have "x = I +>\<^bsub>\\<^esub> y" by (simp add:y_def) also have "... = I +>\<^bsub>\\<^esub> (int z)" by (intro int_cosetI, simp add:z_def) also have "... = zfact_iso n z" by (simp add:zfact_iso_def) finally have "x = zfact_iso n z" by simp thus "x \ zfact_iso n ` {.. 0" using assms(1) prime_gt_0_nat by simp have "Factorial_Ring.prime (int p)" using assms by simp moreover have "finite (carrier (ZFact (int p)))" using fin_zfact[OF p_gt_0] by simp ultimately show ?thesis by (intro finite_domains_are_fields ZFact_prime_is_domain, auto) qed definition int_embed :: "_ \ int \ _" where "int_embed R k = add_pow R k \\<^bsub>R\<^esub>" lemma (in ring) add_pow_consistent: fixes i :: "int" assumes "subring K R" assumes "k \ K" shows "add_pow R i k = add_pow (R \ carrier := K \) i k" (is "?lhs = ?rhs") proof - have a:"subgroup K (add_monoid R)" using assms(1) subring.axioms by auto have "add_pow R i k = k [^]\<^bsub>add_monoid R\carrier := K\\<^esub> i" using add.int_pow_consistent[OF a assms(2)] by simp also have "... = ?rhs" unfolding add_pow_def by simp finally show ?thesis by simp qed lemma (in ring) int_embed_consistent: assumes "subring K R" shows "int_embed R i = int_embed (R \ carrier := K \) i" proof - have a:"\ = \\<^bsub>R \ carrier := K \\<^esub>" by simp have b:"\\<^bsub>R\carrier := K\\<^esub> \ K" using assms subringE(3) by auto show ?thesis unfolding int_embed_def a using b add_pow_consistent[OF assms(1)] by simp qed lemma (in ring) int_embed_closed: "int_embed R k \ carrier R" unfolding int_embed_def using add.int_pow_closed by simp lemma (in ring) int_embed_range: assumes "subring K R" shows "int_embed R k \ K" proof - let ?R' = "R \ carrier := K \" interpret x:ring ?R' using subring_is_ring[OF assms] by simp have "int_embed R k = int_embed ?R' k" using int_embed_consistent[OF assms] by simp also have "... \ K" using x.int_embed_closed by simp finally show ?thesis by simp qed lemma (in ring) int_embed_zero: "int_embed R 0 = \\<^bsub>R\<^esub>" by (simp add:int_embed_def add_pow_def) lemma (in ring) int_embed_one: "int_embed R 1 = \\<^bsub>R\<^esub>" by (simp add:int_embed_def) lemma (in ring) int_embed_add: "int_embed R (x+y) = int_embed R x \\<^bsub>R\<^esub> int_embed R y" by (simp add:int_embed_def add.int_pow_mult) lemma (in ring) int_embed_inv: "int_embed R (-x) = \\<^bsub>R\<^esub> int_embed R x" (is "?lhs = ?rhs") proof - have "?lhs = int_embed R (-x) \ (int_embed R x \ int_embed R x)" using int_embed_closed by simp also have "... = int_embed R (-x) \ int_embed R x \ (\ int_embed R x)" using int_embed_closed by (subst a_minus_def, subst a_assoc, auto) also have "... = int_embed R (-x +x) \ (\ int_embed R x)" by (subst int_embed_add, simp) also have "... = ?rhs" using int_embed_closed by (simp add:int_embed_zero) finally show ?thesis by simp qed lemma (in ring) int_embed_diff: "int_embed R (x-y) = int_embed R x \\<^bsub>R\<^esub> int_embed R y" (is "?lhs = ?rhs") proof - have "?lhs = int_embed R (x + (-y))" by simp also have "... = ?rhs" by (subst int_embed_add, simp add:a_minus_def int_embed_inv) finally show ?thesis by simp qed lemma (in ring) int_embed_mult_aux: "int_embed R (x*int y) = int_embed R x \ int_embed R y" proof (induction y) case 0 then show ?case by (simp add:int_embed_closed int_embed_zero) next case (Suc y) have "int_embed R (x * int (Suc y)) = int_embed R (x + x * int y)" by (simp add:algebra_simps) also have "... = int_embed R x \ int_embed R (x * int y)" by (subst int_embed_add, simp) also have "... = int_embed R x \ \ \ int_embed R x \ int_embed R y" using int_embed_closed by (subst Suc, simp) also have "... = int_embed R x \ (int_embed R 1 \ int_embed R y)" using int_embed_closed by (subst r_distr, simp_all add:int_embed_one) also have "... = int_embed R x \ int_embed R (1+int y)" by (subst int_embed_add, simp) also have "... = int_embed R x \ int_embed R (Suc y)" by simp finally show ?case by simp qed lemma (in ring) int_embed_mult: "int_embed R (x*y) = int_embed R x \\<^bsub>R\<^esub> int_embed R y" proof (cases "y \ 0") case True then obtain y' where y_def: "y = int y'" using nonneg_int_cases by auto have "int_embed R (x * y) = int_embed R (x * int y')" unfolding y_def by simp also have "... = int_embed R x \ int_embed R y'" by (subst int_embed_mult_aux, simp) also have "... = int_embed R x \ int_embed R y" unfolding y_def by simp finally show ?thesis by simp next case False then obtain y' where y_def: "y = - int y'" by (meson nle_le nonpos_int_cases) have "int_embed R (x * y) = int_embed R (-(x * int y'))" unfolding y_def by simp also have "... = \ (int_embed R (x * int y'))" by (subst int_embed_inv, simp) also have "... = \ (int_embed R x \ int_embed R y')" by (subst int_embed_mult_aux, simp) also have "... = int_embed R x \ \ int_embed R y'" using int_embed_closed by algebra also have "... = int_embed R x \ int_embed R (-y')" by (subst int_embed_inv, simp) also have "... = int_embed R x \ int_embed R y" unfolding y_def by simp finally show ?thesis by simp qed lemma (in ring) int_embed_ring_hom: "ring_hom_ring int_ring R (int_embed R)" proof (rule ring_hom_ringI) show "ring int_ring" using int.ring_axioms by simp show "ring R" using ring_axioms by simp show "int_embed R x \ carrier R" if "x \ carrier \" for x using int_embed_closed by simp show "int_embed R (x\\<^bsub>\\<^esub>y) = int_embed R x \ int_embed R y" if "x \ carrier \" "y \ carrier \" for x y using int_embed_mult by simp show "int_embed R (x\\<^bsub>\\<^esub>y) = int_embed R x \ int_embed R y" if "x \ carrier \" "y \ carrier \" for x y using int_embed_add by simp show "int_embed R \\<^bsub>\\<^esub> = \" by (simp add:int_embed_one) qed abbreviation char_subring where "char_subring R \ int_embed R ` UNIV" definition char where "char R = card (char_subring R)" text \This is a non-standard definition for the characteristic of a ring. Commonly~\<^cite>\\Definition 1.43\ in "lidl1986"\ it is defined to be the smallest natural number $n$ such that n-times repeated addition of any number is zero. If no such number exists then it is defined to be $0$. In the case of rings with unit elements --- not that the locale @{locale "ring"} requires unit elements --- the above definition can be simplified to the number of times the unit elements needs to be repeatedly added to reach $0$. The following three lemmas imply that the definition of the characteristic here coincides with the latter definition.\ lemma (in ring) char_bound: assumes "x > 0" assumes "int_embed R (int x) = \" shows "char R \ x" "char R > 0" proof - have "char_subring R \ int_embed R ` ({0.. UNIV" define u where "u = y div (int x)" define v where "v = y mod (int x)" have "int x > 0" using assms by simp hence y_exp: "y = u * int x + v" "v \ 0" "v < int x" unfolding u_def v_def by simp_all have "int_embed R y = int_embed R v" using int_embed_closed unfolding y_exp by (simp add:int_embed_mult int_embed_add assms(2)) also have "... \ int_embed R ` ({0.. int_embed R ` {0.. card {0.. x" by simp have "1 = card {int_embed R 0}" by simp also have "... \ card (int_embed R ` {0.. 0" by simp qed lemma (in ring) embed_char_eq_0: "int_embed R (int (char R)) = \" proof (cases "finite (char_subring R)") case True interpret h: ring_hom_ring "int_ring" R "(int_embed R)" using int_embed_ring_hom by simp define A where "A = {0..int (char R)}" have "card (int_embed R ` A) \ card (char_subring R)" by (intro card_mono[OF True] image_subsetI, simp) also have "... = char R" unfolding char_def by simp also have "... < card A" unfolding A_def by simp finally have "card (int_embed R ` A) < card A" by simp hence "\inj_on (int_embed R) A" using pigeonhole by simp then obtain x y where xy: "x \ A" "y \ A" "x \ y" "int_embed R x = int_embed R y" unfolding inj_on_def by auto define v where "v = nat (max x y - min x y)" have a:"int_embed R v = \" using xy int_embed_closed by (cases "x < y", simp_all add:int_embed_diff v_def) moreover have "v > 0" using xy by (cases "x < y", simp_all add:v_def) ultimately have "char R \ v" using char_bound by simp moreover have "v \ char R" using xy v_def A_def by (cases "x < y", simp_all) ultimately have "char R = v" by simp then show ?thesis using a by simp next case False hence "char R = 0" unfolding char_def by simp then show ?thesis by (simp add:int_embed_zero) qed lemma (in ring) embed_char_eq_0_iff: fixes n :: int shows "int_embed R n = \ \ char R dvd n" proof (cases "char R > 0") case True define r where "r = n mod char R" define s where "s = n div char R" have rs: "r < char R" "r \ 0" "n = r + s * char R" using True by (simp_all add:r_def s_def) have "int_embed R n = int_embed R r" using int_embed_closed unfolding rs(3) by (simp add: int_embed_add int_embed_mult embed_char_eq_0) moreover have "nat r < char R" using rs by simp hence "int_embed R (nat r) \ \ \ nat r = 0" using True char_bound not_less by blast hence "int_embed R r \ \ \ r = 0" using rs by simp ultimately have "int_embed R n = \ \ r = 0" using int_embed_zero by auto also have "r = 0 \ char R dvd n" using r_def by auto finally show ?thesis by simp next case False hence "char R = 0" by simp hence a:"x > 0 \ int_embed R (int x) \ \" for x using char_bound by auto have c:"int_embed R (abs x) \ \ \ int_embed R x \ \" for x using int_embed_closed by (cases "x > 0", simp, simp add:int_embed_inv) have "int_embed R x \ \" if b:"x \ 0" for x proof - have "nat (abs x) > 0" using b by simp hence "int_embed R (nat (abs x)) \ \" using a by blast hence "int_embed R (abs x) \ \" by simp thus ?thesis using c by simp qed hence "int_embed R n = \ \ n = 0" using int_embed_zero by auto also have "n = 0 \ char R dvd n" using False by simp finally show ?thesis by simp qed text \This result can be found in \<^cite>\\Theorem 1.44\ in "lidl1986"\.\ lemma (in domain) characteristic_is_prime: assumes "char R > 0" shows "prime (char R)" proof (rule ccontr) have "\(char R = 1)" using embed_char_eq_0 int_embed_one by auto hence "\(char R dvd 1)" using assms(1) by simp moreover assume "\(prime (char R))" hence "\(irreducible (char R))" using irreducible_imp_prime_elem_gcd prime_elem_nat_iff by blast ultimately obtain p q where pq_def: "p * q = char R" "p > 1" "q > 1" using assms unfolding Factorial_Ring.irreducible_def by auto have "int_embed R p \ int_embed R q = \" using embed_char_eq_0 pq_def by (subst int_embed_mult[symmetric]) (metis of_nat_mult) hence "int_embed R p = \ \ int_embed R q = \" using integral int_embed_closed by simp hence "p*q \ p \ p*q \ q" using char_bound pq_def by auto thus "False" using pq_def(2,3) by simp qed lemma (in ring) char_ring_is_subring: "subring (char_subring R) R" proof - have "subring (int_embed R ` carrier int_ring) R" by (intro ring.carrier_is_subring int.ring_axioms ring_hom_ring.img_is_subring[OF int_embed_ring_hom]) thus ?thesis by simp qed lemma (in cring) char_ring_is_subcring: "subcring (char_subring R) R" using subcringI'[OF char_ring_is_subring] by auto lemma (in domain) char_ring_is_subdomain: "subdomain (char_subring R) R" using subdomainI'[OF char_ring_is_subring] by auto lemma image_set_eqI: assumes "\x. x \ A \ f x \ B" assumes "\x. x \ B \ g x \ A \ f (g x) = x" shows "f ` A = B" using assms by force text \This is the binomial expansion theorem for commutative rings.\ lemma (in cring) binomial_expansion: fixes n :: nat assumes [simp]: "x \ carrier R" "y \ carrier R" shows "(x \ y) [^] n = (\k \ {..n}. int_embed R (n choose k) \ x [^] k \ y [^] (n-k))" proof - define A where "A = (\k. {A. A \ {.. card A = k})" have fin_A: "finite (A i)" for i unfolding A_def by simp have disj_A: "pairwise (\i j. disjnt (A i) (A j)) {..n}" unfolding pairwise_def disjnt_def A_def by auto have card_A: "B \ A i \ card B = i" if " i \ {..n}" for i B unfolding A_def by simp have card_A2: "card (A i) = (n choose i)" if "i \ {..n}" for i unfolding A_def using n_subsets[where A="{.. n" if "A \ {.. {..<(n::nat)}" for n A using finite_subset that by (subst card_insert_disjoint, auto) have embed_distr: "[m] \ y = int_embed R (int m) \ y" if "y \ carrier R" for m y unfolding int_embed_def add_pow_def using that by (simp add:add_pow_def[symmetric] int_pow_int add_pow_ldistr) have "(x \ y) [^] n = (\A \ Pow {.. y [^] (n-card A))" proof (induction n) case 0 then show ?case by simp next case (Suc n) have s1: "insert n ` Pow {.. {.. n \ A}" by (intro image_set_eqI[where g="\x. x \ {.. {.. n \ A}" using lessThan_Suc by auto have "(x \ y) [^] Suc n = (x \ y) [^] n \ (x \ y)" by simp also have "... = (\A \ Pow {.. y [^] (n-card A)) \ (x \ y)" by (subst Suc, simp) also have "... = (\A \ Pow {.. y [^] (n-card A)) \ x \ (\A \ Pow {.. y [^] (n-card A)) \ y" by (subst r_distr, auto) also have "... = (\A \ Pow {.. y [^] (n-card A) \ x) \ (\A \ Pow {.. y [^] (n-card A) \ y)" by (simp add:finsum_ldistr) also have "... = (\A \ Pow {.. y [^] (n-card A)) \ (\A \ Pow {.. y [^] (n-card A+1))" using m_assoc m_comm by (intro arg_cong2[where f="(\)"] finsum_cong', auto) also have "... = (\A \ Pow {.. y [^] (n+1-card (insert n A))) \ (\A \ Pow {.. y [^] (n+1-card A))" using finite_subset card_bound card_insert Suc_diff_le by (intro arg_cong2[where f="(\)"] finsum_cong', simp_all) also have "... = (\A \ insert n ` Pow {.. y [^] (n+1-card A)) \ (\A \ Pow {.. y [^] (n+1-card A))" by (subst finsum_reindex, auto simp add:inj_on_def) also have "... = (\A \ {A. A \ {.. n \ A}. x [^] (card A) \ y [^] (n+1-card A)) \ (\A \ {A. A \ {.. n \ A}. x [^] (card A) \ y [^] (n+1-card A))" by (intro arg_cong2[where f="(\)"] finsum_cong' s1 s2, simp_all) also have "... = (\A \ {A. A \ {.. n \ A} \ {A. A \ {.. n \ A}. x [^] (card A) \ y [^] (n+1-card A))" by (subst finsum_Un_disjoint, auto) also have "... = (\A \ Pow {.. y [^] (n+1-card A))" by (intro finsum_cong', auto) finally show ?case by simp qed also have "... = (\A \ (\ (A ` {..n})). x [^] (card A) \ y [^] (n-card A))" using card_bound by (intro finsum_cong', auto simp add:A_def) also have "... = (\ k \ {..n}. (\ A \ A k. x [^] (card A) \ y [^] (n-card A)))" using fin_A disj_A by (subst add.finprod_UN_disjoint, auto) also have "... = (\ k \ {..n}. (\ A \ A k. x [^] k \ y [^] (n-k)))" using card_A by (intro finsum_cong', auto) also have "... = (\ k \ {..n}. int_embed R (card (A k)) \ x [^] k \ y [^] (n-k))" using int_embed_closed by (subst add.finprod_const, simp_all add:embed_distr m_assoc) also have "... = (\ k \ {..n}. int_embed R (n choose k) \ x [^] k \ y [^] (n-k))" using int_embed_closed card_A2 by (intro finsum_cong', simp_all) finally show ?thesis by simp qed lemma bin_prime_factor: assumes "prime p" assumes "k > 0" "k < p" shows "p dvd (p choose k)" proof - have "p dvd fact p" using assms(1) prime_dvd_fact_iff by auto hence "p dvd fact k * fact (p - k) * (p choose k)" using binomial_fact_lemma assms by simp hence "p dvd fact k \ p dvd fact (p-k) \ p dvd (p choose k)" by (simp add: assms(1) prime_dvd_mult_eq_nat) thus "p dvd (p choose k)" using assms(1,2,3) prime_dvd_fact_iff by auto qed theorem (in domain) freshmans_dream: assumes "char R > 0" assumes [simp]: "x \ carrier R" "y \ carrier R" shows "(x \ y) [^] (char R) = x [^] char R \ y [^] char R" (is "?lhs = ?rhs") proof - have c:"prime (char R)" using assms(1) characteristic_is_prime by auto have a:"int_embed R (char R choose i) = \" if "i \ {..char R} - {0, char R}" for i proof - have "i > 0" "i < char R" using that by auto hence "char R dvd char R choose i" using c bin_prime_factor by simp thus ?thesis using embed_char_eq_0_iff by simp qed have "?lhs = (\k \ {..char R}. int_embed R (char R choose k) \ x [^] k \ y [^] (char R-k))" using binomial_expansion[OF assms(2,3)] by simp also have "... = (\k \ {0,char R}.int_embed R (char R choose k) \ x [^] k \ y [^] (char R-k))" using a int_embed_closed by (intro add.finprod_mono_neutral_cong_right, simp, simp_all) also have "... = ?rhs" using int_embed_closed assms(1) by (simp add:int_embed_one a_comm) finally show ?thesis by simp qed text \The following theorem is somtimes called Freshman's dream for obvious reasons, it can be found in Lidl and Niederreiter~\<^cite>\\Theorem 1.46\ in "lidl1986"\.\ lemma (in domain) freshmans_dream_ext: fixes m assumes "char R > 0" assumes [simp]: "x \ carrier R" "y \ carrier R" defines "n \ char R^m" shows "(x \ y) [^] n = x [^] n \ y [^] n" (is "?lhs = ?rhs") unfolding n_def proof (induction m) case 0 then show ?case by simp next case (Suc m) have "(x \ y) [^] (char R^(m+1)) = (x \ y) [^] (char R^m * char R)" by (simp add:mult.commute) also have "... = ((x \ y) [^] (char R^m)) [^] char R" using nat_pow_pow by simp also have "... = (x [^] (char R^m) \ y [^] (char R^m)) [^] char R" by (subst Suc, simp) also have "... = (x [^] (char R^m)) [^] char R \ (y [^] (char R^m)) [^] char R" by (subst freshmans_dream[OF assms(1), symmetric], simp_all) also have "... = x [^] (char R^m * char R) \ y [^] (char R^m * char R)" by (simp add:nat_pow_pow) also have "... = x [^] (char R^Suc m) \ y [^] (char R^Suc m)" by (simp add:mult.commute) finally show ?case by simp qed text \The following is a generalized version of the Frobenius homomorphism. The classic version of the theorem is the case where @{term "(k::nat) = 1"}.\ theorem (in domain) frobenius_hom: assumes "char R > 0" assumes "m = char R ^ k" shows "ring_hom_cring R R (\x. x [^] m)" proof - have a:"(x \ y) [^] m = x [^] m \ y [^] m" if b:"x \ carrier R" "y \ carrier R" for x y using b nat_pow_distrib by simp have b:"(x \ y) [^] m = x [^] m \ y [^] m" if b:"x \ carrier R" "y \ carrier R" for x y unfolding assms(2) freshmans_dream_ext[OF assms(1) b] by simp have "ring_hom_ring R R (\x. x [^] m)" by (intro ring_hom_ringI a b ring_axioms, simp_all) thus "?thesis" using RingHom.ring_hom_cringI is_cring by blast qed lemma (in domain) char_ring_is_subfield: assumes "char R > 0" shows "subfield (char_subring R) R" proof - interpret d:domain "R \ carrier := char_subring R \" using char_ring_is_subdomain subdomain_is_domain by simp have "finite (char_subring R)" using char_def assms by (metis card_ge_0_finite) hence "Units (R \ carrier := char_subring R \) = char_subring R - {\}" using d.finite_domain_units by simp thus ?thesis using subfieldI[OF char_ring_is_subcring] by simp qed lemma card_lists_length_eq': fixes A :: "'a set" shows "card {xs. set xs \ A \ length xs = n} = card A ^ n" proof (cases "finite A") case True then show ?thesis using card_lists_length_eq by auto next case False hence inf_A: "infinite A" by simp show ?thesis proof (cases "n = 0") case True hence "card {xs. set xs \ A \ length xs = n} = card {([] :: 'a list)}" by (intro arg_cong[where f="card"], auto simp add:set_eq_iff) also have "... = 1" by simp also have "... = card A^n" using True inf_A by simp finally show ?thesis by simp next case False hence "inj (replicate n)" by (meson inj_onI replicate_eq_replicate) hence "inj_on (replicate n) A" using inj_on_subset by (metis subset_UNIV) hence "infinite (replicate n ` A)" using inf_A finite_image_iff by auto moreover have "replicate n ` A \ {xs. set xs \ A \ length xs = n}" by (intro image_subsetI, auto) ultimately have "infinite {xs. set xs \ A \ length xs = n}" using infinite_super by auto hence "card {xs. set xs \ A \ length xs = n} = 0" by simp then show ?thesis using inf_A False by simp qed qed lemma (in ring) card_span: assumes "subfield K R" assumes "independent K w" assumes "set w \ carrier R" shows "card (Span K w) = card K^(length w)" proof - define A where "A = {x. set x \ K \ length x = length w}" define f where "f = (\x. combine x w)" have "x \ f ` A" if a:"x \ Span K w" for x proof - obtain y where "y \ A" "x = f y" unfolding A_def f_def using unique_decomposition[OF assms(1,2) a] by auto thus ?thesis by simp qed moreover have "f x \ Span K w" if a: "x \ A" for x using Span_eq_combine_set[OF assms(1,3)] a unfolding A_def f_def by auto ultimately have b:"Span K w = f ` A" by auto have "False" if a: "x \ A" "y \ A" "f x = f y" "x \ y" for x y proof - have "f x \ Span K w" using b a by simp thus "False" using a unique_decomposition[OF assms(1,2)] unfolding f_def A_def by blast qed hence f_inj: "inj_on f A" unfolding inj_on_def by auto have "card (Span K w) = card (f ` A)" using b by simp also have "... = card A" by (intro card_image f_inj) also have "... = card K^length w" unfolding A_def by (intro card_lists_length_eq') finally show ?thesis by simp qed lemma (in ring) finite_carr_imp_char_ge_0: assumes "finite (carrier R)" shows "char R > 0" proof - have "char_subring R \ carrier R" using int_embed_closed by auto hence "finite (char_subring R)" using finite_subset assms by auto hence "card (char_subring R) > 0" using card_range_greater_zero by simp thus "char R > 0" unfolding char_def by simp qed lemma (in ring) char_consistent: assumes "subring H R" shows "char (R \ carrier := H \) = char R" proof - show ?thesis using int_embed_consistent[OF assms(1)] unfolding char_def by simp qed lemma (in ring_hom_ring) char_consistent: assumes "inj_on h (carrier R)" shows "char R = char S" proof - have a:"h (int_embed R (int n)) = int_embed S (int n)" for n using R.int_embed_range[OF R.carrier_is_subring] using R.int_embed_range[OF R.carrier_is_subring] using S.int_embed_one R.int_embed_one using S.int_embed_zero R.int_embed_zero using S.int_embed_add R.int_embed_add by (induction n, simp_all) have b:"h (int_embed R (-(int n))) = int_embed S (-(int n))" for n using R.int_embed_range[OF R.carrier_is_subring] using S.int_embed_range[OF S.carrier_is_subring] a by (simp add:R.int_embed_inv S.int_embed_inv) have c:"h (int_embed R n) = int_embed S n" for n proof (cases "n \ 0") case True then obtain m where "n = int m" using nonneg_int_cases by auto then show ?thesis by (simp add:a) next case False hence "n \ 0" by simp then obtain m where "n = -int m" using nonpos_int_cases by auto then show ?thesis by (simp add:b) qed have "char S = card (h ` char_subring R)" unfolding char_def image_image c by simp also have "... = card (char_subring R)" using R.int_embed_range[OF R.carrier_is_subring] by (intro card_image inj_on_subset[OF assms(1)]) auto also have "... = char R" unfolding char_def by simp finally show ?thesis by simp qed definition char_iso :: "_ \ int set \ 'a" where "char_iso R x = the_elem (int_embed R ` x)" text \The function @{term "char_iso R"} denotes the isomorphism between @{term "ZFact (char R)"} and the characteristic subring.\ lemma (in ring) char_iso: "char_iso R \ ring_iso (ZFact (char R)) (R\carrier := char_subring R\)" proof - interpret h: ring_hom_ring "int_ring" "R" "int_embed R" using int_embed_ring_hom by simp have "a_kernel \ R (int_embed R) = {x. int_embed R x = \}" unfolding a_kernel_def kernel_def by simp also have "... = {x. char R dvd x}" using embed_char_eq_0_iff by simp also have "... = PIdl\<^bsub>\\<^esub> (int (char R))" unfolding cgenideal_def by auto also have "... = Idl\<^bsub>\\<^esub> {int (char R)}" using int.cgenideal_eq_genideal by simp finally have a:"a_kernel \ R (int_embed R) = Idl\<^bsub>\\<^esub> {int (char R)}" by simp show "?thesis" unfolding char_iso_def ZFact_def a[symmetric] by (intro h.FactRing_iso_set_aux) qed text \The size of a finite field must be a prime power. This can be found in Ireland and Rosen~\<^cite>\\Proposition 7.1.3\ in "ireland1982"\.\ theorem (in finite_field) finite_field_order: "\n. order R = char R ^ n \ n > 0" proof - have a:"char R > 0" using finite_carr_imp_char_ge_0[OF finite_carrier] by simp let ?CR = "char_subring R" obtain v where v_def: "set v = carrier R" using finite_carrier finite_list by auto hence b:"set v \ carrier R" by auto have "carrier R = set v" using v_def by simp also have "... \ Span ?CR v" using Span_base_incl[OF char_ring_is_subfield[OF a] b] by simp finally have "carrier R \ Span ?CR v" by simp moreover have "Span ?CR v \ carrier R" using int_embed_closed v_def by (intro Span_in_carrier, auto) ultimately have Span_v: "Span ?CR v = carrier R" by simp obtain w where w_def: "set w \ carrier R" "independent ?CR w" "Span ?CR v = Span ?CR w" using b filter_base[OF char_ring_is_subfield[OF a]] by metis have Span_w: "Span ?CR w = carrier R" using w_def(3) Span_v by simp hence "order R = card (Span ?CR w)" by (simp add:order_def) also have "... = card ?CR^length w" by (intro card_span char_ring_is_subfield[OF a] w_def(1,2)) finally have c: "order R = char R^(length w)" by (simp add:char_def) have "length w > 0" using finite_field_min_order c by auto thus ?thesis using c by auto qed end diff --git a/thys/Finitely_Generated_Abelian_Groups/Finitely_Generated_Abelian_Groups.thy b/thys/Finitely_Generated_Abelian_Groups/Finitely_Generated_Abelian_Groups.thy --- a/thys/Finitely_Generated_Abelian_Groups/Finitely_Generated_Abelian_Groups.thy +++ b/thys/Finitely_Generated_Abelian_Groups/Finitely_Generated_Abelian_Groups.thy @@ -1,982 +1,984 @@ (* File: Finitely_Generated_Abelian_Groups.thy Author: Joseph Thommes, TU München *) section \Fundamental Theorem of Finitely Generated Abelian Groups\ theory Finitely_Generated_Abelian_Groups imports DirProds Group_Relations begin notation integer_mod_group ("Z") locale fin_gen_comm_group = comm_group + fixes gen :: "'a set" assumes gens_closed: "gen \ carrier G" and fin_gen: "finite gen" and generators: "carrier G = generate G gen" text \Every finite abelian group is also finitely generated.\ sublocale finite_comm_group \ fin_gen_comm_group G "carrier G" using generate_incl generate_sincl by (unfold_locales, auto) text \This lemma contains the proof of Kemper from his lecture notes on algebra~\<^cite>\"kemper"\. However, the proof is not done in the context of a finitely generated group but for a finitely generated subgroup in a commutative group.\ lemma (in comm_group) ex_idirgen: fixes A :: "'a set" assumes "finite A" "A \ carrier G" shows "\gs. set gs \ generate G A \ distinct gs \ is_idirprod (generate G A) (\g. generate G {g}) (set gs) \ successively (dvd) (map ord gs) \ card (set gs) \ card A" (is "?t A") using assms proof (induction "card A" arbitrary: A rule: nat_less_induct) case i: 1 show ?case proof (cases "relations A = {restrict (\_. 0::int) A}") (* only trivial relation *) case True have fi: "finite A" by fact then obtain gs where gs: "set gs = A" "distinct gs" by (meson finite_distinct_list) have o: "ord g = 0" if "g \ set gs" for g by (intro relations_zero_imp_ord_zero[OF that], use i(3) that True gs in auto) have m: "map ord gs = replicate (length gs) 0" using o by (induction gs; auto) show ?thesis proof(rule, safe) show "\x. x \ set gs \ x \ generate G A" using gs generate.incl[of _ A G] by blast show "distinct gs" by fact show "is_idirprod (generate G A) (\g. generate G {g}) (set gs)" proof(unfold is_idirprod_def, intro conjI, rule) show "generate G {g} \ G" if "g \ set gs" for g by (intro subgroup_imp_normal, use that generate_is_subgroup i(3) gs in auto) show "generate G A = IDirProds G (\g. generate G {g}) (set gs)" unfolding IDirProds_def by (subst gs(1), use generate_idem_Un i(3) in blast) show "compl_fam (\g. generate G {g}) (set gs)" using compl_fam_iff_relations_triv[OF i(2, 3)] o gs(1) True by blast qed show "successively (dvd) (map ord gs)" using m proof (induction gs) case c: (Cons a gs) thus ?case by(cases gs; simp) qed simp show "card (set gs) \ card A" using gs by blast qed next case ntrel: False then have Ane: "A \ {}" using i(2) triv_rel[of A] unfolding relations_def extensional_def by fastforce from ntrel obtain a where a: "a \ A" "\r \relations A. r a \ 0" using i(2) triv_rel[of A] unfolding relations_def extensional_def by fastforce hence ac: "a \ carrier G" using i(3) by blast have iH: "\B.\card B < card A; finite B; B \ carrier G\ \ ?t B" using i(1) by blast have iH2: "\B. \?t B; generate G A = generate G B; card B < card A\ \ ?t A" by fastforce show ?thesis proof(cases "inv a \ (A - {a})") case True have "generate G A = generate G (A - {a})" proof(intro generate_subset_eqI[OF i(3)]) show "A - (A - {a}) \ generate G (A - {a})" proof - have "A - (A - {a}) = {a}" using a True by auto also have "\ \ generate G {inv a}" using generate.inv[of "inv a" "{inv a}" G] ac by simp also have "\ \ generate G (A - {a})" by (intro mono_generate, use True in simp) finally show ?thesis . qed qed simp moreover have "?t (A - {a})" by (intro iH[of "A - {a}"], use i(2, 3) a(1) in auto, meson Ane card_gt_0_iff diff_Suc_less) ultimately show ?thesis using card.remove[OF i(2) a(1)] by fastforce next case inv: False define n where n: "n = card A" define all_gens where "all_gens = {gs\Pow (generate G A). finite gs \ card gs \ n \ generate G gs = generate G A}" define exps where "exps = (\gs'\all_gens. \rel\relations gs'. nat ` {e\rel`gs'. e > 0})" define min_exp where "min_exp = Inf exps" have "exps \ {}" proof - let ?B = "A - {a} \ {inv a}" have "A \ all_gens" unfolding all_gens_def using generate.incl n i(2) by fast moreover have "?B \ all_gens" proof - have "card (A - {a}) = n - 1" using a n by (meson card_Diff_singleton_if i(2)) hence "card ?B = n" using inv i(2, 3) n a(1) by (metis Un_empty_right Un_insert_right card.remove card_insert_disjoint finite_Diff) moreover have "generate G A = generate G ?B" proof(intro generate_one_switched_eqI[OF i(3) a(1), of _ "inv a"]) show "inv a \ generate G A" using generate.inv[OF a(1), of G] . show "a \ generate G ?B" proof - have "a \ generate G {inv a}" using generate.inv[of "inv a" "{inv a}" G] ac by simp also have "\ \ generate G ?B" by (intro mono_generate, blast) finally show ?thesis . qed qed simp moreover hence "?B \ generate G A" using generate_sincl by simp ultimately show ?thesis unfolding all_gens_def using i(2) by blast qed moreover have "(\r \ relations A. r a > 0) \ (\r \ relations ?B. r (inv a) > 0)" proof(cases "\r \ relations A. r a > 0") case True then show ?thesis by blast next case False with a obtain r where r: "r \ relations A" "r a < 0" by force have rc: "(\x. x [^] r x) \ A \ carrier G" using i(3) int_pow_closed by fast let ?r = "restrict (r(inv a := - r a)) ?B" have "?r \ relations ?B" proof have "finprod G (\x. x [^] ?r x) ?B = finprod G (\x. x [^] r x) A" proof - have "finprod G (\x. x [^] ?r x) ?B = finprod G (\x. x [^] ?r x) (insert (inv a) (A - {a}))" by simp also have "\ = (inv a) [^] ?r (inv a) \ finprod G (\x. x [^] ?r x) (A - {a})" proof(intro finprod_insert[OF _ inv]) show "finite (A - {a})" using i(2) by fast show "inv a [^] ?r (inv a) \ carrier G" using int_pow_closed[OF inv_closed[OF ac]] by fast show "(\x. x [^] ?r x) \ A - {a} \ carrier G" using int_pow_closed i(3) by fast qed also have "\ = a [^] r a \ finprod G (\x. x [^] r x) (A - {a})" proof - have "(inv a) [^] ?r (inv a) = a [^] r a" by (simp add: int_pow_inv int_pow_neg ac) moreover have "finprod G (\x. x [^] r x) (A - {a}) = finprod G (\x. x [^] ?r x) (A - {a})" proof(intro finprod_cong) show "((\x. x [^] r x) \ A - {a} \ carrier G) = True" using rc by blast have "i [^] r i = i [^] ?r i" if "i \ A - {a}" for i using that inv by auto thus "\i. i \ A - {a} =simp=> i [^] r i = i [^] restrict (r(inv a := - r a)) (A - {a} \ {inv a}) i" by algebra qed simp ultimately show ?thesis by argo qed also have "\ = finprod G (\x. x [^] r x) A" by (intro finprod_minus[symmetric, OF a(1) rc i(2)]) finally show ?thesis . qed also have "\ = \" using r unfolding relations_def by fast finally show "finprod G (\x. x [^] ?r x) ?B = \" . qed simp then show ?thesis using r by fastforce qed ultimately show ?thesis unfolding exps_def using a by blast qed hence me: "min_exp \ exps" unfolding min_exp_def using Inf_nat_def1 by force from cInf_lower min_exp_def have le: "\x. x \ exps \ min_exp \ x" by blast from me obtain gs rel g where gr: "gs \ all_gens" "rel \ relations gs" "g \ gs" "rel g = min_exp" "min_exp > 0" unfolding exps_def by auto from gr(1) have cgs: "card gs \ card A" unfolding all_gens_def n by blast with gr(3) have cgsg: "card (gs - {g}) < card A" by (metis Ane card.infinite card_Diff1_less card_gt_0_iff finite.emptyI finite.insertI finite_Diff2 i.prems(1) le_neq_implies_less less_trans) from gr(1) have fgs: "finite gs" and gsg: "generate G gs = generate G A" unfolding all_gens_def n using i(2) card.infinite Ane by force+ from gsg have gsc: "gs \ carrier G" unfolding all_gens_def using generate_incl[OF i(3)] generate_sincl[of gs] by simp hence gc: "g \ carrier G" using gr(3) by blast have ihgsg: "?t (gs - {g})" by (intro iH, use cgs fgs gsc gr(3) cgsg in auto) then obtain hs where hs: "set hs \ generate G (gs - {g})" "distinct hs" "is_idirprod (generate G (gs - {g})) (\g. generate G {g}) (set hs)" "successively (dvd) (map ord hs)" "card (set hs) \ card (gs - {g})" by blast hence hsc: "set hs \ carrier G" using generate_sincl[of "set hs"] generate_incl[of "gs - {g}"] gsc by blast from hs(3) have ghs: "generate G (gs - {g}) = generate G (set hs)" unfolding is_idirprod_def IDirProds_def using generate_idem_Un[OF hsc] by argo have dvot: "?t A \ (\e\rel`gs. rel g dvd e)" proof(intro disjCI) assume na: "\ (\e\rel ` gs. rel g dvd e)" have "\x. \x \ gs; \rel g dvd rel x\ \ ?t A" proof - fix x assume x: "x \ gs" "\ rel g dvd rel x" hence xng: "x \ g" by auto from x have xc: "x \ carrier G" using gsc by blast have rg: "rel g > 0" using gr by simp define r::int where r: "r = rel x mod rel g" define q::int where q: "q = rel x div rel g" from r rg x have "r > 0" using mod_int_pos_iff[of "rel x" "rel g"] mod_eq_0_iff_dvd by force - moreover have "r < rel g" using r rg Euclidean_Division.pos_mod_bound by blast + moreover have "r < rel g" using r rg by simp moreover have "rel x = q * rel g + r" using r q by presburger ultimately have rq: "rel x = q * (rel g) + r" "0 < r" "r < rel g" by auto define t where t: "t = g \ x [^] q" hence tc: "t \ carrier G" using gsc gr(3) x by fast define s where s: "s = gs - {g} \ {t}" hence fs: "finite s" using fgs by blast have sc: "s \ carrier G" using s tc gsc by blast have g: "generate G gs = generate G s" proof(unfold s, intro generate_one_switched_eqI[OF gsc gr(3), of _ t]) show "t \ generate G gs" proof(unfold t, intro generate.eng) show "g \ generate G gs" using gr(3) generate.incl by fast show "x [^] q \ generate G gs" using x generate_pow[OF xc] generate_sincl[of "{x}"] mono_generate[of "{x}" gs] by fast qed show "g \ generate G (gs - {g} \ {t})" proof - have gti: "g = t \ inv (x [^] q)" using inv_solve_right[OF gc tc int_pow_closed[OF xc, of q]] t by blast moreover have "t \ generate G (gs - {g} \ {t})" by (intro generate.incl[of t], simp) moreover have "inv (x [^] q) \ generate G (gs - {g})" proof - have "x [^] q \ generate G {x}" using generate_pow[OF xc] by blast from generate_m_inv_closed[OF _ this] xc have "inv (x [^] q) \ generate G {x}" by blast moreover have "generate G {x} \ generate G (gs - {g})" by (intro mono_generate, use x a in force) finally show ?thesis . qed ultimately show ?thesis using generate.eng mono_generate[of "gs - {g}" "gs - {g} \ {t}"] by fast qed qed simp show "\x \ gs; \ rel g dvd rel x\ \ ?t A" proof (cases "t \ gs - {g}") case xt: True from xt have gts: "s = gs - {g}" using x s by auto moreover have "card (gs - {g}) < card gs" using fgs gr(3) by (meson card_Diff1_less) ultimately have "card (set hs) < card A" using hs(5) cgs by simp moreover have "set hs \ generate G (set hs)" using generate_sincl by simp moreover have "distinct hs" by fact moreover have "is_idirprod (generate G (set hs)) (\g. generate G {g}) (set hs)" using hs ghs unfolding is_idirprod_def by blast moreover have "generate G A = generate G (set hs)" using g gts ghs gsg by argo moreover have "successively (dvd) (map ord hs)" by fact ultimately show "?t A" using iH2 by blast next case tngsg: False hence xnt: "x \ t" using x xng by blast have "rel g dvd rel x" proof (rule ccontr) have "nat r \ exps" unfolding exps_def proof show "s \ all_gens" unfolding all_gens_def using gsg g fgs generate_sincl[of s] switch_elem_card_le[OF gr(3), of t] cgs n s by auto have xs: "x \ s" using s xng x(1) by blast have ts: "t \ s" using s by fast show "nat r \ (\rel\relations s. nat ` {e \ rel ` s. 0 < e})" proof let ?r = "restrict (rel(x := r, t := rel g)) s" show "?r \ relations s" proof have "finprod G (\x. x [^] ?r x) s = finprod G (\x. x [^] rel x) gs" proof - have "finprod G (\x. x [^] ?r x) s = x [^] r \ (t [^] rel g \ finprod G (\x. x [^] rel x) (gs - {g} - {x}))" proof - have "finprod G (\x. x [^] ?r x) s = x [^] ?r x \ finprod G (\x. x [^] ?r x) (s - {x})" by (intro finprod_minus[OF xs _ fs], use sc in auto) moreover have "finprod G (\x. x [^] ?r x) (s - {x}) = t [^] ?r t \ finprod G (\x. x [^] ?r x) (s - {x} - {t})" by (intro finprod_minus, use ts xnt fs sc in auto) moreover have "finprod G (\x. x [^] ?r x) (s - {x} - {t}) = finprod G (\x. x [^] rel x) (s - {x} - {t})" unfolding s by (intro finprod_cong', use gsc in auto) moreover have "s - {x} - {t} = gs - {g} - {x}" unfolding s using tngsg by blast moreover hence "finprod G (\x. x [^] rel x) (s - {x} - {t}) = finprod G (\x. x [^] rel x) (gs - {g} - {x})" by simp moreover have "x [^] ?r x = x [^] r" using xs xnt by auto moreover have "t [^] ?r t = t [^] rel g" using ts by simp ultimately show ?thesis by argo qed also have "\ = x [^] r \ t [^] rel g \ finprod G (\x. x [^] rel x) (gs - {g} - {x})" by (intro m_assoc[symmetric], use xc tc in simp_all, intro finprod_closed, use gsc in fast) also have "\ = g [^] rel g \ x [^] rel x \ finprod G (\x. x [^] rel x) (gs - {g} - {x})" proof - have "x [^] r \ t [^] rel g = g [^] rel g \ x [^] rel x" proof - have "x [^] r \ t [^] rel g = x [^] r \ (g \ x [^] q) [^] rel g" using t by blast also have "\ = x [^] r \ x [^] (q * rel g) \ g [^] rel g" proof - have "(g \ x [^] q) [^] rel g = g [^] rel g \ (x [^] q) [^] rel g" using gc xc int_pow_distrib by auto moreover have "(x [^] q) [^] rel g = x [^] (q * rel g)" using xc int_pow_pow by auto moreover have "g [^] rel g \ x [^] (q * rel g) = x [^] (q * rel g) \ g [^] rel g" using m_comm[OF int_pow_closed[OF xc] int_pow_closed[OF gc]] by simp ultimately have "(g \ x [^] q) [^] rel g = x [^] (q * rel g) \ g [^] rel g" by argo thus ?thesis by (simp add: gc m_assoc xc) qed also have "\ = x [^] rel x \ g [^] rel g" proof - have "x [^] r \ x [^] (q * rel g) = x [^] (q * rel g + r)" by (simp add: add.commute int_pow_mult xc) also have "\ = x [^] rel x" using rq by argo finally show ?thesis by argo qed finally show ?thesis by (simp add: gc m_comm xc) qed thus ?thesis by simp qed also have "\ = g [^] rel g \ (x [^] rel x \ finprod G (\x. x [^] rel x) (gs - {g} - {x}))" by (intro m_assoc, use xc gc in simp_all, intro finprod_closed, use gsc in fast) also have "\ = g [^] rel g \ finprod G (\x. x [^] rel x) (gs - {g})" proof - have "finprod G (\x. x [^] rel x) (gs - {g}) = x [^] rel x \ finprod G (\x. x [^] rel x) (gs - {g} - {x})" by (intro finprod_minus, use xng x(1) fgs gsc in auto) thus ?thesis by argo qed also have "\ = finprod G (\x. x [^] rel x) gs" by (intro finprod_minus[symmetric, OF gr(3) _ fgs], use gsc in auto) finally show ?thesis . qed thus "finprod G (\x. x [^] ?r x) s = \" using gr(2) unfolding relations_def by simp qed auto show "nat r \ nat ` {e \ ?r ` s. 0 < e}" using xs xnt rq(2) by fastforce qed qed from le[OF this] rq(3) gr(4, 5) show False by linarith qed thus "\x \ gs; \ rel g dvd rel x\ \ ?t A" by blast qed qed thus "?t A" using na by blast qed show "?t A" proof (cases "\e\rel`gs. rel g dvd e") case dv: True define tau where "tau = finprod G (\x. x [^] ((rel x) div rel g)) gs" have tc: "tau \ carrier G" by (subst tau_def, intro finprod_closed[of "(\x. x [^] ((rel x) div rel g))" gs], use gsc in fast) have gts: "generate G gs = generate G (gs - {g} \ {tau})" proof(intro generate_one_switched_eqI[OF gsc gr(3), of _ tau]) show "tau \ generate G gs" by (subst generate_eq_finprod_Pi_int_image[OF fgs gsc], unfold tau_def, fast) show "g \ generate G (gs - {g} \ {tau})" proof - have "tau = g \ finprod G (\x. x [^] ((rel x) div rel g)) (gs - {g})" proof - have "finprod G (\x. x [^] ((rel x) div rel g)) gs = g [^] (rel g div rel g) \ finprod G (\x. x [^] ((rel x) div rel g)) (gs - {g})" by (intro finprod_minus[OF gr(3) _ fgs], use gsc in fast) moreover have "g [^] (rel g div rel g) = g" using gr gsc by auto ultimately show ?thesis unfolding tau_def by argo qed hence gti: "g = tau \ inv finprod G (\x. x [^] ((rel x) div rel g)) (gs - {g})" using inv_solve_right[OF gc tc finprod_closed[of "(\x. x [^] ((rel x) div rel g))" "gs - {g}"]] gsc by fast have "tau \ generate G (gs - {g} \ {tau})" by (intro generate.incl[of tau], simp) moreover have "inv finprod G (\x. x [^] ((rel x) div rel g)) (gs - {g}) \ generate G (gs - {g})" proof - have "finprod G (\x. x [^] ((rel x) div rel g)) (gs - {g}) \ generate G (gs - {g})" using generate_eq_finprod_Pi_int_image[of "gs - {g}"] fgs gsc by fast from generate_m_inv_closed[OF _ this] gsc show ?thesis by blast qed ultimately show ?thesis by (subst gti, intro generate.eng, use mono_generate[of "gs - {g}"] in auto) qed qed simp with gr(1) have gt: "generate G (gs - {g} \ {tau}) = generate G A" unfolding all_gens_def by blast have trgo: "tau [^] rel g = \" proof - have "tau [^] rel g = finprod G (\x. x [^] ((rel x) div rel g)) gs [^] rel g" unfolding tau_def by blast also have "\ = finprod G ((\x. x [^] rel g) \ (\x. x [^] ((rel x) div rel g))) gs" by (intro finprod_exp, use gsc in auto) also have "\ = finprod G (\a. a [^] rel a) gs" proof(intro finprod_cong') show "((\x. x [^] rel g) \ (\x. x [^] ((rel x) div rel g))) x = x [^] rel x" if "x \ gs" for x proof - have "((\x. x [^] rel g) \ (\x. x [^] ((rel x) div rel g))) x = x [^] (((rel x) div rel g) * rel g)" using that gsc int_pow_pow by auto also have "\ = x [^] rel x" using dv that by auto finally show ?thesis . qed qed (use gsc in auto) also have "\ = \" using gr(2) unfolding relations_def by blast finally show ?thesis . qed hence otdrg: "ord tau dvd rel g" using tc int_pow_eq_id by force have ot: "ord tau = rel g" proof - from gr(4, 5) have "rel g > 0" by simp with otdrg have "ord tau \ rel g" by (meson zdvd_imp_le) moreover have "\ord tau < rel g" proof assume a: "int (ord tau) < rel g" define T where T: "T = gs - {g} \ {tau}" hence tT: "tau \ T" by blast let ?r = "restrict ((\_.(0::int))(tau := int(ord tau))) T" from T have "T \ all_gens" using gt generate_sincl[of "gs - {g} \ {tau}"] switch_elem_card_le[OF gr(3), of tau] fgs cgs n unfolding all_gens_def by auto moreover have "?r \ relations T" proof(intro in_relationsI finprod_one_eqI) have "tau [^] int (ord tau) = \" using tc pow_ord_eq_1[OF tc] int_pow_int by metis thus "x [^] ?r x = \" if "x \ T" for x using tT that by(cases "\x = tau", auto) qed auto moreover have "?r tau = ord tau" using tT by auto moreover have "ord tau > 0" using dvd_nat_bounds gr(4) gr(5) int_dvd_int_iff otdrg by presburger ultimately have "ord tau \ exps" unfolding exps_def using tT by (auto, force) with le a gr(4) show False by force qed ultimately show ?thesis by auto qed hence otnz: "ord tau \ 0" using gr me exps_def by linarith define l where l: "l = tau#hs" hence ls: "set l = set hs \ {tau}" by auto with hsc tc have slc: "set l \ carrier G" by auto have gAhst: "generate G A = generate G (set hs \ {tau})" proof - have "generate G A = generate G (gs - {g} \ {tau})" using gt by simp also have "\ = generate G (set hs \ {tau})" by (rule generate_subset_change_eqI, use hsc gsc tc ghs in auto) finally show ?thesis . qed have glgA: "generate G (set l) = generate G A" using gAhst ls by simp have lgA: "set l \ generate G A" using ls gt gts hs(1) mono_generate[of "gs - {g}" gs] generate.incl[of tau "gs - {g} \ {tau}"] by fast show ?thesis proof (cases "ord tau = 1") case True hence "tau = \" using ord_eq_1 tc by blast hence "generate G A = generate G (gs - {g})" using gAhst generate_one_irrel hs(3) ghs by auto from iH2[OF ihgsg this cgsg] show ?thesis . next case otau: False consider (nd) "\distinct l" | (ltn) "length l < n \ distinct l" | (dn) "length l = n \ distinct l" proof - have "length l \ n" proof - have "length l = length hs + 1" using l by simp moreover have "length hs \ card (gs - {g})" using hs(2, 5) by (metis distinct_card) moreover have "card (gs - {g}) + 1 \ n" using n cgsg gr(3) fgs Ane i(2) by (simp add: card_gt_0_iff) ultimately show ?thesis by linarith qed thus "\\ distinct l \ thesis; length l < n \ distinct l \ thesis; length l = n \ distinct l \ thesis\ \ thesis" by linarith qed thus ?thesis proof(cases) case nd with hs(2) l have ths: "set hs = set hs \ {tau}" by auto hence "set l = set hs" using l by auto hence "generate G (gs - {g}) = generate G A" using gAhst ths ghs by argo moreover have "card (set hs) \ card A" by (metis Diff_iff card_mono cgs dual_order.trans fgs hs(5) subsetI) ultimately show ?thesis using hs by auto next case ltn then have cl: "card (set l) < card A" using n by (metis distinct_card) from iH[OF this] hsc tc ls have "?t (set l)" by blast thus ?thesis by (subst (1 2) gAhst, use cl ls in fastforce) next case dn hence ln: "length l = n" and dl: "distinct l" by auto have c: "complementary (generate G {tau}) (generate G (gs - {g}))" proof - have "x = \" if "x \ generate G {tau} \ generate G (set hs)" for x proof - from that generate_incl[OF hsc] have xc: "x \ carrier G" by blast from that have xgt: "x \ generate G {tau}" and xgs: "x \ generate G (set hs)" by auto from generate_nat_pow[OF otnz tc] xgt have "\a. a \ 0 \ a < ord tau \ x = tau [^] a" unfolding atLeastAtMost_def atLeast_def atMost_def by (auto, metis Suc_pred less_Suc_eq_le neq0_conv otnz) then obtain a where a: "0 \ a" "a < ord tau" "x = tau [^] a" by blast then have ix: "inv x \ generate G (set hs)" using xgs generate_m_inv_closed ghs hsc by blast with generate_eq_finprod_Pi_int_image[OF _ hsc] obtain f where f: "f \ Pi (set hs) (\_. (UNIV::int set))" "inv x = finprod G (\g. g [^] f g) (set hs)" by blast let ?f = "restrict (f(tau := a)) (set l)" have fr: "?f \ relations (set l)" proof(intro in_relationsI) from ls dl l have sh: "set hs = set l - {tau}" by auto have "finprod G (\a. a [^] ?f a) (set l) = tau [^] ?f tau \ finprod G (\a. a [^] ?f a) (set hs)" by (subst sh, intro finprod_minus, use l slc in auto) moreover have "tau [^] ?f tau = x" using a l int_pow_int by fastforce moreover have "finprod G (\a. a [^] ?f a) (set hs) = finprod G (\g. g [^] f g) (set hs)" by (intro finprod_cong', use slc dl l in auto) ultimately have "finprod G (\a. a [^] ?f a) (set l) = x \ inv x" using f by argo thus "finprod G (\a. a [^] ?f a) (set l) = \" using xc by auto qed blast have "\a > 0" proof assume ag: "0 < a" have "set l \ all_gens" unfolding all_gens_def using glgA lgA dn distinct_card by fastforce moreover have "int a = ?f tau" using l by auto moreover have "tau \ set l" using l by simp ultimately have "a \ exps" using fr ag unfolding exps_def by (auto, force) from le[OF this] a(2) ot gr(4) show False by simp qed hence "a = 0" using a by blast thus "x = \" using tc a by force qed thus ?thesis unfolding complementary_def using generate.one ghs by blast qed moreover have idl: "is_idirprod (generate G A) (\g. generate G {g}) (set l)" proof - have "is_idirprod (generate G (set hs \ {tau})) (\g. generate G {g}) (set hs \ {tau})" by (intro idirprod_generate_ind, use tc hsc hs(3) ghs c in auto) thus ?thesis using ls gAhst by auto qed moreover have "\?t A \ successively (dvd) (map ord l)" proof (cases hs) case Nil thus ?thesis using l by simp next case (Cons a list) hence ac: "a \ carrier G" using hsc by auto assume nA: "\?t A" have "ord tau dvd ord a" proof (rule ccontr) assume nd: "\ ord tau dvd ord a" + then have \0 < ord a mod ord tau\ + using mod_eq_0_iff_dvd by auto have "int (ord tau) > 0" using otnz by simp - with nd obtain r q::int where rq: "ord a = q * (ord tau) + r" "0 < r" "r < ord tau" - using Euclidean_Division.pos_mod_bound div_mult_mod_eq mod_int_pos_iff mod_0_imp_dvd - by (metis linorder_not_le of_nat_le_0_iff of_nat_mod) + obtain r q :: int where rq: "ord a = q * (ord tau) + r" "0 < r" "r < ord tau" + by (rule that [of \ord a div ord tau\ \ord a mod ord tau\]) + (use otnz \0 < ord a mod ord tau\ in \simp_all add: div_mult_mod_eq flip: of_nat_mult of_nat_add\) define b where b: "b = tau \ a [^] q" hence bc: "b \ carrier G" using hsc tc Cons by auto have g: "generate G (set (b#hs)) = generate G (set l)" proof - have se: "set (b#hs) = set l - {tau} \ {b}" using l Cons dl by auto show ?thesis proof(subst se, intro generate_one_switched_eqI[symmetric, of _ tau _ b]) show "b \ generate G (set l)" proof - have "tau \ generate G (set l)" using l generate.incl[of tau "set l"] by auto moreover have "a [^] q \ generate G (set l)" using mono_generate[of "{a}" "set l"] generate_pow[OF ac] Cons l by auto ultimately show ?thesis using b generate.eng by fast qed show "tau \ generate G (set l - {tau} \ {b})" proof - have "tau = b \ inv(a [^] q)" by (simp add: ac b m_assoc tc) moreover have "b \ generate G (set l - {tau} \ {b})" using generate.incl[of b "set l - {tau} \ {b}"] by blast moreover have "inv(a [^] q) \ generate G (set l - {tau} \ {b})" proof - have "generate G {a} \ generate G (set l - {tau} \ {b})" using mono_generate[of "{a}" "set l - {tau} \ {b}"] dl Cons l by auto moreover have "inv(a [^] q) \ generate G {a}" by (subst generate_pow[OF ac], subst int_pow_neg[OF ac, of q, symmetric], blast) ultimately show ?thesis by fast qed ultimately show ?thesis using generate.eng by fast qed qed (use bc tc hsc dl Cons l in auto) qed show False proof (cases "card (set (b#hs)) \ n") case True hence cln: "card (set (b#hs)) < n" using l Cons ln by (metis card_length list.size(4) nat_less_le) hence seq: "set (b#hs) = set hs" proof - from dn l Cons True have "b \ set hs" by (metis distinct.simps(2) distinct_card list.size(4)) thus ?thesis by auto qed with cln have clA: "card (set hs) < card A" using n by auto moreover have "set hs \ generate G (set hs)" using generate_sincl by simp moreover have "distinct hs" by fact moreover have "is_idirprod (generate G (set hs)) (\g. generate G {g}) (set hs)" by (intro is_idirprod_generate, use hs[unfolded is_idirprod_def] hsc in auto) moreover have "generate G A = generate G (set hs)" using glgA g seq by argo moreover have "successively (dvd) (map ord hs)" by fact ultimately show False using iH2 nA by blast next case False hence anb: "a \ b" by (metis card_distinct distinct_length_2_or_more l list.size(4) ln local.Cons) have "nat r \ exps" unfolding exps_def proof(rule) show "set (b#hs) \ all_gens" unfolding all_gens_def using gAhst g ls generate_sincl[of "set (b#hs)"] False by simp let ?r = "restrict ((\_. 0::int)(b := ord tau, a := r)) (set (b#hs))" have "?r \ relations (set (b#hs))" proof(intro in_relationsI) show "finprod G (\x. x [^] ?r x) (set (b#hs)) = \" proof - have "finprod G (\x. x [^] ?r x) (set (b#hs)) = b [^] ?r b \ finprod G (\x. x[^] ?r x) (set (b#hs) - {b})" by (intro finprod_minus, use hsc Cons bc in auto) moreover have "finprod G (\x. x[^] ?r x) (set (b#hs) - {b}) = a [^] ?r a \ finprod G (\x. x[^] ?r x) (set (b#hs) - {b} - {a})" by (intro finprod_minus, use hsc Cons False n anb in auto) moreover have "finprod G (\x. x[^] ?r x) (set (b#hs) - {b} - {a}) = \" by (intro finprod_one_eqI, simp) ultimately have "finprod G (\x. x [^] ?r x) (set (b#hs)) = b [^] ?r b \ (a [^] ?r a \ \)" by argo also have "\ = b [^] ?r b \ a [^] ?r a" using Cons hsc by simp also have "\ = b [^] int(ord tau) \ a [^] r" using anb Cons by simp also have "\ = \" proof - have "b [^] int (ord tau) = tau [^] int (ord tau) \ (a [^] q) [^] int (ord tau)" using b bc hsc int_pow_distrib local.Cons tc by force also have "\ = (a [^] q) [^] int (ord tau)" using trgo hsc local.Cons ot by force finally have "b [^] int (ord tau) \ a [^] r = (a [^] q) [^] int (ord tau) \ a [^] r" by argo also have "\ = a [^] (q * int (ord tau) + r)" using Cons hsc by (metis comm_group_axioms comm_group_def group.int_pow_pow int_pow_mult list.set_intros(1) subsetD) also have "\ = a [^] int (ord a)" using rq by argo finally show ?thesis using Cons hsc int_pow_eq_id by simp qed finally show ?thesis . qed qed simp moreover have "r \ {e \ ?r ` set (b # hs). 0 < e}" proof (rule, rule, rule) show "0 < r" by fact show "a \ set (b#hs)" using Cons by simp thus "r = ?r a" by auto qed ultimately show "nat r \ (\rel\relations (set (b # hs)). nat ` {e \ rel ` set (b # hs). 0 < e})" by fast qed moreover have "nat r < min_exp" using ot rq(2, 3) gr(4) by linarith ultimately show False using le by fastforce qed qed thus ?thesis using hs(4) Cons l by simp qed ultimately show ?thesis using lgA n dn by (metis card_length) qed qed qed (use dvot in blast) qed qed qed lemma (in comm_group) fundamental_subgr: fixes A :: "'a set" assumes "finite A" "A \ carrier G" obtains gs where "set gs \ generate G A" "distinct gs" "is_idirprod (generate G A) (\g. generate G {g}) (set gs)" "successively (dvd) (map ord gs)" "card (set gs) \ card A" using assms ex_idirgen by meson text \As every group is a subgroup of itself, the theorem follows directly. However, for reasons of convenience and uniqueness (although not completely proved), we strengthen the result by proving that the decomposition can be done without having the trivial factor in the product. We formulate the theorem in various ways: firstly, the invariant factor decomposition.\ theorem (in fin_gen_comm_group) invariant_factor_decomposition_idirprod: obtains gs where "set gs \ carrier G" "distinct gs" "is_idirprod (carrier G) (\g. generate G {g}) (set gs)" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" "\ \ set gs" proof - from fundamental_subgr[OF fin_gen gens_closed] obtain gs where gs: "set gs \ carrier G" "distinct gs" "is_idirprod (carrier G) (\g. generate G {g}) (set gs)" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" using generators by auto hence cf: "compl_fam (\g. generate G {g}) (set gs)" by simp let ?r = "remove1 \ gs" have r: "set ?r = set gs - {\}" using gs by auto have "set ?r \ carrier G" using gs by auto moreover have "distinct ?r" using gs by auto moreover have "is_idirprod (carrier G) (\g. generate G {g}) (set ?r)" proof (intro is_idirprod_generate) show "set ?r \ carrier G" using gs by auto show "compl_fam (\g. generate G {g}) (set (remove1 \ gs))" by (rule compl_fam_generate_subset[OF cf gs(1)], use set_remove1_subset in fastforce) show "carrier G = generate G (set ?r)" proof - have "generate G (set ?r) = generate G (set gs)" using generate_one_irrel' r by simp with gs(3) show ?thesis by simp qed qed moreover have "successively (dvd) (map ord ?r)" proof (cases gs) case (Cons a list) have r: "(map ord (remove1 \ gs)) = remove1 1 (map ord gs)" using gs(1) proof(induction gs) case (Cons a gs) hence "a \ carrier G" by simp with Cons ord_eq_1[OF this] show ?case by auto qed simp show ?thesis by (unfold r, rule transp_successively_remove1[OF _ gs(4), unfolded transp_def], auto) qed simp moreover have "card (set ?r) \ card gen" using gs(5) r by (metis List.finite_set card_Diff1_le dual_order.trans) moreover have "\ \ set ?r" using gs(2) by auto ultimately show ?thesis using that by blast qed corollary (in fin_gen_comm_group) invariant_factor_decomposition_dirprod: obtains gs where "set gs \ carrier G" "distinct gs" "DirProds (\g. G\carrier := generate G {g}\) (set gs) \ G" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" "compl_fam (\g. generate G {g}) (set gs)" "\ \ set gs" proof - from invariant_factor_decomposition_idirprod obtain gs where gs: "set gs \ carrier G" "distinct gs" "is_idirprod (carrier G) (\g. generate G {g}) (set gs)" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" "\ \ set gs" by blast with cong_DirProds_IDirProds[OF gs(3)] gs have "DirProds (\g. G\carrier := generate G {g}\) (set gs) \ G" by blast with gs that show ?thesis by auto qed corollary (in fin_gen_comm_group) invariant_factor_decomposition_dirprod_fam: obtains Hs where "\H. H \ set Hs \ subgroup H G" "distinct Hs" "DirProds (\H. G\carrier := H\) (set Hs) \ G" "successively (dvd) (map card Hs)" "card (set Hs) \ card gen" "compl_fam id (set Hs)" "{\} \ set Hs" proof - from invariant_factor_decomposition_dirprod obtain gs where gs: "set gs \ carrier G" "distinct gs" "DirProds (\g. G\carrier := generate G {g}\) (set gs) \ G" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" "compl_fam (\g. generate G {g}) (set gs)" "\ \ set gs" by blast let ?gen = "(\g. generate G {g})" let ?Hs = "map (\g. ?gen g) gs" have "subgroup H G" if "H \ set ?Hs" for H using that gs by (auto intro: generate_is_subgroup) moreover have "distinct ?Hs" using compl_fam_imp_generate_inj[OF gs(1)] gs distinct_map by blast moreover have "DirProds (\H. G\carrier := H\) (set ?Hs) \ G" proof - have gg: "group (G\carrier := ?gen g\)" if "g \ set gs" for g by (use gs that in \auto intro: subgroup.subgroup_is_group generate_is_subgroup\) then interpret og: group "DirProds (\g. G\carrier := ?gen g\) (set gs)" using DirProds_group_iff by blast have "DirProds (\g. G\carrier := ?gen g\) (set gs) \ DirProds (\H. G\carrier := H\) (set ?Hs)" proof (intro DirProds_iso[of ?gen]) show "bij_betw ?gen (set gs) (set ?Hs)" using \distinct ?Hs\ gs(2) compl_fam_imp_generate_inj[OF gs(1, 6)] by (simp add: bij_betw_def) show "G\carrier := ?gen g\ \ G\carrier := ?gen g\" if "g \ set gs" for g by simp show "group (G\carrier := ?gen g\)" if "g \ set gs" for g using that by fact show "Group.group (G\carrier := H\)" if "H \ set ?Hs" for H by (use gs that in \auto intro: subgroup.subgroup_is_group generate_is_subgroup\) qed from group.iso_sym[OF og.is_group this] show ?thesis using gs iso_trans by blast qed moreover have "successively (dvd) (map card ?Hs)" proof - have "card (generate G {g}) = ord g" if "g \ set gs" for g using generate_pow_card that gs(1) by auto hence "map card ?Hs = map ord gs" by simp thus ?thesis using gs(4) by argo qed moreover have "card (set ?Hs) \ card gen" using gs by (metis \distinct ?Hs\ distinct_card length_map) moreover have "compl_fam id (set ?Hs)" using compl_fam_cong[OF _ compl_fam_imp_generate_inj[OF gs(1, 6)], of id] using gs by auto moreover have "{\} \ set ?Hs" using generate_singleton_one gs by auto ultimately show ?thesis using that by blast qed text \Here, the invariant factor decomposition in its classical form.\ corollary (in fin_gen_comm_group) invariant_factor_decomposition_Zn: obtains ns where "DirProds (\n. Z (ns!n)) {.. G" "successively (dvd) ns" "length ns \ card gen" proof - from invariant_factor_decomposition_dirprod obtain gs where gs: "set gs \ carrier G" "distinct gs" "DirProds (\g. G\carrier := generate G {g}\) (set gs) \ G" "successively (dvd) (map ord gs)" "card (set gs) \ card gen" "compl_fam (\g. generate G {g}) (set gs)" "\ \ set gs" by blast let ?DP = "DirProds (\g. G\carrier := generate G {g}\) (set gs)" have "\ns. DirProds (\n. Z (ns!n)) {.. G \ successively (dvd) ns \ length ns \ card gen" proof (cases gs, rule) case Nil from gs(3) Nil have co: "carrier ?DP = {\\<^bsub>?DP\<^esub>}" unfolding DirProds_def by auto let ?ns = "[]" have "DirProds (\n. Z ([] ! n)) {} \ ?DP" proof(intro triv_iso DirProds_is_group) show "carrier (DirProds (\n. Z ([] ! n)) {}) = {\\<^bsub>DirProds (\n. Z ([] ! n)) {}\<^esub>}" using DirProds_empty by blast qed (use co group_integer_mod_group Nil in auto) from that[of ?ns] gs co iso_trans[OF this gs(3)] show "DirProds (\n. Z (?ns ! n)) {.. G \ successively (dvd) ?ns \ length ?ns \ card gen" unfolding lessThan_def by simp next case c: (Cons a list) let ?l = "map ord gs" from c have l: "length ?l > 0" by auto have "DirProds (\n. Z (?l ! n)) {.. G" proof - have "DirProds (\n. Z (?l ! n)) {.. ?DP" proof(intro DirProds_iso[where ?f = "\n. gs!n"]) show "bij_betw ((!) gs) {.. G\carrier := generate G {gs ! i}\" if "i \ {..carrier := generate G {gs ! i}\) = map ord gs ! i" unfolding order_def using that generate_pow_card[of "gs ! i"] gs(1) by force qed (use gs(1) that in auto) show "Group.group (Z (map ord gs ! i))" if "i \ {..carrier := generate G {g}\)" if "g \ set gs" for g using that gs(1) subgroup.subgroup_is_group[OF generate_is_subgroup] by auto qed from iso_trans[OF this gs(3)] show ?thesis . qed moreover have "length ?l \ card gen" using gs by (metis distinct_card length_map) ultimately show ?thesis using gs c by fastforce qed thus ?thesis using that by blast qed text \As every \integer_mod_group\ can be decomposed into a product of prime power groups, we obtain (by using the fact that the direct product does not care about nestedness) the primary decomposition.\ lemma Zn_iso_DirProds_prime_powers: assumes "n \ 0" shows "Z n \ DirProds (\p. Z (p ^ multiplicity p n)) (prime_factors n)" (is "Z n \ ?DP") proof (cases "n = 1") case True show ?thesis by (intro triv_iso[OF group_integer_mod_group DirProds_is_group], use DirProds_empty carrier_integer_mod_group True in auto) next case nno: False interpret DP: group ?DP by (intro DirProds_is_group, use group_integer_mod_group in blast) have "order ?DP = prod (order \ (\p. Z (p ^ multiplicity p n))) (prime_factors n)" by (intro DirProds_order, blast) also have "\ = prod (\p. p ^ multiplicity p n) (prime_factors n)" using Zn_order by simp also have n: "\ = n" using prod_prime_factors[OF assms] by simp finally have oDP: "order ?DP = n" . then interpret DP: finite_group ?DP by (unfold_locales, unfold order_def, metis assms card.infinite) let ?f = "\p\(prime_factors n). 1" have fc: "?f \ carrier ?DP" proof - have p: "0 < multiplicity p n" if "p \ prime_factors n" for p using prime_factors_multiplicity that by auto have pk: "1 < p ^ k" if "Factorial_Ring.prime p" "0 < k" for p k::nat using that one_less_power prime_gt_1_nat by blast show ?thesis unfolding DirProds_def PiE_def by(use carrier_integer_mod_group assms nno pk p in auto, metis in_prime_factors_iff nat_int of_nat_power one_less_nat_eq) qed have of: "DP.ord ?f = n" proof - have "n dvd j" if j: "?f [^]\<^bsub>?DP\<^esub> j = \\<^bsub>?DP\<^esub>" for j proof (intro pairwise_coprime_dvd'[OF _ _ n[symmetric]]) show "finite (prime_factors n)" by simp show "\a\#prime_factorization n. a ^ multiplicity a n dvd j" proof show "p ^ multiplicity p n dvd j" if "p \ prime_factors n" for p proof - from j have "(?f [^]\<^bsub>?DP\<^esub> j) p = 0" using that unfolding DirProds_def one_integer_mod_group by auto hence "?f p [^]\<^bsub>Z (p ^ multiplicity p n)\<^esub> j = 0" using comp_exp_nat[OF that] by metis hence "group.ord (Z (p ^ multiplicity p n)) (?f p) dvd j" using comp_in_carr[OF fc that] by (metis group.pow_eq_id group_integer_mod_group one_integer_mod_group) moreover have "group.ord (Z (p ^ multiplicity p n)) (?f p) = p ^ multiplicity p n" by (metis (no_types, lifting) Zn_neq1_cyclic_group Zn_order comp_in_carr cyclic_group.ord_gen_is_group_order fc integer_mod_group_1 restrict_apply' that) ultimately show ?thesis by simp qed qed show "coprime (i ^ multiplicity i n) (j ^ multiplicity j n)" if "i \# prime_factorization n" "j \# prime_factorization n" "i \ j" for i j using that diff_prime_power_imp_coprime by blast qed thus ?thesis using fc DP.ord_dvd_group_order gcd_nat.asym oDP by force qed interpret DP: cyclic_group ?DP ?f by (intro DP.element_ord_generates_cyclic, use of oDP fc in auto) show ?thesis using DP.iso_sym[OF DP.Zn_iso[OF oDP]] . qed lemma Zn_iso_DirProds_prime_powers': assumes "n \ 0" shows "Z n \ DirProds (\p. Z p) ((\p. p ^ multiplicity p n) ` (prime_factors n))" (is "Z n \ ?DP") proof - have cp: "(\p. Z (p ^ multiplicity p n)) = (\p. Z p) \ (\p. p ^ multiplicity p n)" by auto have "DirProds (\p. Z (p ^ multiplicity p n)) (prime_factors n) \ ?DP" proof(subst cp, intro DirProds_iso2) show "inj_on (\p. p ^ multiplicity p n) (prime_factors n)" by (intro inj_onI; simp add: prime_factors_multiplicity prime_power_inj'') show "group (DirProds Z ((\p. p ^ multiplicity p n) ` prime_factors n))" by (intro DirProds_is_group, use group_integer_mod_group in auto) qed with Zn_iso_DirProds_prime_powers[OF assms] show ?thesis using Group.iso_trans by blast qed corollary (in fin_gen_comm_group) primary_decomposition_Zn: obtains ns where "DirProds (\n. Z (ns!n)) {.. G" "\n\set ns. n = 0 \ (\p k. Factorial_Ring.prime p \ k > 0 \ n = p ^ k)" proof - from invariant_factor_decomposition_Zn obtain ms where ms: "DirProds (\m. Z (ms!m)) {.. G" "successively (dvd) ms" "length ms \ card gen" by blast let ?I = "{..m. Z (ms!m)) {.. DirProds ?f {.. ?f (id i)" if "i \ {..i. i \ {.. group (Z (ms ! i))" by auto show "group (?f j)" if "j \ {.. \ DirProds (\(i,j). ?G i j) (Sigma ?I ?J)" by(rule DirProds_Sigma) finally have G1: "G \ DirProds (\(i,j). ?G i j) (Sigma ?I ?J)" using ms(1) by (metis (no_types, lifting) DirProds_is_group Group.iso_trans group.iso_sym group_integer_mod_group) have "\ps. set ps = Sigma ?I ?J \ distinct ps" by(intro finite_distinct_list, auto) then obtain ps where ps: "set ps = Sigma ?I ?J" "distinct ps" by blast define ns where ns: "ns = map snd ps" have "DirProds (\n. Z (ns!n)) {.. DirProds (\(i,j). ?G i j) (Sigma ?I ?J)" proof - obtain b::"nat \ (nat \ nat)" where b: "\i (case b i of (i, x) \ Z x)" if "i \ {..n. Z (ns!n)) {.. G" using Group.iso_trans iso_sym by blast moreover have "n = 0 \ (\p k. Factorial_Ring.prime p \ k > 0 \ n = p ^ k)" if "n\set ns" for n proof - have "k = 0 \ (\p\prime_factors (ms!i). k = p ^ multiplicity p (ms!i))" if "k \ ?J i" for k i by (cases "ms!i = 0", use that in auto) with that ns ps show ?thesis by (auto, metis (no_types, lifting) mem_Collect_eq neq0_conv prime_factors_multiplicity) qed ultimately show "(\ns. \DirProds (\n. Z (ns ! n)) {.. G; \n\set ns. n = 0 \ (\p k. Factorial_Ring.prime p \ k > 0 \ n = p ^ k)\ \ thesis) \ thesis" by blast qed text \As every finite group is also finitely generated, it follows that a finite group can be decomposed in a product of finite cyclic groups.\ lemma (in finite_comm_group) cyclic_product: obtains ns where "DirProds (\n. Z (ns!n)) {.. G" "\n\set ns. n\0" proof - from primary_decomposition_Zn obtain ns where ns: "DirProds (\n. Z (ns ! n)) {.. G" "\n\set ns. n = 0 \ (\p k. normalization_semidom_class.prime p \ 0 < k \ n = p ^ k)" by blast have "False" if "n \ {..n. Z (ns ! n)) {..n. Z (ns!n)"] Zn_order by auto with fin iso_same_card[OF ns(1)] show False unfolding order_def by auto qed hence "\n\set ns. n\0" by (metis in_set_conv_nth lessThan_iff) with ns show ?thesis using that by blast qed no_notation integer_mod_group ("Z") end diff --git a/thys/Formula_Derivatives/Presburger_Formula.thy b/thys/Formula_Derivatives/Presburger_Formula.thy --- a/thys/Formula_Derivatives/Presburger_Formula.thy +++ b/thys/Formula_Derivatives/Presburger_Formula.thy @@ -1,705 +1,705 @@ section \Concrete Atomic Presburger Formulas\ (*<*) theory Presburger_Formula imports Abstract_Formula "HOL-Library.Code_Target_Int" "List-Index.List_Index" begin (*>*) declare [[coercion "of_bool :: bool \ nat"]] declare [[coercion int]] declare [[coercion_map map]] declare [[coercion_enabled]] fun len :: "nat \ nat" where \ \FIXME yet another logarithm\ "len 0 = 0" | "len (Suc 0) = 1" | "len n = Suc (len (n div 2))" lemma len_eq0_iff: "len n = 0 \ n = 0" by (induct n rule: len.induct) auto lemma len_mult2[simp]: "len (2 * x) = (if x = 0 then 0 else Suc (len x))" proof (induct x rule: len.induct) show "len (2 * Suc 0) = (if Suc 0 = 0 then 0 else Suc (len (Suc 0)))" by (simp add: numeral_eq_Suc) qed auto lemma len_mult2'[simp]: "len (x * 2) = (if x = 0 then 0 else Suc (len x))" using len_mult2 [of x] by (simp add: ac_simps) lemma len_Suc_mult2[simp]: "len (Suc (2 * x)) = Suc (len x)" proof (induct x rule: len.induct) show "len (Suc (2 * Suc 0)) = Suc (len (Suc 0))" by (metis div_less One_nat_def div2_Suc_Suc len.simps(3) lessI mult.right_neutral numeral_2_eq_2) qed auto lemma len_le_iff: "len x \ l \ x < 2 ^ l" proof (induct x arbitrary: l rule: len.induct) fix l show "(len (Suc 0) \ l) = (Suc 0 < 2 ^ l)" proof (cases l) case Suc then show ?thesis using le_less by fastforce qed simp next fix v l assume "\l. (len (Suc (Suc v) div 2) \ l) = (Suc (Suc v) div 2 < 2 ^ l)" then show "(len (Suc (Suc v)) \ l) = (Suc (Suc v) < 2 ^ l)" by (cases l) (simp_all, linarith) qed simp lemma len_pow2[simp]: "len (2 ^ x) = Suc x" by (induct x) auto lemma len_div2[simp]: "len (x div 2) = len x - 1" by (induct x rule: len.induct) auto lemma less_pow2_len[simp]: "x < 2 ^ len x" by (induct x rule: len.induct) auto lemma len_alt: "len x = (LEAST i. x < 2 ^ i)" proof (rule antisym) show "len x \ (LEAST i. x < 2 ^ i)" unfolding len_le_iff by (rule LeastI) (rule less_pow2_len) qed (auto intro: Least_le) lemma len_mono[simp]: "x \ y \ len x \ len y" unfolding len_le_iff using less_pow2_len[of y] by linarith lemma len_div_pow2[simp]: "len (x div 2 ^ m) = len x - m" by (induct m arbitrary: x) (auto simp: div_mult2_eq) lemma len_mult_pow2[simp]: "len (x * 2 ^ m) = (if x = 0 then 0 else len x + m)" by (induct m arbitrary: x) (auto simp: div_mult2_eq mult.assoc[symmetric] mult.commute[of _ 2]) lemma map_index'_Suc[simp]: "map_index' (Suc i) f xs = map_index' i (\i. f (Suc i)) xs" by (induct xs arbitrary: i) auto abbreviation (input) "zero n \ replicate n False" abbreviation (input) "SUC \ \_::unit. Suc" definition "test_bit m n \ (m :: nat) div 2 ^ n mod 2 = 1" lemma test_bit_eq_iff: \test_bit = bit\ by (simp add: fun_eq_iff test_bit_def bit_iff_odd_drop_bit mod_2_eq_odd flip: drop_bit_eq_div) definition "downshift m \ (m :: nat) div 2" definition "upshift m \ (m :: nat) * 2" lemma set_bit_def: "set_bit n m \ m + (if \ test_bit m n then 2 ^ n else (0 :: nat))" apply (rule eq_reflection) apply (rule bit_eqI) apply (subst disjunctive_add) apply (auto simp add: bit_simps test_bit_eq_iff) done definition "cut_bits n m \ (m :: nat) mod 2 ^ n" typedef interp = "{(n :: nat, xs :: nat list). \x \ set xs. len x \ n}" by (force intro: exI[of _ "[]"]) setup_lifting type_definition_interp type_synonym atom = "bool list" type_synonym "value" = "nat" datatype presb = Eq (tm: "int list") (const: int) (offset: "int") derive linorder list derive linorder presb type_synonym formula = "(presb, unit) aformula" lift_definition assigns :: "nat \ interp \ unit \ value" ("_\<^bsup>_\<^esup>_" [900, 999, 999] 999) is "\n (_, I) _. if n < length I then I ! n else 0" . lift_definition nvars :: "interp \ nat" ("#\<^sub>V _" [1000] 900) is "\(_, I). length I" . lift_definition Length :: "interp \ nat" is "\(n, _). n" . lift_definition Extend :: "unit \ nat \ interp \ value \ interp" is "\_ i (n, I) m. (max n (len m), insert_nth i m I)" by (force simp: max_def dest: in_set_takeD in_set_dropD) lift_definition CONS :: "atom \ interp \ interp" is "\bs (n, I). (Suc n, map_index (\i n. 2 * n + (if bs ! i then 1 else 0)) I)" by (auto simp: set_zip) lift_definition SNOC :: "atom \ interp \ interp" is "\bs (n, I). (Suc n, map_index (\i m. m + (if bs ! i then 2 ^ n else 0)) I)" by (auto simp: all_set_conv_all_nth len_le_iff) definition extend :: "unit \ bool \ atom \ atom" where "extend _ b bs \ b # bs" abbreviation (input) size_atom :: "atom \ nat" where "size_atom \ length" definition FV0 :: "unit \ presb \ nat set" where "FV0 _ fm = (case fm of Eq is _ _ \ {n. n < length is \ is!n \ 0})" lemma FV0_code[code]: "FV0 x (Eq is i off) = Option.these (set (map_index (\i x. if x = 0 then None else Some i) is))" unfolding FV0_def by (force simp: Option.these_def image_iff) primrec wf0 :: "nat \ presb \ bool" where "wf0 idx (Eq is _ _) = (length is = idx)" fun find0 where "find0 (_::unit) n (Eq is _ _) = (is ! n \ 0)" primrec decr0 where "decr0 (_::unit) k (Eq is i d) = Eq (take k is @ drop (Suc k) is) i d" definition scalar_product :: "nat list \ int list \ int" where "scalar_product ns is = sum_list (map_index (\i b. (if i < length ns then ns ! i else 0) * b) is)" lift_definition eval_tm :: "interp \ int list \ int" is "\(_, I). scalar_product I" . primrec satisfies0 where "satisfies0 I (Eq is i d) = (eval_tm I is = i - (2 ^ Length I) * d)" inductive lformula0 where "lformula0 (Eq is i 0)" code_pred lformula0 . fun lderiv0 :: "bool list \ presb \ formula" where "lderiv0 bs (Eq is i d) = (if d \ 0 then undefined else (let v = i - scalar_product bs is in if v mod 2 = 0 then FBase (Eq is (v div 2) 0) else FBool False))" fun rderiv0 :: "bool list \ presb \ formula" where "rderiv0 bs (Eq is i d) = (let l = - sum_list [i. i \ is, i < 0]; h = - sum_list [i. i \ is, i > 0]; d' = scalar_product bs is + 2 * d in if d' \ {min h i .. max l i} then FBase (Eq is i d') else FBool False)" primrec nullable0 where "nullable0 (Eq is i off) = (i = off)" definition \ :: "nat \ atom list" where "\ n = List.n_lists n [True, False]" named_theorems Presb_simps lemma nvars_Extend[Presb_simps]: "#\<^sub>V (Extend () i \ P) = Suc (#\<^sub>V \)" by (transfer, auto) lemma Length_Extend[Presb_simps]: "Length (Extend () i \ P) = max (Length \) (len P)" by (transfer, auto) lemma Length0_inj[Presb_simps]: "Length \ = 0 \ Length \ = 0 \ #\<^sub>V \ = #\<^sub>V \ \ \ = \" by transfer (auto intro: nth_equalityI simp: all_set_conv_all_nth len_eq0_iff) lemma ex_Length0[Presb_simps]: "\\. Length \ = 0 \ #\<^sub>V \ = idx" by (transfer fixing: idx) (auto intro: exI[of _ "replicate idx 0"]) lemma Extend_commute_safe[Presb_simps]: "\j \ i; i < Suc (#\<^sub>V \)\ \ Extend k j (Extend k i \ P) Q = Extend k (Suc i) (Extend k j \ Q) P" by transfer (auto simp add: min_def take_Cons take_drop le_imp_diff_is_add split: nat.splits) lemma Extend_commute_unsafe[Presb_simps]: "k \ k' \ Extend k j (Extend k' i \ P) Q = Extend k' i (Extend k j \ Q) P" by transfer auto lemma assigns_Extend[Presb_simps]: "i < Suc (#\<^sub>V \) \ m\<^bsup>Extend k i \ P\<^esup>k' = (if k = k' then if m = i then P else dec i m\<^bsup>\\<^esup>k else m\<^bsup>\\<^esup>k')" by transfer (auto simp: nth_append dec_def min_def) lemma assigns_SNOC_zero[Presb_simps]: "m < #\<^sub>V \ \ m\<^bsup>SNOC (zero (#\<^sub>V \)) \\<^esup>k = m\<^bsup>\\<^esup>k" by transfer auto lemma Length_CONS[Presb_simps]: "Length (CONS x \) = Suc (Length \)" by transfer auto lemma Length_SNOC[Presb_simps]: "Length (SNOC x \) = Suc (Length \)" by transfer auto lemma nvars_CONS[Presb_simps]: "#\<^sub>V (CONS x \) = #\<^sub>V \" by transfer auto lemma nvars_SNOC[Presb_simps]: "#\<^sub>V (SNOC x \) = #\<^sub>V \" by transfer auto lemma Extend_CONS[Presb_simps]: "#\<^sub>V \ = length x \ Extend k 0 (CONS x \) P = CONS (extend k (test_bit P 0) x) (Extend k 0 \ (downshift P))" by transfer (auto simp: extend_def downshift_def test_bit_def, presburger+) lemma Extend_SNOC[Presb_simps]: "\#\<^sub>V \ = length x; len P \ Length (SNOC x \)\ \ Extend k 0 (SNOC x \) P = SNOC (extend k (test_bit P (Length \)) x) (Extend k 0 \ (cut_bits (Length \) P))" apply transfer apply (auto simp: cut_bits_def extend_def test_bit_def nth_Cons' max_absorb1 len_le_iff split: if_splits cong del: if_weak_cong) - apply (metis add.commute mod_less mod_mult2_eq mult_numeral_1_right numeral_1_eq_Suc_0 power_commuting_commutes) - apply (metis Euclidean_Division.div_eq_0_iff div_0 less_mult_imp_div_less mod_less nat_dvd_not_less semiring_normalization_rules(7)) + apply (metis add.commute div_mult_mod_eq less_mult_imp_div_less mod_less mult_numeral_1 numeral_1_eq_Suc_0) + apply (metis div_eq_0_iff less_mult_imp_div_less mod2_eq_if mod_less power_not_zero zero_neq_numeral) done lemma odd_neq_even: "Suc (2 * x) = 2 * y \ False" "2 * y = Suc (2 * x) \ False" by presburger+ lemma CONS_inj[Presb_simps]: "size x = #\<^sub>V \ \ size y = #\<^sub>V \ \ #\<^sub>V \ = #\<^sub>V \ \ CONS x \ = CONS y \ \ (x = y \ \ = \)" by transfer (auto simp: list_eq_iff_nth_eq odd_neq_even split: if_splits) lemma mod_2_Suc_iff: "x mod 2 = Suc 0 \ x = Suc (2 * (x div 2))" by presburger+ lemma CONS_surj[Presb_simps]: "Length \ \ 0 \ \x \. \ = CONS x \ \ #\<^sub>V \ = #\<^sub>V \ \ size x = #\<^sub>V \" by transfer (auto simp: gr0_conv_Suc list_eq_iff_nth_eq len_le_iff split: if_splits intro!: exI[of _ "map (\n. n mod 2 \ 0) _"] exI[of _ "map (\n. n div 2) _"]; auto simp: mod_2_Suc_iff) lemma [Presb_simps]: "length (extend k b x) = Suc (length x)" "downshift (upshift P) = P" "downshift (set_bit 0 P) = downshift P" "test_bit (set_bit n P) n" "\ test_bit (upshift P) 0" "len P \ p \ \ test_bit P p" "len (cut_bits n P) \ n" "len P \ n \ cut_bits n P = P" "len (upshift P) = (case len P of 0 \ 0 | Suc n \ Suc (Suc n))" "len (downshift P) = (case len P of 0 \ 0 | Suc n \ n)" by (auto simp: extend_def set_bit_def cut_bits_def upshift_def downshift_def test_bit_def len_le_iff len_eq0_iff div_add_self2 split: nat.split) lemma Suc0_div_pow2_eq: "Suc 0 div 2 ^ i = (if i = 0 then 1 else 0)" by (induct i) (auto simp: div_mult2_eq) lemma set_unset_bit_preserves_len: assumes "x div 2 ^ m = 2 * q" "m < len x" shows "x + 2 ^ m < 2 ^ len x" using assms proof (induct m arbitrary: x) case 0 then show ?case by (auto simp: div_mult2_eq len_Suc_mult2[symmetric] simp del: len_Suc_mult2 power_Suc split: if_splits) next case (Suc m) with Suc(1)[of "x div 2"] show ?case by (cases "len x") (auto simp: div_mult2_eq) qed lemma len_set_bit[Presb_simps]: "len (set_bit m P) = max (Suc m) (len P)" proof (rule antisym) show "len (set_bit m P) \ max (Suc m) (len P)" by (auto simp: set_bit_def test_bit_def max_def Suc_le_eq not_less len_le_iff set_unset_bit_preserves_len simp del: One_nat_def) next have "P < 2 ^ len (P + 2 ^ m)" by (rule order.strict_trans2[OF less_pow2_len]) auto moreover have "m < len (P + 2 ^ m)" by (rule order.strict_trans2[OF _ len_mono[of "2 ^ m"]]) auto ultimately show "max (Suc m) (len P) \ len (set_bit m P)" by (auto simp: set_bit_def test_bit_def max_def Suc_le_eq not_less len_le_iff) qed lemma mod_pow2_div_pow2: fixes p m n :: nat shows "m < n \ p mod 2 ^ n div 2 ^ m = p div 2 ^ m mod 2 ^ (n - m)" by (induct m arbitrary: p n) (auto simp: div_mult2_eq mod_mult2_eq Suc_less_eq2) lemma irrelevant_set_bit[simp]: fixes p m n :: nat assumes "n \ m" shows "(p + 2 ^ m) mod 2 ^ n = p mod 2 ^ n" proof - from assms obtain q :: nat where "2 ^ m = q * 2 ^ n" by (metis le_add_diff_inverse mult.commute power_add) then show ?thesis by simp qed lemma mod_lemma: "\ (0::nat) < c; r < b \ \ b * (q mod c) + r < b * c" by (metis add_gr_0 div_le_mono div_mult_self1_is_m less_imp_add_positive mod_less_divisor not_less split_div) lemma relevant_set_bit[simp]: fixes p m n :: nat assumes "m < n" "p div 2 ^ m = 2 * q" shows "(p + 2 ^ m) mod 2 ^ n = p mod 2 ^ n + 2 ^ m" proof - have "p mod 2 ^ n + 2 ^ m < 2 ^ n" using assms proof (induct m arbitrary: p n) case 0 then show ?case by (auto simp: gr0_conv_Suc) (metis One_nat_def Suc_eq_plus1 lessI mod_lemma numeral_2_eq_2 zero_less_numeral zero_less_power) next case (Suc m) from Suc(1)[of "n - 1" "p div 2"] Suc(2,3) show ?case by (auto simp: div_mult2_eq mod_mult2_eq Suc_less_eq2) qed with \m < n\ show ?thesis by (subst mod_add_eq [symmetric]) auto qed lemma cut_bits_set_bit[Presb_simps]: "cut_bits n (set_bit m p) = (if n \ m then cut_bits n p else set_bit m (cut_bits n p))" unfolding cut_bits_def set_bit_def test_bit_def by (auto simp: not_le mod_pow2_div_pow2 mod_mod_cancel simp del: One_nat_def) lemma wf0_decr0[Presb_simps]: "wf0 (Suc idx) a \ l < Suc idx \ \ find0 k l a \ wf0 idx (decr0 k l a)" by (induct a) auto lemma lformula0_decr0[Presb_simps]: "lformula0 a \ lformula0 (decr0 k l a)" by (induct a) (auto elim: lformula0.cases intro: lformula0.intros) abbreviation sat0_syn (infix "\0" 65) where "sat0_syn \ satisfies0" abbreviation sat_syn (infix "\" 65) where "sat_syn \ Formula_Operations.satisfies Extend Length satisfies0" abbreviation sat_bounded_syn (infix "\\<^sub>b" 65) where "sat_bounded_syn \ Formula_Operations.satisfies_bounded Extend Length len satisfies0" lemma scalar_product_Nil[simp]: "scalar_product [] xs = 0" by (induct xs) (auto simp: scalar_product_def) lemma scalar_product_Nil2[simp]: "scalar_product xs [] = 0" by (induct xs) (auto simp: scalar_product_def) lemma scalar_product_Cons[simp]: "scalar_product xs (y # ys) = (case xs of x # xs \ x * y + scalar_product xs ys | [] \ 0)" by (cases xs) (simp, auto simp: scalar_product_def cong del: if_weak_cong) lemma scalar_product_append[simp]: "scalar_product ns (xs @ ys) = scalar_product (take (length xs) ns) xs + scalar_product (drop (length xs) ns) ys" by (induct xs arbitrary: ns) (auto split: list.splits) lemma scalar_product_trim: "scalar_product ns xs = scalar_product (take (length xs) ns) xs" by (induct xs arbitrary: ns) (auto split: list.splits) lemma Extend_satisfies0_decr0[Presb_simps]: assumes "\ find0 k i a" "i < Suc (#\<^sub>V \)" "lformula0 a \ len P \ Length \" shows "Extend k i \ P \0 a = \ \0 decr0 k i a" proof - { fix "is" :: "int list" assume "is ! i = 0" with assms(1,2) have "eval_tm (Extend k i \ P) is = eval_tm \ (take i is @ drop (Suc i) is)" by (cases a, transfer) (force intro: trans[OF scalar_product_trim] simp: min_def arg_cong2[OF refl id_take_nth_drop, of i _ scalar_product "take i xs @ _" for i x xs]) } note * = this from assms show ?thesis by (cases a) (auto dest!: * simp: Length_Extend max_def elim: lformula0.cases) qed lemma scalar_product_eq0: "\c\set ns. c = 0 \ scalar_product ns is = 0" proof (induct "is" arbitrary: "ns") case Cons then show ?case by (cases ns) (auto simp: scalar_product_def cong del: if_weak_cong) qed (simp add: scalar_product_def) lemma nullable0_satisfies0[Presb_simps]: "Length \ = 0 \ nullable0 a = \ \0 a" proof (induct a) case Eq then show ?case unfolding nullable0.simps satisfies0.simps by transfer (auto simp: len_eq0_iff scalar_product_eq0) qed lemma satisfies0_cong: "wf0 (#\<^sub>V \) a \ #\<^sub>V \ = #\<^sub>V \ \ lformula0 a \ (\m k. m < #\<^sub>V \ \ m\<^bsup>\\<^esup>k = m\<^bsup>\\<^esup>k) \ \ \0 a = \ \0 a" proof (induct a) case Eq then show ?case unfolding satisfies0.simps by transfer (auto simp: scalar_product_def intro!: arg_cong[of _ _ sum_list] map_index_cong elim!: lformula0.cases) qed lemma wf_lderiv0[Presb_simps]: "wf0 idx a \ lformula0 a \ Formula_Operations.wf (\_. Suc) wf0 idx (lderiv0 x a)" by (induct a) (auto elim: lformula0.cases simp: Formula_Operations.wf.simps Let_def) lemma lformula_lderiv0[Presb_simps]: "lformula0 a \ Formula_Operations.lformula lformula0 (lderiv0 x a)" by (induct a) (auto elim: lformula0.cases intro: lformula0.intros simp: Let_def Formula_Operations.lformula.simps) lemma wf_rderiv0[Presb_simps]: "wf0 idx a \ Formula_Operations.wf (\_. Suc) wf0 idx (rderiv0 x a)" by (induct a) (auto elim: lformula0.cases simp: Formula_Operations.wf.simps Let_def) lemma find0_FV0[Presb_simps]: "\wf0 idx a; l < idx\ \ find0 k l a = (l \ FV0 k a)" by (induct a) (auto simp: FV0_def) lemma FV0_less[Presb_simps]: "wf0 idx a \ v \ FV0 k a \ v < idx" by (induct a) (auto simp: FV0_def) lemma finite_FV0[Presb_simps]: "finite (FV0 k a)" by (induct a) (auto simp: FV0_def) lemma finite_lderiv0[Presb_simps]: assumes "lformula0 a" shows "finite {\. \xs. \ = fold (Formula_Operations.deriv extend lderiv0) xs (FBase a)}" proof - define d where "d = Formula_Operations.deriv extend lderiv0" define l where "l is = sum_list [i. i \ is, i < 0]" for "is" :: "int list" define h where "h is = sum_list [i. i \ is, i > 0]" for "is" :: "int list" define \ where "\ a = (case a of Eq is n z \ {FBase (Eq is i 0) | i . i \ {min (- h is) n .. max (- l is) n}} \ {FBool False :: formula})" for a { fix xs note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \_def[simp] from \lformula0 a\ have "FBase a \ \ a" by (auto simp: elim!: lformula0.cases) moreover have "\x \. \ \ \ a \ d x \ \ \ a" proof (induct a, unfold \_def presb.case, elim UnE CollectE insertE emptyE exE conjE) fix "is" :: "int list" and bs :: "bool list" and i n :: int and \ :: formula assume "i \ {min (- h is) n..max (- l is) n}" "\ = FBase (presb.Eq is i 0)" moreover have "scalar_product bs is \ h is" proof (induct "is" arbitrary: bs) case (Cons x xs) from Cons[of "tl bs"] show ?case by (cases bs) (auto simp: h_def) qed (auto simp: h_def scalar_product_def) moreover have "l is \ scalar_product bs is" proof (induct "is" arbitrary: bs) case (Cons x xs) from Cons[of "tl bs"] show ?case by (cases bs) (auto simp: l_def) qed (auto simp: l_def scalar_product_def) ultimately show "d bs \ \ {FBase (presb.Eq is i 0) |i. i \ {min (- h is) n..max (- l is) n}} \ {FBool False}" by (auto simp: d_def Let_def) qed (auto simp: d_def) then have "\\. \ \ \ a \ fold d xs \ \ \ a" by (induct xs) auto ultimately have "fold d xs (FBase a) \ \ a" by blast } moreover have "finite (\ a)" unfolding \_def by (auto split: presb.splits) ultimately show "?thesis" unfolding d_def by (blast intro: finite_subset) qed lemma finite_rderiv0[Presb_simps]: "finite {\. \xs. \ = fold (Formula_Operations.deriv extend rderiv0) xs (FBase a)}" proof - define d where "d = Formula_Operations.deriv extend rderiv0" define l where "l is = sum_list [i. i \ is, i < 0]"for "is" :: "int list" define h where "h is = sum_list [i. i \ is, i > 0]"for "is" :: "int list" define \ where "\ a = (case a of Eq is n z \ {FBase (Eq is n i) | i . i \ {min (- h is) (min n z) .. max (- l is) (max n z)}} \ {FBool False :: formula})" for a { fix xs note Formula_Operations.fold_deriv_FBool[simp] Formula_Operations.deriv.simps[simp] \_def[simp] have "FBase a \ \ a" by (auto split: presb.splits) moreover have "\x \. \ \ \ a \ d x \ \ \ a" proof (induct a, unfold \_def presb.case, elim UnE CollectE insertE emptyE exE conjE) fix "is" :: "int list" and bs :: "bool list" and i n m :: int and \ :: formula assume "i \ {min (- h is) (min n m)..max (- l is) (max n m)}" "\ = FBase (presb.Eq is n i)" moreover have "scalar_product bs is \ h is" proof (induct "is" arbitrary: bs) case (Cons x xs) from Cons[of "tl bs"] show ?case by (cases bs) (auto simp: h_def) qed (auto simp: h_def scalar_product_def) moreover have "l is \ scalar_product bs is" proof (induct "is" arbitrary: bs) case (Cons x xs) from Cons[of "tl bs"] show ?case by (cases bs) (auto simp: l_def) qed (auto simp: l_def scalar_product_def) ultimately show "d bs \ \ {FBase (presb.Eq is n i) |i. i \ {min (- h is) (min n m)..max (- l is) (max n m)}} \ {FBool False}" by (auto 0 1 simp: d_def Let_def h_def l_def) qed (auto simp: d_def) then have "\\. \ \ \ a \ fold d xs \ \ \ a" by (induct xs) auto ultimately have "fold d xs (FBase a) \ \ a" by blast } moreover have "finite (\ a)" unfolding \_def by (auto split: presb.splits) ultimately show "?thesis" unfolding d_def by (blast intro: finite_subset) qed lemma scalar_product_CONS: "length xs = length (bs :: bool list) \ scalar_product (map_index (\i n. 2 * n + bs ! i) xs) is = scalar_product bs is + 2 * scalar_product xs is" by (induct "is" arbitrary: bs xs) (auto split: list.splits simp: algebra_simps) lemma eval_tm_CONS[simp]: "\length is \ #\<^sub>V \; #\<^sub>V \ = length x\ \ eval_tm (CONS x \) is = scalar_product x is + 2 * eval_tm \ is" by transfer (auto simp: scalar_product_CONS[symmetric] intro!: arg_cong2[of _ _ _ _ scalar_product] nth_equalityI) lemma satisfies_lderiv0[Presb_simps]: "\wf0 (#\<^sub>V \) a; #\<^sub>V \ = length x; lformula0 a\ \ \ \ lderiv0 x a \ CONS x \ \0 a" by (auto simp: Let_def Formula_Operations.satisfies_gen.simps split: if_splits elim!: lformula0.cases) lemma satisfies_bounded_lderiv0[Presb_simps]: "\wf0 (#\<^sub>V \) a; #\<^sub>V \ = length x; lformula0 a\ \ \ \\<^sub>b lderiv0 x a \ CONS x \ \0 a" by (auto simp: Let_def Formula_Operations.satisfies_gen.simps split: if_splits elim!: lformula0.cases) lemma scalar_product_SNOC: "length xs = length (bs :: bool list) \ scalar_product (map_index (\i m. m + 2 ^ a * bs ! i) xs) is = scalar_product xs is + 2 ^ a * scalar_product bs is" by (induct "is" arbitrary: bs xs) (auto split: list.splits simp: algebra_simps) lemma eval_tm_SNOC[simp]: "\length is \ #\<^sub>V \; #\<^sub>V \ = length x\ \ eval_tm (SNOC x \) is = eval_tm \ is + 2 ^ Length \ * scalar_product x is" by transfer (auto simp: scalar_product_SNOC[symmetric] intro!: arg_cong2[of _ _ _ _ scalar_product] nth_equalityI) lemma Length_eq0_eval_tm_eq0[simp]: "Length \ = 0 \ eval_tm \ is = 0" by transfer (auto simp: len_eq0_iff scalar_product_eq0) lemma less_pow2: "x < 2 ^ a \ int x < 2 ^ a" by (metis of_nat_less_iff of_nat_numeral of_nat_power [symmetric]) lemma scalar_product_upper_bound: "\x\set b. len x \ a \ scalar_product b is \ (2 ^ a - 1) * sum_list [i. i \ is, i > 0]" proof (induct "is" arbitrary: b) case (Cons i "is") then have "scalar_product (tl b) is \ (2 ^ a - 1) * sum_list [i. i \ is, i > 0]" by (auto simp: in_set_tlD) with Cons(2) show ?case by (auto 0 3 split: list.splits simp: len_le_iff mult_le_0_iff distrib_left add.commute[of _ "(2 ^ a - 1) * i"] less_pow2 intro: add_mono elim: order_trans[OF add_mono[OF order_refl]]) qed simp lemma scalar_product_lower_bound: "\x\set b. len x \ a \ scalar_product b is \ (2 ^ a - 1) * sum_list [i. i \ is, i < 0]" proof (induct "is" arbitrary: b) case (Cons i "is") then have "scalar_product (tl b) is \ (2 ^ a - 1) * sum_list [i. i \ is, i < 0]" by (auto simp: in_set_tlD) with Cons(2) show ?case by (auto 0 3 split: list.splits simp: len_le_iff mult_le_0_iff distrib_left add.commute[of _ "(2 ^ a - 1) * i"] less_pow2 intro: add_mono elim: order_trans[OF add_mono[OF order_refl]] order_trans) qed simp lemma eval_tm_upper_bound: "eval_tm \ is \ (2 ^ Length \ - 1) * sum_list [i. i \ is, i > 0]" by transfer (auto simp: scalar_product_upper_bound) lemma eval_tm_lower_bound: "eval_tm \ is \ (2 ^ Length \ - 1) * sum_list [i. i \ is, i < 0]" by transfer (auto simp: scalar_product_lower_bound) lemma satisfies_bounded_rderiv0[Presb_simps]: "\wf0 (#\<^sub>V \) a; #\<^sub>V \ = length x\ \ \ \\<^sub>b rderiv0 x a \ SNOC x \ \0 a" proof (induct a) case (Eq "is" n d) let ?l = "Length \" define d' where "d' = scalar_product x is + 2 * d" define l where "l = sum_list [i. i \ is, i < 0]" define h where "h = sum_list [i. i \ is, i > 0]" from Eq show ?case unfolding wf0.simps satisfies0.simps rderiv0.simps Let_def proof (split if_splits, simp only: Formula_Operations.satisfies_gen.simps satisfies0.simps Length_SNOC eval_tm_SNOC simp_thms(13) de_Morgan_conj not_le min_less_iff_conj max_less_iff_conj d'_def[symmetric] h_def[symmetric] l_def[symmetric], safe) assume "eval_tm \ is + 2 ^ ?l * scalar_product x is = n - 2 ^ Suc ?l * d" with eval_tm_upper_bound[of \ "is"] eval_tm_lower_bound[of \ "is"] have *: "n + h \ 2 ^ ?l * (h + d')" "2 ^ ?l * (l + d') \ n + l" by (auto simp: algebra_simps h_def l_def d'_def) assume **: "d' \ {min (- h) n..max (- l) n}" { assume "0 \ n + h" from order_trans[OF this *(1)] have "0 \ h + d'" unfolding zero_le_mult_iff using zero_less_power[of 2] by presburger } moreover { assume "n + h < 0" with *(1) have "n \ d'" by (auto dest: order_trans[OF _ mult_right_mono_neg[of 1]]) } moreover { assume "n + l < 0" from le_less_trans[OF *(2) this] have "l + d' < 0" unfolding mult_less_0_iff by auto } moreover { assume "0 \ n + l" then have "0 \ l + d'" using ** calculation(1-2) by force with *(2) have "d' \ n" by (force dest: order_trans[OF mult_right_mono[of 1], rotated]) } ultimately have "min (- h) n \ d'" "d' \ max (- l) n" by (auto simp: min_def max_def) with ** show False by auto qed (auto simp: algebra_simps d'_def) qed declare [[goals_limit = 100]] global_interpretation Presb: Formula where SUC = SUC and LESS = "\_. (<)" and Length = Length and assigns = assigns and nvars = nvars and Extend = Extend and CONS = CONS and SNOC = SNOC and extend = extend and size = size_atom and zero = zero and alphabet = \ and eval = test_bit and downshift = downshift and upshift = upshift and add = set_bit and cut = cut_bits and len = len and restrict = "\_ _. True" and Restrict = "\_ _. FBool True" and lformula0 = lformula0 and FV0 = FV0 and find0 = find0 and wf0 = wf0 and decr0 = decr0 and satisfies0 = satisfies0 and nullable0 = nullable0 and lderiv0 = lderiv0 and rderiv0 = rderiv0 and TYPEVARS = undefined defines norm = "Formula_Operations.norm find0 decr0" and nFOr = "Formula_Operations.nFOr :: formula \ _" and nFAnd = "Formula_Operations.nFAnd :: formula \ _" and nFNot = "Formula_Operations.nFNot find0 decr0 :: formula \ _" and nFEx = "Formula_Operations.nFEx find0 decr0" and nFAll = "Formula_Operations.nFAll find0 decr0" and decr = "Formula_Operations.decr decr0 :: _ \ _ \ formula \ _" and find = "Formula_Operations.find find0 :: _ \ _ \ formula \ _" and FV = "Formula_Operations.FV FV0" and RESTR = "Formula_Operations.RESTR (\_ _. FBool True) :: _ \ formula" and RESTRICT = "Formula_Operations.RESTRICT (\_ _. FBool True) FV0" and deriv = "\d0 (a :: atom) (\ :: formula). Formula_Operations.deriv extend d0 a \" and nullable = "\\ :: formula. Formula_Operations.nullable nullable0 \" and fut_default = "Formula.fut_default extend zero rderiv0" and fut = "Formula.fut extend zero find0 decr0 rderiv0" and finalize = "Formula.finalize SUC extend zero find0 decr0 rderiv0" and final = "Formula.final SUC extend zero find0 decr0 nullable0 rderiv0 :: nat \ formula \ _" and presb_wf = "Formula_Operations.wf SUC (wf0 :: nat \ presb \ _)" and presb_lformula = "Formula_Operations.lformula (lformula0 :: presb \ _) :: formula \ _" and check_eqv = "\idx. DAs.check_eqv (\ idx) (\\. norm (RESTRICT \) :: formula) (\a \. norm (deriv (lderiv0 :: _ \ _ \ formula) (a :: atom) \)) (final idx) (\\ :: formula. presb_wf idx \ \ presb_lformula \) (\ idx) (\\. norm (RESTRICT \) :: formula) (\a \. norm (deriv (lderiv0 :: _ \ _ \ formula) (a :: atom) \)) (final idx) (\\ :: formula. presb_wf idx \ \ presb_lformula \) (=)" and bounded_check_eqv = "\idx. DAs.check_eqv (\ idx) (\\. norm (RESTRICT \) :: formula) (\a \. norm (deriv (lderiv0 :: _ \ _ \ formula) (a :: atom) \)) nullable (\\ :: formula. presb_wf idx \ \ presb_lformula \) (\ idx) (\\. norm (RESTRICT \) :: formula) (\a \. norm (deriv (lderiv0 :: _ \ _ \ formula) (a :: atom) \)) nullable (\\ :: formula. presb_wf idx \ \ presb_lformula \) (=)" and automaton = "DA.automaton (\a \. norm (deriv lderiv0 (a :: atom) \ :: formula))" apply standard apply (auto simp: Presb_simps \_def set_n_lists distinct_n_lists Formula_Operations.lformula.simps Formula_Operations.satisfies_gen.simps Formula_Operations.wf.simps dest: satisfies0_cong split: presb.splits if_splits) apply (simp add: downshift_def) done (*Workaround for code generation*) lemma check_eqv_code[code]: "check_eqv idx r s = ((presb_wf idx r \ presb_lformula r) \ (presb_wf idx s \ presb_lformula s) \ (case rtrancl_while (\(p, q). final idx p = final idx q) (\(p, q). map (\a. (norm (deriv lderiv0 a p), norm (deriv lderiv0 a q))) (\ idx)) (norm (RESTRICT r), norm (RESTRICT s)) of None \ False | Some ([], x) \ True | Some (a # list, x) \ False))" unfolding check_eqv_def Presb.check_eqv_def Presb.step_alt .. definition while where [code del, code_abbrev]: "while idx \ = while_default (fut_default idx \)" declare while_default_code[of "fut_default idx \" for idx \, folded while_def, code] lemma check_eqv_sound: "\#\<^sub>V \ = idx; check_eqv idx \ \\ \ (Presb.sat \ \ \ Presb.sat \ \)" unfolding check_eqv_def by (rule Presb.check_eqv_soundness) lemma bounded_check_eqv_sound: "\#\<^sub>V \ = idx; bounded_check_eqv idx \ \\ \ (Presb.sat\<^sub>b \ \ \ Presb.sat\<^sub>b \ \)" unfolding bounded_check_eqv_def by (rule Presb.bounded_check_eqv_soundness) method_setup check_equiv = \ let fun tac ctxt = let val conv = @{computation_check terms: Trueprop "0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc "plus :: nat \ _" "minus :: nat \ _" "times :: nat \ _" "divide :: nat \ _" "modulo :: nat \ _" "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" check_eqv datatypes: formula "int list" integer bool} ctxt in CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 conv)) ctxt) THEN' resolve_tac ctxt [TrueI] end in Scan.succeed (SIMPLE_METHOD' o tac) end \ end diff --git a/thys/IMP2/doc/Examples.thy b/thys/IMP2/doc/Examples.thy --- a/thys/IMP2/doc/Examples.thy +++ b/thys/IMP2/doc/Examples.thy @@ -1,1873 +1,1873 @@ theory Examples imports "../IMP2" "../lib/IMP2_Aux_Lemmas" begin section \Examples\ lemmas nat_distribs = nat_add_distrib nat_diff_distrib Suc_diff_le nat_mult_distrib nat_div_distrib subsection \Common Loop Patterns\ subsubsection \Count Up\ text \ Counter \c\ counts from \0\ to \n\, such that loop is executed \n\ times. The result is computed in an accumulator \a\. \ text \The invariant states that we have computed the function for the counter value \c\\ text \The variant is the difference between \n\ and \c\, i.e., the number of loop iterations that we still have to do\ program_spec exp_count_up assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; c = 0; while (cn-c\ @invariant \0\c \ c\n \ a=2^nat c\ { G_par = a; scope { a = G_par; a=2*a; G_ret = a }; a = G_ret; c=c+1 } \ apply vcg by (auto simp: algebra_simps nat_distribs) program_spec sum_prog assumes "n \ 0" ensures "s = \{0..n\<^sub>0}" defines \ s = 0; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ s = \{0..i}\ { i = i + 1; s = s + i } \ apply vcg_cs by (simp_all add: intvs_incdec) program_spec sq_prog assumes "n \ 0" ensures "a = n\<^sub>0 * n\<^sub>0" defines \ a = 0; z = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n\<^sub>0 = n \ 0 \ i \ i \ n \ a = i * i \ z = 2 * i + 1\ { a = a + z; z = z + 2; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) fun factorial :: "int \ int" where "factorial i = (if i \ 0 then 1 else i * factorial (i - 1))" program_spec factorial_prog assumes "n \ 0" ensures "a = factorial n\<^sub>0" defines \ a = 1; i = 1; while (i \ n) @variant \n\<^sub>0 + 1 - i\ @invariant \n\<^sub>0 = n \ 1 \ i \ i \ n + 1 \ a = factorial (i - 1)\ { a = a * i; i = i + 1 } \ by vcg (simp add: antisym_conv)+ fun fib :: "int \ int" where "fib i = (if i \ 0 then 0 else if i = 1 then 1 else fib (i - 2) + fib (i - 1))" lemma fib_simps[simp]: "i \ 0 \ fib i = 0" "i = 1 \ fib i = 1" "i > 1 \ fib i = fib (i - 2) + fib (i - 1)" by simp+ lemmas [simp del] = fib.simps text \With precondition\ program_spec fib_prog assumes "n \ 0" ensures "a = fib n" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ 0 \ i \ i \ n \ a = fib i \ b = fib (i + 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (simp add: algebra_simps) text \Without precondition, returning \0\ for negative numbers\ program_spec fib_prog' assumes True ensures "a = fib n\<^sub>0" defines \ a = 0; b = 1; i = 0; while (i < n) @variant \n\<^sub>0 - i\ @invariant \n = n\<^sub>0 \ (0 \ i \ i \ n \ n\<^sub>0 < 0 \ i = 0) \ a = fib i \ b = fib (i+ 1)\ { c = b; b = a + b; a = c; i = i + 1 } \ by vcg_cs (auto simp: algebra_simps) subsubsection \Count down\ text \Essentially the same as count up, but we (ab)use the input variable as a counter.\ text \The invariant is the same as for count-up. Only that we have to compute the actual number of loop iterations by \n\<^sub>0 - n\. We locally introduce the name \c\ for that. \ program_spec exp_count_down assumes "0\n" ensures "a = 2^nat n\<^sub>0" defines \ a = 1; while (n>0) @variant \n\ @invariant \let c = n\<^sub>0-n in 0\n \ n\n\<^sub>0 \ a=2^nat c\ { a=2*a; n=n-1 } \ apply vcg_cs by (auto simp: algebra_simps nat_distribs) subsubsection \Approximate from Below\ text \Used to invert a monotonic function. We count up, until we overshoot the desired result, then we subtract one. \ text \The invariant states that the \r-1\ is not too big. When the loop terminates, \r-1\ is not too big, but \r\ is already too big, so \r-1\ is the desired value (rounding down). \ text \The variant measures the gap that we have to the correct result. Note that the loop will do a final iteration, when the result has been reached exactly. We account for that by adding one, such that the measure also decreases in this case. \ program_spec sqr_approx_below assumes "0\n" ensures "0\r \ r\<^sup>2 \ n\<^sub>0 \ n\<^sub>0 < (r+1)\<^sup>2" defines \ r=1; while (r*r \ n) @variant \n + 1 - r*r\ @invariant \0\r \ (r-1)\<^sup>2 \ n\<^sub>0\ { r = r + 1 }; r = r - 1 \ apply vcg by (auto simp: algebra_simps power2_eq_square) subsubsection \Bisection\ text \A more efficient way of inverting monotonic functions is by bisection, that is, one keeps track of a possible interval for the solution, and halfs the interval in each step. The program will need $O(\log n)$ iterations, and is thus very efficient in practice. Although the final algorithm looks quite simple, getting it right can be quite tricky. \ text \The invariant is surprisingly easy, just stating that the solution is in the interval \l... \ lemma "\h l n\<^sub>0 :: int. \\''invar-final''; 0 \ n\<^sub>0; \ 1 + l < h; 0 \ l; l < h; l * l \ n\<^sub>0; n\<^sub>0 < h * h\ \ n\<^sub>0 < 1 + (l * l + l * 2)" by (smt mult.commute semiring_normalization_rules(3)) program_spec sqr_bisect assumes "0\n" ensures "r\<^sup>2\n\<^sub>0 \ n\<^sub>0<(r+1)\<^sup>2" defines \ l=0; h=n+1; while (l+1 < h) @variant \h-l\ @invariant \0\l \ l l\<^sup>2\n \ n < h\<^sup>2\ { m = (l + h) / 2; if (m*m \ n) l=m else h=m }; r=l \ apply vcg text \We use quick-and-dirty apply style proof to discharge the VCs\ apply (auto simp: power2_eq_square algebra_simps add_pos_pos) apply (smt not_sum_squares_lt_zero) by (smt mult.commute semiring_normalization_rules(3)) subsection \Debugging\ subsubsection \Testing Programs\ text \Stepwise\ schematic_goal "Map.empty: (sqr_approx_below,<''n'':=\_. 4>) \ ?s" unfolding sqr_approx_below_def apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' apply big_step' done text \Or all steps at once\ schematic_goal "Map.empty: (sqr_bisect,<''n'':=\_. 4900000001>) \ ?s" unfolding sqr_bisect_def by big_step subsection \More Numeric Algorithms\ subsubsection \Euclid's Algorithm (with subtraction)\ (* Crucial Lemmas *) thm gcd.commute gcd_diff1 program_spec euclid1 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (a\b) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ (a>0 \ b>0)\ @variant \a+b\ { if (a apply vcg_cs apply (metis gcd.commute gcd_diff1) apply (metis gcd_diff1) done subsubsection \Euclid's Algorithm (with mod)\ (* Crucial Lemmas *) thm gcd_red_int[symmetric] program_spec euclid2 assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0" defines \ while (b\0) @invariant \gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0\ @variant \b\ { t = a; a = b; b = t mod b } \ apply vcg_cs by (simp add: gcd_red_int[symmetric]) subsubsection \Extended Euclid's Algorithm\ (* Provided by Simon Wimmer. *) locale extended_euclid_aux_lemmas begin lemma aux2: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "gcd a b = gcd a\<^sub>0 b\<^sub>0" shows "gcd b (a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q))) = gcd a\<^sub>0 b\<^sub>0" proof - have "a - (a\<^sub>0 * (s * q) + b\<^sub>0 * (t * q)) = a - b * q" unfolding \b = _\ by (simp add: algebra_simps) also have "a - b * q = a mod b" unfolding \q = _\ by (simp add: algebra_simps) finally show ?thesis unfolding \gcd a b = _\[symmetric] by (simp add: gcd_red_int[symmetric]) qed lemma aux3: fixes a b :: int assumes "b = t * b\<^sub>0 + s * a\<^sub>0" "q = a div b" "b > 0" shows "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) \ a" proof - have "t * (b\<^sub>0 * q) + s * (a\<^sub>0 * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) then show ?thesis using \b > 0\ - by (simp add: algebra_simps \q = _\) - (smt Euclidean_Division.pos_mod_sign cancel_div_mod_rules(1) mult.commute) + by (simp add: algebra_simps \q = a div b\) + (simp flip: minus_mod_eq_mult_div) qed end text \The following is a direct translation of the pseudocode for the Extended Euclidean algorithm as described by the English version of Wikipedia (\<^url>\https://en.wikipedia.org/wiki/Extended_Euclidean_algorithm\): \ program_spec euclid_extended assumes "a>0 \ b>0" ensures "old_r = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * old_s + b\<^sub>0 * old_t" defines \ s = 0; old_s = 1; t = 1; old_t = 0; r = b; old_r = a; while (r\0) @invariant \ gcd old_r r = gcd a\<^sub>0 b\<^sub>0 \ r\0 \ old_r>0 \ a\<^sub>0 * old_s + b\<^sub>0 * old_t = old_r \ a\<^sub>0 * s + b\<^sub>0 * t = r \ @variant \r\ { quotient = old_r / r; temp = old_r; old_r = r; r = temp - quotient * r; temp = old_s; old_s = s; s = temp - quotient * s; temp = old_t; old_t = t; t = temp - quotient * t } \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: algebra_simps) apply (simp add: aux2 aux3 minus_div_mult_eq_mod)+ done qed text \Non-Wikipedia version\ context extended_euclid_aux_lemmas begin lemma aux: fixes a b q x y:: int assumes "a = old_y * b\<^sub>0 + old_x * a\<^sub>0" "b = y * b\<^sub>0 + x * a\<^sub>0" "q = a div b" shows "a mod b + (a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q)) = a" proof - have *: "a\<^sub>0 * (x * q) + b\<^sub>0 * (y * q) = q * b" unfolding \b = _\ by (simp add: algebra_simps) show ?thesis unfolding * unfolding \q = _\ by simp qed end program_spec euclid_extended' assumes "a>0 \ b>0" ensures "a = gcd a\<^sub>0 b\<^sub>0 \ gcd a\<^sub>0 b\<^sub>0 = a\<^sub>0 * x + b\<^sub>0 * y" defines \ x = 0; y = 1; old_x = 1; old_y = 0; while (b\0) @invariant \ gcd a b = gcd a\<^sub>0 b\<^sub>0 \ b\0 \ a>0 \ a = a\<^sub>0 * old_x + b\<^sub>0 * old_y \ b = a\<^sub>0 * x + b\<^sub>0 * y \ @variant \b\ { q = a / b; t = a; a = b; b = t mod b; t = x; x = old_x - q * x; old_x = t; t = y; y = old_y - q * y; old_y = t }; x = old_x; y = old_y \ proof - interpret extended_euclid_aux_lemmas . show ?thesis apply vcg_cs apply (simp add: gcd_red_int[symmetric]) apply (simp add: algebra_simps) apply (rule aux; simp add: algebra_simps) done qed subsubsection \Exponentiation by Squaring\ (* Contributed by Simon Wimmer *) lemma ex_by_sq_aux: fixes x :: int and n :: nat assumes "n mod 2 = 1" shows "x * (x * x) ^ (n div 2) = x ^ n" proof - have "n > 0" using assms by presburger have "2 * (n div 2) = n - 1" using assms by presburger then have "(x * x) ^ (n div 2) = x ^ (n - 1)" by (simp add: semiring_normalization_rules) with \0 < n\ show ?thesis by simp (metis Suc_pred power.simps(2)) qed text \A classic algorithm for computing \x\<^sup>n\ works by repeated squaring, using the following recurrence: \<^item> \x\<^sup>n = x * x\<^bsup>(n-1)/2\<^esup>\ if \n\ is odd \<^item> \x\<^sup>n = x\<^bsup>n/2\<^esup>\ if \n\ is even \ program_spec ex_by_sq assumes "n\0" ensures "r = x\<^sub>0 ^ nat n\<^sub>0" defines \ r = 1; while (n\0) @invariant \ n \ 0 \ r * x ^ nat n = x\<^sub>0 ^ nat n\<^sub>0 \ @variant \n\ { if (n mod 2 == 1) { r = r * x }; x = x * x; n = n / 2 } \ apply vcg_cs subgoal apply (subst (asm) ex_by_sq_aux[symmetric]) apply (smt Suc_1 nat_1 nat_2 nat_mod_distrib nat_one_as_int) by (simp add: div_int_pos_iff int_div_less_self nat_div_distrib) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) apply (simp add: algebra_simps semiring_normalization_rules nat_mult_distrib; fail) done subsubsection \Power-Tower of 2s\ fun tower2 where "tower2 0 = 1" | "tower2 (Suc n) = 2 ^ tower2 n" definition "tower2' n = int (tower2 (nat n))" program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; a = 1; while (n>0) @variant \n\ @invariant \True\ \ \This will get ugly, there is no \n\<^sub>0\ that we could use!\ { a=2*a; n=n-1 }; m=m-1 } \ oops text \We prove the inner loop separately instead! (It happens to be exactly our \<^const>\exp_count_down\ program.) \ program_spec tower2_imp assumes \m>0\ ensures \a = tower2' m\<^sub>0\ defines \ a=1; while (m>0) @variant \m\ @invariant \0\m \ m\m\<^sub>0 \ a = tower2' (m\<^sub>0-m)\ { n=a; inline exp_count_down; m=m-1 } \ apply vcg_cs by (auto simp: algebra_simps tower2'_def nat_distribs) subsection \Array Algorithms\ (* Presentation in lecture: introduce array syntax and semantics CONVENTIONS: * Every variable is an array. * Arrays are modeled as int\int. Negative indexes allowed. * Syntactic sugar to use it at index 0 by default V vname \ Vidx vname aexp Syntax: x[i] abbreviation "V x = Vidx x 0" Assign vname aexp \ AssignIdx vname aexp aexp Syntax: x[i] = a abbreviation: Assign x a = AssignIdx x 0 a NEW COMMAND: ArrayCpy vname vname Syntax: x[] = y copy two arrays as a whole. Assign x (V y) only copies index 0 semantics: Obvious aval (Vidx x i) s = s x (aval i s) (x[i] = a, s) \ s(x := (s x)(aval i s := aval a s)) (x[] = y, s) \ s(x := s y) weakest preconditions: Also obvious \ Reasoning with arrays: Usually: Model as int\int is appropriate. Common idioms: * {l..i Warning: Sledgehammer does not perform well on these! Data Refinement Sometimes, abstraction to list may be useful. Idea: Model an algorithm on lists first. Prove it correct. Then implement algorithm. Show that it implements the model. By transitivity, get: Algorithm is correct! Sometimes, abstraction to list may be useful *) subsubsection \Summation\ program_spec array_sum assumes "l\h" ensures "r = (\i=l\<^sub>0..0. a\<^sub>0 i)" defines \ r = 0; while (ll\<^sub>0\l \ l\h \ r = (\i=l\<^sub>0..0 i)\ @variant \h-l\ { r = r + a[l]; l=l+1 } \ apply vcg_cs by (auto simp: intvs_incdec) (* Exercise: Remove l\h precondition! *) subsubsection \Finding Least Index of Element\ program_spec find_least_idx assumes \l\h\ ensures \if l=h\<^sub>0 then x\<^sub>0 \ a\<^sub>0`{l\<^sub>0..0} else l\{l\<^sub>0..0} \ a\<^sub>0 l = x\<^sub>0 \ x\<^sub>0\a\<^sub>0`{l\<^sub>0.. defines \ while (l a[l] \ x) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ x\a`{l\<^sub>0.. l=l+1 \ apply vcg_cs by (smt atLeastLessThan_iff imageI) subsubsection \Check for Sortedness\ term ran_sorted program_spec check_sorted assumes \l\h\ ensures \r\0 \ ran_sorted a\<^sub>0 l\<^sub>0 h\<^sub>0\ defines \ if (l==h) r=1 else { l=l+1; while (l a[l-1] \ a[l]) @variant \h-l\ @invariant \l\<^sub>0 l\h \ ran_sorted a l\<^sub>0 l\ l=l+1; if (l==h) r=1 else r=0 } \ apply vcg_cs apply (auto simp: ran_sorted_def) by (smt atLeastLessThan_iff) subsubsection \Find Equilibrium Index\ definition "is_equil a l h i \ l\i \ i (\j=l..j=i..l\h\ ensures \is_equil a l h i \ i=h \ (\i. is_equil a l h i)\ defines \ usum=0; i=l; while (ih-i\ @invariant \l\i \ i\h \ usum = (\j=l.. { usum = usum + a[i]; i=i+1 }; i=l; lsum=0; while (usum \ lsum \ ih-i\ @invariant \l\i \ i\h \ lsum=(\j=l.. usum=(\j=i.. (\jis_equil a l h j) \ { lsum = lsum + a[i]; usum = usum - a[i]; i=i+1 } \ apply vcg_cs apply (auto simp: intvs_incdec is_equil_def) apply (metis atLeastLessThan_iff eq_iff finite_atLeastLessThan_int sum_diff1) apply force by force subsubsection \Rotate Right\ program_spec rotate_right assumes "0i\{0..0 ((i-1) mod n)" defines \ i = 0; prev = a[n - 1]; while (i < n) @invariant \ 0 \ i \ i \ n \ (\j\{0..0 ((j-1) mod n)) \ (\j\{i..0 j) \ prev = a\<^sub>0 ((i-1) mod n) \ @variant \n - i\ { temp = a[i]; a[i] = prev; prev = temp; i = i + 1 } \ apply vcg_cs by (simp add: zmod_minus1) subsubsection \Binary Search, Leftmost Element\ text \We first specify the pre- and postcondition\ definition "bin_search_pre a l h \ l\h \ ran_sorted a l h" definition "bin_search_post a l h x i \ l\i \ i\h \ (\i\{l.. (\i\{i.. a i)" text \Then we prove that the program is correct\ program_spec binsearch assumes \bin_search_pre a l h\ ensures \bin_search_post a\<^sub>0 l\<^sub>0 h\<^sub>0 x\<^sub>0 l\ defines \ while (l < h) @variant \h-l\ @invariant \l\<^sub>0\l \ l\h \ h\h\<^sub>0 \ (\i\{l\<^sub>0.. (\i\{h..0}. x \ a i)\ { m = (l + h) / 2; if (a[m] < x) l = m + 1 else h = m } \ apply vcg_cs apply (auto simp: algebra_simps bin_search_pre_def bin_search_post_def) txt \Driving sledgehammer to its limits ...\ apply (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) by (smt div_add_self1 even_succ_div_two odd_two_times_div_two_succ ran_sorted_alt) text \Next, we show that our postcondition (which was easy to prove) implies the expected properties of the algorithm. \ lemma assumes "bin_search_pre a l h" "bin_search_post a l h x i" shows bin_search_decide_membership: "x\a`{l.. (i x = a i)" and bin_search_leftmost: "x\a`{l..Naive String Search\ (* By Simon Wimmer *) program_spec match_string assumes "l1 \ h1" ensures "(\j \ {0.. (i < h1 - l1 \ a (l + i) \ b (l1 + i)) \ 0 \ i \ i \ h1 - l1" defines \ i = 0; while (l1 + i < h1 \ a[l + i] == b[l1 + i]) @invariant \(\j \ {0.. 0 \ i \ i \ h1 - l1\ @variant \(h1 - (l1 + i))\ { i = i + 1 } \ by vcg_cs auto lemma lran_eq_iff': "lran a l1 (l1 + (h - l)) = lran a' l h \ (\i. 0 \ i \ i < h - l \ a (l1 + i) = a' (l + i))" if "l \ h" using that proof (induction "nat (h - l)" arbitrary: h) case 0 then show ?case by auto next case (Suc x) then have *: "x = nat (h - 1 - l)" "l \ h - 1" by auto note IH = Suc.hyps(1)[OF *] from * have 1: "lran a l1 (l1 + (h - l)) = lran a l1 (l1 + (h - 1 - l)) @ [a (l1 + (h - 1 - l))]" "lran a' l h = lran a' l (h - 1) @ [a' (h - 1)]" by (auto simp: lran_bwd_simp algebra_simps lran_butlast[symmetric]) from IH * Suc.hyps(2) Suc.prems show ?case unfolding 1 apply auto subgoal for i by (cases "i = h - 1 - l") auto done qed program_spec match_string' assumes "l1 \ h1" ensures "i = h1 - l1 \ lran a l (l + (h1 - l1)) = lran b l1 h1" for i h1 l1 l a[] b[] defines \inline match_string\ by vcg_cs (auto simp: lran_eq_iff') program_spec substring assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; while (l < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } \ by vcg_cs auto program_spec substring' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..h\<^sub>0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { h = h - (h1 - l1) + 1; inline substring } \ by vcg_cs auto (* Or combined: *) program_spec substring'' assumes "l \ h \ l1 \ h1" ensures "match = 1 \ (\ j \ {l\<^sub>0..0-(h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1)" for a[] b[] defines \ match = 0; if (l + (h1 - l1) \ h) { while (l + (h1 - l1) < h \ match == 0) @invariant\l\<^sub>0 \ l \ l \ h - (h1 - l1) \ match \ {0,1} \ (if match = 1 then lran a l (l + (h1 - l1)) = lran b l1 h1 \ l < h - (h1 - l1) else (\j \ {l\<^sub>0.. lran b l1 h1))\ @variant\(h - l) * (1 - match)\ { inline match_string'; if (i == h1 - l1) {match = 1} else {l = l + 1} } } \ by vcg_cs auto lemma lran_split: "lran a l h = lran a l p @ lran a p h" if "l \ p" "p \ h" using that by (induction p; simp; simp add: lran.simps) lemma lran_eq_append_iff: "lran a l h = as @ bs \ (\ i. l \ i \ i \ h \ as = lran a l i \ bs = lran a i h)" if "l \ h" apply safe subgoal using that proof (induction as arbitrary: l) case Nil then show ?case by auto next case (Cons x as) from this(2-) have "lran a (l + 1) h = as @ bs" "a l = x" "l + 1 \ h" apply - subgoal by simp subgoal by (smt append_Cons list.inject lran.simps lran_append1) subgoal using add1_zle_eq by fastforce done from Cons.IH[OF this(1,3)] obtain i where IH: "l + 1 \ i" "i \ h" "as = lran a (l + 1) i" "bs = lran a i h" by auto with \a l = x\ show ?case apply (intro exI[where x = i]) apply auto by (smt IH(3) lran_prepend1) qed apply (subst lran_split; simp) done lemma lran_split': "(\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1) = (\as bs. lran a l h = as @ lran b l1 h1 @ bs)" if "l \ h" "l1 \ h1" proof safe fix j assume j: "j \ {l..h - (h1 - l1)}" and match: "lran a j (j + (h1 - l1)) = lran b l1 h1" with \l1 \ h1\ have "lran a l h = lran a l j @ lran a j (j + (h1 - l1)) @ lran a (j + (h1 - l1)) h" apply (subst lran_split[where p = j], simp, simp) apply (subst (2) lran_split[where p = "j + (h1 - l1)"]; simp) done then show "\as bs. lran a l h = as @ lran b l1 h1 @ bs" by (auto simp: match) next fix as bs assume "lran a l h = as @ lran b l1 h1 @ bs" with that lran_eq_append_iff[of l h a as "lran b l1 h1 @ bs"] obtain i where "as = lran a l i" "lran a i h = lran b l1 h1 @ bs" "l \ i" "i \ h" by auto with lran_eq_append_iff[of i h a "lran b l1 h1" bs] obtain j where j: "lran b l1 h1 = lran a i j" "bs = lran a j h" "i \ j" "j \ h" by auto moreover have "j = i + (h1 - l1)" proof - have "length (lran b l1 h1) = nat (h1 - l1)" "length (lran a i j) = nat (j - i)" by auto with j(1,3) that show ?thesis by auto qed ultimately show "\j\{l..h - (h1 - l1)}. lran a j (j + (h1 - l1)) = lran b l1 h1" using \l \ i\ by auto qed program_spec substring_final assumes "l \ h \ 0 \ len" ensures "match = 1 \ (\as bs. lran a l\<^sub>0 h\<^sub>0 = as @ lran b 0 len @ bs)" for l h len match a[] b[] defines \l1 = 0; h1 = len; inline substring'\ supply [simp] = lran_split'[symmetric] apply vcg_cs done subsubsection \Insertion Sort\ text \ We first prove the inner loop. The specification here specifies what the algorithm does as closely as possible, such that it becomes easier to prove. In this case, sortedness is not a precondition for the inner loop to move the key element backwards over all greater elements. \ definition "insort_insert_post l j a\<^sub>0 a i \ (let key = a\<^sub>0 j in i\{l-1.. \\i\ is in range\ \ \Content of new array\ \ (\k\{l..i}. a k = a\<^sub>0 k) \ a (i+1) = key \ (\k\{i+2..j}. a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ \Placement of \key\\ \ (i\l \ a i \ key) \ \Element at \i\ smaller than \key\, if it exists \ \ (\k\{i+2..j}. a k > key) \ \Elements \\i+2\ greater than \key\\ ) " for l j i :: int and a\<^sub>0 a :: "int \ int" program_spec insort_insert assumes "l0 a i" defines \ key = a[j]; i = j-1; while (i\l \ a[i]>key) @variant \i-l+1\ @invariant \l-1\i \ i (\k\{l..i}. a k = a\<^sub>0 k) \ (\k\{i+2..j}. a k > key \ a k = a\<^sub>0 (k-1)) \ a = a\<^sub>0 on -{l..j} \ { a[i+1] = a[i]; i=i-1 }; a[i+1] = key \ unfolding insort_insert_post_def Let_def apply vcg apply auto by (smt atLeastAtMost_iff) text \Next, we show that our specification that was easy to prove implies the specification that the outer loop expects:\ text \Invoking \insort_insert\ will sort in the element\ lemma insort_insert_sorted: assumes "lInvoking \insort_insert\ will only mutate the elements\ lemma insort_insert_ran1: assumes "li+1" "i+1\j" unfolding insort_insert_post_def Let_def by auto have ranshift: "mset_ran (a o (+)(-1)) {i+2..j} = mset_ran a {i+1..j-1}" by (simp add: mset_ran_shift algebra_simps) have "mset_ran a' {l..j} = mset_ran a' {l..i} + {# a' (i+1) #} + mset_ran a' {i+2..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a'"]) also have "\ = mset_ran a {l..i} + {# a j #} + mset_ran (a o (+)(-1)) {i+2..j}" using EQS(1,3)[THEN mset_ran_cong] EQS(2) by simp also have "\ = mset_ran a {l..i} + mset_ran a {i+1..j-1} + {# a j #}" by (simp add: mset_ran_shift algebra_simps) also have "\ = mset_ran a {l..j}" using \l \l\i+1\ \i+1\j\ apply (simp add: mset_ran_combine) by (auto intro: arg_cong[where f="mset_ran a"]) finally show ?thesis . qed text \The property @{thm insort_insert_ran1} extends to the whole array to be sorted\ lemma insort_insert_ran2: assumes "linsort_insert_post l j a a' i\ have "a' = a on {j<..Finally, we specify and prove correct the outer loop\ program_spec insort assumes "l mset_ran a {l..0 {l.. a=a\<^sub>0 on -{l.. j = l + 1; while (jh-j\ @invariant \ l j\h \ \\j\ in range\ \ ran_sorted a l j \ \Array is sorted up to \j\\ \ mset_ran a {l..0 {l.. \Elements in range only permuted\ \ a=a\<^sub>0 on -{l.. { inline insort_insert; j=j+1 } \ apply vcg_cs apply (intro conjI) subgoal by (rule insort_insert_sorted) subgoal using insort_insert_ran2(1) by auto subgoal apply (frule (2) insort_insert_ran2(2)) by (auto simp: eq_on_def) done subsubsection \Quicksort\ procedure_spec partition_aux(a,l,h,p) returns (a,i) assumes "l\h" ensures "mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0..0) \ (\j\{i..0}. a j \ p\<^sub>0) \ l\<^sub>0\i \ i\h\<^sub>0 \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\k\{l.. (\k\{i.. p) \ (\k\{j..0 k = a k) \ a\<^sub>0 = a on -{l\<^sub>0..0} \ @variant \(h-j)\ { if (a[j] supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs subgoal by (simp add: mset_ran_swap[unfolded swap_def]) subgoal by auto done procedure_spec partition(a,l,h,p) returns (a,i) assumes "l0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ (\j\{l\<^sub>0.. (\j\{i..0}. a j \ a i) \ l\<^sub>0\i \ i0 \ a\<^sub>0 (h\<^sub>0-1) = a i \ a\<^sub>0 = a on -{l\<^sub>0..0} " defines \ p = a[h-1]; (a,i) = partition_aux(a,l,h-1,p); a[h-1] = a[i]; a[i] = p \ apply vcg_cs apply (auto simp: eq_on_def mset_ran_swap[unfolded swap_def] intvs_incdec intro: mset_ran_combine_eq_diff) done lemma quicksort_sorted_aux: assumes BOUNDS: "l \ i" "i < h" assumes LESS: "\j\{l..1 j < a\<^sub>1 i" assumes GEQ: "\j\{i..1 i \ a\<^sub>1 j" assumes R1: "mset_ran a\<^sub>1 {l..2 {l..1 = a\<^sub>2 on - {l..2 l i" assumes R2: "mset_ran a\<^sub>2 {i + 1..3 {i + 1..2 = a\<^sub>3 on - {i + 1..3 (i + 1) h" shows "ran_sorted a\<^sub>3 l h" proof - have [simp]: "{l.. - {i + 1..1 i = a\<^sub>3 i" using E1 E2 by (auto simp: eq_on_def) note X1 = mset_ran_xfer_pointwise[where P = \\x. x < p\ for p, OF R1, simplified] note X2 = eq_on_xfer_pointwise[where P = \\x. x < p\ for p, OF E2, of "{l..j\{l..3 j < a\<^sub>3 i" by (simp add: X1 X2) from GEQ have GEQ1: "\j\{i+1..1 i \ a\<^sub>1 j" by auto have [simp]: "{i + 1.. - {l..\x. x \ p\ for p, OF E1, of "{i+1..\x. x \ p\ for p, OF R2, simplified] from GEQ1 have GEQ': "\j\{i+1..3 i \ a\<^sub>3 j" by (simp add: X3 X4) from SL eq_on_xfer_ran_sorted[OF E2, of l i] have SL': "ran_sorted a\<^sub>3 l i" by simp show ?thesis using combine_sorted_pivot[OF BOUNDS SL' SH LESS' GEQ'] . qed lemma quicksort_mset_aux: assumes B: "l\<^sub>0\i" "i0" assumes R1: "mset_ran a {l\<^sub>0..0..0..0} = mset_ran ab {i + 1..0}" assumes E2: "aa = ab on - {i + 1..0}" shows "mset_ran a {l\<^sub>0..0} = mset_ran ab {l\<^sub>0..0}" apply (rule trans) apply (rule mset_ran_eq_extend[OF R1 E1]) using B apply auto [2] apply (rule mset_ran_eq_extend[OF R2 E2]) using B apply auto [2] done recursive_spec quicksort(a,l,h) returns a assumes "True" ensures "ran_sorted a l\<^sub>0 h\<^sub>0 \ mset_ran a\<^sub>0 {l\<^sub>0..0} = mset_ran a {l\<^sub>0..0} \ a\<^sub>0=a on -{l\<^sub>0..0}" variant "h-l" defines \ if (l apply (vcg_cs; (intro conjI)?) subgoal using quicksort_sorted_aux by metis subgoal using quicksort_mset_aux by metis subgoal by (smt ComplD ComplI atLeastLessThan_iff eq_on_def) subgoal by (auto simp: ran_sorted_def) done subsection \Data Refinement\ subsubsection \Filtering\ program_spec array_filter_negative assumes "l\h" ensures "lran a l\<^sub>0 i = filter (\x. x\0) (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)" defines \ i=l; j=l; while (j l\i \ i\j \ j\h \ lran a l i = filter (\x. x\0) (lran a\<^sub>0 l j) \ lran a j h = lran a\<^sub>0 j h \ @variant \h-j\ { if (a[j]\0) {a[i] = a[j]; i=i+1}; j=j+1 } \ supply lran_eq_iff[simp] lran_tail[simp del] apply vcg_cs done (* Exercise: Use swap to modify this algorithm to partitioning! *) subsubsection \Merge Two Sorted Lists\ text \ We define the merge function abstractly first, as a functional program on lists. \ fun merge where "merge [] ys = ys" | "merge xs [] = xs" | "merge (x#xs) (y#ys) = (if xIt's straightforward to show that this produces a sorted list with the same elements.\ lemma merge_sorted: assumes "sorted xs" "sorted ys" shows "sorted (merge xs ys) \ set (merge xs ys) = set xs \ set ys" using assms apply (induction xs ys rule: merge.induct) apply auto done lemma merge_mset: "mset (merge xs ys) = mset xs + mset ys" by (induction xs ys rule: merge.induct) auto text \Next, we prove an equation that characterizes one step of the while loop, on the list level.\ lemma merge_eq: "xs\[] \ ys\[] \ merge xs ys = ( if ys=[] \ (xs\[] \ hd xs < hd ys) then hd xs # merge (tl xs) ys else hd ys # merge xs (tl ys) )" by (cases xs; cases ys; simp) text \ We do a first proof that our merge implementation on the arrays and indexes behaves like the functional merge on the corresponding lists. The annotations use the @{const lran} function to map from the implementation level to the list level. Moreover, the invariant of the implementation, \l\h\, is carried through explicitly. \ program_spec merge_imp' assumes "l1\h1 \ l2\h2" ensures "let ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in j\0 \ ms = merge xs\<^sub>0 ys\<^sub>0" defines \ j=0; while (l1\h1 \ l2\h2) @variant \h1 + h2 - l1 - l2\ @invariant \let xs=lran a1 l1 h1; ys = lran a2 l2 h2; ms = lran m 0 j; xs\<^sub>0 = lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0; ys\<^sub>0 = lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0 in l1\h1 \ l2\h2 \ 0\j \ merge xs\<^sub>0 ys\<^sub>0 = ms@merge xs ys \ { if (l2==h2 \ (l1\h1 \ a1[l1] < a2[l2])) { m[j] = a1[l1]; l1=l1+1 } else { m[j] = a2[l2]; l2=l2+1 }; j=j+1 } \ text \Given the @{thm [source] merge_eq} theorem, which captures the essence of a loop step, and the theorems @{thm lran_append1}, @{thm lran_tail}, and @{thm hd_lran}, which convert from the operations on arrays and indexes to operations on lists, the proof is straightforward \ apply vcg_cs subgoal apply (subst merge_eq) by auto subgoal by linarith subgoal apply (subst merge_eq) by auto done text \ In a next step, we refine our proof to combine it with the abstract properties we have proved about merge. The program does not change (we simply inline the original one here). \ procedure_spec merge_imp (a1,l1,h1,a2,l2,h2) returns (m,j) assumes "l1\h1 \ l2\h2 \ sorted (lran a1 l1 h1) \ sorted (lran a2 l2 h2)" ensures "let ms = lran m 0 j in j\0 \ sorted ms \ mset ms = mset (lran a1\<^sub>0 l1\<^sub>0 h1\<^sub>0) + mset (lran a2\<^sub>0 l2\<^sub>0 h2\<^sub>0)" for l1 h1 l2 h2 a1[] a2[] m[] j defines \inline merge_imp'\ apply vcg_cs apply (auto simp: Let_def merge_mset dest: merge_sorted) done thm merge_imp_spec thm merge_imp_def lemma [named_ss vcg_bb]: "UNIV \ a = UNIV" "a \ UNIV = UNIV" by auto lemma merge_msets_aux: "\l\m; m\h\ \ mset (lran a l m) + mset (lran a m h) = mset (lran a l h)" by (auto simp: mset_lran mset_ran_combine ivl_disj_un_two) recursive_spec mergesort (a,l,h) returns (b,j) assumes "l\h" ensures \0\j \ sorted (lran b 0 j) \ mset (lran b 0 j) = mset (lran a\<^sub>0 l\<^sub>0 h\<^sub>0)\ variant \h-l\ for a[] b[] defines \ if (l==h) j=0 else if (l+1==h) { b[0] = a[l]; j=1 } else { m = (h+l) / 2; (a1,h1) = rec mergesort (a,l,m); (a2,h2) = rec mergesort (a,m,h); (b,j) = merge_imp (a1,0,h1,a2,0,h2) } \ apply vcg apply auto [] apply (auto simp: lran.simps) [] apply auto [] apply auto [] apply auto [] apply (auto simp: Let_def merge_msets_aux) [] done print_theorems subsubsection \Remove Duplicates from Array, using Bitvector Set\ text \We use an array to represent a set of integers. If we only insert elements in range \{0.., this representation is called bit-vector (storing a single bit per index is enough). \ definition set_of :: "(int \ int) \ int set" where "set_of a \ {i. a i \ 0}" context notes [simp] = set_of_def begin lemma set_of_empty[simp]: "set_of (\_. 0) = {}" by auto lemma set_of_insert[simp]: "x\0 \ set_of (a(i:=x)) = insert i (set_of a)" by auto lemma set_of_remove[simp]: "set_of (a(i:=0)) = set_of a - {i}" by auto lemma set_of_mem[simp]: "i\set_of a \ a i \ 0" by auto end program_spec dedup assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l h) \ distinct (lran a l i)\ defines \ i=l; j=l; clear b[]; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { if (b[a[j]] == 0) { a[i] = a[j]; i=i+1; b[a[j]] = 1 }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) [] apply (auto simp: lran_eq_iff) [] done procedure_spec bv_init () returns b assumes True ensures \set_of b = {}\ defines \clear b[]\ by vcg_cs procedure_spec bv_insert (x, b) returns b assumes True ensures \set_of b = insert x\<^sub>0 (set_of b\<^sub>0)\ defines \b[x] = 1\ by vcg_cs procedure_spec bv_remove (x, b) returns b assumes True ensures \set_of b = set_of b\<^sub>0 - {x\<^sub>0}\ defines \b[x] = 0\ by vcg_cs procedure_spec bv_elem (x,b) returns r assumes True ensures \r\0 \ x\<^sub>0\set_of b\<^sub>0\ defines \r = b[x]\ by vcg_cs procedure_spec dedup' (a,l,h) returns (a,l,i) assumes \l\h\ ensures \set (lran a l i) = set (lran a\<^sub>0 l\<^sub>0 h\<^sub>0) \ distinct (lran a l i) \ for b[] defines \ b = bv_init(); i=l; j=l; while (jh-j\ @invariant \l\i \ i\j \ j\h \ set (lran a l i) = set (lran a\<^sub>0 l j) \ distinct (lran a l i) \ lran a j h = lran a\<^sub>0 j h \ set_of b = set (lran a l i) \ { mem = bv_elem (a[j],b); if (mem == 0) { a[i] = a[j]; i=i+1; b = bv_insert(a[j],b) }; j=j+1 } \ apply vcg_cs apply (auto simp: lran_eq_iff lran_upd_inside intro: arg_cong[where f=tl]) done subsection \Recursion\ subsubsection \Recursive Fibonacci\ recursive_spec fib_imp (i) returns r assumes True ensures \r = fib i\<^sub>0\ variant \i\ defines \ if (i\0) r=0 else if (i==1) r=1 else { r1=rec fib_imp (i-2); r2=rec fib_imp (i-1); r = r1+r2 } \ by vcg_cs subsubsection \Homeier's Cycling Termination\ text \A contrived example from Homeier's thesis. Only the termination proof is done.\ (* TODO: Also show correctness: pedal (a,b) will multiply its arguments! correctness proof may be a good test how to handle specs with logical vars! *) recursive_spec pedal (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m\ defines \ if (n\0 \ m\0) { G = G + m; if (n and coast (n,m) returns () assumes \n\0 \ m\0\ ensures True variant \n+m+1\ defines \ G = G + n; if (n by vcg_cs subsubsection \Ackermann\ fun ack :: "nat \ nat \ nat" where "ack 0 n = n+1" | "ack m 0 = ack (m-1) 1" | "ack m n = ack (m-1) (ack m (n-1))" lemma ack_add_simps[simp]: "m\0 \ ack m 0 = ack (m-1) 1" "\m\0; n\0\ \ ack m n = ack (m-1) (ack m (n-1))" subgoal by (cases m) auto subgoal by (cases "(m,n)" rule: ack.cases) (auto) done recursive_spec relation "less_than <*lex*> less_than" ack_imp (m,n) returns r assumes "m\0 \ n\0" ensures "r=int (ack (nat m\<^sub>0) (nat n\<^sub>0))" variant "(nat m, nat n)" defines \ if (m==0) r = n+1 else if (n==0) r = rec ack_imp (m-1,1) else { t = rec ack_imp (m,n-1); r = rec ack_imp (m-1,t) } \ supply nat_distribs[simp] by vcg_cs subsubsection \McCarthy's 91 Function\ text \A standard benchmark for verification of recursive functions. We use Homeier's version with a global variable.\ recursive_spec p91(y) assumes True ensures "if 1000 then G = y\<^sub>0-10 else G = 91" variant "101-y" for G defines \ if (100 apply vcg_cs apply (auto split: if_splits) done subsubsection \Odd/Even\ recursive_spec odd_imp (a) returns b assumes True ensures "b\0 \ odd a\<^sub>0" variant "\a\" defines \ if (a==0) b=0 else if (a<0) b = rec even_imp (a+1) else b = rec even_imp (a-1) \ and even_imp (a) returns b assumes True ensures "b\0 \ even a\<^sub>0" variant "\a\" defines \ if (a==0) b=1 else if (a<0) b = rec odd_imp (a+1) else b = rec odd_imp (a-1) \ apply vcg apply auto done thm even_imp_spec subsubsection \Pandya and Joseph's Product Producers\ text \Again, taking the version from Homeier's thesis, but with a modification to also terminate for negative \y\. \ recursive_spec relation \measure nat <*lex*> less_than\ product () assumes True ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,1::nat)" for GX GY GZ defines \ e = even_imp (GY); if (e\0) rec evenproduct() else rec oddproduct() \ and oddproduct() assumes \odd GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY<0) { GY = GY + 1; GZ = GZ - GX } else { GY = GY - 1; GZ = GZ + GX }; rec evenproduct() \ and evenproduct() assumes \even GY\ ensures \GZ = GZ\<^sub>0 + GX\<^sub>0*GY\<^sub>0\ variant "(\GY\,0::nat)" for GX GY GZ defines \ if (GY\0) { GX = 2*GX; GY = GY / 2; rec product() } \ apply vcg_cs apply (auto simp: algebra_simps) apply presburger+ done (* TODO: * Infrastructure to prove different specification for same program (DONE) We can use inline, and just give program a new name. * ghost variables. Keep them existentially qualified, with naming tag. ABS_GHOST ''name'' (\x. P x) \ \x. P x GHOST ''name'' value = True assumes ABS_GHOST name (\x. P x) obtains x where "GHOST name x" "P x" assumes "GHOST name x" and "P x" shows "ABS_GHOST name (\x. P x)" Let_Ghost name (\x s. C x s) = skip assumes "\x. C x s" assumes "\x. \ GHOST name x; C x s \ \ Q s" shows "wp (Let_Ghost name (\x s. C x s)) Q s" Let specification command also abstract over ghost variables in program. *) (* IDEAS: Extend arrays to multiple dimensions: int list \ int Vidx vname "aexp list" AssignIdx vname "aexp list" aexp ArrayCpy vname vname -- copy all dimensions (it's simpler) ArrayInit vname -- Init all dimensions plain values on dimension []! then we can easily encode matrices and adjacency lists! *) subsection \Graph Algorithms\ subsubsection \DFS\ (* By Simon Wimmer *) text \A graph is stored as an array of integers. Each node is an index into this array, pointing to a size-prefixed list of successors. Example for node \i\, which has successors \s1... sn\: \<^verbatim>\ Indexes: ... | i | i+1 | ... | i+n | ... Data: ... | n | s1 | ... | sn | ... \ \ definition succs where "succs a i \ a ` {i+1.. int" definition Edges where "Edges a \ {(i, j). j \ succs a i}" procedure_spec push' (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \lran stack 0 ptr = lran stack\<^sub>0 0 ptr\<^sub>0 @ [x\<^sub>0] \ ptr = ptr\<^sub>0 + 1\ defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs procedure_spec push (x, stack, ptr) returns (stack, ptr) assumes "ptr \ 0" ensures \stack ` {0..0} \ stack\<^sub>0 ` {0..0} \ ptr = ptr\<^sub>0 + 1\ for stack[] defines \stack[ptr] = x; ptr = ptr + 1\ by vcg_cs (auto simp: fun_upd_image) program_spec get_succs assumes "j \ stop \ stop = a (j - 1) \ 0 \ i" ensures " stack ` {0..0 - 1, x) \ Edges a \ x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ i \ i\<^sub>0" for i j stop stack[] a[] visited[] defines \ while (j < stop) @invariant \stack ` {0.. a ` {j\<^sub>0.. x \ set_of visited} \ stack\<^sub>0 ` {0..0} \ j \ stop \ i\<^sub>0 \ i \ j\<^sub>0 \ j \ @variant \(stop - j)\ { succ = a[j]; is_elem = bv_elem(succ,visited); if (is_elem == 0) { (stack, i) = push (succ, stack, i) }; j = j + 1 } \ by vcg_cs (auto simp: intvs_incr_h Edges_def succs_def) procedure_spec pop (stack, ptr) returns (x, ptr) assumes "ptr \ 1" ensures \stack\<^sub>0 ` {0..0} = stack\<^sub>0 ` {0.. {x} \ ptr\<^sub>0 = ptr + 1\ for stack[] defines \ptr = ptr - 1; x = stack[ptr]\ by vcg_cs (simp add: intvs_upper_decr) procedure_spec stack_init () returns i assumes True ensures \i = 0\ defines \i = 0\ by vcg_cs lemma Edges_empty: "Edges a `` {i} = {}" if "i + 1 \ a i" using that unfolding Edges_def succs_def by auto text \This is one of the main insights of the algorithm: if a set of visited states is closed w.r.t.\ to the edge relation, then it is guaranteed to contain all the states that are reachable from any state within the set.\ lemma reachability_invariant: assumes reachable: "(s, x) \ (Edges a)\<^sup>*" and closed: "\v\visited. Edges a `` {v} \ visited" and start: "s \ visited" shows "x \ visited" using reachable start closed by induction auto program_spec (partial) dfs assumes "0 \ x \ 0 \ s" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" defines \ b = 0; clear stack[]; i = stack_init(); (stack, i) = push (s, stack, i); clear visited[]; while (b == 0 \ i \ 0) @invariant \0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ { (next, i) = pop(stack, i); \\Take the top most element from the stack.\ visited = bv_insert(next, visited); \\Mark it as visited,\ if (next == x) { b = 1 \\If it is the target, we are done.\ } else { \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } \ apply vcg_cs subgoal by (auto simp: set_of_def) subgoal using intvs_lower_incr by (auto simp: Edges_empty) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done text \Assuming that the input graph is finite, we can also prove that the algorithm terminates. We will thus use an \Isabelle context\ to fix a certain finite graph and a start state: \ context fixes start :: int and edges assumes finite_graph[intro!]: "finite ((Edges edges)\<^sup>* `` {start})" begin lemma sub_insert_same_iff: "s \ insert x s \ x\s" by auto program_spec dfs1 assumes "0 \ x \ 0 \ s \ start = s \ edges = a" ensures "b = 1 \ x \ (Edges a)\<^sup>* `` {s}" for visited[] defines \ b = 0; \\\i\ will point to the next free space in the stack (i.e. it is the size of the stack)\ i = 1; \\Initially, we put \s\ on the stack.\ stack[0] = s; visited = bv_init(); while (b == 0 \ i \ 0) @invariant \ 0 \ i \ (s \ stack ` {0.. s \ set_of visited) \ (b = 0 \ b = 1) \ set_of visited \ (Edges edges)\<^sup>* `` {start} \ ( if b = 0 then stack ` {0.. (Edges a)\<^sup>* `` {s} \ (\v \ set_of visited. (Edges a) `` {v} \ set_of visited \ stack ` {0.. (x \ set_of visited) else x \ (Edges a)\<^sup>* `` {s}) \ @relation \finite_psupset ((Edges edges)\<^sup>* `` {start}) <*lex*> less_than\ @variant \ (set_of visited, nat i)\ { \\Take the top most element from the stack.\ (next, i) = pop(stack, i); if (next == x) { \\If it is the target, we are done.\ visited = bv_insert(next, visited); b = 1 } else { is_elem = bv_elem(next, visited); if (is_elem == 0) { visited = bv_insert(next, visited); \\Else, put its successors on the stack if they are not yet visited.\ stop = a[next]; j = next + 1; if (j \ stop) { inline get_succs } } } } \ apply vcg_cs subgoal by auto subgoal by (auto simp add: image_constant_conv) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: Edges_empty) subgoal by (clarsimp simp: finite_psupset_def sub_insert_same_iff) subgoal by (auto simp: set_of_def) subgoal by auto (fastforce simp: set_of_def dest!: reachability_invariant) done end end \ No newline at end of file diff --git a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Signed_Modulo.thy b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Signed_Modulo.thy --- a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Signed_Modulo.thy +++ b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Signed_Modulo.thy @@ -1,109 +1,119 @@ section \Signed Modulo Operation\ theory Signed_Modulo imports Berlekamp_Zassenhaus.Poly_Mod Sqrt_Babylonian.Sqrt_Babylonian_Auxiliary begin text \The upcoming definition of symmetric modulo is different to the HOL-Library-Signed\_Division.smod, since here the modulus will be in range $\{-m/2,...,m/2\}$, whereas there -1 symmod m = m - 1. The advantage of have range $\{-m/2,...,m/2\}$ is that small negative numbers are represented by small numbers. One limitation is that the symmetric modulo is only working properly, if the modulus is a positive number.\ definition sym_mod :: "int \ int \ int" (infixl "symmod" 70) where "sym_mod x y = poly_mod.inv_M y (x mod y)" lemma sym_mod_code[code]: "sym_mod x y = (let m = x mod y in if m + m \ y then m else m - y)" unfolding sym_mod_def poly_mod.inv_M_def Let_def .. lemma sym_mod_zero[simp]: "n symmod 0 = n" "n > 0 \ 0 symmod n = 0" unfolding sym_mod_def poly_mod.inv_M_def by auto -lemma sym_mod_range: "y > 0 \ x symmod y \ {- ((y - 1) div 2) .. y div 2}" - unfolding sym_mod_def poly_mod.inv_M_def using pos_mod_bound[of y x] - by (cases "x mod y \ y", auto) - (smt (verit) Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign half_nonnegative_int_iff)+ +lemma sym_mod_range: + \x symmod y \ {- ((y - 1) div 2) .. y div 2}\ if \y > 0\ +proof - + from that have \x mod y < y\ + by (rule pos_mod_bound) + then have \x mod y - y < 0\ + by simp + moreover from that have \- ((y - 1) div 2) \ 0\ \0 \ x mod y\ + by simp_all + then have \- ((y - 1) div 2) \ x mod y\ + by (rule order_trans) + ultimately show ?thesis + by (auto simp add: sym_mod_def poly_mod.inv_M_def) +qed text \The range is optimal in the sense that exactly y elements can be represented.\ lemma card_sym_mod_range: "y > 0 \ card {- ((y - 1) div 2) .. y div 2} = y" by simp lemma sym_mod_abs: "y > 0 \ \x symmod y\ < y" "y \ 1 \ \x symmod y\ \ y div 2" using sym_mod_range[of y x] by auto - lemma sym_mod_sym_mod[simp]: "x symmod y symmod y = x symmod (y :: int)" unfolding sym_mod_def using poly_mod.M_def poly_mod.M_inv_M_id by auto lemma sym_mod_diff_eq: "(a symmod c - b symmod c) symmod c = (a - b) symmod c" unfolding sym_mod_def by (metis mod_diff_cong mod_mod_trivial poly_mod.M_def poly_mod.M_inv_M_id) lemma sym_mod_sym_mod_cancel: "c dvd b \ a symmod b symmod c = a symmod c" using mod_mod_cancel[of c b] unfolding sym_mod_def by (metis poly_mod.M_def poly_mod.M_inv_M_id) lemma sym_mod_diff_right_eq: "(a - b symmod c) symmod c = (a - b) symmod c" using sym_mod_diff_eq by (metis sym_mod_sym_mod) lemma sym_mod_mult_right_eq: "a * (b symmod c) symmod c = a * b symmod c" unfolding sym_mod_def by (metis poly_mod.M_def poly_mod.M_inv_M_id mod_mult_right_eq) lemma dvd_imp_sym_mod_0 [simp]: "b symmod a = 0" if "a > 0" "a dvd b" unfolding sym_mod_def poly_mod.inv_M_def using that by simp lemma sym_mod_0_imp_dvd [dest!]: "b dvd a" if "a symmod b = 0" - using that unfolding sym_mod_def poly_mod.inv_M_def - by (smt (verit) Euclidean_Division.pos_mod_bound dvd_eq_mod_eq_0) + using that apply (simp add: sym_mod_def poly_mod.inv_M_def not_le split: if_splits) + using pos_mod_bound [of b a] apply auto + done definition sym_div :: "int \ int \ int" (infixl "symdiv" 70) where "sym_div x y = (let d = x div y; m = x mod y in if m + m \ y then d else d + 1)" lemma of_int_mod_integer: "(of_int (x mod y) :: integer) = (of_int x :: integer) mod (of_int y)" using integer_of_int_eq_of_int modulo_integer.abs_eq by presburger lemma sym_div_code[code]: "sym_div x y = (let yy = integer_of_int y in (case divmod_integer (integer_of_int x) yy of (d, m) \ if m + m \ yy then int_of_integer d else (int_of_integer (d + 1))))" unfolding sym_div_def Let_def divmod_integer_def split apply (rule if_cong, subst of_int_le_iff[symmetric], unfold of_int_add) by (subst (1 2) of_int_mod_integer, auto) lemma sym_mod_sym_div: assumes y: "y > 0" shows "x symmod y = x - sym_div x y * y" proof - let ?z = "x - y * (x div y)" let ?u = "y * (x div y)" have "x = y * (x div y) + x mod y" using y by simp hence id: "x mod y = ?z" by linarith have "x symmod y = poly_mod.inv_M y ?z" unfolding sym_mod_def id by auto also have "\ = (if ?z + ?z \ y then ?z else ?z - y)" unfolding poly_mod.inv_M_def .. also have "\ = x - (if (x mod y) + (x mod y) \ y then x div y else x div y + 1) * y" by (simp add: algebra_simps id) also have "(if (x mod y) + (x mod y) \ y then x div y else x div y + 1) = sym_div x y" unfolding sym_div_def Let_def .. finally show ?thesis . qed lemma dvd_sym_div_mult_right [simp]: "(a symdiv b) * b = a" if "b > 0" "b dvd a" using sym_mod_sym_div[of b a] that by simp lemma dvd_sym_div_mult_left [simp]: "b * (a symdiv b) = a" if "b > 0" "b dvd a" using dvd_sym_div_mult_right[OF that] by (simp add: ac_simps) end \ No newline at end of file diff --git a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy --- a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy +++ b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy @@ -1,2327 +1,2327 @@ section \Storjohann's basis reduction algorithm (abstract version)\ text \This theory contains the soundness proofs of Storjohann's basis reduction algorithms, both for the normal and the improved-swap-order variant. The implementation of Storjohann's version of LLL uses modular operations throughout. It is an abstract implementation that is already quite close to what the actual implementation will be. In particular, the swap operation here is derived from the computation lemma for the swap operation in the old, integer-only formalization of LLL.\ theory Storjohann imports Storjohann_Mod_Operation LLL_Basis_Reduction.LLL_Number_Bounds Sqrt_Babylonian.NthRoot_Impl begin subsection \Definition of algorithm\ text \In the definition of the algorithm, the first-flag determines, whether only the first vector of the reduced basis should be computed, i.e., a short vector. Then the modulus can be slightly decreased in comparison to the required modulus for computing the whole reduced matrix.\ fun max_list_rats_with_index :: "(int * int * nat) list \ (int * int * nat)" where "max_list_rats_with_index [x] = x" | "max_list_rats_with_index ((n1,d1,i1) # (n2,d2,i2) # xs) = max_list_rats_with_index ((if n1 * d2 \ n2 * d1 then (n2,d2,i2) else (n1,d1,i1)) # xs)" context LLL begin definition "log_base = (10 :: int)" definition bound_number :: "bool \ nat" where "bound_number first = (if first \ m \ 0 then 1 else m)" definition compute_mod_of_max_gso_norm :: "bool \ rat \ int" where "compute_mod_of_max_gso_norm first mn = log_base ^ (log_ceiling log_base (max 2 ( root_rat_ceiling 2 (mn * (rat_of_nat (bound_number first) + 3)) + 1)))" definition g_bnd_mode :: "bool \ rat \ int vec list \ bool" where "g_bnd_mode first b fs = (if first \ m \ 0 then sq_norm (gso fs 0) \ b else g_bnd b fs)" definition d_of where "d_of dmu i = (if i = 0 then 1 :: int else dmu $$ (i - 1, i - 1))" definition compute_max_gso_norm :: "bool \ int mat \ rat \ nat" where "compute_max_gso_norm first dmu = (if m = 0 then (0,0) else case max_list_rats_with_index (map (\ i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< (if first then 1 else m)]) of (num, denom, i) \ (of_int num / of_int denom, i))" context fixes p :: int \ \the modulus\ and first :: bool \ \only compute first vector of reduced basis\ begin definition basis_reduction_mod_add_row :: "int vec list \ int mat \ nat \ nat \ (int vec list \ int mat)" where "basis_reduction_mod_add_row mfs dmu i j = (let c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j)) in (if c = 0 then (mfs, dmu) else (mfs[ i := (map_vec (\ x. x symmod p)) (mfs ! i - c \\<^sub>v mfs ! j)], mat m m (\(i',j'). (if (i' = i \ j' \ j) then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j')) else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d_of dmu j' * d_of dmu (Suc j'))) else (dmu $$ (i',j')))))))" fun basis_reduction_mod_add_rows_loop where "basis_reduction_mod_add_rows_loop mfs dmu i 0 = (mfs, dmu)" | "basis_reduction_mod_add_rows_loop mfs dmu i (Suc j) = ( let (mfs', dmu') = basis_reduction_mod_add_row mfs dmu i j in basis_reduction_mod_add_rows_loop mfs' dmu' i j)" definition basis_reduction_mod_swap_dmu_mod :: "int mat \ nat \ int mat" where "basis_reduction_mod_swap_dmu_mod dmu k = mat m m (\(i, j). ( if j < i \ (j = k \ j = k - 1) then dmu $$ (i, j) symmod (p * d_of dmu j * d_of dmu (Suc j)) else dmu $$ (i, j)))" definition basis_reduction_mod_swap where "basis_reduction_mod_swap mfs dmu k = (mfs[k := mfs ! (k - 1), k - 1 := mfs ! k], basis_reduction_mod_swap_dmu_mod (mat m m (\(i,j). ( if j < i then if i = k - 1 then dmu $$ (k, j) else if i = k \ j \ k - 1 then dmu $$ (k - 1, j) else if i > k \ j = k then ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j)) div (d_of dmu k) else if i > k \ j = k - 1 then (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1))) div (d_of dmu k) else dmu $$ (i, j) else if i = j then if i = k - 1 then ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1)) div (d_of dmu k) else (d_of dmu (Suc i)) else dmu $$ (i, j)) )) k)" fun basis_reduction_adjust_mod where "basis_reduction_adjust_mod mfs dmu = (let (b,g_idx) = compute_max_gso_norm first dmu; p' = compute_mod_of_max_gso_norm first b in if p' < p then let mfs' = map (map_vec (\x. x symmod p')) mfs; d_vec = vec (Suc m) (\ i. d_of dmu i); dmu' = mat m m (\ (i,j). if j < i then dmu $$ (i,j) symmod (p' * d_vec $ j * d_vec $ (Suc j)) else dmu $$ (i,j)) in (p', mfs', dmu', g_idx) else (p, mfs, dmu, g_idx))" definition basis_reduction_adjust_swap_add_step where "basis_reduction_adjust_swap_add_step mfs dmu g_idx i = ( let i1 = i - 1; (mfs1, dmu1) = basis_reduction_mod_add_row mfs dmu i i1; (mfs2, dmu2) = basis_reduction_mod_swap mfs1 dmu1 i in if i1 = g_idx then basis_reduction_adjust_mod mfs2 dmu2 else (p, mfs2, dmu2, g_idx))" definition basis_reduction_mod_step where "basis_reduction_mod_step mfs dmu g_idx i (j :: int) = (if i = 0 then (p, mfs, dmu, g_idx, Suc i, j) else let di = d_of dmu i; (num, denom) = quotient_of \ in if di * di * denom \ num * d_of dmu (i - 1) * d_of dmu (Suc i) then (p, mfs, dmu, g_idx, Suc i, j) else let (p', mfs', dmu', g_idx') = basis_reduction_adjust_swap_add_step mfs dmu g_idx i in (p', mfs', dmu', g_idx', i - 1, j + 1))" primrec basis_reduction_mod_add_rows_outer_loop where "basis_reduction_mod_add_rows_outer_loop mfs dmu 0 = (mfs, dmu)" | "basis_reduction_mod_add_rows_outer_loop mfs dmu (Suc i) = (let (mfs', dmu') = basis_reduction_mod_add_rows_outer_loop mfs dmu i in basis_reduction_mod_add_rows_loop mfs' dmu' (Suc i) (Suc i))" end text \the main loop of the normal Storjohann algorithm\ partial_function (tailrec) basis_reduction_mod_main where "basis_reduction_mod_main p first mfs dmu g_idx i (j :: int) = ( (if i < m then case basis_reduction_mod_step p first mfs dmu g_idx i j of (p', mfs', dmu', g_idx', i', j') \ basis_reduction_mod_main p' first mfs' dmu' g_idx' i' j' else (p, mfs, dmu)))" definition compute_max_gso_quot:: "int mat \ (int * int * nat)" where "compute_max_gso_quot dmu = max_list_rats_with_index (map (\i. ((d_of dmu (i+1)) * (d_of dmu (i+1)), (d_of dmu (i+2)) * (d_of dmu i), Suc i)) [0..<(m-1)])" text \the main loop of Storjohann's algorithm with improved swap order\ partial_function (tailrec) basis_reduction_iso_main where "basis_reduction_iso_main p first mfs dmu g_idx (j :: int) = ( (if m > 1 then (let (max_gso_num, max_gso_denum, indx) = compute_max_gso_quot dmu; (num, denum) = quotient_of \ in (if (max_gso_num * denum > num * max_gso_denum) then case basis_reduction_adjust_swap_add_step p first mfs dmu g_idx indx of (p', mfs', dmu', g_idx') \ basis_reduction_iso_main p' first mfs' dmu' g_idx' (j + 1) else (p, mfs, dmu))) else (p, mfs, dmu)))" definition compute_initial_mfs where "compute_initial_mfs p = map (map_vec (\x. x symmod p)) fs_init" definition compute_initial_dmu where "compute_initial_dmu p dmu = mat m m (\(i',j'). if j' < i' then dmu $$ (i', j') symmod (p * d_of dmu j' * d_of dmu (Suc j')) else dmu $$ (i', j'))" definition "dmu_initial = (let dmu = d\_impl fs_init in mat m m (\ (i,j). if j \ i then d\_impl fs_init !! i !! j else 0))" definition "compute_initial_state first = (let dmu = dmu_initial; (b, g_idx) = compute_max_gso_norm first dmu; p = compute_mod_of_max_gso_norm first b in (p, compute_initial_mfs p, compute_initial_dmu p dmu, g_idx))" text \Storjohann's algorithm\ definition reduce_basis_mod :: "int vec list" where "reduce_basis_mod = ( let first = False; (p0, mfs0, dmu0, g_idx) = compute_initial_state first; (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0; (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1) in mfs'')" text \Storjohann's algorithm with improved swap order\ definition reduce_basis_iso :: "int vec list" where "reduce_basis_iso = ( let first = False; (p0, mfs0, dmu0, g_idx) = compute_initial_state first; (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0; (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1) in mfs'')" text \Storjohann's algorithm for computing a short vector\ definition "short_vector_mod = ( let first = True; (p0, mfs0, dmu0, g_idx) = compute_initial_state first; (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0 in hd mfs')" text \Storjohann's algorithm (iso-variant) for computing a short vector\ definition "short_vector_iso = ( let first = True; (p0, mfs0, dmu0, g_idx) = compute_initial_state first; (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0 in hd mfs')" end subsection \Towards soundness of Storjohann's algorithm\ lemma max_list_rats_with_index_in_set: assumes max: "max_list_rats_with_index xs = (nm, dm, im)" and len: "length xs \ 1" shows "(nm, dm, im) \ set xs" using assms proof (induct xs rule: max_list_rats_with_index.induct) case (2 n1 d1 i1 n2 d2 i2 xs) have "1 \ length ((if n1 * d2 \ n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs)" by simp moreover have "max_list_rats_with_index ((if n1 * d2 \ n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs) = (nm, dm, im)" using 2 by simp moreover have "(if n1 * d2 \ n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) \ set ((n1, d1, i1) # (n2, d2, i2) # xs)" by simp moreover then have "set ((if n1 * d2 \ n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs) \ set ((n1, d1, i1) # (n2, d2, i2) # xs)" by auto ultimately show ?case using 2(1) by auto qed auto lemma max_list_rats_with_index: assumes "\ n d i. (n,d,i) \ set xs \ d > 0" and max: "max_list_rats_with_index xs = (nm, dm, im)" and "(n,d,i) \ set xs" shows "rat_of_int n / of_int d \ of_int nm / of_int dm" using assms proof (induct xs arbitrary: n d i rule: max_list_rats_with_index.induct) case (2 n1 d1 i1 n2 d2 i2 xs n d i) let ?r = "rat_of_int" from 2(2) have "d1 > 0" "d2 > 0" by auto hence d: "?r d1 > 0" "?r d2 > 0" by auto have "(n1 * d2 \ n2 * d1) = (?r n1 * ?r d2 \ ?r n2 * ?r d1)" unfolding of_int_mult[symmetric] by presburger also have "\ = (?r n1 / ?r d1 \ ?r n2 / ?r d2)" using d by (smt divide_strict_right_mono leD le_less_linear mult.commute nonzero_mult_div_cancel_left not_less_iff_gr_or_eq times_divide_eq_right) finally have id: "(n1 * d2 \ n2 * d1) = (?r n1 / ?r d1 \ ?r n2 / ?r d2)" . obtain n' d' i' where new: "(if n1 * d2 \ n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) = (n',d',i')" by force have nd': "(n',d',i') \ {(n1,d1,i1), (n2, d2, i2)}" using new[symmetric] by auto from 2(3) have res: "max_list_rats_with_index ((n',d',i') # xs) = (nm, dm, im)" using new by auto note 2 = 2[unfolded new] show ?case proof (cases "(n,d,i) \ set xs") case True show ?thesis by (rule 2(1)[of n d, OF 2(2) res], insert True nd', force+) next case False with 2(4) have "n = n1 \ d = d1 \ n = n2 \ d = d2" by auto hence "?r n / ?r d \ ?r n' / ?r d'" using new[unfolded id] by (metis linear prod.inject) also have "?r n' / ?r d' \ ?r nm / ?r dm" by (rule 2(1)[of n' d', OF 2(2) res], insert nd', force+) finally show ?thesis . qed qed auto context LLL begin lemma log_base: "log_base \ 2" unfolding log_base_def by auto definition LLL_invariant_weak' :: "nat \ int vec list \ bool" where "LLL_invariant_weak' i fs = ( gs.lin_indpt_list (RAT fs) \ lattice_of fs = L \ weakly_reduced fs i \ i \ m \ length fs = m )" lemma LLL_invD_weak: assumes "LLL_invariant_weak' i fs" shows "lin_indep fs" "length (RAT fs) = m" "set fs \ carrier_vec n" "\ i. i < m \ fs ! i \ carrier_vec n" "\ i. i < m \ gso fs i \ carrier_vec n" "length fs = m" "lattice_of fs = L" "weakly_reduced fs i" "i \ m" proof (atomize (full), goal_cases) case 1 interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs" by (standard) (use assms LLL_invariant_weak'_def gs.lin_indpt_list_def in auto) show ?case using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier by (auto simp add: LLL_invariant_weak'_def gram_schmidt_fs.reduced_def) qed lemma LLL_invI_weak: assumes "set fs \ carrier_vec n" "length fs = m" "lattice_of fs = L" "i \ m" "lin_indep fs" "weakly_reduced fs i" shows "LLL_invariant_weak' i fs" unfolding LLL_invariant_weak'_def Let_def using assms by auto lemma LLL_invw'_imp_w: "LLL_invariant_weak' i fs \ LLL_invariant_weak fs" unfolding LLL_invariant_weak'_def LLL_invariant_weak_def by auto lemma basis_reduction_add_row_weak: assumes Linvw: "LLL_invariant_weak' i fs" and i: "i < m" and j: "j < i" and fs': "fs' = fs[ i := fs ! i - c \\<^sub>v fs ! j]" shows "LLL_invariant_weak' i fs'" "g_bnd B fs \ g_bnd B fs'" proof (atomize(full), goal_cases) case 1 note Linv = LLL_invw'_imp_w[OF Linvw] note main = basis_reduction_add_row_main[OF Linv i j fs'] have bnd: "g_bnd B fs \ g_bnd B fs'" using main(6) unfolding g_bnd_def by auto note new = LLL_inv_wD[OF main(1)] note old = LLL_invD_weak[OF Linvw] have red: "weakly_reduced fs' i" using \weakly_reduced fs i\ main(6) \i < m\ unfolding gram_schmidt_fs.weakly_reduced_def by auto have inv: "LLL_invariant_weak' i fs'" using LLL_inv_wD[OF main(1)] \i < m\ by (intro LLL_invI_weak, auto intro: red) show ?case using inv red main bnd by auto qed lemma LLL_inv_weak_m_impl_i: assumes inv: "LLL_invariant_weak' m fs" and i: "i \ m" shows "LLL_invariant_weak' i fs" proof - have "weakly_reduced fs i" using LLL_invD_weak(8)[OF inv] by (meson assms(2) gram_schmidt_fs.weakly_reduced_def le_trans less_imp_le_nat linorder_not_less) then show ?thesis using LLL_invI_weak[of fs i, OF LLL_invD_weak(3,6,7)[OF inv] _ LLL_invD_weak(1)[OF inv]] LLL_invD_weak(2,4,5,8-)[OF inv] i by simp qed definition mod_invariant where "mod_invariant b p first = (b \ rat_of_int (p - 1)^2 / (rat_of_nat (bound_number first) + 3) \ (\ e. p = log_base ^ e))" lemma compute_mod_of_max_gso_norm: assumes mn: "mn \ 0" and m: "m = 0 \ mn = 0" and p: "p = compute_mod_of_max_gso_norm first mn" shows "p > 1" "mod_invariant mn p first" proof - let ?m = "bound_number first" define p' where "p' = root_rat_ceiling 2 (mn * (rat_of_nat ?m + 3)) + 1" define p'' where "p'' = max 2 p'" define q where "q = real_of_rat (mn * (rat_of_nat ?m + 3))" have *: "-1 < (0 :: real)" by simp also have "0 \ root 2 (real_of_rat (mn * (rat_of_nat ?m + 3)))" using mn by auto finally have "p' \ 0 + 1" unfolding p'_def by (intro plus_left_mono, simp) hence p': "p' > 0" by auto have p'': "p'' > 1" unfolding p''_def by auto have pp'': "p \ p''" unfolding compute_mod_of_max_gso_norm_def p p'_def[symmetric] p''_def[symmetric] using log_base p'' log_ceiling_sound by auto hence pp': "p \ p'" unfolding p''_def by auto show "p > 1" using pp'' p'' by auto have q0: "q \ 0" unfolding q_def using mn m by auto have "(mn \ rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)) = (real_of_rat mn \ real_of_rat (rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)))" using of_rat_less_eq by blast also have "\ = (real_of_rat mn \ real_of_rat (rat_of_int (p' - 1)^2) / real_of_rat (rat_of_nat ?m + 3))" by (simp add: of_rat_divide) also have "\ = (real_of_rat mn \ ((real_of_int (p' - 1))^2) / real_of_rat (rat_of_nat ?m + 3))" by (metis of_rat_of_int_eq of_rat_power) also have "\ = (real_of_rat mn \ (real_of_int \sqrt q\)^2 / real_of_rat (rat_of_nat ?m + 3))" unfolding p'_def sqrt_def q_def by simp also have "\" proof - have "real_of_rat mn \ q / real_of_rat (rat_of_nat ?m + 3)" unfolding q_def using m by (auto simp: of_rat_mult) also have "\ \ (real_of_int \sqrt q\)^2 / real_of_rat (rat_of_nat ?m + 3)" proof (rule divide_right_mono) have "q = (sqrt q)^2" using q0 by simp also have "\ \ (real_of_int \sqrt q\)^2" by (rule power_mono, auto simp: q0) finally show "q \ (real_of_int \sqrt q\)^2" . qed auto finally show ?thesis . qed finally have "mn \ rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)" . also have "\ \ rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)" unfolding power2_eq_square by (intro divide_right_mono mult_mono, insert p' pp', auto) finally have "mn \ rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)" . moreover have "\ e. p = log_base ^ e" unfolding p compute_mod_of_max_gso_norm_def by auto ultimately show "mod_invariant mn p first" unfolding mod_invariant_def by auto qed lemma g_bnd_mode_cong: assumes "\ i. i < m \ gso fs i = gso fs' i" shows "g_bnd_mode first b fs = g_bnd_mode first b fs'" using assms unfolding g_bnd_mode_def g_bnd_def by auto definition LLL_invariant_mod :: "int vec list \ int vec list \ int mat \ int \ bool \ rat \ nat \ bool" where "LLL_invariant_mod fs mfs dmu p first b i = ( length fs = m \ length mfs = m \ i \ m \ lattice_of fs = L \ gs.lin_indpt_list (RAT fs) \ weakly_reduced fs i \ (map (map_vec (\x. x symmod p)) fs = mfs) \ (\i' < m. \ j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j')) \ (\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j')) \ p > 1 \ g_bnd_mode first b fs \ mod_invariant b p first )" lemma LLL_invD_mod: assumes "LLL_invariant_mod fs mfs dmu p first b i" shows "length mfs = m" "i \ m" "length fs = m" "lattice_of fs = L" "gs.lin_indpt_list (RAT fs)" "weakly_reduced fs i" "(map (map_vec (\x. x symmod p)) fs = mfs)" "(\i' < m. \j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j'))" "(\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j'))" "\ i. i < m \ fs ! i \ carrier_vec n" "set fs \ carrier_vec n" "\ i. i < m \ gso fs i \ carrier_vec n" "\ i. i < m \ mfs ! i \ carrier_vec n" "set mfs \ carrier_vec n" "p > 1" "g_bnd_mode first b fs" "mod_invariant b p first" proof (atomize (full), goal_cases) case 1 interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs" using assms LLL_invariant_mod_def gs.lin_indpt_list_def by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro) have allfs: "\i < m. fs ! i \ carrier_vec n" using assms gs'.f_carrier by (simp add: LLL.LLL_invariant_mod_def) then have setfs: "set fs \ carrier_vec n" by (metis LLL_invariant_mod_def assms in_set_conv_nth subsetI) have allgso: "(\i < m. gso fs i \ carrier_vec n)" using assms gs'.gso_carrier by (simp add: LLL.LLL_invariant_mod_def) show ?case using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso LLL_invariant_mod_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce qed lemma LLL_invI_mod: assumes "length mfs = m" "i \ m" "length fs = m" "lattice_of fs = L" "gs.lin_indpt_list (RAT fs)" "weakly_reduced fs i" "map (map_vec (\x. x symmod p)) fs = mfs" "(\i' < m. \j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j'))" "(\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j'))" "p > 1" "g_bnd_mode first b fs" "mod_invariant b p first" shows "LLL_invariant_mod fs mfs dmu p first b i" unfolding LLL_invariant_mod_def using assms by blast definition LLL_invariant_mod_weak :: "int vec list \ int vec list \ int mat \ int \ bool \ rat \ bool" where "LLL_invariant_mod_weak fs mfs dmu p first b = ( length fs = m \ length mfs = m \ lattice_of fs = L \ gs.lin_indpt_list (RAT fs) \ (map (map_vec (\x. x symmod p)) fs = mfs) \ (\i' < m. \ j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j')) \ (\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j')) \ p > 1 \ g_bnd_mode first b fs \ mod_invariant b p first )" lemma LLL_invD_modw: assumes "LLL_invariant_mod_weak fs mfs dmu p first b" shows "length mfs = m" "length fs = m" "lattice_of fs = L" "gs.lin_indpt_list (RAT fs)" "(map (map_vec (\x. x symmod p)) fs = mfs)" "(\i' < m. \j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j'))" "(\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j'))" "\ i. i < m \ fs ! i \ carrier_vec n" "set fs \ carrier_vec n" "\ i. i < m \ gso fs i \ carrier_vec n" "\ i. i < m \ mfs ! i \ carrier_vec n" "set mfs \ carrier_vec n" "p > 1" "g_bnd_mode first b fs" "mod_invariant b p first" proof (atomize (full), goal_cases) case 1 interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs" using assms LLL_invariant_mod_weak_def gs.lin_indpt_list_def by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro) have allfs: "\i < m. fs ! i \ carrier_vec n" using assms gs'.f_carrier by (simp add: LLL.LLL_invariant_mod_weak_def) then have setfs: "set fs \ carrier_vec n" by (metis LLL_invariant_mod_weak_def assms in_set_conv_nth subsetI) have allgso: "(\i < m. gso fs i \ carrier_vec n)" using assms gs'.gso_carrier by (simp add: LLL.LLL_invariant_mod_weak_def) show ?case using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso LLL_invariant_mod_weak_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce qed lemma LLL_invI_modw: assumes "length mfs = m" "length fs = m" "lattice_of fs = L" "gs.lin_indpt_list (RAT fs)" "map (map_vec (\x. x symmod p)) fs = mfs" "(\i' < m. \j' < i'. \d\ fs i' j'\ < p * d fs j' * d fs (Suc j'))" "(\i' < m. \j' < m. d\ fs i' j' = dmu $$ (i',j'))" "p > 1" "g_bnd_mode first b fs" "mod_invariant b p first" shows "LLL_invariant_mod_weak fs mfs dmu p first b" unfolding LLL_invariant_mod_weak_def using assms by blast lemma dd\: assumes i: "i < m" shows "d fs (Suc i) = d\ fs i i" proof- have "\ fs i i = 1" using i by (simp add: gram_schmidt_fs.\.simps) then show ?thesis using d\_def by simp qed lemma d_of_main: assumes "(\i' < m. d\ fs i' i' = dmu $$ (i',i'))" and "i \ m" shows "d_of dmu i = d fs i" proof (cases "i = 0") case False with assms have "i - 1 < m" by auto from assms(1)[rule_format, OF this] dd\[OF this, of fs] False show ?thesis by (simp add: d_of_def) next case True thus ?thesis unfolding d_of_def True d_def by simp qed lemma d_of: assumes inv: "LLL_invariant_mod fs mfs dmu p b first j" and "i \ m" shows "d_of dmu i = d fs i" by (rule d_of_main[OF _ assms(2)], insert LLL_invD_mod(9)[OF inv], auto) lemma d_of_weak: assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b" and "i \ m" shows "d_of dmu i = d fs i" by (rule d_of_main[OF _ assms(2)], insert LLL_invD_modw(7)[OF inv], auto) lemma compute_max_gso_norm: assumes dmu: "(\i' < m. d\ fs i' i' = dmu $$ (i',i'))" and Linv: "LLL_invariant_weak fs" shows "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs" "fst (compute_max_gso_norm first dmu) \ 0" "m = 0 \ fst (compute_max_gso_norm first dmu) = 0" proof - show gbnd: "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs" proof (cases "first \ m \ 0") case False have "?thesis = (g_bnd (fst (compute_max_gso_norm first dmu)) fs)" unfolding g_bnd_mode_def using False by auto also have \ unfolding g_bnd_def proof (intro allI impI) fix i assume i: "i < m" have id: "(if first then 1 else m) = m" using False i by auto define list where "list = map (\ i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< m ]" obtain num denom j where ml: "max_list_rats_with_index list = (num, denom, j)" by (metis prod_cases3) have dpos: "d fs i > 0" using LLL_d_pos[OF Linv, of i] i by auto have pos: "(n, d, i) \ set list \ 0 < d" for n d i using LLL_d_pos[OF Linv] unfolding list_def using d_of_main[OF dmu] by auto from i have "list ! i \ set list" using i unfolding list_def by auto also have "list ! i = (d_of dmu (Suc i), d_of dmu i, i)" unfolding list_def using i by auto also have "\ = (d fs (Suc i), d fs i, i)" using d_of_main[OF dmu] i by auto finally have "(d fs (Suc i), d fs i, i) \ set list" . from max_list_rats_with_index[OF pos ml this] have "of_int (d fs (Suc i)) / of_int (d fs i) \ fst (compute_max_gso_norm first dmu)" unfolding compute_max_gso_norm_def list_def[symmetric] ml id split using i by auto also have "of_int (d fs (Suc i)) / of_int (d fs i) = sq_norm (gso fs i)" using LLL_d_Suc[OF Linv i] dpos by auto finally show "sq_norm (gso fs i) \ fst (compute_max_gso_norm first dmu)" . qed finally show ?thesis . next case True thus ?thesis unfolding g_bnd_mode_def compute_max_gso_norm_def using d_of_main[OF dmu] LLL_d_Suc[OF Linv, of 0] LLL_d_pos[OF Linv, of 0] LLL_d_pos[OF Linv, of 1] by auto qed show "fst (compute_max_gso_norm first dmu) \ 0" proof (cases "m = 0") case True thus ?thesis unfolding compute_max_gso_norm_def by simp next case False hence 0: "0 < m" by simp have "0 \ sq_norm (gso fs 0)" by blast also have "\ \ fst (compute_max_gso_norm first dmu)" using gbnd[unfolded g_bnd_mode_def g_bnd_def] using 0 by metis finally show ?thesis . qed qed (auto simp: LLL.compute_max_gso_norm_def) lemma increase_i_mod: assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i" and i: "i < m" and red_i: "i \ 0 \ sq_norm (gso fs (i - 1)) \ \ * sq_norm (gso fs i)" shows "LLL_invariant_mod fs mfs dmu p first b (Suc i)" "LLL_measure i fs > LLL_measure (Suc i) fs" proof - note inv = LLL_invD_mod[OF Linv] from inv have red: "weakly_reduced fs i" by (auto) from red red_i i have red: "weakly_reduced fs (Suc i)" unfolding gram_schmidt_fs.weakly_reduced_def by (intro allI impI, rename_tac ii, case_tac "Suc ii = i", auto) show "LLL_invariant_mod fs mfs dmu p first b (Suc i)" by (intro LLL_invI_mod, insert inv red i, auto) show "LLL_measure i fs > LLL_measure (Suc i) fs" unfolding LLL_measure_def using i by auto qed lemma basis_reduction_mod_add_row_main: assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b" and i: "i < m" and j: "j < i" and c: "c = round (\ fs i j)" and mfs': "mfs' = mfs[ i := (map_vec (\ x. x symmod p)) (mfs ! i - c \\<^sub>v mfs ! j)]" and dmu': "dmu' = mat m m (\(i',j'). (if (i' = i \ j' \ j) then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j')) else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))) else (dmu $$ (i',j'))))" shows "(\fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \ LLL_measure i fs' = LLL_measure i fs \ (\_small_row i fs (Suc j) \ \_small_row i fs' j) \ (\k < m. gso fs' k = gso fs k) \ (\ii \ m. d fs' ii = d fs ii) \ \\ fs' i j\ \ 1 / 2 \ (\i' j'. i' < i \ j' \ i' \ \ fs' i' j' = \ fs i' j') \ (LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs' mfs' dmu' p first b i))" proof - define fs' where "fs' = fs[ i := fs ! i - c \\<^sub>v fs ! j]" from LLL_invD_modw[OF Linvmw] have gbnd: "g_bnd_mode first b fs" and p1: "p > 1" and pgtz: "p > 0" by auto have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp have Linvw': "LLL_invariant_weak fs'" and 01: "c = round (\ fs i j) \ \_small_row i fs (Suc j) \ \_small_row i fs' j" and 02: "LLL_measure i fs' = LLL_measure i fs" and 03: "\ i. i < m \ gso fs' i = gso fs i" and 04: "\ i' j'. i' < m \ j' < m \ \ fs' i' j' = (if i' = i \ j' \ j then \ fs i j' - of_int c * \ fs j j' else \ fs i' j')" and 05: "\ ii. ii \ m \ d fs' ii = d fs ii" and 06: "\\ fs' i j\ \ 1 / 2" and 061: "(\i' j'. i' < i \ j' \ i' \ \ fs i' j' = \ fs' i' j')" using basis_reduction_add_row_main[OF Linvww i j fs'_def] c i by auto have 07: "lin_indep fs'" and 08: "length fs' = m" and 09: "lattice_of fs' = L" using LLL_inv_wD Linvw' by auto have 091: "fs_int_indpt n fs'" using 07 using Gram_Schmidt_2.fs_int_indpt.intro by simp define I where "I = {(i',j'). i' = i \ j' < j}" have 10: "I \ {(i',j'). i' < m \ j' < i'}" "(i,j)\ I" "\j' \ j. (i,j') \ I" using I_def i j by auto obtain fs'' where 11: "lattice_of fs'' = L" and 12: "map (map_vec (\ x. x symmod p)) fs'' = map (map_vec (\ x. x symmod p)) fs'" and 13: "lin_indep fs''" and 14: "length fs'' = m" and 15: "(\ k < m. gso fs'' k = gso fs' k)" and 16: "(\ k \ m. d fs'' k = d fs' k)" and 17: "(\ i' < m. \ j' < m. d\ fs'' i' j' = (if (i',j') \ I then d\ fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\ fs' i' j'))" using mod_finite_set[OF 07 08 10(1) 09 pgtz] by blast have 171: "(\i' j'. i' < i \ j' \ i' \ \ fs'' i' j' = \ fs' i' j')" proof - { fix i' j' assume i'j': "i' < i" "j' \ i'" have "rat_of_int (d\ fs'' i' j') = rat_of_int (d\ fs' i' j')" using "17" I_def i i'j' by auto then have "rat_of_int (int_of_rat (rat_of_int (d fs'' (Suc j')) * \ fs'' i' j')) = rat_of_int (int_of_rat (rat_of_int (d fs' (Suc j')) * \ fs' i' j'))" using d\_def i'j' j by auto then have "rat_of_int (d fs'' (Suc j')) * \ fs'' i' j' = rat_of_int (d fs' (Suc j')) * \ fs' i' j'" by (smt "08" "091" "13" "14" d_def dual_order.strict_trans fs_int.d_def fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro i i'j'(1) i'j'(2) int_of_rat(2)) then have "\ fs'' i' j' = \ fs' i' j'" by (smt "16" LLL_d_pos[OF Linvw'] Suc_leI int_of_rat(1) dual_order.strict_trans fs'_def i i'j' j le_neq_implies_less nonzero_mult_div_cancel_left of_int_hom.hom_zero) } then show ?thesis by simp qed then have 172: "(\i' j'. i' < i \ j' \ i' \ \ fs'' i' j' = \ fs i' j')" using 061 by simp (* goal *) have 18: "LLL_measure i fs'' = LLL_measure i fs'" using 16 LLL_measure_def logD_def D_def by simp have 19: "(\k < m. gso fs'' k = gso fs k)" using 03 15 by simp have "\j' \ {j..(m-1)}. j' < m" using j i by auto then have 20: "\j' \ {j..(m-1)}. d\ fs'' i j' = d\ fs' i j'" using 10(3) 17 Suc_lessD less_trans_Suc by (meson atLeastAtMost_iff i) have 21: "\j' \ {j..(m-1)}. \ fs'' i j' = \ fs' i j'" proof - { fix j' assume j': "j' \ {j..(m-1)}" define \'' :: rat where "\'' = \ fs'' i j'" define \' :: rat where "\' = \ fs' i j'" have "rat_of_int (d\ fs'' i j') = rat_of_int (d\ fs' i j')" using 20 j' by simp moreover have "j' < length fs'" using i j' 08 by auto ultimately have "rat_of_int (d fs' (Suc j')) * gram_schmidt_fs.\ n (map of_int_hom.vec_hom fs') i j' = rat_of_int (d fs'' (Suc j')) * gram_schmidt_fs.\ n (map of_int_hom.vec_hom fs'') i j'" using 20 08 091 13 14 fs_int_indpt.d\_def fs_int.d_def fs_int_indpt.d\ d\_def d_def i fs_int_indpt.intro j' by metis then have "rat_of_int (d fs' (Suc j')) * \'' = rat_of_int (d fs' (Suc j')) * \'" using 16 i j' \'_def \''_def unfolding d\_def by auto moreover have "0 < d fs' (Suc j')" using LLL_d_pos[OF Linvw', of "Suc j'"] i j' by auto ultimately have "\ fs'' i j' = \ fs' i j'" using \'_def \''_def by simp } then show ?thesis by simp qed then have 22: "\ fs'' i j = \ fs' i j" using i j by simp then have 23: "\\ fs'' i j\ \ 1 / 2" using 06 by simp (* goal *) have 24: "LLL_measure i fs'' = LLL_measure i fs" using 02 18 by simp (* goal *) have 25: "(\ k \ m. d fs'' k = d fs k)" using 16 05 by simp (* goal *) have 26: "(\ k < m. gso fs'' k = gso fs k)" using 15 03 by simp (* goal *) have 27: "\_small_row i fs (Suc j) \ \_small_row i fs'' j" using 21 01 \_small_row_def i j c by auto (* goal *) have 28: "length fs = m" "length mfs = m" using LLL_invD_modw[OF Linvmw] by auto have 29: "map (map_vec (\x. x symmod p)) fs = mfs" using assms LLL_invD_modw by simp have 30: "\ i. i < m \ fs ! i \ carrier_vec n" "\ i. i < m \ mfs ! i \ carrier_vec n" using LLL_invD_modw[OF Linvmw] by auto have 31: "\ i. i < m \ fs' ! i \ carrier_vec n" using fs'_def 30(1) using "08" "091" fs_int_indpt.f_carrier by blast have 32: "\ i. i < m \ mfs' ! i \ carrier_vec n" unfolding mfs' using 30(2) 28(2) by (metis (no_types, lifting) Suc_lessD j less_trans_Suc map_carrier_vec minus_carrier_vec nth_list_update_eq nth_list_update_neq smult_closed) have 33: "length mfs' = m" using 28(2) mfs' by simp (* invariant goal *) then have 34: "map (map_vec (\x. x symmod p)) fs' = mfs'" proof - { fix i' j' have j2: "j < m" using j i by auto assume i': "i' < m" assume j': "j' < n" then have fsij: "(fs ! i' $ j') symmod p = mfs ! i' $ j'" using 30 i' j' 28 29 by fastforce have "mfs' ! i $ j' = (mfs ! i $ j'- (c \\<^sub>v mfs ! j) $ j') symmod p" unfolding mfs' using 30(2) j' 28 j2 by (metis (no_types, lifting) carrier_vecD i index_map_vec(1) index_minus_vec(1) index_minus_vec(2) index_smult_vec(2) nth_list_update_eq) then have mfs'ij: "mfs' ! i $ j' = (mfs ! i $ j'- c * mfs ! j $ j') symmod p" unfolding mfs' using 30(2) i' j' 28 j2 by fastforce have "(fs' ! i' $ j') symmod p = mfs' ! i' $ j'" proof(cases "i' = i") case True show ?thesis using fs'_def mfs' True 28 fsij proof - have "fs' ! i' $ j' = (fs ! i' - c \\<^sub>v fs ! j) $ j'" using fs'_def True i' j' 28(1) by simp also have "\ = fs ! i' $ j' - (c \\<^sub>v fs ! j) $ j'" using i' j' 30(1) by (metis Suc_lessD carrier_vecD i index_minus_vec(1) index_smult_vec(2) j less_trans_Suc) finally have "fs' ! i' $ j' = fs ! i' $ j' - (c \\<^sub>v fs ! j) $ j'" by auto then have "(fs' ! i' $ j') symmod p = (fs ! i' $ j' - (c \\<^sub>v fs ! j) $ j') symmod p" by auto also have "\ = ((fs ! i' $ j') symmod p - ((c \\<^sub>v fs ! j) $ j') symmod p) symmod p" by (simp add: sym_mod_diff_eq) also have "(c \\<^sub>v fs ! j) $ j' = c * (fs ! j $ j')" using i' j' True 28 30(1) j by (metis Suc_lessD carrier_vecD index_smult_vec(1) less_trans_Suc) also have "((fs ! i' $ j') symmod p - (c * (fs ! j $ j')) symmod p) symmod p = ((fs ! i' $ j') symmod p - c * ((fs ! j $ j') symmod p)) symmod p" using i' j' True 28 30(1) j by (metis sym_mod_diff_right_eq sym_mod_mult_right_eq) also have "((fs ! j $ j') symmod p) = mfs ! j $ j'" using 30 i' j' 28 29 j2 by fastforce also have "((fs ! i' $ j') symmod p - c * mfs ! j $ j') symmod p = (mfs ! i' $ j' - c * mfs ! j $ j') symmod p" using fsij by simp finally show ?thesis using mfs'ij by (simp add: True) qed next case False show ?thesis using fs'_def mfs' False 28 fsij by simp qed } then have "\i' < m. (map_vec (\x. x symmod p)) (fs' ! i') = mfs' ! i'" using 31 32 33 08 by fastforce then show ?thesis using 31 32 33 08 by (simp add: map_nth_eq_conv) qed then have 35: "map (map_vec (\x. x symmod p)) fs'' = mfs'" using 12 by simp (* invariant req. *) have 36: "lin_indep fs''" using 13 by simp (* invariant req. *) have Linvw'': "LLL_invariant_weak fs''" using LLL_invariant_weak_def 11 13 14 by simp have 39: "(\i' < m. \j' < i'. \d\ fs'' i' j'\ < p * d fs'' j' * d fs'' (Suc j'))" (* invariant req. *) proof - { fix i' j' assume i': "i' < m" assume j': "j' < i'" define pdd where "pdd = (p * d fs'' j' * d fs'' (Suc j'))" then have pddgtz: "pdd > 0" using pgtz j' LLL_d_pos[OF Linvw', of "Suc j'"] LLL_d_pos[OF Linvw', of j'] j' i' 16 by simp have "\d\ fs'' i' j'\ < p * d fs'' j' * d fs'' (Suc j')" proof(cases "i' = i") case i'i: True then show ?thesis proof (cases "j' < j") case True then have eq'': "d\ fs'' i' j' = d\ fs' i' j' symmod (p * d fs'' j' * d fs'' (Suc j'))" using 16 17 10 I_def True i' j' i'i by simp have "0 < pdd" using pddgtz by simp then show ?thesis unfolding eq'' unfolding pdd_def[symmetric] using sym_mod_abs by blast next case fls: False then have "(i',j') \ I" using I_def i'i by simp then have dmufs''fs': "d\ fs'' i' j' = d\ fs' i' j'" using 17 i' j' by simp show ?thesis proof (cases "j' = j") case True define \'' where "\'' = \ fs'' i' j'" define d'' where "d'' = d fs'' (Suc j')" have pge1: "p \ 1" using pgtz by simp have lh: "\\''\ \ 1 / 2" using 23 True i'i \''_def by simp moreover have eq: "d\ fs'' i' j' = \'' * d''" using d\_def i' j' \''_def d''_def by (smt "14" "36" LLL.d_def Suc_lessD fs_int.d_def fs_int_indpt.d\ fs_int_indpt.intro int_of_rat(1) less_trans_Suc mult_of_int_commute of_rat_mult of_rat_of_int_eq) moreover have Sj': "Suc j' \ m" "j' \ m" using True j' i i' by auto moreover then have gtz: "0 < d''" using LLL_d_pos[OF Linvw''] d''_def by simp moreover have "rat_of_int \d\ fs'' i' j'\ = \\'' * (rat_of_int d'')\" using eq by (metis of_int_abs of_rat_hom.injectivity of_rat_mult of_rat_of_int_eq) moreover then have "\\'' * rat_of_int d'' \ = \\''\ * rat_of_int \d''\" by (metis (mono_tags, opaque_lifting) abs_mult of_int_abs) moreover have "\ = \\''\ * rat_of_int d'' " using gtz by simp moreover have "\ < rat_of_int d''" using lh gtz by simp ultimately have "rat_of_int \d\ fs'' i' j'\ < rat_of_int d''" by simp then have "\d\ fs'' i' j'\ < d fs'' (Suc j')" using d''_def by simp then have "\d\ fs'' i' j'\ < p * d fs'' (Suc j')" using pge1 by (smt mult_less_cancel_right2) then show ?thesis using pge1 LLL_d_pos[OF Linvw'' Sj'(2)] gtz unfolding d''_def by (smt mult_less_cancel_left2 mult_right_less_imp_less) next case False have "j' < m" using i' j' by simp moreover have "j' > j" using False fls by simp ultimately have "\ fs' i' j' = \ fs i' j'" using i' 04 i by simp then have "d\ fs' i' j' = d\ fs i' j'" using d\_def i' j' 05 by simp then have "d\ fs'' i' j' = d\ fs i' j'" using dmufs''fs' by simp then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp qed qed next case False then have "(i',j') \ I" using I_def by simp then have dmufs''fs': "d\ fs'' i' j' = d\ fs' i' j'" using 17 i' j' by simp have "\ fs' i' j' = \ fs i' j'" using i' 04 j' False by simp then have "d\ fs' i' j' = d\ fs i' j'" using d\_def i' j' 05 by simp moreover then have "d\ fs'' i' j' = d\ fs i' j'" using dmufs''fs' by simp then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp qed } then show ?thesis by simp qed have 40: "(\i' < m. \j' < m. i' \ i \ j' > j \ d\ fs' i' j' = dmu $$ (i',j'))" proof - { fix i' j' assume i': "i' < m" and j': "j' < m" assume assm: "i' \ i \ j' > j" have "d\ fs' i' j' = dmu $$ (i',j')" proof (cases "i' \ i") case True then show ?thesis using fs'_def LLL_invD_modw[OF Linvmw] d\_def i i' j j' 04 28(1) LLL_invI_weak basis_reduction_add_row_main(8)[OF Linvww] by auto next case False then show ?thesis using 05 LLL_invD_modw[OF Linvmw] d\_def i j j' 04 assm by simp qed } then show ?thesis by simp qed have 41: "\j' \ j. d\ fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')" proof - { let ?oi = "of_int :: _ \ rat" fix j' assume j': "j' \ j" define dj' \i \j where "dj' = d fs (Suc j')" and "\i = \ fs i j'" and "\j = \ fs j j'" have "?oi (d\ fs' i j') = ?oi (d fs (Suc j')) * (\ fs i j' - ?oi c * \ fs j j')" using j' 04 d\_def by (smt "05" "08" "091" Suc_leI d_def diff_diff_cancel fs_int.d_def fs_int_indpt.fs_int_mu_d_Z i int_of_rat(2) j less_imp_diff_less less_imp_le_nat) also have "\ = (?oi dj') * (\i - of_int c * \j)" using dj'_def \i_def \j_def by (simp add: of_rat_mult) also have "\ = (rat_of_int dj') * \i - of_int c * (rat_of_int dj') * \j" by algebra also have "\ = rat_of_int (d\ fs i j') - ?oi c * rat_of_int (d\ fs j j')" unfolding dj'_def \i_def \j_def using i j j' d\_def using "28"(1) LLL.LLL_invD_modw(4) Linvmw d_def fs_int.d_def fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro by auto also have "\ = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))" using LLL_invD_modw(7)[OF Linvmw] d\_def j' i j by auto finally have "?oi (d\ fs' i j') = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))" by simp then have "d\ fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')" using of_int_eq_iff by fastforce } then show ?thesis by simp qed have 42: "(\i' < m. \j' < m. d\ fs'' i' j' = dmu' $$ (i',j'))" proof - { fix i' j' assume i': "i' < m" and j': "j' < m" have "d\ fs'' i' j' = dmu' $$ (i',j')" proof (cases "i' = i") case i'i: True then show ?thesis proof (cases "j' > j") case True then have "(i',j')\I" using I_def by simp moreover then have "d\ fs' i' j' = d\ fs i' j'" using "04" "05" True Suc_leI d\_def i' j' by simp moreover have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' True i' j' by simp ultimately show ?thesis using "17" "40" True i' j' by auto next case False then have j'lej: "j' \ j" by simp then have eq': "d\ fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')" using 41 by simp have id: "d_of dmu j' = d fs j'" "d_of dmu (Suc j') = d fs (Suc j')" using d_of_weak[OF Linvmw] \j' < m\ by auto show ?thesis proof (cases "j' \ j") case True then have j'ltj: "j' < j" using True False by simp then have "(i',j') \ I" using I_def True i'i by simp then have "d\ fs'' i' j' = (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs' j' * d fs' (Suc j'))" using 17 i' 41 j'lej by (simp add: j' i'i) also have "\ = (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs j' * d fs (Suc j'))" using 05 i j'ltj j by simp also have "\ = dmu' $$ (i,j')" unfolding dmu' index_mat(1)[OF \i < m\ \j' < m\] split id using j'lej True by auto finally show ?thesis using i'i by simp next case False then have j'j: "j' = j" by simp then have "d\ fs'' i j' = d\ fs' i j'" using 20 j' by simp also have "\ = dmu $$ (i,j') - c * dmu $$ (j,j')" using eq' by simp also have "\ = dmu' $$ (i,j')" using dmu' j'j i j' by simp finally show ?thesis using i'i by simp qed qed next case False then have "(i',j')\I" using I_def by simp moreover then have "d\ fs' i' j' = d\ fs i' j'" by (simp add: "04" "05" False Suc_leI d\_def i' j') moreover then have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' False i' j' by simp ultimately show ?thesis using "17" "40" False i' j' by auto qed } then show ?thesis by simp qed from gbnd 26 have gbnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs'' fs] by simp { assume Linv: "LLL_invariant_mod fs mfs dmu p first b i" have Linvw: "LLL_invariant_weak' i fs" using Linv LLL_invD_mod LLL_invI_weak by simp note Linvww = LLL_invw'_imp_w[OF Linvw] have 00: "LLL_invariant_weak' i fs'" using Linvw basis_reduction_add_row_weak[OF Linvw i j fs'_def] by auto have 37: "weakly_reduced fs'' i" using 15 LLL_invD_weak(8)[OF 00] gram_schmidt_fs.weakly_reduced_def by (smt Suc_lessD i less_trans_Suc) (* invariant req. *) have 38: "LLL_invariant_weak' i fs''" using 00 11 14 36 37 i 31 12 LLL_invariant_weak'_def by blast have "LLL_invariant_mod fs'' mfs' dmu' p first b i" using LLL_invI_mod[OF 33 _ 14 11 13 37 35 39 42 p1 gbnd LLL_invD_mod(17)[OF Linv]] i by simp } moreover have "LLL_invariant_mod_weak fs'' mfs' dmu' p first b" using LLL_invI_modw[OF 33 14 11 13 35 39 42 p1 gbnd LLL_invD_modw(15)[OF Linvmw]] by simp ultimately show ?thesis using 27 23 24 25 26 172 by auto qed definition D_mod :: "int mat \ nat" where "D_mod dmu = nat (\ i < m. d_of dmu i)" definition logD_mod :: "int mat \ nat" where "logD_mod dmu = (if \ = 4/3 then (D_mod dmu) else nat (floor (log (1 / of_rat reduction) (D_mod dmu))))" end locale fs_int'_mod = fixes n m fs_init \ i fs mfs dmu p first b assumes LLL_inv_mod: "LLL.LLL_invariant_mod n m fs_init \ fs mfs dmu p first b i" context LLL_with_assms begin lemma basis_reduction_swap_weak': assumes Linvw: "LLL_invariant_weak' i fs" and i: "i < m" and i0: "i \ 0" and mu_F1_i: "\\ fs i (i-1)\ \ 1 / 2" and norm_ineq: "sq_norm (gso fs (i - 1)) > \ * sq_norm (gso fs i)" and fs'_def: "fs' = fs[i := fs ! (i - 1), i - 1 := fs ! i]" shows "LLL_invariant_weak' (i - 1) fs'" proof - note inv = LLL_invD_weak[OF Linvw] note invw = LLL_invw'_imp_w[OF Linvw] note main = basis_reduction_swap_main[OF invw disjI2[OF mu_F1_i] i i0 norm_ineq fs'_def] note inv' = LLL_inv_wD[OF main(1)] from \weakly_reduced fs i\ have "weakly_reduced fs (i - 1)" unfolding gram_schmidt_fs.weakly_reduced_def by auto also have "weakly_reduced fs (i - 1) = weakly_reduced fs' (i - 1)" unfolding gram_schmidt_fs.weakly_reduced_def by (intro all_cong, insert i0 i main(5), auto) finally have red: "weakly_reduced fs' (i - 1)" . show "LLL_invariant_weak' (i - 1) fs'" using i by (intro LLL_invI_weak red inv', auto) qed lemma basis_reduction_add_row_done_weak: assumes Linv: "LLL_invariant_weak' i fs" and i: "i < m" and mu_small: "\_small_row i fs 0" shows "\_small fs i" proof - note inv = LLL_invD_weak[OF Linv] from mu_small have mu_small: "\_small fs i" unfolding \_small_row_def \_small_def by auto show ?thesis using i mu_small LLL_invI_weak[OF inv(3,6,7,9,1)] by auto qed lemma LLL_invariant_mod_to_weak_m_to_i: assumes inv: "LLL_invariant_mod fs mfs dmu p first b m" and i: "i \ m" shows "LLL_invariant_mod fs mfs dmu p first b i" "LLL_invariant_weak' m fs" "LLL_invariant_weak' i fs" proof - show "LLL_invariant_mod fs mfs dmu p first b i" proof - have "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp then have "LLL_invariant_weak' i fs" using LLL_inv_weak_m_impl_i i by simp then have "weakly_reduced fs i" using i LLL_invD_weak(8) by simp then show ?thesis using LLL_invD_mod[OF inv] LLL_invI_mod i by simp qed then show fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod LLL_invI_weak by simp show "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp qed lemma basis_reduction_mod_swap_main: assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b" and k: "k < m" and k0: "k \ 0" and mu_F1_i: "\\ fs k (k-1)\ \ 1 / 2" and norm_ineq: "sq_norm (gso fs (k - 1)) > \ * sq_norm (gso fs k)" and mfs'_def: "mfs' = mfs[k := mfs ! (k - 1), k - 1 := mfs ! k]" and dmu'_def: "dmu' = (mat m m (\(i,j). ( if j < i then if i = k - 1 then dmu $$ (k, j) else if i = k \ j \ k - 1 then dmu $$ (k - 1, j) else if i > k \ j = k then ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j)) div (d_of dmu k) else if i > k \ j = k - 1 then (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1))) div (d_of dmu k) else dmu $$ (i, j) else if i = j then if i = k - 1 then ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1)) div (d_of dmu k) else (d_of dmu (Suc i)) else dmu $$ (i, j)) ))" and dmu'_mod_def: "dmu'_mod = mat m m (\(i, j). ( if j < i \ (j = k \ j = k - 1) then dmu' $$ (i, j) symmod (p * (d_of dmu' j) * (d_of dmu' (Suc j))) else dmu' $$ (i, j)))" shows "(\fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \ LLL_measure (k-1) fs' < LLL_measure k fs \ (LLL_invariant_mod fs mfs dmu p first b k \ LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))" proof - define fs' where "fs' = fs[k := fs ! (k - 1), k - 1 := fs ! k]" have pgtz: "p > 0" and p1: "p > 1" using LLL_invD_modw[OF Linvmw] by auto have invw: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp note swap_main = basis_reduction_swap_main(3-)[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def] note dd\_swap = d_d\_swap[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def] have invw': "LLL_invariant_weak fs'" using fs'_def assms invw basis_reduction_swap_main(1) by simp have 02: "LLL_measure k fs > LLL_measure (k - 1) fs'" by fact have 03: "\ i j. i < m \ j < i \ d\ fs' i j = ( if i = k - 1 then d\ fs k j else if i = k \ j \ k - 1 then d\ fs (k - 1) j else if i > k \ j = k then (d fs (Suc k) * d\ fs i (k - 1) - d\ fs k (k - 1) * d\ fs i j) div d fs k else if i > k \ j = k - 1 then (d\ fs k (k - 1) * d\ fs i j + d\ fs i k * d fs (k - 1)) div d fs k else d\ fs i j)" using dd\_swap by auto have 031: "\i. i < k-1 \ gso fs' i = gso fs i" using swap_main(2) k k0 by auto have 032: "\ ii. ii \ m \ of_int (d fs' ii) = (if ii = k then sq_norm (gso fs' (k - 1)) / sq_norm (gso fs (k - 1)) * of_int (d fs k) else of_int (d fs ii))" by fact have gbnd: "g_bnd_mode first b fs'" proof (cases "first \ m \ 0") case True have "sq_norm (gso fs' 0) \ sq_norm (gso fs 0)" proof (cases "k - 1 = 0") case False thus ?thesis using 031[of 0] by simp next case *: True have k_1: "k - 1 < m" using k by auto from * k0 have k1: "k = 1" by simp (* this is a copy of what is done in LLL.swap-main, should be made accessible in swap-main *) have "sq_norm (gso fs' 0) \ abs (sq_norm (gso fs' 0))" by simp also have "\ = abs (sq_norm (gso fs 1) + \ fs 1 0 * \ fs 1 0 * sq_norm (gso fs 0))" by (subst swap_main(3)[OF k_1, unfolded *], auto simp: k1) also have "\ \ sq_norm (gso fs 1) + abs (\ fs 1 0) * abs (\ fs 1 0) * sq_norm (gso fs 0)" by (simp add: sq_norm_vec_ge_0) also have "\ \ sq_norm (gso fs 1) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)" using mu_F1_i[unfolded k1] by (intro plus_right_mono mult_mono, auto) also have "\ < 1 / \ * sq_norm (gso fs 0) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)" by (intro add_strict_right_mono, insert norm_ineq[unfolded mult.commute[of \], THEN mult_imp_less_div_pos[OF \0(1)]] k1, auto) also have "\ = reduction * sq_norm (gso fs 0)" unfolding reduction_def using \0 by (simp add: ring_distribs add_divide_distrib) also have "\ \ 1 * sq_norm (gso fs 0)" using reduction(2) by (intro mult_right_mono, auto) finally show ?thesis by simp qed thus ?thesis using LLL_invD_modw(14)[OF Linvmw] True unfolding g_bnd_mode_def by auto next case False from LLL_invD_modw(14)[OF Linvmw] False have "g_bnd b fs" unfolding g_bnd_mode_def by auto hence "g_bnd b fs'" using g_bnd_swap[OF k k0 invw mu_F1_i norm_ineq fs'_def] by simp thus ?thesis using False unfolding g_bnd_mode_def by auto qed note d_of = d_of_weak[OF Linvmw] have 033: "\ i. i < m \ d\ fs' i i = ( if i = k - 1 then ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1)) div (d_of dmu k) else (d_of dmu (Suc i)))" proof - fix i assume i: "i < m" have "d\ fs' i i = d fs' (Suc i)" using dd\ i by simp also have "\ = (if i = k - 1 then (d fs (Suc k) * d fs (k - 1) + d\ fs k (k - 1) * d\ fs k (k - 1)) div d fs k else d fs (Suc i))" by (subst dd\_swap, insert dd\ k0 i, auto) also have "\ = (if i = k - 1 then ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1)) div (d_of dmu k) else (d_of dmu (Suc i)))" (is "_ = ?r") using d_of i k LLL_invD_modw(7)[OF Linvmw] by auto finally show "d\ fs' i i = ?r" . qed have 04: "lin_indep fs'" "length fs' = m" "lattice_of fs' = L" using LLL_inv_wD[OF invw'] by auto define I where "I = {(i, j). i < m \ j < i \ (j = k \ j = k - 1)}" then have Isubs: "I \ {(i,j). i < m \ j < i}" using k k0 by auto obtain fs'' where 05: "lattice_of fs'' = L" and 06: "map (map_vec (\ x. x symmod p)) fs'' = map (map_vec (\ x. x symmod p)) fs'" and 07: "lin_indep fs''" and 08: "length fs'' = m" and 09: "(\ k < m. gso fs'' k = gso fs' k)" and 10: "(\ k \ m. d fs'' k = d fs' k)" and 11: "(\ i' < m. \ j' < m. d\ fs'' i' j' = (if (i',j') \ I then d\ fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\ fs' i' j'))" using mod_finite_set[OF 04(1) 04(2) Isubs 04(3) pgtz] by blast have 13: "length mfs' = m" using mfs'_def LLL_invD_modw(1)[OF Linvmw] by simp (* invariant requirement *) have 14: "map (map_vec (\ x. x symmod p)) fs'' = mfs'" (* invariant requirement *) using 06 fs'_def k k0 04(2) LLL_invD_modw(5)[OF Linvmw] by (metis (no_types, lifting) length_list_update less_imp_diff_less map_update mfs'_def nth_map) have "LLL_measure (k - 1) fs'' = LLL_measure (k - 1) fs'" using 10 LLL_measure_def logD_def D_def by simp then have 15: "LLL_measure (k - 1) fs'' < LLL_measure k fs" using 02 by simp (* goal *) { fix i' j' assume i'j': "i' k" "j' \ k - 1" hence j'k: "j' \ k" "Suc j' \ k" using k0 by auto hence "d fs'' j' = d fs j'" "d fs'' (Suc j') = d fs (Suc j')" using \k < m\ i'j' k0 10[rule_format, of j'] 032[rule_format, of j'] 10[rule_format, of "Suc j'"] 032[rule_format, of "Suc j'"] by auto } note d_id = this have 16: "\i'j'd\ fs'' i' j'\ < p * d fs'' j' * d fs'' (Suc j')" (* invariant requirement *) proof - { fix i' j' assume i'j': "i'd\ fs'' i' j'\ < p * d fs'' j' * d fs'' (Suc j')" proof (cases "(i',j') \ I") case True define pdd where "pdd = (p * d fs' j' * d fs' (Suc j'))" have pdd_pos: "pdd > 0" using pgtz i'j' LLL_d_pos[OF invw'] pdd_def by simp have "d\ fs'' i' j' = d\ fs' i' j' symmod pdd" using True 11 i'j' pdd_def by simp then have "\d\ fs'' i' j'\ < pdd" using True 11 i'j' pdd_pos sym_mod_abs by simp then show ?thesis unfolding pdd_def using 10 i'j' by simp next case False from False[unfolded I_def] i'j' have neg: "j' \ k" "j' \ k - 1" by auto consider (1) "i' = k - 1 \ i' = k" | (2) "\ (i' = k - 1 \ i' = k)" using False i'j' unfolding I_def by linarith thus ?thesis proof cases case **: 1 let ?i'' = "if i' = k - 1 then k else k -1" from ** neg i'j' have i'': "?i'' < m" "j' < ?i''" using k0 k by auto have "d\ fs'' i' j' = d\ fs' i' j'" using 11 False i'j' by simp also have "\ = d\ fs ?i'' j'" unfolding 03[OF \i' < m\ \j' < i'\] using ** neg by auto finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i''] unfolding d_id[OF i'j' neg] by auto next case **: 2 hence neq: "j' \ k" "j' \ k - 1" using False k k0 i'j' unfolding I_def by auto have "d\ fs'' i' j' = d\ fs' i' j'" using 11 False i'j' by simp also have "\ = d\ fs i' j'" unfolding 03[OF \i' < m\ \j' < i'\] using ** neq by auto finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i'j'] using d_id[OF i'j' neq] by auto qed qed } then show ?thesis by simp qed have 17: "\i'j' fs'' i' j' = dmu'_mod $$ (i', j')" (* invariant requirement *) proof - { fix i' j' assume i'j': "i'j' < m. d fs' (Suc j') = dmu' $$ (j', j')" using dd\ dmu'_def 033 by simp have eq': "d\ fs' i' j' = dmu' $$ (i', j')" proof - have t00: "d\ fs k j' = dmu $$ (k, j')" and t01: "d\ fs (k - 1) j' = dmu $$ (k - 1, j')" and t04: "d\ fs k (k - 1) = dmu $$ (k, k - 1)" and t05: "d\ fs i' k = dmu $$ (i', k)" using LLL_invD_modw(7)[OF Linvmw] i'j' k dd\ k0 by auto have t03: "d fs k = d\ fs (k-1) (k-1)" using k0 k by (metis LLL.dd\ Suc_diff_1 lessI not_gr_zero) have t06: "d fs (k - 1) = (d_of dmu (k-1))" using d_of k by auto have t07: "d fs k = (d_of dmu k)" using d_of k by auto have j': "j' < m" using i'j' by simp have "d\ fs' i' j' = (if i' = k - 1 then dmu $$ (k, j') else if i' = k \ j' \ k - 1 then dmu $$ (k - 1, j') else if i' > k \ j' = k then (dmu $$ (k, k) * dmu $$ (i', k - 1) - dmu $$ (k, k - 1) * dmu $$ (i', j')) div (d_of dmu k) else if i' > k \ j' = k - 1 then (dmu $$ (k, k - 1) * dmu $$ (i', j') + dmu $$ (i', k) * d fs (k - 1)) div (d_of dmu k) else dmu $$ (i', j'))" using dd\ k t00 t01 t03 LLL_invD_modw(7)[OF Linvmw] k i'j' j' 03 t07 by simp then show ?thesis using dmu'_def i'j' j' t06 t07 by (simp add: d_of_def) qed have "d\ fs'' i' j' = dmu'_mod $$ (i', j')" proof (cases "(i',j') \ I") case i'j'I: True have j': "j' < m" using i'j' by simp show ?thesis proof - have "dmu'_mod $$ (i',j') = dmu' $$ (i',j') symmod (p * (d_of dmu' j') * (d_of dmu' (Suc j')))" using dmu'_mod_def i'j' i'j'I I_def by simp also have "d_of dmu' j' = d fs' j'" using j' d'dmu' d_def Suc_diff_1 less_imp_diff_less unfolding d_of_def by (cases j', auto) finally have "dmu'_mod $$ (i',j') = dmu' $$ (i',j') symmod (p * d fs' j' * d fs' (Suc j'))" using dd\[OF j'] d'dmu' j' by (auto simp: d_of_def) then show ?thesis using i'j'I 11 i'j' eq' by simp qed next case False have "d\ fs'' i' j' = d\ fs' i' j'" using False 11 i'j' by simp also have "\ = dmu' $$ (i', j')" unfolding eq' .. finally show ?thesis unfolding dmu'_mod_def using False[unfolded I_def] i'j' by auto qed } moreover have "\i' j'. i' < m \ j' < m \ i' = j' \ d\ fs'' i' j' = dmu'_mod $$ (i', j')" using dd\ dmu'_def 033 10 dmu'_mod_def 11 I_def by simp moreover { fix i' j' assume i'j'': "i' < m" "j' < m" "i' < j'" then have \z: "\ fs'' i' j' = 0" by (simp add: gram_schmidt_fs.\.simps) have "dmu'_mod $$ (i',j') = dmu' $$ (i',j')" using dmu'_mod_def i'j'' by auto also have "\ = d\ fs i' j'" using LLL_invD_modw(7)[OF Linvmw] i'j'' dmu'_def by simp also have "\ = 0" using d\_def i'j'' by (simp add: gram_schmidt_fs.\.simps) finally have "d\ fs'' i' j' = dmu'_mod $$ (i',j')" using \z d_def i'j'' d\_def by simp } ultimately show ?thesis by (meson nat_neq_iff) qed from gbnd 09 have g_bnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs' fs''] by auto { assume Linv: "LLL_invariant_mod fs mfs dmu p first b k" have 00: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp note swap_weak' = basis_reduction_swap_weak'[OF 00 k k0 mu_F1_i norm_ineq fs'_def] have 01: "LLL_invariant_weak' (k - 1) fs'" by fact have 12: "weakly_reduced fs'' (k-1)" (* invariant requirement *) using 031 09 k LLL_invD_weak(8)[OF 00] unfolding gram_schmidt_fs.weakly_reduced_def by simp have "LLL_invariant_mod fs'' mfs' dmu'_mod p first b (k-1)" using LLL_invI_mod[OF 13 _ 08 05 07 12 14 16 17 p1 g_bnd LLL_invD_mod(17)[OF Linv]] k by simp } moreover have "LLL_invariant_mod_weak fs'' mfs' dmu'_mod p first b" using LLL_invI_modw[OF 13 08 05 07 14 16 17 p1 g_bnd LLL_invD_modw(15)[OF Linvmw]] by simp ultimately show ?thesis using 15 by auto qed lemma dmu_quot_is_round_of_\: assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i'" and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))" and i: "i < m" and j: "j < i" shows "c = round(\ fs i j)" proof - have Linvw: "LLL_invariant_weak' i' fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp have j2: "j < m" using i j by simp then have j3: "Suc j \ m" by simp have \1: "\ fs j j = 1" using i j by (meson gram_schmidt_fs.\.elims less_irrefl_nat) have inZ: "rat_of_int (d fs (Suc j)) * \ fs i j \ \" using fs_int_indpt.fs_int_mu_d_Z_m_m i j LLL_invD_mod(5)[OF Linv] LLL_invD_weak(2) Linvw d_def fs_int.d_def fs_int_indpt.intro by auto have "c = round(rat_of_int (d\ fs i j) / rat_of_int (d\ fs j j))" using LLL_invD_mod(9) Linv i j c by (simp add: round_num_denom d_of_def) then show ?thesis using LLL_d_pos[OF LLL_invw'_imp_w[OF Linvw] j3] j i inZ d\_def \1 by simp qed lemma dmu_quot_is_round_of_\_weak: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))" and i: "i < m" and j: "j < i" shows "c = round(\ fs i j)" proof - have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linv] LLL_invariant_weak_def by simp have j2: "j < m" using i j by simp then have j3: "Suc j \ m" by simp have \1: "\ fs j j = 1" using i j by (meson gram_schmidt_fs.\.elims less_irrefl_nat) have inZ: "rat_of_int (d fs (Suc j)) * \ fs i j \ \" using fs_int_indpt.fs_int_mu_d_Z_m_m i j LLL_invD_modw[OF Linv] d_def fs_int.d_def fs_int_indpt.intro by auto have "c = round(rat_of_int (d\ fs i j) / rat_of_int (d\ fs j j))" using LLL_invD_modw(7) Linv i j c by (simp add: round_num_denom d_of_def) then show ?thesis using LLL_d_pos[OF Linvww j3] j i inZ d\_def \1 by simp qed lemma basis_reduction_mod_add_row: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and res: "basis_reduction_mod_add_row p mfs dmu i j = (mfs', dmu')" and i: "i < m" and j: "j < i" and igtz: "i \ 0" shows "(\fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \ LLL_measure i fs' = LLL_measure i fs \ (\_small_row i fs (Suc j) \ \_small_row i fs' j) \ \\ fs' i j\ \ 1 / 2 \ (\i' j'. i' < i \ j' \ i' \ \ fs' i' j' = \ fs i' j') \ (LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs' mfs' dmu' p first b i) \ (\ii \ m. d fs' ii = d fs ii))" proof - define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))" then have c: "c = round(\ fs i j)" using dmu_quot_is_round_of_\_weak[OF Linv c_def i j] by simp show ?thesis proof (cases "c = 0") case True then have pair_id: "(mfs', dmu') = (mfs, dmu)" using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto moreover have "\\ fs i j\ \ inverse 2" using c[symmetric, unfolded True] by (simp add: round_def, linarith) moreover then have "(\_small_row i fs (Suc j) \ \_small_row i fs j)" unfolding \_small_row_def using Suc_leI le_neq_implies_less by blast ultimately show ?thesis using Linv pair_id by auto next case False then have pair_id: "(mfs', dmu') = (mfs[i := map_vec (\x. x symmod p) (mfs ! i - c \\<^sub>v mfs ! j)], mat m m (\(i', j'). if i' = i \ j' \ j then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j') else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))) else dmu $$ (i', j')))" using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto then have mfs': "mfs' = mfs[i := map_vec (\x. x symmod p) (mfs ! i - c \\<^sub>v mfs ! j)]" and dmu': "dmu' = mat m m (\(i', j'). if i' = i \ j' \ j then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j') else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))) else dmu $$ (i', j'))" by auto show ?thesis using basis_reduction_mod_add_row_main[OF Linv i j c mfs' dmu'] by blast qed qed lemma basis_reduction_mod_swap: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and mu: "\\ fs k (k-1)\ \ 1 / 2" and res: "basis_reduction_mod_swap p mfs dmu k = (mfs', dmu'_mod)" and cond: "sq_norm (gso fs (k - 1)) > \ * sq_norm (gso fs k)" and i: "k < m" "k \ 0" shows "(\fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \ LLL_measure (k - 1) fs' < LLL_measure k fs \ (LLL_invariant_mod fs mfs dmu p first b k \ LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))" using res[unfolded basis_reduction_mod_swap_def basis_reduction_mod_swap_dmu_mod_def] basis_reduction_mod_swap_main[OF Linv i mu cond] by blast lemma basis_reduction_adjust_mod: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and res: "basis_reduction_adjust_mod p first mfs dmu = (p', mfs', dmu', g_idx')" shows "(\fs' b'. (LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs' mfs' dmu' p' first b' i) \ LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \ LLL_measure i fs' = LLL_measure i fs)" proof (cases "\ g_idx. basis_reduction_adjust_mod p first mfs dmu = (p, mfs, dmu, g_idx)") case True thus ?thesis using res Linv by auto next case False obtain b' g_idx where norm: "compute_max_gso_norm first dmu = (b', g_idx)" by force define p'' where "p'' = compute_mod_of_max_gso_norm first b'" define d_vec where "d_vec = vec (Suc m) (\i. d_of dmu i)" define mfs'' where "mfs'' = map (map_vec (\x. x symmod p'')) mfs" define dmu'' where "dmu'' = mat m m (\(i, j). if j < i then dmu $$ (i, j) symmod (p'' * d_vec $ j * d_vec $ Suc j) else dmu $$ (i, j))" note res = res False note res = res[unfolded basis_reduction_adjust_mod.simps Let_def norm split, folded p''_def, folded d_vec_def mfs''_def, folded dmu''_def] from res have pp': "p'' < p" and id: "dmu' = dmu''" "mfs' = mfs''" "p' = p''" "g_idx' = g_idx" by (auto split: if_splits) define I where "I = {(i',j'). i' < m \ j' < i'}" note inv = LLL_invD_modw[OF Linv] from inv(4) have lin: "gs.lin_indpt_list (RAT fs)" . from inv(3) have lat: "lattice_of fs = L" . from inv(2) have len: "length fs = m" . have weak: "LLL_invariant_weak fs" using Linv by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_weak_def) from compute_max_gso_norm[OF _ weak, of dmu first, unfolded norm] inv(7) have bnd: "g_bnd_mode first b' fs" and b': "b' \ 0" "m = 0 \ b' = 0" by auto from compute_mod_of_max_gso_norm[OF b' p''_def] have p'': "0 < p''" "1 < p''" "mod_invariant b' p'' first" by auto obtain fs' where 01: "lattice_of fs' = L" and 02: "map (map_vec (\ x. x symmod p'')) fs' = map (map_vec (\ x. x symmod p'')) fs" and 03: "lin_indep fs'" and 04: "length fs' = m" and 05: "(\ k < m. gso fs' k = gso fs k)" and 06: "(\ k \ m. d fs' k = d fs k)" and 07: "(\ i' < m. \ j' < m. d\ fs' i' j' = (if (i',j') \ I then d\ fs i' j' symmod (p'' * d fs j' * d fs (Suc j')) else d\ fs i' j'))" using mod_finite_set[OF lin len _ lat, of I] I_def p'' by blast from bnd 05 have bnd: "g_bnd_mode first b' fs'" using g_bnd_mode_cong[of fs fs'] by auto have D: "D fs = D fs'" unfolding D_def using 06 by auto have Linv': "LLL_invariant_mod_weak fs' mfs'' dmu'' p'' first b'" proof (intro LLL_invI_modw p'' 04 03 01 bnd) { have "mfs'' = map (map_vec (\x. x symmod p'')) mfs" by fact also have "\ = map (map_vec (\x. x symmod p'')) (map (map_vec (\x. x symmod p)) fs)" using inv by simp also have "\ = map (map_vec (\x. x symmod p symmod p'')) fs" by auto also have "(\ x. x symmod p symmod p'') = (\ x. x symmod p'')" proof (intro ext) fix x from \mod_invariant b p first\[unfolded mod_invariant_def] obtain e where p: "p = log_base ^ e" by auto from p''[unfolded mod_invariant_def] obtain e' where p'': "p'' = log_base ^ e'" by auto from pp'[unfolded p p''] log_base have "e' \ e" by simp hence dvd: "p'' dvd p" unfolding p p'' using log_base by (metis le_imp_power_dvd) thus "x symmod p symmod p'' = x symmod p''" by (intro sym_mod_sym_mod_cancel) qed finally show "map (map_vec (\x. x symmod p'')) fs' = mfs''" unfolding 02 .. } thus "length mfs'' = m" using 04 by auto show "\i'j'd\ fs' i' j'\ < p'' * d fs' j' * d fs' (Suc j')" proof - { fix i' j' assume i'j': "i' < m" "j' < i'" then have "d\ fs' i' j' = d\ fs i' j' symmod (p'' * d fs' j' * d fs' (Suc j'))" using 07 06 unfolding I_def by simp then have "\d\ fs' i' j'\ < p'' * d fs' j' * d fs' (Suc j')" using sym_mod_abs p'' LLL_d_pos[OF weak] mult_pos_pos by (smt "06" i'j' less_imp_le_nat less_trans_Suc nat_SN.gt_trans) } then show ?thesis by simp qed from inv(7) have dmu: "i' < m \ j' < m \ dmu $$ (i', j') = d\ fs i' j'" for i' j' by auto note d_of = d_of_weak[OF Linv] have dvec: "i \ m \ d_vec $ i = d fs i" for i unfolding d_vec_def using d_of by auto show "\i'j' fs' i' j' = dmu'' $$ (i', j')" using 07 unfolding dmu''_def I_def by (auto simp: dmu dvec) qed moreover { assume linv: "LLL_invariant_mod fs mfs dmu p first b i" note inv = LLL_invD_mod[OF linv] hence i: "i \ m" by auto have norm: "j < m \ \gso fs j\\<^sup>2 = \gso fs' j\\<^sup>2" for j using 05 by auto have "weakly_reduced fs i = weakly_reduced fs' i" unfolding gram_schmidt_fs.weakly_reduced_def using i by (intro all_cong arg_cong2[where f = "(\)"] arg_cong[where f = "\ x. _ * x"] norm, auto) with inv have "weakly_reduced fs' i" by auto hence "LLL_invariant_mod fs' mfs'' dmu'' p'' first b' i" using inv by (intro LLL_invI_mod LLL_invD_modw[OF Linv']) } moreover have "LLL_measure i fs' = LLL_measure i fs" unfolding LLL_measure_def logD_def D .. ultimately show ?thesis unfolding id by blast qed lemma alpha_comparison: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and alph: "quotient_of \ = (num, denom)" and i: "i < m" and i0: "i \ 0" shows "(d_of dmu i * d_of dmu i * denom \ num * d_of dmu (i - 1) * d_of dmu (Suc i)) = (sq_norm (gso fs (i - 1)) \ \ * sq_norm (gso fs i))" proof - note inv = LLL_invD_modw[OF Linv] interpret fs_indep: fs_int_indpt n fs by (unfold_locales, insert inv, auto) from inv(2) i have ifs: "i < length fs" by auto note d_of_fs = d_of_weak[OF Linv] show ?thesis unfolding fs_indep.d_sq_norm_comparison[OF alph ifs i0, symmetric] by (subst (1 2 3 4) d_of_fs, use i d_def fs_indep.d_def in auto) qed lemma basis_reduction_adjust_swap_add_step: assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b" and res: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i = (p', mfs', dmu', g_idx')" and alph: "quotient_of \ = (num, denom)" and ineq: "\ (d_of dmu i * d_of dmu i * denom \ num * d_of dmu (i - 1) * d_of dmu (Suc i))" and i: "i < m" and i0: "i \ 0" shows "\fs' b'. LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \ LLL_measure (i - 1) fs' < LLL_measure i fs \ LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs \ (LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1))" proof - obtain mfs0 dmu0 where add: "basis_reduction_mod_add_row p mfs dmu i (i-1) = (mfs0, dmu0)" by force obtain mfs1 dmu1 where swap: "basis_reduction_mod_swap p mfs0 dmu0 i = (mfs1, dmu1)" by force note res = res[unfolded basis_reduction_adjust_swap_add_step_def Let_def add split swap] from i0 have ii: "i - 1 < i" by auto from basis_reduction_mod_add_row[OF Linv add i ii i0] obtain fs0 where Linv0: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p first b" and meas0: "LLL_measure i fs0 = LLL_measure i fs" and small: "\\ fs0 i (i - 1)\ \ 1 / 2" and Linv0': "LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs0 mfs0 dmu0 p first b i" by blast { have id: "d_of dmu0 i = d_of dmu i" "d_of dmu0 (i - 1) = d_of dmu (i - 1)" "d_of dmu0 (Suc i) = d_of dmu (Suc i)" using i i0 add[unfolded basis_reduction_mod_add_row_def Let_def] by (auto split: if_splits simp: d_of_def) from ineq[folded id, unfolded alpha_comparison[OF Linv0 alph i i0]] have "\gso fs0 (i - 1)\\<^sup>2 > \ * \gso fs0 i\\<^sup>2" by simp } note ineq = this from Linv have "LLL_invariant_weak fs" by (auto simp: LLL_invariant_weak_def LLL_invariant_mod_weak_def) from basis_reduction_mod_swap[OF Linv0 small swap ineq i i0, unfolded meas0] Linv0' obtain fs1 where Linv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p first b" and meas1: "LLL_measure (i - 1) fs1 < LLL_measure i fs" and Linv1': "LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs1 mfs1 dmu1 p first b (i - 1)" by auto show ?thesis proof (cases "i - 1 = g_idx") case False with res have id: "p' = p" "mfs' = mfs1" "dmu' = dmu1" "g_idx' = g_idx" by auto show ?thesis unfolding id using Linv1' meas1 Linv1 by (intro exI[of _ fs1] exI[of _ b], auto simp: LLL_measure_def) next case True with res have adjust: "basis_reduction_adjust_mod p first mfs1 dmu1 = (p', mfs', dmu', g_idx')" by simp from basis_reduction_adjust_mod[OF Linv1 adjust, of "i - 1"] Linv1' obtain fs' b' where Linvw: "LLL_invariant_mod_weak fs' mfs' dmu' p' first b'" and Linv: "LLL_invariant_mod fs mfs dmu p first b i \ LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1)" and meas: "LLL_measure (i - 1) fs' = LLL_measure (i - 1) fs1" by blast note meas = meas1[folded meas] from meas have meas': "LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs" unfolding LLL_measure_def using i by auto show ?thesis by (intro exI conjI impI, rule Linvw, rule meas, rule meas', rule Linv) qed qed lemma basis_reduction_mod_step: assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i" and res: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')" and i: "i < m" shows "\fs' b'. LLL_measure i' fs' < LLL_measure i fs \ LLL_invariant_mod fs' mfs' dmu' p' first b' i'" proof - note res = res[unfolded basis_reduction_mod_step_def Let_def] from Linv have Linvw: "LLL_invariant_mod_weak fs mfs dmu p first b" by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def) show ?thesis proof (cases "i = 0") case True then have ids: "mfs' = mfs" "dmu' = dmu" "i' = Suc i" "p' = p" using res by auto have "LLL_measure i' fs < LLL_measure i fs \ LLL_invariant_mod fs mfs' dmu' p first b i'" using increase_i_mod[OF Linv i] True res ids inv by simp then show ?thesis using res ids inv by auto next case False hence id: "(i = 0) = False" by auto obtain num denom where alph: "quotient_of \ = (num, denom)" by force note res = res[unfolded id if_False alph split] let ?comp = "d_of dmu i * d_of dmu i * denom \ num * d_of dmu (i - 1) * d_of dmu (Suc i)" show ?thesis proof (cases ?comp) case False hence id: "?comp = False" by simp note res = res[unfolded id if_False] let ?step = "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i" from res have step: "?step = (p', mfs', dmu', g_idx')" and i': "i' = i - 1" by (cases ?step, auto)+ from basis_reduction_adjust_swap_add_step[OF Linvw step alph False i \i \ 0\] Linv show ?thesis unfolding i' by blast next case True hence id: "?comp = True" by simp note res = res[unfolded id if_True] from res have ids: "p' = p" "mfs' = mfs" "dmu' = dmu" "i' = Suc i" by auto from True alpha_comparison[OF Linvw alph i False] have ineq: "sq_norm (gso fs (i - 1)) \ \ * sq_norm (gso fs i)" by simp from increase_i_mod[OF Linv i ineq] show ?thesis unfolding ids by auto qed qed qed lemma basis_reduction_mod_main: assumes "LLL_invariant_mod fs mfs dmu p first b i" and res: "basis_reduction_mod_main p first mfs dmu g_idx i j = (p', mfs', dmu')" shows "\fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m" using assms proof (induct "LLL_measure i fs" arbitrary: i mfs dmu j p b fs g_idx rule: less_induct) case (less i fs mfs dmu j p b g_idx) hence fsinv: "LLL_invariant_mod fs mfs dmu p first b i" by auto note res = less(3)[unfolded basis_reduction_mod_main.simps[of p first mfs dmu g_idx i j]] note inv = less(2) note IH = less(1) show ?case proof (cases "i < m") case i: True obtain p' mfs' dmu' g_idx' i' j' where step: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')" (is "?step = _") by (cases ?step, auto) then obtain fs' b' where Linv: "LLL_invariant_mod fs' mfs' dmu' p' first b' i'" and decr: "LLL_measure i' fs' < LLL_measure i fs" using basis_reduction_mod_step[OF fsinv step i] i fsinv by blast note res = res[unfolded step split] from res i show ?thesis using IH[OF decr Linv] by auto next case False with LLL_invD_mod[OF fsinv] res have i: "i = m" "p' = p" by auto then obtain fs' b' where "LLL_invariant_mod fs' mfs' dmu' p first b' m" using False res fsinv by simp then show ?thesis using i by auto qed qed lemma compute_max_gso_quot_alpha: assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b" and max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)" and alph: "quotient_of \ = (num, denum)" and cmp: "(msq_num * denum > num * msq_denum) = cmp" and m: "m > 1" shows "cmp \ idx \ 0 \ idx < m \ \ (d_of dmu idx * d_of dmu idx * denum \ num * d_of dmu (idx - 1) * d_of dmu (Suc idx))" and "\ cmp \ LLL_invariant_mod fs mfs dmu p first b m" proof - from inv have fsinv: "LLL_invariant_weak fs" by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def) define qt where "qt = (\i. ((d_of dmu (i + 1)) * (d_of dmu (i + 1)), (d_of dmu (i + 2)) * (d_of dmu i), Suc i))" define lst where "lst = (map (\i. qt i) [0..<(m-1)])" have msqlst: "(msq_num, msq_denum, idx) = max_list_rats_with_index lst" using max lst_def qt_def unfolding compute_max_gso_quot_def by simp have nz: "\n d i. (n, d, i) \ set lst \ d > 0" unfolding lst_def qt_def using d_of_weak[OF inv] LLL_d_pos[OF fsinv] by auto have geq: "\(n, d, i) \ set lst. rat_of_int msq_num / of_int msq_denum \ rat_of_int n / of_int d" using max_list_rats_with_index[of lst] nz msqlst by (metis (no_types, lifting) case_prodI2) have len: "length lst \ 1" using m unfolding lst_def by simp have inset: "(msq_num, msq_denum, idx) \ set lst" using max_list_rats_with_index_in_set[OF msqlst[symmetric] len] nz by simp then have idxm: "idx \ {1.. 0" and idx: "idx < m" by auto have 00: "(msq_num, msq_denum, idx) = qt (idx - 1)" using lst_def inset qt_def by auto then have id_qt: "msq_num = d_of dmu idx * d_of dmu idx" "msq_denum = d_of dmu (Suc idx) * d_of dmu (idx - 1)" unfolding qt_def by auto have "msq_denum = (d_of dmu (idx + 1)) * (d_of dmu (idx - 1))" using 00 unfolding qt_def by simp then have dengt0: "msq_denum > 0" using d_of_weak[OF inv] idxm LLL_d_pos[OF fsinv] by auto have \dengt0: "denum > 0" using alph by (metis quotient_of_denom_pos) from cmp[unfolded id_qt] have cmp: "cmp = (\ (d_of dmu idx * d_of dmu idx * denum \ num * d_of dmu (idx - 1) * d_of dmu (Suc idx)))" by (auto simp: ac_simps) { assume cmp from this[unfolded cmp] show "idx \ 0 \ idx < m \ \ (d_of dmu idx * d_of dmu idx * denum \ num * d_of dmu (idx - 1) * d_of dmu (Suc idx))" using idx0 idx by auto } { assume "\ cmp" from this[unfolded cmp] have small: "d_of dmu idx * d_of dmu idx * denum \ num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto note d_pos = LLL_d_pos[OF fsinv] have gso: "k < m \ sq_norm (gso fs k) = of_int (d fs (Suc k)) / of_int (d fs k)" for k using LLL_d_Suc[OF fsinv, of k] d_pos[of k] by simp have gso_pos: "k < m \ sq_norm (gso fs k) > 0" for k using gso[of k] d_pos[of k] d_pos[of "Suc k"] by auto from small[unfolded alpha_comparison[OF inv alph idx idx0]] have alph: "sq_norm (gso fs (idx - 1)) \ \ * sq_norm (gso fs idx)" . with gso_pos[OF idx] have alph: "sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx) \ \" by (metis mult_imp_div_pos_le) have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def proof (intro allI impI, goal_cases) case (1 i) from idx have idx1: "idx - 1 < m" by auto from geq[unfolded lst_def] have mem: "(d_of dmu (Suc i) * d_of dmu (Suc i), d_of dmu (Suc (Suc i)) * d_of dmu i, Suc i) \ set lst" unfolding lst_def qt_def using 1 by auto have "sq_norm (gso fs i) / sq_norm (gso fs (Suc i)) = of_int (d_of dmu (Suc i) * d_of dmu (Suc i)) / of_int (d_of dmu (Suc (Suc i)) * d_of dmu i)" using gso idx0 d_of_weak[OF inv] 1 by auto also have "\ \ rat_of_int msq_num / rat_of_int msq_denum" using geq[rule_format, OF mem, unfolded split] by auto also have "\ = sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx)" unfolding id_qt gso[OF idx] gso[OF idx1] using idx0 d_of_weak[OF inv] idx by auto also have "\ \ \" by fact finally show "sq_norm (gso fs i) \ \ * sq_norm (gso fs (Suc i))" using gso_pos[OF 1] using pos_divide_le_eq by blast qed with inv show "LLL_invariant_mod fs mfs dmu p first b m" by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def) } qed lemma small_m: assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b" and m: "m \ 1" shows "LLL_invariant_mod fs mfs dmu p first b m" proof - have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def using m by auto with inv show "LLL_invariant_mod fs mfs dmu p first b m" by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def) qed lemma basis_reduction_iso_main: assumes "LLL_invariant_mod_weak fs mfs dmu p first b" and res: "basis_reduction_iso_main p first mfs dmu g_idx j = (p', mfs', dmu')" shows "\fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m" using assms proof (induct "LLL_measure (m-1) fs" arbitrary: fs mfs dmu j p b g_idx rule: less_induct) case (less fs mfs dmu j p b g_idx) have inv: "LLL_invariant_mod_weak fs mfs dmu p first b" using less by auto hence fsinv: "LLL_invariant_weak fs" by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def) note res = less(3)[unfolded basis_reduction_iso_main.simps[of p first mfs dmu g_idx j]] note IH = less(1) obtain msq_num msq_denum idx where max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)" by (metis prod_cases3) obtain num denum where alph: "quotient_of \ = (num, denum)" by force note res = res[unfolded max alph Let_def split] consider (small) "m \ 1" | (final) "m > 1" "\ (num * msq_denum < msq_num * denum)" | (step) "m > 1" "num * msq_denum < msq_num * denum" by linarith thus ?case proof cases case *: step obtain p1 mfs1 dmu1 g_idx1 where step: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx idx = (p1, mfs1, dmu1, g_idx1)" by (metis prod_cases4) from res[unfolded step split] * have res: "basis_reduction_iso_main p1 first mfs1 dmu1 g_idx1 (j + 1) = (p', mfs', dmu')" by auto from compute_max_gso_quot_alpha(1)[OF inv max alph refl *] have idx0: "idx \ 0" and idx: "idx < m" and cmp: "\ d_of dmu idx * d_of dmu idx * denum \ num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto from basis_reduction_adjust_swap_add_step[OF inv step alph cmp idx idx0] obtain fs1 b1 where inv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p1 first b1" and meas: "LLL_measure (m - 1) fs1 < LLL_measure (m - 1) fs" by auto from IH[OF meas inv1 res] show ?thesis . next case small with res small_m[OF inv] show ?thesis by auto next case final from compute_max_gso_quot_alpha(2)[OF inv max alph refl final] final show ?thesis using res by auto qed qed lemma basis_reduction_mod_add_rows_loop_inv': assumes fsinv: "LLL_invariant_mod fs mfs dmu p first b m" and res: "basis_reduction_mod_add_rows_loop p mfs dmu i i = (mfs', dmu')" and i: "i < m" shows "\fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \ (\i' j'. i' < i \ j' \ i' \ \ fs i' j' = \ fs' i' j') \ \_small fs' i" proof - { fix j assume j: "j \ i" and mu_small: "\_small_row i fs j" and resj: "basis_reduction_mod_add_rows_loop p mfs dmu i j = (mfs', dmu')" have "\fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \ (\i' j'. i' < i \ j' \ i' \ \ fs i' j' = \ fs' i' j') \ (\_small fs' i)" proof (insert fsinv mu_small resj i j, induct j arbitrary: fs mfs dmu mfs' dmu') case (0 fs) then have "(mfs', dmu') = (mfs, dmu)" by simp then show ?case using LLL_invariant_mod_to_weak_m_to_i(3) basis_reduction_add_row_done_weak 0 by auto next case (Suc j) hence j: "j < i" by auto have in0: "i \ 0" using Suc(6) by simp define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))" have c2: "c = round (\ fs i j)" using dmu_quot_is_round_of_\[OF _ _ i j] c_def Suc by simp define mfs'' where "mfs'' = (if c=0 then mfs else mfs[ i := (map_vec (\ x. x symmod p)) (mfs ! i - c \\<^sub>v mfs ! j)])" define dmu'' where "dmu'' = (if c=0 then dmu else mat m m (\(i',j'). (if (i' = i \ j' \ j) then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j')) else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))) else (dmu $$ (i',j')))))" have 00: "basis_reduction_mod_add_row p mfs dmu i j = (mfs'', dmu'')" using mfs''_def dmu''_def unfolding basis_reduction_mod_add_row_def c_def[symmetric] by simp then have 01: "basis_reduction_mod_add_rows_loop p mfs'' dmu'' i j = (mfs', dmu')" using basis_reduction_mod_add_rows_loop.simps(2)[of p mfs dmu i j] Suc by simp have fsinvi: "LLL_invariant_mod fs mfs dmu p first b i" using LLL_invariant_mod_to_weak_m_to_i[OF Suc(2)] i by simp then have fsinvmw: "LLL_invariant_mod_weak fs mfs dmu p first b" using LLL_invD_mod LLL_invI_modw by simp obtain fs'' where fs''invi: "LLL_invariant_mod fs'' mfs'' dmu'' p first b i" and \_small': "(\_small_row i fs (Suc j) \ \_small_row i fs'' j)" and \s: "(\i' j'. i' < i \ j' \ i' \ \ fs'' i' j' = \ fs i' j')" using Suc basis_reduction_mod_add_row[OF fsinvmw 00 i j] fsinvi by auto moreover then have \sm: "\_small_row i fs'' j" using Suc by simp have fs''invwi: "LLL_invariant_weak' i fs''" using LLL_invD_mod[OF fs''invi] LLL_invI_weak by simp have fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod[OF fsinvi] LLL_invI_weak by simp note invw = LLL_invw'_imp_w[OF fsinvwi] note invw'' = LLL_invw'_imp_w[OF fs''invwi] have "LLL_invariant_mod fs'' mfs'' dmu'' p first b m" proof - have "(\ l. Suc l < m \ sq_norm (gso fs'' l) \ \ * sq_norm (gso fs'' (Suc l)))" proof - { fix l assume l: "Suc l < m" have "sq_norm (gso fs'' l) \ \ * sq_norm (gso fs'' (Suc l))" proof (cases "i \ Suc l") case True have deq: "\k. k < m \ d fs (Suc k) = d fs'' (Suc k)" using dd\ LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp { fix k assume k: "k < m" then have "d fs (Suc k) = d fs'' (Suc k)" using dd\ LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp have "d fs 0 = 1" "d fs'' 0 = 1" using d_def by auto moreover have sqid: "sq_norm (gso fs'' k) = rat_of_int (d fs'' (Suc k)) / rat_of_int (d fs'' k)" using LLL_d_Suc[OF invw''] LLL_d_pos[OF invw''] k by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero) moreover have "sq_norm (gso fs k) = rat_of_int (d fs (Suc k)) / rat_of_int (d fs k)" using LLL_d_Suc[OF invw] LLL_d_pos[OF invw] k by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero) ultimately have "sq_norm (gso fs k) = sq_norm (gso fs'' k)" using k deq LLL_d_pos[OF invw] LLL_d_pos[OF invw''] by (metis (no_types, lifting) Nat.lessE Suc_lessD old.nat.inject zero_less_Suc) } then show ?thesis using LLL_invD_mod(6)[OF Suc(2)] by (simp add: gram_schmidt_fs.weakly_reduced_def l) next case False then show ?thesis using LLL_invD_mod(6)[OF fs''invi] gram_schmidt_fs.weakly_reduced_def by (metis less_or_eq_imp_le nat_neq_iff) qed } then show ?thesis by simp qed then have "weakly_reduced fs'' m" using gram_schmidt_fs.weakly_reduced_def by blast then show ?thesis using LLL_invD_mod[OF fs''invi] LLL_invI_mod by simp qed then show ?case using "01" Suc.hyps i j less_imp_le_nat \sm \s by metis qed } then show ?thesis using \_small_row_refl res by auto qed lemma basis_reduction_mod_add_rows_outer_loop_inv: assumes inv: "LLL_invariant_mod fs mfs dmu p first b m" and "(mfs', dmu') = basis_reduction_mod_add_rows_outer_loop p mfs dmu i" and i: "i < m" shows "(\fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \ (\j. j \ i \ \_small fs' j))" proof(insert assms, induct i arbitrary: fs mfs dmu mfs' dmu') case (0 fs) then show ?case using \_small_def by auto next case (Suc i fs mfs dmu mfs' dmu') obtain mfs'' dmu'' where mfs''dmu'': "(mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p mfs dmu i" by (metis surj_pair) then obtain fs'' where fs'': "LLL_invariant_mod fs'' mfs'' dmu'' p first b m" and 00: "(\j. j \ i \ \_small fs'' j)" using Suc by fastforce have "(mfs', dmu') = basis_reduction_mod_add_rows_loop p mfs'' dmu'' (Suc i) (Suc i)" using Suc(3,4) mfs''dmu'' by (smt basis_reduction_mod_add_rows_outer_loop.simps(2) case_prod_conv) then obtain fs' where 01: "LLL_invariant_mod fs' mfs' dmu' p first b m" and 02: "\i' j'. i' < (Suc i) \ j' \ i' \ \ fs'' i' j' = \ fs' i' j'" and 03: "\_small fs' (Suc i)" using fs'' basis_reduction_mod_add_rows_loop_inv' Suc by metis moreover have "\j. j \ (Suc i) \ \_small fs' j" using 02 00 03 \_small_def by (simp add: le_Suc_eq) ultimately show ?case by blast qed lemma basis_reduction_mod_fs_bound: assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k" and mu_small: "\_small fs i" and i: "i < m" and nFirst: "\ first" shows "fs ! i = mfs ! i" proof - from LLL_invD_mod(16-17)[OF Linv] nFirst g_bnd_mode_def have gbnd: "g_bnd b fs" and bp: "b \ (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m + 3)" by (auto simp: mod_invariant_def bound_number_def) have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp then interpret fs: fs_int_indpt n fs using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp have "\gso fs 0\\<^sup>2 \ b" using gbnd i unfolding g_bnd_def by blast then have b0: "0 \ b" using sq_norm_vec_ge_0 dual_order.trans by auto have 00: "of_int \fs ! i\\<^sup>2 = (\j\[0.. fs i j)\<^sup>2 * \gso fs j\\<^sup>2)" using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro i by simp have 01: "\j < i. (\ fs i j)\<^sup>2 * \gso fs j\\<^sup>2 \ (1 / rat_of_int 4) * \gso fs j\\<^sup>2" proof - { fix j assume j: "j < i" then have "\fs.gs.\ i j\ \ 1 / (rat_of_int 2)" using mu_small Power.linordered_idom_class.abs_square_le_1 j unfolding \_small_def by simp moreover have "\\ fs i j\ \ 0" by simp ultimately have "\\ fs i j\\<^sup>2 \ (1 / rat_of_int 2)\<^sup>2" using Power.linordered_idom_class.abs_le_square_iff by fastforce also have "\ = 1 / (rat_of_int 4)" by (simp add: field_simps) finally have "\\ fs i j\\<^sup>2 \ 1 / rat_of_int 4" by simp } then show ?thesis using fs.gs.\.simps by (metis mult_right_mono power2_abs sq_norm_vec_ge_0) qed then have 0111: "\j. j \ set [0.. (\ fs i j)\<^sup>2 * \gso fs j\\<^sup>2 \ (1 / rat_of_int 4) * \gso fs j\\<^sup>2" by simp { fix j assume j: "j < n" have 011: "(\ fs i i)\<^sup>2 * \gso fs i\\<^sup>2 = 1 * \gso fs i\\<^sup>2" using fs.gs.\.simps by simp have 02: "\j < Suc i. \gso fs j\\<^sup>2 \ b" using gbnd i unfolding g_bnd_def by simp have 03: "length [0..fs ! i\\<^sup>2 = (\j\[0.. fs i j)\<^sup>2 * \gso fs j\\<^sup>2) + \gso fs i\\<^sup>2" unfolding 00 using 011 by simp also have "(\j\[0.. fs i j)\<^sup>2 * \gso fs j\\<^sup>2) \ (\j\[0..gso fs j\\<^sup>2))" using Groups_List.sum_list_mono[OF 0111] by fast finally have "of_int \fs ! i\\<^sup>2 \ (\j\[0..gso fs j\\<^sup>2)) + \gso fs i\\<^sup>2" by simp also have "(\j\[0..gso fs j\\<^sup>2)) \ (\j\[0..gso fs i\\<^sup>2 \ b" using 02 by simp finally have "of_int \fs ! i\\<^sup>2 \ (\j\[0.. = (rat_of_nat i) * ((1 / rat_of_int 4) * b) + b" using 03 sum_list_triv[of "(1 / rat_of_int 4) * b" "[0.. = (rat_of_nat i) / 4 * b + b" by simp also have "\ = ((rat_of_nat i) / 4 + 1)* b" by algebra also have "\ = (rat_of_nat i + 4) / 4 * b" by simp finally have "of_int \fs ! i\\<^sup>2 \ (rat_of_nat i + 4) / 4 * b" by simp also have "\ \ (rat_of_nat (m + 3)) / 4 * b" using i b0 times_left_mono by fastforce finally have "of_int \fs ! i\\<^sup>2 \ rat_of_nat (m+3) / 4 * b" by simp moreover have "\fs ! i $ j\\<^sup>2 \ \fs ! i\\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] i j by blast ultimately have 04: "of_int (\fs ! i $ j\\<^sup>2) \ rat_of_nat (m+3) / 4 * b" using ge_trans i by linarith then have 05: "real_of_int (\fs ! i $ j\\<^sup>2) \ real_of_rat (rat_of_nat (m+3) / 4 * b)" proof - from j have "rat_of_int (\fs ! i $ j\\<^sup>2) \ rat_of_nat (m+3) / 4 * b" using 04 by simp then have "real_of_int (\fs ! i $ j\\<^sup>2) \ real_of_rat (rat_of_nat (m+3) / 4 * b)" using j of_rat_less_eq by (metis of_rat_of_int_eq) then show ?thesis by simp qed define rhs where "rhs = real_of_rat (rat_of_nat (m+3) / 4 * b)" have rhs0: "rhs \ 0" using b0 i rhs_def by simp have fsij: "real_of_int \fs ! i $ j\ \ 0" by simp have "real_of_int (\fs ! i $ j\\<^sup>2) = (real_of_int \fs ! i $ j\)\<^sup>2" by simp then have "(real_of_int \fs ! i $ j\)\<^sup>2 \ rhs" using 05 j rhs_def by simp then have g1: "real_of_int \fs ! i $ j\ \ sqrt rhs" using NthRoot.real_le_rsqrt by simp have pbnd: "2 * \fs ! i $ j\ < p" proof - have "rat_of_nat (m+3) / 4 * b \ (rat_of_nat (m +3) / 4) * (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m+3)" using bp b0 i times_left_mono SN_Orders.of_nat_ge_zero gs.m_comm times_divide_eq_right by (smt gs.l_null le_divide_eq_numeral1(1)) also have "\ = (rat_of_int (p - 1))\<^sup>2 / 4 * (rat_of_nat (m + 3) / rat_of_nat (m + 3))" by (metis (no_types, lifting) gs.m_comm of_nat_add of_nat_numeral times_divide_eq_left) finally have "rat_of_nat (m+3) / 4 * b \ (rat_of_int (p - 1))\<^sup>2 / 4" by simp then have "sqrt rhs \ sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))" unfolding rhs_def using of_rat_less_eq by fastforce then have two_ineq: "2 * \fs ! i $ j\ \ 2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))" using g1 by linarith have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) = sqrt (real_of_rat (4 * ((rat_of_int (p - 1))\<^sup>2 / 4)))" by (metis (no_types, opaque_lifting) real_sqrt_mult of_int_numeral of_rat_hom.hom_mult of_rat_of_int_eq real_sqrt_four times_divide_eq_right) also have "\ = sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2))" using i by simp also have "(real_of_rat ((rat_of_int (p - 1))\<^sup>2)) = (real_of_rat (rat_of_int (p - 1)))\<^sup>2" using Rat.of_rat_power by blast also have "sqrt ((real_of_rat (rat_of_int (p - 1)))\<^sup>2) = real_of_rat (rat_of_int (p - 1))" using LLL_invD_mod(15)[OF Linv] by simp finally have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) = real_of_rat (rat_of_int (p - 1))" by simp then have "2 * \fs ! i $ j\ \ real_of_rat (rat_of_int (p - 1))" using two_ineq by simp then show ?thesis by (metis of_int_le_iff of_rat_of_int_eq zle_diff1_eq) qed have p1: "p > 1" using LLL_invD_mod[OF Linv] by blast interpret pm: poly_mod_2 p by (unfold_locales, rule p1) from LLL_invD_mod[OF Linv] have len: "length fs = m" and fs: "set fs \ carrier_vec n" by auto from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! i $ j mod p) = fs ! i $ j" . also have "pm.inv_M (fs ! i $ j mod p) = mfs ! i $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def using i j len fs by auto finally have "fs ! i $ j = mfs ! i $ j" .. } thus "fs ! i = mfs ! i" using LLL_invD_mod(10,13)[OF Linv i] by auto qed lemma basis_reduction_mod_fs_bound_first: assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k" and m0: "m > 0" and first: "first" shows "fs ! 0 = mfs ! 0" proof - from LLL_invD_mod(16-17)[OF Linv] first g_bnd_mode_def m0 have gbnd: "sq_norm (gso fs 0) \ b" and bp: "b \ (rat_of_int (p - 1))\<^sup>2 / 4" by (auto simp: mod_invariant_def bound_number_def) from LLL_invD_mod[OF Linv] have p1: "p > 1" by blast have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp then interpret fs: fs_int_indpt n fs using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp from gbnd have b0: "0 \ b" using sq_norm_vec_ge_0 dual_order.trans by auto have "of_int \fs ! 0\\<^sup>2 = (\ fs 0 0)\<^sup>2 * \gso fs 0\\<^sup>2" using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro m0 by simp also have "\ = \gso fs 0\\<^sup>2" unfolding fs.gs.\.simps by (simp add: gs.\.simps) also have "\ \ (rat_of_int (p - 1))\<^sup>2 / 4" using gbnd bp by auto finally have one: "of_int (sq_norm (fs ! 0)) \ (rat_of_int (p - 1))\<^sup>2 / 4" . { fix j assume j: "j < n" have leq: "\fs ! 0 $ j\\<^sup>2 \ \fs ! 0\\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] m0 j by blast have "rat_of_int ((2 * \fs ! 0 $ j\)^2) = rat_of_int (4 * \fs ! 0 $ j\\<^sup>2)" by simp also have "\ \ 4 * of_int \fs ! 0\\<^sup>2" using leq by simp also have "\ \ 4 * (rat_of_int (p - 1))\<^sup>2 / 4" using one by simp also have "\ = (rat_of_int (p - 1))\<^sup>2" by simp also have "\ = rat_of_int ((p - 1)\<^sup>2)" by simp finally have "(2 * \fs ! 0 $ j\)^2 \ (p - 1)\<^sup>2" by linarith hence "2 * \fs ! 0 $ j\ \ p - 1" using p1 by (smt power_mono_iff zero_less_numeral) hence pbnd: "2 * \fs ! 0 $ j\ < p" by simp interpret pm: poly_mod_2 p by (unfold_locales, rule p1) from LLL_invD_mod[OF Linv] m0 have len: "length fs = m" "length mfs = m" and fs: "fs ! 0 \ carrier_vec n" "mfs ! 0 \ carrier_vec n" by auto from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! 0 $ j mod p) = fs ! 0 $ j" . also have "pm.inv_M (fs ! 0 $ j mod p) = mfs ! 0 $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def using m0 j len fs by auto finally have "mfs ! 0 $ j = fs ! 0 $ j" . } thus "fs ! 0 = mfs ! 0" using LLL_invD_mod(10,13)[OF Linv m0] by auto qed lemma dmu_initial: "dmu_initial = mat m m (\ (i,j). d\ fs_init i j)" proof - interpret fs: fs_int_indpt n fs_init by (unfold_locales, intro lin_dep) show ?thesis unfolding dmu_initial_def Let_def proof (intro cong_mat refl refl, unfold split, goal_cases) case (1 i j) show ?case proof (cases "j \ i") case False thus ?thesis by (auto simp: d\_def gs.\.simps) next case True hence id: "d\_impl fs_init !! i !! j = fs.d\ i j" unfolding fs.d\_impl by (subst of_fun_nth, use 1 len in force, subst of_fun_nth, insert True, auto) also have "\ = d\ fs_init i j" unfolding fs.d\_def d\_def fs.d_def d_def by simp finally show ?thesis using True by auto qed qed qed lemma LLL_initial_invariant_mod: assumes res: "compute_initial_state first = (p, mfs, dmu', g_idx)" shows "\fs b. LLL_invariant_mod fs mfs dmu' p first b 0" proof - from dmu_initial have dmu: "(\i' < m. \j' < m. d\ fs_init i' j' = dmu_initial $$ (i',j'))" by auto obtain b g_idx where norm: "compute_max_gso_norm first dmu_initial = (b,g_idx)" by force note res = res[unfolded compute_initial_state_def Let_def norm split] from res have p: "p = compute_mod_of_max_gso_norm first b" by auto then have p0: "p > 0" unfolding compute_mod_of_max_gso_norm_def using log_base by simp then have p1: "p \ 1" by simp note res = res[folded p] from res[unfolded compute_initial_mfs_def] have mfs: "mfs = map (map_vec (\x. x symmod p)) fs_init" by auto from res[unfolded compute_initial_dmu_def] have dmu': "dmu' = mat m m (\(i',j'). if j' < i' then dmu_initial $$ (i', j') symmod (p * d_of dmu_initial j' * d_of dmu_initial (Suc j')) else dmu_initial $$ (i',j'))" by auto have lat: "lattice_of fs_init = L" by (auto simp: L_def) define I where "I = {(i',j'). i' < m \ j' < i'}" obtain fs where 01: "lattice_of fs = L" and 02: "map (map_vec (\ x. x symmod p)) fs = map (map_vec (\ x. x symmod p)) fs_init" and 03: "lin_indep fs" and 04: "length fs = m" and 05: "(\ k < m. gso fs k = gso fs_init k)" and 06: "(\ k \ m. d fs k = d fs_init k)" and 07: "(\ i' < m. \ j' < m. d\ fs i' j' = (if (i',j') \ I then d\ fs_init i' j' symmod (p * d fs_init j' * d fs_init (Suc j')) else d\ fs_init i' j'))" using mod_finite_set[OF lin_dep len _ lat p0, of I] I_def by blast have inv: "LLL_invariant_weak fs_init" by (intro LLL_inv_wI lat len lin_dep fs_init) have "\i' fs_init i' i' = dmu_initial $$ (i', i')" unfolding dmu_initial by auto from compute_max_gso_norm[OF this inv, of first, unfolded norm] have gbnd: "g_bnd_mode first b fs_init" and b0: "0 \ b" and mb0: "m = 0 \ b = 0" by auto from gbnd 05 have gbnd: "g_bnd_mode first b fs" using g_bnd_mode_cong[of fs fs_init] by auto have d\dmu': "\i'j' fs i' j' = dmu' $$ (i', j')" using 07 dmu d_of_main[of fs_init dmu_initial] unfolding I_def dmu' by simp have wred: "weakly_reduced fs 0" by (simp add: gram_schmidt_fs.weakly_reduced_def) have fs_carr: "set fs \ carrier_vec n" using 03 unfolding gs.lin_indpt_list_def by force have m0: "m \ 0" using len by auto have Linv: "LLL_invariant_weak' 0 fs" by (intro LLL_invI_weak 03 04 01 wred fs_carr m0) note Linvw = LLL_invw'_imp_w[OF Linv] from compute_mod_of_max_gso_norm[OF b0 mb0 p] have p: "mod_invariant b p first" "p > 1" by auto from len mfs have len': "length mfs = m" by auto have modbnd: "\i'j'd\ fs i' j'\ < p * d fs j' * d fs (Suc j')" proof - have "\ i' < m. \ j' < i'. d\ fs i' j' = d\ fs i' j' symmod (p * d fs j' * d fs (Suc j'))" using I_def 07 06 by simp moreover have "\j' < m. p * d fs j' * d fs (Suc j') > 0" using p(2) LLL_d_pos[OF Linvw] by simp ultimately show ?thesis using sym_mod_abs - by (smt Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign less_trans) + by auto (smt (verit, del_insts) Suc_lessD less_trans_Suc) qed have "LLL_invariant_mod fs mfs dmu' p first b 0" using LLL_invI_mod[OF len' m0 04 01 03 wred _ modbnd d\dmu' p(2) gbnd p(1)] 02 mfs by simp then show ?thesis by auto qed subsection \Soundness of Storjohann's algorithm\ text \For all of these abstract algorithms, we actually formulate their soundness proofs by linking to the LLL-invariant (which implies that @{term fs} is reduced (@{term "LLL_invariant True m fs"}) or that the first vector of @{term fs} is short (@{term "LLL_invariant_weak fs \ weakly_reduced fs m"}).\ text \Soundness of Storjohann's algorithm\ lemma reduce_basis_mod_inv: assumes res: "reduce_basis_mod = fs" shows "LLL_invariant True m fs" proof (cases "m = 0") case True from True have *: "fs_init = []" using len by simp moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1) unfolding reduce_basis_mod_def Let_def basis_reduction_mod_main.simps[of _ _ _ _ _ 0] compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def unfolding True * by (auto split: prod.splits) ultimately show ?thesis using True LLL_inv_initial_state by blast next case False let ?first = False obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4) from LLL_initial_invariant_mod[OF init] obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast note res = res[unfolded reduce_basis_mod_def init Let_def split] obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0" by (metis prod.exhaust) obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m" using basis_reduction_mod_main[OF fs0 mfs1dmu1[symmetric]] by auto obtain mfs2 dmu2 where mfs2dmu2: "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust) obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m" and \s: "((\j. j < m \ \_small fs2 j))" using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto have rbd: "LLL_invariant_weak' m fs2" "\j < m. \_small fs2 j" using LLL_invD_mod[OF fs2] LLL_invI_weak \s by auto have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \_small_def by blast have fs: "fs = mfs2" using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] .. have "\i < m. fs2 ! i = fs ! i" proof (intro allI impI) fix i assume i: "i < m" then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i" using fs2 LLL_invariant_mod_to_weak_m_to_i by simp have \si: "\_small fs2 i" using \s i by simp show "fs2 ! i = fs ! i" using basis_reduction_mod_fs_bound(1)[OF fs2i \si i] fs by simp qed then have "fs2 = fs" using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI) then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak LLL_invariant_def by simp qed text \Soundness of Storjohann's algorithm for computing a short vector.\ lemma short_vector_mod_inv: assumes res: "short_vector_mod = v" and m: "m > 0" shows "\ fs. LLL_invariant_weak fs \ weakly_reduced fs m \ v = hd fs" proof - let ?first = True obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4) from LLL_initial_invariant_mod[OF init] obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast obtain p1 mfs1 dmu1 where main: "basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0 = (p1, mfs1, dmu1)" by (metis prod.exhaust) obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m" using basis_reduction_mod_main[OF fs0 main] by auto have "v = hd mfs1" using res[unfolded short_vector_mod_def Let_def init split main] .. with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m have v: "v = hd fs1" by (cases fs1; cases mfs1; auto) from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m" unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto show ?thesis by (intro exI[of _ fs1] conjI Linv1 red v) qed text \Soundness of Storjohann's algorithm with improved swap order\ lemma reduce_basis_iso_inv: assumes res: "reduce_basis_iso = fs" shows "LLL_invariant True m fs" proof (cases "m = 0") case True then have *: "fs_init = []" using len by simp moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1) unfolding reduce_basis_iso_def Let_def basis_reduction_iso_main.simps[of _ _ _ _ _ 0] compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def unfolding True * by (auto split: prod.splits) ultimately show ?thesis using True LLL_inv_initial_state by blast next case False let ?first = False obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4) from LLL_initial_invariant_mod[OF init] obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp note res = res[unfolded reduce_basis_iso_def init Let_def split] obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0" by (metis prod.exhaust) obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m" using basis_reduction_iso_main[OF fs0w mfs1dmu1[symmetric]] by auto obtain mfs2 dmu2 where mfs2dmu2: "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust) obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m" and \s: "((\j. j < m \ \_small fs2 j))" using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto have rbd: "LLL_invariant_weak' m fs2" "\j < m. \_small fs2 j" using LLL_invD_mod[OF fs2] LLL_invI_weak \s by auto have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \_small_def by blast have fs: "fs = mfs2" using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] .. have "\i < m. fs2 ! i = fs ! i" proof (intro allI impI) fix i assume i: "i < m" then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i" using fs2 LLL_invariant_mod_to_weak_m_to_i by simp have \si: "\_small fs2 i" using \s i by simp show "fs2 ! i = fs ! i" using basis_reduction_mod_fs_bound(1)[OF fs2i \si i] fs by simp qed then have "fs2 = fs" using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI) then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak LLL_invariant_def by simp qed text \Soundness of Storjohann's algorithm to compute short vectors with improved swap order\ lemma short_vector_iso_inv: assumes res: "short_vector_iso = v" and m: "m > 0" shows "\ fs. LLL_invariant_weak fs \ weakly_reduced fs m \ v = hd fs" proof - let ?first = True obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4) from LLL_initial_invariant_mod[OF init] obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp obtain p1 mfs1 dmu1 where main: "basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0 = (p1, mfs1, dmu1)" by (metis prod.exhaust) obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m" using basis_reduction_iso_main[OF fs0w main] by auto have "v = hd mfs1" using res[unfolded short_vector_iso_def Let_def init split main] .. with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m have v: "v = hd fs1" by (cases fs1; cases mfs1; auto) from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m" unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto show ?thesis by (intro exI[of _ fs1] conjI Linv1 red v) qed end text \From the soundness results of these abstract versions of the algorithms, one just needs to derive actual implementations that may integrate low-level optimizations.\ end diff --git a/thys/Multi_Party_Computation/ETP_RSA_OT.thy b/thys/Multi_Party_Computation/ETP_RSA_OT.thy --- a/thys/Multi_Party_Computation/ETP_RSA_OT.thy +++ b/thys/Multi_Party_Computation/ETP_RSA_OT.thy @@ -1,705 +1,718 @@ subsubsection \ RSA instantiation \ text\It is known that the RSA collection forms an ETP. Here we instantitate our proof of security for OT that uses a general ETP for RSA. We use the proof of the general construction of OT. The main proof effort here is in showing the RSA collection meets the requirements of an ETP, mainly this involves showing the RSA mapping is a bijection.\ theory ETP_RSA_OT imports ETP_OT Number_Theory_Aux Uniform_Sampling begin type_synonym index = "(nat \ nat)" type_synonym trap = nat type_synonym range = nat type_synonym domain = nat type_synonym viewP1 = "((bool \ bool) \ nat \ nat) spmf" type_synonym viewP2 = "(bool \ index \ (bool \ bool)) spmf" type_synonym dist2 = "(bool \ index \ bool \ bool) \ bool spmf" type_synonym advP2 = "index \ bool \ bool \ dist2 \ bool spmf" locale rsa_base = fixes prime_set :: "nat set" \ \the set of primes used\ and B :: "index \ nat \ bool" assumes prime_set_ass: "prime_set \ {x. prime x \ x > 2}" and finite_prime_set: "finite prime_set" and prime_set_gt_2: "card prime_set > 2" begin lemma prime_set_non_empty: "prime_set \ {}" using prime_set_gt_2 by auto definition coprime_set :: "nat \ nat set" where "coprime_set N \ {x. coprime x N \ x > 1 \ x < N}" lemma coprime_set_non_empty: assumes "N > 2" shows "coprime_set N \ {}" by(simp add: coprime_set_def; metis assms(1) Suc_lessE coprime_Suc_right_nat lessI numeral_2_eq_2) definition sample_coprime :: "nat \ nat spmf" where "sample_coprime N = spmf_of_set (coprime_set (N))" lemma sample_coprime_e_gt_1: assumes "e \ set_spmf (sample_coprime N)" shows "e > 1" using assms by(simp add: sample_coprime_def coprime_set_def) lemma lossless_sample_coprime: assumes "\ prime N" and "N > 2" shows "lossless_spmf (sample_coprime N)" proof- have "coprime_set N \ {}" by(simp add: coprime_set_non_empty assms) also have "finite (coprime_set N)" by(simp add: coprime_set_def) ultimately show ?thesis by(simp add: sample_coprime_def) qed lemma set_spmf_sample_coprime: shows "set_spmf (sample_coprime N) = {x. coprime x N \ x > 1 \ x < N}" by(simp add: sample_coprime_def coprime_set_def) definition sample_primes :: "nat spmf" where "sample_primes = spmf_of_set prime_set" lemma lossless_sample_primes: shows "lossless_spmf sample_primes" by(simp add: sample_primes_def prime_set_non_empty finite_prime_set) lemma set_spmf_sample_primes: shows "set_spmf sample_primes \ {x. prime x \ x > 2}" by(auto simp add: sample_primes_def prime_set_ass finite_prime_set) lemma mem_samp_primes_gt_2: shows "x \ set_spmf sample_primes \ x > 2" apply (simp add: finite_prime_set sample_primes_def) using prime_set_ass by blast lemma mem_samp_primes_prime: shows "x \ set_spmf sample_primes \ prime x" apply (simp add: finite_prime_set sample_primes_def prime_set_ass) using prime_set_ass by blast definition sample_primes_excl :: "nat set \ nat spmf" where "sample_primes_excl P = spmf_of_set (prime_set - P)" lemma lossless_sample_primes_excl: shows "lossless_spmf (sample_primes_excl {P})" apply(simp add: sample_primes_excl_def finite_prime_set) using prime_set_gt_2 subset_singletonD by fastforce definition sample_set_excl :: "nat set \ nat set \ nat spmf" where "sample_set_excl Q P = spmf_of_set (Q - P)" lemma set_spmf_sample_set_excl [simp]: assumes "finite (Q - P)" shows "set_spmf (sample_set_excl Q P) = (Q - P)" unfolding sample_set_excl_def by (metis set_spmf_of_set assms)+ lemma lossless_sample_set_excl: assumes "finite Q" and "card Q > 2" shows "lossless_spmf (sample_set_excl Q {P})" unfolding sample_set_excl_def using assms subset_singletonD by fastforce lemma mem_samp_primes_excl_gt_2: shows "x \ set_spmf (sample_set_excl prime_set {y}) \ x > 2" apply(simp add: finite_prime_set sample_set_excl_def prime_set_ass ) using prime_set_ass by blast lemma mem_samp_primes_excl_prime : shows "x \ set_spmf (sample_set_excl prime_set {y}) \ prime x" apply (simp add: finite_prime_set sample_set_excl_def) using prime_set_ass by blast lemma sample_coprime_lem: assumes "x \ set_spmf sample_primes" and " y \ set_spmf (sample_set_excl prime_set {x}) " shows "lossless_spmf (sample_coprime ((x - Suc 0) * (y - Suc 0)))" proof- have gt_2: "x > 2" "y > 2" using mem_samp_primes_gt_2 assms mem_samp_primes_excl_gt_2 by auto have "\ prime ((x-1)*(y-1))" proof- have "prime x" "prime y" using mem_samp_primes_prime mem_samp_primes_excl_prime assms by auto then show ?thesis using prod_not_prime gt_2 by simp qed also have "((x-1)*(y-1)) > 2" by (metis (no_types, lifting) gt_2 One_nat_def Suc_diff_1 assms(1) assms(2) calculation divisors_zero less_2_cases nat_1_eq_mult_iff nat_neq_iff not_numeral_less_one numeral_2_eq_2 prime_gt_0_nat rsa_base.mem_samp_primes_excl_prime rsa_base.mem_samp_primes_prime rsa_base_axioms two_is_prime_nat) ultimately show ?thesis using lossless_sample_coprime by simp qed definition I :: "(index \ trap) spmf" where "I = do { P \ sample_primes; Q \ sample_set_excl prime_set {P}; let N = P*Q; let N' = (P-1)*(Q-1); e \ sample_coprime N'; let d = nat ((fst (bezw e N')) mod N'); return_spmf ((N, e), d)}" lemma lossless_I: "lossless_spmf I" by(auto simp add: I_def lossless_sample_primes lossless_sample_set_excl finite_prime_set prime_set_gt_2 Let_def sample_coprime_lem) lemma set_spmf_I_N: assumes "((N,e),d) \ set_spmf I" obtains P Q where "N = P * Q" and "P \ Q" and "prime P" and "prime Q" and "coprime e ((P - 1)*(Q - 1))" and "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" using assms apply(auto simp add: I_def Let_def) using finite_prime_set mem_samp_primes_prime sample_set_excl_def rsa_base_axioms sample_primes_def by (simp add: set_spmf_sample_coprime) lemma set_spmf_I_e_d: - assumes "((N,e),d) \ set_spmf I" - shows "e > 1" and "d > 1" - using assms sample_coprime_e_gt_1 - apply(auto simp add: I_def Let_def) - by (smt Euclidean_Division.pos_mod_sign Num.of_nat_simps(5) Suc_diff_1 bezw_inverse cong_def coprime_imp_gcd_eq_1 gr0I less_1_mult less_numeral_extra(2) mem_Collect_eq mod_by_0 mod_less more_arith_simps(6) nat_0 nat_0_less_mult_iff nat_int nat_neq_iff numerals(2) of_nat_0_le_iff of_nat_1 rsa_base.mem_samp_primes_gt_2 rsa_base_axioms set_spmf_sample_coprime zero_less_diff) + \e > 1\ \d > 1\ if \((N, e), d) \ set_spmf I\ +proof - + from that obtain M where + e: \e \ set_spmf (sample_coprime M)\ + and d: \d = nat (fst (bezw e M) mod M)\ + by (auto simp add: I_def Let_def) + from e set_spmf_sample_coprime [of M] + have \coprime e M\ \1 < e\ \e < M\ + by simp_all + then have \2 < M\ + by simp + from \1 < e\ show \e > 1\. + from d \coprime e M\ bezw_inverse [of e M] + have \[e * d = 1] (mod M)\ + by simp + with \e > 1\ \2 < M\ show \d > 1\ + by (cases \d = 0 \ d = 1\) (auto simp add: \e < M\ cong_def) +qed definition domain :: "index \ nat set" where "domain index \ {..< fst index}" definition range :: "index \ nat set" where "range index \ {..< fst index}" lemma finite_range: "finite (range index)" by(simp add: range_def) lemma dom_eq_ran: "domain index = range index" by(simp add: range_def domain_def) definition F :: "index \ (nat \ nat)" where "F index x = x ^ (snd index) mod (fst index)" definition F\<^sub>i\<^sub>n\<^sub>v :: "index \ trap \ nat \ nat" where "F\<^sub>i\<^sub>n\<^sub>v \ \ y = y ^ \ mod (fst \)" text \ We must prove the RSA function is a bijection \ lemma rsa_bijection: assumes coprime: "coprime e ((P-1)*(Q-1))" and prime_P: "prime (P::nat)" and prime_Q: "prime Q" and P_neq_Q: "P \ Q" and x_lt_pq: "x < P * Q" and y_lt_pd: "y < P * Q" and rsa_map_eq: "x ^ e mod (P * Q) = y ^ e mod (P * Q)" shows "x = y" proof- have flt_xP: "[x ^ P = x] (mod P)" using fermat_little prime_P by blast have flt_yP: "[y ^ P = y] (mod P)" using fermat_little prime_P by blast have flt_xQ: "[x ^ Q = x] (mod Q)" using fermat_little prime_Q by blast have flt_yQ: "[y ^ Q = y] (mod Q)" using fermat_little prime_Q by blast show ?thesis proof(cases "y \ x") case True hence ye_gt_xe: "y^e \ x^e" by (simp add: power_mono) have x_y_exp_e: "[x^e = y^e] (mod P)" using cong_modulus_mult_nat cong_altdef_nat True ye_gt_xe cong_sym cong_def assms by blast obtain d where d: "[e*d = 1] (mod (P-1)) \ d \ 0" using ex_inverse assms by blast then obtain k where k: "e*d = 1 + k*(P-1)" using ex_k_mod assms by blast hence xk_yk: "[x^(1 + k*(P-1)) = y^(1 + k*(P-1))] (mod P)" by(metis k power_mult x_y_exp_e cong_pow) have xk_x: "[x^(1 + k*(P-1)) = x] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[x ^ (1 + k * (P - 1)) = x] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k * (P - 1)) = x] (mod P)" using asm by simp hence "[x ^ (k * (P - 1)) * x ^ P = x] (mod P)" using flt_xP by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k * (P - 1) + P) = x] (mod P)" by (simp add: power_add) hence "[x ^ (1 + (k + 1) * (P - 1)) = x] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y: "[y^(1 + k*(P-1)) = y] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[y ^ (1 + k * (P - 1)) = y] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k * (P - 1)) = y] (mod P)" using asm by simp hence "[y ^ (k * (P - 1)) * y ^ P = y] (mod P)" using flt_yP by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k * (P - 1) + P) = y] (mod P)" by (simp add: power_add) hence "[y ^ (1 + (k + 1) * (P - 1)) = y] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have "[x^e = y^e] (mod Q)" by (metis rsa_map_eq cong_modulus_mult_nat cong_def mult.commute) obtain d' where d': "[e*d' = 1] (mod (Q-1)) \ d' \ 0" by (metis mult.commute ex_inverse prime_P prime_Q P_neq_Q coprime) then obtain k' where k': "e*d' = 1 + k'*(Q-1)" by(metis ex_k_mod mult.commute prime_P prime_Q P_neq_Q coprime) hence xk_yk': "[x^(1 + k'*(Q-1)) = y^(1 + k'*(Q-1))] (mod Q)" by(metis k' power_mult \[x ^ e = y ^ e] (mod Q)\ cong_pow) have xk_x': "[x^(1 + k'*(Q-1)) = x] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[x ^ (1 + k' * (Q - 1)) = x] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k' * (Q - 1)) = x] (mod Q)" using asm by simp hence "[x ^ (k' * (Q - 1)) * x ^ Q = x] (mod Q)" using flt_xQ by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k' * (Q - 1) + Q) = x] (mod Q)" by (simp add: power_add) hence "[x ^ (1 + (k' + 1) * (Q - 1)) = x] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y': "[y^(1 + k'*(Q-1)) = y] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[y ^ (1 + k' * (Q - 1)) = y] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k' * (Q - 1)) = y] (mod Q)" using asm by simp hence "[y ^ (k' * (Q - 1)) * y ^ Q = y] (mod Q)" using flt_yQ by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k' * (Q - 1) + Q) = y] (mod Q)" by (simp add: power_add) hence "[y ^ (1 + (k' + 1) * (Q - 1)) = y] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have P_dvd_xy: "P dvd (y - x)" proof- have "[x = y] (mod P)" using xk_x yk_y xk_yk by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym True by blast qed have Q_dvd_xy: "Q dvd (y - x)" proof- have "[x = y] (mod Q)" using xk_x' yk_y' xk_yk' by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym True by blast qed show ?thesis proof- have "P*Q dvd (y - x)" using P_dvd_xy Q_dvd_xy by (simp add: assms divides_mult primes_coprime) then have "[x = y] (mod P*Q)" by (simp add: cong_altdef_nat cong_sym True) hence "x mod P*Q = y mod P*Q" using cong_def xk_x xk_yk yk_y by metis then show ?thesis using \[x = y] (mod P * Q)\ cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast qed next case False hence ye_gt_xe: "x^e \ y^e" by (simp add: power_mono) have pow_eq: "[x^e = y^e] (mod P*Q)" by (simp add: cong_def assms) then have PQ_dvd_ye_xe: "(P*Q) dvd (x^e - y^e)" using cong_altdef_nat False ye_gt_xe cong_sym by blast then have "[x^e = y^e] (mod P)" using cong_modulus_mult_nat pow_eq by blast obtain d where d: "[e*d = 1] (mod (P-1)) \ d \ 0" using ex_inverse assms by blast then obtain k where k: "e*d = 1 + k*(P-1)" using ex_k_mod assms by blast have xk_yk: "[x^(1 + k*(P-1)) = y^(1 + k*(P-1))] (mod P)" proof- have "[(x^e)^d = (y^e)^d] (mod P)" using \[x ^ e = y ^ e] (mod P)\ cong_pow by blast then have "[x^(e*d) = y^(e*d)] (mod P)" by (simp add: power_mult) thus ?thesis using k by simp qed have xk_x: "[x^(1 + k*(P-1)) = x] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[x ^ (1 + k * (P - 1)) = x] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k * (P - 1)) = x] (mod P)" using asm by simp hence "[x ^ (k * (P - 1)) * x ^ P = x] (mod P)" using flt_xP by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k * (P - 1) + P) = x] (mod P)" by (simp add: power_add) hence "[x ^ (1 + (k + 1) * (P - 1)) = x] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y: "[y^(1 + k*(P-1)) = y] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[y ^ (1 + k * (P - 1)) = y] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k * (P - 1)) = y] (mod P)" using asm by simp hence "[y ^ (k * (P - 1)) * y ^ P = y] (mod P)" using flt_yP by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k * (P - 1) + P) = y] (mod P)" by (simp add: power_add) hence "[y ^ (1 + (k + 1) * (P - 1)) = y] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have P_dvd_xy: "P dvd (x - y)" proof- have "[x = y] (mod P)" using xk_x yk_y xk_yk by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym False by simp qed have "[x^e = y^e] (mod Q)" using cong_modulus_mult_nat pow_eq PQ_dvd_ye_xe cong_dvd_modulus_nat dvd_triv_right by blast obtain d' where d': "[e*d' = 1] (mod (Q-1)) \ d' \ 0" by (metis mult.commute ex_inverse prime_P prime_Q coprime P_neq_Q) then obtain k' where k': "e*d' = 1 + k'*(Q-1)" by(metis ex_k_mod mult.commute prime_P prime_Q coprime P_neq_Q) have xk_yk': "[x^(1 + k'*(Q-1)) = y^(1 + k'*(Q-1))] (mod Q)" proof- have "[(x^e)^d' = (y^e)^d'] (mod Q)" using \[x ^ e = y ^ e] (mod Q)\ cong_pow by blast then have "[x^(e*d') = y^(e*d')] (mod Q)" by (simp add: power_mult) thus ?thesis using k' by simp qed have xk_x': "[x^(1 + k'*(Q-1)) = x] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[x ^ (1 + k' * (Q - 1)) = x] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k' * (Q - 1)) = x] (mod Q)" using asm by simp hence "[x ^ (k' * (Q - 1)) * x ^ Q = x] (mod Q)" using flt_xQ by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k' * (Q - 1) + Q) = x] (mod Q)" by (simp add: power_add) hence "[x ^ (1 + (k' + 1) * (Q - 1)) = x] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y': "[y^(1 + k'*(Q-1)) = y] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[y ^ (1 + k' * (Q - 1)) = y] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k' * (Q - 1)) = y] (mod Q)" using asm by simp hence "[y ^ (k' * (Q - 1)) * y ^ Q = y] (mod Q)" using flt_yQ by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k' * (Q - 1) + Q) = y] (mod Q)" by (simp add: power_add) hence "[y ^ (1 + (k' + 1) * (Q - 1)) = y] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have Q_dvd_xy: "Q dvd (x - y)" proof- have "[x = y] (mod Q)" using xk_x' yk_y' xk_yk' by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym False by simp qed show ?thesis proof- have "P*Q dvd (x - y)" using P_dvd_xy Q_dvd_xy by (simp add: assms divides_mult primes_coprime) hence 1: "[x = y] (mod P*Q)" using False cong_altdef_nat linear by blast hence "x mod P*Q = y mod P*Q" using cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast thus ?thesis using 1 cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast qed qed qed lemma rsa_bij_betw: assumes "coprime e ((P - 1)*(Q - 1))" and "prime P" and "prime Q" and "P \ Q" shows "bij_betw (F ((P * Q), e)) (range ((P * Q), e)) (range ((P * Q), e))" proof- have PQ_not_0: "prime P \ prime Q \ P * Q \ 0" using assms by auto have "inj_on (\x. x ^ snd (P * Q, e) mod fst (P * Q, e)) {..x. x ^ snd (P * Q, e) mod fst (P * Q, e)) ` {.. set_spmf I" shows "bij_betw (F ((N), e)) (range ((N), e)) (range ((N), e))" proof- obtain P Q where "N = P * Q" and "bij_betw (F ((P*Q), e)) (range ((P*Q), e)) (range ((P*Q), e))" proof- obtain P Q where "prime P" and "prime Q" and "N = P * Q" and "P \ Q" and "coprime e ((P - 1)*(Q - 1))" using set_spmf_I_N assms by blast then show ?thesis using rsa_bij_betw that by blast qed thus ?thesis by blast qed lemma assumes "P dvd x" shows "[x = 0] (mod P)" using assms cong_def by force lemma rsa_inv: assumes d: "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" and coprime: "coprime e ((P-1)*(Q-1))" and prime_P: "prime (P::nat)" and prime_Q: "prime Q" and P_neq_Q: "P \ Q" and e_gt_1: "e > 1" and d_gt_1: "d > 1" shows "((((x) ^ e) mod (P*Q)) ^ d) mod (P*Q) = x mod (P*Q)" proof(cases "x = 0 \ x = 1") case True then show ?thesis by (metis assms(6) assms(7) le_simps(1) nat_power_eq_Suc_0_iff neq0_conv not_one_le_zero numeral_nat(7) power_eq_0_iff power_mod) next case False hence x_gt_1: "x > 1" by simp define z where "z = (x ^ e) ^ d - x" hence z_gt_0: "z > 0" proof- have "(x ^ e) ^ d - x = x ^ (e * d) - x" by (simp add: power_mult) also have "... > 0" by (metis x_gt_1 e_gt_1 d_gt_1 le_neq_implies_less less_one linorder_not_less n_less_m_mult_n not_less_zero numeral_nat(7) power_increasing_iff power_one_right zero_less_diff) ultimately show ?thesis using z_def by argo qed hence "[z = 0] (mod P)" proof(cases "[x = 0] (mod P)") case True then show ?thesis proof - have "0 \ d * e" by (metis (no_types) assms assms mult_is_0 not_one_less_zero) then show ?thesis by (metis (no_types) Groups.add_ac(2) True add_diff_inverse_nat cong_def cong_dvd_iff cong_mult_self_right dvd_0_right dvd_def dvd_trans mod_add_self2 more_arith_simps(5) nat_diff_split power_eq_if power_mult semiring_normalization_rules(7) z_def) qed next case False have "[e * d = 1] (mod ((P - 1) * (Q - 1)))" by (metis d bezw_inverse coprime coprime_imp_gcd_eq_1 nat_int) hence "[e * d = 1] (mod (P - 1))" using assms cong_modulus_mult_nat by blast then obtain k where k: "e*d = 1 + k*(P-1)" using ex_k_mod assms by force hence "x ^ (e * d) = x * ((x ^ (P - 1)) ^ k)" by (metis power_add power_one_right mult.commute power_mult) hence "[x ^ (e * d) = x * ((x ^ (P - 1)) ^ k)] (mod P)" using cong_def by simp moreover have "[x ^ (P - 1) = 1] (mod P)" using prime_P fermat_theorem False by (simp add: cong_0_iff) moreover have "[x ^ (e * d) = x * ((1) ^ k)] (mod P)" by (metis \x ^ (e * d) = x * (x ^ (P - 1)) ^ k\ calculation(2) cong_pow cong_scalar_left) hence "[x ^ (e * d) = x] (mod P)" by simp thus ?thesis using z_def z_gt_0 by (simp add: cong_diff_iff_cong_0_nat power_mult) qed moreover have "[z = 0] (mod Q)" proof(cases "[x = 0] (mod Q)") case True then show ?thesis proof - have "0 \ d * e" by (metis (no_types) assms mult_is_0 not_one_less_zero) then show ?thesis by (metis (no_types) Groups.add_ac(2) True add_diff_inverse_nat cong_def cong_dvd_iff cong_mult_self_right dvd_0_right dvd_def dvd_trans mod_add_self2 more_arith_simps(5) nat_diff_split power_eq_if power_mult semiring_normalization_rules(7) z_def) qed next case False have "[e * d = 1] (mod ((P - 1) * (Q - 1)))" by (metis d bezw_inverse coprime coprime_imp_gcd_eq_1 nat_int) hence "[e * d = 1] (mod (Q - 1))" using assms cong_modulus_mult_nat mult.commute by metis then obtain k where k: "e*d = 1 + k*(Q-1)" using ex_k_mod assms by force hence "x ^ (e * d) = x * ((x ^ (Q - 1)) ^ k)" by (metis power_add power_one_right mult.commute power_mult) hence "[x ^ (e * d) = x * ((x ^ (Q - 1)) ^ k)] (mod P)" using cong_def by simp moreover have "[x ^ (Q - 1) = 1] (mod Q)" using prime_Q fermat_theorem False by (simp add: cong_0_iff) moreover have "[x ^ (e * d) = x * ((1) ^ k)] (mod Q)" by (metis \x ^ (e * d) = x * (x ^ (Q - 1)) ^ k\ calculation(2) cong_pow cong_scalar_left) hence "[x ^ (e * d) = x] (mod Q)" by simp thus ?thesis using z_def z_gt_0 by (simp add: cong_diff_iff_cong_0_nat power_mult) qed ultimately have "Q dvd (x ^ e) ^ d - x" "P dvd (x ^ e) ^ d - x" using z_def assms cong_0_iff by blast + hence "P * Q dvd ((x ^ e) ^ d - x)" using assms divides_mult primes_coprime_nat by blast hence "[(x ^ e) ^ d = x] (mod (P * Q))" using z_gt_0 cong_altdef_nat z_def by auto thus ?thesis by (simp add: cong_def power_mod) qed lemma rsa_inv_set_spmf_I: assumes "((N, e), d) \ set_spmf I" shows "((((x::nat) ^ e) mod N) ^ d) mod N = x mod N" proof- obtain P Q where "N = P * Q" and "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" and "prime P" and "prime Q" and "coprime e ((P - 1)*(Q - 1))" and "P \ Q" using assms set_spmf_I_N by blast moreover have "e > 1" and "d > 1" using set_spmf_I_e_d assms by auto ultimately show ?thesis using rsa_inv by blast qed sublocale etp_rsa: etp I domain range F F\<^sub>i\<^sub>n\<^sub>v unfolding etp_def apply(auto simp add: etp_def dom_eq_ran finite_range bij_betw1 lossless_I) apply (metis fst_conv lessThan_iff mem_simps(2) nat_0_less_mult_iff prime_gt_0_nat range_def set_spmf_I_N) apply(auto simp add: F_def F\<^sub>i\<^sub>n\<^sub>v_def) using rsa_inv_set_spmf_I by (simp add: range_def) sublocale etp: ETP_base I domain range B F F\<^sub>i\<^sub>n\<^sub>v unfolding ETP_base_def by (simp add: etp_rsa.etp_axioms) text\After proving the RSA collection is an ETP the proofs of security come easily from the general proofs.\ lemma correctness_rsa: "etp.OT_12.correctness m1 m2" by (rule local.etp.correct) lemma P1_security_rsa: "etp.OT_12.perfect_sec_P1 m1 m2" by(rule local.etp.P1_security_inf_the) lemma P2_security_rsa: assumes "\ a. lossless_spmf (D a)" and "\b\<^sub>\. local.etp_rsa.HCP_adv etp.\ m2 b\<^sub>\ D \ HCP_ad" shows "etp.OT_12.adv_P2 m1 m2 D \ 2 * HCP_ad" by(simp add: local.etp.P2_security assms) end locale rsa_asym = fixes prime_set :: "nat \ nat set" and B :: "index \ nat \ bool" assumes rsa_proof_assm: "\ n. rsa_base (prime_set n)" begin sublocale rsa_base "(prime_set n)" B using local.rsa_proof_assm by simp lemma correctness_rsa_asymp: shows "etp.OT_12.correctness n m1 m2" by(rule correctness_rsa) lemma P1_sec_asymp: "etp.OT_12.perfect_sec_P1 n m1 m2" by(rule local.P1_security_rsa) lemma P2_sec_asym: assumes "\ a. lossless_spmf (D a)" and HCP_adv_neg: "negligible (\ n. hcp_advantage n)" and hcp_adv_bound: "\b\<^sub>\ n. local.etp_rsa.HCP_adv n etp.\ m2 b\<^sub>\ D \ hcp_advantage n" shows "negligible (\ n. etp.OT_12.adv_P2 n m1 m2 D)" proof- have "negligible (\ n. 2 * hcp_advantage n)" using HCP_adv_neg by (simp add: negligible_cmultI) moreover have "\etp.OT_12.adv_P2 n m1 m2 D\ = etp.OT_12.adv_P2 n m1 m2 D" for n unfolding sim_det_def.adv_P2_def local.etp.OT_12.adv_P2_def by linarith moreover have "etp.OT_12.adv_P2 n m1 m2 D \ 2 * hcp_advantage n" for n using P2_security_rsa assms by blast ultimately show ?thesis using assms negligible_le by presburger qed end end \ No newline at end of file diff --git a/thys/Number_Theoretic_Transform/Butterfly.thy b/thys/Number_Theoretic_Transform/Butterfly.thy --- a/thys/Number_Theoretic_Transform/Butterfly.thy +++ b/thys/Number_Theoretic_Transform/Butterfly.thy @@ -1,1323 +1,1323 @@ (* Title: Butterfly Algorithm for Number Theoretic Transform Author: Thomas Ammer *) theory Butterfly imports NTT "HOL-Library.Discrete" begin text \\pagebreak\ section \Butterfly Algorithms\ text \\label{Butterfly}\ text \\indent Several recursive algorithms for $FFT$ based on the divide and conquer principle have been developed in order to speed up the transform. A method for reducing complexity is the butterfly scheme. In this formalization, we consider the butterfly algorithm by Cooley and Tukey~\parencite{Good1997} adapted to the setting of \textit{NTT}. \ text \\noindent We additionally assume that $n$ is power of two.\ locale butterfly = ntt + fixes N assumes n_two_pot: "n = 2^N" begin subsection \Recursive Definition\ text \Let's recall the definition of a transformed vector element: \begin{equation*} \mathsf{NTT}(\vec{x})_i = \sum _{j = 0} ^{n-1} x_j \cdot \omega ^{i\cdot j} \end{equation*} We assume $n = 2^N$ and obtain: \begin{align*} \sum _{j = 0} ^{< 2^N} x_j \cdot \omega ^{i\cdot j} \\ &= \sum _{j = 0} ^{< 2^{N-1}} x_{2j} \cdot \omega ^{i\cdot 2j} + \sum _{j = 0} ^{< 2^{N-1}} x_{2j+1} \cdot \omega ^{i\cdot (2j+1)} \\ & = \sum _{j = 0} ^{< 2^{N-1}} x_{2j} \cdot (\omega^2) ^{i\cdot j} + \omega^i \cdot \sum _{j = 0} ^{< 2^{N-1}} x_{2j+1} \cdot (\omega^2) ^{i\cdot j}\\ & = (\sum _{j = 0} ^{< 2^{N-2}} x_{4j} \cdot (\omega^4) ^{i\cdot j} + \omega^i \cdot \sum _{j = 0} ^{< 2^{N-2}} x_{4j+2} \cdot (\omega^4) ^{i\cdot j}) \\ & \hspace{1cm}+ \omega^i \cdot (\sum _{j = 0} ^{< 2^{N-2}} x_{4j+1} \cdot (\omega^4) ^{i\cdot j} + \omega^i \cdot \sum _{j = 0} ^{< 2^{N-2}} x_{4j+3} \cdot (\omega^4) ^{i\cdot j}) \text{ etc.} \end{align*} which gives us a recursive algorithm: \begin{itemize} \item Compose vectors consisting of elements at even and odd indices respectively \item Compute a transformation of these vectors recursively where the dimensions are halved. \item Add results after scaling the second subresult by $\omega^i$ \end{itemize} \ text \Now we give a functional definition of the analogue to $FFT$ adapted to finite fields. A gentle introduction to $FFT$ can be found in~\parencite{10.5555/1614191}. For the fast implementation of Number Theoretic Transform in particular, have a look at~\parencite{cryptoeprint:2016/504}.\ text \(The following lemma is needed to obtain an automated termination proof of $FNTT$.)\ lemma FNTT_termination_aux [simp]: "length (filter P [0..Please note that we closely adhere to the textbook definition which just talks about elements at even and odd indices. We model the informal definition by predefined functions, since this seems to be more handy during proofs. An algorithm splitting the elements smartly will be presented afterwards.\ fun FNTT::"('a mod_ring) list \ ('a mod_ring) list" where "FNTT [] = []"| "FNTT [a] = [a]"| "FNTT nums = (let nn = length nums; nums1 = [nums!i. i \ filter even [0.. filter odd [0.. x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]); sum2 = map2 (-) fntt1 (map2 ( \ x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]) in sum1@sum2)" lemmas [simp del] = FNTT_termination_aux text \ Finally, we want to prove correctness, i.e. $FNTT\; xs = NTT\;xs$. Since we consider a recursive algorithm, some kind of induction is appropriate: Assume the claim for $\frac{2^d}{2} = 2^{d-1}$ and prove it for $2^d$, where $2^d$ is the vector length. This implies that we have to talk about \textit{NTT}s with respect to some powers of $\omega$. In particular, we decide to annotate \textit{NTT} with a degree $degr$ indicating the referred vector length. There is a correspondence to the current level $l$ of recursion: \begin{equation*} degr = 2^{N-l} \end{equation*} \ text \\noindent A generalized version of \textit{NTT} keeps track of all levels during recursion:\ definition "ntt_gen numbers degr i = (\j=0..<(length numbers). (numbers ! j) * \^((n div degr)*i*j)) " definition "NTT_gen degr numbers = map (ntt_gen numbers (degr)) [0..< length numbers]" text \Whenever generalized \textit{NTT} is applied to a list of full length, then its actually equal to the defined \textit{NTT}.\ lemma NTT_gen_NTT_full_length: assumes "length numbers =n" shows "NTT_gen n numbers = NTT numbers" unfolding NTT_gen_def ntt_gen_def NTT_def ntt_def using assms by simp subsection \Arguments on Correctness\ text \First some general lemmas on list operations.\ lemma length_even_filter: "length [f i . i <- (filter even [0.. i < length ys \ (map2 f xs ys) ! i = f (xs ! i) (ys ! i)" by (induction xs arbitrary: ys i) auto lemma filter_last_not: "\ P x \ filter P (xs@[x]) = filter P xs" by simp lemma filter_even_map: "filter even [0..<2*(x::nat)] = map ((*) (2::nat)) [0.. 2*x = l \ (filter even [0.. y. (2::nat)*y +1) [0.. 2*x = l \ (filter odd [0..\noindent Lemmas by using the assumption $n = 2^N$.\ text \\noindent ($-1$ denotes the additive inverse of $1$ in the finite field.)\ lemma n_min1_2: "n = 2 \ \ = -1" using omega_properties(1) omega_properties(2) power2_eq_1_iff by blast lemma n_min1_gr2: assumes "n > 2" shows "\^(n div 2) = -1" proof- have "\^(n div 2) \ -1 \ False" proof- assume "\^(n div 2) \ -1" hence False - proof(cases "\^(n div 2) = 1") -case True - then show ?thesis using omega_properties(3) - by (metis Euclidean_Division.div_eq_0_iff div_less_dividend leD less_le_trans n_lst2 one_less_numeral_iff pos2 semiring_norm(76) zero_neq_numeral) -next - case False - hence "(\^(n div 2)) ^ (2::nat) \ 1" - by (smt (verit, ccfv_threshold) n_two_pot One_nat_def \\ ^ (n div 2) \ - 1\ diff_zero leD n_lst2 not_less_eq omega_properties(1) one_less_numeral_iff one_power2 power2_eq_square power_mult power_one_right power_strict_increasing_iff semiring_norm(76) square_eq_iff two_powr_div two_powrs_div) - moreover have "(n div 2) * 2 = n" using n_two_pot n_lst2 - by (metis One_nat_def Suc_lessD assms div_by_Suc_0 one_less_numeral_iff power_0 power_one_right power_strict_increasing_iff semiring_norm(76) two_powrs_div) - ultimately show ?thesis using omega_properties(1) - by (metis power_mult) -qed - thus False by simp -qed + proof (cases \\ ^ (n div 2) = 1\) + case True + then show ?thesis using omega_properties(3) assms + by auto + next + case False + hence "(\^(n div 2)) ^ (2::nat) \ 1" + by (smt (verit, ccfv_threshold) n_two_pot One_nat_def \\ ^ (n div 2) \ - 1\ diff_zero leD n_lst2 not_less_eq omega_properties(1) one_less_numeral_iff one_power2 power2_eq_square power_mult power_one_right power_strict_increasing_iff semiring_norm(76) square_eq_iff two_powr_div two_powrs_div) + moreover have "(n div 2) * 2 = n" using n_two_pot n_lst2 + by (metis One_nat_def Suc_lessD assms div_by_Suc_0 one_less_numeral_iff power_0 power_one_right power_strict_increasing_iff semiring_norm(76) two_powrs_div) + ultimately show ?thesis using omega_properties(1) + by (metis power_mult) + qed + thus False by simp + qed then show ?thesis by auto qed lemma div_exp_sub: "2^l < n \ n div (2^l) = 2^(N-l)"using n_two_pot by (smt (z3) One_nat_def diff_is_0_eq diff_le_diff_pow div_if div_le_dividend eq_imp_le le_0_eq le_Suc_eq n_lst2 nat_less_le not_less_eq_eq numeral_2_eq_2 power_0 two_powr_div) lemma omega_div_exp_min1: assumes "2^(Suc l) \ n" shows "(\ ^(n div 2^(Suc l)))^(2^l) = -1" proof- have "(\ ^(n div 2^(Suc l)))^(2^l) = \ ^((n div 2^(Suc l))*2^l)" by (simp add: power_mult) moreover have "(n div 2^(Suc l)) = 2^(N - Suc l)" using assms div_exp_sub by (metis n_two_pot eq_imp_le le_neq_implies_less one_less_numeral_iff power_diff power_inject_exp semiring_norm(76) zero_neq_numeral) moreover have "N \ Suc l" using assms n_two_pot by (metis diff_is_0_eq diff_le_diff_pow gr0I leD le_refl) moreover hence "(2::nat)^(N - Suc l)*2^l = 2^(N- 1)" by (metis Nat.add_diff_assoc diff_Suc_1 diff_diff_cancel diff_le_self le_add1 le_add_diff_inverse plus_1_eq_Suc power_add) ultimately show ?thesis by (metis n_two_pot One_nat_def \n div 2 ^ Suc l = 2 ^ (N - Suc l)\ diff_Suc_1 div_exp_sub n_lst2 n_min1_2 n_min1_gr2 nat_less_le nat_power_eq_Suc_0_iff one_less_numeral_iff power_inject_exp power_one_right semiring_norm(76)) qed lemma omg_n_2_min1: "\^(n div 2) = -1" by (metis n_lst2 n_min1_2 n_min1_gr2 nat_less_le numeral_Bit0_div_2 numerals(1) power_one_right) lemma neg_cong: "-(x::('a mod_ring)) = - y \ x = y" by simp text \Generalized \textit{NTT} indeed describes all recursive levels, and thus, it is actually equivalent to the ordinary \textit{NTT} definition.\ theorem FNTT_NTT_gen_eq: "length numbers = 2^l \ 2^l \ n \ FNTT numbers = NTT_gen (length numbers) numbers" proof(induction l arbitrary: numbers) case 0 then show ?case unfolding NTT_gen_def ntt_gen_def by (auto simp: length_Suc_conv) next case (Suc l) text \We define some lists that are used during the recursive call.\ define numbers1 where "numbers1 = [numbers!i . i <- (filter even [0.. x k. x*(\^( (n div (length numbers)) * k))) fntt2 [0..<((length numbers) div 2)])" define sum2 where "sum2 = map2 (-) fntt1 (map2 ( \ x k. x*(\^( (n div (length numbers)) * k))) fntt2 [0..<((length numbers) div 2)])" define l1 where "l1 = length numbers1" define l2 where "l2 = length numbers2" define llen where "llen = length numbers" text \Properties of those lists.\ have numbers1_even: "length numbers1 = 2^l" using numbers1_def length_even_filter Suc by simp have numbers2_even: "length numbers2 = 2^l" using numbers2_def length_odd_filter Suc by simp have numbers1_fntt: "fntt1 = NTT_gen (2^l) numbers1" using fntt1_def Suc.IH[of numbers1] numbers1_even Suc(3) by simp hence fntt1_by_index: "fntt1 ! i = ntt_gen numbers1 (2^l) i" if "i < 2^l" for i unfolding NTT_gen_def by (simp add: numbers1_even that) have numbers2_fntt: "fntt2 = NTT_gen (2^l) numbers2" using fntt2_def Suc.IH[of numbers2] numbers2_even Suc(3) by simp hence fntt2_by_index: "fntt2 ! i = ntt_gen numbers2 (2^l) i" if "i < 2^l" for i unfolding NTT_gen_def by (simp add: numbers2_even that) have fntt1_length: "length fntt1 = 2^l" unfolding numbers1_fntt NTT_gen_def numbers1_def using numbers1_def numbers1_even by force have fntt2_length: "length fntt2 = 2^l" unfolding numbers2_fntt NTT_gen_def numbers2_def using numbers2_def numbers2_even by force text \We show that the list resulting from $FNTT$ is equal to the $NTT$ list. First, we prove $FNTT$ and $NTT$ to be equal concerning their first halves.\ have before_half: "map (ntt_gen numbers llen) [0..<(llen div 2)] = sum1" proof- text \Length is important, since we want to use list lemmas later on.\ have 00:"length (map (ntt_gen numbers llen) [0..<(llen div 2)]) = length sum1" unfolding sum1_def llen_def using Suc(2) map2_length[of _ fntt2 "[0..x y. x * \ ^ (n div length numbers * y)) fntt2 [0..We show equality by extensionality w.r.t. indices.\ have 02:"(map (ntt_gen numbers llen) [0..<(llen div 2)]) ! i = sum1 ! i" if "i < 2^l" for i proof- text \First simplify this term.\ have 000:"(map (ntt_gen numbers llen) [0..<(llen div 2)]) ! i = ntt_gen numbers llen i" using "00" "01" that by auto text \Expand the definition of $sum1$ and massage the result.\ moreover have 001:"sum1 ! i = (fntt1!i) + (fntt2!i) * (\^((n div llen) * i))" unfolding sum1_def using map2_index "00" "01" NTT_gen_def add.left_neutral diff_zero fntt1_length length_map length_upt map2_map_map map_nth nth_upt numbers2_even numbers2_fntt that llen_def by force moreover have 002:"(fntt1!i) = (\j=0..^((n div (2^l))*i*j))" unfolding l1_def using fntt1_by_index[of i] that unfolding ntt_gen_def by simp have 003:"... = (\j=0..^((n div llen)*i*(2*j)))" apply (rule sum_rules(2)) subgoal for j unfolding numbers1_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l1 " hence "map ((!) numbers) (filter even [0.. ^ (n div 2 ^ l * i * j) = numbers ! (2 * j) * \ ^ (n div llen * i * (2 * j))" unfolding llen_def l1_def l2_def by (metis (mono_tags, lifting) mult.assoc mult.left_commute) qed done moreover have 004: "(fntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div (2^l))*i*j+ (n div llen) * i))" apply(rule trans[where s = "(\j = 0.. ^ (n div 2 ^ l * i * j) * \ ^ (n div llen * i))"]) subgoal unfolding l2_def llen_def using fntt2_by_index[of i] that sum_in[of _ "(\^((n div llen) * i))" "l2"] comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding ntt_gen_def using sum_rules apply presburger done apply (rule sum_rules(2)) subgoal for j using fntt2_by_index[of i] that sum_in[of _ "(\^((n div llen) * i))" "l2"] comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding ntt_gen_def apply auto done done have 005: "\ = (\j=0..^((n div llen)*i*(2*j+1))))" apply (rule sum_rules(2)) subgoal for j unfolding numbers2_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l2 " hence "map ((!) numbers) (filter odd [0.. ^ (n div 2 ^ l * i * j + n div llen * i) = numbers ! (2 * j + 1) * \ ^ (n div llen * i * (2 * j + 1))" unfolding llen_def by (smt (z3) Groups.mult_ac(2) distrib_left mult.right_neutral mult_2 mult_cancel_left) qed done then show ?thesis using 000 001 002 003 004 005 unfolding sum1_def llen_def l1_def l2_def using sum_splice_other_way_round[of "\ d. numbers ! d * \ ^ (n div length numbers * i * d)" "2^l"] Suc(2) unfolding ntt_gen_def by (smt (z3) Groups.mult_ac(2) numbers1_even numbers2_even power_Suc2) qed then show ?thesis by (metis "00" "01" nth_equalityI) qed text \We show equality for the indices in the second halves.\ have after_half: "map (ntt_gen numbers llen) [(llen div 2)..Equality for every index.\ have 02:"(map (ntt_gen numbers llen) [(llen div 2)..x y. x * \ ^ (n div llen * y)) fntt2 [0.. ^ (n div llen * i)" using Suc(2) that by (simp add: fntt2_length llen_def) have 003: "- fntt2 ! i * \ ^ (n div llen * i) = fntt2 ! i * \ ^ (n div llen * (i+ llen div 2))" using Suc(2) omega_div_exp_min1[of l] unfolding llen_def by (smt (z3) Suc.prems(2) mult.commute mult.left_commute mult_1s_ring_1(2) neq0_conv nonzero_mult_div_cancel_left numeral_One pos2 power_Suc power_add power_mult) hence 004:"sum2 ! i = (fntt1!i) - (fntt2!i) * (\^((n div llen) * i))" unfolding sum2_def llen_def by (simp add: Suc.prems(1) fntt1_length fntt2_length that) have 005:"(fntt1!i) = (\j=0..^((n div (2^l))*i*j))" using fntt1_by_index that unfolding ntt_gen_def l1_def by simp have 006:"\ =(\j=0..^((n div llen)*i*(2*j)))" apply (rule sum_rules(2)) subgoal for j unfolding numbers1_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l1 " hence "map ((!) numbers) (filter even [0.. ^ (n div 2 ^ l * i * j) = numbers ! (2 * j) * \ ^ (n div llen * i * (2 * j))" by (metis (mono_tags, lifting) mult.assoc mult.left_commute) qed done have 007:"\ = (\j=0..^((n div llen)*(2^l + i)*(2*j))) " apply (rule sum_rules(2)) subgoal for j using Suc(2) Suc(3) omega_div_exp_min1[of l] llen_def l1_def numbers1_def apply(smt (verit, del_insts) add.commute minus_power_mult_self mult_2 mult_minus1_right power_add power_mult) done done moreover have 008: "(fntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div (2^l))*i*j+ (n div llen) * i))" apply(rule trans[where s = "(\j = 0.. ^ (n div 2 ^ l * i * j) * \ ^ (n div llen * i))"]) subgoal using fntt2_by_index[of i] that sum_in comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding ntt_gen_def using sum_rules l2_def apply presburger done apply (rule sum_rules(2)) subgoal for j using fntt2_by_index[of i] that sum_in comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding ntt_gen_def apply auto done done have 009: "\ = (\j=0..^((n div llen)*i*(2*j+1))))" apply (rule sum_rules(2)) subgoal for j unfolding numbers2_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l2 " hence "map ((!) numbers) (filter odd [0.. ^ (n div 2 ^ l * i * j + n div llen * i) = numbers ! (2 * j + 1) * \ ^ (n div llen * i * (2 * j + 1))" by (smt (z3) Groups.mult_ac(2) distrib_left mult.right_neutral mult_2 mult_cancel_left) qed done have 010: " (fntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div llen)*i*(2*j+1)))) " using 008 009 by presburger have 011: " - (fntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div llen)*i*(2*j+1)))) " apply(rule neg_cong) apply(rule trans[of _ "fntt2 ! i * \ ^ (n div llen * i)"]) subgoal by simp apply(rule trans[where s="(\j=0..^((n div llen)*i*(2*j+1))))"]) subgoal using 008 009 by simp apply(rule sym) using sum_neg_in[of _ "l2"] apply simp done have 012: "\ = (\j=0..^((n div llen)*(2^l+i)*(2*j+1))))" apply(rule sum_rules(2)) subgoal for j using Suc(2) Suc(3) omega_div_exp_min1[of l] llen_def l2_def apply (smt (z3) add.commute exp_rule mult.assoc mult_minus1_right plus_1_eq_Suc power_add power_minus1_odd power_mult) done done have 013:"fntt1 ! i = (\j = 0..<2 ^ l. numbers!(2*j) * \ ^ (n div llen * (2^l + i) * (2*j)))" using 005 006 007 numbers1_even llen_def l1_def by auto have 014: "(\j = 0..<2 ^ l. numbers ! (2*j + 1) * \ ^ (n div llen* (2^l + i) * (2*j + 1))) = - fntt2 ! i * \ ^ (n div llen * i)" using trans[OF l2_def numbers2_even] sym[OF 012] sym[OF 011] by simp have "ntt_gen numbers llen (2 ^ l + i) = (fntt1!i) - (fntt2!i) * (\^((n div llen) * i))" unfolding ntt_gen_def apply(subst Suc(2)) using sum_splice[of "\ d. numbers ! d * \ ^ (n div llen * (2^l+i) * d)" "2^l"] sym[OF 013] 014 Suc(2) by simp thus ?thesis using 000 sym[OF 001] "004" sum2_def by simp qed then show ?thesis by (metis "00" "01" list_eq_iff_nth_eq) qed obtain x y xs where xyxs: "numbers = x#y#xs" using Suc(2) by (metis FNTT.cases add.left_neutral even_Suc even_add length_Cons list.size(3) mult_2 power_Suc power_eq_0_iff zero_neq_numeral) show ?case apply(subst xyxs) apply(subst FNTT.simps(3)) apply(subst xyxs[symmetric])+ unfolding Let_def using map_append[of "ntt_gen numbers llen" " [0..\noindent \textbf{Major Correctness Theorem for Butterfly Algorithm}.\\ We have already shown: \begin{itemize} \item Generalized $NTT$ with degree annotation $2^N$ equals usual $NTT$. \item Generalized $NTT$ tracks all levels of recursion in $FNTT$. \end{itemize} Thus, $FNTT$ equals $NTT$. \ theorem FNTT_correct: assumes "length numbers = n" shows "FNTT numbers = NTT numbers" using FNTT_NTT_gen_eq NTT_gen_NTT_full_length assms n_two_pot by force subsection \Inverse Transform in Butterfly Scheme\ text \We also formalized the inverse transform by using the butterfly scheme. Proofs are obtained by adaption of arguments for $FNTT$.\ lemmas [simp] = FNTT_termination_aux fun IFNTT where "IFNTT [] = []"| "IFNTT [a] = [a]"| "IFNTT nums = (let nn = length nums; nums1 = [nums!i . i <- (filter even [0.. x k. x*(\^( (n div nn) * k))) ifntt2 [0..<(nn div 2)]); sum2 = map2 (-) ifntt1 (map2 ( \ x k. x*(\^( (n div nn) * k))) ifntt2 [0..<(nn div 2)]) in sum1@sum2)" lemmas [simp del] = FNTT_termination_aux definition "intt_gen numbers degr i = (\j=0..<(length numbers). (numbers ! j) * \ ^((n div degr)*i*j)) " definition "INTT_gen degr numbers = map (intt_gen numbers (degr)) [0..< length numbers]" lemma INTT_gen_INTT_full_length: assumes "length numbers =n" shows "INTT_gen n numbers = INTT numbers" unfolding INTT_gen_def intt_gen_def INTT_def intt_def using assms by simp lemma my_div_exp_min1: assumes "2^(Suc l) \ n" shows "(\ ^(n div 2^(Suc l)))^(2^l) = -1" by (metis assms divide_minus1 mult_zero_right mu_properties(1) nonzero_mult_div_cancel_right omega_div_exp_min1 power_one_over zero_neq_one) lemma my_n_2_min1: "\^(n div 2) = -1" by (metis divide_minus1 mult_zero_right mu_properties(1) nonzero_mult_div_cancel_right omg_n_2_min1 power_one_over zero_neq_one) text \Correctness proof by common induction technique. Same strategies as for $FNTT$.\ theorem IFNTT_INTT_gen_eq: "length numbers = 2^l \ 2^l \ n \ IFNTT numbers = INTT_gen (length numbers) numbers" proof(induction l arbitrary: numbers) case 0 hence "local.IFNTT numbers = [numbers ! 0]" by (metis IFNTT.simps(2) One_nat_def Suc_length_conv length_0_conv nth_Cons_0 power_0) then show ?case unfolding INTT_gen_def intt_gen_def using 0 by simp next case (Suc l) text \We define some lists that are used during the recursive call.\ define numbers1 where "numbers1 = [numbers!i . i <- (filter even [0.. x k. x*(\^( (n div (length numbers)) * k))) ifntt2 [0..<((length numbers) div 2)])" define sum2 where "sum2 = map2 (-) ifntt1 (map2 ( \ x k. x*(\^( (n div (length numbers)) * k))) ifntt2 [0..<((length numbers) div 2)])" define l1 where "l1 = length numbers1" define l2 where "l2 = length numbers2" define llen where "llen = length numbers" text \Properties of those lists\ have numbers1_even: "length numbers1 = 2^l" using numbers1_def length_even_filter Suc by simp have numbers2_even: "length numbers2 = 2^l" using numbers2_def length_odd_filter Suc by simp have numbers1_ifntt: "ifntt1 = INTT_gen (2^l) numbers1" using ifntt1_def Suc.IH[of numbers1] numbers1_even Suc(3) by simp hence ifntt1_by_index: "ifntt1 ! i = intt_gen numbers1 (2^l) i" if "i < 2^l" for i unfolding INTT_gen_def by (simp add: numbers1_even that) have numbers2_ifntt: "ifntt2 = INTT_gen (2^l) numbers2" using ifntt2_def Suc.IH[of numbers2] numbers2_even Suc(3) by simp hence ifntt2_by_index: "ifntt2 ! i = intt_gen numbers2 (2^l) i" if "i < 2^l" for i unfolding INTT_gen_def by (simp add: numbers2_even that) have ifntt1_length: "length ifntt1 = 2^l" unfolding numbers1_ifntt INTT_gen_def numbers1_def using numbers1_def numbers1_even by force have ifntt2_length: "length ifntt2 = 2^l" unfolding numbers2_ifntt INTT_gen_def numbers2_def using numbers2_def numbers2_even by force text \Same proof structure as for the \textit{FNTT} proof. $\omega$s are just replaced by $\mu$s.\ have before_half: "map (intt_gen numbers llen) [0..<(llen div 2)] = sum1" proof- text \Length is important, since we want to use list lemmas later on.\ have 00:"length (map (intt_gen numbers llen) [0..<(llen div 2)]) = length sum1" unfolding sum1_def llen_def using Suc(2) map2_length[of _ ifntt2 "[0..x y. x * \ ^ (n div length numbers * y)) ifntt2 [0..We show equality by extensionality on indices.\ have 02:"(map (intt_gen numbers llen) [0..<(llen div 2)]) ! i = sum1 ! i" if "i < 2^l" for i proof- text \First simplify this term.\ have 000:"(map (intt_gen numbers llen) [0..<(llen div 2)]) ! i = intt_gen numbers llen i" using "00" "01" that by auto text \Expand the definition of $sum1$ and massage the result.\ moreover have 001:"sum1 ! i = (ifntt1!i) + (ifntt2!i) * (\^((n div llen) * i))" unfolding sum1_def using map2_index "00" "01" INTT_gen_def add.left_neutral diff_zero ifntt1_length length_map length_upt map2_map_map map_nth nth_upt numbers2_even numbers2_ifntt that llen_def by force moreover have 002:"(ifntt1!i) = (\j=0..^((n div (2^l))*i*j))" unfolding l1_def using ifntt1_by_index[of i] that unfolding intt_gen_def by simp have 003:"... = (\j=0..^((n div llen)*i*(2*j)))" apply (rule sum_rules(2)) subgoal for j unfolding numbers1_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l1 " hence "map ((!) numbers) (filter even [0.. ^ (n div 2 ^ l * i * j) = numbers ! (2 * j) * \ ^ (n div llen * i * (2 * j))" unfolding llen_def l1_def l2_def by (metis (mono_tags, lifting) mult.assoc mult.left_commute) qed done moreover have 004: "(ifntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div (2^l))*i*j+ (n div llen) * i))" apply(rule trans[where s = "(\j = 0.. ^ (n div 2 ^ l * i * j) * \ ^ (n div llen * i))"]) subgoal unfolding l2_def llen_def using ifntt2_by_index[of i] that sum_in[of _ "(\^((n div llen) * i))" "l2"] comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding intt_gen_def using sum_rules apply presburger done apply (rule sum_rules(2)) subgoal for j using ifntt2_by_index[of i] that sum_in[of _ "(\^((n div llen) * i))" "l2"] comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding intt_gen_def apply auto done done have 005: "\ = (\j=0..^((n div llen)*i*(2*j+1))))" apply (rule sum_rules(2)) subgoal for j unfolding numbers2_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l2 " hence "map ((!) numbers) (filter odd [0.. ^ (n div 2 ^ l * i * j + n div llen * i) = numbers ! (2 * j + 1) * \ ^ (n div llen * i * (2 * j + 1))" unfolding llen_def by (smt (z3) Groups.mult_ac(2) distrib_left mult.right_neutral mult_2 mult_cancel_left) qed done then show ?thesis using 000 001 002 003 004 005 unfolding sum1_def llen_def l1_def l2_def using sum_splice_other_way_round[of "\ d. numbers ! d * \ ^ (n div length numbers * i * d)" "2^l"] Suc(2) unfolding intt_gen_def by (smt (z3) Groups.mult_ac(2) numbers1_even numbers2_even power_Suc2) qed then show ?thesis by (metis "00" "01" nth_equalityI) qed text \We show index-wise equality for the second halves\ have after_half: "map (intt_gen numbers llen) [(llen div 2)..Equality for every index\ have 02:"(map (intt_gen numbers llen) [(llen div 2)..x y. x * \ ^ (n div llen * y)) ifntt2 [0.. ^ (n div llen * i)" using Suc(2) that by (simp add: ifntt2_length llen_def) have 003: "- ifntt2 ! i * \ ^ (n div llen * i) = ifntt2 ! i * \ ^ (n div llen * (i+ llen div 2))" using Suc(2) my_div_exp_min1[of l] unfolding llen_def by (smt (z3) Suc.prems(2) mult.commute mult.left_commute mult_1s_ring_1(2) neq0_conv nonzero_mult_div_cancel_left numeral_One pos2 power_Suc power_add power_mult) hence 004:"sum2 ! i = (ifntt1!i) - (ifntt2!i) * (\^((n div llen) * i))" unfolding sum2_def llen_def by (simp add: Suc.prems(1) ifntt1_length ifntt2_length that) have 005:"(ifntt1!i) = (\j=0..^((n div (2^l))*i*j))" using ifntt1_by_index that unfolding intt_gen_def l1_def by simp have 006:"\ =(\j=0..^((n div llen)*i*(2*j)))" apply (rule sum_rules(2)) subgoal for j unfolding numbers1_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l1 " hence "map ((!) numbers) (filter even [0.. ^ (n div 2 ^ l * i * j) = numbers ! (2 * j) * \ ^ (n div llen * i * (2 * j))" by (metis (mono_tags, lifting) mult.assoc mult.left_commute) qed done have 007:"\ = (\j=0.. ^((n div llen)*(2^l + i)*(2*j))) " apply (rule sum_rules(2)) subgoal for j using Suc(2) Suc(3) my_div_exp_min1[of l] llen_def l1_def numbers1_def apply(smt (verit, del_insts) add.commute minus_power_mult_self mult_2 mult_minus1_right power_add power_mult) done done moreover have 008: "(ifntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div (2^l))*i*j+ (n div llen) * i))" apply(rule trans[where s = "(\j = 0.. ^ (n div 2 ^ l * i * j) * \ ^ (n div llen * i))"]) subgoal using ifntt2_by_index[of i] that sum_in comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding intt_gen_def using sum_rules l2_def apply presburger done apply (rule sum_rules(2)) subgoal for j using ifntt2_by_index[of i] that sum_in comm_semiring_1_class.semiring_normalization_rules(26)[of \] unfolding intt_gen_def apply auto done done have 009: "\ = (\j=0..^((n div llen)*i*(2*j+1))))" apply (rule sum_rules(2)) subgoal for j unfolding numbers2_def apply(subst llen_def[symmetric]) proof- assume ass: "j < l2 " hence "map ((!) numbers) (filter odd [0.. ^ (n div 2 ^ l * i * j + n div llen * i) = numbers ! (2 * j + 1) * \ ^ (n div llen * i * (2 * j + 1))" by (smt (z3) Groups.mult_ac(2) distrib_left mult.right_neutral mult_2 mult_cancel_left) qed done have 010: " (ifntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div llen)*i*(2*j+1)))) " using 008 009 by presburger have 011: " - (ifntt2!i) * (\^((n div llen) * i)) = (\j=0..^((n div llen)*i*(2*j+1)))) " apply(rule neg_cong) apply(rule trans[where s="(\j=0..^((n div llen)*i*(2*j+1))))"]) subgoal using 008 009 by simp apply(rule sym) using sum_neg_in[of _ "l2"] apply simp done have 012: "\ = (\j=0..^((n div llen)*(2^l+i)*(2*j+1))))" apply(rule sum_rules(2)) subgoal for j using Suc(2) Suc(3) my_div_exp_min1[of l] llen_def l2_def apply (smt (z3) add.commute exp_rule mult.assoc mult_minus1_right plus_1_eq_Suc power_add power_minus1_odd power_mult) done done have 013:"ifntt1 ! i = (\j = 0..<2 ^ l. numbers!(2*j) * \ ^ (n div llen * (2^l + i) * (2*j)))" using 005 006 007 numbers1_even llen_def l1_def by auto have 014: "(\j = 0..<2 ^ l. numbers ! (2*j + 1) * \ ^ (n div llen* (2^l + i) * (2*j + 1))) = - ifntt2 ! i * \ ^ (n div llen * i)" using trans[OF l2_def numbers2_even] sym[OF 012] sym[OF 011] by simp have "intt_gen numbers llen (2 ^ l + i) = (ifntt1!i) - (ifntt2!i) * (\^((n div llen) * i))" unfolding intt_gen_def apply(subst Suc(2)) using sum_splice[of "\ d. numbers ! d * \ ^ (n div llen * (2^l+i) * d)" "2^l"] sym[OF 013] 014 Suc(2) by simp thus ?thesis using 000 sym[OF 001] "004" sum2_def by simp qed then show ?thesis by (metis "00" "01" list_eq_iff_nth_eq) qed obtain x y xs where xyxs: "numbers = x#y#xs" using Suc(2) by (metis FNTT.cases add.left_neutral even_Suc even_add length_Cons list.size(3) mult_2 power_Suc power_eq_0_iff zero_neq_numeral) show ?case apply(subst xyxs) apply(subst IFNTT.simps(3)) apply(subst xyxs[symmetric])+ unfolding Let_def using map_append[of "intt_gen numbers llen" " [0..Correctness of the butterfly scheme for the inverse \textit{INTT}.\ theorem IFNTT_correct: assumes "length numbers = n" shows "IFNTT numbers = INTT numbers" using IFNTT_INTT_gen_eq INTT_gen_INTT_full_length assms n_two_pot by force text \Also $FNTT$ and $IFNTT$ are mutually inverse\ theorem IFNTT_inv_FNTT: assumes "length numbers = n" shows "IFNTT (FNTT numbers) = map ((*) (of_int_mod_ring (int n))) numbers" by (simp add: FNTT_correct IFNTT_correct assms length_NTT ntt_correct) text \The other way round:\ theorem FNTT_inv_IFNTT: assumes "length numbers = n" shows "FNTT (IFNTT numbers) = map ((*) (of_int_mod_ring (int n))) numbers" by (simp add: FNTT_correct IFNTT_correct assms inv_ntt_correct length_INTT) subsection \An Optimization\ text \Currently, we extract elements on even and odd positions respectively by a list comprehension over even and odd indices. Due to the definition in Isabelle, an index access has linear time complexity. This results in quadratic running time complexity for every level in the recursion tree of the \textit{FNTT}. In order to reach the $\mathcal{O}(n \log n)$ time bound, we have find a better way of splitting the elements at even or odd indices respectively. \ text \A core of this optimization is the $evens\text{-}odds$ function, which splits the vectors in linear time.\ fun evens_odds::"bool \'b list \ 'b list" where "evens_odds _ [] = []"| "evens_odds True (x#xs)= (x# evens_odds False xs)"| "evens_odds False (x#xs) = evens_odds True xs" lemma map_filter_shift: " map f (filter even [0.. x. f (x+1)) (filter odd [0.. x. f (x+1)) (filter even [0..A splitting by the $evens\text{-}odds$ function is equivalent to the more textbook-like list comprehension.\ lemma filter_compehension_evens_odds: "[xs ! i. i <- filter even [0.. [xs ! i. i <- filter odd [0..For automated termination proof.\ lemma [simp]: "length (evens_odds True vc) < Suc (length vc)" "length (evens_odds False vc) < Suc (length vc)" by (metis filter_compehension_evens_odds le_imp_less_Suc length_filter_le length_map map_nth)+ text \The $FNTT$ definition from above was suitable for matters of proof conduction. However, the naive decomposition into elements at odd and even indices induces a complexity of $n^2$ in every recursive step. As mentioned, the $evens\text{-}odds$ function filters for elements on even or odd positions respectively. The list has to be traversed only once which gives \textit{linear} complexity for every recursive step. \ fun FNTT' where "FNTT' [] = []"| "FNTT' [a] = [a]"| "FNTT' nums = (let nn = length nums; nums1 = evens_odds True nums; nums2 = evens_odds False nums; fntt1 = FNTT' nums1; fntt2 = FNTT' nums2; fntt2_omg = (map2 ( \ x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]); sum1 = map2 (+) fntt1 fntt2_omg; sum2 = map2 (-) fntt1 fntt2_omg in sum1@sum2)" text \The optimized \textit{FNTT} is equivalent to the naive \textit{NTT}.\ lemma FNTT'_FNTT: "FNTT' xs = FNTT xs" apply(induction xs rule: FNTT'.induct) subgoal by simp subgoal by simp apply(subst FNTT'.simps(3)) apply(subst FNTT.simps(3)) subgoal for a b xs unfolding Let_def apply (metis filter_compehension_evens_odds) done done text \It is quite surprising that some inaccuracies in the interpretation of informal textbook definitions - even when just considering such a simple algorithm - can indeed affect time complexity.\ subsection \Arguments on Running Time\ text \ $FFT$ is especially known for its $\mathcal{O}(n \log n)$ running time. Unfortunately, Isabelle does not provide a built-in time formalization. Nonetheless we can reason about running time after defining some "reasonable" consumption functions by hand. Our approach loosely follows a general pattern by Nipkow et al.~\parencite{funalgs}. First, we give running times and lemmas for the auxiliary functions used during FNTT.\\ General ideas behind the $\mathcal{O}(n \log n)$ are: \begin{itemize} \item By recursively halving the problem size, we obtain a tree of depth $\mathcal{O}(\log n)$. \item For every level of that tree, we have to process all elements which gives $\mathcal{O}(n)$ time. \end{itemize} \ text \Time for splitting the list according to even and odd indices.\ fun T_\<^sub>e\<^sub>o::"bool \ 'c list \ nat" where " T_\<^sub>e\<^sub>o _ [] = 1"| " T_\<^sub>e\<^sub>o True (x#xs)= (1+ T_\<^sub>e\<^sub>o False xs)"| " T_\<^sub>e\<^sub>o False (x#xs) = (1+ T_\<^sub>e\<^sub>o True xs)" lemma T_eo_linear: "T_\<^sub>e\<^sub>o b xs = length xs + 1" by (induction b xs rule: T_\<^sub>e\<^sub>o.induct) auto text \Time for length.\ fun T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h where "T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h [] = 1 "| "T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h (x#xs) = 1+ T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h xs" lemma T_length_linear: "T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h xs = length xs +1" by (induction xs) auto text \Time for index access.\ fun T\<^sub>n\<^sub>t\<^sub>h where "T\<^sub>n\<^sub>t\<^sub>h [] i = 1 "| "T\<^sub>n\<^sub>t\<^sub>h (x#xs) 0 = 1"| "T\<^sub>n\<^sub>t\<^sub>h (x#xs) (Suc i) = 1 + T\<^sub>n\<^sub>t\<^sub>h xs i" lemma T_nth_linear: "T\<^sub>n\<^sub>t\<^sub>h xs i \ length xs +1" by (induction xs i rule: T\<^sub>n\<^sub>t\<^sub>h.induct) auto text \Time for mapping two lists into one result.\ fun T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 where "T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t [] _ = 1"| "T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t _ [] = 1"| "T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t (x#xs) (y#ys) = (t x y + 1 + T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t xs ys)" lemma T_map_2_linear: "c > 0 \ (\ x y. t x y \ c) \ T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t xs ys \ min (length xs) (length ys) * (c+1) + 1" apply(induction t xs ys rule: T\<^sub>m\<^sub>a\<^sub>p\<^sub>2.induct) subgoal by simp subgoal by simp subgoal for t x xs y ys apply(subst T\<^sub>m\<^sub>a\<^sub>p\<^sub>2.simps, subst length_Cons, subst length_Cons) using min_add_distrib_right[of 1] by (smt (z3) Suc_eq_plus1 add.assoc add.commute add_le_mono le_numeral_extra(4) min_def mult.commute mult_Suc_right) done lemma T_map_2_linear': "c > 0 \ (\ x y. t x y = c) \ T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 t xs ys = min (length xs) (length ys) * (c+1) + 1" by(induction t xs ys rule: T\<^sub>m\<^sub>a\<^sub>p\<^sub>2.induct) simp+ text \Time for append.\ fun T\<^sub>a\<^sub>p\<^sub>p where " T\<^sub>a\<^sub>p\<^sub>p [] _ = 1"| " T\<^sub>a\<^sub>p\<^sub>p (x#xs) ys = 1 + T\<^sub>a\<^sub>p\<^sub>p xs ys" lemma T_app_linear: " T\<^sub>a\<^sub>p\<^sub>p xs ys = length xs +1" by(induction xs) auto text \Running Time of (optimized) $FNTT$.\ fun T\<^sub>F\<^sub>N\<^sub>T\<^sub>T::"('a mod_ring) list \ nat" where "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T [] = 1"| "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T [a] = 1"| "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums = (1 +T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h nums+ 3+ (let nn = length nums; nums1 = evens_odds True nums; nums2 = evens_odds False nums in T_\<^sub>e\<^sub>o True nums + T_\<^sub>e\<^sub>o False nums + 2 + (let fntt1 = FNTT nums1; fntt2 = FNTT nums2 in (T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums1) + (T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums2) + (let sum1 = map2 (+) fntt1 (map2 ( \ x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]); sum2 = map2 (-) fntt1 (map2 ( \ x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]) in 2* T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 (\ x y. 1) fntt2 [0..<(nn div 2)] + 2* T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 (\ x y. 1) fntt1 (map2 ( \ x k. x*(\^( (n div nn) * k))) fntt2 [0..<(nn div 2)]) + T\<^sub>a\<^sub>p\<^sub>p sum1 sum2))))" lemma mono: "((f x)::nat) \ f y \ f y \ fz \ f x \ fz" by simp lemma evens_odds_length: "length (evens_odds True xs) = (length xs+1) div 2 \ length (evens_odds False xs) = (length xs) div 2" by(induction xs) simp+ text \Length preservation during $FNTT$.\ lemma FNTT_length: "length numbers = 2^l \ length (FNTT numbers) = length numbers" proof(induction l arbitrary: numbers) case (Suc l) define numbers1 where "numbers1 = [numbers!i . i <- (filter even [0.. x k. x*(\^( (n div (length numbers)) * k))) fntt2 [0..<((length numbers) div 2)])" define sum1 where "sum1 = map2 (+) fntt1 presum" define sum2 where "sum2 = map2 (-) fntt1 presum" have "length numbers1 = 2^l" by (metis Suc.prems numbers1_def diff_add_inverse2 length_even_filter mult_2 nonzero_mult_div_cancel_left power_Suc zero_neq_numeral) hence "length fntt1 = 2^l" by (simp add: Suc.IH fntt1_def) hence "length presum = 2^l" unfolding presum_def using map2_length Suc.IH Suc.prems fntt2_def length_odd_filter numbers2_def by force hence "length sum1 = 2^l" by (simp add: \length fntt1 = 2 ^ l\ sum1_def) have "length numbers2 = 2^l" by (metis Suc.prems numbers2_def length_odd_filter nonzero_mult_div_cancel_left power_Suc zero_neq_numeral) hence "length fntt2 = 2^l" by (simp add: Suc.IH fntt2_def) hence "length sum2 = 2^l" unfolding sum2_def using \length sum1 = 2 ^ l\ sum1_def by force hence final:"length (sum1@sum2) = 2^(Suc l)" by (simp add: \length sum1 = 2 ^ l\) obtain x y xs where xyxs_Def: "numbers = x#y#xs" by (metis \length numbers2 = 2 ^ l\ evens_odds.elims filter_compehension_evens_odds length_0_conv neq_Nil_conv numbers2_def power_eq_0_iff zero_neq_numeral) show ?case apply(subst xyxs_Def, subst FNTT.simps(3), subst xyxs_Def[symmetric]) unfolding Let_def using final unfolding sum1_def sum2_def presum_def fntt1_def fntt2_def numbers1_def numbers2_def using Suc by (metis xyxs_Def) qed (metis FNTT.simps(2) Suc_length_conv length_0_conv nat_power_eq_Suc_0_iff) lemma add_cong: "(a1::nat) + a2+a3 +a4= b \ a1 +a2+ c + a3+a4= c +b" by simp lemma add_mono:"a \ (b::nat) \ c \ d \ a + c \ b +d" by simp lemma xyz: " Suc (Suc (length xs)) = 2 ^ l \ length (x # evens_odds True xs) = 2 ^ (l - 1)" by (metis (no_types, lifting) Nat.add_0_right Suc_eq_plus1 div2_Suc_Suc div_mult_self2 evens_odds_length length_Cons nat.distinct(1) numeral_2_eq_2 one_div_two_eq_zero plus_1_eq_Suc power_eq_if) lemma zyx:" Suc (Suc (length xs)) = 2 ^ l \ length (y # evens_odds False xs) = 2 ^ (l - 1)" by (smt (z3) One_nat_def Suc_pred diff_Suc_1 div2_Suc_Suc evens_odds_length le_numeral_extra(4) length_Cons nat_less_le neq0_conv power_0 power_diff power_one_right zero_less_diff zero_neq_numeral) text \When $length \; xs = 2^l$, then $length \; (evens\text{-}odds \; xs) = 2^{l-1}$.\ lemma evens_odds_power_2: fixes x::'b and y::'b assumes "Suc (Suc (length (xs::'b list))) = 2 ^ l" shows " Suc(length (evens_odds b xs)) = 2 ^ (l-1)" proof- have "Suc(length (evens_odds b xs)) = length (evens_odds b (x#y#xs))" by (metis (full_types) evens_odds.simps(2) evens_odds.simps(3) length_Cons) have "length (x#y#xs) = 2^l" using assms by simp have "length (evens_odds b (x#y#xs)) = 2^(l-1)" apply (cases b) apply (smt (z3) Suc_eq_plus1 Suc_pred \length (x # y # xs) = 2 ^ l\ add.commute add_diff_cancel_left' assms filter_compehension_evens_odds gr0I le_add1 le_imp_less_Suc length_even_filter mult_2 nat_less_le power_diff power_eq_if power_one_right zero_neq_numeral) by (smt (z3) One_nat_def Suc_inject \length (x # y # xs) = 2 ^ l\ assms evens_odds_length le_zero_eq nat.distinct(1) neq0_conv not_less_eq_eq pos2 power_Suc0_right power_diff_power_eq power_eq_if) then show ?thesis by (metis \Suc (length (evens_odds b xs)) = length (evens_odds b (x # y # xs))\) qed text \ \noindent \textbf{Major Lemma:} We rewrite the Running time of $FNTT$ in this proof and collect constraints for the time bound. Using this, bounds are chosen in a way such that the induction goes through properly. \paragraph \noindent We define: \begin{equation*} T(2^0) = 1 \end{equation*} \begin{equation*} T(2^l) = (2^l - 1)\cdot 14 apply+ 15 \cdot l \cdot 2^{l-1} + 2^l \end{equation*} \paragraph \noindent We want to show: \begin{equation*} T_{FNTT}(2^l) = T(2^l) \end{equation*} (Note that by abuse of types, the $2^l$ denotes a list of length $2^l$.) \paragraph \noindent First, let's informally check that $T$ is indeed an accurate description of the running time: \begin{align*} T_{FNTT}(2^l) & \; = 14 + 15 \cdot 2 ^ {l-1} + 2 \cdot T_{FNTT}(2^{l-1}) \hspace{1cm} \text{by analyzing the running time function}\\ &\overset{I.H.}{=} 14 + 15 \cdot 2 ^ {l-1} + 2 \cdot ((2^{l-1} - 1) \cdot 14 + (l - 1) \cdot 15 \cdot 2^{l-2} + 2^{l-1})\\ & \;= 14 \cdot 2^l - 14 + 15 \cdot 2 ^ {l-1} + 15\cdot l \cdot 2^{l-1} - 15 \cdot 2^{l-1} + 2^l\\ &\; = (2^l - 1)\cdot 14 + 15 \cdot l \cdot 2^{l-1} + 2^l\\ &\overset{def.}{=} T(2^l) \end{align*} The base case is trivially true. \ theorem tight_bound: assumes T_def: "\ numbers l. length numbers = 2^l \ l > 0 \ T numbers = (2^l - 1) * 14 + l *15*2^(l-1) + 2^l" "\ numbers l. l =0 \ length numbers = 2^l \ T numbers = 1" shows " length numbers = 2^l \ T\<^sub>F\<^sub>N\<^sub>T\<^sub>T numbers = T numbers" proof(induction numbers arbitrary: l rule: T\<^sub>F\<^sub>N\<^sub>T\<^sub>T.induct) case (3 x y numbers) text \Some definitions for making term rewriting simpler.\ define nn where "nn = length (x # y # numbers)" define nums1 where "nums1 = evens_odds True (x # y # numbers)" define nums2 where "nums2 = evens_odds False (x # y # numbers)" define fntt1 where "fntt1 = local.FNTT nums1" define fntt2 where "fntt2 = local.FNTT nums2" define sum1 where "sum1 = map2 (+) fntt1 (map2 (\x y. x * \ ^ (n div nn * y)) fntt2 [0..x y. x * \ ^ (n div nn * y)) fntt2 [0..Unfolding the running time function and combining it with the definitions above.\ have TFNNT_simp: " T\<^sub>F\<^sub>N\<^sub>T\<^sub>T (x # y # numbers) = 1 + T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h (x # y # numbers) + 3 + T_\<^sub>e\<^sub>o True (x # y # numbers) + T_\<^sub>e\<^sub>o False (x # y # numbers) + 2 + local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums1 + local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums2 + 2 * T\<^sub>m\<^sub>a\<^sub>p\<^sub>2 (\x y. 1) fntt2 [0..m\<^sub>a\<^sub>p\<^sub>2 (\x y. 1) fntt1 (map2 (\x y. x * \ ^ (n div nn * y)) fntt2 [0..a\<^sub>p\<^sub>p sum1 sum2" apply(subst T\<^sub>F\<^sub>N\<^sub>T\<^sub>T.simps(3)) unfolding Let_def unfolding sum2_def sum1_def fntt1_def fntt2_def nums1_def nums2_def nn_def apply simp done text \Application of lemmas related to running times of auxiliary functions.\ have length_nums1: "length nums1 = (2::nat)^(l-1)" unfolding nums1_def using evens_odds_length[of "x # y # numbers"] 3(3) xyz by fastforce have length_nums2: "length nums2 = (2::nat)^(l-1)" unfolding nums2_def using evens_odds_length[of "x # y # numbers"] 3(3) by (metis One_nat_def le_0_eq length_Cons lessI list.size(4) neq0_conv not_add_less2 not_less_eq_eq pos2 power_Suc0_right power_diff_power_eq power_eq_if) have length_simp: "T\<^sub>l\<^sub>e\<^sub>n\<^sub>g\<^sub>t\<^sub>h (x # y # numbers) = (2::nat) ^l +1" using T_length_linear[of "x#y#numbers"] 3(3) by simp have even_odd_simp: " T_\<^sub>e\<^sub>o b (x # y # numbers) = (2::nat)^l + 1" for b by (metis "3.prems" T_eo_linear)+ have 02: "(length fntt2) = (length [0..m\<^sub>a\<^sub>p\<^sub>2 (\x y. 1) fntt2 [0..m\<^sub>a\<^sub>p\<^sub>2 (\x y. 1) fntt1 (map2 (\x y. x * \ ^ (n div nn * y)) fntt2 [0..a\<^sub>p\<^sub>p sum1 sum2 = (2::nat)^(l-1) + 1" by(subst T_app_linear, subst sum1_simp, simp) let ?T1 = "(2^(l-1) - 1) * 14 + (l-1) *15*2^(l-1 -1) + 2^(l-1)" text \Induction hypotheses\ have IH_pluged1: "local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums1 = ?T1" apply(subst "3.IH"(1)[of nn nums1 nums2 fntt1 fntt2 "l-1", OF nn_def nums1_def nums2_def fntt1_def fntt2_def length_nums1]) apply(cases "l \ 1") subgoal apply(subst T_def(2)[of "l-1"]) subgoal by simp apply(rule length_nums1) apply simp done apply(subst T_def(1)[OF length_nums1]) subgoal by simp subgoal by simp done have IH_pluged2: "local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums2 = ?T1" apply(subst "3.IH"(2)[of nn nums1 _ fntt1 fntt2 "l-1", OF nn_def nums1_def nums2_def fntt1_def fntt2_def length_nums2 ]) apply(cases "l \ 1") subgoal apply(subst T_def(2)[of "l-1"]) subgoal by simp apply(rule length_nums2) apply simp done apply(subst T_def(1)[OF length_nums2]) subgoal by simp subgoal by simp done have " T\<^sub>F\<^sub>N\<^sub>T\<^sub>T (x # y # numbers) = 14 + (3 * 2 ^ l + (local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums1 + (local.T\<^sub>F\<^sub>N\<^sub>T\<^sub>T nums2 + (5 * 2^(l-1) + 4 * (2 ^ l div 2)))))" apply(subst TFNNT_simp, subst map21_simp, subst map22_simp, subst length_simp, subst app_simp, subst even_odd_simp, subst even_odd_simp) apply(auto simp add: algebra_simps power_eq_if[of 2 l]) done text \Proof that the term $T\text{-}def$ indeed fulfills the recursive properties, i.e. $t(2^l) = 2 \cdot t(2^{l-1}) + s$\ also have "\ = 14 + (3 * 2 ^ l + (?T1 + (?T1 + (5 * 2^(l-1) + 4 * (2 ^ l div 2)))))" apply(subst IH_pluged1, subst IH_pluged2) by simp also have "\ = 14 + (6 * 2 ^ (l-1) + 2*((2 ^ (l - 1) - 1) * 14 + (l - 1) * 15 * 2 ^ (l - 1 - 1) + 2 ^ (l - 1)) + (5 * 2 ^ (l - 1) + 4 * (2 ^ l div 2)))" by (smt (verit) "3"(3) add.assoc div_less evens_odds_length left_add_twice length_nums2 lessI mult.assoc mult_2_right nat_1_add_1 numeral_Bit0 nums2_def plus_1_eq_Suc power_eq_if power_not_zero zero_neq_numeral) also have "\ = 14 + 15 * 2 ^ (l-1) + 2*((2 ^ (l - 1) - 1) * 14 + (l - 1) * 15 * 2 ^ (l - 1 - 1) + 2 ^ (l - 1))" by (smt (z3) "3"(3) add.assoc add.commute calculation diff_diff_left distrib_left div2_Suc_Suc evens_odds_length left_add_twice length_Cons length_nums2 mult.assoc mult.commute mult_2 mult_2_right numeral_Bit0 numeral_Bit1 numeral_plus_numeral nums2_def one_add_one) also have "... = 14 + 15 * 2 ^ (l-1) + (2 ^ l - 2) * 14 + (l - 1) * 15 * 2 ^ (l - 1) + 2 ^ l" apply(cases "l > 1") apply (smt (verit, del_insts) add.assoc diff_is_0_eq distrib_left_numeral left_diff_distrib' less_imp_le_nat mult.assoc mult_2 mult_2_right nat_1_add_1 not_le not_one_le_zero power_eq_if) by (smt (z3) "3"(3) add.commute add.right_neutral cancel_comm_monoid_add_class.diff_cancel diff_add_inverse2 diff_is_0_eq div_less_dividend evens_odds_length length_nums2 mult_2 mult_eq_0_iff nat_1_add_1 not_le nums2_def power_eq_if) also have "\ = 15 * 2 ^ (l - 1) + (2 ^ l - 1) * 14 + (l - 1) * 15 * 2 ^ (l - 1) + 2 ^ l" by (smt (z3) "3"(3) One_nat_def add.commute combine_common_factor diff_add_inverse2 diff_diff_left list.size(4) nat_1_add_1 nat_mult_1) also have "\ = (2^l - 1) * 14 + l *15*2^(l-1) + 2^l" apply(cases "l > 0") subgoal using group_cancel.add1 group_cancel.add2 less_numeral_extra(3) mult.assoc mult_eq_if by auto[1] using "3"(3) by fastforce text \By the previous proposition, we can conclude that $T$ is indeed a suitable term for describing the running time\ finally have "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T (x # y # numbers) = T (x # y # numbers)" using T_def(1)[of "x#y#numbers" l] by (metis "3.prems" bits_1_div_2 diff_is_0_eq' evens_odds_length length_nums2 neq0_conv nums2_def power_0 zero_le_one zero_neq_one) thus ?case by simp qed (auto simp add: assms) text \We can finally state that $FNTT$ has $\mathcal{O}(n \log n)$ time complexity.\ theorem log_lin_time: assumes "length numbers = 2^l" shows "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T numbers \ 30 * l * length numbers + 1" proof- have 00: "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T numbers = (2 ^ l - 1) * 14 + l * 15 * 2 ^ (l - 1) + 2 ^ l" using tight_bound[of "\ xs. (length xs - 1) * 14 + (Discrete.log (length xs)) * 15 * 2 ^ ( (Discrete.log (length xs)) - 1) + length xs" numbers l] assms by simp have " l * 15 * 2 ^ (l - 1) \ 15 * l * length numbers" using assms by simp moreover have "(2 ^ l - 1) * 14 + 2^l\ 15 * length numbers " using assms by linarith moreover hence "(2 ^ l - 1) * 14 + 2^l \ 15 * l * length numbers +1" using assms apply(cases l) subgoal by simp by (metis (no_types) add.commute le_add1 mult.assoc mult.commute mult_le_mono nat_mult_1 plus_1_eq_Suc trans_le_add2) ultimately have " (2 ^ l - 1) * 14 + l * 15 * 2 ^ (l - 1) + 2 ^ l \ 30 * l * length numbers +1" by linarith then show ?thesis using 00 by simp qed theorem log_lin_time_explicitly: assumes "length numbers = 2^l" shows "T\<^sub>F\<^sub>N\<^sub>T\<^sub>T numbers \ 30 * Discrete.log (length numbers) * length numbers + 1" using log_lin_time[of numbers l] assms by simp end end diff --git a/thys/Number_Theoretic_Transform/Preliminary_Lemmas.thy b/thys/Number_Theoretic_Transform/Preliminary_Lemmas.thy --- a/thys/Number_Theoretic_Transform/Preliminary_Lemmas.thy +++ b/thys/Number_Theoretic_Transform/Preliminary_Lemmas.thy @@ -1,453 +1,456 @@ (* Title: Preliminary Lemmas for Number Theoretic Transform Author: Thomas Ammer *) theory Preliminary_Lemmas imports Berlekamp_Zassenhaus.Finite_Field "HOL-Number_Theory.Number_Theory" begin section \Preliminary Lemmas\ subsection \A little bit of Modular Arithmetic\ text \An obvious lemma. Just for simplification.\ lemma two_powrs_div: assumes "j < (i::nat) " shows "((2^i) div ((2::nat)^(Suc j)))*2 = ((2^i) div (2^j))" proof- have "((2::nat)^i) div (2^(Suc j)) = 2^(i -1) div(2^ j)" using assms by (smt (z3) One_nat_def add_le_cancel_left diff_Suc_Suc div_by_Suc_0 div_if less_nat_zero_code plus_1_eq_Suc power_diff_power_eq zero_neq_numeral) thus ?thesis by (metis Suc_diff_Suc Suc_leI assms less_imp_le_nat mult.commute power_Suc power_diff_power_eq zero_neq_numeral) qed lemma two_powr_div: assumes "j < (i::nat) " shows "((2^i) div ((2::nat)^j)) = 2^(i-j)" by (simp add: assms less_or_eq_imp_le power_diff) text \ The order of an element is the same whether we consider it as an integer or as a natural number. \ (* TODO: Move *) lemma ord_int: "ord (int p) (int x) = ord p x" proof (cases "coprime p x") case False thus ?thesis by (auto simp: ord_def) next case True have "(LEAST d. 0 < d \ [int x ^ d = 1] (mod int p)) = ord p x" proof (intro Least_equality conjI) show "[int x ^ ord p x = 1] (mod int p)" using True by (metis cong_int_iff of_nat_1 of_nat_power ord_works) show "ord p x \ y" if "0 < y \ [int x ^ y = 1] (mod int p)" for y using that by (metis cong_int_iff int_ops(2) linorder_not_less of_nat_power ord_minimal) qed (use True in auto) thus ?thesis by (auto simp: ord_def) qed lemma not_residue_primroot_1: assumes "n > 2" shows "\residue_primroot n 1" using assms totient_gt_1[of n] by (auto simp: residue_primroot_def) lemma residue_primroot_not_cong_1: assumes "residue_primroot n g" "n > 2" shows "[g \ 1] (mod n)" using residue_primroot_cong not_residue_primroot_1 assms by metis text \ We want to show the existence of a generating element of $\mathbb{Z}_p$ where $p$ is prime. \label{primroot1} \ text \Non-trivial order of an element $g$ modulo $p$ in a ring implies $g\neq1$. Although this lemma applies to all rings, it's only intended to be used in connection with $nat$s or $int$s \ lemma prime_not_2_order_not_1: assumes "prime p" "p > 2 " "ord p g > 2" shows "g \ 1" proof assume "g = 1" hence "ord p g = 1" unfolding ord_def by (simp add: Least_equality) then show False using assms by auto qed text \The same for modular arithmetic.\ lemma prime_not_2_order_not_1_mod: assumes "prime p " "p > 2 " "ord p g > 2" shows "[g \ 1] (mod p)" proof assume "[g = 1] (mod p)" hence "ord p g = 1" unfolding ord_def by(split if_split, metis assms(1) assms(2) assms(3) ord_cong prime_not_2_order_not_1) then show False using assms by auto qed text \ Now we formulate our lemma about generating elements in residue classes: There is an element $g \in \mathbb{Z}_p$ such that for any $x \in \mathbb{Z}_p$ there is a natural $i$ such that $g^i \equiv x \; (\mod p)$.\ lemma generator_exists: assumes "prime (p::nat)" "p > 2" shows "\ g. [g \ 1] (mod p) \ (\ x. (0 x < p )\ (\ i. [g^i = x] (mod p)))" proof- obtain g where g_prim_root:"residue_primroot p g" using assms prime_gt_1_nat prime_primitive_root_exists by (metis One_nat_def) have g_not_1: "[g \ 1] (mod p)" using residue_primroot_not_cong_1 assms g_prim_root by blast have "\i. [g ^ i = x] (mod p)" if x_bounds: "x > 0" "x < p" for x proof - have 1:"coprime p x" using assms prime_nat_iff'' x_bounds by blast have 2:"ord p g = p-1" by (metis assms(1) g_prim_root residue_primroot_def totient_prime) hence bij: "bij_betw (\i. g ^ i mod p) {.. totatives p" by (simp add: "1" coprime_commute in_totatives_iff order_le_less x_bounds) have " {.. {}" by (metis assms(1) lessThan_empty_iff prime_nat_iff'' totient_0_iff) then obtain i where "g^i mod p = x mod p" using bij_betw_inv[of "(\i. g ^ i mod p)" "{..General Lemmas in a Finite Field\ text \ \label{primroot2} We make certain assumptions: From now on, we will calculate in a finite field which is the ring of integers modulo a prime $p$. Let $n$ be the length of vectors to be transformed. By Dirichlet's theorem on arithmetic progressions we can assume that there is a natural number $k$ and a prime $p$ with $p = k\cdot n + 1$. In order to avoid some special cases and even contradictions, we additionally assume that $p \geq 3$ and $n \geq 2$. \ text \\label{prelim}\ locale preliminary = fixes a_type::"('a::prime_card) itself" and p::nat and n::nat and k::nat assumes p_def: "p= CARD('a)" and p_lst3: "p > 2" and p_fact: "p = k*n +1" and n_lst2: "n \ 2" begin lemma exp_rule: "((c::('a) mod_ring) * d )^e= (c^e) * (d^e)" by (simp add: power_mult_distrib) lemma "\ y. x \ 0 \ (x::(('a) mod_ring)) * y = 1" by (metis dvd_field_iff unit_dvdE) lemma test: "prime p" by (simp add: p_def prime_card) lemma k_bound: "k > 0" using p_fact prime_nat_iff'' test by force text \We show some homomorphisms.\ lemma homomorphism_add: "(of_int_mod_ring x)+(of_int_mod_ring y) = ((of_int_mod_ring (x+y)) ::(('a::prime_card) mod_ring))" by (metis of_int_hom.hom_add of_int_of_int_mod_ring) lemma homomorphism_mul_on_ring: "(of_int_mod_ring x)*(of_int_mod_ring y) = ((of_int_mod_ring (x*y)) ::(('a::prime_card) mod_ring))" by (metis of_int_mult of_int_of_int_mod_ring) lemma exp_homo:"(of_int_mod_ring (x^i)) = ((of_int_mod_ring x)^i ::(('a::prime_card) mod_ring))" by (induction i) (metis of_int_of_int_mod_ring of_int_power)+ lemma mod_homo: "((of_int_mod_ring x)::(('a::prime_card) mod_ring)) = of_int_mod_ring (x mod p)" using p_def unfolding of_int_mod_ring_def by simp lemma int_exp_hom: "int x ^i = int (x^i)" by simp lemma coprime_nat_int: "coprime (int p) (to_int_mod_ring pr) \ coprime p (nat(to_int_mod_ring pr))" unfolding coprime_def to_int_mod_ring_def by (smt (z3) Rep_mod_ring atLeastLessThan_iff dvd_trans int_dvd_int_iff int_nat_eq int_ops(2) prime_divisor_exists prime_nat_int_transfer primes_dvd_imp_eq test to_int_mod_ring.rep_eq to_int_mod_ring_def) lemma nat_int_mod:"[nat (to_int_mod_ring pr) ^ d = 1] (mod p) = [ (to_int_mod_ring pr) ^ d = 1] (mod (int p)) " unfolding to_int_mod_ring_def by (metis Rep_mod_ring atLeastLessThan_iff cong_int_iff int_exp_hom int_nat_eq int_ops(2) to_int_mod_ring.rep_eq to_int_mod_ring_def) text \Order of $p$ doesn't change when interpreting it as an integer.\ lemma ord_lift: "ord (int p) (to_int_mod_ring pr) = ord p (nat (to_int_mod_ring pr))" proof - have "to_int_mod_ring pr = int (nat (to_int_mod_ring pr))" by (metis Rep_mod_ring atLeastLessThan_iff int_nat_eq to_int_mod_ring.rep_eq) thus ?thesis using ord_int by metis qed text \A primitive root has order $p-1$.\ lemma primroot_ord: "residue_primroot p g \ ord p g = p -1" by (simp add: residue_primroot_def test totient_prime) text \If $x^l = 1$ in $\mathbb{Z}_p$, then $l$ is an upper bound for the order of $x$ in $\mathbb{Z}_ p$.\ lemma ord_max: assumes "l \ 0" "(x :: (('a::prime_card) mod_ring))^l = 1" shows " ord p (to_int_mod_ring x) \ l" proof- have "[(to_int_mod_ring x)^l = 1] (mod p)" by (metis assms(2) cong_def exp_homo of_int_mod_ring.rep_eq of_int_mod_ring_to_int_mod_ring one_mod_card_int one_mod_ring.rep_eq p_def) thus ?thesis unfolding ord_def using assms by (smt (z3) Least_le less_imp_le_nat not_gr0) qed subsection \Existence of $n$-th Roots of Unity in the Finite Field\ text \ \label{primroot3} We obtain an element in the finite field such that its reinterpretation as a $nat$ will be a primitive root in the residue class modulo $p$. The difference between residue classes, their representatives in the Integers and elements of the finite field is notable. When conducting informal proofs, this distinction is usually blurred, but Isabelle enforces the explicit conversion between those structures. \ lemma primroot_ex: obtains primroot::"('a::prime_card) mod_ring" where "primroot^(p-1) = 1" "primroot \ 1" "residue_primroot p (nat (to_int_mod_ring primroot))" proof- obtain g where g_Def: "residue_primroot p g \ g \ 1" using prime_nat_iff' prime_primitive_root_exists test by (metis bigger_prime euler_theorem ord_1_right power_one_right prime_nat_iff'' residue_primroot.cases residue_primroot_cong) hence "[g \ 1] (mod p)" using prime_not_2_order_not_1_mod[of p g] by (metis One_nat_def p_lst3 less_numeral_extra(4) ord_eq_Suc_0_iff residue_primroot.cases totient_gt_1) hence "[g^(p-1) = 1] (mod p)" using g_Def by (metis coprime_commute euler_theorem residue_primroot_def test totient_prime) moreover hence "int (g ^ (p - 1)) mod int p = (1::int)" by (metis cong_def int_ops(2) mod_less of_nat_mod prime_gt_1_nat test) moreover hence "of_int_mod_ring (int (g ^ (p - 1)) mod int p) = ((of_int_mod_ring 1) ::(('a::prime_card) mod_ring))" by simp ultimately have "(of_int_mod_ring (g^(p-1))) = (1 ::(('a::prime_card) mod_ring))" using mod_homo[of "g^(p-1)"] by (metis exp_homo power_0) hence "((of_int_mod_ring g)^(p-1) ::(('a::prime_card) mod_ring)) = 1" using exp_homo[of "int g" "p-1"] by simp moreover have "((of_int_mod_ring g) ::(('a::prime_card) mod_ring)) \ 1" proof assume "((of_int_mod_ring g) ::(('a::prime_card) mod_ring)) = 1" hence "[int g = 1] (mod p)" using p_def unfolding of_int_mod_ring_def by (metis \of_int_mod_ring (int g) = 1\ cong_def of_int_mod_ring.rep_eq one_mod_card_int one_mod_ring.rep_eq) hence "[g=1] (mod p)" by (metis cong_int_iff int_ops(2)) thus False using \[g \ 1] (mod p)\ by auto - qed - thus ?thesis using that g_Def calculation - by (metis Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign mod_homo nat_int of_nat_0_less_iff of_nat_mod p_def residue_primroot_mod to_int_mod_ring_of_int_mod_ring zero_less_card_finite) + qed + moreover have \residue_primroot p (g mod p)\ + using g_Def by simp + then have \residue_primroot p (nat (to_int_mod_ring (of_int_mod_ring (int g) :: 'a mod_ring)))\ + by (transfer fixing: p) (simp add: p_def nat_mod_distrib) + ultimately show thesis .. qed text \From this, we obtain an $n$-th root of unity $\omega$ in the finite field of characteristic $p$. Note that in this step we will use the assumption $p = k \cdot n +1$ from locale $preliminary$: The $k$-th power of a primitive root $pr$ modulo $p$ will have the property $(pr^k)^n \equiv 1 \mod p$. \ lemma omega_properties_ex: obtains \ ::"(('a::prime_card) mod_ring)" where "\^n = 1" "\ \ 1" "\ m. \^m = 1 \ m\0 \ m \ n" proof- obtain pr::"(('a::prime_card) mod_ring)" where a: "pr^(p-1) = 1 " and b: "pr \ 1" and c: "residue_primroot p (nat( to_int_mod_ring pr))" using primroot_ex by blast moreover hence "(pr^k)^n =1" by (simp add: p_fact power_mult) moreover have "pr^k \ 1" proof assume " pr ^ k = 1" hence "(to_int_mod_ring pr)^k mod p = 1" by (metis exp_homo of_int_mod_ring.rep_eq of_int_mod_ring_to_int_mod_ring one_mod_ring.rep_eq p_def) hence "ord p (to_int_mod_ring pr) \ k" by (simp add: \pr ^ k = 1\ k_bound ord_max) hence "ord p (nat (to_int_mod_ring pr)) \ k" by (metis ord_lift) also have "ord p (nat (to_int_mod_ring pr)) = p - 1" using c primroot_ord[of "(nat (to_int_mod_ring pr))"] by blast also have "\ = k * n" using p_fact by simp finally have "n \ 1" using k_bound by simp thus False using n_lst2 by linarith qed moreover have "\ m. (pr^k)^m = 1 \ m\0 \ m \ n" proof(rule ccontr) assume "\ (\m. (pr ^ k) ^ m = 1 \ m\0 \ n \ m) " then obtain m where "(pr^k)^m = 1 \ m\0 \ m < n" by force hence "ord p (to_int_mod_ring pr) \ k * m" using ord_max[of "k*m" pr] by (metis calculation(5) mult_is_0 power_mult) moreover have "ord p (nat (to_int_mod_ring pr)) = p-1" using c primroot_ord ord_lift by simp ultimately show False by (metis \(pr ^ k) ^ m = 1 \ m \ 0 \ m < n\ add_diff_cancel_right' nat_0_less_mult_iff nat_mult_le_cancel_disj not_less ord_lift p_def p_fact prime_card prime_gt_1_nat zero_less_diff) qed ultimately show ?thesis using that by simp qed text \We define an $n$-th root of unity $\omega$ for $NTT$.\ theorem omega_exists: "\ \ ::(('a::prime_card) mod_ring) . \^n = 1 \ \ \ 1 \ (\ m. \^m = 1 \ m\0 \ m \ n)" using omega_properties_ex by metis definition "(omega::(('a::prime_card) mod_ring)) = (SOME \ . (\^n = 1 \ \ \ 1\ (\ m. \^m = 1 \ m\0 \ m \ n)))" lemma omega_properties: "omega^n = 1" "omega \ 1" "(\ m. omega^m = 1 \ m\0 \ m \ n)" unfolding omega_def using omega_exists by (smt (verit, best) verit_sko_ex')+ text \We define the multiplicative inverse $\mu$ of $\omega$.\ definition "mu = omega ^ (n - 1)" lemma mu_properties: "mu * omega = 1" "mu \ 1" proof - have "omega ^ (n - 1) * omega = omega ^ Suc (n - 1)" by simp also have "Suc (n - 1) = n" using n_lst2 by simp also have "omega ^ n = 1" using omega_properties(1) by auto finally show "mu * omega = 1" by (simp add: mu_def) next show "mu \ 1" using omega_properties n_lst2 by (auto simp: mu_def) qed subsection \Some Lemmas on Sums\ text \\label{sums}\ text \The following lemmas concern sums over a finite field. Most of the propositions are intuitive.\ lemma sum_in: "(\i=0..<(x::nat). f i * (y ::('a mod_ring))) = (\i=0.. i. i < x \ f i = g i) \ (\i=0..<(x::nat). f i) = (\i=0..i=0..<(x::nat). (f i)::('a mod_ring)) - (\i=0..i=0..i=0..<(x::nat). \j=0..<(y::nat). f i j) = (\j=0..<(y::nat). \i=0..<(x::nat). f i j ) " using Groups_Big.comm_monoid_add_class.sum.swap by fast lemma sum_const: "(\i=0..<(x::nat). (c::('a::prime_card) mod_ring)) = (of_int_mod_ring x) * c" by(induction x, simp add: algebra_simps, simp add: algebra_simps) (metis distrib_left mult.right_neutral of_int_of_int_mod_ring of_int_of_nat_eq of_nat_Suc) lemma sum_split: "(r1::nat) < r2 \ (\l = 0..l = r1..l = 0..l = (a::nat)..< b. f(l+c)) = (\l = (a+c)..< (b+c). f l )" by(induction a arbitrary: b c) (metis sum.shift_bounds_nat_ivl)+ text \One may sum over even and odd indices independently. The lemma statement was taken from a formalization of FFT~\parencite{FFT-AFP}. We give an alternative proof adapted to the finite field $\mathbb{Z}_p$. \ lemma sum_splice: "(\i::nat = 0..<2*nn. f i) = (\i = 0..i = 0..i::nat = 0..<2*(n+1). f i) = (\i::nat = 0..<(2*n). f i) + f(2*n+1) + f (2*n)" by( simp add: algebra_simps) also have "\ = (\i::nat = 0..i::nat = 0.. = (\i::nat = 0..<(Suc n). f (2*i)) + (\i::nat = 0..<(Suc n). f (2*i+1))" by( simp add: algebra_simps) finally show ?case by simp qed simp lemma sum_even_odd_split: "even (a::nat) \ (\j=0..<(a div 2). f (2*j))+ (\j=0..<(a div 2). f (2*j+1)) = (\j=0..j=(0::nat)..j=0..j=(0::nat)..<2*i. f j )" by (metis sum_splice) lemma sum_neg_in: "- (\j = 0..j = 0..Geometric Sums\ text \\label{geosum}\ text \This lemma will be important for proving properties on $\mathsf{NTT}$. At first, an informal proof sketch: \begin{align*} (1-x) \cdot \sum \limits _{l = 0} ^ {r-1} x^l &= \sum \limits _{l = 0} ^ {r-1} x^l - x \cdot \sum \limits _{l = 0} ^{r-1} x^l \\ &= \sum \limits _{l = 0} ^ {r-1} x^l - \sum \limits _{l = 1} ^{r} x^l \\ & = 1 - x^r \end{align*} The same lemma for integers can be found in~\parencite{Dirichlet_Series-AFP}. Our version is adapted to finite fields. \ lemma geo_sum: assumes "x \ 1" shows "(1-x)*(\l = 0..l = 0..l = 0.. l. x^l" x r] by(simp add: algebra_simps) have 1:"(\l = 0..l = 0..l = 0..l = 0..This theory is intended to develop the necessary background on subsets of powers of a $p$-adic field to prove Macintyre's quantifier elimination theorem. In particular, we define semi-algebraic subsets of $\mathbb{Q}_p^n$, semi-algebraic functions $\mathbb{Q}_p^n \to \mathbb{Q}_p$, and semi- algebraic mappings $\mathbb{Q}_p^n \to \mathbb{Q}_p^m$ for arbitrary $n, m \in \mathbb{N}$. In addition we prove that many common sets and functions are semi-algebraic. We are closely following the paper \<^cite>\"denef1986"\ by Denef, where an algebraic proof of Mactinyre's theorem is developed.\ (**************************************************************************************************) (**************************************************************************************************) section\Cartesian Powers of $p$-adic Fields\ (**************************************************************************************************) (**************************************************************************************************) lemma list_tl: "tl (t#x) = x" using List.list.sel(3) by auto lemma list_hd: "hd (t#x) = t" unfolding List.list.sel(1) by auto sublocale padic_fields < cring_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding cring_coord_rings_axioms_def cring_coord_rings_def using Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.is_UP_cring) by auto sublocale padic_fields < Qp: domain_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding domain_coord_rings_def cring_coord_rings_axioms_def cring_coord_rings_def using Qp.domain_axioms Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.UP_cring_axioms) by auto context padic_fields begin no_notation Zp.to_fun (infixl\\\ 70) no_notation ideal_prod (infixl "\\" 80) notation evimage (infixr "\\" 90) and euminus_set ("_ \<^sup>c\" 70) type_synonym padic_tuple = "padic_number list" type_synonym padic_function = "padic_number \ padic_number" type_synonym padic_nary_function = "padic_tuple \ padic_number" type_synonym padic_function_tuple = "padic_nary_function list" type_synonym padic_nary_function_poly = "nat \ padic_nary_function" (**************************************************************************************************) (**************************************************************************************************) subsection\Polynomials over $\mathbb{Q}_p$ and Polynomial Maps\ (**************************************************************************************************) (**************************************************************************************************) lemma last_closed': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "t \ carrier Q\<^sub>p" using assms last_closed[of n "x@[t]" Q\<^sub>p] by (metis (full_types) cartesian_power_car_memE gr0I last_snoc length_append_singleton less_not_refl zero_less_Suc) lemma segment_in_car': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" proof- have 0: "length x = n" by (metis Suc_inject assms cartesian_power_car_memE length_append_singleton) have "set x \ set (x@[t])" by (metis rotate1.simps(2) set_rotate1 set_subset_Cons) then have 1: "set x \ carrier Q\<^sub>p" using assms cartesian_power_car_memE''[of "x@[t]" Q\<^sub>p "Suc n"] by blast show ?thesis using 0 1 assms cartesian_power_car_memI[of x n Q\<^sub>p] by blast qed lemma Qp_zero: "Q\<^sub>p\<^bsup>0\<^esup> = nil_ring" unfolding cartesian_power_def by simp lemma Qp_zero_carrier: "carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" by (simp add: Qp_zero) text\Abbreviation for constant polynomials\ abbreviation(input) Qp_to_IP where "Qp_to_IP k \ Qp.indexed_const k" lemma Qp_to_IP_car: assumes "k \ carrier Q\<^sub>p" shows "Qp_to_IP k \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Qp.indexed_const_closed by blast lemma(in cring_coord_rings) smult_closed: assumes "a \ carrier R" assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q \ carrier (R[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Pring_smult_closed by (simp add: R.Pring_smult_closed) lemma Qp_poly_smult_cfs: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) m = a \ (P m)" using assms unfolding coord_ring_def using Qp.Pring_smult_cfs by blast lemma Qp_smult_r_distr: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q) = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q)" using assms unfolding coord_ring_def using Qp.Pring_smult_r_distr by blast lemma Qp_smult_l_distr: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \ b) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (b \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P)" using assms unfolding coord_ring_def using Qp.Pring_smult_l_distr by blast abbreviation(input) Qp_funs where "Qp_funs n \ Fun\<^bsub>n\<^esub> Q\<^sub>p" (**************************************************************************************************) (**************************************************************************************************) subsection\Evaluation of Polynomials in $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) Qp_ev where "Qp_ev P q \ (eval_at_point Q\<^sub>p q P)" lemma Qp_ev_one: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \" unfolding coord_ring_def by (metis Qp.Pring_one eval_at_point_const Qp.one_closed assms) lemma Qp_ev_zero: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \"unfolding coord_ring_def by (metis Qp.Pring_zero eval_at_point_const Qp.zero_closed assms) lemma Qp_eval_pvar_pow: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "k < n" assumes "(m::nat) \ 0" shows "Qp_ev ((pvar Q\<^sub>p k)[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> m) a = ((a!k)[^]m)" by (metis eval_at_point_nat_pow eval_pvar assms(1) assms(2) pvar_closed) text\composition of polynomials over $\mathbb{Q}_p$\ definition Qp_poly_comp where "Qp_poly_comp m fs = poly_compose (length fs) m fs" text\lemmas about polynomial maps\ lemma Qp_is_poly_tupleI: assumes "\i. i < length fs\ fs!i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_poly_tuple m fs" unfolding is_poly_tuple_def using assms using cartesian_power_car_memE'' cartesian_power_car_memI' by blast lemma Qp_is_poly_tuple_append: assumes "is_poly_tuple m fs" assumes "is_poly_tuple m gs" shows "is_poly_tuple m (fs@gs)" proof(rule Qp_is_poly_tupleI) show "\i. i < length (fs @ gs) \ (fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" proof- fix i assume A: "i < length (fs @ gs)" show "(fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" apply(cases "i < length fs") using assms is_poly_tupleE[of m fs i] nth_append[of fs gs i] apply presburger proof- assume B: "\ i < length fs" then have C: "length fs \ i \ i < length (fs @ gs)" using A not_le by blast then have "i - length fs < length gs" using length_append[of fs gs] by linarith then show ?thesis using A assms is_poly_tupleE[of m gs "i - length fs"] nth_append[of fs gs i] B by presburger qed qed qed lemma Qp_poly_mapE: assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j \ carrier Q\<^sub>p" using assms poly_map_closed cartesian_power_car_memE' by blast lemma Qp_poly_mapE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = length fs" unfolding poly_map_def using Qp.cring_axioms poly_tuple_evalE' by (metis assms restrict_def) lemma Qp_poly_mapE'': assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "n \ 0" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j = (Qp_ev (fs!j) as)" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) nth_map restrict_apply') lemma poly_map_apply: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n fs as = poly_tuple_eval fs as" unfolding poly_map_def restrict_def by (simp add: assms) lemma poly_map_pullbackI: assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ poly_map n fs \\<^bsub>n\<^esub> S" using assms poly_map_apply by blast lemma poly_map_pullbackI': assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ ((poly_map n fs) -` S)" by (simp add: assms(3)) text\lemmas about polynomial composition\ lemma poly_compose_ring_hom: assumes "is_poly_tuple m fs" assumes "length fs = n" shows "(ring_hom_ring (Q\<^sub>p[\\<^bsub>n\<^esub>]) (Q\<^sub>p[\\<^bsub>m\<^esub>]) (Qp_poly_comp m fs))" unfolding Qp_poly_comp_def by (simp add: assms(1) assms(2) poly_compose_ring_hom) lemma poly_compose_closed: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(Qp_poly_comp m fs f) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" using Qp.cring_axioms assms unfolding Qp_poly_comp_def using poly_compose_closed by blast lemma poly_compose_add: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_add unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_mult: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_mult unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_const: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "a \ carrier Q\<^sub>p" shows "Qp_poly_comp m fs (Qp_to_IP a) = Qp_to_IP a" using Qp.cring_axioms assms poly_compose_const unfolding is_poly_tuple_def Qp_poly_comp_def by metis lemma Qp_poly_comp_eval: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" shows "Qp_ev (Qp_poly_comp m fs f) as = Qp_ev f (poly_map m fs as)" proof- have "(restrict (poly_tuple_eval fs) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) as) = poly_tuple_eval fs as" unfolding restrict_def by (simp add: assms) thus ?thesis using assms Qp.cring_axioms poly_compose_eval unfolding Qp_poly_comp_def poly_map_def by presburger qed (**************************************************************************************************) (**************************************************************************************************) subsection\Mapping Univariate Polynomials to Multivariable Polynomials in One Variable\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) to_Qp_x where "to_Qp_x \ (IP_to_UP (0::nat) :: (nat multiset \ padic_number) \ nat \ padic_number)" abbreviation(input) from_Qp_x where "from_Qp_x \ UP_to_IP Q\<^sub>p (0::nat)" lemma from_Qp_x_closed: assumes "q \ carrier Q\<^sub>p_x" shows "from_Qp_x q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" using assms UP_to_IP_closed unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_closed: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "to_Qp_x q \ carrier Q\<^sub>p_x" using assms Qp.IP_to_UP_closed[of q "0::nat"] Qp.cring_axioms unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_from_Qp_x: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "from_Qp_x (to_Qp_x q) = q" using assms UP_to_IP_inv[of q "0::nat"] Qp.Pring_car unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_to_Qp_x: assumes "q \ carrier Q\<^sub>p_x" shows "to_Qp_x (from_Qp_x q) = q" by (meson UPQ.IP_to_UP_inv assms) text\ring hom properties of these maps\ lemma to_Qp_x_ring_hom: "to_Qp_x \ ring_hom (Q\<^sub>p[\\<^bsub>1\<^esub>]) Q\<^sub>p_x" using IP_to_UP_ring_hom[of "0::nat"] ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_ring_hom: "from_Qp_x \ ring_hom Q\<^sub>p_x (Q\<^sub>p[\\<^bsub>1\<^esub>])" using UP_to_IP_ring_hom ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_add: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis (mono_tags, lifting) assms(1) assms(2) from_Qp_x_ring_hom ring_hom_add) lemma from_Qp_x_mult: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis assms(1) assms(2) from_Qp_x_ring_hom ring_hom_mult) text\equivalence of evaluation maps\ lemma Qp_poly_Qp_x_eval: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" shows "Qp_ev P a = (to_Qp_x P)\(Qp.to_R a)" proof- have 0: "(IP_to_UP 0 P) \ (a ! 0) = ((IP_to_UP 0 P) \ (if 0 < length a then a ! 0 else \))" using assms by (metis (mono_tags, lifting) cartesian_power_car_memE gr0I zero_neq_one) have 1: "closed_fun Q\<^sub>p (\n. if n < length a then a ! n else \)" proof fix n show "(if n < length a then a ! n else \) \ carrier Q\<^sub>p" apply(cases "n < length a") using assms apply (metis cartesian_power_car_memE cartesian_power_car_memE') by (meson Qp.cring_axioms cring.cring_simprules(2)) qed have 2: " P \ Pring_set Q\<^sub>p {0::nat}" using assms unfolding coord_ring_def by (metis Qp.Pring_car UPQ.UP_to_IP_closed assms(1) to_Qp_x_closed to_Qp_x_from_Qp_x) have 3: "total_eval Q\<^sub>p (\i. if i < length a then a ! i else \) P = IP_to_UP 0 P \ (if 0 < length a then a ! 0 else \)" using 1 2 assms IP_to_UP_poly_eval[of P "0::nat" "(\i. if i < length a then a ! i else \)" ] UPQ.to_fun_def by presburger then show ?thesis using 0 unfolding eval_at_point_def by blast qed lemma Qp_x_Qp_poly_eval: assumes "P \ carrier Q\<^sub>p_x" assumes "a \ carrier Q\<^sub>p" shows "P \ a = Qp_ev (from_Qp_x P) (to_R1 a)" proof- have "Qp_ev (from_Qp_x P) (to_R1 a) = (to_Qp_x (from_Qp_x P))\(Qp.to_R (Qp.to_R1 a))" using Qp_poly_Qp_x_eval assms(1) assms(2) from_Qp_x_closed Qp.to_R1_closed by blast then show ?thesis using assms by (metis UPQ.IP_to_UP_inv Qp.to_R_to_R1) qed (**************************************************************************************************) (**************************************************************************************************) subsection\$n^{th}$-Power Sets over $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) definition P_set where "P_set (n::nat) = {a \ nonzero Q\<^sub>p. (\y \ carrier Q\<^sub>p . (y[^] n) = a)}" lemma P_set_carrier: "P_set n \ carrier Q\<^sub>p" unfolding P_set_def nonzero_def by blast lemma P_set_memI: assumes "a \ carrier Q\<^sub>p" assumes "a \ \" assumes "b \ carrier Q\<^sub>p" assumes "b[^](n::nat) = a" shows "a \ P_set n" unfolding P_set_def using assms by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) lemma P_set_nonzero: "P_set n \ nonzero Q\<^sub>p" unfolding P_set_def by blast lemma P_set_nonzero': assumes "a \ P_set n" shows "a \ nonzero Q\<^sub>p" "a \ carrier Q\<^sub>p" using assms P_set_nonzero P_set_carrier apply blast using assms P_set_carrier by blast lemma P_set_one: assumes "n \ 0" shows "\ \ P_set (n::nat)" proof- have 0: "\[^]n = \" using Qp.nat_pow_one by blast have 1: "\ \ carrier Q\<^sub>p" by blast then show ?thesis using one_nonzero unfolding P_set_def using 0 by blast qed lemma zeroth_P_set: "P_set 0 = {\}" proof show "P_set 0 \ {\}" unfolding P_set_def proof fix x assume "x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^](0::nat)) = a}" then have "\y\carrier Q\<^sub>p. (y[^](0::nat)) = x" by blast then obtain a where a_def: "a \ carrier Q\<^sub>p \ (a[^](0::nat)) = x" by blast then show "x \ {\}" using Qp.nat_pow_0 by blast qed show "{\} \ P_set 0" using P_set_memI[of \ \ 0] Qp.nat_pow_one Qp.one_closed local.one_neq_zero by blast qed lemma P_set_mult_closed: assumes "n \ 0" assumes "a \ P_set n" assumes "b \ P_set n" shows "a \ b \ P_set n" proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ (a0 [^] n = a)" using assms(2) unfolding P_set_def by blast obtain b0 where b0_def: "b0 \ carrier Q\<^sub>p \ (b0 [^] n = b)" using assms(3) unfolding P_set_def by blast have "(a0 \ b0) [^] n = a0 [^] n \ b0 [^] n" using a0_def b0_def Qp.nat_pow_distrib by blast then have 0: "a \ b = (a0 \ b0) [^] n" using a0_def b0_def by blast have 1: "a0 \ b0 \ carrier Q\<^sub>p" by (meson Qp.cring_axioms a0_def b0_def cring.cring_simprules(5)) have 2: "a \ b \ nonzero Q\<^sub>p" using assms nonzero_is_submonoid unfolding P_set_def by (metis (no_types, lifting) "0" "1" Qp.integral Qp_nat_pow_nonzero a0_def b0_def mem_Collect_eq not_nonzero_Qp) then show ?thesis using 0 1 assms unfolding P_set_def by blast qed lemma P_set_inv_closed: assumes "a \ P_set n" shows "inv a \ P_set n" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis Qp.inv_one singletonD) next case False then show ?thesis proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ a0[^]n = a" using assms P_set_def[of n] by blast have "a0 \ nonzero Q\<^sub>p" apply(rule ccontr) proof- assume "a0 \ nonzero Q\<^sub>p " then have "a0 = \" using a0_def by (meson not_nonzero_Qp) then show False using a0_def assms by (metis (mono_tags, lifting) False P_set_def Qp.cring_axioms \a0 \ nonzero Q\<^sub>p\ cring_def mem_Collect_eq neq0_conv ring.pow_zero) qed then have "(inv a0)[^]n = inv a" using a0_def \a0 \ carrier Q\<^sub>p \ (a0[^]n) = a\ \a0 \ nonzero Q\<^sub>p\ Units_nonzero monoid.nat_pow_of_inv[of Q\<^sub>p a n] Qp.nat_pow_of_inv Units_eq_nonzero by presburger then show ?thesis by (metis P_set_memI Qp.nat_pow_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero \a0 \ nonzero Q\<^sub>p\ a0_def inv_in_frac(1) inv_in_frac(2)) qed qed lemma P_set_val: assumes "a \ P_set (n::nat)" shows "(ord a) mod n = 0" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis mod_by_0 of_nat_0 ord_one singletonD) next case False then show ?thesis proof- obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^] n) = a" using assms P_set_def by blast have an: "a \ nonzero Q\<^sub>p" using P_set_def assms by blast have bn: "b \ nonzero Q\<^sub>p" proof(rule ccontr) assume "b \ nonzero Q\<^sub>p" then have "b = \\<^bsub> Q\<^sub>p\<^esub>" using b_def not_nonzero_Qp by metis then have "(b[^] n) = \" using False Qp.cring_axioms cring_def ring.pow_zero by blast then show False using b_def an Qp.not_nonzero_memI by blast qed then have "ord a = n * (ord b)" using b_def an nonzero_nat_pow_ord by blast then show ?thesis by (metis mod_mult_self1_is_0) qed qed lemma P_set_pow: assumes "n > 0" assumes "s \ P_set n" shows "s[^]k \ P_set (n*k)" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ y[^]n = s" using assms unfolding P_set_def by blast then have 0: "y \ nonzero Q\<^sub>p" using assms P_set_nonzero'(1) Qp_nonzero_nat_pow by blast have 1: "y[^](n*k) = s[^] k" using 0 y_def assms Qp.nat_pow_pow by blast hence 2: "s[^]k \ nonzero Q\<^sub>p" using 0 by (metis Qp_nat_pow_nonzero) thus ?thesis unfolding P_set_def using 1 y_def by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) text\ In this section we introduce the notion of a $p$-adic semialgebraic set. Intuitively, these are the subsets of $\mathbb{Q}_p^n$ which are definable by first order quantifier-free formulas in the standard first-order language of rings, with an additional relation symbol included for the relation $\text{ val}(x) \leq \text{ val}(y)$, interpreted according to the definiton of the $p$-adic valuation on $\mathbb{Q}_p$. In fact, by Macintyre's quantifier elimination theorem for the first-order theory of $\mathbb{Q}_p$ in this language, one can equivalently remove the ``quantifier-free" clause from the latter definition. The definition we give here is also equivalent, and due to Denef in \<^cite>\"denef1986"\. The given definition here is desirable mainly for its utility in producing a proof of Macintyre's theorem, which is our overarching goal. \ (********************************************************************************************) (********************************************************************************************) subsubsection\Defining Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition basic_semialg_set where "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). \y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)}" lemma basic_semialg_set_zero_set: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "Qp_ev P q = \" assumes "n \ 0" shows "q \ basic_semialg_set (m::nat) (n::nat) P" proof- have "\ = (\[^]n)" using assms(4) Qp.nat_pow_zero by blast then show ?thesis unfolding basic_semialg_set_def using assms Qp.cring_axioms cring.cring_simprules(2) by blast qed lemma basic_semialg_set_def': assumes "n \ 0" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ (P_set n)}" proof show "basic_semialg_set m n P \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" proof fix x assume A: "x \ basic_semialg_set m n P" show "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" apply(cases "Qp_ev P x = \") using A basic_semialg_set_def apply blast unfolding basic_semialg_set_def P_set_def proof assume A0: "Qp_ev P x \ \" have A1: " \y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n)" using A basic_semialg_set_def by blast have A2: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A basic_semialg_set_def by blast show " x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (Qp_ev P x = \ \ Qp_ev P x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a})" by (metis (mono_tags, lifting) A1 A2 Qp.nonzero_memI assms(2) eval_at_point_closed mem_Collect_eq) qed qed show "{q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} \ basic_semialg_set m n P" proof fix x assume A: " x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" then have A':"x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by blast show "x \ basic_semialg_set m n P" using A A' apply(cases "Qp_ev P x = \") using assms basic_semialg_set_zero_set[of P m x n] apply blast proof- assume B: "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} " assume B': "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assume B'': "Qp_ev P x \ \ " show "x \ basic_semialg_set m n P" unfolding basic_semialg_set_def P_set_def proof have "\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n) " using A nonzero_def [of Q\<^sub>p] unfolding P_set_def proof - assume "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a}}" then have "Qp_ev P x \ nonzero Q\<^sub>p \ (\r. r \ carrier Q\<^sub>p \ (r[^]n) = Qp_ev P x)" using B'' by blast then show ?thesis by blast qed then show "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n))" using B' by blast qed qed qed qed lemma basic_semialg_set_memI: assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "y \ carrier Q\<^sub>p" assumes "Qp_ev P q = (y[^]n)" shows "q \ basic_semialg_set m n P" using assms(1) assms(2) assms(3) basic_semialg_set_def by blast lemma basic_semialg_set_memE: assumes "q \ basic_semialg_set m n P" shows "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" "\y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)" using assms basic_semialg_set_def apply blast using assms basic_semialg_set_def by blast definition is_basic_semialg :: "nat \ ((nat \ int) \ (nat \ int)) set list set \ bool" where "is_basic_semialg m S \ (\ (n::nat) \ 0. (\ P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]). S = basic_semialg_set m n P))" abbreviation(input) basic_semialgs where "basic_semialgs m \ {S. (is_basic_semialg m S)}" definition semialg_sets where "semialg_sets n = gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" lemma carrier_is_semialg: "(carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ semialg_sets n " unfolding semialg_sets_def using gen_boolean_algebra.universe by blast lemma empty_set_is_semialg: " {} \ semialg_sets n" using carrier_is_semialg[of n] unfolding semialg_sets_def using gen_boolean_algebra.complement by (metis Diff_cancel) lemma semialg_intersect: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms(1) assms(2) gen_boolean_algebra_intersect semialg_sets_def by blast lemma semialg_union: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms gen_boolean_algebra.union semialg_sets_def by blast lemma semialg_complement: assumes "A \ semialg_sets n" shows "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) \ semialg_sets n " using assms gen_boolean_algebra.complement semialg_sets_def by blast lemma semialg_zero: assumes "A \ semialg_sets 0" shows "A = {[]} \ A = {}" using assms unfolding semialg_sets_def cartesian_power_def proof- assume A0: " A \ gen_boolean_algebra (carrier (RDirProd_list (R_list 0 Q\<^sub>p))) (basic_semialgs 0)" show " A = {[]} \ A = {}" proof- have "A \ {[]} \ A = {}" proof assume A1: "A \ {[]}" show "A = {}" proof- have "(R_list 0 Q\<^sub>p) = []" by simp then have "(carrier (RDirProd_list (R_list 0 Q\<^sub>p))) = {[]}" using RDirProd_list_nil by simp then show ?thesis using A0 A1 by (metis gen_boolean_algebra_subset subset_singletonD) qed qed then show ?thesis by linarith qed qed lemma basic_semialg_is_semialg: assumes "is_basic_semialg n A" shows "A \ semialg_sets n" by (metis (no_types, lifting) assms gen_boolean_algebra.simps inf_absorb1 is_basic_semialg_def mem_Collect_eq basic_semialg_set_def semialg_sets_def subsetI) lemma basic_semialg_is_semialg': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "A \ semialg_sets n" using assms basic_semialg_is_semialg is_basic_semialg_def by blast definition is_semialgebraic where "is_semialgebraic n S = (S \ semialg_sets n)" lemma is_semialgebraicE: assumes "is_semialgebraic n S" shows "S \ semialg_sets n" using assms is_semialgebraic_def by blast lemma is_semialgebraic_closed: assumes "is_semialgebraic n S" shows "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using is_semialgebraicE[of n S] unfolding semialg_sets_def using assms gen_boolean_algebra_subset is_semialgebraicE semialg_sets_def by blast lemma is_semialgebraicI: assumes "S \ semialg_sets n" shows "is_semialgebraic n S" by (simp add: assms is_semialgebraic_def) lemma basic_semialg_is_semialgebraic: assumes "is_basic_semialg n A" shows "is_semialgebraic n A" using assms basic_semialg_is_semialg is_semialgebraicI by blast lemma basic_semialg_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "is_semialgebraic n A" using assms(1) assms(2) assms(3) basic_semialg_is_semialg' is_semialgebraicI by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Algebraic Sets over $p$-adic Fields\ (********************************************************************************************) (********************************************************************************************) lemma p_times_square_not_square: assumes "a \ nonzero Q\<^sub>p" shows "\ \ (a [^] (2::nat)) \ P_set (2::nat)" proof assume A: "\ \ (a[^](2::nat)) \ P_set (2::nat)" then have "\ \ (a[^](2::nat)) \ nonzero Q\<^sub>p" unfolding P_set_def by blast then obtain b where b_def: "b \ carrier Q\<^sub>p \ b[^](2::nat) = \ \ (a[^](2::nat))" using A P_set_def by blast have "b \ nonzero Q\<^sub>p" apply(rule ccontr) using b_def assms by (metis A P_set_nonzero'(1) Qp.nat_pow_zero not_nonzero_Qp zero_neq_numeral) then have LHS: "ord (b[^](2::nat)) = 2* (ord b)" using nonzero_nat_pow_ord by presburger have "ord( \ \ (a[^](2::nat))) = 1 + 2* ord a" using assms nonzero_nat_pow_ord Qp_nat_pow_nonzero ord_mult ord_p p_nonzero by presburger then show False using b_def LHS by presburger qed lemma p_times_square_not_square': assumes "a \ carrier Q\<^sub>p" shows "\ \ (a [^] (2::nat)) = \ \ a = \" by (metis Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero assms p_nonzero) lemma zero_set_semialg_set: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev q a = \ \( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev q a) [^] (2::nat)) = y[^](2::nat)) " proof show "Qp_ev q a = \ \ \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" proof- assume "Qp_ev q a = \" then have "\ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" by (metis Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null zero_neq_numeral) then have "\ \ carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" using Qp.cring_axioms cring.cring_simprules(2) by blast then show "\y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" by blast qed show " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat)) \ Qp_ev q a = \" proof- assume A: " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat))" then obtain b where b_def: "b\carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (b[^](2::nat))" by blast show "Qp_ev q a = \" proof(rule ccontr) assume " Qp_ev q a \ \" then have " Qp_ev q a \ nonzero Q\<^sub>p" using assms eval_at_point_closed[of a n q] nonzero_def proof - have "Qp_ev q a \ carrier Q\<^sub>p" using \\a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>); q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ \ Qp_ev q a \ carrier Q\<^sub>p\ \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ by fastforce then have "Qp_ev q a \ {r \ carrier Q\<^sub>p. r \ \}" using \Qp_ev q a \ \\ by force then show ?thesis by (metis nonzero_def ) qed then have "\ \ (Qp_ev q a[^](2::nat)) \ nonzero Q\<^sub>p" by (metis Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero not_nonzero_Qp p_nonzero p_times_square_not_square') then have "\ \ (Qp_ev q a[^](2::nat)) \ P_set (2::nat)" using b_def unfolding P_set_def by blast then show False using \Qp_ev q a \ nonzero Q\<^sub>p\ p_times_square_not_square by blast qed qed qed lemma alg_as_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" shows "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" proof have 00: "\x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ Qp_ev q x = \ \ (Qp_ev P x) [^] (2::nat)" using assms eval_at_point_smult MP.nat_pow_closed Qp.int_inc_closed eval_at_point_nat_pow by presburger show "V\<^bsub>Q\<^sub>p\<^esub> n P \ basic_semialg_set n 2 q" proof fix x assume A: "x \ V\<^bsub>Q\<^sub>p\<^esub> n P " show "x \ basic_semialg_set n (2::nat) q " proof- have P: "Qp_ev P x = \" using A zero_setE(2) by blast have "Qp_ev q x = \" proof- have "Qp_ev q x = \ \ (Qp_ev (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) x)" using assms eval_at_point_smult[of x n "(P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" \] basic_semialg_set_def by (meson A MP.nat_pow_closed Qp.int_inc_closed zero_setE(1)) then show ?thesis by (metis A P Qp.int_inc_closed Qp.integral_iff Qp.nat_pow_zero Qp.zero_closed assms(1) eval_at_point_nat_pow neq0_conv zero_less_numeral zero_setE(1)) qed then have 0: "Qp_ev q x = \ \ Qp_ev q x \ P_set (2::nat)" by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A zero_setE(1) by blast then show ?thesis using 0 basic_semialg_set_def' by (metis (no_types, opaque_lifting) Qp.nat_pow_zero Qp.zero_closed \eval_at_point Q\<^sub>p x q = \\ basic_semialg_set_memI zero_neq_numeral) qed qed show "basic_semialg_set n 2 q \ V\<^bsub>Q\<^sub>p\<^esub> n P" proof fix x assume A: "x \ basic_semialg_set n 2 q" have 0: "\ Qp_ev q x \ P_set 2" proof assume "Qp_ev q x \ P_set 2" then have 0: "Qp_ev q x \ nonzero Q\<^sub>p \ (\y \ carrier Q\<^sub>p . (y[^] (2::nat)) = Qp_ev q x)" using P_set_def by blast have "( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev P x) [^] (2::nat)) = y[^](2::nat))" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ (y[^] (2::nat)) = Qp_ev q x" using 0 by blast have "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_memE(1) by blast then have "Qp_ev q x = \ \ ((Qp_ev P x) [^] (2::nat))" using assms eval_at_point_scalar_mult 00 by blast then have "(y[^] (2::nat)) = \ \ ((Qp_ev P x) [^] (2::nat))" using y_def by blast then show ?thesis using y_def by blast qed then have "Qp_ev P x = \" by (metis (no_types, lifting) A assms(1) basic_semialg_set_def mem_Collect_eq zero_set_semialg_set) then have "Qp_ev q x = \" using assms eval_at_point_smult by (metis "00" A Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null basic_semialg_set_memE(1) zero_neq_numeral) then show False using 0 Qp.not_nonzero_memI by blast qed show " x \ V\<^bsub>Q\<^sub>p\<^esub> n P" apply(rule zero_setI) using A basic_semialg_set_memE(1) apply blast using A 0 00[of x] by (metis assms(1) basic_semialg_set_memE(1) basic_semialg_set_memE(2) zero_set_semialg_set) qed qed lemma is_zero_set_imp_basic_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_basic_semialg n S" unfolding is_basic_semialg_def proof- obtain q where q_def: "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" by blast have 0: "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" using alg_as_semialg[of P n q] q_def assms(1) by linarith have "(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms(1) by blast then have "\ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms q_def Qp.int_inc_closed local.smult_closed by blast then have 1: "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" by (metis q_def ) then show "\m. m \ 0 \ (\P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). S = basic_semialg_set n m P)" using 0 assms by (metis zero_neq_numeral) qed lemma is_zero_set_imp_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_semialgebraic n S" using assms(1) assms(2) basic_semialg_is_semialg is_semialgebraicI is_zero_set_imp_basic_semialg by blast text\Algebraic sets are semialgebraic\ lemma is_algebraic_imp_is_semialg: assumes "is_algebraic Q\<^sub>p n S" shows "is_semialgebraic n S" proof(rule is_semialgebraicI) obtain ps where ps_def: "finite ps \ ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ S = affine_alg_set Q\<^sub>p n ps" using is_algebraicE by (metis assms) have "ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n ps \ semialg_sets n" apply(rule finite.induct[of ps]) apply (simp add: ps_def) using affine_alg_set_empty[of n] apply (simp add: carrier_is_semialg) proof fix A a assume IH: "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n A \ semialg_sets n" assume P: "insert a A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" have "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using P by blast then show "affine_alg_set Q\<^sub>p n (insert a A) \ semialg_sets n" using IH P semialg_intersect[of "affine_alg_set Q\<^sub>p n A" n "affine_alg_set Q\<^sub>p n {a}" ] is_zero_set_imp_semialg affine_alg_set_insert[of n a A] by (metis Int_commute affine_alg_set_singleton insert_subset is_semialgebraicE) qed then show "S \ semialg_sets n" using ps_def by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\Basic Lemmas about the Semialgebraic Predicate\ (********************************************************************************************) (********************************************************************************************) text\Finite and cofinite sets are semialgebraic\ lemma finite_is_semialg: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "finite F" shows "is_semialgebraic n F" using Qp.finite_sets_are_algebraic is_algebraic_imp_is_semialg[of n F] assms(1) assms(2) by blast definition is_cofinite where "is_cofinite n F = finite (ring_pow_comp Q\<^sub>p n F)" lemma is_cofiniteE: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "finite (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)" using assms(2) is_cofinite_def by (simp add: ring_pow_comp_def) lemma complement_is_semialg: assumes "is_semialgebraic n F" shows "is_semialgebraic n ((carrier (Q\<^sub>p\<^bsup>n\<^esup>)) - F)" using assms is_semialgebraic_def semialg_complement by blast lemma cofinite_is_semialgebraic: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "is_semialgebraic n F" using assms ring_pow_comp_inv[of F Q\<^sub>p n] complement_is_semialg[of n "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] finite_is_semialg[of "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] is_cofiniteE[of F] by (simp add: ring_pow_comp_def) lemma diff_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A - B)" apply(rule is_semialgebraicI) using assms unfolding semialg_sets_def using gen_boolean_algebra_diff is_semialgebraicE semialg_sets_def by blast lemma intersection_is_semialg: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_intersect by blast lemma union_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_union by blast lemma carrier_is_semialgebraic: "is_semialgebraic n (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using carrier_is_semialg by (simp add: carrier_is_semialg is_semialgebraic_def) lemma empty_is_semialgebraic: "is_semialgebraic n {}" by (simp add: empty_set_is_semialg is_semialgebraic_def) (********************************************************************************************) (********************************************************************************************) subsubsection\One-Dimensional Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition one_var_semialg where "one_var_semialg S = ((to_R1 ` S) \ (semialg_sets 1))" definition univ_basic_semialg_set where "univ_basic_semialg_set (m::nat) P = {a \ carrier Q\<^sub>p. (\y \ carrier Q\<^sub>p. (P \ a = (y[^]m)))}" text\Equivalence of univ\_basic\_semialg\_sets and semialgebraic subsets of $\mathbb{Q}^1$ \ lemma univ_basic_semialg_set_to_semialg_set: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "to_R1 ` (univ_basic_semialg_set m P) = basic_semialg_set 1 m (from_Qp_x P)" proof show "(\a. [a]) ` univ_basic_semialg_set m P \ basic_semialg_set 1 m (from_Qp_x P)" proof fix x assume A: "x \ (\a. [a]) ` univ_basic_semialg_set m P" then obtain b y where by_def:"b \ carrier Q\<^sub>p \ y \ carrier Q\<^sub>p \ (P \ b) = (y[^]m) \ x = [b]" unfolding univ_basic_semialg_set_def by blast then have "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A Qp.to_R1_closed[of b] unfolding univ_basic_semialg_set_def by blast then show "x \ basic_semialg_set 1 m (from_Qp_x P)" using by_def Qp_x_Qp_poly_eval assms unfolding basic_semialg_set_def by blast qed show "basic_semialg_set 1 m (from_Qp_x P) \ (\a. [a]) ` univ_basic_semialg_set m P" proof fix x assume A: "x \ basic_semialg_set 1 m (from_Qp_x P)" then obtain b where b_def: "b \ carrier Q\<^sub>p \ x = [b]" unfolding basic_semialg_set_def by (metis (mono_tags, lifting) mem_Collect_eq Qp.to_R1_to_R Qp.to_R_pow_closed) obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev (from_Qp_x P) [b] = (y[^]m))" using A b_def unfolding basic_semialg_set_def by blast have " P \ b = (y[^]m)" using assms y_def b_def Qp_x_Qp_poly_eval by blast then show " x \ (\a. [a]) ` univ_basic_semialg_set m P" using y_def b_def unfolding basic_semialg_set_def univ_basic_semialg_set_def by blast qed qed definition is_univ_semialgebraic where "is_univ_semialgebraic S = (S \ carrier Q\<^sub>p \ is_semialgebraic 1 (to_R1 ` S))" lemma is_univ_semialgebraicE: assumes "is_univ_semialgebraic S" shows "is_semialgebraic 1 (to_R1 ` S)" using assms is_univ_semialgebraic_def by blast lemma is_univ_semialgebraicI: assumes "is_semialgebraic 1 (to_R1 ` S)" shows "is_univ_semialgebraic S" proof- have "S \ carrier Q\<^sub>p" proof fix x assume "x \ S" then have "(to_R1 x) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms by (smt Collect_mono_iff gen_boolean_algebra_subset image_def is_semialgebraicE mem_Collect_eq semialg_sets_def Qp.to_R1_carrier) then show "x \ carrier Q\<^sub>p" using assms by (metis nth_Cons_0 Qp.to_R_pow_closed) qed then show ?thesis using assms unfolding is_univ_semialgebraic_def by blast qed lemma univ_basic_semialg_set_is_univ_semialgebraic: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "is_univ_semialgebraic (univ_basic_semialg_set m P)" using assms by (metis (mono_tags, lifting) basic_semialg_is_semialgebraic' from_Qp_x_closed is_univ_semialgebraic_def mem_Collect_eq subsetI univ_basic_semialg_set_def univ_basic_semialg_set_to_semialg_set) lemma intersection_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms intersection_is_semialg[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis le_infI1 Qp.to_R1_intersection) lemma union_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms union_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis Un_subset_iff image_Un) lemma diff_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A - B)" using assms diff_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (smt Diff_subset subset_trans Qp.to_R1_diff) lemma finite_is_univ_semialgebraic: assumes "A \ carrier Q\<^sub>p" assumes "finite A" shows "is_univ_semialgebraic A" using assms finite_is_semialg[of "((\a. [a]) ` A)" ] to_R1_finite[of A] unfolding is_univ_semialgebraic_def by (metis Qp.to_R1_carrier Qp.to_R1_subset) (********************************************************************************************) (********************************************************************************************) subsubsection\Defining the $p$-adic Valuation Semialgebraically\ (********************************************************************************************) (********************************************************************************************) lemma Qp_square_root_criterion0: assumes "p \ 2" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "val a \ val b" assumes "a \ \" assumes "b \ \" assumes "val a \ 0" shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" proof- have 0: "(to_Zp a) \ carrier Z\<^sub>p" using assms(2) to_Zp_closed by blast have 1: "(to_Zp b) \ carrier Z\<^sub>p" using assms(3) to_Zp_closed by blast have 2: "a \ \\<^sub>p" using val_ring_val_criterion assms(2) assms(5) assms(7) by blast have 3: "b \ \\<^sub>p" using assms val_ring_val_criterion[of b] dual_order.trans by blast have 4: "val_Zp (to_Zp b) = val b" using 3 Zp_def \_def padic_fields.to_Zp_val padic_fields_axioms by blast have 5: "val_Zp (to_Zp a) = val a" using Q\<^sub>p_def Zp_def assms(2) assms(7) padic_fields.Qp_val_ringI padic_fields.to_Zp_val padic_fields_axioms by blast have "\y \ carrier Z\<^sub>p. (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

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

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

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

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

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

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

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

b) \ nonzero Q\<^sub>p" using B by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) inv_in_frac(3)) then have "val a \ val b" using F0 F1 A by blast then have "val (a \
b) \ 0" using F0 F1 val_fract assms(1) local.eint_minus_ineq' by presburger obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat)) = (y[^](2::nat))" using Qp_val_ring_alt_def0 by (meson B False Qp.integral Qp.nonzero_closed \(a \
b) \ nonzero Q\<^sub>p\ \0 \ val (a \
b)\ assms(1) assms(2) inv_in_frac(1) inv_in_frac(2) ord_nonneg val_ringI) then have "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = (b[^](4::nat)) \ (y[^](2::nat))" by presburger then have F2: "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis Qp.nat_pow_pow assms(2) mult_2_right numeral_Bit0) have F3: "((b[^](4::nat)) \ \) \\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat))[^] (2::nat)) \ (y[^](2::nat))" proof- have 0: "(\[^](3::nat)) \ (a \
b[^](4::nat)) \ carrier Q\<^sub>p " proof- have "(a \
b[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed Qp_nat_pow_nonzero by presburger then show ?thesis by (meson Qp.cring_axioms cring.cring_simprules(5) p_natpow_closed(1)) qed have 1: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then show ?thesis using 0 F2 ring.ring_simprules(23)[of Q\<^sub>p "\" "(\[^](3::nat))\ ((a \
b)[^](4::nat))" "(b[^](4::nat))"] Qp.cring_axioms Qp.nonzero_mult_closed Qp.ring_axioms Qp_nat_pow_nonzero \(a \
b) \ nonzero Q\<^sub>p\ p_nonzero by blast qed have F4: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then have "((b[^](4::nat)) \ \) = (b[^](4::nat))" using Qp.r_one by blast then have F5: "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F3 by presburger have "((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat)))" proof- have 0: "(b[^](4::nat)) \ carrier Q\<^sub>p" using F4 by blast have 1: "(\[^](3::nat)) \ carrier Q\<^sub>p" by blast have 2: "((a \
b)[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed by blast show ?thesis using 0 1 2 monoid.m_assoc[of Q\<^sub>p] comm_monoid.m_comm[of Q\<^sub>p] using Qp.m_lcomm by presburger qed then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F5 by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b \(a \
b))[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F0 Qp.nat_pow_distrib assms(1) assms(2) fract_closed by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis F0 assms(1) local.fract_cancel_right) then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" using Qp.nat_pow_closed Qp.nat_pow_distrib assms(2) y_def by blast then have "((b[^](2::nat))\ y) \ carrier Q\<^sub>p \ (b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" by (meson Qp.cring_axioms Qp.nat_pow_closed assms(2) cring.cring_simprules(5) y_def) then show ?thesis by blast qed qed qed show "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \ val a \ val b" proof- assume A: "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" show "val a \ val b" proof(cases "a = \") case True then show ?thesis using eint_ord_code(3) local.val_zero by presburger next case False have "b \ \" proof(rule ccontr) assume "\ b \ \" then have "\y \ carrier Q\<^sub>p. (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.nat_pow_closed Qp.nonzero_memE(2) Qp_nonzero_nat_pow assms(1) assms(2) local.monom_term_car not_nonzero_Qp p_natpow_closed(1) zero_less_numeral) then obtain y where y_def: "y \ carrier Q\<^sub>p \ (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" by blast have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" proof- have 00: "(\[^](3::nat)) \ nonzero Q\<^sub>p" using Qp_nat_pow_nonzero p_nonzero by blast have 01: "(a[^](4::nat)) \ nonzero Q\<^sub>p" using False Qp_nat_pow_nonzero assms(1) not_nonzero_Qp Qp.nonzero_memI by presburger then show ?thesis using ord_mult[of "\[^](3::nat)" "a[^](4::nat)"] by (metis (no_types, lifting) "00" False assms(1) nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral ord_p_pow_nat) qed have 1: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 2* (ord y)" proof- have "y \ \" proof(rule ccontr) assume " \ y \ \" then have "(\[^](3::nat))\ (a[^](4::nat)) = \" using y_def Qp.cring_axioms cring_def pos2 ring.pow_zero by blast then show False by (metis False Qp.integral Qp.nat_pow_closed Qp.nonzero_pow_nonzero Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) p_natpow_closed(1) p_nonzero) qed then show ?thesis using y_def by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) qed then show False using 0 by presburger qed then have F0: "b \ nonzero Q\<^sub>p" using assms(2) not_nonzero_Qp by metis have F1: "a \ nonzero Q\<^sub>p" using False assms(1) not_nonzero_Qp by metis obtain y where y_def: "y \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by blast show ?thesis proof(rule ccontr) assume " \ val a \ val b " then have F2: "ord a < ord b" using F0 F1 assms by (metis False \b \ \\ eint_ord_simps(1) leI val_def) have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using F0 ord_mult F1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_p_pow_nat p_natpow_closed(2) by presburger have 1: " ord (b[^](4::nat)) = 4* ord b" using F0 nonzero_nat_pow_ord by presburger have 2: "(4 * (ord b)) > 4 * (ord a)" using F2 by linarith have 3: "(4 * (ord b)) \ 3 + 4* ord a" proof(rule ccontr) assume "\ (4 * (ord b)) \ 3 + 4* ord a" then have "(4 * (ord b)) > 3 + 4* ord a" by linarith then have 30: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using "0" "1" F0 F1 Qp_nat_pow_nonzero Qp.nat_pow_closed assms(1) monom_term_car not_nonzero_Qp ord_ultrametric_noteq p_natpow_closed(1) p_nonzero by (metis Qp.integral) have "y \ nonzero Q\<^sub>p" proof(rule ccontr) assume A: "y \ nonzero Q\<^sub>p" then have "y = \" using y_def Qp.nonzero_memI by blast then have "b [^] 4 \ \ [^] 3 \ a [^] 4 = \" by (smt "0" "1" A F0 False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_mult_closed Qp.nonzero_pow_nonzero Qp.pow_zero assms(1) diff_ord_nonzero not_nonzero_Qp p_nonzero pos2 y_def) then show False by (smt "0" "1" A F0 F1 Qp.integral Qp.nat_pow_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero assms(1) diff_ord_nonzero not_nonzero_Qp p_natpow_closed(1) p_nonzero y_def) qed then have 31: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 2* ord y" using nonzero_nat_pow_ord y_def by presburger then show False using 30 by presburger qed show False using 2 3 by presburger qed qed qed qed text\The polynomial in two variables which semialgebraically defines the valuation relation\ definition Qp_val_poly where "Qp_val_poly = (pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" lemma Qp_val_poly_closed: "Qp_val_poly \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" proof- have "(pvar Q\<^sub>p 1) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed one_less_numeral_iff semiring_norm(76) by blast then have 0: "(pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using ring.Pring_is_ring[of Q\<^sub>p "{0::nat..2-1}"] monoid.nat_pow_closed[of "coord_ring Q\<^sub>p 2"] Qp.cring_axioms cring.axioms(1) ring.Pring_is_monoid by blast have 1: "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed pos2 by blast have 2: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using 1 local.smult_closed p_natpow_closed(1) by blast then show ?thesis unfolding Qp_val_poly_def using 0 by blast qed lemma Qp_val_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "Qp_ev Qp_val_poly [a, b] = (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))" proof- have 0: "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" proof(rule cartesian_power_car_memI) show "length [a, b] = 2" by simp have "set [a,b] = {a,b}" by auto then show "set [a, b] \ carrier Q\<^sub>p" using assms by (simp add: \a \ carrier Q\<^sub>p\ \b \ carrier Q\<^sub>p\) qed obtain f where f_def: "f = ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))" by blast obtain g where g_def: "g = (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" by blast have 1: "Qp_val_poly = f \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> g" unfolding Qp_val_poly_def using f_def g_def by blast have 1: "Qp_ev (pvar Q\<^sub>p (0::nat)) [a,b] = a" using eval_pvar by (metis \[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)\ nth_Cons_0 pos2) have 2: "Qp_ev (pvar Q\<^sub>p (1::nat)) [a,b] = b" using eval_pvar by (metis (no_types, lifting) "0" One_nat_def add_diff_cancel_right' assms(2) cartesian_power_car_memE gr_zeroI less_numeral_extra(1) less_numeral_extra(4) list.size(4) nth_Cons_pos Qp.to_R1_closed Qp.to_R_to_R1 zero_less_diff) have 3: "Qp_ev ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (b[^](4::nat))" by (metis "0" "2" eval_at_point_nat_pow local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 4: "Qp_ev ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (a[^](4::nat))" using "0" "1" eval_at_point_nat_pow local.pvar_closed pos2 by presburger then have 5: "Qp_ev (poly_scalar_mult Q\<^sub>p (\[^](3::nat)) ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))) [a,b] = (\[^](3::nat))\ (a[^](4::nat))" using eval_at_point_smult[of "[a,b]" 2 "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)" "\[^](3::nat)" ] 2 by (metis "0" MP.nat_pow_closed eval_at_point_scalar_mult local.pvar_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis proof- have 00: "[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" by (simp add: "0") have 01: " pvar Q\<^sub>p 1 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 02: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (pvar Q\<^sub>p 0 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat)) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed local.smult_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis unfolding Qp_val_poly_def using 00 01 02 by (metis (no_types, lifting) "3" "4" MP.nat_pow_closed eval_at_point_add eval_at_point_smult local.pvar_closed p_natpow_closed(1) zero_less_numeral) qed qed lemma Qp_2I: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule cartesian_power_car_memI) using assms apply (simp add: assms(1) assms(2)) using assms by (simp add: assms(1) assms(2)) lemma pair_id: assumes "length as = 2" shows "as = [as!0, as!1]" using assms by (smt One_nat_def diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons' nth_equalityI numeral_2_eq_2) lemma Qp_val_semialg: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "val b \ val a \ [a,b] \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val a \ val b \ [a, b] \ basic_semialg_set 2 2 Qp_val_poly" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def by (metis (mono_tags, lifting) assms(1) assms(2) mem_Collect_eq) show "[a, b] \ basic_semialg_set 2 2 Qp_val_poly \ val a \ val b" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def using assms(1) assms(2) by blast qed definition val_relation_set where "val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as!1) \ val (as!0)}" lemma val_relation_setE: assumes "as \ val_relation_set" shows "as!0 \ carrier Q\<^sub>p \ as!1 \ carrier Q\<^sub>p \ as = [as!0,as!1] \ val (as!1) \ val (as!0)" using assms unfolding val_relation_set_def by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff pair_id pos2 semiring_norm(76)) lemma val_relation_setI: assumes "as!0 \ carrier Q\<^sub>p" assumes "as!1 \ carrier Q\<^sub>p" assumes "length as = 2" assumes "val (as!1) \ val(as!0)" shows "as \ val_relation_set" unfolding val_relation_set_def using assms Qp_2I[of "as!0" "as!1"] by (metis (no_types, lifting) mem_Collect_eq pair_id) lemma val_relation_semialg: "val_relation_set = basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val_relation_set \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof fix as assume A: "as \ val_relation_set" have 0: "length as = 2" unfolding val_relation_set_def by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq val_relation_set_def) have 1: "as = [as ! 0, as ! 1]" by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq pair_id val_relation_set_def) show "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" using A 1 val_relation_setE[of as] Qp_val_semialg[of "as!0" "as!1"] by presburger qed show "basic_semialg_set 2 (2::nat) Qp_val_poly \ val_relation_set" proof fix as assume "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" then show "as \ val_relation_set" using val_relation_setI[of as] by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff Qp_val_semialg basic_semialg_set_def val_relation_set_def padic_fields_axioms pair_id pos2 semiring_norm(76)) qed qed lemma val_relation_is_semialgebraic: "is_semialgebraic 2 val_relation_set" proof - have "{rs \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (rs ! 0) \ val (rs ! 1)} = basic_semialg_set (Suc 1) (Suc 1) Qp_val_poly" using Suc_1 val_relation_semialg val_relation_set_def by presburger then show ?thesis by (metis (no_types) Qp_val_poly_closed Suc_1 basic_semialg_is_semialgebraic' val_relation_set_def zero_neq_numeral) qed lemma Qp_val_ring_is_semialg: obtains P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" proof- obtain P where P_def: "P = (\[^](3::nat)) \\<^bsub>Q\<^sub>p_x \<^esub>(X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat) \\<^bsub>Q\<^sub>p_x\<^esub> \\<^bsub>Q\<^sub>p_x\<^esub>" by blast have 0: "P \ carrier Q\<^sub>p_x" proof- have 0: "(X_poly Q\<^sub>p) \ carrier Q\<^sub>p_x" using UPQ.X_closed by blast then show ?thesis using P_def UPQ.P.nat_pow_closed p_natpow_closed(1) by blast qed have 1: "\\<^sub>p = univ_basic_semialg_set 2 P" proof show "\\<^sub>p \ univ_basic_semialg_set 2 P" proof fix x assume A: "x \ \\<^sub>p" show "x \ univ_basic_semialg_set 2 P" proof- have x_car: "x \ carrier Q\<^sub>p" using A val_ring_memE by blast then have "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat)))" using A Qp_val_ring_alt_def[of x] by blast then obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat))" by blast have "y \ carrier Q\<^sub>p \ P \ x = (y[^](2::nat))" proof- have "P \ x = \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat))" proof- have "((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat)) \ carrier Q\<^sub>p_x" using UPQ.monom_closed p_natpow_closed(1) by blast then have "P \ x = (((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using P_def x_car UPQ.to_fun_plus by blast then have 0: "P \ x = (\[^](3::nat)) \(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_smult p_natpow_closed(1) x_car by presburger have "(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) = (x[^](4::nat))" using UPQ.to_fun_X_pow x_car by blast then have "P \ x = (\[^](3::nat)) \(x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using "0" UPQ.to_fun_one x_car by presburger then show ?thesis using y_def Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1) x_car by presburger qed then show ?thesis using y_def by blast qed then show ?thesis unfolding univ_basic_semialg_set_def using x_car by blast qed qed show "univ_basic_semialg_set 2 P \ \\<^sub>p" proof fix x assume A: "x \ univ_basic_semialg_set (2::nat) P" then obtain y where y_def: "y \ carrier Q\<^sub>p \ (P \ x) = (y[^](2::nat))" unfolding univ_basic_semialg_set_def by blast have x_car: "x \ carrier Q\<^sub>p" using A by (metis (no_types, lifting) mem_Collect_eq univ_basic_semialg_set_def) have 0: "(P \ x) = (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using P_def x_car UPQ.UP_one_closed UPQ.monom_closed UPQ.monom_rep_X_pow UPQ.to_fun_monom UPQ.to_fun_one UPQ.to_fun_plus p_natpow_closed(1) by presburger have 1: "y \ carrier Q\<^sub>p \ (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ = (y[^](2::nat))" using "0" y_def by blast then show "x \ \\<^sub>p" using x_car Qp_val_ring_alt_def[of x] y_def by (metis Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1)) qed qed show ?thesis using 0 1 that by blast qed lemma Qp_val_ring_is_univ_semialgebraic: "is_univ_semialgebraic \\<^sub>p" proof- obtain P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" using Qp_val_ring_is_semialg by blast then show ?thesis by (metis univ_basic_semialg_set_is_univ_semialgebraic zero_neq_numeral) qed lemma Qp_val_ring_is_semialgebraic: "is_semialgebraic 1 (to_R1` \\<^sub>p)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Inverse Images of Semialgebraic Sets by Polynomial Maps\ (********************************************************************************************) (********************************************************************************************) lemma basic_semialg_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S = basic_semialg_set k m f" assumes "m \0" shows "poly_map n fs \\<^bsub>n\<^esub> S = basic_semialg_set n m (Qp_poly_comp n fs f)" proof show "poly_map n fs \\<^bsub>n\<^esub> S \ basic_semialg_set n m (Qp_poly_comp n fs f)" proof fix x assume A: "x \ poly_map n fs \\<^bsub>n\<^esub> S" then have 0: "poly_map n fs x \ S" proof - have "\n f. {rs. rs \ S} \ {rs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). \r. r \ carrier Q\<^sub>p \ Qp_ev f rs = (r[^](n::nat))}" by (metis (no_types) Collect_mem_eq \S = basic_semialg_set k m f\ basic_semialg_set_def eq_iff) then show ?thesis using A by blast qed have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A assms by (meson evimage_eq) have "\y \ (carrier Q\<^sub>p). Qp_ev f (poly_map n fs x) = (y[^]m)" using A 0 assms basic_semialg_set_def by blast then have "\y \ (carrier Q\<^sub>p). Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using 1 assms Qp_poly_comp_eval by blast then show "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" using "1" basic_semialg_set_def by blast qed show "basic_semialg_set n m (Qp_poly_comp n fs f) \ poly_map n fs \\<^bsub>n\<^esub> S" proof fix x assume A: "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_def by blast have 1: "(poly_map n fs x) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using "0" poly_map_closed assms(2) assms(3) by blast show "x \ poly_map n fs \\<^bsub>n\<^esub> S" proof- have "\y \ carrier Q\<^sub>p. Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using A basic_semialg_set_def by blast then have 2: "\y \ carrier Q\<^sub>p. Qp_ev f (poly_map n fs x) = (y[^]m)" using assms Qp_poly_comp_eval by (metis (no_types, lifting) A basic_semialg_set_def mem_Collect_eq) have 3: "poly_map n fs x \ S" using assms 0 1 basic_semialg_set_def[of k m f] "2" by blast show ?thesis using "0" "3" by blast qed qed qed lemma basic_semialg_pullback': assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "A \ basic_semialgs k" shows "poly_map n fs \\<^bsub>n\<^esub> A \ (basic_semialgs n)" proof- obtain f m where fm_def: "m \0 \f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>]) \ A = basic_semialg_set k m f" using assms by (metis is_basic_semialg_def mem_Collect_eq) then have "poly_map n fs \\<^bsub>n\<^esub> A = basic_semialg_set n m (Qp_poly_comp n fs f)" using assms basic_semialg_pullback[of f k n fs A m] by linarith then show ?thesis unfolding is_basic_semialg_def by (metis (mono_tags, lifting) assms(1) assms(2) fm_def mem_Collect_eq poly_compose_closed) qed lemma semialg_pullback: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "poly_map n fs \\<^bsub>n\<^esub> S \ semialg_sets n" unfolding semialg_sets_def apply(rule gen_boolean_algebra.induct[of S "(carrier (Q\<^sub>p\<^bsup>k\<^esup>))" "basic_semialgs k"]) using assms semialg_sets_def apply blast apply (metis assms(1) assms(2) carrier_is_semialgebraic evimageI2 extensional_vimage_closed is_semialgebraicE poly_map_closed semialg_sets_def subsetI subset_antisym) apply (metis Int_absorb2 assms(1) assms(2) basic_semialg_is_semialg basic_semialg_is_semialgebraic basic_semialg_pullback' is_semialgebraic_closed mem_Collect_eq semialg_sets_def) apply (metis evimage_Un semialg_sets_def semialg_union) by (metis assms(1) assms(2) carrier_is_semialgebraic diff_is_semialgebraic evimage_Diff extensional_vimage_closed is_semialgebraicE is_semialgebraicI poly_map_closed poly_map_pullbackI semialg_sets_def subsetI subset_antisym) lemma pullback_is_semialg: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "is_semialgebraic n (poly_map n fs \\<^bsub>n\<^esub> S)" using assms(1) assms(2) assms(3) is_semialgebraicI padic_fields_axioms semialg_pullback by blast text\Equality and inequality sets for a pair of polynomials\ definition val_ineq_set where "val_ineq_set n f g = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" lemma poly_map_length : assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = m" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) length_map restrict_apply') lemma val_ineq_set_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g = poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set " proof show "val_ineq_set n f g \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" proof fix x assume "x \ val_ineq_set n f g" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ val (Qp_ev f x) \ val (Qp_ev g x)" by (metis (mono_tags, lifting) mem_Collect_eq val_ineq_set_def) have 1: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" unfolding poly_map_def poly_tuple_eval_def using 0 by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') have 2: "poly_map n [g,f] x \ val_relation_set" apply(rule val_relation_setI) using 1 0 assms apply (metis eval_at_point_closed nth_Cons_0) using 1 0 assms apply (metis One_nat_def eval_at_point_closed diff_Suc_1 less_numeral_extra(1) nth_Cons_pos Qp.to_R_to_R1) using poly_map_length assms 0 apply (metis "1" Qp_2I cartesian_power_car_memE eval_at_point_closed) by (metis "0" "1" One_nat_def nth_Cons_0 nth_Cons_Suc) have 3: "is_poly_tuple n [g, f]" using assms by (smt One_nat_def diff_Suc_1 Qp_is_poly_tupleI length_Suc_conv less_SucE less_one list.size(3) nth_Cons') then show "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" using 0 1 2 by blast qed show "poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set \ val_ineq_set n f g" proof fix x have 0: "is_poly_tuple n [g, f]" using Qp_is_poly_tupleI assms by (metis (no_types, lifting) diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons') assume A: "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ poly_map n [g,f] x \ val_relation_set" using 0 by (meson evimageD extensional_vimage_closed subsetD) have 2: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" by (metis "1" Qp_poly_mapE' length_0_conv poly_map_cons) show "x \ val_ineq_set n f g" using 0 1 2 unfolding val_ineq_set_def val_relation_set_def by (metis (no_types, lifting) "1" list.inject mem_Collect_eq nth_Cons_0 poly_map_apply val_relation_setE) qed qed lemma val_ineq_set_is_semialg: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g \ semialg_sets n" proof- have 0: "val_relation_set \ semialg_sets 2" using val_relation_semialg basic_semialg_is_semialg' by (metis Qp_val_poly_closed zero_neq_numeral) show ?thesis using val_ineq_set_pullback semialg_pullback[of n "[g,f]" 2 "val_relation_set" ] by (metis (no_types, lifting) "0" assms(1) assms(2) diff_Suc_1 Qp_is_poly_tupleI length_Cons less_Suc0 less_SucE list.size(3) nth_Cons_0 nth_Cons_pos numeral_2_eq_2 zero_neq_numeral) qed lemma val_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n (val_ineq_set n f g)" using assms(1) assms(2) is_semialgebraicI val_ineq_set_is_semialg by blast lemma val_ineq_setI: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ (val_ineq_set n f g)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "val (Qp_ev f x) \ val (Qp_ev g x)" using assms unfolding val_ineq_set_def apply blast using assms unfolding val_ineq_set_def by blast lemma val_ineq_setE: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "val (Qp_ev f x) \ val (Qp_ev g x)" shows "x \ (val_ineq_set n f g)" using assms unfolding val_ineq_set_def by blast lemma val_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast lemma val_eq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" proof- have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 1: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 2: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} \ {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule equalityI, rule subsetI , rule IntI) unfolding mem_Collect_eq using le_less apply blast apply (metis order_refl) apply(rule subsetI, erule IntE) unfolding mem_Collect_eq by (meson less_le_trans not_less_iff_gr_or_eq) show ?thesis unfolding 2 apply(rule intersection_is_semialg) using 0 apply blast using 1 by blast qed lemma equalityI'': assumes "\x. A x \ B x" assumes "\x. B x \ A x" shows "{x. A x} = {x. B x}" using assms by blast lemma val_strict_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" proof- have 0: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} - {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule equalityI', rule DiffI) unfolding le_less mem_Collect_eq apply blast unfolding mem_Collect_eq using neq_iff apply blast apply(erule DiffE) unfolding mem_Collect_eq by blast show ?thesis unfolding 0 apply(rule diff_is_semialgebraic) using assms val_ineq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def apply blast using assms val_eq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def by blast qed lemma constant_poly_val_exists: shows "\g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" proof- obtain a where a_def: "a \ carrier Q\<^sub>p \ val a = c" by (meson Qp.minus_closed Qp.nonzero_closed dist_nonempty' p_nonzero) obtain g where g_def: "g = coord_const a" by blast show ?thesis using a_def g_def Qp_to_IP_car by (metis (no_types, opaque_lifting) Qp_to_IP_car a_def eval_at_point_const g_def le_less subset_iff) qed lemma val_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" apply(rule val_ineq_set_is_semialgebraic') using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_ineq_set_is_semialgebraic''': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule val_ineq_set_is_semialgebraic') using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_eq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule val_eq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" apply(rule equalityI'') using g_def apply fastforce using g_def by metis show ?thesis using 0 unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" apply(rule equalityI'') using assms g_def apply fastforce using assms g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma(in cring) R1_memE: assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "x = [(hd x)]" using assms cartesian_power_car_memE by (metis diff_is_0_eq' hd_conv_nth le_eq_less_or_eq length_0_conv length_tl list.exhaust list.sel(3) normalize.cases nth_Cons_0 zero_neq_one) lemma(in cring) R1_memE': assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "hd x \ carrier R" using R1_memE assms cartesian_power_car_memE[of x R 1] cartesian_power_car_memE'[of x R 1 0] by (metis hd_conv_nth less_numeral_extra(1) list.size(3) zero_neq_one) lemma univ_val_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x \ c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c}" apply(rule val_ineq_set_is_semialgebraic'') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c} = to_R1 ` {x \ carrier Q\<^sub>p. val x \ c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a \ c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_strict_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x < c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) < c} = to_R1 ` {x \ carrier Q\<^sub>p. val x < c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a < c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_eq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x = c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c}" apply(rule val_eq_set_is_semialgebraic') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c} = to_R1 ` {x \ carrier Q\<^sub>p. val x = c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" using A 0 unfolding mem_Collect_eq 1 by blast qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a = c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\One Dimensional $p$-adic Balls are Semialgebraic\ (********************************************************************************************) (********************************************************************************************) lemma coord_ring_one_def: "Pring Q\<^sub>p {(0::nat)} = (Q\<^sub>p[\\<^bsub>1\<^esub>])" proof- have "{(0::nat)} = {..<1}" by auto thus ?thesis unfolding coord_ring_def by auto qed lemma times_p_pow_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]n \ a" shows "val b = val a + n" using val_mult[of "\[^]n" a] assms unfolding assms(2) val_p_int_pow by (metis add.commute p_intpow_closed(1)) lemma times_p_pow_neg_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]-n \ a" shows "val b = val a - n" by (metis Qp.m_comm Qp_int_pow_nonzero assms(1) assms(2) p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) lemma eint_minus_int_pos: assumes "a - eint n \ 0" shows "a \ n" using assms apply(induction a) apply (metis diff_ge_0_iff_ge eint_ord_simps(1) idiff_eint_eint zero_eint_def) by simp text\\p\-adic balls as pullbacks of polynomial maps\ lemma balls_as_pullbacks: assumes "c \ carrier Q\<^sub>p" shows "\P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]). to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p)" proof- obtain P0 where P0_def: "P0 = (to_poly (\[^](-n))) \\<^bsub>Q\<^sub>p_x\<^esub>((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c)" by blast have 0: "P0 \ carrier Q\<^sub>p_x" proof- have P0: "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have P1: "(to_poly (\[^](-n))) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast then show ?thesis using P0_def P0 P1 by blast qed have 1: "\x. x \ carrier Q\<^sub>p \ P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" proof- fix x assume A: "x \ carrier Q\<^sub>p" have P0: "(to_poly (\[^](-n))) \ x = (\[^](-n))" using A UPQ.to_fun_to_poly p_intpow_closed(1) by blast have P1: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (x \\<^bsub>Q\<^sub>p\<^esub> c)" by (metis A UPQ.to_fun_X_minus X_poly_minus_def assms) have P2: "to_poly (\[^](-n)) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast have P3: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = to_poly (\[^]- n) \ x \ (((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x)" using A P0_def P0 P1 P2 P3 to_fun_mult[of "to_poly (\[^](-n))" "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c" x] UPQ.to_fun_mult by blast then have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c) " by (metis P0 P1) then show "P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" using P0_def by metis qed have 2: " (\a. [a]) ` B\<^bsub>n\<^esub>[c] = poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof show "(\a. [a]) ` B\<^bsub>n\<^esub>[c] \ poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof fix x assume A: "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" then obtain a where a_def: "x = [a] \ a \ B\<^bsub>n\<^esub>[c]" by blast have P0: "P0 \ a \ \\<^sub>p" proof- have "B\<^bsub>n\<^esub>[c] \ carrier Q\<^sub>p" using c_ball_in_Qp by blast hence a_closed: "a \ carrier Q\<^sub>p" using a_def by blast have P0: "P0 \ a = (\[^](-n)) \ (a \ c)" using 1 a_def c_ballE(1) by blast then have P1: "val (P0 \ a) = val (\[^](-n)) + val (a \ c)" using val_mult[of "\[^]-n" "a \ c"] a_closed assms Qp.minus_closed p_intpow_closed(1) by presburger then have P2: "val (P0 \ a) = val (a \\<^bsub>Q\<^sub>p\<^esub> c) - n" by (metis P0 Qp.m_comm Qp.minus_closed Qp_int_pow_nonzero assms local.a_closed p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) have P3: "val (a \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using a_def c_ballE(2) by blast then have "val (P0 \ a) \ -n + n" using P2 by (metis add.commute diff_conv_add_uminus diff_self local.eint_minus_ineq' zero_eint_def) then have P4: "val (P0 \ a) \ 0" by (metis add.commute add.right_inverse zero_eint_def) have P5: "P0 \ a \ carrier Q\<^sub>p" using "0" UPQ.to_fun_closed local.a_closed by blast then show ?thesis using P4 using val_ring_val_criterion by blast qed have "poly_map 1 [from_Qp_x P0] x = [Qp_ev (from_Qp_x P0) [a]] " using a_def poly_map_def[of 1 "[from_Qp_x P0]"] poly_tuple_eval_def[of ] by (metis Qp_poly_mapE' c_ballE(1) length_0_conv poly_map_cons Qp.to_R1_closed) then have "poly_map 1 [from_Qp_x P0] x = [P0 \ a] " using Qp_x_Qp_poly_eval[of P0 a] by (metis "0" a_def c_ballE(1)) then have P1: "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using P0 by blast have P2: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A c_ballE(1) Qp.to_R1_closed by blast have P3: "is_poly_tuple 1 [from_Qp_x P0]" apply(rule Qp_is_poly_tupleI) by (metis "0" Qp_is_poly_tupleI from_Qp_x_closed gr_implies_not0 is_poly_tupleE is_poly_tuple_Cons list.size(3) zero_neq_one) show "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p" using P3 P2 P1 unfolding evimage_def poly_map_def by blast qed have 20: "is_poly_tuple 1 [from_Qp_x P0]" using 0 UP_to_IP_closed[of P0 "0::nat"] unfolding is_poly_tuple_def by (metis (no_types, lifting) empty_set from_Qp_x_closed list.simps(15) singletonD subset_code(1)) show "poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" proof fix x assume A: "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" have A0: "(\a. [a]) ` \\<^sub>p \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def Qp.to_R1_car_subset Qp_val_ring_is_semialgebraic is_semialgebraic_closed by presburger have "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using A0 A 20 by blast then obtain a where a_def: "a \ \\<^sub>p \ (poly_map 1 [from_Qp_x P0] x) = [a]" by blast have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A by (meson evimage_eq) then obtain y where y_def: "x = [y] \ y \ carrier Q\<^sub>p" using A by (metis Qp.to_R1_to_R Qp.to_R_pow_closed) have "(poly_map 1 [from_Qp_x P0] x) = [(Qp_ev (from_Qp_x P0) [y])]" unfolding poly_map_def poly_tuple_eval_def using x_closed by (smt "20" One_nat_def length_Suc_conv list.size(3) nth_Cons_0 nth_map poly_tuple_eval_closed poly_tuple_eval_def restrict_apply' Qp.to_R1_to_R y_def zero_less_Suc) then have "(poly_map 1 [from_Qp_x P0] x) = [P0 \ y]" by (metis "0" Qp_x_Qp_poly_eval y_def) then have "[a] = [P0 \ y]" using a_def by presburger then have A1: "a = (\[^](-n)) \ (y \\<^bsub>Q\<^sub>p\<^esub> c)" using 1[of y] y_def by blast have "y \ B\<^bsub>n\<^esub>[c]" proof- have B0: "val a = val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n" using A1 y_def Qp.minus_closed assms times_p_pow_neg_val by blast have B1: "val a \ 0" using a_def val_ring_memE by blast then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n \ 0" using B0 by metis then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using eint_minus_int_pos by blast then show "y \ B\<^bsub>n\<^esub>[c]" using c_ballI y_def by blast qed then show "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" using y_def by blast qed qed then show ?thesis by (meson "0" from_Qp_x_closed) qed lemma ball_is_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1` B\<^bsub>n\<^esub>[c])" proof- obtain P where P_def: "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]) \ to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p) " using assms balls_as_pullbacks[of c n] by meson have "is_poly_tuple 1 [P]" using P_def unfolding is_poly_tuple_def by (metis (no_types, opaque_lifting) list.inject list.set_cases neq_Nil_conv subset_code(1)) then show ?thesis using assms P_def pullback_is_semialg[of 1 "[P]" 1 "((\a. [a]) ` \\<^sub>p) "] by (metis (mono_tags, lifting) One_nat_def Qp_val_ring_is_semialgebraic is_semialgebraic_def length_Suc_conv list.distinct(1) list.size(3)) qed lemma ball_is_univ_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (B\<^bsub>n\<^esub>[c])" using assms ball_is_semialgebraic c_ball_in_Qp is_univ_semialgebraic_def by presburger abbreviation Qp_to_R1_set where "Qp_to_R1_set S \ to_R1 ` S" (********************************************************************************************) (********************************************************************************************) subsubsection\Finite Unions and Intersections of Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition are_semialgebraic where "are_semialgebraic n Xs = (\ x. x \ Xs \ is_semialgebraic n x)" lemma are_semialgebraicI: assumes "\x. x \ Xs \ is_semialgebraic n x " shows "are_semialgebraic n Xs" using are_semialgebraic_def assms by blast lemma are_semialgebraicE: assumes "are_semialgebraic n Xs" assumes "x \ Xs" shows "is_semialgebraic n x" using are_semialgebraic_def assms(1) assms(2) by blast definition are_univ_semialgebraic where "are_univ_semialgebraic Xs = (\ x. x \ Xs \ is_univ_semialgebraic x)" lemma are_univ_semialgebraicI: assumes "\x. x \ Xs \ is_univ_semialgebraic x " shows "are_univ_semialgebraic Xs" using are_univ_semialgebraic_def assms by blast lemma are_univ_semialgebraicE: assumes "are_univ_semialgebraic Xs" assumes "x \ Xs" shows "is_univ_semialgebraic x" using are_univ_semialgebraic_def assms(1) assms(2) by blast lemma are_univ_semialgebraic_semialgebraic: assumes "are_univ_semialgebraic Xs" shows "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" apply(rule are_semialgebraicI) using are_univ_semialgebraicE assms image_iff is_univ_semialgebraicE by (metis (no_types, lifting)) lemma to_R1_set_union: "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" using image_iff by blast lemma to_R1_inter: assumes "Xs \ {}" shows "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" proof show "to_R1 ` (\ Xs) \ \ (Qp_to_R1_set ` Xs)" by blast show "\ (Qp_to_R1_set ` Xs) \ to_R1 ` (\ Xs)" proof fix x assume A: "x \ \ (Qp_to_R1_set ` Xs)" then have 0: "\S. S \ Xs \ x \ (Qp_to_R1_set S)" by blast obtain S where "S \ Xs \ x \ (Qp_to_R1_set S)" using assms 0 by blast then obtain b where b_def: "b \ S \ x = [b]" by blast have "b \ (\ Xs)" using "0" b_def by blast then show "x \ to_R1 ` (\ Xs)" using b_def by blast qed qed lemma finite_union_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply (simp add: empty_is_semialgebraic) by (metis Sup_insert insert_subset is_semialgebraicI union_is_semialgebraic) lemma finite_union_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n " shows "is_semialgebraic n (\ Xs)" using assms(1) assms(2) finite_union_is_semialgebraic by blast lemma(in padic_fields) finite_union_is_semialgebraic'': assumes "finite S" assumes "\x. x \ S \ is_semialgebraic m (F x)" shows "is_semialgebraic m (\ x \ S. F x)" using assms finite_union_is_semialgebraic[of "F`S" m] unfolding is_semialgebraic_def by blast lemma finite_union_is_univ_semialgebraic': assumes "finite Xs" assumes "are_univ_semialgebraic Xs" shows "is_univ_semialgebraic (\ Xs)" proof- have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms finite_union_is_semialgebraic'[of "((`) (\a. [a]) ` Xs)"] to_R1_set_union[of Xs] by (metis (no_types, lifting) are_semialgebraicE are_univ_semialgebraic_semialgebraic finite_imageI is_semialgebraicE subsetI) then show ?thesis using is_univ_semialgebraicI by blast qed lemma finite_intersection_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ Xs \{} \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply auto[1] proof fix A::"((nat \ int) \ (nat \ int)) set list set set" fix a assume 0: "finite A" assume 1: "A \ semialg_sets n \ A \ {} \ is_semialgebraic n (\ A) " assume 2: "insert a A \ semialg_sets n \ insert a A \ {}" show "is_semialgebraic n (\ (insert a A))" proof(cases "A = {}") case True then have "insert a A = {a}" by simp then show ?thesis by (metis "2" cInf_singleton insert_subset is_semialgebraicI) next case False then have "A \ semialg_sets n \ A \ {}" using "2" by blast then have "is_semialgebraic n (\ A) " using "1" by linarith then show ?thesis using 0 1 2 intersection_is_semialg by (metis Inf_insert insert_subset is_semialgebraicI) qed qed lemma finite_intersection_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (simp add: assms(1) assms(2) finite_intersection_is_semialgebraic) lemma finite_intersection_is_semialgebraic'': assumes "finite Xs" assumes "are_semialgebraic n Xs \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (meson are_semialgebraicE assms(1) assms(2) finite_intersection_is_semialgebraic' is_semialgebraicE subsetI) lemma finite_intersection_is_univ_semialgebraic: assumes "finite Xs" assumes "are_univ_semialgebraic Xs" assumes "Xs \ {}" shows "is_univ_semialgebraic (\ Xs)" proof- have "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" using are_univ_semialgebraic_semialgebraic assms(2) by blast then have "is_semialgebraic 1 (\ (Qp_to_R1_set ` Xs))" using assms finite_intersection_is_semialgebraic''[of "Qp_to_R1_set ` Xs" 1] by blast then have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms to_R1_inter[of Xs] by presburger then show ?thesis using is_univ_semialgebraicI by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Cartesian Products of Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) lemma Qp_times_basic_semialg_right: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" proof show "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ basic_semialg_set (n + m) k a" proof fix x assume "x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" then obtain as bs where as_bs_def: "as \ (basic_semialg_set n k a) \ bs \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ x = as@bs" using cartesian_product_memE[of x "basic_semialg_set n k a" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p n] basic_semialg_set_def by (metis (no_types, lifting) append_take_drop_id basic_semialg_set_memE(1) subsetI) have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using as_bs_def basic_semialg_set_memE(1) cartesian_product_closed' by blast have 1: "(Qp_ev a x = Qp_ev a as)" using as_bs_def poly_eval_cartesian_prod[of as n bs m a] assms basic_semialg_set_memE(1) by blast obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev a as = (y[^]k))" using as_bs_def using basic_semialg_set_memE(2)[of as n k a] by blast show "x \ basic_semialg_set (n + m) k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "0") apply (simp add: y_def) using "1" y_def by blast qed show "basic_semialg_set (n + m) k a \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" proof fix x assume A: "x \ basic_semialg_set (n + m) k a" have A0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A basic_semialg_set_memE(1) by blast have A1: "set x \ carrier Q\<^sub>p" using A0 by (metis (no_types, lifting) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth subsetI) have A2: "length x = n + m" using A0 cartesian_power_car_memE by blast obtain as where as_def: "as = take n x" by blast obtain bs where bs_def: "bs = drop n x" by blast have 0: "x = as@bs" using A as_def bs_def by (metis append_take_drop_id) have 1: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) using as_def A2 apply (simp add: A2 min.absorb2) by (metis (no_types, lifting) A1 as_def dual_order.trans set_take_subset) have 2: "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule cartesian_power_car_memI) using bs_def A2 apply (simp add: A2) by (metis A1 bs_def order.trans set_drop_subset) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)" using basic_semialg_set_memE A by meson have 3: "as \ basic_semialg_set n k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using \y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)\ apply blast using y_def A 1 0 2 assms(1) poly_eval_cartesian_prod by blast show " x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" using 3 2 "0" by (metis (mono_tags, lifting) as_def basic_semialg_set_memE(1) bs_def cartesian_product_memI subsetI) qed qed lemma Qp_times_basic_semialg_right_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: " a \ carrier (Q\<^sub>p[\\<^bsub>n+m\<^esub>])" using assms poly_ring_car_mono'(2) by blast have 2: "is_semialgebraic (n + m) (basic_semialg_set (n + m) k a)" using assms basic_semialg_is_semialgebraic'[of a "n+m" k "basic_semialg_set (n + m) k a"] "1" by blast show ?thesis using 0 2 by metis qed lemma Qp_times_basic_semialg_right_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_right_is_semialgebraic[of k P] using assms(1) by blast qed lemma cartesian_product_memE': assumes "x \ cartesian_product A B" obtains a b where "a \ A \ b \ B \ x = a@b" using assms unfolding cartesian_product_def by blast lemma Qp_times_basic_semialg_left: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) = basic_semialg_set (n+m) k (shift_vars n m a)" proof show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) \ basic_semialg_set (n + m) k (shift_vars n m a)" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" then obtain as bs where as_bs_def: "as \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ bs \ (basic_semialg_set n k a) \ x = as@bs " using cartesian_product_memE' by blast have 0: "Qp_ev (shift_vars n m a) x = Qp_ev a bs" using A as_bs_def assms shift_vars_eval[of a n as m bs ] by (metis (no_types, lifting) basic_semialg_set_memE(1)) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a bs = (y[^]k)" using as_bs_def basic_semialg_set_memE(2) by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A as_bs_def by (metis (no_types, lifting) add.commute basic_semialg_set_memE(1) cartesian_product_closed') show "x \ basic_semialg_set (n + m) k (shift_vars n m a)" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using y_def apply blast using "0" y_def by blast qed show "basic_semialg_set (n + m) k (shift_vars n m a) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) " proof fix x assume A: "x \ basic_semialg_set (n + m) k (shift_vars n m a)" then obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev (shift_vars n m a) x = (y[^]k)" using assms basic_semialg_set_memE[of x "n + m" k "shift_vars n m a"] shift_vars_closed[of a m] Qp.cring_axioms by blast have "x \ carrier (Q\<^sub>p\<^bsup>m+n\<^esup>)" using A basic_semialg_set_memE(1) by (metis add.commute) then have "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using cartesian_product_carrier by blast then obtain as bs where as_bs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson cartesian_product_memE') have "bs \ (basic_semialg_set n k a)" apply(rule basic_semialg_set_memI[of _ _ y]) using as_bs_def apply blast apply (simp add: y_def) using y_def shift_vars_eval[of a n as m bs ] as_bs_def assms(1) by metis then show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" using as_bs_def unfolding cartesian_product_def by blast qed qed lemma Qp_times_basic_semialg_left_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a))" using basic_semialg_is_semialgebraic'[of a "n+m" k] Qp_times_basic_semialg_left by (metis assms(1) assms(2) basic_semialg_is_semialgebraic is_basic_semialg_def neq0_conv shift_vars_closed) lemma Qp_times_basic_semialg_left_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) A)" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic[of k P] using assms(1) by blast qed lemma product_of_basic_semialgs_is_semialg: assumes "k > 0" assumes "l > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: "cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b) = basic_semialg_set (m + n) l (shift_vars m n b)" using Qp_times_basic_semialg_left assms by blast have 2: "(cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b)) = cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b)" proof- have 0: "basic_semialg_set n k a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using basic_semialg_set_memE(1) by blast have 1: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by simp have 2: "carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by simp have 3: "basic_semialg_set m l b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using basic_semialg_set_memE(1) by blast show ?thesis using 0 1 2 3 cartesian_product_intersection[of "(basic_semialg_set n k a)" Q\<^sub>p n "(carrier (Q\<^sub>p\<^bsup>m\<^esup>))" m "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))" "(basic_semialg_set m l b)"] by (smt Collect_cong inf.absorb_iff1 inf.absorb_iff2) qed then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic Qp_times_basic_semialg_right_is_semialgebraic assms by (metis (no_types, lifting) add.commute intersection_is_semialg) qed lemma product_of_basic_semialgs_is_semialg': assumes "A \ (basic_semialgs n)" assumes "B \ (basic_semialgs m)" shows "is_semialgebraic (n + m) (cartesian_product A B)" proof- obtain k a where ka_def: "k > 0 \ a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ A = (basic_semialg_set n k a)" using assms by (metis is_basic_semialg_def mem_Collect_eq neq0_conv) obtain l b where lb_def: "l > 0 \ b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]) \ B = (basic_semialg_set m l b)" by (metis assms(2) gr_zeroI is_basic_semialg_def mem_Collect_eq) show ?thesis using ka_def lb_def assms product_of_basic_semialgs_is_semialg by blast qed lemma car_times_semialg_is_semialg: assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms is_semialgebraic_def semialg_sets_def apply blast apply (simp add: carrier_is_semialgebraic cartesian_product_carrier) proof- show "\A. A \ basic_semialgs m \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- fix A assume A: "A \ basic_semialgs m " then have " (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = A" by (metis basic_semialg_set_memE(1) inf_absorb1 is_basic_semialg_def mem_Collect_eq subsetI) then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" using add.commute[of n m] assms A Qp_times_basic_semialg_left_is_semialgebraic' by (simp add: \n + m = m + n\) qed show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- fix A C assume A: " A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" " is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" then have B: "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" using union_is_semialgebraic by blast show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- have 0: "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(1) gen_boolean_algebra_subset by blast have 1: " C \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(3) gen_boolean_algebra_subset by blast then show ?thesis using 0 A B using cartesian_product_binary_union_right[of A Q\<^sub>p m C "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))"] unfolding is_semialgebraic_def semialg_sets_def by presburger qed qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" then have "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using gen_boolean_algebra_subset by blast then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" using A cartesian_product_car_complement_right[of A Q\<^sub>p m n] unfolding is_semialgebraic_def semialg_sets_def by (metis (mono_tags, lifting) semialg_complement semialg_sets_def) qed qed lemma basic_semialg_times_semialg_is_semialg: assumes "A \ basic_semialgs n" assumes "is_semialgebraic m B" shows " is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms(2) is_semialgebraic_def semialg_sets_def apply blast using Qp_times_basic_semialg_right_is_semialgebraic' assms(1) apply blast apply (metis assms(1) basic_semialg_is_semialgebraic inf.absorb1 is_semialgebraic_closed mem_Collect_eq product_of_basic_semialgs_is_semialg') apply (metis (no_types, lifting) cartesian_product_binary_union_right is_semialgebraicI is_semialgebraic_closed semialg_sets_def union_is_semialgebraic) proof- show "\Aa. Aa \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product A Aa) \ is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - Aa))" proof- fix B assume A: "B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B))" using A assms cartesian_product_complement_right[of B Q\<^sub>p m A n] add.commute[of n m] proof - have f1: "\n B. \ is_semialgebraic n B \ B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson is_semialgebraic_closed) have "is_basic_semialg n A" using \A \ {S. is_basic_semialg n S}\ by blast then have f2: "is_semialgebraic n A" using padic_fields.basic_semialg_is_semialgebraic padic_fields_axioms by blast have "B \ semialg_sets m" using \B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {S. is_basic_semialg m S}\ semialg_sets_def by blast then have "is_semialgebraic m B" by (meson padic_fields.is_semialgebraicI padic_fields_axioms) then show ?thesis using f2 f1 by (metis (no_types) Qp_times_basic_semialg_right_is_semialgebraic' \A \ {S. is_basic_semialg n S}\ \\B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>); A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \ cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) - cartesian_product A B = cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B)\ \is_semialgebraic (n + m) (cartesian_product A B)\ diff_is_semialgebraic) qed qed qed text\Semialgebraic sets are closed under cartesian products\ lemma cartesian_product_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of A "carrier (Q\<^sub>p\<^bsup>n\<^esup>)""basic_semialgs n"]) using assms is_semialgebraicE semialg_sets_def apply blast using assms car_times_semialg_is_semialg apply blast using assms basic_semialg_times_semialg_is_semialg apply (simp add: Int_absorb2 basic_semialg_is_semialgebraic is_semialgebraic_closed) proof- show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product C B) \ is_semialgebraic (n + m) (cartesian_product (A \ C) B)" proof- fix A C assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product C B)" show "is_semialgebraic (n + m) (cartesian_product (A \ C) B)" using A cartesian_product_binary_union_left[of A Q\<^sub>p n C B] by (metis (no_types, lifting) gen_boolean_algebra_subset union_is_semialgebraic) qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" using assms A cartesian_product_complement_left[of A Q\<^sub>p n B m] unfolding is_semialgebraic_def semialg_sets_def by (metis car_times_semialg_is_semialg diff_is_semialgebraic is_semialgebraicE is_semialgebraicI is_semialgebraic_closed semialg_sets_def) qed qed (**************************************************************************************************) (**************************************************************************************************) subsection\$N^{th}$ Power Residues\ (**************************************************************************************************) (**************************************************************************************************) definition nth_root_poly where "nth_root_poly (n::nat) a = ((X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> n) \\<^bsub>Q\<^sub>p_x\<^esub> (to_poly a)" lemma nth_root_poly_closed: assumes "a \ carrier Q\<^sub>p" shows "nth_root_poly n a \ carrier Q\<^sub>p_x" using assms unfolding nth_root_poly_def by (meson UPQ.P.minus_closed UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_poly_closed) lemma nth_root_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "(nth_root_poly n a) \ b = (b[^]n) \\<^bsub>Q\<^sub>p\<^esub> a" using assms unfolding nth_root_poly_def using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_X_pow UPQ.to_fun_diff UPQ.to_fun_to_poly UPQ.to_poly_closed by presburger text\Hensel's lemma gives us this criterion for the existence of \n\-th roots\ lemma nth_root_poly_root: assumes "(n::nat) > 1" assumes "a \ \\<^sub>p" assumes "a \ \" assumes "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > 2* val ([n]\\)" shows "(\ b \ \\<^sub>p. ((b[^]n) = a))" proof- obtain \ where alpha_def: "\ \ carrier Z\<^sub>p \ \ \ = a" using assms(2) by blast have 0: "\ \ carrier Z\<^sub>p" by (simp add: alpha_def) have 1: "\ \ \\<^bsub>Z\<^sub>p\<^esub>" using assms alpha_def inc_of_one by blast obtain N where N_def: "N = [n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" by blast have N_closed: "N \ carrier Z\<^sub>p" using N_def Zp_nat_mult_closed by blast have 2: "\ ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = ([n]\ \)" proof(induction n) case 0 have 00: "[(0::nat)] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = \\<^bsub>Z\<^sub>p\<^esub>" using Zp_nat_inc_zero by blast have 01: "[(0::nat)] \\<^bsub>Q\<^sub>p\<^esub> \ = \" using Qp.nat_inc_zero by blast then show ?case using 00 inc_of_nat by blast next case (Suc n) then show ?case using inc_of_nat by blast qed have 3: "val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \) = val (\ \\<^bsub>Q\<^sub>p\<^esub> a)" using alpha_def Zp.one_closed ring_hom_one[of \ Z\<^sub>p Q\<^sub>p] inc_is_hom Zp.ring_hom_minus[of Q\<^sub>p \ "\\<^bsub>Z\<^sub>p\<^esub>" \ ] Qp.ring_axioms unfolding \_def by (metis Q\<^sub>p_def Zp.minus_closed Zp_def padic_fields.val_of_inc padic_fields_axioms) have 4: "([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) \ nonzero Z\<^sub>p" proof- have 40: "int n \ 0" using of_nat_0_le_iff by blast have "nat (int n) = n" using nat_int by blast hence "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = [int n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" using 40 unfolding add_pow_def int_pow_def nat_pow_def proof - have "(if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) 0 else rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n) = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n" by (meson of_nat_less_0_iff) then show "rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n = (let f = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) in if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> f (nat (- int n)) else f (nat (int n)))" using \nat (int n) = n\ by presburger qed thus ?thesis using Zp_char_0[of n] Zp.not_nonzero_memE Zp_char_0' assms(1) gr_implies_not_zero by blast qed then have 5: "val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = val ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))" using 2 ord_of_inc by (metis N_closed N_def val_of_inc) then have 6: "(val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \)) > 2*(val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>))" using assms 3 by presburger have "\ b \ carrier Z\<^sub>p. (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" using Zp_nth_root_lemma[of \ n] assms "0" "1" "6" by blast then obtain b where b_def: "b \ carrier Z\<^sub>p \ (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" by blast then have "\ (b [^]\<^bsub>Z\<^sub>p\<^esub>n) = a" using alpha_def by blast then have "(\ b) [^] n = a" by (metis Qp.nat_inc_zero Q\<^sub>p_def Qp.nat_pow_zero Zp.nat_pow_0 Zp.nat_pow_zero Zp_nat_inc_zero \_def alpha_def assms(3) b_def frac_inc_of_nat inc_of_one inc_pow not_nonzero_Qp) then show ?thesis using b_def by blast qed text\All points sufficiently close to 1 have nth roots\ lemma eint_nat_times_2: "2*(n::nat) = 2*eint n" using times_eint_simps(1) by (metis mult.commute mult_2_right of_nat_add) lemma P_set_of_one: "P_set 1 = nonzero Q\<^sub>p" apply(rule equalityI) apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq apply blast apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq using Qp.nat_pow_eone by blast lemma nth_power_fact: assumes "(n::nat) \ 1" shows "\ (m::nat) > 0. \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof(cases "n = 1") case True have "\ u \ carrier Q\<^sub>p. ac 1 u = 1 \ val u = 0 \ u \ P_set n" unfolding True P_set_of_one by (metis iless_Suc_eq padic_fields.val_ring_memE padic_fields.zero_in_val_ring padic_fields_axioms val_nonzero zero_eint_def) then show ?thesis by blast next case F: False obtain m where m_def: "m = 1 + nat ( 2*(ord ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))))" by blast then have m_pos: "m > 0" by linarith have "\ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof fix u assume A: "u \ carrier Q\<^sub>p" show " ac m u = 1 \ val u = 0 \ u \ P_set n" proof assume B: "ac m u = 1 \ val u = 0" then have 0: "val u = val \" by (smt A ac_def not_nonzero_Qp val_one val_ord zero_eint_def) have 1: "ac m u = ac m \" by (metis B Qp.one_nonzero ac_p ac_p_int_pow_factor angular_component_factors_x angular_component_p inc_of_one m_pos p_nonzero) have 2: "u \ nonzero Q\<^sub>p" proof- have "ac m \ = 0" by (meson ac_def) then have "u \ \" by (metis B zero_neq_one) then show ?thesis using A not_nonzero_Qp Qp.nonzero_memI by presburger qed then have 3: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using m_pos 0 1 ac_ord_prop[of "\" u "0::int" m] by (metis B Qp.one_nonzero add.right_neutral eint.inject val_ord zero_eint_def) show "u \ P_set n" proof(cases "u = \") case True then show ?thesis by (metis P_set_one insert_iff zeroth_P_set) next case False have F0: "u \ \\<^sub>p" apply(rule val_ring_memI, rule A) unfolding 0 val_one by auto have F1: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using False 3 by blast have "ord (\ \ u) \ m" by (metis A F1 False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) val_ord) hence F2: "ord (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using m_def F1 A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) int_nat_eq of_nat_1 of_nat_add val_ord[of "\ \ u"] eint_nat_times_2 by linarith have "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(val ([n]\ \))" proof- have 0: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using F2 val_ord[of "\ \ u"] A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(2) by presburger have "n > 0" using assms by linarith hence "eint (ord ([n] \ \)) = val ([n] \ \)" using val_ord_nat_inc[of n] by blast hence "2*ord ([n]\ \) = 2*val ([n]\ \)" by (metis inc_of_nat times_eint_simps(1)) thus ?thesis using 0 val_ord[of "\ \ u"] assms by presburger qed then have "(\ b \ \\<^sub>p. ((b[^]n) = u))" using m_def False nth_root_poly_root[of n u] F0 assms F by linarith then have "(\ b \ carrier Q\<^sub>p. ((b[^]n) = u))" using val_ring_memE by blast then show "u \ P_set n" using P_set_def[of n] 2 by blast qed qed qed then show ?thesis using m_pos by blast qed definition pow_res where "pow_res (n::nat) x = {a. a \ carrier Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. (a = x \ (y[^]n)))}" lemma nonzero_pow_res: assumes "x \ nonzero Q\<^sub>p" shows "pow_res (n::nat) x \ nonzero Q\<^sub>p" proof fix a assume "a \ pow_res n x" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (a = x \ (y[^]n))" using pow_res_def by blast then show "a \ nonzero Q\<^sub>p" using assms Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero by blast qed lemma pow_res_of_zero: shows "pow_res n \ = {\}" unfolding pow_res_def apply(rule equalityI) apply(rule subsetI) unfolding mem_Collect_eq apply (metis Qp.integral_iff Qp.nat_pow_closed Qp.nonzero_closed Qp.zero_closed insertCI) apply(rule subsetI) unfolding mem_Collect_eq by (metis Qp.nat_pow_one Qp.one_nonzero Qp.r_one Qp.zero_closed equals0D insertE) lemma equal_pow_resI: assumes "x \ carrier Q\<^sub>p" assumes "y \ pow_res n x" shows "pow_res n x = pow_res n y" proof have y_closed: "y \ carrier Q\<^sub>p" using assms unfolding pow_res_def by blast obtain c where c_def: "c \ nonzero Q\<^sub>p \ y = x \ (c[^]n)" using assms pow_res_def by blast have "((inv c)[^]n) = inv (c[^]n)" using c_def Qp.field_axioms Qp.nat_pow_of_inv Units_eq_nonzero by blast then have "y \ ((inv c)[^]n) = x" using y_closed c_def assms Qp.inv_cancelL(2) Qp.nonzero_closed Qp_nat_pow_nonzero Units_eq_nonzero by presburger then have P0: "(inv c) \ nonzero Q\<^sub>p \ x =y \ ((inv c)[^]n) " using c_def nonzero_inverse_Qp by blast show "pow_res n x \ pow_res n y" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = x \ (b[^]n)" using A pow_res_def by blast then have 0: "b \ nonzero Q\<^sub>p \ a = y \ ((inv c)[^]n) \ (b[^]n)" using \y \ inv c [^] n = x\ by blast have "b \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" proof- have "(inv c)[^]n \ (b[^]n) = ((inv c) \ b)[^]n" using c_def b_def assms P0 Qp.nat_pow_distrib Qp.nonzero_closed by presburger then have " y \ (((inv c)[^]n) \ (b[^]n)) = y \ (((inv c) \ b)[^]n)" by presburger then show ?thesis using y_closed 0 P0 Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) by presburger qed then have "((inv c) \ b) \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" by (metis P0 Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed not_nonzero_Qp) then show "a \ pow_res n y" using pow_res_def \a \ carrier Q\<^sub>p\ by blast qed show "pow_res n y \ pow_res n x" proof fix a assume A: "a \ pow_res n y" then have 0: "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = y \ (b[^]n)" using A pow_res_def by blast then have "a = (x \ (c[^]n)) \ (b[^]n)" using c_def by blast then have "a = x \ ((c[^]n) \ (b[^]n))" by (meson Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) b_def c_def) then have "a = x \ ((c \ b)[^]n)" using Qp.nat_pow_distrib Qp.nonzero_closed b_def c_def by presburger then have "(c \ b) \ nonzero Q\<^sub>p \ a = x \ ((c \ b)[^]n)" by (metis Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed b_def c_def not_nonzero_Qp) then show "a \ pow_res n x" using pow_res_def 0 by blast qed qed lemma zeroth_pow_res: assumes "x \ carrier Q\<^sub>p" shows "pow_res 0 x = {x}" apply(rule equalityI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms apply (metis Qp.nat_pow_0 Qp.r_one singletonI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms by (metis Qp.nat_pow_0 Qp.one_nonzero Qp.r_one equals0D insertE) lemma Zp_car_zero_res: assumes "x \ carrier Z\<^sub>p" shows "x 0 = 0" using assms unfolding Zp_def using Zp_def Zp_defs(3) padic_set_zero_res prime by blast lemma zeroth_ac: assumes "x \ carrier Q\<^sub>p" shows "ac 0 x = 0" apply(cases "x = \ ") unfolding ac_def apply presburger using assms angular_component_closed[of x] Zp_car_zero_res unfolding nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_imp_nonzero: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "x \ nonzero Q\<^sub>p" using assms unfolding ac_def nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_val_ord: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "val x = ord x" using nonzero_ac_imp_nonzero assms val_ord by blast lemma pow_res_equal_ord: assumes "n > 0" shows "\m > 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof- obtain m where m_def_0: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n)" using assms nth_power_fact[of n] by (metis less_imp_le_nat less_one linorder_neqE_nat nat_le_linear zero_less_iff_neq_zero) then have m_def: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ ord u = 0 \ u \ P_set n)" by (smt nonzero_ac_val_ord zero_eint_def) have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y " then have 0: "ac m (x \
y) = 1" using ac_inv[of y m] ac_mult by (metis ac_inv'''(1) ac_mult' m_def nonzero_inverse_Qp) have 1: "ord (x \
y) = 0" using A ord_fract by presburger have 2: "(x \
y) \ nonzero Q\<^sub>p" using A by (metis Qp.nonzero_closed Qp.nonzero_mult_closed local.fract_cancel_right nonzero_inverse_Qp not_nonzero_Qp zero_fract) have 3: "(x \
y) \ P_set n" using m_def 0 1 2 nonzero_def by (smt Qp.nonzero_closed) then obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^]n) = (x \
y)" using P_set_def by blast then have "(x \
y) \ y = (b[^]n) \ y" by presburger then have "x = (b[^]n) \ y" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_left) then have "x = y \(b[^]n)" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_right) then have "x \ pow_res n y" unfolding pow_res_def using A b_def by (metis (mono_tags, lifting) "2" Qp.nat_pow_0 Qp.nonzero_closed Qp_nonzero_nat_pow mem_Collect_eq not_gr_zero) then show "pow_res n x = pow_res n y" using A equal_pow_resI[of x y n] unfolding nonzero_def by (metis (mono_tags, lifting) A Qp.nonzero_closed equal_pow_resI) qed qed qed then show ?thesis using m_def by blast qed lemma pow_res_equal: assumes "n > 0" shows "\m> 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = (ord y mod n) \ pow_res n x = pow_res n y" proof- obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms pow_res_equal_ord by meson have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n" show "pow_res n x = pow_res n y" proof- have A0: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p" using A by blast have A1: "ac m x = ac m y" using A by blast have A2: "ord x = ord y mod int n" using A by blast obtain k where k_def: "k = ord x" by blast obtain l where l_def: "ord y = ord x + (l:: int)*(int n)" using assms A2 by (metis A k_def mod_eqE mod_mod_trivial mult_of_nat_commute) have m_def': "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" using m_def by blast have 0: "ord (y \ (\[^](- l*n))) = ord x" proof- have 0: "ord (y \ (\[^](- l*n))) = ord y + (ord (\[^](- l*n)))" using ord_mult p_nonzero A0 Qp_int_pow_nonzero by blast have 1: "ord (\[^](- l*n)) = - l*n" using ord_p_pow_int[of "-l*n"] by blast then have "ord (y \ (\[^](- l*n))) = ord y - l*n" using 0 by linarith then show ?thesis using k_def l_def by linarith qed have 1: "ac m (y \ (\[^](- l*n))) = ac m y" using assms ac_p_int_pow_factor_right[of ] m_def A Qp.nonzero_closed by presburger have 2: "y \ (\[^](- l*n)) \ nonzero Q\<^sub>p" using A0 Qp_int_pow_nonzero[of \ "- l*n"] Qp.cring_axioms nonzero_def cring.cring_simprules(5) fract_cancel_left not_nonzero_Qp p_intpow_inv'' p_nonzero zero_fract Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_memI Qp.nonzero_mult_closed minus_mult_commute mult_minus_right p_intpow_closed(1) p_intpow_closed(2) by presburger then have 3: "pow_res n (y \ (\[^](- l*n))) = pow_res n x" using 2 A0 m_def'[of "y \ (\[^](- l*n))" x] "0" "1" A1 by linarith have 4: "(y \ (\[^](- l*n))) = (y \ ((\[^]- l)[^]n))" using Qp_int_nat_pow_pow[of \ "-l" n] p_nonzero by presburger have "y \ (\[^](- l*n))\ pow_res n y " using "2" "4" Qp_int_pow_nonzero nonzero_def p_nonzero unfolding pow_res_def nonzero_def proof - assume a1: "\x n. x \ {a \ carrier Q\<^sub>p. a \ \} \ x [^] (n::int) \ {a \ carrier Q\<^sub>p. a \ \}" assume a2: "\ \ {a \ carrier Q\<^sub>p. a \ \}" assume a3: "y \ \ [^] (- l * int n) \ {a \ carrier Q\<^sub>p. a \ \}" have f4: "\ [^] (- 1 * l) \ {r \ carrier Q\<^sub>p. r \ \}" using a2 a1 by presburger have f5: "- l = - 1 * l" by linarith then have f6: "y \ \ [^] (- 1 * l * int n) = y \ (\ [^] (- 1 * l)) [^] n" using \y \ \ [^] (- l * int n) = y \ (\ [^] - l) [^] n\ by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. r \ \}" using f5 a3 by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f4 by blast then have "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f6 f5 by presburger then show "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra\{r \ carrier Q\<^sub>p. r \ \}. r = y \ ra [^] n}" by meson qed then have "pow_res n (y \ (\[^](- l*n))) = pow_res n y" using equal_pow_resI[of "(y \ (\[^](- l*n)))" y n] "2" A0 assms Qp.nonzero_mult_closed p_intpow_closed(2) by (metis (mono_tags, opaque_lifting) "3" A Qp.nonzero_closed equal_pow_resI) then show ?thesis using 3 by blast qed qed qed qed then show ?thesis using m_def by blast qed definition pow_res_classes where "pow_res_classes n = {S. \x \ nonzero Q\<^sub>p. S = pow_res n x }" lemma pow_res_semialg_def: assumes "x \ nonzero Q\<^sub>p" assumes "n \ 1" shows "\P \ carrier Q\<^sub>p_x. pow_res n x = (univ_basic_semialg_set n P) - {\}" proof- have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (inv x) \ a = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))" unfolding pow_res_def by blast then obtain y where y_def: "y \ nonzero Q\<^sub>p \a = x \ (y[^]n)" by blast then have "y \ nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" proof - show ?thesis by (metis (no_types, opaque_lifting) Qp.m_assoc Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ y [^] n)\ assms(1) local.fract_cancel_right nonzero_inverse_Qp y_def) qed then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" using assms \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))\ by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)} \ pow_res n x" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" show "a \ pow_res n x" proof- have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using A by blast then obtain y where y_def: "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ a = x \(y[^]n)" by (metis Qp.l_one Qp.m_assoc Qp.nonzero_closed Qp.not_nonzero_memI \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = y [^] n)\ assms(1) field_inv(2) inv_in_frac(1)) then show ?thesis by (metis (mono_tags, lifting) \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))\ mem_Collect_eq pow_res_def) qed qed qed obtain P where P_def: "P = up_ring.monom Q\<^sub>p_x (inv x) 1" by blast have P_closed: "P \ carrier Q\<^sub>p_x" using P_def Qp.nonzero_closed Qp.nonzero_memE(2) UPQ.is_UP_monomE(1) UPQ.is_UP_monomI assms(1) inv_in_frac(1) by presburger have P_eval: "\a. a \ carrier Q\<^sub>p \ (P \ a) = (inv x) \ a" using P_def to_fun_monom[of ] by (metis Qp.nat_pow_eone Qp.nonzero_closed Qp.not_nonzero_memI assms(1) inv_in_frac(1)) have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using 0 by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using P_eval by (metis (mono_tags, lifting) mem_Collect_eq) qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ pow_res n x" proof fix a assume "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" using P_eval \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then show "a \ pow_res n x" using 0 by blast qed qed have 1: "univ_basic_semialg_set n P - {\} = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "univ_basic_semialg_set n P - {\} \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume A: "a \ univ_basic_semialg_set n P - {\}" then have A0: "a \ carrier Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" unfolding univ_basic_semialg_set_def by blast then have A0': "a \ nonzero Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" using A by (metis DiffD2 not_nonzero_Qp singletonI) then obtain y where y_def: "y\carrier Q\<^sub>p \ P \ a = (y[^]n)" by blast have A1: "(P \ a) \ \" using P_eval A0' Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) assms(1) inv_in_frac(1) inv_in_frac(2) by presburger have A2: "y \ nonzero Q\<^sub>p" proof- have A20: "(y[^]n) \\" using A1 y_def by blast have "y \ \" apply(rule ccontr) using A20 assms by (metis Qp.nat_pow_eone Qp.semiring_axioms Qp.zero_closed le_zero_eq semiring.nat_pow_zero) then show ?thesis using y_def A1 not_nonzero_Qp Qp.not_nonzero_memE by blast qed then have "y \ nonzero Q\<^sub>p \ P \ a = (y[^]n)" using y_def by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using A0 nonzero_def by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ univ_basic_semialg_set n P - {\}" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y \\ \ y\ carrier Q\<^sub>p \ P \ a = (y[^]n)" by (metis (mono_tags, opaque_lifting) Qp.nonzero_closed Qp.not_nonzero_memI) then have "a \\" using P_eval by (metis Qp.m_comm Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.zero_closed assms(1) inv_in_frac(1) zero_fract) then show "a \ univ_basic_semialg_set n P - {\}" unfolding univ_basic_semialg_set_def using A \y \ \ \ y \ carrier Q\<^sub>p \ P \ a = (y[^]n)\ by blast qed qed show ?thesis using 0 1 by (metis (no_types, lifting) P_closed) qed lemma pow_res_is_univ_semialgebraic: assumes "x \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (pow_res n x)" proof(cases "n = 0") case True have T0: "pow_res n x = {x}" unfolding True using assms by (simp add: assms zeroth_pow_res) have "[x] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms Qp.to_R1_closed by blast hence "is_semialgebraic 1 {[x]}" using is_algebraic_imp_is_semialg singleton_is_algebraic by blast thus ?thesis unfolding T0 using assms by (simp add: \x \ carrier Q\<^sub>p\ finite_is_univ_semialgebraic) next case False show ?thesis proof(cases "x = \") case True then show ?thesis using finite_is_univ_semialgebraic False pow_res_of_zero by (metis Qp.zero_closed empty_subsetI finite.emptyI finite.insertI insert_subset) next case F: False then show ?thesis using False pow_res_semialg_def[of x n] diff_is_univ_semialgebraic[of _ "{\}"] finite_is_univ_semialgebraic[of "{\}"] by (metis Qp.zero_closed assms empty_subsetI finite.emptyI finite.insertI insert_subset less_one less_or_eq_imp_le linorder_neqE_nat not_nonzero_Qp univ_basic_semialg_set_is_univ_semialgebraic) qed qed lemma pow_res_is_semialg: assumes "x \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1 ` (pow_res n x))" using assms pow_res_is_univ_semialgebraic is_univ_semialgebraicE by blast lemma pow_res_refl: assumes "x \ carrier Q\<^sub>p" shows "x \ pow_res n x" proof- have "x = x \ (\ [^]n)" using assms Qp.nat_pow_one Qp.r_one by presburger thus ?thesis using assms unfolding pow_res_def mem_Collect_eq using Qp.one_nonzero by blast qed lemma equal_pow_resE: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "n > 0" assumes "pow_res n a = pow_res n b" shows "\ s \ P_set n. a = b \ s" proof- have "a \ pow_res n b" using assms pow_res_refl by blast then obtain y where y_def: " y \ nonzero Q\<^sub>p \ a = b \ y[^]n" unfolding pow_res_def by blast thus ?thesis unfolding P_set_def using Qp.nonzero_closed Qp_nat_pow_nonzero by blast qed lemma pow_res_one: assumes "x \ nonzero Q\<^sub>p" shows "pow_res 1 x = nonzero Q\<^sub>p" proof show "pow_res 1 x \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of x 1] by blast show "nonzero Q\<^sub>p \ pow_res 1 x" proof fix y assume A: "y \ nonzero Q\<^sub>p" then have 0: "\ \ nonzero Q\<^sub>p \ y = x \ ((inv x)\ y)[^](1::nat)" using assms Qp.m_comm Qp.nat_pow_eone Qp.nonzero_closed Qp.nonzero_mult_closed Qp.one_nonzero local.fract_cancel_right nonzero_inverse_Qp by presburger have 1: "(inv x)\ y \ nonzero Q\<^sub>p" using A assms by (metis Qp.Units_m_closed Units_eq_nonzero nonzero_inverse_Qp) then show "y \ pow_res 1 x" unfolding pow_res_def using 0 1 A Qp.nonzero_closed by blast qed qed lemma pow_res_zero: assumes "n > 0" shows "pow_res n \ = {\}" proof show "pow_res n \ \ {\}" unfolding pow_res_def using Qp.l_null Qp.nat_pow_closed Qp.nonzero_closed by blast show "{\} \ pow_res n \" using assms unfolding pow_res_def using Qp.l_null Qp.one_closed Qp.one_nonzero empty_subsetI insert_subset by blast qed lemma equal_pow_resI': assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ P_set n" assumes "a = b \ c" assumes "n > 0" shows "pow_res n a = pow_res n b" proof- obtain y where y_def: "c = y[^]n \ y \ carrier Q\<^sub>p" using assms unfolding P_set_def by blast have c_nonzero: "c \ nonzero Q\<^sub>p" using P_set_nonzero'(1) assms(3) by blast have y_nonzero: "y \ nonzero Q\<^sub>p" using y_def c_nonzero Qp_nonzero_nat_pow assms(5) by blast have 0: "a \ pow_res n b" using assms y_nonzero y_def unfolding pow_res_def by blast show ?thesis apply(cases "b = \") using pow_res_zero Qp.l_null Qp.nonzero_closed assms(4) c_nonzero apply presburger by (metis "0" assms(1) assms(2) assms(5) not_nonzero_Qp equal_pow_resI) qed lemma equal_pow_resI'': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "a \ inv b \ P_set n" shows "pow_res n a = pow_res n b" using assms equal_pow_resI'[of a b "a \ inv b" n] Qp.nonzero_closed local.fract_cancel_right by blast lemma equal_pow_resI''': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "c \ nonzero Q\<^sub>p" assumes "pow_res n (c \ a) = pow_res n (c \ b)" shows "pow_res n a = pow_res n b" proof- have 0: "c \a \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) assms(4)) have 1: "c \b \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(3) assms(4)) have "c\a \ pow_res n (c\b)" proof(cases "n = 1") case True then show ?thesis using assms 0 1 pow_res_one[of "c\b"] by blast next case False then have "n \ 2" using assms(1) by linarith then show ?thesis using assms 0 1 pow_res_refl[of "c\a" n] unfolding nonzero_def by blast qed then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (c \ a) = (c \ b)\y[^]n" using assms unfolding pow_res_def by blast then have "a = b\y[^]n" using assms by (metis Qp.m_assoc Qp.m_lcancel Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero) then show ?thesis by (metis P_set_memI Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) assms(3) equal_pow_resI' y_def) qed lemma equal_pow_resI'''': assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "a = b \ u" assumes "u \ P_set n" shows "pow_res n a = pow_res n b" proof(cases "a = \") case True then have "b = \" using assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.not_nonzero_memI mem_Collect_eq) then show ?thesis using pow_res_zero using True by blast next case False then have 0: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast have 1: "b \ nonzero Q\<^sub>p" using 0 assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral_iff Qp.nonzero_closed mem_Collect_eq not_nonzero_Qp) have 2: "a \ (inv b)\ P_set n" using assms 0 1 by (metis P_set_nonzero'(2) Qp.inv_cancelR(1) Qp.m_comm Qp.nonzero_memE(2) Units_eq_nonzero inv_in_frac(1)) then show ?thesis using equal_pow_resI'' by (meson "0" "1" assms(1) equal_pow_resI) qed lemma Zp_Units_ord_zero: assumes "a \ Units Z\<^sub>p" shows "ord_Zp a = 0" proof- have "inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p" apply(rule Zp.nonzero_memI, rule Zp.Units_inv_closed, rule assms) using assms Zp.Units_inverse in_Units_imp_not_zero by blast then have "ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub> a) = ord_Zp a + ord_Zp (inv \<^bsub>Z\<^sub>p\<^esub> a)" using assms ord_Zp_mult Zp.Units_nonzero zero_not_one by (metis Zp.zero_not_one) then show ?thesis by (smt Zp.Units_closed Zp.Units_r_inv Zp.integral_iff Zp.nonzero_closed \inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p\ assms ord_Zp_one ord_pos) qed lemma pow_res_nth_pow: assumes "a \ nonzero Q\<^sub>p" assumes "n > 0" shows "pow_res n (a[^]n) = pow_res n \" proof show "pow_res n (a [^] n) \ pow_res n \" proof fix x assume A: "x \ pow_res n (a [^] n)" then show "x \ pow_res n \" by (metis P_set_memI Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.one_closed assms(1) assms(2) equal_pow_resI') qed show "pow_res n \ \ pow_res n (a [^] n)" proof fix x assume A: "x \ pow_res n \" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ x = \\y[^]n" unfolding pow_res_def by blast then have 0: "x = y[^]n" using Qp.l_one Qp.nonzero_closed by blast have "y[^]n = a[^]n \ (inv a \ y)[^]n" proof- have "a[^]n \ (inv a \ y)[^]n = a[^]n \ (inv a)[^]n \ y[^]n" using Qp.Units_inv_closed Qp.m_assoc Qp.nat_pow_closed Qp.nat_pow_distrib Qp.nonzero_closed Units_eq_nonzero assms(1) y_def by presburger then show ?thesis by (metis Qp.Units_inv_inv Qp.inv_cancelR(1) Qp.nat_pow_distrib Qp.nonzero_closed Qp.nonzero_mult_closed Units_eq_nonzero assms(1) nonzero_inverse_Qp y_def) qed then show "x \ pow_res n (a [^] n)" using y_def A assms unfolding pow_res_def mem_Collect_eq by (metis "0" Qp.integral Qp.m_closed Qp.nonzero_closed Qp.not_nonzero_memI inv_in_frac(1) inv_in_frac(2) not_nonzero_Qp) qed qed lemma pow_res_of_p_pow: assumes "n > 0" shows "pow_res n (\[^]((l::int)*n)) = pow_res n \" proof- have 0: "\[^]((l::int)*n) = (\[^]l)[^]n" using Qp_p_int_nat_pow_pow by blast have "\[^]((l::int)*n) \ P_set n" using P_set_memI[of _ "\[^]l"] by (metis "0" Qp.not_nonzero_memI Qp_int_pow_nonzero p_intpow_closed(1) p_nonzero) thus ?thesis using "0" assms p_intpow_closed(2) pow_res_nth_pow by presburger qed lemma pow_res_nonzero: assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n b" shows "b \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of a n] pow_res_zero[of n] by (metis insert_subset not_nonzero_Qp) lemma pow_res_mult: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "d \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n c" assumes "pow_res n b = pow_res n d" shows "pow_res n (a \ b) = pow_res n (c \ d)" proof(cases "a \ nonzero Q\<^sub>p") case True then have "c \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where alpha_def: "\ \ nonzero Q\<^sub>p \ a = c \ \[^]n" using assms True pow_res_refl[of a n] unfolding assms unfolding pow_res_def by blast show ?thesis proof(cases "b \ nonzero Q\<^sub>p") case T: True then have "d \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where beta_def: "\ \ nonzero Q\<^sub>p \ b = d \ \[^]n" using T pow_res_refl[of b n] unfolding assms unfolding pow_res_def using assms by blast then have "a \ b = (c \ d) \ (\[^]n \ \[^] n)" using Qp.m_assoc Qp.m_lcomm Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero alpha_def assms(3) assms(4) assms(5) by presburger then have 0: "a \ b = (c \ d) \ ((\ \ \)[^] n)" by (metis Qp.nat_pow_distrib Qp.nonzero_closed alpha_def beta_def) show ?thesis apply(intro equal_pow_resI'[of _ _ "(\ \ \)[^] n"] Qp.ring_simprules assms P_set_memI[of _ "\ \ \"] Qp.nat_pow_closed nonzero_memE 0 Qp_nat_pow_nonzero ) using alpha_def beta_def apply auto apply(intro nonzero_memI Qp.nonzero_mult_closed) using alpha_def beta_def nonzero_memE apply auto by (meson Qp.integral_iff) next case False then have "b = \" by (meson assms not_nonzero_Qp) then have "d = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \b = \\ assms by presburger qed next case False then have "a = \" by (meson assms not_nonzero_Qp) then have "c = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \a = \\ assms Qp.l_null by presburger qed lemma pow_res_p_pow_factor: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" shows "pow_res n a = pow_res n (\[^]((l::int)*n) \ a)" proof(cases "a = \") case True then show ?thesis using Qp.r_null p_intpow_closed(1) by presburger next case False then show ?thesis using assms pow_res_of_p_pow using Qp.m_comm Qp.one_closed Qp.r_one p_intpow_closed(1) pow_res_mult by presburger qed lemma pow_res_classes_finite: assumes "n \ 1" shows "finite (pow_res_classes n)" proof(cases "n = 1") case True have "pow_res_classes n = {(nonzero Q\<^sub>p)}" using True pow_res_one unfolding pow_res_classes_def using Qp.one_nonzero by blast then show ?thesis by auto next case False then have n_bound: "n \ 2" using assms by linarith obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms False pow_res_equal_ord n_bound by (metis gr_zeroI le_numeral_extra(2)) obtain f where f_def: "f = (\ \ \. (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)))" by blast have 0: "\x. x \ nonzero Q\<^sub>p \ pow_res n x = f (ac m x) (ord x)" proof- fix x assume A: "x \ nonzero Q\<^sub>p" obtain \ where eta_def: "\ = ac m x" by blast obtain \ where nu_def: "\ = ord x" by blast have "\y \pow_res n x. ac m y = ac m x \ ord y = ord x" using pow_res_refl A assms neq0_conv Qp.nonzero_closed by blast hence "pow_res n x \ (pow_res_classes n) \ (\ y \ (pow_res n x). ac m y = \ \ ord y = \)" unfolding nu_def eta_def using assms unfolding pow_res_classes_def using A by blast then have 0: "(\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" by blast have "f \ \ = (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" using f_def by blast then have 1: "f \ \ \ (pow_res_classes n) \ ((\ y \ (f \ \). ac m y = \ \ ord y = \))" using 0 someI_ex[of "\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)"] unfolding f_def by blast then obtain y where y_def: "y \ (f \ \) \ ac m y = ac m x \ ord y = ord x" unfolding nu_def eta_def by blast obtain a where a_def: "a \ nonzero Q\<^sub>p \ (f \ \) = pow_res n a" using 1 unfolding pow_res_classes_def by blast then have 2: "y \ pow_res n a" using y_def by blast have 3: "y \ nonzero Q\<^sub>p" using y_def nonzero_pow_res[of a n] a_def by blast then have 4: "pow_res n y = pow_res n a" using 3 y_def a_def equal_pow_resI[of y a n] n_bound Qp.nonzero_closed by (metis equal_pow_resI) have 5: "pow_res n y = f \ \" using 4 a_def by blast then show "pow_res n x = f (ac m x) (ord x)" unfolding eta_def nu_def using "3" A m_def y_def by blast qed obtain N where N_def: "N > 0 \ (\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n)" using n_bound nth_power_fact assms by blast have 1: "\x. x \ nonzero Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. ord y \ 0 \ ord y < n \ pow_res n x = pow_res n y)" proof- fix x assume x_def: "x \ nonzero Q\<^sub>p" then obtain k where k_def: "k = ord x mod n" by blast then obtain l where l_def: "ord x = (int n)*l + k" using cancel_div_mod_rules(2)[of n "ord x"0] unfolding k_def by (metis group_add_class.add.right_cancel) have "x = (\[^](ord x)) \ \ (angular_component x)" using x_def angular_component_factors_x by blast then have "x = (\[^](n*l + k)) \ \ (angular_component x)" unfolding l_def by blast hence "x = \[^](int n*l) \ (\[^] k) \ \ (angular_component x)" by (metis p_intpow_add) hence 0: "x = (\[^]l)[^]n \ (\[^] k) \ \ (angular_component x)" using p_pow_factor[of n l k] \x = \ [^] (int n * l + k) \ \ (angular_component x)\ by presburger have 1: "\ (angular_component x) \ carrier Q\<^sub>p" using x_def angular_component_closed inc_closed by blast hence 2: "x = (\[^]l)[^]n \ ((\[^] k) \ \ (angular_component x))" using 0 by (metis Qp.m_assoc Qp.nat_pow_closed p_intpow_closed(1)) obtain a where a_def: "a = (\[^] k) \ \ (angular_component x)" by blast have 30: "angular_component x \ Units Z\<^sub>p" using angular_component_unit x_def by blast then have 3: "\ (angular_component x) \ Units Q\<^sub>p" by (metis Units_eq_nonzero Zp.Units_closed in_Units_imp_not_zero inc_of_nonzero not_nonzero_Qp) have 4: "\ (angular_component x) \ nonzero Q\<^sub>p" using 3 Units_nonzero_Qp by blast have a_nonzero: "a \ nonzero Q\<^sub>p" unfolding a_def 4 by (meson "3" Qp.UnitsI(1) Qp.Units_m_closed Units_nonzero_Qp p_intpow_closed(1) p_intpow_inv) have 5: "x = a \(\[^]l)[^]n" using 2 a_nonzero unfolding a_def using Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed p_intpow_closed(1) by presburger hence "x \ pow_res n a" unfolding pow_res_def using Qp.nonzero_closed Qp_int_pow_nonzero p_nonzero x_def by blast hence 6:"pow_res n a = pow_res n x" using x_def a_def equal_pow_resI[of x a n] a_nonzero n_bound Qp.nonzero_closed equal_pow_resI by blast have 7: "ord (\ (angular_component x)) = 0" proof- have "ord_Zp (angular_component x) = 0" using 30 Zp_Units_ord_zero by blast then have "val_Zp (angular_component x) = 0" using "30" unit_imp_val_Zp0 by blast then have "val (\ (angular_component x)) = 0" by (metis angular_component_closed val_of_inc x_def) then show ?thesis using angular_component_closed x_def by (metis "30" Zp.Units_closed \ord_Zp (angular_component x) = 0\ in_Units_imp_not_zero not_nonzero_Qp ord_of_inc) qed have 8: "ord a = k" unfolding a_def using 3 4 7 ord_mult[of "\ [^] k" "\ (angular_component x)"] ord_p_pow_int[of k] p_pow_nonzero using Qp_int_pow_nonzero p_nonzero by presburger have 9: "k < n" unfolding k_def using assms by auto from 6 8 9 assms have \0 \ ord a\ \ord a < int n\ \pow_res n x = pow_res n a\ by (auto simp add: k_def) with a_nonzero show "\y\nonzero Q\<^sub>p. 0 \ ord y \ ord y < int n \ pow_res n x = pow_res n y" by auto qed have 2: "\x. x \ (pow_res_classes n) \ \ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" proof- fix a assume A: "a \ (pow_res_classes n)" then obtain x where x_def: "x \ nonzero Q\<^sub>p \ a = pow_res n x" unfolding pow_res_classes_def by blast then obtain x' where x'_def: "x' \ nonzero Q\<^sub>p \ ord x' \ 0 \ ord x' < n \ pow_res n x' = a" using 1[of x] unfolding x_def by blast hence 20: "f (ac m x') (ord x') = a" using 0 by blast have 21: "ac m x' \ Units (Zp_res_ring m)" using x'_def ac_units m_def by presburger then have 22: "ac m x' \ Units (Zp_res_ring m) \ (ord x') \ ({0.. a = f (ac m x') (ord x')" using x'_def 20 atLeastLessThan_iff by blast then show "\ \ \. \ \ Units (Zp_res_ring m) \ \ \ {0.. a = f \ \" by blast qed obtain F where F_def: "F = (\ps. f (fst ps) (snd ps))" by blast have 3: "\x. x \ (pow_res_classes n) \ \ ps \ Units (Zp_res_ring m) \ {0.. pow_res_classes n" obtain \ \ where eta_nu_def: " \ \ Units (Zp_res_ring m) \ \ \ {0.. x = f \ \" using 2 A by blast then have "F (\, \) = x" unfolding F_def by (metis fst_conv snd_conv) then show " \ ps \ Units (Zp_res_ring m) \ {0.. F ` (Units (Zp_res_ring m) \ {0.. {0.. pow_res_classes n" shows "is_univ_semialgebraic S" proof- obtain x where x_def: "x\nonzero Q\<^sub>p \ S = pow_res n x" using assms unfolding pow_res_classes_def by blast then show ?thesis using pow_res_is_univ_semialgebraic using Qp.nonzero_closed by blast qed lemma pow_res_classes_semialg: assumes "S \ pow_res_classes n" shows "is_semialgebraic 1 (to_R1` S)" using pow_res_classes_univ_semialg assms(1) is_univ_semialgebraicE by blast definition nth_pow_wits where "nth_pow_wits n = (\ S. (SOME x. x \ (S \ \\<^sub>p)))` (pow_res_classes n)" lemma nth_pow_wits_finite: assumes "n > 0" shows "finite (nth_pow_wits n)" proof- have "n \ 1" by (simp add: assms leI) thus ?thesis unfolding nth_pow_wits_def using assms pow_res_classes_finite[of n] by blast qed lemma nth_pow_wits_covers: assumes "n > 0" assumes "x \ nonzero Q\<^sub>p" shows "\y \ (nth_pow_wits n). y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" proof- have PP: "(pow_res n x) \ pow_res_classes n" unfolding pow_res_classes_def using assms by blast obtain k where k_def: "val x = eint k" using assms val_ord by blast obtain N::int where N_def: "N = (if k < 0 then -k else k)" by blast then have N_nonneg: "N \ 0" unfolding N_def by presburger have 0: "int n \ 1" using assms by linarith have "N*(int n) + k \ 0" proof(cases "k<0") case True then have "N = -k" unfolding N_def by presburger then have "N*n + k = k*(1- int n)" using distrib_left[of k 1 "-int n"] mult_cancel_left2 mult_minus_left by (metis add.inverse_inverse diff_minus_eq_add minus_mult_minus neg_equal_iff_equal uminus_add_conv_diff) then show ?thesis using 0 True zero_less_mult_iff[of k "1 - int n"] proof - have "0 \ N * (int n - 1)" by (meson "0" N_nonneg diff_ge_0_iff_ge zero_le_mult_iff) then show ?thesis by (metis (no_types) \N = - k\ add.commute distrib_left minus_add_cancel mult_minus1_right uminus_add_conv_diff) qed next case False then have "N = k" unfolding N_def by presburger then show ?thesis using 0 False by (metis N_nonneg add_increasing2 mult_nonneg_nonneg of_nat_0_le_iff) qed then have 1: "ord (\[^](N*n)\x) \ 0" using ord_mult k_def val_ord assms by (metis Qp_int_pow_nonzero eint.inject ord_p_pow_int p_nonzero) have 2: "\[^](N*n)\x \ pow_res n x" proof- have "\[^](N*n) = (\[^]N)[^]n" using Qp_p_int_nat_pow_pow by blast then have "\[^]N \ nonzero Q\<^sub>p \ \[^](N*n)\x = x \ (\[^]N)[^]n" by (metis Qp.m_comm Qp.nonzero_closed Qp_int_pow_nonzero assms(2) p_nonzero) then show ?thesis unfolding pow_res_def by (metis (mono_tags, lifting) Qp.m_closed Qp.nonzero_closed assms(2) mem_Collect_eq p_intpow_closed(1)) qed have 3: "\[^](N*n)\x \ \\<^sub>p" using 1 assms by (metis Q\<^sub>p_def Qp.nonzero_mult_closed Qp_int_pow_nonzero Z\<^sub>p_def val_ring_ord_criterion \_def p_nonzero padic_fields.zero_in_val_ring padic_fields_axioms) have 4: "x \ pow_res n (\[^](N*n)\x)" using 2 equal_pow_resI[of x "\[^](N*n)\x" n] pow_res_refl[of "\[^](N*n)\x" n] assms Qp.nonzero_mult_closed p_intpow_closed(2) pow_res_refl Qp.nonzero_closed by metis have 5: "\[^](N*n)\x \ (pow_res n x \ \\<^sub>p)" using 2 3 by blast have 6: "(SOME z. z \ (pow_res n x) \ \\<^sub>p) \ pow_res n x \ \\<^sub>p" using 5 by (meson someI) obtain y where y_def: "y = (SOME z. z \ (pow_res n x) \ \\<^sub>p)" by blast then have A: "y \ pow_res n x" using "6" by blast then have "pow_res n x = pow_res n y" using equal_pow_resI[of x y n] assms y_def Qp.nonzero_closed nonzero_pow_res by blast then have 7: "x \ pow_res n y" using pow_res_refl[of x n] assms unfolding nonzero_def by blast have 8: "y \ nonzero Q\<^sub>p " using y_def PP 6 A nonzero_pow_res[of x n] assms by blast have 9: "y \ \\<^sub>p" using y_def "6" by blast have "y\(\S. SOME x. x \ S \ \\<^sub>p) ` pow_res_classes n \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res n y" using y_def PP 6 7 8 9 A nonzero_pow_res[of x n] assms by blast then show ?thesis unfolding nth_pow_wits_def by blast qed lemma nth_pow_wits_closed: assumes "n > 0" assumes "x \ nth_pow_wits n" shows "x \ carrier Q\<^sub>p" "x \ \\<^sub>p" "x \ nonzero Q\<^sub>p" "\ y \ pow_res_classes n. y = pow_res n x" proof- obtain c where c_def: "c \ pow_res_classes n \ x = (SOME x. x \ (c \ \\<^sub>p))" by (metis (no_types, lifting) assms(2) image_iff nth_pow_wits_def) then obtain y where y_def: "y \ nonzero Q\<^sub>p \ c = pow_res n y" unfolding pow_res_classes_def by blast then obtain a where a_def: "a \ (nth_pow_wits n) \ a \ nonzero Q\<^sub>p \ a \ \\<^sub>p \ y \ pow_res n a" using nth_pow_wits_covers[of n y] assms(1) by blast have 00: "pow_res n a = c" using equal_pow_resI[of a y n] y_def assms a_def unfolding nonzero_def by blast then have P :"a \ c \ \\<^sub>p" using pow_res_refl[of a n] assms a_def unfolding 00 nonzero_def by blast then show 0: "x \ \\<^sub>p" using c_def by (metis Collect_mem_eq Int_Collect tfl_some) then show "x \ carrier Q\<^sub>p" using val_ring_memE by blast have 1: "c \ nonzero Q\<^sub>p" using c_def nonzero_pow_res[of y n] unfolding pow_res_classes_def using assms(1) y_def by blast have "(SOME x. x \ (c \ \\<^sub>p)) \ (c \ \\<^sub>p)" using P tfl_some by (smt Int_def someI_ex) then have 2: "x \ c" using c_def by blast thus "x \ nonzero Q\<^sub>p" using 1 by blast show "\ y \ pow_res_classes n. y = pow_res n x" using 00 2 c_def P a_def equal_pow_resI[of a x n] 0 val_ring_memE assms(1) by blast qed lemma finite_extensional_funcset: assumes "finite A" assumes "finite (B::'b set)" shows "finite (A \\<^sub>E B)" using finite_PiE[of A "\_. B"] assms by blast lemma nth_pow_wits_exists: assumes "m > 0" assumes "c \ pow_res_classes m" shows "\x. x \ c \ \\<^sub>p" proof- obtain x where x_def: "x \ nonzero Q\<^sub>p \ pow_res m x = c" using assms unfolding pow_res_classes_def by blast obtain y where y_def: "y \ (nth_pow_wits m) \ y \ nonzero Q\<^sub>p \ y \ \\<^sub>p \ x \ pow_res m y" using nth_pow_wits_covers assms x_def by blast have 0: "pow_res m x = pow_res m y" using x_def y_def equal_pow_resI Qp.nonzero_closed assms(1) by blast then have 1: "y \ pow_res m x" using pow_res_refl[of y m ] y_def assms unfolding nonzero_def by blast thus ?thesis using x_def y_def assms by blast qed lemma pow_res_classes_mem_eq: assumes "m > 0" assumes "a \ pow_res_classes m" assumes "x \ a" shows "a = pow_res m x" proof- obtain y where y_def: "y \ nonzero Q\<^sub>p \ a = pow_res m y" using assms unfolding pow_res_classes_def by blast then show ?thesis using assms equal_pow_resI[of y x m] by (meson Qp.nonzero_closed nonzero_pow_res equal_pow_resI subsetD) qed lemma nth_pow_wits_neq_pow_res: assumes "m > 0" assumes "x \ nth_pow_wits m" assumes "y \ nth_pow_wits m" assumes "x \ y" shows "pow_res m x \ pow_res m y" proof- obtain a where a_def: "a \ pow_res_classes m \ x = (\ S. (SOME x. x \ (S \ \\<^sub>p))) a" using assms unfolding nth_pow_wits_def by blast obtain b where b_def: "b \ pow_res_classes m \ y = (\ S. (SOME x. x \ (S \ \\<^sub>p))) b" using assms unfolding nth_pow_wits_def by blast have a_neq_b: "a \ b" using assms a_def b_def by blast have 0: "x \ a \ \\<^sub>p" using a_def nth_pow_wits_exists[of m a] assms by (meson someI_ex) have 1: "y \ b \ \\<^sub>p" using b_def nth_pow_wits_exists[of m b] assms by (meson someI_ex) have 2: "pow_res m x = a" using a_def pow_res_classes_mem_eq[of m a x] assms 0 by blast have 3: "pow_res m y = b" using b_def pow_res_classes_mem_eq[of m b y] assms 1 by blast show ?thesis by (simp add: "2" "3" a_neq_b) qed lemma nth_pow_wits_disjoint_pow_res: assumes "m > 0" assumes "x \ nth_pow_wits m" assumes "y \ nth_pow_wits m" assumes "x \ y" shows "pow_res m x \ pow_res m y = {}" using assms nth_pow_wits_neq_pow_res disjoint_iff_not_equal by (metis (no_types, opaque_lifting) nth_pow_wits_closed(4) pow_res_classes_mem_eq) lemma nth_power_fact': assumes "0 < (n::nat)" shows "\m>0. \u\carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" using nth_power_fact[of n] assms by (metis less_one less_or_eq_imp_le linorder_neqE_nat neq0_conv) lemma equal_pow_res_criterion: assumes "N > 0" assumes "n > 0" assumes "\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "a = b \ (\ \ c)" assumes "val c \ N" shows "pow_res n a = pow_res n b" proof(cases "b = \") case True then have "a = \" using assms Qp.add.m_closed Qp.l_null Qp.one_closed by presburger then show ?thesis using True by blast next case False then have F0: "a \
b = \ \ c" by (metis Qp.Units_one_closed Qp.add.m_closed Qp.inv_cancelR(2) Qp.one_closed Qp.unit_factor assms(4) assms(5) assms(6) assms(7) field_inv(2) inv_in_frac(1)) have "0 < eint N" using assms by (metis eint_ord_simps(2) of_nat_0_less_iff zero_eint_def) hence F1: "val \ < val c" using assms less_le_trans[of 0 N "val c"] unfolding val_one by blast hence F2: " val \ = val (\ \ c)" using assms val_one one_nonzero Qp.add.m_comm Qp.one_closed val_ultrametric_noteq by metis have "val \ + eint (int N) \ val (\ \ (\ \ c))" proof- have "val (\ \ (\ \ c)) = val c" using Qp.add.inv_closed Qp.minus_eq Qp.minus_sum Qp.one_closed Qp.r_neg2 assms(6) val_minus by presburger thus ?thesis unfolding val_one using assms F1 by (metis add.left_neutral) qed hence F3: "ac N \ = ac N (\ \ c)" using F2 F1 assms ac_val[of \ "\ \ c" N] by (metis Qp.add.m_closed Qp.one_closed val_nonzero) have F4: "\ \ c \ P_set n" using assms F1 F2 F3 val_one ac_one by (metis Qp.add.m_closed Qp.one_closed Qp.one_nonzero ac_inv'' ac_inv'''(1) ac_one') then show ?thesis using assms(2) assms(4) assms(5) assms(7) equal_pow_resI' by blast qed lemma pow_res_nat_pow: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n b" shows "pow_res n (a[^](k::nat)) = pow_res n (b[^]k)" apply(induction k) using assms apply (metis Group.nat_pow_0) using assms pow_res_mult by (smt Qp.nat_pow_Suc2 Qp.nat_pow_closed) lemma pow_res_mult': assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "d \ carrier Q\<^sub>p" assumes "e \ carrier Q\<^sub>p" assumes "f \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n d" assumes "pow_res n b = pow_res n e" assumes "pow_res n c = pow_res n f" shows "pow_res n (a \ b \ c) = pow_res n (d \ e \ f)" proof- have "pow_res n (a \ b) = pow_res n (d \ e)" using pow_res_mult assms by meson then show ?thesis using pow_res_mult assms by (meson Qp.m_closed) qed lemma pow_res_disjoint: assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "a \ pow_res n \" shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" using assms unfolding pow_res_def using Qp.l_one Qp.nonzero_closed by blast lemma pow_res_disjoint': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "pow_res n a \ pow_res n \" shows "\ (\y \ nonzero Q\<^sub>p. a = y[^]n)" using assms pow_res_disjoint pow_res_refl by (metis pow_res_nth_pow) lemma pow_res_one_imp_nth_pow: assumes "n > 0" assumes "a \ pow_res n \" shows "\y \ nonzero Q\<^sub>p. a = y[^]n" using assms unfolding pow_res_def using Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed by blast lemma pow_res_eq: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ pow_res n a" shows "pow_res n b = pow_res n a" proof(cases "a = \") case True then show ?thesis using assms by (metis pow_res_zero singletonD) next case False then have a_nonzero: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast show ?thesis proof(cases "n = 1") case True then show ?thesis using a_nonzero assms using pow_res_one Q\<^sub>p_def Zp_def padic_fields_axioms by blast next case False then have "n \ 2" using assms(1) by linarith then show ?thesis using False a_nonzero assms Qp.nonzero_closed nonzero_pow_res equal_pow_resI by blast qed qed lemma pow_res_classes_n_eq_one: "pow_res_classes 1 = {nonzero Q\<^sub>p}" unfolding pow_res_classes_def using pow_res_one Qp.one_nonzero by blast lemma nth_pow_wits_closed': assumes "n > 0" assumes "x \ nth_pow_wits n" shows "x \ \\<^sub>p \ x \ nonzero Q\<^sub>p" using nth_pow_wits_closed assms by blast (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Sets Defined by Congruences\ (**************************************************************************************************) (**************************************************************************************************) (********************************************************************************************) (********************************************************************************************) subsubsection\$p$-adic ord Congruence Sets\ (********************************************************************************************) (********************************************************************************************) lemma carrier_is_univ_semialgebraic: "is_univ_semialgebraic (carrier Q\<^sub>p)" apply(rule is_univ_semialgebraicI) using Qp.to_R1_carrier carrier_is_semialgebraic by presburger lemma nonzero_is_univ_semialgebraic: "is_univ_semialgebraic (nonzero Q\<^sub>p)" proof- have "nonzero Q\<^sub>p = carrier Q\<^sub>p - {\}" unfolding nonzero_def by blast then show ?thesis using diff_is_univ_semialgebraic[of "carrier Q\<^sub>p" "{\}"] by (metis Diff_empty Diff_insert0 carrier_is_univ_semialgebraic empty_subsetI finite.emptyI finite.insertI finite_is_univ_semialgebraic insert_subset) qed definition ord_congruence_set where "ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x mod n = a}" lemma ord_congruence_set_nonzero: "ord_congruence_set n a \ nonzero Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq ord_congruence_set_def subsetI) lemma ord_congruence_set_closed: "ord_congruence_set n a \ carrier Q\<^sub>p" using nonzero_def ord_congruence_set_nonzero unfolding nonzero_def by (meson Qp.nonzero_closed ord_congruence_set_nonzero subset_iff) lemma ord_congruence_set_memE: assumes "x \ ord_congruence_set n a" shows "x \ nonzero Q\<^sub>p" "ord x mod n = a" using assms ord_congruence_set_nonzero apply blast by (metis (mono_tags, lifting) assms mem_Collect_eq ord_congruence_set_def) lemma ord_congruence_set_memI: assumes "x \ nonzero Q\<^sub>p" assumes "ord x mod n = a" shows "x \ ord_congruence_set n a" using assms by (metis (mono_tags, lifting) mem_Collect_eq ord_congruence_set_def) text\ We want to prove that ord\_congruence\_set is a finite union of semialgebraic sets, hence is also semialgebraic. \ lemma pow_res_ord_cong: assumes "x \ carrier Q\<^sub>p" assumes "x \ ord_congruence_set n a" shows "pow_res n x \ ord_congruence_set n a" proof fix y assume A: "y \ pow_res n x" show "y \ ord_congruence_set (int n) a" proof- obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]n)" using A pow_res_def[of n x] by blast have 0: "x \ nonzero Q\<^sub>p" using assms(2) ord_congruence_set_memE(1) by blast have 1: "y \ nonzero Q\<^sub>p" using A by (metis "0" Qp.integral Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero a_def not_nonzero_Qp) have 2: "ord y = ord x + n* ord a" using a_def 0 1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_mult by presburger show ?thesis apply(rule ord_congruence_set_memI) using assms ord_congruence_set_memE 2 1 apply blast using "2" assms(2) ord_congruence_set_memE(2) by presburger qed qed lemma pow_res_classes_are_univ_semialgebraic: shows "are_univ_semialgebraic (pow_res_classes n)" apply(rule are_univ_semialgebraicI) using pow_res_classes_univ_semialg by blast lemma ord_congruence_set_univ_semialg: assumes "n \ 0" shows "is_univ_semialgebraic (ord_congruence_set n a)" proof(cases "n = 0") case True have T0: "ord_congruence_set n a = {x \ nonzero Q\<^sub>p. ord x = a}" unfolding ord_congruence_set_def True by presburger have T1: "{x \ nonzero Q\<^sub>p. ord x = a} = {x \ nonzero Q\<^sub>p. val x = a}" apply(rule equalityI'') using val_ord apply blast using val_ord by (metis eint.inject) have T2: "{x \ nonzero Q\<^sub>p. val x = a} = {x \ carrier Q\<^sub>p. val x = a}" apply(rule equalityI'') using Qp.nonzero_closed apply blast by (metis iless_Suc_eq val_nonzero val_val_ring_prod zero_in_val_ring) show ?thesis unfolding T0 T1 T2 using univ_val_eq_set_is_univ_semialgebraic by blast next case False obtain F where F_def: "F = {S \ (pow_res_classes (nat n)). S \(ord_congruence_set n a) }" by blast have 0: "F \ pow_res_classes (nat n)" using F_def by blast have 1: "finite F" using 0 False nat_mono[of 1 n] nat_numeral[] pow_res_classes_finite[of "nat n"] rev_finite_subset by (smt assms nat_one_as_int) have 2: "are_univ_semialgebraic F" apply(rule are_univ_semialgebraicI) using 0 pow_res_classes_are_univ_semialgebraic by (metis (mono_tags) are_univ_semialgebraicE are_univ_semialgebraic_def assms nat_mono nat_numeral subset_iff) have 3: "\ F = (ord_congruence_set n a)" proof show "\ F \ ord_congruence_set n a" using F_def by blast show "ord_congruence_set n a \ \ F" proof fix x assume A: "x \ ord_congruence_set n a" have x_nonzero: "x \ nonzero Q\<^sub>p" using A ord_congruence_set_memE(1) by blast have 0: "pow_res (nat n) x \ F" using A pow_res_classes_def F_def by (smt nonzero_def assms mem_Collect_eq nat_0_le ord_congruence_set_memE(1) pow_res_ord_cong) have 1: "x \ pow_res (nat n) x" using False x_nonzero assms pow_res_refl[of x "nat n"] using Qp.nonzero_closed by blast show "x \ \ F" using 0 1 by blast qed qed then show ?thesis using "1" "2" finite_union_is_univ_semialgebraic' by fastforce qed lemma ord_congruence_set_is_semialg: assumes "n \ 0" shows "is_semialgebraic 1 (Qp_to_R1_set (ord_congruence_set n a))" using assms is_univ_semialgebraicE ord_congruence_set_univ_semialg by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Congruence Sets for the order of the Evaluation of a Polynomial\ (********************************************************************************************) (********************************************************************************************) lemma poly_map_singleton: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n [f] x = [(Qp_ev f x)]" unfolding poly_map_def poly_tuple_eval_def using assms by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') definition poly_cong_set where "poly_cong_set n f m a = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). (Qp_ev f x) \ \ \ (ord (Qp_ev f x) mod m = a)}" lemma poly_cong_set_as_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "poly_cong_set n f m a = poly_map n [f] \\<^bsub>n\<^esub>(Qp_to_R1_set (ord_congruence_set m a))" proof show "poly_cong_set n f m a \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" proof fix x assume A: "x \ poly_cong_set n f m a" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis (no_types, lifting) mem_Collect_eq poly_cong_set_def) have 1: "(Qp_ev f x) \ \ " by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) have 2: "(ord (Qp_ev f x) mod m = a)" by (metis (mono_tags, lifting) A mem_Collect_eq poly_cong_set_def) have 3: "(Qp_ev f x) \ (ord_congruence_set m a)" using "0" "1" "2" eval_at_point_closed assms not_nonzero_Qp ord_congruence_set_memI by metis show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" proof- have 00: "poly_map n [f] x = [(Qp_ev f x)]" using "0" assms poly_map_singleton by blast have 01: "[eval_at_point Q\<^sub>p x f] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using "0" assms eval_at_point_closed Qp.to_R1_closed by blast hence 02: "poly_map n [f] x \ (\a. [a]) ` ord_congruence_set m a" using 3 "00" by blast then show "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` ord_congruence_set m a)" using 0 unfolding evimage_def by blast qed qed show "poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a \ poly_cong_set n f m a" proof fix x assume A: "x \ poly_map n [f] \\<^bsub>n\<^esub> ((\a. [a]) ` (ord_congruence_set m a))" have 0: "((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using ord_congruence_set_closed Qp.to_R1_carrier by blast have "is_poly_tuple n [f]" using assms unfolding is_poly_tuple_def by (simp add: assms) then have 1:"poly_map n [f] \\<^bsub>n\<^esub>((\a. [a]) ` ord_congruence_set m a) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using 0 A assms One_nat_def by (metis extensional_vimage_closed) then have 2: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A unfolding evimage_def by blast then have 3: "poly_map n [f] x \ ((\a. [a]) ` ord_congruence_set m a)" using A assms 0 One_nat_def by blast have "poly_map n [f] x = [(Qp_ev f x)]" using "2" assms poly_map_singleton by blast then have "Qp_ev f x \ ord_congruence_set m a" using 3 by (metis (mono_tags, lifting) image_iff list.inject) then show "x \ poly_cong_set n f m a" unfolding poly_cong_set_def by (metis (mono_tags, lifting) "2" Qp.nonzero_memE(2) mem_Collect_eq ord_congruence_set_memE(1) ord_congruence_set_memE(2)) qed qed lemma singleton_poly_tuple: "is_poly_tuple n [f] \ f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" unfolding is_poly_tuple_def by (metis (no_types, lifting) list.distinct(1) list.set_cases list.set_intros(1) set_ConsD subset_code(1)) lemma poly_cong_set_is_semialgebraic: assumes "m \ 0" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n (poly_cong_set n f m a)" proof- have 0: "(\a. [a]) ` ord_congruence_set m a \ semialg_sets 1" using assms ord_congruence_set_is_semialg[of m a] unfolding is_semialgebraic_def by blast have 1: "length [f] = 1" by simp hence " poly_map n [f] \\<^bsub>n\<^esub> (\a. [a]) ` ord_congruence_set m a \ semialg_sets n" using 0 singleton_poly_tuple[of n f] zero_neq_one assms pullback_is_semialg[of n "[f]" 1 "(\a. [a]) ` ord_congruence_set m a"] unfolding is_semialgebraic_def by blast thus ?thesis using assms poly_cong_set_as_pullback[of f n m a] unfolding is_semialgebraic_def by presburger qed (********************************************************************************************) (********************************************************************************************) subsubsection\Congruence Sets for Angular Components\ (********************************************************************************************) (********************************************************************************************) text\If a set is a union of \n\-th power residues, then it is semialgebraic.\ lemma pow_res_union_imp_semialg: assumes "n \ 1" assumes "S \ nonzero Q\<^sub>p" assumes "\x. x \ S \ pow_res n x \ S" shows "is_univ_semialgebraic S" proof- obtain F where F_def: "F = {T. T \ pow_res_classes n \ T \ S}" by blast have 0: "F \ pow_res_classes n" using F_def by blast have 1: "finite F" using 0 pow_res_classes_finite[of n] assms(1) finite_subset by auto have 2: "are_univ_semialgebraic F" using 0 by (meson are_univ_semialgebraicE are_univ_semialgebraicI assms(1) pow_res_classes_are_univ_semialgebraic padic_fields_axioms subsetD) have 3: "S = \ F" proof show "S \ \ F" proof fix x assume A: "x \ S" then have "pow_res n x \ S" using assms(3) by blast then have "pow_res n x \ F" using A assms(2) F_def pow_res_classes_def by (smt mem_Collect_eq subsetD) then have "pow_res n x \ \ F" by blast then show "x \ \ F" using A assms(1) assms(2) pow_res_refl[of x n] unfolding nonzero_def by blast qed show "\ F \ S" using F_def by blast qed show ?thesis using 1 2 3 finite_union_is_univ_semialgebraic' by blast qed definition ac_cong_set1 where "ac_cong_set1 n y = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = ac n y}" lemma ac_cong_set1_is_univ_semialg: assumes "n > 0" assumes "b \ nonzero Q\<^sub>p" assumes "b \ \\<^sub>p" shows "is_univ_semialgebraic (ac_cong_set1 n b)" proof(cases "n = 1 \ p = 2") case True have "(ac_cong_set1 n b) = nonzero Q\<^sub>p" proof have 0: "Units (Zp_res_ring n) = {1}" proof show "Units (Zp_res_ring n) \ {1}" proof fix x assume A: "x \ Units (Zp_res_ring n)" have 0: "carrier (Zp_res_ring n) = {0..(int 2) - 1}" using True by (metis assms(1) int_ops(3) p_residues power_one_right residues.res_carrier_eq) have 1: "carrier (Zp_res_ring n) = {0..(1::int)}" proof- have "int 2 - 1 = (1::int)" by linarith then show ?thesis using 0 by presburger qed have 15: "{0..(1::int)} = {0, (1::int)}" using atLeastAtMostPlus1_int_conv [of 0 "0::int"] by (smt atLeastAtMost_singleton insert_commute) have 2: "carrier (Zp_res_ring n) = {0,(1::int)}" using "1" "15" by blast have 3: "0 \ Units (Zp_res_ring n)" using True zero_not_in_residue_units by blast have "x \ carrier (Zp_res_ring n)" using A unfolding Units_def by blast then have "x = 1" using A 2 3 by (metis "1" atLeastAtMost_iff atLeastatMost_empty atLeastatMost_empty_iff2 linorder_neqE_linordered_idom mod_by_1 mod_pos_pos_trivial ) then show "x \ {1}" by simp qed show "{1} \ Units (Zp_res_ring n)" by (meson assms(1) empty_subsetI insert_subset residue_1_unit(1)) qed show "ac_cong_set1 n b \ nonzero Q\<^sub>p" by (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) show "nonzero Q\<^sub>p \ ac_cong_set1 n b" proof fix x assume A: "x \ nonzero Q\<^sub>p" then have P0: "ac n x = 1" using 0 ac_units assms(1) by blast have P1: "ac n b = 1" using assms 0 ac_units assms(1) by blast then have "ac n x = ac n b" using P0 by metis then show " x \ ac_cong_set1 n b" unfolding ac_cong_set1_def using A proof - have "x \ {r \ carrier Q\<^sub>p. r \ \}" by (metis (no_types) \x \ nonzero Q\<^sub>p\ nonzero_def ) then show "x \ {r \ carrier Q\<^sub>p. r \ \ \ ac n r = ac n b}" using \ac n x = ac n b\ by force qed qed qed then show "is_univ_semialgebraic (ac_cong_set1 n b)" by (simp add: nonzero_is_univ_semialgebraic) next case F: False have F0: "2 \ card (Units (Zp_res_ring n))" proof(cases "n = 1") case True then have "field (Zp_res_ring n)" using p_res_ring_1_field by blast then have F00: "Units (Zp_res_ring n) = carrier (Zp_res_ring n) - {\\<^bsub>Zp_res_ring n\<^esub>}" using field.field_Units by blast have F01: "\\<^bsub>Zp_res_ring n\<^esub> \ carrier (Zp_res_ring n)" using assms(1) cring.cring_simprules(2) padic_integers.R_cring padic_integers_axioms by blast have F02: "card (carrier (Zp_res_ring n)) = p \ finite (carrier (Zp_res_ring n))" by (smt F01 True nat_eq_iff2 p_res_ring_zero p_residue_ring_car_memE(1) power_one_right residue_ring_card) have F03: "\\<^bsub>residue_ring (p ^ n)\<^esub> \ carrier (residue_ring (p ^ n)) " using F01 by blast have F04: "int (card (carrier (residue_ring (p ^ n)))) \ int (card {\\<^bsub>residue_ring (p ^ n)\<^esub>}) " by (smt F02 F03 nat_int of_nat_0_le_iff of_nat_1 of_nat_power p_res_ring_0 p_res_ring_zero p_residue_ring_car_memE(1) power_increasing power_one_right residue_ring_card) have "card (carrier (residue_ring (p ^ n))) - 1 = p - 1" using F02 prime by (metis Totient.of_nat_eq_1_iff True less_imp_le_nat less_one nat_int nat_less_eq_zless of_nat_1 of_nat_diff of_nat_zero_less_power_iff p_residues pos_int_cases power_0 power_one_right residue_ring_card residues.m_gt_one zero_le_one) hence F05: "card (carrier (residue_ring (p ^ n)) - {\\<^bsub>residue_ring (p ^ n)\<^esub>}) = p - 1" using F02 F03 F04 card_Diff_singleton_if[of "(carrier (Zp_res_ring n))" "\\<^bsub>residue_ring (p^n)\<^esub>"] True int_ops(6)[of "card (carrier (residue_ring (p ^ n)))" "card {\\<^bsub>residue_ring (p ^ n)\<^esub>}"] p_res_ring_zero p_residue_ring_car_memE(1) by (metis) hence F06: "card (Units (Zp_res_ring n)) = p -1" using True F02 F01 F00 by (metis p_res_ring_zero) have F04: "p - 1 \2 " using F prime by (meson True linorder_cases not_less prime_ge_2_int zle_diff1_eq) then show ?thesis using F03 F06 by linarith next case False then show ?thesis by (metis assms(1) less_imp_le_nat mod2_gr_0 mod_less nat_le_linear nat_neq_iff residue_units_card_geq_2) qed show ?thesis apply(rule pow_res_union_imp_semialg[of "card (Units (Zp_res_ring n))"]) using F0 assms apply linarith apply (metis (mono_tags, lifting) ac_cong_set1_def mem_Collect_eq not_nonzero_Qp subsetI) proof- fix x assume AA: "x \ ac_cong_set1 n b" show "pow_res (card (Units (Zp_res_ring n))) x \ ac_cong_set1 n b" proof fix y assume A: "y \ pow_res (card (Units (Zp_res_ring n))) x" show "y \ ac_cong_set1 n b" proof- obtain k where k_def: "k = card (Units (Zp_res_ring n))" by blast have "k \2" using assms k_def F F0 by blast then obtain a where a_def: "a \ nonzero Q\<^sub>p \ y = x \ (a[^]k)" using k_def A pow_res_def[of k x] by blast have 0: "x \ nonzero Q\<^sub>p" using AA ac_cong_set1_def by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) have 1: "y \ nonzero Q\<^sub>p" by (metis "0" Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero \\thesis. (\a. a \ nonzero Q\<^sub>p \ y = x \ a [^] k \ thesis) \ thesis\) have "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)" using a_def 0 1 Qp_nat_pow_nonzero ac_mult' by blast then have 2: "ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> (ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k" proof- have "ac n (a[^]k) = ac n a [^]\<^bsub>Zp_res_ring n\<^esub> k" using a_def assms(1) ac_nat_pow'[of a n k] by linarith then show ?thesis using \ac n y = ac n x \\<^bsub>Zp_res_ring n\<^esub> ac n (a[^]k)\ by presburger qed then have "ac n y = ac n x" proof- have "(ac n a) \ Units (Zp_res_ring n)" by (metis (mono_tags, opaque_lifting) a_def ac_units assms(1) ) then have "(ac n a)^k mod (p^n) = 1" using k_def a_def ac_nat_pow ac_nat_pow' assms(1) residue_units_nilpotent using neq0_conv by presburger then have 00: "(ac n a)[^]\<^bsub>Zp_res_ring n\<^esub> k = 1" by (metis a_def ac_nat_pow ac_nat_pow' mod_by_1 power_0 zero_neq_one) have "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x \\<^bsub>Zp_res_ring n\<^esub> \\<^bsub>Zp_res_ring n\<^esub>" using 00 assms(1) p_res_ring_one by presburger hence "ac n x \\<^bsub>residue_ring (p ^ n)\<^esub> ac n a [^]\<^bsub>residue_ring (p ^ n)\<^esub> k = ac n x" by (metis "0" Qp.nonzero_closed Qp.one_nonzero Qp.r_one ac_mult' ac_one' assms(1)) then show ?thesis using 2 "0" 00 by linarith qed then show ?thesis using "1" AA nonzero_def ac_cong_set1_def[of n b] mem_Collect_eq by smt qed qed qed qed definition ac_cong_set where "ac_cong_set n k = {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" lemma ac_cong_set_is_univ_semialg: assumes "n >0 " assumes "k \ Units (Zp_res_ring n)" shows "is_univ_semialgebraic (ac_cong_set n k)" proof- have "k \ carrier (Zp_res_ring n)" using assms(2) Units_def[of "Zp_res_ring n"] by blast then have k_n: "([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) n = k" using assms by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) obtain b where b_def: "b = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" by blast have 0: "k mod p \ 0" using assms residue_UnitsE[of n k] by (metis le_eq_less_or_eq le_refl less_one nat_le_linear p_residues power_0 power_one_right residues.mod_in_res_units residues_def zero_less_one zero_neq_one zero_not_in_residue_units zero_power) then have "val_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = 0" using val_Zp_p_int_unit by blast then have 1: "val b = 0" by (metis Zp_int_inc_closed b_def val_of_inc) have 2: "b \ \\<^sub>p" using b_def Zp_int_mult_closed by blast have "ord_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0" using 0 ord_Zp_p_int_unit by blast have "ac_Zp ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using "0" Zp_int_inc_closed ac_Zp_of_Unit ord_Zp_p_int_unit \val_Zp ([k] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = 0\ by blast then have "(angular_component b) = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using b_def 1 2 angular_component_ord_zero[of b] by (metis Qp.int_inc_zero Qp.one_closed val_ring_memE Zp.int_inc_zero Zp.one_closed Zp.one_nonzero Zp_int_inc_closed angular_component_of_inclusion inc_closed inc_of_int inc_of_one inc_to_Zp local.val_zero not_nonzero_Qp val_ineq val_one zero_in_val_ring) then have "ac n b = k" using ac_def[of n b] k_n by (metis Qp_char_0_int Zp_defs(1) ac_def b_def inc_of_int inc_of_one) then have 3: "(ac_cong_set n k) = (ac_cong_set1 n b)" unfolding ac_cong_set_def ac_cong_set1_def by meson have 4: "b \ nonzero Q\<^sub>p" using 1 2 val_nonzero by (metis Qp.one_closed val_ring_memE Zp_def \_def local.one_neq_zero not_nonzero_Qp padic_fields.val_ring_memE padic_fields_axioms val_ineq val_one) then show ?thesis using 1 2 3 assms ac_cong_set1_is_univ_semialg[of n b] val_nonzero[of b 1] by presburger qed definition val_ring_constant_ac_set where "val_ring_constant_ac_set n k = {a \ \\<^sub>p. val a = 0 \ ac n a = k}" lemma val_nonzero': assumes "a \ carrier Q\<^sub>p" assumes "val a = eint k" shows "a \ nonzero Q\<^sub>p" using val_nonzero[of a "k + 1"] by (metis Suc_ile_eq assms(1) assms(2) eint_ord_code(3) val_nonzero) lemma val_ord': assumes "a \ carrier Q\<^sub>p" assumes "a \\" shows "val a = ord a" by (meson assms(1) assms(2) not_nonzero_Qp val_ord) lemma val_ring_constant_ac_set_is_univ_semialgebraic: assumes "n > 0" assumes "k \ 0" shows "is_univ_semialgebraic (val_ring_constant_ac_set n k)" proof(cases "val_ring_constant_ac_set n k = {}") case True then show ?thesis by (metis equals0D order_refl pow_res_union_imp_semialg subsetI) next case False then obtain b where b_def: "b \ val_ring_constant_ac_set n k" by blast have 0: "val_ring_constant_ac_set n k = q_ball n k 0 \" proof show "val_ring_constant_ac_set n k \ q_ball n k 0 \" proof fix x assume A: "x \ val_ring_constant_ac_set n k" then show "x \ q_ball n k 0 \" proof- have 0: "x \ \\<^sub>p \ val x = 0 \ ac n x = k" using A unfolding val_ring_constant_ac_set_def by blast then have x_car: "x \ carrier Q\<^sub>p" using val_ring_memE by blast then have 00: "x = x \ \" using Qp.ring_simprules by metis then have 1: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" using 0 by presburger have 2: "val (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" using 0 00 by metis have 3: "x \ nonzero Q\<^sub>p" proof(rule ccontr) assume " x \ nonzero Q\<^sub>p " then have "x = \" using Qp.nonzero_memI x_car by blast then show False using 0 val_zero by (metis ac_def assms(2)) qed have 4: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" proof(rule ccontr) assume "ord (x \ \) \ 0" then have "val (x \ \) \ 0" by (metis "00" "3" Qp.one_closed equal_val_imp_equal_ord(1) ord_one val_one) then show False using "2" by blast qed show ?thesis using 0 1 4 unfolding q_ball_def using x_car by blast qed qed show "q_ball n k 0 \ \ val_ring_constant_ac_set n k" proof fix x assume A: "x \ q_ball n k 0 \" then have 0: "ac n (x \\<^bsub>Q\<^sub>p\<^esub> \) = k" using q_ballE'(1) by blast have 1: "ord (x \\<^bsub>Q\<^sub>p\<^esub> \) = 0" using q_ball_def A by blast have 2: "x \ carrier Q\<^sub>p" using A q_ball_def by blast have 3: "ord x = 0" using 2 1 ring.ring_simprules[of Q\<^sub>p] by (metis Qp.ring_axioms) have 4: "ac n x = k" using 0 2 1 cring.axioms(1)[of Q\<^sub>p] ring.ring_simprules[of Q\<^sub>p] by (metis Qp.ring_axioms) have 5: "x \ \\<^sub>p" using Qp_val_ringI[of x] 2 3 val_ord val_nonzero' by (metis Qp.integral_iff val_ring_memE Zp.nonzero_closed angular_component_closed angular_component_ord_zero image_eqI local.numer_denom_facts(1) local.numer_denom_facts(2) local.numer_denom_facts(4) not_nonzero_Qp) have 6: "x \ \" using 4 assms ac_def[of n x] by meson have 7: "val x = 0" using 6 3 2 assms val_ord' zero_eint_def by presburger show " x \ val_ring_constant_ac_set n k" unfolding val_ring_constant_ac_set_def using 7 6 5 4 by blast qed qed obtain b where b_def: "b \ q_ball n k (0::int) \" using "0" b_def by blast have 1: "b \ carrier Q\<^sub>p \ ac n b = k" using b_def unfolding q_ball_def by (metis (mono_tags, lifting) "0" b_def mem_Collect_eq val_ring_constant_ac_set_def) then have 2: "b \ nonzero Q\<^sub>p" using 1 assms by (metis ac_def not_nonzero_Qp) have "q_ball n k 0 \ = B\<^bsub>0 + int n\<^esub>[b]" using 1 b_def nonzero_def [of Q\<^sub>p] assms 0 2 c_ball_q_ball[of b n k "\" b 0] by (meson Qp.cring_axioms cring.cring_simprules(2)) then have "is_univ_semialgebraic (q_ball n k (0::int) \) " using 1 ball_is_univ_semialgebraic[of b "0 + int n"] by metis then show ?thesis using 0 by presburger qed definition val_ring_constant_ac_sets where "val_ring_constant_ac_sets n = val_ring_constant_ac_set n ` (Units (Zp_res_ring n))" lemma val_ring_constant_ac_sets_are_univ_semialgebraic: assumes "n > 0" shows "are_univ_semialgebraic (val_ring_constant_ac_sets n)" proof(rule are_univ_semialgebraicI) have 0: "\ coprime 0 p" using coprime_0_right_iff[of p] coprime_commute[of p 0] coprime_int_iff[of "nat p" 0] nat_dvd_1_iff_1 prime_gt_1_nat zdvd1_eq by (metis not_prime_unit prime) have "(0::int) \(Units (Zp_res_ring n))" apply(rule ccontr) using 0 assms residues.cring[of "p ^ n"] unfolding residues_def by (smt less_one not_gr_zero power_le_imp_le_exp power_less_imp_less_exp residue_UnitsE) fix x assume A: "x \ val_ring_constant_ac_sets n" then obtain k where k_def: "x = val_ring_constant_ac_set n k \ k \ Units (Zp_res_ring n)" by (metis image_iff val_ring_constant_ac_sets_def) then show "is_univ_semialgebraic x" using assms by (metis \0 \ Units (Zp_res_ring n)\ val_ring_constant_ac_set_is_univ_semialgebraic) qed definition ac_cong_set3 where "ac_cong_set3 n = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = ac n b) \ as = [a, b] }" definition ac_cong_set2 where "ac_cong_set2 n k = {as. \ a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k \ as = [a, b] }" lemma ac_cong_set2_cartesian_product: assumes "k \ Units (Zp_res_ring n)" assumes "n > 0" shows "ac_cong_set2 n k = cartesian_product (to_R1` (ac_cong_set n k)) (to_R1` (val_ring_constant_ac_set n k))" proof show "ac_cong_set2 n k \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" proof fix x assume A: "x \ ac_cong_set2 n k" show "x \ (cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k))" unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 1]) apply (metis (mono_tags, lifting) mem_Collect_eq subsetI Qp.to_R1_car_subset) apply (metis (no_types, lifting) val_ring_memE mem_Collect_eq subsetI Qp.to_R1_car_subset) proof- obtain a b where ab_def: "x = [a,b] \ a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ (ac n a = k) \ (ac n b) = k" using A unfolding ac_cong_set_def val_ring_constant_ac_set_def ac_cong_set2_def by blast have 0: "take 1 x = [a]" by (simp add: ab_def) have 1: "drop 1 x = [b]" by (simp add: ab_def) have 2: "a \ {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" using ab_def nonzero_def by (smt mem_Collect_eq) have 3: "b \ {a \ \\<^sub>p. val a = 0 \ ac n a = k}" using ab_def by blast show "take 1 x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. x \ \ \ ac n x = k}" using 0 2 by blast show "drop 1 x \ (\a. [a]) ` {a \ \\<^sub>p. val a = 0 \ ac n a = k}" using 1 3 by blast qed qed show "cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k) \ ac_cong_set2 n k" proof fix x have 0: "(\a. [a]) ` ac_cong_set n k \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms by (metis (no_types, lifting) ac_cong_set_def mem_Collect_eq subsetI Qp.to_R1_car_subset) have 1: "((\a. [a]) ` val_ring_constant_ac_set n k) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" by (smt val_ring_memE mem_Collect_eq subsetI Qp.to_R1_carrier Qp.to_R1_subset val_ring_constant_ac_set_def) assume A: "x \ cartesian_product ((\a. [a]) ` ac_cong_set n k) ((\a. [a]) ` val_ring_constant_ac_set n k)" then have "length x = 2" using 0 1 A cartesian_product_closed[of "((\a. [a]) ` ac_cong_set n k)" Q\<^sub>p 1 "((\a. [a]) ` val_ring_constant_ac_set n k)" 1] by (metis (no_types, lifting) cartesian_power_car_memE one_add_one subset_iff) then obtain a b where ab_def: "take 1 x = [a] \ drop 1 x = [b]" by (metis One_nat_def add_diff_cancel_left' drop0 drop_Cons_numeral numerals(1) pair_id plus_1_eq_Suc take0 take_Cons_numeral) have 2: " a \ (ac_cong_set n k) \ b \ (val_ring_constant_ac_set n k)" proof- have P0: "take 1 x \ (\a. [a]) ` ac_cong_set n k" using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] by blast have P1: "drop 1 x \ (\a. [a]) ` val_ring_constant_ac_set n k" using 0 A cartesian_product_memE[of x "((\a. [a]) ` ac_cong_set n k) " " ((\a. [a]) ` val_ring_constant_ac_set n k)" Q\<^sub>p 1] by blast have P2: "[a] \ (\a. [a]) ` ac_cong_set n k" using P0 ab_def by metis have P3: "[b] \ (\a. [a]) ` val_ring_constant_ac_set n k" using P1 ab_def by metis show ?thesis using P2 P3 by blast qed have 3: "a \ nonzero Q\<^sub>p" using 2 assms nonzero_def [of Q\<^sub>p] ac_cong_set_def[of n k] by blast have 4: "x = [a,b]" by (metis (no_types, lifting) \length x = 2\ ab_def less_numeral_extra(1) nth_Cons_0 nth_take nth_via_drop pair_id) then have "\a b. a \ nonzero Q\<^sub>p \ b \ \\<^sub>p \ val b = 0 \ ac n a = k \ ac n b = k \ x = [a, b]" using 2 3 ab_def unfolding val_ring_constant_ac_set_def ac_cong_set_def by blast then show "x \ ac_cong_set2 n k" unfolding ac_cong_set2_def val_ring_constant_ac_set_def ac_cong_set_def by blast qed qed lemma ac_cong_set2_is_semialg: assumes "k \ Units (Zp_res_ring n)" assumes "n > 0" shows "is_semialgebraic 2 (ac_cong_set2 n k)" using ac_cong_set_is_univ_semialg ac_cong_set2_cartesian_product[of k n] cartesian_product_is_semialgebraic[of 1 "((\a. [a]) ` ac_cong_set n k)" 1 " ((\a. [a]) ` val_ring_constant_ac_set n k)"] by (metis assms(1) assms(2) is_univ_semialgebraicE less_one less_or_eq_imp_le nat_neq_iff one_add_one val_ring_constant_ac_set_is_univ_semialgebraic zero_not_in_residue_units) lemma ac_cong_set3_as_union: assumes "n > 0" shows "ac_cong_set3 n = \ (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" proof show "ac_cong_set3 n \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" proof fix x assume A: "x \ ac_cong_set3 n" then have 0: "x \ (ac_cong_set2 n (ac n (x!0)))" unfolding ac_cong_set2_def ac_cong_set3_def by (smt mem_Collect_eq nth_Cons_0) have 1: "(ac n (x!0)) \ Units (Zp_res_ring n)" using A unfolding ac_cong_set3_def by (smt ac_units assms mem_Collect_eq nth_Cons_0) then show "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" using 0 by blast qed show "\ (ac_cong_set2 n ` Units (Zp_res_ring n)) \ ac_cong_set3 n" proof fix x assume A: "x \ \ (ac_cong_set2 n ` Units (Zp_res_ring n))" obtain k where k_def: "x \ (ac_cong_set2 n k) \ k \ (Units (Zp_res_ring n))" using A by blast have 0: "k mod p \ 0" using k_def One_nat_def Suc_le_eq assms less_numeral_extra(1) power_one_right residues.m_gt_one residues.mod_in_res_units by (metis p_residues residue_UnitsE zero_not_in_residue_units) obtain b where b_def: "b = ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" by blast have "k \0" using 0 mod_0 by blast then have 1: "b \ nonzero Z\<^sub>p" using 0 b_def int_unit by (metis Zp.Units_nonzero Zp.zero_not_one) have 10: "ord_Zp b = 0" using 0 1 using b_def ord_Zp_p_int_unit by blast have 2: "\ b \ nonzero Q\<^sub>p" using k_def using "1" inc_of_nonzero by blast have 3: "angular_component (\ b) = ac_Zp b" using "1" angular_component_of_inclusion by blast have 4: "ac_Zp b = b" using 1 10 by (metis "3" Zp.r_one ac_Zp_factors' angular_component_closed inc_of_nonzero int_pow_0 mult_comm ord_Zp_def) have 5: "ac_Zp b n = k" proof- have "k \ carrier (Zp_res_ring n)" using k_def unfolding Units_def by blast then show ?thesis using b_def k_def 4 Zp_int_inc_res mod_pos_pos_trivial by (metis p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) qed then have "ac n (\ b) = k" using 10 1 2 3 4 unfolding ac_def using Qp.not_nonzero_memI by metis then show "x \ ac_cong_set3 n" unfolding ac_cong_set3_def using k_def unfolding ac_cong_set2_def by (smt mem_Collect_eq) qed qed lemma ac_cong_set3_is_semialgebraic: assumes "n > 0" shows "is_semialgebraic 2 (ac_cong_set3 n)" proof- have 0: "finite (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" using assms residues.finite_Units[of "p^n"] unfolding residues_def using p_residues residues.finite_Units by blast have 1: "are_semialgebraic 2 (ac_cong_set2 n ` (Units (Zp_res_ring n)) )" apply(rule are_semialgebraicI) using ac_cong_set2_is_semialg assms by blast show ?thesis using 0 1 ac_cong_set3_as_union by (metis (no_types, lifting) are_semialgebraicE assms finite_union_is_semialgebraic' is_semialgebraicE subsetI) qed (**************************************************************************************************) (**************************************************************************************************) subsection\Permutations of indices of semialgebraic sets\ (**************************************************************************************************) (**************************************************************************************************) lemma fun_inv_permute: assumes "\ permutes {.. permutes {.. \ (fun_inv \) = id" "(fun_inv \) \ \ = id" using assms unfolding fun_inv_def using permutes_inv apply blast using assms permutes_inv_o(1) apply blast using assms permutes_inv_o(2) by blast lemma poly_tuple_pullback_eq_poly_map_vimage: assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "S \ carrir (Q\<^sub>p\<^bsup>m\<^esup>)" shows "poly_map n fs \\<^bsub>n\<^esub> S = poly_tuple_pullback n S fs" unfolding poly_map_def poly_tuple_pullback_def evimage_def restrict_def using assms by (smt vimage_inter_cong) lemma permutation_is_semialgebraic: assumes "is_semialgebraic n S" assumes "\ permutes {.. ` S)" proof- have "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms gen_boolean_algebra_subset is_semialgebraic_def semialg_sets_def by blast then have "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" using Qp.cring_axioms assms pullback_by_permutation_of_poly_list'[of \ n S] unfolding poly_map_def by blast then have 0: "(permute_list \ ` S) = poly_tuple_pullback n S (permute_list (fun_inv \) (pvar_list Q\<^sub>p n))" using poly_tuple_pullback_def by blast have 1: "(fun_inv \) permutes {..) (pvar_list Q\<^sub>p n))"] permutation_of_poly_list_is_poly_list[of n "(pvar_list Q\<^sub>p n)" "fun_inv \"] pvar_list_is_poly_tuple[of n] assms poly_tuple_pullback_eq_poly_map_vimage by (metis "0" \S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ is_semialgebraic_def length_permute_list pvar_list_length) qed lemma permute_list_closed: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) using assms cartesian_power_car_memE length_permute_list apply blast using assms cartesian_power_car_memE'' permute_list_set by blast lemma permute_list_closed': assumes "\ permutes {.. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) apply (metis assms(2) cartesian_power_car_memE length_permute_list) using assms cartesian_power_car_memE'[of "permute_list \ a" Q\<^sub>p n] by (metis cartesian_power_car_memE in_set_conv_nth length_permute_list set_permute_list subsetI) lemma permute_list_compose_inv: assumes "\ permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "permute_list \ (permute_list (fun_inv \) a) = a" "permute_list (fun_inv \) (permute_list \ a) = a" using assms apply (metis cartesian_power_car_memE fun_inv_permute(3) permute_list_compose permute_list_id) using assms by (metis cartesian_power_car_memE fun_inv_permute(2) fun_inv_permute(1) permute_list_compose permute_list_id) lemma permutation_is_semialgebraic_imp_is_semialgebraic: assumes "is_semialgebraic n (permute_list \ ` S)" assumes "\ permutes {..) ` (permute_list \ ` S) = S" proof- have 0: "(permute_list \ ` S) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms unfolding is_semialgebraic_def semialg_sets_def using gen_boolean_algebra_subset by blast have 1: "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" proof fix x assume "x \ S" then show "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using 0 assms by (meson image_subset_iff permute_list_closed') qed show ?thesis proof show "permute_list (fun_inv \) ` permute_list \ ` S \ S" using 0 assms permute_list_compose_inv[of \] "1" image_iff image_subset_iff subsetD by smt show "S \ permute_list (fun_inv \) ` permute_list \ ` S" using 0 assms permute_list_compose_inv[of \] by (smt "1" image_iff subset_eq) qed qed then show ?thesis using permutation_is_semialgebraic by (metis assms(1) assms(2) fun_inv_permute(1)) qed lemma split_cartesian_product_is_semialgebraic: assumes "i \ n" assumes "is_semialgebraic n A" assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (split_cartesian_product n m i A B)" using assms cartesian_product_is_semialgebraic scp_permutes[of i n m] permutation_is_semialgebraic[of "n + m" "cartesian_product A B" "(scp_permutation n m i)"] unfolding split_cartesian_product_def by blast definition reverse_val_relation_set where "reverse_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)}" lemma Qp_2_car_memE: assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" shows "x = [x!0, x!1]" proof- have "length x = 2" using assms cartesian_power_car_memE by blast then show ?thesis using pair_id by blast qed definition flip where "flip = (\i::nat. (if i = 0 then 1 else (if i = 1 then 0 else i)))" lemma flip_permutes: "flip permutes {0,1}" unfolding permutes_def flip_def by (smt mem_simps(1)) lemma flip_eval: "flip 0 = 1" "flip 1 = 0" unfolding flip_def by auto lemma flip_x: assumes "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" shows "permute_list flip x = [x!1, x!0]" proof- have 0: "x = [x!0, x!1]" using assms Qp_2_car_memE by blast have 1: "length (permute_list flip x) = length [x!1, x!0]" using 0 unfolding permute_list_def by (metis length_Cons length_map map_nth) have 2: "\i. i < 2 \ permute_list flip x ! i = [x!1, x!0] ! i" proof- fix i::nat assume A: "i < 2" show "permute_list flip x ! i = [x!1, x!0] ! i" using 0 unfolding permute_list_def by (smt flip_eval(1) flip_eval(2) length_Cons length_greater_0_conv list.simps(8) map_upt_Suc numeral_nat(7) upt_rec) qed have "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" proof- have 0: "length x = 2" using assms cartesian_power_car_memE by blast show "\i. i < length x \ permute_list flip x ! i = [x!1, x!0] ! i" using 2 unfolding 0 by blast qed thus ?thesis using 1 by (metis length_permute_list nth_equalityI) qed lemma permute_with_flip_closed: assumes "x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" shows "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2::nat\<^esup>)" apply(rule permute_list_closed) using assms apply blast proof- have "{0::nat, 1} = {..<2::nat}" by auto thus "flip permutes {..<2}" using flip_permutes by auto qed lemma reverse_val_relation_set_semialg: "is_semialgebraic 2 reverse_val_relation_set" proof- have 1: "reverse_val_relation_set = permute_list flip ` val_relation_set" apply(rule equalityI') proof- show " \x. x \ reverse_val_relation_set \ x \ permute_list flip ` val_relation_set" proof- fix x assume A: "x \ reverse_val_relation_set" have 0: "permute_list flip x = [x ! 1, x ! 0]" using flip_x[of x] A unfolding reverse_val_relation_set_def by blast have 1: "permute_list flip x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule permute_with_flip_closed) using A unfolding reverse_val_relation_set_def by blast have 2: "permute_list flip x \ val_relation_set" using 1 A unfolding 0 reverse_val_relation_set_def val_relation_set_def mem_Collect_eq by (metis Qp_2_car_memE list_hd list_tl) show "x \ permute_list flip ` val_relation_set" using flip_x[of x] A unfolding reverse_val_relation_set_def val_relation_set_def mem_Collect_eq by (metis (no_types, lifting) "1" "2" Qp_2_car_memE flip_x image_eqI list_tl nth_Cons_0 val_relation_set_def) qed show "\x. x \ permute_list flip ` val_relation_set \ x \ reverse_val_relation_set" proof- fix x assume a: " x \ permute_list flip ` val_relation_set" then obtain y where y_def: "y \ val_relation_set \x = permute_list flip y" by blast have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using y_def basic_semialg_set_memE(1) val_relation_semialg by blast have y_length: " length y = 2" using y_def basic_semialg_set_memE val_relation_semialg by (metis cartesian_power_car_memE) obtain a b where ab_def: "y = [a,b]" using y_length pair_id by blast have 0: "a = y!0" using ab_def by (metis nth_Cons_0) have 1: "b = y!1" using ab_def by (metis cancel_comm_monoid_add_class.diff_cancel eq_numeral_extra(2) nth_Cons') have a_closed: "a \ carrier Q\<^sub>p" using 0 y_closed unfolding 0 by (meson cartesian_power_car_memE' rel_simps(75) zero_order(5)) have b_closed: "b \ carrier Q\<^sub>p" proof- have "1 < (2::nat)" by linarith thus ?thesis using y_closed unfolding 1 by (meson cartesian_power_car_memE') qed have 2: "x = [b, a]" using flip_x[of y] y_def y_closed unfolding ab_def unfolding 0 1 using \y \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ permute_list flip y = [y ! 1, y ! 0]\ y_closed y_def by presburger have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using y_def unfolding val_relation_set_def using permute_with_flip_closed[of y] by blast show " x \ reverse_val_relation_set" using x_closed y_def unfolding val_relation_set_def reverse_val_relation_set_def mem_Collect_eq 2 0 1 by (metis Qp_2_car_memE list_hd list_tl) qed qed show ?thesis unfolding 1 apply(rule permutation_is_semialgebraic) using val_relation_is_semialgebraic apply blast using flip_permutes by (metis Suc_1 insert_commute lessThan_0 lessThan_Suc numeral_nat(7)) qed definition strict_val_relation_set where "strict_val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) < val (as ! 1)}" definition val_diag where "val_diag = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)}" lemma val_diag_semialg: "is_semialgebraic 2 val_diag" proof- have "val_diag = val_relation_set \reverse_val_relation_set" apply(rule equalityI') apply(rule IntI) unfolding val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq apply simp apply simp apply(erule IntE) unfolding mem_Collect_eq using basic_trans_rules(24) by blast then show ?thesis using intersection_is_semialg by (simp add: reverse_val_relation_set_semialg val_relation_is_semialgebraic) qed lemma strict_val_relation_set_is_semialg: "is_semialgebraic 2 strict_val_relation_set" proof- have 0: "strict_val_relation_set = reverse_val_relation_set - val_diag" apply(rule equalityI') apply(rule DiffI) unfolding strict_val_relation_set_def val_diag_def val_relation_set_def reverse_val_relation_set_def mem_Collect_eq using order_le_less apply blast proof show "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1) \ x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) = val (x ! 1) \ False" using order_less_le by blast show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) \ val (as ! 1)} - {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as ! 0) = val (as ! 1)} \ x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ val (x ! 0) < val (x ! 1)" apply(erule DiffE) unfolding mem_Collect_eq using order_le_less by blast qed show ?thesis unfolding 0 apply(rule diff_is_semialgebraic ) using reverse_val_relation_set_semialg apply blast using val_diag_semialg by blast qed lemma singleton_length: "length [a] = 1" by auto lemma take_closed': assumes "m > 0" assumes "x \ carrier (Q\<^sub>p\<^bsup>m+l\<^esup>)" shows "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule take_closed[of m "m+l"]) apply simp using assms by blast lemma triple_val_ineq_set_semialg: shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}" proof- have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule equalityI') show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!0, x!1]" by blast have a_length: "length a = 2" proof- have "a = x!0 #[x!1]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] by presburger qed obtain b where b_def: "b = [x!2]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule cartesian_power_car_memI') apply (simp add: a_length) unfolding 0 using A unfolding mem_Collect_eq using cartesian_power_car_memE' by fastforce show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] by simp have 2: "x = a@b" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of a b] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (a @ b) ! i" apply(cases "i = 0") apply (metis a_def append.simps(2) nth_Cons_0) apply(cases "(i:: nat) = 1") apply (simp add: a_def) proof- assume a: "i \0" "i \ 1" then have "i = 2" using A1 by presburger thus ?thesis by (metis a_length b_def nth_append_length) qed qed have 3: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) apply blast using 3 a_closed apply blast proof- have "drop 2 x = b" unfolding 2 unfolding 3 using 0 by simp then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using b_closed by blast qed qed show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast have "(0::nat)< 2" by presburger hence 0: "x!0 = a!0" unfolding ab_def using a_length by (metis append.simps(2) nth_Cons_0 pair_id) have "(1::nat)< 2" by presburger hence 1: "x!1 = a!1" unfolding ab_def using a_length by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) qed have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) using reverse_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) qed have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) \ val (as!2)}= {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" by blast show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast qed lemma triple_val_ineq_set_semialg': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}" proof- have 0: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} = cartesian_product (reverse_val_relation_set) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule equalityI') show " \x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)} \ x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!0, x!1]" by blast have a_length: "length a = 2" proof- have "a = x!0 #[x!1]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!0" "[x!1]"] unfolding singleton_length[of "x!1"] by presburger qed obtain b where b_def: "b = [x!2]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def 0 A unfolding mem_Collect_eq by (meson Qp_2I cartesian_power_car_memE' rel_simps(49) rel_simps(51) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "2::nat"] by simp have 2: "x = a@b" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of a b] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (a @ b) ! i" apply(cases "i = 0") apply (metis a_def append.simps(2) nth_Cons_0) apply(cases "(i:: nat) = 1") apply (simp add: a_def) proof- assume a: "i \0" "i \ 1" then have "i = 2" using A1 by presburger thus ?thesis by (metis a_length b_def nth_append_length) qed qed have 3: "a = take 2 x" apply(rule nth_equalityI) unfolding a_length 0 length_take[of 2 x] apply linarith proof- fix i::nat assume a: "i < 2" show "a ! i = take 2 x ! i " apply(cases "i = 0") apply (metis a_def nth_Cons_0 nth_take zero_less_numeral) by (smt "0" \length (take 2 x) = min (length x) 2\ a_def linorder_neqE_nat min.commute min.strict_order_iff nth_take numeral_eq_iff one_less_numeral_iff pair_id pos2 rel_simps(22) rel_simps(48) rel_simps(9) semiring_norm(81)) qed show " x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p 2 _ 1]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) apply blast using 3 a_closed apply blast proof- have "drop 2 x = b" unfolding 2 unfolding 3 using 0 by simp then show "drop 2 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using b_closed by blast qed qed show "\x. x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 0) \ val (as ! 1)}" proof fix x assume A: "x \ cartesian_product reverse_val_relation_set (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = a@b" using cartesian_product_memE'[of x reverse_val_relation_set "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast have "(0::nat)< 2" by presburger hence 0: "x!0 = a!0" unfolding ab_def using a_length by (metis append.simps(2) nth_Cons_0 pair_id) have "(1::nat)< 2" by presburger hence 1: "x!1 = a!1" unfolding ab_def using a_length by (metis append.simps(2) less_2_cases nth_Cons_0 nth_Cons_Suc pair_id) obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" using ab_def cartesian_power_append[of a Q\<^sub>p 2 b'] b'_def b'_closed unfolding b'_def ab_def(3) reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 0) \ val (x ! 1)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (simp add: carrier_is_semialgebraic reverse_val_relation_set_semialg) qed have 1: "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ strict_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) using strict_val_relation_set_def apply blast using take_closed[of 1 3 x Q\<^sub>p] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding strict_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding strict_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) qed have 2: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1) \ val (as!1) < val (as!2)}= {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!0) \ val (as!1)} \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" by blast show ?thesis using intersection_is_semialg 0 1 unfolding 2 by blast qed lemma triple_val_ineq_set_semialg'': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) < val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (strict_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ strict_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def strict_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed strict_val_relation_set_is_semialg) using strict_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq using one_le_numeral apply blast using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) < val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) strict_val_relation_set " then obtain a b where ab_def: "a \ strict_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" strict_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding strict_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding strict_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) < val (x ! 2)" using x_closed ab_def unfolding strict_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral strict_val_relation_set_is_semialg) qed lemma triple_val_ineq_set_semialg''': shows "is_semialgebraic 3 {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)}" proof- have 0: "{as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as!1) \ val (as!2)} = cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) (reverse_val_relation_set)" proof(rule equalityI') show "\x. x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)} \ x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" proof- fix x assume A: " x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" then have 0: "length x = 3" unfolding mem_Collect_eq using cartesian_power_car_memE by blast obtain a where a_def: "a = [x!1, x!2]" by blast have a_length: "length a = 2" proof- have "a = x!1 #[x!2]" unfolding a_def by blast thus ?thesis using length_Cons[of "x!1" "[x!2]"] unfolding singleton_length[of "x!2"] by presburger qed obtain b where b_def: "b = [x!0]" by blast have b_length: "length b = 1" unfolding b_def singleton_length by auto have a_closed: "a \ reverse_val_relation_set" proof- have 0: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed have 1: "a \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" using a_def A drop_closed[of 1 3 x Q\<^sub>p] unfolding 0 mem_Collect_eq by (metis One_nat_def Suc_1 diff_Suc_1 numeral_3_eq_3 rel_simps(49) semiring_norm(77)) show ?thesis using 1 A unfolding a_def reverse_val_relation_set_def A mem_Collect_eq by (metis Qp_2_car_memE list_tl nth_Cons_0) qed have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" apply(rule cartesian_power_car_memI) unfolding b_length apply blast apply(rule subsetI) unfolding b_def using A unfolding mem_Collect_eq using cartesian_power_car_memE'[of x Q\<^sub>p "3::nat" "0::nat"] by (metis b_def b_length in_set_conv_nth less_one Qp.to_R_to_R1 zero_less_numeral) have 2: "x = b@a" apply(rule nth_equalityI) using 0 unfolding a_length b_length length_append[of b a] apply presburger proof- fix i assume A: "i < length x" then have A1: "i < 3" unfolding 0 by blast show "x ! i = (b @ a) ! i" apply(cases "i = 0") apply (metis append.simps(2) b_def nth_Cons_0) apply(cases "(i:: nat) = (1::nat)") using append.simps a_def nth_Cons apply (metis b_length nth_append_length) apply(cases "(i:: nat) = (2::nat)") using A unfolding 0 apply (metis a_def a_length arith_special(3) b_length list.inject nth_append_length_plus pair_id) proof- assume A0: "i \0" "i \ 1" "i \2" then have "i \ 3" by presburger then show "x ! i = (b @ a) ! i" using A unfolding 0 by presburger qed qed have 3: "a = drop 1 x" apply(rule nth_equalityI) unfolding a_length 0 length_drop[of 1 x] apply linarith proof- fix i::nat assume a: "i < 2" show " a ! i = drop 1 x ! i" apply(cases "i = 0") unfolding a_def using nth_drop[of 1 x i] apply (metis (no_types, opaque_lifting) "0" a_def arith_extra_simps(6) diff_is_0_eq' eq_imp_le eq_numeral_extra(1) flip_def flip_eval(1) less_numeral_extra(1) less_one less_or_eq_imp_le nat_add_left_cancel_le nat_le_linear nat_less_le nth_Cons_0 nth_drop numeral_neq_zero trans_less_add2 zero_less_diff) apply(cases "i = 1") using nth_drop[of 1 x i] unfolding 0 apply (metis "0" a_def a_length list.simps(1) nat_1_add_1 nth_drop one_le_numeral pair_id semiring_norm(3)) using a by presburger qed show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set" apply(rule cartesian_product_memI[of _ Q\<^sub>p 1 _ 2]) apply (simp add: is_semialgebraic_closed reverse_val_relation_set_semialg) using reverse_val_relation_set_def apply blast using take_closed[of 1 3 x] A unfolding mem_Collect_eq apply auto[1] using a_closed unfolding 3 by blast qed show "\x. x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set \ x \ {as \ carrier (Q\<^sub>p\<^bsup>3\<^esup>). val (as ! 1) \ val (as ! 2)}" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) reverse_val_relation_set " then obtain a b where ab_def: "a \ reverse_val_relation_set" "b \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" "x = b@a" using cartesian_product_memE'[of x "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" reverse_val_relation_set] by metis have a_length: "length a = 2" using ab_def unfolding reverse_val_relation_set_def using cartesian_power_car_memE by blast obtain b' where b'_def: "b = [b']" using ab_def cartesian_power_car_memE by (metis (no_types, opaque_lifting) append_Cons append_Nil append_eq_append_conv min_list.cases singleton_length) have b'_closed: "b' \ carrier Q\<^sub>p" using b'_def ab_def cartesian_power_car_memE by (metis Qp.R1_memE' list_hd) have b_length: "length b = 1" by (simp add: b'_def) have x_id: "x = b'#a" unfolding ab_def b'_def by auto have "(1::nat)< 2" by presburger hence 0: "x!1 = a!0" unfolding ab_def b'_def using a_length by (metis b'_def b_length nth_append_length pair_id) have 00: "2 = Suc 1" by auto have 1: "x!2 = a!1" using a_length nth_Cons[of b' a "2::nat"] unfolding x_id 00 by (meson nth_Cons_Suc) have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>)" unfolding x_id b'_def using b'_closed cartesian_power_cons[of a Q\<^sub>p 2 b'] ab_def unfolding reverse_val_relation_set_def mem_Collect_eq by simp show "x \ carrier (Q\<^sub>p\<^bsup>3\<^esup>) \ val (x ! 1) \ val (x ! 2)" using x_closed ab_def unfolding reverse_val_relation_set_def mem_Collect_eq 0 1 by blast qed qed show ?thesis unfolding 0 using cartesian_product_is_semialgebraic[of 2 reverse_val_relation_set 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] by (metis add_num_simps(2) car_times_semialg_is_semialg one_plus_numeral reverse_val_relation_set_semialg) qed (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Functions\ (**************************************************************************************************) (**************************************************************************************************) text\ The most natural way to define a semialgebraic function $f: \mathbb{Q}_p^n \to \mathbb{Q}_p$ is a function whose graph is a semialgebraic subset of $\mathbb{Q}_p^{n+1}$. However, the definition given here is slightly different, and devised by Denef in \<^cite>\"denef1986"\ in order to prove Macintyre's theorem. As Denef notes, we can use Macintyre's theorem to deduce that the given definition perfectly aligns with the intuitive one. \ (********************************************************************************************) (********************************************************************************************) subsubsection\Defining Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Apply a function f to the tuple consisting of the first n indices, leaving the remaining indices unchanged\ definition partial_image where "partial_image m f xs = (f (take m xs))#(drop m xs)" definition partial_pullback where "partial_pullback m f l S = (partial_image m f) \\<^bsub>m+l\<^esub> S " lemma partial_pullback_memE: assumes "as \ partial_pullback m f l S" shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "partial_image m f as \ S" using assms apply (metis evimage_eq partial_pullback_def) using assms unfolding partial_pullback_def by blast lemma partial_pullback_closed: "partial_pullback m f l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using partial_pullback_memE(1) by blast lemma partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" assumes "(f (take m as))#(drop m as) \ S" shows "as \ partial_pullback m f k S" using assms unfolding partial_pullback_def partial_image_def evimage_def by blast lemma partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" shows "partial_image n f x = (f as)#bs" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(3) cartesian_power_car_memE) have 1: "drop n x = bs" by (metis "0" append_take_drop_id assms(3) same_append_eq) show ?thesis using 0 1 unfolding partial_image_def by blast qed lemma partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" assumes "x \ partial_pullback n f k S" shows "(f as)#bs \ S" using partial_pullback_memE[of x n f k S] partial_image_def[of n f x] by (metis assms(1) assms(2) assms(3) assms(4) partial_image_eq) text\Partial pullbacks have the same algebraic properties as pullbacks\ lemma partial_pullback_intersect: "partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" unfolding partial_pullback_def by simp lemma partial_pullback_union: "partial_pullback m f l (S1 \ S2) = (partial_pullback m f l S1) \ (partial_pullback m f l S2)" unfolding partial_pullback_def by simp lemma cartesian_power_drop: assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "drop n x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" apply(rule cartesian_power_car_memI) using assms cartesian_power_car_memE apply (metis add_diff_cancel_left' length_drop) using assms cartesian_power_car_memE'' by (metis order.trans set_drop_subset) lemma partial_pullback_complement: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (partial_pullback m f l S) " apply(rule equalityI) using partial_pullback_def[of m f l "(carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S)"] partial_pullback_def[of m f l S] apply (smt Diff_iff evimage_Diff partial_pullback_memE(1) subsetI) proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - partial_pullback m f l S" show " x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) - S) " apply(rule partial_pullback_memI) using A apply blast proof have 00: "Suc l = l + 1" by auto have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" by (meson A DiffD1 cartesian_power_drop) have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A by (meson DiffD1 le_add1 take_closed) have "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>l+1\<^esup>) " using assms 0 1 00 cartesian_power_cons[of "drop m x" Q\<^sub>p l "f (take m x)"] by blast thus "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>) " using 00 by metis show "f (take m x) # drop m x \ S" using A unfolding partial_pullback_def partial_image_def by blast qed qed lemma partial_pullback_carrier: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" apply(rule equalityI) using partial_pullback_memE(1) apply blast proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" show "x \ partial_pullback m f l (carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>))" apply(rule partial_pullback_memI) using A cartesian_power_drop[of x m l] assms apply blast proof- have "f (take m x) \ carrier Q\<^sub>p" using A assms take_closed[of m "m+l" x Q\<^sub>p] by (meson Pi_mem le_add1) then show "f (take m x) # drop m x \ carrier (Q\<^sub>p\<^bsup>Suc l\<^esup>)" using cartesian_power_drop[of x m l] by (metis A add.commute cartesian_power_cons plus_1_eq_Suc) qed qed text\Definition 1.4 from Denef\ definition is_semialg_function where "is_semialg_function m f = ((f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p) \ (\l \ 0. \S \ semialg_sets (1 + l). is_semialgebraic (m + l) (partial_pullback m f l S)))" lemma is_semialg_function_closed: assumes "is_semialg_function m f" shows "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using is_semialg_function_def assms by blast lemma is_semialg_functionE: assumes "is_semialg_function m f" assumes "is_semialgebraic (1 + k) S" shows " is_semialgebraic (m + k) (partial_pullback m f k S)" using is_semialg_function_def assms by (meson is_semialgebraicE le0) lemma is_semialg_functionI: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" shows "is_semialg_function m f" using assms unfolding is_semialg_function_def by blast text\Semialgebraicity for functions can be verified on basic semialgebraic sets \ lemma is_semialg_functionI': assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" shows "is_semialg_function m f" apply(rule is_semialg_functionI) using assms(1) apply blast proof- show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m f k S)" proof- fix k S assume A: "S \ semialg_sets (1 + k)" show "is_semialgebraic (m + k) (partial_pullback m f k S)" apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" "basic_semialgs (1 + k)"]) using A unfolding semialg_sets_def apply blast using partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger apply (metis assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg partial_pullback_carrier partial_pullback_intersect plus_1_eq_Suc) using partial_pullback_union union_is_semialgebraic apply presburger using assms(1) complement_is_semialg partial_pullback_complement plus_1_eq_Suc by presburger qed qed text\Graphs of semialgebraic functions are semialgebraic\ abbreviation graph where "graph \ fun_graph Q\<^sub>p" lemma graph_memE: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "x \ graph m f" shows "f (take m x) = x!m" "x = (take m x)@[f (take m x)]" "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" proof- obtain a where a_def: "a\carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ x = a @ [f a]" using assms unfolding fun_graph_def by blast then have 0: "a = take m x" by (metis append_eq_conv_conj cartesian_power_car_memE) then show "f (take m x) = x!m" by (metis a_def cartesian_power_car_memE nth_append_length) show "x = (take m x)@[f (take m x)]" using "0" a_def by blast show "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using "0" a_def by blast qed lemma graph_memI: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "f (take m x) = x!m" assumes "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" shows "x \ graph m f" proof- have 0: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule take_closed[of _ "m + 1"]) apply simp using assms(3) by blast have "x = (take m x)@[x!m]" by (metis \take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)\ add.commute assms(3) cartesian_power_car_memE length_append_singleton lessI nth_equalityI nth_take plus_1_eq_Suc take_Suc_conv_app_nth) then have "x = (take m x)@[f (take m x)]" using assms(2) by presburger then show ?thesis using assms 0 unfolding fun_graph_def by blast qed lemma graph_mem_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "x \ graph m f" shows "x \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" proof(rule cartesian_power_car_memI') show "length x = m + 1" using assms graph_memE[of f m x] by (smt Groups.add_ac(2) cartesian_power_car_memE fun_graph_def length_append_singleton mem_Collect_eq plus_1_eq_Suc) show "\i. i < m + 1 \ x ! i \ carrier Q\<^sub>p" proof- fix i assume A: "i < m + 1" then show "x ! i \ carrier Q\<^sub>p" proof(cases "i = m") case True then show ?thesis using graph_memE[of f m x] by (metis PiE assms(1) assms(2)) next case False then show ?thesis using graph_memE[of f m x] by (metis \i < m + 1\ add.commute assms(1) assms(2) cartesian_power_car_memE' less_SucE nth_take plus_1_eq_Suc) qed qed qed lemma graph_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" shows "graph m f \ carrier (Q\<^sub>p\<^bsup>m+1\<^esup>)" using assms graph_mem_closed by blast text\The \m\-dimensional diagonal set is semialgebraic\ notation diagonal ("\ ") lemma diag_is_algebraic: shows "is_algebraic Q\<^sub>p (n + n) (\ n)" using Qp.cring_axioms diagonal_is_algebraic by blast lemma diag_is_semialgebraic: shows "is_semialgebraic (n + n) (\ n)" using diag_is_algebraic is_algebraic_imp_is_semialg by blast text\Transposition permutations\ definition transpose where "transpose i j = (Fun.swap i j id)" lemma transpose_permutes: assumes "i< n" assumes "j < n" shows "transpose i j permutes {..x. x \ {.. Fun.swap i j id x = x" using assms by (auto simp: Transposition.transpose_def) show "\y. \!x. Fun.swap i j id x = y" proof fix y show "\!x. Fun.swap i j id x = y" using swap_id_eq[of i j y] by (metis eq_id_iff swap_apply(1) swap_apply(2) swap_id_eq swap_self) qed qed lemma transpose_alt_def: "transpose a b x = (if x = a then b else if x = b then a else x)" using swap_id_eq by (simp add: transpose_def) definition last_to_first where "last_to_first n = (\i. if i = (n-1) then 0 else if i < n-1 then i + 1 else i)" definition first_to_last where "first_to_last n = fun_inv (last_to_first n)" lemma last_to_first_permutes: assumes "(n::nat) > 0" shows "last_to_first n permutes {..x. x \ {.. last_to_first n x = x" proof fix x show " x \ {.. last_to_first n x = x" proof assume A: "x \ {.. x < n" by blast then have "x \ n" by linarith then show "last_to_first n x = x" unfolding last_to_first_def using assms by auto qed qed show "\y. \!x. last_to_first n x = y" proof fix y show "\!x. last_to_first n x = y" proof(cases "y = 0") case True then have 0: "last_to_first n (n-1) = y" using last_to_first_def by (simp add: last_to_first_def) have 1: "\x. last_to_first n x = y \ x = n-1" unfolding last_to_first_def using True by (metis add_gr_0 less_numeral_extra(1) not_gr_zero) show ?thesis using 0 1 by blast next case False then show ?thesis proof(cases "y < n") case True then have 0: "last_to_first n (y-1) = y" using False True unfolding last_to_first_def using add.commute by auto have 1: "\x. last_to_first n x = y \ x =(y-1)" unfolding last_to_first_def using True False by auto show ?thesis using 0 1 by blast next case F: False then have 0: "y \ n" using not_less by blast then have 1: "last_to_first n y = y" by (simp add: \\x. x \ {.. last_to_first n x = x\) have 2: "\x. last_to_first n x = y \ x =y" using 0 unfolding last_to_first_def using False by presburger then show ?thesis using 1 2 by blast qed qed qed qed definition graph_swap where "graph_swap n f = permute_list ((first_to_last (n+1))) ` (graph n f)" lemma last_to_first_eq: assumes "length as = n" shows "permute_list (last_to_first (n+1)) (a#as) = (as@[a])" proof- have 0: "\i. i < (n+1) \ permute_list (last_to_first (n + 1)) (a # as) ! i = (as@[a]) ! i" proof- fix i assume A: "i < n+1" show "permute_list (last_to_first (n + 1)) (a # as) ! i = (as @ [a]) ! i" proof(cases "i = n") case True have 0: "(as @ [a]) ! i = a" by (metis True assms nth_append_length) have 1: "length (a#as) = n + 1" by (simp add: assms) have 2: "i < length (a # as)" using "1" A by linarith have 3: "last_to_first (n + 1) permutes {.. carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "a \ carrier Q\<^sub>p" shows "permute_list (first_to_last (n+1)) (as@[a]) = (a#as)" proof- have "length as = n" using assms(1) cartesian_power_car_memE by blast then show ?thesis using last_to_first_eq last_to_first_permutes[of n] permute_list_compose_inv(2)[of "(last_to_first (n + 1))" n "a # as"] unfolding first_to_last_def by (metis add_gr_0 assms(1) assms(2) cartesian_power_append last_to_first_permutes less_one permute_list_closed' permute_list_compose_inv(2)) qed lemma graph_swapI: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "(f as)#as \ graph_swap n f" proof- have 0: "as@[f as] \ graph n f" using assms using graph_memI[of f n] fun_graph_def by blast have 1: "f as \ carrier Q\<^sub>p" using assms by blast then show ?thesis using assms 0 first_to_last_eq[of as "n" "f as"] unfolding graph_swap_def by (metis image_eqI) qed lemma graph_swapE: assumes "x \ graph_swap n f" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "hd x = f (tl x)" proof- obtain y where y_def: "y \ graph n f \ x = permute_list (first_to_last (n+1)) y" using assms graph_swap_def by (smt image_def mem_Collect_eq) then have "take n y \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(2) graph_memE(3) by blast then show "hd x = f (tl x)" by (metis (no_types, lifting) add.commute assms(2) cartesian_power_car_memE' first_to_last_eq graph_memE(1) graph_memE(2) graph_mem_closed lessI list.sel(1) list.sel(3) plus_1_eq_Suc y_def) qed text\Semialgebraic functions have semialgebraic graphs\ lemma graph_as_partial_pullback: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "partial_pullback n f 1 (\ 1) = graph n f" proof show "partial_pullback n f 1 (\ 1) \ graph n f" proof fix x assume A: "x \ partial_pullback n f 1 (\ 1)" then have 0: "f (take n x) # drop n x \ \ 1" by (metis local.partial_image_def partial_pullback_memE(2)) then have 1: "length (f (take n x) # drop n x) = 2" using diagonal_def by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq one_add_one) then obtain b where b_def: "[b] = drop n x" by (metis list.inject pair_id) then have "[f (take n x), b] \ \ 1" using "0" by presburger then have "b = f (take n x)" using 0 by (smt One_nat_def Qp.cring_axioms diagonal_def drop0 drop_Suc_Cons list.inject mem_Collect_eq take_Suc_Cons) then have "x = (take n x)@[f (take n x)]" by (metis append_take_drop_id b_def) then show "x \ graph n f" using graph_memI[of f n x] by (metis (no_types, lifting) A \b = f (take n x)\ assms b_def nth_via_drop partial_pullback_memE(1)) qed show "graph n f \ partial_pullback n f 1 (\ 1)" proof fix x assume A: "x \ graph n f " then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" using assms graph_mem_closed by blast have "x = (take n x) @ [f (take n x)]" using A graph_memE(2)[of f n x] assms by blast then have "partial_image n f x = [f (take n x), f (take n x)]" by (metis append_take_drop_id local.partial_image_def same_append_eq) then have "partial_image n f x \ \ 1" using assms 0 diagonal_def[of 1] Qp.cring_axioms diagonalI[of "partial_image n f x"] by (metis (no_types, lifting) A append_Cons append_eq_conv_conj cartesian_power_car_memE cartesian_power_car_memE' graph_memE(1) less_add_one self_append_conv2 Qp.to_R1_closed) then show "x \ partial_pullback n f 1 (\ 1)" unfolding partial_pullback_def using 0 by blast qed qed lemma semialg_graph: assumes "is_semialg_function n f" shows "is_semialgebraic (n + 1) (graph n f)" using assms graph_as_partial_pullback[of f n] unfolding is_semialg_function_def by (metis diag_is_semialgebraic is_semialgebraicE less_imp_le_nat less_numeral_extra(1)) text\Functions induced by polynomials are semialgebraic\ definition var_list_segment where "var_list_segment i j = map (\i. pvar Q\<^sub>p i) [i..< j]" lemma var_list_segment_length: assumes "i \ j" shows "length (var_list_segment i j) = j - i" using assms var_list_segment_def by fastforce lemma var_list_segment_entry: assumes "k < j - i" assumes "i \ j" shows "var_list_segment i j ! k = pvar Q\<^sub>p (i + k)" using assms var_list_segment_length unfolding var_list_segment_def using nth_map_upt by blast lemma var_list_segment_is_poly_tuple: assumes "i \j" assumes "j \ n" shows "is_poly_tuple n (var_list_segment i j)" apply(rule Qp_is_poly_tupleI) using assms var_list_segment_entry var_list_segment_length Qp.cring_axioms pvar_closed[of _ n] by (metis (no_types, opaque_lifting) add.commute add_lessD1 diff_add_inverse le_Suc_ex less_diff_conv) lemma map_by_var_list_segment: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j \ n" assumes "i \ j" shows "poly_map n (var_list_segment i j) as = list_segment i j as" apply(rule nth_equalityI ) unfolding poly_map_def var_list_segment_def list_segment_def restrict_def poly_tuple_eval_def apply (metis (full_types) assms(1) length_map) using assms eval_pvar[of _ n as] Qp.cring_axioms length_map add.commute length_upt less_diff_conv less_imp_add_positive nth_map nth_upt trans_less_add2 by (smt le_add_diff_inverse2) lemma map_by_var_list_segment_to_length: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i \ n" shows "poly_map n (var_list_segment i n) as = drop i as" apply(rule nth_equalityI ) apply (metis Qp_poly_mapE' assms(1) assms(2) cartesian_power_car_memE length_drop var_list_segment_length) using assms map_by_var_list_segment[of as n n i] list_segment_drop[of i as] cartesian_power_car_memE[of as Q\<^sub>p n] map_nth[of ] nth_drop nth_map[of _ "[i..p)" ] nth_map[of _ "map (pvar Q\<^sub>p) [i..p as"] unfolding poly_map_def poly_tuple_eval_def var_list_segment_def restrict_def list_segment_def by (smt add.commute add_eq_self_zero drop_map drop_upt le_Suc_ex le_refl) lemma map_tail_by_var_list_segment: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "a \ carrier Q\<^sub>p" assumes "i < n" shows "poly_map (n+1) (var_list_segment 1 (n+1)) (a#as) = as" proof- have 0: "(a#as) \ carrier (Q\<^sub>p\<^bsup>n+1\<^esup>)" using assms by (meson cartesian_power_cons) have 1: "length as = n" using assms cartesian_power_car_memE by blast have 2: "drop 1 (a # as) = as" using 0 1 using list_segment_drop[of 1 "a#as"] by (metis One_nat_def drop0 drop_Suc_Cons ) have "1 \n + 1" by auto then show ?thesis using 0 2 map_by_var_list_segment_to_length[of "a#as" "n+1" 1] by presburger qed lemma Qp_poly_tuple_Cons: assumes "is_poly_tuple n fs" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" assumes "k \n" shows "is_poly_tuple n (f#fs)" using is_poly_tuple_Cons[of n fs f] poly_ring_car_mono[of k n] assms by blast lemma poly_map_Cons: assumes "is_poly_tuple n fs" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n (f#fs) a = (Qp_ev f a)#poly_map n fs a" using assms poly_map_cons by blast lemma poly_map_append': assumes "is_poly_tuple n fs" assumes "is_poly_tuple n gs" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n (fs@gs) a = poly_map n fs a @ poly_map n gs a" using assms(3) poly_map_append by blast lemma partial_pullback_by_poly: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" shows "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n+k) S (f# (var_list_segment n (n+k)))" proof show "partial_pullback n (Qp_ev f) k S \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" proof fix x assume A: " x \ partial_pullback n (Qp_ev f) k S" then obtain as bs where as_bs_def: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ x = as @ bs" using partial_pullback_memE(1)[of x n "(Qp_ev f)" k S] cartesian_power_decomp by metis then have 0: "(Qp_ev f as#bs) \ S" using A partial_pullback_memE' by blast have 1: "Qp_ev f as = Qp_ev f (as@bs)" using assms as_bs_def poly_eval_cartesian_prod[of as n bs k f] Qp.cring_axioms [of ] by metis then have 2: "((Qp_ev f x) #bs) \ S" using "0" as_bs_def by presburger have 3: "bs = list_segment n (n+k) x" using as_bs_def list_segment_drop[of n x] by (metis (no_types, lifting) add_cancel_right_right add_diff_cancel_left' append_eq_append_conv append_take_drop_id cartesian_power_car_memE length_0_conv length_append length_map length_upt linorder_neqE_nat list_segment_def not_add_less1) have 4: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" using Qp_poly_tuple_Cons var_list_segment_is_poly_tuple by (metis add.commute assms(1) dual_order.refl le_add2) have 5: "f \ carrier (Q\<^sub>p [\\<^bsub>n + k\<^esub>])" using poly_ring_car_mono[of n "n + k"] assms le_add1 by blast have 6: "is_poly_tuple (n + k) (var_list_segment n (n + k))" by (simp add: var_list_segment_is_poly_tuple) have 7: "x \ carrier (Q\<^sub>p\<^bsup>n + k\<^esup>)" using as_bs_def cartesian_power_concat(1) by blast hence 8: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#poly_map (n+k) (var_list_segment n (n + k)) x" using 5 6 7 A poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 4 unfolding partial_pullback_def evimage_def by blast hence 6: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#bs" using 3 "7" le_add1 le_refl map_by_var_list_segment by presburger show " x \ poly_tuple_pullback (n+k) S (f # var_list_segment n (n + k))" unfolding poly_tuple_pullback_def using 6 by (metis "2" "7" IntI poly_map_apply vimage_eq) qed show "poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k)) \ partial_pullback n (Qp_ev f) k S" proof fix x assume A: "x \ poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" have 0: "is_poly_tuple (n+k) (f # var_list_segment n (n + k))" using Qp_poly_tuple_Cons assms(1) le_add1 var_list_segment_is_poly_tuple by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using A unfolding poly_tuple_pullback_def by blast have 2: "poly_map (n+k) (f # var_list_segment n (n + k)) x \ S" using 1 assms A unfolding poly_map_def poly_tuple_pullback_def restrict_def by (metis (no_types, opaque_lifting) Int_commute add.commute evimage_def evimage_eq) have 3: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f x)#(drop n x)" using poly_map_Cons[of "n + k" "var_list_segment n (n + k)" f x] 1 assms(1) map_by_var_list_segment_to_length le_add1 poly_map_cons by presburger have 4: "poly_map (n+k) (f # var_list_segment n (n + k)) x = (Qp_ev f (take n x))#(drop n x)" using assms 1 3 eval_at_points_higher_pow[of f n "n + k" "x"] le_add1 by (metis nat_le_iff_add) show "x \ partial_pullback n (Qp_ev f) k S" apply(rule partial_pullback_memI) using 1 apply blast using 2 3 4 by metis qed qed lemma poly_is_semialg: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialg_function n (Qp_ev f)" proof(rule is_semialg_functionI) show "Qp_ev f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" using assms by (meson Pi_I eval_at_point_closed) show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" proof- fix k::nat fix S assume A: "S \ semialg_sets (1 + k)" have 0: "is_poly_tuple (n + k) (f # var_list_segment n (n + k))" by (metis add.commute assms le_add2 order_refl Qp_poly_tuple_Cons var_list_segment_is_poly_tuple) have 1: "length (f # var_list_segment n (n + k)) = k + 1" by (metis add.commute add_diff_cancel_left' le_add1 length_Cons plus_1_eq_Suc var_list_segment_length) have 2: "partial_pullback n (Qp_ev f) k S = poly_tuple_pullback (n + k) S (f # var_list_segment n (n + k))" using A assms partial_pullback_by_poly[of f n S k] unfolding semialg_sets_def using gen_boolean_algebra_subset by blast then show "is_semialgebraic (n + k) (partial_pullback n (Qp_ev f) k S)" using add.commute[of 1 k] 0 1 assms(1) pullback_is_semialg[of "n+k" "(f # var_list_segment n (n + k))" "k+1" S] by (metis A is_semialgebraicI is_semialgebraic_closed poly_tuple_pullback_eq_poly_map_vimage) qed qed text\Families of polynomials defined by semialgebraic coefficient functions\ lemma semialg_function_on_carrier: assumes "is_semialg_function n f" assumes "restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict g (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" shows "is_semialg_function n g" proof(rule is_semialg_functionI) have 0: "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" using assms(1) is_semialg_function_closed by blast show "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" then show " g x \ carrier Q\<^sub>p" using assms(2) 0 by (metis (no_types, lifting) PiE restrict_Pi_cancel) qed show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (n + k) (partial_pullback n g k S)" proof- fix k S assume A: "S \ semialg_sets (1 + k)" have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" using A assms(1) is_semialg_functionE is_semialgebraicI by blast have 2: "(partial_pullback n f k S) = (partial_pullback n g k S)" unfolding partial_pullback_def partial_image_def evimage_def proof show "(\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" proof fix x assume "x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) " have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms by (meson \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ inf_le2 le_add1 subset_iff take_closed) then have "f (take n x) = g (take n x)" using assms unfolding restrict_def by meson then show " x \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using assms \x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)\ by blast qed show "(\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>) \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" proof fix x assume A: "x \ (\xs. g (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms by (meson A inf_le2 le_add1 subset_iff take_closed) then have "f (take n x) = g (take n x)" using assms unfolding restrict_def by meson then show "x \ (\xs. f (take n xs) # drop n xs) -` S \ carrier (Q\<^sub>p\<^bsup>n+k\<^esup>)" using A by blast qed qed then show "is_semialgebraic (n + k) (partial_pullback n g k S)" using 1 by auto qed qed lemma semialg_function_on_carrier': assumes "is_semialg_function n f" assumes "\a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = g a" shows "is_semialg_function n g" using assms semialg_function_on_carrier unfolding restrict_def by (meson restrict_ext semialg_function_on_carrier) lemma constant_function_is_semialg: assumes "n > 0" assumes "x \ carrier Q\<^sub>p" assumes "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = x" shows "is_semialg_function n f" proof(rule semialg_function_on_carrier[of _ "Qp_ev (Qp_to_IP x)"]) show "is_semialg_function n (Qp_ev (Qp_to_IP x))" using assms poly_is_semialg[of "(Qp_to_IP x)"] Qp_to_IP_car by blast have 0: "\ a. a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ f a = Qp_ev (Qp_to_IP x) a" using eval_at_point_const assms by blast then show "restrict (Qp_ev (Qp_to_IP x)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) = restrict f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" by (metis (no_types, lifting) restrict_ext) qed lemma cartesian_product_singleton_factor_projection_is_semialg: assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A {b})" shows "is_semialgebraic m A" proof- obtain f where f_def: "f = map (pvar Q\<^sub>p) [0..p" "[0.. ((nat \ int) \ (nat \ int)) set) list) = map (\i::nat. Qp.indexed_const (b ! i)) [(0::nat)..i. i \ set [0::nat..< n] \ b!i \ carrier Q\<^sub>p" using assms(2) cartesian_power_car_memE'[of b Q\<^sub>p n] by blast hence 1: "\i. i \ set [0::nat..< n] \ Qp.indexed_const (b ! i) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" using assms Qp_to_IP_car by blast show ?thesis unfolding is_poly_tuple_def g_def apply(rule subsetI) using set_map[of "\i. Qp.indexed_const (b ! i)" "[0..x. x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ poly_tuple_eval (f@g) x = x@b" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" have 30: "poly_tuple_eval f x = x" proof- have 300: "length (poly_tuple_eval f x) = length x" unfolding poly_tuple_eval_def using cartesian_power_car_memE by (metis "4" A length_map) have "\i. i < length x \ poly_tuple_eval f x ! i = x ! i" unfolding f_def poly_tuple_eval_def using nth_map by (metis "4" A add_cancel_right_left cartesian_power_car_memE eval_pvar f_def length_map nth_upt) thus ?thesis using 300 by (metis nth_equalityI) qed have 31: "poly_tuple_eval g x = b" proof- have 310: "length (poly_tuple_eval g x) = length b" unfolding poly_tuple_eval_def g_def using cartesian_power_car_memE by (metis assms(2) length_map map_nth) have 311: "length b = n" using assms cartesian_power_car_memE by blast hence "\i. i < n \ poly_tuple_eval g x ! i = b ! i" proof- fix i assume "i < n" thus "poly_tuple_eval g x ! i = b ! i" unfolding g_def poly_tuple_eval_def using eval_at_point_const[of "b!i" x m] 310 nth_map by (metis "311" A assms(2) cartesian_power_car_memE' length_map map_nth) qed thus ?thesis using 311 310 nth_equalityI by (metis list_eq_iff_nth_eq) qed have 32: "poly_tuple_eval (f @ g) x = poly_map m (f@g) x" unfolding poly_map_def restrict_def using A by (simp add: A) have 33: "poly_tuple_eval f x = poly_map m f x" unfolding poly_map_def restrict_def using A by (simp add: A) have 34: "poly_tuple_eval g x = poly_map m g x" unfolding poly_map_def restrict_def using A by (simp add: A) show "poly_tuple_eval (f @ g) x = x @ b" using assms 1 2 30 31 poly_map_append[of x m f g] A unfolding 32 33 34 by (simp add: A \b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\) qed have 4: "A = (poly_tuple_eval (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b}))" proof show "A \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" proof(rule subsetI) fix x assume A: "x \ A" then have 0: "poly_tuple_eval (f@g) x = x@b" using 3 assms by blast then show " x \ poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" using A cartesian_product_memE by (smt Un_upper1 assms(1) assms(2) cartesian_product_memI' evimageI2 in_mono insert_is_Un mk_disjoint_insert singletonI) qed show "poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" proof(rule subsetI) fix x assume A: "x \ (poly_tuple_eval (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b})" then have "poly_tuple_eval (f @ g) x \ cartesian_product A {b}" by blast then have "x@b \ cartesian_product A {b}" using A 3 by (metis evimage_eq) then show "x \ A" using A by (metis append_same_eq cartesian_product_memE' singletonD) qed qed have 5: "A = poly_map m (f@g) \\<^bsub>m\<^esub> (cartesian_product A {b})" proof show "A \ poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b}" unfolding poly_map_def evimage_def restrict_def using 4 by (smt IntI assms(1) evimageD in_mono subsetI vimageI) show "poly_map m (f @ g) \\<^bsub>m\<^esub> cartesian_product A {b} \ A" unfolding poly_map_def evimage_def restrict_def using 4 by (smt Int_iff evimageI2 subsetI vimage_eq) qed have 6: "length (f @ g) = m + n" unfolding f_def g_def by (metis index_list_length length_append length_map map_nth) show ?thesis using 2 5 6 assms pullback_is_semialg[of m "f@g" "m+n" "cartesian_product A {b}"] by (metis is_semialgebraicE zero_eq_add_iff_both_eq_0) qed lemma cartesian_product_factor_projection_is_semialg: assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "B \ {}" assumes "is_semialgebraic (m+n) (cartesian_product A B)" shows "is_semialgebraic m A" proof- obtain b where b_def: "b \ B" using assms by blast have "is_semialgebraic n {b}" using assms b_def is_algebraic_imp_is_semialg singleton_is_algebraic by blast hence 0: "is_semialgebraic (m+n) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b})" using car_times_semialg_is_semialg assms(4) by blast have "(cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {b}) \ (cartesian_product A B) = (cartesian_product A {b})" using assms b_def cartesian_product_intersection[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m "{b}" n A B] by (metis (no_types, lifting) Int_absorb1 Int_empty_left Int_insert_left_if1 \is_semialgebraic n {b}\ is_semialgebraic_closed set_eq_subset) hence "is_semialgebraic (m+n) (cartesian_product A {b})" using assms 0 intersection_is_semialg by metis thus ?thesis using assms cartesian_product_singleton_factor_projection_is_semialg by (meson \is_semialgebraic n {b}\ insert_subset is_semialgebraic_closed) qed lemma partial_pullback_cartesian_product: assumes "\ \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" shows "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) = partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) " proof show "cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" proof fix x assume A: "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" then obtain y t where yt_def: "x = y@[t] \ y \ partial_pullback m \ 0 S \ t \ carrier Q\<^sub>p" by (metis cartesian_product_memE' Qp.to_R1_to_R Qp.to_R_pow_closed) then have "[\ y] \ S" using partial_pullback_memE unfolding partial_image_def by (metis (no_types, lifting) add.right_neutral append.right_neutral cartesian_power_drop le_zero_eq take_closed partial_pullback_memE' take_eq_Nil) then have 0: "[\ y]@[t] \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using cartesian_product_memI' yt_def by (metis assms(2) carrier_is_semialgebraic is_semialgebraic_closed Qp.to_R1_closed) have 1: " x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" using A yt_def by (metis add.right_neutral cartesian_power_append partial_pullback_memE(1)) show "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" apply(rule partial_pullback_memI) using "1" apply blast using yt_def 0 by (smt Cons_eq_appendI add.right_neutral local.partial_image_def partial_image_eq partial_pullback_memE(1) self_append_conv2 Qp.to_R1_closed) qed show "partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))) \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof(rule subsetI) fix x assume A: "x \ partial_pullback m \ 1 (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>m + 1\<^esup>)" using assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] by blast have 1: "\ (take m x) # drop m x \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using A assms partial_pullback_memE[of x m \ 1 "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] unfolding partial_image_def by blast have 2: "\ (take m (take m x)) # drop m (take m x) = [\ (take m x)]" using 0 1 by (metis add.commute add.right_neutral append.right_neutral append_take_drop_id take0 take_drop) show "x \ cartesian_product (partial_pullback m \ 0 S) (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" apply(rule cartesian_product_memI[of _ Q\<^sub>p m _ 1]) apply (metis add_cancel_right_right partial_pullback_closed) apply blast apply(rule partial_pullback_memI[of _ m 0 \ S]) using 0 apply (metis Nat.add_0_right le_iff_add take_closed) using 2 apply (metis (no_types, lifting) "1" add.commute add.right_neutral assms(2) cartesian_product_memE(1) list.inject plus_1_eq_Suc take_Suc_Cons take_drop) using 0 cartesian_power_drop by blast qed qed lemma cartesian_product_swap: assumes "A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A B)" shows "is_semialgebraic (m+n) (cartesian_product B A)" proof- obtain f where f_def: "f = (\i. (if i < m then n + i else (if i < m+n then i - m else i)))" by blast have 0: "\i. i \ {.. f i \ {n..i. i \ {m.. f i \ {..i. i \ {.. f i \ {..x. x \ {.. f x = x" unfolding f_def by simp show "\y. \!x. f x = y" proof fix y show "\!x. f x = y" proof(cases "y < n") case True have T0: "f (y+m) = y" unfolding f_def using True by simp have "\i. f i = y \ i \ {m..i. f i = y \ i = y+m" using T0 unfolding f_def by auto thus ?thesis using T0 by blast next case False show ?thesis proof(cases "y \ {n..i. f i = y \ i \ {..i. f i = y \ i = y- n" using f_def by force then show ?thesis using T0 by blast next case F: False then show ?thesis using 0 1 2 unfolding f_def using False add_diff_inverse_nat lessThan_iff by auto qed qed qed qed have "permute_list f ` (cartesian_product A B) = (cartesian_product B A)" proof show "permute_list f ` cartesian_product A B \ cartesian_product B A" proof fix x assume A: " x \ permute_list f ` cartesian_product A B" then obtain a b where ab_def: "a \ A \b \ B \ x = permute_list f (a@b)" by (metis (mono_tags, lifting) cartesian_product_memE' image_iff) have 0: "x = permute_list f (a@b)" using ab_def by blast have 1: "length a = n" using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast have 2: "length b = m" using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast have 3: "length x = m + n" using 1 2 0 f_permutes by simp have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" unfolding 0 using permute_list_nth by (metis "0" "3" f_permutes length_permute_list trans_less_add1) hence 5: "\i. i < m \ x ! i = b!i" unfolding f_def using 1 2 by (metis "4" f_def nth_append_length_plus) have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" unfolding 0 using f_def permute_list_nth f_permutes by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 ordered_cancel_comm_monoid_diff_class.diff_add) have 7: "x = b@a" proof(rule nth_equalityI) show "length x = length (b @ a)" using 1 2 3 by simp show "\i. i < length x \ x ! i = (b @ a) ! i" unfolding 3 using 1 2 4 5 by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) qed show "x \ cartesian_product B A" unfolding 7 using ab_def unfolding cartesian_product_def by blast qed show "cartesian_product B A \ permute_list f ` cartesian_product A B" proof fix y assume A: "y \ cartesian_product B A" then obtain b a where ab_def: "b \ B \ a \ A \ y = b@a" using cartesian_product_memE' by blast obtain x where 0: "x = permute_list f (a@b)" by blast have 1: "length a = n" using ab_def assms cartesian_power_car_memE[of a Q\<^sub>p n] by blast have 2: "length b = m" using ab_def assms cartesian_power_car_memE[of b Q\<^sub>p m] by blast have 3: "length x = m + n" using 1 2 0 f_permutes by simp have 4: "\i. i < m \ x ! i = (a@b) ! (f i)" unfolding 0 using permute_list_nth by (metis "0" "3" f_permutes length_permute_list trans_less_add1) hence 5: "\i. i < m \ x ! i = b!i" unfolding f_def using 1 2 by (metis "4" f_def nth_append_length_plus) have 6: "\i. i \ {m.. x ! i = (a@b) ! (i - m)" unfolding 0 using f_def permute_list_nth f_permutes by (metis (no_types, lifting) "0" "3" atLeastLessThan_iff length_permute_list not_add_less2 ordered_cancel_comm_monoid_diff_class.diff_add) have 7: "x = b@a" proof(rule nth_equalityI) show "length x = length (b @ a)" using 1 2 3 by simp show "\i. i < length x \ x ! i = (b @ a) ! i" unfolding 3 using 1 2 4 5 by (smt "0" add.commute add_diff_inverse_nat f_def f_permutes length_append nat_add_left_cancel_less nth_append permute_list_nth) qed show "y \ permute_list f ` cartesian_product A B" using ab_def 7 cartesian_product_memI'[of _ Q\<^sub>p] unfolding 0 by (metis assms(1) assms(2) image_eqI) qed qed thus ?thesis using assms f_permutes permutation_is_semialgebraic by metis qed lemma Qp_zero_subset_is_semialg: assumes "S \ carrier (Q\<^sub>p\<^bsup>0\<^esup>)" shows "is_semialgebraic 0 S" proof(cases "S = {}") case True then show ?thesis by (simp add: empty_is_semialgebraic) next case False then have "S = carrier (Q\<^sub>p\<^bsup>0\<^esup>)" using assms unfolding Qp_zero_carrier by blast then show ?thesis by (simp add: carrier_is_semialgebraic) qed lemma cartesian_product_empty_list: "cartesian_product A {[]} = A" "cartesian_product {[]} A = A" proof show "cartesian_product A {[]} \ A" apply(rule subsetI) unfolding cartesian_product_def by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) show "A \ cartesian_product A {[]}" apply(rule subsetI) unfolding cartesian_product_def by (smt append_Nil2 empty_iff insert_iff mem_Collect_eq) show "cartesian_product {[]} A = A" proof show "cartesian_product {[]} A \ A" apply(rule subsetI) unfolding cartesian_product_def by (smt append_self_conv2 bex_empty insert_compr mem_Collect_eq) show "A \ cartesian_product {[]} A" apply(rule subsetI) unfolding cartesian_product_def by blast qed qed lemma cartesian_product_singleton_factor_projection_is_semialg': assumes "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_semialgebraic (m+n) (cartesian_product A {b})" shows "is_semialgebraic m A" proof(cases "n > 0") case True show ?thesis proof(cases "m > 0") case T: True then show ?thesis using assms True cartesian_product_singleton_factor_projection_is_semialg by blast next case False then show ?thesis using Qp_zero_subset_is_semialg assms by blast qed next case False then have F0: "b = []" using assms Qp_zero_carrier by blast have "cartesian_product A {b} = A" unfolding F0 by (simp add: cartesian_product_empty_list(1)) then show ?thesis using assms False by (metis add.right_neutral gr0I) qed (**************************************************************************************************) (**************************************************************************************************) subsection \More on graphs of functions\ (**************************************************************************************************) (**************************************************************************************************) text\This section lays the groundwork for showing that semialgebraic functions are closed under various algebraic operations\ text\The take and drop functions on lists are polynomial maps\ lemma function_restriction: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ S" assumes "n \ k" shows "(g \ (take n)) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ S" proof fix x assume "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" then have "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(2) take_closed by blast then show "(g \ take n) x \ S" using assms comp_apply by (metis Pi_iff comp_def) qed lemma partial_pullback_restriction: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" shows "partial_pullback k (g \ take n) m S = split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" proof(rule equalityI) show "partial_pullback k (g \ take n) m S \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" proof fix x assume A: "x \ partial_pullback k (g \ take n) m S" obtain as bs where asbs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using partial_pullback_memE[of x k "g \ take n" m S] A cartesian_power_decomp[of x Q\<^sub>p k m] by metis have 0: "((g \ (take n)) as)#bs \ S" using asbs_def partial_pullback_memE'[of as k bs m x] A by blast have 1: "(g (take n as))#bs \ S" using 0 by (metis comp_apply) have 2: "take n as @ bs \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" by (meson asbs_def assms(2) cartesian_power_concat(1) less_imp_le_nat take_closed) have 3: "(take n as)@bs \ (partial_pullback n g m S)" using 1 2 partial_pullback_memI[of "(take n as)@bs" n m g S] by (metis (mono_tags, opaque_lifting) asbs_def assms(2) local.partial_image_def nat_less_le partial_image_eq subsetD subset_refl take_closed) have 4: "drop n as \ (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" using asbs_def assms(2) drop_closed by blast show " x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" using split_cartesian_product_memI[of "take n as" bs "partial_pullback n g m S" "drop n as" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p "n + m" "k - n" n ] 4 by (metis (no_types, lifting) "3" append.assoc append_take_drop_id asbs_def assms(2) cartesian_power_car_memE less_imp_le_nat partial_pullback_memE(1) subsetI take_closed) qed show "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ partial_pullback k (g \ take n) m S" proof fix x assume A: "x \ split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>))" show "x \ partial_pullback k (g \ take n) m S" proof(rule partial_pullback_memI) have 0: "(partial_pullback n g m S) \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using partial_pullback_closed by blast then have "split_cartesian_product (n + m) (k - n) n (partial_pullback n g m S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)) \ carrier (Q\<^sub>p\<^bsup>n + m + (k - n)\<^esup>)" using assms A split_cartesian_product_closed[of "partial_pullback n g m S" Q\<^sub>p "n + m" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" "k - n" n] using le_add1 by blast then show P: "x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)" by (smt A Nat.add_diff_assoc2 add.commute add_diff_cancel_left' assms(2) le_add1 less_imp_le_nat subsetD) have "take n x @ drop (n + (k - n)) x \ partial_pullback n g m S" using 0 A split_cartesian_product_memE[of x "n + m" "k - n" n "partial_pullback n g m S" "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)" Q\<^sub>p] le_add1 by blast have 1: "g (take n x) # drop k x \ S" using partial_pullback_memE by (metis (no_types, lifting) \take n x @ drop (n + (k - n)) x \ partial_pullback n g m S\ \x \ carrier (Q\<^sub>p\<^bsup>k+m\<^esup>)\ add.assoc assms(2) cartesian_power_drop le_add1 le_add_diff_inverse less_imp_le_nat partial_pullback_memE' take_closed) have 2: "g (take n x) = (g \ take n) (take k x)" using assms P comp_apply[of g "take n" "take k x"] by (metis add.commute append_same_eq append_take_drop_id less_imp_add_positive take_add take_drop) then show "(g \ take n) (take k x) # drop k x \ S" using "1" by presburger qed qed qed lemma comp_take_is_semialg: assumes "is_semialg_function n g" assumes "n < k" assumes "0 < n" shows "is_semialg_function k (g \ (take n))" proof(rule is_semialg_functionI) show "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" using assms function_restriction[of g n "carrier Q\<^sub>p" k] dual_order.strict_implies_order is_semialg_function_closed by blast show "\ka S. S \ semialg_sets (1 + ka) \ is_semialgebraic (k + ka) (partial_pullback k (g \ take n) ka S)" proof- fix l S assume A: "S \ semialg_sets (1 + l)" have 0: "is_semialgebraic (n + l) (partial_pullback n g l S) " using assms A is_semialg_functionE is_semialgebraicI by blast have "is_semialgebraic (n + l + (k - n)) (split_cartesian_product (n + l) (k - n) n (partial_pullback n g l S) (carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)))" using A 0 split_cartesian_product_is_semialgebraic[of _ _ "partial_pullback n g l S" _ "carrier (Q\<^sub>p\<^bsup>k - n\<^esup>)"] add_gr_0 assms(2) assms(3) carrier_is_semialgebraic le_add1 zero_less_diff by presburger then show "is_semialgebraic (k + l) (partial_pullback k (g \ take n) l S)" using partial_pullback_restriction[of g n k l S] by (metis (no_types, lifting) add.assoc add.commute assms(1) assms(2) is_semialg_function_closed le_add_diff_inverse less_imp_le_nat) qed qed text\Restriction of a graph to a semialgebraic domain\ lemma graph_formula: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "graph n g = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). g (take n as) = as!n}" using assms graph_memI fun_graph_def[of Q\<^sub>p n g] by (smt Collect_cong Suc_eq_plus1 graph_memE(1) graph_mem_closed mem_Collect_eq) definition restricted_graph where "restricted_graph n g S = {as \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>). take n as \ S \ g (take n as) = as!n }" lemma restricted_graph_closed: "restricted_graph n g S \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" by (metis (no_types, lifting) mem_Collect_eq restricted_graph_def subsetI) lemma restricted_graph_memE: assumes "a \ restricted_graph n g S" shows "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" "take n a \ S" "g (take n a) = a!n" using assms using restricted_graph_closed apply blast apply (metis (no_types, lifting) assms mem_Collect_eq restricted_graph_def) using assms unfolding restricted_graph_def by blast lemma restricted_graph_mem_formula: assumes "a \ restricted_graph n g S" shows "a = (take n a)@[g (take n a)]" proof- have "length a = Suc n" using assms by (metis (no_types, lifting) cartesian_power_car_memE mem_Collect_eq restricted_graph_def) then have "a = (take n a)@[a!n]" by (metis append_eq_append_conv_if hd_drop_conv_nth lessI take_hd_drop) then show ?thesis by (metis assms restricted_graph_memE(3)) qed lemma restricted_graph_memI: assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" assumes "take n a \ S" assumes "g (take n a) = a!n" shows "a \ restricted_graph n g S" using assms restricted_graph_def by blast lemma restricted_graph_memI': assumes "a \ S" assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "(a@[g a]) \ restricted_graph n g S" proof- have "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms(1) assms(3) by blast then have "g a \ carrier Q\<^sub>p" using assms by blast then have 0: "a @ [g a] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" using assms by (metis (no_types, lifting) add.commute cartesian_power_append plus_1_eq_Suc subsetD) have 1: "take n (a @ [g a]) \ S" using assms by (metis (no_types, lifting) append_eq_conv_conj cartesian_power_car_memE subsetD) show ?thesis using assms restricted_graph_memI[of "a@[g a]" n S g] by (metis "0" \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ append_eq_conv_conj cartesian_power_car_memE nth_append_length) qed lemma restricted_graph_subset: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S \ graph n g" proof fix x assume A: "x \ restricted_graph n g S" show "x \ graph n g" apply(rule graph_memI) using assms(1) apply blast using A restricted_graph_memE(3) apply blast by (metis A add.commute plus_1_eq_Suc restricted_graph_memE(1)) qed lemma restricted_graph_subset': assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" proof fix a assume A: "a \ restricted_graph n g S" then have "a = (take n a)@[g (take n a)]" using restricted_graph_mem_formula by blast then show "a \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using cartesian_product_memI' A unfolding restricted_graph_def by (metis (mono_tags, lifting) assms(2) last_closed' mem_Collect_eq subsetI Qp.to_R1_closed) qed lemma restricted_graph_intersection: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" proof show "restricted_graph n g S \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" using assms restricted_graph_subset restricted_graph_subset' by (meson Int_subset_iff) show "graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)) \ restricted_graph n g S" proof fix x assume A: " x \ graph n g \ cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))" show "x \ restricted_graph n g S" apply(rule restricted_graph_memI) using A graph_memE[of g n x] apply (metis (no_types, lifting) Int_iff add.commute assms(1) graph_mem_closed plus_1_eq_Suc) using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] using assms(2) apply blast using A graph_memE[of g n x] cartesian_product_memE[of x S "carrier (Q\<^sub>p\<^bsup>1\<^esup>)" Q\<^sub>p n] using assms(1) by blast qed qed lemma restricted_graph_is_semialgebraic: assumes "is_semialg_function n g" assumes "is_semialgebraic n S" shows "is_semialgebraic (n+1) (restricted_graph n g S)" proof- have 0: "restricted_graph n g S = graph n g \ (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" using assms is_semialg_function_closed is_semialgebraic_closed restricted_graph_intersection by presburger have 1: "is_semialgebraic (n + 1) (graph n g)" using assms semialg_graph by blast have 2: "is_semialgebraic (n + 1) (cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>)))" using cartesian_product_is_semialgebraic[of n S 1 "carrier (Q\<^sub>p\<^bsup>1\<^esup>)"] assms carrier_is_semialgebraic less_one by presburger then show ?thesis using 0 1 2 intersection_is_semialg[of "n+1" "graph n g" "cartesian_product S (carrier (Q\<^sub>p\<^bsup>1\<^esup>))"] by presburger qed lemma take_closed: assumes "n \ k" assumes "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" shows "take n x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms take_closed by blast lemma take_compose_closed: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" shows "g \ take n \ carrier (Q\<^sub>p\<^bsup>k\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" then have "(take n x) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using assms less_imp_le_nat take_closed by blast then have "g (take n x) \ carrier Q\<^sub>p" using assms(1) by blast then show "(g \ take n) x \ carrier Q\<^sub>p" using comp_apply[of g "take n" x] by presburger qed lemma take_graph_formula: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "n < k" assumes "0 < n" shows "graph k (g \ (take n)) = {as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>). g (take n as) = as!k}" proof- have "\as. as \ carrier (Q\<^sub>p\<^bsup>k+1\<^esup>) \ (g \ take n) (take k as) = g (take n as) " using assms comp_apply take_take[of n k] proof - fix as :: "((nat \ int) \ (nat \ int)) set list" show "(g \ take n) (take k as) = g (take n as)" by (metis (no_types) \n < k\ comp_eq_dest_lhs min.strict_order_iff take_take) qed then show ?thesis using take_compose_closed[of g n k] assms comp_apply[of g "take n"] graph_formula[of "g \ (take n)" k] by (smt Collect_cong Suc_eq_plus1) qed lemma graph_memI': assumes "a \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" assumes "take n a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "g (take n a) = a!n" shows "a \ graph n g" using assms fun_graph_def[of Q\<^sub>p n g] by (smt cartesian_power_car_memE eq_imp_le lessI mem_Collect_eq take_Suc_conv_app_nth take_all) lemma graph_memI'': assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "(a@[g a]) \ graph n g " using assms fun_graph_def by blast lemma graph_as_restricted_graph: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "graph n f = restricted_graph n f (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" apply(rule equalityI) apply (metis Suc_eq_plus1 assms graph_memE(1) graph_memE(3) graph_mem_closed restricted_graph_memI subsetI) by (simp add: assms restricted_graph_subset) definition double_graph where "double_graph n f g = {as \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>). f (take n as) = as!n \ g (take n as) = as!(n + 1)}" lemma double_graph_rep: assumes "g \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" shows "double_graph n f g = restricted_graph (n + 1) (g \ take n) (graph n f)" proof show "double_graph n f g \ restricted_graph (n + 1) (g \ take n) (graph n f)" proof fix x assume A: "x \ double_graph n f g" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>) \ f (take n x) = x!n \ g (take n x) = x!(n + 1)" using double_graph_def by blast have 1: "take (n+1) x \ graph n f" apply(rule graph_memI) using assms(2) apply blast apply (metis "0" append_eq_conv_conj cartesian_power_car_memE le_add1 length_take less_add_same_cancel1 less_numeral_extra(1) min.absorb2 nth_take take_add) by (metis (no_types, opaque_lifting) "0" Suc_eq_plus1 Suc_n_not_le_n add_cancel_right_right dual_order.antisym le_iff_add not_less_eq_eq one_add_one plus_1_eq_Suc take_closed) show " x \ restricted_graph (n + 1) (g \ take n) (graph n f)" apply(rule restricted_graph_memI) apply (metis "0" One_nat_def add_Suc_right numeral_2_eq_2) using "1" apply blast using 0 take_take[of n "n + 1" x] comp_apply by (metis le_add1 min.absorb1) qed show "restricted_graph (n + 1) (g \ take n) (graph n f) \ double_graph n f g" proof fix x assume A: "x \ restricted_graph (n + 1) (g \ take n) (graph n f)" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc (n + 1)\<^esup>) \ take (n + 1) x \ graph n f \ (g \ take n) (take (n + 1) x) = x ! (n + 1)" using restricted_graph_memE[of x "n+1" "(g \ take n)" "graph n f" ] by blast then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" using 0 by (metis Suc_1 add_Suc_right) have 2: " f (take n x) = x ! n" using 0 take_take[of n "n + 1" x] graph_memE[of f n "take (n + 1) x"] by (metis assms(2) le_add1 less_add_same_cancel1 less_numeral_extra(1) min.absorb1 nth_take) have 3: "g (take n x) = x ! (n + 1)" using 0 comp_apply take_take[of n "n + 1" x] by (metis le_add1 min.absorb1) then show "x \ double_graph n f g" unfolding double_graph_def using 1 2 3 by blast qed qed lemma double_graph_is_semialg: assumes "n > 0" assumes "is_semialg_function n f" assumes "is_semialg_function n g" shows "is_semialgebraic (n+2) (double_graph n f g)" using double_graph_rep[of g n f] assms restricted_graph_is_semialgebraic[of n "g \ take n" "graph n f"] by (metis (no_types, lifting) Suc_eq_plus1 add_Suc_right is_semialg_function_closed less_add_same_cancel1 less_numeral_extra(1) one_add_one restricted_graph_is_semialgebraic comp_take_is_semialg semialg_graph) definition add_vars :: "nat \ nat \ padic_tuple \ padic_number" where "add_vars i j as = as!i \\<^bsub>Q\<^sub>p\<^esub> as!j" lemma add_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" assumes "j < n" shows "add_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" unfolding add_vars_def using assms eval_at_point_add[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] eval_pvar by (metis pvar_closed) lemma add_vars_is_semialg: assumes "i < n" assumes "j < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (add_vars i j)" proof- have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] by blast then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] by blast then show ?thesis using assms add_vars_rep semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "add_vars i j" ] by (metis (no_types, lifting) restrict_ext) qed definition mult_vars :: "nat \ nat \ padic_tuple \ padic_number" where "mult_vars i j as = as!i \ as!j" lemma mult_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" assumes "j < n" shows "mult_vars i j as = Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)) as" unfolding mult_vars_def using assms eval_at_point_mult[of as n "pvar Q\<^sub>p i" "pvar Q\<^sub>p j"] eval_pvar[of i n as] eval_pvar[of j n as ] by (metis pvar_closed) lemma mult_vars_is_semialg: assumes "i < n" assumes "j < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (mult_vars i j)" proof- have "pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] by blast then have "is_semialg_function n (Qp_ev (pvar Q\<^sub>p i \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> pvar Q\<^sub>p j))" using assms poly_is_semialg[of "(pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j)"] by blast then show ?thesis using assms mult_vars_rep semialg_function_on_carrier[of n "Qp_ev ((pvar Q\<^sub>p i) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (pvar Q\<^sub>p j))" "mult_vars i j" ] by (metis (no_types, lifting) restrict_ext) qed definition minus_vars :: "nat \ padic_tuple \ padic_number" where "minus_vars i as = \\<^bsub>Q\<^sub>p\<^esub> as!i" lemma minus_vars_rep: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "i < n" shows "minus_vars i as = Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)) as" unfolding minus_vars_def using assms eval_pvar[of i n as] eval_at_point_a_inv[of as n "pvar Q\<^sub>p i"] by (metis pvar_closed) lemma minus_vars_is_semialg: assumes "i < n" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "is_semialg_function n (minus_vars i)" proof- have 0: "pvar Q\<^sub>p i \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms pvar_closed[of ] Qp.cring_axioms by presburger have "is_semialg_function n (Qp_ev (\\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(pvar Q\<^sub>p i)))" apply(rule poly_is_semialg ) using "0" by blast then show ?thesis using assms minus_vars_rep[of a i n] semialg_function_on_carrier[of n _ "minus_vars i" ] by (metis (no_types, lifting) minus_vars_rep restrict_ext) qed definition extended_graph where "extended_graph n f g h = {as \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>). f (take n as) = as!n \ g (take n as) = as! (n + 1) \ h [(f (take n as)),(g (take n as))] = as! (n + 2) }" lemma extended_graph_rep: "extended_graph n f g h = restricted_graph (n + 2) (h \ (drop n)) (double_graph n f g)" proof show "extended_graph n f g h \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" proof fix x assume "x \ extended_graph n f g h" then have A: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \f (take n x) = x!n \ g (take n x) = x! (n + 1) \ h [(f (take n x)),(g (take n x))] = x! (n + 2)" unfolding extended_graph_def by blast then have 0: "take (n + 2) x \ carrier (Q\<^sub>p\<^bsup>n+2\<^esup>)" proof - have "Suc (Suc n) \ n + numeral (num.One + num.Bit0 num.One)" by simp then show ?thesis by (metis (no_types) \x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>) \ f (take n x) = x ! n \ g (take n x) = x ! (n + 1) \ h [f (take n x), g (take n x)] = x ! (n + 2)\ add_2_eq_Suc' add_One_commute semiring_norm(5) take_closed) qed have 1: "f (take n (take (n + 2) x)) = (take (n + 2) x) ! n" using A by (metis Suc_1 add.commute append_same_eq append_take_drop_id less_add_same_cancel1 nth_take take_add take_drop zero_less_Suc) have 2: " g (take n (take (n + 2) x)) = (take (n + 2) x) ! (n + 1)" using A by (smt add.assoc add.commute append_same_eq append_take_drop_id less_add_same_cancel1 less_numeral_extra(1) nth_take one_add_one take_add take_drop) then have 3: "take (n + 2) x \ double_graph n f g" unfolding double_graph_def using 0 1 2 by blast have 4: "drop n (take (n + 2) x) = [(f (take n x)),(g (take n x))]" proof- have 40: "take (n + 2) x ! (n + 1) = x! (n + 1)" by (metis add.commute add_2_eq_Suc' lessI nth_take plus_1_eq_Suc) have 41: "take (n + 2) x ! n = x! n" by (metis Suc_1 less_SucI less_add_same_cancel1 less_numeral_extra(1) nth_take) have 42: "take (n + 2) x ! (n + 1) = g (take n x)" using 40 A by blast have 43: "take (n + 2) x ! n = f (take n x)" using 41 A by blast show ?thesis using A 42 43 by (metis "0" add_cancel_right_right cartesian_power_car_memE cartesian_power_drop le_add_same_cancel1 nth_drop pair_id zero_le_numeral) qed then have 5: "(h \ drop n) (take (n + 2) x) = x ! (n + 2)" using 3 A by (metis add_2_eq_Suc' comp_eq_dest_lhs) show "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" using restricted_graph_def A 3 5 by (metis (no_types, lifting) One_nat_def Suc_1 add_Suc_right numeral_3_eq_3 restricted_graph_memI) qed show "restricted_graph (n + 2) (h \ drop n) (double_graph n f g) \ extended_graph n f g h" proof fix x assume A: "x \ restricted_graph (n + 2) (h \ drop n) (double_graph n f g)" then have 0: "take (n+2) x \ double_graph n f g" using restricted_graph_memE(2) by blast have 1: "(h \ drop n) (take (n+2) x) = x! (n+2) " by (meson A restricted_graph_memE(3) padic_fields_axioms) have 2: "x \ carrier (Q\<^sub>p\<^bsup>n+3\<^esup>)" using A by (metis (no_types, opaque_lifting) Suc3_eq_add_3 add.commute add_2_eq_Suc' restricted_graph_closed subsetD) have 3: "length x = n + 3" using "2" cartesian_power_car_memE by blast have 4: "drop n (take (n+2) x) = [x!n, x!(n+1)]" proof- have "length (take (n+2) x) = n+2" by (simp add: "3") then have 40:"length (drop n (take (n+2) x)) = 2" by (metis add_2_eq_Suc' add_diff_cancel_left' length_drop) have 41: "(drop n (take (n+2) x))!0 = x!n" using 3 by (metis Nat.add_0_right \length (take (n + 2) x) = n + 2\ add_gr_0 le_add1 less_add_same_cancel1 less_numeral_extra(1) nth_drop nth_take one_add_one) have 42: "(drop n (take (n+2) x))!1 = x!(n+1)" using 3 nth_take nth_drop A by (metis add.commute le_add1 less_add_same_cancel1 less_numeral_extra(1) one_add_one take_drop) show ?thesis using 40 41 42 by (metis pair_id) qed have "(take n x) = take n (take (n+2) x)" using take_take 3 by (metis le_add1 min.absorb1) then have 5: "f (take n x) = x ! n" using 0 double_graph_def[of n f g] 3 by (smt Suc_1 less_add_same_cancel1 mem_Collect_eq nth_take zero_less_Suc) have 6: "g (take n x) = x ! (n + 1) " using 0 double_graph_def[of n f g] 3 take_take[of n "n+2" x] by (smt Suc_1 \take n x = take n (take (n + 2) x)\ add_Suc_right lessI mem_Collect_eq nth_take) have 7: " h [f (take n x), g (take n x)] = x ! (n + 2)" using 4 A comp_apply by (metis "1" "5" "6") show " x \ extended_graph n f g h" unfolding extended_graph_def using 2 5 6 7 A by blast qed qed lemma function_tuple_eval_closed: assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "function_tuple_eval Q\<^sub>p n fs x \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" using function_tuple_eval_closed[of Q\<^sub>p n fs x] assms by blast definition k_graph where "k_graph n fs = {x \ carrier (Q\<^sub>p\<^bsup>n + length fs\<^esup>). x = (take n x)@ (function_tuple_eval Q\<^sub>p n fs (take n x)) }" lemma k_graph_memI: assumes "is_function_tuple Q\<^sub>p n fs" assumes "x = as@function_tuple_eval Q\<^sub>p n fs as" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "x \ k_graph n fs" proof- have "take n x = as" using assms by (metis append_eq_conv_conj cartesian_power_car_memE) then show ?thesis unfolding k_graph_def using assms by (smt append_eq_conv_conj cartesian_power_car_memE cartesian_power_car_memI'' length_append local.function_tuple_eval_closed mem_Collect_eq) qed text\composing a function with a function tuple\ lemma Qp_function_tuple_comp_closed: assumes "f \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier Q\<^sub>p" assumes "length fs = n" assumes "is_function_tuple Q\<^sub>p m fs" shows "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms function_tuple_comp_closed by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Tuples of Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Predicate for a tuple of semialgebraic functions\ definition is_semialg_function_tuple where "is_semialg_function_tuple n fs = (\ f \ set fs. is_semialg_function n f)" lemma is_semialg_function_tupleI: assumes "\ f. f \ set fs \ is_semialg_function n f" shows "is_semialg_function_tuple n fs" using assms is_semialg_function_tuple_def by blast lemma is_semialg_function_tupleE: assumes "is_semialg_function_tuple n fs" assumes "i < length fs" shows "is_semialg_function n (fs ! i)" by (meson assms(1) assms(2) in_set_conv_nth is_semialg_function_tuple_def padic_fields_axioms) lemma is_semialg_function_tupleE': assumes "is_semialg_function_tuple n fs" assumes "f \ set fs" shows "is_semialg_function n f" using assms(1) assms(2) is_semialg_function_tuple_def by blast lemma semialg_function_tuple_is_function_tuple: assumes "is_semialg_function_tuple n fs" shows "is_function_tuple Q\<^sub>p n fs" apply(rule is_function_tupleI) using assms is_semialg_function_closed is_semialg_function_tupleE' by blast lemma const_poly_function_tuple_comp_is_semialg: assumes "n > 0" assumes "is_semialg_function_tuple n fs" assumes "a \ carrier Q\<^sub>p" shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (Qp_to_IP a))" apply(rule semialg_function_on_carrier[of n "Qp_ev (Qp_to_IP a)"]) using poly_is_semialg[of "(Qp_to_IP a)"] using assms(1) assms(3) Qp_to_IP_car apply blast using poly_function_tuple_comp_eq[of n fs "(Qp_to_IP a)"] assms unfolding restrict_def by (metis (no_types, opaque_lifting) eval_at_point_const poly_function_tuple_comp_constant semialg_function_tuple_is_function_tuple) lemma pvar_poly_function_tuple_comp_is_semialg: assumes "n > 0" assumes "is_semialg_function_tuple n fs" assumes "i < length fs" shows "is_semialg_function n (poly_function_tuple_comp Q\<^sub>p n fs (pvar Q\<^sub>p i))" apply(rule semialg_function_on_carrier[of n "fs!i"]) using assms(2) assms(3) is_semialg_function_tupleE apply blast by (metis assms(2) assms(3) poly_function_tuple_comp_pvar restrict_ext semialg_function_tuple_is_function_tuple) text\Polynomial functions with semialgebraic coefficients\ definition point_to_univ_poly :: "nat \ padic_tuple \ padic_univ_poly" where "point_to_univ_poly n a = ring_cfs_to_univ_poly n a" definition tuple_partial_image where "tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs (take m x))@(drop m x)" lemma tuple_partial_image_closed: assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "tuple_partial_image n fs x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using assms unfolding tuple_partial_image_def by (meson cartesian_power_concat(1) cartesian_power_drop function_tuple_eval_closed le_add1 take_closed) lemma tuple_partial_image_indices: assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" assumes "i < length fs" shows "(tuple_partial_image n fs x) ! i = (fs!i) (take n x)" proof- have 0: "(function_tuple_eval Q\<^sub>p n fs (take n x)) ! i = (fs!i) (take n x)" using assms unfolding function_tuple_eval_def by (meson nth_map) have 1: "length (function_tuple_eval Q\<^sub>p n fs (take n x)) > i" by (metis assms(4) function_tuple_eval_def length_map) show ?thesis using 0 1 assms unfolding tuple_partial_image_def by (metis nth_append) qed lemma tuple_partial_image_indices': assumes "length fs > 0" assumes "is_function_tuple Q\<^sub>p n fs" assumes "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" assumes "i < l" shows "(tuple_partial_image n fs x) ! (length fs + i) = x!(n + i)" using assms unfolding tuple_partial_image_def by (metis (no_types, lifting) cartesian_power_car_memE function_tuple_eval_closed le_add1 nth_append_length_plus nth_drop take_closed) definition tuple_partial_pullback where "tuple_partial_pullback n fs l S = ((tuple_partial_image n fs)-`S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" lemma tuple_partial_pullback_memE: assumes "as \ tuple_partial_pullback m fs l S" shows "as \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" "tuple_partial_image m fs as \ S" using assms apply (metis (no_types, opaque_lifting) Int_iff add.commute tuple_partial_pullback_def) using assms unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_closed: "tuple_partial_pullback m fs l S \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using tuple_partial_pullback_memE by blast lemma tuple_partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" assumes "is_function_tuple Q\<^sub>p m fs" assumes "((function_tuple_eval Q\<^sub>p m fs) (take m as))@(drop m as) \ S" shows "as \ tuple_partial_pullback m fs k S" using assms unfolding tuple_partial_pullback_def tuple_partial_image_def by blast lemma tuple_partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" shows "tuple_partial_image n fs x = ((function_tuple_eval Q\<^sub>p n fs) as)@bs" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(3) cartesian_power_car_memE) have 1: "drop n x = bs" by (metis "0" append_take_drop_id assms(3) same_append_eq) show ?thesis using assms 0 1 unfolding tuple_partial_image_def by presburger qed lemma tuple_partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" assumes "x = as @ bs" assumes "x \ tuple_partial_pullback n fs k S" shows "(function_tuple_eval Q\<^sub>p n fs as)@bs \ S" using tuple_partial_pullback_memE[of x n fs k S] tuple_partial_image_def[of n fs x] by (metis assms(1) assms(2) assms(3) assms(4) tuple_partial_image_eq) text\tuple partial pullbacks have the same algebraic properties as pullbacks\ lemma tuple_partial_pullback_intersect: "tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_union: "tuple_partial_pullback m f l (S1 \ S2) = (tuple_partial_pullback m f l S1) \ (tuple_partial_pullback m f l S2)" unfolding tuple_partial_pullback_def by blast lemma tuple_partial_pullback_complement: assumes "is_function_tuple Q\<^sub>p m fs" shows "tuple_partial_pullback m fs l ((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - (tuple_partial_pullback m fs l S) " apply(rule equalityI) using tuple_partial_pullback_def[of m fs l "((carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) - S)"] tuple_partial_pullback_def[of m fs l S] apply blast proof fix x assume A: " x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) - tuple_partial_pullback m fs l S" show " x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>) - S) " apply(rule tuple_partial_pullback_memI) using A apply blast using assms apply blast proof have 0: "drop m x \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" by (meson A DiffD1 cartesian_power_drop) have 1: "take m x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A by (meson DiffD1 le_add1 take_closed) show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using 0 1 assms using cartesian_power_concat(1) function_tuple_eval_closed by blast show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ S" using A unfolding tuple_partial_pullback_def tuple_partial_image_def by blast qed qed lemma tuple_partial_pullback_carrier: assumes "is_function_tuple Q\<^sub>p m fs" shows "tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)) = carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" apply(rule equalityI) using tuple_partial_pullback_memE(1) apply blast proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" show "x \ tuple_partial_pullback m fs l (carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>))" apply(rule tuple_partial_pullback_memI) using A cartesian_power_drop[of x m l] take_closed assms apply blast using assms apply blast proof- have "function_tuple_eval Q\<^sub>p m fs (take m x) \ carrier (Q\<^sub>p\<^bsup>length fs\<^esup>)" using A take_closed assms function_tuple_eval_closed le_add1 by blast then show "function_tuple_eval Q\<^sub>p m fs (take m x) @ drop m x \ carrier (Q\<^sub>p\<^bsup>length fs + l\<^esup>)" using cartesian_power_drop[of x m l] A cartesian_power_concat(1) by blast qed qed definition is_semialg_map_tuple where "is_semialg_map_tuple m fs = (is_function_tuple Q\<^sub>p m fs \ (\l \ 0. \S \ semialg_sets ((length fs) + l). is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)))" lemma is_semialg_map_tuple_closed: assumes "is_semialg_map_tuple m fs" shows "is_function_tuple Q\<^sub>p m fs" using is_semialg_map_tuple_def assms by blast lemma is_semialg_map_tupleE: assumes "is_semialg_map_tuple m fs" assumes "is_semialgebraic ((length fs) + l) S" shows " is_semialgebraic (m + l) (tuple_partial_pullback m fs l S)" using is_semialg_map_tuple_def[of m fs] assms is_semialgebraicE[of "((length fs) + l)" S] by blast lemma is_semialg_map_tupleI: assumes "is_function_tuple Q\<^sub>p m fs" assumes "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" shows "is_semialg_map_tuple m fs" using assms unfolding is_semialg_map_tuple_def by blast text\Semialgebraicity for maps can be verified on basic semialgebraic sets\ lemma is_semialg_map_tupleI': assumes "is_function_tuple Q\<^sub>p m fs" assumes "\k S. S \ basic_semialgs ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" shows "is_semialg_map_tuple m fs" apply(rule is_semialg_map_tupleI) using assms(1) apply blast proof- show "\k S. S \ semialg_sets ((length fs) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- fix k S assume A: "S \ semialg_sets ((length fs) + k)" show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" apply(rule gen_boolean_algebra.induct[of S "carrier (Q\<^sub>p\<^bsup>length fs + k\<^esup>)" "basic_semialgs ((length fs) + k)"]) using A unfolding semialg_sets_def apply blast using tuple_partial_pullback_carrier assms carrier_is_semialgebraic plus_1_eq_Suc apply presburger using assms(1) assms(2) carrier_is_semialgebraic intersection_is_semialg tuple_partial_pullback_carrier tuple_partial_pullback_intersect apply presburger using tuple_partial_pullback_union union_is_semialgebraic apply presburger using assms(1) complement_is_semialg tuple_partial_pullback_complement plus_1_eq_Suc by presburger qed qed text\ The goal of this section is to show that tuples of semialgebraic functions are semialgebraic maps. \ text\The function $(x_0, x, y) \mapsto (x_0, f(x), y)$\ definition twisted_partial_image where "twisted_partial_image n m f xs = (take n xs)@ partial_image m f (drop n xs)" text\The set ${(x_0, x, y) \mid (x_0, f(x), y) \in S}$\ text\Convention: a function which produces a subset of (Qp (i + j +k)) will receive the 3 arity parameters in sequence, at the very beginning of the function\ definition twisted_partial_pullback where "twisted_partial_pullback n m l f S = ((twisted_partial_image n m f)-`S) \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" lemma twisted_partial_pullback_memE: assumes "as \ twisted_partial_pullback n m l f S" shows "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" "twisted_partial_image n m f as \ S" using assms apply (metis (no_types, opaque_lifting) Int_iff add.commute twisted_partial_pullback_def subset_iff) using assms unfolding twisted_partial_pullback_def by blast lemma twisted_partial_pullback_closed: "twisted_partial_pullback n m l f S \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" using twisted_partial_pullback_memE(1) by blast lemma twisted_partial_pullback_memI: assumes "as \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" assumes "(take n as)@((f (take m (drop n as)))#(drop (n + m) as)) \ S" shows "as \ twisted_partial_pullback n m l f S" using assms unfolding twisted_partial_pullback_def twisted_partial_image_def by (metis (no_types, lifting) IntI add.commute drop_drop local.partial_image_def vimageI) lemma twisted_partial_image_eq: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" assumes "x = as @ bs @ cs" shows "twisted_partial_image n m f x = as@((f bs)#cs)" proof- have 0: "(take n x) = as" by (metis append_eq_conv_conj assms(1) assms(4) cartesian_power_car_memE) have 1: "twisted_partial_image n m f x = as@(partial_image m f (bs@cs))" using 0 assms twisted_partial_image_def by (metis append_eq_conv_conj cartesian_power_car_memE) have 2: "(partial_image m f (bs@cs)) = (f bs)#cs" using partial_image_eq[of bs m cs l "bs@cs" f] assms by blast show ?thesis using assms 0 1 2 by (metis ) qed lemma twisted_partial_pullback_memE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "cs \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" assumes "x = as @ bs @ cs" assumes "x \ twisted_partial_pullback n m l f S" shows "as@((f bs)#cs) \ S" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) twisted_partial_image_eq twisted_partial_pullback_memE(2)) text\partial pullbacks have the same algebraic properties as pullbacks\ text\permutation which moves the entry at index \i\ to 0\ definition twisting_permutation where "twisting_permutation (i::nat) = (\j. if j < i then j + 1 else (if j = i then 0 else j))" lemma twisting_permutation_permutes: assumes "i < n" shows "twisting_permutation i permutes {..x. x > i \ twisting_permutation i x = x" unfolding twisting_permutation_def by auto have 1: "(\x. x \ {.. twisting_permutation i x = x)" using 0 assms by auto have 2: "(\y. \!x. twisting_permutation i x = y)" proof fix y show " \!x. twisting_permutation i x = y" proof(cases "y = 0") case True show "\!x. twisting_permutation i x = y" by (metis Suc_eq_plus1 True add_eq_0_iff_both_eq_0 less_nat_zero_code nat_neq_iff twisting_permutation_def zero_neq_one) next case False show ?thesis proof(cases "y \i") case True show ?thesis proof show "twisting_permutation i (y - 1) = y" using True by (metis False add.commute add_diff_inverse_nat diff_less gr_zeroI le_eq_less_or_eq less_imp_diff_less less_one twisting_permutation_def) show "\x. twisting_permutation i x = y \ x = y - 1" using True False twisting_permutation_def by force qed next case F: False then show ?thesis using False unfolding twisting_permutation_def by (metis add_leD1 add_leD2 add_le_same_cancel2 discrete le_numeral_extra(3) less_imp_not_less ) qed qed qed show ?thesis using 1 2 by (simp add: permutes_def) qed lemma twisting_permutation_action: assumes "length as = i" shows "permute_list (twisting_permutation i) (b#(as@bs)) = as@(b#bs)" proof- have 0: "length (permute_list (twisting_permutation i) (b#(as@bs))) = length (as@(b#bs))" by (metis add.assoc length_append length_permute_list list.size(4)) have "\j. j < length (as@(b#bs)) \ (permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" proof- fix j assume A: "j < length (as@(b#bs))" show "(permute_list (twisting_permutation i) (b#(as@bs))) ! j = (as@(b#bs)) ! j" proof(cases "j < i") case True then have T0: "twisting_permutation i j = j + 1" using twisting_permutation_def by auto then have T1: "(b # as @ bs) ! twisting_permutation i j = as!j" using assms by (simp add: assms True nth_append) have T2: "(permute_list (twisting_permutation i) (b # as @ bs)) ! j = as!j" proof- have "twisting_permutation i permutes {..length as = i\) then have "permute_list (fun_inv TI) (as@(b#bs)) = permute_list ((fun_inv TI) \ TI) (b#(as@bs))" using 0 1 by (metis TI_def fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv length_permute_list permute_list_compose) then show ?thesis by (metis "0" Nil_is_append_conv TI_def fun_inv_permute(3) length_greater_0_conv list.distinct(1) permute_list_id) qed lemma twisting_semialg: assumes "is_semialgebraic n S" assumes "n > i" shows "is_semialgebraic n ((permute_list ((twisting_permutation i)) ` S))" proof- obtain TI where TI_def: "TI = twisting_permutation i" by blast have 0: "TI permutes {..<(n::nat)}" using assms TI_def twisting_permutation_permutes[of i n] by blast have "(TI) permutes {.. i" shows "is_semialgebraic n ((permute_list (fun_inv (twisting_permutation i)) ` S))" proof- obtain TI where TI_def: "TI = twisting_permutation i" by blast have 0: "TI permutes {..<(n::nat)}" using assms TI_def twisting_permutation_permutes[of i n] by blast have "(fun_inv TI) permutes {..Defining a permutation that does: $(x0, x1, y) \mapsto (x_1, x0, y)$\ definition tp_1 where "tp_1 i j = (\ n. (if n n \ n < i + j then n - i else n)))" lemma permutes_I: assumes "\x. x \ S \ f x = x" assumes "\y. y \ S \ \!x \ S. f x = y" assumes "\x. x \ S \ f x \ S" shows "f permutes S" proof- have 0 : "(\x. x \ S \ f x = x) " using assms(1) by blast have 1: "(\y. \!x. f x = y)" proof fix y show "\!x. f x = y" apply(cases "y \ S") apply (metis "0" assms(2)) proof- assume "y \ S" then show "\!x. f x = y" by (metis assms(1) assms(3)) qed qed show ?thesis using assms 1 unfolding permutes_def by blast qed lemma tp_1_permutes: "(tp_1 (i::nat) j) permutes {..< i + j}" proof(rule permutes_I) show "\x. x \ {.. tp_1 i j x = x" proof- fix x assume A: "x \ {..y. y \ {.. \!x. x \ {.. tp_1 i j x = y" proof- fix y assume A: "y \ {..!x. x \ {.. tp_1 i j x = y" proof(cases "y < j") case True then have 0:"tp_1 i j (y + i) = y" by (simp add: tp_1_def) have 1: "\x. x \ y + i \ tp_1 i j x \ y" proof- fix x assume A: " x \ y + i" show "tp_1 i j x \ y" apply(cases "x < j") apply (metis A True add.commute le_add_diff_inverse le_eq_less_or_eq nat_neq_iff not_add_less1 tp_1_def trans_less_add2) by (metis A True add.commute le_add_diff_inverse less_not_refl tp_1_def trans_less_add1) qed show ?thesis using 0 1 by (metis A \\x. x \ {.. tp_1 i j x = x\) next case False then have "y - j < i" using A by auto then have "tp_1 i j (y - j) = y" using False tp_1_def by (simp add: tp_1_def) then show ?thesis by (smt A False \\x. x \ {.. tp_1 i j x = x\ add.commute add_diff_inverse_nat add_left_imp_eq less_diff_conv2 not_less tp_1_def padic_fields_axioms) qed qed show "\x. x \ {.. tp_1 i j x \ {.. {.. carrier (Q\<^sub>p\<^bsup>i\<^esup>)" assumes "b \ carrier (Q\<^sub>p\<^bsup>j\<^esup>)" assumes "c \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "permute_list (tp_1 i j) (b@a@c)= a@b@c" proof- have 0:"length (permute_list (tp_1 i j) (b@a@c))= length (a@b@c)" by (metis add.commute append.assoc length_append length_permute_list) have "\x. x < length (a@b@c) \ (permute_list (tp_1 i j) (b@a@c)) ! x= (a@b@c) ! x" proof- fix x assume A: "x < length (a@b@c)" have B: "length (a @ b @ c) = i + j + length c" using add.assoc assms(1) assms(2) cartesian_power_car_memE length_append by metis have C: "tp_1 i j permutes {..length (permute_list (tp_1 i j) (b @ a @ c)) = length (a @ b @ c)\ length_permute_list nth_append permute_list_nth) next case False show ?thesis proof(cases "x < i + j") case True then have T0: "(tp_1 i j x) = x - i" by (meson False not_less tp_1_def) have "x - i < length b" using E False True by linarith then have T1: "permute_list (tp_1 i j) (b@ a @ c) ! x = b!(x-i)" using nth_append by (metis A C T0 \length (permute_list (tp_1 i j) (b @ a @ c)) = length (a @ b @ c)\ length_permute_list permute_list_nth) then show ?thesis by (metis D False \x - i < length b\ nth_append) next case False then have "(tp_1 i j x) = x" by (meson tp_1_def trans_less_add1) then show ?thesis by (smt A C D E False add.commute add_diff_inverse_nat append.assoc length_append nth_append_length_plus permute_list_nth) qed qed qed then show ?thesis using 0 by (metis list_eq_iff_nth_eq) qed definition tw where "tw i j = permute_list (tp_1 j i)" lemma tw_is_semialg: assumes "n > 0" assumes "is_semialgebraic n S" assumes "n \ i + j" shows "is_semialgebraic n ((tw i j)`S)" unfolding tw_def using assms tp_1_permutes'[of j i "n - (j + i)"] permutation_is_semialgebraic[of n S] by (metis add.commute le_add_diff_inverse) lemma twisted_partial_pullback_factored: assumes "f \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ carrier Q\<^sub>p" assumes "S \ carrier (Q\<^sub>p\<^bsup>n+1+ l\<^esup>)" assumes "Y = partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" shows "twisted_partial_pullback n m l f S = (tw m n) ` Y" proof show "twisted_partial_pullback n m l f S \ tw m n ` Y" proof fix x assume A: "x \ twisted_partial_pullback n m l f S" then have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>n+m+l\<^esup>)" using twisted_partial_pullback_memE(1) by blast obtain a where a_def: "a = take n x" by blast obtain b where b_def: "b = take m (drop n x)" by blast obtain c where c_def: "c = (drop (n + m) x)" by blast have x_eq:"x = a@(b@c)" by (metis a_def append.assoc append_take_drop_id b_def c_def take_add) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis (no_types, lifting) a_def dual_order.trans le_add1 take_closed x_closed) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" proof- have "drop n x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" by (metis (no_types, lifting) add.assoc cartesian_power_drop x_closed) then show ?thesis using b_def le_add1 take_closed by blast qed have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using c_def cartesian_power_drop x_closed by blast have B: "a@((f b)#c) \ S" using A twisted_partial_pullback_memE' by (smt a_closed a_def add.commute append_take_drop_id b_closed b_def c_closed c_def drop_drop) have "permute_list (fun_inv (twisting_permutation n)) (a@((f b)#c)) = (f b)#(a@c)" using assms twisting_permutation_action'[of a n "f b" c] a_closed cartesian_power_car_memE by blast then have C: "(f b)#(a@c) \ (permute_list (fun_inv (twisting_permutation n)) ` S)" by (metis B image_eqI) have C: "b@(a@c) \ partial_pullback m f (n + l) (permute_list (fun_inv (twisting_permutation n)) ` S)" proof(rule partial_pullback_memI) show "b @ a @ c \ carrier (Q\<^sub>p\<^bsup>m + (n + l)\<^esup>)" using a_closed b_closed c_closed cartesian_power_concat(1) by blast have 0: "(take m (b @ a @ c)) = b" by (metis append.right_neutral b_closed cartesian_power_car_memE diff_is_0_eq diff_self_eq_0 take0 take_all take_append) have 1: "drop m (b @ a @ c) = a@c" by (metis "0" append_take_drop_id same_append_eq) show "f (take m (b @ a @ c)) # drop m (b @ a @ c) \ permute_list (fun_inv (twisting_permutation n)) ` S" using 0 1 C by presburger qed have D: "tw m n (b@(a@c)) = a@(b@c)" using assms tw_def a_closed b_closed c_closed by (metis tp_1_permutation_action x_eq) then show "x \ tw m n ` Y" using x_eq C assms by (metis image_eqI) qed show "tw m n ` Y \ twisted_partial_pullback n m l f S" proof fix x assume A: "x \ tw m n ` Y" then obtain y where y_def: "x = tw m n y \ y \ Y" by blast obtain as where as_def: "as \ (permute_list (fun_inv (twisting_permutation n)) ` S) \ as = partial_image m f y" using partial_pullback_memE by (metis assms(3) y_def) obtain s where s_def: "s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as" using as_def by blast obtain b where b_def: "b = take m y" by blast obtain a where a_def: "a = take n (drop m y)" by blast obtain c where c_def: "c = drop (n + m) y" by blast have y_closed: "y \ carrier (Q\<^sub>p\<^bsup>m + n + l\<^esup>)" by (metis add.assoc assms(3) partial_pullback_memE(1) y_def) then have y_eq: "y = b@a@c" using a_def b_def c_def by (metis append_take_drop_id drop_drop) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (metis a_def add.commute cartesian_power_drop le_add1 take_closed take_drop y_closed) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using add_leD2 b_def le_add1 take_closed y_closed by (meson trans_le_add1) have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using c_def cartesian_power_drop y_closed by (metis add.commute) have ac_closed: "a@c \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using a_closed c_closed cartesian_power_concat(1) by blast then have C: " local.partial_image m f y = f b # a @ c" using b_closed y_eq partial_image_eq[of b m "a@c" "n + l" y f] by blast then have as_eq: "as = (f b)#(a@c)" using as_def by force have B: "(tw m n) y = a@b@c" using y_eq tw_def[of n m] tp_1_permutation_action by (smt a_closed b_closed c_closed tw_def) then have "x = a@(b@c)" by (simp add: y_def) then have "twisted_partial_image n m f x = a@((f b)# c)" using a_closed b_closed c_closed twisted_partial_image_eq by blast then have D: "permute_list (twisting_permutation n) as = twisted_partial_image n m f x" using as_eq twisting_permutation_action[of a n "f b" c ] by (metis a_closed cartesian_power_car_memE) have "permute_list (twisting_permutation n) as \ S" proof- have S: "length s > n" using s_def assms cartesian_power_car_memE le_add1 le_neq_implies_less le_trans less_add_same_cancel1 less_one not_add_less1 by (metis (no_types, lifting) subset_iff) have "permute_list (twisting_permutation n) as = permute_list (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)" using fun_inv_def s_def by blast then have "permute_list (twisting_permutation n) as = permute_list ((twisting_permutation n) \ (fun_inv (twisting_permutation n))) s" using fun_inv_permute(2) fun_inv_permute(3) length_greater_0_conv length_permute_list twisting_permutation_permutes[of n "length s"] permute_list_compose[of "fun_inv (twisting_permutation n)" s "twisting_permutation n"] by (metis S permute_list_compose) then have "permute_list (twisting_permutation n) as = permute_list (id) s" by (metis S \permute_list (twisting_permutation n) as = permute_list (twisting_permutation n) (permute_list (fun_inv (twisting_permutation n)) s)\ fun_inv_permute(3) length_greater_0_conv length_permute_list permute_list_compose twisting_permutation_permutes) then have "permute_list (twisting_permutation n) as = s" by simp then show ?thesis using s_def by (simp add: \s \ S \ permute_list (fun_inv (twisting_permutation n)) s = as\) qed then show "x \ twisted_partial_pullback n m l f S" unfolding twisted_partial_pullback_def using D by (smt \x = a @ b @ c\ a_closed append.assoc append_eq_conv_conj b_closed c_closed cartesian_power_car_memE cartesian_power_concat(1) length_append list.inject local.partial_image_def twisted_partial_image_def twisted_partial_pullback_def twisted_partial_pullback_memI) qed qed lemma twisted_partial_pullback_is_semialgebraic: assumes "is_semialg_function m f" assumes "is_semialgebraic (n + 1 + l) S" shows "is_semialgebraic (n + m + l)(twisted_partial_pullback n m l f S)" proof- have "(fun_inv (twisting_permutation n)) permutes {.. carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" shows "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" apply(rule cartesian_power_car_memI) apply (smt ab_semigroup_add_class.add_ac(1) add.commute append_take_drop_id assms augment_def cartesian_power_car_memE cartesian_power_drop length_append) using assms cartesian_power_car_memE'' unfolding augment_def by (metis (no_types, opaque_lifting) append_take_drop_id cartesian_power_concat(2) nat_le_iff_add take_closed) lemma tuple_partial_image_factor: assumes "is_function_tuple Q\<^sub>p m fs" assumes "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" assumes "length fs = n" assumes "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" shows "tuple_partial_image m (fs@[f]) x = twisted_partial_image n m f (tuple_partial_image m fs (augment m x))" proof- obtain a where a_def: "a = take m x" by blast obtain b where b_def: "b = drop m x" by blast have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using a_def assms(4) le_add1 take_closed by (meson dual_order.trans) have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using assms(4) b_def cartesian_power_drop by (metis (no_types, lifting)) have A: "(augment m x) = a @ (a @ b)" using a_def augment_def b_def by blast have 0: "tuple_partial_image m fs (augment m x) = ((function_tuple_eval Q\<^sub>p m fs) a) @ a @ b" using A a_closed b_closed tuple_partial_image_eq[of a m "a@b" "m + l" "augment m x" fs] cartesian_power_concat(1) by blast have 1: "tuple_partial_image m (fs@[f]) x = ((function_tuple_eval Q\<^sub>p m fs a) @ [f a])@b" using 0 tuple_partial_image_eq[of a m b l x "fs@[f]"] unfolding function_tuple_eval_def by (metis (no_types, lifting) a_closed a_def append_take_drop_id b_closed b_def list.simps(8) list.simps(9) map_append) have 2: "tuple_partial_image m (fs@[f]) x = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" using 1 by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv2 same_append_eq) have 3: "tuple_partial_image m fs x = (function_tuple_eval Q\<^sub>p m fs a) @ b" using a_def b_def 2 tuple_partial_image_eq[of a m b l x fs ] assms tuple_partial_image_def by blast have 4: "twisted_partial_image n m f (tuple_partial_image m fs (augment m x)) = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)#b)" using twisted_partial_image_eq[of _ n _ m _ l] 0 assms(1) assms(3) b_closed local.a_closed local.function_tuple_eval_closed by blast show ?thesis using 2 4 by presburger qed definition diagonalize where "diagonalize n m S = S \ cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" lemma diagaonlize_is_semiaglebraic: assumes "is_semialgebraic (n + n + m) S" shows "is_semialgebraic (n + n + m) (diagonalize n m S)" proof(cases "m = 0") case True then have 0: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) = {[]}" unfolding cartesian_power_def by simp have 1: "\ n \ carrier (Q\<^sub>p\<^bsup>n+n\<^esup>)" using Qp.cring_axioms assms diagonalE(2) by blast then have "cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = \ n" using 0 cartesian_product_empty_right[of "\ n" Q\<^sub>p "n + n" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] by linarith then have "diagonalize n m S = S \ (\ n)" using diagonalize_def by presburger then show ?thesis using intersection_is_semialg True assms diag_is_semialgebraic by auto next case False have "is_semialgebraic (n + n + m) (cartesian_product (\ n) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" using carrier_is_semialgebraic[of m] cartesian_product_is_semialgebraic[of "n + n" "\ n" m "carrier (Q\<^sub>p\<^bsup>m\<^esup>)"] diag_is_semialgebraic[of n] False by blast then show ?thesis using intersection_is_semialg assms(1) diagonalize_def by presburger qed lemma list_segment_take: assumes "length a \n" shows "list_segment 0 n a = take n a" proof- have 0: "length (list_segment 0 n a) = length (take n a)" using assms unfolding list_segment_def by (metis (no_types, lifting) Groups.add_ac(2) add_diff_cancel_left' append_take_drop_id le_Suc_ex length_append length_drop length_map map_nth) have "\i. i < n \ list_segment 0 n a !i = take n a ! i" unfolding list_segment_def using assms by (metis add.left_neutral diff_zero nth_map_upt nth_take) then show ?thesis using 0 by (metis assms diff_zero le0 list_segment_length nth_equalityI) qed lemma augment_inverse_is_semialgebraic: assumes "is_semialgebraic (n+n+l) S" shows "is_semialgebraic (n+l) ((augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>))" proof- obtain Ps where Ps_def: "Ps = (var_list_segment 0 n)" by blast obtain Qs where Qs_def: "Qs = (var_list_segment n (n+l))" by blast obtain Fs where Fs_def: "Fs = Ps@Ps@Qs" by blast have 0: "is_poly_tuple (n+l) Ps" by (simp add: Ps_def var_list_segment_is_poly_tuple) have 1: "is_poly_tuple (n+l) Qs" by (simp add: Qs_def var_list_segment_is_poly_tuple) have 2: "is_poly_tuple (n+l) (Ps@Qs)" using Qp_is_poly_tuple_append[of "n+l" Ps Qs] by (metis (no_types, opaque_lifting) "0" "1" add.commute) have "is_poly_tuple (n+l) Fs" using 0 2 Qp_is_poly_tuple_append[of "n + l" Ps "Ps@Qs"] Fs_def assms by blast have 3: "\x. x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ augment n x = poly_map (n + l) Fs x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" have 30: "poly_map (n+l) Ps x = take n x" using Ps_def map_by_var_list_segment[of x "n + l" n 0] list_segment_take[of n x] cartesian_power_car_memE[of x Q\<^sub>p "n+l"] by (simp add: A) have 31: "poly_map (n + l) Qs x = drop n x" using Qs_def map_by_var_list_segment_to_length[of x "n + l" n] A le_add1 by blast have 32: "poly_map (n + l) (Ps@Qs) x = take n x @ drop n x" using poly_map_append[of x "n+l" Ps Qs ] by (metis "30" "31" A append_take_drop_id) show "augment n x = poly_map (n + l) Fs x" using 30 32 poly_map_append by (metis A Fs_def poly_map_append augment_def) qed have 4: "(augment n -` S) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) = poly_tuple_pullback (n + l) S Fs" proof show "augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>) \ poly_tuple_pullback (n + l) S Fs" proof fix x assume A: "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" then have 40: "augment n x \ S" by blast have 41: "augment n x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" using 40 assms unfolding augment_def using is_semialgebraic_closed by blast have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" proof- have "take n x @ x \ carrier (Q\<^sub>p\<^bsup>n+n+ l\<^esup>)" using augment_def A by (metis "41" append_take_drop_id) then have 0: "drop n (take n x @ x) \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" by (metis (no_types, lifting) add.assoc cartesian_power_drop) have "drop n (take n x @ x) = x" proof- have "length x \ n" using A by (metis IntD2 cartesian_power_car_memE le_add1) then have "length (take n x) = n" by (metis add_right_cancel append_take_drop_id le_add_diff_inverse length_append length_drop) then show ?thesis by (metis append_eq_conv_conj) qed then show ?thesis using 0 by presburger qed then show "x \ poly_tuple_pullback (n + l) S Fs" using 41 3 unfolding poly_tuple_pullback_def by (metis (no_types, opaque_lifting) "40" add.commute cartesian_power_car_memE evimageI evimage_def poly_map_apply) qed show "poly_tuple_pullback (n + l) S Fs \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" proof fix x assume A: "x \ poly_tuple_pullback (n + l) S Fs" have "x \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using A unfolding poly_tuple_pullback_def by blast then show "x \ augment n -` S \ carrier (Q\<^sub>p\<^bsup>n+l\<^esup>)" using 3 by (metis (no_types, lifting) A poly_map_apply poly_tuple_pullback_def vimage_inter_cong) qed qed then show ?thesis using assms pullback_is_semialg[of "n + l" Fs] using poly_tuple_pullback_eq_poly_map_vimage unfolding restrict_def evimage_def Fs_def by (smt "4" Ex_list_of_length Fs_def Ps_def Qs_def \is_poly_tuple (n + l) Fs\ add.commute add_diff_cancel_left' append_assoc diff_zero is_semialgebraic_closed le_add2 length_append not_add_less1 not_gr_zero padic_fields.is_semialgebraicE padic_fields_axioms var_list_segment_length zero_le) qed lemma tuple_partial_pullback_is_semialg_map_tuple_induct: assumes "is_semialg_map_tuple m fs" assumes "is_semialg_function m f" assumes "length fs = n" shows "is_semialg_map_tuple m (fs@[f])" proof(rule is_semialg_map_tupleI) have 0: "is_function_tuple Q\<^sub>p m fs" using assms is_semialg_map_tuple_def by blast show "is_function_tuple Q\<^sub>p m (fs @ [f])" proof(rule is_function_tupleI) have A0: "set (fs @ [f]) = insert f (set fs)" by simp have A1: "set fs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using 0 is_function_tuple_def by blast then show "set (fs @ [f]) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms 0 by (metis (no_types, lifting) A0 is_semialg_function_closed list.simps(15) set_ConsD subset_code(1)) qed show "\k S. S \ semialg_sets (length (fs @ [f]) + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m (fs @ [f]) k S)" proof- fix l S assume A: "S \ semialg_sets (length (fs @ [f]) + l)" then have B: "S \ semialg_sets (n + l + 1)" using assms by (metis (no_types, lifting) add.commute add_Suc_right add_diff_cancel_left' append_Nil2 diff_Suc_1 length_Suc_conv length_append) show "is_semialgebraic (m + l) (tuple_partial_pullback m (fs @ [f]) l S)" proof- obtain S0 where S0_def: "S0 = tuple_partial_pullback m fs (l+1) S" by blast have 0: "is_semialgebraic (m + l + 1) S0" using B assms is_semialg_map_tupleE[of m fs "l + 1" S] by (metis S0_def add.assoc is_semialgebraicI) obtain S1 where S1_def: "S1 = twisted_partial_pullback m m l f S0" by blast then have "is_semialgebraic (m + m + l) S1" using S1_def assms(1) 0 twisted_partial_pullback_is_semialgebraic[of m f m l S0] by (simp add: assms(2)) then have L: "is_semialgebraic (m + m + l) (diagonalize m l S1)" using assms diagaonlize_is_semiaglebraic by blast have 1: "(tuple_partial_pullback m (fs @ [f]) l S) = (augment m -` (diagonalize m l S1)) \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof show "tuple_partial_pullback m (fs @ [f]) l S \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof fix x assume P0: "x \ tuple_partial_pullback m (fs @ [f]) l S " show "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" proof show "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" using tuple_partial_pullback_closed P0 by blast show "x \ augment m -` diagonalize m l S1" proof- obtain a where a_def: "a = take m x" by blast then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ le_add1 take_closed by blast obtain b where b_def: "b = drop m x" by blast then have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ cartesian_power_drop by blast have x_eq: "x = a@b" using a_def b_def by (metis append_take_drop_id) have X0: "a @ a @ b = augment m x" by (metis a_def augment_def b_def) have "a @ a @ b \ diagonalize m l S1" proof- have "length (a@a) = m + m" using a_closed cartesian_power_car_memE length_append by blast then have "take (m + m) (a @ a @ b) = a@a" by (metis append.assoc append_eq_conv_conj) then have X00: "take (m + m) (a @ a @ b) \ \ m" using diagonalI[of "a@a"] a_def a_closed by (metis append_eq_conv_conj cartesian_power_car_memE) then have X01: "a @ a @ b \ cartesian_product (\ m) (carrier (Q\<^sub>p\<^bsup>l\<^esup>))" using a_closed b_closed cartesian_product_memI[of "\ m" Q\<^sub>p "m+m" "carrier (Q\<^sub>p\<^bsup>l\<^esup>)" l "a @ a @ b"] unfolding diagonal_def by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed cartesian_power_drop mem_Collect_eq subsetI) have X02: "twisted_partial_image m m f (a @ a @ b) = a @ ((f a)# b)" using twisted_partial_image_eq[of a m a m b l _ f] a_closed b_closed by blast have "a @ a @ b \ S1" proof- have "twisted_partial_image m m f (a @ a @ b) \ S0" proof- have X020:"tuple_partial_image m fs (a @ ((f a)# b)) = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" using tuple_partial_image_eq[of a m "(f a)# b" "l + 1" _ fs] by (metis (no_types, lifting) a_closed append_Cons append_eq_conv_conj cartesian_power_car_memE self_append_conv2 tuple_partial_image_def) have X021: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b \ S" proof- have X0210: "(function_tuple_eval Q\<^sub>p m fs a)@[f a]@b = (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" unfolding function_tuple_eval_def by (metis (mono_tags, lifting) append.assoc list.simps(8) list.simps(9) map_append) have X0211: "(function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b = tuple_partial_image m (fs @ [f]) x" using x_eq tuple_partial_image_eq[of a m b l x "fs@[f]"] by (simp add: a_closed b_closed) have "tuple_partial_image m (fs @ [f]) x \ S" using P0 tuple_partial_pullback_memE(2) by blast then show ?thesis using X0211 X0210 by presburger qed have X022: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = (function_tuple_eval Q\<^sub>p m fs a)@[f a]@b" proof- have X0220: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = tuple_partial_image m fs (a @ ((f a)# b))" using X02 by presburger have X0221: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) = (function_tuple_eval Q\<^sub>p m fs a) @ ((f a)# b)" using tuple_partial_image_eq by (metis X02 X020 append_Cons self_append_conv2) then show ?thesis unfolding function_tuple_eval_def by (metis X02 X020 X0221 append_same_eq) qed have X023: "tuple_partial_image m fs (twisted_partial_image m m f (a @ a @ b)) \ S" using X02 X020 X021 by presburger have "twisted_partial_image m m f (a @ a @ b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" proof- have "a @ ((f a)# b) \ carrier (Q\<^sub>p\<^bsup>m + (l+1)\<^esup>)" apply(rule cartesian_power_car_memI) apply (metis a_closed add.commute b_closed cartesian_power_car_memE length_Cons length_append plus_1_eq_Suc) proof- have "f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using assms(2) is_semialg_function_closed by blast then have "f a \ carrier Q\<^sub>p" using a_closed assms by blast then show "set (a @ f a # b) \ carrier Q\<^sub>p" using assms a_closed b_closed by (meson cartesian_power_car_memE'' cartesian_power_concat(1) cartesian_power_cons) qed then show ?thesis using X02 by presburger qed then show ?thesis using S0_def X023 tuple_partial_pullback_def[of m fs "l+1" S ] by blast qed then show ?thesis using X02 S1_def twisted_partial_pullback_def by (metis (no_types, lifting) X0 \x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)\ augment_closed drop_drop local.partial_image_def twisted_partial_image_def twisted_partial_pullback_memI) qed then show ?thesis using X01 diagonalize_def[of m l S1] by blast qed then show ?thesis by (metis X0 vimageI) qed qed qed show "augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>) \ tuple_partial_pullback m (fs @ [f]) l S" proof fix x assume A: "x \ augment m -` diagonalize m l S1 \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" then have X0: "x \ carrier (Q\<^sub>p\<^bsup>m + l\<^esup>)" by blast obtain a where a_def: "a = take m x" by blast then have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using X0 le_add1 take_closed by blast obtain b where b_def: "b = drop m x" by blast then have a_closed: "b \ carrier (Q\<^sub>p\<^bsup>l\<^esup>)" using X0 cartesian_power_drop by blast have X1: "augment m x = a@a@b" using a_def augment_def b_def by blast have X2: "a@a@b \ diagonalize m l S1" using A X1 by (metis Int_iff vimage_eq) have X3: "a@a@b \ S1" using X2 diagonalize_def by blast have X4: "twisted_partial_image m m f (a@a@b) \ S0" using X3 S1_def twisted_partial_pullback_memE(2) by blast have X5: "a@((f a)#b) \ S0" using X4 twisted_partial_image_eq[of a m a m b l _ f] by (metis X0 a_closed a_def le_add1 take_closed) have X6: "tuple_partial_image m fs (a@((f a)#b)) \ S" using S0_def X5 tuple_partial_pullback_memE(2) by blast have X7: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) \ S" using X6 using tuple_partial_image_eq by (metis X0 a_def append_eq_conv_conj cartesian_power_car_memE le_add1 take_closed tuple_partial_image_def) have X8: "((function_tuple_eval Q\<^sub>p m fs a)@((f a)#b)) = tuple_partial_image m (fs @ [f]) x" proof- have X80: "tuple_partial_image m (fs @ [f]) x = (function_tuple_eval Q\<^sub>p m (fs@[f]) a)@b" using tuple_partial_image_def a_def b_def by blast then show ?thesis unfolding function_tuple_eval_def by (metis (no_types, lifting) append_Cons append_eq_append_conv2 list.simps(9) map_append self_append_conv2) qed show "x \ tuple_partial_pullback m (fs @ [f]) l S" using X8 X7 tuple_partial_pullback_def by (metis X0 \is_function_tuple Q\<^sub>p m (fs @ [f])\ tuple_partial_image_def tuple_partial_pullback_memI) qed qed then show ?thesis using augment_inverse_is_semialgebraic by (simp add: L) qed qed qed lemma singleton_tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" assumes "length fs = 1" shows "is_semialg_map_tuple m fs" proof(rule is_semialg_map_tupleI) show "is_function_tuple Q\<^sub>p m fs" by (simp add: assms(1) semialg_function_tuple_is_function_tuple) show "\k S. S \ semialg_sets (length fs + k) \ is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- fix k S assume A: "S \ semialg_sets (length fs + k)" show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" proof- obtain f where f_def: "fs = [f]" using assms by (metis One_nat_def length_0_conv length_Suc_conv) have 0: "is_semialg_function m f" using f_def assms is_semialg_function_tupleE'[of m fs f] by simp have 1: "\x. tuple_partial_image m fs x = partial_image m f x" unfolding function_tuple_eval_def tuple_partial_image_def partial_image_def by (metis (no_types, lifting) append_Cons append_Nil2 append_eq_append_conv_if f_def list.simps(8) list.simps(9)) have 2: "tuple_partial_pullback m fs k S = partial_pullback m f k S" proof show "tuple_partial_pullback m fs k S \ partial_pullback m f k S" using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def by (metis (no_types, lifting) set_eq_subset vimage_inter_cong) show "partial_pullback m f k S \ tuple_partial_pullback m fs k S" using 1 unfolding tuple_partial_pullback_def partial_pullback_def evimage_def by blast qed then show ?thesis by (metis "0" A assms(2) is_semialg_functionE is_semialgebraicI) qed qed qed lemma empty_tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" assumes "length fs = 0" shows "is_semialg_map_tuple m fs" apply(rule is_semialg_map_tupleI) using assms(1) semialg_function_tuple_is_function_tuple apply blast proof- fix k S assume A: "S \ semialg_sets (length fs + k)" then have 0: "is_semialgebraic k S" by (metis add.left_neutral assms(2) is_semialgebraicI) have 1: "tuple_partial_pullback m fs k S = cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" proof have 1: "\x. function_tuple_eval Q\<^sub>p m fs (take m x) = []" using assms unfolding function_tuple_eval_def by blast show "tuple_partial_pullback m fs k S \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S" apply(rule subsetI) apply(rule cartesian_product_memI[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k]) apply blast using 0 is_semialgebraic_closed apply blast using 0 assms unfolding 1 tuple_partial_pullback_def tuple_partial_image_def apply (meson IntD2 le_add1 take_closed) by (metis append_Nil evimageD evimage_def) have 2: "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" using is_semialgebraic_closed[of k S] 0 assms cartesian_product_closed[of "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p m S k] by blast show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) S \ tuple_partial_pullback m fs k S" apply(rule subsetI) apply(rule tuple_partial_pullback_memI) using 2 apply blast using assms semialg_function_tuple_is_function_tuple apply blast unfolding 1 by (metis carrier_is_semialgebraic cartesian_product_memE(2) is_semialgebraic_closed self_append_conv2) qed show "is_semialgebraic (m + k) (tuple_partial_pullback m fs k S)" unfolding 1 using "0" car_times_semialg_is_semialg by blast qed lemma tuple_partial_pullback_is_semialg_map_tuple: assumes "is_semialg_function_tuple m fs" shows "is_semialg_map_tuple m fs" proof- have "\n fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" proof- fix n show " \ fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs" apply(induction n) using singleton_tuple_partial_pullback_is_semialg_map_tuple empty_tuple_partial_pullback_is_semialg_map_tuple apply blast proof- fix n fs assume IH: "(\fs. is_semialg_function_tuple m fs \ length fs = n \ is_semialg_map_tuple m fs)" assume A: "is_semialg_function_tuple m fs \ length fs = Suc n" then obtain gs f where gs_f_def: "fs = gs@[f]" by (metis length_Suc_conv list.discI rev_exhaust) have gs_length: "length gs = n" using gs_f_def by (metis A length_append_singleton nat.inject) have 0: "set gs \ set fs" by (simp add: gs_f_def subsetI) have 1: "is_semialg_function_tuple m gs" apply(rule is_semialg_function_tupleI) using 0 A gs_f_def is_semialg_function_tupleE'[of m fs] by blast then have 2: "is_semialg_map_tuple m gs" using IH gs_length by blast have 3: "is_semialg_function m f" using gs_f_def A by (metis gs_length is_semialg_function_tupleE lessI nth_append_length) then show "is_semialg_map_tuple m fs" using assms 2 gs_f_def tuple_partial_pullback_is_semialg_map_tuple_induct by blast qed qed then show ?thesis using assms by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\Semialgebraic Functions are Closed under Composition with Semialgebraic Tuples\ (********************************************************************************************) (********************************************************************************************) lemma function_tuple_comp_partial_pullback: assumes "is_semialg_function_tuple m fs" assumes "length fs = n" assumes "is_semialg_function n f" assumes "S \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" shows "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" proof- have 0: "\x. partial_image m (function_tuple_comp Q\<^sub>p fs f) x = partial_image n f (tuple_partial_image m fs x)" unfolding partial_image_def function_tuple_comp_def tuple_partial_image_def using comp_apply[of f "function_tuple_eval Q\<^sub>p 0 fs"] unfolding function_tuple_eval_def proof - fix x :: "((nat \ int) \ (nat \ int)) set list" assume a1: "\x. (f \ (\x. map (\f. f x) fs)) x = f (map (\f. f x) fs)" have f2: "\f rs. drop n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = rs" by (simp add: assms(2)) have "\f rs. take n (map f fs @ (rs::((nat \ int) \ (nat \ int)) set list)) = map f fs" by (simp add: assms(2)) then show "(f \ (\rs. map (\f. f rs) fs)) (take m x) # drop m x = f (take n (map (\f. f (take m x)) fs @ drop m x)) # drop n (map (\f. f (take m x)) fs @ drop m x)" using f2 a1 by presburger qed show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" proof show "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S \ tuple_partial_pullback m fs k (partial_pullback n f k S)" proof fix x assume A: "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" then have 1: "partial_image m (function_tuple_comp Q\<^sub>p fs f) x \ S" using partial_pullback_memE(2) by blast have 2: " partial_image n f (tuple_partial_image m fs x) \ S" using 0 1 by presburger have 3: "x \ carrier (Q\<^sub>p\<^bsup>m + k\<^esup>)" using A assms by (metis partial_pullback_memE(1)) have 4: "tuple_partial_image m fs x \ partial_pullback n f k S" apply(rule partial_pullback_memI) apply (metis "0" "3" add_cancel_left_left assms(1) assms(2) cartesian_power_drop drop0 list.inject local.partial_image_def not_gr_zero semialg_function_tuple_is_function_tuple tuple_partial_image_closed) by (metis "2" local.partial_image_def) show " x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" apply(rule tuple_partial_pullback_memI) apply (simp add: "3") using assms(1) semialg_function_tuple_is_function_tuple apply blast by (metis "4" tuple_partial_image_def) qed show " tuple_partial_pullback m fs k (partial_pullback n f k S) \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S" proof fix x assume A: "x \ tuple_partial_pullback m fs k (partial_pullback n f k S)" show "x \ partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S " proof- have "partial_image n f (tuple_partial_image m fs x) \ S" using A partial_pullback_memE(2) tuple_partial_pullback_memE(2) by blast show ?thesis apply(rule partial_pullback_memI) apply (meson A subset_eq tuple_partial_pullback_closed) by (metis "0" \local.partial_image n f (tuple_partial_image m fs x) \ S\ local.partial_image_def) qed qed qed qed lemma semialg_function_tuple_comp: assumes "is_semialg_function_tuple m fs" assumes "length fs = n" assumes "is_semialg_function n f" shows "is_semialg_function m (function_tuple_comp Q\<^sub>p fs f)" proof(rule is_semialg_functionI) show "function_tuple_comp Q\<^sub>p fs f \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier Q\<^sub>p" using function_tuple_comp_closed[of f Q\<^sub>p n fs] assms(1) assms(2) assms(3) is_semialg_function_closed semialg_function_tuple_is_function_tuple by blast show "\k S. S \ semialg_sets (1 + k) \ is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" proof- fix k S assume A0: "S \ semialg_sets (1 + k)" show "is_semialgebraic (m + k) (partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S)" proof- have 0: "partial_pullback m (function_tuple_comp Q\<^sub>p fs f) k S = tuple_partial_pullback m fs k (partial_pullback n f k S)" using function_tuple_comp_partial_pullback[of m fs n f S k] assms \S \ semialg_sets (1 + k)\ is_semialgebraicI is_semialgebraic_closed by blast have 1: "is_semialgebraic (n + k) (partial_pullback n f k S)" using assms A0 is_semialg_functionE is_semialgebraicI by blast have 2: "is_semialgebraic (m + k) (tuple_partial_pullback m fs k (partial_pullback n f k S))" using 1 0 assms tuple_partial_pullback_is_semialg_map_tuple[of m fs] is_semialg_map_tupleE[of m fs k "partial_pullback n f k S"] by blast then show ?thesis using 0 by simp qed qed qed (********************************************************************************************) (********************************************************************************************) subsubsection\Algebraic Operations on Semialgebraic Functions\ (********************************************************************************************) (********************************************************************************************) text\Defining the set of extensional semialgebraic functions\ definition Qp_add_fun where "Qp_add_fun xs = xs!0 \\<^bsub>Q\<^sub>p\<^esub> xs!1" definition Qp_mult_fun where "Qp_mult_fun xs = xs!0 \ xs!1" text\Inversion function on first coordinates of Qp tuples. Arbitrarily redefined at 0 to map to 0\ definition Qp_invert where "Qp_invert xs = (if ((xs!0) = \) then \ else (inv (xs!0)))" text\Addition is semialgebraic\ lemma addition_is_semialg: "is_semialg_function 2 Qp_add_fun" proof- have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = (Qp_ev (pvar Q\<^sub>p 0) x) \\<^bsub>Q\<^sub>p\<^esub> (Qp_ev (pvar Q\<^sub>p 1) x)" by (metis A One_nat_def eval_at_point_add pvar_closed less_Suc_eq numeral_2_eq_2) then show " Qp_add_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" by (metis A Qp_add_fun_def add_vars_def add_vars_rep one_less_numeral_iff pos2 semiring_norm(76)) qed then have 1: "restrict Qp_add_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" by (meson restrict_ext) have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] by (meson MP.add.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) then show ?thesis using 1 semialg_function_on_carrier[of 2 "Qp_add_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] semialg_function_on_carrier by presburger qed text\Multiplication is semialgebraic:\ lemma multiplication_is_semialg: "is_semialg_function 2 Qp_mult_fun" proof- have 0: "\x. x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>) \ Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" proof- fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" have "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x = (Qp_ev (pvar Q\<^sub>p 0) x) \ (Qp_ev (pvar Q\<^sub>p 1) x)" by (metis A One_nat_def eval_at_point_mult pvar_closed less_Suc_eq numeral_2_eq_2) then show " Qp_mult_fun x = Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1) x" by (metis A Qp_mult_fun_def mult_vars_def mult_vars_rep one_less_numeral_iff pos2 semiring_norm(76)) qed then have 1: "restrict Qp_mult_fun (carrier (Q\<^sub>p\<^bsup>2\<^esup>)) = restrict (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)) (carrier (Q\<^sub>p\<^bsup>2\<^esup>))" by (meson restrict_ext) have "is_semialg_function 2 (Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1))" using poly_is_semialg[of "pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1"] by (meson MP.m_closed local.pvar_closed one_less_numeral_iff pos2 semiring_norm(76)) thus ?thesis using 1 semialg_function_on_carrier[of 2 "Qp_mult_fun" "Qp_ev (pvar Q\<^sub>p 0 \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> pvar Q\<^sub>p 1)"] semialg_function_on_carrier by presburger qed text\Inversion is semialgebraic:\ lemma(in field) field_nat_pow_inv: assumes "a \ carrier R" assumes "a \ \" shows "inv (a [^] (n::nat)) = (inv a) [^] (n :: nat)" apply(induction n) using inv_one local.nat_pow_0 apply presburger using assms nat_pow_of_inv by (metis Units_one_closed field_inv(2) field_inv(3) unit_factor) lemma Qp_invert_basic_semialg: assumes "is_basic_semialg (1 + k) S" shows "is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" proof- obtain P n where P_n_def: "(n::nat) \ 0 \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>]) \ S = basic_semialg_set (1+k) n P \ P \ carrier (Q\<^sub>p[\\<^bsub>1+k\<^esub>])" using assms is_basic_semialg_def by meson obtain d::nat where d_def: "d = deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P)" by auto obtain l where l_def: "l = ((- d) mod n)" by blast + have 10: "n > 0" + using P_n_def + by blast + have 11: "l \ 0" + using 10 by (simp add: l_def) have 1: "(l + d) mod n = 0" by (metis add_eq_0_iff equation_minus_iff l_def mod_0 mod_add_cong mod_minus_eq zmod_int) - then obtain m::int where m_def: " (l + d) = m*n " + then obtain m::int where m_def: "(l + int d) = m * int n" using d_def l_def by (metis mult_of_nat_commute zmod_eq_0D) - have 2: "m \0" - proof- - have 10: "n > 0" - using P_n_def - by blast - have 11: "l \ 0" - using l_def 10 Euclidean_Division.pos_mod_sign of_nat_0_less_iff - by blast - then show ?thesis - using m_def - by (metis "10" le_add_same_cancel1 minus_add_cancel negative_zle - neq0_conv of_nat_le_0_iff zero_le_imp_eq_int zero_le_mult_iff) - qed + with 10 have \m = (l + d) div n\ + by simp + with 10 11 have 2: "m \ 0" + by (simp add: div_int_pos_iff) obtain N where N_def: "N = m*n" by blast - have 3: "N \ d" - proof- - have "l \ 0" - using l_def d_def m_def Euclidean_Division.pos_mod_sign[of n "-d"] P_n_def - by linarith - then show ?thesis - using d_def N_def m_def - by linarith - qed + from 11 have 3: "N \ d" + by (simp add: N_def flip: m_def) have 4: "deg (coord_ring Q\<^sub>p k) (to_univ_poly (Suc k) 0 P) \ nat N" using d_def N_def 3 by linarith have 5: " P \ carrier (coord_ring Q\<^sub>p (Suc k))" by (metis P_n_def plus_1_eq_Suc) have 6: " \q\carrier (coord_ring Q\<^sub>p (Suc k)). \x\carrier Q\<^sub>p - {\}. \a\carrier (Q\<^sub>p\<^bsup>k\<^esup>). Qp_ev q (insert_at_index a x 0) = (x[^]nat N) \ Qp_ev P (insert_at_index a (inv x) 0)" using 3 4 d_def to_univ_poly_one_over_poly''[of 0 k P "nat N"] "5" Qp.field_axioms by blast obtain q where q_def: "q \ carrier (coord_ring Q\<^sub>p (Suc k)) \ ( \ x \ carrier Q\<^sub>p - {\}. ( \ a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). eval_at_point Q\<^sub>p (insert_at_index a x 0) q = (x[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index a (inv x) 0) P)))" using 6 by blast obtain T where T_def: "T = basic_semialg_set (1+k) n q" by auto have "is_basic_semialg (1 + k) T" proof- have "q \ carrier ( Q\<^sub>p[\\<^bsub>Suc k\<^esub>])" using q_def by presburger then show ?thesis using T_def is_basic_semialg_def by (metis P_n_def plus_1_eq_Suc) qed then have T_semialg: "is_semialgebraic (1+k) T" using T_def basic_semialg_is_semialg[of "1+k" T] is_semialgebraicI by blast obtain Nz where Nz_def: "Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 \ \}" by blast have Nz_semialg: "is_semialgebraic (1+k) Nz" proof- obtain Nzc where Nzc_def: "Nzc = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" by blast have 0: "Nzc = zero_set Q\<^sub>p (Suc k) (pvar Q\<^sub>p 0)" unfolding zero_set_def using Nzc_def by (metis (no_types, lifting) Collect_cong eval_pvar zero_less_Suc) have 1: "is_algebraic Q\<^sub>p (1+k) Nzc" using 0 pvar_closed[of ] by (metis is_algebraicI' plus_1_eq_Suc zero_less_Suc) then have 2: "is_semialgebraic (1+k) Nzc" using is_algebraic_imp_is_semialg by blast have 3: "Nz = carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>) - Nzc" using Nz_def Nzc_def by blast then show ?thesis using 2 by (simp add: complement_is_semialg) qed have 7: "(partial_pullback 1 Qp_invert k S) \ Nz = T \ Nz" proof show "partial_pullback 1 Qp_invert k S \ Nz \ T \ Nz" proof fix c assume A: "c \ partial_pullback 1 Qp_invert k S \ Nz" show "c \ T \ Nz" proof- have c_closed: "c \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using A partial_pullback_closed[of 1 Qp_invert k S] by blast obtain x a where xa_def: "c = (x#a)" using c_closed by (metis Suc_eq_plus1 add.commute cartesian_power_car_memE length_Suc_conv) have x_closed: "x \ carrier Q\<^sub>p" using xa_def c_closed by (metis (no_types, lifting) append_Cons cartesian_power_decomp list.inject Qp.to_R1_to_R Qp.to_R_pow_closed) have a_closed: "a \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using xa_def c_closed by (metis One_nat_def cartesian_power_drop drop0 drop_Suc_Cons) have 0: "c \ Nz" using A by blast then have "x \ \" using Nz_def xa_def by (metis (mono_tags, lifting) mem_Collect_eq nth_Cons_0) have 1: "Qp_invert [x] = inv x" unfolding Qp_invert_def by (metis \x \ \\ nth_Cons_0) have 2: "partial_image 1 Qp_invert c \ S" using A partial_pullback_memE[of c 1 "Qp_invert" k S] by blast have 3: "inv x # a \ S" proof- have 30: "[x] = take 1 c" by (simp add: xa_def) have 31: "a = drop 1 c" by (simp add: xa_def) show ?thesis using 1 30 31 partial_image_def[of 1 "Qp_invert" c] xa_def "2" by metis qed obtain y where y_def: "y \ carrier Q\<^sub>p \ eval_at_point Q\<^sub>p (inv x # a) P = y [^] n" using 3 P_n_def basic_semialg_set_memE(2) by blast then have 4: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P = x [^] (nat N) \ y [^] n" by presburger have 5: "x [^] (nat N) \ y [^] n = ((x[^]m)\y)[^]n" proof- have 50: "x [^] (N) \ y [^] n = x [^] (m*n) \ y [^] n" using N_def by blast have 51: "x [^] (m*n) = (x[^]m)[^]n" using Qp_int_nat_pow_pow \x \ \\ not_nonzero_Qp x_closed by metis have 52: "x [^] (m*n)\ y [^] n = ((x[^]m) \ y) [^] n" proof- have 0: "x [^] (m*n)\ y [^] n= (x[^]m)[^]n \ (y[^] n)" using "51" by presburger have 1: "(x[^]m)[^]n \ (y[^] n) = ((x[^]m) \ y) [^] n" apply(induction n) using Qp.nat_pow_0 Qp.one_closed Qp.r_one apply presburger using x_closed y_def by (metis Qp.nat_pow_distrib Qp.nonzero_closed Qp_int_pow_nonzero \x \ \\ not_nonzero_Qp) then show ?thesis using "0" by blast qed have 53: "x [^] N = x [^] (nat N)" proof- - have "N \ 0" - by (metis (full_types) Euclidean_Division.pos_mod_sign N_def P_n_def - add_increasing2 int.lless_eq l_def m_def of_nat_0_le_iff of_nat_le_0_iff ) + from 11 m_def N_def [symmetric] have "N \ 0" + by simp then show ?thesis by (metis pow_nat) qed then show ?thesis using 50 52 by presburger qed have 6: "x [^] (nat N) \ eval_at_point Q\<^sub>p (inv x # a) P = ((x[^]m)\y)[^]n" using "4" "5" by blast have 7: "eval_at_point Q\<^sub>p c q = ((x[^]m)\y)[^]n" proof- have 70: "(insert_at_index a (inv x) 0) = inv x # a" using insert_at_index.simps by (metis (no_types, lifting) append_eq_append_conv2 append_same_eq append_take_drop_id drop0 same_append_eq) have 71: "(insert_at_index a x) 0 = x # a" by simp then show ?thesis using 6 q_def by (metis "70" DiffI \x \ \\ a_closed empty_iff insert_iff x_closed xa_def) qed have 8: "(x[^]m)\y \ carrier Q\<^sub>p" proof- have 80: "x[^]m \ carrier Q\<^sub>p" using \x \ \\ x_closed Qp_int_pow_nonzero[of x m] unfolding nonzero_def by blast then show ?thesis using y_def by blast qed then have "c \ T" using T_def basic_semialg_set_def "7" c_closed by blast then show ?thesis by (simp add: \c \ T\ "0") qed qed show "T \ Nz \ partial_pullback 1 Qp_invert k S \ Nz" proof fix x assume A: "x \ T \ Nz" show " x \ partial_pullback 1 Qp_invert k S \ Nz " proof- have " x \ partial_pullback 1 Qp_invert k S" proof(rule partial_pullback_memI) show x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using T_def A by (meson IntD1 basic_semialg_set_memE(1)) show "Qp_invert (take 1 x) # drop 1 x \ S" proof- have 00: "x!0 \ \" using A Nz_def by blast then have 0: "Qp_invert (take 1 x) # drop 1 x = inv (x!0) # drop 1 x" unfolding Qp_invert_def by (smt One_nat_def lessI nth_take) have "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ cartesian_power_drop by blast obtain a where a_def: "a = (x!0)" by blast have a_closed: "a \ carrier Q\<^sub>p" using 00 a_def A Nz_def cartesian_power_car_memE'[of x Q\<^sub>p "Suc k" 0] inv_in_frac(1) by blast have a_nz: "a \ \" using a_def Nz_def A by blast obtain b where b_def: "b = drop 1 x" by blast have b_closed: "b \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using b_def A Nz_def \drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)\ by blast have abx: "x = a#b" using a_def b_def x_closed by (metis (no_types, lifting) One_nat_def append_Cons append_Nil append_eq_conv_conj cartesian_power_car_memE cartesian_power_decomp lessI nth_take Qp.to_R1_to_R) have 1: "Qp_invert (take 1 x) # drop 1 x = (inv a)#b" using "0" a_def b_def by blast have 22: "eval_at_point Q\<^sub>p (insert_at_index b a 0) q = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" using q_def a_closed a_nz b_closed by blast obtain c where c_def: "c \ carrier Q\<^sub>p \ Qp_ev q x = (c[^]n)" using A T_def unfolding basic_semialg_set_def by blast obtain c' where c'_def: "c' = (inv a)[^]m \ c" by blast have c'_closed: "c' \ carrier Q\<^sub>p" using c_def a_def a_closed a_nz Qp_int_pow_nonzero nonzero_def c'_def inv_in_frac(3) Qp.m_closed Qp.nonzero_closed by presburger have 3: "(eval_at_point Q\<^sub>p ((inv a) # b) P) = (c'[^]n)" proof- have 30: "x = insert_at_index b a 0" using abx by simp have 31: "(c[^]n) = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p (insert_at_index b (inv a) 0) P)" using 22 30 c_def by blast have 32: "insert_at_index b (inv a) 0 = (inv a) # b" using insert_at_index.simps by (metis drop0 self_append_conv2 take0) have 33: "(c[^]n) = (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using "31" "32" by presburger have 34: "(inv a) # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" apply(rule cartesian_power_car_memI'') apply (metis b_closed cartesian_power_car_memE length_Suc_conv plus_1_eq_Suc) using a_closed a_nz b_closed apply (metis One_nat_def inv_in_frac(1) take0 take_Suc_Cons Qp.to_R1_closed) by (metis abx b_closed b_def drop_Cons' not_Cons_self2) have 35: "(eval_at_point Q\<^sub>p ((inv a) # b) P) \ carrier Q\<^sub>p" using 34 P_n_def eval_at_point_closed by blast have "inv(a[^] (nat N)) \ (c[^]n) = inv(a[^] (nat N)) \ ((a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P))" using 31 "33" by presburger then have 6: "inv(a[^] (nat N)) \ (c[^]n) = inv(a[^] (nat N)) \ (a[^] (nat N)) \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 35 monoid.m_assoc[of Q\<^sub>p] Qp.monoid_axioms Qp.nat_pow_closed Qp.nonzero_pow_nonzero a_nz inv_in_frac(1) local.a_closed by presburger have 37:"inv(a[^] (nat N)) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" proof- have "inv(a[^] (nat N)) \ (a[^] (nat N)) = \ " using a_closed a_nz Qp.nat_pow_closed Qp.nonzero_pow_nonzero field_inv(1) by blast then have "inv(a[^] (nat N)) \ (c[^]n) = \ \ (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 6 by presburger then show ?thesis using 35 Qp.l_one by blast qed have 38:"(inv a)[^] (nat N) \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 37 group.nat_pow_inv[of Q\<^sub>p a "nat N"] a_closed Qp.field_axioms field.field_nat_pow_inv[of Q\<^sub>p] by (metis a_nz) have 39:"((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" using 2 38 monoid.nat_pow_pow[of Q\<^sub>p "inv a" ] N_def by (smt "3" Qp_int_nat_pow_pow a_closed a_nz inv_in_frac(3) of_nat_0_le_iff pow_nat) have 310:"((((inv a)[^]m) \ c)[^]n) = (eval_at_point Q\<^sub>p ((inv a) # b) P)" proof- have AA: "(inv a)[^]m \ carrier Q\<^sub>p" using Qp_int_pow_nonzero nonzero_def a_closed a_nz inv_in_frac(3) Qp.nonzero_closed by presburger have "((inv a)[^]m) [^] \<^bsub>Q\<^sub>p\<^esub> n \ (c[^]n) = ((((inv a)[^]m) \ c)[^]n)" using Qp.nat_pow_distrib[of "(inv a)[^]m" c n] a_closed a_def c_def AA by blast then show ?thesis using "39" by blast qed then show ?thesis using c'_def by blast qed have 4: "inv a # b \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" by (metis a_closed a_nz add.commute b_closed cartesian_power_cons inv_in_frac(1)) then have 5: "((inv a) # b) \ S" using 3 P_n_def c'_closed basic_semialg_set_memI[of "(inv a) # b" "1 + k" c' P n] by blast have 6: "Qp_invert (take 1 x) # drop 1 x = inv a # b" using a_def b_def unfolding Qp_invert_def using "1" Qp_invert_def by blast show ?thesis using 5 6 by presburger qed qed then show ?thesis using A by blast qed qed qed have 8: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) \ Nz)" using "7" Nz_semialg T_semialg intersection_is_semialg by auto have 9: "(partial_pullback 1 Qp_invert k S) - Nz = {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S" proof show "partial_pullback 1 Qp_invert k S - Nz \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" proof fix x assume A: " x \ partial_pullback 1 Qp_invert k S - Nz" have 0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" using A by (metis DiffD1 partial_pullback_memE(1) plus_1_eq_Suc) have 1: "take 1 x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" by (metis "0" le_add1 plus_1_eq_Suc take_closed) have 2: "drop 1 x \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using "0" cartesian_power_drop plus_1_eq_Suc by presburger have 3: " x = take 1 x @ drop 1 x " using 0 by (metis append_take_drop_id) have 4: "Qp_invert (take 1 x) # drop 1 x \ S" using A partial_pullback_memE'[of "take 1 x" 1 "drop 1 x" k x Qp_invert S] 1 2 3 by blast have 5: "x!0 = \" using A 0 Nz_def by blast have 6: "Qp_invert (take 1 x) # drop 1 x = x" proof- have "(take 1 x) =[x!0]" using 0 by (metis "1" "3" append_Cons nth_Cons_0 Qp.to_R1_to_R) then have "Qp_invert (take 1 x) = \" unfolding Qp_invert_def using 5 by (metis less_one nth_take) then show ?thesis using 0 5 by (metis "3" Cons_eq_append_conv \take 1 x = [x ! 0]\ self_append_conv2) qed have "x \ S" using 6 4 by presburger then show "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" using Nz_def A 0 by blast qed show "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S \ partial_pullback 1 Qp_invert k S - Nz" proof fix x assume A: "x \ {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs ! 0 = \} \ S" have A0: "x \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>)" using A by blast have A1: "x!0 = \" using A by blast have A2: "x \ S" using A by blast show " x \ partial_pullback 1 Qp_invert k S - Nz" proof show "x \ Nz" using Nz_def A1 by blast show " x \ partial_pullback 1 Qp_invert k S" proof(rule partial_pullback_memI) show "x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)" using A0 by (simp add: A0) show "Qp_invert (take 1 x) # drop 1 x \ S" proof- have "Qp_invert (take 1 x) = \" unfolding Qp_invert_def using A0 A1 by (metis less_numeral_extra(1) nth_take) then have "Qp_invert (take 1 x) # drop 1 x = x" using A0 A1 A2 by (metis (no_types, lifting) Cons_eq_append_conv Qp_invert_def \x \ carrier (Q\<^sub>p\<^bsup>1+k\<^esup>)\ append_take_drop_id inv_in_frac(2) le_add_same_cancel1 self_append_conv2 take_closed Qp.to_R1_to_R Qp.to_R_pow_closed zero_le) then show ?thesis using A2 by presburger qed qed qed qed qed have 10: "is_semialgebraic (1+k) {xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \}" proof- have "{xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} = V\<^bsub>Q\<^sub>p\<^esub> (Suc k) (pvar Q\<^sub>p 0)" unfolding zero_set_def using eval_pvar[of 0 "Suc k"] Qp.cring_axioms by blast then show ?thesis using is_zero_set_imp_basic_semialg pvar_closed[of 0 "Suc k"] Qp.cring_axioms is_zero_set_imp_semialg plus_1_eq_Suc zero_less_Suc by presburger qed have 11: "is_semialgebraic (1+k) ({xs \ carrier (Q\<^sub>p\<^bsup>Suc k\<^esup>). xs!0 = \} \S)" using 10 assms basic_semialg_is_semialgebraic intersection_is_semialg by blast have 12: "(partial_pullback 1 Qp_invert k S) = ((partial_pullback 1 Qp_invert k S) \ Nz) \ ((partial_pullback 1 Qp_invert k S) - Nz)" by blast have 13: "is_semialgebraic (1+k) ((partial_pullback 1 Qp_invert k S) - Nz)" using 11 9 by metis show ?thesis using 8 12 13 by (metis "7" Int_Diff_Un Int_commute plus_1_eq_Suc union_is_semialgebraic) qed lemma Qp_invert_is_semialg: "is_semialg_function 1 Qp_invert" proof(rule is_semialg_functionI') show 0: "Qp_invert \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ carrier Q\<^sub>p" proof fix x assume A: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" then obtain a where a_def: "x = [a]" by (metis Qp.to_R1_to_R) have a_closed: "a \ carrier Q\<^sub>p" using a_def A cartesian_power_concat(1) last_closed' by blast show " Qp_invert x \ carrier Q\<^sub>p" apply(cases "a = \") unfolding Qp_invert_def using a_def a_closed apply (meson Qp.to_R_to_R1) by (metis a_closed a_def inv_in_frac(1) Qp.to_R_to_R1) qed show "\k S. S \ basic_semialgs (1 + k) \ is_semialgebraic (1 + k) (partial_pullback 1 Qp_invert k S)" using Qp_invert_basic_semialg by blast qed lemma Taylor_deg_1_expansion'': assumes "f \ carrier Q\<^sub>p_x" assumes "\n. f n \ \\<^sub>p" assumes "a \ \\<^sub>p " assumes "b \ \\<^sub>p" shows "\c c' c''. c = to_fun f a \ c' = deriv f a \ c \ \\<^sub>p \ c' \ \\<^sub>p \c'' \ \\<^sub>p \ to_fun f (b) = c \ c' \ (b \ a) \ (c'' \ (b \ a)[^](2::nat))" proof- obtain S where S_def: "S = (Q\<^sub>p \ carrier := \\<^sub>p \)" by blast have 1: "f \ carrier (UP S)" unfolding S_def using val_ring_subring UPQ.poly_cfs_subring[of \\<^sub>p f] assms by blast have 2: " f \ carrier (UP (Q\<^sub>p\carrier := \\<^sub>p\))" using val_ring_subring 1 assms poly_cfs_subring[of \\<^sub>p] by blast have 3: "\c\\\<^sub>p. f \ b = f \ a \ UPQ.deriv f a \ (b \ a) \ c \ (b \ a) [^] (2::nat)" using UP_subring_taylor_appr'[of \\<^sub>p f b a] UP_subring_taylor_appr[of \\<^sub>p f b a] val_ring_subring 1 2 assms by blast then show ?thesis using UP_subring_taylor_appr[of \\<^sub>p f b a] assms UP_subring_deriv_closed[of \\<^sub>p f a] UP_subring_eval_closed[of \\<^sub>p f a] 2 val_ring_subring by blast qed end (**************************************************************************************************) (**************************************************************************************************) subsection\Sets Defined by Residues of Valuation Ring Elements\ (**************************************************************************************************) (**************************************************************************************************) sublocale padic_fields < Res: cring "Zp_res_ring (Suc n)" using p_residues residues.cring by blast context padic_fields begin definition Qp_res where "Qp_res x n = to_Zp x n " lemma Qp_res_closed: assumes "x \ \\<^sub>p" shows "Qp_res x n \ carrier (Zp_res_ring n)" unfolding Qp_res_def using assms val_ring_memE residue_closed to_Zp_closed by blast lemma Qp_res_add: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_sum to_Zp_add by presburger lemma Qp_res_mult: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_prod to_Zp_mult by presburger lemma Qp_res_diff: assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" shows "Qp_res (x \ y) n = Qp_res x n \\<^bsub>Zp_res_ring n\<^esub> Qp_res y n" unfolding Qp_res_def using assms residue_of_diff to_Zp_minus by (meson val_ring_res) lemma Qp_res_zero: shows "Qp_res \ n = 0" unfolding Qp_res_def to_Zp_zero using residue_of_zero(2) by blast lemma Qp_res_one: assumes "n > 0" shows "Qp_res \ n = (1::int)" using assms unfolding Qp_res_def to_Zp_one using residue_of_one(2) by blast lemma Qp_res_nat_inc: shows "Qp_res ([(n::nat)]\\) n = n mod p^n" unfolding Qp_res_def unfolding to_Zp_nat_inc using Zp_nat_inc_res by blast lemma Qp_res_int_inc: shows "Qp_res ([(k::int)]\\) n = k mod p^n" unfolding Qp_res_def unfolding to_Zp_int_inc using Zp_int_inc_res by blast lemma Qp_poly_res_monom: assumes "a \ \\<^sub>p" assumes "x \ \\<^sub>p" assumes "Qp_res a n = 0" assumes "k > 0" shows "Qp_res (up_ring.monom (UP Q\<^sub>p) a k \ x) n = 0" proof- have 0: "up_ring.monom (UP Q\<^sub>p) a k \ x = a \ x [^] k" apply(rule UPQ.to_fun_monom[of a x k]) using assms val_ring_memE apply blast using assms val_ring_memE by blast have 1: "x[^]k \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast show ?thesis unfolding 0 using Qp_res_mult[of a "x[^]k" n] assms using "1" residue_times_zero_r by presburger qed lemma Qp_poly_res_zero: assumes "q \ carrier (UP Q\<^sub>p)" assumes "\i. q i \ \\<^sub>p" assumes "\i. Qp_res (q i) n = 0" assumes "x \ \\<^sub>p" shows "Qp_res (q \ x) n = 0" proof- have "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0" proof(rule UPQ.poly_induct[of q], rule assms, rule ) fix p assume A: "p \ carrier (UP Q\<^sub>p)" " deg Q\<^sub>p p = 0" " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" have 0: "p \ x = p 0" using assms by (metis A(1) A(2) val_ring_memE UPQ.ltrm_deg_0 UPQ.to_fun_ctrm) show "Qp_res (p \ x) n = 0" unfolding 0 using A by blast next fix p assume A0: "(\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p p \ (\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0) \ Qp_res (q \ x) n = 0)" "p \ carrier (UP Q\<^sub>p)" "0 < deg Q\<^sub>p p" show "(\i. p i \ \\<^sub>p \ Qp_res (p i) n = 0) \ Qp_res (p \ x) n = 0" proof assume A1: " \i. p i \ \\<^sub>p \ Qp_res (p i) n = 0" obtain k where k_def: "k = deg Q\<^sub>p p" by blast obtain q where q_def: "q = UPQ.trunc p" by blast have q_closed: "q \ carrier (UP Q\<^sub>p)" unfolding q_def using A0(2) UPQ.trunc_closed by blast have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p p" unfolding q_def using A0(2) A0(3) UPQ.trunc_degree by blast have 9: "\i. i < deg Q\<^sub>p p \ q i = p i" unfolding q_def using A0(2) UPQ.trunc_cfs by blast have 90: "\i. \ i < deg Q\<^sub>p p \ q i = \" unfolding q_def proof - fix i :: nat assume "\ i < deg Q\<^sub>p p" then have "deg Q\<^sub>p q < i" using q_deg by linarith then show "Cring_Poly.truncate Q\<^sub>p p i = \" using UPQ.deg_gtE q_closed q_def by blast qed have 10: "(\i. q i \ \\<^sub>p \ Qp_res (q i) n = 0)" proof fix i show "q i \ \\<^sub>p \ Qp_res (q i) n = 0" apply(cases "i < deg Q\<^sub>p p") using A1 9[of i] apply presburger unfolding q_def using Qp_res_zero 90 by (metis q_def zero_in_val_ring) qed have 11: "Qp_res (q \ x) n = 0" using 10 A1 A0 q_closed q_deg by blast have 12: "p = q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k" unfolding k_def q_def using A0(2) UPQ.trunc_simps(1) by blast have 13: "p \ x = q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x" proof- have 0: " (q \\<^bsub>UP Q\<^sub>p\<^esub> up_ring.monom (UP Q\<^sub>p) (p k) k) \ x = q \ x \ up_ring.monom (UP Q\<^sub>p) (p k) k \ x" apply(rule UPQ.to_fun_plus) using A0(2) UPQ.ltrm_closed k_def apply blast unfolding q_def apply(rule UPQ.trunc_closed, rule A0) using assms val_ring_memE by blast show ?thesis using 0 12 by metis qed have 14: "(up_ring.monom (UP Q\<^sub>p) (p k) k) \ x \ \\<^sub>p" apply(rule val_ring_poly_eval) using A0(2) UPQ.ltrm_closed k_def apply blast using UPQ.cfs_monom[of "p k" k ] A1 zero_in_val_ring using A0(2) UPQ.ltrm_cfs k_def apply presburger using assms(4) by blast have 15: "Qp_res ((up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" apply(rule Qp_poly_res_monom) using A1 apply blast using assms apply blast using A1 apply blast unfolding k_def using A0 by blast have 16: "Qp_res (q \ x) n = 0" using A0 10 11 by blast have 17: "q \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule q_closed) using 10 apply blast by(rule assms) have 18: "Qp_res (q \ x \ (up_ring.monom (UP Q\<^sub>p) (p k) k) \ x) n = 0" using Qp_res_add[of "q \ x" "up_ring.monom (UP Q\<^sub>p) (p k) k \ x" n] 14 17 unfolding 15 16 by (metis "10" Qp_res_add UPQ.cfs_add UPQ.coeff_of_sum_diff_degree0 q_closed q_deg) show "Qp_res (p \ x) n = 0" using 13 18 by metis qed qed thus ?thesis using assms by blast qed lemma Qp_poly_res_eval_0: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "\i. g i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" shows "Qp_res (f \ x) n = Qp_res (g \ x) n" proof- obtain F where F_def: "F = f \\<^bsub>UP Q\<^sub>p\<^esub>g" by blast have F_closed: "F \ carrier (UP Q\<^sub>p)" unfolding F_def using assms by blast have F_cfs: "\i. F i = (f i) \ (g i)" unfolding F_def using assms UPQ.cfs_minus by blast have F_cfs_res: "\i. Qp_res (F i) n = Qp_res (f i) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g i) n" unfolding F_cfs apply(rule Qp_res_diff) using assms apply blast using assms by blast have 0: "\i. Qp_res (f i) n = Qp_res (g i) n" using assms by blast have F_cfs_res': "\i. Qp_res (F i) n = 0" unfolding F_cfs_res 0 by (metis diff_self mod_0 residue_minus) have 1: "\i. F i \ \\<^sub>p" unfolding F_cfs using assms using val_ring_minus_closed by blast have 2: "Qp_res (F \ x) n = 0" by(rule Qp_poly_res_zero, rule F_closed, rule 1, rule F_cfs_res', rule assms) have 3: "F \ x = f \ x \ g \ x" unfolding F_def using assms by (meson assms UPQ.to_fun_diff val_ring_memE) have 4: "Qp_res (F \ x) n = Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n" unfolding 3 apply(rule Qp_res_diff, rule val_ring_poly_eval, rule assms) using assms apply blast using assms apply blast apply(rule val_ring_poly_eval, rule assms) using assms apply blast by(rule assms) have 5: "f \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule assms) using assms apply blast using assms by blast have 6: "g \ x \ \\<^sub>p" apply(rule val_ring_poly_eval, rule assms) using assms apply blast by(rule assms) show "Qp_res (f \ x) n = Qp_res (g \ x) n" using 5 6 2 Qp_res_closed[of "f \ x" n] Qp_res_closed[of "g \ x" n] unfolding 4 proof - assume "Qp_res (f \ x) n \\<^bsub>Zp_res_ring n\<^esub> Qp_res (g \ x) n = 0" then show ?thesis by (metis (no_types) Qp_res_def 5 6 res_diff_zero_fact(1) residue_of_diff to_Zp_closed val_ring_memE) qed qed lemma Qp_poly_res_eval_1: assumes "f \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "Qp_res x n = Qp_res y n" shows "Qp_res (f \ x) n = Qp_res (f \ y) n" proof- have "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" apply(rule UPQ.poly_induct[of f], rule assms) proof fix f assume A: "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f = 0" "\i. f i \ \\<^sub>p" show "Qp_res (f \ x) n = Qp_res (f \ y) n" proof- obtain a where a_def: "a \ carrier Q\<^sub>p \ f = to_polynomial Q\<^sub>p a" using assms by (metis A(1) A(2) UPQ.lcf_closed UPQ.to_poly_inverse) have a_eq: "f = to_polynomial Q\<^sub>p a" using a_def by blast have 0: "f \ x = a" using a_def assms unfolding a_eq by (meson UPQ.to_fun_to_poly val_ring_memE) have 1: "f \ y = a" using a_def assms unfolding a_eq by (meson UPQ.to_fun_to_poly val_ring_memE) show " Qp_res (f \ x) n = Qp_res (f \ y) n" unfolding 0 1 by blast qed next fix f assume A: " (\q. q \ carrier (UP Q\<^sub>p) \ deg Q\<^sub>p q < deg Q\<^sub>p f \ (\i. q i \ \\<^sub>p) \ Qp_res (q \ x) n = Qp_res (q \ y) n)" "f \ carrier (UP Q\<^sub>p)" " 0 < deg Q\<^sub>p f" show "(\i. f i \ \\<^sub>p) \ Qp_res (f \ x) n = Qp_res (f \ y) n" proof assume A1: "\i. f i \ \\<^sub>p" obtain q where q_def: "q = UPQ.trunc f" by blast have q_closed: "q \ carrier (UP Q\<^sub>p)" using q_def A UPQ.trunc_closed by presburger have q_deg: "deg Q\<^sub>p q < deg Q\<^sub>p f" using q_def A UPQ.trunc_degree by blast have q_cfs: "\i. q i \ \\<^sub>p" proof fix i show "q i \ \\<^sub>p" apply(cases "i < deg Q\<^sub>p f") unfolding q_def using A A1 UPQ.trunc_cfs apply presburger using q_deg q_closed proof - assume "\ i < deg Q\<^sub>p f" then have "deg Q\<^sub>p f \ i" by (meson diff_is_0_eq neq0_conv zero_less_diff) then show "Cring_Poly.truncate Q\<^sub>p f i \ \\<^sub>p" by (metis (no_types) UPQ.deg_eqI diff_is_0_eq' le_trans nat_le_linear neq0_conv q_closed q_def q_deg zero_in_val_ring zero_less_diff) qed qed hence 0: "Qp_res (q \ x) n = Qp_res (q \ y) n" using A q_closed q_deg by blast have 1: "Qp_res (UPQ.ltrm f \ x) n = Qp_res (UPQ.ltrm f \ y) n" proof- have 10: "UPQ.ltrm f \ x = (f (deg Q\<^sub>p f)) \ x[^](deg Q\<^sub>p f)" using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger have 11: "UPQ.ltrm f \ y = (f (deg Q\<^sub>p f)) \ y[^](deg Q\<^sub>p f)" using A assms A1 UPQ.to_fun_monom val_ring_memE by presburger obtain d where d_def: "d = deg Q\<^sub>p f" by blast have 12: "Qp_res (x[^]d) n = Qp_res (y[^]d) n" apply(induction d) using Qp.nat_pow_0 apply presburger using Qp_res_mult assms Qp.nat_pow_Suc val_ring_nat_pow_closed by presburger hence 13: "Qp_res (x [^] deg Q\<^sub>p f) n = Qp_res (y [^] deg Q\<^sub>p f) n" unfolding d_def by blast have 14: "x [^] deg Q\<^sub>p f \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast have 15: "y [^] deg Q\<^sub>p f \ \\<^sub>p" using assms val_ring_nat_pow_closed by blast have 16: "Qp_res (f (deg Q\<^sub>p f) \ x [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (x [^] deg Q\<^sub>p f) n" apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " x[^](deg Q\<^sub>p f)" n]) using A1 apply blast by(rule 14) have 17: "Qp_res (f (deg Q\<^sub>p f) \ y [^] deg Q\<^sub>p f) n = Qp_res (f (deg Q\<^sub>p f)) n \\<^bsub>residue_ring (p ^ n)\<^esub> Qp_res (y [^] deg Q\<^sub>p f) n" apply(rule Qp_res_mult[of "f (deg Q\<^sub>p f)" " y[^](deg Q\<^sub>p f)" n]) using A1 apply blast by(rule 15) show ?thesis unfolding 10 11 16 17 13 by blast qed have f_decomp: "f = q \\<^bsub>UP Q\<^sub>p\<^esub> UPQ.ltrm f" using A unfolding q_def using UPQ.trunc_simps(1) by blast have 2: "f \ x = q \ x \ (UPQ.ltrm f \ x)" using A f_decomp q_closed q_cfs by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(2)) have 3: "f \ y = q \ y \ (UPQ.ltrm f \ y)" using A f_decomp q_closed q_cfs by (metis val_ring_memE UPQ.ltrm_closed UPQ.to_fun_plus assms(3)) show 4: " Qp_res (f \ x) n = Qp_res (f \ y) n " unfolding 2 3 using assms q_cfs Qp_res_add 0 1 by (metis (no_types, opaque_lifting) "2" "3" A(2) A1 Qp_res_def poly_eval_cong) qed qed thus ?thesis using assms by blast qed lemma Qp_poly_res_eval_2: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "x \ \\<^sub>p" assumes "y \ \\<^sub>p" assumes "\i. f i \ \\<^sub>p" assumes "\i. g i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" assumes "Qp_res x n = Qp_res y n" shows "Qp_res (f \ x) n = Qp_res (g \ y) n" proof- have 0: "Qp_res (f \ x) n = Qp_res (g \ x) n" using Qp_poly_res_eval_0 assms by blast have 1: "Qp_res (g \ x) n = Qp_res (g \ y) n" using Qp_poly_res_eval_1 assms by blast show ?thesis unfolding 0 1 by blast qed definition poly_res_class where "poly_res_class n d f = {q \ carrier (UP Q\<^sub>p). deg Q\<^sub>p q \ d \ (\i. q i \ \\<^sub>p \ Qp_res (f i) n = Qp_res (q i) n) }" lemma poly_res_class_closed: assumes "f \ carrier (UP Q\<^sub>p)" assumes "g \ carrier (UP Q\<^sub>p)" assumes "deg Q\<^sub>p f \ d" assumes "deg Q\<^sub>p g \ d" assumes "g \ poly_res_class n d f" shows "poly_res_class n d f = poly_res_class n d g" unfolding poly_res_class_def apply(rule equalityI) apply(rule subsetI) unfolding mem_Collect_eq apply(rule conjI, blast, rule conjI, blast) using assms unfolding poly_res_class_def mem_Collect_eq apply presburger apply(rule subsetI) unfolding mem_Collect_eq apply(rule conjI, blast, rule conjI, blast) using assms unfolding poly_res_class_def mem_Collect_eq by presburger lemma poly_res_class_memE: assumes "f \ poly_res_class n d g" shows "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f \ d" "f i \ \\<^sub>p" "Qp_res (g i) n = Qp_res (f i) n" using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq apply blast using assms unfolding poly_res_class_def mem_Collect_eq by blast definition val_ring_polys where "val_ring_polys = {f \ carrier (UP Q\<^sub>p). (\i. f i \ \\<^sub>p)} " lemma val_ring_polys_closed: "val_ring_polys \ carrier (UP Q\<^sub>p)" unfolding val_ring_polys_def by blast lemma val_ring_polys_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "\i. f i \ \\<^sub>p" shows "f \ val_ring_polys" using assms unfolding val_ring_polys_def by blast lemma val_ring_polys_memE: assumes "f \ val_ring_polys" shows "f \ carrier (UP Q\<^sub>p)" "f i \ \\<^sub>p" using assms unfolding val_ring_polys_def apply blast using assms unfolding val_ring_polys_def by blast definition val_ring_polys_grad where "val_ring_polys_grad d = {f \ val_ring_polys. deg Q\<^sub>p f \ d}" lemma val_ring_polys_grad_closed: "val_ring_polys_grad d \ val_ring_polys" unfolding val_ring_polys_grad_def by blast lemma val_ring_polys_grad_closed': "val_ring_polys_grad d \ carrier (UP Q\<^sub>p)" unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma val_ring_polys_grad_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "\i. f i \ \\<^sub>p" assumes "deg Q\<^sub>p f \ d" shows "f \ val_ring_polys_grad d" using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma val_ring_polys_grad_memE: assumes "f \ val_ring_polys_grad d" shows "f \ carrier (UP Q\<^sub>p)" "deg Q\<^sub>p f \ d" "f i \ \\<^sub>p" using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast using assms unfolding val_ring_polys_grad_def val_ring_polys_def apply blast using assms unfolding val_ring_polys_grad_def val_ring_polys_def by blast lemma poly_res_classes_in_val_ring_polys_grad: assumes "f \ val_ring_polys_grad d" shows "poly_res_class n d f \ val_ring_polys_grad d" apply(rule subsetI, rule val_ring_polys_grad_memI) apply(rule poly_res_class_memE[of _ n d f], blast) apply(rule poly_res_class_memE[of _ n d f], blast) by(rule poly_res_class_memE[of _ n d f], blast) lemma poly_res_class_disjoint: assumes "f \ val_ring_polys_grad d" assumes "f \ poly_res_class n d g" shows "poly_res_class n d f \ poly_res_class n d g = {}" apply(rule equalityI) apply(rule subsetI) using assms unfolding poly_res_class_def mem_Collect_eq Int_iff apply (metis val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3)) by blast lemma poly_res_class_refl: assumes "f \ val_ring_polys_grad d" shows "f \ poly_res_class n d f" unfolding poly_res_class_def mem_Collect_eq using assms val_ring_polys_grad_memE(1) val_ring_polys_grad_memE(2) val_ring_polys_grad_memE(3) by blast lemma poly_res_class_memI: assumes "f \ carrier (UP Q\<^sub>p)" assumes "deg Q\<^sub>p f \ d" assumes "\i. f i \ \\<^sub>p" assumes "\i. Qp_res (f i) n = Qp_res (g i) n" shows "f \ poly_res_class n d g" unfolding poly_res_class_def mem_Collect_eq using assms by metis definition poly_res_classes where "poly_res_classes n d = poly_res_class n d ` val_ring_polys_grad d" lemma poly_res_classes_disjoint: assumes "A \ poly_res_classes n d" assumes "B \ poly_res_classes n d" assumes "g \ A - B" shows "A \ B = {}" proof- obtain a where a_def: "a \ val_ring_polys_grad d \ A = poly_res_class n d a" using assms unfolding poly_res_classes_def by blast obtain b where b_def: "b \ val_ring_polys_grad d \ B = poly_res_class n d b" using assms unfolding poly_res_classes_def by blast have 0: "\f. f \ A \ B \ False" proof- fix f assume A: "f \ A \ B" have 1: "\i. Qp_res (g i) n \ Qp_res (f i) n" proof(rule ccontr) assume B: "\i. Qp_res (g i) n \ Qp_res (f i) n" then have 2: "\i. Qp_res (g i) n = Qp_res (f i) n" by blast have 3: "g \ poly_res_class n d a" using a_def assms by blast have 4: "\i. Qp_res (b i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using assms A b_def by blast have 5: "\i. Qp_res (a i) n = Qp_res (g i) n" apply(rule poly_res_class_memE[of _ n d]) using assms A a_def by blast have 6: "g \ poly_res_class n d b" apply(rule poly_res_class_memI, rule poly_res_class_memE[of _ n d a], rule 3, rule poly_res_class_memE[of _ n d a], rule 3, rule poly_res_class_memE[of _ n d a], rule 3) unfolding 2 4 by blast show False using 6 b_def assms by blast qed then obtain i where i_def: "Qp_res (g i) n \ Qp_res (f i) n" by blast have 2: "\i. Qp_res (a i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using A a_def by blast have 3: "\i. Qp_res (b i) n = Qp_res (f i) n" apply(rule poly_res_class_memE[of _ n d]) using A b_def by blast have 4: "\i. Qp_res (a i) n = Qp_res (g i) n" apply(rule poly_res_class_memE[of _ n d]) using assms a_def by blast show False using i_def 2 unfolding 4 2 by blast qed show ?thesis using 0 by blast qed definition int_fun_to_poly where "int_fun_to_poly (f::nat \ int) i = [(f i)]\\" lemma int_fun_to_poly_closed: assumes "\i. i > d \ f i = 0" shows "int_fun_to_poly f \ carrier (UP Q\<^sub>p)" apply(rule UPQ.UP_car_memI[of d]) using assms unfolding int_fun_to_poly_def using Qp.int_inc_zero apply presburger by(rule Qp.int_inc_closed) lemma int_fun_to_poly_deg: assumes "\i. i > d \ f i = 0" shows "deg Q\<^sub>p (int_fun_to_poly f) \ d" apply(rule UPQ.deg_leqI, rule int_fun_to_poly_closed, rule assms, blast) unfolding int_fun_to_poly_def using assms using Qp.int_inc_zero by presburger lemma Qp_res_mod_triv: assumes "a \ \\<^sub>p" shows "Qp_res a n mod p ^ n = Qp_res a n" using assms Qp_res_closed[of a n] by (meson mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) lemma int_fun_to_poly_is_class_wit: assumes "f \ poly_res_class n d g" shows "(int_fun_to_poly (\i::nat. Qp_res (f i) n)) \ poly_res_class n d g" proof(rule poly_res_class_memI[of ], rule int_fun_to_poly_closed[of d]) show 0: "\i. d < i \ Qp_res (f i) n = 0" proof- fix i assume A: "d < i" hence 0: "deg Q\<^sub>p f < i" using A assms poly_res_class_memE(2)[of f n d g] by linarith have 1: "f i = \" using 0 assms poly_res_class_memE[of f n d g] using UPQ.UP_car_memE(2) by blast show "Qp_res (f i) n = 0" unfolding 1 Qp_res_zero by blast qed show "deg Q\<^sub>p (int_fun_to_poly (\i. Qp_res (f i) n)) \ d" by(rule int_fun_to_poly_deg, rule 0, blast) show "\i. int_fun_to_poly (\i. Qp_res (f i) n) i \ \\<^sub>p" unfolding int_fun_to_poly_def using Qp.int_mult_closed Qp_val_ringI val_of_int_inc by blast show "\i. Qp_res (int_fun_to_poly (\i. Qp_res (f i) n) i) n = Qp_res (g i) n" unfolding int_fun_to_poly_def Qp_res_int_inc using Qp_res_mod_triv assms poly_res_class_memE(4) Qp_res_closed UPQ.cfs_closed by (metis poly_res_class_memE(3)) qed lemma finite_support_funs_finite: "finite (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" proof- have 0: "finite (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n))" apply(rule finite_PiE, blast) using residue_ring_card[of n] by blast obtain g where g_def: "g = (\f. (\i::nat. if i \ {..d} then f i else (0::int)))" by blast have 1: "g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) = (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0})" proof(rule equalityI, rule subsetI) fix x assume A: "x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" obtain f where f_def: "f \ (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)) \ x = g f" using A by blast have x_eq: "x = g f" using f_def by blast show "x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" proof(rule, rule) fix i assume A: "i \ {..d}" show "x i \ carrier (Zp_res_ring n)" proof(cases "i \ {..d}") case True then have T0: "f i \ carrier (Zp_res_ring n)" using f_def by blast have "x i = f i" unfolding x_eq g_def using True by metis thus ?thesis using T0 by metis next case False then have F0: "x i = 0" unfolding x_eq g_def by metis show ?thesis unfolding F0 by (metis residue_mult_closed residue_times_zero_r) qed next show "x \ {f. \i>d. f i = 0}" proof(rule, rule, rule) fix i assume A: "d < i" then have 0: "i \ {..d}" by simp thus "x i = 0" unfolding x_eq g_def by metis qed qed next show "({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" proof(rule subsetI) fix x assume A: " x \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" show " x \ g ` ({..d} \\<^sub>E carrier (residue_ring (p ^ n)))" proof- obtain h where h_def: "h = restrict x {..d}" by blast have 0: "\i. i \ {..d} \ h i = x i" unfolding h_def restrict_apply by metis have 1: "\i. i \ {..d} \ h i = undefined" unfolding h_def restrict_apply by metis have 2: "\i. i \ {..d} \ h i \ carrier (Zp_res_ring n)" using A 0 unfolding 0 by blast have 3: "h \ {..d} \\<^sub>E carrier (residue_ring (p ^ n))" by(rule, rule 2, blast, rule 1, blast) have 4: "\i. i \ {..d} \ x i = 0" using A unfolding Int_iff mem_Collect_eq by (metis atMost_iff eq_imp_le le_simps(1) linorder_neqE_nat) have 5: "x = g h" proof fix i show "x i = g h i" unfolding g_def apply(cases "i \ {..d}") using 0 apply metis unfolding 4 by metis qed show ?thesis unfolding 5 using 3 by blast qed qed qed have 2: "finite (g ` (\\<^sub>E i \ {..d}.carrier (Zp_res_ring n)))" using 0 by blast show ?thesis using 2 unfolding 1 by blast qed lemma poly_res_classes_finite: "finite (poly_res_classes n d)" proof- have 0: "poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}) = poly_res_classes n d" proof(rule equalityI, rule subsetI) fix x assume A: " x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" then obtain f where f_def: "f \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0} \ x = poly_res_class n d (int_fun_to_poly f)" by blast have x_eq: "x = poly_res_class n d (int_fun_to_poly f)" using f_def by blast show "x \ poly_res_classes n d" proof- have 0: "int_fun_to_poly f \ val_ring_polys_grad d" apply(rule val_ring_polys_grad_memI, rule int_fun_to_poly_closed[of d]) using f_def apply blast using int_fun_to_poly_def apply (metis Qp.int_inc_closed padic_fields.int_fun_to_poly_def padic_fields.val_of_int_inc padic_fields_axioms val_ring_memI) apply(rule int_fun_to_poly_deg) using f_def by blast show ?thesis unfolding poly_res_classes_def x_eq using 0 by blast qed next show "poly_res_classes n d \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" proof(rule subsetI) fix x assume A: " x \ poly_res_classes n d" show "x \ poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0})" proof- obtain f where f_def: "f \ val_ring_polys_grad d \ x = poly_res_class n d f" using A unfolding poly_res_classes_def by blast have x_eq: "x = poly_res_class n d f" using f_def by blast obtain h where h_def: "h = (\i::nat. Qp_res (f i) n)" by blast have 0: "\i. i > d \ f i = \" proof- fix i assume A: "i > d" have "i > deg Q\<^sub>p f" apply(rule le_less_trans[of _ d]) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast by(rule A) then show "f i = \" using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq using UPQ.deg_leE by blast qed have 1: "\i. i > d \ h i = 0" unfolding h_def 0 Qp_res_zero by blast have 2: "x = poly_res_class n d (int_fun_to_poly h)" unfolding x_eq apply(rule poly_res_class_closed) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast apply(rule int_fun_to_poly_closed[of d], rule 1, blast) using f_def unfolding val_ring_polys_grad_def val_ring_polys_def mem_Collect_eq apply blast apply(rule int_fun_to_poly_deg, rule 1, blast) unfolding h_def apply(rule int_fun_to_poly_is_class_wit, rule poly_res_class_refl) using f_def by blast have 3: "h \ ({..d} \ carrier (residue_ring (p ^ n))) \ {f. \i>d. f i = 0}" apply(rule , rule ) unfolding h_def apply(rule Qp_res_closed, rule val_ring_polys_grad_memE[of _ d]) using f_def apply blast unfolding mem_Collect_eq apply(rule, rule) unfolding 0 Qp_res_zero by blast show ?thesis unfolding 2 using 3 by blast qed qed qed have 1: "finite (poly_res_class n d ` int_fun_to_poly ` (({..d} \ carrier (Zp_res_ring n)) \ {(f::nat \ int). \i > d. f i = 0}))" using finite_support_funs_finite by blast show ?thesis using 1 unfolding 0 by blast qed lemma Qp_res_eq_zeroI: assumes "a \ \\<^sub>p" assumes "val a \ n" shows "Qp_res a n = 0" proof- have 0: "val_Zp (to_Zp a) \ n" using assms to_Zp_val by presburger have 1: "to_Zp a n = 0" apply(rule zero_below_val_Zp, rule to_Zp_closed) using val_ring_closed assms apply blast by(rule 0) thus ?thesis unfolding Qp_res_def by blast qed lemma Qp_res_eqI: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "Qp_res (a \ b) n = 0" shows "Qp_res a n = Qp_res b n" using assms by (metis Qp_res_def val_ring_memE res_diff_zero_fact(1) to_Zp_closed to_Zp_minus) lemma Qp_res_eqI': assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "val (a \ b) \ n" shows "Qp_res a n = Qp_res b n" apply(rule Qp_res_eqI, rule assms, rule assms, rule Qp_res_eq_zeroI) using assms Q\<^sub>p_def Zp_def \_def padic_fields.val_ring_minus_closed padic_fields_axioms apply blast by(rule assms) lemma Qp_res_eqE: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "Qp_res a n = Qp_res b n" shows "val (a \ b) \ n" proof- have 0: "val (a \ b) = val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b)" using assms by (metis to_Zp_minus to_Zp_val val_ring_minus_closed) have 1: "(to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) n = 0" using assms unfolding Qp_res_def by (meson val_ring_memE res_diff_zero_fact'' to_Zp_closed) have 2: "val_Zp (to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b) \ n" apply(cases "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>") proof - assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b = \\<^bsub>Z\<^sub>p\<^esub>" have "\n. eint (int n) \ val_Zp \\<^bsub>Z\<^sub>p\<^esub>" by (metis (no_types) Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) then show ?thesis using a1 by presburger next assume a1: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ \\<^bsub>Z\<^sub>p\<^esub>" have 00: "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b \ carrier Z\<^sub>p" using assms by (meson val_ring_memE Zp.cring_simprules(4) to_Zp_closed) show ?thesis using 1 a1 ord_Zp_geq[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b" n] 00 val_ord_Zp[of "to_Zp a \\<^bsub>Z\<^sub>p\<^esub> to_Zp b"] eint_ord_code by metis qed thus ?thesis unfolding 0 by blast qed lemma notin_closed: "(\ ((c::eint) \ x \ x \ d)) = (x < c \ d < x)" by auto lemma Qp_res_neqI: assumes "a \ \\<^sub>p" assumes "b \ \\<^sub>p" assumes "val (a \ b) < n" shows "Qp_res a n \ Qp_res b n" apply(rule ccontr) using Qp_res_eqE[of a b n] assms using notin_closed by blast lemma Qp_res_equal: assumes "a \ \\<^sub>p" assumes "l = Qp_res a n" shows "Qp_res a n = Qp_res ([l]\\) n " unfolding Qp_res_int_inc assms using assms Qp_res_mod_triv by presburger definition Qp_res_class where "Qp_res_class n b = {a \ \\<^sub>p. Qp_res a n = Qp_res b n}" definition Qp_res_classes where "Qp_res_classes n = Qp_res_class n ` \\<^sub>p" lemma val_ring_int_inc_closed: "[(k::int)]\\ \ \\<^sub>p" proof- have 0: "[(k::int)]\\ = \ ([(k::int)]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using inc_of_int by blast thus ?thesis by blast qed lemma val_ring_nat_inc_closed: "[(k::nat)]\\ \ \\<^sub>p" proof- have 0: "[k]\\ = \ ([k]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" using inc_of_nat by blast thus ?thesis by blast qed lemma Qp_res_classes_wits: "Qp_res_classes n = (\l::int. Qp_res_class n ([l]\\)) ` (carrier (Zp_res_ring n))" proof- obtain F where F_def: "F = (\l::int. Qp_res_class n ([l]\\))" by blast have 0: "Qp_res_classes n = F ` (carrier (Zp_res_ring n))" proof(rule equalityI, rule subsetI) fix x assume A: "x \ Qp_res_classes n" then obtain a where a_def: "a \ \\<^sub>p \ x = Qp_res_class n a" unfolding Qp_res_classes_def by blast have 1: "Qp_res a n = Qp_res ([(Qp_res a n)]\\) n " apply(rule Qp_res_equal) using a_def apply blast by blast have 2: "Qp_res_class n a = Qp_res_class n ([(Qp_res a n)]\\)" unfolding Qp_res_class_def using 1 by metis have 3: "x = Qp_res_class n ([(Qp_res a n)]\\)" using a_def unfolding 2 by blast have 4: "a \ \\<^sub>p" using a_def by blast show " x \ F ` carrier (Zp_res_ring n)" unfolding F_def 3 using Qp_res_closed[of a n] 4 by blast next show "F ` carrier (residue_ring (p ^ n)) \ Qp_res_classes n" proof(rule subsetI) fix x assume A: "x \ F ` (carrier (Zp_res_ring n))" then obtain l where l_def: "l \ carrier (Zp_res_ring n) \ x = F l" using A by blast have 0: "x = F l" using l_def by blast show "x \ Qp_res_classes n" unfolding 0 F_def Qp_res_classes_def using val_ring_int_inc_closed by blast qed qed then show ?thesis unfolding F_def by blast qed lemma Qp_res_classes_finite: "finite (Qp_res_classes n)" by (metis Qp_res_classes_wits finite_atLeastLessThan_int finite_imageI p_res_ring_car) definition Qp_cong_set where "Qp_cong_set \ a = {x \ \\<^sub>p. to_Zp x \ = a \}" lemma Qp_cong_set_as_ball: assumes "a \ carrier Z\<^sub>p" assumes "a \ = 0" shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[\]" proof- have 0: "\ a \ carrier Q\<^sub>p" using assms inc_closed[of a] by blast show ?thesis proof show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\]" proof fix x assume A: "x \ Qp_cong_set \ a" show "x \ B\<^bsub>\ \<^esub>[\]" proof(rule c_ballI) show t0: "x \ carrier Q\<^sub>p" using A unfolding Qp_cong_set_def using val_ring_memE by blast show "eint (int \) \ val (x \ \)" proof- have t1: "to_Zp x \ = 0" using A unfolding Qp_cong_set_def by (metis (mono_tags, lifting) assms(2) mem_Collect_eq) have t2: "val_Zp (to_Zp x) \ \" apply(cases "to_Zp x = \\<^bsub>Z\<^sub>p\<^esub>") apply (metis Zp.r_right_minus_eq Zp.zero_closed val_Zp_dist_def val_Zp_dist_res_eq2) using ord_Zp_geq[of "to_Zp x" \] A unfolding Qp_cong_set_def by (metis (no_types, lifting) val_ring_memE eint_ord_simps(1) t1 to_Zp_closed to_Zp_def val_ord_Zp) then show ?thesis using A unfolding Qp_cong_set_def mem_Collect_eq using val_ring_memE by (metis Qp_res_eqE Qp_res_eq_zeroI Qp_res_zero to_Zp_val zero_in_val_ring) qed qed qed show "B\<^bsub>int \\<^esub>[\] \ Qp_cong_set \ a" proof fix x assume A: "x \ B\<^bsub>int \\<^esub>[\]" then have 0: "val x \ \" using assms c_ballE[of x \ \] by (smt Qp.minus_closed Qp.r_right_minus_eq Qp_diff_diff) have 1: "to_Zp x \ carrier Z\<^sub>p" using A 0 assms c_ballE(1) to_Zp_closed by blast have 2: "x \ \\<^sub>p" using 0 A val_ringI c_ballE by (smt Q\<^sub>p_def Zp_def \_def eint_ord_simps(1) of_nat_0 of_nat_le_0_iff val_ring_ord_criterion padic_fields_axioms val_ord' zero_in_val_ring) then have "val_Zp (to_Zp x) \ \" using 0 1 A assms c_ballE[of x \ \] to_Zp_val by presburger then have "to_Zp x \ = 0" using 1 zero_below_val_Zp by blast then show " x \ Qp_cong_set \ a" unfolding Qp_cong_set_def using assms(2) 2 by (metis (mono_tags, lifting) mem_Collect_eq) qed qed qed lemma Qp_cong_set_as_ball': assumes "a \ carrier Z\<^sub>p" assumes "val_Zp a < eint (int \)" shows "Qp_cong_set \ a = B\<^bsub>\\<^esub>[(\ a)]" proof show "Qp_cong_set \ a \ B\<^bsub>\\<^esub>[\ a]" proof fix x assume A: "x \ Qp_cong_set \ a" then have 0: "to_Zp x \ = a \" unfolding Qp_cong_set_def by blast have 1: "x \ \\<^sub>p" using A unfolding Qp_cong_set_def by blast have 2: "to_Zp x \ carrier Z\<^sub>p" using 1 val_ring_memE to_Zp_closed by blast have 3: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" using 0 assms 2 val_Zp_dist_def val_Zp_dist_res_eq2 by presburger have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) > val_Zp a" using 3 assms(2) less_le_trans[of "val_Zp a" "eint (int \)" "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a)" ] by blast then have 5: "val_Zp (to_Zp x) = val_Zp a" using assms 2 equal_val_Zp by blast have 7: "val (x \ (\ a)) \ \" using 3 5 1 by (metis "2" Zp.minus_closed assms(1) inc_of_diff to_Zp_inc val_of_inc) then show "x \ B\<^bsub>int \\<^esub>[\ a]" using c_ballI[of x \ "\ a"] 1 assms val_ring_memE by blast qed show "B\<^bsub>int \\<^esub>[\ a] \ Qp_cong_set \ a" proof fix x assume A: "x \ B\<^bsub>int \\<^esub>[\ a]" then have 0: "val (x \ \ a) \ \" using c_ballE by blast have 1: "val (\ a) = val_Zp a" using assms Zp_def \_def padic_fields.val_of_inc padic_fields_axioms by metis then have 2: "val (x \ \ a) > val (\ a)" using 0 assms less_le_trans[of "val (\ a)" "eint (int \)" "val (x \ \ a)"] by metis have "\ a \ carrier Q\<^sub>p" using assms(1) inc_closed by blast then have B: "val x = val (\ a)" using 2 A assms c_ballE(1)[of x \ "\ a"] by (metis ultrametric_equal_eq) have 3: "val_Zp (to_Zp x) = val_Zp a" by (metis "1" A \val x = val (\ a)\ assms(1) c_ballE(1) to_Zp_val val_pos val_ringI) have 4: "val_Zp (to_Zp x \\<^bsub>Z\<^sub>p\<^esub> a) \ \" using 0 A 3 by (metis B Zp.minus_closed assms(1) c_ballE(1) inc_of_diff to_Zp_closed to_Zp_inc val_of_inc val_pos val_ring_val_criterion) then have 5: "to_Zp x \ = a \" by (meson A Zp.minus_closed assms(1) c_ballE(1) res_diff_zero_fact(1) to_Zp_closed zero_below_val_Zp) have 6: "x \ \\<^sub>p" proof- have "val x \ 0" using B assms 1 val_pos by presburger then show ?thesis using A c_ballE(1) val_ringI by blast qed then show "x \ Qp_cong_set \ a" unfolding Qp_cong_set_def using "5" by blast qed qed lemma Qp_cong_set_is_univ_semialgebraic: assumes "a \ carrier Z\<^sub>p" shows "is_univ_semialgebraic (Qp_cong_set \ a)" proof(cases "a \ = 0") case True then show ?thesis using ball_is_univ_semialgebraic[of \ \] Qp.zero_closed Qp_cong_set_as_ball assms by metis next case False then have "\ \ 0" using assms residues_closed[of a 0] by (meson p_res_ring_0') then obtain n where n_def: "Suc n = \" by (metis lessI less_Suc_eq_0_disj) then have "val_Zp a < eint (int \)" using below_val_Zp_zero[of a n] by (smt False assms eint_ile eint_ord_simps(1) eint_ord_simps(2) zero_below_val_Zp) then show ?thesis using ball_is_univ_semialgebraic[of "\ a" \] Qp.zero_closed Qp_cong_set_as_ball'[of a \] assms inc_closed by presburger qed lemma constant_res_set_semialg: assumes "l \ carrier (Zp_res_ring n)" shows "is_univ_semialgebraic {x \ \\<^sub>p. Qp_res x n = l}" proof- have 0: "{x \ \\<^sub>p. Qp_res x n = l} = Qp_cong_set n ([l]\\<^bsub>Z\<^sub>p\<^esub>\\<^bsub>Z\<^sub>p\<^esub>)" apply(rule equalityI') unfolding mem_Collect_eq Qp_cong_set_def Qp_res_def apply (metis val_ring_memE Zp_int_inc_rep nat_le_linear p_residue_padic_int to_Zp_closed) using assms by (metis Zp_int_inc_res mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) show ?thesis unfolding 0 apply(rule Qp_cong_set_is_univ_semialgebraic) by(rule Zp.int_inc_closed) qed end end diff --git a/thys/Padic_Ints/Padic_Integers.thy b/thys/Padic_Ints/Padic_Integers.thy --- a/thys/Padic_Ints/Padic_Integers.thy +++ b/thys/Padic_Ints/Padic_Integers.thy @@ -1,2505 +1,2504 @@ theory Padic_Integers imports Padic_Construction Extended_Int Supplementary_Ring_Facts "HOL-Algebra.Subrings" "HOL-Number_Theory.Residues" "HOL-Algebra.Multiplicative_Group" begin text\ In what follows we establish a locale for reasoning about the ring of $p$-adic integers for a fixed prime $p$. We will elaborate on the basic metric properties of $\mathbb{Z}_p$ and construct the angular component maps to the residue rings. \ section\A Locale for $p$-adic Integer Rings\ locale padic_integers = fixes Zp:: "_ ring" (structure) fixes p defines "Zp \ padic_int p" assumes prime: "prime p" sublocale padic_integers < UPZ?: UP Zp "UP Zp" by simp sublocale padic_integers < Zp?:UP_cring Zp "UP Zp" unfolding UP_cring_def by(auto simp add: Zp_def padic_int_is_cring prime) sublocale padic_integers < Zp?:ring Zp using Zp_def cring.axioms(1) padic_int_is_cring prime by blast sublocale padic_integers < Zp?:cring Zp by (simp add: Zp_def padic_int_is_cring prime) sublocale padic_integers < Zp?:domain Zp by (simp add: Zp_def padic_int_is_domain padic_integers.prime padic_integers_axioms) context padic_integers begin lemma Zp_defs: "\ = padic_one p" "\ = padic_zero p" "carrier Zp = padic_set p" "(\) = padic_mult p" "(\) = padic_add p" unfolding Zp_def using padic_int_simps by auto end (*************************************************************************************************) (*************************************************************************************************) (***********************)section \Residue Rings\(*************************************) (*************************************************************************************************) (*************************************************************************************************) lemma(in field) field_inv: assumes "a \ carrier R" assumes "a \\" shows "inv\<^bsub>R\<^esub> a \ a = \" "a \ inv\<^bsub>R\<^esub> a = \" "inv \<^bsub>R\<^esub> a \ carrier R" proof- have "a \ Units R" using assms by (simp add: local.field_Units) then show "inv\<^bsub>R\<^esub> a \ a = \" by simp show "a \ inv a = \" using \a \ Units R\ by auto show "inv \<^bsub>R\<^esub> a \ carrier R" by (simp add: \a \ Units R\) qed text\$p_residue$ defines the standard projection maps between residue rings:\ definition(in padic_integers) p_residue:: "nat \ int \ _" where "p_residue m n \ residue (p^m) n" lemma(in padic_integers) p_residue_alt_def: "p_residue m n = n mod (p^m)" using residue_def by (simp add: p_residue_def) lemma(in padic_integers) p_residue_range: "p_residue m n \ {0.. k" shows "p_residue k (p_residue m n) = p_residue k n" using assms unfolding p_residue_def residue_def by (simp add: le_imp_power_dvd mod_mod_cancel) text\Compatibility of p\_residue with elements of $\mathbb{Z}_p$:\ lemma(in padic_integers) p_residue_padic_int: assumes "x \ carrier Zp" assumes "m \ k" shows "p_residue k (x m) = x k" using Zp_def assms(1) assms(2) padic_set_res_coherent prime by (simp add: p_residue_def padic_int_simps(5)) text\Definition of residue rings:\ abbreviation(in padic_integers) (input) Zp_res_ring:: "nat \ _ ring" where "(Zp_res_ring n) \ residue_ring (p^n)" lemma (in padic_integers) p_res_ring_zero: "\\<^bsub>Zp_res_ring k\<^esub> = 0" by (simp add: residue_ring_def) lemma (in padic_integers) p_res_ring_one: assumes "k > 0" shows "\\<^bsub>Zp_res_ring k\<^esub> = 1" using assms by (simp add: residue_ring_def) lemma (in padic_integers) p_res_ring_car: "carrier (Zp_res_ring k) = {0.. carrier (Zp_res_ring m)" using p_residue_range residue_ring_def prime prime_gt_0_nat p_residue_def by fastforce text\First residue ring is a field:\ lemma(in padic_integers) p_res_ring_1_field: "field (Zp_res_ring 1)" by (metis int_nat_eq power_one_right prime prime_ge_0_int prime_nat_int_transfer residues_prime.intro residues_prime.is_field) text\$0^{th}$ residue ring is the zero ring:\ lemma(in padic_integers) p_res_ring_0: "carrier (Zp_res_ring 0) = {0}" by (simp add: residue_ring_def) lemma(in padic_integers) p_res_ring_0': assumes "x \ carrier (Zp_res_ring 0)" shows "x = 0" using p_res_ring_0 assms by blast text\If $m>0$ then $Zp\_res\_ring m$ is an instance of the residues locale:\ lemma(in padic_integers) p_residues: assumes "m >0" shows "residues (p^m)" proof- have "p^m > 1" using assms by (simp add: prime prime_gt_1_int) then show "residues (p^m)" using less_imp_of_nat_less residues.intro by fastforce qed text\If $m>0$ then $Zp\_res\_ring m$ is a commutative ring:\ lemma(in padic_integers) R_cring: assumes "m >0" shows "cring (Zp_res_ring m)" using p_residues assms residues.cring by auto lemma(in padic_integers) R_comm_monoid: assumes "m >0" shows "comm_monoid (Zp_res_ring m)" by (simp add: assms p_residues residues.comm_monoid) lemma(in padic_integers) zero_rep: "\ = (\m. (p_residue m 0))" unfolding p_residue_def using Zp_defs(2) padic_zero_simp(1) residue_def residue_ring_def by auto text\The operations on residue rings are just the standard operations on the integers $\mod p^n$. This means that the basic closure properties and algebraic properties of operations on these rings hold for all integers, not just elements of the ring carrier:\ lemma(in padic_integers) residue_add: shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x + y) mod p^k" unfolding residue_ring_def by simp lemma(in padic_integers) residue_add_closed: shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" using p_residue_def p_residue_range residue_def residue_ring_def by auto lemma(in padic_integers) residue_add_closed': shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ {0..\<^bsub>Zp_res_ring k\<^esub> y) = (x * y) mod p^k" unfolding residue_ring_def by simp lemma(in padic_integers) residue_mult_closed: shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" using p_residue_def p_residue_range residue_def residue_ring_def by auto lemma(in padic_integers) residue_mult_closed': shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ {0..\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> z = x \\<^bsub>Zp_res_ring k\<^esub> (y \\<^bsub>Zp_res_ring k\<^esub> z)" using residue_add by (simp add: add.commute add.left_commute mod_add_right_eq) lemma(in padic_integers) residue_mult_comm: shows "x \\<^bsub>Zp_res_ring k\<^esub> y = y \\<^bsub>Zp_res_ring k\<^esub> x" using residue_mult by (simp add: mult.commute) lemma(in padic_integers) residue_mult_assoc: shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> z = x \\<^bsub>Zp_res_ring k\<^esub> (y \\<^bsub>Zp_res_ring k\<^esub> z)" using residue_mult by (simp add: mod_mult_left_eq mod_mult_right_eq mult.assoc) lemma(in padic_integers) residue_add_comm: shows "x \\<^bsub>Zp_res_ring k\<^esub> y = y \\<^bsub>Zp_res_ring k\<^esub> x" using residue_add by presburger lemma(in padic_integers) residue_minus_car: assumes "y \ carrier (Zp_res_ring k)" shows "(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x - y) mod p^k" proof(cases "k = 0") case True then show ?thesis using residue_ring_def a_minus_def by(simp add: a_minus_def residue_ring_def) next case False have "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = x \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> y \\<^bsub>Zp_res_ring k\<^esub> y)" by (simp add: a_minus_def residue_add_assoc) then have 0: "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = x mod p^k" using assms False by (smt cring.cring_simprules(9) prime residue_add residues.cring residues.res_zero_eq residues_n) have 1: "x mod p ^ k = ((x - y) mod p ^ k + y) mod p ^ k" proof - have f1: "x - y = x + - 1 * y" by auto have "y + (x + - 1 * y) = x" by simp then show ?thesis using f1 by presburger qed have "(x \\<^bsub>Zp_res_ring k\<^esub> y) \\<^bsub>Zp_res_ring k\<^esub> y = (x - y) mod p^k \\<^bsub>Zp_res_ring k\<^esub> y" using residue_add[of k "(x - y) mod p^k" y] 0 1 by linarith then show ?thesis using assms residue_add_closed by (metis False a_minus_def cring.cring_simprules(10) cring.cring_simprules(19) prime residues.cring residues.mod_in_carrier residues_n) qed lemma(in padic_integers) residue_a_inv: shows "\\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" proof- have "y \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) = (y mod p^k) \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) " using residue_minus_car[of "\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" k y] residue_add by (simp add: mod_add_left_eq) then have 0: "y \\<^bsub>Zp_res_ring k\<^esub> (\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) = \\<^bsub>Zp_res_ring k\<^esub>" by (metis cring.cring_simprules(17) p_res_ring_zero padic_integers.p_res_ring_0' padic_integers_axioms prime residue_add_closed residues.cring residues.mod_in_carrier residues_n) have 1: "(\\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)) \\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub>" using residue_add_comm 0 by auto have 2: "\x. x \ carrier (Zp_res_ring k) \ x \\<^bsub>Zp_res_ring k\<^esub> y = \\<^bsub>Zp_res_ring k\<^esub> \ y \\<^bsub>Zp_res_ring k\<^esub> x = \\<^bsub>Zp_res_ring k\<^esub> \ x = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" using 0 1 by (metis cring.cring_simprules(3) cring.cring_simprules(8) mod_by_1 padic_integers.p_res_ring_0' padic_integers.p_res_ring_zero padic_integers_axioms power_0 prime residue_1_prop residue_add_assoc residues.cring residues.mod_in_carrier residues_n) have 3: "carrier (add_monoid (residue_ring (p ^ k))) = carrier (Zp_res_ring k)" by simp have 4: "(\\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub>) = (\\<^bsub>Zp_res_ring k\<^esub>)" by simp have 5: "\x. x \ carrier (add_monoid (residue_ring (p ^ k))) \ x \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> y = \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> \ y \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> x = \\<^bsub>add_monoid (residue_ring (p ^ k))\<^esub> \ x = \\<^bsub>Zp_res_ring k\<^esub> (y mod p^k)" using 0 1 2 3 4 by simp show ?thesis unfolding a_inv_def m_inv_def apply(rule the_equality) using 1 2 3 4 5 unfolding a_inv_def m_inv_def apply (metis (no_types, lifting) "0" "1" cring.cring_simprules(3) mod_by_1 monoid.select_convs(2) padic_integers.p_res_ring_zero padic_integers_axioms power_0 prime residue_1_prop residue_add_closed residues.cring residues.mod_in_carrier residues_n) using 1 2 3 4 5 unfolding a_inv_def m_inv_def by blast qed lemma(in padic_integers) residue_a_inv_closed: "\\<^bsub>Zp_res_ring k\<^esub> y \ carrier (Zp_res_ring k)" apply(cases "k = 0") apply (metis add.comm_neutral add.commute atLeastLessThanPlusOne_atLeastAtMost_int insert_iff mod_by_1 p_res_ring_car p_res_ring_zero padic_integers.p_res_ring_0 padic_integers_axioms power_0 residue_1_prop residue_a_inv) by (simp add: prime residues.mod_in_carrier residues.res_neg_eq residues_n) lemma(in padic_integers) residue_minus: "(x \\<^bsub>Zp_res_ring k\<^esub> y) = (x - y) mod p^k" using residue_minus_car[of "y mod p^k" k x] residue_a_inv[of k y] unfolding a_minus_def by (metis a_minus_def mod_diff_right_eq p_residue_alt_def p_residue_range') lemma(in padic_integers) residue_minus_closed: "(x \\<^bsub>Zp_res_ring k\<^esub> y) \ carrier (Zp_res_ring k)" by (simp add: a_minus_def residue_add_closed) lemma (in padic_integers) residue_plus_zero_r: "0 \\<^bsub>Zp_res_ring k\<^esub> y = y mod p^k" by (simp add: residue_add) lemma (in padic_integers) residue_plus_zero_l: "y \\<^bsub>Zp_res_ring k\<^esub> 0 = y mod p^k" by (simp add: residue_add) lemma (in padic_integers) residue_times_zero_r: "0 \\<^bsub>Zp_res_ring k\<^esub> y = 0" by (simp add: residue_mult) lemma (in padic_integers) residue_times_zero_l: "y \\<^bsub>Zp_res_ring k\<^esub> 0 = 0" by (simp add: residue_mult) lemma (in padic_integers) residue_times_one_r: "1 \\<^bsub>Zp_res_ring k\<^esub> y = y mod p^k" by (simp add: residue_mult) lemma (in padic_integers) residue_times_one_l: "y \\<^bsub>Zp_res_ring k\<^esub> 1 = y mod p^k" by (simp add: residue_mult_comm residue_times_one_r) text\Similarly to the previous lemmas, many identities about taking residues of $p$-adic integers hold even for elements which lie outside the carrier of $\mathbb{Z}_p$:\ lemma (in padic_integers) residue_of_sum: "(a \ b) k = (a k) \\<^bsub>Zp_res_ring k\<^esub> (b k)" using Zp_def residue_ring_def[of "p^k"] Zp_defs(5) padic_add_res by auto lemma (in padic_integers) residue_of_sum': "(a \ b) k = ((a k) + (b k)) mod p^k" using residue_add residue_of_sum by auto lemma (in padic_integers) residue_closed[simp]: assumes "b \ carrier Zp" shows "b k \ carrier (Zp_res_ring k)" using Zp_def assms padic_integers.Zp_defs(3) padic_integers_axioms padic_set_res_closed by auto lemma (in padic_integers) residue_of_diff: assumes "b \ carrier Zp" shows "(a \ b) k = (a k) \\<^bsub>Zp_res_ring k\<^esub> (b k)" unfolding a_minus_def using Zp_def add.inv_closed assms(1) padic_a_inv prime residue_of_sum by auto lemma (in padic_integers) residue_of_prod: "(a \ b) k = (a k) \ \<^bsub>Zp_res_ring k\<^esub> (b k)" by (simp add: Zp_defs(4) padic_mult_def) lemma (in padic_integers) residue_of_prod': "(a \ b) k = ((a k) * (b k)) mod (p^k)" by (simp add: residue_mult residue_of_prod) lemma (in padic_integers) residue_of_one: assumes "k > 0" shows "\ k = \\<^bsub>Zp_res_ring k\<^esub>" "\ k = 1" apply (simp add: Zp_defs(1) assms padic_one_simp(1)) by (simp add: Zp_def assms padic_int_simps(1) padic_one_simp(1) residue_ring_def) lemma (in padic_integers) residue_of_zero: shows "\ k = \\<^bsub>Zp_res_ring k\<^esub>" "\ k = 0" apply (simp add: Zp_defs(2) padic_zero_simp(1)) by (simp add: p_residue_alt_def zero_rep) lemma(in padic_integers) Zp_residue_mult_zero: assumes "a k = 0" shows "(a \ b) k = 0" "(b \ a) k = 0" using assms residue_of_prod' by auto lemma(in padic_integers) Zp_residue_add_zero: assumes "b \ carrier Zp" assumes "(a:: padic_int) k = 0" shows "(a \ b) k = b k" "(b \ a) k = b k" apply (metis Zp_def assms(1) assms(2) cring.cring_simprules(8) mod_by_1 padic_int_is_cring power.simps(1) prime residue_add_closed residue_of_sum residue_of_sum' residues.cring residues.res_zero_eq residues_n) by (metis Zp_def assms(1) assms(2) cring.cring_simprules(16) mod_by_1 padic_int_is_cring power.simps(1) prime residue_add_closed residue_of_sum residue_of_sum' residues.cring residues.res_zero_eq residues_n) text\P-adic addition and multiplication are globally additive and associative:\ lemma padic_add_assoc0: assumes "prime p" shows "padic_add p (padic_add p x y) z = padic_add p x (padic_add p y z)" using assms unfolding padic_add_def by (simp add: padic_integers.residue_add_assoc padic_integers_def) lemma(in padic_integers) add_assoc: "a \ b \ c = a \ (b \ c)" using padic_add_assoc0[of p a b c] prime Zp_defs by auto lemma padic_add_comm0: assumes "prime p" shows "(padic_add p x y)= (padic_add p y x)" using assms unfolding padic_add_def using padic_integers.residue_add_comm[of p] by (simp add: padic_integers_def) lemma(in padic_integers) add_comm: "a \ b = b \ a" using padic_add_comm0[of p a b] prime Zp_defs by auto lemma padic_mult_assoc0: assumes "prime p" shows "padic_mult p (padic_mult p x y) z = padic_mult p x (padic_mult p y z)" using assms unfolding padic_mult_def by (simp add: padic_integers.residue_mult_assoc padic_integers_def) lemma(in padic_integers) mult_assoc: "a \ b \ c = a \ (b \ c)" using padic_mult_assoc0[of p a b c] prime Zp_defs by auto lemma padic_mult_comm0: assumes "prime p" shows "(padic_mult p x y)= (padic_mult p y x)" using assms unfolding padic_mult_def using padic_integers.residue_mult_comm[of p] by (simp add: padic_integers_def) lemma(in padic_integers) mult_comm: "a \ b = b \ a" using padic_mult_comm0[of p a b] prime Zp_defs by auto lemma(in padic_integers) mult_zero_l: "a \ \ = \" proof fix x show "(a \ \) x = \ x" using Zp_residue_mult_zero[of \ x a] by (simp add: residue_of_zero(2)) qed lemma(in padic_integers) mult_zero_r: "\ \ a = \" using mult_zero_l mult_comm by auto lemma (in padic_integers) p_residue_ring_car_memI: assumes "(m::int) \0" assumes "m < p^k" shows "m \ carrier (Zp_res_ring k)" using residue_ring_def[of "p^k"] assms(1) assms(2) by auto lemma (in padic_integers) p_residue_ring_car_memE: assumes "m \ carrier (Zp_res_ring k)" shows "m < p^k" "m \ 0" using assms residue_ring_def by auto lemma (in padic_integers) residues_closed: assumes "a \ carrier Zp" shows "a k \ carrier (Zp_res_ring k)" using Zp_def assms padic_integers.Zp_defs(3) padic_integers_axioms padic_set_res_closed by blast lemma (in padic_integers) mod_in_carrier: "a mod (p^n) \ carrier (Zp_res_ring n)" using p_residue_alt_def p_residue_range' by auto lemma (in padic_integers) Zp_residue_a_inv: assumes "a \ carrier Zp" shows "(\ a) k = \\<^bsub>Zp_res_ring k\<^esub> (a k)" "(\ a) k = (- (a k)) mod (p^k)" using Zp_def assms padic_a_inv prime apply auto[1] using residue_a_inv by (metis Zp_def assms mod_by_1 p_res_ring_zero padic_a_inv padic_integers.prime padic_integers_axioms power_0 residue_1_prop residues.res_neg_eq residues_n) lemma (in padic_integers) residue_of_diff': assumes "b \ carrier Zp" shows "(a \ b) k = ((a k) - (b k)) mod (p^k)" by (simp add: assms residue_minus_car residue_of_diff residues_closed) lemma (in padic_integers) residue_UnitsI: assumes "n \ 1" assumes "(k::int) > 0" assumes "k < p^n" assumes "coprime k p" shows "k \ Units (Zp_res_ring n)" using residues.res_units_eq assms by (metis (mono_tags, lifting) coprime_power_right_iff mem_Collect_eq not_one_le_zero prime residues_n) lemma (in padic_integers) residue_UnitsE: assumes "n \ 1" assumes "k \ Units (Zp_res_ring n)" shows "coprime k p" using residues.res_units_eq assms by (simp add: p_residues) lemma(in padic_integers) residue_units_nilpotent: assumes "m > 0" assumes "k = card (Units (Zp_res_ring m))" assumes "x \ (Units (Zp_res_ring m))" shows "x[^]\<^bsub>Zp_res_ring m\<^esub> k = 1" proof- have 0: "group (units_of (Zp_res_ring m))" using assms(1) cring_def monoid.units_group padic_integers.R_cring padic_integers_axioms ring_def by blast have 1: "finite (Units (Zp_res_ring m))" using p_residues assms(1) residues.finite_Units by auto have 2: "x[^]\<^bsub>units_of (Zp_res_ring m)\<^esub> (order (units_of (Zp_res_ring m))) = \\<^bsub>units_of (Zp_res_ring m)\<^esub>" by (metis "0" assms(3) group.pow_order_eq_1 units_of_carrier) then show ?thesis by (metis "1" assms(1) assms(2) assms(3) cring.units_power_order_eq_one padic_integers.R_cring padic_integers.p_residues padic_integers_axioms residues.res_one_eq) qed lemma (in padic_integers) residue_1_unit: assumes "m > 0" shows "1 \ Units (Zp_res_ring m)" "\\<^bsub>Zp_res_ring m\<^esub> \ Units (Zp_res_ring m)" proof- have 0: "group (units_of (Zp_res_ring m))" using assms(1) cring_def monoid.units_group padic_integers.R_cring padic_integers_axioms ring_def by blast have 1: "1 = \\<^bsub>units_of (Zp_res_ring m)\<^esub>" by (simp add: residue_ring_def units_of_def) have "\\<^bsub>units_of (Zp_res_ring m)\<^esub> \ carrier (units_of (Zp_res_ring m))" using 0 Group.monoid.intro[of "units_of (Zp_res_ring m)"] by (simp add: group.is_monoid) then show "1 \ Units (Zp_res_ring m)" using 1 by (simp add: units_of_carrier) then show " \\<^bsub>Zp_res_ring m\<^esub> \ Units (Zp_res_ring m) " by (simp add: "1" units_of_one) qed lemma (in padic_integers) zero_not_in_residue_units: assumes "n \ 1" shows "(0::int) \ Units (Zp_res_ring n)" using assms p_residues residues.res_units_eq by auto text\Cardinality bounds on the units of residue rings:\ lemma (in padic_integers) residue_units_card_geq_2: assumes "n \2" shows "card (Units (Zp_res_ring n)) \ 2" proof(cases "p = 2") case True then have "(3::int) \ Units (Zp_res_ring n)" proof- have "p \2" by (simp add: True) then have "p^n \ 2^n" using assms by (simp add: True) then have "p^n \ 4" using assms power_increasing[of 2 n 2] by (simp add: True) then have "(3::int) < p^n" by linarith then have 0: "(3::int) \ carrier (Zp_res_ring n)" by (simp add: residue_ring_def) have 1: "coprime 3 p" by (simp add: True numeral_3_eq_3) show ?thesis using residue_UnitsI[of n "3::int"] using "1" \3 < p ^ n\ assms by linarith qed then have 0: "{(1::int), 3} \ Units (Zp_res_ring n)" using assms padic_integers.residue_1_unit padic_integers_axioms by auto have 1: "finite (Units (Zp_res_ring n))" using assms padic_integers.p_residues padic_integers_axioms residues.finite_Units by auto have 2: "{(1::int),3}\Units (Zp_res_ring n)" using "0" by auto have 3: "card {(1::int),3} = 2" by simp show ?thesis using 2 1 by (metis "3" card_mono) next case False have "1 \ Units (Zp_res_ring n)" using assms padic_integers.residue_1_unit padic_integers_axioms by auto have "2 \ Units (Zp_res_ring n)" apply(rule residue_UnitsI) using assms apply linarith apply simp proof- show "2 < p^n" proof- have "p^n > p" by (metis One_nat_def assms le_simps(3) numerals(2) power_one_right power_strict_increasing_iff prime prime_gt_1_int) then show ?thesis using False prime prime_gt_1_int[of p] by auto qed show "coprime 2 p" using False by (metis of_nat_numeral prime prime_nat_int_transfer primes_coprime two_is_prime_nat) qed then have 0: "{(1::int), 2} \ Units (Zp_res_ring n)" using \1 \ Units (Zp_res_ring n)\ by blast have 1: "card {(1::int),2} = 2" by simp then show ?thesis using residues.finite_Units 0 by (metis One_nat_def assms card_mono dual_order.trans le_simps(3) one_le_numeral padic_integers.p_residues padic_integers_axioms) qed lemma (in padic_integers) residue_ring_card: "finite (carrier (Zp_res_ring n)) \ card (carrier (Zp_res_ring n)) = nat (p^n)" using p_res_ring_car[of n] by simp lemma(in comm_monoid) UnitsI: assumes "a \ carrier G" assumes "b \ carrier G" assumes "a \ b = \" shows "a \ Units G" "b \ Units G" unfolding Units_def using comm_monoid_axioms_def assms m_comm[of a b] by auto lemma(in comm_monoid) is_invI: assumes "a \ carrier G" assumes "b \ carrier G" assumes "a \ b = \" shows "inv\<^bsub>G\<^esub> b = a" "inv\<^bsub>G\<^esub> a = b" using assms inv_char m_comm by auto lemma (in padic_integers) residue_of_Units: assumes "k > 0" assumes "a \ Units Zp" shows "a k \ Units (Zp_res_ring k)" proof- have 0: "a k \\<^bsub>Zp_res_ring k\<^esub> (inv \<^bsub>Zp\<^esub> a) k = 1" by (metis R.Units_r_inv assms(1) assms(2) residue_of_one(2) residue_of_prod) have 1: "a k \ carrier (Zp_res_ring k)" by (simp add: R.Units_closed assms(2) residues_closed) have 2: "(inv \<^bsub>Zp\<^esub> a) k \ carrier (Zp_res_ring k)" by (simp add: assms(2) residues_closed) show ?thesis using 0 1 2 comm_monoid.UnitsI[of "Zp_res_ring k"] using assms(1) p_residues residues.comm_monoid residues.res_one_eq by presburger qed (*************************************************************************************************) (*************************************************************************************************) (**********************)section\$int$ and $nat$ inclusions in $\mathbb{Z}_p$.\ (****************************) (*************************************************************************************************) (*************************************************************************************************) lemma(in ring) int_inc_zero: "[(0::int)]\ \ = \" by (simp add: add.int_pow_eq_id) lemma(in ring) int_inc_zero': assumes "x \ carrier R" shows "[(0::int)] \ x = \" by (simp add: add.int_pow_eq_id assms) lemma(in ring) nat_inc_zero: "[(0::nat)]\ \ = \" by auto lemma(in ring) nat_mult_zero: "[(0::nat)]\ x = \" by simp lemma(in ring) nat_inc_closed: fixes n::nat shows "[n] \ \ \ carrier R" by simp lemma(in ring) nat_mult_closed: fixes n::nat assumes "x \ carrier R" shows "[n] \ x \ carrier R" by (simp add: assms) lemma(in ring) int_inc_closed: fixes n::int shows "[n] \ \ \ carrier R" by simp lemma(in ring) int_mult_closed: fixes n::int assumes "x \ carrier R" shows "[n] \ x \ carrier R" by (simp add: assms) lemma(in ring) nat_inc_prod: fixes n::nat fixes m::nat shows "[m]\([n] \ \) = [(m*n)]\\" by (simp add: add.nat_pow_pow mult.commute) lemma(in ring) nat_inc_prod': fixes n::nat fixes m::nat shows "[(m*n)]\\ = [m]\ \ \ ([n] \ \)" by (simp add: add.nat_pow_pow add_pow_rdistr) lemma(in padic_integers) Zp_nat_inc_zero: shows "[(0::nat)] \ x = \" by simp lemma(in padic_integers) Zp_int_inc_zero': shows "[(0::int)] \ x = \" using Zp_nat_inc_zero[of x] unfolding add_pow_def int_pow_def by auto lemma(in padic_integers) Zp_nat_inc_closed: fixes n::nat shows "[n] \ \ \ carrier Zp" by simp lemma(in padic_integers) Zp_nat_mult_closed: fixes n::nat assumes "x \ carrier Zp" shows "[n] \ x \ carrier Zp" by (simp add: assms) lemma(in padic_integers) Zp_int_inc_closed: fixes n::int shows "[n] \ \ \ carrier Zp" by simp lemma(in padic_integers) Zp_int_mult_closed: fixes n::int assumes "x \ carrier Zp" shows "[n] \ x \ carrier Zp" by (simp add: assms) text\The following lemmas give a concrete description of the inclusion of integers and natural numbers into $\mathbb{Z}_p$:\ lemma(in padic_integers) Zp_nat_inc_rep: fixes n::nat shows "[n] \ \ = (\ m. p_residue m n)" apply(induction n) apply (simp add: zero_rep) proof- case (Suc n) fix n assume A: "[n] \ \ = (\m. p_residue m (int n))" then have 0: "[Suc n] \ \ = [n]\\ \ \" by auto show "[Suc n] \ \ = (\m. p_residue m (Suc n))" proof fix m show "([Suc n] \ \) m = p_residue m (int (Suc n)) " proof(cases "m=0") case True have 0: "([Suc n] \ \) m \ carrier (Zp_res_ring m)" using Zp_nat_inc_closed padic_set_res_closed by (simp add: residues_closed) then have "([Suc n] \ \) m = 0" using p_res_ring_0 True by blast then show ?thesis by (metis True p_res_ring_0' p_residue_range') next case False then have R: "residues (p^m)" by (simp add: prime residues_n) have "([Suc n] \ \) m = ([n]\\ \ \) m" by (simp add: "0") then have P0: "([Suc n] \ \) m = p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> \\<^bsub>Zp_res_ring m\<^esub>" using A False Zp_def padic_add_res padic_one_def Zp_defs(5) padic_integers.residue_of_one(1) padic_integers_axioms by auto then have P1:"([Suc n] \ \) m = p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m (1::int)" by (metis R ext p_residue_alt_def residue_add_assoc residue_add_comm residue_plus_zero_r residue_times_one_r residue_times_zero_l residues.res_one_eq) have P2: "p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m 1 = ((int n) mod (p^m)) \\<^bsub>Zp_res_ring m\<^esub> 1" using R P0 P1 residue_def residues.res_one_eq by (simp add: residues.res_one_eq p_residue_alt_def) have P3:"((int n) mod (p^m)) \\<^bsub>Zp_res_ring m\<^esub> 1 = ((int n) + 1) mod (p^m)" using R residue_ring_def by (simp add: mod_add_left_eq) have "p_residue m (int n) \\<^bsub>Zp_res_ring m\<^esub> p_residue m 1 = (int (Suc n)) mod (p^m)" by (metis P2 P3 add.commute of_nat_Suc p_residue_alt_def residue_add) then show ?thesis using False R P1 p_residue_def p_residue_alt_def by auto qed qed qed lemma(in padic_integers) Zp_nat_inc_res: fixes n::nat shows "([n] \ \) k = n mod (p^k)" using Zp_nat_inc_rep p_residue_def by (simp add: p_residue_alt_def) lemma(in padic_integers) Zp_int_inc_rep: fixes n::int shows "[n] \ \ = (\ m. p_residue m n )" proof(induction n) case (nonneg n) then show ?case using Zp_nat_inc_rep by (simp add: add_pow_int_ge) next case (neg n) show "\n. [(- int (Suc n))] \ \ = (\m. p_residue m (- int (Suc n)))" proof fix n fix m show "([(- int (Suc n))] \ \) m = p_residue m (- int (Suc n))" proof- have "[(- int (Suc n))] \ \ = \ ([(int (Suc n))] \ \)" using a_inv_def abelian_group.a_group add_pow_def cring.axioms(1) domain_def negative_zless_0 ring_def R.add.int_pow_neg R.one_closed by blast then have "([(- int (Suc n))] \ \) m = (\ ([(int (Suc n))] \ \)) m" by simp have "\ \ carrier Zp" using cring.cring_simprules(6) domain_def by blast have "([(int (Suc n))] \ \) = ([(Suc n)] \ \)" by (metis add_pow_def int_pow_int) then have "([(int (Suc n))] \ \) \ carrier Zp" using Zp_nat_inc_closed by simp then have P0: "([(- int (Suc n))] \ \) m = \\<^bsub>Zp_res_ring m\<^esub> (([(int (Suc n))] \ \) m)" using Zp_def prime using \[(- int (Suc n))] \ \ = \ ([int (Suc n)] \ \)\ padic_integers.Zp_residue_a_inv(1) padic_integers_axioms by auto have "(([(int (Suc n))] \ \) m) = (p_residue m (Suc n))" using Zp_nat_inc_rep by (simp add: add_pow_int_ge) then have P1: "([(- int (Suc n))] \ \) m = \\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n))" using P0 by auto have "\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) = p_residue m (- int (Suc n))" proof(cases "m=0") case True then have 0:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub>(p_residue 0 (Suc n))" by blast then have 1:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub> (p_residue 1 (Suc n))" by (metis p_res_ring_0' residue_a_inv_closed) then have 2:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =\\<^bsub>Zp_res_ring 0\<^esub> 0" by (metis p_res_ring_0' residue_a_inv_closed) then have 3:"\\<^bsub>Zp_res_ring m\<^esub>(p_residue m (Suc n)) =0" using residue_1_prop p_res_ring_0' residue_a_inv_closed by presburger have 4: "p_residue m (- int (Suc n)) \ carrier (Zp_res_ring 0)" using p_res_ring_0 True residue_1_zero p_residue_range' by blast then show ?thesis using "3" True residue_1_zero by (simp add: p_res_ring_0') next case False then have R: "residues (p^m)" using padic_integers.p_residues padic_integers_axioms by blast have "\\<^bsub>Zp_res_ring m\<^esub> p_residue m (int (Suc n)) = \\<^bsub>Zp_res_ring m\<^esub> (int (Suc n)) mod (p^m) " using R residue_def residues.neg_cong residues.res_neg_eq p_residue_alt_def by auto then have "\\<^bsub>Zp_res_ring m\<^esub> p_residue m (int (Suc n)) = (-(int (Suc n))) mod (p^m)" using R residues.res_neg_eq by auto then show ?thesis by (simp add: p_residue_alt_def) qed then show ?thesis using P1 by linarith qed qed qed lemma(in padic_integers) Zp_int_inc_res: fixes n::int shows "([n] \ \) k = n mod (p^k)" using Zp_int_inc_rep p_residue_def by (simp add: p_residue_alt_def) abbreviation(in padic_integers)(input) \

where "\

\ [p] \ \" lemma(in padic_integers) p_natpow_prod: "\

[^](n::nat) \ \

[^](m::nat) = \

[^](n + m)" by (simp add: R.nat_pow_mult) lemma(in padic_integers) p_natintpow_prod: assumes "(m::int) \ 0" shows "\

[^](n::nat) \ \

[^]m = \

[^](n + m)" using p_natpow_prod[of n "nat m"] assms int_pow_def[of Zp \

m] int_pow_def[of Zp \

"n + m"] by (metis (no_types, lifting) int_nat_eq int_pow_int of_nat_add) lemma(in padic_integers) p_intnatpow_prod: assumes "(n::int) \ 0" shows "\

[^]n \ \

[^](m::nat) = \

[^](m + n)" using p_natintpow_prod[of n m] assms mult_comm[of "\

[^]n" "\

[^]m"] by simp lemma(in padic_integers) p_int_pow_prod: assumes "(n::int) \ 0" assumes "(m::int) \ 0" shows "\

[^]n \ \

[^]m = \

[^](m + n)" proof- have "nat n + nat m = nat (n + m)" using assms by (simp add: nat_add_distrib) then have "\

[^] (nat n + nat m) = \

[^](n + m)" using assms by (simp add: \nat n + nat m = nat (n + m)\) then show ?thesis using assms p_natpow_prod[of "nat n" "nat m"] by (smt pow_nat) qed lemma(in padic_integers) p_natpow_prod_Suc: "\

\ \

[^](m::nat) = \

[^](Suc m)" "\

[^](m::nat) \ \

= \

[^](Suc m)" using R.nat_pow_Suc2 R.nat_pow_Suc by auto lemma(in padic_integers) power_residue: assumes "a \ carrier Zp" assumes "k > 0" shows "(a[^]\<^bsub>Zp\<^esub> (n::nat)) k = (a k)^n mod (p^k)" apply(induction n) using p_residues assms(2) residue_of_one(1) residues.one_cong apply auto[1] by (simp add: assms(1) mod_mult_left_eq power_commutes residue_of_prod') (*************************************************************************************************) (*************************************************************************************************) (*****************************) section\The Valuation on $\mathbb{Z}_p$\ (**********************************) (*************************************************************************************************) (*************************************************************************************************) (**********************************************************************) (**********************************************************************) subsection\The Integer-Valued and Extended Integer-Valued Valuations\ (**********************************************************************) (**********************************************************************) fun fromeint :: "eint \ int" where "fromeint (eint x) = x" text\The extended-integer-valued $p$-adic valuation on $\mathbb{Z}_p$:\ definition(in padic_integers) val_Zp where "val_Zp = (\x. (if (x = \) then (\::eint) else (eint (padic_val p x))))" text\We also define an integer-valued valuation on the nonzero elements of $\mathbb{Z}_p$, for simplified reasoning\ definition(in padic_integers) ord_Zp where "ord_Zp = padic_val p" text\Ord of additive inverse\ lemma(in padic_integers) ord_Zp_of_a_inv: assumes "a \ nonzero Zp" shows "ord_Zp a = ord_Zp (\a)" using ord_Zp_def Zp_def assms padic_val_a_inv prime by (simp add: domain.nonzero_memE(1) padic_int_is_domain) lemma(in padic_integers) val_Zp_of_a_inv: assumes "a \ carrier Zp" shows "val_Zp a = val_Zp (\a)" using R.add.inv_eq_1_iff Zp_def assms padic_val_a_inv prime val_Zp_def by auto text\Ord-based criterion for being nonzero:\ lemma(in padic_integers) ord_of_nonzero: assumes "x \carrier Zp" assumes "ord_Zp x \0" shows "x \ \" "x \ nonzero Zp" proof- show "x \ \" proof assume "x = \" then have "ord_Zp x = -1" using ord_Zp_def padic_val_def Zp_def Zp_defs(2) by auto then show False using assms(2) by auto qed then show "x \ nonzero Zp" using nonzero_def assms(1) by (simp add: nonzero_def) qed lemma(in padic_integers) not_nonzero_Zp: assumes "x \ carrier Zp" assumes "x \ nonzero Zp" shows "x = \" using assms(1) assms(2) nonzero_def by fastforce lemma(in padic_integers) not_nonzero_Qp: assumes "x \ carrier Q\<^sub>p" assumes "x \ nonzero Q\<^sub>p" shows "x = \\<^bsub>Q\<^sub>p\<^esub>" using assms(1) assms(2) nonzero_def by force text\Relationship between val and ord\ lemma(in padic_integers) val_ord_Zp: assumes "a \ \" shows "val_Zp a = eint (ord_Zp a)" by (simp add: assms ord_Zp_def val_Zp_def) lemma(in padic_integers) ord_pos: assumes "x \ carrier Zp" assumes "x \ \" shows "ord_Zp x \ 0" proof- have "x \padic_zero p" using Zp_def assms(2) Zp_defs(2) by auto then have "ord_Zp x = int (LEAST k. x (Suc k) \ \\<^bsub>residue_ring (p^Suc k)\<^esub>)" using ord_Zp_def padic_val_def by auto then show ?thesis by linarith qed lemma(in padic_integers) val_pos: assumes "x \ carrier Zp" shows "val_Zp x \ 0" unfolding val_Zp_def using assms by (metis (full_types) eint_0 eint_ord_simps(1) eint_ord_simps(3) ord_Zp_def ord_pos) text\For passing between nat and int castings of ord\ lemma(in padic_integers) ord_nat: assumes "x \ carrier Zp" assumes "x \\" shows "int (nat (ord_Zp x)) = ord_Zp x" using ord_pos by (simp add: assms(1) assms(2)) lemma(in padic_integers) zero_below_ord: assumes "x \ carrier Zp" assumes "n \ ord_Zp x" shows "x n = 0" proof- have "x n = \\<^bsub>residue_ring (p^n)\<^esub>" using ord_Zp_def zero_below_val Zp_def assms(1) assms(2) prime padic_int_simps(5) by auto then show ?thesis using residue_ring_def by simp qed lemma(in padic_integers) zero_below_val_Zp: assumes "x \ carrier Zp" assumes "n \ val_Zp x" shows "x n = 0" by (metis assms(1) assms(2) eint_ord_simps(1) ord_Zp_def residue_of_zero(2) val_Zp_def zero_below_ord) lemma(in padic_integers) below_ord_zero: assumes "x \ carrier Zp" assumes "x (Suc n) \ 0" shows "n \ ord_Zp x" proof- have 0: "x \ padic_set p" using Zp_def assms(1) Zp_defs(3) by auto have 1: "x (Suc n) \ \\<^bsub>residue_ring (p^(Suc n))\<^esub>" using residue_ring_def assms(2) by auto have "of_nat n \ (padic_val p x )" using 0 1 below_val_zero prime by auto then show ?thesis using ord_Zp_def by auto qed lemma(in padic_integers) below_val_Zp_zero: assumes "x \ carrier Zp" assumes "x (Suc n) \ 0" shows "n \ val_Zp x" by (metis Zp_def assms(1) assms(2) eint_ord_simps(1) padic_integers.below_ord_zero padic_integers.residue_of_zero(2) padic_integers.val_ord_Zp padic_integers_axioms) lemma(in padic_integers) nonzero_imp_ex_nonzero_res: assumes "x \ carrier Zp" assumes "x \ \" shows "\k. x (Suc k) \ 0" proof- have 0: "x 0 = 0" using Zp_def assms(1) padic_int_simps(5) padic_set_zero_res prime by auto have "\k. k > 0 \ x k \ 0" apply(rule ccontr) using 0 Zp_defs unfolding padic_zero_def by (metis assms(2) ext neq0_conv) then show ?thesis using not0_implies_Suc by blast qed lemma(in padic_integers) ord_suc_nonzero: assumes "x \ carrier Zp" assumes "x \\" assumes "ord_Zp x = n" shows "x (Suc n) \ 0" proof- obtain k where k_def: "x (Suc k) \ 0" using assms(1) nonzero_imp_ex_nonzero_res assms(2) by blast then show ?thesis using assms LeastI nonzero_imp_ex_nonzero_res unfolding ord_Zp_def padic_val_def by (metis (mono_tags, lifting) Zp_defs(2) k_def of_nat_eq_iff padic_zero_def padic_zero_simp(1)) qed lemma(in padic_integers) above_ord_nonzero: assumes "x \ carrier Zp" assumes "x \\" assumes "n > ord_Zp x" shows "x n \ 0" proof- have P0: "n \ (Suc (nat (ord_Zp x)))" by (simp add: Suc_le_eq assms(1) assms(2) assms(3) nat_less_iff ord_pos) then have P1: "p_residue (Suc (nat (ord_Zp x))) (x n) = x (Suc (nat (ord_Zp x)))" using assms(1) p_residue_padic_int by blast then have P2: "p_residue (Suc (nat (ord_Zp x))) (x n) \ 0" using Zp_def assms(1) assms(2) ord_nat padic_integers.ord_suc_nonzero padic_integers_axioms by auto then show ?thesis using P0 P1 assms(1) p_residue_padic_int[of x "(Suc (nat (ord_Zp x)))" n] p_residue_def by (metis ord_Zp_def padic_int_simps(2) padic_integers.zero_rep padic_integers_axioms padic_zero_simp(2)) qed lemma(in padic_integers) ord_Zp_geq: assumes "x \ carrier Zp" assumes "x n = 0" assumes "x \\" shows "ord_Zp x \ n" proof(rule ccontr) assume "\ int n \ ord_Zp x" then show False using assms using above_ord_nonzero by auto qed lemma(in padic_integers) ord_equals: assumes "x \ carrier Zp" assumes "x (Suc n) \ 0" assumes "x n = 0" shows "ord_Zp x = n" using assms(1) assms(2) assms(3) below_ord_zero ord_Zp_geq residue_of_zero(2) by fastforce lemma(in padic_integers) ord_Zp_p: "ord_Zp \

= (1::int)" proof- have "ord_Zp \

= int 1" apply(rule ord_equals[of \

]) using Zp_int_inc_res[of p] prime_gt_1_int prime by auto then show ?thesis by simp qed lemma(in padic_integers) ord_Zp_one: "ord_Zp \ = 0" proof- have "ord_Zp ([(1::int)]\\) = int 0" apply(rule ord_equals) using Zp_int_inc_res[of 1] prime_gt_1_int prime by auto then show ?thesis by simp qed text\ord is multiplicative on nonzero elements of Zp\ lemma(in padic_integers) ord_Zp_mult: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" shows "(ord_Zp (x \\<^bsub>Zp\<^esub> y)) = (ord_Zp x) + (ord_Zp y)" using val_prod[of p x y] prime assms Zp_defs Zp_def nonzero_memE(2) ord_Zp_def nonzero_closed nonzero_memE(2) by auto lemma(in padic_integers) ord_Zp_pow: assumes "x \ nonzero Zp" shows "ord_Zp (x[^](n::nat)) = n*(ord_Zp x)" proof(induction n) case 0 have "x[^](0::nat) = \" using assms(1) nonzero_def by simp then show ?case by (simp add: ord_Zp_one) next case (Suc n) fix n assume IH: "ord_Zp (x [^] n) = int n * ord_Zp x " have N: "(x [^] n) \ nonzero Zp" proof- have "ord_Zp x \ 0" using assms by (simp add: nonzero_closed nonzero_memE(2) ord_pos) then have "ord_Zp (x [^] n) \ 0" using IH assms by simp then have 0: "(x [^] n) \ \" using ord_of_nonzero(1) by force have 1: "(x [^] n) \ carrier Zp" by (simp add: nonzero_closed assms) then show ?thesis using "0" not_nonzero_Zp by blast qed have "x[^](Suc n) = x \(x[^]n)" using nonzero_closed assms R.nat_pow_Suc2 by blast then have "ord_Zp (x[^](Suc n)) =(ord_Zp x) + ord_Zp (x[^]n)" using N Zp_def assms padic_integers.ord_Zp_mult padic_integers_axioms by auto then have "ord_Zp (x[^](Suc n)) =(ord_Zp x) +(int n * ord_Zp x)" by (simp add: IH) then have "ord_Zp (x[^](Suc n)) =(1*(ord_Zp x)) +(int n) * (ord_Zp x)" by simp then have "ord_Zp (x[^](Suc n)) =(1+ (int n)) * ord_Zp x" by (simp add: comm_semiring_class.distrib) then show "ord_Zp (x[^](Suc n)) = int (Suc n)*ord_Zp x" by simp qed lemma(in padic_integers) val_Zp_pow: assumes "x \ nonzero Zp" shows "val_Zp (x[^](n::nat)) = (n*(ord_Zp x))" using Zp_def domain.nat_pow_nonzero[of Zp] domain_axioms nonzero_memE assms ord_Zp_def padic_integers.ord_Zp_pow padic_integers_axioms val_Zp_def nonzero_memE(2) by fastforce lemma(in padic_integers) val_Zp_pow': assumes "x \ nonzero Zp" shows "val_Zp (x[^](n::nat)) = n*(val_Zp x)" by (metis Zp_def assms not_nonzero_memI padic_integers.val_Zp_pow padic_integers.val_ord_Zp padic_integers_axioms times_eint_simps(1)) lemma(in padic_integers) ord_Zp_p_pow: "ord_Zp (\

[^](n::nat)) = n" using ord_Zp_pow ord_Zp_p Zp_def Zp_nat_inc_closed ord_of_nonzero(2) padic_integers_axioms int_inc_closed Zp_int_inc_closed by auto lemma(in padic_integers) ord_Zp_p_int_pow: assumes "n \0" shows "ord_Zp (\

[^](n::int)) = n" by (metis assms int_nat_eq int_pow_int ord_Zp_def ord_Zp_p_pow) lemma(in padic_integers) val_Zp_p: "(val_Zp \

) = 1" using Zp_def ord_Zp_def padic_val_def val_Zp_def ord_Zp_p Zp_defs(2) one_eint_def by auto lemma(in padic_integers) val_Zp_p_pow: "val_Zp (\

[^](n::nat)) = eint n" proof- have "(\

[^](n::nat)) \ \" by (metis mult_zero_l n_not_Suc_n of_nat_eq_iff ord_Zp_def ord_Zp_p_pow p_natpow_prod_Suc(1)) then show ?thesis using ord_Zp_p_pow by (simp add: ord_Zp_def val_Zp_def) qed lemma(in padic_integers) p_pow_res: assumes "(n::nat) \ m" shows "(\

[^]n) m = 0" by (simp add: assms ord_Zp_p_pow zero_below_ord) lemma(in padic_integers) p_pow_factor: assumes "(n::nat) \ m" shows "(h \ (\

[^]n)) m = 0" "(h \ (\

[^]n)) m = \\<^bsub>Zp_res_ring n\<^esub>" using assms p_pow_res p_res_ring_zero by(auto simp: residue_of_zero Zp_residue_mult_zero(2)) (**********************************************************************) (**********************************************************************) subsection\The Ultrametric Inequality\ (**********************************************************************) (**********************************************************************) text\Ultrametric inequality for ord\ lemma(in padic_integers) ord_Zp_ultrametric: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "x \ y \ nonzero Zp" shows "ord_Zp (x \ y) \ min (ord_Zp x) (ord_Zp y)" unfolding ord_Zp_def using padic_val_ultrametric[of p x y] Zp_defs assms nonzero_memE Zp_def prime nonzero_closed nonzero_memE(2) by auto text\Variants of the ultrametric inequality\ lemma (in padic_integers) ord_Zp_ultrametric_diff: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "x \ y " shows "ord_Zp (x \ y) \ min (ord_Zp x) (ord_Zp y)" using assms ord_Zp_ultrametric[of x "\ y"] unfolding a_minus_def by (metis (no_types, lifting) R.a_transpose_inv R.add.inv_closed R.add.m_closed R.l_neg nonzero_closed ord_Zp_of_a_inv ord_of_nonzero(2) ord_pos) lemma(in padic_integers) ord_Zp_not_equal_imp_notequal: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "ord_Zp x \ (ord_Zp y)" shows "x \ y" "x \ y \\" "x \ y \\" using assms apply blast using nonzero_closed assms(1) assms(2) assms(3) apply auto[1] using nonzero_memE assms using R.minus_equality nonzero_closed Zp_def padic_integers.ord_Zp_of_a_inv padic_integers_axioms by auto lemma(in padic_integers) ord_Zp_ultrametric_eq: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "ord_Zp x > (ord_Zp y)" shows "ord_Zp (x \ y) = ord_Zp y" proof- have 0: "ord_Zp (x \ y) \ ord_Zp y" using assms ord_Zp_not_equal_imp_notequal[of x y] ord_Zp_ultrametric[of x y] nonzero_memE not_nonzero_Zp nonzero_closed by force have 1: "ord_Zp y \ min (ord_Zp(x \ y)) (ord_Zp x)" proof- have 0: "x \ y \ x" using assms nonzero_memE by (simp add: nonzero_closed nonzero_memE(2)) have 1: "x \ y \ nonzero Zp" using ord_Zp_not_equal_imp_notequal[of x y] nonzero_closed assms(1) assms(2) assms(3) not_nonzero_Zp by force then show ?thesis using 0 assms(1) assms(2) assms(3) ord_Zp_ultrametric_diff[of "x \ y" x] by (simp add: R.minus_eq nonzero_closed R.r_neg1 add_comm) qed then show ?thesis using 0 assms(3) by linarith qed lemma(in padic_integers) ord_Zp_ultrametric_eq': assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "ord_Zp x > (ord_Zp y)" shows "ord_Zp (x \ y) = ord_Zp y" using assms ord_Zp_ultrametric_eq[of x "\ y"] unfolding a_minus_def by (metis R.add.inv_closed R.add.inv_eq_1_iff nonzero_closed not_nonzero_Zp ord_Zp_of_a_inv) lemma(in padic_integers) ord_Zp_ultrametric_eq'': assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "ord_Zp x > (ord_Zp y)" shows "ord_Zp (y \ x) = ord_Zp y" by (metis R.add.inv_closed R.minus_eq nonzero_closed Zp_def add_comm assms(1) assms(2) assms(3) ord_Zp_of_a_inv ord_of_nonzero(2) ord_pos padic_integers.ord_Zp_ultrametric_eq padic_integers_axioms) lemma(in padic_integers) ord_Zp_not_equal_ord_plus_minus: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" assumes "ord_Zp x \ (ord_Zp y)" shows "ord_Zp (x \ y) = ord_Zp (x \ y)" apply(cases "ord_Zp x > ord_Zp y") using assms apply (simp add: ord_Zp_ultrametric_eq ord_Zp_ultrametric_eq') using assms nonzero_memI by (smt add_comm ord_Zp_ultrametric_eq ord_Zp_ultrametric_eq'') text\val is multiplicative on nonzero elements\ lemma(in padic_integers) val_Zp_mult0: assumes "x \ carrier Zp" assumes "x \\" assumes "y \ carrier Zp" assumes "y \\" shows "(val_Zp (x \\<^bsub>Zp\<^esub> y)) = (val_Zp x) + (val_Zp y)" apply(cases "x \\<^bsub>Zp\<^esub> y = \") using assms(1) assms(2) assms(3) assms(4) integral_iff val_ord_Zp ord_Zp_mult nonzero_memI apply (simp add: integral_iff) using assms ord_Zp_mult[of x y] val_ord_Zp by (simp add: nonzero_memI) text\val is multiplicative everywhere\ lemma(in padic_integers) val_Zp_mult: assumes "x \ carrier Zp" assumes "y \ carrier Zp" shows "(val_Zp (x \\<^bsub>Zp\<^esub> y)) = (val_Zp x) + (val_Zp y)" using assms(1) assms(2) integral_iff val_ord_Zp ord_Zp_mult nonzero_memI val_Zp_mult0 val_Zp_def by simp lemma(in padic_integers) val_Zp_ultrametric0: assumes "x \ carrier Zp" assumes "x \\" assumes "y \ carrier Zp" assumes "y \\" assumes "x \ y \ \" shows "min (val_Zp x) (val_Zp y) \ val_Zp (x \ y) " apply(cases "x \ y = \") using assms apply blast using assms ord_Zp_ultrametric[of x y] nonzero_memI val_ord_Zp[of x] val_ord_Zp[of y] val_ord_Zp[of "x \ y"] by simp text\Unconstrained ultrametric inequality\ lemma(in padic_integers) val_Zp_ultrametric: assumes "x \ carrier Zp" assumes "y \ carrier Zp" shows " min (val_Zp x) (val_Zp y) \ val_Zp (x \ y)" apply(cases "x = \") apply (simp add: assms(2)) apply(cases "y = \") apply (simp add: assms(1)) apply(cases "x \ y = \") apply (simp add: val_Zp_def) using assms val_Zp_ultrametric0[of x y] by simp text\Variants of the ultrametric inequality\ lemma (in padic_integers) val_Zp_ultrametric_diff: assumes "x \ carrier Zp" assumes "y \ carrier Zp" shows "val_Zp (x \ y) \ min (val_Zp x) (val_Zp y)" using assms val_Zp_ultrametric[of x "\y"] unfolding a_minus_def by (metis R.add.inv_closed R.add.inv_eq_1_iff nonzero_memI ord_Zp_def ord_Zp_of_a_inv val_Zp_def) lemma(in padic_integers) val_Zp_not_equal_imp_notequal: assumes "x \ carrier Zp" assumes "y \ carrier Zp" assumes "val_Zp x \ val_Zp y" shows "x \ y" "x \ y \\" "x \ y \\" using assms(3) apply auto[1] using assms(1) assms(2) assms(3) R.r_right_minus_eq apply blast by (metis R.add.inv_eq_1_iff assms(1) assms(2) assms(3) R.minus_zero R.minus_equality not_nonzero_Zp ord_Zp_def ord_Zp_of_a_inv val_ord_Zp) lemma(in padic_integers) val_Zp_ultrametric_eq: assumes "x \ carrier Zp" assumes "y \ carrier Zp" assumes "val_Zp x > val_Zp y" shows "val_Zp (x \ y) = val_Zp y" apply(cases "x \ \ \ y \ \ \ x \ y") using assms ord_Zp_ultrametric_eq[of x y] val_ord_Zp nonzero_memE using not_nonzero_memE val_Zp_not_equal_imp_notequal(3) apply force unfolding val_Zp_def using assms(2) assms(3) val_Zp_def by force lemma(in padic_integers) val_Zp_ultrametric_eq': assumes "x \ carrier Zp" assumes "y \ carrier Zp" assumes "val_Zp x > (val_Zp y)" shows "val_Zp (x \ y) = val_Zp y" using assms val_Zp_ultrametric_eq[of x "\ y"] unfolding a_minus_def by (metis R.add.inv_closed R.r_neg val_Zp_not_equal_imp_notequal(3)) lemma(in padic_integers) val_Zp_ultrametric_eq'': assumes "x \ carrier Zp" assumes "y \ carrier Zp" assumes "val_Zp x > (val_Zp y)" shows "val_Zp (y \ x) = val_Zp y" proof- have 0: "y \ x = \ (x \ y)" using assms(1,2) unfolding a_minus_def by (simp add: R.add.m_comm R.minus_add) have 1: "val_Zp (x \ y) = val_Zp y" using assms val_Zp_ultrametric_eq' by blast have 2: "val_Zp (x \ y) = val_Zp (y \ x)" unfolding 0 unfolding a_minus_def by(rule val_Zp_of_a_inv, rule R.ring_simprules, rule assms, rule R.ring_simprules, rule assms) show ?thesis using 1 unfolding 2 by blast qed lemma(in padic_integers) val_Zp_not_equal_ord_plus_minus: assumes "x \ carrier Zp" assumes "y \ carrier Zp" assumes "val_Zp x \ (val_Zp y)" shows "val_Zp (x \ y) = val_Zp (x \ y)" by (metis R.add.inv_closed R.minus_eq R.r_neg R.r_zero add_comm assms(1) assms(2) assms(3) not_nonzero_Zp ord_Zp_def ord_Zp_not_equal_ord_plus_minus val_Zp_def val_Zp_not_equal_imp_notequal(3)) (**********************************************************************) (**********************************************************************) subsection\Units of $\mathbb{Z}_p$\ (**********************************************************************) (**********************************************************************) text\Elements with valuation 0 in Zp are the units\ lemma(in padic_integers) val_Zp_0_criterion: assumes "x \ carrier Zp" assumes "x 1 \ 0" shows "val_Zp x = 0" unfolding val_Zp_def using Zp_def assms(1) assms(2) ord_equals padic_set_zero_res prime by (metis One_nat_def Zp_defs(3) of_nat_0 ord_Zp_def residue_of_zero(2) zero_eint_def) text\Units in Zp have val 0\ lemma(in padic_integers) unit_imp_val_Zp0: assumes "x \ Units Zp" shows "val_Zp x = 0" apply(rule val_Zp_0_criterion) apply (simp add: R.Units_closed assms) using assms residue_of_prod[of x "inv x" 1] residue_of_one(2)[of 1] R.Units_r_inv[of x] comm_monoid.UnitsI[of "R 1"] p_res_ring_1_field by (metis le_eq_less_or_eq residue_of_prod residue_times_zero_r zero_le_one zero_neq_one) text\Elements in Zp with ord 0 are units\ lemma(in padic_integers) val_Zp0_imp_unit0: assumes "val_Zp x = 0" assumes "x \ carrier Zp" fixes n::nat shows "(x (Suc n)) \ Units (Zp_res_ring (Suc n))" unfolding val_Zp_def proof- have p_res_ring: "residues (p^(Suc n))" using p_residues by blast have "\ n. coprime (x (Suc n)) p" proof- fix n show "coprime (x (Suc n)) p" proof- have "\ \ coprime (x (Suc n)) p" proof assume "\ coprime (x (Suc n)) p" then have "p dvd (x (Suc n))" using prime by (meson coprime_commute prime_imp_coprime prime_nat_int_transfer) then obtain k where "(x (Suc n)) = k*p" by fastforce then have S:"x (Suc n) mod p = 0" by simp have "x 1 = 0" proof- have "Suc n \ 1" by simp then have "x 1 = p_residue 1 (x (Suc n))" using p_residue_padic_int assms(2) by presburger then show ?thesis using S by (simp add: p_residue_alt_def) qed have "x \\" proof- have "ord_Zp x \ ord_Zp \" using Zp_def ord_Zp_def padic_val_def assms(1) ord_of_nonzero(1) R.zero_closed Zp_defs(2) val_Zp_def by auto then show ?thesis by blast qed then have "x 1 \ 0" using assms(1) assms(2) ord_suc_nonzero unfolding val_Zp_def by (simp add: ord_Zp_def zero_eint_def) then show False using \x 1 = 0\ by blast qed then show ?thesis by auto qed qed then have "\ n. coprime (x (Suc n)) (p^(Suc n))" by simp then have "coprime (x (Suc n)) (p^(Suc n))" by blast then show ?thesis using assms residues.res_units_eq p_res_ring by (metis (no_types, lifting) mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2) residues.m_gt_one residues.mod_in_res_units residues_closed) qed lemma(in padic_integers) val_Zp0_imp_unit0': assumes "val_Zp x = 0" assumes "x \ carrier Zp" assumes "(n::nat) > 0" shows "(x n) \ Units (Zp_res_ring n)" using assms val_Zp0_imp_unit0 gr0_implies_Suc by blast lemma(in cring) ring_hom_Units_inv: assumes "a \ Units R" assumes "cring S" assumes "h \ ring_hom R S" shows "h (inv a) = inv\<^bsub>S\<^esub> h a" "h a \ Units S" proof- have 0:"h (inv a) \\<^bsub>S\<^esub> h a = \\<^bsub>S\<^esub>" using assms Units_closed Units_inv_closed by (metis (no_types, lifting) Units_l_inv ring_hom_mult ring_hom_one) then show 1: "h (inv a) = inv\<^bsub>S\<^esub> h a" by (metis Units_closed Units_inv_closed assms(1) assms(2) assms(3) comm_monoid.is_invI(1) cring_def ring_hom_closed) show "h a \ Units S" apply(rule comm_monoid.UnitsI[of S "inv\<^bsub>S\<^esub> h a"]) using 0 1 assms using cring.axioms(2) apply blast apply (metis "1" Units_inv_closed assms(1) assms(3) ring_hom_closed) apply (meson Units_closed assms(1) assms(3) ring_hom_closed) using "0" "1" by auto qed lemma(in padic_integers) val_Zp_0_imp_unit: assumes "val_Zp x = 0" assumes "x \ carrier Zp" shows "x \ Units Zp" proof- obtain y where y_def: " y= (\n. (if n=0 then 0 else (m_inv (Zp_res_ring n) (x n))))" by blast have 0: "\m. m > 0 \ y m = inv \<^bsub>Zp_res_ring m\<^esub> (x m)" using y_def by auto have 1: "\m. m > 0 \ inv\<^bsub>Zp_res_ring m\<^esub> (x m) \ carrier (Zp_res_ring m)" proof- fix m::nat assume A: "m > 0" then show "inv\<^bsub>Zp_res_ring m\<^esub> (x m) \ carrier (Zp_res_ring m)" using assms val_Zp0_imp_unit0' monoid.Units_inv_closed[of "Zp_res_ring m" "x m"] by (smt One_nat_def Zp_def Zp_defs(2) cring.axioms(1) of_nat_0 ord_Zp_def padic_integers.R_cring padic_integers.ord_suc_nonzero padic_integers.val_Zp_0_criterion padic_integers_axioms padic_val_def ring_def) qed have 2: "y \ padic_set p" proof(rule padic_set_memI) show 20: "\m. y m \ carrier (residue_ring (p ^ m))" proof- fix m show "y m \ carrier (residue_ring (p ^ m))" apply(cases "m = 0") using y_def 0[of m] 1[of m] by(auto simp: residue_ring_def y_def) qed show "\m n. m < n \ residue (p ^ m) (y n) = y m" proof- fix m n::nat assume A: "m < n" show "residue (p ^ m) (y n) = y m" proof(cases "m = 0") case True then show ?thesis by (simp add: residue_1_zero y_def) next case False have hom: "residue (p ^ m) \ ring_hom (Zp_res_ring n) (Zp_res_ring m)" using A False prime residue_hom_p by auto have inv: "y n = inv\<^bsub>Zp_res_ring n\<^esub> x n" using A by (simp add: False y_def) have unit: "x n \ Units (Zp_res_ring n)" using A False Zp_def assms(1) assms(2) val_Zp0_imp_unit0' prime by (metis gr0I gr_implies_not0) have F0: "residue (p ^ m) (x n) = x m" using A Zp_defs(3) assms(2) padic_set_res_coherent prime by auto have F1: "residue (p ^ m) (y n) = inv\<^bsub>Zp_res_ring m\<^esub> x m" using F0 R_cring A hom inv unit cring.ring_hom_Units_inv[of "Zp_res_ring n" "x n" "Zp_res_ring m" "residue (p ^ m)"] False by auto then show ?thesis by (simp add: False y_def) qed qed qed have 3: "y \ x = \" proof fix m show "(y \ x) m = \ m" proof(cases "m=0") case True then have L: "(y \ x) m = 0" using Zp_def "1" assms(2) Zp_residue_mult_zero(1) y_def by auto have R: "\ m = 0" by (simp add: True cring.cring_simprules(6) domain.axioms(1) ord_Zp_one zero_below_ord) then show ?thesis using L R by auto next case False have P: "(y \ x) m = (y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m)" using Zp_def residue_of_prod by auto have "(y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m) = 1" proof- have "p^m > 1" using False prime prime_gt_1_int by auto then have "residues (p^m)" using less_imp_of_nat_less residues.intro by fastforce have "cring (residue_ring (p^m))" using residues.cring \residues (p ^ m)\ by blast then have M: "monoid (residue_ring (p^m))" using cring_def ring_def by blast have U: "(x m) \ Units (residue_ring (p^m))" using False Zp_def assms(1) assms(2) padic_integers.val_Zp0_imp_unit0' padic_integers_axioms by auto have I: "y m = m_inv (residue_ring (p^m)) (x m)" by (simp add: False y_def) have "(y m) \\<^bsub>residue_ring (p^m)\<^esub> (x m) = \\<^bsub>residue_ring (p^m)\<^esub>" using M U I by (simp add: monoid.Units_l_inv) then show ?thesis using residue_ring_def by simp qed then show ?thesis using P Zp_def False residue_of_one(2) by auto qed qed have 4: "y \ carrier Zp" using 2 Zp_defs by auto show ?thesis apply(rule R.UnitsI[of y]) using assms 4 3 by auto qed text\Definition of ord on a fraction is independent of the choice of representatives\ lemma(in padic_integers) ord_Zp_eq_frac: assumes "a \ nonzero Zp" assumes "b \ nonzero Zp" assumes "c \ nonzero Zp" assumes "d \ nonzero Zp" assumes "a \ d = b \ c" shows "(ord_Zp a) - (ord_Zp b) = (ord_Zp c) - (ord_Zp d)" proof- have "ord_Zp (a \ d) = ord_Zp (b \ c)" using assms by presburger then have "(ord_Zp a) + (ord_Zp d) = (ord_Zp b) + (ord_Zp c)" using assms(1) assms(2) assms(3) assms(4) ord_Zp_mult by metis then show ?thesis by linarith qed lemma(in padic_integers) val_Zp_eq_frac_0: assumes "a \ nonzero Zp" assumes "b \ nonzero Zp" assumes "c \ nonzero Zp" assumes "d \ nonzero Zp" assumes "a \ d = b \ c" shows "(val_Zp a) - (val_Zp b) = (val_Zp c) - (val_Zp d)" proof- have 0:"(val_Zp a) - (val_Zp b) = (ord_Zp a) - (ord_Zp b)" using assms nonzero_memE Zp_defs(2) ord_Zp_def val_Zp_def by auto have 1: "(val_Zp c) - (val_Zp d) = (ord_Zp c) - (ord_Zp d)" using assms nonzero_memE val_ord_Zp[of c] val_ord_Zp[of d] by (simp add: nonzero_memE(2)) then show ?thesis using "0" assms(1) assms(2) assms(3) assms(4) assms(5) ord_Zp_eq_frac by presburger qed (*************************************************************************************************) (*************************************************************************************************) (*****************) section\Angular Component Maps on $\mathbb{Z}_p$\ (*********************) (*************************************************************************************************) (*************************************************************************************************) text\The angular component map on $\mathbb{Z}_p$ is just the map which normalizes a point $x \in \mathbb{Z}_p$ by mapping it to a point with valuation $0$. It is explicitly defined as the mapping $x \mapsto p^{-ord (p)}*x$ for nonzero $x$, and $0 \mapsto 0$. By composing these maps with reductions mod $p^n$ we get maps which are equal to the standard residue maps on units of $\mathbb{Z}_p$, but in general unequal elsewhere. Both the angular component map and the angular component map mod $p^n$ are homomorpshims from the multiplicative group of units of $\mathbb{Z}_p$ to the multiplicative group of units of the residue rings, and play a key role in first-order model-theoretic formalizations of the $p$-adics (see, for example \<^cite>\"10.2307/2274477"\, or \<^cite>\"Denef1986"\). \ lemma(in cring) int_nat_pow_rep: "[(k::int)]\\ [^] (n::nat) = [(k^n)]\\" apply(induction n) by (auto simp: add.int_pow_pow add_pow_rdistr_int mult.commute) lemma(in padic_integers) p_pow_rep0: fixes n::nat shows "\

[^]n = [(p^n)]\\" using R.int_nat_pow_rep by auto lemma(in padic_integers) p_pow_nonzero: shows "(\

[^](n::nat)) \ carrier Zp" "(\

[^](n::nat)) \ \" apply simp using Zp_def nat_pow_nonzero domain_axioms nonzero_memE int_inc_closed ord_Zp_p padic_integers.ord_of_nonzero(2) padic_integers_axioms Zp_int_inc_closed nonzero_memE(2) by (metis ord_of_nonzero(2) zero_le_one) lemma(in padic_integers) p_pow_nonzero': shows "(\

[^](n::nat)) \ nonzero Zp" using nonzero_def p_pow_nonzero by (simp add: nonzero_def) lemma(in padic_integers) p_pow_rep: fixes n::nat shows "(\

[^]n) k = (p^n) mod (p^k)" by (simp add: R.int_nat_pow_rep Zp_int_inc_res) lemma(in padic_integers) p_pow_car: assumes " (n::int)\ 0" shows "(\

[^]n) \ carrier Zp" proof- have "(\

[^]n) = (\

[^](nat n))" by (metis assms int_nat_eq int_pow_int) then show ?thesis by simp qed lemma(in padic_integers) p_int_pow_nonzero: assumes "(n::int) \0" shows "(\

[^]n) \ nonzero Zp" by (metis assms not_nonzero_Zp ord_Zp_p_int_pow ord_of_nonzero(1) p_pow_car) lemma(in padic_integers) p_nonzero: shows "\

\ nonzero Zp" using p_int_pow_nonzero[of 1] by (simp add: ord_Zp_p ord_of_nonzero(2)) text\Every element of Zp is a unit times a power of p.\ lemma(in padic_integers) residue_factor_unique: assumes "k>0" assumes "x \ carrier Zp" assumes "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" shows "u = (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k)))" proof- obtain P where P_def: "P = (\ u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k)))" by simp have 0: "P u" using P_def assms(3) by blast have 1: "\ v. P v \ v = u" by (metis P_def assms(3) mult_cancel_right not_prime_0 power_not_zero prime) have "u = (THE u. P u)" by (metis 0 1 the_equality) then show ?thesis using P_def by blast qed lemma(in padic_integers) residue_factor_exists: assumes "m = nat (ord_Zp x)" assumes "k > 0" assumes "x \ carrier Zp" assumes "x \\" obtains u where "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" proof- have X0: "(x (m+k)) \ carrier (Zp_res_ring (m+k)) " using Zp_def assms(3) padic_set_res_closed residues_closed by blast then have X1: "(x (m+k)) \ 0" using p_residues assms(2) residues.res_carrier_eq by simp then have X2: "(x (m+k)) > 0" using assms(1) assms(2) assms(3) assms(4) above_ord_nonzero by (metis add.right_neutral add_cancel_right_right add_gr_0 int_nat_eq less_add_same_cancel1 less_imp_of_nat_less not_gr_zero of_nat_0_less_iff of_nat_add ord_pos) have 0: "x m = 0" using Zp_def assms(1) assms(3) zero_below_val ord_nat zero_below_ord[of x m] assms(4) ord_Zp_def by auto then have 1: "x (m +k) mod p^m = 0" using assms(2) assms(3) p_residue_padic_int residue_def by (simp add: p_residue_alt_def) then have "\ u. u*(p^m) = (x (m+k))" by auto then obtain u where U0: " u*(p^m) = (x (m+k))" by blast have I: "(p^m) > 0 " using prime by (simp add: prime_gt_0_int) then have U1: "(u* p^m) = (x (m+k))" by (simp add: U0) have U2: "u \ 0" using I U1 X1 by (metis U0 less_imp_triv mult.right_neutral mult_less_cancel_left of_nat_zero_less_power_iff power.simps(1) times_int_code(1)) have X3: "(x (m+k)) < p^(m+k)" using assms(3) X0 p_residues assms(2) residues.res_carrier_eq by auto have U3: "u < p^k" proof(rule ccontr) assume "\ u < (p ^ k)" then have "(p^k) \ u" by simp then have " (p^k * p^m) \ u*(p^m)" using I by simp then have "p^(m + k) \ (x (m+k))" by (simp add: U0 add.commute semiring_normalization_rules(26)) then show False using X3 by linarith qed then have "u \ carrier (Zp_res_ring k)" using assms(2) p_residues residues.res_carrier_eq U3 U2 by auto then show ?thesis using U1 that by blast qed definition(in padic_integers) normalizer where "normalizer m x = (\k. if (k=0) then 0 else (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m + k)) ) )" definition(in padic_integers) ac_Zp where "ac_Zp x = normalizer (nat (ord_Zp x)) x" lemma(in padic_integers) ac_Zp_equation: assumes "x \ carrier Zp" assumes "x \\" assumes "k > 0" assumes "m = nat (ord_Zp x)" shows "(ac_Zp x k) \ carrier (Zp_res_ring k) \ (ac_Zp x k)*(p^m) = (x (m+k))" proof- have K0: "k >0" using assms nat_neq_iff by blast have KM: "m+ k > m" using assms(3) assms(4) by linarith obtain u where U0: "u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m+k))" using assms(1) assms(2) assms(3) assms(4) residue_factor_exists by blast have RHS: "ac_Zp x k = (THE u. u \ carrier (Zp_res_ring k) \ u*(p^m) = (x (m+k)))" proof- have K: "k \0" by (simp add: K0) have "ac_Zp x k = normalizer (nat (ord_Zp x)) x k" using ac_Zp_def by presburger then have "ac_Zp x k = normalizer m x k" using assms by blast then show "ac_Zp x k = (THE u. u \ carrier (Zp_res_ring k) \ (u* p^m) = (x (m + k)) ) " using K unfolding normalizer_def p_residue_def by simp qed have LHS:"u = (THE u. u \ carrier (Zp_res_ring k) \ u*(p^m) = (x (m+k)))" using assms U0 K0 assms(1) residue_factor_unique[of k x u m] by metis then have "u = ac_Zp x k" by (simp add: RHS) then show ?thesis using U0 by auto qed lemma(in padic_integers) ac_Zp_res: assumes "m >k" assumes "x \ carrier Zp" assumes "x \\" shows "p_residue k (ac_Zp x m) = (ac_Zp x k)" proof(cases "k =0") case True then show ?thesis unfolding ac_Zp_def normalizer_def by (meson p_res_ring_0' p_residue_range') next case False obtain n where n_def: "n = nat (ord_Zp x)" by blast have K0: "k >0" using False by simp obtain uk where Uk0: "uk = (ac_Zp x k)" by simp obtain um where Um0: "um = (ac_Zp x m)" by simp have Uk1: "uk \ carrier (Zp_res_ring k) \ uk*(p^n) = (x (n+k))" using K0 Uk0 ac_Zp_equation assms(2) assms(3) n_def by metis have Um1: "um \ carrier (Zp_res_ring m) \ um*(p^n) = (x (n+m))" using Uk1 Um0 ac_Zp_equation assms(1) assms(3) n_def assms(2) by (metis neq0_conv not_less0) have "um mod (p^k) = uk" proof- have "(x (n+m)) mod (p^(n + k)) = (x (n+k))" using assms(1) assms(3) p_residue_padic_int p_residue_def n_def by (simp add: assms(2) p_residue_alt_def) then have "(p^(n + k)) dvd (x (n+m)) - (x (n+k))" by (metis dvd_minus_mod) then obtain d where "(x (n+m)) - (x (n+k)) = (p^(n+k))*d" using dvd_def by blast then have "((um*(p^n)) - (uk*(p^n))) = p^(n+k)*d" using Uk1 Um1 by auto then have "((um - uk)*(p^n)) = p^(n+k)*d" by (simp add: left_diff_distrib) then have "((um - uk)*(p^n)) = ((p^k)*d)*(p^n)" by (simp add: power_add) then have "(um - uk) = ((p^k)*d)" using prime by auto then have "um mod p^k = uk mod p^k" by (simp add: mod_eq_dvd_iff) then show ?thesis using Uk1 by (metis mod_pos_pos_trivial p_residue_ring_car_memE(1) p_residue_ring_car_memE(2)) qed then show ?thesis by (simp add: Uk0 Um0 p_residue_alt_def) qed lemma(in padic_integers) ac_Zp_in_Zp: assumes "x \ carrier Zp" assumes "x \\" shows "ac_Zp x \ carrier Zp" proof- have "ac_Zp x \ padic_set p" proof(rule padic_set_memI) show "\m. ac_Zp x m \ carrier (residue_ring (p ^ m))" proof- fix m show "ac_Zp x m \ carrier (residue_ring (p ^ m))" proof(cases "m = 0") case True then have "ac_Zp x m = 0" unfolding ac_Zp_def normalizer_def by auto then show ?thesis by (simp add: True residue_ring_def) next case False then have "m>0" by blast then show ?thesis using ac_Zp_equation by (metis assms(1) assms(2)) qed qed show "\m n. m < n \ residue (p ^ m) (ac_Zp x n) = ac_Zp x m" using ac_Zp_res by (simp add: assms(1) assms(2) p_residue_def) qed then show ?thesis by (simp add: Zp_defs(3)) qed lemma(in padic_integers) ac_Zp_is_Unit: assumes "x \ carrier Zp" assumes "x \\" shows "ac_Zp x \ Units Zp" proof(rule val_Zp_0_imp_unit) show "ac_Zp x \ carrier Zp" by (simp add: ac_Zp_in_Zp assms(1) assms(2)) obtain m where M: "m = nat (ord_Zp x)" by blast have AC1: "(ac_Zp x 1)*(p^m) = (x (m+1))" using M ac_Zp_equation assms(1) assms(2) by (metis One_nat_def lessI) have "(x (m+1)) \0" using M assms by (metis Suc_eq_plus1 Suc_le_eq nat_int nat_mono nat_neq_iff ord_Zp_geq) then have "(ac_Zp x 1) \ 0" using AC1 by auto then show "val_Zp (ac_Zp x) = 0" using \ac_Zp x \ carrier Zp\ val_Zp_0_criterion by blast qed text\The typical defining equation for the angular component map.\ lemma(in padic_integers) ac_Zp_factors_x: assumes "x \ carrier Zp" assumes "x \\" shows "x = (\

[^](nat (ord_Zp x))) \ (ac_Zp x)" "x = (\

[^](ord_Zp x)) \ (ac_Zp x)" proof- show "x = (\

[^](nat (ord_Zp x)))\ (ac_Zp x)" proof fix k show "x k = ((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k" proof(cases "k=0") case True then show ?thesis using Zp_def Zp_defs(3) Zp_residue_mult_zero(1) ac_Zp_in_Zp assms(1) assms(2) mult_comm padic_set_zero_res prime by auto next case False show ?thesis proof(cases "k \ ord_Zp x") case True have 0: "x k = 0" using True assms(1) zero_below_ord by blast have 1: "(\

[^](nat (ord_Zp x))) k = 0" using True assms(1) assms(2) ord_Zp_p_pow ord_nat p_pow_nonzero(1) zero_below_ord by presburger have "((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k = (\

[^](nat (ord_Zp x))) k * (ac_Zp x) k mod p^k" using Zp_def padic_mult_res residue_ring_def using residue_of_prod' by blast then have "((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) k = 0" by (simp add: "1") then show ?thesis using 0 by metis next case False obtain n where N: "n = nat (ord_Zp x)" by metis obtain m where M0: "k = n + m" using False N le_Suc_ex ord_Zp_def by fastforce have M1: "m >0" using M0 False N assms(1) assms(2) ord_nat by (metis Nat.add_0_right gr0I le_refl less_eq_int_code(1) nat_eq_iff2 neq0_conv of_nat_eq_0_iff of_nat_mono) have E1: "(ac_Zp x m)*(p^n) = (x k)" using M0 M1 N ac_Zp_equation assms(1) assms(2) by blast have E2: "(ac_Zp x k)*(p^n) = (x (n + k))" using M0 M1 N ac_Zp_equation assms(1) assms(2) add_gr_0 by presburger have E3: "((ac_Zp x k) mod (p^k))*((p^n) mod p^k) mod (p^k) = (x (n + k)) mod (p^k)" by (metis E2 mod_mult_left_eq mod_mult_right_eq) have E4: "((ac_Zp x k) mod (p^k))*(p^n) mod (p^k) = (x k)" using E2 assms(1) le_add2 mod_mult_left_eq p_residue_padic_int p_residue_def by (metis Zp_int_inc_rep Zp_int_inc_res) have E5: "(ac_Zp x k)*((p^n) mod p^k) mod (p^k) = (x k)" using E2 assms(1) p_residue_padic_int p_residue_def by (metis E3 E4 mod_mult_left_eq) have E6: "(ac_Zp x k) \\<^bsub>(Zp_res_ring k)\<^esub> ((p^n) mod p^k) = (x k)" using E5 M0 M1 p_residues residues.res_mult_eq by auto have E7: " ((p^n) mod p^k) \\<^bsub>(Zp_res_ring k)\<^esub>(ac_Zp x k) = (x k)" by (simp add: E6 residue_mult_comm) have E8: "((\

[^](nat (ord_Zp x))) k) \\<^bsub>(Zp_res_ring k)\<^esub> (ac_Zp x k) = (x k)" using E7 N p_pow_rep by metis then show ?thesis by (simp add: residue_of_prod) qed qed qed then show "x = (\

[^](ord_Zp x)) \ (ac_Zp x)" by (metis assms(1) assms(2) int_pow_int ord_nat) qed lemma(in padic_integers) ac_Zp_factors': assumes "x \ nonzero Zp" shows "x = [p] \ \ [^] ord_Zp x \ ac_Zp x" using assms nonzero_memE by (simp add: nonzero_closed nonzero_memE(2) ac_Zp_factors_x(2)) lemma(in padic_integers) ac_Zp_mult: assumes "x \ nonzero Zp" assumes "y \ nonzero Zp" shows "ac_Zp (x \ y) = (ac_Zp x) \ (ac_Zp y)" proof- have P0: "x = (\

[^](nat (ord_Zp x))) \ (ac_Zp x)" using nonzero_memE ac_Zp_factors_x assms(1) by (simp add: nonzero_closed nonzero_memE(2)) have P1: "y = (\

[^](nat (ord_Zp y))) \ (ac_Zp y)" using nonzero_memE ac_Zp_factors_x assms(2) by (simp add: nonzero_closed nonzero_memE(2)) have "x \ y = (\

[^](nat (ord_Zp (x \ y)))) \ (ac_Zp (x \ y))" proof- have "x \ y \ nonzero Zp" by (simp add: assms(1) assms(2) nonzero_mult_closed) then show ?thesis using nonzero_closed nonzero_memE(2) Zp_def padic_integers.ac_Zp_factors_x(1) padic_integers_axioms by blast qed then have "x \ y = (\

[^](nat ((ord_Zp x) + (ord_Zp y)))) \ (ac_Zp (x \ y))" using assms ord_Zp_mult[of x y] by (simp add: Zp_def) then have "x \ y = (\

[^]((nat (ord_Zp x)) + nat (ord_Zp y))) \ (ac_Zp (x \ y))" using nonzero_closed nonzero_memE(2) assms(1) assms(2) nat_add_distrib ord_pos by auto then have "x \ y = (\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y))" using p_natpow_prod by metis then have P2: "(\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y)) = ((\

[^](nat (ord_Zp x))) \ (ac_Zp x)) \ ((\

[^](nat (ord_Zp y))) \ (ac_Zp y))" using P0 P1 by metis have "(\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y))) \ (ac_Zp (x \ y)) = ((\

[^](nat (ord_Zp x))) \ ((\

[^](nat (ord_Zp y))) \ (ac_Zp x)) \ (ac_Zp y))" by (metis P0 P1 Zp_def \x \ y = [p] \ \ [^] nat (ord_Zp x) \ [p] \ \ [^] nat (ord_Zp y) \ ac_Zp (x \ y)\ mult_comm padic_integers.mult_assoc padic_integers_axioms) then have "((\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y)))) \ (ac_Zp (x \ y)) =((\

[^](nat (ord_Zp x))) \ (\

[^](nat(ord_Zp y)))) \ ((ac_Zp x) \ (ac_Zp y))" using Zp_def mult_assoc by auto then show ?thesis by (metis (no_types, lifting) R.m_closed \x \ y = [p] \ \ [^] nat (ord_Zp x) \ [p] \ \ [^] nat (ord_Zp y) \ ac_Zp (x \ y)\ ac_Zp_in_Zp assms(1) assms(2) integral_iff m_lcancel nonzero_closed nonzero_memE(2) p_pow_nonzero(1)) qed lemma(in padic_integers) ac_Zp_one: "ac_Zp \ = \" by (metis R.one_closed Zp_def ac_Zp_factors_x(2) int_pow_0 ord_Zp_one padic_integers.ac_Zp_in_Zp padic_integers_axioms padic_one_id prime zero_not_one) lemma(in padic_integers) ac_Zp_inv: assumes "x \ Units Zp" shows "ac_Zp (inv\<^bsub>Zp\<^esub> x) = inv\<^bsub>Zp\<^esub> (ac_Zp x)" proof- have "x \ (inv\<^bsub>Zp\<^esub> x) = \" using assms by simp then have "(ac_Zp x) \ (ac_Zp (inv\<^bsub>Zp\<^esub> x)) = ac_Zp \" using ac_Zp_mult[of x "(inv x)"] R.Units_nonzero assms zero_not_one by auto then show ?thesis using R.invI(2)[of "(ac_Zp x)" "(ac_Zp (inv\<^bsub>Zp\<^esub> x))"] assms ac_Zp_in_Zp ac_Zp_one by (metis (no_types, lifting) R.Units_closed R.Units_inv_closed R.Units_r_inv integral_iff R.inv_unique ac_Zp_is_Unit) qed lemma(in padic_integers) ac_Zp_of_Unit: assumes "val_Zp x = 0" assumes "x \ carrier Zp" shows "ac_Zp x = x" using assms unfolding val_Zp_def by (metis R.one_closed Zp_def ac_Zp_factors_x(2) ac_Zp_one eint.inject infinity_ne_i0 mult_assoc ord_Zp_def ord_Zp_one padic_integers.ac_Zp_in_Zp padic_integers_axioms padic_one_id prime zero_eint_def zero_not_one) lemma(in padic_integers) ac_Zp_p: "(ac_Zp \

) = \" proof- have 0: "\

= \

[^] nat (ord_Zp \

) \ ac_Zp \

" using ac_Zp_factors_x[of \

] ord_Zp_p ord_of_nonzero(1) by auto have 1: "\

[^] nat (ord_Zp \

) = \

" by (metis One_nat_def nat_1 ord_Zp_p p_pow_rep0 power_one_right) then have 2: "\

= \

\ ac_Zp \

" using "0" by presburger have "ac_Zp \

\ carrier Zp" using ac_Zp_in_Zp[of \

] by (simp add: ord_Zp_p ord_of_nonzero(1)) then show ?thesis by (metis "1" "2" m_lcancel R.one_closed R.r_one Zp_int_inc_closed p_pow_nonzero(2)) qed lemma(in padic_integers) ac_Zp_p_nat_pow: "(ac_Zp (\

[^] (n::nat))) = \" apply(induction n) apply (simp add: ac_Zp_one) using ac_Zp_mult ac_Zp_p int_nat_pow_rep nat_pow_Suc2 R.nat_pow_one R.one_closed p_natpow_prod_Suc(1) p_nonzero p_pow_nonzero' p_pow_rep0 by auto text\Facts for reasoning about integer powers in an arbitrary commutative monoid:\ lemma(in monoid) int_pow_add: fixes n::int fixes m::int assumes "a \ Units G" shows "a [^] (n + m) = (a [^] n) \ (a [^] m)" proof- have 0: "group (units_of G)" by (simp add: units_group) have 1: "a \ carrier (units_of G)" by (simp add: assms units_of_carrier) have "\n::int. a [^] n = a [^]\<^bsub>units_of G\<^esub> n" proof- fix k::int show "a [^] k = a [^]\<^bsub>units_of G\<^esub> k" using 1 assms units_of_pow by (metis Units_pow_closed int_pow_def nat_pow_def units_of_inv units_of_pow) qed have 2: "a [^]\<^bsub>units_of G\<^esub> (n + m) = (a [^]\<^bsub>units_of G\<^esub> n) \\<^bsub>units_of G\<^esub> (a [^]\<^bsub>units_of G\<^esub> m)" by (simp add: "1" group.int_pow_mult units_group) show ?thesis using 0 1 2 by (simp add: \\n. a [^] n = a [^]\<^bsub>units_of G\<^esub> n\ units_of_mult) qed lemma(in monoid) int_pow_unit_closed: fixes n::int assumes "a \ Units G" shows "a[^] n \ Units G" apply(cases "n \ 0") using units_of_def[of G] units_group Units_inv_Units[of a] Units_pow_closed[of "inv a"] Units_pow_closed[of a] apply (metis assms pow_nat) using units_of_def[of G] units_group Units_inv_Units[of a] Units_pow_closed[of "inv a"] Units_pow_closed[of a] by (simp add: assms int_pow_def nat_pow_def) lemma(in monoid) nat_pow_of_inv: fixes n::nat assumes "a \ Units G" shows "inv a[^] n = inv (a[^] n)" by (metis (no_types, opaque_lifting) Units_inv_Units Units_inv_closed Units_inv_inv Units_pow_closed Units_r_inv assms inv_unique' nat_pow_closed nat_pow_one pow_mult_distrib) lemma(in monoid) int_pow_of_inv: fixes n::int assumes "a \ Units G" shows "inv a[^] n = inv (a[^] n)" apply(cases "n \0") apply (metis assms nat_pow_of_inv pow_nat) by (metis assms int_pow_def2 nat_pow_of_inv) lemma(in monoid) int_pow_inv: fixes n::int assumes "a \ Units G" shows "a[^] -n = inv a[^] n" apply(cases "n =0") apply simp apply(cases "n > 0") using int_pow_def2[of G a "-n"] int_pow_of_inv apply (simp add: assms) using assms int_pow_def2[of G a "-n"] int_pow_def2[of G a "n"] int_pow_def2[of G "inv a"] int_pow_of_inv[of a n] Units_inv_Units[of a] Units_inv_inv Units_pow_closed[of a] by (metis linorder_not_less nat_0_iff nat_eq_iff2 nat_zero_as_int neg_0_less_iff_less) lemma(in monoid) int_pow_inv': fixes n::int assumes "a \ Units G" shows "a[^] -n = inv (a[^] n)" by (simp add: assms int_pow_inv int_pow_of_inv) lemma(in comm_monoid) inv_of_prod: assumes "a \ Units G" assumes "b \ Units G" shows "inv (a \ b) = (inv a) \ (inv b)" by (metis Units_m_closed assms(1) assms(2) comm_monoid.m_comm comm_monoid_axioms group.inv_mult_group monoid.Units_inv_closed monoid_axioms units_group units_of_carrier units_of_inv units_of_mult) (*************************************************************************************************) (*************************************************************************************************) (************)section\Behaviour of $val\_Zp$ and $ord\_Zp$ on Natural Numbers and Integers\ (***********) (*************************************************************************************************) (*************************************************************************************************) text\If f and g have an equal residue at k, then they differ by a multiple of $p^k$.\ lemma(in padic_integers) eq_residue_mod: assumes "f \ carrier Zp" assumes "g \ carrier Zp" assumes "f k = g k" shows "\h. h \ carrier Zp \ f = g \ (\

[^]k)\h" proof(cases "f = g") case True then show ?thesis using Zp_int_inc_zero' assms(1) by auto next case False have "(f \ g) k = 0" using assms by (metis R.r_right_minus_eq residue_of_diff residue_of_zero(2)) then have "ord_Zp (f \ g) \ k" using False assms by (simp add: ord_Zp_geq) then obtain m::int where m_def: "m \ 0 \ ord_Zp (f \ g) = k + m" using zle_iff_zadd by auto have "f \ g = \

[^](k + m) \ ac_Zp (f \ g)" using ac_Zp_factors_x(2)[of "f \ g"] False m_def assms(1) assms(2) by auto then have 0: "f \ g = \

[^]k \ \

[^] m \ ac_Zp (f \ g)" by (simp add: Zp_def m_def padic_integers.p_natintpow_prod padic_integers_axioms) have "\

[^]k \ \

[^] m \ ac_Zp (f \ g) \ carrier Zp" using assms "0" by auto then have "f = g \ \

[^]k \ \

[^] m \ ac_Zp (f \ g)" using 0 assms R.ring_simprules by simp then show ?thesis using mult_assoc by (metis "0" False R.m_closed R.r_right_minus_eq \[p] \ \ [^] k \ [p] \ \ [^] m \ ac_Zp (f \ g) \ carrier Zp\ ac_Zp_in_Zp assms(1) assms(2) m_def p_pow_car) qed lemma(in padic_integers) eq_residue_mod': assumes "f \ carrier Zp" assumes "g \ carrier Zp" assumes "f k = g k" obtains h where "h \ carrier Zp \ f = g \ (\

[^]k)\h" using assms eq_residue_mod by meson text\Valuations of integers which do not divide $p$:\ lemma(in padic_integers) ord_Zp_p_nat_unit: assumes "(n::nat) mod p \ 0" shows "ord_Zp ([n]\\) = 0" using ord_equals[of "[n]\\" "0::nat"] by (simp add: Zp_nat_inc_res assms) lemma(in padic_integers) val_Zp_p_nat_unit: assumes "(n::nat) mod p \ 0" shows "val_Zp ([n]\\) = 0" unfolding val_Zp_def using assms ord_Zp_def ord_Zp_p_nat_unit ord_of_nonzero(1) zero_eint_def by auto lemma(in padic_integers) nat_unit: assumes "(n::nat) mod p \ 0" shows "([n]\\) \ Units Zp " using Zp_nat_mult_closed val_Zp_p_nat_unit by (simp add: assms val_Zp_0_imp_unit ord_Zp_p_nat_unit) lemma(in padic_integers) ord_Zp_p_int_unit: assumes "(n::int) mod p \ 0" shows "ord_Zp ([n]\\) = 0" by (metis One_nat_def Zp_int_inc_closed Zp_int_inc_res assms mod_by_1 of_nat_0 ord_equals power_0 power_one_right) lemma(in padic_integers) val_Zp_p_int_unit: assumes "(n::int) mod p \ 0" shows "val_Zp ([n]\\) = 0" unfolding val_Zp_def using assms ord_Zp_def ord_Zp_p_int_unit ord_of_nonzero(1) zero_eint_def by auto lemma(in padic_integers) int_unit: assumes "(n::int) mod p \ 0" shows "([n]\\) \ Units Zp " by (simp add: assms val_Zp_0_imp_unit val_Zp_p_int_unit) lemma(in padic_integers) int_decomp_ord: assumes "n = l*(p^k)" assumes "l mod p \ 0" shows "ord_Zp ([n]\\) = k" proof- have 0: "n = l * (p^k)" using assms(1) by simp then have "(l * (p ^ k) mod (p ^ (Suc k))) \ 0" using Zp_def Zp_nat_inc_zero assms(2) p_nonzero nonzero_memE padic_integers_axioms R.int_inc_zero nonzero_memE(2) by auto then have 3: "(l * p ^ k) mod (p ^ (Suc k)) \ 0" by presburger show ?thesis using "0" "3" Zp_int_inc_res ord_equals by auto qed lemma(in padic_integers) int_decomp_val: assumes "n = l*(p^k)" assumes "l mod p \ 0" shows "val_Zp ([n]\\) = k" using Zp_def assms(1) assms(2) R.int_inc_closed ord_of_nonzero(1) int_decomp_ord padic_integers_axioms val_ord_Zp by auto text\$\mathbb{Z}_p$ has characteristic zero:\ lemma(in padic_integers) Zp_char_0: assumes "(n::int) > 0" shows "[n]\\ \ \" proof- have "prime (nat p)" using prime prime_nat_iff_prime by blast then obtain l0 k where 0: "nat n = l0*((nat p)^k) \ \ (nat p) dvd l0 " using prime assms prime_power_canonical[of "nat p" "nat n"] by auto obtain l where l_def: "l = int l0" by blast have 1: "n = l*(p^k) \ \ p dvd l " using 0 l_def by (smt assms int_dvd_int_iff int_nat_eq of_nat_mult of_nat_power prime prime_gt_0_int) show ?thesis apply(cases "l = 1") using 1 p_pow_nonzero(2) p_pow_rep0 apply auto[1] using 1 by (simp add: dvd_eq_mod_eq_0 int_decomp_ord ord_of_nonzero(1)) qed lemma(in padic_integers) Zp_char_0': assumes "(n::nat) > 0" shows "[n]\\ \ \" proof- have "[n]\\ = [(int n)]\\" using assms by (simp add: add_pow_def int_pow_int) then show ?thesis using assms Zp_char_0[of "int n"] by simp qed lemma (in domain) not_eq_diff_nonzero: assumes "a \ b" assumes "a \ carrier R" assumes "b \carrier R" shows "a \ b \ nonzero R" by (simp add: nonzero_def assms(1) assms(2) assms(3)) lemma (in domain) minus_a_inv: assumes "a \ carrier R" assumes "b \ carrier R" shows "a \ b = \ (b \ a)" by (simp add: add.inv_mult_group assms(1) assms(2) minus_eq) lemma(in ring) plus_diff_simp: assumes "a \ carrier R" assumes "b \ carrier R" assumes "c \ carrier R" assumes "X = a \ b" assumes "Y = c \ a" shows "X \ Y = c \ b" using assms unfolding a_minus_def using ring_simprules by (simp add: r_neg r_neg2) lemma (in padic_integers) Zp_residue_eq: assumes "a \ carrier Zp" assumes "b \ carrier Zp" assumes "val_Zp (a \ b) > k" shows "(a k) = (b k)" proof- have 0: "(a \ b) k = a k \\<^bsub>Zp_res_ring k\<^esub> b k" using assms by (simp add: residue_of_diff) have 1: "(a \ b) k = 0" using assms zero_below_val by (smt R.minus_closed Zp_def eint_ord_simps(2) padic_integers.p_res_ring_zero padic_integers.residue_of_zero(1) padic_integers.val_ord_Zp padic_integers.zero_below_ord padic_integers_axioms) show ?thesis apply(cases "k = 0") apply (metis assms(1) assms(2) p_res_ring_0' residues_closed) using 0 1 assms p_residues R_cring Zp_def assms(1) assms(2) cring_def padic_set_res_closed residues.res_zero_eq ring.r_right_minus_eq by (metis Zp_defs(3) linorder_neqE_nat not_less0 p_res_ring_zero) qed lemma (in padic_integers) Zp_residue_eq2: assumes "a \ carrier Zp" assumes "b \ carrier Zp" assumes "(a k) = (b k)" assumes "a \ b" shows "val_Zp (a \ b) \ k" proof- have "(a \ b) k = 0" using assms residue_of_diff by (simp add: Zp_def padic_integers.residue_of_diff' padic_integers_axioms) then show ?thesis using assms(1) assms(2) ord_Zp_def ord_Zp_geq val_Zp_def by auto qed lemma (in padic_integers) equal_val_Zp: assumes "a \ carrier Zp" assumes "b \ carrier Zp" assumes "c \ carrier Zp" assumes "val_Zp a = val_Zp b" assumes "val_Zp (c \ a) > val_Zp b" shows "val_Zp c = val_Zp b" proof- have 0: "val_Zp c = val_Zp (c \ a \ a)" using assms by (simp add: R.l_neg R.minus_eq add_assoc) have "val_Zp c \ min (val_Zp (c \ a)) (val_Zp a)" using val_Zp_ultrametric[of "(c \ a)" a] assms(1) assms(3) ord_Zp_ultrametric_eq'' by (simp add: "0") then have 1: "val_Zp c \ (val_Zp a)" by (metis assms(4) assms(5) dual_order.order_iff_strict less_le_trans min_le_iff_disj) have "val_Zp c = (val_Zp a)" proof(rule ccontr) assume A: "val_Zp c \ val_Zp a" then have 0: "val_Zp c > val_Zp a" using 1 A by auto then have "val_Zp (c \ (a \ c)) \ min (val_Zp c) (val_Zp (a \ c))" by (simp add: assms(1) assms(3) val_Zp_ultrametric) then have 1: "val_Zp a \ min (val_Zp c) (val_Zp (a \ c))" using assms(1) assms(3) assms(4) assms(5) val_Zp_ultrametric_eq' 0 by auto have 2: "val_Zp (a \ c) > val_Zp a" using "0" assms(1) assms(3) assms(4) assms(5) val_Zp_ultrametric_eq' by auto then have "val_Zp a > val_Zp a" using 0 1 2 val_Zp_of_a_inv by (metis assms(1) assms(3) assms(4) assms(5) val_Zp_ultrametric_eq') then show False by blast qed then show ?thesis using assms(4) by simp qed lemma (in padic_integers) equal_val_Zp': assumes "a \ carrier Zp" assumes "b \ carrier Zp" assumes "c \ carrier Zp" assumes "val_Zp a = val_Zp b" assumes "val_Zp c > val_Zp b" shows "val_Zp (a \ c) = val_Zp b" proof- have 0: "val_Zp b < val_Zp (a \ c \ a)" by (simp add: R.minus_eq nonzero_closed R.r_neg1 add_comm assms(1) assms(3) assms(5)) have 1: "val_Zp a \ val_Zp (\ c)" using assms(3) assms(4) assms(5) by (metis eq_iff not_less val_Zp_of_a_inv) then show ?thesis by (meson "0" R.semiring_axioms assms(1) assms(2) assms(3) assms(4) equal_val_Zp semiring.semiring_simprules(1)) qed lemma (in padic_integers) val_Zp_of_minus: assumes "a \ carrier Zp" shows "val_Zp a = val_Zp (\ a)" using assms not_nonzero_Zp ord_Zp_def ord_Zp_of_a_inv val_Zp_def by auto end diff --git a/thys/Polynomial_Factorization/Prime_Factorization.thy b/thys/Polynomial_Factorization/Prime_Factorization.thy --- a/thys/Polynomial_Factorization/Prime_Factorization.thy +++ b/thys/Polynomial_Factorization/Prime_Factorization.thy @@ -1,801 +1,801 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Prime Factorization\ text \This theory contains not-completely naive algorithms to test primality and to perform prime factorization. More precisely, it corresponds to prime factorization algorithm A in Knuth's textbook \<^cite>\"Knuth"\.\ theory Prime_Factorization imports "HOL-Computational_Algebra.Primes" Missing_List Missing_Multiset begin subsection \Definitions\ definition primes_1000 :: "nat list" where "primes_1000 = [2, 3, 5, 7, 11, 13, 17, 19, 23, 29, 31, 37, 41, 43, 47, 53, 59, 61, 67, 71, 73, 79, 83, 89, 97, 101, 103, 107, 109, 113, 127, 131, 137, 139, 149, 151, 157, 163, 167, 173, 179, 181, 191, 193, 197, 199, 211, 223, 227, 229, 233, 239, 241, 251, 257, 263, 269, 271, 277, 281, 283, 293, 307, 311, 313, 317, 331, 337, 347, 349, 353, 359, 367, 373, 379, 383, 389, 397, 401, 409, 419, 421, 431, 433, 439, 443, 449, 457, 461, 463, 467, 479, 487, 491, 499, 503, 509, 521, 523, 541, 547, 557, 563, 569, 571, 577, 587, 593, 599, 601, 607, 613, 617, 619, 631, 641, 643, 647, 653, 659, 661, 673, 677, 683, 691, 701, 709, 719, 727, 733, 739, 743, 751, 757, 761, 769, 773, 787, 797, 809, 811, 821, 823, 827, 829, 839, 853, 857, 859, 863, 877, 881, 883, 887, 907, 911, 919, 929, 937, 941, 947, 953, 967, 971, 977, 983, 991, 997]" lemma primes_1000: "primes_1000 = filter prime [0..<1001]" by eval definition next_candidates :: "nat \ nat \ nat list" where "next_candidates n = (if n = 0 then (1001,primes_1000) else (n + 30, [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]))" definition "candidate_invariant n = (n = 0 \ n mod 30 = (11 :: nat))" partial_function (tailrec) remove_prime_factor :: "nat \ nat \ nat list \ nat \ nat list " where - [code]: "remove_prime_factor p n ps = (case Euclidean_Division.divmod_nat n p of (n',m) \ + [code]: "remove_prime_factor p n ps = (case Euclidean_Rings.divmod_nat n p of (n',m) \ if m = 0 then remove_prime_factor p n' (p # ps) else (n,ps))" partial_function (tailrec) prime_factorization_nat_main :: "nat \ nat \ nat list \ nat list \ nat list" where [code]: "prime_factorization_nat_main n j is ps = (case is of [] \ (case next_candidates j of (j,is) \ prime_factorization_nat_main n j is ps) - | (i # is) \ (case Euclidean_Division.divmod_nat n i of (n',m) \ + | (i # is) \ (case Euclidean_Rings.divmod_nat n i of (n',m) \ if m = 0 then case remove_prime_factor i n' (i # ps) of (n',ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' j is ps' else if i * i \ n then prime_factorization_nat_main n j is ps else (n # ps)))" partial_function (tailrec) prime_nat_main :: "nat \ nat \ nat list \ bool" where [code]: "prime_nat_main n j is = (case is of [] \ (case next_candidates j of (j,is) \ prime_nat_main n j is) | (i # is) \ (if i dvd n then i \ n else if i * i \ n then prime_nat_main n j is else True))" definition prime_nat :: "nat \ bool" where "prime_nat n \ if n < 2 then False else \ \TODO: integrate precomputed map\ case next_candidates 0 of (j,is) \ prime_nat_main n j is" definition prime_factorization_nat :: "nat \ nat list" where "prime_factorization_nat n \ rev (if n < 2 then [] else case next_candidates 0 of (j,is) \ prime_factorization_nat_main n j is [])" definition divisors_nat :: "nat \ nat list" where "divisors_nat n \ if n = 0 then [] else remdups_adj (sort (map prod_list (subseqs (prime_factorization_nat n))))" definition divisors_int_pos :: "int \ int list" where "divisors_int_pos x \ map int (divisors_nat (nat (abs x)))" definition divisors_int :: "int \ int list" where "divisors_int x \ let xs = divisors_int_pos x in xs @ (map uminus xs)" subsection \Proofs\ lemma remove_prime_factor: assumes res: "remove_prime_factor i n ps = (m,qs)" and i: "i > 1" and n: "n \ 0" shows "\ rs. qs = rs @ ps \ n = m * prod_list rs \ \ i dvd m \ set rs \ {i}" using res n proof (induct n arbitrary: ps rule: less_induct) case (less n ps) - obtain n' mo where dm: "Euclidean_Division.divmod_nat n i = (n',mo)" by force - hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: Euclidean_Division.divmod_nat_def) + obtain n' mo where dm: "Euclidean_Rings.divmod_nat n i = (n',mo)" by force + hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: Euclidean_Rings.divmod_nat_def) from less(2)[unfolded remove_prime_factor.simps[of i n] dm] have res: "(if mo = 0 then remove_prime_factor i n' (i # ps) else (n, ps)) = (m, qs)" by auto from less(3) have n: "n \ 0" by auto with n' i have "n' < n" by auto note IH = less(1)[OF this] show ?case proof (cases "mo = 0") case True with mo n' have n: "n = n' * i" by auto with \n \ 0\ have n': "n' \ 0" by auto from True res have "remove_prime_factor i n' (i # ps) = (m,qs)" by auto from IH[OF this n'] obtain rs where "qs = rs @ i # ps" and "n' = m * prod_list rs \ \ i dvd m \ set rs \ {i}" by auto thus ?thesis by (intro exI[of _ "rs @ [i]"], unfold n, auto) next case False with mo have i_n: "\ i dvd n" by auto from False res have id: "m = n" "qs = ps" by auto show ?thesis unfolding id using i_n by auto qed qed lemma prime_sqrtI: assumes n: "n \ 2" and small: "\ j. 2 \ j \ j < i \ \ j dvd n" and i: "\ i * i \ n" shows "prime (n::nat)" unfolding prime_nat_iff proof (intro conjI impI allI) show "1 < n" using n by auto fix j assume jn: "j dvd n" from jn obtain k where njk: "n = j * k" unfolding dvd_def by auto with \1 < n\ have jn: "j \ n" by (metis dvd_imp_le jn neq0_conv not_less0) show "j = 1 \ j = n" proof (rule ccontr) assume "\ ?thesis" with njk n have j1: "j > 1 \ j \ n" by simp have "\ j k. 1 < j \ j \ k \ n = j * k" proof (cases "j \ k") case True thus ?thesis unfolding njk using j1 by blast next case False show ?thesis by (rule exI[of _ k], rule exI[of _ j], insert \1 < n\ j1 njk False, auto) (metis Suc_lessI mult_0_right neq0_conv) qed then obtain j k where j1: "1 < j" and jk: "j \ k" and njk: "n = j * k" by auto with small[of j] have ji: "j \ i" unfolding dvd_def by force from mult_mono[OF ji ji] have "i * i \ j * j" by auto with i have "j * j > n" by auto from this[unfolded njk] have "k < j" by auto with jk show False by auto qed qed lemma candidate_invariant_0: "candidate_invariant 0" unfolding candidate_invariant_def by auto lemma next_candidates: assumes res: "next_candidates n = (m,ps)" and n: "candidate_invariant n" shows "candidate_invariant m" "sorted ps" "{i. prime i \ n \ i \ i < m} \ set ps" "set ps \ {2..} \ {n.. []" "n < m" unfolding candidate_invariant_def proof - note res = res[unfolded next_candidates_def] note n = n[unfolded candidate_invariant_def] show "m = 0 \ m mod 30 = 11" using res n by (auto split: if_splits) show "sorted ps" using res n by (auto split: if_splits simp: primes_1000_def sorted2_simps simp del: sorted_wrt.simps(2)) show "set ps \ {2..} \ {n.. []" using res n by (auto split: if_splits simp: primes_1000_def) show "n < m" using res by (auto split: if_splits) show "{i. prime i \ n \ i \ i < m} \ set ps" proof (cases "n = 0") case True hence *: "m = 1001" "ps = primes_1000" using res by auto show ?thesis unfolding * True primes_1000 by auto next case False hence n: "n mod 30 = 11" and m: "m = n + 30" and ps: "ps = [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]" using res n by auto { fix i assume *: "prime i" "n \ i" "i < n + 30" "i \ set ps" from n * have i11: "i \ 11" by auto define j where "j = i - n" have i: "i = n + j" using \n \ i\ j_def by auto have "i mod 30 = (j + n) mod 30" using \n \ i\ unfolding j_def by simp also have "\ = (j mod 30 + n mod 30) mod 30" by (simp add: mod_simps) also have "\ = (j mod 30 + 11) mod 30" unfolding n by simp finally have i30: "i mod 30 = (j mod 30 + 11) mod 30" by simp have 2: "2 dvd (30 :: nat)" and 112: "11 mod (2 :: nat) = 1" by simp_all have "(j + 11) mod 2 = (j + 1) mod 2" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\j. j mod 2"] have 2: "i mod 2 = (j mod 2 + 1) mod 2" by (simp add: mod_simps mod_mod_cancel [OF 2]) have 3: "3 dvd (30 :: nat)" and 113: "11 mod (3 :: nat) = 2" by simp_all have "(j + 11) mod 3 = (j + 2) mod 3" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 3"] have 3: "i mod 3 = (j mod 3 + 2) mod 3" by (simp add: mod_simps mod_mod_cancel [OF 3]) have 5: "5 dvd (30 :: nat)" and 115: "11 mod (5 :: nat) = 1" by simp_all have "(j + 11) mod 5 = (j + 1) mod 5" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 5"] have 5: "i mod 5 = (j mod 5 + 1) mod 5" by (simp add: mod_simps mod_mod_cancel [OF 5]) from n *(2-)[unfolded ps i, simplified] have "j \ {1,3,5,7,9,11,13,15,17,19,21,23,25,27,29} \ j \ {4,10,16,22,28} \ j \ {14,24}" (is "j \ ?j2 \ j \ ?j3 \ j \ ?j5") by simp presburger moreover { assume "j \ ?j2" hence "j mod 2 = 1" by auto with 2 have "i mod 2 = 0" by auto with i11 have "2 dvd i" "i \ 2" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j3" hence "j mod 3 = 1" by auto with 3 have "i mod 3 = 0" by auto with i11 have "3 dvd i" "i \ 3" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j5" hence "j mod 5 = 4" by auto with 5 have "i mod 5 = 0" by auto with i11 have "5 dvd i" "i \ 5" by auto with *(1) have False unfolding prime_nat_iff by auto } ultimately have False by blast } thus ?thesis unfolding m ps by auto qed qed lemma prime_test_iterate2: assumes small: "\ j. 2 \ j \ j < (i :: nat) \ \ j dvd n" and odd: "odd n" and n: "n \ 3" and i: "i \ 3" "odd i" and mod: "\ i dvd n" and j: "2 \ j" "j < i + 2" shows "\ j dvd n" proof assume dvd: "j dvd n" with small[OF j(1)] have "j \ i" by linarith with dvd mod have "j > i" by (cases "i = j", auto) with j have "j = Suc i" by simp with i have "even j" by auto with dvd j(1) have "2 dvd n" by (metis dvd_trans) with odd show False by auto qed lemma prime_divisor: assumes "j \ 2" and "j dvd n" shows "\ p :: nat. prime p \ p dvd j \ p dvd n" proof - let ?pf = "prime_factors j" from assms have "j > 0" by auto from prime_factorization_nat[OF this] have "j = (\p\?pf. p ^ multiplicity p j)" by auto with \j \ 2\ have "?pf \ {}" by auto then obtain p where p: "p \ ?pf" by auto hence pr: "prime p" by auto define rem where "rem = (\p\?pf - {p}. p ^ multiplicity p j)" from p have mult: "multiplicity p j \ 0" by (auto simp: prime_factors_multiplicity) have "finite ?pf" by simp have "j = (\p\?pf. p ^ multiplicity p j)" by fact also have "?pf = (insert p (?pf - {p}))" using p by auto also have "(\p\insert p (?pf - {p}). p ^ multiplicity p j) = p ^ multiplicity p j * rem" unfolding rem_def by (subst prod.insert, auto) also have "\ = p * (p ^ (multiplicity p j - 1) * rem)" using mult by (cases "multiplicity p j", auto) finally have pj: "p dvd j" unfolding dvd_def by blast with \j dvd n\ have "p dvd n" by (metis dvd_trans) with pj pr show ?thesis by blast qed lemma prime_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_nat_main n jj is \ res = prime n" proof (induct ni arbitrary: n i "is" jj res rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_nat_main n jjj iis" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) with res[unfolded simps] have res: "res = (if i' dvd n then n \ i' else if i' * i' \ n then prime_nat_main n jj iis else True)" by simp from 1(11) Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) have "set iis \ {i'..}" by(auto simp: Cons) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by(auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 n _ _ i'(3) sd_iis 1(10) iis] show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_nat_main n jj iis" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis by (rule IH[OF _ dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = True" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True have "i' \ 2" using i i' by auto from \i' dvd n\ obtain k where "n = i' * k" .. with n have "k \ 0" by (cases "k = 0", auto) with \n = i' * k\ have *: "i' < n \ i' = n" by auto with True res have "res \ i' = n" by auto also have "\ = prime n" using * proof assume "i' < n" with \i' \ 2\ \i' dvd n\ have "\ prime n" by (auto simp add: prime_nat_iff) with \i' < n\ show ?thesis by auto next assume "i' = n" with dvd n have "prime n" by (simp add: prime_nat_iff') with \i' = n\ show ?thesis by auto qed finally show ?thesis . qed qed qed lemma prime_factorization_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_factorization_nat_main n jj is ps \ \ qs. res = qs @ ps \ Ball (set qs) prime \ n = prod_list qs" proof (induct ni arbitrary: n i "is" jj res ps rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res ps) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_factorization_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_factorization_nat_main n jjj iis ps" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) - obtain n' m where dm: "Euclidean_Division.divmod_nat n i' = (n',m)" by force - hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: Euclidean_Division.divmod_nat_def) + obtain n' m where dm: "Euclidean_Rings.divmod_nat n i' = (n',m)" by force + hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: Euclidean_Rings.divmod_nat_def) have m: "(m = 0) = (i' dvd n)" unfolding m by auto from Cons res[unfolded simps] dm m n' have res: "res = ( if i' dvd n then case remove_prime_factor i' (n div i') (i' # ps) of (n', ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' jj iis ps' else if i' * i' \ n then prime_factorization_nat_main n jj iis ps else n # ps)" by simp from 1(11) i Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" "i' > 1" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) Cons have "set iis \ {i'..}" by(auto) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by (auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 _ _ _ i'(3) sd_iis 1(10) iis] { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_factorization_nat_main n jj iis ps" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast show ?thesis by (rule IH[OF _ n dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = n # ps" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True note i_n = this obtain n'' qs where rp: "remove_prime_factor i' (n div i') (i' # ps) = (n'',qs)" by force with res True have res: "res = (if n'' = 1 then qs else prime_factorization_nat_main n'' jj iis qs)" by auto have pi: "prime i'" unfolding prime_nat_iff proof (intro conjI allI impI) show "1 < i'" using i' i by auto fix j assume ji: "j dvd i'" with i' i have j0: "j \ 0" by (cases "j = 0", auto) from ji i_n have jn: "j dvd n" by (metis dvd_trans) with dvd[of j] have j: "2 > j \ j \ i'" by linarith from ji \1 < i'\ have "j \ i'" unfolding dvd_def by (simp add: dvd_imp_le ji) with j j0 show "j = 1 \ j = i'" by linarith qed from True n' have id: "n = n' * i'" by auto from n id have "n' \ 0" by (cases "n = 0", auto) with id have "i' \ n" by auto from remove_prime_factor[OF rp[folded n'] \1 < i'\ \n' \ 0\] obtain rs where qs: "qs = rs @ i' # ps" and n': "n' = n'' * prod_list rs" and i_n'': "\ i' dvd n''" and rs: "set rs \ {i'}" by auto { fix j assume "j dvd n''" hence "j dvd n" unfolding id n' by auto } note dvd' = this show ?thesis proof (cases "n'' = 1") case False with res have res: "res = prime_factorization_nat_main n'' jj iis qs" by simp from i i' have "i' \ 2" by simp from False n' \n' \ 0\ have n2: "n'' \ 2" by (cases "n'' = 0"; auto) have lrs: "prod_list rs \ 0" using n' \n' \ 0\ by (cases "prod_list rs = 0", auto) with \i' \ 2\ have "prod_list rs * i' \ 2" by (cases "prod_list rs", auto) hence nn'': "n > n''" unfolding id n' using n2 by simp have "i' \ n" unfolding id n' using pi False by fastforce with \i' \ n\ i' have "n > i" by auto with nn'' i i' have less: "n - i > n'' - Suc i'" by simp { fix j assume 2: "2 \ j" and ji: "j < Suc i'" have "\ j dvd n''" proof (cases "j = i'") case False with ji have "j < i'" by auto from dvd' dvd[OF 2 this] show ?thesis by blast qed (insert i_n'', auto) } from IH[OF _ n2 this iis res] less obtain ss where res: "res = ss @ qs \ Ball (set ss) prime \ n'' = prod_list ss" by auto thus ?thesis unfolding id n' qs using pi rs by auto next case True with res have res: "res = qs" by auto show ?thesis unfolding id n' res qs True using rs \prime i'\ by (intro exI[of _ "rs @ [i']"], auto) qed qed qed qed lemma prime_nat[simp]: "prime_nat n = prime n" proof (cases "n < 2") case True thus ?thesis unfolding prime_nat_def prime_nat_iff by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) from n can have res: "prime_nat n = prime_nat_main n jj is" unfolding prime_nat_def by auto show ?thesis using prime_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub res] cann(4) by auto qed lemma prime_factorization_nat: fixes n :: nat defines "pf \ prime_factorization_nat n" shows "Ball (set pf) prime" and "n \ 0 \ prod_list pf = n" and "n = 0 \ pf = []" proof - note pf = pf_def[unfolded prime_factorization_nat_def] have "Ball (set pf) prime \ (n \ 0 \ prod_list pf = n) \ (n = 0 \ pf = [])" proof (cases "n < 2") case True thus ?thesis using pf by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) let ?pfm = "prime_factorization_nat_main n jj is []" from pf[unfolded can] False have res: "pf = rev ?pfm" by simp from prime_factorization_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub refl, of Nil] cann(4) have "Ball (set ?pfm) prime" "n = prod_list ?pfm" by auto thus ?thesis unfolding res using n by auto qed thus "Ball (set pf) prime" "n \ 0 \ prod_list pf = n" "n = 0 \ pf = []" by auto qed lemma prod_mset_multiset_prime_factorization_nat [simp]: "(x::nat) \ 0 \ prod_mset (prime_factorization x) = x" by simp (* TODO Move *) lemma prime_factorization_unique'': fixes A :: "'a :: {factorial_semiring_multiplicative} multiset" assumes "\p. p \# A \ prime p" assumes "prod_mset A = normalize x" shows "prime_factorization x = A" proof - have "prod_mset A \ 0" by (auto dest: assms(1)) with assms(2) have "x \ 0" by simp hence "prod_mset (prime_factorization x) = prod_mset A" by (simp add: assms prod_mset_prime_factorization) with assms show ?thesis by (intro prime_factorization_unique') auto qed lemma multiset_prime_factorization_nat_correct: "prime_factorization n = mset (prime_factorization_nat n)" proof - note pf = prime_factorization_nat[of n] show ?thesis proof (cases "n = 0") case True thus ?thesis using pf(3) by simp next case False note pf = pf(1) pf(2)[OF False] show ?thesis proof (rule prime_factorization_unique'') show "prime p" if "p \# mset (prime_factorization_nat n)" for p using pf(1) that by simp let ?l = "\i\#prime_factorization n. i" let ?r = "\i\#mset (prime_factorization_nat n). i" show "prod_mset (mset (prime_factorization_nat n)) = normalize n" by (simp add: pf(2) prod_mset_prod_list) qed qed qed lemma multiset_prime_factorization_code[code_unfold]: "prime_factorization = (\n. mset (prime_factorization_nat n))" by (intro ext multiset_prime_factorization_nat_correct) lemma divisors_nat: "n \ 0 \ set (divisors_nat n) = {p. p dvd n}" "distinct (divisors_nat n)" "divisors_nat 0 = []" proof - show "distinct (divisors_nat n)" "divisors_nat 0 = []" unfolding divisors_nat_def by auto assume n: "n \ 0" from n have "n > 0" by auto { fix x have "(x dvd n) = (x \ 0 \ (\p. multiplicity p x \ multiplicity p n))" proof (cases "x = 0") case False with \n > 0\ show ?thesis by (auto simp: dvd_multiplicity_eq) next case True with n show ?thesis by auto qed } note dvd = this let ?dn = "set (divisors_nat n)" let ?mf = "\ (n :: nat). prime_factorization n" have "?dn = prod_list ` set (subseqs (prime_factorization_nat n))" unfolding divisors_nat_def using n by auto also have "\ = prod_mset ` mset ` set (subseqs (prime_factorization_nat n))" by (force simp: prod_mset_prod_list) also have "mset ` set (subseqs (prime_factorization_nat n)) = { ps. ps \# mset (prime_factorization_nat n)}" unfolding multiset_of_subseqs by simp also have "\ = { ps. ps \# ?mf n}" thm multiset_prime_factorization_code[symmetric] unfolding multiset_prime_factorization_nat_correct[symmetric] by auto also have "prod_mset ` \ = {p. p dvd n}" (is "?l = ?r") proof - { fix x assume "x dvd n" from this[unfolded dvd] have x: "x \ 0" by auto from \x dvd n\ \x \ 0\ \n \ 0\ have sub: "?mf x \# ?mf n" by (subst prime_factorization_subset_iff_dvd) auto have "prod_mset (?mf x) = x" using x by (simp add: prime_factorization_nat) hence "x \ ?l" using sub by force } moreover { fix x assume "x \ ?l" then obtain ps where x: "x = prod_mset ps" and sub: "ps \# ?mf n" by auto have "x dvd n" using prod_mset_subset_imp_dvd[OF sub] n x by simp } ultimately show ?thesis by blast qed finally show "set (divisors_nat n) = {p. p dvd n}" . qed lemma divisors_int_pos: "x \ 0 \ set (divisors_int_pos x) = {i. i dvd x \ i > 0}" "distinct (divisors_int_pos x)" "divisors_int_pos 0 = []" proof - show "divisors_int_pos 0 = []" by code_simp show "distinct (divisors_int_pos x)" unfolding divisors_int_pos_def using divisors_nat(2)[of "nat (abs x)"] by (simp add: distinct_map inj_on_def) assume x: "x \ 0" let ?x = "nat (abs x)" from x have xx: "?x \ 0" by auto from x have 0: "\ y. y dvd x \ y \ 0" by auto have id: "int ` {p. int p dvd x} = {i. i dvd x \ 0 < i}" (is "?l = ?r") proof - { fix y assume "y \ ?l" then obtain p where y: "y = int p" and dvd: "int p dvd x" by auto have "y \ ?r" unfolding y using dvd 0[OF dvd] by auto } moreover { fix y assume "y \ ?r" hence dvd: "y dvd x" and y0: "y > 0" by auto define n where "n = nat y" from y0 have y: "y = int n" unfolding n_def by auto with dvd have "y \ ?l" by auto } ultimately show ?thesis by blast qed from xx show "set (divisors_int_pos x) = {i. i dvd x \ i > 0}" by (simp add: divisors_int_pos_def divisors_nat id) qed lemma divisors_int: "x \ 0 \ set (divisors_int x) = {i. i dvd x}" "distinct (divisors_int x)" "divisors_int 0 = []" proof - show "divisors_int 0 = []" by code_simp show "distinct (divisors_int x)" proof (cases "x = 0") case True show ?thesis unfolding True by code_simp next case False from divisors_int_pos(1)[OF False] divisors_int_pos(2) show ?thesis unfolding divisors_int_def Let_def distinct_append distinct_map inj_on_def by auto qed assume x: "x \ 0" show "set (divisors_int x) = {i. i dvd x}" unfolding divisors_int_def Let_def set_append set_map divisors_int_pos(1)[OF x] using x by auto (metis (no_types, lifting) dvd_mult_div_cancel image_eqI linorder_neqE_linordered_idom mem_Collect_eq minus_dvd_iff minus_minus mult_zero_left neg_less_0_iff_less) qed definition divisors_fun :: "('a \ ('a :: {comm_monoid_mult,zero}) list) \ bool" where "divisors_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x }) \ (\ x. distinct (df x))" lemma divisors_funD: "divisors_fun df \ x \ 0 \ d dvd x \ d \ set (df x)" unfolding divisors_fun_def by auto definition divisors_pos_fun :: "('a \ ('a :: {comm_monoid_mult,zero,ord}) list) \ bool" where "divisors_pos_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x \ d > 0}) \ (\ x. distinct (df x))" lemma divisors_pos_funD: "divisors_pos_fun df \ x \ 0 \ d dvd x \ d > 0 \ d \ set (df x)" unfolding divisors_pos_fun_def by auto lemma divisors_fun_nat: "divisors_fun divisors_nat" unfolding divisors_fun_def using divisors_nat by auto lemma divisors_fun_int: "divisors_fun divisors_int" unfolding divisors_fun_def using divisors_int by auto lemma divisors_pos_fun_int: "divisors_pos_fun divisors_int_pos" unfolding divisors_pos_fun_def using divisors_int_pos by auto end diff --git a/thys/Polynomial_Interpolation/Improved_Code_Equations.thy b/thys/Polynomial_Interpolation/Improved_Code_Equations.thy --- a/thys/Polynomial_Interpolation/Improved_Code_Equations.thy +++ b/thys/Polynomial_Interpolation/Improved_Code_Equations.thy @@ -1,73 +1,73 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Improved Code Equations\ text \This theory contains improved code equations for certain algorithms.\ theory Improved_Code_Equations imports "HOL-Computational_Algebra.Polynomial" "HOL-Library.Code_Target_Nat" begin subsection \@{const divmod_integer}.\ text \We improve @{thm divmod_integer_code} by deleting @{const sgn}-expressions.\ text \We guard the application of divmod-abs' with the condition @{term "x \ 0 \ y \ 0"}, so that application can be ensured on non-negative values. Hence, one can drop "abs" in target language setup.\ definition divmod_abs' where "x \ 0 \ y \ 0 \ divmod_abs' x y = Code_Numeral.divmod_abs x y" (* led to an another 10 % improvement on factorization example *) lemma divmod_integer_code''[code]: "divmod_integer k l = (if k = 0 then (0, 0) else if l > 0 then (if k > 0 then divmod_abs' k l else case divmod_abs' (- k) l of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, l - s)) else if l = 0 then (0, k) else apsnd uminus (if k < 0 then divmod_abs' (-k) (-l) else case divmod_abs' k (-l) of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, - l - s)))" unfolding divmod_integer_code by (cases "l = 0"; cases "l < 0"; cases "l > 0"; auto split: prod.splits simp: divmod_abs'_def divmod_abs_def) code_printing \ \FIXME illusion of partiality\ constant divmod_abs' \ (SML) "IntInf.divMod/ ( _,/ _ )" and (Eval) "Integer.div'_mod/ ( _ )/ ( _ )" and (OCaml) "Z.div'_rem" and (Haskell) "divMod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k '/% l))" -subsection \@{const Euclidean_Division.divmod_nat}.\ -text \We implement @{const Euclidean_Division.divmod_nat} via @{const divmod_integer} +subsection \@{const Euclidean_Rings.divmod_nat}.\ +text \We implement @{const Euclidean_Rings.divmod_nat} via @{const divmod_integer} instead of invoking both division and modulo separately, and we further simplify the case-analysis which is performed in @{thm divmod_integer_code''}.\ -lemma divmod_nat_code'[code]: "Euclidean_Division.divmod_nat m n = ( +lemma divmod_nat_code'[code]: "Euclidean_Rings.divmod_nat m n = ( let k = integer_of_nat m; l = integer_of_nat n in map_prod nat_of_integer nat_of_integer (if k = 0 then (0, 0) else if l = 0 then (0,k) else divmod_abs' k l))" using divmod_nat_code [of m n] by (simp add: divmod_abs'_def integer_of_nat_eq_of_nat Let_def) subsection \@{const binomial}\ lemma binomial_code[code]: "n choose k = (if k \ n then fact n div (fact k * fact (n - k)) else 0)" using binomial_eq_0[of n k] binomial_altdef_nat[of k n] by simp end diff --git a/thys/Projective_Measurements/Linear_Algebra_Complements.thy b/thys/Projective_Measurements/Linear_Algebra_Complements.thy --- a/thys/Projective_Measurements/Linear_Algebra_Complements.thy +++ b/thys/Projective_Measurements/Linear_Algebra_Complements.thy @@ -1,2269 +1,2268 @@ (* Author: Mnacho Echenim, Université Grenoble Alpes *) theory Linear_Algebra_Complements imports "Isabelle_Marries_Dirac.Tensor" "Isabelle_Marries_Dirac.More_Tensor" "QHLProver.Gates" "HOL-Types_To_Sets.Group_On_With" "HOL-Probability.Probability" begin hide_const(open) S section \Preliminaries\ subsection \Misc\ lemma mult_real_cpx: fixes a::complex fixes b::complex assumes "a\ Reals" shows "a* (Re b) = Re (a * b)" using assms by (metis Reals_cases complex.exhaust complex.sel(1) complex_of_real_mult_Complex of_real_mult) lemma fct_bound: fixes f::"complex\ real" assumes "f(-1) + f 1 = 1" and "0 \ f 1" and "0 \ f (-1)" shows "-1 \ f 1 - f(-1) \ f 1 - f(-1) \ 1" proof have "f 1 - f(-1) = 1 - f(-1) - f(-1)" using assms by simp also have "...\ -1" using assms by simp finally show "-1 \ f 1 - f(-1)" . next have "f(-1) - f 1 = 1 - f 1 - f 1 " using assms by simp also have "... \ -1" using assms by simp finally have "-1 \ f(-1) - f 1" . thus "f 1 - f (-1) \ 1" by simp qed lemma fct_bound': fixes f::"complex\ real" assumes "f(-1) + f 1 = 1" and "0 \ f 1" and "0 \ f (-1)" shows "\f 1 - f(-1)\ \ 1" using assms fct_bound by auto lemma pos_sum_1_le: assumes "finite I" and "\ i \ I. (0::real) \ f i" and "(\i\ I. f i) = 1" and "j\ I" shows "f j \ 1" proof (rule ccontr) assume "\ f j \ 1" hence "1 < f j" by simp hence "1 < (\i\ I. f i)" using assms by (metis \\ f j \ 1\ sum_nonneg_leq_bound) thus False using assms by simp qed lemma last_subset: assumes "A \ {a,b}" and "a\ b" and "A \ {a, b}" and "A\ {}" and "A \ {a}" shows "A = {b}" using assms by blast lemma disjoint_Un: assumes "disjoint_family_on A (insert x F)" and "x\ F" shows "(A x) \ (\ a\ F. A a) = {}" proof - have "(A x) \ (\ a\ F. A a) = (\i\F. (A x) \ A i)" using Int_UN_distrib by simp also have "... = (\i\F. {})" using assms disjoint_family_onD by fastforce also have "... = {}" using SUP_bot_conv(2) by simp finally show ?thesis . qed lemma sum_but_one: assumes "\j < (n::nat). j \i \ f j = (0::'a::ring)" and "i < n" shows "(\ j \ {0 ..< n}. f j * g j) = f i * g i" proof - have "sum (\x. f x * g x) (({0 ..< n} - {i}) \ {i}) = sum (\x. f x * g x) ({0 ..< n} - {i}) + sum (\x. f x * g x) {i}" by (rule sum.union_disjoint, auto) also have "... = sum (\x. f x * g x) {i}" using assms by auto also have "... = f i * g i" by simp finally have "sum (\x. f x * g x) (({0 ..< n} - {i}) \ {i}) = f i * g i" . moreover have "{0 ..< n} = ({0 ..< n} - {i}) \ {i}" using assms by auto ultimately show ?thesis by simp qed lemma sum_2_elems: assumes "I = {a,b}" and "a\ b" shows "(\a\I. f a) = f a + f b" proof - have "(\a\I. f a) = (\a\(insert a {b}). f a)" using assms by simp also have "... = f a + (\a\{b}. f a)" proof (rule sum.insert) show "finite {b}" by simp show "a\ {b}" using assms by simp qed also have "... = f a + f b" by simp finally show ?thesis . qed lemma sum_4_elems: shows "(\i<(4::nat). f i) = f 0 + f 1 + f 2 + f 3" proof - have "(\i<(4::nat). f i) = (\i<(3::nat). f i) + f 3" by (metis Suc_numeral semiring_norm(2) semiring_norm(8) sum.lessThan_Suc) moreover have "(\i<(3::nat). f i) = (\i<(2::nat). f i) + f 2" by (metis Suc_1 add_2_eq_Suc' nat_1_add_1 numeral_code(3) numerals(1) one_plus_numeral_commute sum.lessThan_Suc) moreover have "(\i<(2::nat). f i) = (\i<(1::nat). f i) + f 1" by (metis Suc_1 sum.lessThan_Suc) ultimately show ?thesis by simp qed lemma disj_family_sum: shows "finite I \ disjoint_family_on A I \ (\i. i \ I \ finite (A i)) \ (\ i \ (\n \ I. A n). f i) = (\ n\ I. (\ i \ A n. f i))" proof (induct rule:finite_induct) case empty then show ?case by simp next case (insert x F) hence "disjoint_family_on A F" by (meson disjoint_family_on_mono subset_insertI) have "(\n \ (insert x F). A n) = A x \ (\n \ F. A n)" using insert by simp hence "(\ i \ (\n \ (insert x F). A n). f i) = (\ i \ (A x \ (\n \ F. A n)). f i)" by simp also have "... = (\ i \ A x. f i) + (\ i \ (\n \ F. A n). f i)" by (rule sum.union_disjoint, (simp add: insert disjoint_Un)+) also have "... = (\ i \ A x. f i) + (\n\F. sum f (A n))" using \disjoint_family_on A F\ by (simp add: insert) also have "... = (\n\(insert x F). sum f (A n))" using insert by simp finally show ?case . qed lemma integrable_real_mult_right: fixes c::real assumes "integrable M f" shows "integrable M (\w. c * f w)" proof (cases "c = 0") case True thus ?thesis by simp next case False thus ?thesis using integrable_mult_right[of c] assms by simp qed subsection \Unifying notions between Isabelle Marries Dirac and QHLProver\ lemma mult_conj_cmod_square: fixes z::complex shows "z * conjugate z = (cmod z)\<^sup>2" proof - have "z * conjugate z = (Re z)\<^sup>2 + (Im z)\<^sup>2" using complex_mult_cnj by auto also have "... = (cmod z)\<^sup>2" unfolding cmod_def by simp finally show ?thesis . qed lemma vec_norm_sq_cpx_vec_length_sq: shows "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" proof - have "(vec_norm v)\<^sup>2 = inner_prod v v" unfolding vec_norm_def using power2_csqrt by blast also have "... = (\i2)" unfolding Matrix.scalar_prod_def proof - have "\i. i < dim_vec v \ Matrix.vec_index v i * conjugate (Matrix.vec_index v i) = (cmod (Matrix.vec_index v i))\<^sup>2" using mult_conj_cmod_square by simp thus "(\i = 0..i2)" by (simp add: lessThan_atLeast0) qed finally show "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" unfolding cpx_vec_length_def by (simp add: sum_nonneg) qed lemma vec_norm_eq_cpx_vec_length: shows "vec_norm v = cpx_vec_length v" using vec_norm_sq_cpx_vec_length_sq by (metis cpx_vec_length_inner_prod inner_prod_csqrt power2_csqrt vec_norm_def) lemma cpx_vec_length_square: shows "\v\\<^sup>2 = (\i = 0..2)" unfolding cpx_vec_length_def by (simp add: lessThan_atLeast0 sum_nonneg) lemma state_qbit_norm_sq: assumes "v\ state_qbit n" shows "(cpx_vec_length v)\<^sup>2 = 1" proof - have "cpx_vec_length v = 1" using assms unfolding state_qbit_def by simp thus ?thesis by simp qed lemma dagger_adjoint: shows "dagger M = Complex_Matrix.adjoint M" unfolding dagger_def Complex_Matrix.adjoint_def by (simp add: cong_mat) subsection \Types to sets lemmata transfers\ context ab_group_add_on_with begin context includes lifting_syntax assumes ltd: "\(Rep::'s \ 'a) (Abs::'a \ 's). type_definition Rep Abs S" begin interpretation local_typedef_ab_group_add_on_with pls z mns um S "TYPE('s)" by unfold_locales fact lemmas lt_sum_union_disjoint = sum.union_disjoint [var_simplified explicit_ab_group_add, unoverload_type 'c, OF type.comm_monoid_add_axioms, untransferred] lemmas lt_disj_family_sum = disj_family_sum [var_simplified explicit_ab_group_add, unoverload_type 'd, OF type.comm_monoid_add_axioms, untransferred] lemmas lt_sum_reindex_cong = sum.reindex_cong [var_simplified explicit_ab_group_add, unoverload_type 'd, OF type.comm_monoid_add_axioms, untransferred] end lemmas sum_with_union_disjoint = lt_sum_union_disjoint [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] lemmas disj_family_sum_with = lt_disj_family_sum [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] lemmas sum_with_reindex_cong = lt_sum_reindex_cong [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] end lemma (in comm_monoid_add_on_with) sum_with_cong': shows "finite I \ (\i. i\ I \ A i = B i) \ (\i. i\ I \ A i \ S) \ (\i. i\ I \ B i \ S) \ sum_with pls z A I = sum_with pls z B I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "sum_with pls z A (insert x F) = pls (A x) (sum_with pls z A F)" using insert sum_with_insert[of A] by (simp add: image_subset_iff) also have "... = pls (B x) (sum_with pls z B F)" using insert by simp also have "... = sum_with pls z B (insert x F)" using insert sum_with_insert[of B] by (simp add: image_subset_iff) finally show ?case . qed section \Linear algebra complements\ subsection \Additional properties of matrices\ lemma smult_one: shows "(1::'a::monoid_mult) \\<^sub>m A = A" by (simp add:eq_matI) lemma times_diag_index: fixes A::"'a::comm_ring Matrix.mat" assumes "A \ carrier_mat n n" and "B\ carrier_mat n n" and "diagonal_mat B" and "j < n" and "i < n" shows "Matrix.vec_index (Matrix.row (A*B) j) i = diag_mat B ! i *A $$ (j, i)" proof - have "Matrix.vec_index (Matrix.row (A*B) j) i = (A*B) $$ (j,i)" using Matrix.row_def[of "A*B" ] assms by simp also have "... = Matrix.scalar_prod (Matrix.row A j) (Matrix.col B i)" using assms times_mat_def[of A] by simp also have "... = Matrix.scalar_prod (Matrix.col B i) (Matrix.row A j)" using comm_scalar_prod[of "Matrix.row A j" n] assms by auto also have "... = (Matrix.vec_index (Matrix.col B i) i) * (Matrix.vec_index (Matrix.row A j) i)" unfolding Matrix.scalar_prod_def proof (rule sum_but_one) show "i < dim_vec (Matrix.row A j)" using assms by simp show "\ia i \ Matrix.vec_index (Matrix.col B i) ia = 0" using assms by (metis carrier_matD(1) carrier_matD(2) diagonal_mat_def index_col index_row(2)) qed also have "... = B $$(i,i) * (Matrix.vec_index (Matrix.row A j) i)" using assms by auto also have "... = diag_mat B ! i * (Matrix.vec_index (Matrix.row A j) i)" unfolding diag_mat_def using assms by simp also have "... = diag_mat B ! i * A $$ (j, i)" using assms by simp finally show ?thesis . qed lemma inner_prod_adjoint_comp: assumes "(U::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "(V::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "i < n" and "j < n" shows "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" proof - have "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) = Matrix.scalar_prod (Matrix.col U j) (Matrix.row (Complex_Matrix.adjoint V) i)" using adjoint_row[of i V] assms by simp also have "... = Matrix.scalar_prod (Matrix.row (Complex_Matrix.adjoint V) i) (Matrix.col U j)" by (metis adjoint_row assms(1) assms(2) assms(3) carrier_matD(1) carrier_matD(2) Matrix.col_dim conjugate_vec_sprod_comm) also have "... = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" using assms by (simp add:times_mat_def) finally show ?thesis . qed lemma mat_unit_vec_col: assumes "(A::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "i < n" shows "A *\<^sub>v (unit_vec n i) = Matrix.col A i" proof show "dim_vec (A *\<^sub>v unit_vec n i) = dim_vec (Matrix.col A i)" using assms by simp show "\j. j < dim_vec (Matrix.col A i) \ Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.vec_index (Matrix.col A i) j" proof - fix j assume "j < dim_vec (Matrix.col A i)" hence "Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.scalar_prod (Matrix.row A j) (unit_vec n i)" unfolding mult_mat_vec_def by simp also have "... = Matrix.scalar_prod (unit_vec n i) (Matrix.row A j)" using comm_scalar_prod assms by auto also have "... = (Matrix.vec_index (unit_vec n i) i) * (Matrix.vec_index (Matrix.row A j) i)" unfolding Matrix.scalar_prod_def proof (rule sum_but_one) show "i < dim_vec (Matrix.row A j)" using assms by auto show "\ia i \ Matrix.vec_index (unit_vec n i) ia = 0" using assms unfolding unit_vec_def by auto qed also have "... = (Matrix.vec_index (Matrix.row A j) i)" using assms by simp also have "... = A $$ (j, i)" using assms \j < dim_vec (Matrix.col A i)\ by simp also have "... = Matrix.vec_index (Matrix.col A i) j" using assms \j < dim_vec (Matrix.col A i)\ by simp finally show "Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.vec_index (Matrix.col A i) j" . qed qed lemma mat_prod_unit_vec_cong: assumes "(A::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "B\ carrier_mat n n" and "\i. i < n \ A *\<^sub>v (unit_vec n i) = B *\<^sub>v (unit_vec n i)" shows "A = B" proof show "dim_row A = dim_row B" using assms by simp show "dim_col A = dim_col B" using assms by simp show "\i j. i < dim_row B \ j < dim_col B \ A $$ (i, j) = B $$ (i, j)" proof - fix i j assume ij: "i < dim_row B" "j < dim_col B" hence "A $$ (i,j) = Matrix.vec_index (Matrix.col A j) i" using assms by simp also have "... = Matrix.vec_index (A *\<^sub>v (unit_vec n j)) i" using mat_unit_vec_col[of A] ij assms by simp also have "... = Matrix.vec_index (B *\<^sub>v (unit_vec n j)) i" using assms ij by simp also have "... = Matrix.vec_index (Matrix.col B j) i" using mat_unit_vec_col ij assms by simp also have "... = B $$ (i,j)" using assms ij by simp finally show "A $$ (i, j) = B $$ (i, j)" . qed qed lemma smult_smult_times: fixes a::"'a::semigroup_mult" shows "a\\<^sub>m (k \\<^sub>m A) = (a * k)\\<^sub>m A" proof show r:"dim_row (a \\<^sub>m (k \\<^sub>m A)) = dim_row (a * k \\<^sub>m A)" by simp show c:"dim_col (a \\<^sub>m (k \\<^sub>m A)) = dim_col (a * k \\<^sub>m A)" by simp show "\i j. i < dim_row (a * k \\<^sub>m A) \ j < dim_col (a * k \\<^sub>m A) \ (a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = (a * k \\<^sub>m A) $$ (i, j)" proof - fix i j assume "i < dim_row (a * k \\<^sub>m A)" and "j < dim_col (a * k \\<^sub>m A)" note ij = this hence "(a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = a * (k \\<^sub>m A) $$ (i, j)" by simp also have "... = a * (k * A $$ (i,j))" using ij by simp also have "... = (a * k) * A $$ (i,j)" by (simp add: semigroup_mult_class.mult.assoc) also have "... = (a * k \\<^sub>m A) $$ (i, j)" using r c ij by simp finally show "(a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = (a * k \\<^sub>m A) $$ (i, j)" . qed qed lemma mat_minus_minus: fixes A :: "'a :: ab_group_add Matrix.mat" assumes "A \ carrier_mat n m" and "B\ carrier_mat n m" and "C\ carrier_mat n m" shows "A - (B - C) = A - B + C" proof show "dim_row (A - (B - C)) = dim_row (A - B + C)" using assms by simp show "dim_col (A - (B - C)) = dim_col (A - B + C)" using assms by simp show "\i j. i < dim_row (A - B + C) \ j < dim_col (A - B + C) \ (A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" proof - fix i j assume "i < dim_row (A - B + C)" and "j < dim_col (A - B + C)" note ij = this have "(A - (B - C)) $$ (i, j) = (A $$ (i,j) - B $$ (i,j) + C $$ (i,j))" using ij assms by simp also have "... = (A - B + C) $$ (i, j)" using assms ij by simp finally show "(A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" . qed qed subsection \Complements on complex matrices\ lemma hermitian_square: assumes "hermitian M" shows "M \ carrier_mat (dim_row M) (dim_row M)" proof - have "dim_col M = dim_row M" using assms unfolding hermitian_def adjoint_def by (metis adjoint_dim_col) thus ?thesis by auto qed lemma hermitian_add: assumes "A\ carrier_mat n n" and "B\ carrier_mat n n" and "hermitian A" and "hermitian B" shows "hermitian (A + B)" unfolding hermitian_def by (metis adjoint_add assms hermitian_def) lemma hermitian_minus: assumes "A\ carrier_mat n n" and "B\ carrier_mat n n" and "hermitian A" and "hermitian B" shows "hermitian (A - B)" unfolding hermitian_def by (metis adjoint_minus assms hermitian_def) lemma hermitian_smult: fixes a::real fixes A::"complex Matrix.mat" assumes "A \ carrier_mat n n" and "hermitian A" shows "hermitian (a \\<^sub>m A)" proof - have dim: "Complex_Matrix.adjoint A \ carrier_mat n n" using assms by (simp add: adjoint_dim) { fix i j assume "i < n" and "j < n" hence "Complex_Matrix.adjoint (a \\<^sub>m A) $$ (i,j) = a * (Complex_Matrix.adjoint A $$ (i,j))" using adjoint_scale[of a A] assms by simp also have "... = a * (A $$ (i,j))" using assms unfolding hermitian_def by simp also have "... = (a \\<^sub>m A) $$ (i,j)" using \i < n\ \j < n\ assms by simp finally have "Complex_Matrix.adjoint (a \\<^sub>m A) $$ (i,j) = (a \\<^sub>m A) $$ (i,j)" . } thus ?thesis using dim assms unfolding hermitian_def by auto qed lemma unitary_eigenvalues_norm_square: fixes U::"complex Matrix.mat" assumes "unitary U" and "U \ carrier_mat n n" and "eigenvalue U k" shows "conjugate k * k = 1" proof - have "\v. eigenvector U v k" using assms unfolding eigenvalue_def by simp from this obtain v where "eigenvector U v k" by auto define vn where "vn = vec_normalize v" have "eigenvector U vn k" using normalize_keep_eigenvector \eigenvector U v k\ using assms(2) eigenvector_def vn_def by blast have "vn \ carrier_vec n" using \eigenvector U v k\ assms(2) eigenvector_def normalized_vec_dim vn_def by blast have "Complex_Matrix.inner_prod vn vn = 1" using \vn = vec_normalize v\ \eigenvector U v k\ eigenvector_def normalized_vec_norm by blast hence "conjugate k * k = conjugate k * k * Complex_Matrix.inner_prod vn vn" by simp also have "... = conjugate k * Complex_Matrix.inner_prod vn (k \\<^sub>v vn)" proof - have "k * Complex_Matrix.inner_prod vn vn = Complex_Matrix.inner_prod vn (k \\<^sub>v vn)" using inner_prod_smult_left[of vn n vn k] \vn \ carrier_vec n\ by simp thus ?thesis by simp qed also have "... = Complex_Matrix.inner_prod (k \\<^sub>v vn) (k \\<^sub>v vn)" using inner_prod_smult_right[of vn n _ k] by (simp add: \vn \ carrier_vec n\) also have "... = Complex_Matrix.inner_prod (U *\<^sub>v vn) (U *\<^sub>v vn)" using \eigenvector U vn k\ unfolding eigenvector_def by simp also have "... = Complex_Matrix.inner_prod (Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn)) vn" using adjoint_def_alter[of "U *\<^sub>v vn" n vn n U] assms by (metis \eigenvector U vn k\ carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec eigenvector_def) also have "... = Complex_Matrix.inner_prod vn vn" proof - have "Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn) = (Complex_Matrix.adjoint U * U) *\<^sub>v vn" using assms by (metis \eigenvector U vn k\ adjoint_dim assoc_mult_mat_vec carrier_matD(1) eigenvector_def) also have "... = vn" using assms unfolding unitary_def inverts_mat_def by (metis \eigenvector U vn k\ assms(1) eigenvector_def one_mult_mat_vec unitary_simps(1)) finally show ?thesis by simp qed also have "... = 1" using \vn = vec_normalize v\ \eigenvector U v k\ eigenvector_def normalized_vec_norm by blast finally show ?thesis . qed lemma outer_prod_smult_left: fixes v::"complex Matrix.vec" shows "outer_prod (a \\<^sub>v v) w = a \\<^sub>m outer_prod v w" proof - define paw where "paw = outer_prod (a \\<^sub>v v) w" define apw where "apw = a \\<^sub>m outer_prod v w" have "paw = apw" proof have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_vec(2)) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim using carrier_vec_dim_vec by blast also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = a * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_smult_right: fixes v::"complex Matrix.vec" shows "outer_prod v (a \\<^sub>v w) = (conjugate a) \\<^sub>m outer_prod v w" proof - define paw where "paw = outer_prod v (a \\<^sub>v w)" define apw where "apw = (conjugate a) \\<^sub>m outer_prod v w" have "paw = apw" proof have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim using carrier_vec_dim_vec by (metis carrier_matD(2) index_smult_vec(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (conjugate a) * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_add_left: fixes v::"complex Matrix.vec" assumes "dim_vec v = dim_vec x" shows "outer_prod (v + x) w = outer_prod v w + (outer_prod x w)" proof - define paw where "paw = outer_prod (v+x) w" define apw where "apw = outer_prod v w + (outer_prod x w)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec add_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (Matrix.vec_index v i + Matrix.vec_index x i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) + Matrix.vec_index x i * cnj (Matrix.vec_index w j)" by (simp add: ring_class.ring_distribs(2)) also have "... = (outer_prod v w) $$ (i,j) + (outer_prod x w) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_add_right: fixes v::"complex Matrix.vec" assumes "dim_vec w = dim_vec x" shows "outer_prod v (w + x) = outer_prod v w + (outer_prod v x)" proof - define paw where "paw = outer_prod v (w+x)" define apw where "apw = outer_prod v w + (outer_prod v x)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2) index_add_vec(2) paw_def) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis assms carrier_matD(2) carrier_vec_dim_vec index_add_mat(3)) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = Matrix.vec_index v i * (cnj (Matrix.vec_index w j + (Matrix.vec_index x j)))" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) + Matrix.vec_index v i * cnj (Matrix.vec_index x j)" by (simp add: ring_class.ring_distribs(1)) also have "... = (outer_prod v w) $$ (i,j) + (outer_prod v x) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_minus_left: fixes v::"complex Matrix.vec" assumes "dim_vec v = dim_vec x" shows "outer_prod (v - x) w = outer_prod v w - (outer_prod x w)" proof - define paw where "paw = outer_prod (v-x) w" define apw where "apw = outer_prod v w - (outer_prod x w)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec minus_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (Matrix.vec_index v i - Matrix.vec_index x i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) - Matrix.vec_index x i * cnj (Matrix.vec_index w j)" by (simp add: ring_class.ring_distribs) also have "... = (outer_prod v w) $$ (i,j) - (outer_prod x w) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_minus_right: fixes v::"complex Matrix.vec" assumes "dim_vec w = dim_vec x" shows "outer_prod v (w - x) = outer_prod v w - (outer_prod v x)" proof - define paw where "paw = outer_prod v (w-x)" define apw where "apw = outer_prod v w - (outer_prod v x)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2) index_minus_vec(2) paw_def) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis assms carrier_matD(2) carrier_vec_dim_vec index_minus_mat(3)) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = Matrix.vec_index v i * (cnj (Matrix.vec_index w j - (Matrix.vec_index x j)))" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) - Matrix.vec_index v i * cnj (Matrix.vec_index x j)" by (simp add: ring_class.ring_distribs) also have "... = (outer_prod v w) $$ (i,j) - (outer_prod v x) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_minus_minus: fixes a::"complex Matrix.vec" assumes "dim_vec a = dim_vec b" and "dim_vec u = dim_vec v" shows "outer_prod (a - b) (u - v) = outer_prod a u - outer_prod a v - outer_prod b u + outer_prod b v" proof - have "outer_prod (a - b) (u - v) = outer_prod a (u - v) - outer_prod b (u - v)" using outer_prod_minus_left assms by simp also have "... = outer_prod a u - outer_prod a v - outer_prod b (u - v)" using assms outer_prod_minus_right by simp also have "... = outer_prod a u - outer_prod a v - (outer_prod b u - outer_prod b v)" using assms outer_prod_minus_right by simp also have "... = outer_prod a u - outer_prod a v - outer_prod b u + outer_prod b v" proof (rule mat_minus_minus) show "outer_prod b u \ carrier_mat (dim_vec b) (dim_vec u)" by simp show "outer_prod b v \ carrier_mat (dim_vec b) (dim_vec u)" using assms by simp show "outer_prod a u - outer_prod a v \ carrier_mat (dim_vec b) (dim_vec u)" using assms by (metis carrier_vecI minus_carrier_mat outer_prod_dim) qed finally show ?thesis . qed lemma outer_trace_inner: assumes "A \ carrier_mat n n" and "dim_vec u = n" and "dim_vec v = n" shows "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.inner_prod v (A *\<^sub>v u)" proof - have "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.trace (A * outer_prod u v)" proof (rule trace_comm) show "A \ carrier_mat n n" using assms by simp show "outer_prod u v \ carrier_mat n n" using assms by (metis carrier_vec_dim_vec outer_prod_dim) qed also have "... = Complex_Matrix.inner_prod v (A *\<^sub>v u)" using trace_outer_prod_right[of A n u v] assms carrier_vec_dim_vec by metis finally show ?thesis . qed lemma zero_hermitian: shows "hermitian (0\<^sub>m n n)" unfolding hermitian_def by (metis adjoint_minus hermitian_def hermitian_one minus_r_inv_mat one_carrier_mat) lemma trace_1: shows "Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat) =(n::complex)" using one_mat_def by (simp add: Complex_Matrix.trace_def Matrix.mat_def) lemma trace_add: assumes "square_mat A" and "square_mat B" and "dim_row A = dim_row B" shows "Complex_Matrix.trace (A + B) = Complex_Matrix.trace A + Complex_Matrix.trace B" using assms by (simp add: Complex_Matrix.trace_def sum.distrib) lemma bra_vec_carrier: shows "bra_vec v \ carrier_mat 1 (dim_vec v)" proof - have "dim_row (ket_vec v) = dim_vec v" unfolding ket_vec_def by simp thus ?thesis using bra_bra_vec[of v] bra_def[of "ket_vec v"] by simp qed lemma mat_mult_ket_carrier: assumes "A\ carrier_mat n m" shows "A * |v\ \ carrier_mat n 1" using assms by (metis bra_bra_vec bra_vec_carrier carrier_matD(1) carrier_matI dagger_of_ket_is_bra dim_row_of_dagger index_mult_mat(2) index_mult_mat(3)) lemma mat_mult_ket: assumes "A \ carrier_mat n m" and "dim_vec v = m" shows "A * |v\ = |A *\<^sub>v v\" proof - have rn: "dim_row (A * |v\) = n" unfolding times_mat_def using assms by simp have co: "dim_col |A *\<^sub>v v\ = 1" using assms unfolding ket_vec_def by simp have cov: "dim_col |v\ = 1" using assms unfolding ket_vec_def by simp have er: "dim_row (A * |v\) = dim_row |A *\<^sub>v v\" using assms by (metis bra_bra_vec bra_vec_carrier carrier_matD(2) dagger_of_ket_is_bra dim_col_of_dagger dim_mult_mat_vec index_mult_mat(2)) have ec: "dim_col (A * |v\) = dim_col |A *\<^sub>v v\" using assms by (metis carrier_matD(2) index_mult_mat(3) mat_mult_ket_carrier) { fix i::nat fix j::nat assume "i < n" and "j < 1" hence "j = 0" by simp have "(A * |v\) $$ (i,0) = Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\ 0)" using times_mat_def[of A] \i < n\ rn cov by simp also have "... = Matrix.scalar_prod (Matrix.row A i) v" using ket_vec_col by simp also have "... = |A *\<^sub>v v\ $$ (i,j)" unfolding mult_mat_vec_def using \i < n\ \j = 0\ assms(1) by auto } note idx = this have "A * |v\ = Matrix.mat n 1 (\(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\ j))" using assms unfolding times_mat_def ket_vec_def by simp also have "... = |A *\<^sub>v v\" using er ec idx rn co by auto finally show ?thesis . qed lemma unitary_density: assumes "density_operator R" and "unitary U" and "R\ carrier_mat n n" and "U\ carrier_mat n n" shows "density_operator (U * R * (Complex_Matrix.adjoint U))" unfolding density_operator_def proof (intro conjI) show "Complex_Matrix.positive (U * R * Complex_Matrix.adjoint U)" proof (rule positive_close_under_left_right_mult_adjoint) show "U \ carrier_mat n n" using assms by simp show "R \ carrier_mat n n" using assms by simp show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp qed have "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = Complex_Matrix.trace (Complex_Matrix.adjoint U * U * R)" using trace_comm[of "U * R" n "Complex_Matrix.adjoint U"] assms by (metis adjoint_dim mat_assoc_test(10)) also have "... = Complex_Matrix.trace R" using assms by simp also have "... = 1" using assms unfolding density_operator_def by simp finally show "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = 1" . qed subsection \Tensor product complements\ lemma tensor_vec_dim[simp]: shows "dim_vec (tensor_vec u v) = dim_vec u * (dim_vec v)" proof - have "length (mult.vec_vec_Tensor (*) (list_of_vec u) (list_of_vec v)) = length (list_of_vec u) * length (list_of_vec v)" using mult.vec_vec_Tensor_length[of "1::real" "(*)" "list_of_vec u" "list_of_vec v"] by (simp add: Matrix_Tensor.mult_def) thus ?thesis unfolding tensor_vec_def by simp qed lemma index_tensor_vec[simp]: assumes "0 < dim_vec v" and "i < dim_vec u * dim_vec v" shows "vec_index (tensor_vec u v) i = vec_index u (i div (dim_vec v)) * vec_index v (i mod dim_vec v)" proof - have m: "Matrix_Tensor.mult (1::complex) (*)" by (simp add: Matrix_Tensor.mult_def) have "length (list_of_vec v) = dim_vec v" using assms by simp hence "vec_index (tensor_vec u v) i = (*) (vec_index u (i div dim_vec v)) (vec_index v (i mod dim_vec v))" unfolding tensor_vec_def using mult.vec_vec_Tensor_elements assms m by (metis (mono_tags, lifting) length_greater_0_conv length_list_of_vec list_of_vec_index mult.vec_vec_Tensor_elements vec_of_list_index) thus ?thesis by simp qed lemma outer_prod_tensor_comm: fixes a::"complex Matrix.vec" fixes u::"complex Matrix.vec" assumes "0 < dim_vec a" and "0 < dim_vec b" shows "outer_prod (tensor_vec u v) (tensor_vec a b) = tensor_mat (outer_prod u a) (outer_prod v b)" proof - define ot where "ot = outer_prod (tensor_vec u v) (tensor_vec a b)" define to where "to = tensor_mat (outer_prod u a) (outer_prod v b)" define dv where "dv = dim_vec v" define db where "db = dim_vec b" have "ot = to" proof have ro: "dim_row ot = dim_vec u * dim_vec v" unfolding ot_def outer_prod_def by simp have "dim_row to = dim_row (outer_prod u a) * dim_row (outer_prod v b)" unfolding to_def by simp also have "... = dim_vec u * dim_vec v" using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec) finally have rt: "dim_row to = dim_vec u * dim_vec v" . show "dim_row ot = dim_row to" using ro rt by simp have co: "dim_col ot = dim_vec a * dim_vec b" unfolding ot_def outer_prod_def by simp have "dim_col to = dim_col (outer_prod u a) * dim_col (outer_prod v b)" unfolding to_def by simp also have "... = dim_vec a * dim_vec b" using outer_prod_dim by (metis carrier_matD(2) carrier_vec_dim_vec) finally have ct: "dim_col to = dim_vec a * dim_vec b" . show "dim_col ot = dim_col to" using co ct by simp show "\i j. i < dim_row to \ j < dim_col to \ ot $$ (i, j) = to $$ (i, j)" proof - fix i j assume "i < dim_row to" and "j < dim_col to" note ij = this have "ot $$ (i,j) = Matrix.vec_index (tensor_vec u v) i * (conjugate (Matrix.vec_index (tensor_vec a b) j))" unfolding ot_def outer_prod_def using ij rt ct by simp also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (Matrix.vec_index (tensor_vec a b) j))" using ij rt assms unfolding dv_def by (metis index_tensor_vec less_nat_zero_code nat_0_less_mult_iff neq0_conv) also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (vec_index a (j div db) * vec_index b (j mod db)))" using ij ct assms unfolding db_def by simp also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (vec_index a (j div db))) * (conjugate (vec_index b (j mod db)))" by simp also have "... = vec_index u (i div dv) * (conjugate (vec_index a (j div db))) * vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" by simp also have "... = (outer_prod u a) $$ (i div dv, j div db) * vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" proof - have "i div dv < dim_vec u" using ij rt unfolding dv_def by (simp add: less_mult_imp_div_less) moreover have "j div db < dim_vec a" using ij ct assms unfolding db_def by (simp add: less_mult_imp_div_less) ultimately have "vec_index u (i div dv) * (conjugate (vec_index a (j div db))) = (outer_prod u a) $$ (i div dv, j div db)" unfolding outer_prod_def by simp thus ?thesis by simp qed also have "... = (outer_prod u a) $$ (i div dv, j div db) * (outer_prod v b) $$ (i mod dv, j mod db)" proof - have "i mod dv < dim_vec v" using ij rt unfolding dv_def using assms mod_less_divisor by (metis less_nat_zero_code mult.commute neq0_conv times_nat.simps(1)) moreover have "j mod db < dim_vec b" using ij ct assms unfolding db_def by (simp add: less_mult_imp_div_less) ultimately have "vec_index v (i mod dv) * (conjugate (vec_index b (j mod db))) = (outer_prod v b) $$ (i mod dv, j mod db)" unfolding outer_prod_def by simp thus ?thesis by simp qed also have "... = tensor_mat (outer_prod u a) (outer_prod v b) $$ (i, j)" proof (rule index_tensor_mat[symmetric]) show "dim_row (outer_prod u a) = dim_vec u" unfolding outer_prod_def by simp show "dim_row (outer_prod v b) = dv" unfolding outer_prod_def dv_def by simp show "dim_col (outer_prod v b) = db" unfolding db_def outer_prod_def by simp show "i < dim_vec u * dv" unfolding dv_def using ij rt by simp show "dim_col (outer_prod u a) = dim_vec a" unfolding outer_prod_def by simp show "j < dim_vec a * db" unfolding db_def using ij ct by simp show "0 < dim_vec a" using assms by simp show "0 < db" unfolding db_def using assms by simp qed finally show "ot $$ (i, j) = to $$ (i, j)" unfolding to_def . qed qed thus ?thesis unfolding ot_def to_def by simp qed lemma tensor_mat_adjoint: assumes "m1 \ carrier_mat r1 c1" and "m2 \ carrier_mat r2 c2" and "0 < c1" and "0 < c2" and "0 < r1" and "0 < r2" shows "Complex_Matrix.adjoint (tensor_mat m1 m2) = tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)" apply (rule eq_matI, auto) proof - fix i j assume "i < dim_col m1 * dim_col m2" and "j < dim_row m1 * dim_row m2" note ij = this have c1: "dim_col m1 = c1" using assms by simp have r1: "dim_row m1 = r1" using assms by simp have c2: "dim_col m2 = c2" using assms by simp have r2: "dim_row m2 = r2" using assms by simp have "Complex_Matrix.adjoint (m1 \ m2) $$ (i, j) = conjugate ((m1 \ m2) $$ (j, i))" using ij by (simp add: adjoint_eval) also have "... = conjugate (m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2))" proof - have "(m1 \ m2) $$ (j, i) = m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2)" proof (rule index_tensor_mat[of m1 r1 c1 m2 r2 c2 j i], (auto simp add: assms ij c1 c2 r1 r2)) show "j < r1 * r2" using ij r1 r2 by simp show "i < c1 * c2" using ij c1 c2 by simp qed thus ?thesis by simp qed also have "... = conjugate (m1 $$ (j div r2, i div c2)) * conjugate ( m2 $$ (j mod r2, i mod c2))" by simp also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) * conjugate ( m2 $$ (j mod r2, i mod c2))" by (metis adjoint_eval c2 ij less_mult_imp_div_less r2) also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) * (Complex_Matrix.adjoint m2) $$ (i mod c2, j mod r2)" - by (metis Euclidean_Division.div_eq_0_iff adjoint_eval assms(4) bits_mod_div_trivial c2 - gr_implies_not_zero ij(2) mult_not_zero r2) + using \0 < c2\ \0 < r2\ by (simp add: adjoint_eval c2 r2) also have "... = (tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)) $$ (i,j)" proof (rule index_tensor_mat[symmetric], (simp add: ij c1 c2 r1 r2) +) show "i < c1 * c2" using ij c1 c2 by simp show "j < r1 * r2" using ij r1 r2 by simp show "0 < r1" using assms by simp show "0 < r2" using assms by simp qed finally show "Complex_Matrix.adjoint (m1 \ m2) $$ (i, j) = (Complex_Matrix.adjoint m1 \ Complex_Matrix.adjoint m2) $$ (i, j)" . qed lemma index_tensor_mat': assumes "0 < dim_col A" and "0 < dim_col B" and "i < dim_row A * dim_row B" and "j < dim_col A * dim_col B" shows "(A \ B) $$ (i, j) = A $$ (i div (dim_row B), j div (dim_col B)) * B $$ (i mod (dim_row B), j mod (dim_col B))" by (rule index_tensor_mat, (simp add: assms)+) lemma tensor_mat_carrier: shows "tensor_mat U V \ carrier_mat (dim_row U * dim_row V) (dim_col U * dim_col V)" by auto lemma tensor_mat_id: assumes "0 < d1" and "0 < d2" shows "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) = 1\<^sub>m (d1 * d2)" proof (rule eq_matI, auto) show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, i) = 1" if "i < (d1 * d2)" for i using that index_tensor_mat'[of "1\<^sub>m d1" "1\<^sub>m d2"] by (simp add: assms less_mult_imp_div_less) next show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, j) = 0" if "i < d1 * d2" "j < d1 * d2" "i \ j" for i j using that index_tensor_mat[of "1\<^sub>m d1" d1 d1 "1\<^sub>m d2" d2 d2 i j] by (metis assms(1) assms(2) index_one_mat(1) index_one_mat(2) index_one_mat(3) less_mult_imp_div_less mod_less_divisor mult_div_mod_eq mult_not_zero) qed lemma tensor_mat_hermitian: assumes "A \ carrier_mat n n" and "B \ carrier_mat n' n'" and "0 < n" and "0 < n'" and "hermitian A" and "hermitian B" shows "hermitian (tensor_mat A B)" using assms by (metis hermitian_def tensor_mat_adjoint) lemma tensor_mat_unitary: assumes "Complex_Matrix.unitary U" and "Complex_Matrix.unitary V" and "0 < dim_row U" and "0 < dim_row V" shows "Complex_Matrix.unitary (tensor_mat U V)" proof - define UI where "UI = tensor_mat U V" have "Complex_Matrix.adjoint UI = tensor_mat (Complex_Matrix.adjoint U) (Complex_Matrix.adjoint V)" unfolding UI_def proof (rule tensor_mat_adjoint) show "U \ carrier_mat (dim_row U) (dim_row U)" using assms unfolding Complex_Matrix.unitary_def by simp show "V \ carrier_mat (dim_row V) (dim_row V)" using assms unfolding Complex_Matrix.unitary_def by simp show "0 < dim_row V" using assms by simp show "0 < dim_row U" using assms by simp show "0 < dim_row V" using assms by simp show "0 < dim_row U" using assms by simp qed hence "UI * (Complex_Matrix.adjoint UI) = tensor_mat (U * Complex_Matrix.adjoint U) (V * Complex_Matrix.adjoint V)" using mult_distr_tensor[of U "Complex_Matrix.adjoint U" "V" "Complex_Matrix.adjoint V"] unfolding UI_def by (metis (no_types, lifting) Complex_Matrix.unitary_def adjoint_dim_col adjoint_dim_row assms carrier_matD(2) ) also have "... = tensor_mat (1\<^sub>m (dim_row U)) (1\<^sub>m (dim_row V))" using assms unitary_simps(2) by (metis Complex_Matrix.unitary_def) also have "... = (1\<^sub>m (dim_row U * dim_row V))" using tensor_mat_id assms by simp finally have "UI * (Complex_Matrix.adjoint UI) = (1\<^sub>m (dim_row U * dim_row V))" . hence "inverts_mat UI (Complex_Matrix.adjoint UI)" unfolding inverts_mat_def UI_def by simp thus ?thesis using assms unfolding Complex_Matrix.unitary_def UI_def by (metis carrier_matD(2) carrier_matI dim_col_tensor_mat dim_row_tensor_mat) qed subsection \Fixed carrier matrices locale\ text \We define a locale for matrices with a fixed number of rows and columns, and define a finite sum operation on this locale. The \verb+Type_To_Sets+ transfer tools then permits to obtain lemmata on the locale for free. \ locale fixed_carrier_mat = fixes fc_mats::"'a::field Matrix.mat set" fixes dimR dimC assumes fc_mats_carrier: "fc_mats = carrier_mat dimR dimC" begin sublocale semigroup_add_on_with fc_mats "(+)" proof (unfold_locales) show "\a b. a \ fc_mats \ b \ fc_mats \ a + b \ fc_mats" using fc_mats_carrier by simp show "\a b c. a \ fc_mats \ b \ fc_mats \ c \ fc_mats \ a + b + c = a + (b + c)" using fc_mats_carrier by simp qed sublocale ab_semigroup_add_on_with fc_mats "(+)" proof (unfold_locales) show "\a b. a \ fc_mats \ b \ fc_mats \ a + b = b + a" using fc_mats_carrier by (simp add: comm_add_mat) qed sublocale comm_monoid_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" proof (unfold_locales) show "0\<^sub>m dimR dimC \ fc_mats" using fc_mats_carrier by simp show "\a. a \ fc_mats \ 0\<^sub>m dimR dimC + a = a" using fc_mats_carrier by simp qed sublocale ab_group_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" "(-)" "uminus" proof (unfold_locales) show "\a. a \ fc_mats \ - a + a = 0\<^sub>m dimR dimC" using fc_mats_carrier by simp show "\a b. a \ fc_mats \ b \ fc_mats \ a - b = a + - b" using fc_mats_carrier by (simp add: add_uminus_minus_mat) show "\a. a \ fc_mats \ - a \ fc_mats" using fc_mats_carrier by simp qed end lemma (in fixed_carrier_mat) smult_mem: assumes "A \ fc_mats" shows "a \\<^sub>m A \ fc_mats" using fc_mats_carrier assms by auto definition (in fixed_carrier_mat) sum_mat where "sum_mat A I = sum_with (+) (0\<^sub>m dimR dimC) A I" lemma (in fixed_carrier_mat) sum_mat_empty[simp]: shows "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp lemma (in fixed_carrier_mat) sum_mat_carrier: shows "(\i. i \ I \ (A i)\ fc_mats) \ sum_mat A I \ carrier_mat dimR dimC" unfolding sum_mat_def using sum_with_mem[of A I] fc_mats_carrier by auto lemma (in fixed_carrier_mat) sum_mat_insert: assumes "A x \ fc_mats" "A ` I \ fc_mats" and A: "finite I" and x: "x \ I" shows "sum_mat A (insert x I) = A x + sum_mat A I" unfolding sum_mat_def using assms sum_with_insert[of A x I] by simp subsection \A locale for square matrices\ locale cpx_sq_mat = fixed_carrier_mat "fc_mats::complex Matrix.mat set" for fc_mats + assumes dim_eq: "dimR = dimC" and npos: "0 < dimR" lemma (in cpx_sq_mat) one_mem: shows "1\<^sub>m dimR \ fc_mats" using fc_mats_carrier dim_eq by simp lemma (in cpx_sq_mat) square_mats: assumes "A \ fc_mats" shows "square_mat A" using fc_mats_carrier dim_eq assms by simp lemma (in cpx_sq_mat) cpx_sq_mat_mult: assumes "A \ fc_mats" and "B \ fc_mats" shows "A * B \ fc_mats" proof - have "dim_row (A * B) = dimR" using assms fc_mats_carrier by simp moreover have "dim_col (A * B) = dimR" using assms fc_mats_carrier dim_eq by simp ultimately show ?thesis using fc_mats_carrier carrier_mat_def dim_eq by auto qed lemma (in cpx_sq_mat) sum_mat_distrib_left: shows "finite I \ R\ fc_mats \ (\i. i \ I \ (A i)\ fc_mats) \ sum_mat (\i. R * (A i)) I = R * (sum_mat A I)" proof (induct rule: finite_induct) case empty hence a: "sum_mat (\i. R * (A i)) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp hence "R * (sum_mat A {}) = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R dimR dimC dimC] empty dim_eq by simp thus ?case using a by simp next case (insert x F) hence "sum_mat (\i. R * A i) (insert x F) = R * (A x) + sum_mat (\i. R * A i) F" using sum_mat_insert[of "\i. R * A i" x F] by (simp add: image_subsetI fc_mats_carrier dim_eq) also have "... = R * (A x) + R * (sum_mat A F)" using insert by simp also have "... = R * (A x + (sum_mat A F))" by (metis dim_eq fc_mats_carrier insert.prems(1) insert.prems(2) insertCI mult_add_distrib_mat sum_mat_carrier) also have "... = R * sum_mat A (insert x F)" proof - have "A x + (sum_mat A F) = sum_mat A (insert x F)" by (rule sum_mat_insert[symmetric], (auto simp add: insert)) thus ?thesis by simp qed finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_distrib_right: shows "finite I \ R\ fc_mats \ (\i. i \ I \ (A i)\ fc_mats) \ sum_mat (\i. (A i) * R) I = (sum_mat A I) * R" proof (induct rule: finite_induct) case empty hence a: "sum_mat (\i. (A i) * R) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp hence "(sum_mat A {}) * R = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R ] dim_eq empty by simp thus ?case using a by simp next case (insert x F) have a: "(\i. A i * R) ` F \ fc_mats" using insert cpx_sq_mat_mult by (simp add: image_subsetI) have "A x * R \ fc_mats" using insert by (metis insertI1 local.fc_mats_carrier mult_carrier_mat dim_eq) hence "sum_mat (\i. A i * R) (insert x F) = (A x) * R + sum_mat (\i. A i * R) F" using insert a using sum_mat_insert[of "\i. A i * R" x F] by (simp add: image_subsetI local.fc_mats_carrier) also have "... = (A x) * R + (sum_mat A F) * R" using insert by simp also have "... = (A x + (sum_mat A F)) * R" proof (rule add_mult_distrib_mat[symmetric]) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier by blast show "R \ carrier_mat dimC dimC" using insert fc_mats_carrier dim_eq by simp qed also have "... = sum_mat A (insert x F) * R" proof - have "A x + (sum_mat A F) = sum_mat A (insert x F)" by (rule sum_mat_insert[symmetric], (auto simp add: insert)) thus ?thesis by simp qed finally show ?case . qed lemma (in cpx_sq_mat) trace_sum_mat: fixes A::"'b \ complex Matrix.mat" shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ Complex_Matrix.trace (sum_mat A I) = (\ i\ I. Complex_Matrix.trace (A i))" unfolding sum_mat_def proof (induct rule: finite_induct) case empty then show ?case using trace_zero dim_eq by simp next case (insert x F) have "Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A (insert x F)) = Complex_Matrix.trace (A x + sum_with (+) (0\<^sub>m dimR dimC) A F)" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff dim_eq) also have "... = Complex_Matrix.trace (A x) + Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A F)" using trace_add square_mats insert by (metis carrier_matD(1) fc_mats_carrier image_subset_iff insert_iff sum_with_mem) also have "... = Complex_Matrix.trace (A x) + (\ i\ F. Complex_Matrix.trace (A i))" using insert by simp also have "... = (\ i\ (insert x F). Complex_Matrix.trace (A i))" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff) finally show ?case . qed lemma (in cpx_sq_mat) cpx_sq_mat_smult: assumes "A \ fc_mats" shows "x \\<^sub>m A \ fc_mats" using assms fc_mats_carrier by auto lemma (in cpx_sq_mat) mult_add_distrib_right: assumes "A\ fc_mats" "B\ fc_mats" "C\ fc_mats" shows "A * (B + C) = A * B + A * C" using assms fc_mats_carrier mult_add_distrib_mat dim_eq by simp lemma (in cpx_sq_mat) mult_add_distrib_left: assumes "A\ fc_mats" "B\ fc_mats" "C\ fc_mats" shows "(B + C) * A = B * A + C * A" using assms fc_mats_carrier add_mult_distrib_mat dim_eq by simp lemma (in cpx_sq_mat) mult_sum_mat_distrib_left: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ B \ fc_mats \ (sum_mat (\i. B * (A i)) I) = B * (sum_mat A I)" proof (induct rule: finite_induct) case empty hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp hence "B * (sum_mat A {}) = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq) moreover have "sum_mat (\i. B * (A i)) {} = 0\<^sub>m dimR dimC" using sum_mat_empty[of "\i. B * (A i)"] by simp ultimately show ?case by simp next case (insert x F) have "sum_mat (\i. B * (A i)) (insert x F) = B * (A x) + sum_mat (\i. B * (A i)) F" using sum_with_insert[of "\i. B * (A i)" x F] insert by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult) also have "... = B * (A x) + B * (sum_mat A F)" using insert by simp also have "... = B * (A x + (sum_mat A F))" proof (rule mult_add_distrib_right[symmetric]) show "B\ fc_mats" using insert by simp show "A x \ fc_mats" using insert by simp show "sum_mat A F \ fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier) qed also have "... = B * (sum_mat A (insert x F))" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff sum_mat_def) finally show ?case . qed lemma (in cpx_sq_mat) mult_sum_mat_distrib_right: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ B \ fc_mats \ (sum_mat (\i. (A i) * B) I) = (sum_mat A I) * B" proof (induct rule: finite_induct) case empty hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp hence "(sum_mat A {}) * B = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq) moreover have "sum_mat (\i. (A i) * B) {} = 0\<^sub>m dimR dimC" by simp ultimately show ?case by simp next case (insert x F) have "sum_mat (\i. (A i) * B) (insert x F) = (A x) * B + sum_mat (\i. (A i) * B) F" using sum_with_insert[of "\i. (A i) * B" x F] insert by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult) also have "... = (A x) * B + (sum_mat A F) * B" using insert by simp also have "... = (A x + (sum_mat A F)) * B" proof (rule mult_add_distrib_left[symmetric]) show "B\ fc_mats" using insert by simp show "A x \ fc_mats" using insert by simp show "sum_mat A F \ fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier) qed also have "... = (sum_mat A (insert x F)) * B" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff sum_mat_def) finally show ?case . qed lemma (in cpx_sq_mat) trace_sum_mat_mat_distrib: assumes "finite I" and "\i. i\ I \ B i \ fc_mats" and "A\ fc_mats" and "C \ fc_mats" shows "(\ i\ I. Complex_Matrix.trace(A * (B i) * C)) = Complex_Matrix.trace (A * (sum_mat B I) * C)" proof - have seq: "sum_mat (\i. A * (B i) * C) I = A * (sum_mat B I) * C" proof - have "sum_mat (\i. A * (B i) * C) I = (sum_mat (\i. A * (B i)) I) * C" proof (rule mult_sum_mat_distrib_right) show "finite I" using assms by simp show "C\ fc_mats" using assms by simp show "\i. i \ I \ A * B i \ fc_mats" using assms cpx_sq_mat_mult by simp qed moreover have "sum_mat (\i. A * (B i)) I = A * (sum_mat B I)" by (rule mult_sum_mat_distrib_left, (auto simp add: assms)) ultimately show "sum_mat (\i. A * (B i) * C) I = A * (sum_mat B I) * C" by simp qed have "(\ i\ I. Complex_Matrix.trace(A * (B i) * C)) = Complex_Matrix.trace (sum_mat (\i. A * (B i) * C) I)" proof (rule trace_sum_mat[symmetric]) show "finite I" using assms by simp fix i assume "i\ I" thus "A * B i * C \ fc_mats" using assms by (simp add: cpx_sq_mat_mult) qed also have "... = Complex_Matrix.trace (A * (sum_mat B I) * C)" using seq by simp finally show ?thesis . qed definition (in cpx_sq_mat) zero_col where "zero_col U = (\i. if i < dimR then Matrix.col U i else 0\<^sub>v dimR)" lemma (in cpx_sq_mat) zero_col_dim: assumes "U \ fc_mats" shows "dim_vec (zero_col U i) = dimR" using assms fc_mats_carrier unfolding zero_col_def by simp lemma (in cpx_sq_mat) zero_col_col: assumes "i < dimR" shows "zero_col U i = Matrix.col U i" using assms unfolding zero_col_def by simp lemma (in cpx_sq_mat) sum_mat_index: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ i < dimR \ j < dimC \ (sum_mat (\k. (A k)) I) $$ (i,j) = (\ k\I. (A k) $$ (i,j))" proof (induct rule: finite_induct) case empty thus ?case unfolding sum_mat_def by simp next case (insert x F) hence "(sum_mat (\k. (A k)) (insert x F)) $$ (i,j) = (A x + (sum_mat (\k. (A k)) F)) $$ (i,j)" using insert sum_mat_insert[of A] by (simp add: image_subsetI local.fc_mats_carrier) also have "... = (A x) $$ (i,j) + (sum_mat (\k. (A k)) F) $$ (i,j)" using insert sum_mat_carrier[of F A] fc_mats_carrier by simp also have "... = (A x) $$ (i,j) + (\ k\F. (A k) $$ (i,j))" using insert by simp also have "... = (\ k\(insert x F). (A k) $$ (i,j))" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_cong: shows "finite I \ (\i. i\ I \ A i = B i) \ (\i. i\ I \ A i \ fc_mats) \ (\i. i\ I \ B i \ fc_mats) \ sum_mat A I = sum_mat B I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "sum_mat A (insert x F) = A x + sum_mat A F" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = B x + sum_mat B F" using insert by simp also have "... = sum_mat B (insert x F)" using insert sum_mat_insert[of B] by (simp add: image_subset_iff) finally show ?case . qed lemma (in cpx_sq_mat) smult_sum_mat: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ a \\<^sub>m sum_mat A I = sum_mat (\i. a \\<^sub>m (A i)) I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "a \\<^sub>m sum_mat A (insert x F) = a \\<^sub>m (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = a \\<^sub>m A x + a \\<^sub>m (sum_mat A F)" using insert by (metis add_smult_distrib_left_mat fc_mats_carrier insert_iff sum_mat_carrier) also have "... = a \\<^sub>m A x + sum_mat (\i. a \\<^sub>m (A i)) F" using insert by simp also have "... = sum_mat (\i. a \\<^sub>m (A i)) (insert x F)" using insert sum_mat_insert[of "(\i. a \\<^sub>m (A i))"] by (simp add: image_subset_iff cpx_sq_mat_smult) finally show ?case . qed lemma (in cpx_sq_mat) zero_sum_mat: shows "finite I \ sum_mat (\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) I = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" proof (induct rule:finite_induct) case empty then show ?case using dim_eq sum_mat_empty by auto next case (insert x F) have "sum_mat (\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) (insert x F) = 0\<^sub>m dimR dimR + sum_mat (\i. 0\<^sub>m dimR dimR) F" using insert dim_eq zero_mem sum_mat_insert[of "\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)"] by fastforce also have "... = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" using insert by auto finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_adjoint: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\i. Complex_Matrix.adjoint (A i)) I" proof (induct rule: finite_induct) case empty then show ?case using zero_hermitian[of dimR] by (metis (no_types) dim_eq hermitian_def sum_mat_empty) next case (insert x F) have "Complex_Matrix.adjoint (sum_mat A (insert x F)) = Complex_Matrix.adjoint (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = Complex_Matrix.adjoint (A x) + Complex_Matrix.adjoint (sum_mat A F)" proof (rule adjoint_add) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier[of F] by simp qed also have "... = Complex_Matrix.adjoint (A x) + sum_mat (\i. Complex_Matrix.adjoint (A i)) F" using insert by simp also have "... = sum_mat (\i. Complex_Matrix.adjoint (A i)) (insert x F)" proof (rule sum_mat_insert[symmetric], (auto simp add: insert)) show "Complex_Matrix.adjoint (A x) \ fc_mats" using insert fc_mats_carrier dim_eq by (simp add: adjoint_dim) show "\i. i \ F \ Complex_Matrix.adjoint (A i) \ fc_mats" using insert fc_mats_carrier dim_eq by (simp add: adjoint_dim) qed finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_hermitian: assumes "finite I" and "\i\ I. hermitian (A i)" and "\i\ I. A i\ fc_mats" shows "hermitian (sum_mat A I)" proof - have "Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\i. Complex_Matrix.adjoint (A i)) I" using assms sum_mat_adjoint[of I] by simp also have "... = sum_mat A I" proof (rule sum_mat_cong, (auto simp add: assms)) show "\i. i \ I \ Complex_Matrix.adjoint (A i) = A i" using assms unfolding hermitian_def by simp show "\i. i \ I \ Complex_Matrix.adjoint (A i) \ fc_mats" using assms fc_mats_carrier dim_eq by (simp add: adjoint_dim) qed finally show ?thesis unfolding hermitian_def . qed lemma (in cpx_sq_mat) sum_mat_positive: shows "finite I \ (\i. i\ I \ Complex_Matrix.positive (A i)) \ (\i. i \ I \ A i \ fc_mats) \ Complex_Matrix.positive (sum_mat A I)" proof (induct rule: finite_induct) case empty then show ?case using positive_zero[of dimR] by (metis (no_types) dim_eq sum_mat_empty) next case (insert x F) hence "sum_mat A (insert x F) = A x + (sum_mat A F)" using sum_mat_insert[of A] by (simp add: image_subset_iff) moreover have "Complex_Matrix.positive (A x + (sum_mat A F))" proof (rule positive_add, (auto simp add: insert)) show "A x \ carrier_mat dimR dimR" using insert fc_mats_carrier dim_eq by simp show "sum_mat A F \ carrier_mat dimR dimR" using insert sum_mat_carrier dim_eq by (metis insertCI) qed ultimately show "Complex_Matrix.positive (sum_mat A (insert x F))" by simp qed lemma (in cpx_sq_mat) sum_mat_left_ortho_zero: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ (B \ fc_mats) \ (\ i. i\ I \ A i * B = (0\<^sub>m dimR dimR)) \ (sum_mat A I) * B = 0\<^sub>m dimR dimR" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_right) next case (insert x F) have "(sum_mat A (insert x F)) * B = (A x + sum_mat A F) * B" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = A x * B + sum_mat A F * B" proof (rule add_mult_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) show "B \ carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp qed also have "... = A x * B + (0\<^sub>m dimR dimR)" using insert by simp also have "... = 0\<^sub>m dimR dimR" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_right_ortho_zero: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ (B \ fc_mats) \ (\ i. i\ I \ B * A i = (0\<^sub>m dimR dimR)) \ B * (sum_mat A I) = 0\<^sub>m dimR dimR" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_left) next case (insert x F) have "B * (sum_mat A (insert x F)) = B * (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = B * A x + B * sum_mat A F" proof (rule mult_add_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) show "B \ carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp qed also have "... = B * A x + (0\<^sub>m dimR dimR)" using insert by simp also have "... = 0\<^sub>m dimR dimR" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_ortho_square: shows "finite I \ (\i. i\ I \ ((A i)::complex Matrix.mat) * (A i) = A i) \ (\i. i\ I \ A i \ fc_mats) \ (\ i j. i\ I \ j\ I \ i\ j \ A i * (A j) = (0\<^sub>m dimR dimR)) \ (sum_mat A I) * (sum_mat A I) = (sum_mat A I)" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis fc_mats_carrier right_mult_zero_mat sum_mat_empty zero_mem) next case (insert x F) have "(sum_mat A (insert x F)) * (sum_mat A (insert x F)) = (A x + sum_mat A F) * (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: \\i. i \ insert x F \ A i * A i = A i\ image_subset_iff) also have "... = A x * (A x + sum_mat A F) + sum_mat A F * (A x + sum_mat A F)" proof (rule add_mult_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) thus "A x + sum_mat A F \ carrier_mat dimC dimC" using insert dim_eq by simp qed also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * (A x + sum_mat A F)" proof - have "A x * (A x + sum_mat A F) = A x * A x + A x * (sum_mat A F)" using dim_eq insert.prems(2) mult_add_distrib_right sum_mat_carrier by (metis fc_mats_carrier insertI1 subsetD subset_insertI) thus ?thesis by simp qed also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * A x + sum_mat A F * (sum_mat A F)" proof - have "sum_mat A F * (A x + local.sum_mat A F) = sum_mat A F * A x + local.sum_mat A F * local.sum_mat A F" using insert dim_eq add_assoc add_mem mult_add_distrib_right cpx_sq_mat_mult sum_mat_carrier by (metis fc_mats_carrier insertI1 subsetD subset_insertI) hence "A x * A x + A x * sum_mat A F + sum_mat A F * (A x + sum_mat A F) = A x * A x + A x * sum_mat A F + (sum_mat A F * A x + sum_mat A F * sum_mat A F)" by simp also have "... = A x * A x + A x * sum_mat A F + sum_mat A F * A x + sum_mat A F * sum_mat A F" proof (rule assoc_add_mat[symmetric]) show "A x * A x + A x * sum_mat A F \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis add_mem cpx_sq_mat_mult insertCI) show "sum_mat A F * A x \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI) show "sum_mat A F * sum_mat A F \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI) qed finally show ?thesis . qed also have "... = A x + sum_mat A F" proof - have "A x * A x = A x" using insert by simp moreover have "sum_mat A F * sum_mat A F = sum_mat A F" using insert by simp moreover have "A x * (sum_mat A F) = 0\<^sub>m dimR dimR" proof - have "A x * (sum_mat A F) = sum_mat (\i. A x * (A i)) F" by (rule sum_mat_distrib_left[symmetric], (simp add: insert)+) also have "... = sum_mat (\i. 0\<^sub>m dimR dimR) F" proof (rule sum_mat_cong, (auto simp add: insert zero_mem)) show "\i. i \ F \ A x * A i = 0\<^sub>m dimR dimR" using insert by auto show "\i. i \ F \ A x * A i \ fc_mats" using insert cpx_sq_mat_mult by auto show "\i. i \ F \ 0\<^sub>m dimR dimR \ fc_mats" using zero_mem dim_eq by simp qed also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp finally show ?thesis . qed moreover have "sum_mat A F * A x = 0\<^sub>m dimR dimR" proof - have "sum_mat A F * A x = sum_mat (\i. A i * (A x)) F" by (rule sum_mat_distrib_right[symmetric], (simp add: insert)+) also have "... = sum_mat (\i. 0\<^sub>m dimR dimR) F" proof (rule sum_mat_cong, (auto simp add: insert zero_mem)) show "\i. i \ F \ A i * A x = 0\<^sub>m dimR dimR" using insert by auto show "\i. i \ F \ A i * A x \ fc_mats" using insert cpx_sq_mat_mult by auto show "\i. i \ F \ 0\<^sub>m dimR dimR \ fc_mats" using zero_mem dim_eq by simp qed also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp finally show ?thesis . qed ultimately show ?thesis using add_commute add_zero insert.prems(2) zero_mem dim_eq by auto qed also have "... = sum_mat A (insert x F)" using insert sum_mat_insert[of A x F] by (simp add: \\i. i \ insert x F \ A i * A i = A i\ image_subsetI) finally show ?case . qed lemma diagonal_unit_vec: assumes "B \ carrier_mat n n" and "diagonal_mat (B::complex Matrix.mat)" shows "B *\<^sub>v (unit_vec n i) = B $$ (i,i) \\<^sub>v (unit_vec n i)" proof - define v::"complex Matrix.vec" where "v = unit_vec n i" have "B *\<^sub>v v = Matrix.vec n (\ i. Matrix.scalar_prod (Matrix.row B i) v)" using assms unfolding mult_mat_vec_def by simp also have "... = Matrix.vec n (\ i. B $$(i,i) * Matrix.vec_index v i)" proof - have "\i < n. (Matrix.scalar_prod (Matrix.row B i) v = B $$(i,i) * Matrix.vec_index v i)" proof (intro allI impI) fix i assume "i < n" have "(Matrix.scalar_prod (Matrix.row B i) v) = (\ j \ {0 ..< n}. Matrix.vec_index (Matrix.row B i) j * Matrix.vec_index v j)" using assms unfolding Matrix.scalar_prod_def v_def by simp also have "... = Matrix.vec_index (Matrix.row B i) i * Matrix.vec_index v i" proof (rule sum_but_one) show "\j < n. j \ i \ Matrix.vec_index (Matrix.row B i) j = 0" proof (intro allI impI) fix j assume "j < n" and "j \ i" hence "Matrix.vec_index (Matrix.row B i) j = B $$ (i,j)" using \i < n\ \j < n\ assms by auto also have "... = 0" using assms \i < n\ \j < n\ \j\ i\ unfolding diagonal_mat_def by simp finally show "Matrix.vec_index (Matrix.row B i) j = 0" . qed show "i < n" using \i < n\ . qed also have "... = B $$(i,i) * Matrix.vec_index v i" using assms \i < n\ by auto finally show "(Matrix.scalar_prod (Matrix.row B i) v) = B $$(i,i) * Matrix.vec_index v i" . qed thus ?thesis by auto qed also have "... = B $$ (i,i) \\<^sub>v v" unfolding v_def unit_vec_def by auto finally have "B *\<^sub>v v = B $$ (i,i) \\<^sub>v v" . thus ?thesis unfolding v_def by simp qed lemma mat_vec_mult_assoc: assumes "A \ carrier_mat n p" and "B\ carrier_mat p q" and "dim_vec v = q" shows "A *\<^sub>v (B *\<^sub>v v) = (A * B) *\<^sub>v v" using assms by auto lemma (in cpx_sq_mat) similar_eigenvectors: assumes "A\ fc_mats" and "B\ fc_mats" and "P\ fc_mats" and "similar_mat_wit A B P (Complex_Matrix.adjoint P)" and "diagonal_mat B" and "i < n" shows "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = B $$ (i,i) \\<^sub>v (P *\<^sub>v (unit_vec dimR i))" proof - have "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = (P * B * (Complex_Matrix.adjoint P)) *\<^sub>v (P *\<^sub>v (unit_vec dimR i))" using assms unfolding similar_mat_wit_def by metis also have "... = P * B * (Complex_Matrix.adjoint P) * P *\<^sub>v (unit_vec dimR i)" proof (rule mat_vec_mult_assoc[of _ dimR dimR], (auto simp add: assms fc_mats_carrier)) show "P \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "P * B * Complex_Matrix.adjoint P \ carrier_mat dimR dimR" using assms fc_mats_carrier by auto qed also have "... = P * B * ((Complex_Matrix.adjoint P) * P) *\<^sub>v (unit_vec dimR i)" using assms dim_eq by (smt fc_mats_carrier mat_assoc_test(1) similar_mat_witD2(6) similar_mat_wit_sym) also have "... = P * B *\<^sub>v (unit_vec dimR i)" proof - have "(Complex_Matrix.adjoint P) * P = 1\<^sub>m dimR" using assms dim_eq unfolding similar_mat_wit_def by (simp add: fc_mats_carrier) thus ?thesis using assms(2) local.fc_mats_carrier dim_eq by auto qed also have "... = P *\<^sub>v (B *\<^sub>v (unit_vec dimR i))" using mat_vec_mult_assoc assms fc_mats_carrier dim_eq by simp also have "... = P *\<^sub>v (B $$ (i,i) \\<^sub>v (unit_vec dimR i))" using assms diagonal_unit_vec fc_mats_carrier dim_eq by simp also have "... = B $$ (i,i) \\<^sub>v (P *\<^sub>v (unit_vec dimR i))" proof (rule mult_mat_vec) show "P \ carrier_mat dimR dimC" using assms fc_mats_carrier by simp show "unit_vec dimR i \ carrier_vec dimC" using dim_eq by simp qed finally show ?thesis . qed subsection \Projectors\ definition projector where "projector M \ (hermitian M \ M * M = M)" lemma projector_hermitian: assumes "projector M" shows "hermitian M" using assms unfolding projector_def by simp lemma zero_projector[simp]: shows "projector (0\<^sub>m n n)" unfolding projector_def proof show "hermitian (0\<^sub>m n n)" using zero_hermitian[of n] by simp show "0\<^sub>m n n * 0\<^sub>m n n = 0\<^sub>m n n" by simp qed lemma projector_square_eq: assumes "projector M" shows "M * M = M" using assms unfolding projector_def by simp lemma projector_positive: assumes "projector M" shows "Complex_Matrix.positive M" proof (rule positive_if_decomp) show "M \ carrier_mat (dim_row M) (dim_row M)" using assms projector_hermitian hermitian_square by auto next have "M = Complex_Matrix.adjoint M" using assms projector_hermitian[of M] unfolding hermitian_def by simp hence "M * Complex_Matrix.adjoint M = M * M" by simp also have "... = M" using assms projector_square_eq by auto finally have "M * Complex_Matrix.adjoint M = M" . thus "\Ma. Ma * Complex_Matrix.adjoint Ma = M" by auto qed lemma projector_collapse_trace: assumes "projector (P::complex Matrix.mat)" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Complex_Matrix.trace (P * R * P) = Complex_Matrix.trace (R * P)" proof - have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R)" using trace_comm assms by auto also have "... = Complex_Matrix.trace ((P * P) * R)" using assms projector_square_eq[of P] by simp also have "... = Complex_Matrix.trace (P * (P * R))" using assms by auto also have "... = Complex_Matrix.trace (P * R * P)" using trace_comm[of P n "P * R"] assms by auto finally have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R * P)" . thus ?thesis by simp qed lemma positive_proj_trace: assumes "projector (P::complex Matrix.mat)" and "Complex_Matrix.positive R" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Complex_Matrix.trace (R * P) \ 0" proof - have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace ((P * R) * P)" using assms projector_collapse_trace by auto also have "... = Complex_Matrix.trace ((P * R) * (Complex_Matrix.adjoint P))" using assms projector_hermitian[of P] unfolding hermitian_def by simp also have "... \ 0" proof (rule positive_trace) show " P * R * Complex_Matrix.adjoint P \ carrier_mat n n" using assms by auto show "Complex_Matrix.positive (P * R * Complex_Matrix.adjoint P)" by (rule positive_close_under_left_right_mult_adjoint[of _ n], (auto simp add: assms)) qed finally show ?thesis . qed lemma trace_proj_pos_real: assumes "projector (P::complex Matrix.mat)" and "Complex_Matrix.positive R" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)" proof - have "Complex_Matrix.trace (R * P) \ 0" using assms positive_proj_trace by simp thus ?thesis by (simp add: complex_eqI less_eq_complex_def) qed lemma (in cpx_sq_mat) trace_sum_mat_proj_pos_real: fixes f::"'a \ real" assumes "finite I" and "\ i\ I. projector (P i)" and "Complex_Matrix.positive R" and "\i\ I. P i \ fc_mats" and "R \ fc_mats" shows "Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)) = Re (Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)))" proof - have sm: "\x. x \ I \ Complex_Matrix.trace (f x \\<^sub>m (R * P x)) = f x * Complex_Matrix.trace (R * P x)" proof - fix i assume "i\ I" show "Complex_Matrix.trace (f i \\<^sub>m (R * P i)) = f i * Complex_Matrix.trace (R * P i)" proof (rule trace_smult) show "R * P i \ carrier_mat dimR dimR" using assms cpx_sq_mat_mult fc_mats_carrier \i\ I\ dim_eq by simp qed qed have sw: "\x. x \ I \ R * (f x \\<^sub>m P x) = f x \\<^sub>m (R * P x)" proof - fix i assume "i \ I" show "R * (f i \\<^sub>m P i) = f i \\<^sub>m (R * P i)" proof (rule mult_smult_distrib) show "R\ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "P i \ carrier_mat dimR dimR" using assms \i\ I\ fc_mats_carrier dim_eq by simp qed qed have dr: "Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)) = Complex_Matrix.trace (sum_mat (\i. (R * (f i \\<^sub>m (P i)))) I)" using sum_mat_distrib_left[of I] assms by (simp add: cpx_sq_mat_smult) also have trs: "... = (\ i\ I. Complex_Matrix.trace (R * (f i \\<^sub>m (P i))))" proof (rule trace_sum_mat, (simp add: assms)) show "\i. i \ I \ R * (f i \\<^sub>m P i) \ fc_mats" using assms by (simp add: cpx_sq_mat_smult cpx_sq_mat_mult) qed also have "... = (\ i\ I. Complex_Matrix.trace (f i \\<^sub>m (R * (P i))))" by (rule sum.cong, (simp add: sw)+) also have "... = (\ i\ I. f i * Complex_Matrix.trace (R * (P i)))" by (rule sum.cong, (simp add: sm)+) also have "... = (\ i\ I. complex_of_real (f i * Re (Complex_Matrix.trace (R * (P i)))))" proof (rule sum.cong, simp) show "\x. x \ I \ complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" proof - fix x assume "x\ I" have "complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x) * complex_of_real (Re (Complex_Matrix.trace (R * P x)))" using assms sum.cong[of I I] fc_mats_carrier trace_proj_pos_real \x \ I\ dim_eq by auto also have "... = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" by simp finally show "complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" . qed qed also have "... = (\ i\ I. f i * Re (Complex_Matrix.trace (R * (P i))))" by simp also have "... = (\ i\ I. Re (Complex_Matrix.trace (f i \\<^sub>m (R * (P i)))))" proof - have "(\ i\ I. f i * Re (Complex_Matrix.trace (R * (P i)))) = (\ i\ I. Re (Complex_Matrix.trace (f i \\<^sub>m (R * (P i)))))" by (rule sum.cong, (simp add: sm)+) thus ?thesis by simp qed also have "... = (\ i\ I. Re (Complex_Matrix.trace (R * (f i \\<^sub>m (P i)))))" proof - have "\i. i \ I \ f i \\<^sub>m (R * (P i)) = R * (f i \\<^sub>m (P i))" using sw by simp thus ?thesis by simp qed also have "... = Re (\ i\ I. (Complex_Matrix.trace (R * (f i \\<^sub>m (P i)))))" by simp also have "... = Re (Complex_Matrix.trace (sum_mat (\i. R * (f i \\<^sub>m (P i))) I))" using trs by simp also have "... = Re (Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i))) I))" using dr by simp finally show ?thesis . qed subsection \Rank 1 projection\ definition rank_1_proj where "rank_1_proj v = outer_prod v v" lemma rank_1_proj_square_mat: shows "square_mat (rank_1_proj v)" using outer_prod_dim unfolding rank_1_proj_def by (metis carrier_matD(1) carrier_matD(2) carrier_vec_dim_vec square_mat.simps) lemma rank_1_proj_dim[simp]: shows "dim_row (rank_1_proj v) = dim_vec v" using outer_prod_dim unfolding rank_1_proj_def using carrier_vecI by blast lemma rank_1_proj_carrier[simp]: shows "rank_1_proj v \ carrier_mat (dim_vec v) (dim_vec v)" using outer_prod_dim unfolding rank_1_proj_def using carrier_vecI by blast lemma rank_1_proj_coord: assumes "i < dim_vec v" and "j < dim_vec v" shows "(rank_1_proj v) $$ (i, j) = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" using assms unfolding rank_1_proj_def outer_prod_def by auto lemma rank_1_proj_adjoint: shows "Complex_Matrix.adjoint (rank_1_proj (v::complex Matrix.vec)) = rank_1_proj v" proof show "dim_row (Complex_Matrix.adjoint (rank_1_proj v)) = dim_row (rank_1_proj v)" using rank_1_proj_square_mat by auto thus "dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)" by auto fix i j assume "i < dim_row (rank_1_proj v)" and "j < dim_col (rank_1_proj v)" note ij = this have "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = conjugate ((rank_1_proj v) $$ (j, i))" using ij \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_eval by fastforce also have "... = conjugate (Matrix.vec_index v j * (cnj (Matrix.vec_index v i)))" using rank_1_proj_coord ij by (metis \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_dim_col rank_1_proj_dim) also have "... = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" by simp also have "... = rank_1_proj v $$ (i, j)" using ij rank_1_proj_coord by (metis \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_dim_col rank_1_proj_dim) finally show "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = rank_1_proj v $$ (i, j)". qed lemma rank_1_proj_hermitian: shows "hermitian (rank_1_proj (v::complex Matrix.vec))" using rank_1_proj_adjoint unfolding hermitian_def by simp lemma rank_1_proj_trace: assumes "\v\ = 1" shows "Complex_Matrix.trace (rank_1_proj v) = 1" proof - have "Complex_Matrix.trace (rank_1_proj v) = inner_prod v v" using trace_outer_prod unfolding rank_1_proj_def using carrier_vecI by blast also have "... = (vec_norm v)\<^sup>2" unfolding vec_norm_def using power2_csqrt by presburger also have "... = \v\\<^sup>2" using vec_norm_sq_cpx_vec_length_sq by simp also have "... = 1" using assms by simp finally show ?thesis . qed lemma rank_1_proj_mat_col: assumes "A \ carrier_mat n n" and "i < n" and "j < n" and "k < n" shows "(rank_1_proj (Matrix.col A i)) $$ (j, k) = A $$ (j, i) * conjugate (A $$ (k,i))" proof - have "(rank_1_proj (Matrix.col A i)) $$ (j, k) = Matrix.vec_index (Matrix.col A i) j * conjugate (Matrix.vec_index (Matrix.col A i) k)" using index_outer_prod[of "Matrix.col A i" n "Matrix.col A i" n] assms unfolding rank_1_proj_def by simp also have "... = A $$ (j, i) * conjugate (A $$ (k,i))" using assms by simp finally show ?thesis . qed lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary_index: assumes "A \ fc_mats" and "B \ fc_mats" and "diagonal_mat B" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (A * B * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (\ i\ {..< dimR}. ((diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) $$ (j,k))" proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms)) show "\i. i < dimR \ (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" using rank_1_proj_carrier assms fc_mats_carrier dim_eq by (metis smult_carrier_mat zero_col_col zero_col_dim) show "k < dimC" using assms dim_eq by simp qed also have "... = (\ i\ {..< dimR}. (diag_mat B)!i* A $$ (j, i) * conjugate (A $$ (k,i)))" proof (rule sum.cong, simp) have idx: "\x. x \ {.. (rank_1_proj (Matrix.col A x)) $$ (j, k) = A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms fc_mats_carrier dim_eq by blast show "\x. x \ {.. ((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" proof - fix x assume "x\ {..< dimR}" have "((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * (rank_1_proj (Matrix.col A x)) $$ (j, k)" proof (rule index_smult_mat) show "j < dim_row (rank_1_proj (Matrix.col A x))" using \x\ {..< dimR}\ assms fc_mats_carrier by simp show "k < dim_col (rank_1_proj (Matrix.col A x))" using \x\ {..< dimR}\ assms fc_mats_carrier rank_1_proj_carrier[of "Matrix.col A x"] by simp qed thus "((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" using idx \x\ {..< dimR}\ by simp qed qed also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j) " unfolding Matrix.scalar_prod_def proof (rule sum.cong) show "{..x. x \ {0.. diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row (A*B) j) x" proof - fix x assume "x\ {0.. {0.. carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "B \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "x < dimR" using \x\{0..< dimR}\ by simp qed moreover have "conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \x\ {0.. assms fc_mats_carrier dim_eq by (simp add: adjoint_col) ultimately show "diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row (A*B) j) x" by simp qed qed also have "... = (A*B * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "Matrix.mat (dim_row (A*B)) (dim_col (Complex_Matrix.adjoint A)) (\(i, j). Matrix.scalar_prod (Matrix.row (A*B) i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$ (j, k) = Matrix.scalar_prod (Matrix.row (A*B) j) (Matrix.col (Complex_Matrix.adjoint A) k)" using assms fc_mats_carrier by simp also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j)" using assms comm_scalar_prod[of "Matrix.row (A*B) j" dimR] fc_mats_carrier dim_eq by (metis adjoint_dim Matrix.col_carrier_vec row_carrier_vec cpx_sq_mat_mult) thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp qed finally show ?thesis . qed lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary: assumes "A \ fc_mats" and "B \ fc_mats" and "diagonal_mat B" and "Complex_Matrix.unitary A" shows "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) = (A * B * (Complex_Matrix.adjoint A))" proof show "dim_row (sum_mat (\i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..i j. i < dim_row (A * B * Complex_Matrix.adjoint A) \ j < dim_col (A * B * Complex_Matrix.adjoint A) \ local.sum_mat (\i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..v\ = 1" shows "projector (rank_1_proj v)" proof - have "Complex_Matrix.adjoint (rank_1_proj v) = rank_1_proj v" using rank_1_proj_adjoint by simp hence a: "hermitian (rank_1_proj v)" unfolding hermitian_def by simp have "rank_1_proj v * rank_1_proj v = inner_prod v v \\<^sub>m outer_prod v v" unfolding rank_1_proj_def using outer_prod_mult_outer_prod assms using carrier_vec_dim_vec by blast also have "... = (vec_norm v)\<^sup>2 \\<^sub>m outer_prod v v" unfolding vec_norm_def using power2_csqrt by presburger also have "... = (cpx_vec_length v)\<^sup>2 \\<^sub>m outer_prod v v " using vec_norm_sq_cpx_vec_length_sq by simp also have "... = outer_prod v v" using assms state_qbit_norm_sq smult_one[of "outer_prod v v"] by simp also have "... = rank_1_proj v" unfolding rank_1_proj_def by simp finally show ?thesis using a unfolding projector_def by simp qed lemma rank_1_proj_positive: assumes "\v\ = 1" shows "Complex_Matrix.positive (rank_1_proj v)" proof - have "projector (rank_1_proj v)" using assms rank_1_proj_projector by simp thus ?thesis using projector_positive by simp qed lemma rank_1_proj_density: assumes "\v\ = 1" shows "density_operator (rank_1_proj v)" unfolding density_operator_def using assms by (simp add: rank_1_proj_positive rank_1_proj_trace) lemma (in cpx_sq_mat) sum_rank_1_proj_unitary_index: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (1\<^sub>m dimR) $$ (j,k)" proof - have "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (\ i\ {..< dimR}. (rank_1_proj (Matrix.col A i)) $$ (j,k))" proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms)) show "\i. i < dimR \ rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" proof - fix i assume "i < dimR" hence "dim_vec (Matrix.col A i) = dimR" using assms fc_mats_carrier by simp thus "rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" using rank_1_proj_carrier assms fc_mats_carrier dim_eq fc_mats_carrier by (metis dim_col fc_mats_carrier) qed show "k < dimC" using assms dim_eq by simp qed also have "... = (\ i\ {..< dimR}. A $$ (j, i) * conjugate (A $$ (k,i)))" proof (rule sum.cong, simp) show "\x. x \ {.. rank_1_proj (Matrix.col A x) $$ (j, k) = A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms using local.fc_mats_carrier dim_eq by blast qed also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j) " unfolding Matrix.scalar_prod_def proof (rule sum.cong) show "{..x. x \ {0.. A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x" proof - fix x assume "x\ {0.. {0..x\ {0.. assms fc_mats_carrier dim_eq by simp moreover have "conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \x\ {0.. assms fc_mats_carrier dim_eq by (simp add: adjoint_col) ultimately show "A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x" by simp qed qed also have "... = (A * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "Matrix.mat (dim_row A) (dim_col (Complex_Matrix.adjoint A)) (\(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$ (j, k) = Matrix.scalar_prod (Matrix.row A j) (Matrix.col (Complex_Matrix.adjoint A) k)" using assms fc_mats_carrier by simp also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j)" using assms comm_scalar_prod[of "Matrix.row A j" dimR] fc_mats_carrier by (simp add: adjoint_dim dim_eq) thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp qed also have "... = (1\<^sub>m dimR) $$ (j,k)" using assms dim_eq by (simp add: fc_mats_carrier) finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_sum_density: assumes "finite I" and "\i\ I. \u i\ = 1" and "\i\ I. dim_vec (u i) = dimR" and "\i\ I. 0 \ p i" and "(\i\ I. p i) = 1" shows "density_operator (sum_mat (\i. p i \\<^sub>m (rank_1_proj (u i))) I)" unfolding density_operator_def proof have "Complex_Matrix.trace (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I) = (\i\ I. Complex_Matrix.trace (p i \\<^sub>m rank_1_proj (u i)))" proof (rule trace_sum_mat, (simp add: assms)) show "\i. i \ I \ p i \\<^sub>m rank_1_proj (u i) \ fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto qed also have "... = (\i\ I. p i * Complex_Matrix.trace (rank_1_proj (u i)))" proof - { fix i assume "i \ I" hence "Complex_Matrix.trace (p i \\<^sub>m rank_1_proj (u i)) = p i * Complex_Matrix.trace (rank_1_proj (u i))" using trace_smult[of "rank_1_proj (u i)" dimR] assms fc_mats_carrier dim_eq by auto } thus ?thesis by simp qed also have "... = 1" using assms rank_1_proj_trace assms by auto finally show "Complex_Matrix.trace (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I) = 1" . next show "Complex_Matrix.positive (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I)" proof (rule sum_mat_positive, (simp add: assms)) fix i assume "i\ I" thus "p i \\<^sub>m rank_1_proj (u i) \ fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto show "Complex_Matrix.positive (p i \\<^sub>m rank_1_proj (u i))" using assms \i\ I\ rank_1_proj_positive positive_smult[of "rank_1_proj (u i)" dimR] dim_eq by auto qed qed lemma (in cpx_sq_mat) sum_rank_1_proj_unitary: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" shows "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR})= (1\<^sub>m dimR)" proof show "dim_row (sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR)" using assms fc_mats_carrier by (metis carrier_matD(1) dim_col dim_eq index_one_mat(2) rank_1_proj_carrier sum_mat_carrier) show "dim_col (sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR)" using assms fc_mats_carrier rank_1_proj_carrier sum_mat_carrier by (metis carrier_matD(2) dim_col dim_eq index_one_mat(3) square_mat.simps square_mats) show "\i j. i < dim_row (1\<^sub>m dimR) \ j < dim_col (1\<^sub>m dimR) \ sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR $$ (i, j)" using assms sum_rank_1_proj_unitary_index by simp qed lemma (in cpx_sq_mat) rank_1_proj_unitary: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (1\<^sub>m dimR) $$ (j,k) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" proof - have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = Complex_Matrix.inner_prod (Matrix.col A j) (Matrix.col A k) \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)" using outer_prod_mult_outer_prod assms Matrix.col_dim local.fc_mats_carrier unfolding rank_1_proj_def by blast also have "... = (Complex_Matrix.adjoint A * A) $$ (j, k)\\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using inner_prod_adjoint_comp[of A dimR A] assms fc_mats_carrier dim_eq by simp also have "... = (1\<^sub>m dimR) $$ (j,k) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms unfolding Complex_Matrix.unitary_def by (metis add_commute assms(2) index_add_mat(2) index_one_mat(2) one_mem unitary_simps(1)) finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_unitary_ne: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" and "j\ k" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0\<^sub>m dimR dimR)" proof - have "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (outer_prod (Matrix.col A j) (Matrix.col A k))" by simp also have "... = dimR" using assms fc_mats_carrier dim_eq by auto finally have rw: "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" . have "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (outer_prod (Matrix.col A j) (Matrix.col A k))" by simp also have "... = dimR" using assms fc_mats_carrier dim_eq by auto finally have cl: "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" . have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0::complex) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms rank_1_proj_unitary by simp also have "... = (0\<^sub>m dimR dimR)" proof show "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (0\<^sub>m dimR dimR)" using rw by simp next show "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (0\<^sub>m dimR dimR)" using cl by simp next show "\i p. i < dim_row (0\<^sub>m dimR dimR) \ p < dim_col (0\<^sub>m dimR dimR) \ (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) $$ (i, p) = 0\<^sub>m dimR dimR $$ (i, p)" using rw cl by auto qed finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_unitary_eq: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = rank_1_proj (Matrix.col A j)" proof - have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = (1::complex) \\<^sub>m (rank_1_proj (Matrix.col A j))" using assms rank_1_proj_unitary unfolding rank_1_proj_def by simp also have "... = (rank_1_proj (Matrix.col A j))" by (simp add: smult_one) finally show ?thesis . qed end \ No newline at end of file diff --git a/thys/Registers/Finite_Tensor_Product_Matrices.thy b/thys/Registers/Finite_Tensor_Product_Matrices.thy --- a/thys/Registers/Finite_Tensor_Product_Matrices.thy +++ b/thys/Registers/Finite_Tensor_Product_Matrices.thy @@ -1,251 +1,250 @@ section \Tensor products as matrices\ theory Finite_Tensor_Product_Matrices imports Finite_Tensor_Product begin definition tensor_pack :: "nat \ nat \ (nat \ nat) \ nat" where "tensor_pack X Y = (\(x, y). x * Y + y)" definition tensor_unpack :: "nat \ nat \ nat \ (nat \ nat)" where "tensor_unpack X Y xy = (xy div Y, xy mod Y)" lemma tensor_unpack_inj: assumes "i < A * B" and "j < A * B" shows "tensor_unpack A B i = tensor_unpack A B j \ i = j" by (metis div_mult_mod_eq prod.sel(1) prod.sel(2) tensor_unpack_def) lemma tensor_unpack_bound1[simp]: "i < A * B \ fst (tensor_unpack A B i) < A" unfolding tensor_unpack_def apply auto using less_mult_imp_div_less by blast lemma tensor_unpack_bound2[simp]: "i < A * B \ snd (tensor_unpack A B i) < B" unfolding tensor_unpack_def apply auto by (metis mod_less_divisor mult.commute mult_zero_left nat_neq_iff not_less0) lemma tensor_unpack_fstfst: \fst (tensor_unpack A B (fst (tensor_unpack (A * B) C i))) = fst (tensor_unpack A (B * C) i)\ unfolding tensor_unpack_def apply auto by (metis div_mult2_eq mult.commute) lemma tensor_unpack_sndsnd: \snd (tensor_unpack B C (snd (tensor_unpack A (B * C) i))) = snd (tensor_unpack (A * B) C i)\ unfolding tensor_unpack_def apply auto by (meson dvd_triv_right mod_mod_cancel) lemma tensor_unpack_fstsnd: \fst (tensor_unpack B C (snd (tensor_unpack A (B * C) i))) = snd (tensor_unpack A B (fst (tensor_unpack (A * B) C i)))\ unfolding tensor_unpack_def apply auto - by (metis (no_types, lifting) Euclidean_Division.div_eq_0_iff add_0_iff bits_mod_div_trivial div_mult_self4 mod_mult2_eq mod_mult_self1_is_0 mult.commute) - + by (cases \C = 0\) (simp_all add: mult.commute [of B C] mod_mult2_eq [of i C B]) definition "tensor_state_jnf \ \ = (let d1 = dim_vec \ in let d2 = dim_vec \ in vec (d1*d2) (\i. let (i1,i2) = tensor_unpack d1 d2 i in (vec_index \ i1) * (vec_index \ i2)))" lemma tensor_state_jnf_dim[simp]: \dim_vec (tensor_state_jnf \ \) = dim_vec \ * dim_vec \\ unfolding tensor_state_jnf_def Let_def by simp lemma enum_prod_nth_tensor_unpack: assumes \i < CARD('a) * CARD('b)\ shows "(Enum.enum ! i :: 'a::enum\'b::enum) = (let (i1,i2) = tensor_unpack CARD('a) CARD('b) i in (Enum.enum ! i1, Enum.enum ! i2))" using assms by (simp add: enum_prod_def card_UNIV_length_enum product_nth tensor_unpack_def) lemma vec_of_basis_enum_tensor_state_index: fixes \ :: \'a::enum ell2\ and \ :: \'b::enum ell2\ assumes [simp]: \i < CARD('a) * CARD('b)\ shows \vec_of_basis_enum (\ \\<^sub>s \) $ i = (let (i1,i2) = tensor_unpack CARD('a) CARD('b) i in vec_of_basis_enum \ $ i1 * vec_of_basis_enum \ $ i2)\ proof - define i1 i2 where "i1 = fst (tensor_unpack CARD('a) CARD('b) i)" and "i2 = snd (tensor_unpack CARD('a) CARD('b) i)" have [simp]: "i1 < CARD('a)" "i2 < CARD('b)" using assms i1_def tensor_unpack_bound1 apply presburger using assms i2_def tensor_unpack_bound2 by presburger have \vec_of_basis_enum (\ \\<^sub>s \) $ i = Rep_ell2 (\ \\<^sub>s \) (enum_class.enum ! i)\ by (simp add: vec_of_basis_enum_ell2_component) also have \\ = Rep_ell2 \ (Enum.enum!i1) * Rep_ell2 \ (Enum.enum!i2)\ apply (transfer fixing: i i1 i2) by (simp add: enum_prod_nth_tensor_unpack case_prod_beta i1_def i2_def) also have \\ = vec_of_basis_enum \ $ i1 * vec_of_basis_enum \ $ i2\ by (simp add: vec_of_basis_enum_ell2_component) finally show ?thesis by (simp add: case_prod_beta i1_def i2_def) qed lemma vec_of_basis_enum_tensor_state: fixes \ :: \'a::enum ell2\ and \ :: \'b::enum ell2\ shows \vec_of_basis_enum (\ \\<^sub>s \) = tensor_state_jnf (vec_of_basis_enum \) (vec_of_basis_enum \)\ apply (rule eq_vecI, simp_all) apply (subst vec_of_basis_enum_tensor_state_index, simp_all) by (simp add: tensor_state_jnf_def case_prod_beta Let_def) lemma mat_of_cblinfun_tensor_op_index: fixes a :: \'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2\ and b :: \'c::enum ell2 \\<^sub>C\<^sub>L 'd::enum ell2\ assumes [simp]: \i < CARD('b) * CARD('d)\ assumes [simp]: \j < CARD('a) * CARD('c)\ shows \mat_of_cblinfun (tensor_op a b) $$ (i,j) = (let (i1,i2) = tensor_unpack CARD('b) CARD('d) i in let (j1,j2) = tensor_unpack CARD('a) CARD('c) j in mat_of_cblinfun a $$ (i1,j1) * mat_of_cblinfun b $$ (i2,j2))\ proof - define i1 i2 j1 j2 where "i1 = fst (tensor_unpack CARD('b) CARD('d) i)" and "i2 = snd (tensor_unpack CARD('b) CARD('d) i)" and "j1 = fst (tensor_unpack CARD('a) CARD('c) j)" and "j2 = snd (tensor_unpack CARD('a) CARD('c) j)" have [simp]: "i1 < CARD('b)" "i2 < CARD('d)" "j1 < CARD('a)" "j2 < CARD('c)" using assms i1_def tensor_unpack_bound1 apply presburger using assms i2_def tensor_unpack_bound2 apply blast using assms(2) j1_def tensor_unpack_bound1 apply blast using assms(2) j2_def tensor_unpack_bound2 by presburger have \mat_of_cblinfun (tensor_op a b) $$ (i,j) = Rep_ell2 (tensor_op a b *\<^sub>V ket (Enum.enum!j)) (Enum.enum ! i)\ by (simp add: mat_of_cblinfun_ell2_component) also have \\ = Rep_ell2 ((a *\<^sub>V ket (Enum.enum!j1)) \\<^sub>s (b *\<^sub>V ket (Enum.enum!j2))) (Enum.enum!i)\ by (simp add: tensor_op_ell2 enum_prod_nth_tensor_unpack[where i=j] Let_def case_prod_beta j1_def[symmetric] j2_def[symmetric] flip: tensor_ell2_ket) also have \\ = vec_of_basis_enum ((a *\<^sub>V ket (Enum.enum!j1)) \\<^sub>s b *\<^sub>V ket (Enum.enum!j2)) $ i\ by (simp add: vec_of_basis_enum_ell2_component) also have \\ = vec_of_basis_enum (a *\<^sub>V ket (enum_class.enum ! j1)) $ i1 * vec_of_basis_enum (b *\<^sub>V ket (enum_class.enum ! j2)) $ i2\ by (simp add: case_prod_beta vec_of_basis_enum_tensor_state_index i1_def[symmetric] i2_def[symmetric]) also have \\ = Rep_ell2 (a *\<^sub>V ket (enum_class.enum ! j1)) (enum_class.enum ! i1) * Rep_ell2 (b *\<^sub>V ket (enum_class.enum ! j2)) (enum_class.enum ! i2)\ by (simp add: vec_of_basis_enum_ell2_component) also have \\ = mat_of_cblinfun a $$ (i1, j1) * mat_of_cblinfun b $$ (i2, j2)\ by (simp add: mat_of_cblinfun_ell2_component) finally show ?thesis by (simp add: i1_def[symmetric] i2_def[symmetric] j1_def[symmetric] j2_def[symmetric] case_prod_beta) qed definition "tensor_op_jnf A B = (let r1 = dim_row A in let c1 = dim_col A in let r2 = dim_row B in let c2 = dim_col B in mat (r1 * r2) (c1 * c2) (\(i,j). let (i1,i2) = tensor_unpack r1 r2 i in let (j1,j2) = tensor_unpack c1 c2 j in (A $$ (i1,j1)) * (B $$ (i2,j2))))" lemma tensor_op_jnf_dim[simp]: \dim_row (tensor_op_jnf a b) = dim_row a * dim_row b\ \dim_col (tensor_op_jnf a b) = dim_col a * dim_col b\ unfolding tensor_op_jnf_def Let_def by simp_all lemma mat_of_cblinfun_tensor_op: fixes a :: \'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2\ and b :: \'c::enum ell2 \\<^sub>C\<^sub>L 'd::enum ell2\ shows \mat_of_cblinfun (tensor_op a b) = tensor_op_jnf (mat_of_cblinfun a) (mat_of_cblinfun b)\ apply (rule eq_matI, simp_all add: ) apply (subst mat_of_cblinfun_tensor_op_index, simp_all) by (simp add: tensor_op_jnf_def case_prod_beta Let_def) lemma mat_of_cblinfun_assoc_ell2'[simp]: \mat_of_cblinfun (assoc_ell2' :: (('a::enum\('b::enum\'c::enum)) ell2 \\<^sub>C\<^sub>L _)) = one_mat (CARD('a)*CARD('b)*CARD('c))\ (is "mat_of_cblinfun ?assoc = _") proof (rule mat_eq_iff[THEN iffD2], intro conjI allI impI) show \dim_row (mat_of_cblinfun ?assoc) = dim_row (1\<^sub>m (CARD('a) * CARD('b) * CARD('c)))\ by (simp) show \dim_col (mat_of_cblinfun ?assoc) = dim_col (1\<^sub>m (CARD('a) * CARD('b) * CARD('c)))\ by (simp) fix i j let ?i = "Enum.enum ! i :: (('a\'b)\'c)" and ?j = "Enum.enum ! j :: ('a\('b\'c))" assume \i < dim_row (1\<^sub>m (CARD('a) * CARD('b) * CARD('c)))\ then have iB[simp]: \i < CARD('a) * CARD('b) * CARD('c)\ by simp then have iB'[simp]: \i < CARD('a) * (CARD('b) * CARD('c))\ by linarith assume \j < dim_col (1\<^sub>m (CARD('a) * CARD('b) * CARD('c)))\ then have jB[simp]: \j < CARD('a) * CARD('b) * CARD('c)\ by simp then have jB'[simp]: \j < CARD('a) * (CARD('b) * CARD('c))\ by linarith define i1 i23 i2 i3 where "i1 = fst (tensor_unpack CARD('a) (CARD('b)*CARD('c)) i)" and "i23 = snd (tensor_unpack CARD('a) (CARD('b)*CARD('c)) i)" and "i2 = fst (tensor_unpack CARD('b) CARD('c) i23)" and "i3 = snd (tensor_unpack CARD('b) CARD('c) i23)" define j12 j1 j2 j3 where "j12 = fst (tensor_unpack (CARD('a)*CARD('b)) CARD('c) j)" and "j1 = fst (tensor_unpack CARD('a) CARD('b) j12)" and "j2 = snd (tensor_unpack CARD('a) CARD('b) j12)" and "j3 = snd (tensor_unpack (CARD('a)*CARD('b)) CARD('c) j)" have [simp]: "j12 < CARD('a)*CARD('b)" "i23 < CARD('b)*CARD('c)" using j12_def jB tensor_unpack_bound1 apply presburger using i23_def iB' tensor_unpack_bound2 by blast have j1': \fst (tensor_unpack CARD('a) (CARD('b) * CARD('c)) j) = j1\ by (simp add: j1_def j12_def tensor_unpack_fstfst) let ?i1 = "Enum.enum ! i1 :: 'a" and ?i2 = "Enum.enum ! i2 :: 'b" and ?i3 = "Enum.enum ! i3 :: 'c" let ?j1 = "Enum.enum ! j1 :: 'a" and ?j2 = "Enum.enum ! j2 :: 'b" and ?j3 = "Enum.enum ! j3 :: 'c" have i: \?i = ((?i1,?i2),?i3)\ by (auto simp add: enum_prod_nth_tensor_unpack case_prod_beta tensor_unpack_fstfst tensor_unpack_fstsnd tensor_unpack_sndsnd i1_def i2_def i23_def i3_def) have j: \?j = (?j1,(?j2,?j3))\ by (auto simp add: enum_prod_nth_tensor_unpack case_prod_beta tensor_unpack_fstfst tensor_unpack_fstsnd tensor_unpack_sndsnd j1_def j2_def j12_def j3_def) have ijeq: \(?i1,?i2,?i3) = (?j1,?j2,?j3) \ i = j\ unfolding i1_def i2_def i3_def j1_def j2_def j3_def apply simp apply (subst enum_inj, simp, simp) apply (subst enum_inj, simp, simp) apply (subst enum_inj, simp, simp) apply (subst tensor_unpack_inj[symmetric, where i=i and j=j and A="CARD('a)" and B="CARD('b)*CARD('c)"], simp, simp) unfolding prod_eq_iff apply (subst tensor_unpack_inj[symmetric, where i=\snd (tensor_unpack CARD('a) (CARD('b) * CARD('c)) i)\ and A="CARD('b)" and B="CARD('c)"], simp, simp) by (simp add: i1_def[symmetric] j1_def[symmetric] i2_def[symmetric] j2_def[symmetric] i3_def[symmetric] j3_def[symmetric] i23_def[symmetric] j12_def[symmetric] j1' prod_eq_iff tensor_unpack_fstsnd tensor_unpack_sndsnd) have \mat_of_cblinfun ?assoc $$ (i, j) = Rep_ell2 (assoc_ell2' *\<^sub>V ket ?j) ?i\ by (subst mat_of_cblinfun_ell2_component, auto) also have \\ = Rep_ell2 ((ket ?j1 \\<^sub>s ket ?j2) \\<^sub>s ket ?j3) ?i\ by (simp add: j assoc_ell2'_tensor flip: tensor_ell2_ket) also have \\ = (if (?i1,?i2,?i3) = (?j1,?j2,?j3) then 1 else 0)\ by (auto simp add: ket.rep_eq i) also have \\ = (if i=j then 1 else 0)\ using ijeq by simp finally show \mat_of_cblinfun ?assoc $$ (i, j) = 1\<^sub>m (CARD('a) * CARD('b) * CARD('c)) $$ (i, j)\ by auto qed lemma assoc_ell2'_inv: "assoc_ell2 o\<^sub>C\<^sub>L assoc_ell2' = id_cblinfun" apply (rule equal_ket, case_tac x, hypsubst) by (simp flip: tensor_ell2_ket add: cblinfun_apply_cblinfun_compose assoc_ell2'_tensor assoc_ell2_tensor) lemma assoc_ell2_inv: "assoc_ell2' o\<^sub>C\<^sub>L assoc_ell2 = id_cblinfun" apply (rule equal_ket, case_tac x, hypsubst) by (simp flip: tensor_ell2_ket add: cblinfun_apply_cblinfun_compose assoc_ell2'_tensor assoc_ell2_tensor) lemma mat_of_cblinfun_assoc_ell2[simp]: \mat_of_cblinfun (assoc_ell2 :: ((('a::enum\'b::enum)\'c::enum) ell2 \\<^sub>C\<^sub>L _)) = one_mat (CARD('a)*CARD('b)*CARD('c))\ (is "mat_of_cblinfun ?assoc = _") proof - let ?assoc' = "assoc_ell2' :: (('a::enum\('b::enum\'c::enum)) ell2 \\<^sub>C\<^sub>L _)" have "one_mat (CARD('a)*CARD('b)*CARD('c)) = mat_of_cblinfun (?assoc o\<^sub>C\<^sub>L ?assoc')" by (simp add: mult.assoc mat_of_cblinfun_id) also have \\ = mat_of_cblinfun ?assoc * mat_of_cblinfun ?assoc'\ using mat_of_cblinfun_compose by blast also have \\ = mat_of_cblinfun ?assoc * one_mat (CARD('a)*CARD('b)*CARD('c))\ by simp also have \\ = mat_of_cblinfun ?assoc\ apply (rule right_mult_one_mat') by (simp) finally show ?thesis by simp qed end diff --git a/thys/Sigma_Commit_Crypto/Rivest.thy b/thys/Sigma_Commit_Crypto/Rivest.thy --- a/thys/Sigma_Commit_Crypto/Rivest.thy +++ b/thys/Sigma_Commit_Crypto/Rivest.thy @@ -1,363 +1,362 @@ subsection\Rivest Commitment Scheme\ text\The Rivest commitment scheme was first introduced in \<^cite>\"rivest1999"\. We note however the original scheme did not allow for perfect hiding. This was pointed out by Blundo and Masucci in \<^cite>\"DBLP:journals/dcc/BlundoMSW02"\ who alightly ammended the commitment scheme so that is provided perfect hiding. The Rivest commitment scheme uses a trusted initialiser to provide correlated randomness to the two parties before an execution of the protocol. In our framework we set these as keys that held by the respective parties.\ theory Rivest imports Commitment_Schemes "HOL-Number_Theory.Cong" CryptHOL.CryptHOL Cyclic_Group_Ext Discrete_Log Number_Theory_Aux Uniform_Sampling begin locale rivest = fixes q :: nat assumes prime_q: "prime q" begin lemma q_gt_0 [simp]: "q > 0" by (simp add: prime_q prime_gt_0_nat) type_synonym ck = "nat \ nat" type_synonym vk = "nat \ nat" type_synonym plain = "nat" type_synonym commit = "nat" type_synonym "opening" = "nat \ nat" definition key_gen :: "(ck \ vk) spmf" where "key_gen = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; x1 :: nat \ sample_uniform q; let y1 = (a * x1 + b) mod q; return_spmf ((a,b), (x1,y1))}" definition commit :: "ck \ plain \ (commit \ opening) spmf" where "commit ck m = do { let (a,b) = ck; return_spmf ((m + a) mod q, (a,b))}" fun verify :: "vk \ plain \ commit \ opening \ bool" where "verify (x1,y1) m c (a,b) = (((c = (m + a) mod q)) \ (y1 = (a * x1 + b) mod q))" definition valid_msg :: "plain \ bool" where "valid_msg msg \ msg \ {..< q}" sublocale rivest_commit: abstract_commitment key_gen commit verify valid_msg . lemma abstract_correct: "rivest_commit.correct" unfolding abstract_commitment.correct_def abstract_commitment.correct_game_def by(simp add: key_gen_def commit_def bind_spmf_const lossless_weight_spmfD) lemma rivest_hiding: "(spmf (rivest_commit.hiding_game_ind_cpa \) True - 1/2 = 0)" including monad_normalisation proof- note [simp] = Let_def split_def obtain \1 \2 where [simp]: "\ = (\1, \2)" by(cases \) have "rivest_commit.hiding_game_ind_cpa (\1, \2) = TRY do { a :: nat \ sample_uniform q; x1 :: nat \ sample_uniform q; y1 \ map_spmf (\ b. (a * x1 + b) mod q) (sample_uniform q); ((m0, m1), \) \ \1 (x1,y1); _ :: unit \ assert_spmf (valid_msg m0 \ valid_msg m1); d \ coin_spmf; let c = ((if d then m0 else m1) + a) mod q; b' \ \2 c \; return_spmf (b' = d)} ELSE coin_spmf" unfolding abstract_commitment.hiding_game_ind_cpa_def by(simp add: commit_def key_gen_def o_def bind_map_spmf) also have "... = TRY do { a :: nat \ sample_uniform q; x1 :: nat \ sample_uniform q; y1 \ sample_uniform q; ((m0, m1), \) \ \1 (x1,y1); _ :: unit \ assert_spmf (valid_msg m0 \ valid_msg m1); d \ coin_spmf; let c = ((if d then m0 else m1) + a) mod q; b' \ \2 c \; return_spmf (b' = d)} ELSE coin_spmf" by(simp add: samp_uni_plus_one_time_pad) also have "... = TRY do { x1 :: nat \ sample_uniform q; y1 \ sample_uniform q; ((m0, m1), \) \ \1 (x1,y1); _ :: unit \ assert_spmf (valid_msg m0 \ valid_msg m1); d \ coin_spmf; c \ map_spmf (\ a. ((if d then m0 else m1) + a) mod q) (sample_uniform q); b' \ \2 c \; return_spmf (b' = d)} ELSE coin_spmf" by(simp add: o_def bind_map_spmf) also have "... = TRY do { x1 :: nat \ sample_uniform q; y1 \ sample_uniform q; ((m0, m1), \) \ \1 (x1,y1); _ :: unit \ assert_spmf (valid_msg m0 \ valid_msg m1); d \ coin_spmf; c \ sample_uniform q; b' :: bool \ \2 c \; return_spmf (b' = d)} ELSE coin_spmf" by(simp add: samp_uni_plus_one_time_pad) also have "... = TRY do { x1 :: nat \ sample_uniform q; y1 \ sample_uniform q; ((m0, m1), \) \ \1 (x1,y1); _ :: unit \ assert_spmf (valid_msg m0 \ valid_msg m1); c :: nat \ sample_uniform q; guess :: bool \ \2 c \; map_spmf((=) guess) coin_spmf} ELSE coin_spmf" by(simp add: map_spmf_conv_bind_spmf) also have "... = coin_spmf" by(simp add: map_eq_const_coin_spmf bind_spmf_const try_bind_spmf_lossless2' scale_bind_spmf weight_spmf_le_1 scale_scale_spmf) ultimately show ?thesis by(simp add: spmf_of_set) qed lemma rivest_perfect_hiding: "rivest_commit.perfect_hiding_ind_cpa \" unfolding abstract_commitment.perfect_hiding_ind_cpa_def abstract_commitment.hiding_advantage_ind_cpa_def by(simp add: rivest_hiding) lemma samp_uni_break': assumes fst_cond: "m \ m' \ valid_msg m \ valid_msg m'" and c: "c = (m + a) mod q \ y1 = (a * x1 + b) mod q" and c': "c = (m' + a') mod q \ y1 = (a' * x1 + b') mod q" and x1: "x1 < q" shows "x1 = (if (a mod q > a' mod q) then nat ((int b'- int b) * (inverse (nat ((int a mod q - int a' mod q) mod q)) q) mod q) else nat ((int b- int b') * (inverse (nat ((int a' mod q - int a mod q) mod q)) q) mod q))" proof- have m: "m < q \ m' < q" using fst_cond valid_msg_def by simp have a_a': "\ [a = a'] (mod q)" proof- have "[m + a = m' + a'] (mod q)" using assms cong_def by blast thus ?thesis by (metis m fst_cond c c' add.commute cong_less_modulus_unique_nat cong_add_rcancel_nat cong_mod_right) qed have cong_y1: "[int a * int x1 + int b = int a' * int x1 + int b'] (mod q)" by (metis c c' cong_def Num.of_nat_simps(4) Num.of_nat_simps(5) cong_int_iff) show ?thesis - proof(cases "a mod q > a' mod q") + proof (cases "a mod q > a' mod q") case True - hence gcd: "gcd (nat ((int a mod q - int a' mod q) mod q)) q = 1" - proof- - have "((int a mod q - int a' mod q) mod q) \ 0" - by (metis True comm_monoid_add_class.add_0 diff_add_cancel mod_add_left_eq mod_diff_eq nat_mod_as_int order_less_irrefl) - moreover have "((int a mod q - int a' mod q) mod q) < q" by simp - ultimately show ?thesis - using prime_field[of q "nat ((int a mod int q - int a' mod int q) mod int q)"] prime_q - by (smt Euclidean_Division.pos_mod_sign coprime_imp_gcd_eq_1 int_nat_eq nat_less_iff of_nat_0_less_iff q_gt_0) - qed + moreover have \((int a mod q - int a' mod q) mod q) \ 0\ + by (metis True comm_monoid_add_class.add_0 diff_add_cancel mod_add_left_eq mod_diff_eq nat_mod_as_int order_less_irrefl) + moreover have "((int a mod q - int a' mod q) mod q) < q" by simp + ultimately have \coprime (nat ((int a mod q - int a' mod q) mod q)) q\ + using prime_field [of q \nat ((int a mod int q - int a' mod int q) mod int q)\] prime_q + by (simp flip: of_nat_mod of_nat_diff) + then have gcd: "gcd (nat ((int a mod q - int a' mod q) mod q)) q = 1" + by simp hence "[int a * int x1 - int a' * int x1 = int b'- int b] (mod q)" by (smt cong_diff_iff_cong_0 cong_y1 cong_diff cong_diff) hence "[int a mod q * int x1 - int a' mod q * int x1 = int b'- int b] (mod q)" proof - have "[int x1 * (int a mod int q - int a' mod int q) = int x1 * (int a - int a')] (mod int q)" by (meson cong_def cong_mult cong_refl mod_diff_eq) then show ?thesis by (metis (no_types, opaque_lifting) Groups.mult_ac(2) \[int a * int x1 - int a' * int x1 = int b' - int b] (mod int q)\ cong_def mod_diff_left_eq mod_diff_right_eq mod_mult_right_eq) qed hence "[int x1 * (int a mod q - int a' mod q) = int b'- int b] (mod q)" by(metis int_distrib(3) mult.commute) hence "[int x1 * (int a mod q - int a' mod q) mod q = int b'- int b] (mod q)" using cong_def by simp hence "[int x1 * nat ((int a mod q - int a' mod q) mod q) = int b'- int b] (mod q)" by (simp add: True cong_def mod_mult_right_eq) hence "[int x1 * nat ((int a mod q - int a' mod q) mod q) * inverse (nat ((int a mod q - int a' mod q) mod q)) q = (int b'- int b) * inverse (nat ((int a mod q - int a' mod q) mod q)) q] (mod q)" using cong_scalar_right by blast hence "[int x1 * (nat ((int a mod q - int a' mod q) mod q) * inverse (nat ((int a mod q - int a' mod q) mod q)) q) = (int b'- int b) * inverse (nat ((int a mod q - int a' mod q) mod q)) q] (mod q)" by (simp add: more_arith_simps(11)) hence "[int x1 * 1 = (int b'- int b) * inverse (nat ((int a mod q - int a' mod q) mod q)) q] (mod q)" using inverse gcd by (meson cong_scalar_left cong_sym_eq cong_trans) hence "[int x1 = (int b'- int b) * inverse (nat ((int a mod q - int a' mod q) mod q)) q] (mod q)" by simp hence "int x1 mod q = ((int b'- int b) * inverse (nat ((int a mod q - int a' mod q) mod q)) q) mod q" using cong_def by fast thus ?thesis using x1 True by simp next case False hence aa': "a mod q < a' mod q" using a_a' cong_refl nat_neq_iff by (simp add: cong_def) - hence gcd: "gcd (nat ((int a' mod q - int a mod q) mod q)) q = 1" - proof- - have "((int a' mod q - int a mod q) mod q) \ 0" - by (metis aa' comm_monoid_add_class.add_0 diff_add_cancel mod_add_left_eq mod_diff_eq nat_mod_as_int order_less_irrefl) - moreover have "((int a' mod q - int a mod q) mod q) < q" by simp - ultimately show ?thesis - using prime_field[of q "nat ((int a' mod int q - int a mod int q) mod int q)"] prime_q - by (smt Euclidean_Division.pos_mod_sign coprime_imp_gcd_eq_1 int_nat_eq nat_less_iff of_nat_0_less_iff q_gt_0) - qed + moreover have "((int a' mod q - int a mod q) mod q) \ 0" + by (metis aa' comm_monoid_add_class.add_0 diff_add_cancel mod_add_left_eq mod_diff_eq nat_mod_as_int order_less_irrefl) + moreover have "((int a' mod q - int a mod q) mod q) < q" by simp + ultimately have \coprime (nat ((int a' mod q - int a mod q) mod q)) q\ + using prime_field [of q \nat ((int a' mod int q - int a mod int q) mod int q)\] prime_q + by (simp flip: of_nat_mod of_nat_diff) + then have gcd: "gcd (nat ((int a' mod q - int a mod q) mod q)) q = 1" + by simp have "[int b - int b' = int a' * int x1 - int a * int x1] (mod q)" by (smt cong_diff_iff_cong_0 cong_y1 cong_diff cong_diff) hence "[int b - int b' = int x1 * (int a' - int a)] (mod q)" using int_distrib mult.commute by metis hence "[int b - int b' = int x1 * (int a' mod q - int a mod q)] (mod q)" by (metis (no_types, lifting) cong_def mod_diff_eq mod_mult_right_eq) hence "[int b - int b' = int x1 * (int a' mod q - int a mod q) mod q] (mod q)" using cong_def by simp hence "[(int b - int b') * inverse (nat ((int a' mod q - int a mod q) mod q)) q = int x1 * (int a' mod q - int a mod q) mod q * inverse (nat ((int a' mod q - int a mod q) mod q)) q ] (mod q)" using cong_scalar_right by blast hence "[(int b - int b') * inverse (nat ((int a' mod q - int a mod q) mod q)) q = int x1 * ((int a' mod q - int a mod q) mod q * inverse (nat ((int a' mod q - int a mod q) mod q)) q)] (mod q)" by (metis (mono_tags, lifting) cong_def mod_mult_left_eq mod_mult_right_eq more_arith_simps(11)) hence *: "[int x1 * ((int a' mod q - int a mod q) mod q * inverse (nat ((int a' mod q - int a mod q) mod q)) q) = (int b - int b') * inverse (nat ((int a' mod q - int a mod q) mod q)) q] (mod q)" using cong_sym_eq by auto hence "[int x1 * 1 = (int b - int b') * inverse (nat ((int a' mod q - int a mod q) mod q)) q] (mod q)" proof - have "[(int a' mod int q - int a mod int q) mod int q * Number_Theory_Aux.inverse (nat ((int a' mod int q - int a mod int q) mod int q)) q = 1] (mod int q)" - by (metis (no_types) Euclidean_Division.pos_mod_sign inverse gcd int_nat_eq of_nat_0_less_iff q_gt_0) + using inverse [of \nat ((int a' mod int q - int a mod int q) mod int q)\ q, OF gcd] + by simp then show ?thesis by (meson * cong_scalar_left cong_sym_eq cong_trans) qed hence "[int x1 = (int b - int b') * inverse (nat ((int a' mod q - int a mod q) mod q)) q] (mod q)" by simp hence "int x1 mod q = (int b - int b') * (inverse (nat ((int a' mod q - int a mod q) mod q)) q) mod q" using cong_def by auto thus ?thesis using x1 aa' by simp qed qed lemma samp_uni_spmf_mod_q: shows "spmf (sample_uniform q) (x mod q) = 1/q" proof- have "indicator {..< q} (x mod q) = 1" using q_gt_0 by auto moreover have "real (card {..< q}) = q" by simp ultimately show ?thesis by(auto simp add: spmf_of_set sample_uniform_def) qed lemma spmf_samp_uni_eq_return_bool_mod: shows "spmf (do { x1 \ sample_uniform q; return_spmf (int x1 = y mod q)}) True = 1/q" proof- have "spmf (do { x1 \ sample_uniform q; return_spmf (x1 = y mod q)}) True = spmf (sample_uniform q \ (\ x1. return_spmf x1)) (y mod q)" apply(simp only: spmf_bind) apply(rule Bochner_Integration.integral_cong[OF refl])+ proof - fix x :: nat have "y mod q = x \ indicator {True} (x = (y mod q)) = (indicator {(y mod q)} x::real)" by simp then have "indicator {True} (x = y mod q) = (indicator {y mod q} x::real)" by fastforce then show "spmf (return_spmf (x = y mod q)) True = spmf (return_spmf x) (y mod q)" by (metis pmf_return spmf_of_pmf_return_pmf spmf_spmf_of_pmf) qed thus ?thesis using samp_uni_spmf_mod_q by simp qed lemma bind_game_le_inv_q: shows "spmf (rivest_commit.bind_game \) True \ 1 / q" proof - let ?eq = "\a a' b b'. (=) (if (a mod q > a' mod q) then nat ((int b'- int b) * (inverse (nat ((int a mod q - int a' mod q) mod q)) q) mod q) else nat ((int b - int b') * (inverse (nat ((int a' mod q - int a mod q) mod q)) q) mod q))" have "spmf (rivest_commit.bind_game \) True = spmf (do { (ck,(x1,y1)) \ key_gen; (c, m, (a,b), m', (a',b')) \ \ ck; _ :: unit \ assert_spmf(m \ m' \ valid_msg m \ valid_msg m'); let b = verify (x1,y1) m c (a,b); let b' = verify (x1,y1) m' c (a',b'); _ :: unit \ assert_spmf (b \ b'); return_spmf True}) True" by(simp add: abstract_commitment.bind_game_alt_def split_def spmf_try_spmf del: verify.simps) also have "... = spmf (do { a' :: nat \ sample_uniform q; b' :: nat \ sample_uniform q; x1 :: nat \ sample_uniform q; let y1 = (a' * x1 + b') mod q; (c, m, (a,b), m', (a',b')) \ \ (a',b'); _ :: unit \ assert_spmf (m \ m' \ valid_msg m \ valid_msg m'); _ :: unit \ assert_spmf (c = (m + a) mod q \ y1 = (a * x1 + b) mod q \ c = (m' + a') mod q \ y1 = (a' * x1 + b') mod q); return_spmf True}) True" by(simp add: key_gen_def Let_def) also have "... = spmf (do { a'' :: nat \ sample_uniform q; b'' :: nat \ sample_uniform q; x1 :: nat \ sample_uniform q; let y1 = (a'' * x1 + b'') mod q; (c, m, (a,b), m', (a',b')) \ \ (a'',b''); _ :: unit \ assert_spmf (m \ m' \ valid_msg m \ valid_msg m'); _ :: unit \ assert_spmf (c = (m + a) mod q \ y1 = (a * x1 + b) mod q \ c = (m' + a') mod q \ y1 = (a' * x1 + b') mod q); return_spmf (?eq a a' b b' x1)}) True" unfolding split_def Let_def by(rule arg_cong2[where f=spmf, OF _ refl] bind_spmf_cong[OF refl])+ (auto simp add: eq_commute samp_uni_break' Let_def split_def valid_msg_def cong: bind_spmf_cong_simp) also have "... \ spmf (do { a'' :: nat \ sample_uniform q; b'' :: nat \ sample_uniform q; (c, m, (a,(b::nat)), m', (a',b')) \ \ (a'',b''); map_spmf (?eq a a' b b') (sample_uniform q)}) True" including monad_normalisation unfolding split_def Let_def assert_spmf_def apply(simp add: map_spmf_conv_bind_spmf) apply(rule ord_spmf_eq_leD ord_spmf_bind_reflI)+ apply auto done also have "... \ 1/q" proof((rule spmf_bind_leI)+, clarify) fix a a' b b' define A where "A = Collect (?eq a a' b b')" define x1 where "x1 = The (?eq a a' b b')" note q_gt_0[simp del] have "A \ {x1}" by(auto simp add: A_def x1_def) hence "card (A \ {.. card {x1}" by(intro card_mono) auto also have "\ = 1" by simp finally have "spmf (map_spmf (\x. x \ A) (sample_uniform q)) True \ 1 / q" using q_gt_0 unfolding sample_uniform_def by(subst map_mem_spmf_of_set)(auto simp add: field_simps) then show "spmf (map_spmf (?eq a a' b b') (sample_uniform q)) True \ 1 / q" unfolding A_def mem_Collect_eq . qed auto finally show ?thesis . qed lemma rivest_bind: shows "rivest_commit.bind_advantage \ \ 1 / q" using bind_game_le_inv_q rivest_commit.bind_advantage_def by simp end locale rivest_asymp = fixes q :: "nat \ nat" assumes rivest: "\\. rivest (q \)" begin sublocale rivest "q \" for \ by(simp add: rivest) theorem rivest_correct: shows "rivest_commit.correct n" using abstract_correct by simp theorem rivest_perfect_hiding_asym: assumes lossless_\: "rivest_commit.lossless (\ n)" shows "rivest_commit.perfect_hiding_ind_cpa n (\ n)" by (simp add: lossless_\ rivest_perfect_hiding) theorem rivest_binding_asym: assumes "negligible (\n. 1 / (q n))" shows "negligible (\n. rivest_commit.bind_advantage n (\ n))" using negligible_le rivest_bind assms rivest_commit.bind_advantage_def by auto end end diff --git a/thys/Sigma_Commit_Crypto/Schnorr_Sigma_Commit.thy b/thys/Sigma_Commit_Crypto/Schnorr_Sigma_Commit.thy --- a/thys/Sigma_Commit_Crypto/Schnorr_Sigma_Commit.thy +++ b/thys/Sigma_Commit_Crypto/Schnorr_Sigma_Commit.thy @@ -1,401 +1,401 @@ subsection\Schnorr \\\-protocol\ text\In this section we show the Schnoor protocol \<^cite>\"DBLP:journals/joc/Schnorr91"\ is a \\\-protocol and then use it to construct a commitment scheme. The security statements for the resulting commitment scheme come for free from our general proof of the construction.\ theory Schnorr_Sigma_Commit imports Commitment_Schemes Sigma_Protocols Cyclic_Group_Ext Discrete_Log Number_Theory_Aux Uniform_Sampling "HOL-Number_Theory.Cong" begin locale schnorr_base = fixes \ :: "'grp cyclic_group" (structure) assumes prime_order: "prime (order \)" begin lemma order_gt_0 [simp]: "order \ > 0" using prime_order prime_gt_0_nat by blast text\The types for the \\\-protocol.\ type_synonym witness = "nat" type_synonym rand = nat type_synonym 'grp' msg = "'grp'" type_synonym response = nat type_synonym challenge = nat type_synonym 'grp' pub_in = "'grp'" definition R_DL :: "('grp pub_in \ witness) set" where "R_DL = {(h, w). h = \<^bold>g [^] w}" definition init :: "'grp pub_in \ witness \ (rand \ 'grp msg) spmf" where "init h w = do { r \ sample_uniform (order \); return_spmf (r, \<^bold>g [^] r)}" lemma lossless_init: "lossless_spmf (init h w)" by(simp add: init_def) definition "response r w c = return_spmf ((w*c + r) mod (order \))" lemma lossless_response: "lossless_spmf (response r w c)" by(simp add: response_def) definition G :: "('grp pub_in \ witness) spmf" where "G = do { w \ sample_uniform (order \); return_spmf (\<^bold>g [^] w, w)}" lemma lossless_G: "lossless_spmf G" by(simp add: G_def) definition "challenge_space = {..< order \}" definition check :: "'grp pub_in \ 'grp msg \ challenge \ response \ bool" where "check h a e z = (a \ (h [^] e) = \<^bold>g [^] z \ a \ carrier \)" definition S2 :: "'grp \ challenge \ ('grp msg, response) sim_out spmf" where "S2 h e = do { c \ sample_uniform (order \); let a = \<^bold>g [^] c \ (inv (h [^] e)); return_spmf (a, c)}" definition ss_adversary :: "'grp \ ('grp msg, challenge, response) conv_tuple \ ('grp msg, challenge, response) conv_tuple \ nat spmf" where "ss_adversary x c1 c2 = do { let (a, e, z) = c1; let (a', e', z') = c2; return_spmf (if (e > e') then (nat ((int z - int z') * inverse ((e - e')) (order \) mod order \)) else (nat ((int z' - int z) * inverse ((e' - e)) (order \) mod order \)))}" definition "valid_pub = carrier \" text\We now use the Schnorr \\\-protocol use Schnorr to construct a commitment scheme.\ type_synonym 'grp' ck = "'grp'" type_synonym 'grp' vk = "'grp' \ nat" type_synonym plain = "nat" type_synonym 'grp' commit = "'grp'" type_synonym "opening" = "nat" text\The adversary we use in the discrete log game to reduce the binding property to the discrete log assumption.\ definition dis_log_\ :: "('grp ck, plain, 'grp commit, opening) bind_adversary \ 'grp ck \ nat spmf" where "dis_log_\ \ h = do { (c, e, z, e', z') \ \ h; _ :: unit \ assert_spmf (e > e' \ \ [e = e'] (mod order \) \ (gcd (e - e') (order \) = 1) \ c \ carrier \); _ :: unit \ assert_spmf (((c \ h [^] e) = \<^bold>g [^] z) \ (c \ h [^] e') = \<^bold>g [^] z'); return_spmf (nat ((int z - int z') * inverse ((e - e')) (order \) mod order \))}" sublocale discrete_log: dis_log \ unfolding dis_log_def by simp end locale schnorr_sigma_protocol = schnorr_base + cyclic_group \ begin sublocale Schnorr_\: \_protocols_base init response check R_DL S2 ss_adversary challenge_space valid_pub apply unfold_locales by(simp add: R_DL_def valid_pub_def; blast) text\The Schnorr \\\-protocol is complete.\ lemma completeness: "Schnorr_\.completeness" proof- have "\<^bold>g [^] y \ (\<^bold>g [^] w') [^] e = \<^bold>g [^] (y + w' * e)" for y e w' :: nat using nat_pow_pow nat_pow_mult by simp then show ?thesis unfolding Schnorr_\.completeness_game_def Schnorr_\.completeness_def by(auto simp add: init_def response_def check_def pow_generator_mod R_DL_def add.commute bind_spmf_const) qed text\The next two lemmas help us rewrite terms in the proof of honest verfier zero knowledge.\ lemma zr_rewrite: assumes z: "z = (x*c + r) mod (order \)" and r: "r < order \" shows "(z + (order \)*x*c - x*c) mod (order \) = r" proof(cases "x = 0") case True then show ?thesis using assms by simp next case x_neq_0: False then show ?thesis proof(cases "c = 0") case True then show ?thesis by (simp add: assms) next case False have cong: "[z + (order \)*x*c = x*c + r] (mod (order \))" by (simp add: cong_def mult.assoc z) hence "[z + (order \)*x*c - x*c = r] (mod (order \))" proof- have "z + (order \)*x*c > x*c" by (metis One_nat_def mult_less_cancel2 n_less_m_mult_n neq0_conv prime_gt_1_nat prime_order trans_less_add2 x_neq_0 False) then show ?thesis by (metis cong add_diff_inverse_nat cong_add_lcancel_nat less_imp_le linorder_not_le) qed then show ?thesis by(simp add: cong_def r) qed qed lemma h_sub_rewrite: assumes "h = \<^bold>g [^] x" and z: "z < order \" shows "\<^bold>g [^] ((z + (order \)*x*c - x*c)) = \<^bold>g [^] z \ inv (h [^] c)" (is "?lhs = ?rhs") proof(cases "x = 0") case True then show ?thesis using assms by simp next case x_neq_0: False then show ?thesis proof- have "(z + order \ * x * c - x * c) = (z + (order \ * x * c - x * c))" using z by (simp add: less_imp_le_nat mult_le_mono) then have lhs: "?lhs = \<^bold>g [^] z \ \<^bold>g [^] ((order \)*x*c - x*c)" by(simp add: nat_pow_mult) have " \<^bold>g [^] ((order \)*x*c - x*c) = inv (h [^] c)" proof(cases "c = 0") case True then show ?thesis by simp next case False hence bound: "((order \)*x*c - x*c) > 0" using assms x_neq_0 prime_gt_1_nat prime_order by auto then have "\<^bold>g [^] ((order \)*x*c- x*c) = \<^bold>g [^] int ((order \)*x*c - x*c)" by (simp add: int_pow_int) also have "... = \<^bold>g [^] int ((order \)*x*c) \ inv (\<^bold>g [^] (x*c))" by (metis bound generator_closed int_ops(6) int_pow_int of_nat_eq_0_iff of_nat_less_0_iff of_nat_less_iff int_pow_diff) also have "... = \<^bold>g [^] ((order \)*x*c) \ inv (\<^bold>g [^] (x*c))" by (metis int_pow_int) also have "... = \<^bold>g [^] ((order \)*x*c) \ inv ((\<^bold>g [^] x) [^] c)" by(simp add: nat_pow_pow) also have "... = \<^bold>g [^] ((order \)*x*c) \ inv (h [^] c)" using assms by simp also have "... = \ \ inv (h [^] c)" using generator_pow_order by (metis generator_closed mult_is_0 nat_pow_0 nat_pow_pow) ultimately show ?thesis by (simp add: assms(1)) qed then show ?thesis using lhs by simp qed qed lemma hvzk_R_rewrite_grp: fixes x c r :: nat assumes "r < order \" shows "\<^bold>g [^] (((x * c + order \ - r) mod order \ + order \ * x * c - x * c) mod order \) = inv \<^bold>g [^] r" (is "?lhs = ?rhs") proof- have "[(x * c + order \ - r) mod order \ + order \ * x * c - x * c = order \ - r] (mod order \)" proof- have "[(x * c + order \ - r) mod order \ + order \ * x * c - x * c = x * c + order \ - r + order \ * x * c - x * c] (mod order \)" by (smt cong_def One_nat_def add_diff_inverse_nat cong_diff_nat less_imp_le_nat linorder_not_less mod_add_left_eq mult.assoc n_less_m_mult_n prime_gt_1_nat prime_order trans_less_add2 zero_less_diff) hence "[(x * c + order \ - r) mod order \ + order \ * x * c - x * c = order \ - r + order \ * x * c] (mod order \)" using assms by auto thus ?thesis by (simp add: cong_def mult.assoc) qed hence "\<^bold>g [^] ((x * c + order \ - r) mod order \ + order \ * x * c - x * c) = \<^bold>g [^] (order \ - r)" using finite_carrier pow_generator_eq_iff_cong by blast thus ?thesis using neg_power_inverse by (simp add: assms inverse_pow_pow pow_generator_mod) qed lemma hv_zk: assumes "(h,x) \ R_DL" shows "Schnorr_\.R h x c = Schnorr_\.S h c" including monad_normalisation proof- have "Schnorr_\.R h x c = do { r \ sample_uniform (order \); let z = (x*c + r) mod (order \); let a = \<^bold>g [^] ((z + (order \)*x*c - x*c) mod (order \)); return_spmf (a,c,z)}" apply(simp add: Let_def Schnorr_\.R_def init_def response_def) using assms zr_rewrite R_DL_def by(simp cong: bind_spmf_cong_simp) also have "... = do { z \ map_spmf (\ r. (x*c + r) mod (order \)) (sample_uniform (order \)); let a = \<^bold>g [^] ((z + (order \)*x*c - x*c) mod (order \)); return_spmf (a,c,z)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { z \ (sample_uniform (order \)); let a = \<^bold>g [^] ((z + (order \)*x*c - x*c)); return_spmf (a,c,z)}" by(simp add: samp_uni_plus_one_time_pad pow_generator_mod) also have "... = do { z \ (sample_uniform (order \)); let a = \<^bold>g [^] z \ inv (h [^] c); return_spmf (a,c,z)}" using h_sub_rewrite assms R_DL_def by(simp cong: bind_spmf_cong_simp) ultimately show ?thesis by(simp add: Schnorr_\.S_def S2_def map_spmf_conv_bind_spmf) qed text\We can now prove that honest verifier zero knowledge holds for the Schnorr \\\-protocol.\ lemma honest_verifier_ZK: shows "Schnorr_\.HVZK" unfolding Schnorr_\.HVZK_def by(auto simp add: hv_zk R_DL_def S2_def check_def valid_pub_def challenge_space_def cyclic_group_assoc) text\It is left to prove the special soundness property. First we prove a lemma we use to rewrite a term in the special soundness proof and then prove the property itself.\ lemma ss_rewrite: assumes "e' < e" and "e < order \" and a_mem:"a \ carrier \" and h_mem: "h \ carrier \" and a: "a \ h [^] e = \<^bold>g [^] z" and a': "a \ h [^] e' = \<^bold>g [^] z'" shows "h = \<^bold>g [^] ((int z - int z') * inverse ((e - e')) (order \) mod int (order \))" proof- have gcd: "gcd (nat (int e - int e') mod (order \)) (order \) = 1" using prime_field by (metis Primes.prime_nat_def assms(1) assms(2) coprime_imp_gcd_eq_1 diff_is_0_eq less_imp_diff_less mod_less nat_minus_as_int not_less schnorr_base.prime_order schnorr_base_axioms) have "a = \<^bold>g [^] z \ inv (h [^] e)" using a a_mem by (simp add: h_mem group.inv_solve_right) moreover have "a = \<^bold>g [^] z' \ inv (h [^] e')" using a' a_mem by (simp add: h_mem group.inv_solve_right) ultimately have "\<^bold>g [^] z \ h [^] e' = \<^bold>g [^] z' \ h [^] e" using h_mem by (metis (no_types, lifting) a a' h_mem a_mem cyclic_group_assoc cyclic_group_commute nat_pow_closed) moreover obtain t :: nat where t: "h = \<^bold>g [^] t" using h_mem generatorE by blast ultimately have "\<^bold>g [^] (z + t * e') = \<^bold>g [^] (z' + t * e) " by (simp add: monoid.nat_pow_mult nat_pow_pow) hence "[z + t * e' = z' + t * e] (mod order \)" using group_eq_pow_eq_mod order_gt_0 by blast hence "[int z + int t * int e' = int z' + int t * int e] (mod order \)" using cong_int_iff by force hence "[int z - int z' = int t * int e - int t * int e'] (mod order \)" by (smt cong_iff_lin) hence "[int z - int z' = int t * (int e - int e')] (mod order \)" by (simp add: \[int z - int z' = int t * int e - int t * int e'] (mod int (order \))\ right_diff_distrib) hence "[int z - int z' = int t * (int e - int e')] (mod order \)" by (meson cong_diff cong_mod_left cong_mult cong_refl cong_trans) hence *: "[int z - int z' = int t * (int e - int e')] (mod order \)" using assms by (simp add: int_ops(9) of_nat_diff) hence "[int z - int z' = int t * nat (int e - int e')] (mod order \)" using assms by auto hence **: "[(int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)) = int t * (nat (int e - int e') * fst (bezw ((nat (int e - int e'))) (order \)))] (mod order \)" by (smt \[int z - int z' = int t * (int e - int e')] (mod int (order \))\ assms(1) assms(2) cong_scalar_right int_nat_eq less_imp_of_nat_less mod_less more_arith_simps(11) nat_less_iff of_nat_0_le_iff) hence "[(int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)) = int t * 1] (mod order \)" by (metis (no_types, opaque_lifting) gcd inverse assms(2) cong_scalar_left cong_trans less_imp_diff_less mod_less mult.comm_neutral nat_minus_as_int) hence "[(int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)) = t] (mod order \)" by simp hence "[ ((int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)))mod order \ = t] (mod order \)" using cong_mod_left by blast hence **: "[nat (((int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)))mod order \) = t] (mod order \)" - by (metis Euclidean_Division.pos_mod_sign cong_int_iff int_nat_eq of_nat_0_less_iff order_gt_0) + by (metis cong_def mod_mod_trivial nat_int of_nat_mod) hence "\<^bold>g [^] (nat (((int z - int z') * fst (bezw ((nat (int e - int e'))) (order \)))mod order \)) = \<^bold>g [^] t" using cyclic_group.pow_generator_eq_iff_cong cyclic_group_axioms order_gt_0 order_gt_0_iff_finite by blast - thus ?thesis using t - by (smt Euclidean_Division.pos_mod_sign discrete_log.order_gt_0 int_pow_def2 nat_minus_as_int of_nat_0_less_iff) + thus ?thesis using t + by (simp add: nat_minus_as_int) qed text\The special soundness property for the Schnorr \\\-protocol.\ lemma special_soundness: shows "Schnorr_\.special_soundness" unfolding Schnorr_\.special_soundness_def by(auto simp add: valid_pub_def ss_rewrite challenge_space_def split_def ss_adversary_def check_def R_DL_def Let_def) text\We are now able to prove that the Schnorr \\\-protocol is a \\\-protocol, the proof comes from the properties of completeness, HVZK and special soundness we have previously proven.\ theorem sigma_protocol: shows "Schnorr_\.\_protocol" by(simp add: Schnorr_\.\_protocol_def completeness honest_verifier_ZK special_soundness) text\Having proven the \\\-protocol property is satisfied we can show the commitment scheme we construct from the Schnorr \\\-protocol has the desired properties. This result comes with very little proof effort as we can instantiate our general proof.\ sublocale Schnorr_\_commit: \_protocols_to_commitments init response check R_DL S2 ss_adversary challenge_space valid_pub G unfolding \_protocols_to_commitments_def \_protocols_to_commitments_axioms_def apply(auto simp add: \_protocols_base_def) apply(simp add: R_DL_def valid_pub_def) apply(auto simp add: sigma_protocol lossless_G lossless_init lossless_response) by(simp add: R_DL_def G_def) lemma "Schnorr_\_commit.abstract_com.correct" by(fact Schnorr_\_commit.commit_correct) lemma "Schnorr_\_commit.abstract_com.perfect_hiding_ind_cpa \" by(fact Schnorr_\_commit.perfect_hiding) lemma rel_adv_eq_dis_log_adv: "Schnorr_\_commit.rel_advantage \ = discrete_log.advantage \" proof- have "Schnorr_\_commit.rel_game \ = discrete_log.dis_log \" unfolding Schnorr_\_commit.rel_game_def discrete_log.dis_log_def by(auto intro: try_spmf_cong bind_spmf_cong[OF refl] simp add: G_def R_DL_def cong_less_modulus_unique_nat group_eq_pow_eq_mod finite_carrier pow_generator_eq_iff_cong) thus ?thesis using Schnorr_\_commit.rel_advantage_def discrete_log.advantage_def by simp qed lemma bind_advantage_bound_dis_log: "Schnorr_\_commit.abstract_com.bind_advantage \ \ discrete_log.advantage (Schnorr_\_commit.adversary \)" using Schnorr_\_commit.bind_advantage rel_adv_eq_dis_log_adv by simp end locale schnorr_asymp = fixes \ :: "nat \ 'grp cyclic_group" assumes schnorr: "\\. schnorr_sigma_protocol (\ \)" begin sublocale schnorr_sigma_protocol "\ \" for \ by(simp add: schnorr) text\The \\\-protocol statement comes easily in the asymptotic setting.\ theorem sigma_protocol: shows "Schnorr_\.\_protocol n" by(simp add: sigma_protocol) text\We now show the statements of security for the commitment scheme in the asymptotic setting, the main difference is that we are able to show the binding advantage is negligible in the security parameter.\ lemma asymp_correct: "Schnorr_\_commit.abstract_com.correct n" using Schnorr_\_commit.commit_correct by simp lemma asymp_perfect_hiding: "Schnorr_\_commit.abstract_com.perfect_hiding_ind_cpa n (\ n)" using Schnorr_\_commit.perfect_hiding by blast lemma asymp_computational_binding: assumes "negligible (\ n. discrete_log.advantage n (Schnorr_\_commit.adversary n (\ n)))" shows "negligible (\ n. Schnorr_\_commit.abstract_com.bind_advantage n (\ n))" using Schnorr_\_commit.bind_advantage assms Schnorr_\_commit.abstract_com.bind_advantage_def negligible_le bind_advantage_bound_dis_log by auto end end diff --git a/thys/Solidity/Constant_Folding.thy b/thys/Solidity/Constant_Folding.thy --- a/thys/Solidity/Constant_Folding.thy +++ b/thys/Solidity/Constant_Folding.thy @@ -1,6016 +1,6016 @@ section\Constant Folding\ theory Constant_Folding imports Solidity_Main begin text\ The following function optimizes expressions w.r.t. gas consumption. \ -fun eupdate :: "E \ E" (* Processing of this definition takes several minutes. *) +primrec eupdate :: "E \ E" and lupdate :: "L \ L" where "lupdate (Id i) = Id i" | "lupdate (Ref i exp) = Ref i (map eupdate exp)" | "eupdate (E.INT b v) = (if (b\vbits) then if v \ 0 then E.INT b (-(2^(b-1)) + (v+2^(b-1)) mod (2^b)) else E.INT b (2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1) else E.INT b v)" | "eupdate (UINT b v) = (if (b\vbits) then UINT b (v mod (2^b)) else UINT b v)" | "eupdate (ADDRESS a) = ADDRESS a" | "eupdate (BALANCE a) = BALANCE a" | "eupdate THIS = THIS" | "eupdate SENDER = SENDER" | "eupdate VALUE = VALUE" | "eupdate TRUE = TRUE" | "eupdate FALSE = FALSE" | "eupdate (LVAL l) = LVAL (lupdate l)" | "eupdate (PLUS ex1 ex2) = (case (eupdate ex1) of E.INT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of E.INT b2 v2 \ if b2\vbits then let v=v1+v2 in if v \ 0 then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) else (PLUS (E.INT b1 v1) (E.INT b2 v2)) | UINT b2 v2 \ if b2\vbits \ b2 < b1 then let v=v1+v2 in if v \ 0 then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) else PLUS (E.INT b1 v1) (UINT b2 v2) | _ \ PLUS (E.INT b1 v1) (eupdate ex2)) else PLUS (E.INT b1 v1) (eupdate ex2) | UINT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of UINT b2 v2 \ if b2 \ vbits then UINT (max b1 b2) ((v1 + v2) mod (2^(max b1 b2))) else (PLUS (UINT b1 v1) (UINT b2 v2)) | E.INT b2 v2 \ if b2\vbits \ b1 < b2 then let v=v1+v2 in if v \ 0 then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) else PLUS (UINT b1 v1) (E.INT b2 v2) | _ \ PLUS (UINT b1 v1) (eupdate ex2)) else PLUS (UINT b1 v1) (eupdate ex2) | _ \ PLUS (eupdate ex1) (eupdate ex2))" | "eupdate (MINUS ex1 ex2) = (case (eupdate ex1) of E.INT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of E.INT b2 v2 \ if b2\vbits then let v=v1-v2 in if v \ 0 then E.INT (max b1 b2) (-(2^((max b1 b2)-1)) + (v+2^((max b1 b2)-1)) mod (2^(max b1 b2))) else E.INT (max b1 b2) (2^((max b1 b2)-1) - (-v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1) else (MINUS (E.INT b1 v1) (E.INT b2 v2)) | UINT b2 v2 \ if b2\vbits \ b2 < b1 then let v=v1-v2 in if v \ 0 then E.INT b1 (-(2^(b1-1)) + (v+2^(b1-1)) mod (2^b1)) else E.INT b1 (2^(b1-1) - (-v+2^(b1-1)-1) mod (2^b1) - 1) else MINUS (E.INT b1 v1) (UINT b2 v2) | _ \ MINUS (E.INT b1 v1) (eupdate ex2)) else MINUS (E.INT b1 v1) (eupdate ex2) | UINT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of UINT b2 v2 \ if b2 \ vbits then UINT (max b1 b2) ((v1 - v2) mod (2^(max b1 b2))) else (MINUS (UINT b1 v1) (UINT b2 v2)) | E.INT b2 v2 \ if b2\vbits \ b1 < b2 then let v=v1-v2 in if v \ 0 then E.INT b2 (-(2^(b2-1)) + (v+2^(b2-1)) mod (2^b2)) else E.INT b2 (2^(b2-1) - (-v+2^(b2-1)-1) mod (2^b2) - 1) else MINUS (UINT b1 v1) (E.INT b2 v2) | _ \ MINUS (UINT b1 v1) (eupdate ex2)) else MINUS (UINT b1 v1) (eupdate ex2) | _ \ MINUS (eupdate ex1) (eupdate ex2))" | "eupdate (EQUAL ex1 ex2) = (case (eupdate ex1) of E.INT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of E.INT b2 v2 \ if b2\vbits then if v1 = v2 then TRUE else FALSE else EQUAL (E.INT b1 v1) (E.INT b2 v2) | UINT b2 v2 \ if b2\vbits \ b2 < b1 then if v1 = v2 then TRUE else FALSE else EQUAL (E.INT b1 v1) (UINT b2 v2) | _ \ EQUAL (E.INT b1 v1) (eupdate ex2)) else EQUAL (E.INT b1 v1) (eupdate ex2) | UINT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of UINT b2 v2 \ if b2 \ vbits then if v1 = v2 then TRUE else FALSE else EQUAL (E.INT b1 v1) (UINT b2 v2) | E.INT b2 v2 \ if b2\vbits \ b1 < b2 then if v1 = v2 then TRUE else FALSE else EQUAL (UINT b1 v1) (E.INT b2 v2) | _ \ EQUAL (UINT b1 v1) (eupdate ex2)) else EQUAL (UINT b1 v1) (eupdate ex2) | _ \ EQUAL (eupdate ex1) (eupdate ex2))" | "eupdate (LESS ex1 ex2) = (case (eupdate ex1) of E.INT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of E.INT b2 v2 \ if b2\vbits then if v1 < v2 then TRUE else FALSE else LESS (E.INT b1 v1) (E.INT b2 v2) | UINT b2 v2 \ if b2\vbits \ b2 < b1 then if v1 < v2 then TRUE else FALSE else LESS (E.INT b1 v1) (UINT b2 v2) | _ \ LESS (E.INT b1 v1) (eupdate ex2)) else LESS (E.INT b1 v1) (eupdate ex2) | UINT b1 v1 \ if b1 \ vbits then (case (eupdate ex2) of UINT b2 v2 \ if b2 \ vbits then if v1 < v2 then TRUE else FALSE else LESS (E.INT b1 v1) (UINT b2 v2) | E.INT b2 v2 \ if b2\vbits \ b1 < b2 then if v1 < v2 then TRUE else FALSE else LESS (UINT b1 v1) (E.INT b2 v2) | _ \ LESS (UINT b1 v1) (eupdate ex2)) else LESS (UINT b1 v1) (eupdate ex2) | _ \ LESS (eupdate ex1) (eupdate ex2))" | "eupdate (AND ex1 ex2) = (case (eupdate ex1) of TRUE \ (case (eupdate ex2) of TRUE \ TRUE | FALSE \ FALSE | _ \ AND TRUE (eupdate ex2)) | FALSE \ (case (eupdate ex2) of TRUE \ FALSE | FALSE \ FALSE | _ \ AND FALSE (eupdate ex2)) | _ \ AND (eupdate ex1) (eupdate ex2))" | "eupdate (OR ex1 ex2) = (case (eupdate ex1) of TRUE \ (case (eupdate ex2) of TRUE \ TRUE | FALSE \ TRUE | _ \ OR TRUE (eupdate ex2)) | FALSE \ (case (eupdate ex2) of TRUE \ TRUE | FALSE \ FALSE | _ \ OR FALSE (eupdate ex2)) | _ \ OR (eupdate ex1) (eupdate ex2))" | "eupdate (NOT ex1) = (case (eupdate ex1) of TRUE \ FALSE | FALSE \ TRUE | _ \ NOT (eupdate ex1))" | "eupdate (CALL i xs) = CALL i xs" | "eupdate (ECALL e i xs r) = ECALL e i xs r" value "eupdate (UINT 8 250)" lemma "eupdate (UINT 8 250) =UINT 8 250" by(simp) lemma "eupdate (UINT 8 500) = UINT 8 244" by(simp) lemma "eupdate (E.INT 8 (-100)) = E.INT 8 (- 100)" by(simp) lemma "eupdate (E.INT 8 (-150)) = E.INT 8 106" by(simp) lemma "eupdate (PLUS (UINT 8 100) (UINT 8 100)) = UINT 8 200" by(simp) lemma "eupdate (PLUS (UINT 8 257) (UINT 16 100)) = UINT 16 101" by(simp) lemma "eupdate (PLUS (E.INT 8 100) (UINT 8 250)) = PLUS (E.INT 8 100) (UINT 8 250)" by(simp) lemma "eupdate (PLUS (E.INT 8 250) (UINT 8 500)) = PLUS (E.INT 8 (- 6)) (UINT 8 244)" by(simp) lemma "eupdate (PLUS (E.INT 16 250) (UINT 8 500)) = E.INT 16 494" by(simp) lemma "eupdate (EQUAL (UINT 16 250) (UINT 8 250)) = TRUE" by(simp) lemma "eupdate (EQUAL (E.INT 16 100) (UINT 8 100)) = TRUE" by(simp) lemma "eupdate (EQUAL (E.INT 8 100) (UINT 8 100)) = EQUAL (E.INT 8 100) (UINT 8 100)" by(simp) lemma update_bounds_int: assumes "eupdate ex = (E.INT b v)" and "b\vbits" shows "(v < 2^(b-1)) \ v \ -(2^(b-1))" proof (cases ex) case (INT b' v') then show ?thesis proof cases assume "b'\vbits" show ?thesis proof cases let ?x="-(2^(b'-1)) + (v'+2^(b'-1)) mod 2^b'" assume "v'\0" with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp with assms have "b=b'" and "v=?x" using INT by (simp,simp) moreover from `b'\vbits` have "b'>0" by auto hence "?x < 2 ^(b'-1)" using upper_bound2[of b' "(v' + 2 ^ (b' - 1)) mod 2^b'"] by simp moreover have "?x \ -(2^(b'-1))" by simp ultimately show ?thesis by simp next let ?x="2^(b'-1) - (-v'+2^(b'-1)-1) mod (2^b') - 1" assume "\v'\0" with `b'\vbits` have "eupdate (E.INT b' v') = E.INT b' ?x" by simp with assms have "b=b'" and "v=?x" using INT by (simp,simp) moreover have "(-v'+2^(b'-1)-1) mod (2^b')\0" by simp hence "?x < 2 ^(b'-1)" by arith moreover from `b'\vbits` have "b'>0" by auto hence "?x \ -(2^(b'-1))" using lower_bound2[of b' v'] by simp ultimately show ?thesis by simp qed next assume "\ b'\vbits" with assms show ?thesis using INT by simp qed next case (UINT b' v') with assms show ?thesis proof cases assume "b'\vbits" with assms show ?thesis using UINT by simp next assume "\ b'\vbits" with assms show ?thesis using UINT by simp qed next case (ADDRESS x3) with assms show ?thesis by simp next case (BALANCE x4) with assms show ?thesis by simp next case THIS with assms show ?thesis by simp next case SENDER with assms show ?thesis by simp next case VALUE with assms show ?thesis by simp next case TRUE with assms show ?thesis by simp next case FALSE with assms show ?thesis by simp next case (LVAL x7) with assms show ?thesis by simp next case p: (PLUS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" assume "?v\0" with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x < 2 ^(max b1 b2 - 1)" using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp moreover have "?x \ -(2^(max b1 b2-1))" by simp ultimately show ?thesis by simp next let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" assume "\?v\0" with `b1\vbits` `b2\vbits` i i2 have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" by simp with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp hence "?x < 2 ^(max b1 b2-1)" by arith moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp ultimately show ?thesis by simp qed next assume "b2\vbits" with p i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b20" with `b1\vbits` `b2\vbits` `b2vbits` have "b1>0" by auto hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp moreover have "?x \ -(2^(b1-1))" by simp ultimately show ?thesis by simp next let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" assume "\?v\0" with `b1\vbits` `b2\vbits` `b20" by simp hence "?x < 2 ^(b1-1)" by arith moreover from `b1\vbits` have "b1>0" by auto hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp ultimately show ?thesis by simp qed next assume "\ b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with p i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with p i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with p i `b1\vbits` show ?thesis using assms by simp next case THIS with p i `b1\vbits` show ?thesis using assms by simp next case SENDER with p i `b1\vbits` show ?thesis using assms by simp next case VALUE with p i `b1\vbits` show ?thesis using assms by simp next case TRUE with p i `b1\vbits` show ?thesis using assms by simp next case FALSE with p i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with p i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with p i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with p i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with p i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with p i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with p i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with p i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with p i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with p i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with p i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b10" with `b1\vbits` `b2\vbits` `b1vbits` have "b2>0" by auto hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp moreover have "?x \ -(2^(b2-1))" by simp ultimately show ?thesis by simp next let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" assume "\?v\0" with `b1\vbits` `b2\vbits` `b10" by simp hence "?x < 2 ^(b2-1)" by arith moreover from `b2\vbits` have "b2>0" by auto hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp ultimately show ?thesis by simp qed next assume "\ b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with p i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" with `b1\vbits` u u2 p show ?thesis using assms by simp next assume "\b2\vbits" with p u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with p u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with p u `b1\vbits` show ?thesis using assms by simp next case THIS with p u `b1\vbits` show ?thesis using assms by simp next case SENDER with p u `b1\vbits` show ?thesis using assms by simp next case VALUE with p u `b1\vbits` show ?thesis using assms by simp next case TRUE with p u `b1\vbits` show ?thesis using assms by simp next case FALSE with p u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with p u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with p u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with p u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with p u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with p u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with p u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with p u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with p u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with p u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with p u show ?thesis using assms by simp qed next case (ADDRESS x3) with p show ?thesis using assms by simp next case (BALANCE x4) with p show ?thesis using assms by simp next case THIS with p show ?thesis using assms by simp next case SENDER with p show ?thesis using assms by simp next case VALUE with p show ?thesis using assms by simp next case TRUE with p show ?thesis using assms by simp next case FALSE with p show ?thesis using assms by simp next case (LVAL x7) with p show ?thesis using assms by simp next case (PLUS x81 x82) with p show ?thesis using assms by simp next case (MINUS x91 x92) with p show ?thesis using assms by simp next case (EQUAL x101 x102) with p show ?thesis using assms by simp next case (LESS x111 x112) with p show ?thesis using assms by simp next case (AND x121 x122) with p show ?thesis using assms by simp next case (OR x131 x132) with p show ?thesis using assms by simp next case (NOT x131) with p show ?thesis using assms by simp next case (CALL x181 x182) with p show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p show ?thesis using assms by simp qed next case m: (MINUS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) with m show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" with `b1 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT (max b1 b2) (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) else E.INT (max b1 b2) (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" using i i2 eupdate.simps(11)[of e1 e2] by simp show ?thesis proof cases let ?x="-(2^((max b1 b2)-1)) + (?v+2^((max b1 b2)-1)) mod 2^(max b1 b2)" assume "?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" by simp with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x < 2 ^(max b1 b2 - 1)" using upper_bound2[of "max b1 b2" "(?v + 2 ^ (max b1 b2 - 1)) mod 2^max b1 b2"] by simp moreover have "?x \ -(2^(max b1 b2-1))" by simp ultimately show ?thesis by simp next let ?x="2^((max b1 b2)-1) - (-?v+2^((max b1 b2)-1)-1) mod (2^(max b1 b2)) - 1" assume "\?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using u_def by simp with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) moreover have "(-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2)\0" by simp hence "?x < 2 ^(max b1 b2-1)" by arith moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x \ -(2^(max b1 b2-1))" using lower_bound2[of "max b1 b2" ?v] by simp ultimately show ?thesis by simp qed next assume "b2\vbits" with m i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" show ?thesis proof cases assume "b2 vbits` `b2 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" using i u eupdate.simps(11)[of e1 e2] by simp then show ?thesis proof cases let ?x="(-(2^(b1-1)) + (?v+2^(b1-1)) mod (2^b1))" assume "?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp with assms have "b=b1" and "v=?x" using m by (simp,simp) moreover from `b1\vbits` have "b1>0" by auto hence "?x < 2 ^(b1 - 1)" using upper_bound2[of b1] by simp moreover have "?x \ -(2^(b1-1))" by simp ultimately show ?thesis by simp next let ?x="2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" assume "\?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" by simp with assms have "b=b1" and "v=?x" using m i u by (simp,simp) moreover have "(-?v+2^(b1-1)-1) mod 2^b1\0" by simp hence "?x < 2 ^(b1-1)" by arith moreover from `b1\vbits` have "b1>0" by auto hence "?x \ -(2^(b1-1))" using lower_bound2[of b1 ?v] by simp ultimately show ?thesis by simp qed next assume "\ b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with m i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with m i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with m i `b1\vbits` show ?thesis using assms by simp next case THIS with m i `b1\vbits` show ?thesis using assms by simp next case SENDER with m i `b1\vbits` show ?thesis using assms by simp next case VALUE with m i `b1\vbits` show ?thesis using assms by simp next case TRUE with m i `b1\vbits` show ?thesis using assms by simp next case FALSE with m i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with m i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with m i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with m i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with m i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with m i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with m i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with m i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with m i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with m i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with m i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" show ?thesis proof cases assume "b1 vbits` `b2 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" using i u eupdate.simps(11)[of e1 e2] by simp then show ?thesis proof cases let ?x="(-(2^(b2-1)) + (?v+2^(b2-1)) mod (2^b2))" assume "?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp with assms have "b=b2" and "v=?x" using m by (simp,simp) moreover from `b2\vbits` have "b2>0" by auto hence "?x < 2 ^(b2 - 1)" using upper_bound2[of b2] by simp moreover have "?x \ -(2^(b2-1))" by simp ultimately show ?thesis by simp next let ?x="2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" assume "\?v\0" with u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" by simp with assms have "b=b2" and "v=?x" using m i u by (simp,simp) moreover have "(-?v+2^(b2-1)-1) mod 2^b2\0" by simp hence "?x < 2 ^(b2-1)" by arith moreover from `b2\vbits` have "b2>0" by auto hence "?x \ -(2^(b2-1))" using lower_bound2[of b2 ?v] by simp ultimately show ?thesis by simp qed next assume "\ b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with m i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" with `b1\vbits` u u2 m show ?thesis using assms by simp next assume "\b2\vbits" with m u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with m u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with m u `b1\vbits` show ?thesis using assms by simp next case THIS with m u `b1\vbits` show ?thesis using assms by simp next case SENDER with m u `b1\vbits` show ?thesis using assms by simp next case VALUE with m u `b1\vbits` show ?thesis using assms by simp next case TRUE with m u `b1\vbits` show ?thesis using assms by simp next case FALSE with m u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with m u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with m u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with m u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with m u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with m u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with m u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with m u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with m u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with m u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with m u show ?thesis using assms by simp qed next case (ADDRESS x3) with m show ?thesis using assms by simp next case (BALANCE x4) with m show ?thesis using assms by simp next case THIS with m show ?thesis using assms by simp next case SENDER with m show ?thesis using assms by simp next case VALUE with m show ?thesis using assms by simp next case TRUE with m show ?thesis using assms by simp next case FALSE with m show ?thesis using assms by simp next case (LVAL x7) with m show ?thesis using assms by simp next case (PLUS x81 x82) with m show ?thesis using assms by simp next case (MINUS x91 x92) with m show ?thesis using assms by simp next case (EQUAL x101 x102) with m show ?thesis using assms by simp next case (LESS x111 x112) with m show ?thesis using assms by simp next case (AND x121 x122) with m show ?thesis using assms by simp next case (OR x131 x132) with m show ?thesis using assms by simp next case (NOT x131) with m show ?thesis using assms by simp next case (CALL x181 x182) with m show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m show ?thesis using assms by simp qed next case e: (EQUAL e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1=v2" with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp next assume "\ v1=v2" with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp qed next assume "b2\vbits" with e i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b2vbits` `b2\vbits` `b2 v1=v2" with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with e i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with e i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with e i `b1\vbits` show ?thesis using assms by simp next case THIS with e i `b1\vbits` show ?thesis using assms by simp next case SENDER with e i `b1\vbits` show ?thesis using assms by simp next case VALUE with e i `b1\vbits` show ?thesis using assms by simp next case TRUE with e i `b1\vbits` show ?thesis using assms by simp next case FALSE with e i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with e i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with e i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with e i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with e i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with e i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with e i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with e i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with e i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with e i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with e i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b1vbits` `b2\vbits` `b1 v1=v2" with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with e i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1=v2" with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp next assume "\ v1=v2" with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp qed next assume "\b2\vbits" with e u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with e u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with e u `b1\vbits` show ?thesis using assms by simp next case THIS with e u `b1\vbits` show ?thesis using assms by simp next case SENDER with e u `b1\vbits` show ?thesis using assms by simp next case VALUE with e u `b1\vbits` show ?thesis using assms by simp next case TRUE with e u `b1\vbits` show ?thesis using assms by simp next case FALSE with e u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with e u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with e u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with e u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with e u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with e u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with e u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with e u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with e u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with e u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with e u show ?thesis using assms by simp qed next case (ADDRESS x3) with e show ?thesis using assms by simp next case (BALANCE x4) with e show ?thesis using assms by simp next case THIS with e show ?thesis using assms by simp next case SENDER with e show ?thesis using assms by simp next case VALUE with e show ?thesis using assms by simp next case TRUE with e show ?thesis using assms by simp next case FALSE with e show ?thesis using assms by simp next case (LVAL x7) with e show ?thesis using assms by simp next case (PLUS x81 x82) with e show ?thesis using assms by simp next case (MINUS x91 x92) with e show ?thesis using assms by simp next case (EQUAL x101 x102) with e show ?thesis using assms by simp next case (LESS x111 x112) with e show ?thesis using assms by simp next case (AND x121 x122) with e show ?thesis using assms by simp next case (OR x131 x132) with e show ?thesis using assms by simp next case (NOT x131) with e show ?thesis using assms by simp next case (CALL x181 x182) with e show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e show ?thesis using assms by simp qed next case l: (LESS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1vbits` `b2\vbits` by simp next assume "\ v1vbits` `b2\vbits` by simp qed next assume "b2\vbits" with l i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with l i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with l i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with l i `b1\vbits` show ?thesis using assms by simp next case THIS with l i `b1\vbits` show ?thesis using assms by simp next case SENDER with l i `b1\vbits` show ?thesis using assms by simp next case VALUE with l i `b1\vbits` show ?thesis using assms by simp next case TRUE with l i `b1\vbits` show ?thesis using assms by simp next case FALSE with l i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with l i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with l i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with l i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with l i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with l i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with l i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with l i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with l i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with l i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with l i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with l i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1vbits` `b2\vbits` by simp next assume "\ v1vbits` `b2\vbits` by simp qed next assume "\b2\vbits" with l u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with l u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with l u `b1\vbits` show ?thesis using assms by simp next case THIS with l u `b1\vbits` show ?thesis using assms by simp next case SENDER with l u `b1\vbits` show ?thesis using assms by simp next case VALUE with l u `b1\vbits` show ?thesis using assms by simp next case TRUE with l u `b1\vbits` show ?thesis using assms by simp next case FALSE with l u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with l u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with l u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with l u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with l u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with l u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with l u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with l u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with l u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with l u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with l u show ?thesis using assms by simp qed next case (ADDRESS x3) with l show ?thesis using assms by simp next case (BALANCE x4) with l show ?thesis using assms by simp next case THIS with l show ?thesis using assms by simp next case SENDER with l show ?thesis using assms by simp next case VALUE with l show ?thesis using assms by simp next case TRUE with l show ?thesis using assms by simp next case FALSE with l show ?thesis using assms by simp next case (LVAL x7) with l show ?thesis using assms by simp next case (PLUS x81 x82) with l show ?thesis using assms by simp next case (MINUS x91 x92) with l show ?thesis using assms by simp next case (EQUAL x101 x102) with l show ?thesis using assms by simp next case (LESS x111 x112) with l show ?thesis using assms by simp next case (AND x121 x122) with l show ?thesis using assms by simp next case (OR x131 x132) with l show ?thesis using assms by simp next case (NOT x131) with l show ?thesis using assms by simp next case (CALL x181 x182) with l show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l show ?thesis using assms by simp qed next case a: (AND e1 e2) show ?thesis proof (cases "eupdate e1") case (INT x11 x12) with a show ?thesis using assms by simp next case (UINT x21 x22) with a show ?thesis using assms by simp next case (ADDRESS x3) with a show ?thesis using assms by simp next case (BALANCE x4) with a show ?thesis using assms by simp next case THIS with a show ?thesis using assms by simp next case SENDER with a show ?thesis using assms by simp next case VALUE with a show ?thesis using assms by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with a t show ?thesis using assms by simp next case (UINT x21 x22) with a t show ?thesis using assms by simp next case (ADDRESS x3) with a t show ?thesis using assms by simp next case (BALANCE x4) with a t show ?thesis using assms by simp next case THIS with a t show ?thesis using assms by simp next case SENDER with a t show ?thesis using assms by simp next case VALUE with a t show ?thesis using assms by simp next case TRUE with a t show ?thesis using assms by simp next case FALSE with a t show ?thesis using assms by simp next case (LVAL x7) with a t show ?thesis using assms by simp next case (PLUS x81 x82) with a t show ?thesis using assms by simp next case (MINUS x91 x92) with a t show ?thesis using assms by simp next case (EQUAL x101 x102) with a t show ?thesis using assms by simp next case (LESS x111 x112) with a t show ?thesis using assms by simp next case (AND x121 x122) with a t show ?thesis using assms by simp next case (OR x131 x132) with a t show ?thesis using assms by simp next case (NOT x131) with a t show ?thesis using assms by simp next case (CALL x181 x182) with a t show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a t show ?thesis using assms by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with a f show ?thesis using assms by simp next case (UINT x21 x22) with a f show ?thesis using assms by simp next case (ADDRESS x3) with a f show ?thesis using assms by simp next case (BALANCE x4) with a f show ?thesis using assms by simp next case THIS with a f show ?thesis using assms by simp next case SENDER with a f show ?thesis using assms by simp next case VALUE with a f show ?thesis using assms by simp next case TRUE with a f show ?thesis using assms by simp next case FALSE with a f show ?thesis using assms by simp next case (LVAL x7) with a f show ?thesis using assms by simp next case (PLUS x81 x82) with a f show ?thesis using assms by simp next case (MINUS x91 x92) with a f show ?thesis using assms by simp next case (EQUAL x101 x102) with a f show ?thesis using assms by simp next case (LESS x111 x112) with a f show ?thesis using assms by simp next case (AND x121 x122) with a f show ?thesis using assms by simp next case (OR x131 x132) with a f show ?thesis using assms by simp next case (NOT x131) with a f show ?thesis using assms by simp next case (CALL x181 x182) with a f show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a f show ?thesis using assms by simp qed next case (LVAL x7) with a show ?thesis using assms by simp next case (PLUS x81 x82) with a show ?thesis using assms by simp next case (MINUS x91 x92) with a show ?thesis using assms by simp next case (EQUAL x101 x102) with a show ?thesis using assms by simp next case (LESS x111 x112) with a show ?thesis using assms by simp next case (AND x121 x122) with a show ?thesis using assms by simp next case (OR x131 x132) with a show ?thesis using assms by simp next case (NOT x131) with a show ?thesis using assms by simp next case (CALL x181 x182) with a show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a show ?thesis using assms by simp qed next case o: (OR e1 e2) show ?thesis proof (cases "eupdate e1") case (INT x11 x12) with o show ?thesis using assms by simp next case (UINT x21 x22) with o show ?thesis using assms by simp next case (ADDRESS x3) with o show ?thesis using assms by simp next case (BALANCE x4) with o show ?thesis using assms by simp next case THIS with o show ?thesis using assms by simp next case SENDER with o show ?thesis using assms by simp next case VALUE with o show ?thesis using assms by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with o t show ?thesis using assms by simp next case (UINT x21 x22) with o t show ?thesis using assms by simp next case (ADDRESS x3) with o t show ?thesis using assms by simp next case (BALANCE x4) with o t show ?thesis using assms by simp next case THIS with o t show ?thesis using assms by simp next case SENDER with o t show ?thesis using assms by simp next case VALUE with o t show ?thesis using assms by simp next case TRUE with o t show ?thesis using assms by simp next case FALSE with o t show ?thesis using assms by simp next case (LVAL x7) with o t show ?thesis using assms by simp next case (PLUS x81 x82) with o t show ?thesis using assms by simp next case (MINUS x91 x92) with o t show ?thesis using assms by simp next case (EQUAL x101 x102) with o t show ?thesis using assms by simp next case (LESS x111 x112) with o t show ?thesis using assms by simp next case (AND x121 x122) with o t show ?thesis using assms by simp next case (OR x131 x132) with o t show ?thesis using assms by simp next case (NOT x131) with o t show ?thesis using assms by simp next case (CALL x181 x182) with o t show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o t show ?thesis using assms by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with o f show ?thesis using assms by simp next case (UINT x21 x22) with o f show ?thesis using assms by simp next case (ADDRESS x3) with o f show ?thesis using assms by simp next case (BALANCE x4) with o f show ?thesis using assms by simp next case THIS with o f show ?thesis using assms by simp next case SENDER with o f show ?thesis using assms by simp next case VALUE with o f show ?thesis using assms by simp next case TRUE with o f show ?thesis using assms by simp next case FALSE with o f show ?thesis using assms by simp next case (LVAL x7) with o f show ?thesis using assms by simp next case (PLUS x81 x82) with o f show ?thesis using assms by simp next case (MINUS x91 x92) with o f show ?thesis using assms by simp next case (EQUAL x101 x102) with o f show ?thesis using assms by simp next case (LESS x111 x112) with o f show ?thesis using assms by simp next case (AND x121 x122) with o f show ?thesis using assms by simp next case (OR x131 x132) with o f show ?thesis using assms by simp next case (NOT x131) with o f show ?thesis using assms by simp next case (CALL x181 x182) with o f show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o f show ?thesis using assms by simp qed next case (LVAL x7) with o show ?thesis using assms by simp next case (PLUS x81 x82) with o show ?thesis using assms by simp next case (MINUS x91 x92) with o show ?thesis using assms by simp next case (EQUAL x101 x102) with o show ?thesis using assms by simp next case (LESS x111 x112) with o show ?thesis using assms by simp next case (AND x121 x122) with o show ?thesis using assms by simp next case (OR x131 x132) with o show ?thesis using assms by simp next case (NOT x131) with o show ?thesis using assms by simp next case (CALL x181 x182) with o show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis using assms by simp qed next case o: (NOT e1) show ?thesis proof (cases "eupdate e1") case (INT x11 x12) with o show ?thesis using assms by simp next case (UINT x21 x22) with o show ?thesis using assms by simp next case (ADDRESS x3) with o show ?thesis using assms by simp next case (BALANCE x4) with o show ?thesis using assms by simp next case THIS with o show ?thesis using assms by simp next case SENDER with o show ?thesis using assms by simp next case VALUE with o show ?thesis using assms by simp next case t: TRUE with o show ?thesis using assms by simp next case f: FALSE with o show ?thesis using assms by simp next case (LVAL x7) with o show ?thesis using assms by simp next case (PLUS x81 x82) with o show ?thesis using assms by simp next case (MINUS x91 x92) with o show ?thesis using assms by simp next case (EQUAL x101 x102) with o show ?thesis using assms by simp next case (LESS x111 x112) with o show ?thesis using assms by simp next case (AND x121 x122) with o show ?thesis using assms by simp next case (OR x131 x132) with o show ?thesis using assms by simp next case (NOT x131) with o show ?thesis using assms by simp next case (CALL x181 x182) with o show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis using assms by simp qed next case (CALL x181 x182) with assms show ?thesis by simp next case (ECALL x191 x192 x193 x194) with assms show ?thesis by simp qed lemma update_bounds_uint: assumes "eupdate ex = UINT b v" and "b\vbits" shows "v < 2^b \ v \ 0" proof (cases ex) case (INT b' v') with assms show ?thesis proof cases assume "b'\vbits" show ?thesis proof cases assume "v'\0" with INT show ?thesis using assms `b'\vbits` by simp next assume "\ v'\0" with INT show ?thesis using assms `b'\vbits` by simp qed next assume "\ b'\vbits" with INT show ?thesis using assms by simp qed next case (UINT b' v') then show ?thesis proof cases assume "b'\vbits" with UINT show ?thesis using assms by auto next assume "\ b'\vbits" with UINT show ?thesis using assms by auto qed next case (ADDRESS x3) with assms show ?thesis by simp next case (BALANCE x4) with assms show ?thesis by simp next case THIS with assms show ?thesis by simp next case SENDER with assms show ?thesis by simp next case VALUE with assms show ?thesis by simp next case TRUE with assms show ?thesis by simp next case FALSE with assms show ?thesis by simp next case (LVAL x7) with assms show ?thesis by simp next case p: (PLUS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) with p show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "?v\0" with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp next assume "\?v\0" with assms show ?thesis using p i i2 `b1\vbits` `b2\vbits` by simp qed next assume "b2\vbits" with p i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b20" with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2?v\0" with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with p i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with p i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with p i `b1\vbits` show ?thesis using assms by simp next case THIS with p i `b1\vbits` show ?thesis using assms by simp next case SENDER with p i `b1\vbits` show ?thesis using assms by simp next case VALUE with p i `b1\vbits` show ?thesis using assms by simp next case TRUE with p i `b1\vbits` show ?thesis using assms by simp next case FALSE with p i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with p i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with p i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with p i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with p i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with p i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with p i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with p i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with p i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with p i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with p i show ?thesis using assms by simp qed next case u: (UINT b1 v1) with p show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b10" with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1?v\0" with assms show ?thesis using p i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with p i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases let ?x="((v1 + v2) mod (2^(max b1 b2)))" assume "b2\vbits" with `b1\vbits` u u2 have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" by simp with assms have "b=max b1 b2" and "v=?x" using p by (simp,simp) moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x < 2 ^(max b1 b2)" by simp moreover have "?x \ 0" by simp ultimately show ?thesis by simp next assume "\b2\vbits" with p u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with p u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with p u `b1\vbits` show ?thesis using assms by simp next case THIS with p u `b1\vbits` show ?thesis using assms by simp next case SENDER with p u `b1\vbits` show ?thesis using assms by simp next case VALUE with p u `b1\vbits` show ?thesis using assms by simp next case TRUE with p u `b1\vbits` show ?thesis using assms by simp next case FALSE with p u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with p u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with p u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with p u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with p u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with p u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with p u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with p u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with p u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with p u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with p u show ?thesis using assms by simp qed next case (ADDRESS x3) with p show ?thesis using assms by simp next case (BALANCE x4) with p show ?thesis using assms by simp next case THIS with p show ?thesis using assms by simp next case SENDER with p show ?thesis using assms by simp next case VALUE with p show ?thesis using assms by simp next case TRUE with p show ?thesis using assms by simp next case FALSE with p show ?thesis using assms by simp next case (LVAL x7) with p show ?thesis using assms by simp next case (PLUS x81 x82) with p show ?thesis using assms by simp next case (MINUS x91 x92) with p show ?thesis using assms by simp next case (EQUAL x101 x102) with p show ?thesis using assms by simp next case (LESS x111 x112) with p show ?thesis using assms by simp next case (AND x121 x122) with p show ?thesis using assms by simp next case (OR x131 x132) with p show ?thesis using assms by simp next case (NOT x131) with p show ?thesis using assms by simp next case (CALL x181 x182) with p show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with p show ?thesis using assms by simp qed next case m: (MINUS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) with m show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" show ?thesis proof cases assume "?v\0" with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp next assume "\?v\0" with assms show ?thesis using m i i2 `b1\vbits` `b2\vbits` by simp qed next assume "b2\vbits" with m i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" show ?thesis proof cases assume "b20" with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2?v\0" with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with m i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with m i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with m i `b1\vbits` show ?thesis using assms by simp next case THIS with m i `b1\vbits` show ?thesis using assms by simp next case SENDER with m i `b1\vbits` show ?thesis using assms by simp next case VALUE with m i `b1\vbits` show ?thesis using assms by simp next case TRUE with m i `b1\vbits` show ?thesis using assms by simp next case FALSE with m i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with m i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with m i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with m i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with m i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with m i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with m i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with m i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with m i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with m i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with m i show ?thesis using assms by simp qed next case u: (UINT b1 v1) with m show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1-v2" assume "b2\vbits" show ?thesis proof cases assume "b10" with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1?v\0" with assms show ?thesis using m i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with m i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases let ?x="((v1 - v2) mod (2^(max b1 b2)))" assume "b2\vbits" with `b1\vbits` u u2 have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" by simp with assms have "b=max b1 b2" and "v=?x" using m by (simp,simp) moreover from `b1\vbits` have "max b1 b2>0" by auto hence "?x < 2 ^(max b1 b2)" by simp moreover have "?x \ 0" by simp ultimately show ?thesis by simp next assume "\b2\vbits" with m u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with m u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with m u `b1\vbits` show ?thesis using assms by simp next case THIS with m u `b1\vbits` show ?thesis using assms by simp next case SENDER with m u `b1\vbits` show ?thesis using assms by simp next case VALUE with m u `b1\vbits` show ?thesis using assms by simp next case TRUE with m u `b1\vbits` show ?thesis using assms by simp next case FALSE with m u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with m u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with m u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with m u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with m u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with m u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with m u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with m u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with m u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with m u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with m u show ?thesis using assms by simp qed next case (ADDRESS x3) with m show ?thesis using assms by simp next case (BALANCE x4) with m show ?thesis using assms by simp next case THIS with m show ?thesis using assms by simp next case SENDER with m show ?thesis using assms by simp next case VALUE with m show ?thesis using assms by simp next case TRUE with m show ?thesis using assms by simp next case FALSE with m show ?thesis using assms by simp next case (LVAL x7) with m show ?thesis using assms by simp next case (PLUS x81 x82) with m show ?thesis using assms by simp next case (MINUS x91 x92) with m show ?thesis using assms by simp next case (EQUAL x101 x102) with m show ?thesis using assms by simp next case (LESS x111 x112) with m show ?thesis using assms by simp next case (AND x121 x122) with m show ?thesis using assms by simp next case (OR x131 x132) with m show ?thesis using assms by simp next case (NOT x131) with m show ?thesis using assms by simp next case (CALL x181 x182) with m show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with m show ?thesis using assms by simp qed next case e: (EQUAL e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1=v2" with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp next assume "\ v1=v2" with assms show ?thesis using e i i2 `b1\vbits` `b2\vbits` by simp qed next assume "b2\vbits" with e i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b2vbits` `b2\vbits` `b2 v1=v2" with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with e i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with e i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with e i `b1\vbits` show ?thesis using assms by simp next case THIS with e i `b1\vbits` show ?thesis using assms by simp next case SENDER with e i `b1\vbits` show ?thesis using assms by simp next case VALUE with e i `b1\vbits` show ?thesis using assms by simp next case TRUE with e i `b1\vbits` show ?thesis using assms by simp next case FALSE with e i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with e i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with e i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with e i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with e i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with e i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with e i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with e i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with e i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with e i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with e i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b1vbits` `b2\vbits` `b1 v1=v2" with assms show ?thesis using e i u `b1\vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with e i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1=v2" with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp next assume "\ v1=v2" with assms show ?thesis using e u u2 `b1\vbits` `b2\vbits` by simp qed next assume "\b2\vbits" with e u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with e u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with e u `b1\vbits` show ?thesis using assms by simp next case THIS with e u `b1\vbits` show ?thesis using assms by simp next case SENDER with e u `b1\vbits` show ?thesis using assms by simp next case VALUE with e u `b1\vbits` show ?thesis using assms by simp next case TRUE with e u `b1\vbits` show ?thesis using assms by simp next case FALSE with e u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with e u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with e u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with e u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with e u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with e u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with e u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with e u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with e u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with e u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with e u show ?thesis using assms by simp qed next case (ADDRESS x3) with e show ?thesis using assms by simp next case (BALANCE x4) with e show ?thesis using assms by simp next case THIS with e show ?thesis using assms by simp next case SENDER with e show ?thesis using assms by simp next case VALUE with e show ?thesis using assms by simp next case TRUE with e show ?thesis using assms by simp next case FALSE with e show ?thesis using assms by simp next case (LVAL x7) with e show ?thesis using assms by simp next case (PLUS x81 x82) with e show ?thesis using assms by simp next case (MINUS x91 x92) with e show ?thesis using assms by simp next case (EQUAL x101 x102) with e show ?thesis using assms by simp next case (LESS x111 x112) with e show ?thesis using assms by simp next case (AND x121 x122) with e show ?thesis using assms by simp next case (OR x131 x132) with e show ?thesis using assms by simp next case (NOT x131) with e show ?thesis using assms by simp next case (CALL x181 x182) with e show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with e show ?thesis using assms by simp qed next case l: (LESS e1 e2) show ?thesis proof (cases "eupdate e1") case i: (INT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1vbits` `b2\vbits` by simp next assume "\ v1vbits` `b2\vbits` by simp qed next assume "b2\vbits" with l i i2 `b1\vbits` show ?thesis using assms by simp qed next case u: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "b2vbits` `b2\vbits` `b2 v1vbits` `b2\vbits` `b2 b2vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with l i u `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with l i `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with l i `b1\vbits` show ?thesis using assms by simp next case THIS with l i `b1\vbits` show ?thesis using assms by simp next case SENDER with l i `b1\vbits` show ?thesis using assms by simp next case VALUE with l i `b1\vbits` show ?thesis using assms by simp next case TRUE with l i `b1\vbits` show ?thesis using assms by simp next case FALSE with l i `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with l i `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with l i `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with l i `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with l i `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with l i `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with l i `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with l i `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with l i `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with l i `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l i `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with l i show ?thesis using assms by simp qed next case u: (UINT b1 v1) show ?thesis proof cases assume "b1\vbits" show ?thesis proof (cases "eupdate e2") case i: (INT b2 v2) then show ?thesis proof cases let ?v="v1+v2" assume "b2\vbits" show ?thesis proof cases assume "b1vbits` `b2\vbits` `b1 v1vbits` `b2\vbits` `b1 b1vbits` show ?thesis using assms by simp qed next assume "b2\vbits" with l i u `b1\vbits` show ?thesis using assms by simp qed next case u2: (UINT b2 v2) then show ?thesis proof cases assume "b2\vbits" show ?thesis proof cases assume "v1vbits` `b2\vbits` by simp next assume "\ v1vbits` `b2\vbits` by simp qed next assume "\b2\vbits" with l u u2 `b1\vbits` show ?thesis using assms by simp qed next case (ADDRESS x3) with l u `b1\vbits` show ?thesis using assms by simp next case (BALANCE x4) with l u `b1\vbits` show ?thesis using assms by simp next case THIS with l u `b1\vbits` show ?thesis using assms by simp next case SENDER with l u `b1\vbits` show ?thesis using assms by simp next case VALUE with l u `b1\vbits` show ?thesis using assms by simp next case TRUE with l u `b1\vbits` show ?thesis using assms by simp next case FALSE with l u `b1\vbits` show ?thesis using assms by simp next case (LVAL x7) with l u `b1\vbits` show ?thesis using assms by simp next case (PLUS x81 x82) with l u `b1\vbits` show ?thesis using assms by simp next case (MINUS x91 x92) with l u `b1\vbits` show ?thesis using assms by simp next case (EQUAL x101 x102) with l u `b1\vbits` show ?thesis using assms by simp next case (LESS x111 x112) with l u `b1\vbits` show ?thesis using assms by simp next case (AND x121 x122) with l u `b1\vbits` show ?thesis using assms by simp next case (OR x131 x132) with l u `b1\vbits` show ?thesis using assms by simp next case (NOT x131) with l u `b1\vbits` show ?thesis using assms by simp next case (CALL x181 x182) with l u `b1\vbits` show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l u `b1\vbits` show ?thesis using assms by simp qed next assume "\ b1\vbits" with l u show ?thesis using assms by simp qed next case (ADDRESS x3) with l show ?thesis using assms by simp next case (BALANCE x4) with l show ?thesis using assms by simp next case THIS with l show ?thesis using assms by simp next case SENDER with l show ?thesis using assms by simp next case VALUE with l show ?thesis using assms by simp next case TRUE with l show ?thesis using assms by simp next case FALSE with l show ?thesis using assms by simp next case (LVAL x7) with l show ?thesis using assms by simp next case (PLUS x81 x82) with l show ?thesis using assms by simp next case (MINUS x91 x92) with l show ?thesis using assms by simp next case (EQUAL x101 x102) with l show ?thesis using assms by simp next case (LESS x111 x112) with l show ?thesis using assms by simp next case (AND x121 x122) with l show ?thesis using assms by simp next case (OR x131 x132) with l show ?thesis using assms by simp next case (NOT x131) with l show ?thesis using assms by simp next case (CALL x181 x182) with l show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with l show ?thesis using assms by simp qed next case a: (AND e1 e2) show ?thesis proof (cases "eupdate e1") case (INT x11 x12) with a show ?thesis using assms by simp next case (UINT x21 x22) with a show ?thesis using assms by simp next case (ADDRESS x3) with a show ?thesis using assms by simp next case (BALANCE x4) with a show ?thesis using assms by simp next case THIS with a show ?thesis using assms by simp next case SENDER with a show ?thesis using assms by simp next case VALUE with a show ?thesis using assms by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with a t show ?thesis using assms by simp next case (UINT x21 x22) with a t show ?thesis using assms by simp next case (ADDRESS x3) with a t show ?thesis using assms by simp next case (BALANCE x4) with a t show ?thesis using assms by simp next case THIS with a t show ?thesis using assms by simp next case SENDER with a t show ?thesis using assms by simp next case VALUE with a t show ?thesis using assms by simp next case TRUE with a t show ?thesis using assms by simp next case FALSE with a t show ?thesis using assms by simp next case (LVAL x7) with a t show ?thesis using assms by simp next case (PLUS x81 x82) with a t show ?thesis using assms by simp next case (MINUS x91 x92) with a t show ?thesis using assms by simp next case (EQUAL x101 x102) with a t show ?thesis using assms by simp next case (LESS x111 x112) with a t show ?thesis using assms by simp next case (AND x121 x122) with a t show ?thesis using assms by simp next case (OR x131 x132) with a t show ?thesis using assms by simp next case (NOT x131) with a t show ?thesis using assms by simp next case (CALL x181 x182) with a t show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a t show ?thesis using assms by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with a f show ?thesis using assms by simp next case (UINT x21 x22) with a f show ?thesis using assms by simp next case (ADDRESS x3) with a f show ?thesis using assms by simp next case (BALANCE x4) with a f show ?thesis using assms by simp next case THIS with a f show ?thesis using assms by simp next case SENDER with a f show ?thesis using assms by simp next case VALUE with a f show ?thesis using assms by simp next case TRUE with a f show ?thesis using assms by simp next case FALSE with a f show ?thesis using assms by simp next case (LVAL x7) with a f show ?thesis using assms by simp next case (PLUS x81 x82) with a f show ?thesis using assms by simp next case (MINUS x91 x92) with a f show ?thesis using assms by simp next case (EQUAL x101 x102) with a f show ?thesis using assms by simp next case (LESS x111 x112) with a f show ?thesis using assms by simp next case (AND x121 x122) with a f show ?thesis using assms by simp next case (OR x131 x132) with a f show ?thesis using assms by simp next case (NOT x131) with a f show ?thesis using assms by simp next case (CALL x181 x182) with a f show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a f show ?thesis using assms by simp qed next case (LVAL x7) with a show ?thesis using assms by simp next case (PLUS x81 x82) with a show ?thesis using assms by simp next case (MINUS x91 x92) with a show ?thesis using assms by simp next case (EQUAL x101 x102) with a show ?thesis using assms by simp next case (LESS x111 x112) with a show ?thesis using assms by simp next case (AND x121 x122) with a show ?thesis using assms by simp next case (OR x131 x132) with a show ?thesis using assms by simp next case (NOT x131) with a show ?thesis using assms by simp next case (CALL x181 x182) with a show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with a show ?thesis using assms by simp qed next case o: (OR e1 e2) show ?thesis proof (cases "eupdate e1") case (INT x11 x12) with o show ?thesis using assms by simp next case (UINT x21 x22) with o show ?thesis using assms by simp next case (ADDRESS x3) with o show ?thesis using assms by simp next case (BALANCE x4) with o show ?thesis using assms by simp next case THIS with o show ?thesis using assms by simp next case SENDER with o show ?thesis using assms by simp next case VALUE with o show ?thesis using assms by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with o t show ?thesis using assms by simp next case (UINT x21 x22) with o t show ?thesis using assms by simp next case (ADDRESS x3) with o t show ?thesis using assms by simp next case (BALANCE x4) with o t show ?thesis using assms by simp next case THIS with o t show ?thesis using assms by simp next case SENDER with o t show ?thesis using assms by simp next case VALUE with o t show ?thesis using assms by simp next case TRUE with o t show ?thesis using assms by simp next case FALSE with o t show ?thesis using assms by simp next case (LVAL x7) with o t show ?thesis using assms by simp next case (PLUS x81 x82) with o t show ?thesis using assms by simp next case (MINUS x91 x92) with o t show ?thesis using assms by simp next case (EQUAL x101 x102) with o t show ?thesis using assms by simp next case (LESS x111 x112) with o t show ?thesis using assms by simp next case (AND x121 x122) with o t show ?thesis using assms by simp next case (OR x131 x132) with o t show ?thesis using assms by simp next case (NOT x131) with o t show ?thesis using assms by simp next case (CALL x181 x182) with o t show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o t show ?thesis using assms by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with o f show ?thesis using assms by simp next case (UINT x21 x22) with o f show ?thesis using assms by simp next case (ADDRESS x3) with o f show ?thesis using assms by simp next case (BALANCE x4) with o f show ?thesis using assms by simp next case THIS with o f show ?thesis using assms by simp next case SENDER with o f show ?thesis using assms by simp next case VALUE with o f show ?thesis using assms by simp next case TRUE with o f show ?thesis using assms by simp next case FALSE with o f show ?thesis using assms by simp next case (LVAL x7) with o f show ?thesis using assms by simp next case (PLUS x81 x82) with o f show ?thesis using assms by simp next case (MINUS x91 x92) with o f show ?thesis using assms by simp next case (EQUAL x101 x102) with o f show ?thesis using assms by simp next case (LESS x111 x112) with o f show ?thesis using assms by simp next case (AND x121 x122) with o f show ?thesis using assms by simp next case (OR x131 x132) with o f show ?thesis using assms by simp next case (NOT x131) with o f show ?thesis using assms by simp next case (CALL x181 x182) with o f show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o f show ?thesis using assms by simp qed next case (LVAL x7) with o show ?thesis using assms by simp next case (PLUS x81 x82) with o show ?thesis using assms by simp next case (MINUS x91 x92) with o show ?thesis using assms by simp next case (EQUAL x101 x102) with o show ?thesis using assms by simp next case (LESS x111 x112) with o show ?thesis using assms by simp next case (AND x121 x122) with o show ?thesis using assms by simp next case (OR x131 x132) with o show ?thesis using assms by simp next case (NOT x131) with o show ?thesis using assms by simp next case (CALL x181 x182) with o show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis using assms by simp qed next case o: (NOT x) show ?thesis proof (cases "eupdate x") case (INT x11 x12) with o show ?thesis using assms by simp next case (UINT x21 x22) with o show ?thesis using assms by simp next case (ADDRESS x3) with o show ?thesis using assms by simp next case (BALANCE x4) with o show ?thesis using assms by simp next case THIS with o show ?thesis using assms by simp next case SENDER with o show ?thesis using assms by simp next case VALUE with o show ?thesis using assms by simp next case t: TRUE with o show ?thesis using assms by simp next case f: FALSE with o show ?thesis using assms by simp next case (LVAL x7) with o show ?thesis using assms by simp next case (PLUS x81 x82) with o show ?thesis using assms by simp next case (MINUS x91 x92) with o show ?thesis using assms by simp next case (EQUAL x101 x102) with o show ?thesis using assms by simp next case (LESS x111 x112) with o show ?thesis using assms by simp next case (AND x121 x122) with o show ?thesis using assms by simp next case (OR x131 x132) with o show ?thesis using assms by simp next case (NOT x131) with o show ?thesis using assms by simp next case (CALL x181 x182) with o show ?thesis using assms by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis using assms by simp qed next case (CALL x181 x182) with assms show ?thesis by simp next case (ECALL x191 x192 x193 x194) with assms show ?thesis by simp qed lemma no_gas: assumes "\ gas st > 0" shows "expr ex ep env cd st = Exception Gas" proof (cases ex) case (INT x11 x12) with assms show ?thesis by simp next case (UINT x21 x22) with assms show ?thesis by simp next case (ADDRESS x3) with assms show ?thesis by simp next case (BALANCE x4) with assms show ?thesis by simp next case THIS with assms show ?thesis by simp next case SENDER with assms show ?thesis by simp next case VALUE with assms show ?thesis by simp next case TRUE with assms show ?thesis by simp next case FALSE with assms show ?thesis by simp next case (LVAL x10) with assms show ?thesis by simp next case (PLUS x111 x112) with assms show ?thesis by simp next case (MINUS x121 x122) with assms show ?thesis by simp next case (EQUAL x131 x132) with assms show ?thesis by simp next case (LESS x141 x142) with assms show ?thesis by simp next case (AND x151 x152) with assms show ?thesis by simp next case (OR x161 x162) with assms show ?thesis by simp next case (NOT x17) with assms show ?thesis by simp next case (CALL x181 x182) with assms show ?thesis by simp next case (ECALL x191 x192 x193 x194) with assms show ?thesis by simp qed lemma lift_eq: assumes "expr e1 ep env cd st = expr e1' ep env cd st" and "\st' rv. expr e1 ep env cd st = Normal (rv, st') \ expr e2 ep env cd st'= expr e2' ep env cd st'" shows "lift expr f e1 e2 ep env cd st=lift expr f e1' e2' ep env cd st" proof (cases "expr e1 ep env cd st") case s1: (n a st') then show ?thesis proof (cases a) case f1:(Pair a b) then show ?thesis proof (cases a) case k1:(KValue x1) then show ?thesis proof (cases b) case v1: (Value x1) then show ?thesis proof (cases "expr e2 ep env cd st'") case s2: (n a' st'') then show ?thesis proof (cases a') case f2:(Pair a' b') then show ?thesis proof (cases a') case (KValue x1') with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto next case (KCDptr x2) with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto next case (KMemptr x2') with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto next case (KStoptr x3') with s1 f1 k1 v1 assms(1) assms(2) show ?thesis by auto qed qed next case (e e) then show ?thesis using k1 s1 v1 assms(1) assms(2) f1 by auto qed next case (Calldata x2) then show ?thesis using k1 s1 assms(1) f1 by auto next case (Memory x2) then show ?thesis using k1 s1 assms(1) f1 by auto next case (Storage x3) then show ?thesis using k1 s1 assms(1) f1 by auto qed next case (KCDptr x2) then show ?thesis using s1 assms(1) f1 by fastforce next case (KMemptr x2) then show ?thesis using s1 assms(1) f1 by fastforce next case (KStoptr x3) then show ?thesis using s1 assms(1) f1 by fastforce qed qed next case (e e) then show ?thesis using assms(1) by simp qed lemma ssel_eq_ssel: "(\i st. i \ set ix \ expr i ep env cd st = expr (f i) ep env cd st) \ ssel tp loc ix ep env cd st = ssel tp loc (map f ix) ep env cd st" proof (induction ix arbitrary: tp loc ep env cd st) case Nil then show ?case by simp next case c1: (Cons i ix) then show ?case proof (cases tp) case tp1: (STArray al tp) then show ?thesis proof (cases "expr i ep env cd st") case s1: (n a st') then show ?thesis proof (cases a) case f1: (Pair a b) then show ?thesis proof (cases a) case k1: (KValue v) then show ?thesis proof (cases b) case v1: (Value t) then show ?thesis proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") case None with v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp next case s2: (Some a) then show ?thesis proof (cases a) case p1: (Pair a b) then show ?thesis proof (cases b) case (TSInt x1) with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp next case (TUInt x2) with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp next case b1: TBool show ?thesis proof cases assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" from c1.IH[OF c1.prems] have "ssel tp (hash loc v) ix ep env cd st' = ssel tp (hash loc v) (map f ix) ep env cd st'" by simp with mp s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp next assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp qed next case TAddr with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis by simp qed qed qed next case (Calldata x2) with k1 tp1 s1 c1.prems f1 show ?thesis by simp next case (Memory x2) with k1 tp1 s1 c1.prems f1 show ?thesis by simp next case (Storage x3) with k1 tp1 s1 c1.prems f1 show ?thesis by simp qed next case (KCDptr x2) with tp1 s1 c1.prems f1 show ?thesis by simp next case (KMemptr x2) with tp1 s1 c1.prems f1 show ?thesis by simp next case (KStoptr x3) with tp1 s1 c1.prems f1 show ?thesis by simp qed qed next case (e e) with tp1 c1.prems show ?thesis by simp qed next case tp1: (STMap _ t) then show ?thesis proof (cases "expr i ep env cd st") case s1: (n a s) then show ?thesis proof (cases a) case f1: (Pair a b) then show ?thesis proof (cases a) case k1: (KValue v) from c1.IH[OF c1.prems] have "ssel tp (hash loc v) ix ep env cd st = ssel tp (hash loc v) (map f ix) ep env cd st" by simp with k1 tp1 s1 c1 f1 show ?thesis by simp next case (KCDptr x2) with tp1 s1 c1.prems f1 show ?thesis by simp next case (KMemptr x2) with tp1 s1 c1.prems f1 show ?thesis by simp next case (KStoptr x3) with tp1 s1 c1.prems f1 show ?thesis by simp qed qed next case (e e) with tp1 c1.prems show ?thesis by simp qed next case (STValue x2) then show ?thesis by simp qed qed lemma msel_eq_msel: "(\i st. i \ set ix \ expr i ep env cd st = expr (f i) ep env cd st) \ msel c tp loc ix ep env cd st = msel c tp loc (map f ix) ep env cd st" proof (induction ix arbitrary: c tp loc ep env cd st) case Nil then show ?case by simp next case c1: (Cons i ix) then show ?case proof (cases tp) case tp1: (MTArray al tp) then show ?thesis proof (cases ix) case Nil thus ?thesis using tp1 c1.prems by auto next case c2: (Cons a list) then show ?thesis proof (cases "expr i ep env cd st") case s1: (n a st') then show ?thesis proof (cases a) case f1: (Pair a b) then show ?thesis proof (cases a) case k1: (KValue v) then show ?thesis proof (cases b) case v1: (Value t) then show ?thesis proof (cases "less t (TUInt 256) v (ShowL\<^sub>i\<^sub>n\<^sub>t al)") case None with v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case s2: (Some a) then show ?thesis proof (cases a) case p1: (Pair a b) then show ?thesis proof (cases b) case (TSInt x1) with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case (TUInt x2) with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case b1: TBool show ?thesis proof cases assume "a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" then show ?thesis proof (cases c) case True then show ?thesis proof (cases "accessStore (hash loc v) (memory st')") case None with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp next case mp: (MPointer l) from c1.IH[OF c1.prems] have "msel c tp l ix ep env cd st' = msel c tp l (map f ix) ep env cd st'" by simp with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 True show ?thesis using c2 by simp qed qed next case False then show ?thesis proof (cases "accessStore (hash loc v) cd") case None with s2 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp next case s3: (Some a) then show ?thesis proof (cases a) case (MValue x1) with s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp next case mp: (MPointer l) from c1.IH[OF c1.prems] have "msel c tp l ix ep env cd st' = msel c tp l (map f ix) ep env cd st'" by simp with mp s2 s3 b1 p1 v1 k1 tp1 s1 c1.prems f1 False show ?thesis using c2 by simp qed qed qed next assume "\ a = ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l True" with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp qed next case TAddr with s2 p1 v1 k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp qed qed qed next case (Calldata x2) with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case (Memory x2) with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case (Storage x3) with k1 tp1 s1 c1.prems f1 show ?thesis using c2 by simp qed next case (KCDptr x2) with tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case (KMemptr x2) with tp1 s1 c1.prems f1 show ?thesis using c2 by simp next case (KStoptr x3) with tp1 s1 c1.prems f1 show ?thesis using c2 by simp qed qed next case (e e) with tp1 c1.prems show ?thesis using c2 by simp qed qed next case (MTValue x2) then show ?thesis by simp qed qed lemma ref_eq: assumes "\e st. e \ set ex \ expr e ep env cd st = expr (f e) ep env cd st" shows "rexp (Ref i ex) ep env cd st=rexp (Ref i (map f ex)) ep env cd st" proof (cases "fmlookup (denvalue env) i") case None then show ?thesis by simp next case s1: (Some a) then show ?thesis proof (cases a) case p1: (Pair tp b) then show ?thesis proof (cases b) case k1: (Stackloc l) then show ?thesis proof (cases "accessStore l (stack st)") case None with s1 p1 k1 show ?thesis by simp next case s2: (Some a') then show ?thesis proof (cases a') case (KValue _) with s1 s2 p1 k1 show ?thesis by simp next case cp: (KCDptr cp) then show ?thesis proof (cases tp) case (Value x1) with mp s1 s2 p1 k1 show ?thesis by simp next case mt: (Calldata ct) from msel_eq_msel have "msel False ct cp ex ep env cd st=msel False ct cp (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 s2 p1 k1 mt cp by simp next case mt: (Memory mt) from msel_eq_msel have "msel True mt cp ex ep env cd st=msel True mt cp (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 s2 p1 k1 mt cp by simp next case (Storage x3) with cp s1 s2 p1 k1 show ?thesis by simp qed next case mp: (KMemptr mp) then show ?thesis proof (cases tp) case (Value x1) with mp s1 s2 p1 k1 show ?thesis by simp next case mt: (Calldata ct) from msel_eq_msel have "msel True ct mp ex ep env cd st=msel True ct mp (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 s2 p1 k1 mt mp by simp next case mt: (Memory mt) from msel_eq_msel have "msel True mt mp ex ep env cd st=msel True mt mp (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 s2 p1 k1 mt mp by simp next case (Storage x3) with mp s1 s2 p1 k1 show ?thesis by simp qed next case sp: (KStoptr sp) then show ?thesis proof (cases tp) case (Value x1) then show ?thesis using s1 s2 p1 k1 sp by simp next case (Calldata x2) then show ?thesis using s1 s2 p1 k1 sp by simp next case (Memory x2) then show ?thesis using s1 s2 p1 k1 sp by simp next case st: (Storage stp) from ssel_eq_ssel have "ssel stp sp ex ep env cd st=ssel stp sp (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 s2 p1 k1 st sp by simp qed qed qed next case sl:(Storeloc sl) then show ?thesis proof (cases tp) case (Value x1) then show ?thesis using s1 p1 sl by simp next case (Calldata x2) then show ?thesis using s1 p1 sl by simp next case (Memory x2) then show ?thesis using s1 p1 sl by simp next case st: (Storage stp) from ssel_eq_ssel have "ssel stp sl ex ep env cd st=ssel stp sl (map f ex) ep env cd st" using assms by blast thus ?thesis using s1 sl p1 st by simp qed qed qed qed text\ The following theorem proves that the update function preserves the semantics of expressions. \ theorem update_correctness: "\st lb lv. expr ex ep env cd st = expr (eupdate ex) ep env cd st" "\st. rexp lv ep env cd st = rexp (lupdate lv) ep env cd st" proof (induction ex and lv) case (Id x) then show ?case by simp next case (Ref d ix) then show ?case using ref_eq[where f="eupdate"] by simp next case (INT b v) then show ?case proof (cases "gas st > 0") case True then show ?thesis proof cases assume "b\vbits" show ?thesis proof cases let ?m_def = "(-(2^(b-1)) + (v+2^(b-1)) mod (2^b))" assume "v \ 0" from `b\vbits` True have "expr (E.INT b v) ep env cd st = Normal ((KValue (createSInt b v), Value (TSInt b)), st)" by simp also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `v \ 0` by auto also from `v \ 0` `b\vbits` True have "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),st) = expr (eupdate (E.INT b v)) ep env cd st" by simp finally show "expr (E.INT b v) ep env cd st = expr (eupdate (E.INT b v)) ep env cd st" by simp next let ?m_def = "(2^(b-1) - (-v+2^(b-1)-1) mod (2^b) - 1)" assume "\ v \ 0" from `b\vbits` True have "expr (E.INT b v) ep env cd st = Normal ((KValue (createSInt b v), Value (TSInt b)), st)" by simp also have "createSInt b v = createSInt b ?m_def" using `b\vbits` `\ v \ 0` by auto also from `\ v \ 0` `b\vbits` True have "Normal ((KValue (createSInt b ?m_def), Value (TSInt b)),st) =expr (eupdate (E.INT b v)) ep env cd st" by simp finally show "expr (E.INT b v) ep env cd st = expr (eupdate (E.INT b v)) ep env cd st" by simp qed next assume "\ b\vbits" thus ?thesis by auto qed next case False then show ?thesis using no_gas by simp qed next case (UINT x1 x2) then show ?case by simp next case (ADDRESS x) then show ?case by simp next case (BALANCE x) then show ?case by simp next case THIS then show ?case by simp next case SENDER then show ?case by simp next case VALUE then show ?case by simp next case TRUE then show ?case by simp next case FALSE then show ?case by simp next case (LVAL x) then show ?case by simp next case p: (PLUS e1 e2) show ?case proof (cases "eupdate e1") case i: (INT b1 v1) with p.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp moreover from i `b1 \ vbits` have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" using createSInt_id[of v1 b1] by simp thus ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) with p.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i2 `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp thus ?thesis proof (cases) let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" assume "?v\0" hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp ultimately have "expr (PLUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` `?v\0` have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound3 by simp moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" assume "\ ?v\0" hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "add (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" using Read_ShowL_id add_def olift.simps(1)[of "(+)" b1 b2] by simp ultimately have "expr (PLUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using True r1 r2 by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` `\?v\0` have "eupdate (PLUS e1 e2) = E.INT (max b1 b2) ?x" using i i2 by simp moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign not_exp_less_eq_0_int) + by (simp add: algebra_simps flip: zle_diff1_eq) moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b2 \ vbits" with p i i2 show ?thesis by simp qed next case u2: (UINT b2 v2) with p.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u2 `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" using createUInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b20" hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b1 ?v, TSInt b1)" using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` `b2 \ vbits` `?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" assume "\ ?v\0" hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "add (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b1 ?v, TSInt b1)" using Read_ShowL_id add_def olift.simps(3)[of "(+)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign zero_less_power) + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b2 < b1" with p i u2 show ?thesis by simp qed next assume "\ b2 \ vbits" with p i u2 show ?thesis by simp qed next case (ADDRESS _) with p i show ?thesis by simp next case (BALANCE _) with p i show ?thesis by simp next case THIS with p i show ?thesis by simp next case SENDER with p i show ?thesis by simp next case VALUE with p i show ?thesis by simp next case TRUE with p i show ?thesis by simp next case FALSE with p i show ?thesis by simp next case (LVAL _) with p i show ?thesis by simp next case (PLUS _ _) with p i show ?thesis by simp next case (MINUS _ _) with p i show ?thesis by simp next case (EQUAL _ _) with p i show ?thesis by simp next case (LESS _ _) with p i show ?thesis by simp next case (AND _ _) with p i show ?thesis by simp next case (OR _ _) with p i show ?thesis by simp next case (NOT _) with p i show ?thesis by simp next case (CALL x181 x182) with p i show ?thesis by simp next case (ECALL x191 x192 x193 x194) with p i show ?thesis by simp qed next assume "\ b1 \ vbits" with p i show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case u: (UINT b1 v1) with p.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp moreover from u `b1 \ vbits` have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" by simp thus ?thesis proof (cases "eupdate e2") case u2: (UINT b2 v2) with p.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" let ?x="?v mod 2 ^ max b1 b2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u2 `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" by simp moreover have "add (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" using Read_ShowL_id add_def olift.simps(2)[of "(+)" b1 b2] by simp ultimately have "expr (PLUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" using r1 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "eupdate (PLUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp moreover have "expr (UINT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (UINT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (max b1 b2)" by simp moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next assume "\ b2 \ vbits" with p u u2 show ?thesis by simp qed next case i2: (INT b2 v2) with p.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i2 `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b10" hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b2 ?v, TSInt b2)" using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b1 \ vbits` `b2 \ vbits` `?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" assume "\ ?v\0" hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "add (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b2 ?v, TSInt b2)" using Read_ShowL_id add_def olift.simps(4)[of "(+)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp moreover have "expr (eupdate (PLUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b1 \ vbits` `b2 \ vbits` `\?v\0` `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign zero_less_power) + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b1 < b2" with p u i2 show ?thesis by simp qed next assume "\ b2 \ vbits" with p u i2 show ?thesis by simp qed next case (ADDRESS _) with p u show ?thesis by simp next case (BALANCE _) with p u show ?thesis by simp next case THIS with p u show ?thesis by simp next case SENDER with p u show ?thesis by simp next case VALUE with p u show ?thesis by simp next case TRUE with p u show ?thesis by simp next case FALSE with p u show ?thesis by simp next case (LVAL _) with p u show ?thesis by simp next case (PLUS _ _) with p u show ?thesis by simp next case (MINUS _ _) with p u show ?thesis by simp next case (EQUAL _ _) with p u show ?thesis by simp next case (LESS _ _) with p u show ?thesis by simp next case (AND _ _) with p u show ?thesis by simp next case (OR _ _) with p u show ?thesis by simp next case (NOT _) with p u show ?thesis by simp next case (CALL x181 x182) with p u show ?thesis by simp next case (ECALL x191 x192 x193 x194) with p u show ?thesis by simp qed next assume "\ b1 \ vbits" with p u show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case (ADDRESS x3) with p show ?thesis by simp next case (BALANCE x4) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with p show ?thesis by simp next case SENDER with p show ?thesis by simp next case VALUE with p show ?thesis by simp next case TRUE with p show ?thesis by simp next case FALSE with p show ?thesis by simp next case (LVAL x7) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (PLUS x81 x82) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (ECALL x191 x192 x193 x194) with p show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto qed next case m: (MINUS e1 e2) show ?case proof (cases "eupdate e1") case i: (INT b1 v1) with m.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp moreover from i `b1 \ vbits` have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" using createSInt_id[of v1 b1] by simp thus ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) with m.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 - v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i2 `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp from `b1 \ vbits` `b2 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT (max b1 b2) (- (2 ^ (max b1 b2 - 1)) + (v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2) else E.INT (max b1 b2) (2 ^ (max b1 b2 - 1) - (- v + 2 ^ (max b1 b2 - 1) - 1) mod 2 ^ max b1 b2 - 1))" using i i2 eupdate.simps(11)[of e1 e2] by simp show ?thesis proof (cases) let ?x="- (2 ^ (max b1 b2 - 1)) + (?v + 2 ^ (max b1 b2 - 1)) mod 2 ^ max b1 b2" assume "?v\0" hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp ultimately have "expr (MINUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `?v\0` by simp moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (max b1 b2 - 1)" using upper_bound2 by simp moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1" assume "\ ?v\0" hence "createSInt (max b1 b2) ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "sub (TSInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt (max b1 b2) ?v, TSInt (max b1 b2))" using Read_ShowL_id sub_def olift.simps(1)[of "(-)" b1 b2] by simp ultimately have "expr (MINUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT (max b1 b2) ?x" using `\ ?v\0` by simp moreover have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (E.INT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createSInt (max b1 b2) ?x), Value (TSInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x \ - (2 ^ (max b1 b2 - 1))" using lower_bound2[of "max b1 b2" ?v] by simp moreover from `b1 > 0` have "2^(max b1 b2 -1) > (0::nat)" by simp hence "2^(max b1 b2 -1) - (-?v+2^(max b1 b2-1)-1) mod (2^max b1 b2) - 1 < 2 ^ (max b1 b2 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign not_exp_less_eq_0_int) + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis using createSInt_id[of ?x "max b1 b2"] by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b2 \ vbits" with m i i2 show ?thesis by simp qed next case u: (UINT b2 v2) with m.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 - v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" using createUInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b2 vbits` `b2 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT b1 (- (2 ^ (b1 - 1)) + (v + 2 ^ (b1 - 1)) mod 2 ^ b1) else E.INT b1 (2 ^ (b1 - 1) - (- v + 2 ^ (b1 - 1) - 1) mod 2 ^ b1 - 1))" using i u eupdate.simps(11)[of e1 e2] by simp show ?thesis proof (cases) let ?x="- (2 ^ (b1 - 1)) + (?v + 2 ^ (b1 - 1)) mod 2 ^ b1" assume "?v\0" hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b2i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b1 ?v, TSInt b1)" using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `?v\0` by simp moreover have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (b1 - 1)" using upper_bound2 by simp ultimately show ?thesis using createSInt_id[of ?x "b1"] `0 < b1` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(b1 -1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1" assume "\ ?v\0" hence "createSInt b1 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "sub (TSInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b1 ?v, TSInt b1)" using Read_ShowL_id sub_def olift.simps(3)[of "(-)" b1 b2] `b2i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT b1 ?x" using `\ ?v\0` by simp moreover have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b1)),st)" proof - from `b1 \ vbits` True have "expr (E.INT b1 ?x) ep env cd st = Normal ((KValue (createSInt b1 ?x), Value (TSInt b1)),st)" by simp moreover from `0 < b1` have "?x \ - (2 ^ (b1 - 1))" using upper_bound2 by simp moreover have "2^(b1-1) - (-?v+2^(b1-1)-1) mod (2^b1) - 1 < 2 ^ (b1 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign zero_less_power) + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) ultimately show ?thesis using createSInt_id[of ?x b1] `0 < b1` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b2 < b1" with m i u show ?thesis by simp qed next assume "\ b2 \ vbits" with m i u show ?thesis by simp qed next case (ADDRESS _) with m i show ?thesis by simp next case (BALANCE _) with m i show ?thesis by simp next case THIS with m i show ?thesis by simp next case SENDER with m i show ?thesis by simp next case VALUE with m i show ?thesis by simp next case TRUE with m i show ?thesis by simp next case FALSE with m i show ?thesis by simp next case (LVAL _) with m i show ?thesis by simp next case (PLUS _ _) with m i show ?thesis by simp next case (MINUS _ _) with m i show ?thesis by simp next case (EQUAL _ _) with m i show ?thesis by simp next case (LESS _ _) with m i show ?thesis by simp next case (AND _ _) with m i show ?thesis by simp next case (OR _ _) with m i show ?thesis by simp next case (NOT _) with m i show ?thesis by simp next case (CALL x181 x182) with m i show ?thesis by simp next case (ECALL x191 x192 x193 x194) with m i show ?thesis by simp qed next assume "\ b1 \ vbits" with m i show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case u: (UINT b1 v1) with m.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp moreover from u `b1 \ vbits` have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" by simp thus ?thesis proof (cases "eupdate e2") case u2: (UINT b2 v2) with m.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 - v2" let ?x="?v mod 2 ^ max b1 b2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u2 `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" by simp moreover have "sub (TUInt b1) (TUInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createUInt (max b1 b2) ?v, TUInt (max b1 b2))" using Read_ShowL_id sub_def olift.simps(2)[of "(-)" b1 b2] by simp ultimately have "expr (MINUS e1 e2) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" using r1 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "eupdate (MINUS e1 e2) = UINT (max b1 b2) ?x" using u u2 by simp moreover have "expr (UINT (max b1 b2) ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TUInt (max b1 b2))),st)" proof - from `b1 \ vbits` `b2 \ vbits` have "max b1 b2 \ vbits" using vbits_max by simp with True have "expr (UINT (max b1 b2) ?x) ep env cd st = Normal ((KValue (createUInt (max b1 b2) ?x), Value (TUInt (max b1 b2))),st)" by simp moreover from `0 < b1` have "?x < 2 ^ (max b1 b2)" by simp moreover from `0 < b1` have "0 < max b1 b2" using max_def by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next assume "\ b2 \ vbits" with m u u2 show ?thesis by simp qed next case i: (INT b2 v2) with m.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 - v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b1 vbits` `b2 \ vbits` have u_def: "eupdate (MINUS e1 e2) = (let v = v1 - v2 in if 0 \ v then E.INT b2 (- (2 ^ (b2 - 1)) + (v + 2 ^ (b2 - 1)) mod 2 ^ b2) else E.INT b2 (2 ^ (b2 - 1) - (- v + 2 ^ (b2 - 1) - 1) mod 2 ^ b2 - 1))" using u i by simp show ?thesis proof (cases) let ?x="- (2 ^ (b2 - 1)) + (?v + 2 ^ (b2 - 1)) mod 2 ^ b2" assume "?v\0" hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" using `b1i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b2 ?v, TSInt b2)" using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `?v\0` by simp moreover have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp moreover from `0 < b2` have "?x < 2 ^ (b2 - 1)" using upper_bound2 by simp ultimately show ?thesis using createSInt_id[of ?x "b2"] `0 < b2` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp next let ?x="2^(b2 -1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1" assume "\ ?v\0" hence "createSInt b2 ?v = (ShowL\<^sub>i\<^sub>n\<^sub>t ?x)" by simp moreover have "sub (TUInt b1) (TSInt b2) (ShowL\<^sub>i\<^sub>n\<^sub>t v1) (ShowL\<^sub>i\<^sub>n\<^sub>t v2) = Some (createSInt b2 ?v, TSInt b2)" using Read_ShowL_id sub_def olift.simps(4)[of "(-)" b1 b2] `b1i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" using r1 r2 True by simp moreover have "expr (eupdate (MINUS e1 e2)) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from u_def have "eupdate (MINUS e1 e2) = E.INT b2 ?x" using `\ ?v\0` by simp moreover have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t ?x), Value (TSInt b2)),st)" proof - from `b2 \ vbits` True have "expr (E.INT b2 ?x) ep env cd st = Normal ((KValue (createSInt b2 ?x), Value (TSInt b2)),st)" by simp moreover from `0 < b2` have "?x \ - (2 ^ (b2 - 1))" using upper_bound2 by simp moreover have "2^(b2-1) - (-?v+2^(b2-1)-1) mod (2^b2) - 1 < 2 ^ (b2 - 1)" - by (smt (verit, best) Euclidean_Division.pos_mod_sign zero_less_power) + by (simp add: algebra_simps flip: int_one_le_iff_zero_less) ultimately show ?thesis using createSInt_id[of ?x b2] `0 < b2` by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed next assume "\ b1 < b2" with m u i show ?thesis by simp qed next assume "\ b2 \ vbits" with m u i show ?thesis by simp qed next case (ADDRESS _) with m u show ?thesis by simp next case (BALANCE _) with m u show ?thesis by simp next case THIS with m u show ?thesis by simp next case SENDER with m u show ?thesis by simp next case VALUE with m u show ?thesis by simp next case TRUE with m u show ?thesis by simp next case FALSE with m u show ?thesis by simp next case (LVAL _) with m u show ?thesis by simp next case (PLUS _ _) with m u show ?thesis by simp next case (MINUS _ _) with m u show ?thesis by simp next case (EQUAL _ _) with m u show ?thesis by simp next case (LESS _ _) with m u show ?thesis by simp next case (AND _ _) with m u show ?thesis by simp next case (OR _ _) with m u show ?thesis by simp next case (NOT _) with m u show ?thesis by simp next case (CALL x181 x182) with m u show ?thesis by simp next case (ECALL x191 x192 x193 x194) with m u show ?thesis by simp qed next assume "\ b1 \ vbits" with m u show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case (ADDRESS x3) with m show ?thesis by simp next case (BALANCE x4) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with m show ?thesis by simp next case SENDER with m show ?thesis by simp next case VALUE with m show ?thesis by simp next case TRUE with m show ?thesis by simp next case FALSE with m show ?thesis by simp next case (LVAL x7) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (PLUS x81 x82) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp next case (ECALL x191 x192 x193 x194) with m show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp qed next case e: (EQUAL e1 e2) show ?case proof (cases "eupdate e1") case i: (INT b1 v1) with e.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp moreover from i `b1 \ vbits` have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" using createSInt_id[of v1 b1] by simp thus ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) with e.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i2 `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp with r1 True have "expr (EQUAL e1 e2) ep env cd st= Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using equal_def plift.simps(1)[of "(=)"] by simp hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" using Read_ShowL_id by simp with `b1 \ vbits` `b2 \ vbits` True show ?thesis using i i2 by simp next assume "\ b2 \ vbits" with e i i2 show ?thesis by simp qed next case u: (UINT b2 v2) with e.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" using createUInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using equal_def plift.simps(3)[of "(=)"] by simp hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" using Read_ShowL_id by simp with `b1 \ vbits` `b2 \ vbits` `b2 b2 < b1" with e i u show ?thesis by simp qed next assume "\ b2 \ vbits" with e i u show ?thesis by simp qed next case (ADDRESS _) with e i show ?thesis by simp next case (BALANCE _) with e i show ?thesis by simp next case THIS with e i show ?thesis by simp next case SENDER with e i show ?thesis by simp next case VALUE with e i show ?thesis by simp next case TRUE with e i show ?thesis by simp next case FALSE with e i show ?thesis by simp next case (LVAL _) with e i show ?thesis by simp next case (PLUS _ _) with e i show ?thesis by simp next case (MINUS _ _) with e i show ?thesis by simp next case (EQUAL _ _) with e i show ?thesis by simp next case (LESS _ _) with e i show ?thesis by simp next case (AND _ _) with e i show ?thesis by simp next case (OR _ _) with e i show ?thesis by simp next case (NOT _) with e i show ?thesis by simp next case (CALL x181 x182) with e i show ?thesis by simp next case (ECALL x191 x192 x193 x194) with e i show ?thesis by simp qed next assume "\ b1 \ vbits" with e i show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case u: (UINT b1 v1) with e.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp moreover from u `b1 \ vbits` have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" by simp thus ?thesis proof (cases "eupdate e2") case u2: (UINT b2 v2) with e.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u2 `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" by simp with r1 True have "expr (EQUAL e1 e2) ep env cd st= Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using equal_def plift.simps(2)[of "(=)"] by simp hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" using Read_ShowL_id by simp with `b1 \ vbits` `b2 \ vbits` show ?thesis using u u2 True by simp next assume "\ b2 \ vbits" with e u u2 show ?thesis by simp qed next case i: (INT b2 v2) with e.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))=((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using equal_def plift.simps(4)[of "(=)"] by simp hence "expr (EQUAL e1 e2) ep env cd st=Normal ((KValue (createBool (v1=v2)), Value TBool),st)" using Read_ShowL_id by simp with `b1 \ vbits` `b2 \ vbits` `b1 b1 < b2" with e u i show ?thesis by simp qed next assume "\ b2 \ vbits" with e u i show ?thesis by simp qed next case (ADDRESS _) with e u show ?thesis by simp next case (BALANCE _) with e u show ?thesis by simp next case THIS with e u show ?thesis by simp next case SENDER with e u show ?thesis by simp next case VALUE with e u show ?thesis by simp next case TRUE with e u show ?thesis by simp next case FALSE with e u show ?thesis by simp next case (LVAL _) with e u show ?thesis by simp next case (PLUS _ _) with e u show ?thesis by simp next case (MINUS _ _) with e u show ?thesis by simp next case (EQUAL _ _) with e u show ?thesis by simp next case (LESS _ _) with e u show ?thesis by simp next case (AND _ _) with e u show ?thesis by simp next case (OR _ _) with e u show ?thesis by simp next case (NOT _) with e u show ?thesis by simp next case (CALL x181 x182) with e u show ?thesis by simp next case (ECALL x191 x192 x193 x194) with e u show ?thesis by simp qed next assume "\ b1 \ vbits" with e u show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case (ADDRESS x3) with e show ?thesis by simp next case (BALANCE x4) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with e show ?thesis by simp next case SENDER with e show ?thesis by simp next case VALUE with e show ?thesis by simp next case TRUE with e show ?thesis by simp next case FALSE with e show ?thesis by simp next case (LVAL x7) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (PLUS x81 x82) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp next case (ECALL x191 x192 x193 x194) with e show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp qed next case l: (LESS e1 e2) show ?case proof (cases "eupdate e1") case i: (INT b1 v1) with l.IH have expr1: "expr e1 ep env cd st = expr (E.INT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createSInt b1 v1), Value (TSInt b1)),st)" by simp moreover from i `b1 \ vbits` have "v1 < 2^(b1-1)" and "v1 \ -(2^(b1-1))" using update_bounds_int by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TSInt b1)),st)" using createSInt_id[of v1 b1] by simp thus ?thesis proof (cases "eupdate e2") case i2: (INT b2 v2) with l.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i2 `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp with r1 True have "expr (LESS e1 e2) ep env cd st= Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using less_def plift.simps(1)[of "(<)"] by simp hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using i i2 True by simp next assume "\ b2 \ vbits" with l i i2 show ?thesis by simp qed next case u: (UINT b2 v2) with l.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" using createUInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b2i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using less_def plift.simps(3)[of "(<)"] by simp hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b2 b2 < b1" with l i u show ?thesis by simp qed next assume "\ b2 \ vbits" with l i u show ?thesis by simp qed next case (ADDRESS _) with l i show ?thesis by simp next case (BALANCE _) with l i show ?thesis by simp next case THIS with l i show ?thesis by simp next case SENDER with l i show ?thesis by simp next case VALUE with l i show ?thesis by simp next case TRUE with l i show ?thesis by simp next case FALSE with l i show ?thesis by simp next case (LVAL _) with l i show ?thesis by simp next case (PLUS _ _) with l i show ?thesis by simp next case (MINUS _ _) with l i show ?thesis by simp next case (EQUAL _ _) with l i show ?thesis by simp next case (LESS _ _) with l i show ?thesis by simp next case (AND _ _) with l i show ?thesis by simp next case (OR _ _) with l i show ?thesis by simp next case (NOT _) with l i show ?thesis by simp next case (CALL x181 x182) with l i show ?thesis by simp next case (ECALL x191 x192 x193 x194) with l i show ?thesis by simp qed next assume "\ b1 \ vbits" with l i show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case u: (UINT b1 v1) with l.IH have expr1: "expr e1 ep env cd st = expr (UINT b1 v1) ep env cd st" by simp then show ?thesis proof (cases "gas st > 0") case True then show ?thesis proof (cases) assume "b1 \ vbits" with expr1 True have "expr e1 ep env cd st=Normal ((KValue (createUInt b1 v1), Value (TUInt b1)),st)" by simp moreover from u `b1 \ vbits` have "v1 < 2^b1" and "v1 \ 0" using update_bounds_uint by auto moreover from `b1 \ vbits` have "0 < b1" by auto ultimately have r1: "expr e1 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v1), Value (TUInt b1)),st)" by simp thus ?thesis proof (cases "eupdate e2") case u2: (UINT b2 v2) with l.IH have expr2: "expr e2 ep env cd st = expr (UINT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createUInt b2 v2), Value (TUInt b2)),st)" by simp moreover from u2 `b2 \ vbits` have "v2 < 2^b2" and "v2 \ 0" using update_bounds_uint by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TUInt b2)),st)" by simp with r1 True have "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool ((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using less_def plift.simps(2)[of "(<)"] by simp hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` show ?thesis using u u2 True by simp next assume "\ b2 \ vbits" with l u u2 show ?thesis by simp qed next case i: (INT b2 v2) with l.IH have expr2: "expr e2 ep env cd st = expr (E.INT b2 v2) ep env cd st" by simp then show ?thesis proof (cases) let ?v="v1 + v2" assume "b2 \ vbits" with expr2 True have "expr e2 ep env cd st=Normal ((KValue (createSInt b2 v2), Value (TSInt b2)),st)" by simp moreover from i `b2 \ vbits` have "v2 < 2^(b2-1)" and "v2 \ -(2^(b2-1))" using update_bounds_int by auto moreover from `b2 \ vbits` have "0 < b2" by auto ultimately have r2: "expr e2 ep env cd st = Normal ((KValue (ShowL\<^sub>i\<^sub>n\<^sub>t v2), Value (TSInt b2)),st)" using createSInt_id[of v2 b2] by simp thus ?thesis proof (cases) assume "b1i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v1))<((ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t v2))))), Value TBool),st)" using less_def plift.simps(4)[of "(<)"] by simp hence "expr (LESS e1 e2) ep env cd st=Normal ((KValue (createBool (v1 vbits` `b2 \ vbits` `b1 b1 < b2" with l u i show ?thesis by simp qed next assume "\ b2 \ vbits" with l u i show ?thesis by simp qed next case (ADDRESS _) with l u show ?thesis by simp next case (BALANCE _) with l u show ?thesis by simp next case THIS with l u show ?thesis by simp next case SENDER with l u show ?thesis by simp next case VALUE with l u show ?thesis by simp next case TRUE with l u show ?thesis by simp next case FALSE with l u show ?thesis by simp next case (LVAL _) with l u show ?thesis by simp next case (PLUS _ _) with l u show ?thesis by simp next case (MINUS _ _) with l u show ?thesis by simp next case (EQUAL _ _) with l u show ?thesis by simp next case (LESS _ _) with l u show ?thesis by simp next case (AND _ _) with l u show ?thesis by simp next case (OR _ _) with l u show ?thesis by simp next case (NOT _) with l u show ?thesis by simp next case (CALL x181 x182) with l u show ?thesis by simp next case (ECALL x191 x192 x193 x194) with l u show ?thesis by simp qed next assume "\ b1 \ vbits" with l u show ?thesis by simp qed next case False then show ?thesis using no_gas by simp qed next case (ADDRESS x3) with l show ?thesis by simp next case (BALANCE x4) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with l show ?thesis by simp next case SENDER with l show ?thesis by simp next case VALUE with l show ?thesis by simp next case TRUE with l show ?thesis by simp next case FALSE with l show ?thesis by simp next case (LVAL x7) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (PLUS x81 x82) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp next case (ECALL x191 x192 x193 x194) with l show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp qed next case a: (AND e1 e2) show ?case proof (cases "eupdate e1") case (INT x11 x12) with a show ?thesis by simp next case (UINT x21 x22) with a show ?thesis by simp next case (ADDRESS x3) with a show ?thesis by simp next case (BALANCE x4) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with a show ?thesis by simp next case SENDER with a show ?thesis by simp next case VALUE with a show ?thesis by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with a t show ?thesis by simp next case (UINT x21 x22) with a t show ?thesis by simp next case (ADDRESS x3) with a t show ?thesis by simp next case (BALANCE x4) with a t show ?thesis by simp next case THIS with a t show ?thesis by simp next case SENDER with a t show ?thesis by simp next case VALUE with a t show ?thesis by simp next case TRUE with a t show ?thesis by simp next case FALSE with a t show ?thesis by simp next case (LVAL x7) with a t show ?thesis by simp next case (PLUS x81 x82) with a t show ?thesis by simp next case (MINUS x91 x92) with a t show ?thesis by simp next case (EQUAL x101 x102) with a t show ?thesis by simp next case (LESS x111 x112) with a t show ?thesis by simp next case (AND x121 x122) with a t show ?thesis by simp next case (OR x131 x132) with a t show ?thesis by simp next case (NOT x131) with a t show ?thesis by simp next case (CALL x181 x182) with a t show ?thesis by simp next case (ECALL x191 x192 x193 x194) with a t show ?thesis by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT b v) with a f show ?thesis by simp next case (UINT b v) with a f show ?thesis by simp next case (ADDRESS x3) with a f show ?thesis by simp next case (BALANCE x4) with a f show ?thesis by simp next case THIS with a f show ?thesis by simp next case SENDER with a f show ?thesis by simp next case VALUE with a f show ?thesis by simp next case TRUE with a f show ?thesis by simp next case FALSE with a f show ?thesis by simp next case (LVAL x7) with a f show ?thesis by simp next case (PLUS x81 x82) with a f show ?thesis by simp next case (MINUS x91 x92) with a f show ?thesis by simp next case (EQUAL x101 x102) with a f show ?thesis by simp next case (LESS x111 x112) with a f show ?thesis by simp next case (AND x121 x122) with a f show ?thesis by simp next case (OR x131 x132) with a f show ?thesis by simp next case (NOT x131) with a f show ?thesis by simp next case (CALL x181 x182) with a f show ?thesis by simp next case (ECALL x191 x192 x193 x194) with a f show ?thesis by simp qed next case (LVAL x7) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case p: (PLUS x81 x82) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp next case (ECALL x191 x192 x193 x194) with a show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp qed next case o: (OR e1 e2) show ?case proof (cases "eupdate e1") case (INT x11 x12) with o show ?thesis by simp next case (UINT x21 x22) with o show ?thesis by simp next case (ADDRESS x3) with o show ?thesis by simp next case (BALANCE x4) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case THIS with o show ?thesis by simp next case SENDER with o show ?thesis by simp next case VALUE with o show ?thesis by simp next case t: TRUE show ?thesis proof (cases "eupdate e2") case (INT x11 x12) with o t show ?thesis by simp next case (UINT x21 x22) with o t show ?thesis by simp next case (ADDRESS x3) with o t show ?thesis by simp next case (BALANCE x4) with o t show ?thesis by simp next case THIS with o t show ?thesis by simp next case SENDER with o t show ?thesis by simp next case VALUE with o t show ?thesis by simp next case TRUE with o t show ?thesis by simp next case FALSE with o t show ?thesis by simp next case (LVAL x7) with o t show ?thesis by simp next case (PLUS x81 x82) with o t show ?thesis by simp next case (MINUS x91 x92) with o t show ?thesis by simp next case (EQUAL x101 x102) with o t show ?thesis by simp next case (LESS x111 x112) with o t show ?thesis by simp next case (AND x121 x122) with o t show ?thesis by simp next case (OR x131 x132) with o t show ?thesis by simp next case (NOT x131) with o t show ?thesis by simp next case (CALL x181 x182) with o t show ?thesis by simp next case (ECALL x191 x192 x193 x194) with o t show ?thesis by simp qed next case f: FALSE show ?thesis proof (cases "eupdate e2") case (INT b v) with o f show ?thesis by simp next case (UINT b v) with o f show ?thesis by simp next case (ADDRESS x3) with o f show ?thesis by simp next case (BALANCE x4) with o f show ?thesis by simp next case THIS with o f show ?thesis by simp next case SENDER with o f show ?thesis by simp next case VALUE with o f show ?thesis by simp next case TRUE with o f show ?thesis by simp next case FALSE with o f show ?thesis by simp next case (LVAL x7) with o f show ?thesis by simp next case (PLUS x81 x82) with o f show ?thesis by simp next case (MINUS x91 x92) with o f show ?thesis by simp next case (EQUAL x101 x102) with o f show ?thesis by simp next case (LESS x111 x112) with o f show ?thesis by simp next case (AND x121 x122) with o f show ?thesis by simp next case (OR x131 x132) with o f show ?thesis by simp next case (NOT x131) with o f show ?thesis by simp next case (CALL x181 x182) with o f show ?thesis by simp next case (ECALL x191 x192 x193 x194) with o f show ?thesis by simp qed next case (LVAL x7) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case p: (PLUS x81 x82) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (MINUS x91 x92) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (EQUAL x101 x102) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (LESS x111 x112) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (AND x121 x122) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (OR x131 x132) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (NOT x131) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by auto next case (CALL x181 x182) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis using lift_eq[of e1 ep env cd st "eupdate e1" e2 "eupdate e2"] by simp qed next case o: (NOT e) show ?case proof (cases "eupdate e") case (INT x11 x12) with o show ?thesis by simp next case (UINT x21 x22) with o show ?thesis by simp next case (ADDRESS x3) with o show ?thesis by simp next case (BALANCE x4) with o show ?thesis by simp next case THIS with o show ?thesis by simp next case SENDER with o show ?thesis by simp next case VALUE with o show ?thesis by simp next case t: TRUE with o show ?thesis by simp next case f: FALSE with o show ?thesis by simp next case (LVAL x7) with o show ?thesis by simp next case p: (PLUS x81 x82) with o show ?thesis by simp next case (MINUS x91 x92) with o show ?thesis by simp next case (EQUAL x101 x102) with o show ?thesis by simp next case (LESS x111 x112) with o show ?thesis by simp next case (AND x121 x122) with o show ?thesis by simp next case (OR x131 x132) with o show ?thesis by simp next case (NOT x131) with o show ?thesis by simp next case (CALL x181 x182) with o show ?thesis by simp next case (ECALL x191 x192 x193 x194) with o show ?thesis by simp qed next case (CALL x181 x182) show ?case by simp next case (ECALL x191 x192 x193 x194) show ?case by simp qed end \ No newline at end of file diff --git a/thys/Solidity/ReadShow.thy b/thys/Solidity/ReadShow.thy --- a/thys/Solidity/ReadShow.thy +++ b/thys/Solidity/ReadShow.thy @@ -1,433 +1,438 @@ section\Converting Types to Strings and Back Again\ theory ReadShow imports Solidity_Symbex begin text\ In the following, we formalize a family of projection (and injection) functions for injecting (projecting) basic types (i.e., @{type \nat\}, @{type \int\}, and @{type \bool\} in (out) of the domains of strings. We provide variants for the two string representations of Isabelle/HOL, namely @{type \string\} and @{type \String.literal\}. \ subsubsection\Bool\ definition \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = ''True'' then True else False)\ definition \Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then ''True'' else ''False'')\ definition \STR_is_bool s = (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ declare Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] lemma Show_Read_bool_id: \STR_is_bool s \ (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ using STR_is_bool_def by fastforce lemma STR_is_bool_split: \STR_is_bool s \ s = ''False'' \ s = ''True''\ by(auto simp: STR_is_bool_def Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) lemma Read_Show_bool_id: \Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l (Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ by(auto simp: Read\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def Show\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) definition ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l::\String.literal \ bool\ (\\_\\) where \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s = (if s = STR ''True'' then True else False)\ definition ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l:: \bool \ String.literal\ (\\_\\) where \ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b = (if b then STR ''True'' else STR ''False'')\ definition \strL_is_bool' s = (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ declare ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def [solidity_symbex] lemma Show_Read_bool'_id: \strL_is_bool' s \ (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l s) = s)\ using strL_is_bool'_def by fastforce lemma strL_is_bool'_split: \strL_is_bool' s \ s = STR ''False'' \ s = STR ''True''\ by(auto simp: strL_is_bool'_def ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) lemma Read_Show_bool'_id: \ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l (ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l b) = b\ by(auto simp: ReadL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def ShowL\<^sub>b\<^sub>o\<^sub>o\<^sub>l_def) subsubsection\Natural Numbers\ definition nat_of_digit :: \char \ nat\ where \nat_of_digit c = (if c = CHR ''0'' then 0 else if c = CHR ''1'' then 1 else if c = CHR ''2'' then 2 else if c = CHR ''3'' then 3 else if c = CHR ''4'' then 4 else if c = CHR ''5'' then 5 else if c = CHR ''6'' then 6 else if c = CHR ''7'' then 7 else if c = CHR ''8'' then 8 else if c = CHR ''9'' then 9 else undefined)\ declare nat_of_digit_def [solidity_symbex] definition digit_of_nat :: \nat \ char\ where \digit_of_nat x = (if x = 0 then CHR ''0'' else if x = 1 then CHR ''1'' else if x = 2 then CHR ''2'' else if x = 3 then CHR ''3'' else if x = 4 then CHR ''4'' else if x = 5 then CHR ''5'' else if x = 6 then CHR ''6'' else if x = 7 then CHR ''7'' else if x = 8 then CHR ''8'' else if x = 9 then CHR ''9'' else undefined)\ declare digit_of_nat_def [solidity_symbex] lemma nat_of_digit_digit_of_nat_id: \x < 10 \ nat_of_digit (digit_of_nat x) = x\ by(simp add:nat_of_digit_def digit_of_nat_def) lemma img_digit_of_nat: \n < 10 \ digit_of_nat n \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ by(simp add:nat_of_digit_def digit_of_nat_def) lemma digit_of_nat_nat_of_digit_id: \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} \ digit_of_nat (nat_of_digit c) = c\ by(code_simp, auto) definition nat_implode :: \'a::{numeral,power,zero} list \ 'a\ where \nat_implode n = foldr (+) (map (\ (p,d) \ 10 ^ p * d) (enumerate 0 (rev n))) 0\ declare nat_implode_def [solidity_symbex] fun nat_explode' :: \nat \ nat list\ where \nat_explode' x = (case x < 10 of True \ [x mod 10] | _ \ (x mod 10 )#(nat_explode' (x div 10)))\ definition nat_explode :: \nat \ nat list\ where \nat_explode x = (rev (nat_explode' x))\ declare nat_explode_def [solidity_symbex] lemma nat_explode'_not_empty: \nat_explode' n \ []\ by (smt (z3) list.simps(3) nat_explode'.simps) lemma nat_explode_not_empty: \nat_explode n \ []\ using nat_explode'_not_empty nat_explode_def by auto lemma nat_explode'_ne_suc: \\ n. nat_explode' (Suc n) \ nat_explode' n\ by(rule exI [of _ \0::nat\], simp) lemma nat_explode'_digit: \hd (nat_explode' n ) < 10\ proof(induct \n\) case 0 then show ?case by simp next case (Suc n) then show ?case proof (cases \n < 9\) case True then show ?thesis by simp next case False then show ?thesis by simp qed qed lemma div_ten_less: \n \ 0 \ ((n::nat) div 10) < n\ by simp lemma unroll_nat_explode': \\ n < 10 \ (case n < 10 of True \ [n mod 10] | False \ n mod 10 # nat_explode' (n div 10)) = (n mod 10 # nat_explode' (n div 10))\ by simp lemma nat_explode_mod_10_ident: \map (\ x. x mod 10) (nat_explode' n) = nat_explode' n\ proof (cases \n < 10\) case True then show ?thesis by simp next case False then show ?thesis proof (induct \n\ rule: nat_less_induct) case (1 n) note * = this then show ?case using div_ten_less apply(simp (no_asm)) using unroll_nat_explode'[of n] * by (smt (z3) list.simps(8) list.simps(9) mod_div_trivial mod_eq_self_iff_div_eq_0 nat_explode'.simps zero_less_numeral) qed qed -lemma nat_explode'_digits: \\ d \ set (nat_explode' n). d < 10\ - using image_set[of \(\ x. x mod 10)\ \(nat_explode' n)\] - nat_explode_mod_10_ident[symmetric] - by (metis (no_types, lifting) Euclidean_Division.div_eq_0_iff - imageE mod_div_trivial zero_neq_numeral) +lemma nat_explode'_digits: + \\d \ set (nat_explode' n). d < 10\ +proof + fix d + assume \d \ set (nat_explode' n)\ + then have \d \ set (map (\m. m mod 10) (nat_explode' n))\ + by (simp only: nat_explode_mod_10_ident) + then show \d < 10\ + by auto +qed -lemma nat_explode_digits: \\ d \ set (nat_explode n). d < 10\ - using nat_explode_def set_rev - by (metis nat_explode'_digits) +lemma nat_explode_digits: + \\d \ set (nat_explode n). d < 10\ + using nat_explode'_digits [of n] by (simp only: nat_explode_def set_rev) value \nat_implode(nat_explode 42) = 42\ value \nat_explode (Suc 21)\ lemma nat_implode_append: \nat_implode (a@[b]) = (1*b + foldr (+) (map (\(p, y). 10 ^ p * y) (enumerate (Suc 0) (rev a))) 0 )\ by(simp add:nat_implode_def) lemma enumerate_suc: \enumerate (Suc n) l = map (\ (a,b). (a+1::nat,b)) (enumerate n l)\ proof (induction \l\) case Nil then show ?case by simp next case (Cons a x) note * = this then show ?case apply(simp only:enumerate_simps) apply(simp only:\enumerate (Suc n) x = map (\a. case a of (a, b) \ (a + 1, b)) (enumerate n x)\[symmetric]) apply(simp) using * by (metis apfst_conv cond_case_prod_eta enumerate_Suc_eq) qed lemma mult_assoc_aux1: \(\(p, y). 10 ^ p * y) \ (\(a, y). (Suc a, y)) = (\(p, y). (10::nat) * (10 ^ p) * y)\ by(auto simp:o_def) lemma fold_map_transfer: \(foldr (+) (map (\(x,y). 10 * (f (x,y))) l) (0::nat)) = 10 * (foldr (+) (map (\x. (f x)) l) (0::nat))\ proof(induct \l\) case Nil then show ?case by(simp) next case (Cons a l) then show ?case by(simp) qed lemma mult_assoc_aux2: \(\(p, y). 10 * 10 ^ p * (y::nat)) = (\(p, y). 10 * (10 ^ p * y))\ by(auto) lemma nat_implode_explode_id: \nat_implode (nat_explode n) = n\ proof (cases \n=0\) case True note * = this then show ?thesis by (simp add: nat_explode_def nat_implode_def) next case False then show ?thesis proof (induct \n\ rule: nat_less_induct) case (1 n) note * = this then have **: \nat_implode (nat_explode (n div 10)) = n div 10\ proof (cases \n div 10 = 0\) case True then show ?thesis by(simp add: nat_explode_def nat_implode_def) next case False then show ?thesis using div_ten_less[of \n\] * by(simp) qed then show ?case proof (cases \n < 10\) case True then show ?thesis by(simp add: nat_explode_def nat_implode_def) next case False note *** = this then show ?thesis apply(simp (no_asm) add:nat_explode_def del:rev_rev_ident) apply(simp only: bool.case(2)) apply(simp del:nat_explode'.simps rev_rev_ident) apply(fold nat_explode_def) apply(simp only:nat_implode_append) apply(simp add:enumerate_suc) apply(simp only:mult_assoc_aux1) using mult_assoc_aux2 apply(simp) using fold_map_transfer[of \\(p, y). 10 ^ p * y\ \(enumerate 0 (rev (nat_explode (n div 10))))\, simplified] apply(simp) apply(fold nat_implode_def) using ** by simp qed qed qed definition Read\<^sub>n\<^sub>a\<^sub>t :: \string \ nat\ where \Read\<^sub>n\<^sub>a\<^sub>t s = nat_implode (map nat_of_digit s)\ definition Show\<^sub>n\<^sub>a\<^sub>t::"nat \ string" where \Show\<^sub>n\<^sub>a\<^sub>t n = map digit_of_nat (nat_explode n)\ declare Read\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] Show\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] definition \STR_is_nat s = (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ value \Read\<^sub>n\<^sub>a\<^sub>t ''10''\ value \Show\<^sub>n\<^sub>a\<^sub>t 10\ value \Read\<^sub>n\<^sub>a\<^sub>t (Show\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ value \Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t (''10'')) = ''10''\ lemma Show_nat_not_neg: \set (Show\<^sub>n\<^sub>a\<^sub>t n) \{CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''}\ unfolding Show\<^sub>n\<^sub>a\<^sub>t_def using nat_explode_digits[of n] img_digit_of_nat imageE image_set subsetI by (smt (verit) imageE image_set subsetI) lemma Show_nat_not_empty: \(Show\<^sub>n\<^sub>a\<^sub>t n) \ []\ by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_def nat_explode_not_empty) lemma not_hd: \L \ [] \ e \ set(L) \ hd L \ e\ by auto lemma Show_nat_not_neg'': \hd (Show\<^sub>n\<^sub>a\<^sub>t n) \ (CHR ''-'')\ using Show_nat_not_neg[of \n\] Show_nat_not_empty[of \n\] not_hd[of \Show\<^sub>n\<^sub>a\<^sub>t n\] by auto lemma Show_Read_nat_id: \STR_is_nat s \ (Show\<^sub>n\<^sub>a\<^sub>t (Read\<^sub>n\<^sub>a\<^sub>t s) = s)\ by(simp add:STR_is_nat_def) lemma bar': \\ d \ set l . d < 10 \ map nat_of_digit (map digit_of_nat l) = l\ using nat_of_digit_digit_of_nat_id by (simp add: map_idI) lemma Read_Show_nat_id: \Read\<^sub>n\<^sub>a\<^sub>t(Show\<^sub>n\<^sub>a\<^sub>t n) = n\ apply(unfold Read\<^sub>n\<^sub>a\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_def) using bar' nat_of_digit_digit_of_nat_id nat_explode_digits using nat_implode_explode_id by presburger definition ReadL\<^sub>n\<^sub>a\<^sub>t :: \String.literal \ nat\ (\\_\\) where \ReadL\<^sub>n\<^sub>a\<^sub>t = Read\<^sub>n\<^sub>a\<^sub>t \ String.explode\ definition ShowL\<^sub>n\<^sub>a\<^sub>t::\nat \ String.literal\ (\\_\\)where \ShowL\<^sub>n\<^sub>a\<^sub>t = String.implode \ Show\<^sub>n\<^sub>a\<^sub>t\ declare ReadL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] ShowL\<^sub>n\<^sub>a\<^sub>t_def [solidity_symbex] definition \strL_is_nat' s = (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ value \\STR ''10''\::nat\ value \ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')\ value \\10::nat\\ value \ShowL\<^sub>n\<^sub>a\<^sub>t 10\ value \ReadL\<^sub>n\<^sub>a\<^sub>t (ShowL\<^sub>n\<^sub>a\<^sub>t (10)) = 10\ value \ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t (STR ''10'')) = STR ''10''\ lemma Show_Read_nat'_id: \strL_is_nat' s \ (ShowL\<^sub>n\<^sub>a\<^sub>t (ReadL\<^sub>n\<^sub>a\<^sub>t s) = s)\ by(simp add:strL_is_nat'_def) lemma digits_are_ascii: \c \ {CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9''} \ String.ascii_of c = c\ by auto lemma Show\<^sub>n\<^sub>a\<^sub>t_ascii: \map String.ascii_of (Show\<^sub>n\<^sub>a\<^sub>t n) = Show\<^sub>n\<^sub>a\<^sub>t n\ using Show_nat_not_neg digits_are_ascii by (meson map_idI subsetD) lemma Read_Show_nat'_id: \ReadL\<^sub>n\<^sub>a\<^sub>t(ShowL\<^sub>n\<^sub>a\<^sub>t n) = n\ apply(unfold ReadL\<^sub>n\<^sub>a\<^sub>t_def ShowL\<^sub>n\<^sub>a\<^sub>t_def, simp) by (simp add: Show\<^sub>n\<^sub>a\<^sub>t_ascii Read_Show_nat_id) subsubsection\Integer\ definition Read\<^sub>i\<^sub>n\<^sub>t :: \string \ int\ where \Read\<^sub>i\<^sub>n\<^sub>t x = (if hd x = (CHR ''-'') then -(int (Read\<^sub>n\<^sub>a\<^sub>t (tl x))) else int (Read\<^sub>n\<^sub>a\<^sub>t x))\ definition Show\<^sub>i\<^sub>n\<^sub>t::\int \ string\ where \Show\<^sub>i\<^sub>n\<^sub>t i = (if i < 0 then (CHR ''-'')#(Show\<^sub>n\<^sub>a\<^sub>t (nat (-i))) else Show\<^sub>n\<^sub>a\<^sub>t (nat i))\ definition \STR_is_int s = (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ declare Read\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] Show\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t 10) = 10\ value \Read\<^sub>i\<^sub>n\<^sub>t (Show\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''10'')) = ''10''\ value \Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t (''-10'')) = ''-10''\ lemma Show_Read_id: \STR_is_int s \ (Show\<^sub>i\<^sub>n\<^sub>t (Read\<^sub>i\<^sub>n\<^sub>t s) = s)\ by(simp add:STR_is_int_def) lemma Read_Show_id: \Read\<^sub>i\<^sub>n\<^sub>t(Show\<^sub>i\<^sub>n\<^sub>t(x)) = x\ apply(code_simp) apply(auto simp:Show_nat_not_neg Read_Show_nat_id)[1] apply(thin_tac \\ x < 0 \) using Show_nat_not_neg'' by blast lemma STR_is_int_Show: \STR_is_int (Show\<^sub>i\<^sub>n\<^sub>t n)\ by(simp add:STR_is_int_def Read_Show_id) definition ReadL\<^sub>i\<^sub>n\<^sub>t :: \String.literal \ int\ (\\_\\) where \ReadL\<^sub>i\<^sub>n\<^sub>t = Read\<^sub>i\<^sub>n\<^sub>t \ String.explode\ definition ShowL\<^sub>i\<^sub>n\<^sub>t::\int \ String.literal\ (\\_\\) where \ShowL\<^sub>i\<^sub>n\<^sub>t =String.implode \ Show\<^sub>i\<^sub>n\<^sub>t\ definition \strL_is_int' s = (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ declare ReadL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] ShowL\<^sub>i\<^sub>n\<^sub>t_def [solidity_symbex] value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t 10) = 10\ value \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t (-10)) = -10\ value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''10'')) = STR ''10''\ value \ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t (STR ''-10'')) = STR ''-10''\ lemma Show_ReadL_id: \strL_is_int' s \ (ShowL\<^sub>i\<^sub>n\<^sub>t (ReadL\<^sub>i\<^sub>n\<^sub>t s) = s)\ by(simp add:strL_is_int'_def) lemma Read_ShowL_id: \ReadL\<^sub>i\<^sub>n\<^sub>t (ShowL\<^sub>i\<^sub>n\<^sub>t x) = x\ proof(cases \x < 0\) case True then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii by (metis (no_types, lifting) Read_Show_id String.ascii_of_Char comp_def implode.rep_eq list.simps(9)) next case False then show ?thesis using ShowL\<^sub>i\<^sub>n\<^sub>t_def ReadL\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>i\<^sub>n\<^sub>t_def Show\<^sub>n\<^sub>a\<^sub>t_ascii by (metis Read_Show_id String.explode_implode_eq comp_apply) qed lemma STR_is_int_ShowL: \strL_is_int' (ShowL\<^sub>i\<^sub>n\<^sub>t n)\ by(simp add:strL_is_int'_def Read_ShowL_id) lemma String_Cancel: "a + (c::String.literal) = b + c \ a = b" using String.plus_literal.rep_eq by (metis append_same_eq literal.explode_inject) end diff --git a/thys/Subresultants/Binary_Exponentiation.thy b/thys/Subresultants/Binary_Exponentiation.thy --- a/thys/Subresultants/Binary_Exponentiation.thy +++ b/thys/Subresultants/Binary_Exponentiation.thy @@ -1,59 +1,59 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Binary Exponentiation\ text \This theory defines the standard algorithm for binary exponentiation, or exponentiation by squaring.\ theory Binary_Exponentiation imports Main begin -declare Euclidean_Division.divmod_nat_def[termination_simp] +declare Euclidean_Rings.divmod_nat_def[termination_simp] context monoid_mult begin fun binary_power :: "'a \ nat \ 'a" where "binary_power x n = (if n = 0 then 1 else - let (d,r) = Euclidean_Division.divmod_nat n 2; + let (d,r) = Euclidean_Rings.divmod_nat n 2; rec = binary_power (x * x) d in if r = 0 then rec else rec * x)" lemma binary_power[simp]: "binary_power = (^)" proof (intro ext) fix x n show "binary_power x n = x ^ n" proof (induct x n rule: binary_power.induct) case (1 x n) show ?case proof (cases "n = 0") case False note IH = 1[OF False] - obtain d r where n2: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force - from Euclidean_Division.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto + obtain d r where n2: "Euclidean_Rings.divmod_nat n 2 = (d,r)" by force + from Euclidean_Rings.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto hence r: "r = 0 \ r = 1" by auto let ?rec = "binary_power (x * x) d" have "binary_power x n = (if r = 0 then ?rec else ?rec * x)" unfolding binary_power.simps[of x n] n2 using False by auto also have "\ = ?rec * x ^ r" using r by (cases "r = 0", auto) also have "?rec = (x * x) ^ d" by (rule IH[OF _ refl], simp add: n2) also have "\ = x ^ (d + d)" unfolding power_add using power2_eq_square power_even_eq power_mult by auto also have "\ * x ^ r = x ^ (d + d + r)" by (simp add: power_add) also have "d + d + r = n" unfolding dr by presburger finally show ?thesis . qed auto qed qed lemma binary_power_code_unfold[code_unfold]: "(^) = binary_power" by simp declare binary_power.simps[simp del] end end diff --git a/thys/Subresultants/Dichotomous_Lazard.thy b/thys/Subresultants/Dichotomous_Lazard.thy --- a/thys/Subresultants/Dichotomous_Lazard.thy +++ b/thys/Subresultants/Dichotomous_Lazard.thy @@ -1,173 +1,173 @@ section \Dichotomous Lazard\ text \This theory contains Lazard's optimization in the computation of the subresultant PRS as described by Ducos \<^cite>\\Section 2\ in "Ducos"\.\ theory Dichotomous_Lazard imports "HOL-Computational_Algebra.Polynomial_Factorial" (* to_fract *) begin lemma power_fract[simp]: "(Fract a b)^n = Fract (a^n) (b^n)" by (induct n, auto simp: fract_collapse) lemma range_to_fract_dvd_iff: assumes b: "b \ 0" shows "Fract a b \ range to_fract \ b dvd a" proof assume "b dvd a" then obtain c where a: "a = b * c" unfolding dvd_def by auto have "Fract a b = Fract c 1" using b unfolding a by (simp add: eq_fract) thus "Fract a b \ range to_fract" unfolding to_fract_def by auto next assume "Fract a b \ range to_fract" then obtain c where "Fract a b = Fract c 1" unfolding to_fract_def by auto hence "a = b * c" using b by (simp add: eq_fract) thus "b dvd a" .. qed lemma Fract_cases_coprime [cases type: fract]: fixes q :: "'a :: factorial_ring_gcd fract" obtains (Fract) a b where "q = Fract a b" "b \ 0" "coprime a b" proof - obtain a b where q: "q = Fract a b" and b0: "b \ 0" by (cases q, auto) define g where g: "g = gcd a b" define A where A: "A = a div g" define B where B: "B = b div g" have a: "a = A * g" unfolding A g by simp have b: "b = B * g" unfolding B g by simp from b0 b have 0: "B \ 0" by auto have q: "q = Fract A B" unfolding q a b by (subst eq_fract, auto simp: b0 0 g) have cop: "coprime A B" unfolding A B g using b0 by (simp add: div_gcd_coprime) assume "\a b. q = Fract a b \ b \ 0 \ coprime a b \ thesis" from this[OF q 0 cop] show ?thesis . qed lemma to_fract_power_le: fixes a :: "'a :: factorial_ring_gcd fract" assumes no_fract: "a * b ^ e \ range to_fract" and a: "a \ range to_fract" and le: "f \ e" shows "a * b ^ f \ range to_fract" proof - obtain bn bd where b: "b = Fract bn bd" and bd: "bd \ 0" and copb: "coprime bn bd" by (cases b, auto) obtain an where a: "a = Fract an 1" using a unfolding to_fract_def by auto have id: "a * b ^ e = Fract (an * bn^e) (bd^e)" unfolding a b power_fract mult_fract by simp have 0: "bd^e \ 0" for e using bd by auto from no_fract[unfolded id range_to_fract_dvd_iff[OF 0]] have dvd: "bd ^ e dvd an * bn ^ e" . from copb have copb: "coprime (bd ^ e) (bn ^ e)" for e by (simp add: ac_simps) from dvd copb [of e] bd have "bd ^ e dvd an" by (simp add: coprime_dvd_mult_left_iff) hence "bd ^ f dvd an" using le by (rule power_le_dvd) hence dvd: "bd ^ f dvd an * bn ^ f" by simp from le obtain g where e: "e = f + g" using le_Suc_ex by blast have id': "a * b ^ f = Fract (an * bn^f) (bd^f)" unfolding a b power_fract mult_fract by simp show ?thesis unfolding id' range_to_fract_dvd_iff[OF 0] by (rule dvd) qed lemma div_divide_to_fract: assumes "x \ range to_fract" and "x = (y :: 'a :: idom_divide fract) / z" and "x' = y' div z'" and "y = to_fract y'" "z = to_fract z'" shows "x = to_fract x'" proof (cases "z' = 0") case True thus ?thesis using assms by auto next case False from assms obtain r where "to_fract y' / to_fract z' = to_fract r" by auto thus ?thesis using False assms by (simp add: eq_fract(1) to_fract_def) qed -declare Euclidean_Division.divmod_nat_def [termination_simp] +declare Euclidean_Rings.divmod_nat_def [termination_simp] fun dichotomous_Lazard :: "'a :: idom_divide \ 'a \ nat \ 'a" where "dichotomous_Lazard x y n = (if n \ 1 then if n = 1 then x else 1 else - let (d,r) = Euclidean_Division.divmod_nat n 2; + let (d,r) = Euclidean_Rings.divmod_nat n 2; rec = dichotomous_Lazard x y d; recsq = rec * rec div y in if r = 0 then recsq else recsq * x div y)" lemma dichotomous_Lazard_main: fixes x :: "'a :: idom_divide" assumes "\ i. i \ n \ (to_fract x)^i / (to_fract y)^(i - 1) \ range to_fract" shows "to_fract (dichotomous_Lazard x y n) = (to_fract x)^n / (to_fract y)^(n-1)" using assms proof (induct x y n rule: dichotomous_Lazard.induct) case (1 x y n) let ?f = to_fract consider (0) "n = 0" | (1) "n = 1" | (n) "\ n \ 1" by linarith thus ?case proof cases case n - obtain d r where n2: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force - from Euclidean_Division.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto + obtain d r where n2: "Euclidean_Rings.divmod_nat n 2 = (d,r)" by force + from Euclidean_Rings.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto hence r: "r = 0 \ r = 1" by auto define rec where "rec = dichotomous_Lazard x y d" let ?sq = "rec * rec div y" have res: "dichotomous_Lazard x y n = (if r = 0 then ?sq else ?sq * x div y)" unfolding dichotomous_Lazard.simps[of x y n] n2 Let_def rec_def using n by auto have ndr: "n = d + d + r" unfolding dr by presburger from ndr r n have d0: "d \ 0" by auto have IH: "?f rec = ?f x ^ d / ?f y ^ (d - 1)" using 1(1)[OF n refl n2[symmetric] 1(2), folded rec_def] ndr by auto have "?f (rec * rec) = ?f x ^ d / ?f y ^ (d - 1) * ?f x ^ d / ?f y ^ (d - 1)" using IH by simp also have "\ = ?f x ^ (d + d) / ?f y ^ (d - 1 + (d - 1))" unfolding power_add by simp also have "d - 1 + (d - 1) = d + d - 2" using d0 by simp finally have id: "?f (rec * rec) = ?f x ^ (d + d) / ?f y ^ (d + d - 2)" . let ?dd = "(?f x ^ (d + d) / ?f y ^ (d + d - 2)) / ?f y" let ?d = "?f x ^ (d + d) / ?f y ^ (d + d - 1)" have dd: "?dd = ?d" using d0 by (cases d, auto) have sq: "?f ?sq = ?d" unfolding dd[symmetric] proof (rule sym, rule div_divide_to_fract[OF _ refl refl id[symmetric] refl], unfold dd) show "?d \ range ?f" by (rule 1(2), insert ndr, auto) qed show ?thesis proof (cases "r = 0") case True with res sq show ?thesis unfolding ndr by auto next case False with r have r: "r = 1" by auto let ?sq' = "?sq * x div y" from False res have res: "dichotomous_Lazard x y n = ?sq'" by simp from sq have id: "?f (?sq * x) = ?f x ^ (d + d + r) / ?f y ^ (d + d - 1)" unfolding r by simp let ?dd = "(?f x ^ (d + d + r) / ?f y ^ (d + d - 1)) / ?f y" let ?d = "?f x ^ (d + d + r) / ?f y ^ (d + d + r - 1)" have dd: "?dd = ?d" using d0 unfolding r by (cases d, auto) have sq': "?f ?sq' = ?d" unfolding dd[symmetric] proof (rule sym, rule div_divide_to_fract[OF _ refl refl id[symmetric] refl], unfold dd) show "?d \ range ?f" by (rule 1(2), unfold ndr, auto) qed show ?thesis unfolding res sq' unfolding ndr by simp qed qed auto qed lemma dichotomous_Lazard: fixes x :: "'a :: factorial_ring_gcd" assumes "(to_fract x)^n / (to_fract y)^(n-1) \ range to_fract" shows "to_fract (dichotomous_Lazard x y n) = (to_fract x)^n / (to_fract y)^(n-1)" proof (rule dichotomous_Lazard_main) fix i assume i: "i \ n" show "to_fract x ^ i / to_fract y ^ (i - 1) \ range to_fract" proof (cases i) case (Suc j) have id: "to_fract x ^ i / to_fract y ^ (i - 1) = to_fract x * (to_fract x / to_fract y) ^ j" unfolding Suc by (simp add: power_divide) from Suc i have "n \ 0" and j: "j \ n - 1" by auto hence idd: "to_fract x * (to_fract x / to_fract y) ^ (n - 1) = (to_fract x)^n / (to_fract y)^(n-1)" by (cases n, auto simp: power_divide) show ?thesis unfolding id by (rule to_fract_power_le[OF _ _ j], unfold idd, insert assms, auto) next case 0 have "1 = to_fract 1" by simp hence "1 \ range to_fract" by blast thus ?thesis using 0 by auto qed qed declare dichotomous_Lazard.simps[simp del] end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_Definitions.thy b/thys/Van_Emde_Boas_Trees/VEBT_Definitions.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_Definitions.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_Definitions.thy @@ -1,759 +1,759 @@ (*by Ammer*) theory VEBT_Definitions imports Main "HOL-Library.Extended_Nat" "HOL-Library.Code_Target_Numeral" "HOL-Library.Code_Target_Nat" begin section \Preliminaries and Preparations\ subsection \Data Type Definition\ datatype VEBT = is_Node: Node (info:"(nat*nat) option")(deg: nat)(treeList: "VEBT list") (summary:VEBT) | is_Leaf: Leaf bool bool hide_const (open) info deg treeList summary locale VEBT_internal begin subsection \Functions for obtaining high and low bits of an input number.\ definition high :: "nat \ nat \ nat" where "high x n = (x div (2^n))" definition low :: "nat \ nat \ nat" where "low x n = (x mod (2^n))" subsection \Some auxiliary lemmata\ lemma inthall[termination_simp]: "(\ x. x \ set xs \ P x) \ n < length xs \ P (xs ! n)" apply(induction xs arbitrary: n) apply auto using less_Suc_eq_0_disj apply auto done lemma intind: "i < n \ P x \ P (replicate n x ! i)" by (metis in_set_replicate inthall length_replicate) lemma concat_inth:"(xs @[x]@ys)! (length xs) = x" by simp lemma pos_n_replace: "n length xs = length (take n xs @ [y] @drop (Suc n) xs)" by simp lemma inthrepl: "i < n \ (replicate n x) ! i = x" by simp lemma nth_repl: "m n m\ n \(take n xs @ [x] @ drop (n+1) xs) ! m = xs ! m" by (metis Suc_eq_plus1 append_Cons append_Nil nth_list_update_neq upd_conv_take_nth_drop) lemma [termination_simp]:assumes "high x deg < length treeList" shows"size (treeList ! high x deg) < Suc (size_list size treeList + size s)" proof- have "treeList ! high x deg \ set treeList" using assms by auto then show ?thesis using not_less_eq size_list_estimation by fastforce qed subsection \ Auxiliary functions for defining valid Van Emde Boas Trees\ text \This function checks whether an element occurs in a Leaf\ fun naive_member :: "VEBT \ nat \ bool" where "naive_member (Leaf a b) x = (if x = 0 then a else if x = 1 then b else False)"| "naive_member (Node _ 0 _ _) _ = False"| "naive_member (Node _ deg treeList s) x = (let pos = high x (deg div 2) in (if pos < length treeList then naive_member (treeList ! pos) (low x (deg div 2)) else False))" text \Test for elements stored by using the provide min-max-fields\ fun membermima :: "VEBT \ nat \ bool" where "membermima (Leaf _ _) _ = False"| "membermima (Node None 0 _ _ )_ =False"| "membermima (Node (Some (mi,ma)) 0 _ _) x = (x = mi \ x = ma)"| "membermima (Node (Some (mi, ma)) deg treeList _) x = (x = mi \ x = ma \ ( let pos = high x ( deg div 2) in (if pos < length treeList then membermima (treeList ! pos) (low x (deg div 2)) else False)))"| "membermima (Node None (deg) treeList _) x = (let pos = high x (deg div 2) in (if pos < length treeList then membermima (treeList ! pos) (low x (deg div 2)) else False))" lemma length_mul_elem:"(\ x \ set xs. length x = n) \ length (concat xs) = (length xs) * n" apply(induction xs) apply auto done text \We combine both auxiliary functions: The following test returns true if and only if an element occurs in the tree with respect to our interpretation no matter where it is stored.\ definition both_member_options :: "VEBT \ nat \ bool" where "both_member_options t x = (naive_member t x \ membermima t x)" end context begin interpretation VEBT_internal . definition set_vebt :: "VEBT \ nat set" where "set_vebt t = {x. both_member_options t x}" end subsection \Inductive Definition of semantically valid Vam Emde Boas Trees\ text \Invariant for verification proofs\ context begin interpretation VEBT_internal . inductive invar_vebt::"VEBT \ nat \ bool" where "invar_vebt (Leaf a b) (Suc 0) "| "( \ t \ set treeList. invar_vebt t n) \ invar_vebt summary m \ length treeList = 2^m \ m = n \ deg = n + m \ (\ i. both_member_options summary i) \(\ t \ set treeList. \ x. both_member_options t x) \ invar_vebt (Node None deg treeList summary) deg"| "( \ t \ set treeList. invar_vebt t n) \ invar_vebt summary m \ length treeList = 2^m \ m = Suc n \ deg = n + m \ (\ i. both_member_options summary i) \ (\ t \ set treeList. \ x. both_member_options t x) \ invar_vebt (Node None deg treeList summary) deg"| "( \ t \ set treeList. invar_vebt t n) \ invar_vebt summary m \ length treeList = 2^m \ m = n \deg = n + m\ (\ i < 2^m. (\ x. both_member_options (treeList ! i) x) \ ( both_member_options summary i)) \ (mi = ma \ (\ t \ set treeList. \ x. both_member_options t x)) \ mi \ ma \ ma < 2^deg \ (mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ x. (high x n = i \ both_member_options (treeList ! i) (low x n) ) \ mi < x \ x \ ma) ) ) \ invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg"| "( \ t \ set treeList. invar_vebt t n) \invar_vebt summary m \ length treeList = 2^m \ m = Suc n \deg = n + m\(\ i < 2^m. (\ x. both_member_options (treeList ! i) x) \ ( both_member_options summary i)) \ (mi = ma \ (\ t \ set treeList. \ x. both_member_options t x)) \ mi \ ma \ ma < 2^deg \ (mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ x. (high x n = i \ both_member_options (treeList ! i) (low x n) ) \ mi < x \ x \ ma))) \ invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" end context VEBT_internal begin definition "in_children n treeList x \ both_member_options (treeList ! high x n) (low x n)" text \functional validness definition\ fun valid' :: "VEBT \ nat \ bool" where "valid' (Leaf _ _) d \ d=1" | "valid' (Node mima deg treeList summary) deg' \ ( deg=deg' \ ( let n = deg div 2; m = deg - n in ( \ t \ set treeList. valid' t n ) \ valid' summary m \ length treeList = 2^m \ ( case mima of None \ (\ i. both_member_options summary i) \ (\ t \ set treeList. \ x. both_member_options t x) | Some (mi,ma) \ mi \ ma \ ma<2^deg \ (\ i < 2^m. (\ x. both_member_options (treeList ! i) x) \ ( both_member_options summary i)) \ (if mi=ma then (\ t \ set treeList. \ x. both_member_options t x) else in_children n treeList ma \ (\x < 2^deg. in_children n treeList x \ mi x\ma) ) ) ) ) " text \equivalence proofs\ lemma high_bound_aux: "ma < 2^(n+m) \ high ma n < 2^m" unfolding high_def by (simp add: add.commute less_mult_imp_div_less power_add) lemma valid_eq1: assumes "invar_vebt t d" shows "valid' t d" using assms apply induction apply simp_all apply (auto simp: in_children_def dest: high_bound_aux) [] subgoal for treeList n summary m deg mi ma apply (intro allI impI conjI) apply (auto simp: in_children_def dest: high_bound_aux) [] apply (metis add_Suc_right high_bound_aux power_Suc) apply (auto simp: in_children_def dest: high_bound_aux) [] apply (metis add_Suc_right high_bound_aux power_Suc) apply (auto simp: in_children_def dest: high_bound_aux) [] apply (metis add_Suc_right high_bound_aux power_Suc) done done lemma even_odd_cases: fixes x :: nat obtains n where "x=n+n" | n where "x = n + Suc n" apply (cases "even x"; simp) apply (metis add_self_div_2 div_add) by (metis add.commute mult_2 oddE plus_1_eq_Suc) lemma valid_eq2: "valid' t d \ invar_vebt t d" apply (induction t d rule: valid'.induct) apply (auto intro: invar_vebt.intros simp: Let_def split: option.splits) subgoal for deg treeList summary apply (rule even_odd_cases[of deg]; simp) apply (rule invar_vebt.intros(2); simp) apply (rule invar_vebt.intros(3); simp add: algebra_simps) by presburger subgoal for deg treeList summary mi ma apply (rule even_odd_cases[of deg]; simp) subgoal apply (rule invar_vebt.intros(4); simp?) apply (auto simp: in_children_def) [] - apply (meson le_less_linear le_less_trans) - apply (metis Euclidean_Division.div_eq_0_iff div_exp_eq gr_implies_not0 high_def) - done + apply (meson le_less_linear le_less_trans) + apply (metis div_eq_0_iff div_exp_eq gr_implies_not0 high_def) + done subgoal apply (rule invar_vebt.intros(5); simp?) apply (auto) [] apply (auto) [] apply (auto simp: in_children_def) [] apply (meson le_less_linear le_less_trans) - by (metis Euclidean_Division.div_eq_0_iff add_Suc_right div_exp_eq high_def power_Suc power_eq_0_iff zero_neq_numeral) + apply (metis div_eq_0_iff add_Suc_right div_exp_eq high_def power_Suc power_eq_0_iff zero_neq_numeral) + done done done lemma valid_eq: "valid' t d \ invar_vebt t d" using valid_eq1 valid_eq2 by auto lemma [termination_simp]: assumes "odd (v::nat)" shows "v div 2 < v" by (simp add: assms odd_pos) lemma [termination_simp]:assumes "n > 1" and " odd n" shows" Suc (n div 2) < n" by (metis Suc_lessI add_diff_cancel_left' assms(1) assms(2) div_eq_dividend_iff div_less_dividend even_Suc even_Suc_div_two odd_pos one_less_numeral_iff plus_1_eq_Suc semiring_norm(76) zero_less_diff) end subsection \Function for generating an empty tree of arbitrary degree respectively order\ context begin interpretation VEBT_internal . fun vebt_buildup :: "nat \ VEBT" where "vebt_buildup 0 = Leaf False False"| "vebt_buildup (Suc 0) = Leaf False False"| "vebt_buildup n = (if even n then (let half = n div 2 in Node None n (replicate (2^half) (vebt_buildup half)) (vebt_buildup half)) else (let half = n div 2 in Node None n ( replicate (2^(Suc half)) (vebt_buildup half)) (vebt_buildup (Suc half))))" end context VEBT_internal begin lemma buildup_nothing_in_leaf: "\ naive_member (vebt_buildup n) x" proof(induction arbitrary: x rule: vebt_buildup.induct) case 1 then show ?case by simp next case (2 v) then show ?case by simp next case (3 n) let ?n = "Suc(Suc n)" show ?case proof(cases "even ?n") case True let ?half = "?n div 2" have "\ naive_member (vebt_buildup ?half) y" for y using "3.IH"(1) True by blast hence 0:"\ t \ set (replicate (2^?half) (vebt_buildup ?half)) . \ naive_member t x" by simp have "naive_member (vebt_buildup ?n) x \ False" proof- assume "naive_member (vebt_buildup ?n) x" hence "high x ?half < 2^?half \ naive_member ((replicate (2^?half) (vebt_buildup ?half)) ! (high x ?half)) (low x ?half)" by (metis True vebt_buildup.simps(3) length_replicate naive_member.simps(3)) hence "\ t \ set (replicate (2^?half) (vebt_buildup ?half)) . naive_member t x " by (metis \\y. \ naive_member (vebt_buildup (Suc (Suc n) div 2)) y\ nth_replicate) then show False using 0 by simp qed then show ?thesis by blast next case False let ?half = "?n div 2" have "\ naive_member (vebt_buildup ?half) y" for y using "3.IH" False by blast hence 0:"\ t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half)) . \ naive_member t x" by simp have "naive_member (vebt_buildup ?n) x \ False" proof- assume "naive_member (vebt_buildup ?n) x" hence "high x ?half < 2^(Suc ?half) \ naive_member ((replicate (2^(Suc ?half)) (vebt_buildup ?half)) ! (high x ?half)) (low x ?half)" by (metis False vebt_buildup.simps(3) length_replicate naive_member.simps(3)) hence "\ t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half)) . naive_member t x " by (metis \\y. \ naive_member (vebt_buildup (Suc (Suc n) div 2)) y\ nth_replicate) then show False using 0 by simp qed then show ?thesis by force qed qed lemma buildup_nothing_in_min_max:"\ membermima (vebt_buildup n) x" proof(induction arbitrary: x rule: vebt_buildup.induct) case 1 then show ?case by simp next case 2 then show ?case by simp next case (3 va) let ?n = "Suc (Suc va)" let ?half = "?n div 2" show ?case proof(cases "even ?n") case True have "\ membermima (vebt_buildup ?half) y" for y using "3.IH"(1) True by blast hence 0:"\ t \ set (replicate (2^?half) (vebt_buildup ?half)) . \ membermima t x" by simp then show ?thesis by (metis "3.IH"(1) True vebt_buildup.simps(3) inthrepl length_replicate membermima.simps(5)) next case False have "\ membermima (vebt_buildup ?half) y" for y using "3.IH" False by blast moreover hence 0:"\ t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half)) . \ membermima t x" by simp ultimately show ?thesis by (metis vebt_buildup.simps(3) inthrepl length_replicate membermima.simps(5)) qed qed text \The empty tree generated by $vebt_buildup$ is indeed a valid tree.\ lemma buildup_gives_valid: "n>0 \ invar_vebt (vebt_buildup n) n" proof( induction n rule: vebt_buildup.induct) case 1 then show ?case by simp next case 2 then show ?case by (simp add: invar_vebt.intros(1)) next case (3 va) let ?n = "Suc (Suc va)" let ?half = "?n div 2" show ?case proof(cases "even ?n") case True hence a:"vebt_buildup ?n = Node None ?n (replicate (2^?half) (vebt_buildup ?half)) (vebt_buildup ?half)" by simp moreover hence "invar_vebt (vebt_buildup ?half) ?half" using "3.IH"(1) True by auto moreover hence "( \ t \ set (replicate (2^?half) (vebt_buildup ?half)). invar_vebt t ?half)" by simp moreover have "length (replicate (2^?half) (vebt_buildup ?half)) = 2^?half" by auto moreover have "?half+?half = ?n" using True by auto moreover have " \ t \ set (replicate (2^?half) (vebt_buildup ?half)). (\ x. both_member_options t x)" proof fix t assume "t \ set (replicate (2^?half) (vebt_buildup ?half))" hence "t = (vebt_buildup ?half)" by simp thus "\ x. both_member_options t x" by (simp add: both_member_options_def buildup_nothing_in_leaf buildup_nothing_in_min_max) qed moreover have " (\ i. both_member_options (vebt_buildup ?half) i)" using both_member_options_def buildup_nothing_in_leaf buildup_nothing_in_min_max by blast ultimately have "invar_vebt (Node None ?n (replicate (2^?half) (vebt_buildup ?half)) (vebt_buildup ?half)) ?n" using invar_vebt.intros(2)[of "replicate (2^?half) (vebt_buildup ?half)" ?half "vebt_buildup ?half" ?half ?n] by simp then show ?thesis using a by auto next case False hence a:"vebt_buildup ?n = Node None ?n (replicate (2^(Suc ?half)) (vebt_buildup ?half)) (vebt_buildup (Suc ?half))" by simp moreover hence "invar_vebt (vebt_buildup (Suc ?half)) (Suc ?half)" using "3.IH" False by auto moreover have "invar_vebt (vebt_buildup ?half) ?half" using "3.IH"(3) False by auto moreover hence "( \ t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half)). invar_vebt t ?half)" by simp moreover have "length (replicate (2^(Suc ?half)) (vebt_buildup ?half)) = 2^(Suc ?half)" by auto moreover have "(Suc ?half)+?half = ?n" using False by presburger moreover have " \ t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half)). (\ x. both_member_options t x)" proof fix t assume "t \ set (replicate (2^(Suc ?half)) (vebt_buildup ?half))" hence "t = (vebt_buildup ?half)" by simp thus "\ x. both_member_options t x" by (simp add: both_member_options_def buildup_nothing_in_leaf buildup_nothing_in_min_max) qed moreover have " (\ i. both_member_options (vebt_buildup (Suc ?half)) i)" using both_member_options_def buildup_nothing_in_leaf buildup_nothing_in_min_max by blast moreover have "?half + Suc ?half = ?n" using calculation(6) by auto ultimately have "invar_vebt (Node None ?n (replicate (2^(Suc ?half)) (vebt_buildup ?half)) (vebt_buildup (Suc ?half))) ?n" using invar_vebt.intros(3)[of "replicate (2^(Suc ?half)) (vebt_buildup ?half)" ?half "vebt_buildup (Suc ?half)" "Suc ?half" ?n ] by simp then show ?thesis using a by auto qed qed lemma mi_ma_2_deg: assumes "invar_vebt (Node (Some (mi, ma)) deg treeList summary) n" shows "mi\ ma \ ma < 2^deg" proof- from assms show ?thesis proof cases qed blast+ qed lemma deg_not_0: "invar_vebt t n \ n > 0" apply(induction t n rule: invar_vebt.induct) apply auto done lemma set_n_deg_not_0:assumes " \t\set treeList. invar_vebt t n"and" length treeList = 2^m "shows " n \ 1" proof- have "length treeList > 0" by (simp add: assms(2)) then obtain t ts where "treeList = t#ts" by (metis list.size(3) neq_Nil_conv not_less0) hence "invar_vebt t n" by (simp add: assms(1)) hence "n \ 1" using deg_not_0 by force thus ?thesis by simp qed lemma both_member_options_ding: assumes"invar_vebt (Node info deg treeList summary) n "and "x<2^deg"and" both_member_options (treeList ! (high x (deg div 2))) (low x (deg div 2))"shows "both_member_options (Node info deg treeList summary) x" proof- from assms(1) show ?thesis proof(induction "(Node info deg treeList summary)" n rule: invar_vebt.induct) case (2 n m) hence "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using assms(3) both_member_options_def by auto moreover hence "deg > 1" using "2.hyps"(2) "2.hyps"(5) "2.hyps"(6) deg_not_0 by force moreover have "high x (deg div 2)<2^m" - by (metis "2.hyps"(5) "2.hyps"(6) Euclidean_Division.div_eq_0_iff add_self_div_2 assms(2) div_exp_eq high_def power_not_zero zero_neq_one) + by (metis "2.hyps"(5) "2.hyps"(6) div_eq_0_iff add_self_div_2 assms(2) div_exp_eq high_def power_not_zero) moreover have "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ membermima (Node info deg treeList summary) x" using membermima.simps(5)[of "deg-1" treeList summary x] using "2.hyps"(4) "2.hyps"(9) \1 < deg\ \high x (deg div 2) < 2 ^ m\ zero_le_one by fastforce moreover have "naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (Node info deg treeList summary) x" by (smt "2.hyps"(4) Suc_diff_Suc \1 < deg\ \high x (deg div 2) < 2 ^ m\ diff_zero le_less_trans naive_member.simps(3) zero_le_one) ultimately show ?case using both_member_options_def by blast next case (3 n m) hence "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using assms(3) both_member_options_def by auto moreover hence "deg > 1" by (metis "3.hyps"(1) "3.hyps"(2) "3.hyps"(4) "3.hyps"(5) "3.hyps"(6) One_nat_def Suc_lessI add_Suc add_gr_0 add_self_div_2 deg_not_0 le_imp_less_Suc plus_1_eq_Suc set_n_deg_not_0) moreover have "high x (deg div 2)<2^m" - by (smt "3.hyps"(5) "3.hyps"(6) Euclidean_Division.div_eq_0_iff add_Suc_right add_self_div_2 assms(2) diff_Suc_1 div_exp_eq div_mult_self1_is_m even_Suc high_def odd_add odd_two_times_div_two_nat one_add_one plus_1_eq_Suc power_not_zero zero_less_Suc) + by (smt "3.hyps"(5) "3.hyps"(6) div_eq_0_iff add_Suc_right add_self_div_2 assms(2) diff_Suc_1 div_exp_eq div_mult_self1_is_m even_Suc high_def odd_add odd_two_times_div_two_nat one_add_one plus_1_eq_Suc power_not_zero zero_less_Suc) moreover have "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ membermima (Node info deg treeList summary) x" using membermima.simps(5)[of "deg-1" treeList summary x] using "3.hyps"(4) "3.hyps"(9) \1 < deg\ \high x (deg div 2) < 2 ^ m\ zero_le_one by fastforce moreover have "naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (Node info deg treeList summary) x" by (smt "3.hyps"(4) Suc_diff_Suc \1 < deg\ \high x (deg div 2) < 2 ^ m\ diff_zero le_less_trans naive_member.simps(3) zero_le_one) ultimately show ?case using both_member_options_def by blast next case (4 n m mi ma) hence "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using assms(3) both_member_options_def by auto moreover hence "deg > 1" using "4.hyps"(2) "4.hyps"(5) "4.hyps"(6) deg_not_0 by force moreover have "high x (deg div 2)<2^m" - by (metis "4.hyps"(5) "4.hyps"(6) Euclidean_Division.div_eq_0_iff add_self_div_2 assms(2) div_exp_eq high_def power_not_zero zero_neq_one) + by (metis "4.hyps"(5) "4.hyps"(6) div_eq_0_iff add_self_div_2 assms(2) div_exp_eq high_def power_not_zero) moreover have "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ membermima (Node info deg treeList summary) x" using membermima.simps(5)[of "deg-1" treeList summary x] by (smt "4.hyps"(12) "4.hyps"(4) Suc_diff_Suc calculation(2) calculation(3) diff_zero le_less_trans membermima.simps(4) zero_le_one) moreover have "naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (Node info deg treeList summary) x" by (metis "4.hyps"(4) calculation(2) calculation(3) gr_implies_not0 naive_member.simps(3) old.nat.exhaust) ultimately show ?case using both_member_options_def by blast next case (5 n m mi ma) hence "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using assms(3) both_member_options_def by auto moreover hence "deg > 1" by (metis "5.hyps"(1) "5.hyps"(2) "5.hyps"(4) "5.hyps"(5) "5.hyps"(6) One_nat_def Suc_lessI add_Suc add_gr_0 add_self_div_2 deg_not_0 le_imp_less_Suc plus_1_eq_Suc set_n_deg_not_0) moreover have "high x (deg div 2)<2^m" - by (metis "5.hyps"(5) "5.hyps"(6) Euclidean_Division.div_eq_0_iff add_Suc_right add_self_div_2 assms(2) div_exp_eq even_Suc_div_two even_add high_def nat.simps(3) power_not_zero) + by (metis "5.hyps"(5) "5.hyps"(6) div_eq_0_iff add_Suc_right add_self_div_2 assms(2) div_exp_eq even_Suc_div_two even_add high_def nat.simps(3) power_not_zero) moreover have "membermima (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ membermima (Node info deg treeList summary) x" using membermima.simps(5)[of "deg-1" treeList summary x] by (smt "5.hyps"(12) "5.hyps"(4) Suc_diff_Suc calculation(2) calculation(3) diff_zero le_less_trans membermima.simps(4) zero_le_one) moreover have "naive_member (treeList ! (high x (deg div 2))) (low x (deg div 2)) \ naive_member (Node info deg treeList summary) x" using "5.hyps"(4) "5.hyps"(5) "5.hyps"(6) calculation(3) by auto ultimately show ?case using both_member_options_def by blast qed qed lemma exp_split_high_low: assumes "x < 2^(n+m)" and "n > 0" and "m> 0" shows "high x n < 2^m" and "low x n < 2^n" - apply (metis Euclidean_Division.div_eq_0_iff assms(1) div_exp_eq high_def nat.simps(3) numeral_2_eq_2 power_not_zero) - by (simp add: low_def) + using assms by (simp_all add: high_bound_aux low_def) lemma low_inv: assumes "x< 2^n " shows "low (y*2^n + x) n = x" unfolding low_def by (simp add: assms) lemma high_inv: assumes "x< 2^n " shows "high (y*2^n + x) n = y" unfolding high_def by (simp add: assms) lemma both_member_options_from_chilf_to_complete_tree: assumes "high x (deg div 2) < length treeList" and "deg \1" and "both_member_options (treeList ! ( high x (deg div 2))) (low x (deg div 2))" shows "both_member_options (Node (Some (mi, ma)) deg treeList summary) x" proof- have "membermima (treeList ! ( high x (deg div 2))) (low x (deg div 2)) \ naive_member (treeList ! ( high x (deg div 2))) (low x (deg div 2))" using assms using both_member_options_def by blast moreover have "membermima (treeList ! ( high x (deg div 2))) (low x (deg div 2)) \ membermima (Node (Some (mi, ma)) deg treeList summary) x" using membermima.simps(4)[of mi ma "deg-1" treeList summary x] by (metis Suc_1 Suc_leD assms(1) assms(2) le_add_diff_inverse plus_1_eq_Suc) moreover have "naive_member (treeList ! ( high x (deg div 2))) (low x (deg div 2)) \ naive_member (Node (Some (mi, ma)) deg treeList summary) x" using naive_member.simps(3)[of "Some (mi, ma)" "deg-1" treeList summary x] by (metis Suc_1 Suc_leD assms(1) assms(2) le_add_diff_inverse plus_1_eq_Suc) ultimately show ?thesis using both_member_options_def by blast qed lemma both_member_options_from_complete_tree_to_child: assumes "deg \1" and "both_member_options (Node (Some (mi, ma)) deg treeList summary) x" shows "both_member_options (treeList ! ( high x (deg div 2))) (low x (deg div 2)) \ x = mi \ x = ma" proof- have "naive_member (Node (Some (mi, ma)) deg treeList summary) x \ membermima (Node (Some (mi, ma)) deg treeList summary) x " using assms(2) both_member_options_def by auto moreover have " naive_member (Node (Some (mi, ma)) deg treeList summary) x \ naive_member (treeList ! ( high x (deg div 2))) (low x (deg div 2))" using naive_member.simps(3)[of "Some (mi, ma)" "deg-1" treeList summary x] by (metis assms(1) le_add_diff_inverse plus_1_eq_Suc) moreover have " membermima (Node (Some (mi, ma)) deg treeList summary) x \ membermima (treeList ! ( high x (deg div 2))) (low x (deg div 2))\ x = mi \ x = ma" by (smt (z3) assms(1) le_add_diff_inverse membermima.simps(4) plus_1_eq_Suc) ultimately show ?thesis using both_member_options_def by presburger qed lemma pow_sum: "(divide::nat \ nat \ nat) ((2::nat) ^((a::nat)+(b::nat))) (2^a) = 2^b" by (induction a) simp+ fun elim_dead::"VEBT \ enat \ VEBT" where "elim_dead (Leaf a b) _ = Leaf a b "| "elim_dead (Node info deg treeList summary) \ = (Node info deg (map (\ t. elim_dead t (enat (2^(deg div 2)))) treeList) (elim_dead summary \))"| "elim_dead (Node info deg treeList summary) (enat l) = (Node info deg (take (l div (2^(deg div 2))) (map (\ t. elim_dead t (enat (2^(deg div 2))))treeList)) (elim_dead summary ((enat (l div (2^(deg div 2))))))) " lemma elimnum: "invar_vebt (Node info deg treeList summary) n \ elim_dead (Node info deg treeList summary) (enat ((2::nat)^n)) = (Node info deg treeList summary)" proof(induction rule: invar_vebt.induct) case (1 a b) then show ?case by simp next case (2 treeList n summary m deg) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "2.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" using "2.IH"(1) by blast qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (simp add: "2.IH"(1) map_idI) have "deg div 2 = n" by (simp add: "2.hyps"(3) "2.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "2.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "2"(4) by simp moreover hence " ( elim_dead summary ((enat ((2^deg) div (2^(deg div 2)))))) = summary" using "2.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(3)[of None deg treeList summary "2^deg"] using \deg div 2 = n\ by metis next case (3 treeList n summary m deg) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "3.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" using "3.IH"(1) by blast qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (simp add: "3.IH"(1) map_idI) have "deg div 2 = n" by (simp add: "3.hyps"(3) "3.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "3.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "3"(4) by simp moreover hence " ( elim_dead summary ((enat ((2^deg) div (2^(deg div 2)))))) = summary" using "3.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(3)[of None deg treeList summary "2^deg"] using \deg div 2 = n\ by metis next case (4 treeList n summary m deg mi ma) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "4.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" using "4.IH"(1) by blast qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (simp add: "4.IH"(1) map_idI) have "deg div 2 = n" by (simp add: "4.hyps"(3) "4.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "4.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "4"(4) by simp moreover hence " ( elim_dead summary ((enat ((2^deg) div (2^(deg div 2)))))) = summary" using "4.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(3)[of "Some (mi, ma)" deg treeList summary "2^deg"] using \deg div 2 = n\ by metis next case (5 treeList n summary m deg mi ma) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "5.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" using "5.IH"(1) by blast qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (simp add: "5.IH"(1) map_idI) have "deg div 2 = n" by (simp add: "5.hyps"(3) "5.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "5.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "5"(4) by simp moreover hence " ( elim_dead summary ((enat ((2^deg) div (2^(deg div 2)))))) = summary" using "5.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(3)[of "Some (mi, ma)" deg treeList summary "2^deg"] using \deg div 2 = n\ by metis qed lemma elimcomplete: "invar_vebt (Node info deg treeList summary) n \ elim_dead (Node info deg treeList summary) \ = (Node info deg treeList summary)" proof(induction rule: invar_vebt.induct) case (1 a b) then show ?case by simp next case (2 treeList n summary m deg) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "2.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" apply(cases "(treeList ! i)") apply (smt (z3) "2.IH"(1) \treeList ! i \ set treeList\ elim_dead.simps(1) elimnum invar_vebt.cases)+ done qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (metis "2.hyps"(2) in_set_conv_nth map_idI) have "deg div 2 = n" by (simp add: "2.hyps"(3) "2.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "2.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "2"(4) by simp moreover hence " ( elim_dead summary \) = summary" using "2.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(2)[of None deg treeList summary] using \deg div 2 = n\ b by presburger next case (3 treeList n summary m deg) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "3.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" apply(cases "(treeList ! i)") apply (smt (z3) "3.IH"(1) \treeList ! i \ set treeList\ elim_dead.simps(1) elimnum invar_vebt.cases)+ done qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (metis "3.hyps"(2) in_set_conv_nth map_idI) have "deg div 2 = n" by (simp add: "3.hyps"(3) "3.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "3.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "3"(4) by simp moreover hence " ( elim_dead summary \) = summary" using "3.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(2)[of None deg treeList summary] using \deg div 2 = n\ b by presburger next case (4 treeList n summary m deg mi ma) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "4.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" apply(cases "(treeList ! i)") apply (smt (z3) "4.IH"(1) \treeList ! i \ set treeList\ elim_dead.simps(1) elimnum invar_vebt.cases)+ done qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (metis "4.hyps"(2) in_set_conv_nth map_idI) have "deg div 2 = n" by (simp add: "4.hyps"(3) "4.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "4.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "4"(4) by simp moreover hence " ( elim_dead summary \) = summary" using "4.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(2)[of "Some (mi, ma)" deg treeList summary] using \deg div 2 = n\ b by presburger next case (5 treeList n summary m deg mi ma) have a:"i < 2^m \ (elim_dead (treeList ! i) (enat( 2^n)) = treeList ! i)" for i proof assume "i < 2^m" hence "treeList ! i \ set treeList" by (simp add: "5.hyps"(2)) thus "elim_dead (treeList ! i) (enat (2 ^ n)) = treeList ! i" apply(cases "(treeList ! i)") apply (smt (z3) "5.IH"(1) \treeList ! i \ set treeList\ elim_dead.simps(1) elimnum invar_vebt.cases)+ done qed hence b:"map (\ t. elim_dead t (enat (2 ^ n))) treeList = treeList" by (metis "5.hyps"(2) in_set_conv_nth map_idI) have "deg div 2 = n" by (simp add: "5.hyps"(3) "5.hyps"(4)) hence "(2^m ::nat) = ( (2^deg) div (2^(deg div 2))::nat) " using "5.hyps"(4) pow_sum by metis hence "take (2^deg div (2^(deg div 2)))(map (\ t. elim_dead t (enat (2 ^ n))) treeList) = treeList" using b "5"(4) by simp moreover hence " ( elim_dead summary \) = summary" using "5.IH"(2) by (metis \2 ^ m = 2 ^ deg div 2 ^ (deg div 2)\) ultimately show ?case using elim_dead.simps(2)[of "Some (mi, ma)" deg treeList summary] using \deg div 2 = n\ b by presburger qed end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_Delete.thy b/thys/Van_Emde_Boas_Trees/VEBT_Delete.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_Delete.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_Delete.thy @@ -1,1791 +1,1791 @@ (*by Ammer*) theory VEBT_Delete imports VEBT_Pred VEBT_Succ begin section \Deletion\ subsection \Function Definition\ context begin interpretation VEBT_internal . fun vebt_delete :: "VEBT \ nat \ VEBT" where "vebt_delete (Leaf a b) 0 = Leaf False b"| "vebt_delete (Leaf a b) (Suc 0) = Leaf a False"| "vebt_delete (Leaf a b) (Suc (Suc n)) = Leaf a b"| "vebt_delete (Node None deg treeList summary) _ = (Node None deg treeList summary)"| "vebt_delete (Node (Some (mi, ma)) 0 trLst smry) x = (Node (Some (mi, ma)) 0 trLst smry) "| "vebt_delete (Node (Some (mi, ma)) (Suc 0) tr sm) x = (Node (Some (mi, ma)) (Suc 0) tr sm) "| "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( if (x < mi \ x > ma) then (Node (Some (mi, ma)) deg treeList summary) else if (x = mi \ x = ma) then (Node None deg treeList summary) else let xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x); minn = (if x = mi then xn else mi); l = low xn (deg div 2); h = high xn (deg div 2) in if h < length treeList then( let newnode = vebt_delete (treeList ! h) l; newlist = treeList[h:= newnode]in if minNull newnode then( let sn = vebt_delete summary h in( Node (Some (minn, if xn = ma then (let maxs = vebt_maxt sn in ( if maxs = None then minn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)))) else ma)) deg newlist sn)) else (Node (Some (minn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary )) else (Node (Some (mi, ma)) deg treeList summary))" end subsection \Auxiliary Lemmas\ context VEBT_internal begin context begin lemma delt_out_of_range: assumes "x < mi \ x > ma" and "deg \ 2" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =(Node (Some (mi, ma)) deg treeList summary)" using vebt_delete.simps(7)[of mi ma "deg-2" treeList summary x] by (metis add_2_eq_Suc assms(1) assms(2) le_add_diff_inverse) lemma del_single_cont: assumes "x = mi \ x = ma" and "deg \ 2" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node None deg treeList summary)" using vebt_delete.simps(7)[of mi ma "deg-2" treeList summary x] by (metis add_2_eq_Suc assms(1) assms(2) le_add_diff_inverse nat_less_le) lemma del_in_range: assumes "x \ mi \ x \ ma" and "mi \ ma" and "deg \ 2" shows " vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ( let xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x); minn = (if x = mi then xn else mi); l = low xn (deg div 2); h = high xn (deg div 2) in if h < length treeList then( let newnode = vebt_delete (treeList ! h) l; newlist = treeList[h:= newnode] in if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (minn, if xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then minn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (minn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary ) )else (Node (Some (mi, ma)) deg treeList summary))" using vebt_delete.simps(7)[of mi ma "deg-2" treeList summary x] by (smt (z3) add_2_eq_Suc assms(1) assms(2) assms(3) leD le_add_diff_inverse) lemma del_x_not_mia: assumes "x > mi \ x \ ma" and "mi \ ma" and "deg \ 2" and "high x (deg div 2) = h" and "low x (deg div 2) = l"and "high x (deg div 2) < length treeList" shows " vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ( let newnode = vebt_delete (treeList ! h) l; newlist = treeList[h:= newnode] in if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (mi, if x = ma then (let maxs = vebt_maxt sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (mi, (if x = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary ) )" using del_in_range[of mi x ma deg treeList summary] unfolding Let_def using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) nat_less_le by fastforce lemma del_x_not_mi: assumes "x > mi \ x \ ma" and "mi \ ma" and "deg \ 2" and "high x (deg div 2) = h" and "low x (deg div 2) = l"and " newnode = vebt_delete (treeList ! h) l" and "newlist = treeList[h:= newnode]" and "high x (deg div 2) < length treeList" shows " vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ( if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (mi, if x = ma then (let maxs = vebt_maxt sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (mi, (if x = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary ) )" using del_x_not_mia[of mi x ma deg h l treeList summary] by (smt (z3) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8)) lemma del_x_not_mi_new_node_nil: assumes "x > mi \ x \ ma" and "mi \ ma" and "deg \ 2" and "high x (deg div 2) = h" and "low x (deg div 2) = l"and " newnode = vebt_delete (treeList ! h) l" and "minNull newnode " and "sn = vebt_delete summary h" and "newlist =treeList[h:= newnode]" and "high x (deg div 2) < length treeList" shows " vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, if x = ma then (let maxs = vebt_maxt sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn)" using del_x_not_mi[of mi x ma deg h l newnode treeList] by (metis assms(1) assms(10) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9)) lemma del_x_not_mi_newnode_not_nil: assumes "x > mi \ x \ ma" and "mi \ ma" and "deg \ 2" and "high x (deg div 2) = h" and "low x (deg div 2) = l"and " newnode = vebt_delete (treeList ! h) l" and "\ minNull newnode " and "newlist = treeList[h:= newnode]" and"high x (deg div 2) < length treeList" shows " vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, (if x = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary )" using del_x_not_mi[of mi x ma deg h l newnode treeList newlist summary] using assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9) by auto lemma del_x_mia: assumes "x = mi \ x < ma" and "mi \ ma" and "deg \ 2" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( let xn = the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))); minn = xn; l = low xn (deg div 2); h = high xn (deg div 2) in if h < length treeList then( let newnode = vebt_delete (treeList ! h) l; newlist = treeList[h:= newnode]in if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (minn, if xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then minn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (minn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary ) )else (Node (Some (mi, ma)) deg treeList summary) )" using del_in_range[of mi x ma deg treeList summary] using assms(1) assms(3) nat_less_le order_refl by fastforce lemma del_x_mi: assumes "x = mi \ x < ma" and "mi \ ma" and "deg \ 2" and "high xn (deg div 2) = h" and "xn = the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) " "low xn (deg div 2) = l"and "high xn (deg div 2) < length treeList" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( let newnode = vebt_delete (treeList ! h) l; newlist = treeList[h:= newnode]in if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (xn, if xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then xn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (xn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary )) " using del_x_mia[of x mi ma deg treeList summary] by (smt (z3) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7)) lemma del_x_mi_lets_in: assumes "x = mi \ x < ma" and "mi \ ma" and "deg \ 2" and "high xn (deg div 2) = h" and "xn = the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) " "low xn (deg div 2) = l"and "high xn (deg div 2) < length treeList" and " newnode = vebt_delete (treeList ! h) l" and " newlist = treeList[h:= newnode]" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( if minNull newnode then( let sn = vebt_delete summary h in (Node (Some (xn, if xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then xn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (xn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary ))" using del_x_mi[of x mi ma deg xn h summary treeList l] by (smt (z3) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9)) lemma del_x_mi_lets_in_minNull: assumes "x = mi \ x < ma" and "mi \ ma" and "deg \ 2" and "high xn (deg div 2) = h" and "xn = the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) " "low xn (deg div 2) = l"and "high xn (deg div 2) < length treeList" and "newnode = vebt_delete (treeList ! h) l" and " newlist =treeList[h:= newnode]" and "minNull newnode " and " sn = vebt_delete summary h" shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (xn, if xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then xn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn)" using del_x_mi_lets_in[of x mi ma deg xn h summary treeList l newnode newlist] by (metis assms(1) assms(10) assms(11) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9)) lemma del_x_mi_lets_in_not_minNull: assumes "x = mi \ x < ma" and "mi \ ma" and "deg \ 2" and "high xn (deg div 2) = h" and "xn = the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) " "low xn (deg div 2) = l"and "high xn (deg div 2) < length treeList" and " newnode = vebt_delete (treeList ! h) l" and " newlist = treeList[h:= newnode]" and "\minNull newnode " shows "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (xn, (if xn = ma then h * 2^(deg div 2) + the( vebt_maxt (newlist ! h)) else ma))) deg newlist summary )" using del_x_mi_lets_in[of x mi ma deg xn h summary treeList l newnode newlist] by (meson assms(1) assms(10) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) assms(8) assms(9)) theorem dele_bmo_cont_corr:"invar_vebt t n \ (both_member_options (vebt_delete t x) y \ x \ y \ both_member_options t y)" proof(induction t n arbitrary: x y rule: invar_vebt.induct) case (1 a b) have "(both_member_options (vebt_delete (Leaf a b) x) y) \ (x \ y \ both_member_options (Leaf a b) y)" by (metis One_nat_def both_member_options_def vebt_buildup.cases vebt_delete.simps(1) vebt_delete.simps(2) vebt_delete.simps(3) membermima.simps(1) naive_member.simps(1)) moreover have "(x \ y \ both_member_options (Leaf a b) y) \(both_member_options (vebt_delete (Leaf a b) x) y)" by (metis One_nat_def both_member_options_def vebt_buildup.cases vebt_delete.simps(1) vebt_delete.simps(2) vebt_delete.simps(3) membermima.simps(1) naive_member.simps(1)) ultimately show ?case by blast next case (2 treeList n summary m deg) hence "deg \ 2" by (metis Suc_leI deg_not_0 dual_order.strict_trans2 less_add_same_cancel1 numerals(2)) hence " (vebt_delete (Node None deg treeList summary) x) = (Node None deg treeList summary)" by simp moreover have "\vebt_member (Node None deg treeList summary) y" by simp moreover hence "\both_member_options (Node None deg treeList summary) y" using invar_vebt.intros(2)[of treeList n summary m deg] 2 by (metis valid_member_both_member_options) moreover hence "\both_member_options (vebt_delete (Node None deg treeList summary) x) y" by simp ultimately show ?case by force next case (3 treeList n summary m deg) hence "deg \ 2" by (metis One_nat_def add_mono le_add1 numeral_2_eq_2 plus_1_eq_Suc set_n_deg_not_0) hence " (vebt_delete (Node None deg treeList summary) x) = (Node None deg treeList summary)" by simp moreover have "\vebt_member (Node None deg treeList summary) y" by simp moreover hence "\both_member_options (Node None deg treeList summary) y" using invar_vebt.intros(3)[of treeList n summary m deg] 3 by (metis valid_member_both_member_options) moreover hence "\both_member_options (vebt_delete (Node None deg treeList summary) x) y" by simp ultimately show ?case by force next case (4 treeList n summary m deg mi ma) hence tvalid: "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(4)[of treeList n summary m deg mi ma] by simp hence "mi \ ma" and "deg div 2 = n" and "ma \ 2^deg" using 4 by (auto simp add: "4.hyps"(3) "4.hyps"(4)) hence dp:"deg \ 2" using "4.hyps"(1) "4.hyps"(3) deg_not_0 div_greater_zero_iff by blast then show ?case proof(cases "x x > ma") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, ma)) deg treeList summary)" using delt_out_of_range[of x mi ma deg treeList summary] \2 \ deg\ by blast then show ?thesis by (metis "4.hyps"(7) True tvalid leD member_inv not_less_iff_gr_or_eq valid_member_both_member_options) next case False hence "mi \ x \ x \ ma" by simp hence "x < 2^deg" using "4.hyps"(8) order.strict_trans1 by blast then show ?thesis proof(cases "x = mi \ x = ma") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node None deg treeList summary)" using del_single_cont[of x mi ma deg treeList summary] \2 \ deg\ by blast moreover hence "invar_vebt (Node None deg treeList summary) deg" using "4"(4) "4.IH"(1) "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) True mi_eq_ma_no_ch tvalid invar_vebt.intros(2) by force moreover hence "\ vebt_member (Node None deg treeList summary) y" by simp moreover hence "\both_member_options (Node None deg treeList summary) y" using calculation(2) valid_member_both_member_options by blast then show ?thesis by (metis True calculation(1) member_inv not_less_iff_gr_or_eq tvalid valid_member_both_member_options) next case False hence mimapr:"mi < ma" by (metis "4.hyps"(7) \mi \ x \ x \ ma\ le_antisym nat_less_le) then show ?thesis proof(cases "x \ mi") case True hence xmi:"x \ mi" by simp let ?h ="high x n" let ?l = "low x n" have "?h < length treeList" using "4"(10) "4"(4) "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) \mi \ x \ x \ ma\ deg_not_0 exp_split_high_low(1) by auto let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:= ?newnode]" have "length treeList = length ?newlist" by simp hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high x n < length treeList\ nth_list_update_eq) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by auto then show ?thesis proof(cases "minNull ?newnode") case True let ?sn = "vebt_delete summary ?h" let ?newma= "(if x = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist ?sn)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_new_node_nil[of mi x ma deg ?h ?l ?newnode treeList ?sn summary ?newlist] by (metis True \2 \ deg\ \deg div 2 = n\ \high x n < length treeList\ \mi < ma\ \mi \ x \ x \ ma\ \x \ mi\ less_not_refl3 order.not_eq_order_implies_strict) moreover have "both_member_options (?delsimp) y \ (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" proof- assume "both_member_options (?delsimp) y" hence "y = mi \ y = ?newma \ (both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist)" using both_member_options_from_complete_tree_to_child[of deg mi ?newma ?newlist ?sn y] dp by (smt (z3) Suc_1 Suc_le_D both_member_options_def membermima.simps(4) naive_member.simps(3)) moreover have "y = mi \ ?thesis" by (meson \x \ mi\ both_member_options_equiv_member vebt_mint.simps(3) mint_member tvalid) moreover have "y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "x = ma") case True let ?maxs = "vebt_maxt ?sn" have "?newma = (if ?maxs = None then mi else 2 ^ (deg div 2) * the ?maxs + the (vebt_maxt ((treeList[(high x n):= vebt_delete (treeList ! (high x n)) (low x n)]) ! the ?maxs)))" using True by force then show ?thesis proof(cases "?maxs = None ") case True then show ?thesis using \(if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma) = (if vebt_maxt (vebt_delete summary (high x n)) = None then mi else 2 ^ (deg div 2) * the (vebt_maxt (vebt_delete summary (high x n))) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! the (vebt_maxt (vebt_delete summary (high x n))))))\ \y = (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma)\ calculation(2) by presburger next case False then obtain maxs where "Some maxs = ?maxs" by force hence "both_member_options ?sn maxs" by (simp add: maxbmo) hence "both_member_options summary maxs \ maxs \ ?h" using "4.IH"(2) by blast hence "?newlist ! the ?maxs = treeList ! maxs" by (metis "4.hyps"(1) \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ option.sel member_bound nothprolist valid_member_both_member_options) have "maxs < 2^m" using "4.hyps"(1) \both_member_options summary maxs \ maxs \ high x n\ member_bound valid_member_both_member_options by blast hence "the (vebt_maxt (?newlist ! the ?maxs)) = the (vebt_maxt (treeList ! maxs))" by (simp add: \treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! the (vebt_maxt (vebt_delete summary (high x n))) = treeList ! maxs\) have "\ z. both_member_options(treeList ! maxs) z" by (simp add: "4.hyps"(5) \both_member_options summary maxs \ maxs \ high x n\ \maxs < 2 ^ m\) moreover have "invar_vebt (treeList ! maxs) n" using 4 by (metis \maxs < 2 ^ m\ inthall member_def) ultimately obtain maxi where "Some maxi = (vebt_maxt (treeList ! maxs))" by (metis empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "maxi < 2^n" by (metis \invar_vebt (treeList ! maxs) n\ maxt_member member_bound) hence "both_member_options (treeList ! maxs) maxi" using \Some maxi = vebt_maxt (treeList ! maxs)\ maxbmo by presburger hence "2 ^ (deg div 2) * the ?maxs + the (vebt_maxt (?newlist ! the ?maxs)) = 2^n * maxs + maxi " by (metis \Some maxi = vebt_maxt (treeList ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \deg div 2 = n\ \the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! the (vebt_maxt (vebt_delete summary (high x n))))) = the (vebt_maxt (treeList ! maxs))\ option.sel) hence "y = 2^n * maxs + maxi" using False True \y = (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma)\ by fastforce hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis "4.hyps"(2) Suc_1 \both_member_options (treeList ! maxs) maxi\ \deg div 2 = n\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ add_leD1 both_member_options_from_chilf_to_complete_tree dp high_inv low_inv mult.commute plus_1_eq_Suc) moreover hence "y \ x" by (metis \both_member_options summary maxs \ maxs \ high x n\ \maxi < 2 ^ n\ \y = 2 ^ n * maxs + maxi\ high_inv mult.commute) ultimately show ?thesis by force qed next case False hence "?newma = ma" by simp moreover hence "y \ x" using False \y = ?newma\ by presburger then show ?thesis by (metis False \y =?newma\ both_member_options_equiv_member vebt_maxt.simps(3) maxt_member tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist) \ ?thesis" proof- assume assm:"both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist" show ?thesis proof(cases "(high y (deg div 2)) = ?h") case True hence "both_member_options ?newnode (low y (deg div 2)) " using hprolist by (metis assm) moreover hence "invar_vebt (treeList ! (high y (deg div 2))) n" by (metis "4.IH"(1) True \high x n < length treeList\ inthall member_def) ultimately have "both_member_options (treeList ! ?h) (low y (deg div 2)) \ (low y (deg div 2)) \ (low x (deg div 2))" by (metis "4.IH"(1) \deg div 2 = n\ \high x n < length treeList\ inthall member_def) then show ?thesis by (metis Suc_1 True \high x n < length treeList\ add_leD1 both_member_options_from_chilf_to_complete_tree dp plus_1_eq_Suc) next case False hence "x \ y" using \deg div 2 = n\ by blast moreover hence "(?newlist ! (high y (deg div 2))) = treeList ! (high y (deg div 2))" using nothprolist using "4.hyps"(2) False \length treeList = length ?newlist\ assm by presburger moreover hence "both_member_options (treeList ! (high y (deg div 2)) ) (low y (deg div 2))" using assm by presburger moreover hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis One_nat_def Suc_leD \length treeList = length ?newlist\ assm both_member_options_from_chilf_to_complete_tree dp numeral_2_eq_2) ultimately show ?thesis by blast qed qed ultimately show ?thesis by fastforce qed moreover have " (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y) \ both_member_options (?delsimp) y" proof- assume "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" hence aa:"x \ y" and bb:"y = mi \ y = ma \ (both_member_options (treeList ! (high y n)) (low y n) \ high y n < length treeList)" apply auto[1] by (metis Suc_1 \deg div 2 = n\ \x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) show "both_member_options (?delsimp) y" proof- have "y = mi \both_member_options (?delsimp) y" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have "y = ma \ both_member_options (?delsimp) y" using aa maxbmo vebt_maxt.simps(3) by presburger moreover have "both_member_options (treeList ! (high y n)) (low y n) \both_member_options (?delsimp) y " proof- assume assmy: "both_member_options (treeList ! (high y n)) (low y n)" then show "both_member_options (?delsimp) y " proof(cases "high y n = ?h") case True moreover hence "?newlist ! (high y n) = ?newnode" using hprolist by auto hence 0:"invar_vebt (treeList !(high y n)) n" using 4 by (metis True \high x n < length treeList\ inthall member_def) moreover have 1:"low y n \ low x n" by (metis True aa bit_split_inv) moreover have 11:" (treeList !(high y n)) \ set treeList" by (metis True \high x n < length treeList\ inthall member_def) ultimately have " (\ xa. both_member_options ?newnode xa = ((low x n) \ xa \ both_member_options (treeList ! ?h) xa))" by (simp add: "4.IH"(1)) hence "((low x n) \ xa \ both_member_options (treeList ! ?h) xa) \ both_member_options ?newnode xa" for xa by blast moreover have "((low x n) \ (low y n) \ both_member_options (treeList ! ?h) (low y n))" using 1 using True assmy by presburger ultimately have "both_member_options ?newnode (low y n)" by blast then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high x n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp hprolist numerals(2)) next case False hence "?newlist ! (high y n) = treeList ! (high y n)" by auto hence "both_member_options (?newlist !(high y n)) (low y n)" using assmy by presburger then show ?thesis by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ \length treeList = length ?newlist\ aa add_leD1 bb both_member_options_def both_member_options_from_chilf_to_complete_tree dp membermima.simps(4) plus_1_eq_Suc) qed qed ultimately show ?thesis using bb by fastforce qed qed ultimately show ?thesis by metis next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if x = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist summary)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_newnode_not_nil[of mi x ma deg ?h ?l ?newnode treeList ?newlist summary] False xmi mimapr using \deg div 2 = n\ \high x n < length treeList\ \mi \ x \ x \ ma\ dp nat_less_le plus_1_eq_Suc by fastforce moreover have "both_member_options ?delsimp y \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- assume ssms: "both_member_options ?delsimp y " hence aaaa: "y = mi \ y = ?newma \ (both_member_options (?newlist ! (high y n)) (low y n) \ high y n < length ?newlist)" by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ both_member_options_def dp membermima.simps(4) naive_member.simps(3)) show " x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- have "y = mi \?thesis" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4) xmi) moreover have " y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "x = ma") case True hence "?newma =?h * 2 ^ (deg div 2) +the(vebt_maxt(?newlist ! ?h))" by metis have "?newlist ! ?h = ?newnode" using hprolist by blast obtain maxi where maxidef:"Some maxi = vebt_maxt(?newlist ! ?h)" by (metis False hprolist vebt_maxt.elims minNull.simps(1) minNull.simps(4)) have aa:"invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high x n < length treeList\ inthall member_def) moreover hence ab:"maxi \ ?l \ both_member_options ?newnode maxi" by (metis "4.IH"(1) \high x n < length treeList\ hprolist inthall maxbmo maxidef member_def) ultimately have ac:"maxi \ ?l \ both_member_options (treeList ! ?h) maxi" by (metis "4.IH"(1) \high x n < length treeList\ inthall member_def) hence ad:"maxi < 2^n" using \invar_vebt (treeList ! high x n) n\ member_bound valid_member_both_member_options by blast then show ?thesis by (metis Suc_1 \(if x = ma then high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n)) else ma) = high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n))\ \deg div 2 = n\ \high x n < length treeList\ \y = (if x = ma then high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n)) else ma)\ ac add_leD1 both_member_options_from_chilf_to_complete_tree dp option.sel high_inv low_inv maxidef plus_1_eq_Suc) next case False then show ?thesis by (simp add: \y = ?newma\ maxbmo) qed qed moreover have "both_member_options (?newlist ! (high y n)) (low y n) \ ?thesis" proof- assume assmy:"both_member_options (?newlist ! (high y n)) (low y n)" then show ?thesis proof(cases "high y n = ?h") case True hence "?newlist ! (high y n) = ?newnode" using hprolist by presburger have "invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high x n < length treeList\ inthall member_def) hence "low y n \ ?l \ both_member_options (treeList ! ?h ) (low y n)" by (metis "4.IH"(1) True \high x n < length treeList\ assmy hprolist inthall member_def) then show ?thesis by (metis Suc_1 True \deg div 2 = n\ \high x n < length treeList\ add_leD1 both_member_options_from_chilf_to_complete_tree dp plus_1_eq_Suc) next case False hence "?newlist ! (high y n) = treeList !(high y n)" by auto then show ?thesis by (metis False Suc_1 \deg div 2 = n\ \length treeList = length ?newlist\ aaaa add_leD1 both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp plus_1_eq_Suc) qed qed ultimately show ?thesis using aaaa by fastforce qed qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have " y = ma \ ?thesis" using assm maxbmo vebt_maxt.simps(3) by presburger moreover have " both_member_options (treeList ! (high y n)) (low y n) \ ?thesis" proof- assume myass: "both_member_options (treeList ! (high y n)) (low y n) " thus ?thesis proof(cases "high y n = ?h") case True hence "low y n \ ?l" by (metis assm bit_split_inv) hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high x n < length treeList\ inthall member_def) hence "both_member_options ?newnode (low y n)" by (metis "4.IH"(1) True \high x n < length treeList\ \low y n \ low x n\ in_set_member inthall myass) then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high x n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp numerals(2) pp) next case False hence pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist by auto then show ?thesis by (metis Suc_1 \deg div 2 = n\ \length treeList = length (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)])\ add_leD1 assm both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) member_inv myass plus_1_eq_Suc tvalid valid_member_both_member_options) qed qed then show ?thesis by (metis Suc_1 Suc_leD \deg div 2 = n\ assm both_member_options_from_complete_tree_to_child calculation(1) calculation(2) dp) qed qed ultimately show ?thesis by metis qed next case False hence "x = mi" by simp have "both_member_options summary (high ma n)" - by (metis "4"(10) "4"(11) "4"(7) "4.hyps"(4) Euclidean_Division.div_eq_0_iff Suc_leI Suc_le_D div_exp_eq dual_order.irrefl high_def mimapr nat.simps(3)) + by (metis "4"(10) "4"(11) "4"(7) "4.hyps"(4) div_eq_0_iff Suc_leI Suc_le_D div_exp_eq dual_order.irrefl high_def mimapr nat.simps(3)) hence "vebt_member summary (high ma n)" using "4.hyps"(1) valid_member_both_member_options by blast obtain summin where "Some summin = vebt_mint summary" by (metis "4.hyps"(1) \vebt_member summary (high ma n)\ empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def) hence "\ z . both_member_options (treeList ! summin) z" by (metis "4.hyps"(1) "4.hyps"(5) both_member_options_equiv_member member_bound mint_member) moreover have "invar_vebt (treeList ! summin) n" by (metis "4"(4) "4.IH"(1) "4.hyps"(1) \Some summin = vebt_mint summary\ member_bound mint_member nth_mem) ultimately obtain lx where "Some lx = vebt_mint (treeList ! summin)" by (metis empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) let ?xn = "summin*2^n + lx" have "?xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)" by (metis False \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \deg div 2 = n\ option.sel) have "vebt_member (treeList ! summin) lx" using \Some lx = vebt_mint (treeList ! summin)\ \invar_vebt (treeList ! summin) n\ mint_member by auto moreover have "summin < 2^m" by (metis "4.hyps"(1) \Some summin = vebt_mint summary\ member_bound mint_member) ultimately have xnin: "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" by (metis "4.hyps"(2) Suc_1 \deg div 2 = n\ \invar_vebt (treeList ! summin) n\ add_leD1 both_member_options_equiv_member both_member_options_from_chilf_to_complete_tree dp high_inv low_inv member_bound plus_1_eq_Suc) let ?h ="high ?xn n" let ?l = "low ?xn n" have "?xn < 2^deg" - by (smt (verit, ccfv_SIG) "4.hyps"(1) "4.hyps"(4) Euclidean_Division.div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) + by (smt (verit, ccfv_SIG) "4.hyps"(1) "4.hyps"(4) div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) hence "?h < length treeList" using "4.hyps"(2) "4.hyps"(3) "4.hyps"(4) \invar_vebt (treeList ! summin) n\ deg_not_0 exp_split_high_low(1) by metis let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:= ?newnode]" have "length treeList = length ?newlist" by simp hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high (summin * 2 ^ n + lx) n < length treeList\ nth_list_update) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp have firstsimp: "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( let newnode = vebt_delete (treeList ! ?h) ?l; newlist = (take ?h treeList @ [ newnode]@drop (?h+1) treeList)in if minNull newnode then( let sn = vebt_delete summary ?h in (Node (Some (?xn, if ?xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (?xn, (if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (newlist ! ?h)) else ma))) deg newlist summary ))" using del_x_mi[of x mi ma deg ?xn ?h summary treeList ?l] by (smt (z3) \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x = mi\ add.commute append_Cons append_Nil dp mimapr nat_less_le plus_1_eq_Suc upd_conv_take_nth_drop) have minxnrel: "?xn \ mi" by (metis "4.hyps"(2) "4.hyps"(9) \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ both_member_options_equiv_member high_inv less_not_refl low_inv member_bound mimapr) then show ?thesis proof(cases "minNull ?newnode") case True let ?sn = "vebt_delete summary ?h" let ?newma= "(if ?xn= ma then (let maxs = vebt_maxt ?sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist ?sn)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist ?sn] False True \deg div 2 = n\ \?h < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ dp less_not_refl3 mimapr by fastforce moreover have "both_member_options (?delsimp) y \ (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" proof- assume "both_member_options (?delsimp) y" hence "y = ?xn \ y = ?newma \ (both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist)" using both_member_options_from_complete_tree_to_child[of deg mi ?newma ?newlist ?sn y] dp by (smt (z3) Suc_1 Suc_le_D both_member_options_def membermima.simps(4) naive_member.simps(3)) moreover have "y = ?xn \ ?thesis" by (metis "4.hyps"(9) False \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ both_member_options_equiv_member high_inv less_not_refl low_inv member_bound mimapr xnin) moreover have "y = ?newma \ ?thesis" proof- assume asmt: "y = ?newma" show ?thesis proof(cases "?xn = ma") case True let ?maxs = "vebt_maxt ?sn" have newmaext:"?newma = (if ?maxs = None then ?xn else 2 ^ (deg div 2) * the ?maxs + the (vebt_maxt ( ?newlist ! the ?maxs)))" using True by force then show ?thesis proof(cases "?maxs = None ") case True hence aa:"?newma = ?xn" using newmaext by auto hence bb: "?newma \ x" using False minxnrel by presburger hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" using xnin newmaext minxnrel asmt by simp moreover have "?xn = y" using aa asmt by simp ultimately have "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by simp then show ?thesis using bb using \summin * 2 ^ n + lx = y\ \y = ?xn \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y\ by blast next case False then obtain maxs where "Some maxs = ?maxs" by force hence "both_member_options ?sn maxs" by (simp add: maxbmo) hence "both_member_options summary maxs \ maxs \ ?h" using "4.IH"(2) by blast hence "?newlist ! the ?maxs = treeList ! maxs" by (metis "4.hyps"(1) \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ option.sel member_bound nothprolist valid_member_both_member_options) have "maxs < 2^m" using "4.hyps"(1) \both_member_options summary maxs \ maxs \ high (summin * 2 ^ n + lx) n\ member_bound valid_member_both_member_options by blast hence "the (vebt_maxt (?newlist ! the ?maxs)) = the (vebt_maxt (treeList ! maxs))" using \?newlist ! the (vebt_maxt ?sn) = treeList ! maxs\ by presburger have "\ z. both_member_options(treeList ! maxs) z" using "4.hyps"(5) \both_member_options summary maxs \ maxs \?h\ \maxs < 2 ^ m\ by blast moreover have "invar_vebt (treeList ! maxs) n" using 4 by (metis \maxs < 2 ^ m\ inthall member_def) ultimately obtain maxi where "Some maxi = (vebt_maxt (treeList ! maxs))" by (metis empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "maxi < 2^n" by (metis \invar_vebt (treeList ! maxs) n\ maxt_member member_bound) hence "both_member_options (treeList ! maxs) maxi" using \Some maxi = vebt_maxt (treeList ! maxs)\ maxbmo by presburger hence "2 ^ (deg div 2) * the ?maxs + the (vebt_maxt (?newlist ! the ?maxs)) = 2^n * maxs + maxi " by (metis \Some maxi = vebt_maxt (treeList ! maxs)\ \Some maxs = vebt_maxt ?sn\ \deg div 2 = n\ \the (vebt_maxt (?newlist ! the (vebt_maxt ?sn))) = the (vebt_maxt (treeList ! maxs))\ option.sel) hence "?newma = 2^n * maxs + maxi" using False True by auto hence "y = 2^n * maxs + maxi" using asmt by simp hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis "4.hyps"(2) Suc_1 \both_member_options (treeList ! maxs) maxi\ \deg div 2 = n\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ add_leD1 both_member_options_from_chilf_to_complete_tree dp high_inv low_inv mult.commute plus_1_eq_Suc) moreover hence "y \ x" by (metis "4.hyps"(9) True \Some maxi = vebt_maxt (treeList ! maxs)\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ \x = mi\ \y = 2 ^ n * maxs + maxi\ high_inv less_not_refl low_inv maxbmo minxnrel mult.commute) ultimately show ?thesis by force qed next case False hence "?newma = ma" by simp moreover hence "mi \ ma" using mimapr by blast moreover hence "y \ x" using False \y = ?newma\ \x = mi\ by auto then show ?thesis by (metis False \y =?newma\ both_member_options_equiv_member vebt_maxt.simps(3) maxt_member tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist) \ ?thesis" proof- assume assm:"both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist" show ?thesis proof(cases "(high y (deg div 2)) = ?h") case True hence 000:"both_member_options ?newnode (low y (deg div 2)) " using hprolist by (metis assm) hence 001:"invar_vebt (treeList ! (high y (deg div 2))) n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger then show ?thesis proof(cases "low y n = ?l") case True hence "y = ?xn" by (metis "000" "4.IH"(1) \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ inthall member_def) then show ?thesis using calculation(2) by blast next case False hence "both_member_options (treeList ! ?h) (low y (deg div 2)) \ (low y (deg div 2)) \ (low ?xn (deg div 2))" using "4.IH"(1) \deg div 2 = n\ \high ?xn n < length treeList\ inthall member_def by (metis "000") then show ?thesis by (metis "4.hyps"(2) "4.hyps"(9) Suc_1 Suc_leD True \deg div 2 = n\ \length treeList = length ?newlist\ \x = mi\ assm both_member_options_from_chilf_to_complete_tree dp less_not_refl mimapr) qed next case False hence "x \ y" by (metis "4.hyps"(2) "4.hyps"(9) \deg div 2 = n\ \length treeList = length ?newlist\ \x = mi\ assm less_not_refl mimapr nothprolist) moreover hence "(?newlist ! (high y (deg div 2))) = treeList ! (high y (deg div 2))" using nothprolist using "4.hyps"(2) False \length treeList = length ?newlist\ assm by presburger moreover hence "both_member_options (treeList ! (high y (deg div 2)) ) (low y (deg div 2))" using assm by presburger moreover hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis One_nat_def Suc_leD \length treeList = length ?newlist\ assm both_member_options_from_chilf_to_complete_tree dp numeral_2_eq_2) ultimately show ?thesis by blast qed qed ultimately show ?thesis by fastforce qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" using False assm by force moreover have " y = ma \ ?thesis" by (smt (z3) Suc_le_D both_member_options_def dp membermima.simps(4) nat_1_add_1 plus_1_eq_Suc) moreover have " both_member_options (treeList ! (high y n)) (low y n) \ ?thesis" proof- assume myass: "both_member_options (treeList ! (high y n)) (low y n) " thus ?thesis proof(cases "high y n = ?h") case True hence "high y n = ?h" by simp then show ?thesis proof(cases "low y n = ?l") case True hence "y = ?xn" by (metis \high y n = high (summin * 2 ^ n + lx) n\ bit_split_inv) then show ?thesis by (metis Suc_le_D both_member_options_def dp membermima.simps(4) nat_1_add_1 plus_1_eq_Suc) next case False hence "low y n \ ?l" by (metis assm bit_split_inv) hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger hence "both_member_options ?newnode (low y n)" using "4.IH"(1) False True \high (summin * 2 ^ n + lx) n < length treeList\ myass by auto then show ?thesis by (metis True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \length treeList = length ?newlist\ add_leD1 both_member_options_from_chilf_to_complete_tree dp nat_1_add_1 pp) qed next case False hence pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist abcv by (metis "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) assm deg_not_0 exp_split_high_low(1) member_bound tvalid valid_member_both_member_options) then show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ \length treeList = length ?newlist\ abcv both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp numerals(2)) qed qed then show ?thesis using abcv calculation(1) calculation(2) by fastforce qed qed ultimately show ?thesis by metis next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist summary)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_not_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist] by (metis "4.hyps"(3) "4.hyps"(4) False \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \high (summin * 2 ^ n + lx) n < length treeList\ \x = mi\ add_self_div_2 dp option.sel less_not_refl mimapr) moreover have "both_member_options ?delsimp y \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- assume ssms: "both_member_options ?delsimp y " hence aaaa: "y = ?xn \ y = ?newma \ (both_member_options (?newlist ! (high y n)) (low y n) \ high y n < length ?newlist)" by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ both_member_options_def dp membermima.simps(4) naive_member.simps(3)) show " x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- have "y = ?xn \?thesis" using \x = mi\ minxnrel xnin by blast moreover have " y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "?xn = ma") case True hence aaa:"?newma =?h * 2 ^ (deg div 2) +the(vebt_maxt(?newlist ! ?h))" by metis have "?newlist ! ?h = ?newnode" using hprolist by blast obtain maxi where maxidef:"Some maxi = vebt_maxt(?newlist ! ?h)" by (metis False hprolist vebt_maxt.elims minNull.simps(1) minNull.simps(4)) have aa:"invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high ?xn n < length treeList\ inthall member_def) moreover hence ab:"maxi \ ?l \ both_member_options ?newnode maxi" by (metis "4.IH"(1) \high ?xn n < length treeList\ hprolist inthall maxbmo maxidef member_def) ultimately have ac:"maxi \ ?l \ both_member_options (treeList ! ?h) maxi" by (metis "4.IH"(1) \high ?xn n < length treeList\ inthall member_def) hence ad:"maxi < 2^n" by (meson aa member_bound valid_member_both_member_options) then show ?thesis using Suc_1 aaa \y = ?newma\ ac add_leD1 by (metis "4.hyps"(2) "4.hyps"(9) Suc_leD \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \x = mi\ both_member_options_from_chilf_to_complete_tree dp option.sel high_inv less_not_refl low_inv maxidef mimapr) next case False then show ?thesis by (metis \mi \ x \ x \ ma\ \x = mi\ \y = ?newma\ both_member_options_equiv_member leD vebt_maxt.simps(3) maxt_member mimapr tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y n)) (low y n)\ high y n < length ?newlist) \ ?thesis" proof- assume assmy:"(both_member_options (?newlist ! (high y n)) (low y n)\ high y n < length ?newlist)" then show ?thesis proof(cases "high y n = ?h") case True hence "?newlist ! (high y n) = ?newnode" using hprolist by presburger have "invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high ?xn n < length treeList\ inthall member_def) then show ?thesis proof(cases "low y n= ?l") case True hence "y = ?xn" using "4.IH"(1) \high (summin * 2 ^ n + lx) n < length treeList\ \treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! high y n = vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)\ assmy by force then show ?thesis using calculation(1) by blast next case False hence "low y n \ ?l \ both_member_options (treeList ! ?h ) (low y n)" using assmy by (metis "4.IH"(1) "4.hyps"(2) \?newlist ! high y n = vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)\ \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ high_inv inthall member_bound member_def) then show ?thesis by (metis "4.hyps"(2) "4.hyps"(9) Suc_1 Suc_leD True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \mi \ x \ x \ ma\ \x = mi\ both_member_options_from_chilf_to_complete_tree dp leD mimapr) qed next case False hence "?newlist ! (high y n) = treeList !(high y n)" by (smt (z3) "4.hyps"(1) "4.hyps"(2) "4.hyps"(3) "4.hyps"(4) "4.hyps"(8) \length treeList = length ?newlist\ \ma \ 2 ^ deg\ aaaa calculation(2) deg_not_0 exp_split_high_low(1) less_le_trans member_inv mimapr nothprolist tvalid valid_member_both_member_options) hence "both_member_options (treeList !(high y n)) (low y n)" using assmy by presburger moreover have "x \ y" by (metis "4.hyps"(1) "4.hyps"(4) "4.hyps"(9) \invar_vebt (treeList ! summin) n\ \x < 2 ^ deg\ \x = mi\ calculation deg_not_0 exp_split_high_low(1) mimapr not_less_iff_gr_or_eq) moreover have "high y n < length ?newlist" using assmy by blast moreover hence "high y n < length treeList" using \length treeList = length ?newlist\ by presburger ultimately show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ both_member_options_from_chilf_to_complete_tree dp numerals(2)) qed qed ultimately show ?thesis using aaaa by fastforce qed qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" using \x = mi\ assm by blast moreover have " y = ma \ ?thesis" by (smt (z3) Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have " ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n)) \ ?thesis" proof- assume myass: "( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n)) " thus ?thesis proof(cases "high y n = ?h") case True then show ?thesis proof(cases "low y n = ?l") case True then show ?thesis by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ \length treeList = length (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)])\ add_leD1 bit_split_inv both_member_options_def both_member_options_from_chilf_to_complete_tree dp membermima.simps(4) myass nth_list_update_neq plus_1_eq_Suc) next case False hence "low y n \ ?l" by simp hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" by (metis "4.IH"(1) \high ?xn n < length treeList\ inthall member_def) hence "both_member_options ?newnode (low y n)" by (metis "4.IH"(1) True \high ?xn n < length treeList\ \low y n \ low ?xn n\ in_set_member inthall myass) then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp numerals(2) pp) qed next case False have pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist[of "high y n"] False by (metis "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) assm deg_not_0 exp_split_high_low(1) member_bound tvalid valid_member_both_member_options) then show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ \length treeList = length ?newlist\ abcv both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp numerals(2)) qed qed then show ?thesis using abcv calculation(1) calculation(2) by fastforce qed qed ultimately show ?thesis by metis qed qed qed qed next case (5 treeList n summary m deg mi ma) hence tvalid: "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(5)[of treeList n summary m deg mi ma] by simp hence "mi \ ma" and "deg div 2 = n" and "ma \ 2^deg" using 5 by (auto simp add: "5.hyps"(3) "5.hyps"(4)) hence dp:"deg \ 2" by (meson vebt_maxt.simps(3) maxt_member member_inv tvalid) hence nmpr:"n\ 1 \ m = Suc n" using "5.hyps"(3) \deg div 2 = n\ by linarith then show ?case proof(cases "x x > ma") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, ma)) deg treeList summary)" using delt_out_of_range[of x mi ma deg treeList summary] \2 \ deg\ by blast then show ?thesis by (metis "5.hyps"(7) True tvalid leD member_inv not_less_iff_gr_or_eq valid_member_both_member_options) next case False hence "mi \ x \ x \ ma" by simp hence xdegrel:"x < 2^deg" using "5.hyps"(8) order.strict_trans1 by blast then show ?thesis proof(cases "x = mi \ x = ma") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node None deg treeList summary)" using del_single_cont[of x mi ma deg treeList summary] \2 \ deg\ by blast moreover hence "invar_vebt (Node None deg treeList summary) deg" using "5"(4) "5.IH"(1) "5.hyps"(1) "5.hyps"(3) "5.hyps"(4) True mi_eq_ma_no_ch tvalid invar_vebt.intros(3) by force moreover hence "\ vebt_member (Node None deg treeList summary) y" by simp moreover hence "\both_member_options (Node None deg treeList summary) y" using calculation(2) valid_member_both_member_options by blast then show ?thesis by (metis True calculation(1) member_inv not_less_iff_gr_or_eq tvalid valid_member_both_member_options) next case False hence mimapr:"mi < ma" by (metis "5.hyps"(7) \mi \ x \ x \ ma\ le_antisym nat_less_le) then show ?thesis proof(cases "x \ mi") case True hence xmi:"x \ mi" by simp let ?h ="high x n" let ?l = "low x n" have "?h < length treeList" using xdegrel 5 by (metis \deg div 2 = n\ deg_not_0 div_greater_zero_iff dp exp_split_high_low(1) zero_less_numeral) let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:=?newnode]" have "length treeList = length ?newlist" by simp hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high x n < length treeList\ nth_list_update_eq) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp then show ?thesis proof(cases "minNull ?newnode") case True let ?sn = "vebt_delete summary ?h" let ?newma= "(if x = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist ?sn)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_new_node_nil[of mi x ma deg ?h ?l ?newnode treeList ?sn summary ?newlist] by (metis True \2 \ deg\ \deg div 2 = n\ \high x n < length treeList\ \mi < ma\ \mi \ x \ x \ ma\ \x \ mi\ less_not_refl3 order.not_eq_order_implies_strict) moreover have "both_member_options (?delsimp) y \ (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" proof- assume "both_member_options (?delsimp) y" hence "y = mi \ y = ?newma \ (both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist)" using both_member_options_from_complete_tree_to_child[of deg mi ?newma ?newlist ?sn y] dp by (smt (z3) Suc_1 Suc_le_D both_member_options_def membermima.simps(4) naive_member.simps(3)) moreover have "y = mi \ ?thesis" by (meson \x \ mi\ both_member_options_equiv_member vebt_mint.simps(3) mint_member tvalid) moreover have "y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "x = ma") case True let ?maxs = "vebt_maxt ?sn" have newmapropy:"?newma = (if ?maxs = None then mi else 2 ^ (deg div 2) * the ?maxs + the (vebt_maxt (?newlist ! the ?maxs)))" using True by force then show ?thesis proof(cases "?maxs = None ") case True then show ?thesis using \y = (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma)\ calculation(2) newmapropy by presburger next case False then obtain maxs where "Some maxs = ?maxs" by force hence "both_member_options ?sn maxs" by (simp add: maxbmo) hence "both_member_options summary maxs \ maxs \ ?h" using "5.IH"(2) by blast hence "?newlist ! the ?maxs = treeList ! maxs" by (metis "5.hyps"(1) \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ option.sel member_bound nothprolist valid_member_both_member_options) have "maxs < 2^m" using "5.hyps"(1) \both_member_options summary maxs \ maxs \ high x n\ member_bound valid_member_both_member_options by blast hence "the (vebt_maxt (?newlist ! the ?maxs)) = the (vebt_maxt (treeList ! maxs))" by (metis \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \both_member_options summary maxs \ maxs \ high x n\ option.sel nth_list_update_neq) have "\ z. both_member_options(treeList ! maxs) z" by (simp add: "5.hyps"(5) \both_member_options summary maxs \ maxs \ high x n\ \maxs < 2 ^ m\) moreover have "invar_vebt (treeList ! maxs) n" using 5 by (metis \maxs < 2 ^ m\ inthall member_def) ultimately obtain maxi where "Some maxi = (vebt_maxt (treeList ! maxs))" by (metis empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "maxi < 2^n" by (metis \invar_vebt (treeList ! maxs) n\ maxt_member member_bound) hence "both_member_options (treeList ! maxs) maxi" using \Some maxi = vebt_maxt (treeList ! maxs)\ maxbmo by presburger hence "2 ^ (deg div 2) * the ?maxs + the (vebt_maxt (?newlist ! the ?maxs)) = 2^n * maxs + maxi " by (metis \Some maxi = vebt_maxt (treeList ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \deg div 2 = n\ \the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! the (vebt_maxt (vebt_delete summary (high x n))))) = the (vebt_maxt (treeList ! maxs))\ option.sel) hence "y = 2^n * maxs + maxi" using False \y = (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma)\ newmapropy by presburger hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis "5.hyps"(2) Suc_1 \both_member_options (treeList ! maxs) maxi\ \deg div 2 = n\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ add_leD1 both_member_options_from_chilf_to_complete_tree dp high_inv low_inv mult.commute plus_1_eq_Suc) moreover hence "y \ x" by (metis \both_member_options summary maxs \ maxs \ high x n\ \maxi < 2 ^ n\ \y = 2 ^ n * maxs + maxi\ high_inv mult.commute) ultimately show ?thesis by force qed next case False hence "?newma = ma" by simp moreover hence "y \ x" using False \y = ?newma\ by presburger then show ?thesis by (metis False \y =?newma\ both_member_options_equiv_member vebt_maxt.simps(3) maxt_member tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist) \ ?thesis" proof- assume assm:"both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist" show ?thesis proof(cases "(high y (deg div 2)) = ?h") case True hence "both_member_options ?newnode (low y (deg div 2)) " using hprolist by (metis assm) moreover hence "invar_vebt (treeList ! (high y (deg div 2))) n" by (metis "5.IH"(1) True \high x n < length treeList\ inthall member_def) ultimately have "both_member_options (treeList ! ?h) (low y (deg div 2)) \ (low y (deg div 2)) \ (low x (deg div 2))" by (metis "5.IH"(1) \deg div 2 = n\ \high x n < length treeList\ inthall member_def) then show ?thesis by (metis Suc_1 True \high x n < length treeList\ add_leD1 both_member_options_from_chilf_to_complete_tree dp plus_1_eq_Suc) next case False hence "x \ y" using \deg div 2 = n\ by blast moreover hence "(?newlist ! (high y (deg div 2))) = treeList ! (high y (deg div 2))" using nothprolist using "5.hyps"(2) False \length treeList = length ?newlist\ assm by presburger moreover hence "both_member_options (treeList ! (high y (deg div 2)) ) (low y (deg div 2))" using assm by presburger moreover hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis One_nat_def Suc_leD \length treeList = length ?newlist\ assm both_member_options_from_chilf_to_complete_tree dp numeral_2_eq_2) ultimately show ?thesis by blast qed qed ultimately show ?thesis by fastforce qed moreover have " (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y) \ both_member_options (?delsimp) y" proof- assume "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" hence aa:"x \ y" and bb:"y = mi \ y = ma \ (both_member_options (treeList ! (high y n)) (low y n) \ high y n < length treeList)" apply auto[1] by (metis Suc_1 \deg div 2 = n\ \x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) show "both_member_options (?delsimp) y" proof- have "y = mi \both_member_options (?delsimp) y" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have "y = ma \ both_member_options (?delsimp) y" using aa maxbmo vebt_maxt.simps(3) by presburger moreover have "both_member_options (treeList ! (high y n)) (low y n) \both_member_options (?delsimp) y " proof- assume assmy: "both_member_options (treeList ! (high y n)) (low y n)" then show "both_member_options (?delsimp) y " proof(cases "high y n = ?h") case True moreover hence "?newlist ! (high y n) = ?newnode" using hprolist by auto hence 0:"invar_vebt (treeList !(high y n)) n" using 5 by (metis True \high x n < length treeList\ inthall member_def) moreover have 1:"low y n \ low x n" by (metis True aa bit_split_inv) moreover have 11:" (treeList !(high y n)) \ set treeList" by (metis True \high x n < length treeList\ inthall member_def) ultimately have " (\ xa. both_member_options ?newnode xa = ((low x n) \ xa \ both_member_options (treeList ! ?h) xa))" by (simp add: "5.IH"(1)) hence "((low x n) \ xa \ both_member_options (treeList ! ?h) xa) \ both_member_options ?newnode xa" for xa by blast moreover have "((low x n) \ (low y n) \ both_member_options (treeList ! ?h) (low y n))" using 1 using True assmy by presburger ultimately have "both_member_options ?newnode (low y n)" by blast then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high x n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp hprolist numerals(2)) next case False hence "?newlist ! (high y n) = treeList ! (high y n)" by auto hence "both_member_options (?newlist !(high y n)) (low y n)" using assmy by presburger then show ?thesis by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ \length treeList = length ?newlist\ aa add_leD1 bb both_member_options_def both_member_options_from_chilf_to_complete_tree dp membermima.simps(4) plus_1_eq_Suc) qed qed ultimately show ?thesis using bb by fastforce qed qed ultimately show ?thesis by metis next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if x = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist summary)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_newnode_not_nil[of mi x ma deg ?h ?l ?newnode treeList ?newlist summary] False xmi mimapr using \deg div 2 = n\ \high x n < length treeList\ \mi \ x \ x \ ma\ dp nat_less_le plus_1_eq_Suc by fastforce moreover have "both_member_options ?delsimp y \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- assume ssms: "both_member_options ?delsimp y " hence aaaa: "y = mi \ y = ?newma \ (both_member_options (?newlist ! (high y n)) (low y n) \ high y n < length ?newlist)" by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ both_member_options_def dp membermima.simps(4) naive_member.simps(3)) show " x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- have "y = mi \?thesis" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4) xmi) moreover have " y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "x = ma") case True hence "?newma =?h * 2 ^ (deg div 2) +the(vebt_maxt(?newlist ! ?h))" by metis have "?newlist ! ?h = ?newnode" using hprolist by blast obtain maxi where maxidef:"Some maxi = vebt_maxt(?newlist ! ?h)" by (metis False hprolist vebt_maxt.elims minNull.simps(1) minNull.simps(4)) have aa:"invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high x n < length treeList\ inthall member_def) moreover hence ab:"maxi \ ?l \ both_member_options ?newnode maxi" by (metis "5.IH"(1) \high x n < length treeList\ hprolist inthall maxbmo maxidef member_def) ultimately have ac:"maxi \ ?l \ both_member_options (treeList ! ?h) maxi" by (metis "5.IH"(1) \high x n < length treeList\ inthall member_def) hence ad:"maxi < 2^n" using \invar_vebt (treeList ! high x n) n\ member_bound valid_member_both_member_options by blast then show ?thesis by (metis Suc_1 \(if x = ma then high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n)) else ma) = high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n))\ \deg div 2 = n\ \high x n < length treeList\ \y = (if x = ma then high x n * 2 ^ (deg div 2) + the (vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! high x n)) else ma)\ ac add_leD1 both_member_options_from_chilf_to_complete_tree dp option.sel high_inv low_inv maxidef plus_1_eq_Suc) next case False then show ?thesis by (simp add: \y = ?newma\ maxbmo) qed qed moreover have "both_member_options (?newlist ! (high y n)) (low y n) \ ?thesis" proof- assume assmy:"both_member_options (?newlist ! (high y n)) (low y n)" then show ?thesis proof(cases "high y n = ?h") case True hence "?newlist ! (high y n) = ?newnode" using hprolist by presburger have "invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high x n < length treeList\ inthall member_def) hence "low y n \ ?l \ both_member_options (treeList ! ?h ) (low y n)" by (metis "5.IH"(1) True \high x n < length treeList\ assmy hprolist inthall member_def) then show ?thesis by (metis Suc_1 True \deg div 2 = n\ \high x n < length treeList\ add_leD1 both_member_options_from_chilf_to_complete_tree dp plus_1_eq_Suc) next case False hence "?newlist ! (high y n) = treeList !(high y n)" by auto then show ?thesis by (metis False Suc_1 \deg div 2 = n\ \length treeList = length ?newlist\ aaaa add_leD1 both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp plus_1_eq_Suc) qed qed ultimately show ?thesis using aaaa by fastforce qed qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" by (metis Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have " y = ma \ ?thesis" using assm maxbmo vebt_maxt.simps(3) by presburger moreover have " both_member_options (treeList ! (high y n)) (low y n) \ ?thesis" proof- assume myass: "both_member_options (treeList ! (high y n)) (low y n) " thus ?thesis proof(cases "high y n = ?h") case True hence "low y n \ ?l" by (metis assm bit_split_inv) hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high x n < length treeList\ inthall member_def) hence "both_member_options ?newnode (low y n)" by (metis "5.IH"(1) True \high x n < length treeList\ \low y n \ low x n\ in_set_member inthall myass) then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high x n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp numerals(2) pp) next case False hence pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist abcv by auto then show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ \length treeList = length ?newlist\ abcv both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp numerals(2)) qed qed then show ?thesis using abcv calculation(1) calculation(2) by fastforce qed qed ultimately show ?thesis by metis qed next case False hence "x = mi" by simp have "both_member_options summary (high ma n)" - by (metis "5"(10) "5"(11) "5"(7) "5.hyps"(4) Euclidean_Division.div_eq_0_iff Suc_leI Suc_le_D div_exp_eq dual_order.irrefl high_def mimapr nat.simps(3)) + by (metis "5"(10) "5"(11) "5"(7) "5.hyps"(4) div_eq_0_iff Suc_leI Suc_le_D div_exp_eq dual_order.irrefl high_def mimapr nat.simps(3)) hence "vebt_member summary (high ma n)" using "5.hyps"(1) valid_member_both_member_options by blast obtain summin where "Some summin = vebt_mint summary" by (metis "5.hyps"(1) \vebt_member summary (high ma n)\ empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def) hence "\ z . both_member_options (treeList ! summin) z" by (metis "5.hyps"(1) "5.hyps"(5) both_member_options_equiv_member member_bound mint_member) moreover have "invar_vebt (treeList ! summin) n" by (metis "5.IH"(1) "5.hyps"(1) "5.hyps"(2) \Some summin = vebt_mint summary\ member_bound mint_member nth_mem) ultimately obtain lx where "Some lx = vebt_mint (treeList ! summin)" by (metis empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) let ?xn = "summin*2^n + lx" have "?xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)" by (metis False \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \deg div 2 = n\ option.sel) have "vebt_member (treeList ! summin) lx" using \Some lx = vebt_mint (treeList ! summin)\ \invar_vebt (treeList ! summin) n\ mint_member by auto moreover have "summin < 2^m" by (metis "5.hyps"(1) \Some summin = vebt_mint summary\ member_bound mint_member) ultimately have xnin: "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" by (metis "5.hyps"(2) Suc_1 \deg div 2 = n\ \invar_vebt (treeList ! summin) n\ add_leD1 both_member_options_equiv_member both_member_options_from_chilf_to_complete_tree dp high_inv low_inv member_bound plus_1_eq_Suc) let ?h ="high ?xn n" let ?l = "low ?xn n" have "?xn < 2^deg" - by (smt (verit, ccfv_SIG) "5.hyps"(1) "5.hyps"(4) Euclidean_Division.div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) + by (smt (verit, ccfv_SIG) "5.hyps"(1) "5.hyps"(4) div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) hence "?h < length treeList" using "5.hyps"(2) \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:= ?newnode]" have "length treeList = length ?newlist" by simp hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high (summin * 2 ^ n + lx) n < length treeList\ nth_list_update) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp have firstsimp: "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x =( let newnode = vebt_delete (treeList ! ?h) ?l; newlist = treeList[?h:= ?newnode] in if minNull newnode then( let sn = vebt_delete summary ?h in (Node (Some (?xn, if ?xn = ma then (let maxs = vebt_maxt sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (newlist ! the maxs)) ) ) else ma)) deg newlist sn) )else (Node (Some (?xn, (if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (newlist ! ?h)) else ma))) deg newlist summary ))" using del_x_mi[of x mi ma deg ?xn ?h summary treeList ?l] \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x = mi\ dp mimapr nat_less_le by smt have minxnrel: "?xn \ mi" by (metis "5.hyps"(2) "5.hyps"(9) \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ both_member_options_equiv_member high_inv less_not_refl low_inv member_bound mimapr) then show ?thesis proof(cases "minNull ?newnode") case True let ?sn = "vebt_delete summary ?h" let ?newma= "(if ?xn= ma then (let maxs = vebt_maxt ?sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist ?sn)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist ?sn] False True \deg div 2 = n\ \?h < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ dp less_not_refl3 mimapr by fastforce moreover have "both_member_options (?delsimp) y \ (x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)" proof- assume "both_member_options (?delsimp) y" hence "y = ?xn \ y = ?newma \ (both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist)" using both_member_options_from_complete_tree_to_child[of deg mi ?newma ?newlist ?sn y] dp by (smt (z3) Suc_1 Suc_le_D both_member_options_def membermima.simps(4) naive_member.simps(3)) moreover have "y = ?xn \ ?thesis" by (metis "5.hyps"(9) False \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ both_member_options_equiv_member high_inv less_not_refl low_inv member_bound mimapr xnin) moreover have "y = ?newma \ ?thesis" proof- assume asmt: "y = ?newma" show ?thesis proof(cases "?xn = ma") case True let ?maxs = "vebt_maxt ?sn" have newmaext:"?newma = (if ?maxs = None then ?xn else 2 ^ (deg div 2) * the ?maxs + the (vebt_maxt ( ?newlist ! the ?maxs)))" using True by force then show ?thesis proof(cases "?maxs = None ") case True hence aa:"?newma = ?xn" using newmaext by auto hence bb: "?newma \ x" using False minxnrel by presburger hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" using xnin newmaext minxnrel asmt by simp moreover have "?xn = y" using aa asmt by simp ultimately have "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by simp then show ?thesis using bb using \summin * 2 ^ n + lx = y\ \y = ?xn \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y\ by blast next case False then obtain maxs where "Some maxs = ?maxs" by force hence "both_member_options ?sn maxs" by (simp add: maxbmo) hence "both_member_options summary maxs \ maxs \ ?h" using "5.IH"(2) by blast hence "?newlist ! the ?maxs = treeList ! maxs" by (metis "5.hyps"(1) \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ option.sel member_bound nothprolist valid_member_both_member_options) have "maxs < 2^m" using "5.hyps"(1) \both_member_options summary maxs \ maxs \ high (summin * 2 ^ n + lx) n\ member_bound valid_member_both_member_options by blast hence "the (vebt_maxt (?newlist ! the ?maxs)) = the (vebt_maxt (treeList ! maxs))" using \?newlist ! the (vebt_maxt ?sn) = treeList ! maxs\ by presburger have "\ z. both_member_options(treeList ! maxs) z" using "5.hyps"(5) \both_member_options summary maxs \ maxs \?h\ \maxs < 2 ^ m\ by blast moreover have "invar_vebt (treeList ! maxs) n" using 5 by (metis \maxs < 2 ^ m\ inthall member_def) ultimately obtain maxi where "Some maxi = (vebt_maxt (treeList ! maxs))" by (metis empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "maxi < 2^n" by (metis \invar_vebt (treeList ! maxs) n\ maxt_member member_bound) hence "both_member_options (treeList ! maxs) maxi" using \Some maxi = vebt_maxt (treeList ! maxs)\ maxbmo by presburger hence "2 ^ (deg div 2) * the ?maxs + the (vebt_maxt (?newlist ! the ?maxs)) = 2^n * maxs + maxi " by (metis \Some maxi = vebt_maxt (treeList ! maxs)\ \Some maxs = vebt_maxt ?sn\ \deg div 2 = n\ \the (vebt_maxt (?newlist ! the (vebt_maxt ?sn))) = the (vebt_maxt (treeList ! maxs))\ option.sel) hence "?newma = 2^n * maxs + maxi" using False True by auto hence "y = 2^n * maxs + maxi" using asmt by simp hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis "5.hyps"(2) Suc_1 \both_member_options (treeList ! maxs) maxi\ \deg div 2 = n\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ add_leD1 both_member_options_from_chilf_to_complete_tree dp high_inv low_inv mult.commute plus_1_eq_Suc) moreover hence "y \ x" by (metis "5.hyps"(9) True \Some maxi = vebt_maxt (treeList ! maxs)\ \maxi < 2 ^ n\ \maxs < 2 ^ m\ \x = mi\ \y = 2 ^ n * maxs + maxi\ high_inv less_not_refl low_inv maxbmo minxnrel mult.commute) ultimately show ?thesis by force qed next case False hence "?newma = ma" by simp moreover hence "mi \ ma" using mimapr by blast moreover hence "y \ x" using False \y = ?newma\ \x = mi\ by auto then show ?thesis by (metis False \y =?newma\ both_member_options_equiv_member vebt_maxt.simps(3) maxt_member tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist) \ ?thesis" proof- assume assm:"both_member_options (?newlist ! (high y (deg div 2))) (low y (deg div 2)) \ (high y (deg div 2)) < length ?newlist" show ?thesis proof(cases "(high y (deg div 2)) = ?h") case True hence 000:"both_member_options ?newnode (low y (deg div 2)) " using hprolist by (metis assm) hence 001:"invar_vebt (treeList ! (high y (deg div 2))) n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger then show ?thesis proof(cases "low y n = ?l") case True hence "y = ?xn" by (metis "000" "5.IH"(1) \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ inthall member_def) then show ?thesis using calculation(2) by blast next case False hence "both_member_options (treeList ! ?h) (low y (deg div 2)) \ (low y (deg div 2)) \ (low ?xn (deg div 2))" using "5.IH"(1) \deg div 2 = n\ \high ?xn n < length treeList\ inthall member_def by (metis "000") then show ?thesis by (metis "5.hyps"(2) "5.hyps"(9) Suc_1 Suc_leD True \deg div 2 = n\ \length treeList = length ?newlist\ \x = mi\ assm both_member_options_from_chilf_to_complete_tree dp less_not_refl mimapr) qed next case False hence "x \ y" by (metis "5.hyps"(2) "5.hyps"(9) \deg div 2 = n\ \length treeList = length ?newlist\ \x = mi\ assm less_not_refl mimapr nothprolist) moreover hence "(?newlist ! (high y (deg div 2))) = treeList ! (high y (deg div 2))" using nothprolist using "5.hyps"(2) False \length treeList = length ?newlist\ assm by presburger moreover hence "both_member_options (treeList ! (high y (deg div 2)) ) (low y (deg div 2))" using assm by presburger moreover hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) y" by (metis One_nat_def Suc_leD \length treeList = length ?newlist\ assm both_member_options_from_chilf_to_complete_tree dp numeral_2_eq_2) ultimately show ?thesis by blast qed qed ultimately show ?thesis by fastforce qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" using False assm by force moreover have " y = ma \ ?thesis" by (smt (z3) Suc_le_D both_member_options_def dp membermima.simps(4) nat_1_add_1 plus_1_eq_Suc) moreover have " both_member_options (treeList ! (high y n)) (low y n) \ ?thesis" proof- assume myass: "both_member_options (treeList ! (high y n)) (low y n) " thus ?thesis proof(cases "high y n = ?h") case True hence "high y n = ?h" by simp then show ?thesis proof(cases "low y n = ?l") case True hence "y = ?xn" by (metis \high y n = high (summin * 2 ^ n + lx) n\ bit_split_inv) then show ?thesis by (metis Suc_le_D both_member_options_def dp membermima.simps(4) nat_1_add_1 plus_1_eq_Suc) next case False hence "low y n \ ?l" by (metis assm bit_split_inv) hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger hence "both_member_options ?newnode (low y n)" using "5.IH"(1) False True \high (summin * 2 ^ n + lx) n < length treeList\ myass by force then show ?thesis by (metis True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \length treeList = length ?newlist\ add_leD1 both_member_options_from_chilf_to_complete_tree dp nat_1_add_1 pp) qed next case False hence pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist abcv by auto then show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ \length treeList = length ?newlist\ abcv both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp numerals(2)) qed qed then show ?thesis using abcv calculation(1) calculation(2) by fastforce qed qed ultimately show ?thesis by metis next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist summary)" have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_not_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist] using False \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x = mi\ dp mimapr nat_less_le by fastforce moreover have "both_member_options ?delsimp y \ x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- assume ssms: "both_member_options ?delsimp y " hence aaaa: "y = ?xn \ y = ?newma \ (both_member_options (?newlist ! (high y n)) (low y n) \ high y n < length ?newlist)" by (smt (z3) Suc_1 Suc_le_D \deg div 2 = n\ both_member_options_def dp membermima.simps(4) naive_member.simps(3)) show " x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" proof- have "y = ?xn \?thesis" using \x = mi\ minxnrel xnin by blast moreover have " y = ?newma \ ?thesis" proof- assume "y = ?newma" show ?thesis proof(cases "?xn = ma") case True hence aaa:"?newma =?h * 2 ^ (deg div 2) +the(vebt_maxt(?newlist ! ?h))" by metis have "?newlist ! ?h = ?newnode" using hprolist by blast obtain maxi where maxidef:"Some maxi = vebt_maxt(?newlist ! ?h)" by (metis False hprolist vebt_maxt.elims minNull.simps(1) minNull.simps(4)) have aa:"invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high ?xn n < length treeList\ inthall member_def) moreover hence ab:"maxi \ ?l \ both_member_options ?newnode maxi" by (metis "5.IH"(1) \high ?xn n < length treeList\ hprolist inthall maxbmo maxidef member_def) ultimately have ac:"maxi \ ?l \ both_member_options (treeList ! ?h) maxi" by (metis "5.IH"(1) \high ?xn n < length treeList\ inthall member_def) hence ad:"maxi < 2^n" using \invar_vebt (treeList ! high ?xn n) n\ member_bound valid_member_both_member_options by blast then show ?thesis using Suc_1 aaa \y = ?newma\ ac add_leD1 both_member_options_from_chilf_to_complete_tree dp option.sel high_inv low_inv maxidef plus_1_eq_Suc by (metis (no_types, lifting) True \Some lx = vebt_mint (treeList ! summin)\ \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ \x = mi\ leD member_bound mimapr mint_corr_help nat_add_left_cancel_le valid_member_both_member_options) next case False then show ?thesis by (metis \mi \ x \ x \ ma\ \x = mi\ \y = ?newma\ both_member_options_equiv_member leD vebt_maxt.simps(3) maxt_member mimapr tvalid) qed qed moreover have "(both_member_options (?newlist ! (high y n)) (low y n)\ high y n < length ?newlist) \ ?thesis" proof- assume assmy:"(both_member_options (?newlist ! (high y n)) (low y n)\ high y n < length ?newlist)" then show ?thesis proof(cases "high y n = ?h") case True hence "?newlist ! (high y n) = ?newnode" using hprolist by presburger have "invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high ?xn n < length treeList\ inthall member_def) then show ?thesis proof(cases "low y n= ?l") case True hence "y = ?xn" using "5.IH"(1) \high (summin * 2 ^ n + lx) n < length treeList\ \treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! high y n = vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)\ assmy by force then show ?thesis using calculation(1) by blast next case False hence "low y n \ ?l \ both_member_options (treeList ! ?h ) (low y n)" using assmy by (metis "5.IH"(1) "5.hyps"(2) \?newlist ! high y n = vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)\ \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ high_inv inthall member_bound member_def) then show ?thesis by (metis "5.hyps"(2) "5.hyps"(9) Suc_1 Suc_leD True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \mi \ x \ x \ ma\ \x = mi\ both_member_options_from_chilf_to_complete_tree dp leD mimapr) qed next case False hence "?newlist ! (high y n) = treeList !(high y n)" using "5.hyps"(2) \length treeList = length ?newlist\ assmy nothprolist by presburger hence "both_member_options (treeList !(high y n)) (low y n)" using assmy by presburger moreover have "x \ y" by (metis "5.hyps"(1) "5.hyps"(4) "5.hyps"(9) \invar_vebt (treeList ! summin) n\ \x < 2 ^ deg\ \x = mi\ calculation deg_not_0 exp_split_high_low(1) mimapr not_less_iff_gr_or_eq) moreover have "high y n < length ?newlist" using assmy by blast moreover hence "high y n < length treeList" using \length treeList = length ?newlist\ by presburger ultimately show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ both_member_options_from_chilf_to_complete_tree dp numerals(2)) qed qed ultimately show ?thesis using aaaa by fastforce qed qed moreover have "(x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y)\ both_member_options ?delsimp y" proof- assume assm: "x \ y \ both_member_options (Node (Some (mi, ma)) deg treeList summary) y" hence abcv:"y = mi \ y = ma \ ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n))" by (metis Suc_1 \deg div 2 = n\ add_leD1 both_member_options_from_complete_tree_to_child member_inv plus_1_eq_Suc tvalid valid_member_both_member_options) thus " both_member_options ?delsimp y" proof- have "y = mi \ ?thesis" using \x = mi\ assm by blast moreover have " y = ma \ ?thesis" by (smt (z3) Suc_1 Suc_le_D both_member_options_def dp membermima.simps(4)) moreover have " ( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n)) \ ?thesis" proof- assume myass: "( high y n < length treeList \ both_member_options (treeList ! (high y n)) (low y n)) " thus ?thesis proof(cases "high y n = ?h") case True then show ?thesis proof(cases "low y n = ?l") case True then show ?thesis by (smt (z3) "5.hyps"(3) "5.hyps"(4) Suc_1 \deg div 2 = n\ \length treeList = length (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)])\ add_Suc_right add_leD1 bit_split_inv both_member_options_def both_member_options_from_chilf_to_complete_tree dp membermima.simps(4) myass nth_list_update_neq plus_1_eq_Suc) next case False hence "low y n \ ?l" by simp hence pp:"?newlist ! ?h = ?newnode" using hprolist by blast hence "invar_vebt (treeList ! ?h) n" by (metis "5.IH"(1) \high ?xn n < length treeList\ inthall member_def) hence "both_member_options ?newnode (low y n)" by (metis "5.IH"(1) True \high ?xn n < length treeList\ \low y n \ low ?xn n\ in_set_member inthall myass) then show ?thesis by (metis One_nat_def Suc_leD True \deg div 2 = n\ \high (summin * 2 ^ n + lx) n < length treeList\ \length treeList = length ?newlist\ both_member_options_from_chilf_to_complete_tree dp numerals(2) pp) qed next case False have pp:"?newlist ! (high y n) = treeList ! (high y n)" using nothprolist[of "high y n"] False "5.hyps"(2) myass by presburger then show ?thesis by (metis One_nat_def Suc_leD \deg div 2 = n\ \length treeList = length ?newlist\ abcv both_member_options_from_chilf_to_complete_tree calculation(1) calculation(2) dp numerals(2)) qed qed then show ?thesis using abcv calculation(1) calculation(2) by fastforce qed qed ultimately show ?thesis by metis qed qed qed qed qed end corollary "invar_vebt t n \ both_member_options t x \ x \ y \ both_member_options (vebt_delete t y) x" using dele_bmo_cont_corr by blast corollary "invar_vebt t n \ both_member_options t x \ \ both_member_options (vebt_delete t x) x " by (simp add: dele_bmo_cont_corr) corollary "invar_vebt t n \ both_member_options (vebt_delete t y) x \ both_member_options t x \ x \ y" using dele_bmo_cont_corr by blast end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_DeleteCorrectness.thy b/thys/Van_Emde_Boas_Trees/VEBT_DeleteCorrectness.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_DeleteCorrectness.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_DeleteCorrectness.thy @@ -1,2265 +1,2265 @@ (*by Ammer*) theory VEBT_DeleteCorrectness imports VEBT_Delete begin context VEBT_internal begin subsection \Validness Preservation\ theorem delete_pres_valid: "invar_vebt t n \ invar_vebt (vebt_delete t x) n" proof(induction t n arbitrary: x rule: invar_vebt.induct) case (1 a b) then show ?case proof(cases x) case 0 then show ?thesis by (simp add: invar_vebt.intros(1)) next case (Suc prex) hence "x = Suc prex" by simp then show ?thesis proof(cases prex) case 0 then show ?thesis by (simp add: Suc invar_vebt.intros(1)) next case (Suc preprex) then show ?thesis by (simp add: \x = Suc prex\ invar_vebt.intros(1)) qed qed next case (2 treeList n summary m deg) then show ?case using invar_vebt.intros(2) by force next case (3 treeList n summary m deg) then show ?case using invar_vebt.intros(3) by auto next case (4 treeList n summary m deg mi ma) hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary m" and 2:"length treeList = 2^m" and 3:" deg = n+m" and 4: "(\ i < 2^m. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" and 8: "n = m" and 9: "deg div 2 = n" using "4" add_self_div_2 by auto hence 10: "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(4)[of treeList n summary m deg mi ma] by blast hence 11:"n \ 1 " and 12: " deg \ 2" by (metis "1" "8" "9" One_nat_def Suc_leI deg_not_0 div_greater_zero_iff)+ then show ?case proof(cases "(x < mi \ x > ma)") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, ma)) deg treeList summary)" using delt_out_of_range[of x mi ma deg treeList summary] using "1" "4.hyps"(3) "9" deg_not_0 div_greater_zero_iff by blast then show ?thesis by (simp add: "10") next case False hence inrg: "mi\ x \ x \ ma" by simp then show ?thesis proof(cases "x = mi \ x = ma") case True hence" (\ t \ set treeList. \ y. both_member_options t y)" using "5" by blast moreover have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node None deg treeList summary)" using del_single_cont[of x mi ma deg treeList summary] "1" "8" "9" True deg_not_0 div_greater_zero_iff by blast moreover have " (\ i. both_member_options summary i)" using "10" True mi_eq_ma_no_ch by blast ultimately show ?thesis using "0" "1" "2" "3" "4.hyps"(3) invar_vebt.intros(2) by force next case False hence "x \ mi \ x \ ma" by simp hence "mi \ ma \ x < 2^deg" by (metis "6" inrg le_antisym le_less_trans) hence "7b": "(\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma))" using "7" by fastforce hence "both_member_options (treeList ! (high ma n)) (low ma n)" using "1" "3" "6" "8" deg_not_0 exp_split_high_low(1) by blast hence yhelper:"both_member_options (treeList ! (high y n)) (low y n) \ high y n < 2^m \ mi < y \ y \ ma \ low y n < 2^n" for y by (simp add: "7b" low_def) then show ?thesis proof(cases "x \ mi") case True hence xnotmi: "x \ mi" by simp let ?h = "high x n" let ?l = "low x n" have hlbound:"?h < 2^m \ ?l < 2^n" using "1" "3" "8" \mi \ ma \ x < 2 ^ deg\ deg_not_0 exp_split_high_low(1) exp_split_high_low(2) by blast let ?newnode = "vebt_delete (treeList ! ?h) ?l" have "treeList ! ?h \ set treeList " by (metis "2" hlbound in_set_member inthall) hence nnvalid: "invar_vebt ?newnode n" by (simp add: "4.IH"(1)) let ?newlist = "treeList[?h:= ?newnode]" have hlist:"?newlist ! ?h = ?newnode" by (simp add: "2" hlbound) have nothlist:"i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp have allvalidinlist:"\ t \ set ?newlist. invar_vebt t n" proof fix t assume "t \ set ?newlist" then obtain i where "i< 2^m \ ?newlist ! i = t" by (metis "2" in_set_conv_nth length_list_update) show "invar_vebt t n" by (metis "0" "2" \i < 2 ^ m \ treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! i = t\ hlist nnvalid nth_list_update_neq nth_mem) qed have newlistlength: "length ?newlist = 2^m" using 2 by auto then show ?thesis proof(cases "minNull ?newnode") case True hence ninNullc: "minNull ?newnode" by simp let ?sn = "vebt_delete summary ?h" let ?newma= "(if x = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist ?sn)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_new_node_nil[of mi x ma deg ?h ?l ?newnode treeList ?sn summary ?newlist] hlbound 9 11 12 True 2 inrg xnotmi by simp have newsummvalid: "invar_vebt ?sn m" by (simp add: "4.IH"(2)) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options ?sn i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hlist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" by (simp add: min_Null_member ninNullc) hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" nnvalid valid_member_both_member_options by auto have 1003: "\ both_member_options ?sn i" using "1" True dele_bmo_cont_corr by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothlist by blast hence "both_member_options (?newlist ! i) y \ both_member_options ?sn i" for y proof- assume "both_member_options (?newlist ! i) y" hence "both_member_options summary i" using "1000" "4" \i < 2 ^ m\ by auto thus "both_member_options ?sn i" using "1" False dele_bmo_cont_corr by blast qed moreover have "both_member_options ?sn i \ \ y. both_member_options (?newlist ! i) y" proof- assume "both_member_options ?sn i " hence "both_member_options summary i" using "1" dele_bmo_cont_corr by auto thus " \ y. both_member_options (?newlist ! i) y" using "1000" "4" \i < 2 ^ m\ by presburger qed then show ?thesis using calculation by blast qed qed qed have 112:" (mi = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"mi = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "x = ma") case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence aa:"\ y. vebt_member ?sn y" using maxt_corr_help_empty newsummvalid set_vebt'_def by auto hence "\ y. both_member_options ?sn y" using newsummvalid valid_member_both_member_options by blast hence "t \ set ?newlist \ \y. both_member_options t y" for t proof- assume "t \ set ?newlist" then obtain i where "?newlist ! i = t \ i< 2^m" by (metis in_set_conv_nth newlistlength) thus " \y. both_member_options t y" using "111" \\y. both_member_options (vebt_delete summary (high x n)) y\ by blast qed then show ?thesis by blast next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 "2" by auto hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ empty_Collect_eq option.sel maxt_corr_help_empty set_vebt'_def valid_member_both_member_options by fastforce hence "maxs = high mi n \ both_member_options (?newlist ! maxs) (low mi n)" by (smt (z3) "9" True \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ aampt option.sel high_inv low_inv maxbmo member_bound mult.commute option.distinct(1) valid_member_both_member_options) hence False by (metis bb nat_less_le nothlist yhelper) then show ?thesis by simp qed next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ mi \ ?newma" proof(cases "x = ma") case True hence "x = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "6" by fastforce next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 "2" by auto hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ empty_Collect_eq maxt_corr_help_empty option_shift.elims set_vebt'_def valid_member_both_member_options) then show ?thesis by (smt (verit, best) "6" "9" \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (?newlist ! maxs) n\ bb option.sel high_inv less_le_trans low_inv maxbmo maxt_member member_bound mult.commute not_less_iff_gr_or_eq nothlist verit_comp_simplify1(3) yhelper) qed next case False then show ?thesis using "6" by auto qed have 115: "mi \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof assume "mi \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "x = ma" ) case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \mi \ (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma)\ \treeList ! high x n \ set treeList\ bit_split_inv dele_bmo_cont_corr hlist newmaassm nth_list_update_neq) next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 2 by auto hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ empty_Collect_eq maxt_corr_help_empty set_vebt'_def valid_member_both_member_options by (smt (z3) VEBT_Member.vebt_member.simps(2) \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ vebt_maxt.elims minNull.simps(1) min_Null_member valid_member_both_member_options) then show ?thesis by (smt (z3) "9" False True \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ option.sel high_inv low_inv maxbmo maxt_member member_bound mult.commute newmaassm) qed next case False then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \treeList ! high x n \ set treeList\ assumption bit_split_inv dele_bmo_cont_corr hlist newmaassm nothlist) qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "mi < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr hlist yassm by auto then show ?thesis by (simp add: assumption yassm yhelper) next case False then show ?thesis using assumption nothlist yassm yhelper by presburger qed moreover have "y \ ?newma" proof(cases "x = ma") case True hence "x= ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using \mi \ ?newma\ \x = ma\ by presburger next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 2 by auto hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis "2" "4.IH"(1) Collect_empty_eq bb both_member_options_equiv_member maxt_corr_help_empty nth_list_update_neq nth_mem option.exhaust set_vebt'_def) hence "maxs < 2^m \ maxi < 2^n" by (metis \invar_vebt (?newlist ! maxs) n\ bb maxt_member member_bound) hence "?newma = 2^n* maxs + maxi" by (smt (z3) "9" False True \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ option.sel) hence "low ?newma n = maxi \ high ?newma n = maxs" by (simp add: \maxs < 2 ^ m \ maxi < 2 ^ n\ high_inv low_inv mult.commute) hence "both_member_options (treeList ! (high y n)) (low y n)" by (metis "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm) hence hleqdraft:"high y n > maxs \ False" proof- assume "high y n > maxs" have "both_member_options summary (high y n)" using "1" "111" assumption dele_bmo_cont_corr yassm by blast moreover have "both_member_options ?sn (high y n)" using "111" assumption yassm by blast ultimately show False by (metis \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \maxs < high y n\ leD maxt_corr_help newsummvalid valid_member_both_member_options) qed hence hleqmaxs: "high y n \ maxs" by presburger then show ?thesis proof(cases "high y n = maxs") case True hence "low y n \ maxi" by (metis \Some maxi = vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs)\ \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ maxt_corr_help valid_member_both_member_options yassm) then show ?thesis by (smt (z3) True \(if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt ((?newlist) ! the maxs)) else ma) = 2 ^ n * maxs + maxi\ add_le_cancel_left bit_concat_def bit_split_inv mult.commute) next case False then show ?thesis by (metis \low (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt ((?newlist) ! the maxs)) else ma) n = maxi \ high (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt ((?newlist) ! the maxs)) else ma) n = maxs\ div_le_mono high_def hleqdraft le_neq_implies_less nat_le_linear) qed qed next case False then show ?thesis by (smt (z3) "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm yhelper) qed ultimately show " mi < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "mi \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (mi, ?newma)) deg ?newlist ?sn) deg" using invar_vebt.intros(4)[of ?newlist n ?sn m deg mi ?newma] using 3 allvalidinlist newlistlength newsummvalid "4.hyps"(3) 111 112 118 117 115 by fastforce show ?thesis using "116" dsimp by presburger next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if x = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist summary)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_newnode_not_nil[of mi x ma deg ?h ?l ?newnode treeList ?newlist summary] by (metis "12" "2" "9" False dual_order.eq_iff hlbound inrg order.not_eq_order_implies_strict xnotmi) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options summary i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hlist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" using nnvalid notemp valid_member_both_member_options by auto hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" notemp by presburger have 1003: "both_member_options summary i" using "0" "1000" "1002" "4" True \i < 2 ^ m\ \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by fastforce then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothlist by blast then show ?thesis using "4" \i < 2 ^ m\ by presburger qed qed qed have 112:" (mi = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"mi = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "x = ma") case True obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis False VEBT_Member.vebt_member.simps(2) hlist vebt_maxt.elims minNull.simps(1) nnvalid notemp valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hlist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by blast hence False by (metis "9" True \both_member_options ?newnode maxi\ \vebt_maxt (?newlist ! high x n) = Some maxi\ aampt option.sel high_inv hlbound low_inv member_bound nnvalid not_less_iff_gr_or_eq valid_member_both_member_options yhelper) then show ?thesis by blast next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ mi \ ?newma" proof(cases "x = ma") case True hence "x = ma" by simp obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis empty_Collect_eq hlist maxt_corr_help_empty nnvalid notemp option.exhaust set_vebt'_def valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hlist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by blast hence "maxi < 2^n" using \both_member_options?newnode maxi\ member_bound nnvalid valid_member_both_member_options by blast show ?thesis - by (smt (z3) "3" "9" Euclidean_Division.div_eq_0_iff True \both_member_options (treeList ! high x n) maxi\ \maxi < 2 ^ n\ \vebt_maxt (?newlist ! high x n) = Some maxi\ add.right_neutral div_exp_eq div_mult_self3 option.sel high_inv hlbound le_0_eq less_imp_le_nat low_inv power_not_zero rel_simps(28) yhelper) + by (smt (z3) "3" "9" div_eq_0_iff True \both_member_options (treeList ! high x n) maxi\ \maxi < 2 ^ n\ \vebt_maxt (?newlist ! high x n) = Some maxi\ add.right_neutral div_exp_eq div_mult_self3 option.sel high_inv hlbound le_0_eq less_imp_le_nat low_inv power_not_zero rel_simps(28) yhelper) next case False then show ?thesis using "6" by auto qed have 115: "mi \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof assume "mi \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "x = ma") case True obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hlist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast then show ?thesis by (smt (z3) "9" True \vebt_maxt (?newlist ! high x n) = Some maxi\ option.sel high_inv hlist low_inv maxt_member member_bound newmaassm nnvalid) next case False then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \treeList ! high x n \ set treeList\ assumption bit_split_inv dele_bmo_cont_corr hlist newmaassm nothlist) qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "mi < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr hlist yassm by auto then show ?thesis by (simp add: assumption yassm yhelper) next case False then show ?thesis using assumption nothlist yassm yhelper by presburger qed moreover have "y \ ?newma" proof(cases "x = ma") case True hence "x= ma" by simp obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hlist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast have "high y n \ ?h" by (metis "7b" True assumption div_le_mono high_def nothlist yassm) then show ?thesis proof(cases "high y n = ?h") case True have "low y n > maxi \ False" by (metis True \vebt_maxt (?newlist ! ?h) = Some maxi\ hlist leD maxt_corr_help nnvalid valid_member_both_member_options yassm) then show ?thesis by (smt (z3) "9" True \vebt_maxt (?newlist ! ?h) = Some maxi\ \x = ma\ add_le_cancel_left div_mult_mod_eq option.sel high_def low_def nat_le_linear nat_less_le) next case False then show ?thesis by (smt (z3) "9" True \both_member_options (?newlist ! high x n) maxi\ \high y n \ high x n\ \vebt_maxt (?newlist ! high x n) = Some maxi\ div_le_mono option.sel high_def high_inv hlist le_antisym member_bound nat_le_linear nnvalid valid_member_both_member_options) qed next case False then show ?thesis by (smt (z3) "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm yhelper) qed ultimately show " mi < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "mi \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (mi, ?newma)) deg ?newlist summary) deg" using invar_vebt.intros(4)[of ?newlist n summary m deg mi ?newma] allvalidinlist 1 newlistlength 8 3 111 112 117 118 115 by fastforce then show ?thesis using dsimp by presburger qed next case False hence xmi:"x = mi" by simp have "both_member_options summary (high ma n)" using "1" "3" "4" "4.hyps"(3) "6" \both_member_options (treeList ! high ma n) (low ma n)\ deg_not_0 exp_split_high_low(1) by blast hence "vebt_member summary (high ma n)" using "4.hyps"(1) valid_member_both_member_options by blast obtain summin where "Some summin = vebt_mint summary" by (metis "4.hyps"(1) \vebt_member summary (high ma n)\ empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def) hence "\ z . both_member_options (treeList ! summin) z" by (metis "4.hyps"(1) "4.hyps"(5) both_member_options_equiv_member member_bound mint_member) moreover have "invar_vebt (treeList ! summin) n" by (metis "0" "1" "2" \Some summin = vebt_mint summary\ member_bound mint_member nth_mem) ultimately obtain lx where "Some lx = vebt_mint (treeList ! summin)" by (metis empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) let ?xn = "summin*2^n + lx" have "?xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)" by (metis False \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \deg div 2 = n\ option.sel) have "vebt_member (treeList ! summin) lx" using \Some lx = vebt_mint (treeList ! summin)\ \invar_vebt (treeList ! summin) n\ mint_member by auto moreover have "summin < 2^m" by (metis "4.hyps"(1) \Some summin = vebt_mint summary\ member_bound mint_member) ultimately have xnin: "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" by (metis "12" "2" "9" \invar_vebt (treeList ! summin) n\ add_leD1 both_member_options_equiv_member both_member_options_from_chilf_to_complete_tree high_inv low_inv member_bound numeral_2_eq_2 plus_1_eq_Suc) let ?h ="high ?xn n" let ?l = "low ?xn n" have "?xn < 2^deg" - by (smt (verit, ccfv_SIG) "4.hyps"(1) "4.hyps"(4) Euclidean_Division.div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) + by (smt (verit, ccfv_SIG) "4.hyps"(1) "4.hyps"(4) div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) hence "?h < length treeList" using "4.hyps"(2) "4.hyps"(3) "4.hyps"(4) \invar_vebt (treeList ! summin) n\ deg_not_0 exp_split_high_low(1) by metis let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:= ?newnode]" have "length treeList = length ?newlist" by simp hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high (summin * 2 ^ n + lx) n < length treeList\ nth_list_update_eq) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp have hlbound:"?h < 2^m \ ?l < 2^n" using "1" "2" "3" "8" \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx < 2 ^ deg\ deg_not_0 exp_split_high_low(2) by presburger hence nnvalid: "invar_vebt ?newnode n" by (metis "4.IH"(1) \high (summin * 2 ^ n + lx) n < length treeList\ inthall member_def) have allvalidinlist:"\ t \ set ?newlist. invar_vebt t n" proof fix t assume "t \ set ?newlist" then obtain i where "i < 2^m \ ?newlist ! i = t" by (metis "2" \length treeList = length (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)])\ in_set_conv_nth) then show "invar_vebt t n" by (metis "0" "2" hprolist nnvalid nth_list_update_neq nth_mem) qed have newlistlength: "length ?newlist = 2^m" by (simp add: "2") then show ?thesis proof(cases "minNull ?newnode") case True hence ninNullc: "minNull ?newnode" by simp let ?sn = "vebt_delete summary ?h" let ?newma= "(if ?xn = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist ?sn)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist ?sn] by (metis "12" "9" \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x = mi\ \x \ mi \ x \ ma\ inrg nat_less_le ninNullc) have newsummvalid: "invar_vebt ?sn m" by (simp add: "4.IH"(2)) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options ?sn i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hprolist by fastforce hence 1001:"\ x. vebt_member (?newlist ! i) x" by (simp add: min_Null_member ninNullc) hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" nnvalid valid_member_both_member_options by auto have 1003: "\ both_member_options ?sn i" using "1" True dele_bmo_cont_corr by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothprolist by blast hence "both_member_options (?newlist ! i) y \ both_member_options ?sn i" for y proof- assume "both_member_options (?newlist ! i) y" hence "both_member_options summary i" using "1000" "4" \i < 2 ^ m\ by auto thus "both_member_options ?sn i" using "1" False dele_bmo_cont_corr by blast qed moreover have "both_member_options ?sn i \ \ y. both_member_options (?newlist ! i) y" proof- assume "both_member_options ?sn i " hence "both_member_options summary i" using "1" dele_bmo_cont_corr by auto thus " \ y. both_member_options (?newlist ! i) y" using "1000" "4" \i < 2 ^ m\ by presburger qed then show ?thesis using calculation by blast qed qed qed have 112:" (?xn = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"?xn = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "?xn = ma") case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence aa:"\ y. vebt_member ?sn y" using maxt_corr_help_empty newsummvalid set_vebt'_def by auto hence "\ y. both_member_options ?sn y" using newsummvalid valid_member_both_member_options by blast hence "t \ set ?newlist \ \y. both_member_options t y" for t proof- assume "t \ set ?newlist" then obtain i where "?newlist ! i = t \ i< 2^m" by (metis "2" \length treeList = length (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)])\ in_set_conv_nth) thus " \y. both_member_options t y" using "111" \\y. both_member_options (vebt_delete summary (high (summin * 2 ^ n + lx) n)) y\ by blast qed then show ?thesis by blast next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs) n\ maxt_corr_help_empty set_vebt'_def valid_member_both_member_options by fastforce hence "maxs = high ?xn n \ both_member_options (?newlist ! maxs) (low ?xn n)" by (smt (z3) "9" False True \Some maxs = vebt_maxt (vebt_delete summary ?h)\ \invar_vebt (?newlist ! maxs) n\ aampt option.sel high_inv low_inv maxbmo maxt_member member_bound mult.commute) hence False using bb by blast then show ?thesis by simp qed next case False hence "?xn \ ?newma" by simp hence False using aampt by simp then show ?thesis by blast qed qed have 114: "?newma < 2^deg \ ?xn \ ?newma" proof(cases "?xn = ma") case True hence "?xn = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "4.hyps"(8) \?xn = ma\ by force next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs) n\ empty_Collect_eq maxt_corr_help_empty not_Some_eq set_vebt'_def valid_member_both_member_options by fastforce hence abc:"?newma = 2^n * maxs + maxi" by (smt (z3) "9" True \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ option.sel not_None_eq) have abd:"maxi < 2^n" by (metis \Some maxi = vebt_maxt (?newlist ! maxs)\ \invar_vebt (?newlist ! maxs) n\ maxt_member member_bound) have "high ?xn n \ maxs" using "1" \Some summin = vebt_mint summary\ \both_member_options summary maxs\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound mint_corr_help valid_member_both_member_options by force then show ?thesis proof(cases "high ?xn n = maxs") case True then show ?thesis using bb by fastforce next case False hence "high ?xn n < maxs" by (simp add: \high (summin * 2 ^ n + lx) n \ maxs\ order.not_eq_order_implies_strict) hence "?newma < 2^deg"using "1" "10" "9" True \both_member_options summary maxs\ \mi \ ma \ x < 2 ^ deg\ equals0D leD maxt_corr_help maxt_corr_help_empty mem_Collect_eq summaxma set_vebt'_def valid_member_both_member_options by (metis option.exhaust_sel) moreover have "high ?xn n < high ?newma n" by (smt (z3) "9" True \Some maxi = vebt_maxt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ \high (summin * 2 ^ n + lx) n < maxs\ abd option.sel high_inv mult.commute option.discI) ultimately show ?thesis by (metis div_le_mono high_def linear not_less) qed qed next case False then show ?thesis by (smt (z3) "12" "4.hyps"(7) "4.hyps"(8) "9" both_member_options_from_complete_tree_to_child dual_order.trans hlbound one_le_numeral xnin yhelper) qed have 115: "?xn \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof assume assumption0:"?xn \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "?xn = ma" ) case True hence bb:"?xn = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence "?newma = ?xn" using assumption Let_def bb by simp hence False using assumption0 by simp then show ?thesis by simp next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs) n\ equals0D maxt_corr_help_empty mem_Collect_eq set_vebt'_def valid_member_both_member_options by (metis option.collapse) then show ?thesis using "1" "10" "9" True \Some summin = vebt_mint summary\ \both_member_options summary maxs\ \vebt_member (treeList ! summin) lx\ \mi \ ma \ x < 2 ^ deg\ \invar_vebt (treeList ! summin) n\ bb equals0D high_inv maxt_corr_help maxt_corr_help_empty mem_Collect_eq member_bound mint_corr_help nat_less_le summaxma set_vebt'_def valid_member_both_member_options verit_comp_simplify1(3) by (metis option.collapse) qed next case False hence ccc:"?newma = ma" by simp then show ?thesis proof(cases "?xn = ma") case True hence "?xn = ?newma" using False by blast hence False using False by auto then show ?thesis by simp next case False hence "both_member_options (?newlist ! high ma n) (low ma n)" by (metis "1" \both_member_options (treeList ! high ma n) (low ma n)\ \vebt_member (treeList ! summin) lx\ \vebt_member summary (high ma n)\ \invar_vebt (treeList ! summin) n\ bit_split_inv dele_bmo_cont_corr high_inv hprolist member_bound nothprolist) moreover have "high ma n = i \ low ma n = low ?newma n" using ccc newmaassm by simp ultimately show ?thesis by simp qed qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "?xn < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by auto then show ?thesis using True hprolist min_Null_member ninNullc nnvalid valid_member_both_member_options yassm by fastforce next case False hence "i \ ?h \ False" by (metis "1" "111" \Some summin = vebt_mint summary\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption dele_bmo_cont_corr high_inv le_antisym member_bound mint_corr_help valid_member_both_member_options yassm) hence "i > ?h" using leI by blast then show ?thesis by (metis div_le_mono high_def not_less yassm) qed moreover have "y \ ?newma" proof(cases "?xn = ma") case True hence "?xn= ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "1" "111" assumption dele_bmo_cont_corr nothprolist yassm yhelper by auto next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption calculation dele_bmo_cont_corr high_inv hprolist leD member_bound nth_list_update_neq yassm yhelper) hence "maxs < 2^m \ maxi < 2^n" by (metis \invar_vebt (?newlist ! maxs) n\ bb maxt_member member_bound) hence "?newma = 2^n* maxs + maxi" by (smt (z3) "9" False True \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high ?xn n))\ option.sel) hence "low ?newma n = maxi \ high ?newma n = maxs" by (simp add: \maxs < 2 ^ m \ maxi < 2 ^ n\ high_inv low_inv mult.commute) hence "both_member_options (treeList ! (high y n)) (low y n)" by (metis "1" "111" assumption dele_bmo_cont_corr nothprolist yassm) hence hleqdraft:"high y n > maxs \ False" proof- assume "high y n > maxs" have "both_member_options summary (high y n)" using "1" "111" assumption dele_bmo_cont_corr yassm by blast moreover have "both_member_options ?sn (high y n)" using "111" assumption yassm by blast ultimately show False using True \both_member_options (treeList ! high y n) (low y n)\ \summin * 2 ^ n + lx < y\ assumption leD yassm yhelper by blast qed hence hleqmaxs: "high y n \ maxs" by presburger then show ?thesis using \both_member_options (treeList ! high y n) (low y n)\ assumption calculation dual_order.strict_trans1 yassm yhelper by auto qed next case False then show ?thesis by (smt (z3) \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption dele_bmo_cont_corr high_inv hprolist member_bound nothprolist yassm yhelper) qed ultimately show " ?xn < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "?xn \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (?xn, ?newma)) deg ?newlist ?sn) deg" using invar_vebt.intros(4)[of ?newlist n ?sn m deg ?xn ?newma] using 3 allvalidinlist newlistlength newsummvalid "4.hyps"(3) 111 112 118 117 115 by fastforce show ?thesis using "116" dsimp by presburger next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist summary)" have dsimp:"vebt_delete (Node (Some (x, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_not_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist] "12" "2" "9" False dual_order.eq_iff hlbound inrg order.not_eq_order_implies_strict xmi by (metis \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x \ mi \ x \ ma\) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options summary i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hprolist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" using nnvalid notemp valid_member_both_member_options by auto hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" notemp by presburger have 1003: "both_member_options summary i" using "4" True \\z. both_member_options (treeList ! summin) z\ \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothprolist by blast then show ?thesis using "4" \i < 2 ^ m\ by presburger qed qed qed have 112:" (?xn = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"?xn = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "?xn = ma") case True obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq False hprolist maxt_corr_help_empty nnvalid not_None_eq not_min_Null_member set_vebt'_def valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hprolist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv member_bound by force hence False by (metis "9" \both_member_options (vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)) maxi\ \vebt_maxt (?newlist ! ?h) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ aampt add_diff_cancel_left' dele_bmo_cont_corr option.sel high_inv low_inv member_bound) then show ?thesis by blast next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ ?xn \ ?newma" proof(cases "?xn = ma") case True hence "?xn = ma" by simp obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis "111" "2" "4" Collect_empty_eq True \both_member_options (treeList ! high ma n) (low ma n)\ \high (summin * 2 ^ n + lx) n < length treeList\ hprolist maxt_corr_help_empty nnvalid not_None_eq set_vebt'_def valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hprolist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv member_bound by force hence "maxi < 2^n" using \both_member_options?newnode maxi\ member_bound nnvalid valid_member_both_member_options by blast show ?thesis - by (smt (verit, ccfv_threshold) "3" "9" Euclidean_Division.div_eq_0_iff True \Some lx = vebt_mint (treeList ! summin)\ \both_member_options (treeList ! high (summin * 2 ^ n + lx) n) maxi\ \vebt_maxt (?newlist ! high (summin * 2 ^ n + lx) n) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add.right_neutral add_left_mono div_mult2_eq div_mult_self3 option.sel high_inv hlbound le_0_eq member_bound mint_corr_help power_add power_not_zero rel_simps(28) valid_member_both_member_options) + by (smt (verit, ccfv_threshold) "3" "9" div_eq_0_iff True \Some lx = vebt_mint (treeList ! summin)\ \both_member_options (treeList ! high (summin * 2 ^ n + lx) n) maxi\ \vebt_maxt (?newlist ! high (summin * 2 ^ n + lx) n) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add.right_neutral add_left_mono div_mult2_eq div_mult_self3 option.sel high_inv hlbound le_0_eq member_bound mint_corr_help power_add power_not_zero rel_simps(28) valid_member_both_member_options) next case False then show ?thesis using "10" "4.hyps"(8) maxt_corr_help valid_member_both_member_options xnin by force qed have 115: "?xn \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof assume xnmassm:"?xn \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "?xn = ma") case True obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hprolist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast then show ?thesis by (smt (z3) "2" "9" True \Some lx = vebt_mint (treeList ! summin)\ \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add_left_mono dele_bmo_cont_corr eq_iff high_inv hprolist low_inv member_bound mint_corr_help valid_member_both_member_options yhelper) next case False hence abcd:"?newma = ma" by simp then show ?thesis proof(cases "high ma n = ?h") case True hence "?newlist ! high ma n = ?newnode" using hprolist by presburger then show ?thesis by (smt (z3) False True \both_member_options (treeList ! high ma n) (low ma n)\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ bit_split_inv dele_bmo_cont_corr high_inv member_bound newmaassm) next case False hence "?newlist ! high ma n = treeList ! high ma n" using "1" \vebt_member summary (high ma n)\ member_bound nothprolist by blast moreover hence "both_member_options (treeList ! high ma n) (low ma n)" using \both_member_options (treeList ! high ma n) (low ma n)\ by blast ultimately show ?thesis using abcd newmaassm by simp qed qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "?xn < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by force moreover have "vebt_mint (treeList ! i) = Some (low ?xn n)" using True \Some lx = vebt_mint (treeList ! summin)\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv low_inv member_bound by presburger moreover hence "low y n \ low ?xn n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ calculation(1) high_inv member_bound mint_corr_help valid_member_both_member_options by auto moreover have "low y n \ low ?xn n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by auto ultimately have "low y n > low ?xn n" by simp show ?thesis by (metis True \low (summin * 2 ^ n + lx) n \ low y n\ \low y n \ low (summin * 2 ^ n + lx) n\ bit_concat_def bit_split_inv leD linorder_neqE_nat nat_add_left_cancel_less yassm) next case False have "Some (high ?xn n) = vebt_mint summary" using \Some summin = vebt_mint summary\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger moreover hence "high y n \ high ?xn n" by (metis "1" "111" assumption mint_corr_help valid_member_both_member_options yassm) ultimately show ?thesis by (metis False div_le_mono high_def leI le_antisym yassm) qed moreover have "y \ ?newma" by (smt (z3) \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption calculation dele_bmo_cont_corr high_inv hprolist leD member_bound nothprolist yassm yhelper) ultimately show " ?xn < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "?xn \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (?xn, ?newma)) deg ?newlist summary) deg" using invar_vebt.intros(4)[of ?newlist n summary m deg ?xn ?newma] allvalidinlist 1 newlistlength 8 3 111 112 117 118 115 by fastforce hence "invar_vebt (?delsimp) deg" by simp moreover obtain delsimp where 118:"delsimp = ?delsimp" by simp ultimately have 119:"invar_vebt delsimp deg" by simp have "vebt_delete (Node (Some (x, ma)) deg treeList summary) x = delsimp" using dsimp 118 by simp hence "delsimp = vebt_delete (Node (Some (x, ma)) deg treeList summary) x" by simp then show ?thesis using 119 using xmi by auto qed qed qed qed next case (5 treeList n summary m deg mi ma) hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary m" and 2:"length treeList = 2^m" and 3:" deg = n+m" and 4: "(\ i < 2^m. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" and 8: "Suc n = m" and 9: "deg div 2 = n" using "5" add_self_div_2 by auto hence 10: "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(5)[of treeList n summary m deg mi ma] by blast hence 11:"n \ 1 " and 12: " deg \ 2" by (metis "0" "2" "9" One_nat_def deg_not_0 div_greater_zero_iff le_0_eq numeral_2_eq_2 set_n_deg_not_0)+ then show ?case proof(cases "(x < mi \ x > ma)") case True hence "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node (Some (mi, ma)) deg treeList summary)" using delt_out_of_range[of x mi ma deg treeList summary] using "12" by fastforce then show ?thesis by (simp add: "10") next case False hence inrg: "mi\ x \ x \ ma" by simp then show ?thesis proof(cases "x = mi \ x = ma") case True hence" (\ t \ set treeList. \ y. both_member_options t y)" using "5" by blast moreover have "vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = (Node None deg treeList summary)" using del_single_cont[of x mi ma deg treeList summary] "1" "8" "9" True deg_not_0 div_greater_zero_iff "12" by fastforce moreover have " (\ i. both_member_options summary i)" using "10" True mi_eq_ma_no_ch by blast ultimately show ?thesis using "0" "1" "2" "3" "8" invar_vebt.intros(3) by force next case False hence "x \ mi \ x \ ma" by simp hence "mi \ ma \ x < 2^deg" by (metis "6" inrg le_antisym le_less_trans) hence "7b": "(\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma))" using "7" by fastforce hence "both_member_options (treeList ! (high ma n)) (low ma n)" by (metis "1" "12" "3" "6" "9" deg_not_0 div_greater_zero_iff exp_split_high_low(1) zero_less_numeral) hence yhelper:"both_member_options (treeList ! (high y n)) (low y n) \ high y n < 2^m \ mi < y \ y \ ma \ low y n < 2^n" for y by (simp add: "7b" low_def) then show ?thesis proof(cases "x \ mi") case True hence xnotmi: "x \ mi" by simp let ?h = "high x n" let ?l = "low x n" have hlbound:"?h < 2^m \ ?l < 2^n" by (metis "1" "11" "3" One_nat_def \mi \ ma \ x < 2 ^ deg\ deg_not_0 dual_order.strict_trans1 exp_split_high_low(1) exp_split_high_low(2) zero_less_Suc) let ?newnode = "vebt_delete (treeList ! ?h) ?l" have "treeList ! ?h \ set treeList " by (metis "2" hlbound in_set_member inthall) hence nnvalid: "invar_vebt ?newnode n" by (simp add: "5.IH"(1)) let ?newlist = "treeList[?h:= ?newnode]" have hlist:"?newlist ! ?h = ?newnode" by (simp add: "2" hlbound) have nothlist:"i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by simp have allvalidinlist:"\ t \ set ?newlist. invar_vebt t n" proof fix t assume "t \ set ?newlist" then obtain i where "i< 2^m \ ?newlist ! i = t" by (metis "2" in_set_conv_nth length_list_update) then show "invar_vebt t n" by (metis "0" "2" hlist nnvalid nth_list_update_neq nth_mem) qed have newlistlength: "length ?newlist = 2^m" by (simp add: "2") then show ?thesis proof(cases "minNull ?newnode") case True hence ninNullc: "minNull ?newnode" by simp let ?sn = "vebt_delete summary ?h" let ?newma= "(if x = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then mi else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist ?sn)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_new_node_nil[of mi x ma deg ?h ?l ?newnode treeList ?sn summary ?newlist] hlbound 9 11 12 True 2 inrg xnotmi by simp have newsummvalid: "invar_vebt ?sn m" by (simp add: "5.IH"(2)) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options ?sn i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hlist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" by (simp add: min_Null_member ninNullc) hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" nnvalid valid_member_both_member_options by auto have 1003: "\ both_member_options ?sn i" using "1" True dele_bmo_cont_corr by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothlist by blast hence "both_member_options (?newlist ! i) y \ both_member_options ?sn i" for y by (metis "1" "4" False \i < 2 ^ m\ dele_bmo_cont_corr) moreover have "both_member_options ?sn i \ \ y. both_member_options (?newlist ! i) y" using "1" "4" \i < 2 ^ m\ dele_bmo_cont_corr by force then show ?thesis using calculation by blast qed qed qed have 112:" (mi = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"mi = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "x = ma") case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence aa:"\ y. vebt_member ?sn y" using maxt_corr_help_empty newsummvalid set_vebt'_def by auto hence "\ y. both_member_options ?sn y" using newsummvalid valid_member_both_member_options by blast hence "t \ set ?newlist \ \y. both_member_options t y" for t proof- assume "t \ set ?newlist" then obtain i where "?newlist ! i = t \ i< 2^m" by (metis in_set_conv_nth newlistlength) thus " \y. both_member_options t y" using "111" \\y. both_member_options (vebt_delete summary (high x n)) y\ by blast qed then show ?thesis by blast next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (metis allvalidinlist newlistlength nth_mem) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis Collect_empty_eq_bot \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ bb bot_empty_eq equals0D maxt_corr_help_empty nth_list_update_neq option_shift.elims set_vebt'_def valid_member_both_member_options) hence "maxs = high mi n \ both_member_options (?newlist ! maxs) (low mi n)" by (smt (z3) "9" False True \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (?newlist ! maxs) n\ aampt option.sel high_inv low_inv maxbmo maxt_member member_bound mult.commute) hence False by (metis bb nat_less_le nothlist yhelper) then show ?thesis by simp qed next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ mi \ ?newma" proof(cases "x = ma") case True hence "x = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "6" by fastforce next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (metis allvalidinlist newlistlength nth_mem) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (smt (z3) VEBT_Member.vebt_member.simps(2) \invar_vebt (?newlist ! maxs) n\ vebt_maxt.elims minNull.simps(1) min_Null_member valid_member_both_member_options) then show ?thesis by (smt (verit, best) "6" "9" \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (?newlist ! maxs) n\ bb option.sel high_inv less_le_trans low_inv maxbmo maxt_member member_bound mult.commute not_less_iff_gr_or_eq nothlist verit_comp_simplify1(3) yhelper) qed next case False then show ?thesis using "6" by auto qed have 115: "mi \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof assume "mi \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "x = ma" ) case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \mi \ (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) else ma)\ \treeList ! high x n \ set treeList\ assumption bit_split_inv dele_bmo_cont_corr hlist newmaassm nothlist) next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 "2" by auto hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (smt (z3) VEBT_Member.vebt_member.simps(2) \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ vebt_maxt.elims minNull.simps(1) min_Null_member valid_member_both_member_options) then show ?thesis by (smt (z3) "9" True \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ option.sel high_inv low_inv maxbmo maxt_member member_bound mult.commute newmaassm option.distinct(1)) qed next case False then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \treeList ! high x n \ set treeList\ assumption bit_split_inv dele_bmo_cont_corr hlist newmaassm nothlist) qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "mi < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr hlist yassm by auto then show ?thesis by (simp add: assumption yassm yhelper) next case False then show ?thesis using assumption nothlist yassm yhelper by presburger qed moreover have "y \ ?newma" proof(cases "x = ma") case True hence "x= ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using \mi \ ?newma\ \x = ma\ by presburger next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 "2" by fastforce hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothlist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ equals0D maxt_corr_help_empty mem_Collect_eq option_shift.elims set_vebt'_def valid_member_both_member_options) hence "maxs < 2^m \ maxi < 2^n" by (metis \invar_vebt (?newlist ! maxs) n\ bb maxt_member member_bound) hence "?newma = 2^n* maxs + maxi" by (smt (z3) "9" False True \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ option.sel) hence "low ?newma n = maxi \ high ?newma n = maxs" by (simp add: \maxs < 2 ^ m \ maxi < 2 ^ n\ high_inv low_inv mult.commute) hence "both_member_options (treeList ! (high y n)) (low y n)" by (metis "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm) hence hleqdraft:"high y n > maxs \ False" proof- assume "high y n > maxs" have "both_member_options summary (high y n)" using "1" "111" assumption dele_bmo_cont_corr yassm by blast moreover have "both_member_options ?sn (high y n)" using "111" assumption yassm by blast ultimately show False by (metis \Some maxs = vebt_maxt (vebt_delete summary (high x n))\ \maxs < high y n\ leD maxt_corr_help newsummvalid valid_member_both_member_options) qed hence hleqmaxs: "high y n \ maxs" by presburger then show ?thesis proof(cases "high y n = maxs") case True hence "low y n \ maxi" by (metis \Some maxi = vebt_maxt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs)\ \invar_vebt (treeList[high x n := vebt_delete (treeList ! high x n) (low x n)] ! maxs) n\ maxt_corr_help valid_member_both_member_options yassm) then show ?thesis by (smt (z3) True \(if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma) = 2 ^ n * maxs + maxi\ add_le_cancel_left bit_concat_def bit_split_inv mult.commute) next case False then show ?thesis by (smt (z3) \low (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma) n = maxi \ high (if x = ma then let maxs = vebt_maxt (vebt_delete summary (high x n)) in if maxs = None then mi else 2 ^ (deg div 2) * the maxs + the (vebt_maxt (treeList [high x n := vebt_delete (treeList ! high x n) (low x n)] ! the maxs)) else ma) n = maxs\ div_le_mono high_def hleqmaxs le_antisym nat_le_linear) qed qed next case False then show ?thesis by (smt (z3) "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm yhelper) qed ultimately show " mi < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "mi \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (mi, ?newma)) deg ?newlist ?sn) deg" using invar_vebt.intros(5)[of ?newlist n ?sn m deg mi ?newma] using 3 allvalidinlist newlistlength newsummvalid "5.hyps"(3) 111 112 118 117 115 by fastforce show ?thesis using "116" dsimp by presburger next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if x = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (mi, ?newma)) deg ?newlist summary)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_not_mi_newnode_not_nil[of mi x ma deg ?h ?l ?newnode treeList ?newlist summary] by (metis "12" "2" "9" False dual_order.eq_iff hlbound inrg order.not_eq_order_implies_strict xnotmi) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options summary i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hlist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" using nnvalid notemp valid_member_both_member_options by auto hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" notemp by presburger have 1003: "both_member_options summary i" using "0" "1000" "1002" "4" True \i < 2 ^ m\ \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by fastforce then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothlist by blast then show ?thesis using "4" \i < 2 ^ m\ by presburger qed qed qed have 112:" (mi = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"mi = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "x = ma") case True obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis False VEBT_Member.vebt_member.simps(2) hlist vebt_maxt.elims minNull.simps(1) nnvalid notemp valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hlist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by blast hence False by (metis "9" True \both_member_options ?newnode maxi\ \vebt_maxt ( ?newlist ! high x n) = Some maxi\ aampt option.sel high_inv hlbound low_inv member_bound nnvalid not_less_iff_gr_or_eq valid_member_both_member_options yhelper) then show ?thesis by blast next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ mi \ ?newma" proof(cases "x = ma") case True hence "x = ma" by simp obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis False VEBT_Member.vebt_member.simps(2) hlist vebt_maxt.elims minNull.simps(1) nnvalid notemp valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hlist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr by blast hence "maxi < 2^n" using \both_member_options?newnode maxi\ member_bound nnvalid valid_member_both_member_options by blast show ?thesis - by (smt (z3) "3" "9" Euclidean_Division.div_eq_0_iff True \both_member_options (treeList ! high x n) maxi\ \maxi < 2 ^ n\ \vebt_maxt ( ?newlist ! high x n) = Some maxi\ add.right_neutral div_exp_eq div_mult_self3 option.sel high_inv hlbound le_0_eq less_imp_le_nat low_inv power_not_zero rel_simps(28) yhelper) + by (smt (z3) "3" "9" div_eq_0_iff True \both_member_options (treeList ! high x n) maxi\ \maxi < 2 ^ n\ \vebt_maxt ( ?newlist ! high x n) = Some maxi\ add.right_neutral div_exp_eq div_mult_self3 option.sel high_inv hlbound le_0_eq less_imp_le_nat low_inv power_not_zero rel_simps(28) yhelper) next case False then show ?thesis using "6" by auto qed have 115: "mi \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof assume "mi \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "x = ma") case True obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hlist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast then show ?thesis by (smt (z3) "9" True \vebt_maxt (?newlist ! high x n) = Some maxi\ option.sel high_inv hlist low_inv maxt_member member_bound newmaassm nnvalid) next case False then show ?thesis by (smt (z3) "0" \both_member_options (treeList ! high ma n) (low ma n)\ \treeList ! high x n \ set treeList\ assumption bit_split_inv dele_bmo_cont_corr hlist newmaassm nothlist) qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ mi < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "mi < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using "0" \treeList ! high x n \ set treeList\ dele_bmo_cont_corr hlist yassm by auto then show ?thesis by (simp add: assumption yassm yhelper) next case False then show ?thesis using assumption nothlist yassm yhelper by presburger qed moreover have "y \ ?newma" proof(cases "x = ma") case True hence "x= ma" by simp obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hlist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast have "high y n \ ?h" by (metis "7b" True assumption div_le_mono high_def nothlist yassm) then show ?thesis proof(cases "high y n = ?h") case True have "low y n > maxi \ False" by (metis True \vebt_maxt (?newlist ! ?h) = Some maxi\ hlist leD maxt_corr_help nnvalid valid_member_both_member_options yassm) then show ?thesis by (smt (z3) "9" True \vebt_maxt (?newlist ! ?h) = Some maxi\ \x = ma\ add_le_cancel_left div_mult_mod_eq option.sel high_def low_def nat_le_linear nat_less_le) next case False then show ?thesis by (smt (z3) "9" True \both_member_options (?newlist ! high x n) maxi\ \high y n \ high x n\ \vebt_maxt (?newlist ! high x n) = Some maxi\ div_le_mono option.sel high_def high_inv hlist le_antisym member_bound nat_le_linear nnvalid valid_member_both_member_options) qed next case False then show ?thesis by (smt (z3) "0" \treeList ! high x n \ set treeList\ assumption dele_bmo_cont_corr hlist nothlist yassm yhelper) qed ultimately show " mi < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "mi \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (mi, ?newma)) deg ?newlist summary) deg" using invar_vebt.intros(5)[of ?newlist n summary m deg mi ?newma] allvalidinlist 1 newlistlength 8 3 111 112 117 118 115 by fastforce then show ?thesis using dsimp by presburger qed next case False hence xmi:"x = mi" by simp have "both_member_options summary (high ma n)" by (metis "1" "11" "3" "4" "6" One_nat_def Suc_le_eq \both_member_options (treeList ! high ma n) (low ma n)\ deg_not_0 exp_split_high_low(1)) hence "vebt_member summary (high ma n)" using "5.hyps"(1) valid_member_both_member_options by blast obtain summin where "Some summin = vebt_mint summary" by (metis "5.hyps"(1) \vebt_member summary (high ma n)\ empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def) hence "\ z . both_member_options (treeList ! summin) z" by (metis "5.hyps"(1) "5.hyps"(5) both_member_options_equiv_member member_bound mint_member) moreover have "invar_vebt (treeList ! summin) n" by (metis "0" "1" "2" \Some summin = vebt_mint summary\ member_bound mint_member nth_mem) ultimately obtain lx where "Some lx = vebt_mint (treeList ! summin)" by (metis empty_Collect_eq mint_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) let ?xn = "summin*2^n + lx" have "?xn = (if x = mi then the (vebt_mint summary) * 2^(deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)" by (metis False \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \deg div 2 = n\ option.sel) have "vebt_member (treeList ! summin) lx" using \Some lx = vebt_mint (treeList ! summin)\ \invar_vebt (treeList ! summin) n\ mint_member by auto moreover have "summin < 2^m" by (metis "5.hyps"(1) \Some summin = vebt_mint summary\ member_bound mint_member) ultimately have xnin: "both_member_options (Node (Some (mi, ma)) deg treeList summary) ?xn" by (metis "12" "2" "9" \invar_vebt (treeList ! summin) n\ add_leD1 both_member_options_equiv_member both_member_options_from_chilf_to_complete_tree high_inv low_inv member_bound numeral_2_eq_2 plus_1_eq_Suc) let ?h ="high ?xn n" let ?l = "low ?xn n" have "?xn < 2^deg" - by (smt (verit, ccfv_SIG) "5.hyps"(1) "5.hyps"(4) Euclidean_Division.div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) + by (smt (verit, ccfv_SIG) "5.hyps"(1) "5.hyps"(4) div_eq_0_iff \Some lx = vebt_mint (treeList ! summin)\ \Some summin = vebt_mint summary\ \invar_vebt (treeList ! summin) n\ div_exp_eq high_def high_inv le_0_eq member_bound mint_member not_numeral_le_zero power_not_zero) hence "?h < length treeList" using "2" \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger let ?newnode = "vebt_delete (treeList ! ?h) ?l" let ?newlist = "treeList[?h:= ?newnode]" have "length treeList = length ?newlist" by auto hence hprolist: "?newlist ! ?h = ?newnode" by (meson \high (summin * 2 ^ n + lx) n < length treeList\ nth_list_update_eq) have nothprolist: "i \ ?h \ i < 2^m \ ?newlist ! i = treeList ! i" for i by auto have hlbound:"?h < 2^m \ ?l < 2^n" using "2" \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ low_inv member_bound by presburger hence nnvalid: "invar_vebt ?newnode n" by (metis "5.IH"(1) \high (summin * 2 ^ n + lx) n < length treeList\ inthall member_def) have allvalidinlist:"\ t \ set ?newlist. invar_vebt t n" proof fix t assume "t \ set ?newlist" then obtain i where "i < 2^m \ ?newlist ! i = t" by (metis "2" in_set_conv_nth length_list_update) then show "invar_vebt t n" by (metis "0" "2" hprolist nnvalid nth_list_update_neq nth_mem) qed have newlistlength: "length ?newlist = 2^m" by (simp add: "2") then show ?thesis proof(cases "minNull ?newnode") case True hence ninNullc: "minNull ?newnode" by simp let ?sn = "vebt_delete summary ?h" let ?newma= "(if ?xn = ma then (let maxs = vebt_maxt ?sn in (if maxs = None then ?xn else 2^(deg div 2) * the maxs + the (vebt_maxt (?newlist ! the maxs)) ) ) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist ?sn)" have dsimp:"vebt_delete (Node (Some (mi, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist ?sn] by (metis "12" "9" \high (summin * 2 ^ n + lx) n < length treeList\ \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x = mi\ \x \ mi \ x \ ma\ inrg nat_less_le ninNullc) have newsummvalid: "invar_vebt ?sn m" by (simp add: "5.IH"(2)) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options ?sn i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options ?sn i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hprolist by fastforce hence 1001:"\ x. vebt_member (?newlist ! i) x" by (simp add: min_Null_member ninNullc) hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" nnvalid valid_member_both_member_options by auto have 1003: "\ both_member_options ?sn i" using "1" True dele_bmo_cont_corr by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothprolist by blast hence "both_member_options (?newlist ! i) y \ both_member_options ?sn i" for y using "1" "4" False \i < 2 ^ m\ dele_bmo_cont_corr by auto moreover have "both_member_options ?sn i \ \ y. both_member_options (?newlist ! i) y" proof- assume "both_member_options ?sn i " hence "both_member_options summary i" using "1" dele_bmo_cont_corr by auto thus " \ y. both_member_options (?newlist ! i) y" using "1000" "4" \i < 2 ^ m\ by presburger qed then show ?thesis using calculation by blast qed qed qed have 112:" (?xn = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"?xn = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "?xn = ma") case True let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence aa:"\ y. vebt_member ?sn y" using maxt_corr_help_empty newsummvalid set_vebt'_def by auto hence "\ y. both_member_options ?sn y" using newsummvalid valid_member_both_member_options by blast hence "t \ set ?newlist \ \y. both_member_options t y" for t proof- assume "t \ set ?newlist" then obtain i where "?newlist ! i = t \ i< 2^m" by (metis "2" \length treeList = length (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)])\ in_set_conv_nth) thus " \y. both_member_options t y" using "111" \\y. both_member_options (vebt_delete summary (high (summin * 2 ^ n + lx) n)) y\ by blast qed then show ?thesis by blast next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" bb \both_member_options summary maxs\ nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (smt (z3) VEBT_Member.vebt_member.simps(2) \invar_vebt (?newlist ! maxs) n\ vebt_maxt.elims minNull.simps(1) min_Null_member valid_member_both_member_options) hence "maxs = high ?xn n \ both_member_options (?newlist ! maxs) (low ?xn n)" by (smt (z3) "9" False True \Some maxs = vebt_maxt (vebt_delete summary ?h)\ \invar_vebt (?newlist ! maxs) n\ aampt option.sel high_inv low_inv maxbmo maxt_member member_bound mult.commute) hence False using bb by blast then show ?thesis by simp qed next case False hence "?xn \ ?newma" by simp hence False using aampt by simp then show ?thesis by blast qed qed have 114: "?newma < 2^deg \ ?xn \ ?newma" proof(cases "?xn = ma") case True hence "?xn = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "5.hyps"(8) \?xn = ma\ by force next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs) n\ maxt_corr_help_empty set_vebt'_def valid_member_both_member_options by fastforce hence abc:"?newma = 2^n * maxs + maxi" by (smt (z3) "9" True \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ option.sel not_None_eq) have abd:"maxi < 2^n" by (metis \Some maxi = vebt_maxt (?newlist ! maxs)\ \invar_vebt (?newlist ! maxs) n\ maxt_member member_bound) have "high ?xn n \ maxs" using "1" \Some summin = vebt_mint summary\ \both_member_options summary maxs\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound mint_corr_help valid_member_both_member_options by force then show ?thesis proof(cases "high ?xn n = maxs") case True then show ?thesis using bb by fastforce next case False hence "high ?xn n < maxs" by (simp add: \high (summin * 2 ^ n + lx) n \ maxs\ order.not_eq_order_implies_strict) hence "?newma < 2^deg" by (smt (z3) "5.hyps"(8) "9" \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ \invar_vebt (?newlist ! maxs) n\ abd bb both_member_options_equiv_member option.sel high_inv less_le_trans low_inv maxt_member mult.commute nothprolist verit_comp_simplify1(3) yhelper) moreover have "high ?xn n < high ?newma n" by (smt (z3) "9" True \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high (summin * 2 ^ n + lx) n))\ \high (summin * 2 ^ n + lx) n < maxs\ abd option.sel high_inv mult.commute option.discI) ultimately show ?thesis by (metis div_le_mono high_def linear not_less) qed qed next case False then show ?thesis by (smt (z3) "12" "5.hyps"(7) "5.hyps"(8) "9" both_member_options_from_complete_tree_to_child dual_order.trans hlbound one_le_numeral xnin yhelper) qed have 115: "?xn \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof assume assumption0:"?xn \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "?xn = ma" ) case True hence bb:"?xn = ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True hence "?newma = ?xn" using assumption Let_def bb by simp hence False using assumption0 by simp then show ?thesis by simp next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" using \invar_vebt (treeList [high (summin * 2 ^ n + lx) n := vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)] ! maxs) n\ maxt_corr_help_empty set_vebt'_def valid_member_both_member_options by fastforce then show ?thesis by (metis "1" "10" "9" True \Some summin = vebt_mint summary\ \both_member_options summary maxs\ \vebt_member (treeList ! summin) lx\ \mi \ ma \ x < 2 ^ deg\ \invar_vebt (treeList ! summin) n\ bb equals0D high_inv le_antisym maxt_corr_help maxt_corr_help_empty mem_Collect_eq member_bound mint_corr_help option.collapse summaxma set_vebt'_def valid_member_both_member_options) qed next case False hence ccc:"?newma = ma" by simp then show ?thesis proof(cases "?xn = ma") case True hence "?xn = ?newma" using False by blast hence False using False by auto then show ?thesis by simp next case False hence "both_member_options (?newlist ! high ma n) (low ma n)" by (metis "1" \both_member_options (treeList ! high ma n) (low ma n)\ \vebt_member (treeList ! summin) lx\ \vebt_member summary (high ma n)\ \invar_vebt (treeList ! summin) n\ bit_split_inv dele_bmo_cont_corr high_inv hprolist member_bound nothprolist) moreover have "high ma n = i \ low ma n = low ?newma n" using ccc newmaassm by simp ultimately show ?thesis by simp qed qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "?xn < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by auto then show ?thesis using True hprolist min_Null_member ninNullc nnvalid valid_member_both_member_options yassm by fastforce next case False hence "i \ ?h \ False" by (metis "1" "111" \Some summin = vebt_mint summary\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption dele_bmo_cont_corr high_inv le_antisym member_bound mint_corr_help valid_member_both_member_options yassm) hence "i > ?h" using leI by blast then show ?thesis by (metis div_le_mono high_def not_less yassm) qed moreover have "y \ ?newma" proof(cases "?xn = ma") case True hence "?xn= ma" by simp let ?maxs = "vebt_maxt ?sn" show ?thesis proof(cases "?maxs = None") case True then show ?thesis using "1" "111" assumption dele_bmo_cont_corr nothprolist yassm yhelper by auto next case False then obtain maxs where "Some maxs = ?maxs" by fastforce hence "both_member_options summary maxs" by (metis "1" dele_bmo_cont_corr maxbmo) have bb:"maxs \ ?h \ maxs < 2^m" by (metis "1" \Some maxs = vebt_maxt ?sn\ dele_bmo_cont_corr maxbmo member_bound valid_member_both_member_options) hence "invar_vebt (?newlist ! maxs) n"using 0 by (simp add: "2" allvalidinlist) hence "\ y. both_member_options (?newlist ! maxs) y" using "4" \both_member_options summary maxs\ bb nothprolist by presburger then obtain maxi where "Some maxi = vebt_maxt (?newlist ! maxs)" by (metis True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption calculation dele_bmo_cont_corr high_inv hprolist leD member_bound nth_list_update_neq yassm yhelper) hence "maxs < 2^m \ maxi < 2^n" by (metis \invar_vebt (?newlist ! maxs) n\ bb maxt_member member_bound) hence "?newma = 2^n* maxs + maxi" by (smt (z3) "9" False True \Some maxi = vebt_maxt (?newlist ! maxs)\ \Some maxs = vebt_maxt (vebt_delete summary (high ?xn n))\ option.sel) hence "low ?newma n = maxi \ high ?newma n = maxs" by (simp add: \maxs < 2 ^ m \ maxi < 2 ^ n\ high_inv low_inv mult.commute) hence "both_member_options (treeList ! (high y n)) (low y n)" by (metis "1" "111" assumption dele_bmo_cont_corr nothprolist yassm) hence hleqdraft:"high y n > maxs \ False" proof- assume "high y n > maxs" have "both_member_options summary (high y n)" using "1" "111" assumption dele_bmo_cont_corr yassm by blast moreover have "both_member_options ?sn (high y n)" using "111" assumption yassm by blast ultimately show False using True \both_member_options (treeList ! high y n) (low y n)\ \summin * 2 ^ n + lx < y\ assumption leD yassm yhelper by blast qed hence hleqmaxs: "high y n \ maxs" by presburger then show ?thesis using \both_member_options (treeList ! high y n) (low y n)\ assumption calculation dual_order.strict_trans1 yassm yhelper by auto qed next case False then show ?thesis by (smt (z3) \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption dele_bmo_cont_corr high_inv hprolist member_bound nothprolist yassm yhelper) qed ultimately show " ?xn < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "?xn \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (?xn, ?newma)) deg ?newlist ?sn) deg" using invar_vebt.intros(5)[of ?newlist n ?sn m deg ?xn ?newma] using 3 allvalidinlist newlistlength newsummvalid "5.hyps"(3) 111 112 118 117 115 by fastforce show ?thesis using "116" dsimp by presburger next case False hence notemp:"\ z. both_member_options ?newnode z" using not_min_Null_member by auto let ?newma = "(if ?xn = ma then ?h * 2^(deg div 2) + the( vebt_maxt (?newlist ! ?h)) else ma)" let ?delsimp =" (Node (Some (?xn, ?newma)) deg ?newlist summary)" have dsimp:"vebt_delete (Node (Some (x, ma)) deg treeList summary) x = ?delsimp" using del_x_mi_lets_in_not_minNull[of x mi ma deg ?xn ?h summary treeList ?l ?newnode ?newlist] "12" "2" "9" False dual_order.eq_iff hlbound inrg order.not_eq_order_implies_strict xmi by (metis \summin * 2 ^ n + lx = (if x = mi then the (vebt_mint summary) * 2 ^ (deg div 2) + the (vebt_mint (treeList ! the (vebt_mint summary))) else x)\ \x \ mi \ x \ ma\) have 111: "(\ i < 2^m. (\ x. both_member_options (?newlist ! i) x) \ ( both_member_options summary i))" proof fix i show " i < 2^m \ ((\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i))" proof assume "i < 2^m" show "(\ x. both_member_options (?newlist ! i) x) = ( both_member_options summary i)" proof(cases "i = ?h") case True hence 1000:"?newlist ! i = ?newnode" using hprolist by blast hence 1001:"\ x. vebt_member (?newlist ! i) x" using nnvalid notemp valid_member_both_member_options by auto hence 1002: "\ x. both_member_options (?newlist ! i) x" using "1000" notemp by presburger have 1003: "both_member_options summary i" using "4" True \\z. both_member_options (treeList ! summin) z\ \vebt_member (treeList ! summin) lx\ \summin < 2 ^ m\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by auto then show ?thesis using "1002" by blast next case False hence 1000:"?newlist ! i = treeList ! i" using \i < 2 ^ m\ nothprolist by blast then show ?thesis using "4" \i < 2 ^ m\ by presburger qed qed qed have 112:" (?xn = ?newma \ (\ t \ set ?newlist. \ x. both_member_options t x))" proof assume aampt:"?xn = ?newma" show "(\ t \ set ?newlist. \ y. both_member_options t y)" proof(cases "?xn = ma") case True obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq False hprolist maxt_corr_help_empty nnvalid not_None_eq not_min_Null_member set_vebt'_def valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hprolist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv member_bound by force hence False by (metis "9" \both_member_options (vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)) maxi\ \vebt_maxt (?newlist ! ?h) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ aampt add_diff_cancel_left' dele_bmo_cont_corr option.sel high_inv low_inv member_bound) then show ?thesis by blast next case False then show ?thesis using \mi \ ma \ x < 2 ^ deg\ aampt by presburger qed qed have 114: "?newma < 2^deg \ ?xn \ ?newma" proof(cases "?xn = ma") case True hence "?xn = ma" by simp obtain maxi where " vebt_maxt (?newlist ! ?h) = Some maxi" by (metis "111" "2" "4" Collect_empty_eq True \both_member_options (treeList ! high ma n) (low ma n)\ \high (summin * 2 ^ n + lx) n < length treeList\ hprolist maxt_corr_help_empty nnvalid not_None_eq set_vebt'_def valid_member_both_member_options) hence "both_member_options ?newnode maxi" using hprolist maxbmo by auto hence "both_member_options (treeList ! ?h) maxi" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv member_bound by force hence "maxi < 2^n" using \both_member_options?newnode maxi\ member_bound nnvalid valid_member_both_member_options by blast show ?thesis - by (smt (verit, ccfv_threshold) "3" "9" Euclidean_Division.div_eq_0_iff True \Some lx = vebt_mint (treeList ! summin)\ \both_member_options (treeList ! high (summin * 2 ^ n + lx) n) maxi\ \vebt_maxt (?newlist ! high (summin * 2 ^ n + lx) n) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add.right_neutral add_left_mono div_mult2_eq div_mult_self3 option.sel high_inv hlbound le_0_eq member_bound mint_corr_help power_add power_not_zero rel_simps(28) valid_member_both_member_options) + by (smt (verit, ccfv_threshold) "3" "9" div_eq_0_iff True \Some lx = vebt_mint (treeList ! summin)\ \both_member_options (treeList ! high (summin * 2 ^ n + lx) n) maxi\ \vebt_maxt (?newlist ! high (summin * 2 ^ n + lx) n) = Some maxi\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add.right_neutral add_left_mono div_mult2_eq div_mult_self3 option.sel high_inv hlbound le_0_eq member_bound mint_corr_help power_add power_not_zero rel_simps(28) valid_member_both_member_options) next case False then show ?thesis using "10" "5.hyps"(8) maxt_corr_help valid_member_both_member_options xnin by force qed have 115: "?xn \ ?newma \ (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof assume xnmassm:"?xn \ ?newma" show " (\ i < 2^m. (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma))" proof fix i show "i < 2^m \ (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof assume assumption:"i < 2^m" show " (high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n)) \ (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof- have "(high ?newma n = i \ both_member_options (?newlist ! i) (low ?newma n))" proof assume newmaassm: "high ?newma n = i" thus " both_member_options (?newlist ! i) (low ?newma n)" proof(cases "?xn = ma") case True obtain maxi where "vebt_maxt (?newlist ! ?h) = Some maxi" by (metis Collect_empty_eq both_member_options_equiv_member hprolist maxt_corr_help_empty nnvalid not_Some_eq notemp set_vebt'_def) hence "both_member_options (?newlist ! ?h) maxi" using maxbmo by blast then show ?thesis by (smt (z3) "2" "9" True \Some lx = vebt_mint (treeList ! summin)\ \high (summin * 2 ^ n + lx) n < length treeList\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ add_left_mono dele_bmo_cont_corr eq_iff high_inv hprolist low_inv member_bound mint_corr_help valid_member_both_member_options yhelper) next case False hence abcd:"?newma = ma" by simp then show ?thesis proof(cases "high ma n = ?h") case True hence "?newlist ! high ma n = ?newnode" using hprolist by presburger then show ?thesis proof(cases "low ma n = ?l") case True hence "?newma = ?xn" by (metis "1" False \?newlist ! high ma n = vebt_delete (treeList ! high (summin * 2 ^ n + lx) n) (low (summin * 2 ^ n + lx) n)\ \both_member_options (treeList ! high ma n) (low ma n)\ \vebt_member (treeList ! summin) lx\ \vebt_member summary (high ma n)\ \invar_vebt (treeList ! summin) n\ bit_split_inv dele_bmo_cont_corr high_inv member_bound nothprolist) hence False using False by presburger then show ?thesis by simp next case False have "both_member_options (treeList ! high ma n) (low ma n)" by (simp add: \both_member_options (treeList ! high ma n) (low ma n)\) hence "both_member_options ?newnode (low ma n)" using False True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv member_bound by force hence "both_member_options (?newlist ! high ma n) (low ma n)" using True hprolist by presburger then show ?thesis using abcd newmaassm by simp qed next case False hence "?newlist ! high ma n = treeList ! high ma n" using "1" \vebt_member summary (high ma n)\ member_bound nothprolist by blast moreover hence "both_member_options (treeList ! high ma n) (low ma n)" using \both_member_options (treeList ! high ma n) (low ma n)\ by blast ultimately show ?thesis using abcd newmaassm by simp qed qed qed moreover have " (\ y. (high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma)" proof fix y show "(high y n = i \ both_member_options (?newlist ! i) (low y n) ) \ ?xn < y \ y \ ?newma" proof assume yassm: "(high y n = i \ both_member_options (?newlist ! i) (low y n) )" hence "?xn < y" proof(cases "i = ?h") case True hence "both_member_options (treeList ! i) (low y n)" using \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by force moreover have "vebt_mint (treeList ! i) = Some (low ?xn n)" using True \Some lx = vebt_mint (treeList ! summin)\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv low_inv member_bound by presburger moreover hence "low y n \ low ?xn n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ calculation(1) high_inv member_bound mint_corr_help valid_member_both_member_options by auto moreover have "low y n \ low ?xn n" using True \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ dele_bmo_cont_corr high_inv hprolist member_bound yassm by auto ultimately have "low y n > low ?xn n" by simp show ?thesis by (metis True \low (summin * 2 ^ n + lx) n \ low y n\ \low y n \ low (summin * 2 ^ n + lx) n\ bit_concat_def bit_split_inv leD linorder_neqE_nat nat_add_left_cancel_less yassm) next case False have "Some (high ?xn n) = vebt_mint summary" using \Some summin = vebt_mint summary\ \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ high_inv member_bound by presburger moreover hence "high y n \ high ?xn n" by (metis "1" "111" assumption mint_corr_help valid_member_both_member_options yassm) ultimately show ?thesis by (metis False div_le_mono high_def leI le_antisym yassm) qed moreover have "y \ ?newma" by (smt (z3) \vebt_member (treeList ! summin) lx\ \invar_vebt (treeList ! summin) n\ assumption calculation dele_bmo_cont_corr high_inv hprolist leD member_bound nothprolist yassm yhelper) ultimately show " ?xn < y \ y \ ?newma" by simp qed qed ultimately show ?thesis by simp qed qed qed qed hence 117: "?newma < 2^deg" and 118: "?xn \ ?newma" using 114 by auto have 116: " invar_vebt (Node (Some (?xn, ?newma)) deg ?newlist summary) deg" using invar_vebt.intros(5)[of ?newlist n summary m deg ?xn ?newma] allvalidinlist 1 newlistlength 8 3 111 112 117 118 115 by fastforce hence "invar_vebt (?delsimp) deg" by simp moreover obtain delsimp where 118:"delsimp = ?delsimp" by simp ultimately have 119:"invar_vebt delsimp deg" by simp have "vebt_delete (Node (Some (x, ma)) deg treeList summary) x = delsimp" using dsimp 118 by simp hence "delsimp = vebt_delete (Node (Some (x, ma)) deg treeList summary) x" by simp then show ?thesis using 119 using xmi by auto qed qed qed qed qed corollary dele_member_cont_corr:"invar_vebt t n \ (vebt_member (vebt_delete t x) y \ x \ y \ vebt_member t y)" by (meson both_member_options_equiv_member dele_bmo_cont_corr delete_pres_valid) subsection \Correctness with Respect to Set Interpretation\ theorem delete_correct': assumes "invar_vebt t n" shows "set_vebt' (vebt_delete t x) = set_vebt' t - {x}" using assms by(auto simp add: set_vebt'_def dele_member_cont_corr) corollary delete_correct: assumes "invar_vebt t n" shows "set_vebt' (vebt_delete t x) = set_vebt t - {x}" using assms delete_correct' set_vebt_set_vebt'_valid by auto end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_Insert.thy b/thys/Van_Emde_Boas_Trees/VEBT_Insert.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_Insert.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_Insert.thy @@ -1,825 +1,825 @@ (*by Ammer*) theory VEBT_Insert imports VEBT_Member begin section \Insert Function\ context begin interpretation VEBT_internal . fun vebt_insert :: "VEBT \ nat \VEBT" where "vebt_insert (Leaf a b) x = (if x=0 then Leaf True b else if x = 1 then Leaf a True else Leaf a b)"| "vebt_insert (Node info 0 ts s) x = (Node info 0 ts s)"| "vebt_insert (Node info (Suc 0) ts s) x = (Node info (Suc 0) ts s)"| "vebt_insert (Node None (Suc deg) treeList summary) x = (Node (Some (x,x)) (Suc deg) treeList summary)"| "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = ( let xn = (if x < mi then mi else x); minn = (if x < mi then x else mi); l = low xn (deg div 2); h = high xn (deg div 2) in ( if h < length treeList \ \ (x = mi \ x = ma) then Node (Some (minn, max xn ma)) deg (treeList[h:= vebt_insert (treeList ! h) l]) (if minNull (treeList ! h) then vebt_insert summary h else summary) else (Node (Some (mi, ma)) deg treeList summary)))" end context VEBT_internal begin lemma insert_simp_norm: assumes "high x (deg div 2) < length treeList " and "(mi::nat)< x" and "deg\ 2" and "x \ ma" shows "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList [(high x (deg div 2)):= vebt_insert (treeList ! (high x (deg div 2))) (low x (deg div 2))]) (if minNull (treeList ! (high x (deg div 2))) then vebt_insert summary (high x (deg div 2)) else summary) " proof- have 11:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = (let xn = (if x < mi then mi else x); minn = (if x< mi then x else mi); l= low xn (deg div 2); h = high xn (deg div 2) in ( if h < length treeList \ \ (x = mi \ x = ma)then Node (Some (minn, max xn ma)) deg (treeList [h:= vebt_insert (treeList ! h) l]) (if minNull (treeList ! h) then vebt_insert summary h else summary) else (Node (Some (mi, ma)) deg treeList summary)))" using assms(3) vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] by (smt add_numeral_left le_add_diff_inverse numerals(1) plus_1_eq_Suc semiring_norm(2)) have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[(high x (deg div 2)) := vebt_insert (treeList ! (high x (deg div 2))) (low x (deg div 2))]) (if minNull (treeList ! (high x (deg div 2))) then vebt_insert summary (high x (deg div 2)) else summary)" using 11 apply (simp add: Let_def) apply (auto simp add: If_def) using assms not_less_iff_gr_or_eq apply blast+ done then show ?thesis by blast qed lemma insert_simp_excp: assumes "high mi (deg div 2) < length treeList " and " (x::nat) < mi" and "deg\ 2" and "x \ ma" shows "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[(high mi (deg div 2)) := vebt_insert (treeList ! (high mi (deg div 2))) (low mi (deg div 2))]) (if minNull (treeList ! (high mi (deg div 2))) then vebt_insert summary (high mi (deg div 2)) else summary) " proof- have 11:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = ( let xn = (if x < mi then mi else x); minn = (if x< mi then x else mi); l= low xn (deg div 2); h = high xn (deg div 2) in ( if h < length treeList \ \ (x = mi \ x = ma)then Node (Some (minn, max xn ma)) deg (treeList[h:=vebt_insert (treeList ! h) l]) (if minNull (treeList ! h) then vebt_insert summary h else summary) else (Node (Some (mi, ma)) deg treeList summary)))" using assms(3) vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] by (smt add_numeral_left le_add_diff_inverse numerals(1) plus_1_eq_Suc semiring_norm(2)) have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg ( treeList[ (high mi (deg div 2)) := vebt_insert (treeList ! (high mi (deg div 2))) (low mi (deg div 2))]) (if minNull (treeList ! (high mi (deg div 2))) then vebt_insert summary (high mi (deg div 2)) else summary)" using 11 apply (simp add: Let_def) apply (auto simp add: If_def) using assms not_less_iff_gr_or_eq apply blast+ done then show ?thesis by blast qed lemma insert_simp_mima: assumes "x = mi \ x = ma" and "deg \ 2" shows "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = (Node (Some (mi,ma)) deg treeList summary)" proof - have 11:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = ( let xn = (if x < mi then mi else x); minn = (if x< mi then x else mi); l= low xn (deg div 2); h = high xn (deg div 2) in ( if h < length treeList \ \ (x = mi \ x = ma)then Node (Some (minn, max xn ma)) deg (treeList[h:= vebt_insert (treeList ! h) l]) (if minNull (treeList ! h) then vebt_insert summary h else summary) else (Node (Some (mi, ma)) deg treeList summary)))" using assms vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] by (smt add_numeral_left le_add_diff_inverse numerals(1) plus_1_eq_Suc semiring_norm(2)) then show ?thesis using assms(1) by auto qed lemma valid_insert_both_member_options_add: "invar_vebt t n \ x< 2^n \ both_member_options (vebt_insert t x) x" proof(induction t n arbitrary: x rule: invar_vebt.induct) case (1 a b) then show ?case apply(cases x) by (auto simp add: both_member_options_def) next case (2 treeList n summary m deg) hence "deg>1" using valid_tree_deg_neq_0 by (metis One_nat_def Suc_lessI add_gr_0 add_self_div_2 neq0_conv one_div_two_eq_zero) then show ?case using vebt_insert.simps(4)[of "deg-2" treeList summary x ] by (smt Suc_1 Suc_leI add_numeral_left both_member_options_def le_add_diff_inverse membermima.simps(4) numerals(1) plus_1_eq_Suc semiring_norm(2)) next case (3 treeList n summary m deg) hence "\t\set treeList. invar_vebt t n" by blast hence "n > 0" using set_n_deg_not_0[of treeList n m] "3"(4) by linarith hence "deg \ 2" by (simp add: "3.hyps"(3) "3.hyps"(4) Suc_leI) then show ?case using vebt_insert.simps(4)[of "deg-2" treeList summary x ] by (smt Suc_1 Suc_leI add_numeral_left both_member_options_def le_add_diff_inverse membermima.simps(4) numerals(1) plus_1_eq_Suc semiring_norm(2)) next case (4 treeList n summary m deg mi ma) hence "length treeList =2^n" by blast hence "high x n < length treeList" using "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) "4.prems" deg_not_0 exp_split_high_low(1) by auto hence "mi < 2^deg" using "4.hyps"(7) "4.hyps"(8) le_less_trans by blast then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] by (smt "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) add_diff_inverse_nat add_numeral_left add_self_div_2 both_member_options_def div_if membermima.simps(4) numerals(1) plus_1_eq_Suc semiring_norm(2) valid_tree_deg_neq_0) next case False hence "\ (x = mi \ x = ma)" by simp then show ?thesis proof(cases "x < mi") case True hence "high mi n < length treeList" using "4.hyps"(1) "4.hyps"(2) "4.hyps"(3) "4.hyps"(4) \mi < 2 ^ deg\ deg_not_0 exp_split_high_low(1) by auto hence "vebt_insert ( Node (Some (mi, ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg ( treeList[(high mi n):= vebt_insert (treeList ! (high mi n)) (low mi n)] ) (if minNull (treeList ! high mi n) then vebt_insert summary (high mi n) else summary)" by (metis "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) False True add_self_div_2 div_if insert_simp_excp not_less valid_tree_deg_neq_0) then show ?thesis by (smt "4.hyps"(1) "4.hyps"(4) Suc_pred add_diff_inverse_nat both_member_options_def membermima.simps(4) valid_tree_deg_neq_0 zero_eq_add_iff_both_eq_0) next case False hence "vebt_insert ( Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)" by (metis "4.hyps"(1) "4.hyps"(3) "4.hyps"(4) \\ (x = mi \ x = ma)\ \high x n < length treeList\ add_self_div_2 div_if insert_simp_norm linorder_neqE_nat not_less valid_tree_deg_neq_0) have "low x n < 2^n \ high x n < 2^n" using "4.hyps"(2) "4.hyps"(3) \high x n < length treeList\ low_def by auto have "invar_vebt (treeList ! (high x n)) n" by (metis "4.IH"(1) \high x n < length treeList\ inthall member_def) hence "both_member_options (vebt_insert (treeList ! (high x n)) (low x n)) (low x n)" by (simp add: "4.IH"(1) \high x n < length treeList\ low_def) have " (treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) ! (high x n) = vebt_insert (treeList ! (high x n)) (low x n)" by (simp add: \high x n < length treeList\) then show ?thesis using both_member_options_ding[of "Some (mi, max x ma)" deg "(take (high x n) treeList @ [vebt_insert (treeList ! (high x n)) (low x n)] @ drop (high x n +1) treeList)" "if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary" n x] by (metis "4.hyps"(2) "4.hyps"(3) "4.hyps"(4) Suc_1 Suc_leD \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)\ \both_member_options (vebt_insert (treeList ! high x n) (low x n)) (low x n)\ \low x n < 2 ^ n \ high x n < 2 ^ n\ \invar_vebt (treeList ! high x n) n\ add_self_div_2 both_member_options_from_chilf_to_complete_tree deg_not_0 div_greater_zero_iff length_list_update) qed qed next case (5 treeList n summary m deg mi ma) hence "length treeList =2^m" by blast hence "high x n < length treeList" - by (metis "5.hyps"(4) "5.prems" Euclidean_Division.div_eq_0_iff div_exp_eq high_def length_0_conv length_greater_0_conv zero_less_numeral zero_less_power) + by (metis "5.hyps"(4) "5.prems" div_eq_0_iff div_exp_eq high_def length_0_conv length_greater_0_conv zero_less_numeral zero_less_power) hence "mi<2^deg" using "5.hyps"(7) "5.hyps"(8) le_less_trans by blast then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] by (smt "5.hyps"(3) "5.hyps"(4) Suc_leI add_Suc_right add_diff_inverse_nat add_numeral_left both_member_options_def diff_is_0_eq' vebt_insert.simps(3) membermima.simps(4) not_add_less1 numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False hence "\ (x = mi \ x = ma)" by simp then show ?thesis proof(cases "x < mi") case True hence "high mi n < length treeList" - by (metis "5.hyps"(2) "5.hyps"(4) Euclidean_Division.div_eq_0_iff \mi < 2 ^ deg\ div_exp_eq high_def length_0_conv length_greater_0_conv zero_less_numeral zero_less_power) + by (metis "5.hyps"(2) "5.hyps"(4) div_eq_0_iff \mi < 2 ^ deg\ div_exp_eq high_def length_0_conv length_greater_0_conv zero_less_numeral zero_less_power) hence "vebt_insert ( Node (Some (mi, ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg ( treeList[ (high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)] ) (if minNull (treeList ! high mi n) then vebt_insert summary (high mi n) else summary)" using insert_simp_excp[of mi deg treeList x ma summary] "5"(1) "5.hyps"(3) "5.hyps"(4) False True add_Suc_right add_self_div_2 append_Cons div_less even_Suc_div_two in_set_conv_decomp not_less odd_add valid_tree_deg_neq_0 by (smt (z3) nth_mem) then show ?thesis by (simp add: "5.hyps"(3) "5.hyps"(4) both_member_options_def) next case False hence "vebt_insert ( Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[(high x n):= vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)" by (smt (z3) "5.IH"(1) "5.hyps"(3) "5.hyps"(4) \\ (x = mi \ x = ma)\ \high x n < length treeList\ add_Suc_right add_self_div_2 deg_not_0 div_greater_zero_iff even_Suc_div_two insert_simp_norm linorder_neqE_nat nth_mem odd_add) have "low x n < 2^n \ high x n < 2^m" using "5.hyps"(2) "5.hyps"(3) \high x n < length treeList\ low_def by auto have "invar_vebt (treeList ! (high x n)) n" by (metis "5.IH"(1) \high x n < length treeList\ inthall member_def) hence "both_member_options (vebt_insert (treeList ! (high x n)) (low x n)) (low x n)" by (metis "5.IH"(1) \high x n < length treeList\ \low x n < 2 ^ n \ high x n < 2 ^ m\ inthall member_def) have " (treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) ! (high x n) = vebt_insert (treeList ! (high x n)) (low x n)" by (meson \high x n < length treeList\ nth_list_update_eq) then show ?thesis using both_member_options_ding[of "Some (mi, max x ma)" deg "(treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)])" "if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary" n x] using "5.hyps"(2) "5.hyps"(3) "5.hyps"(4) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)\ \both_member_options (vebt_insert (treeList ! high x n) (low x n)) (low x n)\ \low x n < 2 ^ n \ high x n < 2 ^ m\ both_member_options_from_chilf_to_complete_tree by auto qed qed qed lemma valid_insert_both_member_options_pres: "invar_vebt t n \ x<2^n \ y < 2^n \ both_member_options t x \ both_member_options (vebt_insert t y) x" proof(induction t n arbitrary: x y rule: invar_vebt.induct) case (1 a b) then show ?case by (simp add: both_member_options_def) next case (2 treeList n summary m deg) then show ?case using vebt_member.simps(2) invar_vebt.intros(2) valid_member_both_member_options by blast next case (3 treeList n summary m deg) then show ?case using vebt_member.simps(2) invar_vebt.intros(3) valid_member_both_member_options by blast next case (4 treeList n summary m deg mi ma) hence 00:"deg = n + m \ length treeList = 2^n \ n = m \ n \ 1 \ deg \ 2" by (metis One_nat_def Suc_leI add_mono_thms_linordered_semiring(1) deg_not_0 one_add_one) hence xyprop: "high x n < 2^m \ high y n < 2^m" by (metis "4.prems"(1) "4.prems"(2) high_def less_mult_imp_div_less mult_2 power2_eq_square power_even_eq) have "low x n <2^n \ low y n< 2^n" by (simp add: low_def) hence "x = mi \ x = ma \ both_member_options (treeList ! (high x n)) (low x n)" by (smt "00" "4.prems"(3) add_Suc_right add_self_div_2 both_member_options_def le_add_diff_inverse membermima.simps(4) naive_member.simps(3) plus_1_eq_Suc) have 001:"invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(4)[of treeList n summary m deg mi ma] "4" by simp then show ?case proof(cases "x = y") case True hence "both_member_options (vebt_insert (Node (Some (mi, ma)) deg treeList summary) y) y" using 001 valid_insert_both_member_options_add[of "(Node (Some (mi, ma)) deg treeList summary)" deg y ] using "4.prems"(2) by blast then show ?thesis by (simp add: True) next case False then show ?thesis proof(cases "y = mi \ y = ma") case True have "Suc (Suc (deg -2)) = deg" using "00" by linarith hence "vebt_insert (Node (Some (mi, ma)) deg treeList summary) y = Node (Some (mi, ma)) deg treeList summary" using vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] 00 True insert_simp_mima by blast then show ?thesis by (simp add: "4.prems"(3)) next case False hence 0:"y \ mi \ y \ ma" by simp then show ?thesis proof(cases "x = mi") case True hence 1:"x = mi" by simp then show ?thesis proof(cases "x < y") case True have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (x, max y ma)) deg (treeList [ (high y n):=vebt_insert (treeList ! (high y n)) (low y n)] ) (if minNull (treeList ! (high y n)) then vebt_insert summary (high y n) else summary)" using "00" "1" False True insert_simp_norm xyprop by auto then show ?thesis by (metis "001" Suc_pred both_member_options_def deg_not_0 membermima.simps(4)) next case False have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (y, max x ma)) deg (treeList [ (high x n) :=vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! (high x n)) then vebt_insert summary (high x n) else summary)" by (metis "0" "00" False True add_self_div_2 insert_simp_excp linorder_neqE_nat xyprop) have 15: " invar_vebt (treeList ! (high x n)) n" by (metis "4"(1) "4.hyps"(2) in_set_member inthall xyprop) hence 16: "both_member_options (vebt_insert (treeList ! high x n) (low x n)) (low x n)" using \low x n < 2 ^ n \ low y n < 2 ^ n\ valid_insert_both_member_options_add by blast then show ?thesis by (metis "00" "14" Suc_1 add_leD1 add_self_div_2 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq plus_1_eq_Suc xyprop) qed next case False hence "mi \ ma" using "001" "4.prems"(3) less_irrefl member_inv valid_member_both_member_options by fastforce hence "both_member_options (treeList !(high x n) ) (low x n) \ x = ma" using False \x = mi \ x = ma \ both_member_options (treeList ! high x n) (low x n)\ by blast have "high ma n < 2^n" by (metis "4.hyps"(3) "4.hyps"(4) "4.hyps"(8) high_def less_mult_imp_div_less mult_2 power2_eq_square power_even_eq) hence "both_member_options (treeList !(high ma n) ) (low ma n)" using "4.hyps"(3) "4.hyps"(9) \mi \ ma\ by blast hence "both_member_options (treeList !(high x n) ) (low x n)" using \both_member_options (treeList ! high x n) (low x n) \ x = ma\ by blast then show ?thesis proof(cases "mi < y") case True have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (mi, max y ma)) deg (treeList[(high y n):=vebt_insert (treeList ! (high y n)) (low y n)]) (if minNull (treeList ! (high y n)) then vebt_insert summary (high y n) else summary)" using "0" "00" True insert_simp_norm xyprop by auto have "invar_vebt (treeList ! (high x n)) n" by (metis "4.IH"(1) "4.hyps"(2) in_set_member inthall xyprop) then show ?thesis proof(cases "high x n = high y n") case True have "both_member_options (vebt_insert (treeList ! (high y n)) (low y n)) (low x n)" using "4.IH"(1) "4.hyps"(2) True \both_member_options (treeList ! high x n) (low x n)\ \low x n < 2 ^ n \ low y n < 2 ^ n\ xyprop by auto then show ?thesis by (metis "00" "14" Suc_1 True add_leD1 add_self_div_2 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq plus_1_eq_Suc xyprop) next case False have "(treeList[ (high y n):=vebt_insert (treeList ! (high y n)) (low y n)]) ! (high x n) = treeList ! (high x n)" using False by auto then show ?thesis by (metis "00" "14" One_nat_def Suc_leD \both_member_options (treeList ! high x n) (low x n)\ add_self_div_2 both_member_options_from_chilf_to_complete_tree length_list_update numeral_2_eq_2 xyprop) qed next case False have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (y, max mi ma)) deg (treeList[ (high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) (if minNull (treeList ! (high mi n)) then vebt_insert summary (high mi n) else summary)" using insert_simp_excp[of mi deg treeList y ma summary] by (smt "0" "00" "4.hyps"(7) "4.hyps"(8) False add_self_div_2 antisym_conv3 high_def le_less_trans less_mult_imp_div_less mult_2 power2_eq_square power_even_eq) have mimaprop: "high mi n < 2^n \ low mi n < 2^n" - by (metis "00" "4.hyps"(7) "4.hyps"(8) Euclidean_Division.div_eq_0_iff div_exp_eq high_def le_less_trans low_def mod_less_divisor zero_less_numeral zero_less_power) + by (metis "00" "4.hyps"(7) "4.hyps"(8) div_eq_0_iff div_exp_eq high_def le_less_trans low_def mod_less_divisor zero_less_numeral zero_less_power) have "invar_vebt (treeList ! (high x n)) n" by (metis "4.IH"(1) "4.hyps"(2) in_set_member inthall xyprop) then show ?thesis proof(cases "high x n = high mi n") case True have "both_member_options (vebt_insert (treeList ! (high mi n)) (low mi n)) (low x n)" by (metis "4.IH"(1) "4.hyps"(2) True \both_member_options (treeList ! high x n) (low x n)\ \low x n < 2 ^ n \ low y n < 2 ^ n\ mimaprop nth_mem xyprop) then show ?thesis by (metis "00" "14" Suc_1 Suc_leD True add_self_div_2 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq xyprop) next case False have "(treeList[(high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) ! (high x n) = treeList ! (high x n)" using False by force then show ?thesis by (metis "00" "14" One_nat_def Suc_leD \both_member_options (treeList ! high x n) (low x n)\ add_self_div_2 both_member_options_from_chilf_to_complete_tree length_list_update numeral_2_eq_2 xyprop) qed qed qed qed qed next case (5 treeList n summary m deg mi ma) hence 00:"deg = n + m \ length treeList = 2^m \ Suc n = m \ n \ 1 \ deg \ 2 \ n = deg div 2" by (metis Suc_1 add_Suc_right add_mono_thms_linordered_semiring(1) add_self_div_2 even_Suc_div_two le_add1 odd_add plus_1_eq_Suc set_n_deg_not_0) hence xyprop: "high x n < 2^m \ high y n < 2^m" by (metis "5.prems"(1) "5.prems"(2) Suc_1 div_exp_eq div_if high_def nat.discI power_not_zero) have "low x n <2^n \ low y n< 2^n" by (simp add: low_def) hence "x = mi \ x = ma \ both_member_options (treeList ! (high x n)) (low x n)" by (smt "00" "5.prems"(3) add_Suc_right add_self_div_2 both_member_options_def le_add_diff_inverse membermima.simps(4) naive_member.simps(3) plus_1_eq_Suc) have 001:"invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" using invar_vebt.intros(5)[of treeList n summary m deg mi ma] "5" by simp then show ?case proof(cases "x = y") case True hence "both_member_options (vebt_insert (Node (Some (mi, ma)) deg treeList summary) y) y" using 001 valid_insert_both_member_options_add[of "(Node (Some (mi, ma)) deg treeList summary)" deg y ] using "5.prems"(2) by blast then show ?thesis by (simp add: True) next case False then show ?thesis proof(cases "y = mi \ y = ma") case True have "Suc (Suc (deg -2)) = deg" using "00" by linarith hence "vebt_insert (Node (Some (mi, ma)) deg treeList summary) y = Node (Some (mi, ma)) deg treeList summary" using vebt_insert.simps(5)[of mi ma "deg-2" treeList summary x] 00 True insert_simp_mima by blast then show ?thesis by (simp add: "5.prems"(3)) next case False hence 0:"y \ mi \ y \ ma" by simp then show ?thesis proof(cases "x = mi") case True hence 1:"x = mi" by simp then show ?thesis proof(cases "x < y") case True have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (x, max y ma)) deg (treeList[ (high y n):=vebt_insert (treeList ! (high y n)) (low y n)] ) (if minNull (treeList ! (high y n)) then vebt_insert summary (high y n) else summary)" using "00" "1" False True insert_simp_norm xyprop by metis then show ?thesis by (metis "001" Suc_pred both_member_options_def deg_not_0 membermima.simps(4)) next case False have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (y, max x ma)) deg (treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! (high x n)) then vebt_insert summary (high x n) else summary)" by (metis "0" "00" False True add_self_div_2 insert_simp_excp linorder_neqE_nat xyprop) have 15: " invar_vebt (treeList ! (high x n)) n" by (metis "5"(1) "5.hyps"(2) in_set_member inthall xyprop) hence 16: "both_member_options (vebt_insert (treeList ! high x n) (low x n)) (low x n)" using \low x n < 2 ^ n \ low y n < 2 ^ n\ valid_insert_both_member_options_add by blast then show ?thesis by (metis "00" "14" Suc_1 add_leD1 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq plus_1_eq_Suc xyprop) qed next case False hence "mi \ ma" using "001" "5.prems"(3) less_irrefl member_inv valid_member_both_member_options by fastforce hence "both_member_options (treeList !(high x n) ) (low x n) \ x = ma" using False \x = mi \ x = ma \ both_member_options (treeList ! high x n) (low x n)\ by blast have "high ma n < 2^m" - by (metis "00" "5.hyps"(8) Euclidean_Division.div_eq_0_iff div_exp_eq high_def nat_zero_less_power_iff power_not_zero zero_power2) + by (metis "00" "5.hyps"(8) div_eq_0_iff div_exp_eq high_def nat_zero_less_power_iff power_not_zero zero_power2) hence "both_member_options (treeList !(high ma n) ) (low ma n)" using "5.hyps"(3) "5.hyps"(9) \mi \ ma\ by blast hence "both_member_options (treeList !(high x n) ) (low x n)" using \both_member_options (treeList ! high x n) (low x n) \ x = ma\ by blast then show ?thesis proof(cases "mi < y") case True have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (mi, max y ma)) deg (treeList[(high y n):= vebt_insert (treeList ! (high y n)) (low y n)]) (if minNull (treeList ! (high y n)) then vebt_insert summary (high y n) else summary)" by (metis "0" "00" True insert_simp_norm xyprop) have "invar_vebt (treeList ! (high x n)) n" by (metis "5.IH"(1) "5.hyps"(2) in_set_member inthall xyprop) then show ?thesis proof(cases "high x n = high y n") case True have "both_member_options (vebt_insert (treeList ! (high y n)) (low y n)) (low x n)" by (metis "5.IH"(1) "5.hyps"(2) True \both_member_options (treeList ! high x n) (low x n)\ \low x n < 2 ^ n \ low y n < 2 ^ n\ nth_mem xyprop) then show ?thesis by (metis "00" "14" Suc_1 True add_leD1 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq plus_1_eq_Suc xyprop) next case False have "(treeList[ (high y n):=vebt_insert (treeList ! (high y n)) (low y n)] ) ! (high x n) = treeList ! (high x n)" using False by force then show ?thesis by (metis "00" "14" One_nat_def Suc_leD \both_member_options (treeList ! high x n) (low x n)\ both_member_options_from_chilf_to_complete_tree length_list_update numeral_2_eq_2 xyprop) qed next case False have 14:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) y = Node (Some (y, max mi ma)) deg (treeList[(high mi n):= vebt_insert (treeList ! (high mi n)) (low mi n)] ) (if minNull (treeList ! (high mi n)) then vebt_insert summary (high mi n) else summary)" using insert_simp_excp[of mi deg treeList y ma summary] - by (metis "0" "00" "5.hyps"(7) "5.hyps"(8) Euclidean_Division.div_eq_0_iff False antisym_conv3 div_exp_eq high_def le_less_trans power_not_zero zero_neq_numeral) + by (metis "0" "00" "5.hyps"(7) "5.hyps"(8) div_eq_0_iff False antisym_conv3 div_exp_eq high_def le_less_trans power_not_zero zero_neq_numeral) have mimaprop: "high mi n < 2^m \ low mi n < 2^n" using exp_split_high_low[of mi n m] 00 "5"(9,10) by force have "invar_vebt (treeList ! (high x n)) n" by (metis "5.IH"(1) "5.hyps"(2) in_set_member inthall xyprop) then show ?thesis proof(cases "high x n = high mi n") case True have "both_member_options (vebt_insert (treeList ! (high mi n)) (low mi n)) (low x n)" by (metis "5.IH"(1) "5.hyps"(2) True \both_member_options (treeList ! high x n) (low x n)\ \low x n < 2 ^ n \ low y n < 2 ^ n\ mimaprop nth_mem) then show ?thesis by (metis "00" "14" Suc_1 True add_leD1 both_member_options_from_chilf_to_complete_tree length_list_update nth_list_update_eq plus_1_eq_Suc xyprop) next case False have "(treeList[ (high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) ! (high x n) = treeList ! (high x n)" by (metis False nth_list_update_neq) then show ?thesis by (metis "00" "14" One_nat_def Suc_leD \both_member_options (treeList ! high x n) (low x n)\ both_member_options_from_chilf_to_complete_tree length_list_update numeral_2_eq_2 xyprop) qed qed qed qed qed qed lemma post_member_pre_member:"invar_vebt t n \ x< 2^n \ y <2^n \ vebt_member (vebt_insert t x) y \ vebt_member t y \ x = y" proof(induction t n arbitrary: x y rule: invar_vebt.induct) case (1 a b) then show ?case by auto next case (2 treeList n summary m deg) hence "deg \ 2" using deg_not_0 by fastforce then show ?case using vebt_insert.simps(4)[of "deg-2" treeList summary x] by (metis (no_types, lifting) "2.prems"(3) vebt_member.simps(5) add_numeral_left le_add_diff_inverse member_inv numerals(1) plus_1_eq_Suc semiring_norm(2)) next case (3 treeList n summary m deg) hence "deg \ 2" by (metis vebt_member.simps(2) One_nat_def Suc_1 Suc_eq_plus1 add_mono_thms_linordered_semiring(1) vebt_insert.simps(3) le_Suc_eq le_add1 plus_1_eq_Suc) then show ?case using vebt_insert.simps(4)[of "deg-2" treeList summary x] by (metis (no_types, lifting) "3.prems"(3) vebt_member.simps(5) add_numeral_left le_add_diff_inverse member_inv numerals(1) plus_1_eq_Suc semiring_norm(2)) next case (4 treeList n summary m deg mi ma) hence 00:"deg = n+m \ n \ 0 \ n = m \ deg \ 2 \ length treeList =2^n" - by (metis Euclidean_Division.div_eq_0_iff add_self_div_2 deg_not_0 not_less zero_le) + by (metis div_eq_0_iff add_self_div_2 deg_not_0 not_less zero_le) hence xyprop: "high x n < 2^n \ high y n < 2^n" using "4.hyps"(1) "4.prems"(1) "4.prems"(2) deg_not_0 exp_split_high_low(1) by blast have "low x n <2^n \ low y n< 2^n" by (simp add: low_def) then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using "00" "4.prems"(3) insert_simp_mima by auto next case False hence mimaxyprop: "\ (x = mi \ x = ma) \ high x n < 2^n \ high mi n < 2^n \ low x n <2^n \ low mi n < 2^n \ length treeList = 2^n" using "00" "4.hyps"(1) "4.hyps"(7) "4.hyps"(8) \low x n < 2 ^ n \ low y n < 2 ^ n\ deg_not_0 exp_split_high_low(1) exp_split_high_low(2) le_less_trans xyprop by blast then show ?thesis proof(cases "mi < x") case True hence "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[(high x n) :=vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! (high x n)) then vebt_insert summary (high x n) else summary)" using insert_simp_norm[of x n treeList mi ma summary] mimaxyprop "00" add_self_div_2 insert_simp_norm by metis then show ?thesis proof(cases "y = mi \ y = max x ma") case True then show ?thesis proof(cases "y = mi") case True then show ?thesis by (metis "00" vebt_member.simps(5) le0 not_less_eq_eq numeral_2_eq_2 old.nat.exhaust) next case False hence "y = max x ma" using True by blast then show ?thesis proof(cases "x < ma") case True then show ?thesis by (metis (no_types, lifting) "00" vebt_member.simps(5) \y = max x ma\ add_numeral_left le_add_diff_inverse max_less_iff_conj not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False then show ?thesis using \y = max x ma\ by linarith qed qed next case False hence "vebt_member ((treeList[(high x n):= vebt_insert (treeList ! (high x n)) (low x n)]) ! (high y n)) (low y n)" by (metis "4.hyps"(3) "4.hyps"(4) "4.prems"(3) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)\ add_self_div_2 member_inv) then show ?thesis proof(cases "high x n = high y n") case True hence 000:"vebt_member (vebt_insert (treeList ! (high x n)) (low x n)) (low y n)" using \vebt_member (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)] ! high y n) (low y n)\ mimaxyprop by auto have 001:"invar_vebt (treeList ! (high x n)) n \ treeList ! (high x n) \ set treeList " by (simp add: "4.IH"(1) mimaxyprop) hence 002:"vebt_member (treeList ! (high x n)) (low y n) \ low y n = low x n" using "000" "4.IH"(1) \low x n < 2 ^ n \ low y n < 2 ^ n\ by fastforce hence 003:"both_member_options (treeList ! (high x n)) (low y n) \ low y n = low x n" using \invar_vebt (treeList ! high x n) n \ treeList ! high x n \ set treeList\ both_member_options_equiv_member by blast have 004:"naive_member (treeList ! (high x n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" by (metis "00" Suc_le_D True add_self_div_2 mimaxyprop naive_member.simps(3) one_add_one plus_1_eq_Suc) hence 005:"both_member_options (Node (Some (mi , ma)) deg treeList summary) y \ x = y" by (metis "00" "001" "002" Suc_le_D True add_self_div_2 bit_split_inv both_member_options_def member_valid_both_member_options membermima.simps(4) mimaxyprop one_add_one plus_1_eq_Suc) then show ?thesis by (smt "00" "001" "002" "003" "4"(11) "4"(8) vebt_member.simps(5) True add_numeral_left add_self_div_2 bit_split_inv le_add_diff_inverse mimaxyprop not_less not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False hence 000:"vebt_member (treeList ! (high y n)) (low y n)" using \vebt_member (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)] ! high y n) (low y n)\ by auto moreover have 004:"naive_member (treeList ! (high y n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" by (metis "00" Suc_le_D add_self_div_2 naive_member.simps(3) one_add_one plus_1_eq_Suc xyprop) moreover have 001:"invar_vebt (treeList ! (high y n)) n \ treeList ! (high y n) \ set treeList " by (metis (full_types) "4.IH"(1) "4.hyps"(2) "4.hyps"(3) inthall member_def xyprop) moreover have " both_member_options (Node (Some (mi , ma)) deg treeList summary) y" by (metis "00" "000" "001" "004" Suc_le_D add_self_div_2 both_member_options_def member_valid_both_member_options membermima.simps(4) one_add_one plus_1_eq_Suc xyprop) moreover have "vebt_member (Node (Some (mi, ma)) deg treeList summary) y" using both_member_options_equiv_member[of "(Node (Some (mi, ma)) deg treeList summary)" deg y] invar_vebt.intros(4)[of treeList n summary m deg mi ma] using "4" calculation(4) by blast then show ?thesis by simp qed qed next case False hence "x < mi" using mimaxyprop nat_neq_iff by blast hence "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[ (high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) (if minNull (treeList ! (high mi n)) then vebt_insert summary (high mi n) else summary)" using insert_simp_excp[of mi n treeList x ma summary] mimaxyprop "00" add_self_div_2 insert_simp_excp by metis then show ?thesis proof(cases "y = x \ y = max mi ma") case True then show ?thesis proof(cases "y = x") case True then show ?thesis by (simp add: "00") next case False hence "y = max mi ma" using True by blast then show ?thesis proof(cases "mi < ma") case True then show ?thesis using "00" vebt_member.simps(5) \y = max mi ma\ add_numeral_left le_add_diff_inverse max_less_iff_conj not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2) by (metis (no_types, lifting)) next case False then show ?thesis by (metis "00" "4.hyps"(7) vebt_member.simps(5) \y = max mi ma\ add_numeral_left le_add_diff_inverse max.absorb2 numerals(1) plus_1_eq_Suc semiring_norm(2)) qed qed next case False hence "vebt_member ((treeList[(high mi n) :=vebt_insert (treeList ! (high mi n)) (low mi n)]) ! (high y n)) (low y n)" by (metis "4.hyps"(3) "4.hyps"(4) "4.prems"(3) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)]) (if minNull (treeList ! high mi n) then vebt_insert summary (high mi n) else summary)\ add_self_div_2 member_inv) then show ?thesis proof(cases "high mi n = high y n") case True hence 000:"vebt_member (vebt_insert (treeList ! (high mi n)) (low mi n)) (low y n)" by (metis \vebt_member (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)] ! high y n) (low y n)\ mimaxyprop nth_list_update_eq) have 001:"invar_vebt (treeList ! (high mi n)) n \ treeList ! (high mi n) \ set treeList" by (simp add: "4.IH"(1) mimaxyprop) hence 002:"vebt_member (treeList ! (high mi n)) (low y n) \ low y n = low mi n" using "000" "4.IH"(1) \low x n < 2 ^ n \ low y n < 2 ^ n\ mimaxyprop by fastforce hence 003:"both_member_options (treeList ! (high mi n)) (low y n) \ low y n = low mi n" using \invar_vebt (treeList ! high mi n) n \ treeList ! high mi n \ set treeList\ both_member_options_equiv_member by blast have 004:"naive_member (treeList ! (high mi n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" using naive_member.simps(3)[of "Some (mi, ma)" "deg-1" treeList summary y] using "00" True mimaxyprop by fastforce hence 005:"both_member_options (Node (Some (mi , ma)) deg treeList summary) y \ x = y" by (metis "00" "001" "002" Suc_le_D True add_self_div_2 bit_split_inv both_member_options_def member_valid_both_member_options membermima.simps(4) mimaxyprop one_add_one plus_1_eq_Suc) then show ?thesis by (smt "00" "001" "002" "003" "4.hyps"(6) "4.hyps"(9) vebt_member.simps(5) True add_numeral_left add_self_div_2 bit_split_inv le_add_diff_inverse mimaxyprop not_less not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False hence 000:"vebt_member (treeList ! (high y n)) (low y n)" using \vebt_member (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)] ! high y n) (low y n)\ by auto moreover have 004:"naive_member (treeList ! (high y n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" by (metis "00" Suc_le_D add_self_div_2 naive_member.simps(3) one_add_one plus_1_eq_Suc xyprop) moreover have 001:"invar_vebt (treeList ! (high y n)) n \ treeList ! (high y n) \ set treeList " by (metis (full_types) "4.IH"(1) "4.hyps"(2) "4.hyps"(3) inthall member_def xyprop) moreover have " both_member_options (Node (Some (mi , ma)) deg treeList summary) y" by (metis "00" "000" "001" "004" Suc_le_D add_self_div_2 both_member_options_def member_valid_both_member_options membermima.simps(4) one_add_one plus_1_eq_Suc xyprop) then show ?thesis using both_member_options_equiv_member[of "(Node (Some (mi, ma)) deg treeList summary)" deg y] invar_vebt.intros(4)[of treeList n summary m deg mi ma] "4" by blast qed qed qed qed next case (5 treeList n summary m deg mi ma) hence 00:"deg = n+m \ n \ 0 \ Suc n = m \ deg \ 2 \ length treeList =2^m \ n \ 1" by (metis Suc_1 add_mono_thms_linordered_semiring(1) le_add1 plus_1_eq_Suc set_n_deg_not_0 zero_le) hence xyprop: "high x n < 2^m \ high y n < 2^m" using "5.prems"(1) "5.prems"(2) exp_split_high_low(1) by auto have "low x n <2^n \ low y n< 2^n" by (simp add: low_def) then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using "00" "5.prems"(3) insert_simp_mima by auto next case False hence mimaxyprop: "\ (x = mi \ x = ma) \ high x n < 2^m \ high mi n < 2^m \ low x n <2^n \ low mi n < 2^n \ length treeList = 2^m" using "00" "5" \low x n < 2 ^ n \ low y n < 2 ^ n\ deg_not_0 exp_split_high_low(1) exp_split_high_low(2) le_less_trans xyprop by (smt less_le_trans less_numeral_extra(1)) then show ?thesis proof(cases "mi < x") case True hence "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) (if minNull (treeList ! (high x n)) then vebt_insert summary (high x n) else summary)" using insert_simp_norm[of x deg treeList mi ma summary] by (smt "00" False add_Suc_right add_self_div_2 even_Suc_div_two odd_add xyprop) then show ?thesis proof(cases "y = mi \ y = max x ma") case True then show ?thesis proof(cases "y = mi") case True then show ?thesis by (metis "00" vebt_member.simps(5) le0 not_less_eq_eq numeral_2_eq_2 old.nat.exhaust) next case False hence "y = max x ma" using True by blast then show ?thesis proof(cases "x < ma") case True then show ?thesis by (metis (no_types, lifting) "00" vebt_member.simps(5) \y = max x ma\ add_numeral_left le_add_diff_inverse max_less_iff_conj not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False then show ?thesis using \y = max x ma\ by linarith qed qed next case False hence "vebt_member ((treeList[ (high x n):=vebt_insert (treeList ! (high x n)) (low x n)]) ! (high y n)) (low y n)" using "5.hyps"(3) "5.hyps"(4) "5.prems"(3) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)]) (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary)\ add_Suc_right add_self_div_2 member_inv by force then show ?thesis proof(cases "high x n = high y n") case True hence 000:"vebt_member (vebt_insert (treeList ! (high x n)) (low x n)) (low y n)" by (metis "5.hyps"(2) \vebt_member (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)] ! high y n) (low y n)\ nth_list_update_eq xyprop) have 001:"invar_vebt (treeList ! (high x n)) n \ treeList ! (high x n) \ set treeList " by (simp add: "5.IH"(1) "5.hyps"(2) xyprop) hence 002:"vebt_member (treeList ! (high x n)) (low y n) \ low y n = low x n" using "000" "5.IH"(1) \low x n < 2 ^ n \ low y n < 2 ^ n\ by fastforce hence 003:"both_member_options (treeList ! (high x n)) (low y n) \ low y n = low x n" using \invar_vebt (treeList ! high x n) n \ treeList ! high x n \ set treeList\ both_member_options_equiv_member by blast have 004:"naive_member (treeList ! (high x n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" using "00" True xyprop by auto hence 005:"both_member_options (Node (Some (mi , ma)) deg treeList summary) y \ x = y" by (smt "00" "001" "002" True add_Suc_right add_self_div_2 bit_split_inv both_member_options_def even_Suc_div_two member_valid_both_member_options membermima.simps(4) odd_add xyprop) then show ?thesis using both_member_options_equiv_member[of "(Node (Some (mi, ma)) deg treeList summary)" deg y] invar_vebt.intros(5)[of treeList n summary m deg mi ma] "5" by blast next case False hence 000:"vebt_member (treeList ! (high y n)) (low y n)" using \vebt_member (treeList[high x n := vebt_insert (treeList ! high x n) (low x n)] ! high y n) (low y n)\ by fastforce moreover have 004:"naive_member (treeList ! (high y n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" using "00" xyprop by auto moreover have 001:"invar_vebt (treeList ! (high y n)) n \ treeList ! (high y n) \ set treeList " by (metis (full_types) "5"inthall member_def xyprop) moreover have " both_member_options (Node (Some (mi , ma)) deg treeList summary) y" using "00" "000" "001" both_member_options_def member_valid_both_member_options xyprop by fastforce moreover have "vebt_member (Node (Some (mi, ma)) deg treeList summary) y" using both_member_options_equiv_member[of "(Node (Some (mi, ma)) deg treeList summary)" deg y] invar_vebt.intros(5)[of treeList n summary m deg mi ma] "5" calculation(4) by blast then show ?thesis by simp qed qed next case False hence "x < mi" using mimaxyprop nat_neq_iff by blast hence "vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[ (high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) (if minNull (treeList ! (high mi n)) then vebt_insert summary (high mi n) else summary)" using insert_simp_excp[of mi n treeList x ma summary] mimaxyprop "00" add_self_div_2 insert_simp_excp by (smt add_Suc_right even_Suc_div_two odd_add) then show ?thesis proof(cases "y = x \ y = max mi ma") case True then show ?thesis proof(cases "y = x") case True then show ?thesis by (simp add: "00") next case False hence "y = max mi ma" using True by blast then show ?thesis proof(cases "mi < ma") case True then show ?thesis using "00" vebt_member.simps(5) \y = max mi ma\ add_numeral_left le_add_diff_inverse max_less_iff_conj not_less_iff_gr_or_eq numerals(1) plus_1_eq_Suc semiring_norm(2) by (metis (no_types, lifting)) next case False then show ?thesis by (metis "00" "5.hyps"(7) vebt_member.simps(5) \y = max mi ma\ add_numeral_left le_add_diff_inverse max.absorb2 numerals(1) plus_1_eq_Suc semiring_norm(2)) qed qed next case False hence "vebt_member ((treeList[(high mi n):=vebt_insert (treeList ! (high mi n)) (low mi n)]) ! (high y n)) (low y n)" using "5.hyps"(3) "5.hyps"(4) "5.prems"(3) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)]) (if minNull (treeList ! high mi n) then vebt_insert summary (high mi n) else summary)\ member_inv by force then show ?thesis proof(cases "high mi n = high y n") case True hence 000:"vebt_member (vebt_insert (treeList ! (high mi n)) (low mi n)) (low y n)" by (metis \vebt_member (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)] ! high y n) (low y n)\ mimaxyprop nth_list_update_eq) have 001:"invar_vebt (treeList ! (high mi n)) n \ treeList ! (high mi n) \ set treeList " by (simp add: "5.IH"(1) mimaxyprop) hence 002:"vebt_member (treeList ! (high mi n)) (low y n) \ low y n = low mi n" using "000" "5.IH"(1) \low x n < 2 ^ n \ low y n < 2 ^ n\ mimaxyprop by fastforce hence 003:"both_member_options (treeList ! (high mi n)) (low y n) \ low y n = low mi n" using \invar_vebt (treeList ! high mi n) n \ treeList ! high mi n \ set treeList\ both_member_options_equiv_member by blast have 004:"naive_member (treeList ! (high mi n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" using naive_member.simps(3)[of "Some (mi, ma)" "deg-1" treeList summary y] using "00" True mimaxyprop by fastforce hence 005:"both_member_options (Node (Some (mi , ma)) deg treeList summary) y \ x = y" by (smt "00" "001" "002" True add_Suc_right add_self_div_2 bit_split_inv both_member_options_def even_Suc_div_two member_valid_both_member_options membermima.simps(4) odd_add xyprop) then show ?thesis using "00" "001" "002" "003" "5"(14) "5.hyps"(6) "5.hyps"(7) "5.hyps"(9) vebt_member.simps(5) True add_Suc_right add_self_div_2 bit_split_inv even_Suc_div_two le_add_diff_inverse max.absorb2 mimaxyprop not_less_iff_gr_or_eq odd_add plus_1_eq_Suc by (smt (z3) \vebt_insert (Node (Some (mi, ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)]) (if minNull (treeList ! high mi n) then vebt_insert summary (high mi n) else summary)\) next case False hence 000:"vebt_member (treeList ! (high y n)) (low y n)" using \vebt_member (treeList[high mi n := vebt_insert (treeList ! high mi n) (low mi n)] ! high y n) (low y n)\ by auto moreover have 004:"naive_member (treeList ! (high y n)) (low y n) \ naive_member (Node (Some (mi , ma)) deg treeList summary) y" using "00" xyprop by auto moreover have 001:"invar_vebt (treeList ! (high y n)) n \ treeList ! (high y n) \ set treeList " by (metis (full_types) "5.IH"(1) "5.hyps"(2) "5.hyps"(3) inthall member_def xyprop) moreover have " both_member_options (Node (Some (mi , ma)) deg treeList summary) y" using "00" "000" "001" both_member_options_def member_valid_both_member_options xyprop by fastforce then show ?thesis using both_member_options_equiv_member[of "(Node (Some (mi, ma)) deg treeList summary)" deg y] invar_vebt.intros(5)[of treeList n summary m deg mi ma] "5" by simp qed qed qed qed qed end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_InsertCorrectness.thy b/thys/Van_Emde_Boas_Trees/VEBT_InsertCorrectness.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_InsertCorrectness.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_InsertCorrectness.thy @@ -1,1091 +1,1091 @@ (*by Ammer*) theory VEBT_InsertCorrectness imports VEBT_Member VEBT_Insert begin context VEBT_internal begin section \Correctness of the Insert Operation\ subsection \Validness Preservation\ theorem valid_pres_insert: "invar_vebt t n \ x< 2^n \ invar_vebt (vebt_insert t x) n" proof(induction t n arbitrary: x rule: invar_vebt.induct) case (1 a b) then show ?case using vebt_insert.simps(1)[of a b x] by (simp add: invar_vebt.intros(1)) next case (2 treeList n summary m deg) hence 0: " ( \ t \ set treeList. invar_vebt t n) " and 1:" invar_vebt summary n" and 2:" length treeList = 2^n" and 3:" deg = 2*n" and 4:" (\ i. both_member_options summary i)" and 5:" (\ t \ set treeList. \ x. both_member_options t x) " and 6: "n\ 1" using "2.prems" by (auto simp add: Suc_leI deg_not_0) let ?t = "Node None deg treeList summary" let ?tnew = "vebt_insert ?t x" have 6:"?tnew = (Node (Some (x,x)) deg treeList summary)" using vebt_insert.simps(4)[of "deg-2" treeList summary x] by (metis "1" "2.hyps"(3) "2.hyps"(4) add_2_eq_Suc add_diff_inverse_nat add_self_div_2 deg_not_0 div_less gr_implies_not0) have 7:"(x = x \ (\ t \ set treeList. \ x. both_member_options t x))" using \\t\set treeList. \x. both_member_options t x\ by blast have 8:"x \ x" by simp have 9:" x < 2^deg" by (simp add: "2.prems") have 10:"(x \ x \ (\ i < 2^(2^n). (high x deg = i \ both_member_options (treeList ! i) (low x deg)) \ (\ y. (high y deg = i \ both_member_options (treeList ! i) (low y deg) ) \ x < y \ y \ x) ))" by simp then show ?case using 0 1 2 3 4 5 6 7 8 9 10 invar_vebt.intros(4)[of treeList n summary m deg x x] by (metis "2.hyps"(3) "2.hyps"(4) nth_mem) next case (3 treeList n summary m deg) hence 0: " ( \ t \ set treeList. invar_vebt t n) " and 1:" invar_vebt summary m" and 2:" length treeList = 2^m" and 3:" deg = n+m" and 4:" (\ i. both_member_options summary i)" and 5:" (\ t \ set treeList. \ x. both_member_options t x) " and 6: "n\ 1" and 7: "Suc n = m" using "3.prems" apply auto by (metis "3.hyps"(2) One_nat_def set_n_deg_not_0) let ?t = "Node None deg treeList summary" let ?tnew = "vebt_insert ?t x" have 6:"?tnew = (Node (Some (x,x)) deg treeList summary)" using vebt_insert.simps(4)[of "deg-2" treeList summary x] by (smt "3" "3.hyps"(3) "6" Nat.add_diff_assoc One_nat_def Suc_le_mono add_diff_inverse_nat add_gr_0 add_numeral_left diff_is_0_eq' not_less not_less_iff_gr_or_eq numeral_2_eq_2 numerals(1) plus_1_eq_Suc semiring_norm(2)) have 7:"(x = x \ (\ t \ set treeList. \ x. both_member_options t x))" using \\t\set treeList. \x. both_member_options t x\ by blast have 8:"x \ x" by simp have 9:" x < 2^deg" by (simp add: "3.prems") have 10:"(x \ x \ (\ i < 2^(2^n). (high x deg = i \ both_member_options (treeList ! i) (low x deg)) \ (\ y. (high y deg = i \ both_member_options (treeList ! i) (low y deg) ) \ x < y \ y \ x) ))" by simp then show ?case using 0 1 2 3 4 5 6 7 8 9 10 invar_vebt.intros(5)[of treeList n summary m deg x x] "3.hyps"(3) nth_mem by force next case (4 treeList n summary m deg mi ma) hence myIHs: "x \ set treeList \ invar_vebt x n \ xa < 2 ^ n \ invar_vebt (vebt_insert x xa) n" for x xa by simp hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary m" and 2:"length treeList = 2^m" and 3:" deg = n+m" and 4: "(\ i < 2^m. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" and 8: "n = m" and 9: "deg div 2 = n" using "4" add_self_div_2 by blast+ then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using insert_simp_mima[of x mi ma deg treeList summary] invar_vebt.intros(4)[of treeList n summary m deg mi ma] by (smt "0" "1" "2" "3" "4" "4.hyps"(3) "4.hyps"(7) "4.hyps"(8) "5" "7" "9" deg_not_0 div_greater_zero_iff) next case False hence mimaxrel: "x \ mi \ x \ ma" by simp then show ?thesis proof(cases "mi < x") case True hence abcdef: "mi < x" by simp let ?h = "high x n" and ?l="low x n" have highlowprop: "high x n < 2^m \ low x n < 2^n" using "1" "3" "4.hyps"(3) "4.prems" deg_not_0 exp_split_high_low(1) exp_split_high_low(2) by blast have 10:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[?h:=vebt_insert (treeList ! ?h) ?l]) (if minNull (treeList ! ?h) then vebt_insert summary ?h else summary) " using "2" "3" False True \high x n < 2 ^ m \ low x n < 2 ^ n\ insert_simp_norm by (metis "1" "4.hyps"(3) "9" deg_not_0 div_greater_zero_iff) let ?maxnew = "max x ma" and ?nextTreeList = "(take ?h treeList @ [vebt_insert (treeList ! ?h) ?l] @ drop (?h+1) treeList)" and ?nextSummary = "(if minNull (treeList ! ?h) then vebt_insert summary ?h else summary)" have 11: "( \ t \ set ?nextTreeList. invar_vebt t n)" proof fix t assume " t \ set ?nextTreeList" hence 111:"t \ set (take ?h treeList) \ t \ set ([vebt_insert (treeList ! ?h) ?l] @ drop (?h+1) treeList)" by auto show "invar_vebt t n" proof(cases "t \ set (take ?h treeList) ") case True then show ?thesis by (meson "0" in_set_takeD) next case False hence 1110: "t = vebt_insert (treeList ! ?h) ?l \ t \ set ( drop (?h+1) treeList)" using "111" by auto then show ?thesis proof(cases "t = vebt_insert (treeList ! ?h) ?l") case True have 11110: "invar_vebt (treeList ! ?h) n" by (simp add: "2" "4.IH"(1) highlowprop) have 11111: "?l < 2^n" by (simp add: low_def) then show ?thesis using myIHs[of "treeList ! ?h"] by (simp add: "11110" "2" True highlowprop) next case False then show ?thesis by (metis "0" "1110" append_assoc append_take_drop_id in_set_conv_decomp) qed qed qed have 12: "invar_vebt ?nextSummary n" proof(cases "minNull (treeList ! high x n)") case True then show ?thesis using "4.IH"(2) "8" highlowprop by auto next case False then show ?thesis by (simp add: "1" "8") qed have 13: "\ i < 2^m. (\ y. both_member_options (?nextTreeList ! i) y) \ ( both_member_options ?nextSummary i)" proof fix i show "i < 2 ^ m \ (\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof assume "i< 2^m" show "(\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof(cases "minNull (treeList ! high x n)") case True hence tc: "minNull (treeList ! high x n)" by simp hence nsprop: "?nextSummary = vebt_insert summary ?h" by simp have insprop:"?nextTreeList ! ?h = vebt_insert (treeList ! ?h) ?l" by (metis "2" Suc_eq_plus1 append_Cons highlowprop nth_list_update_eq self_append_conv2 upd_conv_take_nth_drop) then show ?thesis proof(cases "i = ?h") case True have 161:"\ y. vebt_member (treeList ! ?h) y" by (simp add: min_Null_member tc) hence 162:"\ y. both_member_options (treeList ! ?h) y" by (metis "2" "4.IH"(1) highlowprop nth_mem valid_member_both_member_options) hence 163:"\ both_member_options summary i" using "11" "2" "4" True \i < 2 ^ m\ by blast have 164:"?nextTreeList ! i = vebt_insert (treeList ! ?h) ?l" using True insprop by simp have 165:"invar_vebt (vebt_insert (treeList ! ?h) ?l) n" by (simp add: "11") have 166:"both_member_options (vebt_insert (treeList ! ?h) ?l) ?l" using myIHs[of "treeList ! ?h" ?l] by (metis "0" "2" highlowprop nth_mem valid_insert_both_member_options_add) have 167:"\ y. both_member_options ((?nextTreeList) ! i) y " using "164" "166" by auto then show ?thesis using "1" "11" "2" True nsprop valid_insert_both_member_options_add highlowprop by auto next case False have "?nextTreeList ! i = treeList ! i" by (metis "2" False \i < 2 ^ m\ highlowprop nth_repl) have fstprop:"both_member_options ((?nextTreeList) ! i) y \ both_member_options (?nextSummary) i " for y using "1" "4" \(take (high x n) treeList @ [VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] @ drop (high x n + 1) treeList) ! i = treeList ! i\ \i < 2 ^ m\ highlowprop valid_insert_both_member_options_pres by auto moreover have" both_member_options (?nextSummary) i \ \ y . both_member_options ((?nextTreeList) ! i) y " proof- assume "both_member_options (?nextSummary) i " have "i \ high x n" by (simp add: False) hence "both_member_options summary i" by (smt (z3) "1" "12" \both_member_options (if minNull (treeList ! high x n) then VEBT_Insert.vebt_insert summary (high x n) else summary) i\ \i < 2 ^ m\ both_member_options_equiv_member highlowprop post_member_pre_member) hence "\ y. both_member_options (treeList ! i) y" by (simp add: "4" \i < 2 ^ m\) then show ?thesis using \(take (high x n) treeList @ [VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] @ drop (high x n + 1) treeList) ! i = treeList ! i\ by presburger qed ultimately show ?thesis by auto qed next case False hence "?nextSummary = summary" by simp hence "\ y. both_member_options (treeList ! high x n) y" using not_min_Null_member False by blast hence "both_member_options summary (high x n)" using "4" highlowprop by blast hence " both_member_options (?nextTreeList ! high x n) ?l" by (metis "0" "2" Suc_eq_plus1 append_Cons append_Nil highlowprop nth_list_update_eq nth_mem upd_conv_take_nth_drop valid_insert_both_member_options_add) then show ?thesis by (smt (verit, best) "2" "4" False \both_member_options summary (high x n)\ \i < 2 ^ m\ highlowprop nth_repl) qed qed qed have 14: "(mi = max ma x \ (\ t \ set ?nextTreeList. \ y. both_member_options t y))" using True max_less_iff_conj by blast have 15: "mi \ max ma x \ max ma x < 2^deg" using "4.hyps"(8) "4.prems" abcdef by auto have 16: "(mi \ max ma x \ (\ i < 2^m. (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ mi < y \ y \ max ma x)))" proof assume "mi \ max ma x" show "(\ i < 2^m. (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ mi < y \ y \ max ma x))" proof fix i::nat show "i < 2 ^ m\ (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof assume "i < 2^m" show " (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof show "(high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n))" proof assume "high (max ma x) n = i" show "both_member_options (?nextTreeList ! i) (low (max ma x) n)" proof(cases "high x n = high ma n") case True have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence "?nextTreeList ! i = vebt_insert (treeList ! i) (low x n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] "2" True \high (max ma x) n = i\ \i < 2 ^ m\ concat_inth length_take max_def by (metis Suc_eq_plus1 append_Cons append_Nil nth_list_update_eq upd_conv_take_nth_drop) hence "vebt_member (?nextTreeList ! i) (low x n)" using Un_iff \i < 2 ^ m\ \invar_vebt (treeList ! i) n\ both_member_options_equiv_member highlowprop list.set_intros(1) set_append valid_insert_both_member_options_add by (metis "11" True \high (max ma x) n = i\ max_def) then show ?thesis proof(cases "mi = ma") case True then show ?thesis by (metis \(take (high x n) treeList @ [VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] @ drop (high x n + 1) treeList) ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \mi \ max ma x\ \invar_vebt (treeList ! i) n\ highlowprop max_def valid_insert_both_member_options_add) next case False hence "vebt_member (treeList ! i) (low ma n)" by (metis "7" True \high (max ma x) n = i\ \invar_vebt (treeList ! i) n\ both_member_options_equiv_member highlowprop linorder_cases max.absorb3 max.absorb4 mimaxrel) hence "vebt_member (?nextTreeList ! i) (low ma n) \ (low ma n = low x n)" using post_member_pre_member[of " (treeList ! i)" n "low x n" "low ma n" ] by (metis "2" "4.IH"(1) \(take (high x n) treeList @ [VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] @ drop (high x n + 1) treeList) ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \i < 2 ^ m\ both_member_options_equiv_member highlowprop member_bound nth_mem valid_insert_both_member_options_pres) then show ?thesis by (metis "2" "4.IH"(1) True \(take (high x n) treeList @ [VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] @ drop (high x n + 1) treeList) ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \high (max ma x) n = i\ both_member_options_equiv_member highlowprop max_def nth_mem valid_insert_both_member_options_add) qed next case False then show ?thesis proof(cases "x < ma") case True then show ?thesis by (metis "2" "7" False \high (max ma x) n = i\ \i < 2 ^ m\ abcdef highlowprop less_trans max.strict_order_iff nth_repl) next case False hence "x > ma" using mimaxrel nat_neq_iff by blast then show ?thesis by (metis "2" "4.IH"(1) One_nat_def \high (max ma x) n = i\ add.right_neutral add_Suc_right append_Cons highlowprop max.commute max.strict_order_iff nth_list_update_eq nth_mem self_append_conv2 upd_conv_take_nth_drop valid_insert_both_member_options_add) qed qed qed show "(\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof fix y show "high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x" proof assume bb:"high y n = i \ both_member_options (?nextTreeList ! i) (low y n)" show " mi < y \ y \ max ma x" proof(cases "i = high x n") case True hence cc:" i = high x n" by simp have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = vebt_insert (treeList ! i) (low x n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (metis "2" Suc_eq_plus1 append_Cons append_self_conv2 cc highlowprop nth_list_update_eq upd_conv_take_nth_drop) hence "invar_vebt (?nextTreeList ! i) n" by (simp add: "11" True) hence "vebt_member (treeList ! i) (low y n) \ (low y n) = (low x n)" by (metis \invar_vebt (treeList ! i) n\ aa bb highlowprop member_bound post_member_pre_member valid_member_both_member_options) then show ?thesis proof(cases "low y n = low x n") case True hence "high x n = high y n \ low y n = low x n" by (simp add: bb cc) hence "x = y" by (metis bit_split_inv) then show ?thesis using abcdef by auto next case False hence "vebt_member (treeList ! i) (low y n)" using \vebt_member (treeList ! i) (low y n) \ low y n = low x n\ by blast hence "mi \ ma " using 5 inthall by (metis "2" \i < 2 ^ m\ min_Null_member not_min_Null_member) then show ?thesis using "7" \i < 2 ^ m\ \vebt_member (treeList ! i) (low y n)\ \invar_vebt (treeList ! i) n\ bb both_member_options_equiv_member max.coboundedI1 by blast qed next case False have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = (treeList ! i)" by (metis "2" False \i < 2 ^ m\ highlowprop nth_repl) hence "both_member_options (treeList !i) (low y n)" using bb by auto hence "mi \ ma " using 5 "2" \i < 2 ^ m\ by force then show ?thesis using 7 using \both_member_options (treeList ! i) (low y n)\ \i < 2 ^ m\ bb max.coboundedI1 by blast qed qed qed qed qed qed qed then show ?thesis using invar_vebt.intros(4)[of ?nextTreeList n ?nextSummary m deg mi "max ma x"] by (smt (z3) "10" "11" "12" "13" "15" "2" "3" "8" One_nat_def abcdef add.right_neutral add_Suc_right append_Cons highlowprop leD max.cobounded2 max.commute pos_n_replace self_append_conv2 upd_conv_take_nth_drop) next case False hence abcdef: "x < mi" using antisym_conv3 mimaxrel by blast let ?h = "high mi n" and ?l="low mi n" have highlowprop: "high mi n < 2^m \ low mi n < 2^n" using "1" "3" "4.hyps"(3) "4.hyps"(7) "4.hyps"(8) deg_not_0 exp_split_high_low(1) exp_split_high_low(2) le_less_trans by blast have 10:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[?h:=vebt_insert (treeList ! ?h) ?l]) (if minNull (treeList ! ?h) then vebt_insert summary ?h else summary) " by (metis "1" "2" "4.hyps"(3) "9" abcdef deg_not_0 div_greater_zero_iff highlowprop insert_simp_excp mimaxrel) let ?maxnew = "max mi ma" and ?nextTreeList = "(treeList[ ?h :=vebt_insert (treeList ! ?h) ?l])" and ?nextSummary = "(if minNull (treeList ! ?h) then vebt_insert summary ?h else summary)" have 11: "( \ t \ set ?nextTreeList. invar_vebt t n)" proof fix t assume " t \ set ?nextTreeList" then obtain i where "?nextTreeList ! i = t \ i < 2^m" by (metis "2" in_set_conv_nth length_list_update) show "invar_vebt t n" by (metis "2" "4.IH"(1) \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = t \ i < 2 ^ m\ highlowprop nth_list_update_eq nth_list_update_neq nth_mem) qed have 12: "invar_vebt ?nextSummary n" using "1" "4.IH"(2) "8" highlowprop by presburger have 13: "\ i < 2^m. (\ y. both_member_options (?nextTreeList ! i) y) \ ( both_member_options ?nextSummary i)" proof fix i show "i < 2 ^ m \ (\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof assume "i< 2^m" show "(\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof(cases "minNull (treeList ! high mi n)") case True hence tc: "minNull (treeList ! high mi n)" by simp hence nsprop: "?nextSummary = vebt_insert summary ?h" by simp have insprop:"?nextTreeList ! ?h = vebt_insert (treeList ! ?h) ?l" by (simp add: "2" highlowprop) then show ?thesis proof(cases "i = ?h") case True have 161:"\ y. vebt_member (treeList ! ?h) y" by (simp add: min_Null_member tc) hence 162:"\ y. both_member_options (treeList ! ?h) y" by (metis "2" "4.IH"(1) highlowprop nth_mem valid_member_both_member_options) hence 163:"\ both_member_options summary i" using "11" "2" "4" True \i < 2 ^ m\ by blast have 164:"?nextTreeList ! i = vebt_insert (treeList ! ?h) ?l" using True insprop by simp have 165:"invar_vebt (vebt_insert (treeList ! ?h) ?l) n" by (simp add: "2" "4.IH"(1) highlowprop) have 166:"both_member_options (vebt_insert (treeList ! ?h) ?l) ?l" using myIHs[of "treeList ! ?h" ?l] by (metis "0" "2" highlowprop in_set_member inthall valid_insert_both_member_options_add) have 167:"\ y. both_member_options ((?nextTreeList) ! i) y " using "164" "166" by auto then show ?thesis using "1" "11" "2" True nsprop valid_insert_both_member_options_add highlowprop by auto next case False have "?nextTreeList ! i = treeList ! i" using False by fastforce have fstprop:"both_member_options ((?nextTreeList) ! i) y \ both_member_options (?nextSummary) i " for y using "1" "4" \i < 2 ^ m\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\ highlowprop valid_insert_both_member_options_pres by auto moreover have" both_member_options (?nextSummary) i \ \ y . both_member_options ((?nextTreeList) ! i) y " proof- assume "both_member_options (?nextSummary) i " have "i \ high mi n" by (simp add: False) hence "both_member_options summary i" by (smt (z3) "1" "12" \both_member_options (if minNull (treeList ! high mi n) then VEBT_Insert.vebt_insert summary (high mi n) else summary) i\ \i < 2 ^ m\ both_member_options_equiv_member highlowprop post_member_pre_member) hence "\ y. both_member_options (treeList ! i) y" by (simp add: "4" \i < 2 ^ m\) then show ?thesis by (simp add: \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\) qed ultimately show ?thesis by auto qed next case False hence "?nextSummary = summary" by simp hence "\ y. both_member_options (treeList ! high mi n) y" using not_min_Null_member False by blast hence "both_member_options summary (high mi n)" using "4" highlowprop by blast hence " both_member_options (?nextTreeList ! high mi n) ?l" by (metis "0" "2" highlowprop nth_list_update_eq nth_mem valid_insert_both_member_options_add) then show ?thesis by (metis (full_types, opaque_lifting) "4" False \both_member_options summary (high mi n)\ \i < 2 ^ m\ nth_list_update_neq) qed qed qed have 14: "(x = max ma mi \ (\ t \ set ?nextTreeList. \ y. both_member_options t y))" using mimaxrel by linarith have 15: "x \ max ma mi \ max ma mi < 2^deg" using "6" abcdef by linarith have 16: "(x \ max ma mi \ (\ i < 2^m. (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ x < y \ y \ max ma mi)))" proof assume "x \ max ma mi" show "(\ i < 2^m. (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ x < y \ y \ max ma mi))" proof fix i::nat show "i < 2 ^ m\ (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof assume "i < 2^m" show " (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof show "(high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n))" proof assume "high (max ma mi) n = i" show "both_member_options (?nextTreeList ! i) (low (max ma mi) n)" proof(cases "high mi n = high ma n") case True have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence "?nextTreeList ! i = vebt_insert (treeList ! i) (low mi n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (metis "2" True \high (max ma mi) n = i\ highlowprop max_def nth_list_update_eq) hence "vebt_member (?nextTreeList ! i) (low mi n)" by (metis "11" "2" True \high (max ma mi) n = i\ \invar_vebt (treeList ! i) n\ highlowprop max_def set_update_memI valid_insert_both_member_options_add valid_member_both_member_options) then show ?thesis proof(cases "mi = ma") case True then show ?thesis using \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ \invar_vebt (treeList ! i) n\ highlowprop valid_insert_both_member_options_add by force next case False hence "vebt_member (treeList ! i) (low ma n)" using "6" "7" \high (max ma mi) n = i\ \i < 2 ^ m\ \invar_vebt (treeList ! i) n\ both_member_options_equiv_member by auto hence "vebt_member (?nextTreeList ! i) (low ma n) \ (low ma n = low mi n)" using post_member_pre_member[of " (treeList ! i)" n "low mi n" "low ma n" ] by (metis "11" "2" "7" True \high (max ma mi) n = i\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ \invar_vebt (treeList ! i) n\ highlowprop max_def member_bound set_update_memI valid_insert_both_member_options_pres valid_member_both_member_options) then show ?thesis by (metis "11" "2" "4.hyps"(7) "7" False True \high (max ma mi) n = i\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ both_member_options_equiv_member highlowprop less_irrefl max.commute max_def set_update_memI) qed next case False hence "?nextTreeList ! i = treeList ! i" by (metis "4.hyps"(7) \high (max ma mi) n = i\ max.commute max_def nth_list_update_neq) then show ?thesis by (metis "4.hyps"(7) "7" False \high (max ma mi) n = i\ \i < 2 ^ m\ max.orderE) qed qed show "(\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof fix y show "high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi" proof assume bb:"high y n = i \ both_member_options (?nextTreeList ! i) (low y n)" show " x < y \ y \ max ma mi" proof(cases "i = high mi n") case True hence cc:" i = high mi n" by simp have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = vebt_insert (treeList ! i) (low mi n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (simp add: cc highlowprop) hence "invar_vebt (?nextTreeList ! i) n" by (simp add: "2" "4.IH"(1) cc highlowprop) hence "vebt_member (treeList ! i) (low y n) \ (low y n) = (low mi n)" by (metis \invar_vebt (treeList ! i) n\ aa bb both_member_options_equiv_member highlowprop member_bound post_member_pre_member) then show ?thesis proof(cases "low y n = low mi n") case True hence "high mi n = high y n \ low y n = low mi n" by (simp add: bb cc) hence "mi = y" by (metis bit_split_inv) then show ?thesis using abcdef by auto next case False hence "vebt_member (treeList ! i) (low y n)" using \vebt_member (treeList ! i) (low y n) \ low y n = low mi n\ by blast hence "mi \ ma " using 5 inthall by (metis "2" \i < 2 ^ m\ min_Null_member not_min_Null_member) then show ?thesis using "7" by (metis \i < 2 ^ m\ \vebt_member (treeList ! i) (low y n)\ \invar_vebt (treeList ! i) n\ abcdef bb both_member_options_equiv_member max.absorb1 max.strict_order_iff max_less_iff_conj) qed next case False have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = (treeList ! i)" using False by auto hence "both_member_options (treeList ! i) (low y n)" using bb by auto hence "mi \ ma " using 5 "2" \i < 2 ^ m\ by force then show ?thesis using 7 by (metis \both_member_options (treeList ! i) (low y n)\ \i < 2 ^ m\ abcdef bb max.absorb1 max.strict_order_iff max_less_iff_conj) qed qed qed qed qed qed qed then show ?thesis using invar_vebt.intros(4)[of ?nextTreeList n ?nextSummary m deg x "max ma mi"] by (smt (z3) "10" "11" "12" "13" "14" "15" "2" "3" "4.hyps"(3) "4.hyps"(7) length_list_update max.absorb1 max.absorb2) qed qed next case (5 treeList n summary m deg mi ma) hence myIHs: "x \ set treeList \ invar_vebt x n \ xa < 2 ^ n \ invar_vebt (vebt_insert x xa) n" for x xa by simp hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary m" and 2:"length treeList = 2^m" and 3:" deg = n+m" and 4: "(\ i < 2^m. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" and 8: "Suc n = m" and 9: "deg div 2 = n" using "5" add_self_div_2 apply blast+ by (simp add: "5.hyps"(3) "5.hyps"(4)) then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis using insert_simp_mima[of x mi ma deg treeList summary] invar_vebt.intros(5)[of treeList n summary m deg mi ma] by (smt "0" "1" "2" "3" "4" "5" "5.hyps"(3) "5.hyps"(7) "5.hyps"(8) "7" "9" div_less not_less not_one_le_zero set_n_deg_not_0) next case False hence mimaxrel: "x \ mi \ x \ ma" by simp then show ?thesis proof(cases "mi < x") case True hence abcdef: "mi < x" by simp let ?h = "high x n" and ?l="low x n" have highlowprop: "high x n < 2^m \ low x n < 2^n" - by (metis "1" "2" "3" "5.IH"(1) "5.prems" Euclidean_Division.div_eq_0_iff add_nonneg_eq_0_iff deg_not_0 div_exp_eq exp_split_high_low(2) high_def not_one_le_zero one_add_one power_not_zero set_n_deg_not_0 zero_le_one zero_neq_one) + by (metis "1" "2" "3" "5.IH"(1) "5.prems" div_eq_0_iff add_nonneg_eq_0_iff deg_not_0 div_exp_eq exp_split_high_low(2) high_def not_one_le_zero one_add_one power_not_zero set_n_deg_not_0 zero_le_one zero_neq_one) have 10:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (mi, max x ma)) deg (treeList[?h :=vebt_insert (treeList ! ?h) ?l]) (if minNull (treeList ! ?h) then vebt_insert summary ?h else summary) " using "2" "3" False True \high x n < 2 ^ m \ low x n < 2 ^ n\ insert_simp_norm by (smt "5.IH"(1) "9" div_greater_zero_iff div_if less_Suc_eq_0_disj not_one_le_zero set_n_deg_not_0) let ?maxnew = "max x ma" and ?nextTreeList = "(treeList[ ?h :=vebt_insert (treeList ! ?h) ?l])" and ?nextSummary = "(if minNull (treeList ! ?h) then vebt_insert summary ?h else summary)" have 11: "( \ t \ set ?nextTreeList. invar_vebt t n)" proof fix t assume " t \ set ?nextTreeList" then obtain i where "i <2^m \ ?nextTreeList ! i = t" by (metis "2" in_set_conv_nth length_list_update) show "invar_vebt t n" by (metis "2" "5.IH"(1) \i < 2 ^ m \ treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = t\ highlowprop nth_list_update_eq nth_list_update_neq nth_mem) qed have 12: "invar_vebt ?nextSummary m" by (simp add: "1" "5.IH"(2) highlowprop) have 13: "\ i < 2^m. (\ y. both_member_options (?nextTreeList ! i) y) \ ( both_member_options ?nextSummary i)" proof fix i show "i < 2 ^ m \ (\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof assume "i< 2^m" show "(\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof(cases "minNull (treeList ! high x n)") case True hence tc: "minNull (treeList ! high x n)" by simp hence nsprop: "?nextSummary = vebt_insert summary ?h" by simp have insprop:"?nextTreeList ! ?h = vebt_insert (treeList ! ?h) ?l" by (simp add: "2" highlowprop) then show ?thesis proof(cases "i = ?h") case True have 161:"\ y. vebt_member (treeList ! ?h) y" by (simp add: min_Null_member tc) hence 162:"\ y. both_member_options (treeList ! ?h) y" by (metis "0" "2" highlowprop nth_mem valid_member_both_member_options) hence 163:"\ both_member_options summary i" using "11" "2" "4" True \i < 2 ^ m\ by blast have 164:"?nextTreeList ! i = vebt_insert (treeList ! ?h) ?l" using True insprop by simp have 165:"invar_vebt (vebt_insert (treeList ! ?h) ?l) n" by (simp add: "11" "2" highlowprop set_update_memI) have 166:"both_member_options (vebt_insert (treeList ! ?h) ?l) ?l" using myIHs[of "treeList ! ?h" ?l] by (metis "0" "2" highlowprop in_set_member inthall valid_insert_both_member_options_add) have 167:"\ y. both_member_options ((?nextTreeList) ! i) y " using "164" "166" by auto then show ?thesis using "1" "11" "2" True nsprop valid_insert_both_member_options_add highlowprop by auto next case False have "?nextTreeList ! i = treeList ! i" using False by auto have fstprop:"both_member_options ((?nextTreeList) ! i) y \ both_member_options (?nextSummary) i " for y using "1" "4" \i < 2 ^ m\ \treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = treeList ! i\ highlowprop valid_insert_both_member_options_pres by auto moreover have" both_member_options (?nextSummary) i \ \ y . both_member_options ((?nextTreeList) ! i) y " proof- assume "both_member_options (?nextSummary) i " have "i \ high x n" by (simp add: False) hence "both_member_options summary i" by (smt "1" "12" \both_member_options (if minNull (treeList ! high x n) then vebt_insert summary (high x n) else summary) i\ \i < 2 ^ m\ both_member_options_equiv_member highlowprop post_member_pre_member) hence "\ y. both_member_options (treeList ! i) y" by (simp add: "4" \i < 2 ^ m\) then show ?thesis by (simp add: \treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = treeList ! i\) qed ultimately show ?thesis by auto qed next case False hence "?nextSummary = summary" by simp hence "\ y. both_member_options (treeList ! high x n) y" using not_min_Null_member False by blast hence "both_member_options summary (high x n)" using "4" highlowprop by blast hence " both_member_options (?nextTreeList ! high x n) ?l" by (metis "0" "2" highlowprop nth_list_update_eq nth_mem valid_insert_both_member_options_add) then show ?thesis by (metis (full_types) "4" False \both_member_options summary (high x n)\ \i < 2 ^ m\ nth_list_update_neq) qed qed qed have 14: "(mi = max ma x \ (\ t \ set ?nextTreeList. \ y. both_member_options t y))" using True max_less_iff_conj by blast have 15: "mi \ max ma x \ max ma x < 2^deg" using "5.hyps"(8) "5.prems" abcdef by auto have 16: "(mi \ max ma x \ (\ i < 2^m. (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ mi < y \ y \ max ma x)))" proof assume "mi \ max ma x" show "(\ i < 2^m. (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ mi < y \ y \ max ma x))" proof fix i::nat show "i < 2 ^ m\ (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof assume "i < 2^m" show " (high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof show "(high (max ma x) n = i \ both_member_options (?nextTreeList ! i) (low (max ma x) n))" proof assume "high (max ma x) n = i" show "both_member_options (?nextTreeList ! i) (low (max ma x) n)" proof(cases "high x n = high ma n") case True have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence "?nextTreeList ! i = vebt_insert (treeList ! i) (low x n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (metis "2" False True \high (max ma x) n = i\ highlowprop linorder_neqE_nat max.commute max.strict_order_iff nth_list_update_eq) hence "vebt_member (?nextTreeList ! i) (low x n)" by (metis "11" "2" True \high (max ma x) n = i\ \invar_vebt (treeList ! i) n\ highlowprop max_def set_update_memI valid_insert_both_member_options_add valid_member_both_member_options) then show ?thesis proof(cases "mi = ma") case True then show ?thesis by (metis \treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \invar_vebt (treeList ! i) n\ abcdef highlowprop max.commute max.strict_order_iff valid_insert_both_member_options_add) next case False hence "vebt_member (treeList ! i) (low ma n)" by (metis "7" True \high (max ma x) n = i\ \invar_vebt (treeList ! i) n\ highlowprop max_def valid_member_both_member_options) hence "vebt_member (?nextTreeList ! i) (low ma n) \ (low ma n = low x n)" using post_member_pre_member[of " (treeList ! i)" n "low x n" "low ma n" ] by (metis "1" "11" "2" "3" "5.hyps"(8) "7" False True \treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \invar_vebt (treeList ! i) n\ deg_not_0 exp_split_high_low(2) highlowprop nth_list_update_neq set_update_memI valid_insert_both_member_options_pres valid_member_both_member_options) then show ?thesis by (metis "11" "2" True \high (max ma x) n = i\ \treeList[high x n := VEBT_Insert.vebt_insert (treeList ! high x n) (low x n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low x n)\ \invar_vebt (treeList ! i) n\ both_member_options_equiv_member highlowprop max_def set_update_memI valid_insert_both_member_options_add) qed next case False then show ?thesis by (metis "0" "2" "7" \high (max ma x) n = i\ \i < 2 ^ m\ \mi \ max ma x\ highlowprop max_def nth_list_update_eq nth_list_update_neq nth_mem valid_insert_both_member_options_add) qed qed show "(\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x)" proof fix y show "high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ mi < y \ y \ max ma x" proof assume bb:"high y n = i \ both_member_options (?nextTreeList ! i) (low y n)" show " mi < y \ y \ max ma x" proof(cases "i = high x n") case True hence cc:" i = high x n" by simp have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = vebt_insert (treeList ! i) (low x n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (simp add: cc highlowprop) hence "invar_vebt (?nextTreeList ! i) n" by (simp add: "2" "5.IH"(1) cc highlowprop) hence "vebt_member (treeList ! i) (low y n) \ (low y n) = (low x n)" by (metis \high y n = i \ both_member_options ((treeList[?h:=vebt_insert (treeList ! high x n) (low x n)]) ! i) (low y n)\ \invar_vebt (treeList ! i) n\ aa highlowprop member_bound post_member_pre_member valid_member_both_member_options) then show ?thesis proof(cases "low y n = low x n") case True hence "high x n = high y n \ low y n = low x n" by (simp add: bb cc) hence "x = y" by (metis bit_split_inv) then show ?thesis using abcdef by auto next case False hence "vebt_member (treeList ! i) (low y n)" using \vebt_member (treeList ! i) (low y n) \ low y n = low x n\ by blast hence "mi \ ma " using 5 inthall by (metis "2" \i < 2 ^ m\ min_Null_member not_min_Null_member) then show ?thesis using "7" \i < 2 ^ m\ \vebt_member (treeList ! i) (low y n)\ \invar_vebt (treeList ! i) n\ bb both_member_options_equiv_member max.coboundedI1 by blast qed next case False have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = (treeList ! i)" using False by auto hence "both_member_options (treeList ! i) (low y n)" using bb by auto hence "mi \ ma " using 5 using "2" \i < 2 ^ m\ by fastforce then show ?thesis using 7 using \both_member_options (treeList ! i) (low y n)\ \i < 2 ^ m\ bb max.coboundedI1 by blast qed qed qed qed qed qed qed then show ?thesis using invar_vebt.intros(5)[of ?nextTreeList n ?nextSummary m deg mi "max ma x"] by (smt (z3) "10" "11" "12" "13" "14" "15" "2" "3" "8" length_list_update max.commute) next case False hence abcdef: "x < mi" using antisym_conv3 mimaxrel by blast let ?h = "high mi n" and ?l="low mi n" have highlowprop: "high mi n < 2^m \ low mi n < 2^n" by (metis (full_types) "1" "2" "3" "5.IH"(1) "5.hyps"(7) "5.hyps"(8) deg_not_0 exp_split_high_low(1) exp_split_high_low(2) le_less_trans not_one_le_zero set_n_deg_not_0) have 10:"vebt_insert (Node (Some (mi,ma)) deg treeList summary) x = Node (Some (x, max mi ma)) deg (treeList[ ?h :=vebt_insert (treeList ! ?h) ?l]) (if minNull (treeList ! ?h) then vebt_insert summary ?h else summary) " by (metis "0" "2" "9" abcdef div_less highlowprop insert_simp_excp mimaxrel not_less not_one_le_zero set_n_deg_not_0) let ?maxnew = "max mi ma" and ?nextTreeList = "(treeList[ ?h:=vebt_insert (treeList ! ?h) ?l])" and ?nextSummary = "(if minNull (treeList ! ?h) then vebt_insert summary ?h else summary)" have 11: "( \ t \ set ?nextTreeList. invar_vebt t n)" proof fix t assume " t \ set ?nextTreeList" then obtain i where "i < 2^m \ ?nextTreeList ! i = t" by (metis "2" in_set_conv_nth length_list_update) thus "invar_vebt t n" by (metis "2" "5.IH"(1) highlowprop nth_list_update_eq nth_list_update_neq nth_mem) qed have 12: "invar_vebt ?nextSummary m" by (simp add: "1" "5.IH"(2) highlowprop) have 13: "\ i < 2^m. (\ y. both_member_options (?nextTreeList ! i) y) \ ( both_member_options ?nextSummary i)" proof fix i show "i < 2 ^ m \ (\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof assume "i< 2^m" show "(\y. both_member_options ((?nextTreeList) ! i) y) = both_member_options (?nextSummary) i " proof(cases "minNull (treeList ! high mi n)") case True hence tc: "minNull (treeList ! high mi n)" by simp hence nsprop: "?nextSummary = vebt_insert summary ?h" by simp have insprop:"?nextTreeList ! ?h = vebt_insert (treeList ! ?h) ?l" by (simp add: "2" highlowprop) then show ?thesis proof(cases "i = ?h") case True have 161:"\ y. vebt_member (treeList ! ?h) y" by (simp add: min_Null_member tc) hence 162:"\ y. both_member_options (treeList ! ?h) y" by (metis "0" "2" highlowprop nth_mem valid_member_both_member_options) hence 163:"\ both_member_options summary i" using "11" "2" "4" True \i < 2 ^ m\ by blast have 164:"?nextTreeList ! i = vebt_insert (treeList ! ?h) ?l" using True insprop by simp have 165:"invar_vebt (vebt_insert (treeList ! ?h) ?l) n" by (simp add: "11" "2" highlowprop set_update_memI) have 166:"both_member_options (vebt_insert (treeList ! ?h) ?l) ?l" using myIHs[of "treeList ! ?h" ?l] by (metis "0" "2" highlowprop in_set_member inthall valid_insert_both_member_options_add) have 167:"\ y. both_member_options ((?nextTreeList) ! i) y " using "164" "166" by auto then show ?thesis using "1" "11" "2" True nsprop valid_insert_both_member_options_add highlowprop by auto next case False have "?nextTreeList ! i = treeList ! i" by (metis False nth_list_update_neq) have fstprop:"both_member_options ((?nextTreeList) ! i) y \ both_member_options (?nextSummary) i " for y using "1" "4" \i < 2 ^ m\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\ highlowprop valid_insert_both_member_options_pres by auto moreover have" both_member_options (?nextSummary) i \ \ y . both_member_options ((?nextTreeList) ! i) y " proof- assume "both_member_options (?nextSummary) i " have "i \ high mi n" by (simp add: False) hence "both_member_options summary i" by (smt (z3) "1" "12" \both_member_options (if minNull (treeList ! high mi n) then VEBT_Insert.vebt_insert summary (high mi n) else summary) i\ \i < 2 ^ m\ both_member_options_equiv_member highlowprop post_member_pre_member) hence "\ y. both_member_options (treeList ! i) y" by (simp add: "4" \i < 2 ^ m\) then show ?thesis by (simp add: \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\) qed ultimately show ?thesis by auto qed next case False hence "?nextSummary = summary" by simp hence "\ y. both_member_options (treeList ! high mi n) y" using not_min_Null_member False by blast hence "both_member_options summary (high mi n)" using "4" highlowprop by blast hence " both_member_options (?nextTreeList ! high mi n) ?l" by (metis "0" "2" highlowprop nth_list_update_eq nth_mem valid_insert_both_member_options_add) then show ?thesis by (metis (full_types) "4" False \both_member_options summary (high mi n)\ \i < 2 ^ m\ nth_list_update_neq) qed qed qed have 14: "(x = max ma mi \ (\ t \ set ?nextTreeList. \ y. both_member_options t y))" using mimaxrel by linarith have 15: "x \ max ma mi \ max ma mi < 2^deg" using "6" abcdef by linarith have 16: "(x \ max ma mi \ (\ i < 2^m. (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ x < y \ y \ max ma mi)))" proof assume "x \ max ma mi" show "(\ i < 2^m. (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\ y. (high y n = i \ both_member_options (?nextTreeList ! i) (low y n) ) \ x < y \ y \ max ma mi))" proof fix i::nat show "i < 2 ^ m\ (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof assume "i < 2^m" show " (high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n)) \ (\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof show "(high (max ma mi) n = i \ both_member_options (?nextTreeList ! i) (low (max ma mi) n))" proof assume "high (max ma mi) n = i" show "both_member_options (?nextTreeList ! i) (low (max ma mi) n)" proof(cases "high mi n = high ma n") case True have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence "?nextTreeList ! i = vebt_insert (treeList ! i) (low mi n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (metis "2" "5.hyps"(7) True \high (max ma mi) n = i\ highlowprop max.orderE nth_list_update_eq) hence "vebt_member (?nextTreeList ! i) (low mi n)" by (metis "11" "2" True \high (max ma mi) n = i\ \invar_vebt (treeList ! i) n\ highlowprop max_def set_update_memI valid_insert_both_member_options_add valid_member_both_member_options) then show ?thesis proof(cases "mi = ma") case True then show ?thesis using \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ \invar_vebt (treeList ! i) n\ highlowprop valid_insert_both_member_options_add by auto next case False hence "vebt_member (treeList ! i) (low ma n)" using "6" "7" \high (max ma mi) n = i\ \i < 2 ^ m\ \invar_vebt (treeList ! i) n\ both_member_options_equiv_member by auto hence "vebt_member (?nextTreeList ! i) (low ma n) \ (low ma n = low mi n)" using post_member_pre_member[of " (treeList ! i)" n "low mi n" "low ma n" ] by (metis "1" "11" "2" "3" "5.hyps"(8) "7" True \high (max ma mi) n = i\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ \invar_vebt (treeList ! i) n\ deg_not_0 exp_split_high_low(2) highlowprop max_def set_update_memI valid_insert_both_member_options_pres valid_member_both_member_options) then show ?thesis by (metis "5.hyps"(7) "7" False \high (max ma mi) n = i\ \i < 2 ^ m\ \vebt_member (treeList ! i) (low ma n)\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = VEBT_Insert.vebt_insert (treeList ! i) (low mi n)\ \invar_vebt (treeList ! i) n\ highlowprop max.absorb1 member_bound valid_insert_both_member_options_pres) qed next case False hence "?nextTreeList ! i = treeList ! i" by (metis "5.hyps"(7) \high (max ma mi) n = i\ max.commute max_def nth_list_update_neq) then show ?thesis proof(cases "mi < ma") case True then show ?thesis by (metis "5.hyps"(7) "7" False \high (max ma mi) n = i\ \i < 2 ^ m\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\ max.commute max_def) next case False hence "mi \ ma " by simp hence "mi = ma" by (simp add: "6" eq_iff) hence "\both_member_options (treeList ! i) (low (max ma mi) n)" using 5 "2" \i < 2 ^ m\ by auto then show ?thesis by (metis "11" "2" \high (max ma mi) n = i\ \mi = ma\ \treeList[high mi n := VEBT_Insert.vebt_insert (treeList ! high mi n) (low mi n)] ! i = treeList ! i\ highlowprop max.idem nth_list_update_eq set_update_memI valid_insert_both_member_options_add) qed qed qed show "(\y. high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi)" proof fix y show "high y n = i \ both_member_options (?nextTreeList ! i) (low y n) \ x < y \ y \ max ma mi" proof assume bb:"high y n = i \ both_member_options (?nextTreeList ! i) (low y n)" show " x < y \ y \ max ma mi" proof(cases "i = high mi n") case True hence cc:" i = high mi n" by simp have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = vebt_insert (treeList ! i) (low mi n)" using concat_inth[of "take (high x n) treeList" "vebt_insert (treeList ! i) (low x n)" "drop (high x n + 1) treeList"] by (simp add: cc highlowprop) hence "invar_vebt (?nextTreeList ! i) n" by (simp add: "2" "5.IH"(1) \i < 2 ^ m\ highlowprop) hence "vebt_member (treeList ! i) (low y n) \ (low y n) = (low mi n)" by (metis \invar_vebt (treeList ! i) n\ aa bb both_member_options_equiv_member highlowprop member_bound post_member_pre_member) then show ?thesis proof(cases "low y n = low mi n") case True hence "high mi n = high y n \ low y n = low mi n" by (simp add: bb cc) hence "mi = y" by (metis bit_split_inv) then show ?thesis using abcdef by auto next case False hence "vebt_member (treeList ! i) (low y n)" using \vebt_member (treeList ! i) (low y n) \ low y n = low mi n\ by blast hence "mi \ ma " using 5 inthall by (metis "2" \i < 2 ^ m\ min_Null_member not_min_Null_member) then show ?thesis using "7" by (metis \i < 2 ^ m\ \vebt_member (treeList ! i) (low y n)\ \invar_vebt (treeList ! i) n\ abcdef bb both_member_options_equiv_member max.absorb1 max.strict_order_iff max_less_iff_conj) qed next case False have "invar_vebt (treeList ! i ) n" by (metis "0" "2" \i < 2 ^ m\ in_set_member inthall) have "length ?nextTreeList = 2^m" using "2" highlowprop by auto hence aa:"?nextTreeList ! i = (treeList ! i)" using False by auto hence "both_member_options (treeList ! i) (low y n)" using bb by auto hence "mi \ ma " using 5 "2" \i < 2 ^ m\ by fastforce then show ?thesis using 7 by (metis \both_member_options (treeList ! i) (low y n)\ \i < 2 ^ m\ abcdef bb max.absorb1 max.strict_order_iff max_less_iff_conj) qed qed qed qed qed qed qed then show ?thesis using invar_vebt.intros(5)[of ?nextTreeList n ?nextSummary m deg x "max ma mi"] by (smt (z3) "10" "11" "12" "13" "14" "15" "2" "3" "5.hyps"(7) "8" length_list_update max.absorb2 max.orderE) qed qed qed subsection \Correctness with Respect to Set Interpretation\ theorem insert_corr: assumes "invar_vebt t n " and "x < 2^n " shows " set_vebt' t \ {x} = set_vebt' (vebt_insert t x) " proof show "set_vebt' t \ {x} \ set_vebt' (vebt_insert t x)" proof fix y assume "y \ set_vebt' t \ {x}" show "y \set_vebt' (vebt_insert t x)" proof(cases "x=y") case True then show ?thesis by (metis (full_types) assms(1) assms(2) both_member_options_equiv_member mem_Collect_eq set_vebt'_def valid_insert_both_member_options_add valid_pres_insert) next case False have "vebt_member t y" using False \y \ set_vebt' t \ {x}\ set_vebt'_def by auto hence "vebt_member (vebt_insert t x) y" by (meson assms(1) assms(2) both_member_options_equiv_member member_bound valid_insert_both_member_options_pres valid_pres_insert) then show ?thesis by (simp add: set_vebt'_def) qed qed show " set_vebt' (vebt_insert t x) \ set_vebt' t \ {x} " proof fix y assume "y \ set_vebt' (vebt_insert t x)" show "y \set_vebt' t \ {x}" proof(cases "x=y") case True then show ?thesis by simp next case False hence "vebt_member t y \ x=y" using post_member_pre_member using \y \ set_vebt' (vebt_insert t x)\ assms(1) assms(2) set_vebt'_def member_bound valid_pres_insert by fastforce hence "vebt_member t y" by (simp add: False) hence "y \ set_vebt' t" by (simp add: set_vebt'_def) then show ?thesis by simp qed qed qed corollary insert_correct: assumes "invar_vebt t n " and "x < 2^n " shows " set_vebt t \ {x} = set_vebt (vebt_insert t x) " using assms(1) assms(2) insert_corr set_vebt_set_vebt'_valid valid_pres_insert by blast fun insert'::"VEBT \ nat \ VEBT" where "insert' (Leaf a b) x = vebt_insert (Leaf a b) x"| "insert' (Node info deg treeList summary) x = (if x \ 2^deg then (Node info deg treeList summary ) else vebt_insert (Node info deg treeList summary) x)" theorem insert'_pres_valid: assumes "invar_vebt t n" shows "invar_vebt (insert' t x) n" using assms apply cases apply (metis One_nat_def deg1Leaf insert'.simps(1) vebt_insert.simps(1)) apply (metis assms insert'.simps(2) leI valid_pres_insert)+ done theorem insert'_correct: assumes "invar_vebt t n" shows "set_vebt (insert' t x) = (set_vebt t \ {x})\{0..2^n-1}" proof(cases t) case (Node x11 x12 x13 x14) then show ?thesis proof(cases "x < 2^n") case True hence "set_vebt (insert' t x) = set_vebt(vebt_insert t x)" by (metis Node assms deg_deg_n insert'.simps(2) leD) moreover hence "set_vebt(vebt_insert t x) = set_vebt t \ {x}" using True assms insert_correct by auto moreover hence "set_vebt t \ {x} = (set_vebt t \ {x})\{0..2^n-1} " by (metis Diff_Diff_Int True assms calculation(1) inf_le1 inrange le_inf_iff order_refl subset_antisym set_vebt'_def set_vebt_def set_vebt_set_vebt'_valid valid_pres_insert) ultimately show ?thesis by simp next case False hence "set_vebt (insert' t x) = set_vebt t" by (metis Node assms deg_deg_n insert'.simps(2) leI) moreover hence "set_vebt t = (set_vebt t \ {x})\{0..2^n-1} " by (smt (z3) False Int_commute Int_insert_right_if0 Un_Int_assoc_eq assms atLeastAtMost_iff boolean_algebra_cancel.sup0 inf_bot_right inrange le_add_diff_inverse le_imp_less_Suc one_le_numeral one_le_power plus_1_eq_Suc sup_commute set_vebt_set_vebt'_valid) ultimately show ?thesis by simp qed next case (Leaf x21 x22) then show ?thesis apply(auto simp add: insert'.simps vebt_insert.simps set_vebt_def both_member_options_def) using assms apply cases apply simp+ using assms apply cases apply simp+ using assms apply cases apply simp+ using assms apply cases apply simp+ done qed end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_Member.thy b/thys/Van_Emde_Boas_Trees/VEBT_Member.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_Member.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_Member.thy @@ -1,525 +1,525 @@ (*by Ammer*) theory VEBT_Member imports VEBT_Definitions begin section \Member Function\ context begin interpretation VEBT_internal . fun vebt_member :: "VEBT \ nat \ bool" where "vebt_member (Leaf a b) x = (if x = 0 then a else if x = 1 then b else False)"| "vebt_member (Node None _ _ _) x = False"| "vebt_member (Node _ 0 _ _) x = False"| "vebt_member (Node _ (Suc 0) _ _) x = False"| "vebt_member (Node (Some (mi, ma)) deg treeList summary) x = ( if x = mi then True else if x = ma then True else if x < mi then False else if x > ma then False else ( let h = high x (deg div 2); l = low x (deg div 2) in( if h < length treeList then vebt_member (treeList ! h) l else False))) " end context VEBT_internal begin lemma member_inv: assumes "vebt_member (Node (Some (mi, ma)) deg treeList summary) x " shows "deg \ 2 \ (x = mi \ x = ma \ (x < ma \ x > mi \ high x (deg div 2) < length treeList \ vebt_member (treeList ! ( high x (deg div 2))) (low x (deg div 2))))" proof(cases deg) case 0 then show ?thesis using vebt_member.simps(3)[of "(mi, ma)" treeList summary x] using assms by blast next case (Suc nat) hence "deg = Suc nat" by simp then show ?thesis proof(cases nat) case 0 then show ?thesis using Suc assms by auto next case (Suc nata) hence "deg \ 2" by (simp add: \deg = Suc nat\) then show ?thesis by (metis vebt_member.simps(5) Suc \deg = Suc nat\ assms linorder_neqE_nat) qed qed definition bit_concat::"nat \ nat \ nat \ nat" where "bit_concat h l d = h*2^d + l" lemma bit_split_inv: "bit_concat (high x d) (low x d) d = x" unfolding bit_concat_def high_def low_def by presburger definition set_vebt'::"VEBT \ nat set" where "set_vebt' t = {x. vebt_member t x}" lemma Leaf_0_not: assumes "invar_vebt (Leaf a b) 0 "shows " False" proof- from assms show ?thesis proof(cases) qed qed lemma valid_0_not: "invar_vebt t 0 \ False" proof(induction t) case (Node info deg treeList summary) from this(3) have "length treeList > 0" apply cases apply auto done then obtain t where "t \ set treeList" by fastforce from Node(3) obtain n where "invar_vebt t n" apply cases using Node.IH(2) apply auto done from Node(3) have "n \ 0" apply cases using Node.IH(2) apply auto done hence "n = 0" by blast then show ?case using Node.IH(1) \t \ set treeList\ \invar_vebt t n\ by blast next case (Leaf x1 x2) then show ?case using Leaf_0_not by blast qed theorem valid_tree_deg_neq_0: "(\ invar_vebt t 0)" using valid_0_not by blast lemma deg_1_Leafy: "invar_vebt t n \ n = 1 \ \ a b. t = Leaf a b" apply(induction rule: invar_vebt.induct) apply simp apply presburger apply (metis (full_types) Suc_eq_plus1 add_cancel_right_left in_set_replicate list.map_cong0 map_replicate_const nat_neq_iff not_add_less2 numeral_1_eq_Suc_0 numeral_2_eq_2 numerals(1) order_less_irrefl power_eq_0_iff valid_tree_deg_neq_0 zero_less_numeral) apply (metis odd_add odd_one) by (metis Suc_eq_plus1 add_cancel_right_left in_set_replicate list.map_cong0 map_replicate_const nat_neq_iff not_add_less2 numeral_2_eq_2 power_eq_0_iff valid_tree_deg_neq_0) lemma deg_1_Leaf: "invar_vebt t 1 \ \ a b. t = Leaf a b" using deg_1_Leafy by blast corollary deg1Leaf: "invar_vebt t 1 \ (\ a b. t = Leaf a b)" using deg_1_Leaf invar_vebt.intros(1) by auto lemma deg_SUcn_Node: assumes "invar_vebt tree (Suc (Suc n)) " shows " \ info treeList s. tree = Node info (Suc (Suc n)) treeList s" proof- from assms show ?thesis apply(cases) apply blast+ done qed lemma "invar_vebt (Node info deg treeList summary) deg \ deg > 1" by (metis VEBT.simps(4) deg_1_Leafy less_one linorder_neqE_nat valid_tree_deg_neq_0) lemma deg_deg_n: assumes "invar_vebt (Node info deg treeList summary) n "shows" deg = n" proof- from assms show ?thesis proof(cases) qed blast+ qed lemma member_valid_both_member_options: "invar_vebt tree n \ vebt_member tree x \ (naive_member tree x \ membermima tree x)" proof(induction tree n arbitrary: x rule: invar_vebt.induct) case (1 a b) then show ?case using vebt_member.simps(1) naive_member.simps(1) by blast next case (2 treeList n summary m deg) then show ?case by simp next case (3 treeList n summary m deg) then show ?case using vebt_member.simps(2) by blast next case (4 treeList n summary m deg mi ma) - hence "deg \ 2" - by (metis Euclidean_Division.div_eq_0_iff add_self_div_2 le_less_linear valid_tree_deg_neq_0) + hence "deg \ 2" + using member_inv by blast then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis by (metis (full_types) "4"(12) vebt_member.simps(3) membermima.simps(4) old.nat.exhaust) next case False hence 1:"mi < x \ x < ma \ (high x (deg div 2 )) < length treeList \ vebt_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using member_inv[of mi ma deg treeList summary x] "4"(12) by blast hence " (treeList ! (high x (deg div 2))) \ set treeList" by (metis in_set_member inthall) hence "both_member_options (treeList ! (high x (deg div 2))) (low x (deg div 2))" using "1" "4.IH"(1) both_member_options_def by blast then show ?thesis by (smt "1" "4"(1) "4"(6) \treeList ! high x (deg div 2) \ set treeList\ membermima.simps(4) naive_member.simps(3) old.nat.exhaust valid_tree_deg_neq_0 zero_eq_add_iff_both_eq_0) qed next case (5 treeList n summary m deg mi ma) hence "deg \ 2" using member_inv by presburger then show ?case proof(cases "x = mi \ x = ma") case True then show ?thesis by (metis (full_types) "5"(12) vebt_member.simps(3) membermima.simps(4) old.nat.exhaust) next case False hence 1:"mi < x \ x < ma \ (high x (deg div 2 )) < length treeList \ vebt_member (treeList ! (high x (deg div 2))) (low x (deg div 2))" using member_inv[of mi ma deg treeList summary x] "5"(12) by blast hence " (treeList ! (high x (deg div 2))) \ set treeList" by (metis in_set_member inthall) hence "both_member_options (treeList ! (high x (deg div 2))) (low x (deg div 2))" using "1" "5.IH"(1) both_member_options_def by blast then show ?thesis by (smt "1" "5"(1) "5"(6) \treeList ! high x (deg div 2) \ set treeList\ membermima.simps(4) naive_member.simps(3) old.nat.exhaust valid_tree_deg_neq_0 zero_eq_add_iff_both_eq_0) qed qed lemma member_bound: "vebt_member tree x \ invar_vebt tree n \ x < 2^n" proof(induction tree x arbitrary: n rule: vebt_member.induct) case (1 a b x) then show ?case by (metis vebt_member.simps(1) One_nat_def le_neq_implies_less nat_power_eq_Suc_0_iff numeral_eq_iff numerals(1) one_le_numeral one_le_power semiring_norm(85) valid_tree_deg_neq_0 zero_less_numeral zero_less_power) next case (2 uu uv uw x) then show ?case by simp next case (3 v uy uz x) then show ?case by simp next case (4 v vb vc x) then show ?case by simp next case (5 mi ma va treeList summary x) hence 111: "n = Suc (Suc va)" using deg_deg_n by fastforce hence "ma < 2^n" using "5.prems"(2) mi_ma_2_deg by blast then show ?case by (metis "5.prems"(1) "5.prems"(2) le_less_trans less_imp_le_nat member_inv mi_ma_2_deg) qed theorem inrange: assumes "invar_vebt t n" shows " set_vebt' t \ {0..2^n-1}" proof fix x assume "x \ set_vebt' t" hence "vebt_member t x" using set_vebt'_def by auto hence "x < 2^n" using assms member_bound by blast then show "x \ {0..2^n-1}" by simp qed theorem buildup_gives_empty: "set_vebt' (vebt_buildup n) = {}" unfolding set_vebt'_def by (metis Collect_empty_eq vebt_member.simps(1) vebt_member.simps(2) vebt_buildup.elims) fun minNull::"VEBT \ bool" where "minNull (Leaf False False) = True"| "minNull (Leaf _ _ ) = False"| "minNull (Node None _ _ _) = True"| "minNull (Node (Some _) _ _ _) = False" lemma min_Null_member: "minNull t \ \ vebt_member t x" apply(induction t) using vebt_member.simps(2) minNull.elims(2) apply blast apply auto done lemma not_min_Null_member: "\ minNull t \ \ y. both_member_options t y" proof(induction t) case (Node info deg treeList summary) obtain mi ma where "info = Some(mi , ma)" by (metis Node.prems minNull.simps(4) not_None_eq surj_pair) then show ?case by (metis (full_types) both_member_options_def membermima.simps(3) membermima.simps(4) not0_implies_Suc) next case (Leaf x1 x2) then show ?case by (metis (full_types) both_member_options_def minNull.simps(1) naive_member.simps(1) zero_neq_one) qed lemma valid_member_both_member_options: "invar_vebt t n \ both_member_options t x \ vebt_member t x" proof(induction t n arbitrary: x rule: invar_vebt.induct) case (1 a b) then show ?case by (simp add: both_member_options_def) next case (2 treeList n summary m deg) hence 0: " ( \ t \ set treeList. invar_vebt t n) " and 1:" invar_vebt summary n" and 2:" length treeList = 2^n" and 3:" deg = 2*n" and 4:" (\ i. both_member_options summary i)" and 5:" (\ t \ set treeList. \ y. both_member_options t y) " and 6: "n> 0" apply blast+ apply (auto simp add: "2.hyps"(3) "2.hyps") using "2.hyps"(1) "2.hyps"(3) neq0_conv valid_tree_deg_neq_0 by blast have "both_member_options (Node None deg treeList summary) x \ False" proof- assume "both_member_options (Node None deg treeList summary) x" hence "naive_member (Node None deg treeList summary) x \ membermima (Node None deg treeList summary) x" unfolding both_member_options_def by simp then show False proof(cases "naive_member (Node None deg treeList summary) x") case True hence "high x n < length treeList \ naive_member (treeList ! (high x n)) (low x n)" by (metis "1" "2.hyps"(3) "2.hyps"(4) add_cancel_right_left add_self_div_2 naive_member.simps(3) old.nat.exhaust valid_tree_deg_neq_0) then show ?thesis by (metis "5" both_member_options_def in_set_member inthall) next case False hence "membermima (Node None deg treeList summary) x" using \naive_member (Node None deg treeList summary) x \ membermima (Node None deg treeList summary) x\ by blast moreover have "Suc (deg-1) =deg" by (simp add: "2.hyps"(4) "6") moreover hence "(let pos = high x (deg div 2) in if pos < length treeList then membermima (treeList ! pos) (low x (Suc (deg - 1) div 2)) else False)" using calculation(1) membermima.simps(5) by metis moreover hence " if high x (deg div 2) < length treeList then membermima (treeList ! ( high x (deg div 2))) (low x(deg div 2)) else False" using calculation(2) by metis ultimately have " high x (deg div 2) < length treeList \ membermima (treeList ! (high x n)) (low x n)" by (metis "2.hyps"(3) "2.hyps"(4) add_self_div_2) then show ?thesis using "2.IH" "5" both_member_options_def in_set_member inthall by (metis "2.hyps"(3) "2.hyps"(4) add_self_div_2) qed qed then show ?case by (simp add: "2.prems") next case (3 treeList n summary m deg) hence 0: " ( \ t \ set treeList. invar_vebt t n) " and 1:" invar_vebt summary m" and 2:" length treeList = 2^m" and 3:" deg = n+m" and 4:" (\ i. both_member_options summary i)" and 5:" (\ t \ set treeList. \ y. both_member_options t y) " and 6: "n> 0" and 7: "m> 0" and 8: "n+1 = m" apply blast+ apply (metis (full_types) "3.IH"(1) "3.hyps"(2) in_set_member inthall neq0_conv power_eq_0_iff valid_tree_deg_neq_0 zero_neq_numeral) apply (simp add: "3.hyps"(3)) by (simp add: "3.hyps"(3)) have "both_member_options (Node None deg treeList summary) x \ False" proof- assume "both_member_options (Node None deg treeList summary) x" hence "naive_member (Node None deg treeList summary) x \ membermima (Node None deg treeList summary) x" unfolding both_member_options_def by simp then show False proof(cases "naive_member (Node None deg treeList summary) x") case True hence "high x n < length treeList \ naive_member (treeList ! (high x n)) (low x n)" by (metis "3" "3.hyps"(3) add_Suc_right add_self_div_2 even_Suc_div_two naive_member.simps(3) odd_add) then show ?thesis by (metis "5" both_member_options_def in_set_member inthall) next case False hence "membermima (Node None deg treeList summary) x" using \naive_member (Node None deg treeList summary) x \ membermima (Node None deg treeList summary) x\ by blast moreover have "Suc (deg-1) =deg" by (simp add: "3" "3.hyps"(3)) moreover hence "(let pos = high x (deg div 2) in if pos < length treeList then membermima (treeList ! pos) (low x (Suc (deg - 1) div 2)) else False)" using calculation(1) membermima.simps(5) by metis moreover hence 11:" if high x (deg div 2) < length treeList then membermima (treeList ! ( high x (deg div 2))) (low x(deg div 2)) else False" using calculation(2) by metis ultimately have " high x (deg div 2) < length treeList \ membermima (treeList ! (high x n)) (low x n)" by (metis "3" "3.hyps"(3) add_Suc_right add_self_div_2 even_Suc_div_two odd_add) then show ?thesis using "3.IH" "5" both_member_options_def in_set_member inthall 11 by metis qed qed then show ?case using "3.prems" by blast next case (4 treeList n summary m deg mi ma) hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary n" and 2:"length treeList = 2^n" and 3:" deg = n+m" and "n=m" and 4: "(\ i < 2^n. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" using "4.prems" by auto hence "n>0" by (metis neq0_conv valid_tree_deg_neq_0) then show ?case proof(cases "x = mi \ x = ma") case True hence xmimastmt: "x = mi \ x=ma" by simp then show ?thesis using vebt_member.simps(5)[of mi ma "deg-2" treeList summary x] by (metis "3" "4.hyps"(3) \0 < n\ add_diff_inverse_nat add_numeral_left add_self_div_2 div_if nat_neq_iff numerals(1) plus_1_eq_Suc semiring_norm(2)) next case False hence xmimastmt: "x \ mi \ x\ma" by simp hence "mi = ma \ False" proof- assume "mi = ma" hence astmt: "(\ t \ set treeList. \ y. both_member_options t y)" using 5 by simp have bstmt: "both_member_options (Node (Some (mi, ma)) deg treeList summary) x" by (simp add: "4.prems") then show False proof(cases "naive_member (Node (Some (mi, ma)) deg treeList summary) x") case True hence "high x n < length treeList \ naive_member (treeList ! (high x n)) (low x n)" by (metis (no_types, opaque_lifting) "3" "4.hyps"(1) "4.hyps"(3) add_self_div_2 naive_member.simps(3) old.nat.exhaust valid_0_not zero_eq_add_iff_both_eq_0) then show ?thesis by (metis "5" \mi = ma\ both_member_options_def in_set_member inthall) next case False hence "membermima (Node (Some (mi, ma)) deg treeList summary) x" using bstmt unfolding both_member_options_def by blast hence " x = mi \ x = ma \ (if high x n < length treeList then membermima (treeList ! (high x n)) (low x n) else False)" using membermima.simps(4)[of mi ma "deg-1" treeList summary x] by (metis "3" "4.hyps"(3) One_nat_def Suc_diff_Suc \0 < n\ add_gr_0 add_self_div_2 diff_zero) hence " high x n < length treeList \ membermima (treeList ! (high x n)) (low x n)" using xmimastmt by presburger then show ?thesis using both_member_options_def in_set_member inthall membermima.simps(4)[of mi ma n treeList summary x] astmt by metis qed qed hence "mi \ ma " by blast hence followstmt: "(\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma))" using 7 by simp have 10:"high x n < length treeList \ (naive_member (treeList ! (high x n)) (low x n) \ membermima (treeList ! (high x n)) (low x n) )" by (smt "3" "4.hyps"(3) "4.prems" False One_nat_def Suc_leI \0 < n\ add_gr_0 add_self_div_2 both_member_options_def le_add_diff_inverse membermima.simps(4) naive_member.simps(3) plus_1_eq_Suc) hence 11:"both_member_options (treeList ! (high x n)) (low x n)" by (simp add: both_member_options_def) have 12:"high x n< 2^m" using "10" "4.hyps"(2) by auto hence "mi < x \ x < ma" proof- have "(\ y. (high y n = (high x n) \ both_member_options (treeList ! (high y n)) (low y n) ) \ mi < y \ y \ ma)" using "12" followstmt by auto then show ?thesis using "11" False order.not_eq_order_implies_strict by blast qed have "vebt_member (treeList ! (high x n)) (low x n)" by (metis"10" "11" "4.IH"(1) in_set_member inthall) then show ?thesis by (smt "10" "11" "12" "3" "4.hyps"(3) vebt_member.simps(5) One_nat_def Suc_leI \0 < n\ add_Suc_right add_self_div_2 followstmt le_add_diff_inverse le_imp_less_Suc not_less_eq not_less_iff_gr_or_eq plus_1_eq_Suc) qed next case (5 treeList n summary m deg mi ma) hence 0: "( \ t \ set treeList. invar_vebt t n)" and 1: " invar_vebt summary m" and 2:"length treeList = 2^m" and 3:" deg = n+m" and "Suc n=m" and 4: "(\ i < 2^m. (\ y. both_member_options (treeList ! i) y) \ ( both_member_options summary i))" and 5: "(mi = ma \ (\ t \ set treeList. \ y. both_member_options t y))" and 6:"mi \ ma \ ma < 2^deg" and 7: "(mi \ ma \ (\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma)))" using "5.prems" by auto hence "n>0" by (metis "5.hyps"(3) in_set_member inthall neq0_conv power_Suc0_right valid_tree_deg_neq_0 zero_neq_numeral) then show ?case proof(cases "x = mi \ x = ma") case True hence xmimastmt: "x = mi \ x=ma" by simp then show ?thesis using vebt_member.simps(5)[of mi ma "deg-2" treeList summary x] using "3" "5.hyps"(3) \0 < n\ by auto next case False hence xmimastmt: "x \ mi \ x\ma" by simp hence "mi = ma \ False" proof- assume "mi = ma" hence astmt: "(\ t \ set treeList. \ y. both_member_options t y)" using 5 by simp have bstmt: "both_member_options (Node (Some (mi, ma)) deg treeList summary) x" by (simp add: "5.prems") then show False proof(cases "naive_member (Node (Some (mi, ma)) deg treeList summary) x") case True hence "high x n < length treeList \ naive_member (treeList ! (high x n)) (low x n)" by (metis "3" "5.hyps"(3) add_Suc_right add_self_div_2 even_Suc_div_two naive_member.simps(3) odd_add) then show ?thesis by (metis "5" \mi = ma\ both_member_options_def in_set_member inthall) next case False hence "membermima (Node (Some (mi, ma)) deg treeList summary) x" using bstmt unfolding both_member_options_def by blast hence " x = mi \ x = ma \ (if high x n < length treeList then membermima (treeList ! (high x n)) (low x n) else False)" using membermima.simps(4)[of mi ma "deg-1" treeList summary x] using "3" "5.hyps"(3) by auto hence " high x n < length treeList \ membermima (treeList ! (high x n)) (low x n)" using xmimastmt by presburger then show ?thesis using both_member_options_def in_set_member inthall membermima.simps(4)[of mi ma n treeList summary x] astmt by metis qed qed hence "mi \ ma " by blast hence followstmt: "(\ i < 2^m. (high ma n = i \ both_member_options (treeList ! i) (low ma n)) \ (\ y. (high y n = i \ both_member_options (treeList ! i) (low y n) ) \ mi < y \ y \ ma))" using 7 by simp have 10:"high x n < length treeList \ (naive_member (treeList ! (high x n)) (low x n) \ membermima (treeList ! (high x n)) (low x n) )" by (smt "3" "5.hyps"(3) "5.prems" False add_Suc_right add_self_div_2 both_member_options_def even_Suc_div_two membermima.simps(4) naive_member.simps(3) odd_add) hence 11:"both_member_options (treeList ! (high x n)) (low x n)" by (simp add: both_member_options_def) have 12:"high x n< 2^m" using "10" "5.hyps"(2) by auto hence "mi < x \ x < ma" proof- have "(\ y. (high y n = (high x n) \ both_member_options (treeList ! (high y n)) (low y n) ) \ mi < y \ y \ ma)" using "12" followstmt by auto then show ?thesis using "11" False order.not_eq_order_implies_strict by blast qed have "vebt_member (treeList ! (high x n)) (low x n)" by (metis "10" "11" "5.IH"(1) in_set_member inthall) then show ?thesis by (smt "10" "11" "12" "3" "5.hyps"(3) vebt_member.simps(5) Suc_pred \0 < n\ add_Suc_right add_self_div_2 even_Suc_div_two followstmt le_neq_implies_less not_less_iff_gr_or_eq odd_add) qed qed corollary both_member_options_equiv_member: assumes "invar_vebt t n" shows "both_member_options t x \ vebt_member t x" using assms both_member_options_def member_valid_both_member_options valid_member_both_member_options by blast lemma member_correct: "invar_vebt t n \ vebt_member t x \ x \ set_vebt t " using both_member_options_equiv_member set_vebt_def by auto corollary set_vebt_set_vebt'_valid: assumes "invar_vebt t n" shows "set_vebt t =set_vebt' t" unfolding set_vebt_def set_vebt'_def apply auto using assms valid_member_both_member_options apply auto[1] using assms both_member_options_equiv_member by auto lemma set_vebt_finite: "invar_vebt t n \ finite (set_vebt' t)" using finite_subset inrange by blast lemma mi_eq_ma_no_ch:assumes "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg " and " mi = ma " shows "(\ t \ set treeList. \ x. both_member_options t x ) \ (\ x. both_member_options summary x)" proof- from assms(1) show ?thesis proof(cases) case (4 n m) have 0:"\t\set treeList. \ Ex (both_member_options t)" by (simp add: "4"(7) assms(2)) moreover have "both_member_options summary x \ False" for x proof- assume "both_member_options summary x " hence "vebt_member summary x" using "4"(2) valid_member_both_member_options by auto moreover hence "x < 2^m" using "4"(2) member_bound by auto ultimately have "\ y. both_member_options (treeList ! (high x n)) y" using "0" "4"(3) "4"(4) "4"(6) \both_member_options summary x\ inthall by (metis nth_mem) - then show ?thesis - by (metis "0" "4"(3) "4"(4) Euclidean_Division.div_eq_0_iff \x < 2 ^ m\ high_def nth_mem zero_less_numeral zero_less_power) + then show ?thesis + by (metis "0" "4"(3) "4"(4) div_eq_0_iff \x < 2 ^ m\ high_def nth_mem zero_less_numeral zero_less_power) qed then show ?thesis using calculation by blast next case (5 n m) have 0:"\t\set treeList. \ Ex (both_member_options t)" using "5"(7) assms(2) by blast moreover have "both_member_options summary x \ False" for x proof- assume "both_member_options summary x " hence "vebt_member summary x" using "5"(2) valid_member_both_member_options by auto moreover hence "x < 2^m" using "5"(2) member_bound by auto ultimately have "\ y. both_member_options (treeList ! (high x n)) y" using "0" "5"(3) "5"(4) "5"(6) \both_member_options summary x\ inthall by (metis nth_mem) then show ?thesis by (metis "0" "5"(3) "5"(6) \both_member_options summary x\ \x < 2 ^ m\ nth_mem) qed then show ?thesis using calculation by blast qed qed end end diff --git a/thys/Van_Emde_Boas_Trees/VEBT_MinMax.thy b/thys/Van_Emde_Boas_Trees/VEBT_MinMax.thy --- a/thys/Van_Emde_Boas_Trees/VEBT_MinMax.thy +++ b/thys/Van_Emde_Boas_Trees/VEBT_MinMax.thy @@ -1,326 +1,326 @@ (*by Ammer*) theory VEBT_MinMax imports VEBT_Member begin section \The Minimum and Maximum Operation\ fun vebt_mint :: "VEBT \ nat option" where "vebt_mint (Leaf a b) = (if a then Some 0 else if b then Some 1 else None)"| "vebt_mint (Node None _ _ _) = None"| "vebt_mint (Node (Some (mi,ma)) _ _ _ ) = Some mi" fun vebt_maxt :: "VEBT \ nat option" where "vebt_maxt (Leaf a b) = (if b then Some 1 else if a then Some 0 else None)"| "vebt_maxt (Node None _ _ _) = None"| "vebt_maxt (Node (Some (mi,ma)) _ _ _ ) = Some ma" context VEBT_internal begin fun option_shift::"('a\'a\'a) \'a option \'a option\ 'a option" where "option_shift _ None _ = None"| "option_shift _ _ None = None"| "option_shift f (Some a) (Some b) = Some (f a b)" definition power::"nat option \ nat option \ nat option" (infixl"^\<^sub>o" 81) where "power= option_shift (^)" definition add::"nat option \ nat option \ nat option" (infixl"+\<^sub>o" 79) where "add= option_shift (+)" definition mul::"nat option \ nat option \ nat option" (infixl"*\<^sub>o" 80) where "mul = option_shift (*)" fun option_comp_shift::"('a \ 'a \ bool) \ 'a option \ 'a option \ bool" where "option_comp_shift _ None _ = False"| "option_comp_shift _ _ None = False"| "option_comp_shift f (Some x) (Some y) = f x y" fun less::"nat option \ nat option \ bool" (infixl"<\<^sub>o" 80) where "less x y= option_comp_shift (<) x y" fun lesseq::"nat option \ nat option \ bool" (infixl"\\<^sub>o" 80) where "lesseq x y = option_comp_shift (\) x y" fun greater::"nat option \ nat option \ bool" (infixl">\<^sub>o" 80) where "greater x y = option_comp_shift (>) x y" lemma add_shift:"x+y = z \ Some x +\<^sub>o Some y = Some z" by (simp add: add_def) lemma mul_shift:"x*y = z \ Some x *\<^sub>o Some y = Some z" by (simp add: mul_def) lemma power_shift:"x^y = z \ Some x ^\<^sub>o Some y = Some z" by (simp add: power_def) lemma less_shift: "x < y \ Some x <\<^sub>o Some y" by simp lemma lesseq_shift: "x \ y \ Some x \\<^sub>o Some y" by simp lemma greater_shift: "x > y \ Some x >\<^sub>o Some y" by simp definition max_in_set :: "nat set \ nat \ bool" where "max_in_set xs x \ (x \ xs \ (\ y \ xs. y \ x))" lemma maxt_member: "invar_vebt t n \ vebt_maxt t = Some maxi \ vebt_member t maxi" proof(induction t n arbitrary: maxi rule: invar_vebt.induct) case (1 a b) then show ?case by (metis VEBT_Member.vebt_member.simps(1) vebt_maxt.simps(1) option.distinct(1) option.inject zero_neq_one) next case (2 treeList n summary m deg) then show ?case by simp next case (3 treeList n summary m deg) then show ?case by simp next case (4 treeList n summary m deg mi ma) hence "deg \ 2" by (metis One_nat_def Suc_le_eq add_mono deg_not_0 numeral_2_eq_2 plus_1_eq_Suc) then show ?case by (metis "4.prems" VEBT_Member.vebt_member.simps(5) Suc_diff_Suc Suc_pred lessI less_le_trans vebt_maxt.simps(3) numeral_2_eq_2 option.inject zero_less_Suc) next case (5 treeList n summary m deg mi ma) hence "deg \ 2" by (metis Suc_leI le_add2 less_add_same_cancel2 less_le_trans not_less_iff_gr_or_eq not_one_le_zero numeral_2_eq_2 plus_1_eq_Suc set_n_deg_not_0) then show ?case by (metis "5.prems" VEBT_Member.vebt_member.simps(5) add_2_eq_Suc le_add_diff_inverse vebt_maxt.simps(3) option.inject) qed lemma maxt_corr_help: "invar_vebt t n \ vebt_maxt t = Some maxi \ vebt_member t x \ maxi \ x " by (smt VEBT_Member.vebt_member.simps(1) le_less vebt_maxt.elims member_inv mi_ma_2_deg option.simps(1) option.simps(3) zero_le_one) lemma maxt_corr_help_empty: "invar_vebt t n \ vebt_maxt t = None \ set_vebt' t = {}" by (metis (full_types) VEBT_Member.vebt_member.simps(1) empty_Collect_eq vebt_maxt.elims minNull.simps(4) min_Null_member option.distinct(1) set_vebt'_def) theorem maxt_corr:assumes "invar_vebt t n" and "vebt_maxt t = Some x" shows "max_in_set (set_vebt' t) x" unfolding set_vebt'_def Max_def max_in_set_def using assms(1) assms(2) maxt_corr_help maxt_member by blast theorem maxt_sound:assumes "invar_vebt t n" and "max_in_set (set_vebt' t) x" shows "vebt_maxt t = Some x" by (metis (no_types, opaque_lifting) assms(1) assms(2) empty_Collect_eq le_less max_in_set_def maxt_corr_help maxt_corr_help_empty maxt_member mem_Collect_eq not_le option.exhaust set_vebt'_def) definition min_in_set :: "nat set \ nat \ bool" where "min_in_set xs x \ (x \ xs \ (\ y \ xs. y \ x))" lemma mint_member: "invar_vebt t n \ vebt_mint t = Some maxi \ vebt_member t maxi" proof(induction t n arbitrary: maxi rule: invar_vebt.induct) case (1 a b) then show ?case by (metis VEBT_Member.vebt_member.simps(1) vebt_mint.simps(1) option.distinct(1) option.inject zero_neq_one) next case (2 treeList n summary m deg) then show ?case by simp next case (3 treeList n summary m deg) then show ?case by simp next case (4 treeList n summary m deg mi ma) hence "deg \ 2" by (metis One_nat_def Suc_le_eq add_mono deg_not_0 numeral_2_eq_2 plus_1_eq_Suc) then show ?case by (metis "4.prems" VEBT_Member.vebt_member.simps(5) One_nat_def Suc_diff_Suc Suc_pred dual_order.strict_trans1 le_imp_less_Suc le_numeral_extra(4) vebt_mint.simps(3) numeral_2_eq_2 option.inject zero_le_one) next case (5 treeList n summary m deg mi ma) hence "deg \ 2" by (metis Suc_leI le_add2 less_add_same_cancel2 less_le_trans not_less_iff_gr_or_eq not_one_le_zero numeral_2_eq_2 plus_1_eq_Suc set_n_deg_not_0) then show ?case using "5.prems" VEBT_Member.vebt_member.simps(5) add_2_eq_Suc le_add_diff_inverse vebt_mint.simps(3) by (metis option.inject) qed lemma mint_corr_help: "invar_vebt t n \ vebt_mint t = Some mini \ vebt_member t x \ mini \ x " by (smt VEBT_Member.vebt_member.simps(1) eq_iff option.inject less_imp_le_nat member_inv mi_ma_2_deg vebt_mint.elims of_nat_0 of_nat_0_le_iff of_nat_le_iff option.simps(3)) lemma mint_corr_help_empty: "invar_vebt t n \ vebt_mint t = None \ set_vebt' t = {}" by (metis VEBT_internal.maxt_corr_help_empty option.distinct(1) vebt_maxt.simps(1) vebt_maxt.simps(2) vebt_mint.elims) theorem mint_corr:assumes "invar_vebt t n" and "vebt_mint t = Some x" shows "min_in_set (set_vebt' t) x" using assms(1) assms(2) min_in_set_def mint_corr_help mint_member set_vebt'_def by auto theorem mint_sound:assumes "invar_vebt t n" and "min_in_set (set_vebt' t) x" shows "vebt_mint t = Some x" by (metis assms(1) assms(2) empty_Collect_eq eq_iff mem_Collect_eq min_in_set_def mint_corr_help mint_corr_help_empty mint_member option.exhaust set_vebt'_def) lemma summaxma:assumes "invar_vebt (Node (Some (mi, ma)) deg treeList summary) deg" and "mi \ ma" shows "the (vebt_maxt summary) = high ma (deg div 2)" proof- from assms(1) show ?thesis proof(cases) case (4 n m) have "both_member_options summary (high ma n)" using "4"(10) "4"(2) "4"(4) "4"(5) "4"(6) "4"(9) assms(2) deg_not_0 exp_split_high_low(1) by blast have "high ma n \ the (vebt_maxt summary)" using "4"(2) \both_member_options summary (high ma n)\ empty_Collect_eq option.inject maxt_corr_help maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options by (metis option.exhaust_sel) have "high ma n < the (vebt_maxt summary) \ False" proof- assume "high ma n < the (vebt_maxt summary)" obtain maxs where "Some maxs = vebt_maxt summary" by (metis "4"(2) \both_member_options summary (high ma n)\ empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "\ x. both_member_options (treeList ! maxs) x" by (metis "4"(2) "4"(6) both_member_options_equiv_member maxt_member member_bound) then obtain x where "both_member_options (treeList ! maxs) x" by auto hence "vebt_member (treeList ! maxs) x" by (metis "4"(1) "4"(2) "4"(3) \Some maxs = vebt_maxt summary\ maxt_member member_bound nth_mem valid_member_both_member_options) have "maxs < 2^m" by (metis "4"(2) \Some maxs = vebt_maxt summary\ maxt_member member_bound) have "invar_vebt (treeList ! maxs) n" by (metis "4"(1) "4"(3) \maxs < 2 ^ m\ inthall member_def) hence "x < 2^n" using \vebt_member (treeList ! maxs) x\ member_bound by auto let ?X = "2^n*maxs + x" have "high ?X n = maxs" by (simp add: \x < 2 ^ n\ high_inv mult.commute) hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) (2^n*maxs + x)" by (metis "4"(3) "4"(4) "4"(5) One_nat_def Suc_leI \both_member_options (treeList ! maxs) x\ \maxs < 2 ^ m\ \x < 2 ^ n\ add_self_div_2 assms(1) both_member_options_from_chilf_to_complete_tree deg_not_0 low_inv mult.commute) hence "vebt_member (Node (Some (mi, ma)) deg treeList summary) ?X" using assms(1) both_member_options_equiv_member by auto have "high ?X n> high ma n" by (metis \Some maxs = vebt_maxt summary\ \high (2 ^ n * maxs + x) n = maxs\ \high ma n < the (vebt_maxt summary)\ option.exhaust_sel option.inject option.simps(3)) hence "?X > ma" by (metis div_le_mono high_def not_le) then show ?thesis by (metis "4"(8) \vebt_member (Node (Some (mi, ma)) deg treeList summary) (2 ^ n * maxs + x)\ leD member_inv not_less_iff_gr_or_eq) qed then show ?thesis using "4"(4) "4"(5) \high ma n \ the (vebt_maxt summary)\ by fastforce next case (5 n m) have "both_member_options summary (high ma n)" - by (metis "5"(10) "5"(5) "5"(6) "5"(9) Euclidean_Division.div_eq_0_iff assms(2) div_exp_eq high_def nat.simps(3) numerals(2) power_not_zero) + by (metis "5"(10) "5"(5) "5"(6) "5"(9) div_eq_0_iff assms(2) div_exp_eq high_def nat.simps(3) numerals(2) power_not_zero) have "high ma n \ the (vebt_maxt summary)" by (metis "5"(2) VEBT_Member.vebt_member.simps(2) \both_member_options summary (high ma n)\ vebt_maxt.elims maxt_corr_help minNull.simps(1) min_Null_member option.exhaust_sel option.simps(3) valid_member_both_member_options) have "high ma n < the (vebt_maxt summary) \ False" proof- assume "high ma n < the (vebt_maxt summary)" obtain maxs where "Some maxs = vebt_maxt summary" by (metis "5"(2) \both_member_options summary (high ma n)\ empty_Collect_eq maxt_corr_help_empty not_None_eq set_vebt'_def valid_member_both_member_options) hence "\ x. both_member_options (treeList ! maxs) x" by (metis "5"(2) "5"(6) both_member_options_equiv_member maxt_member member_bound) then obtain x where "both_member_options (treeList ! maxs) x" by auto hence "vebt_member (treeList ! maxs) x" by (metis "5"(1) "5"(2) "5"(3) \Some maxs = vebt_maxt summary\ both_member_options_equiv_member maxt_member member_bound nth_mem) have "maxs < 2^m" by (metis "5"(2) \Some maxs = vebt_maxt summary\ maxt_member member_bound) have "invar_vebt (treeList ! maxs) n" by (metis "5"(1) "5"(3) \maxs < 2 ^ m\ inthall member_def) hence "x < 2^n" using \vebt_member (treeList ! maxs) x\ member_bound by auto let ?X = "2^n*maxs + x" have "high ?X n = maxs" by (simp add: \x < 2 ^ n\ high_inv mult.commute) hence "both_member_options (Node (Some (mi, ma)) deg treeList summary) (2^n*maxs + x)" by (smt (z3) "5"(3) "5"(4) "5"(5) \both_member_options (treeList ! maxs) x\ \maxs < 2 ^ m\ \x < 2 ^ n\ add_Suc_right add_self_div_2 both_member_options_from_chilf_to_complete_tree even_Suc_div_two le_add1 low_inv mult.commute odd_add plus_1_eq_Suc) hence "vebt_member (Node (Some (mi, ma)) deg treeList summary) ?X" using assms(1) both_member_options_equiv_member by auto have "high ?X n> high ma n" by (metis \Some maxs = vebt_maxt summary\ \high (2 ^ n * maxs + x) n = maxs\ \high ma n < the (vebt_maxt summary)\ option.sel) hence "?X > ma" by (metis div_le_mono high_def not_le) then show ?thesis by (metis "5"(8) \vebt_member (Node (Some (mi, ma)) deg treeList summary) (2 ^ n * maxs + x)\ leD member_inv not_less_iff_gr_or_eq) qed then show ?thesis using "5"(4) "5"(5) \high ma n \ the(vebt_maxt summary)\ by fastforce qed qed lemma maxbmo: "vebt_maxt t = Some x \ both_member_options t x" apply(induction t rule: vebt_maxt.induct) apply auto apply (metis both_member_options_def naive_member.simps(1) option.distinct(1) option.sel zero_neq_one) by (metis One_nat_def Suc_le_D both_member_options_def div_by_1 div_greater_zero_iff membermima.simps(3) membermima.simps(4) not_gr0) lemma misiz:"invar_vebt t n \ Some m = vebt_mint t \ m < 2^n" by (metis member_bound mint_member) lemma mintlistlength: assumes "invar_vebt (Node (Some (mi, ma)) deg treeList summary) n " " mi \ ma " shows " ma > mi \ (\ m. Some m = vebt_mint summary \ m < 2^(n - n div 2))" using assms(1) proof cases case (4 n m) hence "both_member_options (treeList ! high ma n) (low ma n)" by (metis assms(2) high_bound_aux) moreover hence "both_member_options summary (high ma n)" using "4"(10) "4"(6) "4"(7) high_bound_aux by blast moreover then obtain mini where "Some mini = vebt_mint summary" by (metis "4"(3) empty_Collect_eq mint_corr_help_empty option.exhaust_sel set_vebt'_def valid_member_both_member_options) moreover hence "mini < 2^m" by (metis "4"(3) mint_member member_bound) moreover have "m = (deg - deg div 2)" using 4(6) 4(5) by auto ultimately show ?thesis using 4(1) assms 4(9) by auto next case (5 n m) hence "both_member_options (treeList ! high ma n) (low ma n)" by (metis assms(2) high_bound_aux) moreover hence "both_member_options summary (high ma n)" using "5"(10) "5"(6) "5"(7) high_bound_aux by blast moreover then obtain mini where "Some mini = vebt_mint summary" by (metis "5"(3) empty_Collect_eq mint_corr_help_empty option.exhaust_sel set_vebt'_def valid_member_both_member_options) moreover hence "mini < 2^m" by (metis "5"(3) mint_member member_bound) moreover have "m = (deg - deg div 2)" using 5(6) 5(5) by auto ultimately show ?thesis using 5(1) assms 5(9) by auto qed lemma power_minus_is_div: "b \ a \ (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b" apply (induct a arbitrary: b) apply simp apply (erule le_SucE) apply (clarsimp simp:Suc_diff_le le_iff_add power_add) apply simp done lemma nested_mint:assumes "invar_vebt (Node (Some (mi, ma)) deg treeList summary) n " "n = Suc (Suc va) "" \ ma < mi "" ma \ mi " shows " high (the (vebt_mint summary) * (2 * 2 ^ (va div 2)) + the (vebt_mint (treeList ! the (vebt_mint summary)))) (Suc (va div 2)) < length treeList" proof- have setprop: "t \ set treeList \ invar_vebt t (n div 2 )" for t using assms(1) by (cases) simp+ have listlength: "length treeList = 2^(n - n div 2)" using assms(1) by (cases) simp+ have sumprop: "invar_vebt summary (n - n div 2)" using assms(1) by (cases) simp+ have mimaxprop: "mi \ ma \ ma \ 2^n" using assms(1) by cases simp+ hence xbound: "mi \ x \ x \ ma \ high x (n div 2) \ length treeList " for x using div_le_dividend div_le_mono high_def listlength power_minus_is_div by auto have contcong:"i < length treeList \ \ x. both_member_options (treeList ! i) x \ both_member_options summary i " for i using assms(1)by cases auto+ obtain m where " Some m = vebt_mint summary \ m < 2^(n - n div 2)" using assms(1) assms(4) mintlistlength by blast then obtain miny where "(vebt_mint (treeList ! the (vebt_mint summary))) =Some miny" by (metis both_member_options_equiv_member contcong empty_Collect_eq listlength mint_corr_help_empty mint_member nth_mem option.exhaust_sel option.sel setprop sumprop set_vebt'_def) hence "miny < 2^(n div 2)" by (metis \\thesis. (\m. Some m = vebt_mint summary \ m < 2 ^ (n - n div 2) \ thesis) \ thesis\ listlength misiz nth_mem option.sel setprop) then show ?thesis by (metis \\thesis. (\m. Some m = vebt_mint summary \ m < 2 ^ (n - n div 2) \ thesis) \ thesis\ \vebt_mint (treeList ! the (vebt_mint summary)) = Some miny\ assms(2) div2_Suc_Suc high_inv listlength option.sel power_Suc) qed lemma minminNull: "vebt_mint t = None \ minNull t" by (metis minNull.simps(1) minNull.simps(4) vebt_mint.elims option.distinct(1)) lemma minNullmin: "minNull t \ vebt_mint t = None" by (metis minNull.elims(2) vebt_mint.simps(1) vebt_mint.simps(2)) end end diff --git a/thys/Word_Lib/More_Word_Operations.thy b/thys/Word_Lib/More_Word_Operations.thy --- a/thys/Word_Lib/More_Word_Operations.thy +++ b/thys/Word_Lib/More_Word_Operations.thy @@ -1,1015 +1,1023 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Misc word operations\ theory More_Word_Operations imports "HOL-Library.Word" Aligned Reversed_Bit_Lists More_Misc Signed_Words Word_Lemmas Word_EqI begin context includes bit_operations_syntax begin definition ptr_add :: "'a :: len word \ nat \ 'a word" where "ptr_add ptr n \ ptr + of_nat n" definition alignUp :: "'a::len word \ nat \ 'a word" where "alignUp x n \ x + 2 ^ n - 1 AND NOT (2 ^ n - 1)" lemma alignUp_unfold: \alignUp w n = (w + mask n) AND NOT (mask n)\ by (simp add: alignUp_def mask_eq_exp_minus_1 add_mask_fold) (* standard notation for blocks of 2^n-1 words, usually aligned; abbreviation so it simplifies directly *) abbreviation mask_range :: "'a::len word \ nat \ 'a word set" where "mask_range p n \ {p .. p + mask n}" definition w2byte :: "'a :: len word \ 8 word" where "w2byte \ ucast" (* Count leading zeros *) definition word_clz :: "'a::len word \ nat" where "word_clz w \ length (takeWhile Not (to_bl w))" (* Count trailing zeros *) definition word_ctz :: "'a::len word \ nat" where "word_ctz w \ length (takeWhile Not (rev (to_bl w)))" lemma word_ctz_unfold: \word_ctz w = length (takeWhile (Not \ bit w) [0.. for w :: \'a::len word\ by (simp add: word_ctz_def rev_to_bl_eq takeWhile_map) lemma word_ctz_unfold': \word_ctz w = Min (insert LENGTH('a) {n. bit w n})\ for w :: \'a::len word\ proof (cases \\n. bit w n\) case True then obtain n where \bit w n\ .. from \bit w n\ show ?thesis apply (simp add: word_ctz_unfold) apply (subst Min_eq_length_takeWhile [symmetric]) apply (auto simp add: bit_imp_le_length) apply (subst Min_insert) apply auto apply (subst min.absorb2) apply (subst Min_le_iff) apply auto apply (meson bit_imp_le_length order_less_le) done next case False then have \bit w = bot\ by auto then have \word_ctz w = LENGTH('a)\ by (simp add: word_ctz_def rev_to_bl_eq bot_fun_def map_replicate_const) with \bit w = bot\ show ?thesis by simp qed lemma word_ctz_le: "word_ctz (w :: ('a::len word)) \ LENGTH('a)" apply (clarsimp simp: word_ctz_def) using length_takeWhile_le apply (rule order_trans) apply simp done lemma word_ctz_less: "w \ 0 \ word_ctz (w :: ('a::len word)) < LENGTH('a)" apply (clarsimp simp: word_ctz_def eq_zero_set_bl) using length_takeWhile_less apply (rule less_le_trans) apply auto done lemma take_bit_word_ctz_eq [simp]: \take_bit LENGTH('a) (word_ctz w) = word_ctz w\ for w :: \'a::len word\ apply (simp add: take_bit_nat_eq_self_iff word_ctz_def to_bl_unfold) using length_takeWhile_le apply (rule le_less_trans) apply simp done lemma word_ctz_not_minus_1: \word_of_nat (word_ctz (w :: 'a :: len word)) \ (- 1 :: 'a::len word)\ if \1 < LENGTH('a)\ proof - note word_ctz_le also from that have \LENGTH('a) < mask LENGTH('a)\ by (simp add: less_mask) finally have \word_ctz w < mask LENGTH('a)\ . then have \word_of_nat (word_ctz w) < (word_of_nat (mask LENGTH('a)) :: 'a word)\ by (simp add: of_nat_word_less_iff) also have \\ = - 1\ by (rule bit_word_eqI) (simp add: bit_simps) finally show ?thesis by simp qed lemma unat_of_nat_ctz_mw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w" by (simp add: unsigned_of_nat) lemma unat_of_nat_ctz_smw: "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len signed word) = word_ctz w" by (simp add: unsigned_of_nat) definition word_log2 :: "'a::len word \ nat" where "word_log2 (w::'a::len word) \ size w - 1 - word_clz w" (* Bit population count. Equivalent of __builtin_popcount. *) definition pop_count :: "('a::len) word \ nat" where "pop_count w \ length (filter id (to_bl w))" (* Sign extension from bit n *) definition sign_extend :: "nat \ 'a::len word \ 'a word" where "sign_extend n w \ if bit w n then w OR NOT (mask n) else w AND mask n" lemma sign_extend_eq_signed_take_bit: \sign_extend = signed_take_bit\ proof (rule ext)+ fix n and w :: \'a::len word\ show \sign_extend n w = signed_take_bit n w\ proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ then show \bit (sign_extend n w) q \ bit (signed_take_bit n w) q\ by (auto simp add: bit_signed_take_bit_iff sign_extend_def bit_and_iff bit_or_iff bit_not_iff bit_mask_iff not_less exp_eq_0_imp_not_bit not_le min_def) qed qed definition sign_extended :: "nat \ 'a::len word \ bool" where "sign_extended n w \ \i. n < i \ i < size w \ bit w i = bit w n" lemma ptr_add_0 [simp]: "ptr_add ref 0 = ref " unfolding ptr_add_def by simp lemma pop_count_0[simp]: "pop_count 0 = 0" by (clarsimp simp:pop_count_def) lemma pop_count_1[simp]: "pop_count 1 = 1" by (clarsimp simp:pop_count_def to_bl_1) lemma pop_count_0_imp_0: "(pop_count w = 0) = (w = 0)" apply (rule iffI) apply (clarsimp simp:pop_count_def) apply (subst (asm) filter_empty_conv) apply (clarsimp simp:eq_zero_set_bl) apply fast apply simp done lemma word_log2_zero_eq [simp]: \word_log2 0 = 0\ by (simp add: word_log2_def word_clz_def word_size) lemma word_log2_unfold: \word_log2 w = (if w = 0 then 0 else Max {n. bit w n})\ for w :: \'a::len word\ proof (cases \w = 0\) case True then show ?thesis by simp next case False then obtain r where \bit w r\ by (auto simp add: bit_eq_iff) then have \Max {m. bit w m} = LENGTH('a) - Suc (length (takeWhile (Not \ bit w) (rev [0.. by (subst Max_eq_length_takeWhile [of _ \LENGTH('a)\]) (auto simp add: bit_imp_le_length) then have \word_log2 w = Max {x. bit w x}\ by (simp add: word_log2_def word_clz_def word_size to_bl_unfold rev_map takeWhile_map) with \w \ 0\ show ?thesis by simp qed lemma word_log2_eqI: \word_log2 w = n\ if \w \ 0\ \bit w n\ \\m. bit w m \ m \ n\ for w :: \'a::len word\ proof - from \w \ 0\ have \word_log2 w = Max {n. bit w n}\ by (simp add: word_log2_unfold) also have \Max {n. bit w n} = n\ using that by (auto intro: Max_eqI) finally show ?thesis . qed lemma bit_word_log2: \bit w (word_log2 w)\ if \w \ 0\ proof - from \w \ 0\ have \\r. bit w r\ by (auto intro: bit_eqI) then obtain r where \bit w r\ .. from \w \ 0\ have \word_log2 w = Max {n. bit w n}\ by (simp add: word_log2_unfold) also have \Max {n. bit w n} \ {n. bit w n}\ using \bit w r\ by (subst Max_in) auto finally show ?thesis by simp qed lemma word_log2_maximum: \n \ word_log2 w\ if \bit w n\ proof - have \n \ Max {n. bit w n}\ using that by (auto intro: Max_ge) also from that have \w \ 0\ by force then have \Max {n. bit w n} = word_log2 w\ by (simp add: word_log2_unfold) finally show ?thesis . qed lemma word_log2_nth_same: "w \ 0 \ bit w (word_log2 w)" by (drule bit_word_log2) simp lemma word_log2_nth_not_set: "\ word_log2 w < i ; i < size w \ \ \ bit w i" using word_log2_maximum [of w i] by auto lemma word_log2_highest: assumes a: "bit w i" shows "i \ word_log2 w" using a by (simp add: word_log2_maximum) lemma word_log2_max: "word_log2 w < size w" apply (cases \w = 0\) apply (simp_all add: word_size) apply (drule bit_word_log2) apply (fact bit_imp_le_length) done lemma word_clz_0[simp]: "word_clz (0::'a::len word) = LENGTH('a)" unfolding word_clz_def by simp lemma word_clz_minus_one[simp]: "word_clz (-1::'a::len word) = 0" unfolding word_clz_def by simp lemma is_aligned_alignUp[simp]: "is_aligned (alignUp p n) n" by (simp add: alignUp_def is_aligned_mask mask_eq_decr_exp word_bw_assocs) lemma alignUp_le[simp]: "alignUp p n \ p + 2 ^ n - 1" unfolding alignUp_def by (rule word_and_le2) lemma alignUp_idem: fixes a :: "'a::len word" assumes "is_aligned a n" "n < LENGTH('a)" shows "alignUp a n = a" using assms unfolding alignUp_def by (metis add_cancel_right_right add_diff_eq and_mask_eq_iff_le_mask mask_eq_decr_exp mask_out_add_aligned order_refl word_plus_and_or_coroll2) lemma alignUp_not_aligned_eq: fixes a :: "'a :: len word" assumes al: "\ is_aligned a n" and sz: "n < LENGTH('a)" shows "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" proof - + from \n < LENGTH('a)\ have \(2::int) ^ n < 2 ^ LENGTH('a)\ + by simp + with take_bit_int_less_exp [of n] + have *: \take_bit n k < 2 ^ LENGTH('a)\ for k :: int + by (rule less_trans) have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) fact+ - - then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz - by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans - unat_less_power word_less_sub_le word_mod_less_divisor) - + then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" + apply (transfer fixing: n) using sz + apply (simp flip: take_bit_eq_mod add: div_eq_0_iff) + apply (subst take_bit_int_eq_self) + using * + apply (auto simp add: diff_less_eq intro: less_imp_le) + apply (simp add: less_le) + done have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1" by (simp add: word_mod_div_equality) also have "\ = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n" by (simp add: field_simps) finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz unfolding alignUp_def apply (subst mask_eq_decr_exp [symmetric]) apply (erule ssubst) apply (subst neg_mask_is_div) apply (simp add: word_arith_nat_div) apply (subst unat_word_ariths(1) unat_word_ariths(2))+ apply (subst uno_simps) apply (subst unat_1) apply (subst mod_add_right_eq) apply simp apply (subst power_mod_div) apply (subst div_mult_self1) apply simp apply (subst um) apply simp apply (subst mod_mod_power) apply simp apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst mult_mod_left) apply (subst power_add [symmetric]) apply simp apply (subst Abs_fnat_hom_1) apply (subst Abs_fnat_hom_add) apply (subst word_unat_power, subst Abs_fnat_hom_mult) apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult) apply simp done qed lemma alignUp_ge: fixes a :: "'a :: len word" assumes sz: "n < LENGTH('a)" and nowrap: "alignUp a n \ 0" shows "a \ alignUp a n" proof (cases "is_aligned a n") case True then show ?thesis using sz by (subst alignUp_idem, simp_all) next case False have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have"2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le) moreover have "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using nowrap sz apply - apply (erule contrapos_nn) apply (subst alignUp_not_aligned_eq [OF False sz]) apply (subst unat_arith_simps) apply (subst unat_word_ariths) apply (subst unat_word_ariths) apply simp apply (subst mult_mod_left) apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power) done ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric]) also have "\ < (a div 2 ^ n + 1) * 2 ^ n" using sz lt apply (simp add: field_simps) apply (rule word_add_less_mono1) apply (rule word_mod_less_divisor) apply (simp add: word_less_nat_alt) apply (subst unat_word_ariths) apply (simp add: unat_div) done also have "\ = alignUp a n" by (rule alignUp_not_aligned_eq [symmetric]) fact+ finally show ?thesis by (rule order_less_imp_le) qed lemma alignUp_le_greater_al: fixes x :: "'a :: len word" assumes le: "a \ x" and sz: "n < LENGTH('a)" and al: "is_aligned x n" shows "alignUp a n \ x" proof (cases "is_aligned a n") case True then show ?thesis using sz le by (simp add: alignUp_idem) next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)" by (auto elim!: is_alignedE) then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)" using sz apply (subst unat_of_nat_eq) apply (erule order_less_le_trans) apply simp apply (subst mult.commute) apply simp apply (rule nat_less_power_trans) apply simp apply simp done have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ also have "\ \ of_nat k * 2 ^ n" proof (rule word_mult_le_mono1 [OF inc_le _ kn]) show "a div 2 ^ n < of_nat k" using kv xk le sz anz by (simp add: alignUp_div_helper) show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz) qed finally show ?thesis using xk by (simp add: field_simps) qed lemma alignUp_is_aligned_nz: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and ax: "a \ x" and az: "a \ 0" shows "alignUp (a::'a :: len word) n \ 0" proof (cases "is_aligned a n") case True then have "alignUp a n = a" using sz by (simp add: alignUp_idem) then show ?thesis using az by simp next case False then have anz: "a mod 2 ^ n \ 0" by (rule not_aligned_mod_nz) { assume asm: "alignUp a n = 0" have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz by (metis le_add_diff_inverse2 less_mult_imp_div_less order_less_imp_le power_add unsigned_less) have leq: "2 ^ n * (unat a div 2 ^ n + 1) \ 2 ^ LENGTH('a)" using sz by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans order_less_imp_le) from al obtain k where kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k" by (auto elim!: is_alignedE) then have "a div 2 ^ n < of_nat k" using ax sz anz by (rule alignUp_div_helper) then have r: "unat a div 2 ^ n < k" using sz by (simp flip: drop_bit_eq_div unat_drop_bit_eq) (metis leI le_unat_uoi unat_mono) have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" by (rule alignUp_not_aligned_eq) fact+ then have "\ = 0" using asm by simp then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)" using sz by (simp add: unat_arith_simps ac_simps) (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd) with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)" by (force elim!: le_SucE) then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1" by (metis (no_types, opaque_lifting) Groups.add_ac(2) add.right_neutral add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0 le_neq_implies_less power_eq_0_iff zero_neq_numeral) then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1" using sz by (simp add: power_sub) then have "2 ^ (LENGTH('a) - n) - 1 < k" using r by simp then have False using kv by simp } then show ?thesis by clarsimp qed lemma alignUp_ar_helper: fixes a :: "'a :: len word" assumes al: "is_aligned x n" and sz: "n < LENGTH('a)" and sub: "{x..x + 2 ^ n - 1} \ {a..b}" and anz: "a \ 0" shows "a \ alignUp a n \ alignUp a n + 2 ^ n - 1 \ b" proof from al have xl: "x \ x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow) from xl sub have ax: "a \ x" by auto show "a \ alignUp a n" proof (rule alignUp_ge) show "alignUp a n \ 0" using al sz ax anz by (rule alignUp_is_aligned_nz) qed fact+ show "alignUp a n + 2 ^ n - 1 \ b" proof (rule order_trans) from xl show tp: "x + 2 ^ n - 1 \ b" using sub by auto from ax have "alignUp a n \ x" by (rule alignUp_le_greater_al) fact+ then have "alignUp a n + (2 ^ n - 1) \ x + (2 ^ n - 1)" using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast then show "alignUp a n + 2 ^ n - 1 \ x + 2 ^ n - 1" by (simp add: field_simps) qed qed lemma alignUp_def2: "alignUp a sz = a + 2 ^ sz - 1 AND NOT (mask sz)" by (simp add: alignUp_def flip: mask_eq_decr_exp) lemma alignUp_def3: "alignUp a sz = 2^ sz + (a - 1 AND NOT (mask sz))" by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned) lemma alignUp_plus: "is_aligned w us \ alignUp (w + a) us = w + alignUp a us" by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps) lemma alignUp_distance: "alignUp (q :: 'a :: len word) sz - q \ mask sz" by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff) lemma is_aligned_diff_neg_mask: "is_aligned p sz \ (p - q AND NOT (mask sz)) = (p - ((alignUp q sz) AND NOT (mask sz)))" apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus) apply (subst mask_out_add_aligned[symmetric]; simp) apply (simp add: eq_neg_iff_add_eq_0) apply (subst add.commute) apply (simp add: alignUp_distance is_aligned_neg_mask_eq mask_out_add_aligned and_mask_eq_iff_le_mask flip: mask_eq_x_eq_0) done lemma word_clz_max: "word_clz w \ size (w::'a::len word)" unfolding word_clz_def by (metis length_takeWhile_le word_size_bl) lemma word_clz_nonzero_max: fixes w :: "'a::len word" assumes nz: "w \ 0" shows "word_clz w < size (w::'a::len word)" proof - { assume a: "word_clz w = size (w::'a::len word)" hence "length (takeWhile Not (to_bl w)) = length (to_bl w)" by (simp add: word_clz_def word_size) hence allj: "\j\set(to_bl w). \ j" by (metis a length_takeWhile_less less_irrefl_nat word_clz_def) hence "to_bl w = replicate (length (to_bl w)) False" using eq_zero_set_bl nz by fastforce hence "w = 0" by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep') with nz have False by simp } thus ?thesis using word_clz_max by (fastforce intro: le_neq_trans) qed (* Sign extension from bit n. *) lemma bin_sign_extend_iff [bit_simps]: \bit (sign_extend e w) i \ bit w (min e i)\ if \i < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: sign_extend_def bit_simps min_def) lemma sign_extend_bitwise_if: "i < size w \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)" by (simp add: word_size bit_simps) lemma sign_extend_bitwise_if' [word_eqI_simps]: \i < LENGTH('a) \ bit (sign_extend e w) i \ (if i < e then bit w i else bit w e)\ for w :: \'a::len word\ using sign_extend_bitwise_if [of i w e] by (simp add: word_size) lemma sign_extend_bitwise_disj: "i < size w \ bit (sign_extend e w) i \ i \ e \ bit w i \ e \ i \ bit w e" by (auto simp: sign_extend_bitwise_if) lemma sign_extend_bitwise_cases: "i < size w \ bit (sign_extend e w) i \ (i \ e \ bit w i) \ (e \ i \ bit w e)" by (auto simp: sign_extend_bitwise_if) lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size] lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size] (* Often, it is easier to reason about an operation which does not overwrite the bit which determines which mask operation to apply. *) lemma sign_extend_def': "sign_extend n w = (if bit w n then w OR NOT (mask (Suc n)) else w AND mask (Suc n))" by (rule bit_word_eqI) (auto simp add: bit_simps sign_extend_eq_signed_take_bit min_def less_Suc_eq_le) lemma sign_extended_sign_extend: "sign_extended n (sign_extend n w)" by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if) lemma sign_extended_iff_sign_extend: "sign_extended n w \ sign_extend n w = w" apply auto apply (auto simp add: bit_eq_iff) apply (simp_all add: bit_simps sign_extend_eq_signed_take_bit not_le min_def sign_extended_def word_size split: if_splits) using le_imp_less_or_eq apply auto done lemma sign_extended_weaken: "sign_extended n w \ n \ m \ sign_extended m w" unfolding sign_extended_def by (cases "n < m") auto lemma sign_extend_sign_extend_eq: "sign_extend m (sign_extend n w) = sign_extend (min m n) w" by (rule bit_word_eqI) (simp add: sign_extend_eq_signed_take_bit bit_simps) lemma sign_extended_high_bits: "\ sign_extended e p; j < size p; e \ i; i < j \ \ bit p i = bit p j" by (drule (1) sign_extended_weaken; simp add: sign_extended_def) lemma sign_extend_eq: "w AND mask (Suc n) = v AND mask (Suc n) \ sign_extend n w = sign_extend n v" by (simp flip: take_bit_eq_mask add: sign_extend_eq_signed_take_bit signed_take_bit_eq_iff_take_bit_eq) lemma sign_extended_add: assumes p: "is_aligned p n" assumes f: "f < 2 ^ n" assumes e: "n \ e" assumes "sign_extended e p" shows "sign_extended e (p + f)" proof (cases "e < size p") case True note and_or = is_aligned_add_or[OF p f] have "\ bit f e" using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f] by (fastforce simp: word_size) hence i: "bit (p + f) e = bit p e" by (simp add: and_or bit_simps) have fm: "f AND mask e = f" by (fastforce intro: subst[where P="\f. f AND mask e = f", OF less_mask_eq[OF f]] simp: mask_twice e) show ?thesis using assms apply (simp add: sign_extended_iff_sign_extend sign_extend_def i) apply (simp add: and_or word_bw_comms[of p f]) apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits) done next case False thus ?thesis by (simp add: sign_extended_def word_size) qed lemma sign_extended_neq_mask: "\sign_extended n ptr; m \ n\ \ sign_extended n (ptr AND NOT (mask m))" by (fastforce simp: sign_extended_def word_size neg_mask_test_bit bit_simps) definition "limited_and (x :: 'a :: len word) y \ (x AND y = x)" lemma limited_and_eq_0: "\ limited_and x z; y AND NOT z = y \ \ x AND y = 0" unfolding limited_and_def apply (subst arg_cong2[where f="(AND)"]) apply (erule sym)+ apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs) done lemma limited_and_eq_id: "\ limited_and x z; y AND z = z \ \ x AND y = x" unfolding limited_and_def by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms) lemma lshift_limited_and: "limited_and x z \ limited_and (x << n) (z << n)" using push_bit_and [of n x z] by (simp add: limited_and_def shiftl_def) lemma rshift_limited_and: "limited_and x z \ limited_and (x >> n) (z >> n)" using drop_bit_and [of n x z] by (simp add: limited_and_def shiftr_def) lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id lemmas is_aligned_limited_and = is_aligned_neg_mask_eq[unfolded mask_eq_decr_exp, folded limited_and_def] lemmas limited_and_simps = limited_and_simps1 limited_and_simps1[OF is_aligned_limited_and] limited_and_simps1[OF lshift_limited_and] limited_and_simps1[OF rshift_limited_and] limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and] not_one_eq definition from_bool :: "bool \ 'a::len word" where "from_bool b \ case b of True \ of_nat 1 | False \ of_nat 0" lemma from_bool_eq: \from_bool = of_bool\ by (simp add: fun_eq_iff from_bool_def) lemma from_bool_0: "(from_bool x = 0) = (\ x)" by (simp add: from_bool_def split: bool.split) lemma from_bool_eq_if': "((if P then 1 else 0) = from_bool Q) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) definition to_bool :: "'a::len word \ bool" where "to_bool \ (\) 0" lemma to_bool_and_1: "to_bool (x AND 1) \ bit x 0" by (simp add: to_bool_def word_and_1) lemma to_bool_from_bool [simp]: "to_bool (from_bool r) = r" unfolding from_bool_def to_bool_def by (simp split: bool.splits) lemma from_bool_neq_0 [simp]: "(from_bool b \ 0) = b" by (simp add: from_bool_def split: bool.splits) lemma from_bool_mask_simp [simp]: "(from_bool r :: 'a::len word) AND 1 = from_bool r" unfolding from_bool_def by (clarsimp split: bool.splits) lemma from_bool_1 [simp]: "(from_bool P = 1) = P" by (simp add: from_bool_def split: bool.splits) lemma ge_0_from_bool [simp]: "(0 < from_bool P) = P" by (simp add: from_bool_def split: bool.splits) lemma limited_and_from_bool: "limited_and (from_bool b) 1" by (simp add: from_bool_def limited_and_def split: bool.split) lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def) lemma to_bool_0 [simp]: "\to_bool 0" by (simp add: to_bool_def) lemma from_bool_eq_if: "(from_bool Q = (if P then 1 else 0)) = (P = Q)" by (cases Q) (simp_all add: from_bool_def) lemma to_bool_eq_0: "(\ to_bool x) = (x = 0)" by (simp add: to_bool_def) lemma to_bool_neq_0: "(to_bool x) = (x \ 0)" by (simp add: to_bool_def) lemma from_bool_all_helper: "(\bool. from_bool bool = val \ P bool) = ((\bool. from_bool bool = val) \ P (val \ 0))" by (auto simp: from_bool_0) lemma fold_eq_0_to_bool: "(v = 0) = (\ to_bool v)" by (simp add: to_bool_def) lemma from_bool_to_bool_iff: "w = from_bool b \ to_bool w = b \ (w = 0 \ w = 1)" by (cases b) (auto simp: from_bool_def to_bool_def) lemma from_bool_eqI: "from_bool x = from_bool y \ x = y" unfolding from_bool_def by (auto split: bool.splits) lemma neg_mask_in_mask_range: "is_aligned ptr bits \ (ptr' AND NOT(mask bits) = ptr) = (ptr' \ mask_range ptr bits)" apply (erule is_aligned_get_word_bits) apply (rule iffI) apply (drule sym) apply (simp add: word_and_le2) apply (subst word_plus_and_or_coroll, word_eqI_solve) apply (metis bit.disj_ac(2) bit.disj_conj_distrib2 le_word_or2 word_and_max word_or_not) apply clarsimp apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le word_and_not) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_offset_in_range: "\ is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \ m; n < LENGTH('a) \ \ (x + y \ {p .. p + mask n}) = (x \ mask_range p n)" apply (subst disjunctive_add) apply (simp add: bit_simps) apply (erule is_alignedE') apply (auto simp add: bit_simps not_le)[1] apply (metis less_2p_is_upper_bits_unset) apply (simp only: is_aligned_add_or word_ao_dist flip: neg_mask_in_mask_range) apply (subgoal_tac \y AND NOT (mask n) = 0\) apply simp apply (metis (full_types) is_aligned_mask is_aligned_neg_mask less_mask_eq word_bw_comms(1) word_bw_lcs(1)) done lemma mask_range_to_bl': "\ is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \ \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (rule set_eqI, rule iffI) apply clarsimp apply (subgoal_tac "\y. x = ptr + y \ y < 2 ^ bits") apply clarsimp apply (subst is_aligned_add_conv) apply assumption apply simp apply simp apply (rule_tac x="x - ptr" in exI) apply (simp add: add_diff_eq[symmetric]) apply (simp only: word_less_sub_le[symmetric]) apply (rule word_diff_ls') apply (simp add: field_simps mask_eq_decr_exp) apply assumption apply simp apply (subgoal_tac "\y. y < 2 ^ bits \ to_bl (ptr + y) = to_bl x") apply clarsimp apply (rule conjI) apply (erule(1) is_aligned_no_wrap') apply (simp only: add_diff_eq[symmetric] mask_eq_decr_exp) apply (rule word_plus_mono_right) apply simp apply (erule is_aligned_no_wrap') apply simp apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI) apply (rule context_conjI) apply (rule order_less_le_trans [OF of_bl_length]) apply simp apply simp apply (subst is_aligned_add_conv) apply assumption apply simp apply (drule sym) apply (simp add: word_rep_drop) done lemma mask_range_to_bl: "is_aligned (ptr :: 'a :: len word) bits \ mask_range ptr bits = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}" apply (erule is_aligned_get_word_bits) apply (erule(1) mask_range_to_bl') apply (rule set_eqI) apply (simp add: power_overflow mask_eq_decr_exp) done lemma aligned_mask_range_cases: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \ \ mask_range p n \ mask_range p' n' = {} \ mask_range p n \ mask_range p' n' \ mask_range p n \ mask_range p' n'" apply (simp add: mask_range_to_bl) apply (rule Meson.disj_comm, rule disjCI) apply auto apply (subgoal_tac "(\n''. LENGTH('a) - n = (LENGTH('a) - n') + n'') \ (\n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')") apply (fastforce simp: take_add) apply arith done lemma aligned_mask_range_offset_subset: assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'" and szv: "sz' \ sz" and xsz: "x < 2 ^ sz" shows "mask_range (ptr+x) sz' \ mask_range ptr sz" using al proof (rule is_aligned_get_word_bits) assume p0: "ptr = 0" and szv': "LENGTH ('a) \ sz" then have "(2 ::'a word) ^ sz = 0" by simp show ?thesis using p0 by (simp add: \2 ^ sz = 0\ mask_eq_decr_exp) next assume szv': "sz < LENGTH('a)" hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)" using szv by auto show ?thesis using szv szv' apply auto using al assms(4) is_aligned_no_wrap' apply blast apply (simp only: flip: add_diff_eq add_mask_fold) apply (subst add.assoc, rule word_plus_mono_right) using al' is_aligned_add_less_t2n xsz apply fastforce apply (simp add: field_simps szv al is_aligned_no_overflow) done qed lemma aligned_mask_ranges_disjoint: "\ is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n'; p AND NOT(mask n') \ p'; p' AND NOT(mask n) \ p \ \ mask_range p n \ mask_range p' n' = {}" using aligned_mask_range_cases by (auto simp: neg_mask_in_mask_range) lemma aligned_mask_ranges_disjoint2: "\ is_aligned p n; is_aligned ptr bits; n \ m; n < size p; m \ bits; (\y < 2 ^ (n - m). p + (y << m) \ mask_range ptr bits) \ \ mask_range p n \ mask_range ptr bits = {}" apply safe apply (simp only: flip: neg_mask_in_mask_range) apply (drule_tac x="x AND mask n >> m" in spec) apply (erule notE[OF mp]) apply (simp flip: take_bit_eq_mask add: shiftr_def drop_bit_take_bit) apply transfer apply simp apply (simp add: word_size and_mask_less_size) apply (subst disjunctive_add) apply (auto simp add: bit_simps word_size intro!: bit_eqI) done lemma word_clz_sint_upper[simp]: "LENGTH('a) \ 3 \ sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \ int (LENGTH('a))" using word_clz_max [of w] apply (simp add: word_size signed_of_nat) apply (subst signed_take_bit_int_eq_self) apply simp_all apply (metis negative_zle of_nat_numeral semiring_1_class.of_nat_power) apply (drule small_powers_of_2) apply (erule le_less_trans) apply simp done lemma word_clz_sint_lower[simp]: "LENGTH('a) \ 3 \ - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \ int (LENGTH('a))" apply (subst sint_eq_uint) using word_clz_max [of w] apply (simp_all add: word_size unsigned_of_nat) apply (rule not_msb_from_less) apply (simp add: word_less_nat_alt unsigned_of_nat) apply (subst take_bit_nat_eq_self) apply (simp add: le_less_trans) apply (drule small_powers_of_2) apply (erule le_less_trans) apply simp done lemma mask_range_subsetD: "\ p' \ mask_range p n; x' \ mask_range p' n'; n' \ n; is_aligned p n; is_aligned p' n' \ \ x' \ mask_range p n" using aligned_mask_step by fastforce lemma add_mult_in_mask_range: "\ is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \ n; x < 2 ^ (n - bits) \ \ base + x * 2^bits \ mask_range base n" by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2 word_plus_mono_right) lemma from_to_bool_last_bit: "from_bool (to_bool (x AND 1)) = x AND 1" by (metis from_bool_to_bool_iff word_and_1) lemma sint_ctz: \0 \ sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word) \ sint (of_nat (word_ctz x) :: 'a signed word) \ int (LENGTH('a))\ (is \?P \ ?Q\) if \LENGTH('a) > 2\ proof have *: \word_ctz x < 2 ^ (LENGTH('a) - Suc 0)\ using word_ctz_le apply (rule le_less_trans) using that small_powers_of_2 [of \LENGTH('a)\] apply simp done have \int (word_ctz x) div 2 ^ (LENGTH('a) - Suc 0) = 0\ apply (rule div_pos_pos_trivial) apply (simp_all add: *) done then show ?P by (simp add: signed_of_nat bit_iff_odd) show ?Q apply (auto simp add: signed_of_nat) apply (subst signed_take_bit_int_eq_self) apply (auto simp add: word_ctz_le * minus_le_iff [of _ \int (word_ctz x)\]) apply (rule order.trans [of _ 0]) apply simp_all done qed lemma unat_of_nat_word_log2: "LENGTH('a) < 2 ^ LENGTH('b) \ unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n" by (metis less_trans unat_of_nat_eq word_log2_max word_size) lemma aligned_mask_diff: "\ is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz; bits \ sz; sz < LENGTH('a); dest < ptr \ \ mask bits + dest < ptr" apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption) apply (elim disjE) apply (drule_tac is_aligned_no_overflow_mask, simp)+ apply (simp add: algebra_split_simps word_le_not_less) apply (drule is_aligned_no_overflow_mask; fastforce) apply (simp add: is_aligned_weaken algebra_split_simps) apply (auto simp add: not_le) using is_aligned_no_overflow_mask leD apply blast apply (meson aligned_add_mask_less_eq is_aligned_weaken le_less_trans) done end end \ No newline at end of file