diff --git a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy --- a/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy +++ b/thys/Affine_Arithmetic/Executable_Euclidean_Space.thy @@ -1,1078 +1,1075 @@ section \Euclidean Space: Executability\ theory Executable_Euclidean_Space imports "HOL-Analysis.Multivariate_Analysis" "List-Index.List_Index" "HOL-Library.Float" "HOL-Library.Code_Cardinality" Affine_Arithmetic_Auxiliarities begin subsection \Ordered representation of Basis and Rounding of Components\ class executable_euclidean_space = ordered_euclidean_space + fixes Basis_list eucl_down eucl_truncate_down eucl_truncate_up assumes eucl_down_def: "eucl_down p b = (\i \ Basis. round_down p (b \ i) *\<^sub>R i)" assumes eucl_truncate_down_def: "eucl_truncate_down q b = (\i \ Basis. truncate_down q (b \ i) *\<^sub>R i)" assumes eucl_truncate_up_def: "eucl_truncate_up q b = (\i \ Basis. truncate_up q (b \ i) *\<^sub>R i)" assumes Basis_list[simp]: "set Basis_list = Basis" assumes distinct_Basis_list[simp]: "distinct Basis_list" begin lemma length_Basis_list: "length Basis_list = card Basis" by (metis Basis_list distinct_Basis_list distinct_card) end lemma eucl_truncate_down_zero[simp]: "eucl_truncate_down p 0 = 0" by (auto simp: eucl_truncate_down_def truncate_down_zero) lemma eucl_truncate_up_zero[simp]: "eucl_truncate_up p 0 = 0" by (auto simp: eucl_truncate_up_def) subsection \Instantiations\ instantiation real::executable_euclidean_space begin definition Basis_list_real :: "real list" where "Basis_list_real = [1]" definition "eucl_down prec b = round_down prec b" definition "eucl_truncate_down prec b = truncate_down prec b" definition "eucl_truncate_up prec b = truncate_up prec b" instance proof qed (auto simp: Basis_list_real_def eucl_down_real_def eucl_truncate_down_real_def eucl_truncate_up_real_def) end instantiation prod::(executable_euclidean_space, executable_euclidean_space) executable_euclidean_space begin definition Basis_list_prod :: "('a \ 'b) list" where "Basis_list_prod = zip Basis_list (replicate (length (Basis_list::'a list)) 0) @ zip (replicate (length (Basis_list::'b list)) 0) Basis_list" definition "eucl_down p a = (eucl_down p (fst a), eucl_down p (snd a))" definition "eucl_truncate_down p a = (eucl_truncate_down p (fst a), eucl_truncate_down p (snd a))" definition "eucl_truncate_up p a = (eucl_truncate_up p (fst a), eucl_truncate_up p (snd a))" instance proof show "set Basis_list = (Basis::('a*'b) set)" by (auto simp: Basis_list_prod_def Basis_prod_def elim!: in_set_zipE) (auto simp: Basis_list[symmetric] in_set_zip in_set_conv_nth simp del: Basis_list) show "distinct (Basis_list::('a*'b)list)" using distinct_Basis_list[where 'a='a] distinct_Basis_list[where 'a='b] by (auto simp: Basis_list_prod_def Basis_list intro: distinct_zipI1 distinct_zipI2 elim!: in_set_zipE) qed (auto simp: eucl_down_prod_def eucl_truncate_down_prod_def eucl_truncate_up_prod_def sum_Basis_prod_eq inner_add_left inner_sum_left inner_Basis eucl_down_def eucl_truncate_down_def eucl_truncate_up_def intro!: euclidean_eqI[where 'a="'a*'b"]) end lemma eucl_truncate_down_Basis[simp]: "i \ Basis \ eucl_truncate_down e x \ i = truncate_down e (x \ i)" by (simp add: eucl_truncate_down_def) lemma eucl_truncate_down_correct: "dist (x::'a::executable_euclidean_space) (eucl_down e x) \ {0..sqrt (DIM('a)) * 2 powr of_int (- e)}" proof - have "dist x (eucl_down e x) = sqrt (\i\Basis. (dist (x \ i) (eucl_down e x \ i))\<^sup>2)" unfolding euclidean_dist_l2[where 'a='a] L2_set_def .. also have "\ \ sqrt (\i\(Basis::'a set). ((2 powr of_int (- e))\<^sup>2))" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_real_def eucl_down_def abs_round_down_le) finally show ?thesis by (simp add: real_sqrt_mult) qed lemma eucl_down: "eucl_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] round_down eucl_down_def) lemma eucl_truncate_down: "eucl_truncate_down e (x::'a::executable_euclidean_space) \ x" by (auto simp add: eucl_le[where 'a='a] truncate_down) lemma eucl_truncate_down_le: "x \ y \ eucl_truncate_down w x \ (y::'a::executable_euclidean_space)" using eucl_truncate_down by (rule order.trans) lemma eucl_truncate_up_Basis[simp]: "i \ Basis \ eucl_truncate_up e x \ i = truncate_up e (x \ i)" by (simp add: eucl_truncate_up_def truncate_up_def) lemma eucl_truncate_up: "x \ eucl_truncate_up e (x::'a::executable_euclidean_space)" by (auto simp add: eucl_le[where 'a='a] round_up truncate_up_def) lemma eucl_truncate_up_le: "x \ y \ x \ eucl_truncate_up e (y::'a::executable_euclidean_space)" using _ eucl_truncate_up by (rule order.trans) lemma eucl_truncate_down_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_down p x \ eucl_truncate_down p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_down_mono) lemma eucl_truncate_up_mono: fixes x::"'a::executable_euclidean_space" shows "x \ y \ eucl_truncate_up p x \ eucl_truncate_up p y" by (auto simp: eucl_le[where 'a='a] intro!: truncate_up_mono) lemma infnorm[code]: fixes x::"'a::executable_euclidean_space" shows "infnorm x = fold max (map (\i. abs (x \ i)) Basis_list) 0" by (auto simp: Max.set_eq_fold[symmetric] infnorm_Max[symmetric] infnorm_pos_le intro!: max.absorb2[symmetric]) declare Inf_real_def[code del] declare Sup_real_def[code del] declare Inf_prod_def[code del] declare Sup_prod_def[code del] declare [[code abort: "Inf::real set \ real"]] declare [[code abort: "Sup::real set \ real"]] declare [[code abort: "Inf::('a::Inf * 'b::Inf) set \ 'a * 'b"]] declare [[code abort: "Sup::('a::Sup * 'b::Sup) set \ 'a * 'b"]] lemma nth_Basis_list_in_Basis[simp]: "n < length (Basis_list::'a::executable_euclidean_space list) \ Basis_list ! n \ (Basis::'a set)" by (metis Basis_list nth_mem) subsection \Representation as list\ lemma nth_eq_iff_index: "distinct xs \ n < length xs \ xs ! n = i \ n = index xs i" using index_nth_id by fastforce lemma in_Basis_index_Basis_list: "i \ Basis \ i = Basis_list ! index Basis_list i" by simp lemmas [simp] = length_Basis_list lemma sum_Basis_sum_nth_Basis_list: "(\i\Basis. f i) = (\i(x, i)\zip xs Basis_list. x *\<^sub>R i)" lemma eucl_of_list_nth: assumes "length xs = DIM('a)" shows "eucl_of_list xs = (\iR ((Basis_list::'a list) ! i))" by (auto simp: eucl_of_list_def sum_list_sum_nth length_Basis_list assms atLeast0LessThan) lemma eucl_of_list_inner: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" by (simp add: eucl_of_list_nth[OF l] inner_sum_left assms inner_Basis nth_eq_iff_index sum.delta if_distrib cong: if_cong) lemma inner_eucl_of_list: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs = DIM('a)" shows "i \ eucl_of_list xs = xs ! (index Basis_list i)" using eucl_of_list_inner[OF assms] by (auto simp: inner_commute) definition "list_of_eucl x = map ((\) x) Basis_list" lemma index_Basis_list_nth[simp]: "i < DIM('a::executable_euclidean_space) \ index Basis_list ((Basis_list::'a list) ! i) = i" by (simp add: index_nth_id) lemma list_of_eucl_eucl_of_list[simp]: "length xs = DIM('a::executable_euclidean_space) \ list_of_eucl (eucl_of_list xs::'a) = xs" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: nth_equalityI) lemma eucl_of_list_list_of_eucl[simp]: "eucl_of_list (list_of_eucl x) = x" by (auto simp: list_of_eucl_def eucl_of_list_inner intro!: euclidean_eqI[where 'a='a]) lemma length_list_of_eucl[simp]: "length (list_of_eucl (x::'a::executable_euclidean_space)) = DIM('a)" by (auto simp: list_of_eucl_def) lemma list_of_eucl_nth[simp]: "n < DIM('a::executable_euclidean_space) \ list_of_eucl x ! n = x \ (Basis_list ! n::'a)" by (auto simp: list_of_eucl_def) lemma nth_ge_len: "n \ length xs \ xs ! n = [] ! (n - length xs)" by (induction xs arbitrary: n) auto lemma list_of_eucl_nth_if: "list_of_eucl x ! n = (if n < DIM('a::executable_euclidean_space) then x \ (Basis_list ! n::'a) else [] ! (n - DIM('a)))" apply (auto simp: list_of_eucl_def ) apply (subst nth_ge_len) apply auto done lemma list_of_eucl_eq_iff: "list_of_eucl (x::'a::executable_euclidean_space) = list_of_eucl (y::'b::executable_euclidean_space) \ (DIM('a) = DIM('b) \ (\i < DIM('b). x \ Basis_list ! i = y \ Basis_list ! i))" by (auto simp: list_eq_iff_nth_eq) lemma eucl_le_Basis_list_iff: "(x::'a::executable_euclidean_space) \ y \ (\i Basis_list ! i \ y \ Basis_list ! i)" apply (auto simp: eucl_le[where 'a='a]) subgoal for i subgoal by (auto dest!: spec[where x="index Basis_list i"]) done done lemma eucl_of_list_inj: "length xs = DIM('a::executable_euclidean_space) \ length ys = DIM('a) \ (eucl_of_list xs::'a) = eucl_of_list (ys) \ xs = ys" apply (auto intro!: nth_equalityI simp: euclidean_eq_iff[where 'a="'a"] eucl_of_list_inner) using nth_Basis_list_in_Basis[where 'a="'a"] by fastforce lemma eucl_of_list_map_plus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x + g x) xs)::'a) = eucl_of_list (map f xs) + eucl_of_list (map g xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_uminus[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. - f x) xs)::'a) = - eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_left[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. r * f x) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_mult_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x * r) xs)::'a) = r *\<^sub>R eucl_of_list (map f xs)" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma eucl_of_list_map_divide_right[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. f x / r) xs)::'a) = eucl_of_list (map f xs) /\<^sub>R r" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner divide_simps) lemma eucl_of_list_map_const[simp]: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" shows "(eucl_of_list (map (\x. c) xs)::'a) = c *\<^sub>R One" by (auto intro!: euclidean_eqI[where 'a='a] simp: algebra_simps eucl_of_list_inner) lemma replicate_eq_list_of_eucl_zero: "replicate DIM('a::executable_euclidean_space) 0 = list_of_eucl (0::'a)" by (auto intro!: nth_equalityI) lemma eucl_of_list_append_zeroes[simp]: "eucl_of_list (xs @ replicate n 0) = eucl_of_list xs" unfolding eucl_of_list_def apply (auto simp: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) by (auto simp: nth_append) lemma Basis_prodD: assumes "(i, j) \ Basis" shows "i \ Basis \ j = 0 \ i = 0 \ j \ Basis" using assms by (auto simp: Basis_prod_def) lemma eucl_of_list_take_DIM[simp]: assumes "d = DIM('b::executable_euclidean_space)" shows "(eucl_of_list (take d xs)::'b) = (eucl_of_list xs)" by (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list sum_list_sum_nth assms dest!: Basis_prodD) lemma eucl_of_list_eqI: assumes "take DIM('a) (xs @ replicate (DIM('a) - length xs) 0) = take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)" shows "eucl_of_list xs = (eucl_of_list ys::'a::executable_euclidean_space)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ replicate (DIM('a) - length xs) 0))" by simp also note assms also have "eucl_of_list (take DIM('a) (ys @ replicate (DIM('a) - length ys) 0)) = (eucl_of_list ys::'a)" by simp finally show ?thesis . qed lemma eucl_of_list_replicate_zero[simp]: "eucl_of_list (replicate E 0) = 0" proof - have "eucl_of_list (replicate E 0) = (eucl_of_list (replicate E 0 @ replicate (DIM('a) - E) 0)::'a)" by simp also have "\ = eucl_of_list (replicate DIM('a) 0)" apply (rule eucl_of_list_eqI) by (auto simp: min_def nth_append intro!: nth_equalityI) also have "\ = 0" by (simp add: replicate_eq_list_of_eucl_zero) finally show ?thesis by simp qed lemma eucl_of_list_Nil[simp]: "eucl_of_list [] = 0" using eucl_of_list_replicate_zero[of 0] by simp lemma fst_eucl_of_list_prod: shows "fst (eucl_of_list xs::'b::executable_euclidean_space \ _) = (eucl_of_list (take DIM('b) xs)::'b)" apply (auto simp: eucl_of_list_inner eucl_of_list_def fst_sum_list dest!: Basis_prodD) apply (simp add: sum_list_sum_nth) apply (rule sum.mono_neutral_cong_right) subgoal by simp subgoal by auto subgoal by (auto simp: Basis_list_prod_def nth_append) subgoal by (auto simp: Basis_list_prod_def nth_append) done lemma index_zip_replicate1[simp]: "index (zip (replicate d a) bs) (a, b) = index bs b" if "d = length bs" using that by (induction bs arbitrary: d) auto lemma index_zip_replicate2[simp]: "index (zip as (replicate d b)) (a, b) = index as a" if "d = length as" using that by (induction as arbitrary: d) auto lemma index_Basis_list_prod[simp]: fixes a::"'a::executable_euclidean_space" and b::"'b::executable_euclidean_space" shows "a \ Basis \ index Basis_list (a, 0::'b) = index Basis_list a" "b \ Basis \ index Basis_list (0::'a, b) = DIM('a) + index Basis_list b" by (auto simp: Basis_list_prod_def index_append in_set_zip zip_replicate index_map_inj dest: spec[where x="index Basis_list a"]) lemma eucl_of_list_eq_takeI: assumes "(eucl_of_list (take DIM('a::executable_euclidean_space) xs)::'a) = x" shows "eucl_of_list xs = x" using eucl_of_list_take_DIM[OF refl, of xs, where 'b='a] assms by auto lemma eucl_of_list_inner_le: fixes i::"'a::executable_euclidean_space" assumes i: "i \ Basis" assumes l: "length xs \ DIM('a)" shows "eucl_of_list xs \ i = xs ! (index Basis_list i)" proof - have "(eucl_of_list xs::'a) = eucl_of_list (take DIM('a) (xs @ (replicate (DIM('a) - length xs) 0)))" by (rule eucl_of_list_eq_takeI) simp also have "\ \ i = xs ! (index Basis_list i)" using assms by (subst eucl_of_list_inner) auto finally show ?thesis . qed lemma eucl_of_list_prod_if: assumes "length xs = DIM('a::executable_euclidean_space) + DIM('b::executable_euclidean_space)" shows "eucl_of_list xs = (eucl_of_list (take DIM('a) xs)::'a, eucl_of_list (drop DIM('a) xs)::'b)" apply (rule euclidean_eqI) using assms apply (auto simp: eucl_of_list_inner dest!: Basis_prodD) apply (subst eucl_of_list_inner_le) apply (auto simp: Basis_list_prod_def index_append in_set_zip) done lemma snd_eucl_of_list_prod: shows "snd (eucl_of_list xs::'b::executable_euclidean_space \ 'c::executable_euclidean_space) = (eucl_of_list (drop DIM('b) xs)::'c)" proof (cases "length xs \ DIM('b)") case True then show ?thesis by (auto simp: eucl_of_list_inner eucl_of_list_def snd_sum_list dest!: Basis_prodD) (simp add: sum_list_sum_nth Basis_list_prod_def nth_append) next case False have "xs = take DIM('b) xs @ drop DIM('b) xs" by simp also have "eucl_of_list \ = (eucl_of_list (\ @ replicate (length xs - DIM('c)) 0)::'b \ 'c)" by simp finally have "eucl_of_list xs = (eucl_of_list (xs @ replicate (DIM('b) + DIM('c) - length xs) 0)::'b \ 'c)" by simp also have "\ = eucl_of_list (take (DIM ('b \ 'c)) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by simp finally have *: "(eucl_of_list xs::'b\'c) = eucl_of_list (take DIM('b \ 'c) (xs @ replicate (DIM('b) + DIM('c) - length xs) 0))" by simp show ?thesis apply (subst *) apply (subst eucl_of_list_prod_if) subgoal by simp subgoal apply simp apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (subst (2) eucl_of_list_take_DIM[OF refl, symmetric]) apply (rule arg_cong[where f=eucl_of_list]) by (auto intro!: nth_equalityI simp: nth_append min_def split: if_splits) done qed lemma eucl_of_list_prod: shows "eucl_of_list xs = (eucl_of_list (take DIM('b) xs)::'b::executable_euclidean_space, eucl_of_list (drop DIM('b) xs)::'c::executable_euclidean_space)" using snd_eucl_of_list_prod[of xs, where 'b='b and 'c='c] using fst_eucl_of_list_prod[of xs, where 'b='b and 'a='c] by (auto simp del: snd_eucl_of_list_prod fst_eucl_of_list_prod simp add: prod_eq_iff) lemma eucl_of_list_real[simp]: "eucl_of_list [x] = (x::real)" by (auto simp: eucl_of_list_def Basis_list_real_def) lemma eucl_of_list_append[simp]: assumes "length xs = DIM('i::executable_euclidean_space)" assumes "length ys = DIM('j::executable_euclidean_space)" shows "eucl_of_list (xs @ ys) = (eucl_of_list xs::'i, eucl_of_list ys::'j)" using assms by (auto simp: eucl_of_list_prod) lemma list_allI: "(\x. x \ set xs \ P x) \ list_all P xs" by (auto simp: list_all_iff) lemma concat_map_nthI: assumes "\x y. x \ set xs \ y \ set (f x) \ P y" assumes "j < length (concat (map f xs))" shows "P (concat (map f xs) ! j)" proof - have "list_all P (concat (map f xs))" by (rule list_allI) (auto simp: assms) then show ?thesis by (auto simp: list_all_length assms) qed lemma map_nth_append1: assumes "length xs = d" shows "map ((!) (xs @ ys)) [0.. y = z" by (metis eucl_of_list_list_of_eucl) lemma length_Basis_list_pos[simp]: "length Basis_list > 0" by (metis length_pos_if_in_set Basis_list SOME_Basis) lemma Basis_list_nth_nonzero: "i < length (Basis_list::'a::executable_euclidean_space list) \ (Basis_list::'a list) ! i \ 0" by (auto dest!: nth_mem simp: nonzero_Basis) lemma nth_Basis_list_prod: "i < DIM('a) + DIM('b) \ (Basis_list::('a::executable_euclidean_space \ 'b::executable_euclidean_space) list) ! i = (if i < DIM('a) then (Basis_list ! i, 0) else (0, Basis_list ! (i - DIM('a))))" by (auto simp: Basis_list_nth_nonzero prod_eq_iff Basis_list_prod_def nth_append not_less) lemma eucl_of_list_if: assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" "distinct xs" shows "eucl_of_list (map (\xa. if xa = x then 1 else 0) (xs::nat list)) = (if x \ set xs then Basis_list ! index xs x else 0::'a)" by (rule euclidean_eqI) (auto simp: eucl_of_list_inner inner_Basis index_nth_id) lemma take_append_take_minus_idem: "take n XS @ map ((!) XS) [n..b\Basis. f b)" by (subst sum_list_distinct_conv_sum_set) (auto simp: Basis_list distinct_Basis_list) lemma hd_Basis_list[simp]: "hd Basis_list \ Basis" unfolding Basis_list[symmetric] by (rule hd_in_set) (auto simp: set_empty[symmetric]) definition "inner_lv_rel a b = sum_list (map2 (*) a b)" lemma eucl_of_list_inner_eq: "(eucl_of_list xs::'a) \ eucl_of_list ys = inner_lv_rel xs ys" if "length xs = DIM('a::executable_euclidean_space)" "length ys = DIM('a)" using that by (subst euclidean_inner[abs_def], subst sum_list_Basis_list[symmetric]) (auto simp: eucl_of_list_inner sum_list_sum_nth index_nth_id inner_lv_rel_def) lemma euclidean_vec_componentwise: "(\(xa::'a::euclidean_space^'b::finite)\Basis. f xa) = (\a\Basis. (\b::'b\UNIV. f (axis b a)))" apply (auto simp: Basis_vec_def) apply (subst sum.swap) apply (subst sum.Union_disjoint) apply auto apply (simp add: axis_eq_axis nonzero_Basis) apply (simp add: axis_eq_axis nonzero_Basis) apply (subst sum.reindex) apply (auto intro!: injI) subgoal apply (auto simp: set_eq_iff) by (metis (full_types) all_not_in_conv inner_axis_axis inner_eq_zero_iff nonempty_Basis nonzero_Basis) apply (rule sum.cong[OF refl]) apply (auto ) apply (rule sum.reindex_cong[OF _ _ refl]) apply (auto intro!: inj_onI) using axis_eq_axis by blast lemma vec_nth_inner_scaleR_craziness: "f (x $ i \ j) *\<^sub>R j = (\xa\UNIV. f (x $ xa \ j) *\<^sub>R axis xa j) $ i" by vector (auto simp: axis_def if_distrib scaleR_vec_def sum.delta' cong: if_cong) instantiation vec :: ("{executable_euclidean_space}", enum) executable_euclidean_space begin definition Basis_list_vec :: "('a, 'b) vec list" where "Basis_list_vec = concat (map (\n. map (axis n) Basis_list) enum_class.enum)" definition eucl_down_vec :: "int \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_down_vec p x = (\ i. eucl_down p (x $ i))" definition eucl_truncate_down_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_down_vec p x = (\ i. eucl_truncate_down p (x $ i))" definition eucl_truncate_up_vec :: "nat \ ('a, 'b) vec \ ('a, 'b) vec" where "eucl_truncate_up_vec p x = (\ i. eucl_truncate_up p (x $ i))" instance proof show *: "set (Basis_list::('a, 'b) vec list) = Basis" unfolding Basis_list_vec_def Basis_vec_def apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map Basis_vec_def intro!: distinct_concat inj_onI split: if_splits) apply (auto simp: Basis_list_vec_def vec_eq_iff distinct_map enum_distinct UNIV_enum[symmetric] intro!: distinct_concat inj_onI split: if_splits) done have "length (Basis_list::('a, 'b) vec list) = CARD('b) * DIM('a)" by (auto simp: Basis_list_vec_def length_concat o_def enum_distinct sum_list_distinct_conv_sum_set UNIV_enum[symmetric]) then show "distinct (Basis_list::('a, 'b) vec list)" using * by (auto intro!: card_distinct) qed (simp_all only: vector_cart[symmetric] vec_eq_iff eucl_down_vec_def eucl_down_def eucl_truncate_down_vec_def eucl_truncate_down_def eucl_truncate_up_vec_def eucl_truncate_up_def, auto simp: euclidean_vec_componentwise inner_axis Basis_list_vec_def vec_nth_inner_scaleR_craziness intro!: sum.cong[OF refl]) end lemma concat_same_lengths_nth: assumes "\xs. xs \ set XS \ length xs = N" assumes "i < length XS * N" "N > 0" shows "concat XS ! i = XS ! (i div N) ! (i mod N)" - using assms - apply (induction XS arbitrary: i) - apply (auto simp: nth_append nth_Cons split: nat.splits) - apply (simp add: div_eq_0_iff) - by (metis Suc_inject div_geq mod_geq) +using assms by (induction XS arbitrary: i) + (auto simp: nth_append nth_Cons div_eq_0_iff le_div_geq le_mod_geq split: nat.splits) lemma concat_map_map_index: shows "concat (map (\n. map (f n) xs) ys) = map (\i. f (ys ! (i div length xs)) (xs ! (i mod length xs))) [0..(x, y)\zip xs (map g xs). f x y) = (\x\set xs. f x (g x))" by (force simp add: sum_list_distinct_conv_sum_set assms distinct_zipI1 split_beta' in_set_zip in_set_conv_nth inj_on_convol_ident intro!: sum.reindex_cong[where l="\x. (x, g x)"]) lemma sum_list_zip_map_of: assumes "distinct bs" assumes "length xs = length bs" shows "(\(x, y)\zip xs bs. f x y) = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" proof - have "(\(x, y)\zip xs bs. f x y) = (\(y, x)\zip bs xs. f x y)" by (subst zip_commute) (auto simp: o_def split_beta') also have "\ = (\(x, y)\zip bs (map (the o map_of (zip bs xs)) bs). f y x)" proof (rule arg_cong, rule map_cong) have "xs = (map (the \ map_of (zip bs xs)) bs)" using assms by (auto intro!: nth_equalityI simp: map_nth map_of_zip_nth) then show "zip bs xs = zip bs (map (the \ map_of (zip bs xs)) bs)" by simp qed auto also have "\ = (\x\set bs. f (the (map_of (zip bs xs) x)) x)" using assms(1) by (subst sum_list_zip_map) (auto simp: o_def) finally show ?thesis . qed lemma vec_nth_matrix: "vec_nth (vec_nth (matrix y) i) j = vec_nth (y (axis j 1)) i" unfolding matrix_def by simp lemma matrix_eqI: assumes "\x. x \ Basis \ A *v x = B *v x" shows "(A::real^'n^'n) = B" apply vector using assms apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis matrix_vector_mul_component) lemma matrix_columnI: assumes "\i. column i A = column i B" shows "(A::real^'n^'n) = B" using assms apply vector apply (auto simp: column_def) apply vector by (metis iso_tuple_UNIV_I vec_lambda_inject) lemma vec_nth_Basis: fixes x::"real^'n" shows "x \ Basis \ vec_nth x i = (if x = axis i 1 then 1 else 0)" apply (auto simp: Basis_vec_def) by (metis cart_eq_inner_axis inner_axis_axis) lemma vec_nth_eucl_of_list_eq: "length M = CARD('n) \ vec_nth (eucl_of_list M::real^'n::enum) i = M ! index Basis_list (axis i (1::real))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i 1::real^'n::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) simp qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n by (metis DIM_cart DIM_real index_Basis_list_nth mult.right_neutral prems(2) prems(3)) done done done lemma index_Basis_list_axis1: "index Basis_list (axis i (1::real)) = index enum_class.enum i" apply (auto simp: Basis_list_vec_def Basis_list_real_def ) apply (subst index_map_inj) by (auto intro!: injI simp: axis_eq_axis) lemma vec_nth_eq_list_of_eucl1: "(vec_nth (M::real^'n::enum) i) = list_of_eucl M ! (index enum_class.enum i)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq) unfolding index_Basis_list_axis1 by auto lemma enum_3[simp]: "(enum_class.enum::3 list) = [0, 1, 2]" by code_simp+ lemma three_eq_zero: "(3::3) = 0" by simp lemma forall_3': "(\i::3. P i) \ P 0 \ P 1 \ P 2" using forall_3 three_eq_zero by auto lemma euclidean_eq_list_of_euclI: "x = y" if "list_of_eucl x = list_of_eucl y" using that by (metis eucl_of_list_list_of_eucl) lemma axis_one_neq_zero[simp]: "axis xa (1::'a::zero_neq_one) \ 0" by (auto simp: axis_def vec_eq_iff) lemma eucl_of_list_vec_nth3[simp]: "(eucl_of_list [g, h, i]::real^3) $ 0 = g" "(eucl_of_list [g, h, i]::real^3) $ 1 = h" "(eucl_of_list [g, h, i]::real^3) $ 2 = i" "(eucl_of_list [g, h, i]::real^3) $ 3 = g" by (auto simp: cart_eq_inner_axis eucl_of_list_inner vec_nth_eq_list_of_eucl1 index_Basis_list_axis1) type_synonym R3 = "real*real*real" lemma Basis_list_R3: "Basis_list = [(1,0,0), (0, 1, 0), (0, 0, 1)::R3]" by (auto simp: Basis_list_prod_def Basis_list_real_def zero_prod_def) lemma Basis_list_vec3: "Basis_list = [axis 0 1::real^3, axis 1 1, axis 2 1]" by (auto simp: Basis_list_vec_def Basis_list_real_def) lemma eucl_of_list3[simp]: "eucl_of_list [a, b, c] = (a, b, c)" by (auto simp: eucl_of_list_inner Basis_list_vec_def zero_prod_def Basis_prod_def Basis_list_vec3 Basis_list_R3 intro!: euclidean_eqI[where 'a=R3]) subsection \Bounded Linear Functions\ subsection \bounded linear functions\ locale blinfun_syntax begin no_notation vec_nth (infixl "$" 90) notation blinfun_apply (infixl "$" 999) end lemma bounded_linear_via_derivative: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" \ \TODO: generalize?\ assumes "\i. ((\x. blinfun_apply (f x) i) has_derivative (\x. f' y x i)) (at y)" shows "bounded_linear (f' y x)" proof - interpret linear "f' y x" proof (unfold_locales, goal_cases) case (1 v w) from has_derivative_unique[OF assms[of "v + w", unfolded blinfun.bilinear_simps] has_derivative_add[OF assms[of v] assms[of w]], THEN fun_cong, of x] show ?case . next case (2 r v) from has_derivative_unique[OF assms[of "r *\<^sub>R v", unfolded blinfun.bilinear_simps] has_derivative_scaleR_right[OF assms[of v], of r], THEN fun_cong, of x] show ?case . qed let ?bnd = "\i\Basis. norm (f' y x i)" { fix v have "f' y x v = (\i\Basis. (v \ i) *\<^sub>R f' y x i)" by (subst euclidean_representation[symmetric]) (simp add: sum scaleR) also have "norm \ \ norm v * ?bnd" by (auto intro!: order.trans[OF norm_sum] sum_mono mult_right_mono simp: sum_distrib_left Basis_le_norm) finally have "norm (f' y x v) \ norm v * ?bnd" . } then show ?thesis by unfold_locales auto qed definition blinfun_scaleR::"('a::real_normed_vector \\<^sub>L real) \ 'b::real_normed_vector \ ('a \\<^sub>L 'b)" where "blinfun_scaleR a b = blinfun_scaleR_left b o\<^sub>L a" lemma blinfun_scaleR_transfer[transfer_rule]: "rel_fun (pcr_blinfun (=) (=)) (rel_fun (=) (pcr_blinfun (=) (=))) (\a b c. a c *\<^sub>R b) blinfun_scaleR" by (auto simp: blinfun_scaleR_def rel_fun_def pcr_blinfun_def cr_blinfun_def OO_def) lemma blinfun_scaleR_rep_eq[simp]: "blinfun_scaleR a b c = a c *\<^sub>R b" by (simp add: blinfun_scaleR_def) lemma bounded_linear_blinfun_scaleR: "bounded_linear (blinfun_scaleR a)" unfolding blinfun_scaleR_def[abs_def] by (auto intro!: bounded_linear_intros) lemma blinfun_scaleR_has_derivative[derivative_intros]: assumes "(f has_derivative f') (at x within s)" shows "((\x. blinfun_scaleR a (f x)) has_derivative (\x. blinfun_scaleR a (f' x))) (at x within s)" using bounded_linear_blinfun_scaleR assms by (rule bounded_linear.has_derivative) lemma blinfun_componentwise: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" shows "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (auto intro!: blinfun_eqI simp: blinfun.sum_left euclidean_representation blinfun.scaleR_right[symmetric] blinfun.sum_right[symmetric]) lemma blinfun_has_derivative_componentwiseI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. i \ Basis \ ((\x. f x i) has_derivative blinfun_apply (f' i)) (at x)" shows "(f has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' i x))) (at x)" by (subst blinfun_componentwise) (force intro: derivative_eq_intros assms simp: blinfun.bilinear_simps) lemma has_derivative_BlinfunI: fixes f::"'a::real_normed_vector \ 'b::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. ((\x. f x i) has_derivative (\x. f' y x i)) (at y)" shows "(f has_derivative (\x. Blinfun (f' y x))) (at y)" proof - have 1: "f = (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f x i))" by (rule blinfun_componentwise) moreover have 2: "(\ has_derivative (\x. \i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i))) (at y)" by (force intro: assms derivative_eq_intros) moreover interpret f': bounded_linear "f' y x" for x by (rule bounded_linear_via_derivative) (rule assms) have 3: "(\i\Basis. blinfun_scaleR (blinfun_inner_left i) (f' y x i)) i = f' y x i" for x i by (auto simp: if_distrib if_distribR blinfun.bilinear_simps f'.scaleR[symmetric] f'.sum[symmetric] euclidean_representation intro!: blinfun_euclidean_eqI) have 4: "blinfun_apply (Blinfun (f' y x)) = f' y x" for x apply (subst bounded_linear_Blinfun_apply) subgoal by unfold_locales subgoal by simp done show ?thesis apply (subst 1) apply (rule 2[THEN has_derivative_eq_rhs]) apply (rule ext) apply (rule blinfun_eqI) apply (subst 3) apply (subst 4) apply (rule refl) done qed lemma has_derivative_Blinfun: assumes "(f has_derivative f') F" shows "(f has_derivative Blinfun f') F" using assms by (subst bounded_linear_Blinfun_apply) auto lift_definition flip_blinfun:: "('a::real_normed_vector \\<^sub>L 'b::real_normed_vector \\<^sub>L 'c::real_normed_vector) \ 'b \\<^sub>L 'a \\<^sub>L 'c" is "\f x y. f y x" using bounded_bilinear.bounded_linear_left bounded_bilinear.bounded_linear_right bounded_bilinear.flip by auto lemma flip_blinfun_apply[simp]: "flip_blinfun f a b = f b a" by transfer simp lemma le_norm_blinfun: shows "norm (blinfun_apply f x) / norm x \ norm f" by transfer (rule le_onorm) lemma norm_flip_blinfun[simp]: "norm (flip_blinfun x) = norm x" (is "?l = ?r") proof (rule antisym) from order_trans[OF norm_blinfun, OF mult_right_mono, OF norm_blinfun, OF norm_ge_zero, of x] show "?l \ ?r" by (auto intro!: norm_blinfun_bound simp: ac_simps) have "norm (x a b) \ norm (flip_blinfun x) * norm a * norm b" for a b proof - have "norm (x a b) / norm a \ norm (flip_blinfun x b)" by (rule order_trans[OF _ le_norm_blinfun]) auto also have "\ \ norm (flip_blinfun x) * norm b" by (rule norm_blinfun) finally show ?thesis by (auto simp add: divide_simps blinfun.bilinear_simps algebra_simps split: if_split_asm) qed then show "?r \ ?l" by (auto intro!: norm_blinfun_bound) qed lemma bounded_linear_flip_blinfun[bounded_linear]: "bounded_linear flip_blinfun" by unfold_locales (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI exI[where x=1]) lemma dist_swap2_swap2[simp]: "dist (flip_blinfun f) (flip_blinfun g) = dist f g" by (metis (no_types) bounded_linear_flip_blinfun dist_blinfun_def linear_simps(2) norm_flip_blinfun) context includes blinfun.lifting begin lift_definition blinfun_of_vmatrix::"(real^'m^'n) \ ((real^('m::finite)) \\<^sub>L (real^('n::finite)))" is "matrix_vector_mult:: ((real, 'm) vec, 'n) vec \ ((real, 'm) vec \ (real, 'n) vec)" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma matrix_blinfun_of_vmatrix[simp]: "matrix (blinfun_of_vmatrix M) = M" apply vector apply (auto simp: matrix_def) apply transfer by (metis cart_eq_inner_axis matrix_vector_mul_component) end lemma blinfun_apply_componentwise: "B = (\i\Basis. blinfun_scaleR (blinfun_inner_left i) (blinfun_apply B i))" using blinfun_componentwise[of "\x. B", unfolded fun_eq_iff] by blast lemma blinfun_apply_eq_sum: assumes [simp]: "length v = CARD('n)" shows "blinfun_apply (B::(real^'n::enum)\\<^sub>L(real^'m::enum)) (eucl_of_list v) = (\ij Basis_list ! i) * v ! j) *\<^sub>R (Basis_list ! i))" apply (subst blinfun_apply_componentwise[of B]) apply (auto intro!: euclidean_eqI[where 'a="(real,'m) vec"] simp: blinfun.bilinear_simps eucl_of_list_inner inner_sum_left inner_Basis if_distrib sum_Basis_sum_nth_Basis_list nth_eq_iff_index if_distribR cong: if_cong) apply (subst sum.swap) by (auto simp: sum.delta algebra_simps) lemma in_square_lemma[intro, simp]: "x * C + y < D * C" if "x < D" "y < C" for x::nat proof - have "x * C + y < (D - 1) * C + C" apply (rule add_le_less_mono) apply (rule mult_right_mono) using that by auto also have "\ \ D * C" using that by (auto simp: algebra_simps) finally show ?thesis . qed lemma less_square_imp_div_less[intro, simp]: "i < E * D \ i div E < D" for i::nat by (metis div_eq_0_iff div_mult2_eq gr_implies_not0 mult_not_zero) lemma in_square_lemma'[intro, simp]: "i < L \ n < N \ i * N + n < N * L" for i n::nat by (metis in_square_lemma mult.commute) lemma distinct_nth_eq_iff: "distinct xs \ x < length xs \ y < length xs \ xs ! x = xs ! y \ x = y" by (drule inj_on_nth[where I="{..i. axis (enum_class.enum ! (i div CARD('i))) (axis (enum_class.enum ! (i mod CARD('i))) 1)) ((index enum_class.enum j) * CARD('i) + index enum_class.enum i)" by (auto simp: index_less_cardi enum_UNIV) note less=in_square_lemma[OF index_less_cardj index_less_cardi, of j i] show ?thesis apply (subst *) apply (subst index_map_inj_on[where S="{.. Basis \ vec_nth (vec_nth x i) j = ((if x = axis i (axis j 1) then 1 else 0))" by (auto simp: Basis_vec_def axis_def) lemma vec_nth_eucl_of_list_eq2: "length M = CARD('n) * CARD('m) \ vec_nth (vec_nth (eucl_of_list M::real^'n::enum^'m::enum) i) j = M ! index Basis_list (axis i (axis j (1::real)))" apply (auto simp: eucl_of_list_def) apply (subst sum_list_zip_map_of) apply (auto intro!: distinct_zipI2 simp: split_beta') apply (subst sum.cong[OF refl]) apply (subst vec_nth_Basis2) apply (force simp: set_zip) apply (rule refl) apply (auto simp: if_distrib sum.delta cong: if_cong) subgoal apply (cases "map_of (zip Basis_list M) (axis i (axis j 1)::real^'n::enum^'m::enum)") subgoal premises prems proof - have "fst ` set (zip Basis_list M) = (Basis::(real^'n::enum^'m::enum) set)" using prems by (auto simp: in_set_zip) then show ?thesis using prems by (subst (asm) map_of_eq_None_iff) auto qed subgoal for a apply (auto simp: in_set_zip) subgoal premises prems for n proof - have "n < card (Basis::(real^'n::_^'m::_) set)" by (simp add: prems(4)) then show ?thesis by (metis index_Basis_list_nth prems(2)) qed done done done lemma vec_nth_eq_list_of_eucl2: "vec_nth (vec_nth (M::real^'n::enum^'m::enum) i) j = list_of_eucl M ! (index enum_class.enum i * CARD('n) + index enum_class.enum j)" apply (subst eucl_of_list_list_of_eucl[of M, symmetric]) apply (subst vec_nth_eucl_of_list_eq2) unfolding index_Basis_list_axis2 by auto theorem eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list: assumes "length M = CARD('n) * CARD('m)" assumes "length v = CARD('n)" shows "(eucl_of_list M::real^'n::enum^'m::enum) *v eucl_of_list v = (\ijR Basis_list ! i)" apply (vector matrix_vector_mult_def) apply auto apply (subst vec_nth_eucl_of_list_eq2) apply (auto simp: assms) apply (subst vec_nth_eucl_of_list_eq) apply (auto simp: assms index_Basis_list_axis2 index_Basis_list_axis1 vec_nth_Basis sum.delta nth_eq_iff_index if_distrib cong: if_cong) subgoal for i apply (rule sum.reindex_cong[where l="nth enum_class.enum"]) apply (auto simp: enum_distinct card_UNIV_length_enum distinct_nth_eq_iff intro!: inj_onI) apply (rule image_eqI[OF ]) apply (rule nth_index[symmetric]) apply (auto simp: enum_UNIV) by (auto simp: algebra_simps enum_UNIV enum_distinct index_nth_id) subgoal for i using index_less[of i "enum_class.enum" "CARD('n)"] by (auto simp: enum_UNIV card_UNIV_length_enum) done lemma index_enum_less[intro, simp]: "index enum_class.enum (i::'n::enum) < CARD('n)" by (auto intro!: index_less simp: enum_UNIV card_UNIV_length_enum) lemmas [intro, simp] = enum_distinct lemmas [simp] = card_UNIV_length_enum[symmetric] enum_UNIV lemma sum_index_enum_eq: "(\(k::'n::enum)\UNIV. f (index enum_class.enum k)) = (\i rat list list \ bool" where "well_def_signs num_polys sign_conds \ \ signs \ set(sign_conds). (length signs = num_polys)" definition satisfies_properties:: "real poly \ real poly list \ nat list list \ rat list list \ rat mat \ bool" where "satisfies_properties p qs subsets signs matrix = ( all_list_constr subsets (length qs) \ well_def_signs (length qs) signs \ distinct signs \ satisfy_equation p qs subsets signs \ invertible_mat matrix \ matrix = matrix_A signs subsets \ set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs) )" section "Base Case" lemma mat_base_case: shows "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1, 1], [1, -1]])" unfolding matrix_A_def mtx_row_def z_def apply (auto) by (simp add: numeral_2_eq_2) lemma base_case_sgas: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto lemma base_case_sgas_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ {[1], [- 1]}" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_sgas nonzero rel_prime apply (auto) using characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def by auto qed lemma base_case_satisfy_equation: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "satisfy_equation p [q] [[],[0]] [[1],[-1]]" proof - have h1: "set (characterize_consistent_signs_at_roots_copr p [q]) \ {[1], [- 1]}" using base_case_sgas assms by auto have h2: " \qa. List.member [q] qa \ coprime p qa" using rel_prime by (simp add: member_rec(1) member_rec(2)) have h3: "all_list_constr [[], [0]] (Suc 0)" unfolding all_list_constr_def by (simp add: list_constr_def member_def) then show ?thesis using matrix_equation[where p = "p", where qs = "[q]", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] nonzero h1 h2 h3 by auto qed lemma base_case_satisfy_equation_alt: fixes p:: "real poly" fixes qs:: "real poly list" assumes len1: "length qs = 1" assumes nonzero: "p \ 0" assumes rel_prime: "\q. (List.member qs q) \ coprime p q" shows "satisfy_equation p qs [[],[0]] [[1],[-1]]" proof - have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then show ?thesis using base_case_satisfy_equation nonzero rel_prime apply (auto) by (simp add: member_rec(1)) qed lemma base_case_matrix_eq: fixes q p:: "real poly" assumes nonzero: "p \ 0" assumes rel_prime: "coprime p q" shows "(mult_mat_vec (mat_of_rows_list 2 [[1, 1], [1, -1]]) (construct_lhs_vector p [q] [[1],[-1]]) = (construct_rhs_vector p [q] [[],[0]]))" using mat_base_case base_case_satisfy_equation unfolding satisfy_equation_def by (simp add: nonzero rel_prime) lemma less_two: shows "j < Suc (Suc 0) \ j = 0 \ j = 1" by auto lemma inverse_mat_base_case: shows "inverts_mat (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat)" unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma inverse_mat_base_case_2: shows "inverts_mat (mat_of_rows_list 2 [[1, 1], [1, -1]]::rat mat) (mat_of_rows_list 2 [[1/2, 1/2], [1/2, -(1/2)]]::rat mat) " unfolding inverts_mat_def mat_of_rows_list_def mat_eq_iff apply auto unfolding less_two by (auto simp add: scalar_prod_def) lemma base_case_invertible_mat: shows "invertible_mat (matrix_A [[1], [- 1]] [[], [0]])" unfolding invertible_mat_def using inverse_mat_base_case inverse_mat_base_case_2 apply (auto) apply (simp add: mat_base_case mat_of_rows_list_def) using mat_base_case by auto section "Inductive Step" subsection "Lemmas on smashing subsets and signs" lemma signs_smash_property: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ n m :: nat. elem = ((nth signs1 n)@(nth signs2 m))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\n m. elem = signs1 ! n @ signs2 ! m" using assum unfolding signs_smash_def apply (auto) by (metis in_set_conv_nth) qed lemma signs_smash_property_set: fixes signs1 signs2 :: "'a list list" fixes a b:: "nat" shows "\ (elem :: 'a list). (elem \ (set (signs_smash signs1 signs2)) \ (\ (elem1::'a list). \ (elem2::'a list). (elem1 \ set(signs1) \ elem2 \ set(signs2) \ elem = (elem1@elem2))))" proof clarsimp fix elem assume assum: "elem \ set (signs_smash signs1 signs2)" show "\elem1. elem1 \ set signs1 \ (\elem2. elem2 \ set signs2 \ elem = elem1 @ elem2)" using assum unfolding signs_smash_def by auto qed lemma subsets_smash_property: fixes subsets1 subsets2 :: "nat list list" fixes n:: "nat" shows "\ (elem :: nat list). (List.member (subsets_smash n subsets1 subsets2) elem) \ (\ (elem1::nat list). \ (elem2::nat list). (elem1 \ set(subsets1) \ elem2 \ set(subsets2) \ elem = (elem1 @ ((map ((+) n) elem2)))))" proof - let ?new_subsets = "(map (map ((+) n)) subsets2)" (* a slightly unorthodox use of signs_smash, but it closes the proof *) have "subsets_smash n subsets1 subsets2 = signs_smash subsets1 ?new_subsets" unfolding subsets_smash_def signs_smash_def apply (auto) by (simp add: comp_def) then show ?thesis by (smt imageE in_set_member set_map signs_smash_property_set) qed (* If subsets for smaller systems are well-defined, then subsets for combined systems are well-defined *) subsection "Well-defined subsets preserved when smashing" lemma list_constr_append: "list_constr a n \ list_constr b n \ list_constr (a@b) n" apply (auto) unfolding list_constr_def using list_all_append by blast lemma well_def_step: fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" assumes well_def_subsets1: "all_list_constr (subsets1) (length qs1)" assumes well_def_subsets2: "all_list_constr (subsets2) (length qs2)" shows "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1 @ qs2))" proof - have h1: "\elem. List.member (subsets_smash (length qs1) subsets1 subsets2) elem \ (\elem1 elem2. elem1 \ set subsets1 \ elem2 \ set subsets2 \ elem = elem1 @ map ((+) (length qs1)) elem2)" using subsets_smash_property by blast have h2: "\elem1 elem2. (elem1 \ set subsets1 \ elem2 \ set subsets2) \ list_constr (elem1 @ map ((+) (length qs1)) elem2) (length (qs1 @ qs2))" proof clarsimp fix elem1 fix elem2 assume e1: "elem1 \ set subsets1" assume e2: "elem2 \ set subsets2" have h1: "list_constr elem1 (length qs1 + length qs2) " proof - have h1: "list_constr elem1 (length qs1)" using e1 well_def_subsets1 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list.pred_mono_strong) qed have h2: "list_constr (map ((+) (length qs1)) elem2) (length qs1 + length qs2)" proof - have h1: "list_constr elem2 (length qs2)" using e2 well_def_subsets2 unfolding all_list_constr_def by (simp add: in_set_member) then show ?thesis unfolding list_constr_def by (simp add: list_all_length) qed show "list_constr (elem1 @ map ((+) (length qs1)) elem2) (length qs1 + length qs2)" using h1 h2 list_constr_append by blast qed then show ?thesis using h1 unfolding all_list_constr_def by auto qed subsection "Well def signs preserved when smashing" lemma well_def_signs_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes "length qs1 > 0" assumes "length qs2 > 0" assumes well_def1: "well_def_signs (length qs1) signs1" assumes well_def2: "well_def_signs (length qs2) signs2" shows "well_def_signs (length (qs1@qs2)) (signs_smash signs1 signs2)" using well_def1 well_def2 unfolding well_def_signs_def using signs_smash_property_set[of signs1 signs2] length_append by auto subsection "Distinct signs preserved when smashing" lemma distinct_map_append: assumes "distinct ls" shows "distinct (map ((@) xs) ls)" unfolding distinct_map inj_on_def using assms by auto lemma length_eq_append: assumes "length y = length b" shows "(x @ y = a @ b) \ x=a \ y = b" by (simp add: assms) lemma distinct_step: fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes well_def1: "well_def_signs n1 signs1" assumes well_def2: "well_def_signs n2 signs2" assumes distinct1: "distinct signs1" assumes distinct2: "distinct signs2" shows "distinct (signs_smash signs1 signs2)" unfolding signs_smash_def using distinct1 proof(induction signs1) case Nil then show ?case by auto next case (Cons a signs1) then show ?case apply (auto simp add: distinct2 distinct_map_append) using assms unfolding well_def_signs_def by (simp add: distinct1 distinct2 length_eq_append) qed (* In this section we will show that if signs1 contains all consistent sign assignments and signs2 contains all consistent sign assignments, then their smash contains all consistent sign assignments. Intuitively, this makes sense because signs1 and signs2 are carrying information about different sets of polynomials, and their smash contains all possible combinations of that information. *) subsection "Consistent sign assignments preserved when smashing" lemma subset_smash_signs: fixes a1 b1 a2 b2:: "rat list list" assumes sub1: "set a1 \ set a2" assumes sub2: "set b1 \ set b2" shows "set (signs_smash a1 b1) \ set (signs_smash a2 b2)" proof - have h1: "\x\set (signs_smash a1 b1). x\set (signs_smash a2 b2)" proof clarsimp fix x assume x_in: "x \ set (signs_smash a1 b1)" have h1: "\ n m :: nat. x = (nth a1 n)@(nth b1 m)" using signs_smash_property[of a1 b1] x_in by blast then have h2: "\ p q::nat. x = (nth a2 p)@(nth b2 q)" using sub1 sub2 apply (auto) by (metis nth_find_first signs_smash_property_set subset_code(1) x_in) then show "x \ set (signs_smash a2 b2)" unfolding signs_smash_def apply (auto) using signs_smash_property_set sub1 sub2 x_in by fastforce qed then show ?thesis by blast qed lemma subset_helper: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" shows "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). x = x1@x2" proof clarsimp fix x assume x_in: "x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2))" have x_in_csv: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p))" using x_in unfolding characterize_consistent_signs_at_roots_copr_def by simp have csv_hyp: "\r. consistent_sign_vec_copr (qs1 @ qs2) r = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r)" unfolding consistent_sign_vec_copr_def by auto have exr_iff: "(\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r) \ (\r \ set (characterize_root_list_p p). x = (consistent_sign_vec_copr qs1 r)@(consistent_sign_vec_copr qs2 r))" using csv_hyp by auto have exr: "x \ set(map (consistent_sign_vec_copr (qs1 @ qs2)) (characterize_root_list_p p)) \ (\r \ set (characterize_root_list_p p). x = consistent_sign_vec_copr (qs1 @ qs2) r)" by auto have mem_list1: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs1 r) \ set(map (consistent_sign_vec_copr qs1) (characterize_root_list_p p))" by simp have mem_list2: "\ r \ set (characterize_root_list_p p). (consistent_sign_vec_copr qs2 r) \ set(map (consistent_sign_vec_copr qs2) (characterize_root_list_p p))" by simp then show "\x1\set (characterize_consistent_signs_at_roots_copr p qs1). \x2\set (characterize_consistent_signs_at_roots_copr p qs2). x = x1 @ x2" using x_in_csv exr mem_list1 mem_list2 characterize_consistent_signs_at_roots_copr_def exr_iff by auto qed lemma subset_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes signs1 signs2 :: "rat list list" assumes csa1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set (signs1)" assumes csa2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set (signs2)" shows "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (signs_smash signs1 signs2)" proof - have h0: "\ x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). \ x1 \ set (characterize_consistent_signs_at_roots_copr p qs1). \ x2 \ set (characterize_consistent_signs_at_roots_copr p qs2). (x = x1@x2)" using subset_helper[of p qs1 qs2] by blast then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash (characterize_consistent_signs_at_roots_copr p qs1) (characterize_consistent_signs_at_roots_copr p qs2))" unfolding signs_smash_def apply (auto) by fastforce then have "\x \ set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)). x \ set (signs_smash signs1 signs2)" using csa1 csa2 subset_smash_signs[of _ signs1 _ signs2] apply (auto) by blast thus ?thesis by blast qed subsection "Main Results" lemma dim_row_mat_of_rows_list[simp]: shows "dim_row (mat_of_rows_list nr ls) = length ls" unfolding mat_of_rows_list_def by auto lemma dim_col_mat_of_rows_list[simp]: shows "dim_col (mat_of_rows_list nr ls) = nr" unfolding mat_of_rows_list_def by auto lemma dim_row_matrix_A[simp]: shows "dim_row (matrix_A signs subsets) = length subsets" unfolding matrix_A_def by auto lemma dim_col_matrix_A[simp]: shows "dim_col (matrix_A signs subsets) = length signs" unfolding matrix_A_def by auto lemma length_signs_smash: shows "length (signs_smash signs1 signs2) = length signs1 * length signs2" unfolding signs_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_subsets_smash: shows "length (subsets_smash n subs1 subs2) = length subs1 * length subs2" unfolding subsets_smash_def length_concat by (auto simp add: o_def sum_list_triv) lemma length_eq_concat: assumes "\x. x \ set ls \ length x = n" assumes "i < n * length ls" shows "concat ls ! i = ls ! (i div n) ! (i mod n)" using assms proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons a ls) - then show ?case - apply (auto simp add: nth_append) - using div_if mod_geq by auto + moreover have \n > 0\ + by (rule ccontr) (use assms in auto) + ultimately show ?case + by (auto simp add: nth_append not_less dest: le_Suc_ex) qed lemma z_append: assumes "\i. i \ set xs \ i < length as" shows "z (xs @ (map ((+) (length as)) ys)) (as @ bs) = z xs as * z ys bs" proof - have 1: "map ((!) (as @ bs)) xs = map ((!) as) xs" unfolding list_eq_iff_nth_eq using assms by (auto simp add:nth_append) have 2: "map ((!) (as @ bs) \ (+) (length as)) ys = map ((!) bs) ys" unfolding list_eq_iff_nth_eq using assms by auto show ?thesis unfolding z_def apply auto unfolding 1 2 by auto qed (* Shows that the matrix of a smashed system is the Kronecker product of the matrices of the two systems being combined *) lemma matrix_construction_is_kronecker_product: fixes qs1 :: "real poly list" fixes subs1 subs2 :: "nat list list" fixes signs1 signs2 :: "rat list list" (* n1 is the number of polynomials in the "1" sets *) assumes "\l i. l \ set subs1 \ i \ set l \ i < n1" assumes "\j. j \ set signs1 \ length j = n1" shows " (matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2)) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2)" unfolding mat_eq_iff dim_row_matrix_A dim_col_matrix_A length_subsets_smash length_signs_smash dim_kronecker proof safe fix i j assume i: "i < length subs1 * length subs2" and j: "j < length signs1 * length signs2" have ld: "i div length subs2 < length subs1" "j div length signs2 < length signs1" using i j less_mult_imp_div_less by auto have lm: "i mod length subs2 < length subs2" "j mod length signs2 < length signs2" using i j less_mult_imp_mod_less by auto have n1: "n1 = length (signs1 ! (j div length signs2))" using assms(2) ld(2) nth_mem by blast have 1: "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = z (subsets_smash n1 subs1 subs2 ! i) (signs_smash signs1 signs2 ! j)" unfolding mat_of_rows_list_def matrix_A_def mtx_row_def using i j by (simp add: length_signs_smash length_subsets_smash) have 2: " ... = z (subs1 ! (i div length subs2) @ map ((+) n1) (subs2 ! (i mod length subs2))) (signs1 ! (j div length signs2) @ signs2 ! (j mod length signs2))" unfolding signs_smash_def subsets_smash_def apply (subst length_eq_concat) using i apply (auto simp add: mult.commute) apply (subst length_eq_concat) using j apply (auto simp add: mult.commute) using ld lm by auto have 3: "... = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding n1 apply (subst z_append) apply (auto simp add: n1[symmetric]) using assms(1) ld(1) nth_mem by blast have 4: "kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i,j) = z (subs1 ! (i div length subs2)) (signs1 ! (j div length signs2)) * z (subs2 ! (i mod length subs2)) (signs2 ! (j mod length signs2))" unfolding kronecker_product_def matrix_A_def mat_of_rows_list_def mtx_row_def using i j apply (auto simp add: Let_def) apply (subst index_mat(1)[OF ld]) apply (subst index_mat(1)[OF lm]) using ld lm by auto show "matrix_A (signs_smash signs1 signs2) (subsets_smash n1 subs1 subs2) $$ (i, j) = kronecker_product (matrix_A signs1 subs1) (matrix_A signs2 subs2) $$ (i, j)" using 1 2 3 4 by auto qed (* Given that two smaller systems satisfy some mild constraints, show that their combined system (after smashing the signs and subsets) satisfies the matrix equation, and that the matrix of the combined system is invertible. *) lemma inductive_step: fixes p:: "real poly" fixes qs1 qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes welldefined_signs1: "well_def_signs (length qs1) signs1" assumes welldefined_signs2: "well_def_signs (length qs2) signs2" assumes distinct_signs1: "distinct signs1" assumes distinct_signs2: "distinct signs2" assumes all_info1: "set (characterize_consistent_signs_at_roots_copr p qs1) \ set(signs1)" assumes all_info2: "set (characterize_consistent_signs_at_roots_copr p qs2) \ set(signs2)" assumes welldefined_subsets1: "all_list_constr (subsets1) (length qs1)" assumes welldefined_subsets2: "all_list_constr (subsets2) (length qs2)" assumes invertibleMtx1: "invertible_mat (matrix_A signs1 subsets1)" assumes invertibleMtx2: "invertible_mat (matrix_A signs2 subsets2)" shows "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2) \ invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" proof - have h1: "invertible_mat (matrix_A (signs_smash signs1 signs2) (subsets_smash (length qs1) subsets1 subsets2))" using matrix_construction_is_kronecker_product kronecker_invertible invertibleMtx1 invertibleMtx2 welldefined_subsets1 welldefined_subsets2 unfolding all_list_constr_def list_constr_def by (smt in_set_conv_nth in_set_member list_all_length well_def_signs_def welldefined_signs1) have h2a: "distinct (signs_smash signs1 signs2)" using distinct_signs1 distinct_signs2 welldefined_signs1 welldefined_signs2 nontriv1 nontriv2 distinct_step by auto have h2ba: "\ q. List.member (qs1 @ qs2) q \ (List.member qs1 q \ List.member qs2 q)" unfolding member_def by auto have h2b: "\q. ((List.member (qs1@qs2) q) \ (coprime p q))" using h2ba pairwise_rel_prime1 pairwise_rel_prime2 by auto have h2c: "all_list_constr ((subsets_smash (length qs1) subsets1 subsets2)) (length (qs1@qs2))" using well_def_step welldefined_subsets1 welldefined_subsets2 by blast have h2d: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set(signs_smash signs1 signs2)" using subset_step all_info1 all_info2 by simp have h2: "satisfy_equation p (qs1@qs2) (subsets_smash (length qs1) subsets1 subsets2) (signs_smash signs1 signs2)" using matrix_equation[where p="p", where qs="qs1@qs2", where subsets = "subsets_smash (length qs1) subsets1 subsets2", where signs = "signs_smash signs1 signs2"] h2a h2b h2c h2d apply (auto) using nonzero by blast show ?thesis using h1 h2 by blast qed section "Reduction Step Proofs" (* Some definitions *) definition get_matrix:: "(rat mat \ (nat list list \ rat list list)) \ rat mat" where "get_matrix data = fst(data)" definition get_subsets:: "(rat mat \ (nat list list \ rat list list)) \ nat list list" where "get_subsets data = fst(snd(data))" definition get_signs:: "(rat mat \ (nat list list \ rat list list)) \ rat list list" where "get_signs data = snd(snd(data))" definition reduction_signs:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ rat list list" where "reduction_signs p qs signs subsets matr = (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets matr)))" definition reduction_subsets:: "real poly \ real poly list \ rat list list \ nat list list \ rat mat \ nat list list" where "reduction_subsets p qs signs subsets matr = (take_indices subsets (rows_to_keep (reduce_mat_cols matr (solve_for_lhs p qs subsets matr))))" (* Some basic lemmas *) lemma reduction_signs_is_get_signs: "reduction_signs p qs signs subsets m = get_signs (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_signs_def get_signs_def by (metis reduce_system.simps reduction_step.elims snd_conv) lemma reduction_subsets_is_get_subsets: "reduction_subsets p qs signs subsets m = get_subsets (reduce_system p (qs, (m, (subsets, signs))))" unfolding reduction_subsets_def get_subsets_def by (metis fst_conv reduce_system.simps reduction_step.elims snd_conv) lemma find_zeros_from_vec_prop: fixes input_vec :: "rat vec" shows "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" proof - have h1: "\n < (dim_vec input_vec). ((input_vec $ n \ 0) \ List.member (find_nonzeros_from_input_vec input_vec) n)" unfolding find_nonzeros_from_input_vec_def List.member_def by auto have h2: "\n < (dim_vec input_vec). (List.member (find_nonzeros_from_input_vec input_vec) n \ (input_vec $ n \ 0) )" unfolding find_nonzeros_from_input_vec_def List.member_def by auto show ?thesis using h1 h2 by auto qed subsection "Showing sign conditions preserved when reducing" lemma take_indices_lem: fixes p:: "real poly" fixes qs :: "real poly list" fixes arb_list :: "'a list list" fixes index_list :: "nat list" fixes n:: "nat" assumes lest: "n < length (take_indices arb_list index_list)" assumes well_def: "\q.(List.member index_list q \ q < length arb_list)" shows "\km (dim_row A)" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis by (metis assms carrier_matI invertible_mat_def mat_inverse(2) matr_option.simps(2) square_mat.elims(2)) qed lemma dim_invertible: fixes A:: "'a::field mat" fixes n:: "nat" assumes assm: "invertible_mat A" assumes dim: "A \ carrier_mat n n" shows "matr_option (dim_row A) (mat_inverse (A)) \ carrier_mat n n" proof (cases "mat_inverse(A)") obtain n where n: "A \ carrier_mat n n" using assms invertible_mat_def square_mat.simps by blast case None then have "A \ Units (ring_mat TYPE('a) n n)" by (simp add: mat_inverse(1) n) thus ?thesis by (meson assms det_non_zero_imp_unit invertible_Units n unit_imp_det_non_zero) next case (Some a) then show ?thesis using dim mat_inverse(2) by auto qed lemma mult_assoc: fixes A B:: "rat mat" fixes v:: "rat vec" fixes n:: "nat" assumes "A \ carrier_mat n n" assumes "B \ carrier_mat n n" assumes "dim_vec v = n" shows "A *\<^sub>v (mult_mat_vec B v) = (A*B)*\<^sub>v v" using assms(1) assms(2) assms(3) by auto lemma size_of_mat: fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(matrix_A signs subsets) \ carrier_mat (length subsets) (length signs)" unfolding matrix_A_def carrier_mat_def by auto lemma size_of_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes signs :: "rat list list" shows "dim_vec (construct_lhs_vector p qs signs) = length signs" unfolding construct_lhs_vector_def by simp lemma size_of_rhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" shows "dim_vec (construct_rhs_vector p qs subsets) = length subsets" unfolding construct_rhs_vector_def by simp lemma same_size: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "length subsets = length signs" using invertible_mat unfolding invertible_mat_def using size_of_mat[of signs subsets] size_of_lhs[of p qs signs] size_of_rhs[of p qs subsets] by simp lemma mat_equal_list_lem: fixes A:: "'a::field mat" fixes B:: "'a::field mat" shows "(dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B) \ A = B" proof - assume hyp: "dim_row A = dim_row B \ dim_col A = dim_col B \ mat_to_list A = mat_to_list B" then have "\i j. i < dim_row A \ j < dim_col A \ B $$ (i, j) = A $$ (i, j)" unfolding mat_to_list_def by auto then show ?thesis using hyp unfolding mat_to_list_def using eq_matI[of A B] by metis qed lemma mat_inverse_same: "mat_inverse_var A = mat_inverse A" unfolding mat_inverse_var_def mat_inverse_def mat_equal_def using mat_equal_list_lem apply (simp) by (smt case_prod_beta index_one_mat(2) index_one_mat(3) mat_equal_list_lem) lemma construct_lhs_matches_solve_for_lhs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" proof - have matrix_equation_hyp: "(mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs) = (construct_rhs_vector p qs subsets))" using match unfolding satisfy_equation_def by blast then have eqn_hyp: " ((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs)) = mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets))" using invertible_mat by (simp add: matrix_equation_hyp) have match_hyp: "length subsets = length signs" using same_size invertible_mat by auto then have dim_hyp1: "matrix_A signs subsets \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible by blast have mult_assoc_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs))" using mult_assoc dim_hyp1 dim_hyp2 size_of_lhs by auto have cancel_helper: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (1\<^sub>m (length signs)) *\<^sub>v (construct_lhs_vector p qs signs)" using invertible_means_mult_id[where A= "matrix_A signs subsets"] dim_hyp1 by (simp add: invertible_mat match_hyp) then have cancel_hyp: "(((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) * (matrix_A signs subsets)) *\<^sub>v (construct_lhs_vector p qs signs)) = (construct_lhs_vector p qs signs)" apply (auto) by (metis carrier_vec_dim_vec one_mult_mat_vec size_of_lhs) then have mult_hyp: "((matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) *\<^sub>v (mult_mat_vec (matrix_A signs subsets) (construct_lhs_vector p qs signs))) = (construct_lhs_vector p qs signs)" using mult_assoc_hyp cancel_hyp by simp then have "(construct_lhs_vector p qs signs) = (mult_mat_vec (matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets))) (construct_rhs_vector p qs subsets)) " using eqn_hyp by simp then show ?thesis unfolding solve_for_lhs_def by (simp add: mat_inverse_same) qed (* If set(A) is a subset of B then for all n, nth c n = 0 means nth a n not in B *) lemma reduction_signs_set_helper_lemma: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n = 0 \ (nth B n) \ A" shows "A \ set (take_indices B (find_nonzeros_from_input_vec C))" proof - have unfold: "\ x. (x \ A \ x \ set (take_indices B (find_nonzeros_from_input_vec C)))" proof - fix x assume in_a: "x \ A" have "x \ set (B)" using sub in_a by blast then have "\ n < length B. nth B n = x" by (simp add: in_set_conv_nth) then have "\ n < length B. (nth B n = x \ (List.member (find_nonzeros_from_input_vec C) n) = True)" using not_in_hyp find_zeros_from_vec_prop[of C] using assms(1) in_a by auto thus "x \ set (take_indices B (find_nonzeros_from_input_vec C))" unfolding take_indices_def using member_def by fastforce qed then show ?thesis by blast qed lemma reduction_signs_set_helper_lemma2: fixes A:: "rat list set" fixes C:: "rat vec" fixes B:: "rat list list" assumes dist: "distinct B" assumes eq_len: "dim_vec C = length B" assumes sub: "A \ set(B)" assumes not_in_hyp: "\ n < dim_vec C. C $ n \ 0 \ (nth B n) \ A" shows "set (take_indices B (find_nonzeros_from_input_vec C)) \ A" proof - have unfold: "\ x. (x \ (set (take_indices B (find_nonzeros_from_input_vec C))) \ x \ A)" proof - fix x assume in_set: "x \ set (take_indices B (find_nonzeros_from_input_vec C))" have h: "\n< dim_vec C. (C $ n \ 0 \ (nth B n) = x)" proof - have h1: "\n < length B.(nth B n) = x" using in_set unfolding take_indices_def find_nonzeros_from_input_vec_def eq_len by auto have h2: "\n< length B. (nth B n = x \ List.member (find_nonzeros_from_input_vec C) n)" proof clarsimp fix n assume leq_hyp: "n < length B" assume x_eq: "x = B ! n" have h: "(B !n) \ set (map ((!) B) (find_nonzeros_from_input_vec C))" using x_eq in_set by (simp add: take_indices_def) then have h2: "List.member (map ((!) B) (find_nonzeros_from_input_vec C)) (B ! n)" using in_set by (meson in_set_member) then have h3: "\k< length B. (List.member (find_nonzeros_from_input_vec C) k \ ((B ! k) = (B ! n)))" by (smt atLeastLessThan_iff eq_len find_nonzeros_from_input_vec_def imageE in_set_member mem_Collect_eq set_filter set_map set_upt) have h4: "\v< length B. ((B ! v) = (B ! n) \ v = n)" using dist apply (auto) using leq_hyp nth_eq_iff_index_eq by blast then show "List.member (find_nonzeros_from_input_vec C) n" using h2 h3 h4 by auto qed then have "\n C $ n \ 0)" using find_zeros_from_vec_prop [of C] by (simp add: eq_len) then show ?thesis using h1 eq_len by auto qed thus "x \ A" using not_in_hyp by blast qed then show ?thesis by blast qed (* Show that dropping columns doesn't affect the consistent sign assignments *) lemma reduction_doesnt_break_things_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix x assume card_asm: "card {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = 0" assume zero_asm: "poly p x = 0" have card_hyp: "{xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x} = {}" using card_eq_0_iff using card_asm nonzero poly_roots_finite by fastforce have "x \ {xa. poly p xa = 0 \ consistent_sign_vec_copr qs xa = consistent_sign_vec_copr qs x}" using zero_asm by auto thus "False" using card_hyp by blast qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ (\List.member (characterize_consistent_signs_at_roots_copr p qs) w))" by (smt (verit, best) characterize_consistent_signs_at_roots_copr_def characterize_root_list_p_def h1 imageE in_set_member mem_Collect_eq nonzero poly_roots_finite set_map set_remdups sorted_list_of_set(1)) then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) = 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by blast qed then have "set (characterize_consistent_signs_at_roots_copr p qs) \ set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using all_info reduction_signs_set_helper_lemma[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed lemma reduction_deletes_bad_sign_conds: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (characterize_consistent_signs_at_roots_copr p qs) = set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have dim_hyp2: "matr_option (dim_row (matrix_A signs subsets)) (mat_inverse (matrix_A signs subsets)) \ carrier_mat (length signs) (length signs)" using invertible_mat dim_invertible using same_size by fastforce have supset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have "(construct_lhs_vector p qs signs) = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by auto then have "(solve_for_lhs p qs subsets (matrix_A signs subsets)) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))" using construct_lhs_vector_cleaner assms by (metis (mono_tags, lifting) list.map_cong map_map o_apply of_int_of_nat_eq) then have "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n \ 0) \ (nth signs n) \ set (characterize_consistent_signs_at_roots_copr p qs))" proof - have h0: "\ n < (dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))). (((solve_for_lhs p qs subsets (matrix_A signs subsets)) $ n = 0) \ rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = (nth signs n)}) = 0)" by (metis (mono_tags, lifting) \construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)\ construct_lhs_vector_clean nonzero of_nat_0_eq_iff of_rat_of_nat_eq size_of_lhs) have h1: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (\ x. poly p x = 0 \ consistent_sign_vec_copr qs x = w))" proof clarsimp fix w assume card_asm: "0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" show "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" by (metis (mono_tags, lifting) Collect_empty_eq card_asm card_eq_0_iff gr_implies_not0) qed have h2: "\ w. (rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ (List.member (characterize_consistent_signs_at_roots_copr p qs) w))" proof clarsimp fix w assume card_asm: " 0 < card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}" have h0: "\x. poly p x = 0 \ consistent_sign_vec_copr qs x = w" using card_asm by (simp add: h1) then show "List.member (characterize_consistent_signs_at_roots_copr p qs) w" unfolding characterize_consistent_signs_at_roots_copr_def using in_set_member nonzero poly_roots_finite characterize_root_list_p_def by fastforce qed then have h3: "\ w. rat_of_nat (card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = w}) \ 0 \ w \ set (characterize_consistent_signs_at_roots_copr p qs)" by (simp add: in_set_member) show ?thesis using h0 h3 by (metis (no_types, lifting) \solve_for_lhs p qs subsets (matrix_A signs subsets) = vec_of_list (map rat_of_nat (map (\s. card {x. poly p x = 0 \ consistent_sign_vec_copr qs x = s}) signs))\ dim_vec_of_list length_map nth_map vec_of_list_index) qed then have "set (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (characterize_consistent_signs_at_roots_copr p qs)" using all_info reduction_signs_set_helper_lemma2[where A = "set (characterize_consistent_signs_at_roots_copr p qs)", where B = "signs", where C = "(solve_for_lhs p qs subsets (matrix_A signs subsets))"] using distinct_signs dim_hyp2 solve_for_lhs_def by (simp add: mat_inverse_same) then show ?thesis unfolding reduction_signs_def by auto qed have subset: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(reduction_signs p qs signs subsets (matrix_A signs subsets))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms by blast then show ?thesis using supset subset by auto qed theorem reduce_system_sign_conditions: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" shows "set (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using reduction_deletes_bad_sign_conds[of p qs signs subsets] apply (auto) apply (simp add: all_info distinct_signs match nonzero reduction_signs_def welldefined_signs1) using nonzero invertible_mat apply (metis snd_conv) by (metis all_info distinct_signs invertible_mat match nonzero reduction_signs_def snd_conv welldefined_signs1) subsection "Showing matrix equation preserved when reducing" lemma rows_to_keep_lem: fixes A:: "('a::field) mat" shows "\ell. ell \ set (rows_to_keep A) \ ell < dim_row A" unfolding rows_to_keep_def apply auto using rref_pivot_positions by (metis carrier_mat_triv gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(3)) lemma reduce_system_matrix_equation_preserved: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs: "well_def_signs (length qs) signs" assumes welldefined_subsets: "all_list_constr (subsets) (length qs)" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes invertible_mat: "invertible_mat (matrix_A signs subsets)" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfy_equation p qs (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - have poly_type_hyp: "p \ 0" using nonzero by auto have distinct_signs_hyp: "distinct (snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" proof - let ?sym = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" have h1: "\ i < length (take_indices signs ?sym). \j < length(take_indices signs ?sym). i \ j \ nth (take_indices signs ?sym) i \ nth (take_indices signs ?sym) j" using distinct_signs unfolding take_indices_def proof clarsimp fix i fix j assume "distinct signs" assume "i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume "j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume neq_hyp: "i \ j" assume "signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)" have h1: "find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i \ find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j" unfolding find_nonzeros_from_input_vec_def using neq_hyp by (metis \i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ distinct_conv_nth distinct_filter distinct_upt find_nonzeros_from_input_vec_def) then show "False" using distinct_signs proof - have f1: "\p ns n. ((n::nat) \ {n \ set ns. p n}) = (n \ set ns \ n \ Collect p)" by simp then have f2: "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..i < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) have "filter (\n. solve_for_lhs p qs subsets (matrix_A signs subsets) $ n \ 0) [0.. set [0..j < length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ construct_lhs_matches_solve_for_lhs find_nonzeros_from_input_vec_def invertible_mat match nth_mem set_filter size_of_lhs) then show ?thesis using f2 by (metis \signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! i) = signs ! (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) ! j)\ atLeastLessThan_iff distinct_conv_nth distinct_signs find_nonzeros_from_input_vec_def h1 set_upt) qed qed then have "distinct (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" using distinct_conv_nth by blast then show ?thesis using get_signs_def reduction_signs_def reduction_signs_is_get_signs by auto qed have all_info_hyp: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(snd (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))))" using reduce_system_sign_conditions assms unfolding get_signs_def by auto have pairwise_rel_prime_hyp: "\q. ((List.member qs q) \ (coprime p q))" using pairwise_rel_prime by auto have welldefined_hyp: "all_list_constr (fst (snd (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))) (length qs)" using welldefined_subsets rows_to_keep_lem unfolding all_list_constr_def List.member_def list_constr_def list_all_def apply (auto simp add: Let_def take_indices_def take_cols_from_matrix_def) using nth_mem by fastforce then show ?thesis using poly_type_hyp distinct_signs_hyp all_info_hyp pairwise_rel_prime_hyp welldefined_hyp using matrix_equation unfolding get_subsets_def get_signs_def by blast qed (* Show that we are tracking the correct matrix in the algorithm *) subsection "Showing matrix preserved" lemma reduce_system_matrix_signs_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "alt_matrix_A (take_indices signs S) subsets = take_cols_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_col (take_cols_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices signs S)" unfolding take_cols_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices signs S). (col (alt_matrix_A (take_indices signs S) subsets ) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices signs S)" have i_lt: "i < length (map ((!) (cols (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: " vec (length subsets) (\j. z (subsets ! j) (map ((!) signs) S ! i)) = vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i)))" using nth_map by (metis \i < length (take_indices signs S)\ length_map take_indices_def) have dim: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_row (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) have well_d: "S ! i < length (signs)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (cols (alt_matrix_A signs subsets))) S) ! i = nth (cols (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (cols (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_row (alt_matrix_A signs subsets))" using col_dim unfolding cols_def using nth_map well_d by (simp add: \dim_col (alt_matrix_A signs subsets) = length signs\) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i " by (simp add: take_cols_from_matrix_def take_indices_def) have h2: "col (mat_of_cols (dim_row (alt_matrix_A signs subsets)) (map ((!) (cols (alt_matrix_A signs subsets))) S)) i = nth (map ((!) (cols (alt_matrix_A signs subsets))) S) i " using dim i_lt asm col_mat_of_cols[where j = "i", where n = "(dim_row (alt_matrix_A signs subsets))", where vs = "(map ((!) (cols (alt_matrix_A signs subsets))) S)"] by blast have h3: "col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i = (col (alt_matrix_A signs subsets) (S !i))" using h1 h2 apply (auto) by (metis alt_matrix_char asm cols_nth dim_col_mat(1) in_set_member length_map mat_of_rows_list_def matrix_A_def nth_map nth_mem take_indices_def well_def_h) have "vec (length subsets) (\j. z (subsets ! j) (signs ! (S ! i))) = (col (alt_matrix_A signs subsets) (S !i))" by (metis asm in_set_member length_map nth_mem signs_are_cols take_indices_def well_def_h) then have "vec (length subsets) (\j. z (subsets ! j) (take_indices signs S ! i)) = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using h0 h3 by (simp add: take_indices_def) then show "col (alt_matrix_A (take_indices signs S) subsets) i = col (take_cols_from_matrix (alt_matrix_A signs subsets) S) i " using asm signs_are_cols[where signs = "(take_indices signs S)", where subsets = "subsets"] by auto qed then show ?thesis unfolding alt_matrix_A_def take_cols_from_matrix_def apply (auto) using h0a mat_col_eqI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 mat_of_cols_def take_cols_from_matrix_def) qed lemma reduce_system_matrix_signs_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes well_def_h: "\x. List.member S x \ x < length signs" assumes nonzero: "p \ 0" shows "matrix_A (take_indices signs S) subsets = take_cols_from_matrix (matrix_A signs subsets) S" using reduce_system_matrix_signs_helper_aux alt_matrix_char assms by auto lemma reduce_system_matrix_subsets_helper_aux: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" assumes nonzero: "p \ 0" shows "alt_matrix_A signs (take_indices subsets S) = take_rows_from_matrix (alt_matrix_A signs subsets) S" proof - have h0a: "dim_row (take_rows_from_matrix (alt_matrix_A signs subsets) S) = length (take_indices subsets S)" unfolding take_rows_from_matrix_def apply (auto) unfolding take_indices_def by auto have h0: "\i < length (take_indices subsets S). (row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" proof clarsimp fix i assume asm: "i < length (take_indices subsets S)" have i_lt: "i < length (map ((!) (rows (alt_matrix_A signs subsets))) S)" using asm apply (auto) unfolding take_indices_def by auto have h0: "row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i = row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i " unfolding take_rows_from_matrix_def take_indices_def by auto have dim: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i \ carrier_vec (dim_col (alt_matrix_A signs subsets))" proof - have "dim_col (alt_matrix_A signs subsets) = length (signs)" by (simp add: alt_matrix_A_def) then have lenh: "dim_col (alt_matrix_A signs subsets) \ length (subsets)" using assms by auto have well_d: "S ! i < length (subsets)" using well_def_h using i_lt in_set_member by fastforce have map_eq: "(map ((!) (rows (alt_matrix_A signs subsets))) S) ! i = nth (rows (alt_matrix_A signs subsets)) (S ! i)" using i_lt by auto have "nth (rows (alt_matrix_A signs subsets)) (S ! i) \ carrier_vec (dim_col (alt_matrix_A signs subsets))" using col_dim unfolding rows_def using nth_map well_d using lenh by (simp add: alt_matrix_A_def) then show ?thesis using map_eq unfolding alt_matrix_A_def by auto qed have h1: " row (mat_of_rows (dim_col (alt_matrix_A signs subsets)) (map ((!) (rows (alt_matrix_A signs subsets))) S)) i = (row (alt_matrix_A signs subsets) (S ! i)) " using dim i_lt mat_of_rows_row[where i = "i", where n = "(dim_col (alt_matrix_A signs subsets))", where vs = "(map ((!) (rows (alt_matrix_A signs subsets))) S)"] using alt_matrix_char in_set_member nth_mem well_def_h by fastforce have "row (alt_matrix_A signs (take_indices subsets S) ) i = (row (alt_matrix_A signs subsets) (S ! i))" using subsets_are_rows proof - have f1: "i < length S" by (metis (no_types) asm length_map take_indices_def) then have "List.member S (S ! i)" by (meson in_set_member nth_mem) then show ?thesis using f1 by (simp add: \\subsets signs. \ij. z (subsets ! i) (signs ! j))\ take_indices_def well_def_h) qed then show "(row (alt_matrix_A signs (take_indices subsets S) ) i = row (take_rows_from_matrix (alt_matrix_A signs subsets) S) i)" using h0 h1 unfolding take_indices_def by auto qed then show ?thesis unfolding alt_matrix_A_def take_rows_from_matrix_def apply (auto) using eq_rowI by (metis alt_matrix_A_def dim_col_mat(1) dim_row_mat(1) h0 length_map mat_of_rows_def take_indices_def take_rows_from_matrix_def) qed lemma reduce_system_matrix_subsets_helper: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes S:: "nat list" assumes nonzero: "p \ 0" assumes inv: "length subsets \ length signs" assumes well_def_h: "\x. List.member S x \ x < length subsets" shows "matrix_A signs (take_indices subsets S) = take_rows_from_matrix (matrix_A signs subsets) S" using assms reduce_system_matrix_subsets_helper_aux alt_matrix_char by auto lemma reduce_system_matrix_match: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" assumes inv: "invertible_mat (matrix_A signs subsets)" shows "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?A = "(matrix_A signs subsets)" let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" let ?red_mtx = "(take_rows_from_matrix (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec) (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) ?lhs_vec)))" have h1: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (reduction_signs p qs signs subsets (matrix_A signs subsets)) (reduction_subsets p qs signs subsets (matrix_A signs subsets))" using reduction_signs_is_get_signs reduction_subsets_is_get_subsets by auto have h1_var: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)))" using h1 reduction_signs_def reduction_subsets_def by auto have h2: "?red_mtx = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" by simp have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3b_a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (subsets)" using assms h30 size_of_lhs same_size unfolding find_nonzeros_from_input_vec_def apply (auto) by (simp add: find_nonzeros_from_input_vec_def h3a) have h3ba: "length (filter (\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using length_filter_le[where P = "(\i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0)", where xs = "[0..i. solve_for_lhs p qs subsets (matrix_A signs subsets) $ i \ 0) [0.. length subsets" using h3ba by auto then have h3b: "length subsets \ length (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec))" unfolding take_indices_def find_nonzeros_from_input_vec_def by auto have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols ?A ?lhs_vec)) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" using f4 by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def by (metis (no_types) dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed have h3: "matrix_A (take_indices signs (find_nonzeros_from_input_vec ?lhs_vec)) (take_indices subsets (rows_to_keep (reduce_mat_cols ?A ?lhs_vec))) = (take_rows_from_matrix (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec)) (rows_to_keep (take_cols_from_matrix ?A (find_nonzeros_from_input_vec ?lhs_vec))))" using inv h3a h3b h3c reduce_system_matrix_subsets_helper reduce_system_matrix_signs_helper assms by auto show ?thesis using h1 h2 h3 by (metis fst_conv get_matrix_def h1_var reduce_system.simps reduction_step.simps) qed (* gauss_jordan_single_rank is crucial in this section *) subsection "Showing invertibility preserved when reducing" (* Overall: Start with a matrix equation. Input a matrix, subsets, and signs. Drop columns of the matrix based on the 0's on the LHS---so extract a list of 0's. Reduce signs accordingly. Then find a list of rows to delete based on using rank (use the transpose result, pivot positions!), and delete those rows. Reduce subsets accordingly. End with a reduced system! *) lemma well_def_find_zeros_from_lhs_vec: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" proof - fix i fix j assume j_in: " j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) " let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)) = (length signs)" using size_of_lhs construct_lhs_matches_solve_for_lhs assms by (metis (full_types)) then have h: "j < (length signs)" using j_in unfolding find_nonzeros_from_input_vec_def by simp then show "j < length (cols (matrix_A signs subsets))" using mat_size by auto qed lemma take_cols_subsets_og_cols: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" proof - have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have "\x. x \ set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ x \ set (cols (matrix_A signs subsets))" proof clarsimp fix x let ?og_list = "(cols (matrix_A signs subsets))" let ?ind_list = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" assume x_in: "x \ set (take_indices ?og_list ?ind_list)" show "x \ set (cols (matrix_A signs subsets))" using x_in unfolding take_indices_def using in_set_member apply (auto) using in_set_conv_nth well_def by fastforce qed then show ?thesis by blast qed lemma reduction_doesnt_break_things_invertibility_step1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "vec_space.rank (length signs) (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets))) = (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?new_mat = "(take_rows_from_matrix (reduce_mat_cols ?og_mat ?lhs) (rows_to_keep (reduce_mat_cols ?og_mat ?lhs)))" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto then have mat_size_alt: "?og_mat \ carrier_mat (length subsets) (length subsets)" using size_of_mat same_size assms by auto have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff mat_size by auto then have dist_cols: "distinct (cols ?og_mat)" using mat_size vec_space.non_distinct_low_rank[where A = ?og_mat, where n = "length signs"] by auto have well_def: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length (cols (matrix_A signs subsets)))" using assms well_def_find_zeros_from_lhs_vec by auto have dist1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def by auto have clear: "set (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ set (cols (matrix_A signs subsets))" using assms take_cols_subsets_og_cols by auto then have "distinct (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" unfolding take_indices_def using dist1 dist_cols well_def conjugatable_vec_space.distinct_map_nth[where ls = "cols (matrix_A signs subsets)", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by auto then have unfold_thesis: "vec_space.rank (length signs) (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) (find_nonzeros_from_input_vec ?lhs))) = (length (find_nonzeros_from_input_vec ?lhs))" using clear inv conjugatable_vec_space.rank_invertible_subset_cols[where A= "matrix_A signs subsets", where B = "(take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))"] by (simp add: len_eq mat_size take_indices_def) then show ?thesis apply (simp) unfolding take_cols_from_matrix_def by auto qed lemma rechar_take_cols: "take_cols_var A B = take_cols_from_matrix A B" unfolding take_cols_var_def take_cols_from_matrix_def take_indices_def by auto lemma rows_and_cols_transpose: "rows M = cols M\<^sup>T" using row_transpose by simp lemma take_rows_and_take_cols: "(take_rows_from_matrix M r) = (take_cols_from_matrix M\<^sup>T r)\<^sup>T" unfolding take_rows_from_matrix_def take_cols_from_matrix_def using transpose_carrier_mat rows_and_cols_transpose apply (auto) by (simp add: transpose_mat_of_cols) lemma reduction_doesnt_break_things_invertibility: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes len_eq: "length subsets = length signs" assumes inv: "invertible_mat (matrix_A signs subsets)" assumes nonzero: "p \ 0" assumes welldefined_signs1: "well_def_signs (length qs) signs" assumes distinct_signs: "distinct signs" assumes all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(signs)" assumes match: "satisfy_equation p qs subsets signs" shows "invertible_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" proof - let ?og_mat = "(matrix_A signs subsets)" let ?lhs = "(solve_for_lhs p qs subsets ?og_mat)" let ?step1_mat = "(reduce_mat_cols ?og_mat ?lhs)" let ?new_mat = "take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)" let ?ind_list = "(find_nonzeros_from_input_vec ?lhs)" have "square_mat (matrix_A signs subsets)" using inv using invertible_mat_def by blast then have og_mat_size: "?og_mat \ carrier_mat (length signs) (length signs)" using size_of_mat by auto have "dim_col (mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list)) = (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_indices_def) then have "mat_of_cols (dim_row ?og_mat) (take_indices (cols ?og_mat) ?ind_list) \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: len_eq mat_of_cols_def) then have step1_mat_size: "?step1_mat \ carrier_mat (length signs) (length (find_nonzeros_from_input_vec ?lhs))" by (simp add: take_cols_from_matrix_def) then have "dim_row ?step1_mat \ dim_col ?step1_mat" by (metis carrier_matD(1) carrier_matD(2) construct_lhs_matches_solve_for_lhs diff_zero find_nonzeros_from_input_vec_def inv length_filter_le length_upt match size_of_lhs) then have gt_eq_assm: "dim_col ?step1_mat\<^sup>T \ dim_row ?step1_mat\<^sup>T" by simp have det_h: "det ?og_mat \ 0" using invertible_det[where A = "matrix_A signs subsets"] og_mat_size using inv by blast then have rank_h: "vec_space.rank (length signs) ?og_mat = (length signs)" using vec_space.det_rank_iff og_mat_size by auto have rank_drop_cols: "vec_space.rank (length signs) ?step1_mat = (dim_col ?step1_mat)" using assms reduction_doesnt_break_things_invertibility_step1 step1_mat_size by auto let ?step1_T = "?step1_mat\<^sup>T" have rank_transpose: "vec_space.rank (length signs) ?step1_mat = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using transpose_rank[of ?step1_mat] using step1_mat_size by auto have obv: "?step1_T \ carrier_mat (dim_row ?step1_T) (dim_col ?step1_T)" by auto have should_have_this:"vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) ?step1_T" using obv gt_eq_assm conjugatable_vec_space.gauss_jordan_single_rank[where A = "?step1_T", where n = "dim_row ?step1_T", where nc = "dim_col ?step1_T"] by (simp add: take_cols_from_matrix_def take_indices_def) then have "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_drop_cols rank_transpose by auto then have with_take_cols_from_matrix: "vec_space.rank (length (find_nonzeros_from_input_vec ?lhs)) (take_cols_from_matrix ?step1_T (map snd (pivot_positions (gauss_jordan_single (?step1_T))))) = dim_col ?step1_mat" using rank_transpose rechar_take_cols conjugatable_vec_space.gjs_and_take_cols_var apply (auto) by (smt conjugatable_vec_space.gjs_and_take_cols_var gt_eq_assm obv rechar_take_cols reduce_mat_cols.simps) have "(take_rows_from_matrix ?step1_mat (rows_to_keep ?step1_mat)) = (take_cols_from_matrix ?step1_T (rows_to_keep ?step1_mat))\<^sup>T" using take_rows_and_take_cols by blast then have rank_new_mat: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_col ?step1_mat" using with_take_cols_from_matrix transpose_rank apply (auto) proof - assume a1: "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))) = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" have f2: "\m. vec_space.rank (dim_row (m::rat mat)) m = vec_space.rank (dim_row m\<^sup>T) m\<^sup>T" by (simp add: transpose_rank) have f3: "dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ by force have "vec_space.rank (length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T))))) = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using a1 by (simp add: take_cols_from_matrix_def) then have "vec_space.rank (dim_row (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T) (mat_of_cols (dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (take_indices (cols (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T) (map snd (pivot_positions (gauss_jordan_single (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T)))))\<^sup>T = dim_row (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))\<^sup>T" using f3 f2 by presburger then show "vec_space.rank (dim_col (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))) (take_cols_from_matrix (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))))\<^sup>T = dim_col (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))" by (simp add: rows_to_keep_def take_cols_from_matrix_def) qed have "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ (length (find_nonzeros_from_input_vec ?lhs))" using obv length_pivot_positions_dim_row[where A = "(gauss_jordan_single (?step1_mat\<^sup>T))"] by (metis carrier_matD(1) carrier_matD(2) gauss_jordan_single(2) gauss_jordan_single(3) index_transpose_mat(2) step1_mat_size) then have len_lt_eq: "length (pivot_positions (gauss_jordan_single (?step1_mat\<^sup>T))) \ dim_col ?step1_mat" using step1_mat_size by simp have len_gt_false: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat) \ False" proof - assume length_lt: "length (rows_to_keep ?step1_mat) < (dim_col ?step1_mat)" have h: "dim_row ?new_mat < (dim_col ?step1_mat)" by (metis Matrix.transpose_transpose index_transpose_mat(3) length_lt length_map mat_of_cols_carrier(3) take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) then show "False" using rank_new_mat by (metis Matrix.transpose_transpose carrier_matI index_transpose_mat(2) nat_less_le not_less_iff_gr_or_eq transpose_rank vec_space.rank_le_nc) qed then have len_gt_eq: "length (rows_to_keep ?step1_mat) \ (dim_col ?step1_mat)" using not_less by blast have len_match: "length (rows_to_keep ?step1_mat) = (dim_col ?step1_mat)" using len_lt_eq len_gt_eq by (simp add: rows_to_keep_def) have mat_match: "matrix_A (get_signs (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) (get_subsets (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs))))) = (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using reduce_system_matrix_match[of p qs signs subsets] assms by blast have f2: "length (get_subsets (take_rows_from_matrix (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) subsets) (rows_to_keep (mat_of_cols (dim_row (matrix_A signs subsets)) (map ((!) (cols (matrix_A signs subsets))) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))), map ((!) signs) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" by (metis (no_types) \dim_col (mat_of_cols (dim_row (matrix_A signs subsets)) (take_indices (cols (matrix_A signs subsets)) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) = length (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ \length (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) = dim_col (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ length_map reduce_mat_cols.simps reduce_system.simps reduction_step.simps reduction_subsets_def reduction_subsets_is_get_subsets take_cols_from_matrix_def take_indices_def) have f3: "\p ps rss nss m. map ((!) rss) (find_nonzeros_from_input_vec (solve_for_lhs p ps nss m)) = get_signs (reduce_system p (ps, m, nss, rss))" by (metis (no_types) reduction_signs_def reduction_signs_is_get_signs take_indices_def) have square_final_mat: "square_mat (get_matrix (reduce_system p (qs, ((matrix_A signs subsets), (subsets, signs)))))" using mat_match assms size_of_mat same_size apply (auto) using f2 f3 by (metis (no_types, lifting) Matrix.transpose_transpose fst_conv get_matrix_def index_transpose_mat(2) len_match length_map mat_of_cols_carrier(2) mat_of_cols_carrier(3) reduce_mat_cols.simps take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have rank_match: "vec_space.rank (dim_row ?new_mat) ?new_mat = dim_row ?new_mat" using len_match rank_new_mat by (simp add: take_cols_from_matrix_def take_indices_def take_rows_and_take_cols) have "invertible_mat ?new_mat" using invertible_det og_mat_size using vec_space.det_rank_iff square_final_mat by (metis dim_col_matrix_A dim_row_matrix_A fst_conv get_matrix_def mat_match rank_match reduce_system.simps reduction_step.simps size_of_mat square_mat.elims(2)) then show ?thesis apply (simp) by (metis fst_conv get_matrix_def) qed subsection "Well def signs preserved when reducing" lemma reduction_doesnt_break_length_signs: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes length_init : "\ x\ set(signs). length x = length qs" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "\x \ set(reduction_signs p qs signs subsets (matrix_A signs subsets)). length x = length qs" proof clarsimp fix x assume x_in_set: "x \ set (reduction_signs p qs signs subsets (matrix_A signs subsets))" have "List.member (reduction_signs p qs signs subsets (matrix_A signs subsets)) x" using x_in_set by (simp add: in_set_member) then have take_ind: "List.member (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x" unfolding reduction_signs_def by auto have find_nz_len: "\y. List.member (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) y \ y < length signs" using size_of_lhs sat_eq inv_mat construct_lhs_matches_solve_for_lhs[of p qs subsets signs] unfolding find_nonzeros_from_input_vec_def by (metis atLeastLessThan_iff filter_is_subset in_set_member set_upt subset_code(1)) then have "\ n < length signs. x = signs ! n" using take_ind by (metis in_set_conv_nth reduction_signs_def take_indices_lem x_in_set) then show "length x = length qs" using length_init take_indices_lem using nth_mem by blast qed subsection "Distinct signs preserved when reducing" lemma reduction_signs_are_distinct: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" assumes distinct_init: "distinct signs" shows "distinct (reduction_signs p qs signs subsets (matrix_A signs subsets))" proof - have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have h1: "distinct (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" unfolding find_nonzeros_from_input_vec_def using distinct_filter using distinct_upt by blast have h2: "(\j. j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ j < length signs)" proof - fix j assume "j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" show "j < length signs" using solve_construct size_of_lhs by (metis \j \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))\ atLeastLessThan_iff filter_is_subset find_nonzeros_from_input_vec_def set_upt subset_iff) qed then show ?thesis unfolding reduction_signs_def unfolding take_indices_def using distinct_init h1 h2 conjugatable_vec_space.distinct_map_nth[where ls = "signs", where inds = "(find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))"] by blast qed subsection "Well def subsets preserved when reducing" lemma reduction_doesnt_break_subsets: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes length_init : "all_list_constr subsets (length qs)" assumes sat_eq: "satisfy_equation p qs subsets signs" assumes inv_mat: "invertible_mat (matrix_A signs subsets)" shows "all_list_constr (reduction_subsets p qs signs subsets (matrix_A signs subsets)) (length qs)" unfolding all_list_constr_def proof clarsimp fix x assume in_red_subsets: "List.member (reduction_subsets p qs signs subsets (matrix_A signs subsets)) x " have solve_construct: "construct_lhs_vector p qs signs = solve_for_lhs p qs subsets (matrix_A signs subsets)" using construct_lhs_matches_solve_for_lhs assms by simp have rows_to_keep_hyp: "\y. y \ set (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) \ y < length subsets" proof clarsimp fix y assume in_set: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))))" have in_set_2: "y \ set (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (construct_lhs_vector p qs signs))))" using in_set solve_construct by simp let ?lhs_vec = "(solve_for_lhs p qs subsets (matrix_A signs subsets))" have h30: "(construct_lhs_vector p qs signs) = ?lhs_vec" using assms construct_lhs_matches_solve_for_lhs by simp have h3a: "\x. List.member (find_nonzeros_from_input_vec ?lhs_vec) x \x < length (signs)" using h30 size_of_lhs unfolding find_nonzeros_from_input_vec_def apply (auto) by (metis atLeastLessThan_iff filter_is_subset member_def set_upt subset_eq) have h3c: "\x. List.member (rows_to_keep (reduce_mat_cols (matrix_A signs subsets) (solve_for_lhs p qs subsets (matrix_A signs subsets)))) x \ x < length (subsets)" proof clarsimp fix x assume x_mem: "List.member (rows_to_keep (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))) x" obtain nn :: "rat list list \ nat list \ nat" where "\x2 x3. (\v4. v4 \ set x3 \ \ v4 < length x2) = (nn x2 x3 \ set x3 \ \ nn x2 x3 < length x2)" by moura then have f4: "nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ set (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ \ nn signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) < length signs \ matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))" using h3a nonzero reduce_system_matrix_signs_helper by auto then have "matrix_A (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets)))) subsets = take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (matrix_A signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (matrix_A signs subsets))))\<^sup>T)))" by (metis h3a in_set_member rows_to_keep_def x_mem) thus "x < length (subsets)" using x_mem unfolding rows_to_keep_def using pivot_positions using dim_row_matrix_A h3a in_set_member nonzero reduce_system_matrix_signs_helper rows_to_keep_def rows_to_keep_lem apply (auto) by (smt (z3) \M_mat (take_indices signs (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets)))) subsets = take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))) \ x \ set (map snd (pivot_positions (gauss_jordan_single (take_cols_from_matrix (M_mat signs subsets) (find_nonzeros_from_input_vec (solve_for_lhs p qs subsets (M_mat signs subsets))))\<^sup>T)))\ dim_row_matrix_A rows_to_keep_def rows_to_keep_lem) qed then show "y < length subsets" using in_set_member apply (auto) using in_set_2 solve_construct by fastforce qed show "list_constr x (length qs)" using in_red_subsets unfolding reduction_subsets_def using take_indices_lem[of _ subsets] rows_to_keep_hyp by (metis all_list_constr_def in_set_conv_nth in_set_member length_init) qed section "Overall Lemmas" lemma combining_to_smash: "combine_systems p (qs1, m1, (sub1, sgn1)) (qs2, m2, (sub2, sgn2)) = smash_systems p qs1 qs2 sub1 sub2 sgn1 sgn2 m1 m2" by simp lemma getter_functions: "calculate_data p qs = (get_matrix (calculate_data p qs), (get_subsets (calculate_data p qs), get_signs (calculate_data p qs))) " unfolding get_matrix_def get_subsets_def get_signs_def by auto subsection "Key properties preserved" subsubsection "Properties preserved when combining and reducing systems" lemma combining_sys_satisfies_properties_helper: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets1 subsets2 :: "nat list list" fixes signs1 signs2 :: "rat list list" fixes matrix1 matrix2:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 subsets1 signs1 matrix1" assumes satisfies_properties_sys2: "satisfies_properties p qs2 subsets2 signs2 matrix2" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_signs (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2))))))) (get_matrix (snd ((combine_systems p (qs1,(matrix1, (subsets1, signs1))) (qs2,(matrix2, (subsets2, signs2)))))))" proof - let ?subsets = "(get_subsets (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?signs = "(get_signs (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" let ?matrix = "(get_matrix (snd (combine_systems p (qs1, matrix1, subsets1, signs1) (qs2, matrix2, subsets2, signs2))))" have h1: "all_list_constr ?subsets (length (qs1 @ qs2))" using well_def_step[of subsets1 qs1 subsets2 qs2] assms by (simp add: nontriv2 get_subsets_def satisfies_properties_def smash_systems_def) have h2: "well_def_signs (length (qs1 @ qs2)) ?signs" using well_def_signs_step[of qs1 qs2 signs1 signs2] using get_signs_def nontriv1 nontriv2 satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 smash_systems_def by auto have h3: "distinct ?signs" using distinct_step[of _ signs1 _ signs2] assms using combine_systems.simps get_signs_def satisfies_properties_def smash_systems_def snd_conv by auto have h4: "satisfy_equation p (qs1 @ qs2) ?subsets ?signs" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] using get_signs_def get_subsets_def satisfies_properties_def smash_systems_def by auto have h5: " invertible_mat ?matrix" using assms inductive_step[of p qs1 qs2 signs1 signs2 subsets1 subsets2] by (metis combining_to_smash fst_conv get_matrix_def kronecker_invertible satisfies_properties_def smash_systems_def snd_conv) have h6: "?matrix = matrix_A ?signs ?subsets" unfolding get_matrix_def combine_systems.simps smash_systems_def get_signs_def get_subsets_def apply simp apply (subst matrix_construction_is_kronecker_product[of subsets1 _ signs1 signs2 subsets2]) apply (metis Ball_set all_list_constr_def in_set_member list_constr_def satisfies_properties_def satisfies_properties_sys1) using satisfies_properties_def satisfies_properties_sys1 well_def_signs_def apply blast using satisfies_properties_def satisfies_properties_sys1 satisfies_properties_sys2 by auto have h7: "set (characterize_consistent_signs_at_roots_copr p (qs1 @ qs2)) \ set (?signs)" using subset_step[of p qs1 signs1 qs2 signs2] assms by (simp add: nonzero get_signs_def satisfies_properties_def smash_systems_def) then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed lemma combining_sys_satisfies_properties: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes satisfies_properties_sys1: "satisfies_properties p qs1 (get_subsets (calculate_data p qs1)) (get_signs (calculate_data p qs1)) (get_matrix (calculate_data p qs1))" assumes satisfies_properties_sys2: "satisfies_properties p qs2 (get_subsets (calculate_data p qs2)) (get_signs (calculate_data p qs2)) (get_matrix (calculate_data p qs2))" shows "satisfies_properties p (qs1@qs2) (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using combining_sys_satisfies_properties_helper by (metis getter_functions nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_sys1 satisfies_properties_sys2) lemma reducing_sys_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" fixes matrix:: "rat mat" assumes nonzero: "p \ 0" assumes nontriv: "length qs > 0" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" assumes satisfies_properties_sys: "satisfies_properties p qs subsets signs matrix" shows "satisfies_properties p qs (get_subsets (reduce_system p (qs,matrix,subsets,signs))) (get_signs (reduce_system p (qs,matrix,subsets,signs))) (get_matrix (reduce_system p (qs,matrix,subsets,signs)))" proof - have h1: " all_list_constr (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (length qs)" using reduction_doesnt_break_subsets assms reduction_subsets_is_get_subsets satisfies_properties_def satisfies_properties_sys by auto have h2: "well_def_signs (length qs) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_length_signs[of signs qs p subsets] assms reduction_signs_is_get_signs satisfies_properties_def well_def_signs_def by auto have h3: "distinct (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_signs_are_distinct[of p qs subsets signs] assms reduction_signs_is_get_signs satisfies_properties_def by auto have h4: "satisfy_equation p qs (get_subsets (reduce_system p (qs, matrix, subsets, signs))) (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_equation_preserved[of p qs signs subsets] assms satisfies_properties_def by auto have h5: "invertible_mat (get_matrix (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_invertibility assms same_size satisfies_properties_def by auto have h6: "get_matrix (reduce_system p (qs, matrix, subsets, signs)) = matrix_A (get_signs (reduce_system p (qs, matrix, subsets, signs))) (get_subsets (reduce_system p (qs, matrix, subsets, signs)))" using reduce_system_matrix_match[of p qs signs subsets] assms satisfies_properties_def by auto have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set (get_signs (reduce_system p (qs, matrix, subsets, signs)))" using reduction_doesnt_break_things_signs[of p qs signs subsets] assms reduction_signs_is_get_signs satisfies_properties_def by auto then show ?thesis unfolding satisfies_properties_def using h1 h2 h3 h4 h5 h6 h7 by blast qed subsubsection "For length 1 qs" lemma length_1_calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof - have h1: "all_list_constr [[],[0]] (length qs)" using len1 unfolding all_list_constr_def list_constr_def apply (auto) by (metis (full_types) length_Cons less_Suc0 list.size(3) list_all_length list_all_simps(2) member_rec(1) member_rec(2) nth_Cons_0) have h2: "well_def_signs (length qs) [[1],[-1]]" unfolding well_def_signs_def using len1 in_set_member by auto have h3: "distinct ([[1],[-1]]::rat list list)" unfolding distinct_def using in_set_member by auto have h4: "satisfy_equation p qs [[],[0]] [[1],[-1]]" using assms base_case_satisfy_equation_alt[of qs p] by auto have h6: "(mat_of_rows_list 2 [[1,1], [1,-1]]::rat mat) = (matrix_A ([[1],[-1]]) ([[],[0]]) :: rat mat)" using mat_base_case by auto then have h5: "invertible_mat (mat_of_rows_list 2 [[1,1], [1,-1]]:: rat mat)" using base_case_invertible_mat by simp have h7: "set (characterize_consistent_signs_at_roots_copr p qs) \ set ([[1],[-1]])" using assms base_case_sgas_alt[of qs p] by simp have base_case_hyp: "satisfies_properties p qs [[],[0]] [[1],[-1]] (mat_of_rows_list 2 [[1,1], [1,-1]])" using h1 h2 h3 h4 h5 h6 h7 using satisfies_properties_def by blast then have key_hyp: "satisfies_properties p qs (get_subsets (reduce_system p (qs,base_case_info))) (get_signs (reduce_system p (qs,base_case_info))) (get_matrix (reduce_system p (qs,base_case_info)))" using reducing_sys_satisfies_properties by (metis base_case_info_def len1 nonzero pairwise_rel_prime nonzero zero_less_one_class.zero_less_one) show ?thesis by (simp add: key_hyp len1) qed subsubsection "For arbitrary qs" lemma append_not_distinct_helper: "(List.member l1 m \ List.member l2 m) \ (distinct (l1@l2) = False)" proof - have h1: "List.member l1 m \ (\ n. n < length l1 \ (nth l1 n) = m)" using member_def nth_find_first by (simp add: member_def in_set_conv_nth) have h2: "\n. n < length l1 \ (nth l1 n) = m \ (nth (l1@l2) n) = m" proof clarsimp fix n assume lt: "n < length l1" assume nth_l1: "m = l1 ! n" show "(l1 @ l2) ! n = l1 ! n" proof (induct l2) case Nil then show ?case by simp next case (Cons a l2) then show ?case by (simp add: lt nth_append) qed qed have h3: "List.member l1 m \ (\n. n < length l1 \ (nth (l1@l2) n) = m)" using h1 h2 by auto have h4: "List.member l2 m \ (\ n. (nth l2 n) = m)" by (meson member_def nth_find_first) have h5: "\n. (nth l2 n) = m \ (nth (l1@l2) (n + length l1)) = m" proof clarsimp fix n assume nth_l2: "m = l2 ! n" show "(l1 @ l2) ! (n + length l1) = l2 ! n" proof (induct l2) case Nil then show ?case by (metis add.commute nth_append_length_plus) next case (Cons a l2) then show ?case by (simp add: nth_append) qed qed have h6: "List.member l2 m \ (\n. (nth (l1@l2) (n + length l1)) = m)" using h4 h5 by blast show ?thesis using h3 h6 by (metis distinct_append equalityI insert_disjoint(1) insert_subset member_def order_refl) qed lemma calculate_data_satisfies_properties: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" shows "(p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" proof (induction "length qs" arbitrary: qs rule: less_induct) case less have len1_h: "length qs = 1 \ ( p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" using length_1_calculate_data_satisfies_properties by blast let ?len = "length qs" let ?q1 = "take (?len div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop (?len div 2) qs" let ?right = "calculate_data p ?q2" let ?comb = "combine_systems p (?q1,?left) (?q2,?right)" let ?red = "reduce_system p ?comb" have h_q1_len: "length qs > 1 \ (length ?q1 > 0)" by auto have h_q2_len: "length qs > 1 \ (length ?q2 > 0)" by auto have h_q1_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q1 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q1 q) \ (List.member qs q)))" by (meson in_set_member in_set_takeD) then show ?thesis by blast qed have h_q2_copr: "(\q. ((List.member qs q) \ (coprime p q))) \ (\q. ((List.member ?q2 q) \ (coprime p q)))" proof - have mem_hyp: "(\q. ((List.member ?q2 q) \ (List.member qs q)))" by (meson in_set_dropD in_set_member) then show ?thesis by blast qed have q1_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q1 (get_subsets (calculate_data p ?q1)) (get_signs (calculate_data p ?q1)) (get_matrix (calculate_data p ?q1))" using less.hyps[of ?q1] h_q1_len h_q1_copr by (auto simp add: Let_def) have q2_help: "length qs > 1 \ length (drop (length qs div 2) qs) < length qs" using h_q1_len by auto then have q2_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p ?q2 (get_subsets (calculate_data p ?q2)) (get_signs (calculate_data p ?q2)) (get_matrix (calculate_data p ?q2))" using less.hyps[of ?q2] h_q2_len h_q2_copr by blast have put_tog: "?q1@?q2 = qs" using append_take_drop_id by blast then have comb_sat_props: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_signs (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))) (get_matrix (snd ((combine_systems p (?q1,calculate_data p ?q1) (?q2,calculate_data p ?q2))))))" using q1_sat_props q2_sat_props combining_sys_satisfies_properties using h_q1_copr h_q1_len h_q2_copr h_q2_len put_tog by metis then have comb_sat: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p (qs) (get_subsets (snd ?comb)) (get_signs (snd ?comb)) (get_matrix (snd ?comb)))" by blast have red_char: "?red = (reduce_system p (qs,(get_matrix (snd ?comb)),(get_subsets (snd ?comb)),(get_signs (snd ?comb))))" using getter_functions by (smt Pair_inject combining_to_smash get_matrix_def get_signs_def get_subsets_def prod.collapse put_tog smash_systems_def) then have "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ (satisfies_properties p qs (get_subsets ?red) (get_signs ?red) (get_matrix ?red))" using reducing_sys_satisfies_properties comb_sat by presburger have len_gt1: "length qs > 1 \ (p \ 0 \ (length qs > 0) \ (\q. ((List.member qs q) \ (coprime p q))) ) \ satisfies_properties p qs (get_subsets (calculate_data p qs)) (get_signs (calculate_data p qs)) (get_matrix (calculate_data p qs))" by (smt \1 < length qs \ p \ 0 \ 0 < length qs \ (\q. List.member qs q \ coprime p q) \ satisfies_properties p qs (get_subsets (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_signs (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs))))) (get_matrix (reduce_system p (combine_systems p (take (length qs div 2) qs, calculate_data p (take (length qs div 2) qs)) (drop (length qs div 2) qs, calculate_data p (drop (length qs div 2) qs)))))\ calculate_data.simps neq0_conv not_less) then show ?case using len1_h len_gt1 by (metis One_nat_def Suc_lessI) qed subsection "Some key results on consistent sign assignments" lemma find_consistent_signs_at_roots_len1: fixes p:: "real poly" fixes qs :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes len1: "length qs = 1" assumes pairwise_rel_prime: "\q. ((List.member qs q) \ (coprime p q))" shows "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - let ?signs = "[[1],[-1]]::rat list list" let ?subsets = "[[],[0]]::nat list list" have mat_help: "matrix_A [[1],[-1]] [[],[0]] = (mat_of_rows_list 2 [[1,1], [1,-1]])" using mat_base_case by auto have well_def_signs: "well_def_signs (length qs) ?signs" unfolding well_def_signs_def using len1 by auto have distinct_signs: "distinct ?signs" unfolding distinct_def by auto have ex_q: "\(q::real poly). qs = [q]" using len1 using length_Suc_conv[of qs 0] by auto then have all_info: "set (characterize_consistent_signs_at_roots_copr p qs) \ set(?signs)" using assms base_case_sgas apply (auto) by (meson in_mono insertE insert_absorb insert_not_empty member_rec(1)) have match: "satisfy_equation p qs ?subsets ?signs" using ex_q base_case_satisfy_equation nonzero pairwise_rel_prime apply (auto) by (simp add: member_rec(1)) have invertible_mat: "invertible_mat (matrix_A ?signs ?subsets)" using inverse_mat_base_case inverse_mat_base_case_2 unfolding invertible_mat_def using mat_base_case by auto have h: "set (get_signs (reduce_system p (qs, ((matrix_A ?signs ?subsets), (?subsets, ?signs))))) = set (characterize_consistent_signs_at_roots_copr p qs)" using nonzero nonzero well_def_signs distinct_signs all_info match invertible_mat reduce_system_sign_conditions[where p = "p", where qs = "qs", where signs = "[[1],[-1]]", where subsets = "[[],[0]]"] by blast then have "set (snd (snd (reduce_system p (qs, ((mat_of_rows_list 2 [[1,1], [1,-1]]), ([[],[0]], [[1],[-1]])))))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding get_signs_def using mat_help by auto then have "set (snd (snd (reduce_system p (qs, base_case_info)))) = set (characterize_consistent_signs_at_roots_copr p qs)" unfolding base_case_info_def by auto then show ?thesis using len1 by (simp add: find_consistent_signs_at_roots_thm) qed lemma smaller_sys_are_good: fixes p:: "real poly" fixes qs1 :: "real poly list" fixes qs2 :: "real poly list" fixes subsets :: "nat list list" fixes signs :: "rat list list" assumes nonzero: "p \ 0" assumes nontriv1: "length qs1 > 0" assumes pairwise_rel_prime1: "\q. ((List.member qs1 q) \ (coprime p q))" assumes nontriv2: "length qs2 > 0" assumes pairwise_rel_prime2: "\q. ((List.member qs2 q) \ (coprime p q))" assumes "set(find_consistent_signs_at_roots p qs1) = set(characterize_consistent_signs_at_roots_copr p qs1)" assumes "set(find_consistent_signs_at_roots p qs2) = set(characterize_consistent_signs_at_roots_copr p qs2)" shows "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" proof - let ?signs = "(get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " let ?subsets = "(get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) " have h0: "satisfies_properties p (qs1@qs2) ?subsets ?signs (get_matrix (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" using calculate_data_satisfies_properties combining_sys_satisfies_properties using nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero by simp then have h1: "set(characterize_consistent_signs_at_roots_copr p (qs1@qs2)) \ set ?signs" unfolding satisfies_properties_def by linarith have h2: "well_def_signs (length (qs1@qs2)) ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h3: "distinct ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h4: "satisfy_equation p (qs1@qs2) ?subsets ?signs" using calculate_data_satisfies_properties using h0 satisfies_properties_def by blast have h5: "invertible_mat (matrix_A ?signs ?subsets) " using calculate_data_satisfies_properties using h0 satisfies_properties_def by auto have h: "set (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) = set(characterize_consistent_signs_at_roots_copr p (qs1@qs2))" using h1 h2 h3 h4 h5 reduction_deletes_bad_sign_conds using nonzero nonzero reduction_signs_def by auto then have h: "set (characterize_consistent_signs_at_roots_copr p (qs1@qs2)) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" unfolding reduction_signs_def get_signs_def by blast have help_h: "reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets) = (take_indices ?signs (find_nonzeros_from_input_vec (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets))))" unfolding reduction_signs_def by auto have clear_signs: "(signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) = (get_signs (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt combining_to_smash get_signs_def getter_functions smash_systems_def snd_conv) have clear_subsets: "(subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = (get_subsets (snd ((combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))))" by (smt Pair_inject combining_to_smash get_subsets_def prod.collapse smash_systems_def) have "well_def_signs (length qs1) (get_signs (calculate_data p qs1))" using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_signs1: "(\j. j \ set (get_signs (calculate_data p qs1)) \ length j = (length qs1))" using well_def_signs_def by blast have "all_list_constr (get_subsets(calculate_data p qs1)) (length qs1) " using calculate_data_satisfies_properties using nontriv1 nonzero pairwise_rel_prime1 nonzero satisfies_properties_def by auto then have well_def_subsets1: "(\l i. l \ set (get_subsets(calculate_data p qs1)) \ i \ set l \ i < (length qs1))" unfolding all_list_constr_def list_constr_def using in_set_member by (metis in_set_conv_nth list_all_length) have extra_matrix_same: "matrix_A (signs_smash (get_signs (calculate_data p qs1)) (get_signs (calculate_data p qs2))) (subsets_smash (length qs1) (get_subsets(calculate_data p qs1)) (get_subsets (calculate_data p qs2))) = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using well_def_signs1 well_def_subsets1 using matrix_construction_is_kronecker_product using calculate_data_satisfies_properties nontriv1 nontriv2 nonzero pairwise_rel_prime1 pairwise_rel_prime2 nonzero satisfies_properties_def by auto then have matrix_same: "matrix_A ?signs ?subsets = kronecker_product (get_matrix (calculate_data p qs1)) (get_matrix (calculate_data p qs2))" using clear_signs clear_subsets by simp have comb_sys_h: "snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2)))) = snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs)))))" unfolding get_signs_def get_subsets_def using matrix_same by (smt combining_to_smash get_signs_def get_subsets_def getter_functions prod.collapse prod.inject smash_systems_def) then have extra_h: " snd(snd(reduce_system p (qs1@qs2, (matrix_A ?signs ?subsets, (?subsets, ?signs))))) = snd(snd(reduction_step (matrix_A ?signs ?subsets) ?signs ?subsets (solve_for_lhs p (qs1@qs2) ?subsets (matrix_A ?signs ?subsets)))) " by simp then have same_h: "set(snd(snd(reduce_system p (combine_systems p (qs1,calculate_data p qs1) (qs2,calculate_data p qs2))))) = set (reduction_signs p (qs1@qs2) ?signs ?subsets (matrix_A ?signs ?subsets ))" using comb_sys_h unfolding reduction_signs_def by (metis get_signs_def help_h reduction_signs_is_get_signs) then show ?thesis using h by blast qed lemma find_consistent_signs_at_roots_1: fixes p:: "real poly" fixes qs :: "real poly list" shows "(p \ 0 \ length qs > 0 \ (\q. ((List.member qs q) \ (coprime p q)))) \ set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" proof (induction "length qs" arbitrary: qs rule: less_induct) case less then show ?case proof clarsimp assume ind_hyp: "(\qsa. length qsa < length qs \ qsa \ [] \ (\q. List.member qsa q \ coprime p q) \ set (find_consistent_signs_at_roots p qsa) = set (characterize_consistent_signs_at_roots_copr p qsa))" assume nonzero: "p \ 0" assume nontriv: "qs \ []" assume copr:" \q. List.member qs q \ coprime p q" let ?len = "length qs" let ?q1 = "take ((?len) div 2) qs" let ?left = "calculate_data p ?q1" let ?q2 = "drop ((?len) div 2) qs" let ?right = "calculate_data p ?q2" have nontriv_q1: "length qs>1 \ length ?q1 > 0" by auto have qs_more_q1: "length qs>1 \ length qs > length ?q1" by auto have pairwise_rel_prime_q1: "\q. ((List.member ?q1 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (take (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (take (length qs div 2) qs) q\ in_set_member in_set_takeD) qed then show "coprime p q" using copr by blast qed have nontriv_q2: "length qs>1 \length ?q2 > 0" by auto have qs_more_q2: "length qs>1 \ length qs > length ?q2" by auto have pairwise_rel_prime_q2: "\q. ((List.member ?q2 q) \ (coprime p q))" proof clarsimp fix q assume mem: "List.member (drop (length qs div 2) qs) q" have "List.member qs q" using mem set_take_subset[where n = "length qs div 2"] proof - show ?thesis by (meson \List.member (drop (length qs div 2) qs) q\ in_set_dropD in_set_member) qed then show "coprime p q" using copr by blast qed have key_h: "set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_len1 : "?len = 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" using find_consistent_signs_at_roots_len1[of p qs] copr nonzero nontriv by (simp add: find_consistent_signs_at_roots_thm) have h_len_gt1 : "?len > 1 \ set (snd (snd (if ?len \ Suc 0 then reduce_system p (qs, base_case_info) else Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h_imp_a: "?len > 1 \ set (snd (snd (reduce_system p (combine_systems p (?q1, ?left) (?q2, ?right))))) = set (characterize_consistent_signs_at_roots_copr p qs)" proof - have h1: "?len > 1 \ set(snd(snd(?left))) = set (characterize_consistent_signs_at_roots_copr p ?q1)" using nontriv_q1 pairwise_rel_prime_q1 ind_hyp[of ?q1] qs_more_q1 by (metis find_consistent_signs_at_roots_thm less_numeral_extra(3) list.size(3)) have h2: "?len > 1 \ set(snd(snd(?right))) = set (characterize_consistent_signs_at_roots_copr p ?q2)" using nontriv_q2 pairwise_rel_prime_q2 ind_hyp[of ?q2] qs_more_q2 by (metis (full_types) find_consistent_signs_at_roots_thm list.size(3) not_less_zero) show ?thesis using nonzero nontriv_q1 pairwise_rel_prime_q1 nontriv_q2 pairwise_rel_prime_q2 h1 h2 smaller_sys_are_good by (metis append_take_drop_id find_consistent_signs_at_roots_thm) qed then have h_imp: "?len > 1 \ set (snd (snd (Let (combine_systems p (?q1, ?left) (?q2, ?right)) (reduce_system p)))) = set (characterize_consistent_signs_at_roots_copr p qs)" by auto then show ?thesis by auto qed show ?thesis using h_len1 h_len_gt1 by (meson \qs \ []\ length_0_conv less_one nat_neq_iff) qed then show "set (find_consistent_signs_at_roots p qs) = set (characterize_consistent_signs_at_roots_copr p qs)" by (smt One_nat_def calculate_data.simps find_consistent_signs_at_roots_thm length_0_conv nontriv) qed qed lemma find_consistent_signs_at_roots_0: fixes p:: "real poly" assumes "p \ 0" shows "set(find_consistent_signs_at_roots p []) = set(characterize_consistent_signs_at_roots_copr p [])" proof - obtain a b c where abc: "reduce_system p ([1], base_case_info) = (a,b,c)" using prod_cases3 by blast have "find_consistent_signs_at_roots p [1] = c" using abc by (simp add: find_consistent_signs_at_roots_thm) have *: "set (find_consistent_signs_at_roots p [1]) = set (characterize_consistent_signs_at_roots_copr p [1])" apply (subst find_consistent_signs_at_roots_1) using assms apply auto by (simp add: member_rec(1) member_rec(2)) have "set(characterize_consistent_signs_at_roots_copr p []) = drop 1 ` set(characterize_consistent_signs_at_roots_copr p [1])" unfolding characterize_consistent_signs_at_roots_copr_def consistent_sign_vec_copr_def apply simp by (smt drop0 drop_Suc_Cons image_cong image_image) thus ?thesis using abc * apply (auto) apply (simp add: find_consistent_signs_at_roots_thm) by (simp add: find_consistent_signs_at_roots_thm) qed lemma find_consistent_signs_at_roots_copr: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots_copr p qs)" by (metis assms find_consistent_signs_at_roots_0 find_consistent_signs_at_roots_1 in_set_member length_greater_0_conv) lemma find_consistent_signs_at_roots: fixes p:: "real poly" fixes qs :: "real poly list" assumes "p \ 0" assumes "\q. q \ set qs \ coprime p q" shows "set(find_consistent_signs_at_roots p qs) = set(characterize_consistent_signs_at_roots p qs)" using assms find_consistent_signs_at_roots_copr csa_list_copr_rel by (simp add: in_set_member) end \ No newline at end of file diff --git a/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Poly_Mod_Finite_Field_Record_Based.thy @@ -1,314 +1,314 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsubsection \Over a Finite Field\ theory Poly_Mod_Finite_Field_Record_Based imports Poly_Mod_Finite_Field Finite_Field_Record_Based Polynomial_Record_Based begin locale arith_ops_record = arith_ops ops + poly_mod m for ops :: "'i arith_ops_record" and m :: int begin definition M_rel_i :: "'i \ int \ bool" where "M_rel_i f F = (arith_ops_record.to_int ops f = M F)" definition Mp_rel_i :: "'i list \ int poly \ bool" where "Mp_rel_i f F = (map (arith_ops_record.to_int ops) f = coeffs (Mp F))" lemma Mp_rel_i_Mp[simp]: "Mp_rel_i f (Mp F) = Mp_rel_i f F" unfolding Mp_rel_i_def by auto lemma Mp_rel_i_Mp_to_int_poly_i: "Mp_rel_i f F \ Mp (to_int_poly_i ops f) = to_int_poly_i ops f" unfolding Mp_rel_i_def to_int_poly_i_def by simp end locale mod_ring_gen = ring_ops ff_ops R for ff_ops :: "'i arith_ops_record" and R :: "'i \ 'a :: nontriv mod_ring \ bool" + fixes p :: int assumes p: "p = int CARD('a)" and of_int: "0 \ x \ x < p \ R (arith_ops_record.of_int ff_ops x) (of_int x)" and to_int: "R y z \ arith_ops_record.to_int ff_ops y = to_int_mod_ring z" and to_int': "0 \ arith_ops_record.to_int ff_ops y \ arith_ops_record.to_int ff_ops y < p \ R y (of_int (arith_ops_record.to_int ff_ops y))" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp sublocale poly_mod_type p "TYPE('a)" by (unfold_locales, rule p) lemma coeffs_to_int_poly: "coeffs (to_int_poly (x :: 'a mod_ring poly)) = map to_int_mod_ring (coeffs x)" by (rule coeffs_map_poly, auto) lemma coeffs_of_int_poly: "coeffs (of_int_poly (Mp x) :: 'a mod_ring poly) = map of_int (coeffs (Mp x))" apply (rule coeffs_map_poly) by (metis M_0 M_M Mp_coeff leading_coeff_0_iff of_int_hom.hom_zero to_int_mod_ring_of_int_M) lemma to_int_poly_i: assumes "poly_rel f g" shows "to_int_poly_i ff_ops f = to_int_poly g" proof - have *: "map (arith_ops_record.to_int ff_ops) f = coeffs (to_int_poly g)" unfolding coeffs_to_int_poly by (rule nth_equalityI, insert assms, auto simp: list_all2_conv_all_nth poly_rel_def to_int) show ?thesis unfolding coeffs_eq_iff to_int_poly_i_def poly_of_list_def coeffs_Poly * strip_while_coeffs.. qed lemma poly_rel_of_int_poly: assumes id: "f' = of_int_poly_i ff_ops (Mp f)" "f'' = of_int_poly (Mp f)" shows "poly_rel f' f''" unfolding id poly_rel_def unfolding list_all2_conv_all_nth coeffs_of_int_poly of_int_poly_i_def length_map by (rule conjI[OF refl], intro allI impI, simp add: nth_coeffs_coeff Mp_coeff M_def, rule of_int, insert p, auto) sublocale arith_ops_record ff_ops p . lemma Mp_rel_iI: "poly_rel f1 f2 \ MP_Rel f3 f2 \ Mp_rel_i f1 f3" unfolding Mp_rel_i_def MP_Rel_def poly_rel_def by (auto simp add: list_all2_conv_all_nth to_int intro: nth_equalityI) lemma M_rel_iI: "R f1 f2 \ M_Rel f3 f2 \ M_rel_i f1 f3" unfolding M_rel_i_def M_Rel_def by (simp add: to_int) lemma M_rel_iI': assumes "R f1 f2" shows "M_rel_i f1 (arith_ops_record.to_int ff_ops f1)" by (rule M_rel_iI[OF assms], simp add: to_int[OF assms] M_Rel_def M_to_int_mod_ring) lemma Mp_rel_iI': assumes "poly_rel f1 f2" shows "Mp_rel_i f1 (to_int_poly_i ff_ops f1)" proof (rule Mp_rel_iI[OF assms], unfold to_int_poly_i[OF assms]) show "MP_Rel (to_int_poly f2) f2" unfolding MP_Rel_def by (simp add: Mp_to_int_poly) qed lemma M_rel_iD: assumes "M_rel_i f1 f3" shows "R f1 (of_int (M f3))" "M_Rel f3 (of_int (M f3))" proof - show "M_Rel f3 (of_int (M f3))" using M_Rel_def to_int_mod_ring_of_int_M by auto from assms show "R f1 (of_int (M f3))" unfolding M_rel_i_def - by (metis int_one_le_iff_zero_less leD linear m1 poly_mod.M_def pos_mod_conj to_int') + by (metis int_one_le_iff_zero_less leD linear m1 poly_mod.M_def pos_mod_sign pos_mod_bound to_int') qed lemma Mp_rel_iD: assumes "Mp_rel_i f1 f3" shows "poly_rel f1 (of_int_poly (Mp f3))" "MP_Rel f3 (of_int_poly (Mp f3))" proof - show Rel: "MP_Rel f3 (of_int_poly (Mp f3))" using MP_Rel_def Mp_Mp Mp_f_representative by auto let ?ti = "arith_ops_record.to_int ff_ops" from assms[unfolded Mp_rel_i_def] have *: "coeffs (Mp f3) = map ?ti f1" by auto { fix x assume "x \ set f1" hence "?ti x \ set (map ?ti f1)" by auto from this[folded *] have "?ti x \ range M" by (metis (no_types, lifting) MP_Rel_def M_to_int_mod_ring Rel coeffs_to_int_poly ex_map_conv range_eqI) hence "?ti x \ 0" "?ti x < p" unfolding M_def using m1 by auto hence "R x (of_int (?ti x))" by (rule to_int') } thus "poly_rel f1 (of_int_poly (Mp f3))" using * unfolding poly_rel_def coeffs_of_int_poly by (auto simp: list_all2_map2 list_all2_same) qed end locale prime_field_gen = field_ops ff_ops R + mod_ring_gen ff_ops R p for ff_ops :: "'i arith_ops_record" and R :: "'i \ 'a :: prime_card mod_ring \ bool" and p :: int begin sublocale poly_mod_prime_type p "TYPE('a)" by (unfold_locales, rule p) end lemma (in mod_ring_locale) mod_ring_rel_of_int: "0 \ x \ x < p \ mod_ring_rel x (of_int x)" unfolding mod_ring_rel_def by (transfer, auto simp: p) context prime_field begin lemma prime_field_finite_field_ops_int: "prime_field_gen (finite_field_ops_int p) mod_ring_rel p" proof - interpret field_ops "finite_field_ops_int p" mod_ring_rel by (rule finite_field_ops_int) show ?thesis by (unfold_locales, rule p, auto simp: finite_field_ops_int_def p mod_ring_rel_def of_int_of_int_mod_ring) qed lemma prime_field_finite_field_ops_integer: "prime_field_gen (finite_field_ops_integer (integer_of_int p)) mod_ring_rel_integer p" proof - interpret field_ops "finite_field_ops_integer (integer_of_int p)" mod_ring_rel_integer by (rule finite_field_ops_integer, simp) have pp: "p = int_of_integer (integer_of_int p)" by auto interpret int: prime_field_gen "finite_field_ops_int p" mod_ring_rel by (rule prime_field_finite_field_ops_int) show ?thesis by (unfold_locales, rule p, auto simp: finite_field_ops_integer_def mod_ring_rel_integer_def[OF pp] urel_integer_def[OF pp] mod_ring_rel_of_int int.to_int[symmetric] finite_field_ops_int_def) qed lemma prime_field_finite_field_ops32: assumes small: "p \ 65535" shows "prime_field_gen (finite_field_ops32 (uint32_of_int p)) mod_ring_rel32 p" proof - let ?pp = "uint32_of_int p" have ppp: "p = int_of_uint32 ?pp" by (subst int_of_uint32_inv, insert small p2, auto) note * = ppp small interpret field_ops "finite_field_ops32 ?pp" mod_ring_rel32 by (rule finite_field_ops32, insert *) interpret int: prime_field_gen "finite_field_ops_int p" mod_ring_rel by (rule prime_field_finite_field_ops_int) show ?thesis proof (unfold_locales, rule p, auto simp: finite_field_ops32_def) fix x assume x: "0 \ x" "x < p" from int.of_int[OF this] have "mod_ring_rel x (of_int x)" by (simp add: finite_field_ops_int_def) thus "mod_ring_rel32 (uint32_of_int x) (of_int x)" unfolding mod_ring_rel32_def[OF *] by (intro exI[of _ x], auto simp: urel32_def[OF *], subst int_of_uint32_inv, insert * x, auto) next fix y z assume "mod_ring_rel32 y z" from this[unfolded mod_ring_rel32_def[OF *]] obtain x where yx: "urel32 y x" and xz: "mod_ring_rel x z" by auto from int.to_int[OF xz] have zx: "to_int_mod_ring z = x" by (simp add: finite_field_ops_int_def) show "int_of_uint32 y = to_int_mod_ring z" unfolding zx using yx unfolding urel32_def[OF *] by simp next fix y show "0 \ int_of_uint32 y \ int_of_uint32 y < p \ mod_ring_rel32 y (of_int (int_of_uint32 y))" unfolding mod_ring_rel32_def[OF *] urel32_def[OF *] by (intro exI[of _ "int_of_uint32 y"], auto simp: mod_ring_rel_of_int) qed qed lemma prime_field_finite_field_ops64: assumes small: "p \ 4294967295" shows "prime_field_gen (finite_field_ops64 (uint64_of_int p)) mod_ring_rel64 p" proof - let ?pp = "uint64_of_int p" have ppp: "p = int_of_uint64 ?pp" by (subst int_of_uint64_inv, insert small p2, auto) note * = ppp small interpret field_ops "finite_field_ops64 ?pp" mod_ring_rel64 by (rule finite_field_ops64, insert *) interpret int: prime_field_gen "finite_field_ops_int p" mod_ring_rel by (rule prime_field_finite_field_ops_int) show ?thesis proof (unfold_locales, rule p, auto simp: finite_field_ops64_def) fix x assume x: "0 \ x" "x < p" from int.of_int[OF this] have "mod_ring_rel x (of_int x)" by (simp add: finite_field_ops_int_def) thus "mod_ring_rel64 (uint64_of_int x) (of_int x)" unfolding mod_ring_rel64_def[OF *] by (intro exI[of _ x], auto simp: urel64_def[OF *], subst int_of_uint64_inv, insert * x, auto) next fix y z assume "mod_ring_rel64 y z" from this[unfolded mod_ring_rel64_def[OF *]] obtain x where yx: "urel64 y x" and xz: "mod_ring_rel x z" by auto from int.to_int[OF xz] have zx: "to_int_mod_ring z = x" by (simp add: finite_field_ops_int_def) show "int_of_uint64 y = to_int_mod_ring z" unfolding zx using yx unfolding urel64_def[OF *] by simp next fix y show "0 \ int_of_uint64 y \ int_of_uint64 y < p \ mod_ring_rel64 y (of_int (int_of_uint64 y))" unfolding mod_ring_rel64_def[OF *] urel64_def[OF *] by (intro exI[of _ "int_of_uint64 y"], auto simp: mod_ring_rel_of_int) qed qed end context mod_ring_locale begin lemma mod_ring_finite_field_ops_int: "mod_ring_gen (finite_field_ops_int p) mod_ring_rel p" proof - interpret ring_ops "finite_field_ops_int p" mod_ring_rel by (rule ring_finite_field_ops_int) show ?thesis by (unfold_locales, rule p, auto simp: finite_field_ops_int_def p mod_ring_rel_def of_int_of_int_mod_ring) qed lemma mod_ring_finite_field_ops_integer: "mod_ring_gen (finite_field_ops_integer (integer_of_int p)) mod_ring_rel_integer p" proof - interpret ring_ops "finite_field_ops_integer (integer_of_int p)" mod_ring_rel_integer by (rule ring_finite_field_ops_integer, simp) have pp: "p = int_of_integer (integer_of_int p)" by auto interpret int: mod_ring_gen "finite_field_ops_int p" mod_ring_rel by (rule mod_ring_finite_field_ops_int) show ?thesis by (unfold_locales, rule p, auto simp: finite_field_ops_integer_def mod_ring_rel_integer_def[OF pp] urel_integer_def[OF pp] mod_ring_rel_of_int int.to_int[symmetric] finite_field_ops_int_def) qed lemma mod_ring_finite_field_ops32: assumes small: "p \ 65535" shows "mod_ring_gen (finite_field_ops32 (uint32_of_int p)) mod_ring_rel32 p" proof - let ?pp = "uint32_of_int p" have ppp: "p = int_of_uint32 ?pp" by (subst int_of_uint32_inv, insert small p2, auto) note * = ppp small interpret ring_ops "finite_field_ops32 ?pp" mod_ring_rel32 by (rule ring_finite_field_ops32, insert *) interpret int: mod_ring_gen "finite_field_ops_int p" mod_ring_rel by (rule mod_ring_finite_field_ops_int) show ?thesis proof (unfold_locales, rule p, auto simp: finite_field_ops32_def) fix x assume x: "0 \ x" "x < p" from int.of_int[OF this] have "mod_ring_rel x (of_int x)" by (simp add: finite_field_ops_int_def) thus "mod_ring_rel32 (uint32_of_int x) (of_int x)" unfolding mod_ring_rel32_def[OF *] by (intro exI[of _ x], auto simp: urel32_def[OF *], subst int_of_uint32_inv, insert * x, auto) next fix y z assume "mod_ring_rel32 y z" from this[unfolded mod_ring_rel32_def[OF *]] obtain x where yx: "urel32 y x" and xz: "mod_ring_rel x z" by auto from int.to_int[OF xz] have zx: "to_int_mod_ring z = x" by (simp add: finite_field_ops_int_def) show "int_of_uint32 y = to_int_mod_ring z" unfolding zx using yx unfolding urel32_def[OF *] by simp next fix y show "0 \ int_of_uint32 y \ int_of_uint32 y < p \ mod_ring_rel32 y (of_int (int_of_uint32 y))" unfolding mod_ring_rel32_def[OF *] urel32_def[OF *] by (intro exI[of _ "int_of_uint32 y"], auto simp: mod_ring_rel_of_int) qed qed lemma mod_ring_finite_field_ops64: assumes small: "p \ 4294967295" shows "mod_ring_gen (finite_field_ops64 (uint64_of_int p)) mod_ring_rel64 p" proof - let ?pp = "uint64_of_int p" have ppp: "p = int_of_uint64 ?pp" by (subst int_of_uint64_inv, insert small p2, auto) note * = ppp small interpret ring_ops "finite_field_ops64 ?pp" mod_ring_rel64 by (rule ring_finite_field_ops64, insert *) interpret int: mod_ring_gen "finite_field_ops_int p" mod_ring_rel by (rule mod_ring_finite_field_ops_int) show ?thesis proof (unfold_locales, rule p, auto simp: finite_field_ops64_def) fix x assume x: "0 \ x" "x < p" from int.of_int[OF this] have "mod_ring_rel x (of_int x)" by (simp add: finite_field_ops_int_def) thus "mod_ring_rel64 (uint64_of_int x) (of_int x)" unfolding mod_ring_rel64_def[OF *] by (intro exI[of _ x], auto simp: urel64_def[OF *], subst int_of_uint64_inv, insert * x, auto) next fix y z assume "mod_ring_rel64 y z" from this[unfolded mod_ring_rel64_def[OF *]] obtain x where yx: "urel64 y x" and xz: "mod_ring_rel x z" by auto from int.to_int[OF xz] have zx: "to_int_mod_ring z = x" by (simp add: finite_field_ops_int_def) show "int_of_uint64 y = to_int_mod_ring z" unfolding zx using yx unfolding urel64_def[OF *] by simp next fix y show "0 \ int_of_uint64 y \ int_of_uint64 y < p \ mod_ring_rel64 y (of_int (int_of_uint64 y))" unfolding mod_ring_rel64_def[OF *] urel64_def[OF *] by (intro exI[of _ "int_of_uint64 y"], auto simp: mod_ring_rel_of_int) qed qed end end diff --git a/thys/CRYSTALS-Kyber/Abs_Qr.thy b/thys/CRYSTALS-Kyber/Abs_Qr.thy --- a/thys/CRYSTALS-Kyber/Abs_Qr.thy +++ b/thys/CRYSTALS-Kyber/Abs_Qr.thy @@ -1,577 +1,577 @@ theory Abs_Qr imports Mod_Plus_Minus Kyber_spec begin context kyber_spec begin section \Re-centered "Norm" Function\ text \We want to show that \abs_infty_q\ is a function induced by the Euclidean norm on the \mod_ring\ using a re-centered representative via \mod+-\. \abs_infty_poly\ is the induced norm by \abs_infty_q\ on polynomials over the polynomial ring over the \mod_ring\. Unfortunately this is not a norm per se, as the homogeneity only holds in inequality, not equality. Still, it fulfils its purpose, since we only need the triangular inequality.\ definition abs_infty_q :: "('a mod_ring) \ int" where "abs_infty_q p = abs ((to_int_mod_ring p) mod+- q)" definition abs_infty_poly :: "'a qr \ int" where "abs_infty_poly p = Max (range (abs_infty_q \ poly.coeff (of_qr p)))" text \Helping lemmas and properties of \Max\, \range\ and \finite\.\ lemma to_int_mod_ring_range: "range (to_int_mod_ring :: 'a mod_ring \ int) = {0 ..< q}" using CARD_a by (simp add: range_to_int_mod_ring) lemma finite_Max: "finite (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa)))" proof - have finite_range: "finite (range (\xa. (poly.coeff (of_qr x) xa)))" using MOST_coeff_eq_0[of "of_qr x"] by auto have "range (\xa. \to_int_mod_ring (poly.coeff (of_qr x) xa) mod+- q\) = (\z. \to_int_mod_ring z mod+- q\) ` range (poly.coeff (of_qr x))" using range_composition[of "(\z. abs (to_int_mod_ring z mod+- q))" "poly.coeff (of_qr x)"] by auto then show ?thesis using finite_range finite_image_set[where f = "(\z. abs (to_int_mod_ring z) mod+- q)"] by (auto simp add: abs_infty_q_def) qed lemma finite_Max_scale: "finite (range (\xa. abs_infty_q (of_int_mod_ring s * poly.coeff (of_qr x) xa)))" proof - have "of_int_mod_ring s * poly.coeff (of_qr x) xa = poly.coeff (of_qr (to_module s * x)) xa" for xa by (metis coeff_smult of_qr_to_qr_smult to_qr_of_qr to_qr_smult_to_module to_module_def) then show ?thesis using finite_Max by presburger qed lemma finite_Max_sum: "finite (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)))" proof - have finite_range: "finite (range (\xa. (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)))" using MOST_coeff_eq_0[of "of_qr x"] by auto have "range (\xa. \to_int_mod_ring (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa) mod+- q\) = (\z. \to_int_mod_ring z mod+- q\) ` range (\xa. poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)" using range_composition[of "(\z. abs (to_int_mod_ring z mod+- q))" "(\xa. poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)"] by auto then show ?thesis using finite_range finite_image_set[where f = "(\z. abs (to_int_mod_ring z) mod+- q)" ] by (auto simp add: abs_infty_q_def) qed lemma finite_range_plus: assumes "finite (range f)" "finite (range g)" shows "finite (range (\x. f x + g x))" proof - have subs: "range (\x. (f x, g x)) \ range f \ range g" by auto have cart: "finite (range f \ range g)" using assms by auto have finite: "finite (range (\x. (f x, g x)))" using rev_finite_subset[OF cart subs] . have "range (\x. f x + g x) = (\(a,b). a+b) ` range (\x. (f x, g x))" using range_composition[of "(\(a,b). a+b)" "(\x. (f x, g x))"] by auto then show ?thesis using finite finite_image_set[where f = "(\(a,b). a+b)"] by auto qed lemma finite_Max_sum': "finite (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa) + abs_infty_q (poly.coeff (of_qr y) xa)))" proof - have finite_range_x: "finite (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa)))" using finite_Max[of x] by auto have finite_range_y: "finite (range (\xa. abs_infty_q (poly.coeff (of_qr y) xa)))" using finite_Max[of y] by auto show ?thesis using finite_range_plus[OF finite_range_x finite_range_y] by auto qed lemma all_impl_Max: assumes "\x. f x \ (a::int)" "finite (range f)" shows "(MAX x. f x) \ a" by (simp add: Max_ge_iff assms(1) assms(2)) lemma Max_scale: "(MAX xa. \s\ * abs_infty_q (poly.coeff (of_qr x) xa)) = \s\ * (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa))" proof - have "(MAX xa. \s\ * abs_infty_q (poly.coeff (of_qr x) xa)) = (Max (range (\xa. \s\ * abs_infty_q (poly.coeff (of_qr x) xa))))" by auto moreover have "\ = (Max ((\a. \s\ * a) ` (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa)))))" by (metis range_composition) moreover have "\ = \s\ * (Max (range (\xa. abs_infty_q (poly.coeff (of_qr x) xa))))" by (subst mono_Max_commute[symmetric]) (auto simp add: finite_Max Rings.mono_mult) moreover have "\ = \s\ * (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa))" by auto ultimately show ?thesis by auto qed lemma Max_mono': assumes "\x. f x \ g x" "finite (range f)" "finite (range g)" shows "(MAX x. f x) \ (MAX x. g x)" using assms by (metis (no_types, lifting) Max_ge_iff Max_in UNIV_not_empty image_is_empty rangeE rangeI) lemma Max_mono_plus: assumes "finite (range (f::_\_::ordered_ab_semigroup_add))" "finite (range g)" shows "(MAX x. f x + g x) \ (MAX x. f x) + (MAX x. g x)" proof - obtain xmax where xmax_def: "f xmax + g xmax = (MAX x. f x + g x)" using finite_range_plus[OF assms] Max_in by fastforce have "(MAX x. f x + g x) = f xmax + g xmax" using xmax_def by auto also have "\ \ (MAX x. f x) + g xmax" using Max_ge[OF assms(1), of "f xmax"] by (auto simp add: add_right_mono[of "f xmax"]) also have "\ \ (MAX x. f x) + (MAX x. g x)" using Max_ge[OF assms(2), of "g xmax"] by (auto simp add: add_left_mono[of "g xmax"]) finally show ?thesis by auto qed text \Show that \abs_infty_q\ is definite, positive and fulfils the triangle inequality.\ lemma abs_infty_q_definite: "abs_infty_q x = 0 \ x = 0" proof (auto simp add: abs_infty_q_def mod_plus_minus_zero'[OF q_gt_zero q_odd]) assume "to_int_mod_ring x mod+- q = 0" then have "to_int_mod_ring x mod q = 0" using mod_plus_minus_zero[of "to_int_mod_ring x" q] by auto then have "to_int_mod_ring x = 0" using to_int_mod_ring_range CARD_a by (metis mod_rangeE range_eqI) then show "x = 0" by force qed lemma abs_infty_q_pos: "abs_infty_q x \ 0" by (auto simp add: abs_infty_q_def) lemma abs_infty_q_minus: "abs_infty_q (- x) = abs_infty_q x" proof (cases "x=0") case True then show ?thesis by auto next case False have minus_x: "to_int_mod_ring (-x) = q - to_int_mod_ring x" proof - have "to_int_mod_ring (-x) = to_int_mod_ring (-x) mod q" by (metis CARD_a Rep_mod_ring_mod to_int_mod_ring.rep_eq) also have "\ = (- to_int_mod_ring x) mod q" by (metis (no_types, opaque_lifting) CARD_a diff_eq_eq mod_add_right_eq plus_mod_ring.rep_eq to_int_mod_ring.rep_eq uminus_add_conv_diff) also have "\ = q - to_int_mod_ring x" proof - have "- to_int_mod_ring x \ {-q<..<0}" using CARD_a range_to_int_mod_ring False by (smt (verit, best) Rep_mod_ring_mod greaterThanLessThan_iff q_gt_zero to_int_mod_ring.rep_eq to_int_mod_ring_hom.eq_iff to_int_mod_ring_hom.hom_zero zmod_trivial_iff) then have "q-to_int_mod_ring x\{0<..to_int_mod_ring (- x) mod+- q\ = \(q - (to_int_mod_ring x)) mod+- q\" by auto also have "\ = \ (- to_int_mod_ring x) mod+- q\" unfolding mod_plus_minus_def by (smt (z3) mod_add_self2) also have "\ = \ - (to_int_mod_ring x mod+- q)\" using neg_mod_plus_minus[OF q_odd q_gt_zero, of "to_int_mod_ring x"] by simp also have "\ = \to_int_mod_ring x mod+- q\" by auto finally show ?thesis unfolding abs_infty_q_def by auto qed lemma to_int_mod_ring_mult: "to_int_mod_ring (a*b) = to_int_mod_ring (a::'a mod_ring) * to_int_mod_ring (b::'a mod_ring) mod q" by (metis (no_types, lifting) CARD_a of_int_hom.hom_mult of_int_mod_ring.rep_eq of_int_mod_ring_to_int_mod_ring of_int_of_int_mod_ring to_int_mod_ring.rep_eq) lemma mod_plus_minus_mult: "s*x mod+- q = (s mod+- q) * (x mod+- q) mod+- q" unfolding mod_plus_minus_def by (smt (verit, del_insts) mod_add_eq mod_diff_left_eq mod_mult_left_eq mod_mult_right_eq) text \Scaling only with inequality not equality! This causes a problem in proof of the Kyber scheme. Needed to add $q\equiv 1 \mod 4$ to change proof.\ lemma mod_plus_minus_leq_mod: "\x mod+- q\ \ \x\" by (smt (verit, best) atLeastAtMost_iff mod_plus_minus_range mod_plus_minus_rangeE' q_gt_zero q_odd) lemma abs_infty_q_scale_pos: assumes "s\0" shows "abs_infty_q ((of_int_mod_ring s :: 'a mod_ring) * x) \ \s\ * (abs_infty_q x)" proof - have "\to_int_mod_ring (of_int_mod_ring s * x) mod+- q\ = \(to_int_mod_ring (of_int_mod_ring s ::'a mod_ring) * to_int_mod_ring x mod q) mod+- q\" using to_int_mod_ring_mult[of "of_int_mod_ring s" x] by simp also have "\ = \(s mod q * to_int_mod_ring x) mod+- q\" by (metis CARD_a mod_add_left_eq mod_plus_minus_def of_int_mod_ring.rep_eq to_int_mod_ring.rep_eq) also have "\ \ \s mod q\ * \to_int_mod_ring x mod+- q\" proof - have "\s mod q * to_int_mod_ring x mod+- q\ = \(s mod q mod+- q) * (to_int_mod_ring x mod+- q) mod+- q\" using mod_plus_minus_mult by auto also have "\ \ \(s mod q mod+- q) * (to_int_mod_ring x mod+- q)\" using mod_plus_minus_leq_mod by blast also have "\ \ \s mod q mod+- q\ * \(to_int_mod_ring x mod+- q)\" by (simp add: abs_mult) also have "\ \ \s mod q\ * \(to_int_mod_ring x mod+- q)\" using mod_plus_minus_leq_mod[of "s mod q"] by (simp add: mult_right_mono) finally show ?thesis by auto qed also have "\ \ \s\ * \to_int_mod_ring x mod+- q\" using assms by (simp add: mult_mono' q_gt_zero zmod_le_nonneg_dividend) finally show ?thesis unfolding abs_infty_q_def by auto qed lemma abs_infty_q_scale_neg: assumes "s<0" shows "abs_infty_q ((of_int_mod_ring s :: 'a mod_ring) * x) \ \s\ * (abs_infty_q x)" using abs_infty_q_minus abs_infty_q_scale_pos by (smt (verit, best) mult_minus_left of_int_minus of_int_of_int_mod_ring) lemma abs_infty_q_scale: "abs_infty_q ((of_int_mod_ring s :: 'a mod_ring) * x) \ \s\ * (abs_infty_q x)" apply (cases "s\0") using abs_infty_q_scale_pos apply presburger using abs_infty_q_scale_neg by force text \Lemmas for porting to \qr\.\ lemma of_qr_mult: "of_qr (a * b) = of_qr a * of_qr b mod qr_poly" by (metis of_qr_to_qr to_qr_mult to_qr_of_qr) lemma of_qr_scale: "of_qr (to_module s * b) = Polynomial.smult (of_int_mod_ring s) (of_qr b)" unfolding to_module_def by (auto simp add: of_qr_mult[of "to_qr [:of_int_mod_ring s:]" "b"] of_qr_to_qr) (simp add: mod_mult_left_eq mod_smult_left of_qr.rep_eq) lemma to_module_mult: "poly.coeff (of_qr (to_module s * a)) x1 = of_int_mod_ring (s) * poly.coeff (of_qr a) x1" using of_qr_scale[of s a] by simp text \Lemmas on \round\ and \floor\.\ lemma odd_round_up: assumes "odd x" shows "round (real_of_int x / 2) = (x+1) div 2" proof - have "round (real_of_int x / 2) = round (real_of_int (x+1) /2)" using assms unfolding round_def by (metis (no_types, opaque_lifting) add.commute add_divide_distrib even_add even_succ_div_2 floor_divide_of_int_eq odd_one of_int_add of_int_hom.hom_one of_int_numeral) also have "\ = (x+1) div 2" by (metis add_divide_distrib calculation floor_divide_of_int_eq of_int_add of_int_hom.hom_one of_int_numeral round_def) finally show ?thesis by blast qed lemma floor_unique: assumes "real_of_int a \ x" "x < a+1" shows "floor x = a" using assms(1) assms(2) by linarith lemma same_floor: assumes "real_of_int a \ x" "real_of_int a \ y" "x < a+1" "y < a+1" shows "floor x = floor y" using assms floor_unique by presburger lemma one_mod_four_round: assumes "x mod 4 = 1" shows "round (real_of_int x / 4) = (x-1) div 4" proof - have leq: "(x-1) div 4 \ real_of_int x / 4 + 1 / 2" using assms by linarith have gr: "real_of_int x / 4 + 1 / 2 < (x-1) div 4 + 1" proof - have "x+2 < 4 * ((x-1) div 4 + 1)" proof - have *: "(x-1) div 4 + 1 = (x+3) div 4" by auto have "4 dvd x + 3" using assms by presburger then have "4 * ((x+3) div 4) = x+3" by (subst dvd_imp_mult_div_cancel_left, auto) then show ?thesis unfolding * by auto qed then show ?thesis by auto qed show "round (real_of_int x / 4) = (x-1) div 4" using floor_unique[OF leq gr] unfolding round_def by auto qed lemma odd_half_floor: assumes "odd x" shows "\real_of_int x / 2\ = (x-1) div 2" using assms by (metis add.commute diff_add_cancel even_add even_succ_div_2 floor_divide_of_int_eq odd_one of_int_numeral) text \Estimation inequality using message bit.\ lemma abs_infty_poly_ineq_pm_1: assumes "\x. poly.coeff (of_qr a) x \ {of_int_mod_ring (-1),1}" shows "abs_infty_poly (to_module (round((real_of_int q)/2)) * a) \ 2 * round (real_of_int q / 4)" proof - let ?x = "to_module (round((real_of_int q)/2)) * a" obtain x1 where x1_def: "poly.coeff (of_qr a) x1 \ {of_int_mod_ring(-1),1}" using assms by auto have "abs_infty_poly (to_module (round((real_of_int q)/2)) * a) \ abs_infty_q (poly.coeff (of_qr (to_module (round (real_of_int q / 2)) * a)) x1)" unfolding abs_infty_poly_def using x1_def by (simp add: finite_Max) also have "abs_infty_q (poly.coeff (of_qr (to_module (round (real_of_int q / 2)) * a)) x1) = abs_infty_q (of_int_mod_ring (round (real_of_int q / 2)) * (poly.coeff (of_qr a) x1))" using to_module_mult[of "round (real_of_int q / 2)" a] by simp also have "\ = abs_infty_q (of_int_mod_ring (round (real_of_int q / 2)))" proof - consider "poly.coeff (of_qr a) x1=1" | "poly.coeff (of_qr a) x1 = of_int_mod_ring (-1)" using x1_def by auto then show ?thesis proof (cases) case 2 then show ?thesis by (metis abs_infty_q_minus mult.right_neutral mult_minus_right of_int_hom.hom_one of_int_minus of_int_of_int_mod_ring) qed (auto) qed also have "\ = \round (real_of_int q / 2) mod+- q\" unfolding abs_infty_q_def using to_int_mod_ring_of_int_mod_ring by (simp add: CARD_a mod_add_left_eq mod_plus_minus_def of_int_mod_ring.rep_eq to_int_mod_ring.rep_eq) also have "\ = \((q + 1) div 2) mod+- q\" using odd_round_up[OF q_odd] by auto also have "\ = \((2 * q) div 2) mod q - (q - 1) div 2\" unfolding mod_plus_minus_def odd_half_floor[OF q_odd] using q_odd by (smt (verit, ccfv_SIG) even_plus_one_iff even_succ_div_2 odd_two_times_div_two_succ) also have "\ = \(q-1) div 2\" using q_odd by (subst nonzero_mult_div_cancel_left[of 2 q], simp) (simp add: abs_div abs_minus_commute) also have "\ = 2 * ((q-1) div 4)" proof - - have "(q-1) div 2 > 0" by (simp add: div_positive_int q_gt_two) + from q_gt_two have "(q-1) div 2 > 0" by simp then have "\(q-1) div 2\ = (q-1) div 2" by auto also have "\ = 2 * ((q-1) div 4)" by (subst div_mult_swap) (use q_mod_4 in \metis dvd_minus_mod\, force) finally show ?thesis by blast qed also have "\ = 2 * round (real_of_int q / 4)" unfolding odd_round_up[OF q_odd] one_mod_four_round[OF q_mod_4] by (simp add: round_def) finally show ?thesis unfolding abs_infty_poly_def by simp qed text \Triangle inequality for \abs_infty_q\.\ lemma abs_infty_q_triangle_ineq: "abs_infty_q (x+y) \ abs_infty_q x + abs_infty_q y" proof - have "to_int_mod_ring (x + y) mod+- q = (to_int_mod_ring x + to_int_mod_ring y) mod q mod+-q" by (simp add: to_int_mod_ring_def CARD_a plus_mod_ring.rep_eq) also have "\ = (to_int_mod_ring x + to_int_mod_ring y) mod+-q" unfolding mod_plus_minus_def by presburger also have "\ = (to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q) mod+- q" unfolding mod_plus_minus_def by (smt (verit, ccfv_threshold) mod_add_eq mod_diff_left_eq) finally have rewrite:"to_int_mod_ring (x + y) mod+- q = (to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q) mod+- q" . then have "\to_int_mod_ring (x + y) mod+- q\ \ \to_int_mod_ring x mod+- q\ + \to_int_mod_ring y mod+- q\" proof (cases "(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q) \ {-\real_of_int q/2\..<\real_of_int q/2\}") case True then have "(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q) mod+- q = to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q" using mod_plus_minus_rangeE[OF True q_odd] by auto then show ?thesis by (simp add: rewrite) next case False then have "\(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q)\ \ \real_of_int q /2\" by auto then have "\(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q)\ \ \(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q) mod+- q\" using mod_plus_minus_range[OF q_gt_zero, of "(to_int_mod_ring x mod+- q + to_int_mod_ring y mod+- q)"] by auto then show ?thesis by (simp add: rewrite) qed then show ?thesis by (auto simp add: abs_infty_q_def mod_plus_minus_def) qed text \Show that \abs_infty_poly\ is definite, positive and fulfils the triangle inequality.\ lemma abs_infty_poly_definite: "abs_infty_poly x = 0 \ x = 0" proof (auto simp add: abs_infty_poly_def abs_infty_q_definite) assume "(MAX xa. abs_infty_q (poly.coeff (of_qr x) xa)) = 0" then have abs_le_zero: "abs_infty_q (poly.coeff (of_qr x) xa) \ 0" for xa using Max_ge[OF finite_Max[of x], of "abs_infty_q (poly.coeff (of_qr x) xa)"] by (auto simp add: Max_ge[OF finite_Max]) have "abs_infty_q (poly.coeff (of_qr x) xa) = 0" for xa using abs_infty_q_pos[of "poly.coeff (of_qr x) xa"] abs_le_zero[of xa] by auto then have "poly.coeff (of_qr x) xa = 0" for xa by (auto simp add: abs_infty_q_definite) then show "x = 0" using leading_coeff_0_iff of_qr_eq_0_iff by blast qed lemma abs_infty_poly_pos: "abs_infty_poly x \ 0" proof (auto simp add: abs_infty_poly_def) have f_ge_zero: "\xa. abs_infty_q (poly.coeff (of_qr x) xa) \ 0" by (auto simp add: abs_infty_q_pos) then show " 0 \ (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa))" using all_impl_Max[OF f_ge_zero finite_Max] by auto qed text \Again, homogeneity is only true for inequality not necessarily equality! Need to add $q\equiv 1\mod 4$ such that proof of crypto scheme works out.\ lemma abs_infty_poly_scale: "abs_infty_poly ((to_module s) * x) \ (abs s) * (abs_infty_poly x)" proof - have fin1: "finite (range (\xa. abs_infty_q (of_int_mod_ring s * poly.coeff (of_qr x) xa)))" using finite_Max_scale by auto have fin2: "finite (range (\xa. \s\ * abs_infty_q (poly.coeff (of_qr x) xa)))" by (metis finite_Max finite_imageI range_composition) have "abs_infty_poly (to_module s * x) = (MAX xa. abs_infty_q ((of_int_mod_ring s) * poly.coeff (of_qr x) xa))" using abs_infty_poly_def to_module_mult by force also have "\ \ (MAX xa. \s\ * abs_infty_q (poly.coeff (of_qr x) xa))" using abs_infty_q_scale fin1 fin2 by (subst Max_mono', auto) also have "\ = \s\ * abs_infty_poly x" unfolding abs_infty_poly_def comp_def using Max_scale by auto finally show ?thesis by blast qed text \Triangle inequality for \abs_infty_poly\.\ lemma abs_infty_poly_triangle_ineq: "abs_infty_poly (x+y) \ abs_infty_poly x + abs_infty_poly y" proof - have "abs_infty_q (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa) \ abs_infty_q (poly.coeff (of_qr x) xa) + abs_infty_q (poly.coeff (of_qr y) xa)" for xa using abs_infty_q_triangle_ineq[of "poly.coeff (of_qr x) xa" "poly.coeff (of_qr y) xa"] by auto then have abs_q_triang: "\xa. abs_infty_q (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa) \ abs_infty_q (poly.coeff (of_qr x) xa) + abs_infty_q (poly.coeff (of_qr y) xa)" by auto have "(MAX xa. abs_infty_q (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)) \ (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa) + abs_infty_q (poly.coeff (of_qr y) xa))" using Max_mono'[OF abs_q_triang finite_Max_sum finite_Max_sum'] by auto also have "\ \ (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa)) + (MAX xb. abs_infty_q (poly.coeff (of_qr y) xb))" using Max_mono_plus[OF finite_Max[of x] finite_Max[of y]] by auto finally have "(MAX xa. abs_infty_q (poly.coeff (of_qr x) xa + poly.coeff (of_qr y) xa)) \ (MAX xa. abs_infty_q (poly.coeff (of_qr x) xa)) + (MAX xb. abs_infty_q (poly.coeff (of_qr y) xb))" by auto then show ?thesis by (auto simp add: abs_infty_poly_def) qed end end \ No newline at end of file diff --git a/thys/Case_Labeling/Examples/Monadic_Language.thy b/thys/Case_Labeling/Examples/Monadic_Language.thy --- a/thys/Case_Labeling/Examples/Monadic_Language.thy +++ b/thys/Case_Labeling/Examples/Monadic_Language.thy @@ -1,310 +1,310 @@ subsection \A labeling VCG for a monadic language\ theory Monadic_Language imports Complex_Main "../Case_Labeling" "HOL-Eisbach.Eisbach" begin ML_file \../util.ML\ ML \ fun vcg_tac nt_rules nt_comb ctxt = let val rules = Named_Theorems.get ctxt nt_rules val comb = Named_Theorems.get ctxt nt_comb in REPEAT_ALL_NEW_FWD ( resolve_tac ctxt rules ORELSE' (resolve_tac ctxt comb THEN' resolve_tac ctxt rules)) end \ text \This language is inspired by the languages used in AutoCorres @{cite greenaway_bridging_2012}\ consts bind :: "'a option \ ('a \ 'b option) \ 'b option" (infixr "|>>" 4) consts return :: "'a \ 'a option" consts while :: "('a \ bool) \ ('a \ bool) \ ('a \ 'a option) \ ('a \ 'a option)" consts valid :: "bool \ 'a option \ ('a \ bool) \ bool" named_theorems vcg named_theorems vcg_comb method_setup vcg = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (FIRSTGOAL (vcg_tac @{named_theorems "vcg"} @{named_theorems "vcg_comb"} ctxt))) \ axiomatization where return[vcg]: "valid (Q x) (return x) Q" and bind[vcg]: "\\x. valid (R x) (c2 x) Q; valid P c1 R\ \ valid P (bind c1 c2) Q" and while[vcg]: "\c. \\x. valid (I x \ b x) (c x) I; \x. I x \ \b x \ Q x\ \ valid (I x) (while b I c x) Q" and cond[vcg]: "\b c1 c2. valid P1 c1 Q \ valid P2 c2 Q \ valid (if b then P1 else P2) (if b then c1 else c2) Q" and case_prod[vcg]: "\P. \\x y. v = (x,y) \ valid (P x y) (B x y) Q\ \ valid (case v of (x,y) \ P x y) (case v of (x,y) \ B x y) Q" and conseq[vcg_comb]: "\valid P' c Q; P \ P'\ \ valid P c Q" text \Labeled rules\ named_theorems vcg_l named_theorems vcg_l_comb named_theorems vcg_elim method_setup vcg_l = \ Scan.succeed (fn ctxt => SIMPLE_METHOD (FIRSTGOAL (vcg_tac @{named_theorems "vcg_l"} @{named_theorems "vcg_l_comb"} ctxt))) \ method vcg_l' = (vcg_l; (elim vcg_elim)?) context begin interpretation Labeling_Syntax . lemma L_return[vcg_l]: "CTXT inp ct (Suc inp) (valid (P x) (return x) P)" unfolding LABEL_simps by (rule return) lemma L_bind[vcg_l]: assumes "\x. CTXT (Suc outp') ((''bind'',outp', [VAR x]) # ct) outp (valid (R x) (c2 x) Q)" assumes "CTXT inp ct outp' (valid P c1 R)" shows "CTXT inp ct outp (valid P (bind c1 c2) Q)" using assms unfolding LABEL_simps by (rule bind) lemma L_while[vcg_l]: fixes inp ct defines "ct' \ \x. (''while'', inp, [VAR x]) # ct" assumes "\x. CTXT (Suc inp) (ct' x) outp' (valid (BIND ''inv_pre'' inp (I x) \ BIND ''lcond'' inp (b x)) (c x) (\x. BIND ''inv_post'' inp (I x)))" assumes "\x. B\''inv_pre'',inp: I x\ \ B\''lcond'',inp: \b x\ \ VC (''post'',outp' , []) (ct' x) (P x)" shows "CTXT inp ct (Suc outp') (valid (I x) (while b I c x) P)" using assms(2-) unfolding LABEL_simps by (rule while) lemma L_cond[vcg_l]: fixes inp ct defines "ct' \ (''if'',inp,[]) # ct" assumes "C\Suc inp, (''then'',inp,[]) # ct',outp: valid P1 c1 Q\" assumes "C\Suc outp, (''else'',outp,[]) # ct',outp': valid P2 c2 Q\" shows "C\inp,ct,outp': valid (if B\''cond'',inp: b\ then B\''then'',inp: P1\ else B\''else'',inp: P2\) (if b then c1 else c2) Q\" using assms(2-) unfolding LABEL_simps by (rule cond) lemma L_case_prod[vcg_l]: assumes "\x y. v = (x,y) \ CTXT inp ct outp (valid (P x y) (B x y) Q)" shows "CTXT inp ct outp (valid (case v of (x,y) \ P x y) (case v of (x,y) \ B x y) Q)" using assms unfolding LABEL_simps by (rule case_prod) lemma L_conseq[vcg_l_comb]: assumes "CTXT (Suc inp) ct outp (valid P' c Q)" assumes "P \ VC (''conseq'',inp,[]) ct P'" shows "CTXT inp ct outp (valid P c Q)" using assms unfolding LABEL_simps by (rule conseq) lemma L_assm_conjE[vcg_elim]: assumes "BIND name inp (P \ Q)" obtains "BIND name inp P" "BIND name inp Q" using assms unfolding LABEL_simps by auto declare conjE[vcg_elim] end lemma dvd_div: fixes a b c :: int assumes "a dvd b" "c dvd b" "coprime a c" shows "a dvd (b div c)" using assms coprime_dvd_mult_left_iff by fastforce lemma divides: " valid (0 < (a :: int)) ( return a |>> (\n. while (\n. even n) (\n. 0 < n \ n dvd a \ (\m. odd m \ m dvd a \ m dvd n)) (\n. return (n div 2)) n ) ) (\r. odd r \ r dvd a \ (\m. odd m \ m dvd a \ m \ r)) " apply vcg - apply (auto simp add: zdvd_imp_le dvd_div div_positive_int elim!: + apply (auto simp add: zdvd_imp_le dvd_div elim!: evenE intro: dvd_mult_right) done lemma L_divides: " valid (0 < (a :: int)) ( return a |>> (\n. while (\n. even n) (\n. 0 < n \ n dvd a \ (\m. odd m \ m dvd a \ m dvd n)) (\n. return (n div 2)) n ) ) (\r. odd r \ r dvd a \ (\m. odd m \ m dvd a \ m \ r)) " apply (rule Initial_Label) apply vcg_l' proof casify print_nested_cases case bind { case (while n) { case post then show ?case by (auto simp: zdvd_imp_le) next case conseq from \0 < n\ \even n\ have "0 < n div 2" by (simp add: pos_imp_zdiv_pos_iff zdvd_imp_le) moreover from \n dvd a\ \even n\ have "n div 2 dvd a" by (metis dvd_div_mult_self dvd_mult_left) moreover { fix m assume "odd m" "m dvd a" then have "m dvd n" using conseq.inv_pre(3) by simp moreover note \even n\ moreover from \odd m\ have "coprime m 2" by (metis dvd_eq_mod_eq_0 invertible_coprime mult_cancel_left2 not_mod_2_eq_1_eq_0) ultimately have "m dvd n div 2" by (rule dvd_div) } ultimately show ?case by auto } next case conseq then show ?case by auto } qed lemma add: " valid True ( while \ \COND:\ (\(r,j). j < (b :: nat)) \ \INV:\ (\(r,j). j \ b \ r = a + j) \ \BODY:\ (\(r,j). return (r + 1, j + 1)) \ \START:\ (a,0) |>> (\(r,_). return r) ) (\r. r = a + b) " by vcg auto lemma mult: " valid True ( while \ \COND:\ (\(r,i). i < (a :: nat)) \ \INV:\ (\(r,i). i \ a \ r = i * b) \ \BODY:\ (\(r,i). while \ \COND:\ (\(r,j). j < b) \ \INV:\ (\(r,j). i < a \ j \ b \ r = i * b + j) \ \BODY:\ (\(r,j). return (r + 1, j + 1)) \ \START:\ (r,0) |>> (\(r,_). return (r, i + 1)) ) \ \START:\ (0,0) |>> (\(r,_). return r) ) (\r. r = a * b) " by vcg auto section \Labeled\ lemma L_mult: " valid True ( while \ \COND:\ (\(r,i). i < (a :: nat)) \ \INV:\ (\(r,i). i \ a \ r = i * b) \ \BODY:\ (\(r,i). while \ \COND:\ (\(r,j). j < b) \ \INV:\ (\(r,j). i < a \ j \ b \ r = i * b + j) \ \BODY:\ (\(r,j). return (r + 1, j + 1)) \ \START:\ (r,0) |>> (\(r,_). return (r, i + 1)) ) \ \START:\ (0,0) |>> (\(r,_). return r) ) (\r. r = a * b) " apply (rule Initial_Label) apply vcg_l' proof casify case while { case while { case post then show ?case by auto next case conseq then show ?case by auto } next case post then show ?case by auto next case conseq then show ?case by auto } next case conseq then show ?case by auto qed lemma L_paths: " valid (path \ []) ( while \ \COND:\ (\(p,r). p \ []) \ \INV:\ (\(p,r). distinct r \ hd (r @ p) = hd path \ last (r @ p) = last path) \ \BODY:\ (\(p,r). return (hd p) |>> (\x. if (r \ [] \ x = hd r) then return [] else (if x \ set r then return (takeWhile (\y. y \ x) r) else return (r)) |>> (\r'. return (tl p, r' @ [x]) ) ) ) \ \START:\ (path, []) |>> (\(_,r). return r) ) (\r. distinct r \ hd r = hd path \ last r = last path) " apply (rule Initial_Label) apply vcg_l' proof casify case conseq then show ?case by auto next case (while p r) { case conseq from conseq have "r = [] \ ?case" by (cases p) auto moreover from conseq have "r \ [] \ hd p = hd r \ ?case" by (cases p) auto moreover { assume A: "r \ []" "hd p \ hd r" have "hd (takeWhile (\y. y \ hd p) r @ hd p # tl p) = hd r" using A by (cases r) auto moreover have "last (takeWhile (\y. y \ hd p) r @ hd p # tl p) = last (r @ p)" using A \p \ []\ by auto moreover have "distinct (takeWhile (\y. y \ hd p) r @ [hd p])" using conseq by (auto dest: set_takeWhileD) ultimately have ?case using A conseq by auto } ultimately show ?case by blast next case post then show ?case by auto } qed end diff --git a/thys/Consensus_Refined/Consensus_Misc.thy b/thys/Consensus_Refined/Consensus_Misc.thy --- a/thys/Consensus_Refined/Consensus_Misc.thy +++ b/thys/Consensus_Refined/Consensus_Misc.thy @@ -1,304 +1,292 @@ (*<*) theory Consensus_Misc imports Main begin (*>*) subsection \Miscellaneous lemmas\ method_setup clarsimp_all = \Method.sections clasimp_modifiers >> K (SIMPLE_METHOD o CHANGED_PROP o PARALLEL_ALLGOALS o clarsimp_tac)\ "clarify simplified, all goals" -lemma div_Suc: - "Suc m div n = (if Suc m mod n = 0 then Suc (m div n) else m div n)" (is "_ = ?rhs") -proof (cases "Suc m mod n = 0") - case True - then show ?thesis - by simp (metis Zero_not_Suc diff_Suc_Suc div_geq minus_mod_eq_mult_div mod_Suc mod_less neq0_conv nonzero_mult_div_cancel_left) -next - case False - then show ?thesis - by simp (metis diff_Suc_Suc div_eq_0_iff minus_mod_eq_mult_div mod_Suc neq0_conv nonzero_mult_div_cancel_left) -qed - definition flip where flip_def: "flip f \ \x y. f y x" lemma option_expand': "\(option = None) = (option' = None); \x y. \option = Some x; option' = Some y\ \ x = y\ \ option = option'" by(rule option.expand, auto) (*********************************) subsection \Argmax\ definition Max_by :: "('a \ 'b :: linorder) \ 'a set \ 'a" where "Max_by f S = (SOME x. x \ S \ f x = Max (f ` S))" lemma Max_by_dest: assumes "finite A" and "A \ {}" shows "Max_by f A \ A \ f (Max_by f A) = Max (f ` A)" (is "?P (Max_by f A)") proof(simp only: Max_by_def some_eq_ex[where P="?P"]) from assms have "finite (f ` A)" "f ` A \ {}" by auto from Max_in[OF this] show "\x. x \ A \ f x = Max (f ` A)" by (metis image_iff) qed lemma Max_by_in: assumes "finite A" and "A \ {}" shows "Max_by f A \ A" using assms by(rule Max_by_dest[THEN conjunct1]) lemma Max_by_ge: assumes "finite A" "x \ A" shows "f x \ f (Max_by f A)" proof- from assms(1) have fin_image: "finite (f ` A)" by simp from assms(2) have non_empty: "A \ {}" by auto have "f x \ f ` A" using assms(2) by simp from Max_ge[OF fin_image this] and Max_by_dest[OF assms(1) non_empty, of f] show ?thesis by(simp) qed lemma finite_UN_D: "finite (\S) \ \A \ S. finite A" by (metis Union_upper finite_subset) lemma Max_by_eqI: assumes fin: "finite A" and "\y. y \ A \ cmp_f y \ cmp_f x" and in_X: "x \ A" and inj: "inj_on cmp_f A" shows "Max_by cmp_f A = x" proof- have in_M: "Max_by cmp_f A \ A" using assms by(auto intro!: Max_by_in) hence "cmp_f (Max_by cmp_f A) \ cmp_f x" using assms by auto also have "cmp_f x \ cmp_f (Max_by cmp_f A)" by (intro Max_by_ge assms) finally show ?thesis using inj in_M in_X by(auto simp add: inj_on_def) qed lemma Max_by_Union_distrib: "\ finite A; A = \S; S \ {}; {} \ S; inj_on cmp_f A \ \ Max_by cmp_f A = Max_by cmp_f (Max_by cmp_f ` S)" proof(rule Max_by_eqI, assumption) fix y assume assms: "finite A" "A = \S" "{} \ S" "y \ A" then obtain B where B_def: "B \ S" "y \ B" by auto hence "cmp_f y \ cmp_f (Max_by cmp_f B)" using assms by (metis Max_by_ge finite_UN_D) also have "... \ cmp_f (Max_by cmp_f (Max_by cmp_f ` S))" using B_def(1) assms by (metis Max_by_ge finite_UnionD finite_imageI imageI) finally show "cmp_f y \ cmp_f (Max_by cmp_f (Max_by cmp_f ` S))" . next assume assms: "finite A" "A = \S" "{} \ S" "S \ {}" hence "Max_by cmp_f ` S \ {}" "finite (Max_by cmp_f ` S)" apply (metis image_is_empty) by (metis assms(1) assms(2) finite_UnionD finite_imageI) hence "Max_by cmp_f (Max_by cmp_f ` S) \ (Max_by cmp_f ` S)" by(blast intro!: Max_by_in) also have "(Max_by cmp_f ` S) \ A" proof - have f1: "\v0 v1. (\ finite v0 \ v0 = {}) \ Max_by (v1::'a \ 'b) v0 \ v0" using Max_by_in by blast have "\v1. v1 \ S \ finite v1" using assms(1) assms(2) finite_UN_D by blast then obtain v2_13 :: "'a set set \ 'a \ 'a set" where "Max_by cmp_f ` S \ \S" using f1 assms(3) by blast thus "Max_by cmp_f ` S \ A" using assms(2) by presburger qed finally show "Max_by cmp_f (Max_by cmp_f ` S) \ A" . qed lemma Max_by_UNION_distrib: "\finite A; A = (\x\S. f x); S \ {}; {} \ f ` S; inj_on cmp_f A\ \ Max_by cmp_f A = Max_by cmp_f (Max_by cmp_f ` (f ` S))" by(force intro!: Max_by_Union_distrib) lemma Max_by_eta: "Max_by f = (\S. (SOME x. x \ S \ f x = Max (f ` S)))" by(auto simp add: Max_by_def) lemma Max_is_Max_by_id: "\ finite S; S \ {} \ \ Max S = Max_by id S " apply(clarsimp simp add: Max_by_def) by (metis (mono_tags, lifting) Max_in someI_ex) definition option_Max_by :: "('a \ 'b :: linorder) \ 'a set \ 'a option" where "option_Max_by cmp_f A \ if A = {} then None else Some (Max_by cmp_f A)" (*********************************) subsection \Function and map graphs\ definition fun_graph where "fun_graph f = {(x, f x)|x. True}" definition map_graph :: "('a, 'b)map \ ('a \ 'b) set" where "map_graph f = {(x, y)|x y. (x, Some y) \ fun_graph f}" lemma map_graph_mem[simp]: "((x, y) \ map_graph f) = (f x = Some y)" by(auto simp add: dom_def map_graph_def fun_graph_def) lemma finite_fun_graph: "finite A \ finite (fun_graph f \ (A \ UNIV))" apply(rule bij_betw_finite[where A=A and f="\x. (x, f x)", THEN iffD1]) by(auto simp add: fun_graph_def bij_betw_def inj_on_def) lemma finite_map_graph: "finite A \ finite (map_graph f \ (A \ UNIV))" by(force simp add: map_graph_def dest!: finite_fun_graph[where f=f] intro!: finite_surj[where A="fun_graph f \ (A \ UNIV)" and f="apsnd the"] ) lemma finite_dom_finite_map_graph: "finite (dom f) \ finite (map_graph f)" apply(simp add: dom_def map_graph_def fun_graph_def) apply(erule finite_surj[where f="\x. (x, the (f x))"]) apply(clarsimp simp add: image_def) by (metis option.sel) (*******************************************************************) lemma ran_map_addD: "x \ ran (m ++ f) \ x \ ran m \ x \ ran f" by(auto simp add: ran_def) subsection \Constant maps\ definition const_map :: "'v \ 'k set \ ('k, 'v)map" where "const_map v S \ (\_. Some v) |` S" lemma const_map_empty[simp]: "const_map v {} = Map.empty" by(auto simp add: const_map_def) lemma const_map_ran[simp]: "x \ ran (const_map v S) = (S \ {} \ x = v)" by(auto simp add: const_map_def ran_def restrict_map_def) lemma const_map_is_None: "(const_map y A x = None) = (x \ A)" by(auto simp add: const_map_def restrict_map_def) lemma const_map_is_Some: "(const_map y A x = Some z) = (z = y \ x \ A)" by(auto simp add: const_map_def restrict_map_def) lemma const_map_in_set: "x \ A \ const_map v A x = Some v" by(auto simp add: const_map_def) lemma const_map_notin_set: "x \ A \ const_map v A x = None" by(auto simp add: const_map_def) lemma dom_const_map: "dom (const_map v S) = S" by(auto simp add: const_map_def) (*******************************************************************) subsection \Votes with maximum timestamps.\ (*******************************************************************) definition vote_set :: "('round \ ('process, 'val)map) \ 'process set \ ('round \ 'val)set" where "vote_set vs Q \ {(r, v)|a r v. ((r, a), v) \ map_graph (case_prod vs) \ a \ Q}" lemma inj_on_fst_vote_set: "inj_on fst (vote_set v_hist {p})" by(clarsimp simp add: inj_on_def vote_set_def) lemma finite_vote_set: assumes "\r'\(r :: nat). v_hist r' = Map.empty" "finite S" shows "finite (vote_set v_hist S)" proof- define vs where "vs = {((r, a), v)|r a v. ((r, a), v) \ map_graph (case_prod v_hist) \ a \ S}" have "vs = (\p\S. ((\(r, v). ((r, p), v)) ` (map_graph (\r. v_hist r p))))" by(auto simp add: map_graph_def fun_graph_def vs_def) also have "... \ (\p\S. (\r. ((r, p), the (v_hist r p))) ` {0.. ('process, 'val)map) \ ('process set, 'round \ 'val)map" where "mru_of_set vs \ \Q. option_Max_by fst (vote_set vs Q)" definition process_mru :: "('round :: linorder \ ('process, 'val)map) \ ('process, 'round \ 'val)map" where "process_mru vs \ \a. mru_of_set vs {a}" lemma process_mru_is_None: "(process_mru v_f a = None) = (vote_set v_f {a} = {})" by(auto simp add: process_mru_def mru_of_set_def option_Max_by_def) lemma process_mru_is_Some: "(process_mru v_f a = Some rv) = (vote_set v_f {a} \ {} \ rv = Max_by fst (vote_set v_f {a}))" by(auto simp add: process_mru_def mru_of_set_def option_Max_by_def) lemma vote_set_upd: "vote_set (v_hist(r := v_f)) {p} = (if p \ dom v_f then insert (r, the (v_f p)) else id ) (if v_hist r p = None then vote_set v_hist {p} else vote_set v_hist {p} - {(r, the (v_hist r p))} ) " by(auto simp add: vote_set_def const_map_is_Some split: if_split_asm) lemma finite_vote_set_upd: " finite (vote_set v_hist {a}) \ finite (vote_set (v_hist(r := v_f)) {a})" by(simp add: vote_set_upd) lemma vote_setD: "rv \ vote_set v_f {a} \ v_f (fst rv) a = Some (snd rv)" by(auto simp add: vote_set_def) lemma process_mru_new_votes: assumes "\r' \ (r :: nat). v_hist r' = Map.empty" shows "process_mru (v_hist(r := v_f)) = (process_mru v_hist ++ (\p. map_option (Pair r) (v_f p)))" proof(rule ext, rule option_expand') fix p show "(process_mru (v_hist(r := v_f)) p = None) = ((process_mru v_hist ++ (\p. map_option (Pair r) (v_f p))) p = None)" using assms by(force simp add: vote_set_def restrict_map_def const_map_is_None process_mru_is_None) next fix p rv rv' assume eqs: "process_mru (v_hist(r := v_f)) p = Some rv" "(process_mru v_hist ++ (\p. map_option (Pair r) (v_f p))) p = Some rv'" moreover have "v_hist (r) p = None" using assms(1) by(auto) ultimately show "rv = rv'" using eqs assms by(auto simp add: map_add_Some_iff const_map_is_Some const_map_is_None process_mru_is_Some vote_set_upd dest!: vote_setD intro!: Max_by_eqI finite_vote_set[OF assms] intro: ccontr inj_on_fst_vote_set) qed end diff --git a/thys/Digit_Expansions/Carries.thy b/thys/Digit_Expansions/Carries.thy --- a/thys/Digit_Expansions/Carries.thy +++ b/thys/Digit_Expansions/Carries.thy @@ -1,321 +1,321 @@ theory Carries imports Bits_Digits begin section \Carries in base-b expansions\ text \Some auxiliary lemmas\ lemma rev_induct[consumes 1, case_names base step]: fixes i k :: nat assumes le: "i \ k" and base: "P k" and step: "\i. i \ k \ P i \ P (i - 1)" shows "P i" proof - have "\i::nat. n = k-i \ i \ k \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = (k - (i + 1))" by arith moreover have k: "i + 1 \ k" using Suc.prems by arith ultimately have "P (i + 1)" by (rule Suc.hyps) from step[OF k this] show ?case by simp qed with le show ?thesis by fast qed subsection \Definition of carry received at position k\ text \When adding two numbers m and n, the carry is \emph{introduced} at position 1 but is \emph{received} at position 2. The function below accounts for the latter case. \begin{center} \begin{verbatim} k: 6 5 4 3 2 1 0 c: 1 - - - - - - - - - - - - m: 1 0 1 0 1 0 n: 1 1 ---------------------- m + n: 0 1 0 1 1 0 0 \end{verbatim} \end{center} \ definition bin_carry :: "nat \ nat \ nat \ nat" where "bin_carry a b k = (a mod 2^k + b mod 2^k) div 2^k" text \Carry in the subtraction of two natural numbers\ definition bin_narry :: "nat \ nat \ nat \ nat" where "bin_narry a b k = (if b mod 2^k > a mod 2^k then 1 else 0)" text \Equivalent definition\ definition bin_narry2 :: "nat \ nat \ nat \ nat" where "bin_narry2 a b k = ((2^k + a mod 2^k - b mod 2^k) div 2^k + 1) mod 2" lemma bin_narry_equiv: "bin_narry a b c = bin_narry2 a b c" apply (auto simp add: bin_narry_def bin_narry2_def) subgoal by (smt add.commute div_less dvd_0_right even_Suc le_add_diff_inverse2 less_add_eq_less mod_greater_zero_iff_not_dvd neq0_conv not_mod2_eq_Suc_0_eq_0 order_le_less zero_less_diff zero_less_numeral zero_less_power) subgoal by (simp add: le_div_geq less_imp_diff_less) done subsection \Properties of carries\ lemma div_sub: fixes a b c :: nat shows "(a - b) div c = (if(a mod c < b mod c) then a div c - b div c - 1 else a div c - b div c)" proof- consider (alb) "ab" by linarith then show ?thesis proof cases case alb then show ?thesis using div_le_mono by auto next case ageb obtain a1 a2 where a1_def: "a1 = a div c" and a2_def: "a2 = a mod c" and a_def: "a=a1*c+a2" using mod_div_decomp by blast obtain b1 b2 where b1_def: "b1 = b div c" and b2_def: "b2 = b mod c" and b_def: "b=b1*c+b2" using mod_div_decomp by blast have a1geb1: "a1\b1" using ageb a1_def b1_def using div_le_mono by blast show ?thesis proof(cases "c=0") assume "c=0" then show ?thesis by simp next assume cneq0: "c \ 0" then show ?thesis proof(cases "a2 < b2") assume a2lb2: "a2 < b2" then show ?thesis proof(cases "a1=b1") case True then show ?thesis using ageb a2lb2 a_def b_def by force next assume "\(a1=b1)" hence a1gb1: "a1>b1" using a1geb1 by auto have boundc: "a2+c-b2 a2 < b2" then have "(a - b) div c = ((a1 - b1) * c + (a2 - b2)) div c" using a1geb1 a_def b_def nat_diff_add_eq1 by auto then show ?thesis using a2geb2 div_add1_eq[of "(a1-b1)*c" "a2-b2" c] by(auto simp add: b2_def a2_def a1_def b1_def less_imp_diff_less) qed qed qed qed lemma dif_digit_formula:"a \ b \ (a - b)\k = (a\k + b\k + bin_narry a b k) mod 2" proof - { presume asm: "a\b" "a mod 2 ^ k < b mod 2 ^ k" then have "Suc((a - b) div 2 ^ k) = a div 2 ^ k - b div 2 ^ k" by (smt Nat.add_diff_assoc One_nat_def Suc_pred add.commute diff_is_0_eq div_add_self1 div_le_mono div_sub mod_add_self1 nat_le_linear neq0_conv plus_1_eq_Suc power_not_zero zero_neq_numeral) then have "(a - b) div 2 ^ k mod 2 = Suc (a div 2 ^ k mod 2 + b div 2 ^ k mod 2) mod 2" by (smt diff_is_0_eq even_Suc even_diff_nat even_iff_mod_2_eq_zero le_less mod_add_eq nat.simps(3) not_mod_2_eq_1_eq_0) } moreover { presume asm2: "\ a mod 2 ^ k < b mod 2 ^ k" "b \ a" then have "(a - b) div 2 ^ k mod 2 = (a div 2 ^ k mod 2 + b div 2 ^ k mod 2) mod 2" using div_sub[of b "2^k" a] div_le_mono even_add even_iff_mod_2_eq_zero le_add_diff_inverse2[of "b div 2 ^ k" "a div 2 ^ k"] mod_mod_trivial[of _ 2] not_less[of "a mod 2 ^ k" "b mod 2 ^ k"] not_mod_2_eq_1_eq_0 div_sub by smt } ultimately show ?thesis by (auto simp add: bin_narry_def nth_bit_def) qed lemma dif_narry_formula: "a\b \ bin_narry a b (k + 1) = (if (a\k < b\k + bin_narry a b k) then 1 else 0)" proof - { presume a1: "a mod (2 * 2 ^ k) < b mod (2 * 2 ^ k)" presume a2: "\ a div 2 ^ k mod 2 < Suc (b div 2 ^ k mod 2)" have f3: "2 ^ k \ (0::nat)" by simp have f4: "a div 2 ^ k mod 2 = 1" using a2 by (meson le_less_trans mod2_eq_if mod_greater_zero_iff_not_dvd not_less zero_less_Suc) then have "b mod (2 * 2 ^ k) = b mod 2 ^ k" using a2 by (metis (no_types) One_nat_def le_simps(3) mod_less_divisor mod_mult2_eq mult.left_neutral neq0_conv not_less semiring_normalization_rules(7)) then have "False" using f4 f3 a1 by (metis One_nat_def add.commute div_add_self1 div_le_mono less_imp_le mod_div_trivial mod_mult2_eq mult.left_neutral not_less plus_1_eq_Suc semiring_normalization_rules(7) zero_less_Suc) } moreover { presume a1: "\ a mod 2 ^ k < b mod 2 ^ k" presume a2: "a mod (2 * 2 ^ k) < b mod (2 * 2 ^ k)" presume a3: "\ a div 2 ^ k mod 2 < b div 2 ^ k mod 2" presume a4: "b \ a" have f6: "a mod 2 ^ Suc k < b mod 2 ^ Suc k" using a2 by simp obtain nn :: "nat \ nat \ nat" where f7: "b + nn a b = a" using a4 le_add_diff_inverse by auto have "(a div 2 ^ k - b div 2 ^ k) div 2 = a div 2 ^ k div 2 - b div 2 ^ k div 2" using a3 div_sub by presburger then have f8: "(a - b) div 2 ^ Suc k = a div 2 ^ Suc k - b div 2 ^ Suc k" using a1 by (metis (no_types) div_mult2_eq div_sub power_Suc power_commutes) have f9: "\n na. Suc (na div n) = (n + na) div n \ 0 = n" by (metis (no_types) add_Suc_right add_cancel_left_right div_add_self1 lessI less_Suc_eq_0_disj less_one zero_neq_one) then have "\n na nb. (na + nb - n) div na = Suc (nb div na) - n div na - 1 \ \ (na + nb) mod na < n mod na \ 0 = na" by (metis (no_types) div_sub) then have f10: "\n na nb. \ (nb::nat) mod na < n mod na \ nb div na - n div na = (na + nb - n) div na \ 0 = na" by (metis (no_types) diff_Suc_Suc diff_commute diff_diff_left mod_add_self1 plus_1_eq_Suc) have "\n. Suc n \ n" by linarith then have "(0::nat) = 2 ^ Suc k" using f10 f9 f8 f7 f6 a4 by (metis add_diff_cancel_left' add_diff_assoc) then have "False" by simp } ultimately show ?thesis - using bin_narry_def apply (auto simp add: nth_bit_def) - subgoal by (smt add_0_left add_less_cancel_left divmod_digit_0(2) le_less_trans mod_less_divisor + apply (auto simp add: nth_bit_def not_less bin_narry_def) + subgoal by (smt add_0_left add_less_cancel_left Divides.divmod_digit_0(2) le_less_trans mod_less_divisor mod_mult2_eq mult_zero_right nat_neq_iff not_less not_mod2_eq_Suc_0_eq_0 semiring_normalization_rules(7) zero_less_numeral zero_less_power) - subgoal by (smt One_nat_def add.left_neutral divmod_digit_0(1) le_less_trans less_imp_le + subgoal by (smt One_nat_def add.left_neutral Divides.divmod_digit_0(1) le_less_trans less_imp_le mod_less_divisor mod_mult2_eq mod_mult_self1_is_0 mult_zero_right not_less not_mod2_eq_Suc_0_eq_0 not_one_le_zero semiring_normalization_rules(7) zero_less_numeral zero_less_power) done qed lemma sum_digit_formula:"(a + b)\k =(a\k + b\k + bin_carry a b k) mod 2" by (simp add: bin_carry_def nth_bit_def) (metis div_add1_eq mod_add_eq) lemma sum_carry_formula:"bin_carry a b (k + 1) =(a\k + b\k + bin_carry a b k) div 2" by (simp add: bin_carry_def nth_bit_def) (smt div_mult2_eq div_mult_self4 mod_mult2_eq power_not_zero semiring_normalization_rules(20) semiring_normalization_rules(34) semiring_normalization_rules(7) zero_neq_numeral) lemma bin_carry_bounded: shows "bin_carry a b k = bin_carry a b k mod 2" proof- have "a mod 2 ^ k < 2 ^k" by simp moreover have "b mod 2 ^ k < 2 ^ k" by simp ultimately have "(a mod 2 ^ k + b mod 2 ^ k) < 2 ^(Suc k)" by (simp add: mult_2 add_strict_mono) then have "(a mod 2 ^ k + b mod 2 ^ k) div 2^k \ 1" using less_mult_imp_div_less by force then have "bin_carry a b k \ 1" using div_le_mono bin_carry_def by fastforce then show ?thesis by auto qed lemma carry_bounded: "bin_carry a b k \ 1" using bin_carry_bounded not_mod_2_eq_1_eq_0[of "bin_carry a b k"] by auto lemma no_carry: "(\r< n.((nth_bit a r) + (nth_bit b r) \ 1)) \ (nth_bit (a + b) n) = (nth_bit a n + nth_bit b n) mod 2" (is "?P \ ?Q n") proof (rule ccontr) assume p: "?P" assume nq: "\?Q n" then obtain k where k1:"\?Q k" and k2:"\r (k-1) + b \ (k-1) < 1" by (smt add.right_neutral c1 calculation diff_le_self k2 leI le_add_diff_inverse2 le_less_trans mod_by_1 mod_if nat_less_le nq one_div_two_eq_zero one_neq_zero p sum_carry_formula) ultimately have "0 = bin_carry a b k" using k2 sum_carry_formula by auto (metis Suc_eq_plus1_left add_diff_inverse_nat less_imp_diff_less mod_0 mod_Suc mod_add_self1 mod_div_trivial mod_less n_not_Suc_n plus_nat.add_0) then show False using c1 by auto qed lemma no_carry_mult_equiv:"(\k. nth_bit a k * nth_bit b k = 0) \ (\k. bin_carry a b k = 0)" (is "?P \ ?Q") proof assume P: ?P { fix k from P have "bin_carry a b k = 0" proof (induction k) case 0 then show ?case using bin_carry_def by (simp) next case (Suc k) then show ?case using sum_carry_formula P by (metis One_nat_def Suc_eq_plus1 add.right_neutral div_less lessI mult_is_0 not_mod_2_eq_0_eq_1 nth_bit_def numeral_2_eq_2 zero_less_Suc) qed } then show ?Q by auto next assume Q: ?Q { fix k from Q have "a \ k * b \ k = 0" proof (induction k) case 0 then show ?case using bin_carry_def nth_bit_def by simp (metis add_self_div_2 not_mod2_eq_Suc_0_eq_0 power_one_right) next case (Suc k) then show ?case using nth_bit_def sum_carry_formula by simp (metis One_nat_def add.right_neutral add_self_div_2 not_mod_2_eq_1_eq_0 power_Suc)+ qed } then show ?P by auto qed (* NEW LEMMAS FROM DIGIT COMPARISON *) lemma carry_digit_impl: "bin_carry a b k \ 0 \ \r r + b \ r = 2" proof(rule ccontr) assume "\ (\r r + b \ r = 2)" hence bound: "\r r + b \ r \ 1" using nth_bit_def by auto assume bk:"bin_carry a b k \ 0" hence base: "bin_carry a b k = 1" using carry_bounded le_less[of "bin_carry a b k" 1] by auto have step: "i \ k \ bin_carry a b i = 1 \ bin_carry a b (i - 1) = 1" for i proof(rule ccontr) assume ik: "i \ k" assume carry: "bin_carry a b i = 1" assume "bin_carry a b (i- 1) \ 1" hence "bin_carry a b (i - 1) = 0" using bin_carry_bounded not_mod_2_eq_1_eq_0[of "bin_carry a b (i - 1)"] by auto then show False using ik carry bound sum_carry_formula[of a b "i-1"] apply simp by (metis Suc_eq_plus1 Suc_pred add_lessD1 bot_nat_0.not_eq_extremum carry diff_is_0_eq' div_le_mono le_eq_less_or_eq less_add_one one_div_two_eq_zero) qed have "\i\k. bin_carry a b i = 1" using rev_induct[where ?P="\c.(bin_carry a b c = 1)"] step base by blast moreover have "bin_carry a b 0 = 0" using bin_carry_def by simp ultimately show False by auto qed end \ No newline at end of file diff --git a/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy b/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy --- a/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy +++ b/thys/Diophantine_Eqns_Lin_Hom/Linear_Diophantine_Equations.thy @@ -1,1162 +1,1161 @@ (* Author: Florian Messner Author: Julian Parsert Author: Jonas Schöpf Author: Christian Sternagel License: LGPL *) section \Homogeneous Linear Diophantine Equations\ theory Linear_Diophantine_Equations imports List_Vector begin (*TODO: move*) lemma lcm_div_le: fixes a :: nat shows "lcm a b div b \ a" by (metis div_by_0 div_le_dividend div_le_mono div_mult_self_is_m lcm_nat_def neq0_conv) (*TODO: move*) lemma lcm_div_le': fixes a :: nat shows "lcm a b div a \ b" by (metis lcm.commute lcm_div_le) (*TODO: move*) lemma lcm_div_gt_0: fixes a :: nat assumes "a > 0" and "b > 0" shows "lcm a b div a > 0" proof - have "lcm a b = (a * b) div (gcd a b)" using lcm_nat_def by blast moreover have "\ > 0" using assms by (metis assms calculation lcm_pos_nat) ultimately show ?thesis using assms - by (metis div_by_0 div_mult2_eq div_positive gcd_le2_nat nat_mult_div_cancel_disj neq0_conv - semiring_normalization_rules(7)) + by simp (metis div_greater_zero_iff div_le_mono2 div_mult_self_is_m gcd_le2_nat not_gr0) qed (*TODO: move*) lemma sum_list_list_update_Suc: assumes "i < length u" shows "sum_list (u[i := Suc (u ! i)]) = Suc (sum_list u)" using assms proof (induct u arbitrary: i) case (Cons x xs) then show ?case by (simp_all split: nat.splits) qed (simp) (*TODO: move*) lemma lessThan_conv: assumes "card A = n" and "\x\A. x < n" shows "A = {.. Given a non-empty list \xs\ of \n\ natural numbers, either there is a value in \xs\ that is \0\ modulo \n\, or there are two values whose moduli coincide. \ lemma list_mod_cases: assumes "length xs = n" and "n > 0" shows "(\x\set xs. x mod n = 0) \ (\ij j \ (xs ! i) mod n = (xs ! j) mod n)" proof - let ?f = "\x. x mod n" and ?X = "set xs" have *: "\x \ ?f ` ?X. x < n" using \n > 0\ by auto consider (eq) "card (?f ` ?X) = card ?X" | (less) "card (?f ` ?X) < card ?X" using antisym_conv2 and card_image_le by blast then show ?thesis proof (cases) case eq show ?thesis proof (cases "distinct xs") assume "distinct xs" with eq have "card (?f ` ?X) = n" using \distinct xs\ by (simp add: assms card_distinct distinct_card) from lessThan_conv [OF this *] and \n > 0\ have "\x\set xs. x mod n = 0" by (metis imageE lessThan_iff) then show ?thesis .. next assume "\ distinct xs" then show ?thesis by (auto) (metis distinct_conv_nth) qed next case less from pigeonhole [OF this] show ?thesis by (auto simp: inj_on_def iff: in_set_conv_nth) qed qed text \ Homogeneous linear Diophantine equations: \a\<^sub>1x\<^sub>1 + \ + a\<^sub>mx\<^sub>m = b\<^sub>1y\<^sub>1 + \ + b\<^sub>ny\<^sub>n\ \ locale hlde_ops = fixes a b :: "nat list" begin abbreviation "m \ length a" abbreviation "n \ length b" \ \The set of all solutions.\ definition Solutions :: "(nat list \ nat list) set" where "Solutions = {(x, y). a \ x = b \ y \ length x = m \ length y = n}" lemma in_Solutions_iff: "(x, y) \ Solutions \ length x = m \ length y = n \ a \ x = b \ y" by (auto simp: Solutions_def) \ \The set of pointwise minimal solutions.\ definition Minimal_Solutions :: "(nat list \ nat list) set" where "Minimal_Solutions = {(x, y) \ Solutions. nonzero x \ \ (\(u, v) \ Solutions. nonzero u \ u @ v <\<^sub>v x @ y)}" definition dij :: "nat \ nat \ nat" where "dij i j = lcm (a ! i) (b ! j) div (a ! i)" definition eij :: "nat \ nat \ nat" where "eij i j = lcm (a ! i) (b ! j) div (b ! j)" definition sij :: "nat \ nat \ (nat list \ nat list)" where "sij i j = ((zeroes m)[i := dij i j], (zeroes n)[j := eij i j])" subsection \Further Constraints on Minimal Solutions\ definition Ej :: "nat \ nat list \ nat set" where "Ej j x = { eij i j - 1 | i. i < length x \ x ! i \ dij i j }" definition Di :: "nat \ nat list \ nat set" where "Di i y = { dij i j - 1 | j. j < length y \ y ! j \ eij i j }" definition Di' :: "nat \ nat list \ nat set" where "Di' i y = { dij i (j + length b - length y) - 1 | j. j < length y \ y ! j \ eij i (j + length b - length y) }" lemma Ej_take_subset: "Ej j (take k x) \ Ej j x" by (auto simp: Ej_def) lemma Di_take_subset: "Di i (take l y) \ Di i y" by (auto simp: Di_def) lemma Di'_drop_subset: "Di' i (drop l y) \ Di' i y" by (auto simp: Di'_def) (metis add.assoc add.commute less_diff_conv) lemma finite_Ej: "finite (Ej j x)" by (rule finite_subset [of _ "(\i. eij i j - 1) ` {0 ..< length x}"]) (auto simp: Ej_def) lemma finite_Di: "finite (Di i y)" by (rule finite_subset [of _ "(\j. dij i j - 1) ` {0 ..< length y}"]) (auto simp: Di_def) lemma finite_Di': "finite (Di' i y)" by (rule finite_subset [of _ "(\j. dij i (j + length b - length y) - 1) ` {0 ..< length y}"]) (auto simp: Di'_def) definition max_y :: "nat list \ nat \ nat" where "max_y x j = (if j < n \ Ej j x \ {} then Min (Ej j x) else Max (set a))" definition max_x :: "nat list \ nat \ nat" where "max_x y i = (if i < m \ Di i y \ {} then Min (Di i y) else Max (set b))" definition max_x' :: "nat list \ nat \ nat" where "max_x' y i = (if i < m \ Di' i y \ {} then Min (Di' i y) else Max (set b))" lemma Min_Ej_le: assumes "j < n" and "e \ Ej j x" and "length x \ m" shows "Min (Ej j x) \ Max (set a)" (is "?m \ _") proof - have "?m \ Ej j x" using assms and finite_Ej and Min_in by blast then obtain i where i: "?m = eij i j - 1" "i < length x" "x ! i \ dij i j" by (auto simp: Ej_def) have "lcm (a ! i) (b ! j) div b ! j \ a ! i" by (rule lcm_div_le) then show ?thesis using i and assms by (auto simp: eij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma Min_Di_le: assumes "i < m" and "e \ Di i y" and "length y \ n" shows "Min (Di i y) \ Max (set b)" (is "?m \ _") proof - have "?m \ Di i y" using assms and finite_Di and Min_in by blast then obtain j where j: "?m = dij i j - 1" "j < length y" "y ! j \ eij i j" by (auto simp: Di_def) have "lcm (a ! i) (b ! j) div a ! i \ b ! j" by (rule lcm_div_le') then show ?thesis using j and assms by (auto simp: dij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma Min_Di'_le: assumes "i < m" and "e \ Di' i y" and "length y \ n" shows "Min (Di' i y) \ Max (set b)" (is "?m \ _") proof - have "?m \ Di' i y" using assms and finite_Di' and Min_in by blast then obtain j where j: "?m = dij i (j + length b - length y) - 1" "j < length y" "y ! j \ eij i (j + length b - length y)" by (auto simp: Di'_def) then have "j + length b - length y < length b" using assms by auto moreover have "lcm (a ! i) (b ! (j + length b - length y)) div a ! i \ b ! (j + length b - length y)" by (rule lcm_div_le') ultimately show ?thesis using j and assms by (auto simp: dij_def) (meson List.finite_set Max_ge diff_le_self le_trans less_le_trans nth_mem) qed lemma max_y_le_take: assumes "length x \ m" shows "max_y x j \ max_y (take k x) j" using assms and Min_Ej_le and Ej_take_subset and Min.subset_imp [OF _ _ finite_Ej] by (auto simp: max_y_def) blast lemma max_x_le_take: assumes "length y \ n" shows "max_x y i \ max_x (take l y) i" using assms and Min_Di_le and Di_take_subset and Min.subset_imp [OF _ _ finite_Di] by (auto simp: max_x_def) blast lemma max_x'_le_drop: assumes "length y \ n" shows "max_x' y i \ max_x' (drop l y) i" using assms and Min_Di'_le and Di'_drop_subset and Min.subset_imp [OF _ _ finite_Di'] by (auto simp: max_x'_def) blast end abbreviation "Solutions \ hlde_ops.Solutions" abbreviation "Minimal_Solutions \ hlde_ops.Minimal_Solutions" abbreviation "dij \ hlde_ops.dij" abbreviation "eij \ hlde_ops.eij" abbreviation "sij \ hlde_ops.sij" declare hlde_ops.dij_def [code] declare hlde_ops.eij_def [code] declare hlde_ops.sij_def [code] lemma Solutions_sym: "(x, y) \ Solutions a b \ (y, x) \ Solutions b a" by (auto simp: hlde_ops.in_Solutions_iff) lemma Minimal_Solutions_imp_Solutions: "(x, y) \ Minimal_Solutions a b \ (x, y) \ Solutions a b" by (auto simp: hlde_ops.Minimal_Solutions_def) lemma Minimal_SolutionsI: assumes "(x, y) \ Solutions a b" and "nonzero x" and "\ (\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y)" shows "(x, y) \ Minimal_Solutions a b" using assms by (auto simp: hlde_ops.Minimal_Solutions_def) lemma minimize_nonzero_solution: assumes "(x, y) \ Solutions a b" and "nonzero x" obtains u and v where "u @ v \\<^sub>v x @ y" and "(u, v) \ Minimal_Solutions a b" using assms proof (induct "x @ y" arbitrary: x y thesis rule: wf_induct [OF wf_less]) case 1 then show ?case proof (cases "(x, y) \ Minimal_Solutions a b") case False then obtain u and v where "nonzero u" and "(u, v) \ Solutions a b" and uv: "u @ v <\<^sub>v x @ y" using 1(3,4) by (auto simp: hlde_ops.Minimal_Solutions_def) with 1(1) [rule_format, of "u @ v" u v] obtain u' and v' where uv': "u' @ v' \\<^sub>v u @ v" and "(u', v') \ Minimal_Solutions a b" by blast moreover have "u' @ v' \\<^sub>v x @ y" using uv and uv' by auto ultimately show ?thesis by (intro 1(2)) qed blast qed lemma Minimal_SolutionsI': assumes "(x, y) \ Solutions a b" and "nonzero x" and "\ (\(u, v) \ Minimal_Solutions a b. u @ v <\<^sub>v x @ y)" shows "(x, y) \ Minimal_Solutions a b" proof (rule Minimal_SolutionsI [OF assms(1,2)]) show "\ (\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y)" proof assume "\(u, v) \ Solutions a b. nonzero u \ u @ v <\<^sub>v x @ y" then obtain u and v where "(u, v) \ Solutions a b" and "nonzero u" and uv: "u @ v <\<^sub>v x @ y" by blast then obtain u' and v' where "(u', v') \ Minimal_Solutions a b" and uv': "u' @ v' \\<^sub>v u @ v" by (blast elim: minimize_nonzero_solution) moreover have "u' @ v' <\<^sub>v x @ y" using uv and uv' by auto ultimately show False using assms by blast qed qed lemma Minimal_Solutions_length: "(x, y) \ Minimal_Solutions a b \ length x = length a \ length y = length b" by (auto simp: hlde_ops.Minimal_Solutions_def hlde_ops.in_Solutions_iff) lemma Minimal_Solutions_gt0: "(x, y) \ Minimal_Solutions a b \ zeroes (length x) <\<^sub>v x" using zero_less by (auto simp: hlde_ops.Minimal_Solutions_def) lemma Minimal_Solutions_sym: assumes "0 \ set a" and "0 \ set b" shows "(xs, ys) \ Minimal_Solutions a b \ (ys, xs) \ Minimal_Solutions b a" using assms by (auto simp: hlde_ops.Minimal_Solutions_def hlde_ops.Solutions_def dest: dotprod_eq_nonzero_iff dest!: less_append_swap [of _ _ ys xs]) locale hlde = hlde_ops + assumes no0: "0 \ set a" "0 \ set b" begin lemma nonzero_Solutions_iff: assumes "(x, y) \ Solutions" shows "nonzero x \ nonzero y" using assms and no0 by (auto simp: in_Solutions_iff dest: dotprod_eq_nonzero_iff) lemma Minimal_Solutions_min: assumes "(x, y) \ Minimal_Solutions" and "u @ v <\<^sub>v x @ y" and "a \ u = b \ v" and [simp]: "length u = m" and non0: "nonzero (u @ v)" shows False proof - have [simp]: "length v = n" using assms by (force dest: less_appendD Minimal_Solutions_length) have "(u, v) \ Solutions" using \a \ u = b \ v\ by (simp add: in_Solutions_iff) moreover from nonzero_Solutions_iff [OF this] have "nonzero u" using non0 by auto ultimately show False using assms by (auto simp: hlde_ops.Minimal_Solutions_def) qed lemma Solutions_snd_not_0: assumes "(x, y) \ Solutions" and "nonzero x" shows "nonzero y" using assms by (metis nonzero_Solutions_iff) end subsection \Pointwise Restricting Solutions\ text \ Constructing the list of \u\ vectors from Huet's proof @{cite "Huet1978"}, satisfying \<^item> \\i y ! i\ and \<^item> \0 < sum_list u \ a\<^sub>k\. \ text \ Given \y\, increment a "previous" \u\ vector at first position starting from \i\ where \u\ is strictly smaller than \y\. If this is not possible, return \u\ unchanged. \ function inc :: "nat list \ nat \ nat list \ nat list" where "inc y i u = (if i < length y then if u ! i < y ! i then u[i := u ! i + 1] else inc y (Suc i) u else u)" by (pat_completeness) auto termination inc by (relation "measure (\(y, i, u). max (length y) (length u) - i)") auto (*inc.simps may cause simplification to loop*) declare inc.simps [simp del] text \ Starting from the 0-vector produce \u\s by iteratively incrementing with respect to \y\. \ definition huets_us :: "nat list \ nat \ nat list" ("\<^bold>u" 1000) where "\<^bold>u y i = ((inc y 0) ^^ Suc i) (zeroes (length y))" lemma huets_us_simps [simp]: "\<^bold>u y 0 = inc y 0 (zeroes (length y))" "\<^bold>u y (Suc i) = inc y 0 (\<^bold>u y i)" by (auto simp: huets_us_def) lemma length_inc [simp]: "length (inc y i u) = length u" by (induct y i u rule: inc.induct) (simp add: inc.simps) lemma length_us [simp]: "length (\<^bold>u y i) = length y" by (induct i) (simp_all) text \ \inc\ produces vectors that are pointwise smaller than \y\ \ lemma inc_le: assumes "length u = length y" and "i < length y" and "u \\<^sub>v y" shows "inc y i u \\<^sub>v y" using assms by (induct y i u rule: inc.induct) (auto simp: inc.simps nth_list_update less_eq_def) lemma us_le: assumes "length y > 0" shows "\<^bold>u y i \\<^sub>v y" using assms by (induct i) (auto simp: inc_le le_length) lemma sum_list_inc_le: "u \\<^sub>v y \ sum_list (inc y i u) \ sum_list y" by (induct y i u rule: inc.induct) (auto simp: inc.simps intro: le_sum_list_mono) lemma sum_list_inc_gt0: assumes "sum_list u > 0" and "length y = length u" shows "sum_list (inc y i u) > 0" using assms proof (induct y i u rule: inc.induct) case (1 y i u) then show ?case by (auto simp add: inc.simps) (meson Suc_neq_Zero gr_zeroI set_update_memI sum_list_eq_0_iff) qed lemma sum_list_inc_gt0': assumes "length u = length y" and "i < length y" and "y ! i > 0" and "j \ i" shows "sum_list (inc y j u) > 0" using assms proof (induct y j u rule: inc.induct) case (1 y i u) then show ?case by (auto simp: inc.simps [of y i] sum_list_update) (metis elem_le_sum_list le_antisym le_zero_eq neq0_conv not_less_eq_eq sum_list_inc_gt0) qed lemma sum_list_us_gt0: assumes "sum_list y \ 0" shows "0 < sum_list (\<^bold>u y i)" using assms by (induct i) (auto simp: in_set_conv_nth sum_list_inc_gt0' sum_list_inc_gt0) lemma sum_list_inc_le': assumes "length u = length y" shows "sum_list (inc y i u) \ sum_list u + 1" using assms by (induct y i u rule: inc.induct) (auto simp: inc.simps sum_list_update) lemma sum_list_us_le: "sum_list (\<^bold>u y i) \ i + 1" proof (induct i) case 0 then show ?case by (auto simp: sum_list_update) (metis Suc_eq_plus1 in_set_replicate length_replicate sum_list_eq_0_iff sum_list_inc_le') next case (Suc i) then show ?case by auto (metis Suc_le_mono add.commute le_trans length_us plus_1_eq_Suc sum_list_inc_le') qed lemma sum_list_us_bounded: assumes "i < k" shows "sum_list (\<^bold>u y i) \ k" using assms and sum_list_us_le [of y i] by force lemma sum_list_inc_eq_sum_list_Suc: assumes "length u = length y" and "i < length y" and "\j\i. j < length y \ u ! j < y ! j" shows "sum_list (inc y i u) = Suc (sum_list u)" using assms by (induct y i u rule: inc.induct) (metis inc.simps Suc_eq_plus1 Suc_leI antisym_conv2 leD sum_list_list_update_Suc) lemma sum_list_us_eq: assumes "i < sum_list y" shows "sum_list (\<^bold>u y i) = i + 1" using assms proof (induct i) case (Suc i) then show ?case by (auto) (metis (no_types, lifting) Suc_eq_plus1 gr_implies_not0 length_pos_if_in_set length_us less_Suc_eq_le less_imp_le_nat antisym_conv2 not_less_eq_eq sum_list_eq_0_iff sum_list_inc_eq_sum_list_Suc sum_list_less_diff_Ex us_le) qed (metis Suc_eq_plus1 Suc_leI antisym_conv gr_implies_not0 sum_list_us_gt0 sum_list_us_le) lemma inc_ge: "length u = length y \ u \\<^sub>v inc y i u" by (induct y i u rule: inc.induct) (auto simp: inc.simps nth_list_update less_eq_def) lemma us_le_mono: assumes "i < j" shows "\<^bold>u y i \\<^sub>v \<^bold>u y j" using assms proof (induct "j - i" arbitrary: j i) case (Suc n) then show ?case by (simp add: Suc.prems inc_ge order.strict_implies_order order_vec.lift_Suc_mono_le) qed simp lemma us_mono: assumes "i < j" and "j < sum_list y" shows "\<^bold>u y i <\<^sub>v \<^bold>u y j" proof - let ?u = "\<^bold>u y i" and ?v = "\<^bold>u y j" have "?u \\<^sub>v ?v" using us_le_mono [OF \i < j\] by simp moreover have "sum_list ?u < sum_list ?v" using assms by (auto simp: sum_list_us_eq) ultimately show ?thesis by (intro le_sum_list_less) (auto simp: less_eq_def) qed context hlde begin lemma max_coeff_bound_right: assumes "(xs, ys) \ Minimal_Solutions" shows "\x \ set xs. x \ maxne0 ys b" (is "\x\set xs. x \ ?m") proof (rule ccontr) assume "\ ?thesis" then obtain k where k_def: "k < length xs \ \ (xs ! k \ ?m)" by (metis in_set_conv_nth) have sol: "(xs, ys) \ Solutions" using assms Minimal_Solutions_def by auto then have len: "m = length xs" by (simp add: in_Solutions_iff) have max_suml: "?m * sum_list ys \ b \ ys" using maxne0_times_sum_list_gt_dotprod sol by (auto simp: in_Solutions_iff) then have is_sol: "b \ ys = a \ xs" using sol by (auto simp: in_Solutions_iff) then have a_ge_ak: "a \ xs \ a ! k * xs ! k" using dotprod_pointwise_le k_def len by auto then have ak_gt_max: "a ! k * xs ! k > a ! k * ?m" using no0 in_set_conv_nth k_def len by fastforce then have sl_ys_g_ak: "sum_list ys > a ! k" by (metis a_ge_ak is_sol less_le_trans max_suml mult.commute mult_le_mono1 not_le) define Seq where Seq_def: "Seq = map (\<^bold>u ys) [0 ..< a ! k]" have ak_n0: "a ! k \ 0" using \a ! k * ?m < a ! k * xs ! k\ by auto have "zeroes (length ys) <\<^sub>v ys" by (intro zero_less) (metis gr_implies_not0 nonzero_iff sl_ys_g_ak sum_list_eq_0_iff) then have "length Seq > 0" using ak_n0 Seq_def by auto have u_in_nton: "\u \ set Seq. length u = length ys" by (simp add: Seq_def) have prop_3: "\u \ set Seq. u \\<^sub>v ys" proof - have "length ys > 0" using sl_ys_g_ak by auto then show ?thesis using us_le [of ys ] less_eq_def Seq_def by (simp) qed have prop_4_1: "\u \ set Seq. sum_list u > 0" by (metis Seq_def sl_ys_g_ak gr_implies_not_zero imageE set_map sum_list_us_gt0) have prop_4_2: "\u \ set Seq. sum_list u \ a ! k" by (simp add: Seq_def sum_list_us_bounded) have prop_5: "\u. length u = length ys \ u \\<^sub>v ys \ sum_list u > 0 \ sum_list u \ a ! k" using \0 < length Seq\ nth_mem prop_3 prop_4_1 prop_4_2 u_in_nton by blast define Us where "Us = {u. length u = length ys \ u \\<^sub>v ys \ sum_list u > 0 \ sum_list u \ a ! k}" have "\u \ Us. b \ u mod a ! k = 0" proof (rule ccontr) assume neg_th: "\ ?thesis" define Seq_p where "Seq_p = map (dotprod b) Seq" have "length Seq = a ! k" by (simp add: Seq_def) then consider (eq_0) "(\x\set Seq_p. x mod (a ! k) = 0)" | (not_0) "(\ij j \ (Seq_p ! i) mod (a!k) = (Seq_p ! j) mod (a!k))" using list_mod_cases[of Seq_p] Seq_p_def ak_n0 by auto force then show False proof (cases) case eq_0 have "\u \ set Seq. b \ u mod a ! k = 0" using Seq_p_def eq_0 by auto then show False by (metis (mono_tags, lifting) Us_def mem_Collect_eq neg_th prop_3 prop_4_1 prop_4_2 u_in_nton) next case not_0 obtain i and j where i_j: "i j" " Seq_p ! i mod a ! k = Seq_p ! j mod a ! k" using not_0 by blast define v where v_def: "v = Seq!i" define w where w_def: "w = Seq!j" have mod_eq: "b \ v mod a!k = b \ w mod a!k" using Seq_p_def i_j w_def v_def i_j by auto have "v <\<^sub>v w \ w <\<^sub>v v" using \i \ j\ and i_j proof (cases "i < j") case True then show ?thesis using Seq_p_def sl_ys_g_ak i_j(2) local.Seq_def us_mono v_def w_def by auto next case False then show ?thesis using Seq_p_def sl_ys_g_ak \i \ j\ i_j(1) local.Seq_def us_mono v_def w_def by auto qed then show False proof assume ass: "v <\<^sub>v w" define u where u_def: "u = w -\<^sub>v v" have "w \\<^sub>v ys" using Seq_p_def w_def i_j(2) prop_3 by force then have prop_3: "less_eq u ys" using vdiff_le ass less_eq_def order_vec.less_imp_le u_def by auto have prop_4_1: "sum_list u > 0" using le_sum_list_mono [of v w] ass u_def sum_list_vdiff_distr [of v w] by (simp add: less_vec_sum_list_less) have prop_4_2: "sum_list u \ a ! k" proof - have "u \\<^sub>v w" using u_def using ass less_eq_def order_vec.less_imp_le vdiff_le by auto then show ?thesis by (metis Seq_p_def i_j(2) length_map le_sum_list_mono less_le_trans not_le nth_mem prop_4_2 w_def) qed have "b \ u mod a ! k = 0" by (metis (mono_tags, lifting) in_Solutions_iff \w \\<^sub>v ys\ u_def ass no0(2) less_eq_def mem_Collect_eq mod_eq mods_with_vec_2 prod.simps(2) sol) then show False using neg_th by (metis (mono_tags, lifting) Us_def less_eq_def mem_Collect_eq prop_3 prop_4_1 prop_4_2) next assume ass: "w <\<^sub>v v" define u where u_def: "u = v -\<^sub>v w" have "v \\<^sub>v ys" using Seq_p_def v_def i_j(1) prop_3 by force then have prop_3: "u \\<^sub>v ys" using vdiff_le ass less_eq_def order_vec.less_imp_le u_def by auto have prop_4_1: "sum_list u > 0" using le_sum_list_mono [of w v] sum_list_vdiff_distr [of w v] \u \ v -\<^sub>v w\ ass less_vec_sum_list_less by auto have prop_4_2: "sum_list u \ a!k" proof - have "u \\<^sub>v v" using u_def using ass less_eq_def order_vec.less_imp_le vdiff_le by auto then show ?thesis by (metis Seq_p_def i_j(1) le_neq_implies_less length_map less_imp_le_nat less_le_trans nth_mem prop_4_2 le_sum_list_mono v_def) qed have "b \ u mod a ! k = 0" by (metis (mono_tags, lifting) in_Solutions_iff \v \\<^sub>v ys\ u_def ass no0(2) less_eq_def mem_Collect_eq mod_eq mods_with_vec_2 prod.simps(2) sol) then show False by (metis (mono_tags, lifting) neg_th Us_def less_eq_def mem_Collect_eq prop_3 prop_4_1 prop_4_2) qed qed qed then obtain u where u3_4: "u \\<^sub>v ys" "sum_list u > 0" "sum_list u \ a ! k" " b \ u mod (a ! k) = 0" "length u = length ys" unfolding Us_def by auto have u_b_len: "length u = n" using less_eq_def u3_4 in_Solutions_iff sol by simp have "b \ u \ maxne0 u b * sum_list u" by (simp add: maxne0_times_sum_list_gt_dotprod u_b_len) also have "... \ ?m * a ! k" by (intro mult_le_mono) (simp_all add: u3_4 maxne0_mono) also have "... < a ! k * xs ! k" using ak_gt_max by auto then obtain zk where zk: "b \ u = zk * a ! k" using u3_4(4) by auto have "length xs > k" by (simp add: k_def) have "zk \ 0" proof - have "\e \ set u. e \ 0" using u3_4 by (metis neq0_conv sum_list_eq_0_iff) then have "b \ u > 0" using assms no0 u3_4 unfolding dotprod_gt0_iff[OF u_b_len [symmetric]] by (fastforce simp add: in_set_conv_nth u_b_len) then have "a ! k > 0" using \a ! k \ 0\ by blast then show ?thesis using \0 < b \ u\ zk by auto qed define z where z_def: "z = (zeroes (length xs))[k := zk]" then have zk_zk: "z ! k = zk" by (auto simp add: \k < length xs\) have "length z = length xs" using assms z_def \k < length xs\ by auto then have bu_eq_akzk: "b \ u = a ! k * z ! k" by (simp add: \b \ u = zk * a ! k\ zk_zk) then have "z!k < xs!k" using ak_gt_max calculation by auto then have z_less_xs: "z <\<^sub>v xs" by (auto simp add: z_def) (metis \k < length xs\ le0 le_list_update less_def less_imp_le order_vec.dual_order.antisym nat_neq_iff z_def zk_zk) then have "z @ u <\<^sub>v xs @ ys" by (intro less_append) (auto simp add: u3_4(1) z_less_xs) moreover have "(z, u) \ Solutions" by (auto simp add: bu_eq_akzk in_Solutions_iff z_def u_b_len \k < length xs\ len) moreover have "nonzero z" using \length z = length xs\ and \zk \ 0\ and k_def and zk_zk by (auto simp: nonzero_iff) ultimately show False using assms by (auto simp: Minimal_Solutions_def) qed text \Proof of Lemma 1 of Huet's paper.\ lemma max_coeff_bound: assumes "(xs, ys) \ Minimal_Solutions" shows "(\x \ set xs. x \ maxne0 ys b) \ (\y \ set ys. y \ maxne0 xs a)" proof - interpret ba: hlde b a by (standard) (auto simp: no0) show ?thesis using assms and Minimal_Solutions_sym [OF no0, of xs ys] by (auto simp: max_coeff_bound_right ba.max_coeff_bound_right) qed lemma max_coeff_bound': assumes "(x, y) \ Minimal_Solutions" shows "\i Max (set b)" and "\j Max (set a)" using max_coeff_bound [OF assms] and maxne0_le_Max by auto (metis le_eq_less_or_eq less_le_trans nth_mem)+ lemma Minimal_Solutions_alt_def: "Minimal_Solutions = {(x, y)\Solutions. (x, y) \ (zeroes m, zeroes n) \ x \\<^sub>v replicate m (Max (set b)) \ y \\<^sub>v replicate n (Max (set a)) \ \ (\(u, v)\Solutions. nonzero u \ u @ v <\<^sub>v x @ y)}" by (auto simp: not_nonzero_iff Minimal_Solutions_imp_Solutions less_eq_def Minimal_Solutions_length max_coeff_bound' intro!: Minimal_SolutionsI' dest: Minimal_Solutions_gt0) (auto simp: Minimal_Solutions_def nonzero_Solutions_iff not_nonzero_iff) subsection \Special Solutions\ definition Special_Solutions :: "(nat list \ nat list) set" where "Special_Solutions = {sij i j | i j. i < m \ j < n}" lemma dij_neq_0: assumes "i < m" and "j < n" shows "dij i j \ 0" proof - have "a ! i > 0" and "b ! j > 0" using assms and no0 by (simp_all add: in_set_conv_nth) then have "dij i j > 0" using lcm_div_gt_0 [of "a ! i" "b ! j"] by (simp add: dij_def) then show ?thesis by simp qed lemma eij_neq_0: assumes "i < m" and "j < n" shows "eij i j \ 0" proof - have "a ! i > 0" and "b ! j > 0" using assms and no0 by (simp_all add: in_set_conv_nth) then have "eij i j > 0" using lcm_div_gt_0[of "b ! j" "a ! i"] by (simp add: eij_def lcm.commute) then show ?thesis by simp qed lemma Special_Solutions_in_Solutions: "x \ Special_Solutions \ x \ Solutions" by (auto simp: in_Solutions_iff Special_Solutions_def sij_def dij_def eij_def) lemma Special_Solutions_in_Minimal_Solutions: assumes "(x, y) \ Special_Solutions" shows "(x, y) \ Minimal_Solutions" proof (intro Minimal_SolutionsI') show "(x, y) \ Solutions" by (fact Special_Solutions_in_Solutions [OF assms]) then have [simp]: "length x = m" "length y = n" by (auto simp: in_Solutions_iff) show "nonzero x" using assms and dij_neq_0 by (auto simp: Special_Solutions_def sij_def nonzero_iff) (metis length_replicate set_update_memI) show "\ (\(u, v)\Minimal_Solutions. u @ v <\<^sub>v x @ y)" proof assume "\(u, v)\Minimal_Solutions. u @ v <\<^sub>v x @ y" then obtain u and v where uv: "(u, v) \ Minimal_Solutions" and "u @ v <\<^sub>v x @ y" and [simp]: "length u = m" "length v = n" and "nonzero u" by (auto simp: Minimal_Solutions_def in_Solutions_iff) then consider "u <\<^sub>v x" and "v \\<^sub>v y" | "v <\<^sub>v y" and "u \\<^sub>v x" by (auto elim: less_append_cases) then show False proof (cases) case 1 then obtain i and j where ij: "i < m" "j < n" and less_dij: "u ! i < dij i j" and "u \\<^sub>v (zeroes m)[i := dij i j]" and "v \\<^sub>v (zeroes n)[j := eij i j]" using assms by (auto simp: Special_Solutions_def sij_def unit_less) then have u: "u = (zeroes m)[i := u ! i]" and v: "v = (zeroes n)[j := v ! j]" by (auto simp: less_eq_def list_eq_iff_nth_eq) (metis le_zero_eq length_list_update length_replicate rep_upd_unit)+ then have "u ! i > 0" using \nonzero u\ and ij by (metis gr_implies_not0 neq0_conv unit_less zero_less) define c where "c = a ! i * u ! i" then have ac: "a ! i dvd c" by simp have "a \ u = b \ v" using uv by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "c = b ! j * v ! j" using ij unfolding c_def by (subst (asm) u, subst (asm)v, subst u, subst v) auto then have bc: "b ! j dvd c" by simp have "a ! i * u ! i < a ! i * dij i j" using less_dij and no0 and ij by (auto simp: in_set_conv_nth) then have "c < lcm (a ! i) (b ! j)" by (auto simp: dij_def c_def) moreover have "lcm (a ! i) (b ! j) dvd c" by (simp add: ac bc) moreover have "c > 0" using \u ! i > 0\ and no0 and ij by (auto simp: c_def in_set_conv_nth) ultimately show False using ac and bc by (auto dest: nat_dvd_not_less) next case 2 then obtain i and j where ij: "i < m" "j < n" and less_dij: "v ! j < eij i j" and "u \\<^sub>v (zeroes m)[i := dij i j]" and "v \\<^sub>v (zeroes n)[j := eij i j]" using assms by (auto simp: Special_Solutions_def sij_def unit_less) then have u: "u = (zeroes m)[i := u ! i]" and v: "v = (zeroes n)[j := v ! j]" by (auto simp: less_eq_def list_eq_iff_nth_eq) (metis le_zero_eq length_list_update length_replicate rep_upd_unit)+ moreover have "nonzero v" using \nonzero u\ and \(u, v) \ Minimal_Solutions\ and Minimal_Solutions_imp_Solutions Solutions_snd_not_0 by blast ultimately have "v ! j > 0" using ij by (metis gr_implies_not0 neq0_conv unit_less zero_less) define c where "c = b ! j * v ! j" then have bc: "b ! j dvd c" by simp have "a \ u = b \ v" using uv by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "c = a ! i * u ! i" using ij unfolding c_def by (subst (asm) u, subst (asm)v, subst u, subst v) auto then have ac: "a ! i dvd c" by simp have "b ! j * v ! j < b ! j * eij i j" using less_dij and no0 and ij by (auto simp: in_set_conv_nth) then have "c < lcm (a ! i) (b ! j)" by (auto simp: eij_def c_def) moreover have "lcm (a ! i) (b ! j) dvd c" by (simp add: ac bc) moreover have "c > 0" using \v ! j > 0\ and no0 and ij by (auto simp: c_def in_set_conv_nth) ultimately show False using ac and bc by (auto dest: nat_dvd_not_less) qed qed qed (*Lemma 2 of Huet*) lemma non_special_solution_non_minimal: assumes "(x, y) \ Solutions - Special_Solutions" and ij: "i < m" "j < n" and "x ! i \ dij i j" and "y ! j \ eij i j" shows "(x, y) \ Minimal_Solutions" proof assume min: "(x, y) \ Minimal_Solutions" moreover have "sij i j \ Solutions" using ij by (intro Special_Solutions_in_Solutions) (auto simp: Special_Solutions_def) moreover have "(case sij i j of (u, v) \ u @ v) <\<^sub>v x @ y" using assms and min apply (cases "sij i j") apply (auto simp: sij_def Special_Solutions_def) by (metis List_Vector.le0 Minimal_Solutions_length le_append le_list_update less_append order_vec.dual_order.strict_iff_order same_append_eq) moreover have "(case sij i j of (u, v) \ nonzero u)" apply (auto simp: sij_def) by (metis dij_neq_0 ij length_replicate nonzero_iff set_update_memI) ultimately show False by (auto simp: Minimal_Solutions_def) qed subsection \Huet's conditions\ (*A*) definition "cond_A xs ys \ (\x\set xs. x \ maxne0 ys b)" (*B*) definition "cond_B x \ (\k\m. take k a \ take k x \ b \ map (max_y (take k x)) [0 ..< n])" (*C*) definition "boundr x y \ (\j max_y x j)" (*D*) definition "cond_D x y \ (\l\n. take l b \ take l y \ a \ x)" subsection \New conditions: facilitating generation of candidates from right to left\ (*condition on right sub-dotproduct*) definition "subdprodr y \ (\l\n. take l b \ take l y \ a \ map (max_x (take l y)) [0 ..< m])" (*condition on left sub-dotproduct*) definition "subdprodl x y \ (\k\m. take k a \ take k x \ b \ y)" (*bound on elements of left vector*) definition "boundl x y \ (\i max_x y i)" lemma boundr: assumes min: "(x, y) \ Minimal_Solutions" and "(x, y) \ Special_Solutions" shows "boundr x y" proof (unfold boundr_def, intro allI impI) fix j assume ass: "j < n" have ln: "m = length x \ n = length y" using assms Minimal_Solutions_def in_Solutions_iff min by auto have is_sol: "(x, y) \ Solutions" using assms Minimal_Solutions_def min by auto have j_less_l: "j < n" using assms ass le_less_trans by linarith consider (notemp) "Ej j x \ {}" | (empty) " Ej j x = {}" by blast then show "y ! j \ max_y x j" proof (cases) case notemp have max_y_def: "max_y x j = Min (Ej j x)" using j_less_l max_y_def notemp by auto have fin_e: "finite (Ej j x)" using finite_Ej [of j x] by auto have e_def': "\e \ Ej j x. (\i dij i j \ eij i j - 1 = e)" using Ej_def [of j x] by auto then have "\i dij i j \ eij i j - 1 = Min (Ej j x)" using notemp Min_in e_def' fin_e by blast then obtain i where i: "i < length x" "x ! i \ dij i j" "eij i j - 1 = Min (Ej j x)" by blast show ?thesis proof (rule ccontr) assume "\ ?thesis" with non_special_solution_non_minimal [of x y i j] and i and ln and assms and is_sol and j_less_l have "case sij i j of (u, v) \ u @ v \\<^sub>v x @ y" by (force simp: max_y_def) then have cs:"case sij i j of (u, v) \ u @ v <\<^sub>v x @ y" using assms by(auto simp: Special_Solutions_def) (metis append_eq_append_conv i(1) j_less_l length_list_update length_replicate sij_def order_vec.le_neq_trans ln prod.sel(1)) then obtain u v where u_v: "sij i j = (u, v)" "u @ v <\<^sub>v x @ y" by blast have dij_gt0: "dij i j > 0" using assms(1) assms(2) dij_neq_0 i(1) j_less_l ln by auto then have not_0_u: "nonzero u" proof (unfold nonzero_iff) have "i < length (zeroes m)" by (simp add: i(1) ln) then show "\i\set u. i \ 0" by (metis (no_types) Pair_inject dij_gt0 set_update_memI sij_def u_v(1) neq0_conv) qed then have "sij i j \ Solutions" by (metis (mono_tags, lifting) Special_Solutions_def i(1) Special_Solutions_in_Solutions j_less_l ln mem_Collect_eq u_v(1)) then show False using assms cs u_v not_0_u Minimal_Solutions_def min by auto qed next case empty have "\y\set y. y \ Max (set a)" using assms and max_coeff_bound and maxne0_le_Max using le_trans by blast then show ?thesis using empty j_less_l ln max_y_def by auto qed qed lemma boundl: assumes min: "(x, y) \ Minimal_Solutions" and "(x, y) \ Special_Solutions" shows "boundl x y" proof (unfold boundl_def, intro allI impI) fix i assume ass: "i < m" have ln: "n = length y \ m = length x" using assms Minimal_Solutions_def in_Solutions_iff min by auto have is_sol: "(x, y) \ Solutions" using assms Minimal_Solutions_def min by auto have i_less_l: "i < m" using assms ass le_less_trans by linarith consider (notemp) "Di i y \ {}" | (empty) " Di i y = {}" by blast then show "x ! i \ max_x y i" proof (cases) case notemp have max_x_def: "max_x y i = Min (Di i y)" using i_less_l max_x_def notemp by auto have fin_e: "finite (Di i y)" using finite_Di [of i y] by auto have e_def': "\e \ Di i y. (\j eij i j \ dij i j - 1 = e)" using Di_def [of i y] by auto then have "\j eij i j \ dij i j - 1 = Min (Di i y)" using notemp Min_in e_def' fin_e by blast then obtain j where j: "j < length y" "y ! j \ eij i j" "dij i j - 1 = Min (Di i y)" by blast show ?thesis proof (rule ccontr) assume "\ ?thesis" with non_special_solution_non_minimal [of x y i j] and j and ln and assms and is_sol and i_less_l have "case sij i j of (u, v) \ u @ v \\<^sub>v x @ y" by (force simp: max_x_def) then have cs: "case sij i j of (u, v) \ u @ v <\<^sub>v x @ y" using assms by(auto simp: Special_Solutions_def) (metis append_eq_append_conv j(1) i_less_l length_list_update length_replicate sij_def order_vec.le_neq_trans ln prod.sel(1)) then obtain u v where u_v: "sij i j = (u, v)" "u @ v <\<^sub>v x @ y" by blast have dij_gt0: "dij i j > 0" using assms(1) assms(2) dij_neq_0 j(1) i_less_l ln by auto then have not_0_u: "nonzero u" proof (unfold nonzero_iff) have "i < length (zeroes m)" using ass by simp then show "\i\set u. i \ 0" by (metis (no_types) Pair_inject dij_gt0 set_update_memI sij_def u_v(1) neq0_conv) qed then have "sij i j \ Solutions" by (metis (mono_tags, lifting) Special_Solutions_def j(1) Special_Solutions_in_Solutions i_less_l ln mem_Collect_eq u_v(1)) then show False using assms cs u_v not_0_u Minimal_Solutions_def min by auto qed next case empty have "\x\set x. x \ Max (set b)" using assms and max_coeff_bound and maxne0_le_Max using le_trans by blast then show ?thesis using empty i_less_l ln max_x_def by auto qed qed lemma Solution_imp_cond_D: assumes "(x, y) \ Solutions" shows "cond_D x y" using assms and dotprod_le_take by (auto simp: cond_D_def in_Solutions_iff) lemma Solution_imp_subdprodl: assumes "(x, y) \ Solutions" shows "subdprodl x y" using assms and dotprod_le_take by (auto simp: subdprodl_def in_Solutions_iff) metis theorem conds: assumes min: "(x, y) \ Minimal_Solutions" shows cond_A: "cond_A x y" and cond_B: "(x, y) \ Special_Solutions \ cond_B x" and "(x, y) \ Special_Solutions \ boundr x y" and cond_D: "cond_D x y" and subdprodr: "(x, y) \ Special_Solutions \ subdprodr y" and subdprodl: "subdprodl x y" proof - have sol: "a \ x = b \ y" and ln: "m = length x \ n = length y" using min by (auto simp: Minimal_Solutions_def in_Solutions_iff) then have "\i maxne0 y b" by (metis min max_coeff_bound_right nth_mem) then show "cond_A x y" using min and le_less_trans by (auto simp: cond_A_def max_coeff_bound) show "(x, y) \ Special_Solutions \ cond_B x" proof (unfold cond_B_def, intro allI impI) fix k assume non_spec: "(x, y) \ Special_Solutions" and k: "k \ m" from k have "take k a \ take k x \ a \ x" using dotprod_le_take ln by blast also have "... = b \ y" by fact also have map_b_dot_p: "... \ b \ map (max_y x) [0.. _ b \ ?nt") using non_spec and less_eq_def and ln and boundr and min by (fastforce intro!: dotprod_le_right simp: boundr_def) also have "... \ b \ map (max_y (take k x)) [0.. _ \ ?t") proof - have "\j ?t!j" using min and ln and max_y_le_take and k by auto then have "?nt \\<^sub>v ?t" using less_eq_def by auto then show ?thesis by (simp add: dotprod_le_right) qed finally show "take k a \ take k x \ b \ map (max_y (take k x)) [0.. Special_Solutions \ subdprodr y" proof (unfold subdprodr_def, intro allI impI) fix l assume non_spec: "(x, y) \ Special_Solutions" and l: "l \ n" from l have "take l b \ take l y \ b \ y" using dotprod_le_take ln by blast also have "... = a \ x" by (simp add: sol) also have map_b_dot_p: "... \ a \ map (max_x y) [0.. _ a \ ?nt") using non_spec and less_eq_def and ln and boundl and min by (fastforce intro!: dotprod_le_right simp: boundl_def) also have "... \ a \ map (max_x (take l y)) [0.. _ \ ?t") proof - have "\i ?t ! i" using min and ln and max_x_le_take and l by auto then have "?nt \\<^sub>v ?t" using less_eq_def by auto then show ?thesis by (simp add: dotprod_le_right) qed finally show "take l b \ take l y \ a \ map (max_x (take l y)) [0.. Special_Solutions \ boundr x y" using boundr [of x y] and min by blast show "cond_D x y" using ln and dotprod_le_take and sol by (auto simp: cond_D_def) show "subdprodl x y" using ln and dotprod_le_take and sol by (force simp: subdprodl_def) qed lemma le_imp_Ej_subset: assumes "u \\<^sub>v x" shows "Ej j u \ Ej j x" using assms and le_trans by (force simp: Ej_def less_eq_def dij_def eij_def) lemma le_imp_max_y_ge: assumes "u \\<^sub>v x" and "length x \ m" shows "max_y u j \ max_y x j" using assms and le_imp_Ej_subset and Min_Ej_le [of j, OF _ _ assms(2)] by (metis Min.subset_imp Min_in emptyE finite_Ej max_y_def order_refl subsetCE) lemma le_imp_Di_subset: assumes "v \\<^sub>v y" shows "Di i v \ Di i y" using assms and le_trans by (force simp: Di_def less_eq_def dij_def eij_def) lemma le_imp_max_x_ge: assumes "v \\<^sub>v y" and "length y \ n" shows "max_x v i \ max_x y i" using assms and le_imp_Di_subset and Min_Di_le [of i, OF _ _ assms(2)] by (metis Min.subset_imp Min_in emptyE finite_Di max_x_def order_refl subsetCE) end end diff --git a/thys/Dirichlet_L/Dirichlet_Characters.thy b/thys/Dirichlet_L/Dirichlet_Characters.thy --- a/thys/Dirichlet_L/Dirichlet_Characters.thy +++ b/thys/Dirichlet_L/Dirichlet_Characters.thy @@ -1,463 +1,463 @@ (* File: Dirichlet_Characters.thy Author: Manuel Eberl, TU München *) section \Dirichlet Characters\ theory Dirichlet_Characters imports Multiplicative_Characters "HOL-Number_Theory.Residues" "Dirichlet_Series.Multiplicative_Function" begin text \ Dirichlet characters are essentially just the characters of the multiplicative group of integer residues $\mathbb{ZZ}/n\mathbb{ZZ}$ for some fixed $n$. For convenience, these residues are usually represented by natural numbers from $0$ to $n - 1$, and we extend the characters to all natural numbers periodically, so that $\chi(k\mod n) = \chi(k)$ holds. Numbers that are not coprime to $n$ are not in the group and therefore are assigned $0$ by all characters. \ subsection \The multiplicative group of residues\ definition residue_mult_group :: "nat \ nat monoid" where "residue_mult_group n = \ carrier = totatives n, monoid.mult = (\x y. (x * y) mod n), one = 1 \" definition principal_dchar :: "nat \ nat \ complex" where "principal_dchar n = (\k. if coprime k n then 1 else 0)" lemma principal_dchar_coprime [simp]: "coprime k n \ principal_dchar n k = 1" and principal_dchar_not_coprime [simp]: "\coprime k n \ principal_dchar n k = 0" by (simp_all add: principal_dchar_def) lemma principal_dchar_1 [simp]: "principal_dchar n 1 = 1" by simp lemma principal_dchar_minus1 [simp]: assumes "n > 0" shows "principal_dchar n (n - Suc 0) = 1" proof (cases "n = 1") case False with assms have "n > 1" by linarith thus ?thesis using coprime_diff_one_left_nat[of n] by (intro principal_dchar_coprime) auto qed auto lemma mod_in_totatives: "n > 1 \ a mod n \ totatives n \ coprime a n" by (auto simp: totatives_def mod_greater_zero_iff_not_dvd dest: coprime_common_divisor_nat) bundle dcharacter_syntax begin notation principal_dchar ("\\<^sub>0\") end locale residues_nat = fixes n :: nat (structure) and G assumes n: "n > 1" defines "G \ residue_mult_group n" begin lemma order [simp]: "order G = totient n" by (simp add: order_def G_def totient_def residue_mult_group_def) lemma totatives_mod [simp]: "x \ totatives n \ x mod n = x" using n by (intro mod_less) (auto simp: totatives_def intro!: order.not_eq_order_implies_strict) lemma principal_dchar_minus1 [simp]: "principal_dchar n (n - Suc 0) = 1" using principal_dchar_minus1[of n] n by simp sublocale finite_comm_group G proof fix x y assume xy: "x \ carrier G" "y \ carrier G" hence "coprime (x * y) n" by (auto simp: G_def residue_mult_group_def totatives_def) with xy and n show "x \\<^bsub>G\<^esub> y \ carrier G" using coprime_common_divisor_nat[of "x * y" n] by (auto simp: G_def residue_mult_group_def totatives_def mod_greater_zero_iff_not_dvd le_Suc_eq simp del: coprime_mult_left_iff) next fix x y z assume xyz: "x \ carrier G" "y \ carrier G" "z \ carrier G" thus "x \\<^bsub>G\<^esub> y \\<^bsub>G\<^esub> z = x \\<^bsub>G\<^esub> (y \\<^bsub>G\<^esub> z)" by (auto simp: G_def residue_mult_group_def mult_ac mod_mult_right_eq) next fix x assume "x \ carrier G" with n have "x < n" by (auto simp: G_def residue_mult_group_def totatives_def intro!: order.not_eq_order_implies_strict) thus " \\<^bsub>G\<^esub> \\<^bsub>G\<^esub> x = x" and "x \\<^bsub>G\<^esub> \\<^bsub>G\<^esub> = x" by (simp_all add: G_def residue_mult_group_def) next have "x \ Units G" if "x \ carrier G" for x unfolding Units_def proof safe from that have "x > 0" "coprime x n" by (auto simp: G_def residue_mult_group_def totatives_def) from \coprime x n\ and n obtain y where y: "y < n" "[x * y = 1] (mod n)" by (subst (asm) coprime_iff_invertible'_nat) auto hence "x * y mod n = 1" using n by (simp add: cong_def mult_ac) moreover from y have "coprime y n" by (subst coprime_iff_invertible_nat) (auto simp: mult.commute) ultimately show "\a\carrier G. a \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub> \ x \\<^bsub>G\<^esub> a = \\<^bsub>G\<^esub>" using y by (intro bexI[of _ y]) (auto simp: G_def residue_mult_group_def totatives_def mult.commute intro!: Nat.gr0I) qed fact+ thus "carrier G \ Units G" .. qed (insert n, auto simp: G_def residue_mult_group_def mult_ac) subsection \Definition of Dirichlet characters\ text \ The following two functions make the connection between Dirichlet characters and the multiplicative characters of the residue group. \ definition c2dc :: "(nat \ complex) \ (nat \ complex)" where "c2dc \ = (\x. \ (x mod n))" definition dc2c :: "(nat \ complex) \ (nat \ complex)" where "dc2c \ = (\x. if x < n then \ x else 0)" lemma dc2c_c2dc [simp]: assumes "character G \" shows "dc2c (c2dc \) = \" proof - interpret character G \ by fact show ?thesis using n by (auto simp: fun_eq_iff dc2c_def c2dc_def char_eq_0_iff G_def residue_mult_group_def totatives_def) qed end locale dcharacter = residues_nat + fixes \ :: "nat \ complex" assumes mult_aux: "a \ totatives n \ b \ totatives n \ \ (a * b) = \ a * \ b" assumes eq_zero: "\coprime a n \ \ a = 0" assumes periodic: "\ (a + n) = \ a" assumes one_not_zero: "\ 1 \ 0" begin lemma zero_eq_0 [simp]: "\ 0 = 0" using n by (intro eq_zero) auto lemma Suc_0 [simp]: "\ (Suc 0) = 1" using n mult_aux[of 1 1] one_not_zero by (simp add: totatives_def) lemma periodic_mult: "\ (a + m * n) = \ a" proof (induction m) case (Suc m) have "a + Suc m * n = a + m * n+ n" by simp also have "\ \ = \ (a + m * n)" by (rule periodic) also have "\ = \ a" by (rule Suc.IH) finally show ?case . qed simp_all lemma minus_one_periodic [simp]: assumes "k > 0" shows "\ (k * n - 1) = \ (n - 1)" proof - have "k * n - 1 = n - 1 + (k - 1) * n" using assms n by (simp add: algebra_simps) also have "\ \ = \ (n - 1)" by (rule periodic_mult) finally show ?thesis . qed lemma cong: assumes "[a = b] (mod n)" shows "\ a = \ b" proof - from assms obtain k1 k2 where *: "b + k1 * n = a + k2 * n" by (subst (asm) cong_iff_lin_nat) auto have "\ a = \ (a + k2 * n)" by (rule periodic_mult [symmetric]) also note * [symmetric] also have "\ (b + k1 * n) = \ b" by (rule periodic_mult) finally show ?thesis . qed lemma mod [simp]: "\ (a mod n) = \ a" by (rule cong) (simp_all add: cong_def) lemma mult [simp]: "\ (a * b) = \ a * \ b" proof (cases "coprime a n \ coprime b n") case True hence "a mod n \ totatives n" "b mod n \ totatives n" using n by (auto simp: totatives_def mod_greater_zero_iff_not_dvd coprime_absorb_right) hence "\ ((a mod n) * (b mod n)) = \ (a mod n) * \ (b mod n)" by (rule mult_aux) also have "\ ((a mod n) * (b mod n)) = \ (a * b)" by (rule cong) (auto simp: cong_def mod_mult_eq) finally show ?thesis by simp next case False hence "\coprime (a * b) n" by simp with False show ?thesis by (auto simp: eq_zero) qed sublocale mult: completely_multiplicative_function \ by standard auto lemma eq_zero_iff: "\ x = 0 \ \coprime x n" proof safe assume "\ x = 0" and "coprime x n" from cong_solve_coprime_nat [OF this(2)] obtain y where "[x * y = Suc 0] (mod n)" by blast hence "\ (x * y) = \ (Suc 0)" by (rule cong) with \\ x = 0\ show False by simp qed (auto simp: eq_zero) lemma minus_one': "\ (n - 1) \ {-1, 1}" proof - define n' where "n' = n - 2" have n: "n = Suc (Suc n')" using n by (simp add: n'_def) have "(n - 1) ^ 2 = 1 + (n - 2) * n" by (simp add: power2_eq_square algebra_simps n) also have "\ \ = 1" by (subst periodic_mult) auto also have "\ ((n - 1) ^ 2) = \ (n - 1) ^ 2" by (rule mult.power) finally show ?thesis by (subst (asm) power2_eq_1_iff) auto qed lemma c2dc_dc2c [simp]: "c2dc (dc2c \) = \" using n by (auto simp: c2dc_def dc2c_def fun_eq_iff intro!: cong simp: cong_def) lemma character_dc2c: "character G (dc2c \)" by standard (insert n, auto simp: G_def residue_mult_group_def dc2c_def totatives_def intro!: eq_zero) sublocale dc2c: character G "dc2c \" by (fact character_dc2c) lemma dcharacter_inv_character [intro]: "dcharacter n (inv_character \)" by standard (auto simp: inv_character_def eq_zero periodic) lemma norm: "norm (\ k) = (if coprime k n then 1 else 0)" proof - have "\ k = \ (k mod n)" by (intro cong) (auto simp: cong_def) also from n have "\ = dc2c \ (k mod n)" by (simp add: dc2c_def) also from n have "norm \ = (if coprime k n then 1 else 0)" by (subst dc2c.norm_char) (auto simp: G_def residue_mult_group_def mod_in_totatives) finally show ?thesis . qed lemma norm_le_1: "norm (\ k) \ 1" by (subst norm) auto end definition dcharacters :: "nat \ (nat \ complex) set" where "dcharacters n = {\. dcharacter n \}" context residues_nat begin lemma character_dc2c: "dcharacter n \ \ character G (dc2c \)" using dcharacter.character_dc2c[of n \] by (simp add: G_def) lemma dcharacter_c2dc: assumes "character G \" shows "dcharacter n (c2dc \)" proof - interpret character G \ by fact show ?thesis proof fix x assume "\coprime x n" thus "c2dc \ x = 0" by (auto simp: c2dc_def char_eq_0_iff G_def residue_mult_group_def totatives_def) qed (insert char_mult char_one n, auto simp: c2dc_def G_def residue_mult_group_def simp del: char_mult char_one) qed lemma principal_dchar_altdef: "principal_dchar n = c2dc (principal_char G)" using n by (auto simp: c2dc_def principal_dchar_def principal_char_def G_def residue_mult_group_def fun_eq_iff mod_in_totatives) sublocale principal: dcharacter n G "principal_dchar n" by (simp add: principal_dchar_altdef dcharacter_c2dc | rule G_def)+ lemma c2dc_principal [simp]: "c2dc (principal_char G) = principal_dchar n" by (simp add: principal_dchar_altdef) lemma dc2c_principal [simp]: "dc2c (principal_dchar n) = principal_char G" proof - have "dc2c (c2dc (principal_char G)) = dc2c (principal_dchar n)" by (subst c2dc_principal) (rule refl) thus ?thesis by (subst (asm) dc2c_c2dc) simp_all qed lemma bij_betw_dcharacters_characters: "bij_betw dc2c (dcharacters n) (characters G)" by (intro bij_betwI[where ?g = c2dc]) (auto simp: characters_def dcharacters_def dcharacter_c2dc character_dc2c dcharacter.c2dc_dc2c) lemma bij_betw_characters_dcharacters: "bij_betw c2dc (characters G) (dcharacters n)" by (intro bij_betwI[where ?g = dc2c]) (auto simp: characters_def dcharacters_def dcharacter_c2dc character_dc2c dcharacter.c2dc_dc2c) lemma finite_dcharacters [intro]: "finite (dcharacters n)" using bij_betw_finite [OF bij_betw_dcharacters_characters] by auto lemma card_dcharacters [simp]: "card (dcharacters n) = totient n" using bij_betw_same_card [OF bij_betw_dcharacters_characters] card_characters by simp end lemma inv_character_eq_principal_dchar_iff [simp]: "inv_character \ = principal_dchar n \ \ = principal_dchar n" by (auto simp add: fun_eq_iff inv_character_def principal_dchar_def) subsection \Sums of Dirichlet characters\ lemma (in dcharacter) sum_dcharacter_totatives: "(\x\totatives n. \ x) = (if \ = principal_dchar n then of_nat (totient n) else 0)" proof - from n have "(\x\totatives n. \ x) = (\x\carrier G. dc2c \ x)" by (intro sum.cong) (auto simp: totatives_def dc2c_def G_def residue_mult_group_def) also have "\ = (if dc2c \ = principal_char G then of_nat (order G) else 0)" by (rule dc2c.sum_character) also have "dc2c \ = principal_char G \ \ = principal_dchar n" by (metis c2dc_dc2c dc2c_principal principal_dchar_altdef) finally show ?thesis by simp qed lemma (in dcharacter) sum_dcharacter_block: "(\x x) = (if \ = principal_dchar n then of_nat (totient n) else 0)" proof - from n have "(\x x) = (\x\totatives n. \ x)" by (intro sum.mono_neutral_right) (auto simp: totatives_def eq_zero_iff intro!: Nat.gr0I order.not_eq_order_implies_strict) also have "\ = (if \ = principal_dchar n then of_nat (totient n) else 0)" by (rule sum_dcharacter_totatives) finally show ?thesis . qed lemma (in dcharacter) sum_dcharacter_block': "sum \ {Suc 0..n} = (if \ = principal_dchar n then of_nat (totient n) else 0)" proof - let ?f = "\k. if k = n then 0 else k" and ?g = "\k. if k = 0 then n else k" have "sum \ {1..n} = sum \ {.. \ principal_dchar n" shows "(\x x) = (\x x)" proof (induction m rule: less_induct) case (less m) show ?case proof (cases "m < n") case True thus ?thesis by simp next case False hence "{.. {n..x\\. \ x) = (\x x) + (\x\{n.. x)" by (intro sum.union_disjoint) auto also from assms have "(\x x) = 0" by (subst sum_dcharacter_block) simp_all also from False have "(\x\{n.. x) = (\x\{.. (x + n))" by (intro sum.reindex_bij_witness[of _ "\x. x + n" "\x. x - n"]) (auto simp: periodic) also have "\ = (\x\{.. x)" by (simp add: periodic) also have "\ = (\x<(m - n) mod n. \ x)" using False and n by (intro less.IH) auto also from False and n have "(m - n) mod n = m mod n" - by (simp add: mod_geq [symmetric]) + by (simp add: le_mod_geq) finally show ?thesis by simp qed qed lemma (in dcharacter) sum_dcharacter_lessThan_le: assumes "\ \ principal_dchar n" shows "norm (\x x) \ totient n" proof - have "(\x x) = (\x x)" by (rule sum_lessThan_dcharacter) fact also have "\ = (\x | x < m mod n \ coprime x n. \ x)" by (intro sum.mono_neutral_right) (auto simp: eq_zero_iff) also have "norm \ \ (\x | x < m mod n \ coprime x n. 1)" by (rule sum_norm_le) (auto simp: norm) also have "\ = card {x. x < m mod n \ coprime x n}" by simp also have "\ \ card (totatives n)" unfolding of_nat_le_iff proof (intro card_mono subsetI) fix x assume x: "x \ {x. x < m mod n \ coprime x n}" hence "x < m mod n" by simp also have "\ < n" using n by simp finally show "x \ totatives n" using x by (auto simp: totatives_def intro!: Nat.gr0I) qed auto also have "\ = totient n" by (simp add: totient_def) finally show ?thesis . qed lemma (in dcharacter) sum_dcharacter_atMost_le: assumes "\ \ principal_dchar n" shows "norm (\x\m. \ x) \ totient n" using sum_dcharacter_lessThan_le[OF assms, of "Suc m"] by (subst (asm) lessThan_Suc_atMost) lemma (in residues_nat) sum_dcharacters: "(\\\dcharacters n. \ x) = (if [x = 1] (mod n) then of_nat (totient n) else 0)" proof (cases "coprime x n") case True with n have x: "x mod n \ totatives n" by (auto simp: mod_in_totatives) have "(\\\dcharacters n. \ x) = (\\\characters G. c2dc \ x)" by (rule sum.reindex_bij_betw [OF bij_betw_characters_dcharacters, symmetric]) also from x have "\ = (\\\characters G. \ (x mod n))" by (simp add: c2dc_def) also from x have "\ = (if x mod n = 1 then order G else 0)" by (subst sum_characters) (unfold G_def residue_mult_group_def, auto) also from n have "x mod n = 1 \ [x = 1] (mod n)" by (simp add: cong_def) finally show ?thesis by simp next case False have "x mod n \ 1" proof assume *: "x mod n = 1" have "gcd (x mod n) n = 1" by (subst *) simp also have "gcd (x mod n) n = gcd x n" by (subst gcd.commute) (simp only: gcd_red_nat [symmetric]) finally show False using \\coprime x n\ unfolding coprime_iff_gcd_eq_1 by contradiction qed from False have "(\\\dcharacters n. \ x) = 0" by (intro sum.neutral) (auto simp: dcharacters_def dcharacter.eq_zero) with \x mod n \ 1\ and n show ?thesis by (simp add: cong_def) qed lemma (in dcharacter) even_dcharacter_linear_sum_eq_0 [simp]: assumes "\ \ principal_dchar n" and "\ (n - 1) = 1" shows "(\k=Suc 0.. k) = 0" proof - have "(\k=1.. k) = (\k=1.. (n - k))" by (intro sum.reindex_bij_witness[where i = "\k. n - k" and j = "\k. n - k"]) (auto simp: of_nat_diff) also have "\ = n * (\k=1.. (n - k)) - (\k=1.. (n - k))" by (simp add: algebra_simps sum_subtractf sum_distrib_left) also have "(\k=1.. (n - k)) = (\k=1.. k)" by (intro sum.reindex_bij_witness[where i = "\k. n - k" and j = "\k. n - k"]) auto also have "\ = (\k k)" by (intro sum.mono_neutral_left) (auto simp: Suc_le_eq) also have "\ = 0" using assms by (simp add: sum_dcharacter_block) also have "(\k=1.. (n - k)) = (\k=1.. k)" proof (intro sum.cong refl) fix k assume k: "k \ {1.. k = of_nat k * \ ((n - 1) * k)" using assms by (subst mult) simp_all also have "(n - 1) * k = n - k + (k - 1) * n" using k by (simp add: algebra_simps) also have "\ \ = \ (n - k)" by (rule periodic_mult) finally show "of_nat k * \ (n - k) = of_nat k * \ k" .. qed finally show ?thesis by simp qed end diff --git a/thys/Gauss_Sums/Gauss_Sums.thy b/thys/Gauss_Sums/Gauss_Sums.thy --- a/thys/Gauss_Sums/Gauss_Sums.thy +++ b/thys/Gauss_Sums/Gauss_Sums.thy @@ -1,1704 +1,1705 @@ (* File: Gauss_Sums.thy Authors: Rodrigo Raya, EPFL; Manuel Eberl, TUM Gauss sums and more on Dirichlet characters: induced moduli, separability, primitive characters *) theory Gauss_Sums imports "HOL-Algebra.Coset" "HOL-Real_Asymp.Real_Asymp" Ramanujan_Sums begin section \Gauss sums\ bundle vec_lambda_notation begin notation vec_lambda (binder "\" 10) end bundle no_vec_lambda_notation begin no_notation vec_lambda (binder "\" 10) end unbundle no_vec_lambda_notation subsection \Definition and basic properties\ context dcharacter begin (* TODO remove when integrating periodic and periodic_function *) lemma dir_periodic_arithmetic: "periodic_arithmetic \ n" unfolding periodic_arithmetic_def by (simp add: periodic) definition "gauss_sum k = (\m = 1..n . \(m) * unity_root n (m*k))" lemma gauss_sum_periodic: "periodic_arithmetic (\n. gauss_sum n) n" proof - have "periodic_arithmetic \ n" using dir_periodic_arithmetic by simp let ?h = "\m k. \(m) * unity_root n (m*k)" {fix m :: nat have "periodic_arithmetic (\k. unity_root n (m*k)) n" using unity_periodic_arithmetic_mult[of n m] by simp have "periodic_arithmetic (?h m) n" using scalar_mult_periodic_arithmetic[OF \periodic_arithmetic (\k. unity_root n (m*k)) n\] by blast} then have per_all: "\m \ {1..n}. periodic_arithmetic (?h m) n" by blast have "periodic_arithmetic (\k. (\m = 1..n . \(m) * unity_root n (m*k))) n" using fin_sum_periodic_arithmetic_set[OF per_all] by blast then show ?thesis unfolding gauss_sum_def by blast qed lemma ramanujan_sum_conv_gauss_sum: assumes "\ = principal_dchar n" shows "ramanujan_sum n k = gauss_sum k" proof - {fix m from assms have 1: "coprime m n \ \(m) = 1" and 2: "\ coprime m n \ \(m) = 0" unfolding principal_dchar_def by auto} note eq = this have "gauss_sum k = (\m = 1..n . \(m) * unity_root n (m*k))" unfolding gauss_sum_def by simp also have "\ = (\m | m \ {1..n} \ coprime m n . \(m) * unity_root n (m*k))" by (rule sum.mono_neutral_right,simp,blast,simp add: eq) also have "\ = (\m | m \ {1..n} \ coprime m n . unity_root n (m*k))" by (simp add: eq) also have "\ = ramanujan_sum n k" unfolding ramanujan_sum_def by blast finally show ?thesis .. qed lemma cnj_mult_self: assumes "coprime k n" shows "cnj (\ k) * \ k = 1" proof - have "cnj (\ k) * \ k = norm (\ k)^2" by (simp add: mult.commute complex_mult_cnj cmod_def) also have "\ = 1" using norm[of k] assms by simp finally show ?thesis . qed text \Theorem 8.9\ theorem gauss_sum_reduction: assumes "coprime k n" shows "gauss_sum k = cnj (\ k) * gauss_sum 1" proof - from n have n_pos: "n > 0" by simp have "gauss_sum k = (\r = 1..n . \(r) * unity_root n (r*k))" unfolding gauss_sum_def by simp also have "\ = (\r = 1..n . cnj (\(k)) * \ k * \ r * unity_root n (r*k))" using assms by (intro sum.cong) (auto simp: cnj_mult_self) also have "\ = (\r = 1..n . cnj (\(k)) * \ (k*r) * unity_root n (r*k))" by (intro sum.cong) auto also have "\ = cnj (\(k)) * (\r = 1..n . \ (k*r) * unity_root n (r*k))" by (simp add: sum_distrib_left algebra_simps) also have "\= cnj (\(k)) * (\r = 1..n . \ r * unity_root n r)" proof - have 1: "periodic_arithmetic (\r. \ r * unity_root n r) n" using dir_periodic_arithmetic unity_periodic_arithmetic mult_periodic_arithmetic by blast have "(\r = 1..n . \ (k*r) * unity_root n (r*k)) = (\r = 1..n . \ (r)* unity_root n r)" using periodic_arithmetic_remove_homothecy[OF assms(1) 1 n_pos] by (simp add: algebra_simps n) then show ?thesis by argo qed also have "\ = cnj (\(k)) * gauss_sum 1" using gauss_sum_def by simp finally show ?thesis . qed text \ The following variant takes an integer argument instead. \ definition "gauss_sum_int k = (\m=1..n. \ m * unity_root n (int m*k))" sublocale gauss_sum_int: periodic_fun_simple gauss_sum_int "int n" proof fix k show "gauss_sum_int (k + int n) = gauss_sum_int k" by (simp add: gauss_sum_int_def ring_distribs unity_root_add) qed lemma gauss_sum_int_cong: assumes "[a = b] (mod int n)" shows "gauss_sum_int a = gauss_sum_int b" proof - from assms obtain k where k: "b = a + int n * k" by (subst (asm) cong_iff_lin) auto thus ?thesis using gauss_sum_int.plus_of_int[of a k] by (auto simp: algebra_simps) qed lemma gauss_sum_conv_gauss_sum_int: "gauss_sum k = gauss_sum_int (int k)" unfolding gauss_sum_def gauss_sum_int_def by auto lemma gauss_sum_int_conv_gauss_sum: "gauss_sum_int k = gauss_sum (nat (k mod n))" proof - have "gauss_sum (nat (k mod n)) = gauss_sum_int (int (nat (k mod n)))" by (simp add: gauss_sum_conv_gauss_sum_int) also have "\ = gauss_sum_int k" using n by (intro gauss_sum_int_cong) (auto simp: cong_def) finally show ?thesis .. qed lemma gauss_int_periodic: "periodic_arithmetic gauss_sum_int n" unfolding periodic_arithmetic_def gauss_sum_int_conv_gauss_sum by simp proposition dcharacter_fourier_expansion: "\ m = (\k=1..n. 1 / n * gauss_sum_int (-k) * unity_root n (m*k))" proof - define g where "g = (\x. 1 / of_nat n * (\m m * unity_root n (- int x * int m)))" have per: "periodic_arithmetic \ n" using dir_periodic_arithmetic by simp have "\ m = (\k = (\k = 1..n. g k * unity_root n (m * int k))" proof - have g_per: "periodic_arithmetic g n" using fourier_expansion_periodic_arithmetic(1)[OF _ per] n by (simp add: g_def) have fact_per: "periodic_arithmetic (\k. g k * unity_root n (int m * int k)) n" using mult_periodic_arithmetic[OF g_per] unity_periodic_arithmetic_mult by auto show ?thesis proof - have "(\kl = 0..n - Suc 0. g l * unity_root n (int m * int l))" using n by (intro sum.cong) auto also have "\ = (\l = Suc 0..n. g l * unity_root n (int m * int l))" using periodic_arithmetic_sum_periodic_arithmetic_shift[OF fact_per, of 1] n by auto finally show ?thesis by simp qed qed also have "\ = (\k = 1..n. (1 / of_nat n) * gauss_sum_int (-k) * unity_root n (m*k))" proof - {fix k :: nat have shift: "(\m m * unity_root n (- int k * int m)) = (\m = 1..n. \ m * unity_root n (- int k * int m))" proof - have per_unit: "periodic_arithmetic (\m. unity_root n (- int k * int m)) n" using unity_periodic_arithmetic_mult by blast then have prod_per: "periodic_arithmetic (\m. \ m * unity_root n (- int k * int m)) n" using per mult_periodic_arithmetic by blast show ?thesis proof - have "(\m m * unity_root n (- int k * int m)) = (\l = 0..n - Suc 0. \ l * unity_root n (- int k * int l))" using n by (intro sum.cong) auto also have "\ = (\m = 1..n. \ m * unity_root n (- int k * int m))" using periodic_arithmetic_sum_periodic_arithmetic_shift[OF prod_per, of 1] n by auto finally show ?thesis by simp qed qed have "g k = 1 / of_nat n * (\m m * unity_root n (- int k * int m))" using g_def by auto also have "\ = 1 / of_nat n * (\m = 1..n. \ m * unity_root n (- int k * int m))" using shift by simp also have "\ = 1 / of_nat n * gauss_sum_int (-k)" unfolding gauss_sum_int_def by (simp add: algebra_simps) finally have "g k = 1 / of_nat n * gauss_sum_int (-k)" by simp} note g_expr = this show ?thesis by (rule sum.cong, simp, simp add: g_expr) qed finally show ?thesis by auto qed subsection \Separability\ definition "separable k \ gauss_sum k = cnj (\ k) * gauss_sum 1" corollary gauss_coprime_separable: assumes "coprime k n" shows "separable k" using gauss_sum_reduction[OF assms] unfolding separable_def by simp text \Theorem 8.10\ theorem global_separability_condition: "(\n>0. separable n) \ (\k>0. \coprime k n \ gauss_sum k = 0)" proof - {fix k assume "\ coprime k n" then have "\(k) = 0" by (simp add: eq_zero) then have "cnj (\ k) = 0" by blast then have "separable k \ gauss_sum k = 0" unfolding separable_def by auto} note not_case = this show ?thesis using gauss_coprime_separable not_case separable_def by blast qed lemma of_real_moebius_mu [simp]: "of_real (moebius_mu k) = moebius_mu k" by (simp add: moebius_mu_def) corollary principal_not_totally_separable: assumes "\ = principal_dchar n" shows "\(\k > 0. separable k)" proof - have n_pos: "n > 0" using n by simp have tot_0: "totient n \ 0" by (simp add: n_pos) have "moebius_mu (n div gcd n n) \ 0" by (simp add: \n > 0\) then have moeb_0: "\k. moebius_mu (n div gcd k n) \ 0" by blast have lem: "gauss_sum k = totient n * moebius_mu (n div gcd k n) / totient (n div gcd k n)" if "k > 0" for k proof - have "gauss_sum k = ramanujan_sum n k" using ramanujan_sum_conv_gauss_sum[OF assms(1)] .. also have "\ = totient n * moebius_mu (n div gcd k n) / (totient (n div gcd k n))" by (simp add: ramanujan_sum_k_n_dirichlet_expr[OF n_pos that]) finally show ?thesis . qed have 2: "\ coprime n n" using n by auto have 3: "gauss_sum n \ 0" using lem[OF n_pos] tot_0 moebius_mu_1 by simp from n_pos 2 3 have "\k>0. \coprime k n \ gauss_sum k \ 0" by blast then obtain k where "k > 0 \ \ coprime k n \ gauss_sum k \ 0" by blast note right_not_zero = this have "cnj (\ k) * gauss_sum 1 = 0" if "\coprime k n" for k using that assms by (simp add: principal_dchar_def) then show ?thesis unfolding separable_def using right_not_zero by auto qed text \Theorem 8.11\ theorem gauss_sum_1_mod_square_eq_k: assumes "(\k. k > 0 \ separable k)" shows "norm (gauss_sum 1) ^ 2 = real n" proof - have "(norm (gauss_sum 1))^2 = gauss_sum 1 * cnj (gauss_sum 1)" using complex_norm_square by blast also have "\ = gauss_sum 1 * (\m = 1..n. cnj (\(m)) * unity_root n (-m))" proof - have "cnj (gauss_sum 1) = (\m = 1..n. cnj (\(m)) * unity_root n (-m))" unfolding gauss_sum_def by (simp add: unity_root_uminus) then show ?thesis by argo qed also have "\ = (\m = 1..n. gauss_sum 1 * cnj (\(m)) * unity_root n (-m))" by (subst sum_distrib_left)(simp add: algebra_simps) also have "\ = (\m = 1..n. gauss_sum m * unity_root n (-m))" proof (rule sum.cong,simp) fix x assume as: "x \ {1..n}" show "gauss_sum 1 * cnj (\ x) * unity_root n (-x) = gauss_sum x * unity_root n (-x)" using assms(1) unfolding separable_def by (rule allE[of _ x]) (use as in auto) qed also have "\ = (\m = 1..n. (\r = 1..n. \ r * unity_root n (r*m) * unity_root n (-m)))" unfolding gauss_sum_def by (rule sum.cong,simp,rule sum_distrib_right) also have "\ = (\m = 1..n. (\r = 1..n. \ r * unity_root n (m*(r-1)) ))" by (intro sum.cong refl) (auto simp: unity_root_diff of_nat_diff unity_root_uminus field_simps) also have "\ = (\r=1..n. (\m=1..n. \(r) *unity_root n (m*(r-1))))" by (rule sum.swap) also have "\ = (\r=1..n. \(r) *(\m=1..n. unity_root n (m*(r-1))))" by (rule sum.cong, simp, simp add: sum_distrib_left) also have "\ = (\r=1..n. \(r) * unity_root_sum n (r-1))" proof (intro sum.cong refl) fix x assume "x \ {1..n}" then have 1: "periodic_arithmetic (\m. unity_root n (int (m * (x - 1)))) n" using unity_periodic_arithmetic_mult[of n "x-1"] by (simp add: mult.commute) have "(\m = 1..n. unity_root n (int (m * (x - 1)))) = (\m = 0..n-1. unity_root n (int (m * (x - 1))))" using periodic_arithmetic_sum_periodic_arithmetic_shift[OF 1 _, of 1] n by simp also have "\ = unity_root_sum n (x-1)" using n unfolding unity_root_sum_def by (intro sum.cong) (auto simp: mult_ac) finally have "(\m = 1..n. unity_root n (int (m * (x - 1)))) = unity_root_sum n (int (x - 1))" . then show "\ x * (\m = 1..n. unity_root n (int (m * (x - 1)))) = \ x * unity_root_sum n (int (x - 1))" by argo qed also have "\ = (\r \ {1}. \ r * unity_root_sum n (int (r - 1)))" using n unity_root_sum_nonzero_iff int_ops(6) by (intro sum.mono_neutral_right) auto also have "\ = \ 1 * n" using n by simp also have "\ = n" by simp finally show ?thesis using of_real_eq_iff by fastforce qed text \Theorem 8.12\ theorem gauss_sum_nonzero_noncoprime_necessary_condition: assumes "gauss_sum k \ 0" "\coprime k n" "k > 0" defines "d \ n div gcd k n" assumes "coprime a n" "[a = 1] (mod d)" shows "d dvd n" "d < n" "\ a = 1" proof - show "d dvd n" unfolding d_def using n by (subst div_dvd_iff_mult) auto from assms(2) have "gcd k n \ 1" by blast then have "gcd k n > 1" using assms(3,4) by (simp add: nat_neq_iff) with n show "d < n" by (simp add: d_def) have "periodic_arithmetic (\r. \ (r)* unity_root n (k*r)) n" using mult_periodic_arithmetic[OF dir_periodic_arithmetic unity_periodic_arithmetic_mult] by auto then have 1: "periodic_arithmetic (\r. \ (r)* unity_root n (r*k)) n" by (simp add: algebra_simps) have "gauss_sum k = (\m = 1..n . \(m) * unity_root n (m*k))" unfolding gauss_sum_def by blast also have "\ = (\m = 1..n . \(m*a) * unity_root n (m*a*k))" using periodic_arithmetic_remove_homothecy[OF assms(5) 1] n by auto also have "\ = (\m = 1..n . \(m*a) * unity_root n (m*k))" proof (intro sum.cong refl) fix m from assms(6) obtain b where "a = 1 + b*d" using \d < n\ assms(5) cong_to_1'_nat by auto then have "m*a*k = m*k+m*b*(n div gcd k n)*k" by (simp add: algebra_simps d_def) also have "\ = m*k+m*b*n*(k div gcd k n)" by (simp add: div_mult_swap dvd_div_mult) also obtain p where "\ = m*k+m*b*n*p" by blast finally have "m*a*k = m*k+m*b*p*n" by simp then have 1: "m*a*k mod n= m*k mod n" using mod_mult_self1 by simp then have "unity_root n (m * a * k) = unity_root n (m * k)" proof - have "unity_root n (m * a * k) = unity_root n ((m * a * k) mod n)" using unity_root_mod[of n] zmod_int by simp also have "\ = unity_root n (m * k)" using unity_root_mod[of n] zmod_int 1 by presburger finally show ?thesis by blast qed then show "\ (m * a) * unity_root n (int (m * a * k)) = \ (m * a) * unity_root n (int (m * k))" by auto qed also have "\ = (\m = 1..n . \(a) * (\(m) * unity_root n (m*k)))" by (rule sum.cong,simp,subst mult,simp) also have "\ = \(a) * (\m = 1..n . \(m) * unity_root n (m*k))" by (simp add: sum_distrib_left[symmetric]) also have "\ = \(a) * gauss_sum k" unfolding gauss_sum_def by blast finally have "gauss_sum k = \(a) * gauss_sum k" by blast then show "\ a = 1" using assms(1) by simp qed subsection \Induced moduli and primitive characters\ definition "induced_modulus d \ d dvd n \ (\a. coprime a n \ [a = 1] (mod d) \ \ a = 1)" lemma induced_modulus_dvd: "induced_modulus d \ d dvd n" unfolding induced_modulus_def by blast lemma induced_modulusI [intro?]: "d dvd n \ (\a. coprime a n \ [a = 1] (mod d) \ \ a = 1) \ induced_modulus d" unfolding induced_modulus_def by auto lemma induced_modulusD: "induced_modulus d \ coprime a n \ [a = 1] (mod d) \ \ a = 1" unfolding induced_modulus_def by blast lemma zero_not_ind_mod: "\induced_modulus 0" unfolding induced_modulus_def using n by simp lemma div_gcd_dvd1: "(a :: 'a :: semiring_gcd) div gcd a b dvd a" by (metis dvd_def dvd_div_mult_self gcd_dvd1) lemma div_gcd_dvd2: "(b :: 'a :: semiring_gcd) div gcd a b dvd b" by (metis div_gcd_dvd1 gcd.commute) lemma g_non_zero_ind_mod: assumes "gauss_sum k \ 0" "\coprime k n" "k > 0" shows "induced_modulus (n div gcd k n)" proof show "n div gcd k n dvd n" by (metis dvd_div_mult_self dvd_triv_left gcd.commute gcd_dvd1) fix a :: nat assume "coprime a n" "[a = 1] (mod n div gcd k n)" thus "\ a = 1" using assms n gauss_sum_nonzero_noncoprime_necessary_condition(3) by auto qed lemma induced_modulus_modulus: "induced_modulus n" unfolding induced_modulus_def proof (rule conjI,simp,safe) fix a assume "[a = 1] (mod n)" then have "a mod n = 1 mod n" using cong_def[of a 1 n] by blast also have "\ = 1" using eq_zero_iff zero_eq_0 by fastforce finally have 1: "a mod n = 1" by simp have "\ a = \ (a mod n)" by simp also have "\ = \ 1" using cong_def 1 by auto also have "\ = 1" by simp finally show "\ a = 1" by blast qed text \Theorem 8.13\ theorem one_induced_iff_principal: "induced_modulus 1 \ \ = principal_dchar n" proof assume "induced_modulus 1" then have "(\a. coprime a n \ \ a = 1)" unfolding induced_modulus_def by simp then show "\ = principal_dchar n" unfolding principal_dchar_def using eq_zero by auto next assume as: "\ = principal_dchar n" {fix a assume "coprime a n" then have "\ a = 1" using principal_dchar_def as by simp} then show "induced_modulus 1" unfolding induced_modulus_def by auto qed end locale primitive_dchar = dcharacter + assumes no_induced_modulus: "\(\ddprimitive_dchar n \" proof assume "primitive_dchar n \" then interpret A: primitive_dchar n "residue_mult_group n" \ by auto from A.no_induced_modulus induced_modulus show False by contradiction qed lemma (in dcharacter) primitive_dchar_iff: "primitive_dchar n \ \ \(\dprimitive_dchar n (principal_dchar n)" unfolding principal.primitive_dchar_iff using principal.one_induced_iff_principal n by auto lemma (in dcharacter) not_primitive_imp_nonprimitive: assumes "\primitive_dchar n \" shows "nonprimitive_dchar n \" using assms dcharacter_axioms unfolding nonprimitive_dchar_def primitive_dchar_def primitive_dchar_axioms_def nonprimitive_dchar_axioms_def by auto text \Theorem 8.14\ theorem (in dcharacter) prime_nonprincipal_is_primitive: assumes "prime n" assumes "\ \ principal_dchar n" shows "primitive_dchar n \" proof - {fix m assume "induced_modulus m" then have "m = n" using assms prime_nat_iff induced_modulus_def one_induced_iff_principal by blast} then show ?thesis using primitive_dchar_iff by blast qed text \Theorem 8.15\ corollary (in primitive_dchar) primitive_encoding: "\k>0. \coprime k n \ gauss_sum k = 0" "\k>0. separable k" "norm (gauss_sum 1) ^ 2 = n" proof safe show 1: "gauss_sum k = 0" if "k > 0" and "\coprime k n" for k proof (rule ccontr) assume "gauss_sum k \ 0" hence "induced_modulus (n div gcd k n)" using that by (intro g_non_zero_ind_mod) auto moreover have "n div gcd k n < n" using n that by (meson coprime_iff_gcd_eq_1 div_eq_dividend_iff le_less_trans linorder_neqE_nat nat_dvd_not_less principal.div_gcd_dvd2 zero_le_one) ultimately show False using no_induced_modulus by blast qed have "(\n>0. separable n)" unfolding global_separability_condition by (auto intro!: 1) thus "separable n" if "n > 0" for n using that by blast thus "norm (gauss_sum 1) ^ 2 = n" using gauss_sum_1_mod_square_eq_k by blast qed text \Theorem 8.16\ lemma (in dcharacter) induced_modulus_altdef1: "induced_modulus d \ d dvd n \ (\a b. coprime a n \ coprime b n \ [a = b] (mod d) \ \ a = \ b)" proof assume 1: "induced_modulus d" with n have d: "d dvd n" "d > 0" by (auto simp: induced_modulus_def intro: Nat.gr0I) show " d dvd n \ (\a b. coprime a n \ coprime b n \ [a = b] (mod d) \ \(a) = \(b))" proof safe fix a b assume 2: "coprime a n" "coprime b n" "[a = b] (mod d)" show "\(a) = \(b)" proof - from 2(1) obtain a' where eq: "[a*a' = 1] (mod n)" using cong_solve by blast from this d have "[a*a' = 1] (mod d)" using cong_dvd_modulus_nat by blast from this 1 have "\(a*a') = 1" unfolding induced_modulus_def by (meson "2"(2) eq cong_imp_coprime cong_sym coprime_divisors gcd_nat.refl one_dvd) then have 3: "\(a)*\(a') = 1" by simp from 2(3) have "[a*a' = b*a'] (mod d)" by (simp add: cong_scalar_right) moreover have 4: "[b*a' = 1] (mod d)" using \[a * a' = 1] (mod d)\ calculation cong_sym cong_trans by blast have "\(b*a') = 1" proof - have "coprime (b*a') n" using "2"(2) cong_imp_coprime[OF cong_sym[OF eq]] by simp then show ?thesis using 4 induced_modulus_def 1 by blast qed then have 4: "\(b)*\(a') = 1" by simp from 3 4 show "\(a) = \(b)" using mult_cancel_left by (cases "\(a') = 0") (fastforce simp add: field_simps)+ qed qed fact+ next assume *: "d dvd n \ (\a b. coprime a n \ coprime b n \ [a = b] (mod d) \ \ a = \ b)" then have "\a . coprime a n \ coprime 1 n \ [a = 1] (mod d) \ \ a = \ 1" by blast then have "\a . coprime a n \ [a = 1] (mod d) \ \ a = 1" using coprime_1_left by simp then show "induced_modulus d" unfolding induced_modulus_def using * by blast qed text \Exercise 8.4\ lemma induced_modulus_altdef2_lemma: fixes n a d q :: nat defines "q \ (\ p | prime p \ p dvd n \ \ (p dvd a). p)" defines "m \ a + q * d" assumes "n > 0" "coprime a d" shows "[m = a] (mod d)" and "coprime m n" proof (simp add: assms(2) cong_add_lcancel_0_nat cong_mult_self_right) have fin: "finite {p. prime p \ p dvd n \ \ (p dvd a)}" by (simp add: assms) { fix p assume 4: "prime p" "p dvd m" "p dvd n" have "p = 1" proof (cases "p dvd a") case True from this assms 4(2) have "p dvd q*d" by (simp add: dvd_add_right_iff) then have a1: "p dvd q \ p dvd d" using 4(1) prime_dvd_mult_iff by blast have a2: "\ (p dvd q)" proof (rule ccontr,simp) assume "p dvd q" then have "p dvd (\ p | prime p \ p dvd n \ \ (p dvd a). p)" unfolding assms by simp then have "\x\{p. prime p \ p dvd n \ \ p dvd a}. p dvd x" using prime_dvd_prod_iff[OF fin 4(1)] by simp then obtain x where c: "p dvd x \ prime x \ \ x dvd a" by blast then have "p = x" using 4(1) by (simp add: primes_dvd_imp_eq) then show "False" using True c by auto qed have a3: "\ (p dvd d)" using True assms "4"(1) coprime_def not_prime_unit by auto from a1 a2 a3 show ?thesis by simp next case False then have "p dvd q" proof - have in_s: "p \ {p. prime p \ p dvd n \ \ p dvd a}" using False 4(3) 4(1) by simp show "p dvd q" unfolding assms using dvd_prodI[OF fin in_s ] by fast qed then have "p dvd q*d" by simp then have "p dvd a" using 4(2) assms by (simp add: dvd_add_left_iff) then show ?thesis using False by auto qed } note lem = this show "coprime m n" proof (subst coprime_iff_gcd_eq_1) {fix a assume "a dvd m" "a dvd n" "a \ 1" {fix p assume "prime p" "p dvd a" then have "p dvd m" "p dvd n" using \a dvd m\ \a dvd n\ by auto from lem have "p = a" using not_prime_1 \prime p\ \p dvd m\ \p dvd n\ by blast} then have "prime a" using prime_prime_factor[of a] \a \ 1\ by blast then have "a = 1" using lem \a dvd m\ \a dvd n\ by blast then have "False" using \a = 1\ \a \ 1\ by blast } then show "gcd m n = 1" by blast qed qed text \Theorem 8.17\ text\The case \d = 1\ is exactly the case described in @{thm dcharacter.one_induced_iff_principal}.\ theorem (in dcharacter) induced_modulus_altdef2: assumes "d dvd n" "d \ 1" defines "\\<^sub>1 \ principal_dchar n" shows "induced_modulus d \ (\\. dcharacter d \ \ (\k. \ k = \ k * \\<^sub>1 k))" proof from n have n_pos: "n > 0" by simp assume as_im: "induced_modulus d" define f where "f \ (\k. k + (if k = 1 then 0 else (prod id {p. prime p \ p dvd n \ \ (p dvd k)})*d) )" have [simp]: "f (Suc 0) = 1" unfolding f_def by simp { fix k assume as: "coprime k d" hence "[f k = k] (mod d)" "coprime (f k) n" using induced_modulus_altdef2_lemma[OF n_pos as] by (simp_all add: f_def) } note m_prop = this define \ where "\ \ (\n. (if \ coprime n d then 0 else \(f n)))" have \_1: "\ 1 = 1" unfolding \_def by simp from assms(1,2) n have "d > 0" by (intro Nat.gr0I) auto from induced_modulus_altdef1 assms(1) \d > 0\ as_im have b: "(\a b. coprime a n \ coprime b n \ [a = b] (mod d) \ \ a = \ b)" by blast have \_periodic: " \a. \ (a + d) = \ a" proof fix a have "gcd (a+d) d = gcd a d" by auto then have cop: "coprime (a+d) d = coprime a d" using coprime_iff_gcd_eq_1 by presburger show "\ (a + d) = \ a" proof (cases "coprime a d") case True from True cop have cop_ad: "coprime (a+d) d" by blast have p1: "[f (a+d) = f a] (mod d)" using m_prop(1)[of "a+d", OF cop_ad] m_prop(1)[of "a",OF True] by (simp add: cong_def) have p2: "coprime (f (a+d)) n" "coprime (f a) n" using m_prop(2)[of "a+d", OF cop_ad] m_prop(2)[of "a", OF True] by blast+ from b p1 p2 have eq: "\ (f (a + d)) = \ (f a)" by blast show ?thesis unfolding \_def by (subst cop,simp,safe, simp add: eq) next case False then show ?thesis unfolding \_def by (subst cop,simp) qed qed have \_mult: "\a b. a \ totatives d \ b \ totatives d \ \ (a * b) = \ a * \ b" proof (safe) fix a b assume "a \ totatives d" "b \ totatives d" consider (ab) "coprime a d \ coprime b d" | (a) "coprime a d \ \ coprime b d" | (b) "coprime b d \ \ coprime a d" | (n) "\ coprime a d \ \ coprime b d" by blast then show "\ (a * b) = \ a * \ b" proof cases case ab then have c_ab: "coprime (a*b) d" "coprime a d" "coprime b d" by simp+ then have p1: "[f (a * b) = a * b] (mod d)" "coprime (f (a * b)) n" using m_prop[of "a*b", OF c_ab(1)] by simp+ moreover have p2: "[f a = a] (mod d)" "coprime (f a) n" "[f b = b] (mod d)" "coprime (f b) n" using m_prop[of "a",OF c_ab(2)] m_prop[of "b",OF c_ab(3) ] by simp+ have p1s: "[f (a * b) = (f a) * (f b)] (mod d)" proof - have "[f (a * b) = a * b] (mod d)" using p1(1) by blast moreover have "[a * b = f(a) * f(b)] (mod d)" using p2(1) p2(3) by (simp add: cong_mult cong_sym) ultimately show ?thesis using cong_trans by blast qed have p2s: "coprime (f a*f b) n" using p2(2) p2(4) by simp have "\ (f (a * b)) = \ (f a * f b)" using p1s p2s p1(2) b by blast then show ?thesis unfolding \_def by (simp add: c_ab) qed (simp_all add: \_def) qed have d_gr_1: "d > 1" using assms(1,2) using \0 < d\ by linarith show "\\. dcharacter d \ \ (\n. \ n = \ n * \\<^sub>1 n)" proof (standard,rule conjI) show "dcharacter d \" unfolding dcharacter_def residues_nat_def dcharacter_axioms_def using d_gr_1 \_def f_def \_mult \_1 \_periodic by simp show "\n. \ n = \ n * \\<^sub>1 n" proof fix k show "\ k = \ k * \\<^sub>1 k" proof (cases "coprime k n") case True then have "coprime k d" using assms(1) by auto then have "\(k) = \(f k)" using \_def by simp moreover have "[f k = k] (mod d)" using m_prop[OF \coprime k d\] by simp moreover have "\\<^sub>1 k = 1" using assms(3) principal_dchar_def \coprime k n\ by auto ultimately show "\(k) = \(k) * \\<^sub>1(k)" proof - assume "\ k = \ (f k)" "[f k = k] (mod d)" "\\<^sub>1 k = 1" then have "\ k = \ (f k)" using \local.induced_modulus d\ induced_modulus_altdef1 assms(1) \d > 0\ True \coprime k d\ m_prop(2) by auto also have "\ = \ k" by (simp add: \\ k = \ (f k)\) also have "\ = \ k * \\<^sub>1 k" by (simp add: \\\<^sub>1 k = 1\) finally show ?thesis by simp qed next case False hence "\ k = 0" using eq_zero_iff by blast moreover have "\\<^sub>1 k = 0" using False assms(3) principal_dchar_def by simp ultimately show ?thesis by simp qed qed qed next assume "(\\. dcharacter d \ \ (\k. \ k = \ k * \\<^sub>1 k))" then obtain \ where 1: "dcharacter d \" "(\k. \ k = \ k * \\<^sub>1 k)" by blast show "induced_modulus d" unfolding induced_modulus_def proof (rule conjI,fact,safe) fix k assume 2: "coprime k n" "[k = 1] (mod d)" then have "\\<^sub>1 k = 1" "\ k = 1" proof (simp add: assms(3) principal_dchar_def) have "\ k = \ (k mod d)" by (simp add: dcharacter.mod[OF 1(1), of k]) also have "\ = \ (1 mod d)" using cong_def[of k 1 d] 2(2) by simp also have "\ = \ 1" using assms(2) "1"(1) dcharacter.mod by blast also have "\ = 1" using dcharacter.Suc_0[OF 1(1)] by simp finally show "\ k = 1" by simp qed then show "\ k = 1" using 1(2) by simp qed qed subsection \The conductor of a character\ context dcharacter begin definition "conductor = Min {d. induced_modulus d}" lemma conductor_fin: "finite {d. induced_modulus d}" proof - let ?A = "{d. induced_modulus d}" have "?A \ {d. d dvd n}" unfolding induced_modulus_def by blast moreover have "finite {d. d dvd n}" using n by simp ultimately show "finite ?A" using finite_subset by auto qed lemma conductor_induced: "induced_modulus conductor" proof - have "{d. induced_modulus d} \ {}" using induced_modulus_modulus by blast then show "induced_modulus conductor" using Min_in[OF conductor_fin ] conductor_def by auto qed lemma conductor_le_iff: "conductor \ a \ (\d\a. induced_modulus d)" unfolding conductor_def using conductor_fin induced_modulus_modulus by (subst Min_le_iff) auto lemma conductor_ge_iff: "conductor \ a \ (\d. induced_modulus d \ d \ a)" unfolding conductor_def using conductor_fin induced_modulus_modulus by (subst Min_ge_iff) auto lemma conductor_leI: "induced_modulus d \ conductor \ d" by (subst conductor_le_iff) auto lemma conductor_geI: "(\d. induced_modulus d \ d \ a) \ conductor \ a" by (subst conductor_ge_iff) auto lemma conductor_dvd: "conductor dvd n" using conductor_induced unfolding induced_modulus_def by blast lemma conductor_le_modulus: "conductor \ n" using conductor_dvd by (rule dvd_imp_le) (use n in auto) lemma conductor_gr_0: "conductor > 0" unfolding conductor_def using zero_not_ind_mod using conductor_def conductor_induced neq0_conv by fastforce lemma conductor_eq_1_iff_principal: "conductor = 1 \ \ = principal_dchar n" proof assume "conductor = 1" then have "induced_modulus 1" using conductor_induced by auto then show "\ = principal_dchar n" using one_induced_iff_principal by blast next assume "\ = principal_dchar n" then have im_1: "induced_modulus 1" using one_induced_iff_principal by auto show "conductor = 1" proof - have "conductor \ 1" using conductor_fin Min_le[OF conductor_fin,simplified,OF im_1] by (simp add: conductor_def[symmetric]) then show ?thesis using conductor_gr_0 by auto qed qed lemma conductor_principal [simp]: "\ = principal_dchar n \ conductor = 1" by (subst conductor_eq_1_iff_principal) lemma nonprimitive_imp_conductor_less: assumes "\primitive_dchar n \" shows "conductor < n" proof - obtain d where d: "induced_modulus d" "d < n" using primitive_dchar_iff assms by blast from d(1) have "conductor \ d" by (rule conductor_leI) also have "\ < n" by fact finally show ?thesis . qed lemma (in nonprimitive_dchar) conductor_less_modulus: "conductor < n" using nonprimitive_imp_conductor_less nonprimitive by metis text \Theorem 8.18\ theorem primitive_principal_form: defines "\\<^sub>1 \ principal_dchar n" assumes "\ \ principal_dchar n" shows "\\. primitive_dchar conductor \ \ (\n. \(n) = \(n) * \\<^sub>1(n))" proof - (* TODO: perhaps residues_nat should be relaxed to allow n = 1. Then we could remove the unnecessary precondition here. It makes no real difference though. *) from n have n_pos: "n > 0" by simp define d where "d = conductor" have induced: "induced_modulus d" unfolding d_def using conductor_induced by blast then have d_not_1: "d \ 1" using one_induced_iff_principal assms by auto hence d_gt_1: "d > 1" using conductor_gr_0 by (auto simp: d_def) from induced obtain \ where \_def: "dcharacter d \ \ (\n. \ n = \ n * \\<^sub>1 n)" using d_not_1 by (subst (asm) induced_modulus_altdef2) (auto simp: d_def conductor_dvd \\<^sub>1_def) have phi_dchars: "\ \ dcharacters d" using \_def dcharacters_def by auto interpret \: dcharacter d "residue_mult_group d" \ using \_def by auto have \_prim: "primitive_dchar d \" proof (rule ccontr) assume "\ primitive_dchar d \" then obtain q where 1: "q dvd d \ q < d \ \.induced_modulus q" unfolding \.induced_modulus_def \.primitive_dchar_iff by blast then have 2: "induced_modulus q" proof - {fix k assume mod_1: "[k = 1] (mod q)" assume cop: "coprime k n" have "\(k) = \(k)*\\<^sub>1(k)" using \_def by auto also have "\ = \(k)" using cop by (simp add: assms principal_dchar_def) also have "\ = 1" using 1 mod_1 \.induced_modulus_def \induced_modulus d\ cop induced_modulus_def by auto finally have "\(k) = 1" by blast} then show ?thesis using induced_modulus_def "1" \induced_modulus d\ by auto qed from 1 have "q < d" by simp moreover have "d \ q" unfolding d_def by (intro conductor_leI) fact ultimately show False by linarith qed from \_def \_prim d_def phi_dchars show ?thesis by blast qed definition primitive_extension :: "nat \ complex" where "primitive_extension = (SOME \. primitive_dchar conductor \ \ (\k. \ k = \ k * principal_dchar n k))" lemma assumes nonprincipal: "\ \ principal_dchar n" shows primitive_primitive_extension: "primitive_dchar conductor primitive_extension" and principal_decomposition: "\ k = primitive_extension k * principal_dchar n k" proof - note * = someI_ex[OF primitive_principal_form[OF nonprincipal], folded primitive_extension_def] from * show "primitive_dchar conductor primitive_extension" by blast from * show "\ k = primitive_extension k * principal_dchar n k" by blast qed end subsection \The connection between primitivity and separability\ lemma residue_mult_group_coset: fixes m n m1 m2 :: nat and f :: "nat \ nat" and G H defines "G \ residue_mult_group n" defines "H \ residue_mult_group m" defines "f \ (\k. k mod m)" assumes "b \ (rcosets\<^bsub>G\<^esub> kernel G H f)" assumes "m1 \ b" "m2 \ b" assumes "n > 1" "m dvd n" shows "m1 mod m = m2 mod m" proof - have h_1: "\\<^bsub>H\<^esub> = 1" using assms(2) unfolding residue_mult_group_def totatives_def by simp from assms(4) obtain a :: nat where cos_expr: "b = (kernel G H f) #>\<^bsub>G\<^esub> a \ a \ carrier G" using RCOSETS_def[of G "kernel G H f"] by blast then have cop: "coprime a n" using assms(1) unfolding residue_mult_group_def totatives_def by auto obtain a' where "[a * a' = 1] (mod n)" using cong_solve_coprime_nat[OF cop] by auto then have a_inv: "(a*a') mod n = 1" using cong_def[of "a*a'" 1 n] assms(7) by simp have "m1 \ (\h\kernel G H f. {h \\<^bsub>G\<^esub> a})" "m2 \ (\h\kernel G H f. {h \\<^bsub>G\<^esub> a})" using r_coset_def[of G "kernel G H f" a] cos_expr assms(5,6) by blast+ then have "m1 \ (\h\kernel G H f. {(h * a) mod n})" "m2 \ (\h\kernel G H f. {(h * a) mod n})" using assms(1) unfolding residue_mult_group_def[of n] by auto then obtain m1' m2' where m_expr: "m1 = (m1'* a) mod n \ m1' \ kernel G H f" "m2 = (m2'* a) mod n \ m2' \ kernel G H f" by blast have eq_1: "m1 mod m = a mod m" proof - have "m1 mod m = ((m1'* a) mod n) mod m" using m_expr by blast also have "\ = (m1' * a) mod m" using euclidean_semiring_cancel_class.mod_mod_cancel assms(8) by blast also have "\ = (m1' mod m) * (a mod m) mod m" by (simp add: mod_mult_eq) also have "\ = (a mod m) mod m" using m_expr(1) h_1 unfolding kernel_def assms(3) by simp also have "\ = a mod m" by auto finally show ?thesis by simp qed have eq_2: "m2 mod m = a mod m" proof - have "m2 mod m = ((m2'* a) mod n) mod m" using m_expr by blast also have "\ = (m2' * a) mod m" using euclidean_semiring_cancel_class.mod_mod_cancel assms(8) by blast also have "\ = (m2' mod m) * (a mod m) mod m" by (simp add: mod_mult_eq) also have "\ = (a mod m) mod m" using m_expr(2) h_1 unfolding kernel_def assms(3) by simp also have "\ = a mod m" by auto finally show ?thesis by simp qed from eq_1 eq_2 show ?thesis by argo qed lemma residue_mult_group_kernel_partition: fixes m n :: nat and f :: "nat \ nat" and G H defines "G \ residue_mult_group n" defines "H \ residue_mult_group m" defines "f \ (\k. k mod m)" assumes "m > 1" "n > 0" "m dvd n" shows "partition (carrier G) (rcosets\<^bsub>G\<^esub> kernel G H f)" and "card (rcosets\<^bsub>G\<^esub> kernel G H f) = totient m" and "card (kernel G H f) = totient n div totient m" and "b \(rcosets\<^bsub>G\<^esub> kernel G H f) \ b \ {}" and "b \(rcosets\<^bsub>G\<^esub> kernel G H f) \ card (kernel G H f) = card b" and "bij_betw (\b. (the_elem (f ` b))) (rcosets\<^bsub>G\<^esub> kernel G H f) (carrier H)" proof - have "1 < m" by fact also have "m \ n" using \n > 0\ \m dvd n\ by (intro dvd_imp_le) auto finally have "n > 1" . note mn = \m > 1\ \n > 1\ \m dvd n\ \m \ n\ interpret n: residues_nat n G using mn by unfold_locales (auto simp: assms) interpret m: residues_nat m H using mn by unfold_locales (auto simp: assms) from mn have subset: "f ` carrier G \ carrier H" by (auto simp: assms(1-3) residue_mult_group_def totatives_def dest: coprime_common_divisor_nat intro!: Nat.gr0I) moreover have super_set: "carrier H \ f ` carrier G" proof safe fix k assume "k \ carrier H" hence k: "k > 0" "k \ m" "coprime k m" by (auto simp: assms(2) residue_mult_group_def totatives_def) from mn \k \ carrier H\ have "k < m" by (simp add: totatives_less assms(2) residue_mult_group_def) define P where "P = {p \ prime_factors n. \(p dvd m)}" define a where "a = \P" have [simp]: "a \ 0" by (auto simp: P_def a_def intro!: Nat.gr0I) have [simp]: "prime_factors a = P" proof - have "prime_factors a = set_mset (sum prime_factorization P)" unfolding a_def using mn by (subst prime_factorization_prod) (auto simp: P_def prime_factors_dvd prime_gt_0_nat) also have "sum prime_factorization P = (\p\P. {#p#})" using mn by (intro sum.cong) (auto simp: P_def prime_factorization_prime prime_factors_dvd) finally show ?thesis by (simp add: P_def) qed from mn have "coprime m a" by (subst coprime_iff_prime_factors_disjoint) (auto simp: P_def) hence "\x. [x = k] (mod m) \ [x = 1] (mod a)" by (intro binary_chinese_remainder_nat) then obtain x where x: "[x = k] (mod m)" "[x = 1] (mod a)" by auto from x(1) mn k have [simp]: "x \ 0" using coprime_common_divisor[of k m] by (auto intro!: Nat.gr0I simp: cong_def) from x(2) have "coprime x a" using cong_imp_coprime cong_sym by force hence "coprime x (a * m)" using k cong_imp_coprime[OF cong_sym[OF x(1)]] by auto also have "?this \ coprime x n" using mn by (intro coprime_cong_prime_factors) (auto simp: prime_factors_product P_def in_prime_factors_iff) finally have "x mod n \ totatives n" using mn by (auto simp: totatives_def intro!: Nat.gr0I) moreover have "f (x mod n) = k" using x(1) k mn \k < m\ by (auto simp: assms(3) cong_def mod_mod_cancel) ultimately show "k \ f ` carrier G" by (auto simp: assms(1) residue_mult_group_def) qed ultimately have image_eq: "f ` carrier G = carrier H" by blast have [simp]: "f (k \\<^bsub>G\<^esub> l) = f k \\<^bsub>H\<^esub> f l" if "k \ carrier G" "l \ carrier G" for k l using that mn by (auto simp: assms(1-3) residue_mult_group_def totatives_def mod_mod_cancel mod_mult_eq) interpret f: group_hom G H f using subset by unfold_locales (auto simp: hom_def) show "bij_betw (\b. (the_elem (f ` b))) (rcosets\<^bsub>G\<^esub> kernel G H f) (carrier H)" unfolding bij_betw_def proof show "inj_on (\b. (the_elem (f ` b))) (rcosets\<^bsub>G\<^esub> kernel G H f)" using f.FactGroup_inj_on unfolding FactGroup_def by auto have eq: "f ` carrier G = carrier H" using subset super_set by blast show "(\b. the_elem (f ` b)) ` (rcosets\<^bsub>G\<^esub> kernel G H f) = carrier H" using f.FactGroup_onto[OF eq] unfolding FactGroup_def by simp qed show "partition (carrier G) (rcosets\<^bsub>G\<^esub> kernel G H f)" proof show "\a. a \ carrier G \ \!b. b \ rcosets\<^bsub>G\<^esub> kernel G H f \ a \ b" proof - fix a assume a_in: "a \ carrier G" show "\!b. b \ rcosets\<^bsub>G\<^esub> kernel G H f \ a \ b" proof - (*exists*) have "\b. b \ rcosets\<^bsub>G\<^esub> kernel G H f \ a \ b" using a_in n.rcosets_part_G[OF f.subgroup_kernel] by blast then show ?thesis using group.rcos_disjoint[OF n.is_group f.subgroup_kernel] by (auto simp: disjoint_def) qed qed next show "\b. b \ rcosets\<^bsub>G\<^esub> kernel G H f \ b \ carrier G" using n.rcosets_part_G f.subgroup_kernel by auto qed (* sizes *) have lagr: "card (carrier G) = card (rcosets\<^bsub>G\<^esub> kernel G H f) * card (kernel G H f)" using group.lagrange_finite[OF n.is_group n.fin f.subgroup_kernel] Coset.order_def[of G] by argo have k_size: "card (kernel G H f) > 0" using f.subgroup_kernel finite_subset n.subgroupE(1) n.subgroupE(2) by fastforce have G_size: "card (carrier G) = totient n" using n.order Coset.order_def[of G] by simp have H_size: " totient m = card (carrier H)" using n.order Coset.order_def[of H] by simp also have "\ = card (carrier (G Mod kernel G H f))" using f.FactGroup_iso[OF image_eq] card_image f.FactGroup_inj_on f.FactGroup_onto image_eq by fastforce also have "\ = card (carrier G) div card (kernel G H f)" proof - have "card (carrier (G Mod kernel G H f)) = card (rcosets\<^bsub>G\<^esub> kernel G H f)" unfolding FactGroup_def by simp also have "\ = card (carrier G) div card (kernel G H f)" by (simp add: lagr k_size) finally show ?thesis by blast qed also have "\ = totient n div card (kernel G H f)" using G_size by argo finally have eq: "totient m = totient n div card (kernel G H f)" by simp show "card (kernel G H f) = totient n div totient m" proof - have "totient m \ 0" using totient_0_iff[of m] assms(4) by blast have "card (kernel G H f) dvd totient n" using lagr \card (carrier G) = totient n\ by auto have "totient m * card (kernel G H f) = totient n" unfolding eq using \card (kernel G H f) dvd totient n\ by auto have "totient n div totient m = totient m * card (kernel G H f) div totient m" using \totient m * card (kernel G H f) = totient n\ by auto also have "\ = card (kernel G H f)" using nonzero_mult_div_cancel_left[OF \totient m \ 0\] by blast finally show ?thesis by auto qed show "card (rcosets\<^bsub>G\<^esub> kernel G H f) = totient m" proof - have H_size: " totient m = card (carrier H)" using n.order Coset.order_def[of H] by simp also have "\ = card (carrier (G Mod kernel G H f))" using f.FactGroup_iso[OF image_eq] card_image f.FactGroup_inj_on f.FactGroup_onto image_eq by fastforce also have "card (carrier (G Mod kernel G H f)) = card (rcosets\<^bsub>G\<^esub> kernel G H f)" unfolding FactGroup_def by simp finally show "card (rcosets\<^bsub>G\<^esub> kernel G H f) = totient m" by argo qed assume "b \ rcosets\<^bsub>G\<^esub> kernel G H f" then show "b \ {}" proof - have "card b = card (kernel G H f)" using \b \ rcosets\<^bsub>G\<^esub> kernel G H f\ f.subgroup_kernel n.card_rcosets_equal n.subgroupE(1) by auto then have "card b > 0" by (simp add: k_size) then show ?thesis by auto qed assume b_cos: "b \ rcosets\<^bsub>G\<^esub> kernel G H f" show "card (kernel G H f) = card b" using group.card_rcosets_equal[OF n.is_group b_cos] f.subgroup_kernel subgroup.subset by blast qed lemma primitive_iff_separable_lemma: assumes prod: "(\n. \ n = \ n * \\<^sub>1 n) \ primitive_dchar d \" assumes \d > 1\ \0 < k\ \d dvd k\ \k > 1\ shows "(\m | m \ {1..k} \ coprime m k. \(m) * unity_root d m) = (totient k div totient d) * (\m | m \ {1..d} \ coprime m d. \(m) * unity_root d m)" proof - from assms interpret \: primitive_dchar d "residue_mult_group d" \ by auto define G where "G = residue_mult_group k" define H where "H = residue_mult_group d" define f where "f = (\t. t mod d)" from residue_mult_group_kernel_partition(2)[OF \d > 1\ \0 < k\ \d dvd k\] have fin_cosets: "finite (rcosets\<^bsub>G\<^esub> kernel G H f)" using \1 < d\ card.infinite by (fastforce simp: G_def H_def f_def) have fin_G: "finite (carrier G)" unfolding G_def residue_mult_group_def by simp have eq: "(\m | m \ {1..k} \ coprime m k. \(m) * unity_root d m) = (\m | m \ carrier G . \(m) * unity_root d m)" unfolding residue_mult_group_def totatives_def G_def by (rule sum.cong,auto) also have "\ = sum (\m. \(m) * unity_root d m) (carrier G)" by simp also have eq': "\ = sum (sum (\m. \ m * unity_root d (int m))) (rcosets\<^bsub>G\<^esub> kernel G H f)" by (rule disjoint_sum [symmetric]) (use fin_G fin_cosets residue_mult_group_kernel_partition(1)[OF \d > 1\ \k > 0\ \d dvd k\] in \auto simp: G_def H_def f_def\) also have "\ = (\b \ (rcosets\<^bsub>G\<^esub> kernel G H f) . (\m \ b. \ m * unity_root d (int m)))" by simp finally have 1: "(\m | m \ {1..k} \ coprime m k. \(m) * unity_root d m) = (\b \ (rcosets\<^bsub>G\<^esub> kernel G H f) . (\m \ b. \ m * unity_root d (int m)))" using eq eq' by auto have eq''': "\ = (\b \ (rcosets\<^bsub>G\<^esub> kernel G H f) . (totient k div totient d) * (\ (the_elem (f ` b)) * unity_root d (int (the_elem (f ` b)))))" proof (rule sum.cong,simp) fix b assume b_in: "b \ (rcosets\<^bsub>G\<^esub> kernel G H f)" note b_not_empty = residue_mult_group_kernel_partition(4) [OF \d > 1\ \0 < k\ \d dvd k\ b_in[unfolded G_def H_def f_def]] { fix m1 m2 assume m_in: "m1 \ b" "m2 \ b" have m_mod: "m1 mod d = m2 mod d" using residue_mult_group_coset[OF b_in[unfolded G_def H_def f_def] m_in \k > 1\ \d dvd k\] by blast } note m_mod = this { fix m1 m2 assume m_in: "m1 \ b" "m2 \ b" have "\ m1 * unity_root d (int m1) = \ m2 * unity_root d (int m2)" proof - have \_periodic: "periodic_arithmetic \ d" using \.dir_periodic_arithmetic by blast have 1: "\ m1 = \ m2" using mod_periodic_arithmetic[OF \periodic_arithmetic \ d\ m_mod[OF m_in]] by simp have 2: "unity_root d m1 = unity_root d m2" using m_mod[OF m_in] by (intro unity_root_cong) (auto simp: cong_def simp flip: zmod_int) from 1 2 show ?thesis by simp qed } note all_eq_in_coset = this from all_eq_in_coset b_not_empty obtain l where l_prop: "l \ b \ (\y \ b. \ y * unity_root d (int y) = \ l * unity_root d (int l))" by blast have "(\m \ b. \ m * unity_root d (int m)) = ((totient k div totient d) * (\ l * unity_root d (int l)))" proof - have "(\m \ b. \ m * unity_root d (int m)) = (\m \ b. \ l * unity_root d (int l))" by (rule sum.cong,simp) (use all_eq_in_coset l_prop in blast) also have "\ = card b * \ l * unity_root d (int l)" by simp also have "\ = (totient k div totient d) * \ l * unity_root d (int l)" using residue_mult_group_kernel_partition(3)[OF \d > 1\ \0 < k\ \d dvd k\] residue_mult_group_kernel_partition(5) [OF \d > 1\ \0 < k\ \d dvd k\ b_in [unfolded G_def H_def f_def]] by argo finally have 2: "(\m \ b. \ m * unity_root d (int m)) = (totient k div totient d) * \ l * unity_root d (int l)" by blast from b_not_empty 2 show ?thesis by auto qed also have "\ = ((totient k div totient d) * (\ (the_elem (f ` b)) * unity_root d (int (the_elem (f ` b)))))" proof - have foral: "(\y. y \ b \ f y = f l)" using m_mod l_prop unfolding f_def by blast have eq: "the_elem (f ` b) = f l" using the_elem_image_unique[of _ f l, OF b_not_empty foral] by simp have per: "periodic_arithmetic \ d" using prod \.dir_periodic_arithmetic by blast show ?thesis unfolding eq using mod_periodic_arithmetic[OF per, of "l mod d" l] by (auto simp: f_def unity_root_mod zmod_int) qed finally show "(\m \ b. \ m * unity_root d (int m)) = ((totient k div totient d) * (\ (the_elem (f ` b)) * unity_root d (int (the_elem (f ` b)))))" by blast qed have "\ = (\b \ (rcosets\<^bsub>G\<^esub> kernel G H f) . (totient k div totient d) * (\ (the_elem (f ` b)) * unity_root d (int (the_elem (f ` b)))))" by blast also have eq'': " \ = (\h \ carrier H . (totient k div totient d) * (\ (h) * unity_root d (int (h))))" unfolding H_def G_def f_def by (rule sum.reindex_bij_betw[OF residue_mult_group_kernel_partition(6)[OF \d > 1\ \0 < k\ \d dvd k\]]) finally have 2: "(\m | m \ {1..k} \ coprime m k. \(m) * unity_root d m) = (totient k div totient d)*(\h \ carrier H . (\ (h) * unity_root d (int (h))))" using 1 by (simp add: eq'' eq''' sum_distrib_left) also have "\ = (totient k div totient d)*(\m | m \ {1..d} \ coprime m d . (\ (m) * unity_root d (int (m))))" unfolding H_def residue_mult_group_def by (simp add: totatives_def Suc_le_eq) finally show ?thesis by simp qed text \Theorem 8.19\ theorem (in dcharacter) primitive_iff_separable: "primitive_dchar n \ \ (\k>0. separable k)" proof (cases "\ = principal_dchar n") case True thus ?thesis using principal_not_primitive principal_not_totally_separable by auto next case False note nonprincipal = this show ?thesis proof assume "primitive_dchar n \" then interpret A: primitive_dchar n "residue_mult_group n" \ by auto show "(\k. k > 0 \ separable k)" using n A.primitive_encoding(2) by blast next assume tot_separable: "\k>0. separable k" { assume as: "\ primitive_dchar n \" have "\r. r \ 0 \ \ coprime r n \ gauss_sum r \ 0" proof - from n have "n > 0" by simp define d where "d = conductor" have "d > 0" unfolding d_def using conductor_gr_0 . then have "d > 1" using nonprincipal d_def conductor_eq_1_iff_principal by auto have "d < n" unfolding d_def using nonprimitive_imp_conductor_less[OF as] . have "d dvd n" unfolding d_def using conductor_dvd by blast define r where "r = n div d" have 0: "r \ 0" unfolding r_def using \0 < n\ \d dvd n\ dvd_div_gt0 by auto have "gcd r n > 1" unfolding r_def proof - have "n div d > 1" using \1 < n\ \d < n\ \d dvd n\ by auto have "n div d dvd n" using \d dvd n\ by force have "gcd (n div d) n = n div d" using gcd_nat.absorb1[OF \n div d dvd n\] by blast then show "1 < gcd (n div d) n" using \n div d > 1\ by argo qed then have 1: "\ coprime r n" by auto define \\<^sub>1 where "\\<^sub>1 = principal_dchar n" from primitive_principal_form[OF nonprincipal] obtain \ where prod: "(\k. \(k) = \(k)*\\<^sub>1(k)) \ primitive_dchar d \" using d_def unfolding \\<^sub>1_def by blast then have prod1: "(\k. \(k) = \(k)*\\<^sub>1(k))" "primitive_dchar d \" by blast+ then interpret \: primitive_dchar d "residue_mult_group d" \ by auto have "gauss_sum r = (\m = 1..n . \(m) * unity_root n (m*r))" unfolding gauss_sum_def by blast also have "\ = (\m = 1..n . \(m)*\\<^sub>1(m) * unity_root n (m*r))" by (rule sum.cong,auto simp add: prod) also have "\ = (\m | m \ {1..n} \ coprime m n. \(m)*\\<^sub>1(m) * unity_root n (m*r))" by (intro sum.mono_neutral_right) (auto simp: \\<^sub>1_def principal_dchar_def) also have "\ = (\m | m \ {1..n} \ coprime m n. \(m)*\\<^sub>1(m) * unity_root d m)" proof (rule sum.cong,simp) fix x assume "x \ {m \ {1..n}. coprime m n}" have "unity_root n (int (x * r)) = unity_root d (int x)" using unity_div_num[OF \n > 0\ \d > 0\ \d dvd n\] by (simp add: algebra_simps r_def) then show "\ x * \\<^sub>1 x * unity_root n (int (x * r)) = \ x * \\<^sub>1 x * unity_root d (int x)" by auto qed also have "\ = (\m | m \ {1..n} \ coprime m n. \(m) * unity_root d m)" by (rule sum.cong,auto simp add: \\<^sub>1_def principal_dchar_def) also have "\ = (totient n div totient d) * (\m | m \ {1..d} \ coprime m d. \(m) * unity_root d m)" using primitive_iff_separable_lemma[OF prod \d > 1\ \n > 0\ \d dvd n\ \n > 1\] by blast also have "\ = (totient n div totient d) * \.gauss_sum 1" proof - have "\.gauss_sum 1 = (\m = 1..d . \ m * unity_root d (int (m )))" by (simp add: \.gauss_sum_def) also have "\ = (\m | m \ {1..d} . \ m * unity_root d (int m))" by (rule sum.cong,auto) also have "\ = (\m | m \ {1..d} \ coprime m d. \(m) * unity_root d m)" by (rule sum.mono_neutral_right) (use \.eq_zero in auto) finally have "\.gauss_sum 1 = (\m | m \ {1..d} \ coprime m d. \(m) * unity_root d m)" by blast then show ?thesis by metis qed finally have g_expr: "gauss_sum r = (totient n div totient d) * \.gauss_sum 1" by blast have t_non_0: "totient n div totient d \ 0" by (simp add: \0 < n\ \d dvd n\ dvd_div_gt0 totient_dvd) have "(norm (\.gauss_sum 1))\<^sup>2 = d" using \.primitive_encoding(3) by simp then have "\.gauss_sum 1 \ 0" using \0 < d\ by auto then have 2: "gauss_sum r \ 0" using g_expr t_non_0 by auto from 0 1 2 show "\r. r \ 0 \ \ coprime r n \ gauss_sum r \ 0" by blast qed } note contr = this show "primitive_dchar n \" proof (rule ccontr) assume "\ primitive_dchar n \" then obtain r where 1: "r \ 0 \ \ coprime r n \ gauss_sum r \ 0" using contr by blast from global_separability_condition tot_separable have 2: "(\k>0. \ coprime k n \ gauss_sum k = 0)" by blast from 1 2 show "False" by blast qed qed qed text\Theorem 8.20\ theorem (in primitive_dchar) fourier_primitive: includes no_vec_lambda_notation fixes \ :: complex defines "\ \ gauss_sum 1 / sqrt n" shows "\ m = \ / sqrt n * (\k=1..n. cnj (\ k) * unity_root n (-m*k))" and "norm \ = 1" proof - have chi_not_principal: "\ \ principal_dchar n" using principal_not_totally_separable primitive_encoding(2) by blast then have case_0: "(\k=1..n. \ k) = 0" proof - have "sum \ {0..n-1} = sum \ {1..n}" using periodic_arithmetic_sum_periodic_arithmetic_shift[OF dir_periodic_arithmetic, of 1] n by auto also have "{0..n-1} = {..n = 1..n . \ n) = 0" using sum_dcharacter_block chi_not_principal by simp qed have "\ m = (\k = 1..n. 1 / of_nat n * gauss_sum_int (- int k) * unity_root n (int (m * k)))" using dcharacter_fourier_expansion[of m] by auto also have "\ = (\k = 1..n. 1 / of_nat n * gauss_sum (nat ((- k) mod n)) * unity_root n (int (m * k)))" by (auto simp: gauss_sum_int_conv_gauss_sum) also have "\ = (\k = 1..n. 1 / of_nat n * cnj (\ (nat ((- k) mod n))) * gauss_sum 1 * unity_root n (int (m * k)))" proof (rule sum.cong,simp) fix k assume "k \ {1..n}" have "gauss_sum (nat (- int k mod int n)) = cnj (\ (nat (- int k mod int n))) * gauss_sum 1" proof (cases "nat ((- k) mod n) > 0") case True then show ?thesis using mp[OF spec[OF primitive_encoding(2)] True] unfolding separable_def by auto next case False then have nat_0: "nat ((- k) mod n) = 0" by blast show ?thesis proof - have "gauss_sum (nat (- int k mod int n)) = gauss_sum 0" using nat_0 by argo also have "\ = (\m = 1..n. \ m)" unfolding gauss_sum_def by (rule sum.cong) auto also have "\ = 0" using case_0 by blast finally have 1: "gauss_sum (nat (- int k mod int n)) = 0" by blast have 2: "cnj (\ (nat (- int k mod int n))) = 0" using nat_0 zero_eq_0 by simp show ?thesis using 1 2 by simp qed qed then show "1 / of_nat n * gauss_sum (nat (- int k mod int n)) * unity_root n (int (m * k)) = 1 / of_nat n * cnj (\ (nat (- int k mod int n))) * gauss_sum 1 * unity_root n (int (m * k))" by auto qed also have "\ = (\k = 1..n. 1 / of_nat n * cnj (\ (nat (- int k mod int n))) * gauss_sum 1 * unity_root n (int (m * (nat (int k mod int n)))))" proof (rule sum.cong,simp) fix x assume "x \ {1..n}" have "unity_root n (m * x) = unity_root n (m * x mod n)" using unity_root_mod_nat[of n "m*x"] by (simp add: nat_mod_as_int) also have "\ = unity_root n (m * (x mod n))" by (rule unity_root_cong) (auto simp: cong_def mod_mult_right_eq simp flip: zmod_int of_nat_mult) finally have "unity_root n (m * x) = unity_root n (m * (x mod n))" by blast then show "1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * unity_root n (int (m * x)) = 1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * unity_root n (int (m * nat (int x mod int n)))" by (simp add: nat_mod_as_int) qed also have "\ = (\k = 0..n-1. 1 / of_nat n * cnj (\ k) * gauss_sum 1 * unity_root n (- int (m * k)))" proof - have b: "bij_betw (\k. nat((-k) mod n)) {1..n} {0..n-1}" unfolding bij_betw_def proof show "inj_on (\k. nat (- int k mod int n)) {1..n}" unfolding inj_on_def proof (safe) fix x y assume a1: "x \ {1..n}" "y \ {1..n}" assume a2: "nat (- x mod n) = nat (- y mod n)" then have "(- x) mod n = - y mod n" using n eq_nat_nat_iff by auto then have "[-int x = - int y] (mod n)" using cong_def by blast then have "[x = y] (mod n)" by (simp add: cong_int_iff cong_minus_minus_iff) then have cong: "x mod n = y mod n" using cong_def by blast then show "x = y" proof (cases "x = n") case True then show ?thesis using cong a1(2) by auto next case False then have "x mod n = x" using a1(1) by auto then have "y \ n" using a1(1) local.cong by fastforce then have "y mod n = y" using a1(2) by auto then show ?thesis using \x mod n = x\ cong by linarith qed qed show "(\k. nat (- int k mod int n)) ` {1..n} = {0..n - 1}" unfolding image_def proof let ?A = "{y. \x\{1..n}. y = nat (- int x mod int n)}" let ?B = "{0..n - 1}" show "?A \ ?B" proof fix y assume "y \ {y. \x\{1..n}. y = nat (- int x mod int n)}" then obtain x where "x\{1..n} \ y = nat (- int x mod int n)" by blast then show "y \ {0..n - 1}" by (simp add: nat_le_iff of_nat_diff) qed show "?A \ ?B" proof fix x assume 1: "x \ {0..n-1}" then have "n - x \ {1..n}" using n by auto have "x = nat (- int (n-x) mod int n)" proof - have "nat (- int (n-x) mod int n) = nat (int x) mod int n" apply(simp add: int_ops(6),rule conjI) using \n - x \ {1..n}\ by force+ also have "\ = x" using 1 n by auto finally show ?thesis by presburger qed then show "x \ {y. \x\{1..n}. y = nat (- int x mod int n)}" using \n - x \ {1..n}\ by blast qed qed qed show ?thesis proof - have 1: "(\k = 1..n. 1 / of_nat n * cnj (\ (nat (- int k mod int n))) * gauss_sum 1 * unity_root n (int (m * nat (int k mod int n)))) = (\x = 1..n. 1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * unity_root n (- int (m * nat (- int x mod int n))))" proof (rule sum.cong,simp) fix x have "(int m * (int x mod int n)) mod n = (m*x) mod n" by (simp add: mod_mult_right_eq zmod_int) also have "\ = (- ((- int (m*x) mod n))) mod n" by (simp add: mod_minus_eq of_nat_mod) have "(int m * (int x mod int n)) mod n = (- (int m * (- int x mod int n))) mod n" apply(subst mod_mult_right_eq,subst add.inverse_inverse[symmetric],subst (5) add.inverse_inverse[symmetric]) by (subst minus_mult_minus,subst mod_mult_right_eq[symmetric],auto) then have "unity_root n (int m * (int x mod int n)) = unity_root n (- (int m * (- int x mod int n)))" using unity_root_mod[of n "int m * (int x mod int n)"] unity_root_mod[of n " - (int m * (- int x mod int n))"] by argo then show "1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * unity_root n (int (m * nat (int x mod int n))) = 1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * - unity_root n (- int (m * nat (- int x mod int n)))" by auto + unity_root n (- int (m * nat (- int x mod int n)))" + by clarsimp qed also have 2: "(\x = 1..n. 1 / of_nat n * cnj (\ (nat (- int x mod int n))) * gauss_sum 1 * unity_root n (- int (m * nat (- int x mod int n)))) = (\md = 0..n - 1. 1 / of_nat n * cnj (\ md) * gauss_sum 1 * unity_root n (- int (m * md)))" using sum.reindex_bij_betw[OF b, of "\md. 1 / of_nat n * cnj (\ md) * gauss_sum 1 * unity_root n (- int (m * md))"] by blast also have 3: "\ = (\k = 0..n - 1. 1 / of_nat n * cnj (\ k) * gauss_sum 1 * unity_root n (- int (m * k)))" by blast finally have "(\k = 1..n. 1 / of_nat n * cnj (\ (nat (- int k mod int n))) * gauss_sum 1 * unity_root n (int (m * nat (int k mod int n)))) = (\k = 0..n - 1. 1 / of_nat n * cnj (\ k) * gauss_sum 1 * unity_root n (- int (m * k)))" using 1 2 3 by argo then show ?thesis by blast qed qed also have "\ = (\k = 1..n. 1 / of_nat n * cnj (\ k) * gauss_sum 1 * unity_root n (- int (m * k)))" proof - let ?f = "(\k. 1 / of_nat n * cnj (\ k) * gauss_sum 1 * unity_root n (- int (m * k)))" have "?f 0 = 0" using zero_eq_0 by auto have "?f n = 0" using zero_eq_0 mod_periodic_arithmetic[OF dir_periodic_arithmetic, of n 0] by simp have "(\n = 0..n - 1. ?f n) = (\n = 1..n - 1. ?f n)" using sum_shift_lb_Suc0_0[of ?f, OF \?f 0 = 0\] by auto also have "\ = (\n = 1..n. ?f n)" proof (rule sum.mono_neutral_left,simp,simp,safe) fix i assume "i \ {1..n}" "i \ {1..n - 1}" then have "i = n" using n by auto then show "1 / of_nat n * cnj (\ i) * gauss_sum 1 * unity_root n (- int (m * i)) = 0" using \?f n = 0\ by blast qed finally show ?thesis by blast qed also have "\ = (\k = 1..n. (\ / sqrt n) * cnj (\ k) * unity_root n (- int (m * k)))" proof (rule sum.cong,simp) fix x assume "x \ {1..n}" have "\ / sqrt (real n) = 1 / of_nat n * gauss_sum 1" proof - have "\ / sqrt (real n) = gauss_sum 1 / sqrt n / sqrt n" using assms by auto also have "\ = gauss_sum 1 / (sqrt n * sqrt n)" by (subst divide_divide_eq_left,subst of_real_mult,blast) also have "\ = gauss_sum 1 / n" using real_sqrt_mult_self by simp finally show ?thesis by simp qed then show "1 / of_nat n * cnj (\ x) * gauss_sum 1 * unity_root n (- int (m * x)) = (\ / sqrt n) * cnj (\ x) * unity_root n (- int (m * x))" by simp qed also have "\ = \ / sqrt (real n) * (\k = 1..n. cnj (\ k) * unity_root n (- int (m * k)))" proof - have "(\k = 1..n. \ / sqrt (real n) * cnj (\ k) * unity_root n (- int (m * k))) = (\k = 1..n. \ / sqrt (real n) * (cnj (\ k) * unity_root n (- int (m * k))))" by (rule sum.cong,simp, simp add: algebra_simps) also have "\ = \ / sqrt (real n) * (\k = 1..n. cnj (\ k) * unity_root n (- int (m * k)))" by (rule sum_distrib_left[symmetric]) finally show ?thesis by blast qed finally show "\ m = (\ / sqrt (real n)) * (\k=1..n. cnj (\ k) * unity_root n (- int m * int k))" by simp have 1: "norm (gauss_sum 1) = sqrt n" using gauss_sum_1_mod_square_eq_k[OF primitive_encoding(2)] by (simp add: cmod_def) from assms have 2: "norm \ = norm (gauss_sum 1) / \sqrt n\" by (simp add: norm_divide) show "norm \ = 1" using 1 2 n by simp qed unbundle vec_lambda_notation end \ No newline at end of file diff --git a/thys/Group-Ring-Module/Algebra4.thy b/thys/Group-Ring-Module/Algebra4.thy --- a/thys/Group-Ring-Module/Algebra4.thy +++ b/thys/Group-Ring-Module/Algebra4.thy @@ -1,6922 +1,6922 @@ (** Algebra4 author Hidetsune Kobayashi Lingjun Chen (part of Chap 4. section 2, with revision by H. Kobayashi) Group You Santo Department of Mathematics Nihon University h_coba@math.cst.nihon-u.ac.jp May 3, 2004. April 6, 2007 (revised) chapter 3. Group Theory. Focused on Jordan Hoelder theorem (continued) section 20. abelian groups subsection 20-1. Homomorphism of abelian groups subsection 20-2 quotient abelian group section 21 direct product and direct sum of abelian groups, in general case chapter 4. Ring theory section 1. Definition of a ring and an ideal section 2. Calculation of elements section 3. ring homomorphisms section 4. quotient rings section 5. primary ideals, prime ideals **) theory Algebra4 imports Algebra3 begin section "Abelian groups" record 'a aGroup = "'a carrier" + pop :: "['a, 'a ] \ 'a" (infixl "\\" 62) mop :: "'a \ 'a" ("(-\<^sub>a\ _)" [64]63 ) zero :: "'a" ("\\") locale aGroup = fixes A (structure) assumes pop_closed: "pop A \ carrier A \ carrier A \ carrier A" and aassoc : "\a \ carrier A; b \ carrier A; c \ carrier A\ \ (a \ b) \ c = a \ (b \ c)" and pop_commute:"\a \ carrier A; b \ carrier A\ \ a \ b = b \ a" and mop_closed:"mop A \ carrier A \ carrier A" and l_m :"a \ carrier A \ (-\<^sub>a a) \ a = \" and ex_zero: "\ \ carrier A" and l_zero:"a \ carrier A \ \ \ a = a" definition b_ag :: "_ \ \carrier:: 'a set, top:: ['a, 'a] \ 'a , iop:: 'a \ 'a, one:: 'a \" where "b_ag A = \carrier = carrier A, top = pop A, iop = mop A, one = zero A \" definition asubGroup :: "[_ , 'a set] \ bool" where "asubGroup A H \ (b_ag A) \ H" definition aqgrp :: "[_ , 'a set] \ \ carrier::'a set set, pop::['a set, 'a set] \ 'a set, mop::'a set \ 'a set, zero :: 'a set \" where "aqgrp A H = \carrier = set_rcs (b_ag A) H, pop = \X. \Y. (c_top (b_ag A) H X Y), mop = \X. (c_iop (b_ag A) H X), zero = H \" definition ag_idmap :: "_ \ ('a \ 'a)" ("(aI\<^bsub>_\<^esub>)") where "aI\<^bsub>A\<^esub> = (\x\carrier A. x)" abbreviation ASubG :: "[('a, 'more) aGroup_scheme, 'a set] => bool" (infixl "+>" 58) where "A +> H == asubGroup A H" definition Ag_ind :: "[_ , 'a \ 'd] \ 'd aGroup" where "Ag_ind A f = \carrier = f`(carrier A), pop = \x \ f`(carrier A). \y \ f`(carrier A). f(((invfun (carrier A) (f`(carrier A)) f) x) \\<^bsub>A\<^esub> ((invfun (carrier A) (f`(carrier A)) f) y)), mop = \x\(f`(carrier A)). f (-\<^sub>a\<^bsub>A\<^esub> ((invfun (carrier A) (f`(carrier A)) f) x)), zero = f (\\<^bsub>A\<^esub>)\" definition Agii :: "[_ , 'a \ 'd] \ ('a \ 'd)" where "Agii A f = (\x\carrier A. f x)" (** Ag_induced_isomorphism **) lemma (in aGroup) ag_carrier_carrier:"carrier (b_ag A) = carrier A" by (simp add:b_ag_def) lemma (in aGroup) ag_pOp_closed:"\x \ carrier A; y \ carrier A\ \ pop A x y \ carrier A" apply (cut_tac pop_closed) apply (frule funcset_mem[of "(\) " "carrier A" "carrier A \ carrier A" "x"], assumption+) apply (rule funcset_mem[of "(\) x" "carrier A" "carrier A" "y"], assumption+) done lemma (in aGroup) ag_mOp_closed:"x \ carrier A \ (-\<^sub>a x) \ carrier A" apply (cut_tac mop_closed) apply (rule funcset_mem[of "mop A" "carrier A" "carrier A" "x"], assumption+) done lemma (in aGroup) asubg_subset:"A +> H \ H \ carrier A" apply (simp add:asubGroup_def) apply (simp add:sg_def, (erule conjE)+) apply (simp add:ag_carrier_carrier) done lemma (in aGroup) ag_pOp_commute:"\x \ carrier A; y \ carrier A\ \ pop A x y = pop A y x" by (simp add:pop_commute) lemma (in aGroup) b_ag_group:"Group (b_ag A)" apply (unfold Group_def) apply (simp add:b_ag_def) apply (simp add:pop_closed mop_closed ex_zero) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:aassoc) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:l_m) apply (rule allI, rule impI) apply (simp add:l_zero) done lemma (in aGroup) agop_gop:"top (b_ag A) = pop A" (*agpop_gtop*) apply (simp add:b_ag_def) done lemma (in aGroup) agiop_giop:"iop (b_ag A) = mop A" (*agmop_giop*) apply (simp add:b_ag_def) done lemma (in aGroup) agunit_gone:"one (b_ag A) = \" apply (simp add:b_ag_def) done lemma (in aGroup) ag_pOp_add_r:"\a \ carrier A; b \ carrier A; c \ carrier A; a = b\ \ a \ c = b \ c" apply simp done lemma (in aGroup) ag_add_commute:"\a \ carrier A; b \ carrier A\ \ a \ b = b \ a" by (simp add:pop_commute) lemma (in aGroup) ag_pOp_add_l:"\a \ carrier A; b \ carrier A; c \ carrier A; a = b\ \ c \ a = c \ b" apply simp done lemma (in aGroup) asubg_pOp_closed:"\asubGroup A H; x \ H; y \ H\ \ pop A x y \ H" apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (frule Group.sg_mult_closed [of "b_ag A" "H" "x" "y"], assumption+) apply (simp only:agop_gop) done lemma (in aGroup) asubg_mOp_closed:"\asubGroup A H; x \ H\ \ -\<^sub>a x \ H" apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (frule Group.sg_i_closed[of "b_ag A" "H" "x"], assumption+) apply (simp add:agiop_giop) done lemma (in aGroup) asubg_subset1:"\asubGroup A H; x \ H\ \ x \ carrier A" apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (frule Group.sg_subset_elem[of "b_ag A" "H" "x"], assumption+) apply (simp add:ag_carrier_carrier) done lemma (in aGroup) asubg_inc_zero:"asubGroup A H \ \ \ H" apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (frule Group.sg_unit_closed[of "b_ag A" "H"], assumption) apply (simp add:b_ag_def) done lemma (in aGroup) ag_inc_zero:"\ \ carrier A" by (simp add:ex_zero) lemma (in aGroup) ag_l_zero:"x \ carrier A \ \ \ x = x" by (simp add:l_zero) lemma (in aGroup) ag_r_zero:"x \ carrier A \ x \ \ = x" apply (cut_tac ex_zero) apply (subst pop_commute, assumption+) apply (rule ag_l_zero, assumption) done lemma (in aGroup) ag_l_inv1:"x \ carrier A \ (-\<^sub>a x) \ x = \" by (simp add:l_m) lemma (in aGroup) ag_r_inv1:"x \ carrier A \ x \ (-\<^sub>a x) = \" by (frule ag_mOp_closed[of "x"], subst ag_pOp_commute, assumption+, simp add:ag_l_inv1) lemma (in aGroup) ag_pOp_assoc:"\x \ carrier A; y \ carrier A; z \ carrier A\ \ (x \ y) \ z = x \ (y \ z)" by (simp add:aassoc) lemma (in aGroup) ag_inv_unique:"\x \ carrier A; y \ carrier A; x \ y = \\ \ y = -\<^sub>a x" apply (frule ag_mOp_closed[of "x"], frule aassoc[of "-\<^sub>a x" "x" "y"], assumption+, simp add:l_m l_zero ag_r_zero) done lemma (in aGroup) ag_inv_inj:"\x \ carrier A; y \ carrier A; x \ y\ \ (-\<^sub>a x) \ (-\<^sub>a y)" apply (rule contrapos_pp, simp+) apply (frule ag_mOp_closed[of "y"], frule aassoc[of "y" "-\<^sub>a y" "x"], assumption+) apply (simp only:ag_r_inv1, frule sym, thin_tac "-\<^sub>a x = -\<^sub>a y", simp add:l_m) apply (simp add:l_zero ag_r_zero) done lemma (in aGroup) pOp_assocTr41:"\a \ carrier A; b \ carrier A; c \ carrier A; d \ carrier A\ \ a \ b \ c \ d = a \ b \ (c \ d)" by (frule ag_pOp_closed[of "a" "b"], assumption+, rule aassoc[of "a \ b" "c" "d"], assumption+) lemma (in aGroup) pOp_assocTr42:"\a \ carrier A; b \ carrier A; c \ carrier A; d \ carrier A\ \ a \ b \ c \ d = a \ (b \ c) \ d" by (simp add:aassoc[THEN sym, of "a" "b" "c"]) lemma (in aGroup) pOp_assocTr43:"\a \ carrier A; b \ carrier A; c \ carrier A; d \ carrier A\ \ a \ b \ (c \ d) = a \ (b \ c) \ d" by (subst pOp_assocTr41[THEN sym], assumption+, rule pOp_assocTr42, assumption+) lemma (in aGroup) pOp_assoc_cancel:"\a \ carrier A; b \ carrier A; c \ carrier A\ \ a \ -\<^sub>a b \ (b \ -\<^sub>a c) = a \ -\<^sub>a c" apply (subst pOp_assocTr43, assumption) apply (simp add:ag_l_inv1 ag_mOp_closed)+ apply (simp add:ag_r_zero) done lemma (in aGroup) ag_p_inv:"\x \ carrier A; y \ carrier A\ \ (-\<^sub>a (x \ y)) = (-\<^sub>a x) \ (-\<^sub>a y)" apply (frule ag_mOp_closed[of "x"], frule ag_mOp_closed[of "y"], frule ag_pOp_closed[of "x" "y"], assumption+) apply (frule aassoc[of "x \ y" "-\<^sub>a x" "-\<^sub>a y"], assumption+, simp add:pOp_assocTr43, simp add:pop_commute[of "y" "-\<^sub>a x"], simp add:aassoc[THEN sym, of "x" "-\<^sub>a x" "y"], simp add:ag_r_inv1 l_zero) apply (frule ag_pOp_closed[of "-\<^sub>a x" "-\<^sub>a y"], assumption+, simp add:pOp_assocTr41, rule ag_inv_unique[THEN sym, of "x \ y" "-\<^sub>a x \ -\<^sub>a y"], assumption+) done lemma (in aGroup) gEQAddcross: "\l1 \ carrier A; l2 \ carrier A; r1 \ carrier A; r1 \ carrier A; l1 = r2; l2 = r1\ \ l1 \ l2 = r1 \ r2" apply (simp add:ag_pOp_commute) done lemma (in aGroup) ag_eq_sol1:"\a \ carrier A; x\ carrier A; b\ carrier A; a \ x = b\ \ x = (-\<^sub>a a) \ b" apply (frule ag_mOp_closed[of "a"]) apply (frule aassoc[of "-\<^sub>a a" "a" "x"], assumption+) apply (simp add:l_m l_zero) done lemma (in aGroup) ag_eq_sol2:"\a \ carrier A; x\ carrier A; b\ carrier A; x \ a = b\ \ x = b \ (-\<^sub>a a)" apply (frule ag_mOp_closed[of "a"], frule aassoc[of "x" "a" "-\<^sub>a a"], assumption+, simp add:ag_r_inv1 ag_r_zero) done lemma (in aGroup) ag_add4_rel:"\a \ carrier A; b \ carrier A; c \ carrier A; d \ carrier A \ \ a \ b \ (c \ d) = a \ c \ (b \ d)" apply (simp add:pOp_assocTr43[of "a" "b" "c" "d"], simp add:ag_pOp_commute[of "b" "c"], simp add:pOp_assocTr43[THEN sym, of "a" "c" "b" "d"]) done lemma (in aGroup) ag_inv_inv:"x \ carrier A \ -\<^sub>a (-\<^sub>a x) = x" by (frule ag_l_inv1[of "x"], frule ag_mOp_closed[of "x"], rule ag_inv_unique[THEN sym, of "-\<^sub>a x" "x"], assumption+) lemma (in aGroup) ag_inv_zero:"-\<^sub>a \ = \" apply (cut_tac ex_zero) apply (frule l_zero[of "\"]) apply (rule ag_inv_unique[THEN sym], assumption+) done lemma (in aGroup) ag_diff_minus:"\a \ carrier A; b \ carrier A; c \ carrier A; a \ (-\<^sub>a b) = c\ \ b \ (-\<^sub>a a) = (-\<^sub>a c)" apply (frule sym, thin_tac "a \ -\<^sub>a b = c", simp, thin_tac "c = a \ -\<^sub>a b") apply (frule ag_mOp_closed[of "b"], frule ag_mOp_closed[of "a"], subst ag_p_inv, assumption+, subst ag_inv_inv, assumption) apply (simp add:ag_pOp_commute) done lemma (in aGroup) pOp_cancel_l:"\a \ carrier A; b \ carrier A; c \ carrier A; c \ a = c \ b \ \ a = b" apply (frule ag_mOp_closed[of "c"], frule aassoc[of "-\<^sub>a c" "c" "a"], assumption+, simp only:l_m l_zero) apply (simp only:aassoc[THEN sym, of "-\<^sub>a c" "c" "b"], simp only:l_m l_zero) done lemma (in aGroup) pOp_cancel_r:"\a \ carrier A; b \ carrier A; c \ carrier A; a \ c = b \ c \ \ a = b" by (simp add:ag_pOp_commute pOp_cancel_l) lemma (in aGroup) ag_eq_diffzero:"\a \ carrier A; b \ carrier A\ \ (a = b) = (a \ (-\<^sub>a b) = \)" apply (rule iffI) apply (simp add:ag_r_inv1) apply (frule ag_mOp_closed[of "b"]) apply (simp add:ag_pOp_commute[of "a" "-\<^sub>a b"]) apply (subst ag_inv_unique[of "-\<^sub>a b" "a"], assumption+, simp add:ag_inv_inv) done lemma (in aGroup) ag_eq_diffzero1:"\a \ carrier A; b \ carrier A\ \ (a = b) = ((-\<^sub>a a) \ b = \)" apply (frule ag_mOp_closed[of a], simp add:ag_pOp_commute) apply (subst ag_eq_diffzero[THEN sym], assumption+) apply (rule iffI, rule sym, assumption) apply (rule sym, assumption) done lemma (in aGroup) ag_neq_diffnonzero:"\a \ carrier A; b \ carrier A\ \ (a \ b) = (a \ (-\<^sub>a b) \ \)" apply (rule iffI) apply (rule contrapos_pp, simp+) apply (simp add:ag_eq_diffzero[THEN sym]) apply (rule contrapos_pp, simp+) apply (simp add:ag_r_inv1) done lemma (in aGroup) ag_plus_zero:"\x \ carrier A; y \ carrier A\ \ (x = -\<^sub>a y) = (x \ y = \)" apply (rule iffI) apply (simp add:ag_l_inv1) apply (simp add:ag_pOp_commute[of "x" "y"]) apply (rule ag_inv_unique[of "y" "x"], assumption+) done lemma (in aGroup) asubg_nsubg:"A +> H \ (b_ag A) \ H" apply (cut_tac b_ag_group) apply (simp add:asubGroup_def) apply (rule Group.cond_nsg[of "b_ag A" "H"], assumption+) apply (rule ballI)+ apply(simp add:agop_gop agiop_giop) apply (frule Group.sg_subset[of "b_ag A" "H"], assumption) apply (simp add:ag_carrier_carrier) apply (frule_tac c = h in subsetD[of "H" "carrier A"], assumption+) apply (subst ag_pOp_commute, assumption+) apply (frule_tac x = a in ag_mOp_closed) apply (subst aassoc, assumption+, simp add:ag_r_inv1 ag_r_zero) done lemma (in aGroup) subg_asubg:"b_ag G \ H \ G +> H" apply (simp add:asubGroup_def) done lemma (in aGroup) asubg_test:"\H \ carrier A; H \ {}; \a\H. \b\H. (a \ (-\<^sub>a b) \ H)\ \ A +> H" apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (rule Group.sg_condition [of "b_ag A" "H"], assumption+) apply (simp add:ag_carrier_carrier) apply assumption apply (rule allI)+ apply (rule impI) apply (simp add:agop_gop agiop_giop) done lemma (in aGroup) asubg_zero:"A +> {\}" apply (rule asubg_test[of "{\}"]) apply (simp add:ag_inc_zero) apply simp apply (simp, cut_tac ag_inc_zero, simp add:ag_r_inv1) done lemma (in aGroup) asubg_whole:"A +> carrier A" apply (rule asubg_test[of "carrier A"]) apply (simp, cut_tac ag_inc_zero, simp add:nonempty) apply ((rule ballI)+, rule ag_pOp_closed, assumption, rule_tac x = b in ag_mOp_closed, assumption) done lemma (in aGroup) Ag_ind_carrier:"bij_to f (carrier A) (D::'d set) \ carrier (Ag_ind A f) = f ` (carrier A)" by (simp add:Ag_ind_def) lemma (in aGroup) Ag_ind_aGroup:"\f \ carrier A \ D; bij_to f (carrier A) (D::'d set)\ \ aGroup (Ag_ind A f)" apply (simp add:bij_to_def, frule conjunct1, frule conjunct2, fold bij_to_def) apply (simp add:aGroup_def) apply (rule conjI) apply (rule Pi_I)+ apply (simp add:Ag_ind_carrier surj_to_def) apply (frule_tac b = x in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = xa in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:Ag_ind_def) apply (rule funcset_mem[of "f" "carrier A" "D"], assumption) apply (simp add:ag_pOp_closed) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add: Ag_ind_carrier surj_to_def) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = b in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = c in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:Ag_ind_def) apply (frule_tac x = "invfun (carrier A) D f a" and y = "invfun (carrier A) D f b" in ag_pOp_closed, assumption+, frule_tac x = "invfun (carrier A) D f b" and y = "invfun (carrier A) D f c" in ag_pOp_closed, assumption+) apply (simp add:Pi_def) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def) apply (simp add:invfun_l) apply (subst injective_iff[of "f" "carrier A", THEN sym], assumption) apply (simp add:ag_pOp_closed)+ apply (simp add:ag_pOp_assoc) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Ag_ind_def) apply (subst injective_iff[of "f" "carrier A", THEN sym], assumption) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = b in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:surj_to_def) apply (simp add:surj_to_def) apply (simp add:surj_to_def) apply (frule_tac b = b in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:ag_pOp_closed) apply (simp add:surj_to_def) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = b in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:ag_pOp_closed) apply (simp add:surj_to_def) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+, frule_tac b = b in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:ag_pOp_commute) apply (rule conjI) apply (rule Pi_I) apply (simp add:Ag_ind_def surj_to_def) apply (rule funcset_mem[of "f" "carrier A" "D"], assumption) apply (frule_tac b = x in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:ag_mOp_closed) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:Ag_ind_def surj_to_def) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (frule_tac x = "invfun (carrier A) D f a" in ag_mOp_closed) apply (simp add:Pi_def) apply (subst injective_iff[of "f" "carrier A", THEN sym], assumption) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def) apply (simp add:invfun_l) apply (simp add:ag_pOp_closed) apply (simp add:ag_inc_zero) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def) apply (simp add:invfun_l l_m) apply (rule conjI) apply (simp add:Ag_ind_def surj_to_def) apply (rule funcset_mem[of "f" "carrier A" "D"], assumption) apply (simp add:ag_inc_zero) apply (rule allI, rule impI) apply (simp add:Ag_ind_def surj_to_def) apply (cut_tac ag_inc_zero, simp add:funcset_mem del:Pi_I) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def) apply (simp add:invfun_l) apply (frule_tac b = a in invfun_mem1[of "f" "carrier A" "D"], assumption+) apply (simp add:l_zero) apply (simp add:invfun_r) done subsection "Homomorphism of abelian groups" definition aHom :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme] \ ('a \ 'b) set" where "aHom A B = {f. f \ carrier A \ carrier B \ f \ extensional (carrier A) \ (\a\carrier A. \b\carrier A. f (a \\<^bsub>A\<^esub> b) = (f a) \\<^bsub>B\<^esub> (f b))}" definition compos :: "[('a, 'm) aGroup_scheme, 'b \ 'c, 'a \ 'b] \ 'a \ 'c" where "compos A g f = compose (carrier A) g f" definition ker :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme] \ ('a \ 'b) \ 'a set" ("(3ker\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "ker\<^bsub>F,G\<^esub> f = {a. a \ carrier F \ f a = (\\<^bsub>G\<^esub>)}" definition injec :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme, 'a \ 'b] \ bool" ("(3injec\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "injec\<^bsub>F,G\<^esub> f \ f \ aHom F G \ ker\<^bsub>F,G\<^esub> f = {\\<^bsub>F\<^esub>}" definition surjec :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme, 'a \ 'b] \ bool" ("(3surjec\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "surjec\<^bsub>F,G\<^esub> f \ f \ aHom F G \ surj_to f (carrier F) (carrier G)" definition bijec :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme, 'a \ 'b] \ bool" ("(3bijec\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "bijec\<^bsub>F,G\<^esub> f \ injec\<^bsub>F,G\<^esub> f \ surjec\<^bsub>F,G\<^esub> f" definition ainvf :: "[('a, 'm) aGroup_scheme, ('b, 'm1) aGroup_scheme, 'a \ 'b] \ ('b \ 'a)" ("(3ainvf\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "ainvf\<^bsub>F,G\<^esub> f = invfun (carrier F) (carrier G) f" lemma aHom_mem:"\aGroup F; aGroup G; f \ aHom F G; a \ carrier F\ \ f a \ carrier G" apply (simp add:aHom_def) apply (erule conjE)+ apply (simp add:Pi_def) done lemma aHom_func:"f \ aHom F G \ f \ carrier F \ carrier G" by (simp add:aHom_def) lemma aHom_add:"\aGroup F; aGroup G; f \ aHom F G; a \ carrier F; b \ carrier F\ \ f (a \\<^bsub>F\<^esub> b) = (f a) \\<^bsub>G\<^esub> (f b)" apply (simp add:aHom_def) done lemma aHom_0_0:"\aGroup F; aGroup G; f \ aHom F G\ \ f (\\<^bsub>F\<^esub>) = \\<^bsub>G\<^esub>" apply (frule aGroup.ag_inc_zero [of "F"]) apply (subst aGroup.ag_l_zero [THEN sym, of "F" "\\<^bsub>F\<^esub>"], assumption+) apply (simp add:aHom_add) apply (frule aGroup.ag_l_zero [THEN sym, of "F" "\\<^bsub>F\<^esub>"], assumption+) apply (subgoal_tac "f (\\<^bsub>F\<^esub>) = f (\\<^bsub>F\<^esub> \\<^bsub>F\<^esub> \\<^bsub>F\<^esub>)") prefer 2 apply simp apply (thin_tac "\\<^bsub>F\<^esub> = \\<^bsub>F\<^esub> \\<^bsub>F\<^esub> \\<^bsub>F\<^esub>") apply (simp add:aHom_add) apply (frule sym) apply (thin_tac "f \\<^bsub>F\<^esub> = f \\<^bsub>F\<^esub> \\<^bsub>G\<^esub> f \\<^bsub>F\<^esub>") apply (frule aHom_mem[of "F" "G" "f" "\\<^bsub>F\<^esub>"], assumption+) apply (frule aGroup.ag_mOp_closed[of "G" "f \\<^bsub>F\<^esub>"], assumption+) apply (frule aGroup.aassoc[of "G" "-\<^sub>a\<^bsub>G\<^esub> (f \\<^bsub>F\<^esub>)" "f \\<^bsub>F\<^esub>" "f \\<^bsub>F\<^esub>"], assumption+) apply (simp add:aGroup.l_m aGroup.l_zero) done lemma ker_inc_zero:"\aGroup F; aGroup G; f \ aHom F G\ \ \\<^bsub>F\<^esub> \ ker\<^bsub>F,G\<^esub> f" by (frule aHom_0_0[of "F" "G" "f"], assumption+, simp add:ker_def, simp add:aGroup.ag_inc_zero [of "F"]) lemma aHom_inv_inv:"\aGroup F; aGroup G; f \ aHom F G; a \ carrier F\ \ f (-\<^sub>a\<^bsub>F\<^esub> a) = -\<^sub>a\<^bsub>G\<^esub> (f a)" apply (frule aGroup.ag_l_inv1 [of "F" "a"], assumption+, frule sym, thin_tac "-\<^sub>a\<^bsub>F\<^esub> a \\<^bsub>F\<^esub> a = \\<^bsub>F\<^esub>", frule aHom_0_0[of "F" "G" "f"], assumption+, frule aGroup.ag_mOp_closed[of "F" "a"], assumption+) apply (simp add:aHom_add, thin_tac "\\<^bsub>F\<^esub> = -\<^sub>a\<^bsub>F\<^esub> a \\<^bsub>F\<^esub> a") apply (frule aHom_mem[of "F" "G" "f" "-\<^sub>a\<^bsub>F\<^esub> a"], assumption+, frule aHom_mem[of "F" "G" "f" "a"], assumption+, simp only:aGroup.ag_pOp_commute[of "G" "f (-\<^sub>a\<^bsub>F\<^esub> a)" "f a"]) apply (rule aGroup.ag_inv_unique[of "G"], assumption+) done lemma aHom_compos:"\aGroup L; aGroup M; aGroup N; f \ aHom L M; g \ aHom M N \ \ compos L g f \ aHom L N" apply (simp add:aHom_def [of "L" "N"]) apply (rule conjI) apply (rule Pi_I) apply (simp add:compos_def compose_def) apply (rule aHom_mem [of "M" "N" "g"], assumption+) apply (simp add:aHom_mem [of "L" "M" "f"]) apply (rule conjI) apply (simp add:compos_def compose_def extensional_def) apply (rule ballI)+ apply (simp add:compos_def compose_def) apply (simp add:aGroup.ag_pOp_closed) apply (simp add:aHom_add) apply (rule aHom_add, assumption+) apply (simp add:aHom_mem)+ done lemma aHom_compos_assoc:"\aGroup K; aGroup L; aGroup M; aGroup N; f \ aHom K L; g \ aHom L M; h \ aHom M N \ \ compos K h (compos K g f) = compos K (compos L h g) f" apply (simp add:compos_def compose_def) apply (rule funcset_eq[of _ "carrier K"]) apply (simp add:restrict_def extensional_def) apply (simp add:restrict_def extensional_def) apply (rule ballI, simp) apply (simp add:aHom_mem) done lemma injec_inj_on:"\aGroup F; aGroup G; injec\<^bsub>F,G\<^esub> f\ \ inj_on f (carrier F)" apply (simp add:inj_on_def) apply (rule ballI)+ apply (rule impI) apply (simp add:injec_def, erule conjE) apply (frule_tac a = x in aHom_mem[of "F" "G" "f"], assumption+, frule_tac a = x in aHom_mem[of "F" "G" "f"], assumption+) apply (frule_tac x = "f x" in aGroup.ag_r_inv1[of "G"], assumption+) apply (simp only:aHom_inv_inv[THEN sym, of "F" "G" "f"]) apply (frule sym, thin_tac "f x = f y", simp) apply (frule_tac x = y in aGroup.ag_mOp_closed[of "F"], assumption+) apply (simp add:aHom_add[THEN sym], simp add:ker_def) apply (subgoal_tac "x \\<^bsub>F\<^esub> -\<^sub>a\<^bsub>F\<^esub> y \ {a \ carrier F. f a = \\<^bsub>G\<^esub>}", simp) apply (subst aGroup.ag_eq_diffzero[of "F"], assumption+) apply (frule_tac x = x and y = "-\<^sub>a\<^bsub>F\<^esub> y" in aGroup.ag_pOp_closed[of "F"], assumption+) apply simp apply blast done lemma surjec_surj_to:"surjec\<^bsub>R,S\<^esub> f \ surj_to f (carrier R) (carrier S)" by (simp add:surjec_def) lemma compos_bijec:"\aGroup E; aGroup F; aGroup G; bijec\<^bsub>E,F\<^esub> f; bijec\<^bsub>F,G\<^esub> g\ \ bijec\<^bsub>E,G\<^esub> (compos E g f)" apply (simp add:bijec_def, (erule conjE)+) apply (rule conjI) apply (simp add:injec_def, (erule conjE)+) apply (simp add:aHom_compos[of "E" "F" "G" "f" "g"]) apply (rule equalityI, rule subsetI, simp add:ker_def, erule conjE) apply (simp add:compos_def compose_def) apply (frule_tac a = x in aHom_mem[of "E" "F" "f"], assumption+) apply (subgoal_tac "(f x) \ {a \ carrier F. g a = \\<^bsub>G\<^esub>}", simp) apply (subgoal_tac "x \ {a \ carrier E. f a = \\<^bsub>F\<^esub>}", simp) apply blast apply blast apply (rule subsetI, simp) apply (simp add:ker_def compos_def compose_def) apply (simp add:aGroup.ag_inc_zero) apply (simp add:aHom_0_0) apply (simp add:surjec_def, (erule conjE)+) apply (simp add:aHom_compos) apply (simp add:aHom_def, (erule conjE)+) apply (simp add:compos_def) apply (rule compose_surj[of "f" "carrier E" "carrier F" "g" "carrier G"], assumption+) done lemma ainvf_aHom:"\aGroup F; aGroup G; bijec\<^bsub>F,G\<^esub> f\ \ ainvf\<^bsub>F,G\<^esub> f \ aHom G F" apply (subst aHom_def, simp) apply (simp add:ainvf_def) apply (simp add:bijec_def, erule conjE) apply (frule injec_inj_on[of "F" "G" "f"], assumption+) apply (simp add:surjec_def, (erule conjE)+) apply (simp add:aHom_def, (erule conjE)+) apply (frule inv_func[of "f" "carrier F" "carrier G"], assumption+, simp) apply (rule conjI) apply (simp add:invfun_def) apply (rule ballI)+ apply (frule_tac x = a in funcset_mem[of "Ifn F G f" "carrier G" "carrier F"], assumption+, frule_tac x = b in funcset_mem[of "Ifn F G f" "carrier G" "carrier F"], assumption+, frule_tac x = a and y = b in aGroup.ag_pOp_closed[of "G"], assumption+, frule_tac x = "a \\<^bsub>G\<^esub> b" in funcset_mem[of "Ifn F G f" "carrier G" "carrier F"], assumption+) apply (frule_tac a = "(Ifn F G f) a" and b = "(Ifn F G f) b" in aHom_add[of "F" "G" "f"], assumption+, simp add:injec_def, assumption+, thin_tac "\a\carrier F. \b\carrier F. f (a \\<^bsub>F\<^esub> b) = f a \\<^bsub>G\<^esub> f b") apply (simp add:invfun_r[of "f" "carrier F" "carrier G"]) apply (frule_tac x = a and y = b in aGroup.ag_pOp_closed[of "G"], assumption+) apply (frule_tac b = "a \\<^bsub>G\<^esub> b" in invfun_r[of "f" "carrier F" "carrier G"], assumption+) apply (simp add:inj_on_def) apply (frule_tac x = "(Ifn F G f) a" and y = "(Ifn F G f) b" in aGroup.ag_pOp_closed, assumption+) apply (frule_tac x = "(Ifn F G f) (a \\<^bsub>G\<^esub> b)" in bspec, assumption, thin_tac "\x\carrier F. \y\carrier F. f x = f y \ x = y") apply (frule_tac x = "(Ifn F G f) a \\<^bsub>F\<^esub> (Ifn F G f) b" in bspec, assumption, thin_tac "\y\carrier F. f ((Ifn F G f) (a \\<^bsub>G\<^esub> b)) = f y \ (Ifn F G f) (a \\<^bsub>G\<^esub> b) = y") apply simp done lemma ainvf_bijec:"\aGroup F; aGroup G; bijec\<^bsub>F,G\<^esub> f\ \ bijec\<^bsub>G,F\<^esub> (ainvf\<^bsub>F,G\<^esub> f)" apply (subst bijec_def) apply (simp add:injec_def surjec_def) apply (simp add:ainvf_aHom) apply (rule conjI) apply (rule equalityI) apply (rule subsetI, simp add:ker_def, erule conjE) apply (simp add:ainvf_def) apply (simp add:bijec_def,(erule conjE)+, simp add:surjec_def, (erule conjE)+, simp add:aHom_def, (erule conjE)+) apply (frule injec_inj_on[of "F" "G" "f"], assumption+) apply (subst invfun_r[THEN sym, of "f" "carrier F" "carrier G"], assumption+) apply (simp add:injec_def, (erule conjE)+, simp add:aHom_0_0) apply (rule subsetI, simp add:ker_def) apply (simp add:aGroup.ex_zero) apply (frule ainvf_aHom[of "F" "G" "f"], assumption+) apply (simp add:aHom_0_0) apply (frule ainvf_aHom[of "F" "G" "f"], assumption+, simp add:aHom_def, (erule conjE)+, rule surj_to_test[of "ainvf\<^bsub>F,G\<^esub> f" "carrier G" "carrier F"], assumption+) apply (rule ballI, thin_tac "\a\carrier G. \b\carrier G. (ainvf\<^bsub>F,G\<^esub> f) (a \\<^bsub>G\<^esub> b) = (ainvf\<^bsub>F,G\<^esub> f) a \\<^bsub>F\<^esub> (ainvf\<^bsub>F,G\<^esub> f) b") apply (simp add:bijec_def, erule conjE) apply (frule injec_inj_on[of "F" "G" "f"], assumption+) apply (simp add:surjec_def aHom_def, (erule conjE)+) apply (subst ainvf_def) apply (frule_tac a = b in invfun_l[of "f" "carrier F" "carrier G"], assumption+, frule_tac x = b in funcset_mem[of "f" "carrier F" "carrier G"], assumption+, blast) done lemma ainvf_l:"\aGroup E; aGroup F; bijec\<^bsub>E,F\<^esub> f; x \ carrier E\ \ (ainvf\<^bsub>E,F\<^esub> f) (f x) = x" apply (simp add:bijec_def, erule conjE) apply (frule injec_inj_on[of "E" "F" "f"], assumption+) apply (simp add:surjec_def aHom_def, (erule conjE)+) apply (frule invfun_l[of "f" "carrier E" "carrier F" "x"], assumption+) apply (simp add:ainvf_def) done lemma (in aGroup) aI_aHom:"aI\<^bsub>A\<^esub> \ aHom A A" by (simp add:aHom_def ag_idmap_def ag_idmap_def ag_pOp_closed) lemma compos_aI_l:"\aGroup A; aGroup B; f \ aHom A B\ \ compos A aI\<^bsub>B\<^esub> f = f" apply (simp add:compos_def) apply (rule funcset_eq[of _ "carrier A"]) apply (simp add:compose_def extensional_def) apply (simp add:aHom_def) apply (rule ballI) apply (frule_tac a = x in aHom_mem[of "A" "B" "f"], assumption+) apply (simp add:compose_def ag_idmap_def) done lemma compos_aI_r:"\aGroup A; aGroup B; f \ aHom A B\ \ compos A f aI\<^bsub>A\<^esub> = f" apply (simp add:compos_def) apply (rule funcset_eq[of _ "carrier A"]) apply (simp add:compose_def extensional_def) apply (simp add:aHom_def) apply (rule ballI) apply (simp add:compose_def ag_idmap_def) done lemma compos_aI_surj:"\aGroup A; aGroup B; f \ aHom A B; g \ aHom B A; compos A g f = aI\<^bsub>A\<^esub>\ \ surjec\<^bsub>B,A\<^esub> g" apply (simp add:surjec_def) apply (rule surj_to_test[of "g" "carrier B" "carrier A"]) apply (simp add:aHom_def) apply (rule ballI) apply (subgoal_tac "compos A g f b = aI\<^bsub>A\<^esub> b", thin_tac "compos A g f = aI\<^bsub>A\<^esub>") apply (simp add:compos_def compose_def ag_idmap_def) apply (frule_tac a = b in aHom_mem[of "A" "B" "f"], assumption+, blast) apply simp done lemma compos_aI_inj:"\aGroup A; aGroup B; f \ aHom A B; g \ aHom B A; compos A g f = aI\<^bsub>A\<^esub>\ \ injec\<^bsub>A,B\<^esub> f" apply (simp add:injec_def) apply (simp add:ker_def) apply (rule equalityI) apply (rule subsetI, simp, erule conjE) apply (subgoal_tac "compos A g f x = aI\<^bsub>A\<^esub> x", thin_tac "compos A g f = aI\<^bsub>A\<^esub>") apply (simp add:compos_def compose_def) apply (simp add:aHom_0_0 ag_idmap_def) apply simp apply (rule subsetI, simp) apply (simp add:aGroup.ag_inc_zero aHom_0_0) done lemma (in aGroup) Ag_ind_aHom:"\f \ carrier A \ D; bij_to f (carrier A) (D::'d set)\ \ Agii A f \ aHom A (Ag_ind A f)" apply (simp add:aHom_def) apply (unfold bij_to_def, frule conjunct1, frule conjunct2, fold bij_to_def) apply (simp add:Ag_ind_carrier surj_to_def) apply (rule conjI) apply (simp add:Agii_def Pi_def) apply (simp add:Agii_def) apply (simp add:Ag_ind_def Pi_def) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def) apply (simp add:invfun_l) apply (simp add:ag_pOp_closed) done lemma (in aGroup) Agii_mem:"\f \ carrier A \ D; x \ carrier A; bij_to f (carrier A) (D::'d set)\ \ Agii A f x \ carrier (Ag_ind A f)" apply (simp add:Agii_def Ag_ind_carrier) done lemma Ag_ind_bijec:"\aGroup A; f \ carrier A \ D; bij_to f (carrier A) (D::'d set)\ \ bijec\<^bsub>A, (Ag_ind A f)\<^esub> (Agii A f)" apply (frule aGroup.Ag_ind_aHom[of "A" "f" "D"], assumption+) apply (frule aGroup.Ag_ind_aGroup[of "A" "f" "D"], assumption+) apply (simp add:bijec_def) apply (rule conjI) apply (simp add:injec_def) apply (rule equalityI) apply (rule subsetI) apply (simp add:ker_def, erule conjE) apply (frule aHom_0_0[of "A" "Ag_ind A f" "Agii A f"], assumption+) apply (rotate_tac -2, frule sym, thin_tac "Agii A f x = \\<^bsub>Ag_ind A f\<^esub>", simp) apply (frule aGroup.ag_inc_zero[of "A"], simp add:Agii_def) apply (unfold bij_to_def, frule conjunct2, fold bij_to_def) apply (frule aGroup.ag_inc_zero[of "A"]) apply (simp add:injective_iff[THEN sym, of "f" "carrier A" "\\<^bsub>A\<^esub>"]) apply (rule subsetI, simp) apply (subst ker_def, simp) apply (simp add:aGroup.ag_inc_zero, simp add:aHom_0_0) apply (subst surjec_def) apply (unfold bij_to_def, frule conjunct1, fold bij_to_def, simp) apply (simp add:aGroup.Ag_ind_carrier surj_to_def Agii_def) done definition aimg :: "[('b, 'm1) aGroup_scheme, _, 'b \ 'a] \ 'a aGroup" ("(3aimg\<^bsub>_,_\<^esub> _)" [82,82,83]82) where "aimg\<^bsub>F,A\<^esub> f = A \ carrier := f ` (carrier F), pop := pop A, mop := mop A, zero := zero A\" lemma ker_subg:"\aGroup F; aGroup G; f \ aHom F G \ \ F +> ker\<^bsub>F,G\<^esub> f" apply (rule aGroup.asubg_test, assumption+) apply (rule subsetI) apply (simp add:ker_def) apply (simp add:ker_def) apply (frule aHom_0_0 [of "F" "G" "f"], assumption+) apply (frule aGroup.ex_zero [of "F"]) apply blast apply (rule ballI)+ apply (simp add:ker_def) apply (erule conjE)+ apply (frule_tac x = b in aGroup.ag_mOp_closed[of "F"], assumption+) apply (rule conjI) apply (rule aGroup.ag_pOp_closed, assumption+) apply (simp add:aHom_add) apply (simp add:aHom_inv_inv) apply (simp add:aGroup.ag_inv_zero[of "G"]) apply (cut_tac aGroup.ex_zero[of "G"], simp add:aGroup.l_zero) apply assumption done subsection "Quotient abelian group" definition ar_coset :: "['a, _ , 'a set] \ 'a set" (** a_rcs **) ("(3_ \\<^bsub>_\<^esub> _)" [66,66,67]66) where "ar_coset a A H = H \\<^bsub>(b_ag A)\<^esub> a" definition set_ar_cos :: "[_ , 'a set] \ 'a set set" where "set_ar_cos A I = {X. \a\carrier A. X = ar_coset a A I}" definition aset_sum :: "[_ , 'a set, 'a set] \ 'a set" where "aset_sum A H K = s_top (b_ag A) H K" abbreviation ASBOP1 (infix "\\" 60) where "H \\<^bsub>A\<^esub> K == aset_sum A H K" lemma (in aGroup) ag_a_in_ar_cos:"\A +> H; a \ carrier A\ \ a \ a \\<^bsub>A\<^esub> H" apply (simp add:ar_coset_def) apply (simp add:asubGroup_def) apply (cut_tac b_ag_group) apply (rule Group.a_in_rcs[of "b_ag A" "H" "a"], assumption+) apply (simp add:ag_carrier_carrier[THEN sym]) done lemma (in aGroup) r_cos_subset:"\A +> H; X \ set_rcs (b_ag A) H\ \ X \ carrier A" apply (simp add:asubGroup_def set_rcs_def) apply (erule bexE) apply (cut_tac b_ag_group) apply (frule_tac a = a in Group.rcs_subset[of "b_ag A" "H"], assumption+) apply (simp add:ag_carrier_carrier) done lemma (in aGroup) asubg_costOp_commute:"\A +> H; x \ set_rcs (b_ag A) H; y \ set_rcs (b_ag A) H\ \ c_top (b_ag A) H x y = c_top (b_ag A) H y x" apply (simp add:set_rcs_def, (erule bexE)+, simp) apply (cut_tac b_ag_group) apply (subst Group.c_top_welldef[THEN sym], assumption+, simp add:asubg_nsubg, (simp add:ag_carrier_carrier)+) apply (subst Group.c_top_welldef[THEN sym], assumption+, simp add:asubg_nsubg, (simp add:ag_carrier_carrier)+) apply (simp add:agop_gop) apply (simp add:ag_pOp_commute) done lemma (in aGroup) Subg_Qgroup:"A +> H \ aGroup (aqgrp A H)" apply (frule asubg_nsubg[of "H"]) apply (cut_tac b_ag_group) apply (simp add:aGroup_def) apply (simp add:aqgrp_def) apply (simp add:Group.Qg_top [of "b_ag A" "H"]) apply (simp add:Group.Qg_iop [of "b_ag A" "H"]) apply (frule Group.nsg_sg[of "b_ag A" "H"], assumption+, simp add:Group.unit_rcs_in_set_rcs[of "b_ag A" "H"]) apply (simp add:Group.Qg_tassoc) apply (simp add:asubg_costOp_commute) apply (simp add:Group.Qg_i[of "b_ag A" "H"]) apply (simp add:Group.Qg_unit[of "b_ag A" "H"]) done lemma (in aGroup) plus_subgs:"\A +> H1; A +> H2\ \ A +> H1 \ H2" apply (simp add:aset_sum_def) apply (frule asubg_nsubg[of "H2"]) apply (simp add:asubGroup_def[of _ "H1"]) apply (cut_tac "b_ag_group") apply (frule Group.smult_sg_nsg[of "b_ag A" "H1" "H2"], assumption+) apply (simp add:asubGroup_def) done lemma (in aGroup) set_sum:"\H \ carrier A; K \ carrier A\ \ H \ K = {x. \h\H. \k\K. x = h \ k}" apply (cut_tac b_ag_group) apply (rule equalityI) apply (rule subsetI) apply (simp add:aset_sum_def) apply (simp add:agop_gop[THEN sym] s_top_def, (erule bexE)+, frule sym, thin_tac "xa \\<^bsub>b_ag A\<^esub> y = x", simp, blast) apply (rule subsetI, simp add:aset_sum_def, (erule bexE)+) apply (frule_tac c = h in subsetD[of H "carrier A"], assumption+, frule_tac c = k in subsetD[of K "carrier A"], assumption+) apply (simp add:agop_gop[THEN sym], simp add:s_top_def, blast) done lemma (in aGroup) mem_set_sum:"\H \ carrier A; K \ carrier A; x \ H \ K \ \ \h\H. \k\K. x = h \ k" by (simp add:set_sum) lemma (in aGroup) mem_sum_subgs:"\A +> H; A +> K; h \ H; k \ K\ \ h \ k \ H \ K" apply (frule asubg_subset[of H], frule asubg_subset[of K], simp add:set_sum, blast) done lemma (in aGroup) aqgrp_carrier:"A +> H \ set_rcs (b_ag A ) H = set_ar_cos A H" apply (simp add:set_ar_cos_def) apply (simp add:ag_carrier_carrier [THEN sym]) apply (simp add:ar_coset_def set_rcs_def) done lemma (in aGroup) unit_in_set_ar_cos:"A +> H \ H \ set_ar_cos A H" apply (simp add:aqgrp_carrier[THEN sym]) apply (cut_tac b_ag_group) apply (simp add:asubGroup_def) apply (simp add:Group.unit_rcs_in_set_rcs[of "b_ag A" "H"]) done lemma (in aGroup) aqgrp_pOp_maps:"\A +> H; a \ carrier A; b \ carrier A\ \ pop (aqgrp A H) (a \\<^bsub>A\<^esub> H) (b \\<^bsub>A\<^esub> H) = (a \ b) \\<^bsub>A\<^esub> H" apply (simp add:aqgrp_def ar_coset_def) apply (cut_tac b_ag_group) apply (frule asubg_nsubg) apply (simp add:ag_carrier_carrier [THEN sym]) apply (subst Group.c_top_welldef [THEN sym], assumption+) apply (simp add:agop_gop) done lemma (in aGroup) aqgrp_mOp_maps:"\A +> H; a \ carrier A\ \ mop (aqgrp A H) (a \\<^bsub>A\<^esub> H) = (-\<^sub>a a) \\<^bsub>A\<^esub> H" apply (simp add:aqgrp_def ar_coset_def) apply (cut_tac b_ag_group) apply (frule asubg_nsubg) apply (simp add:ag_carrier_carrier [THEN sym]) apply (subst Group.c_iop_welldef, assumption+) apply (simp add:agiop_giop) done lemma (in aGroup) aqgrp_zero:"A +> H \ zero (aqgrp A H) = H" apply (simp add:aqgrp_def) done lemma (in aGroup) arcos_fixed:"\A +> H; a \ carrier A; h \ H \ \ a \\<^bsub>A\<^esub> H = (h \ a) \\<^bsub>A\<^esub> H" apply (cut_tac b_ag_group) apply (simp add:agop_gop[THEN sym]) apply (simp add:ag_carrier_carrier[THEN sym]) apply (simp add:ar_coset_def) apply (simp add:asubGroup_def) apply (simp add:Group.rcs_fixed1[of "b_ag A" "H"]) done definition rind_hom :: "[('a, 'more) aGroup_scheme, ('b, 'more1) aGroup_scheme, ('a \ 'b)] \ ('a set \ 'b )" where "rind_hom A B f = (\X\(set_ar_cos A (ker\<^bsub>A,B\<^esub> f)). f (SOME x. x \ X))" abbreviation RIND_HOM ("(3_\\<^bsub>_,_\<^esub>)" [82,82,83]82) where "f\\<^bsub>F,G\<^esub> == rind_hom F G f" (* tOp \ pOp *) section "Direct product and direct sum of abelian groups, in general case" definition Un_carrier :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ 'a set" where "Un_carrier I A = \{X. \i\I. X = carrier (A i)}" definition carr_prodag :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a ) set" where "carr_prodag I A = {f. f \ extensional I \ f \ I \ (Un_carrier I A) \ (\i\I. f i \ carrier (A i))}" definition prod_pOp :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a) \ ('i \ 'a) \ ('i \ 'a)" where "prod_pOp I A = (\f\carr_prodag I A. \g\carr_prodag I A. \x\I. (f x) \\<^bsub>(A x)\<^esub> (g x))" definition prod_mOp :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a) \ ('i \ 'a)" where "prod_mOp I A = (\f\carr_prodag I A. \x\I. (-\<^sub>a\<^bsub>(A x)\<^esub> (f x)))" definition prod_zero :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a)" where "prod_zero I A = (\x\I. \\<^bsub>(A x)\<^esub>)" definition prodag :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a) aGroup" where "prodag I A = \ carrier = carr_prodag I A, pop = prod_pOp I A, mop = prod_mOp I A, zero = prod_zero I A\" definition PRoject :: "['i set, 'i \ ('a, 'more) aGroup_scheme, 'i] \ ('i \ 'a) \ 'a" ("(3\\<^bsub>_,_,_\<^esub>)" [82,82,83]82) where "PRoject I A x = (\f \ carr_prodag I A. f x)" abbreviation PRODag ("(a\\<^bsub>_\<^esub> _)" [72,73]72) where "a\\<^bsub>I\<^esub> A == prodag I A" lemma prodag_comp_i:"\a \ carr_prodag I A; i \ I\ \ (a i) \ carrier (A i)" by (simp add:carr_prodag_def) lemma prod_pOp_func:"\k\I. aGroup (A k) \ prod_pOp I A \ carr_prodag I A \ carr_prodag I A \ carr_prodag I A" apply (rule Pi_I)+ apply(rename_tac a b) apply (subst carr_prodag_def) apply (simp add:CollectI) apply (rule conjI) apply (simp add:prod_pOp_def restrict_def extensional_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:prod_pOp_def) apply (subst Un_carrier_def) apply (simp add:CollectI) apply (frule_tac x = x in bspec, assumption, thin_tac "\k\I. aGroup (A k)") apply (simp add:carr_prodag_def) apply (erule conjE)+ apply (thin_tac "a \ I \ Un_carrier I A") apply (thin_tac "b \ I \ Un_carrier I A") apply (frule_tac x = x in bspec, assumption, thin_tac "\i\I. a i \ carrier (A i)", frule_tac x = x in bspec, assumption, thin_tac "\i\I. b i \ carrier (A i)") apply (frule_tac x = "a x" and y = "b x" in aGroup.ag_pOp_closed, assumption+) apply blast apply (rule ballI) apply (simp add:prod_pOp_def) apply (rule_tac A = "A i" and x = "a i" and y = "b i" in aGroup.ag_pOp_closed) apply simp apply (simp add:carr_prodag_def)+ done lemma prod_pOp_mem:"\\k\I. aGroup (A k); X \ carr_prodag I A; Y \ carr_prodag I A\ \ prod_pOp I A X Y \ carr_prodag I A" apply (frule prod_pOp_func) apply (frule funcset_mem[of "prod_pOp I A" "carr_prodag I A" "carr_prodag I A \ carr_prodag I A" "X"], assumption+) apply (rule funcset_mem[of "prod_pOp I A X" "carr_prodag I A" "carr_prodag I A" "Y"], assumption+) done lemma prod_pOp_mem_i:"\\k\I. aGroup (A k); X \ carr_prodag I A; Y \ carr_prodag I A; i \ I\ \ prod_pOp I A X Y i = (X i) \\<^bsub>(A i)\<^esub> (Y i)" apply (simp add:prod_pOp_def) done lemma prod_mOp_func:"\k\I. aGroup (A k) \ prod_mOp I A \ carr_prodag I A \ carr_prodag I A" apply (rule Pi_I) apply (simp add:prod_mOp_def carr_prodag_def) apply (erule conjE)+ apply (rule conjI) apply (rule Pi_I) apply simp apply (rename_tac f j) apply (frule_tac f = f and x = j in funcset_mem [of _ "I" "Un_carrier I A"], assumption+) apply (thin_tac "f \ I \ Un_carrier I A") apply (frule_tac x = j in bspec, assumption, thin_tac "\k\I. aGroup (A k)", frule_tac x = j in bspec, assumption, thin_tac "\i\I. f i \ carrier (A i)") apply (thin_tac "f j \ Un_carrier I A") apply (simp add:Un_carrier_def) apply (frule aGroup.ag_mOp_closed, assumption+) apply blast apply (rule ballI) apply (rule_tac A = "A i" and x = "x i" in aGroup.ag_mOp_closed) apply simp+ done lemma prod_mOp_mem:"\\j\I. aGroup (A j); X \ carr_prodag I A\ \ prod_mOp I A X \ carr_prodag I A" apply (frule prod_mOp_func) apply (simp add:Pi_def) done lemma prod_mOp_mem_i:"\\j\I. aGroup (A j); X \ carr_prodag I A; i \ I\ \ prod_mOp I A X i = -\<^sub>a\<^bsub>(A i)\<^esub> (X i)" apply (simp add:prod_mOp_def) done lemma prod_zero_func:"\k\I. aGroup (A k) \ prod_zero I A \ carr_prodag I A" apply (simp add:prod_zero_def prodag_def) apply (simp add:carr_prodag_def) apply (rule conjI) apply (rule Pi_I) apply simp apply (subgoal_tac "aGroup (A x)") prefer 2 apply simp apply (thin_tac "\k\I. aGroup (A k)") apply (simp add:Un_carrier_def) apply (frule aGroup.ex_zero) apply auto apply (frule_tac x = i in bspec, assumption, thin_tac "\k\I. aGroup (A k)") apply (simp add:aGroup.ex_zero) done lemma prod_zero_i:"\\k\I. aGroup (A k); i \ I\ \ prod_zero I A i = \\<^bsub>(A i)\<^esub> " by (simp add:prod_zero_def) lemma carr_prodag_mem_eq:"\\k\I. aGroup (A k); X \ carr_prodag I A; Y \ carr_prodag I A; \l\I. (X l) = (Y l) \ \ X = Y" apply (simp add:carr_prodag_def) apply (erule conjE)+ apply (simp add:funcset_eq) done lemma prod_pOp_assoc:"\\k\I. aGroup (A k); a \ carr_prodag I A; b \ carr_prodag I A; c \ carr_prodag I A\ \ prod_pOp I A (prod_pOp I A a b) c = prod_pOp I A a (prod_pOp I A b c)" apply (frule_tac X = a and Y = b in prod_pOp_mem[of "I" "A"], assumption+, frule_tac X = b and Y = c in prod_pOp_mem[of "I" "A"], assumption+, frule_tac X = "prod_pOp I A a b" and Y = c in prod_pOp_mem[of "I" "A"], assumption+, frule_tac X = a and Y = "prod_pOp I A b c" in prod_pOp_mem[of "I" "A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], assumption+, rule ballI) apply (simp add:prod_pOp_mem_i) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)") apply (rule aGroup.ag_pOp_assoc, assumption) apply (simp add:prodag_comp_i)+ done lemma prod_pOp_commute:"\\k\I. aGroup (A k); a \ carr_prodag I A; b \ carr_prodag I A\ \ prod_pOp I A a b = prod_pOp I A b a" apply (frule_tac X = a and Y = b in prod_pOp_mem[of "I" "A"], assumption+, frule_tac X = b and Y = a in prod_pOp_mem[of "I" "A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], assumption+, rule ballI) apply (simp add:prod_pOp_mem_i) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)", rule aGroup.ag_pOp_commute, assumption) apply (simp add:prodag_comp_i)+ done lemma prodag_aGroup:"\k\I. aGroup (A k) \ aGroup (prodag I A)" apply (simp add:aGroup_def [of "(prodag I A)"]) apply (simp add:prodag_def) apply (simp add:prod_pOp_func) apply (simp add:prod_mOp_func) apply (simp add:prod_zero_func) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:prod_pOp_assoc) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:prod_pOp_commute) apply (rule conjI) apply (rule allI, rule impI) apply (frule_tac X = a in prod_mOp_mem [of "I" "A"], assumption+) apply (frule_tac X = "prod_mOp I A a" and Y = a in prod_pOp_mem[of "I" "A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], assumption+) apply (simp add:prod_zero_func) apply (rule ballI) apply (simp add:prod_pOp_mem_i, simp add:prod_zero_i) apply ( simp add:prod_mOp_mem_i) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)", rule aGroup.l_m, assumption+, simp add:prodag_comp_i) apply (rule allI, rule impI) apply (frule_tac prod_zero_func[of "I" "A"], frule_tac Y = a in prod_pOp_mem[of "I" "A" "prod_zero I A"], assumption+) apply (rule carr_prodag_mem_eq[of "I" "A"], assumption+) apply (rule ballI) apply (subst prod_pOp_mem_i[of "I" "A"], assumption+, subst prod_zero_i[of "I" "A"], assumption+) apply (frule_tac x = l in bspec, assumption, rule aGroup.l_zero, assumption+, simp add:prodag_comp_i) done lemma prodag_carrier:"\k\I. aGroup (A k) \ carrier (prodag I A) = carr_prodag I A" by (simp add:prodag_def) lemma prodag_elemfun:"\\k\I. aGroup (A k); f \ carrier (prodag I A)\ \ f \ extensional I" apply (simp add:prodag_carrier) apply (simp add:carr_prodag_def) done lemma prodag_component:"\f \ carrier (prodag I A); i \ I \ \ f i \ carrier (A i)" by (simp add:prodag_def carr_prodag_def) lemma prodag_pOp:"\k\I. aGroup (A k) \ pop (prodag I A) = prod_pOp I A" apply (simp add:prodag_def) done lemma prodag_iOp:"\k\I. aGroup (A k) \ mop (prodag I A) = prod_mOp I A" apply (simp add:prodag_def) done lemma prodag_zero:"\k\I. aGroup (A k) \ zero (prodag I A) = prod_zero I A" apply (simp add:prodag_def) done lemma prodag_sameTr0:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ Un_carrier I A = Un_carrier I B" apply (simp add:Un_carrier_def) done lemma prodag_sameTr1:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ carr_prodag I A = carr_prodag I B" apply (rule equalityI) apply (rule subsetI) apply (simp add:carr_prodag_def, (erule conjE)+) apply (rule Pi_I) apply (subst Un_carrier_def, simp, blast) apply (rule subsetI) apply (simp add:carr_prodag_def, (erule conjE)+) apply (rule Pi_I) apply (subst Un_carrier_def, simp) apply blast done lemma prodag_sameTr2:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ prod_pOp I A = prod_pOp I B" apply (frule prodag_sameTr1 [of "I" "A" "B"], assumption+) apply (simp add:prod_pOp_def) apply (rule bivar_func_eq) apply (rule ballI)+ apply (rule funcset_eq [of _ "I"]) apply (simp add:restrict_def extensional_def)+ done lemma prodag_sameTr3:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ prod_mOp I A = prod_mOp I B" apply (frule prodag_sameTr1 [of "I" "A" "B"], assumption+) apply (simp add:prod_mOp_def) apply (rule funcset_eq [of _ "carr_prodag I B"]) apply (simp add:restrict_def extensional_def) apply (simp add:restrict_def extensional_def) apply (rule ballI) apply (rename_tac g) apply simp apply (rule funcset_eq [of _ "I"]) apply (simp add:restrict_def extensional_def)+ done lemma prodag_sameTr4:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ prod_zero I A = prod_zero I B" apply (simp add:prod_zero_def) apply (rule funcset_eq [of _ "I"]) apply (simp add:restrict_def extensional_def)+ done lemma prodag_same:"\\k\I. aGroup (A k); \k\I. A k = B k\ \ prodag I A = prodag I B" apply (frule prodag_sameTr1, assumption+) apply (frule prodag_sameTr2, assumption+) apply (frule prodag_sameTr3, assumption+) apply (frule prodag_sameTr4, assumption+) apply (simp add:prodag_def) done lemma project_mem:"\\k\I. aGroup (A k); j \ I; x \ carrier (prodag I A)\ \ (PRoject I A j) x \ carrier (A j)" apply (simp add:PRoject_def) apply (simp add:prodag_def) apply (simp add:carr_prodag_def) done lemma project_aHom:"\\k\I. aGroup (A k); j \ I\ \ PRoject I A j \ aHom (prodag I A) (A j)" apply (simp add:aHom_def) apply (rule conjI) apply (simp add:project_mem) apply (rule conjI) apply (simp add:PRoject_def restrict_def extensional_def) apply (rule allI, rule impI, simp add:prodag_def) apply (rule ballI)+ apply (simp add:prodag_def) apply (simp add:prod_pOp_def) apply (frule_tac X = a and Y = b in prod_pOp_mem[of I A], assumption+) apply (simp add:prod_pOp_def) apply (simp add:PRoject_def) done lemma project_aHom1:"\k\I. aGroup (A k) \ \j \ I. PRoject I A j \ aHom (prodag I A) (A j)" apply (rule ballI) apply (rule project_aHom, assumption+) done definition A_to_prodag :: "[('a, 'm) aGroup_scheme, 'i set, 'i \('a \ 'b), 'i \ ('b, 'm1) aGroup_scheme] \ ('a \ ('i \'b))" where "A_to_prodag A I S B = (\a\carrier A. \k\I. S k a)" (* I is an index set, A is an abelian group, S: I \ carrier A \ carrier (prodag I B), s i \ carrier A \ B i *) lemma A_to_prodag_mem:"\aGroup A; \k\I. aGroup (B k); \k\I. (S k) \ aHom A (B k); x \ carrier A \ \ A_to_prodag A I S B x \ carr_prodag I B" apply (simp add:carr_prodag_def) apply (rule conjI) apply (simp add:A_to_prodag_def extensional_def restrict_def) apply (simp add:Pi_def restrict_def A_to_prodag_def) apply (rule conjI) apply (rule allI) apply (rule impI) apply (simp add:Un_carrier_def) apply (rotate_tac 2, frule_tac x = xa in bspec, assumption, thin_tac "\k\I. S k \ aHom A (B k)") apply (simp add:aHom_def) apply (erule conjE)+ apply (frule_tac f = "S xa" and A = "carrier A" and B = "carrier (B xa)" and x = x in funcset_mem, assumption+) apply blast apply (rule ballI) apply (rotate_tac 2, frule_tac x = i in bspec, assumption, thin_tac "\k\I. S k \ aHom A (B k)") apply (simp add:aHom_def) apply (erule conjE)+ apply (simp add:Pi_def) done lemma A_to_prodag_aHom:"\aGroup A; \k\I. aGroup (B k); \k\I. (S k) \ aHom A (B k) \ \ A_to_prodag A I S B \ aHom A (a\\<^bsub>I\<^esub> B)" apply (simp add:aHom_def [of "A" "a\\<^bsub>I\<^esub> B"]) apply (rule conjI) apply (simp add:prodag_def A_to_prodag_mem) apply (rule conjI) apply (simp add:A_to_prodag_def restrict_def extensional_def) apply (rule ballI)+ apply (frule_tac x = a and y = b in aGroup.ag_pOp_closed, assumption+) apply (frule_tac x = "a \\<^bsub>A\<^esub> b" in A_to_prodag_mem [of "A" "I" "B" "S"], assumption+) apply (frule_tac x = a in A_to_prodag_mem [of "A" "I" "B" "S"], assumption+) apply (frule_tac x = b in A_to_prodag_mem [of "A" "I" "B" "S"], assumption+) apply (frule prodag_aGroup [of "I" "B"]) apply (frule_tac x = a in A_to_prodag_mem[of "A" "I" "B" "S"], assumption+, frule_tac x = b in A_to_prodag_mem[of "A" "I" "B" "S"], assumption+, frule_tac x = "a \\<^bsub>A\<^esub> b" in A_to_prodag_mem[of "A" "I" "B" "S"], assumption+) apply (frule prodag_aGroup[of "I" "B"], frule_tac x = "A_to_prodag A I S B a" and y = "A_to_prodag A I S B b" in aGroup.ag_pOp_closed [of "a\\<^bsub>I\<^esub> B"]) apply (simp add:prodag_carrier) apply (simp add:prodag_carrier) apply (rule carr_prodag_mem_eq, assumption+) apply (simp add:prodag_carrier) apply (rule ballI) apply (simp add:A_to_prodag_def prod_pOp_def) apply (rotate_tac 2, frule_tac x = l in bspec, assumption, thin_tac "\k\I. S k \ aHom A (B k)") apply (simp add:prodag_def prod_pOp_def) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (B k)") apply (simp add: aHom_add) done definition finiteHom :: "['i set, 'i \ ('a, 'more) aGroup_scheme, 'i \ 'a] \ bool" where "finiteHom I A f \ f \ carr_prodag I A \ (\H. H \ I \ finite H \ ( \j \ (I - H). (f j) = \\<^bsub>(A j)\<^esub>))" definition carr_dsumag :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a ) set" where "carr_dsumag I A = {f. finiteHom I A f}" definition dsumag :: "['i set, 'i \ ('a, 'more) aGroup_scheme] \ ('i \ 'a) aGroup" where "dsumag I A = \ carrier = carr_dsumag I A, pop = prod_pOp I A, mop = prod_mOp I A, zero = prod_zero I A\" definition dProj :: "['i set, 'i \ ('a, 'more) aGroup_scheme, 'i] \ ('i \ 'a) \ 'a" where "dProj I A x = (\f\carr_dsumag I A. f x)" abbreviation DSUMag ("(a\\<^bsub>_\<^esub> _)" [72,73]72) where "a\\<^bsub>I\<^esub> A == dsumag I A" lemma dsum_pOp_func:"\k\I. aGroup (A k) \ prod_pOp I A \ carr_dsumag I A \ carr_dsumag I A \ carr_dsumag I A" apply (rule Pi_I)+ apply (subst carr_dsumag_def) apply (simp add:CollectI) apply (simp add:finiteHom_def) apply (rule conjI) apply (simp add:carr_dsumag_def) apply (simp add:finiteHom_def) apply (erule conjE)+ apply (simp add:prod_pOp_mem) apply (simp add:carr_dsumag_def finiteHom_def) apply (erule conjE)+ apply ((erule exE)+, (erule conjE)+) apply (frule_tac F = H and G = Ha in finite_UnI, assumption+) apply (subgoal_tac "\j\I - (H \ Ha). prod_pOp I A x xa j = \\<^bsub>A j\<^esub>") apply (frule_tac A = H and B = Ha in Un_least[of _ "I"], assumption+) apply blast apply (rule ballI) apply (simp, (erule conjE)+) apply (frule_tac x = j in bspec, assumption, thin_tac "\k\I. aGroup (A k)", frule_tac x = j in bspec, simp, thin_tac "\j\I - H. x j = \\<^bsub>A j\<^esub>", frule_tac x = j in bspec, simp, thin_tac "\j\I - Ha. xa j = \\<^bsub>A j\<^esub>") apply (simp add:prod_pOp_def) apply (rule aGroup.ag_l_zero) apply simp apply (rule aGroup.ex_zero) apply assumption done lemma dsum_pOp_mem:"\\k\I. aGroup (A k); X \ carr_dsumag I A; Y \ carr_dsumag I A\ \ prod_pOp I A X Y \ carr_dsumag I A" apply (frule dsum_pOp_func[of "I" "A"]) apply (frule funcset_mem[of "prod_pOp I A" "carr_dsumag I A" "carr_dsumag I A \ carr_dsumag I A" "X"], assumption+) apply (rule funcset_mem[of "prod_pOp I A X" "carr_dsumag I A" "carr_dsumag I A" "Y"], assumption+) done lemma dsum_iOp_func:"\k\I. aGroup (A k) \ prod_mOp I A \ carr_dsumag I A \ carr_dsumag I A" apply (rule Pi_I) apply (simp add:carr_dsumag_def) apply (simp add:finiteHom_def) apply (erule conjE)+ apply (simp add:prod_mOp_mem) apply (erule exE, (erule conjE)+) apply (simp add:prod_mOp_def) apply (subgoal_tac "\j\I - H. -\<^sub>a\<^bsub>A j\<^esub> (x j) = \\<^bsub>A j\<^esub>") apply blast apply (rule ballI) apply (frule_tac x = j in bspec, simp, thin_tac "\k\I. aGroup (A k)", frule_tac x = j in bspec, simp, thin_tac "\j\I - H. x j = \\<^bsub>A j\<^esub>", simp add:aGroup.ag_inv_zero) done lemma dsum_iOp_mem:"\\j\I. aGroup (A j); X \ carr_dsumag I A\ \ prod_mOp I A X \ carr_dsumag I A" apply (frule dsum_iOp_func) apply (simp add:Pi_def) done lemma dsum_zero_func:"\k\I. aGroup (A k) \ prod_zero I A \ carr_dsumag I A" apply (simp add:carr_dsumag_def) apply (simp add:finiteHom_def) apply (rule conjI) apply (simp add:prod_zero_func) apply (subgoal_tac "{} \ I") prefer 2 apply simp apply (subgoal_tac "finite {}") prefer 2 apply simp apply (subgoal_tac "\j\I - {}. prod_zero I A j = \\<^bsub>A j\<^esub>") apply blast apply (rule ballI) apply simp apply (simp add:prod_zero_def) done lemma dsumag_sub_prodag:"\k\I. aGroup (A k) \ carr_dsumag I A \ carr_prodag I A" by (rule subsetI, simp add:carr_dsumag_def finiteHom_def) lemma carrier_dsumag:"\k\I. aGroup (A k) \ carrier (dsumag I A) = carr_dsumag I A" apply (simp add:dsumag_def) done lemma dsumag_elemfun:"\\k\I. aGroup (A k); f \ carrier (dsumag I A)\ \ f \ extensional I" apply (simp add:carrier_dsumag) apply (simp add:carr_dsumag_def) apply (simp add:finiteHom_def) apply (erule conjE) apply (simp add:carr_prodag_def) done lemma dsumag_aGroup:"\k\I. aGroup (A k) \ aGroup (dsumag I A)" apply (simp add:aGroup_def [of "dsumag I A"]) apply (simp add:dsumag_def) apply (simp add:dsum_pOp_func) apply (simp add:dsum_iOp_func) apply (simp add:dsum_zero_func) apply (frule dsumag_sub_prodag[of "I" "A"]) apply (rule conjI) apply (rule allI, rule impI)+ apply (frule_tac X = a and Y = b in dsum_pOp_mem, assumption+) apply (frule_tac X = b and Y = c in dsum_pOp_mem, assumption+) apply (frule_tac X = "prod_pOp I A a b" and Y = c in dsum_pOp_mem, assumption+) apply (frule_tac Y = "prod_pOp I A b c" and X = a in dsum_pOp_mem, assumption+) apply (rule carr_prodag_mem_eq [of "I" "A"], assumption+) apply (simp add:subsetD) apply (simp add:subsetD) apply (rule ballI) apply (subst prod_pOp_mem_i, assumption+, (simp add:subsetD)+) apply (subst prod_pOp_mem_i, assumption+) apply (simp add:subsetD)+ apply (subst prod_pOp_mem_i, assumption+, (simp add:subsetD)+) apply (subst prod_pOp_mem_i, assumption+) apply (simp add:subsetD)+ apply (thin_tac "prod_pOp I A a b \ carr_dsumag I A", thin_tac "prod_pOp I A b c \ carr_dsumag I A", thin_tac "prod_pOp I A (prod_pOp I A a b) c \ carr_dsumag I A", thin_tac "prod_pOp I A a (prod_pOp I A b c) \ carr_dsumag I A", thin_tac "carr_dsumag I A \ carr_prodag I A") apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)", simp add:carr_dsumag_def finiteHom_def, (erule conjE)+, simp add:carr_prodag_def, (erule conjE)+) apply (frule_tac x = l in bspec, assumption, thin_tac "\i\I. a i \ carrier (A i)", frule_tac x = l in bspec, assumption, thin_tac "\i\I. b i \ carrier (A i)", frule_tac x = l in bspec, assumption, thin_tac "\i\I. c i \ carrier (A i)") apply (simp add:aGroup.aassoc) apply (rule conjI) apply (rule allI, rule impI)+ apply (rule carr_prodag_mem_eq [of "I" "A"], assumption+) apply (frule_tac X = a and Y = b in prod_pOp_mem[of "I" "A"], (simp add:subsetD)+) apply (frule_tac X = b and Y = a in prod_pOp_mem[of "I" "A"], (simp add:subsetD)+) apply (rule ballI, subst prod_pOp_mem_i, assumption+, (simp add:subsetD)+) apply (subst prod_pOp_mem_i, assumption+, (simp add:subsetD)+) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)") apply (frule_tac c = a in subsetD[of "carr_dsumag I A" "carr_prodag I A"], assumption+, thin_tac "a \ carr_dsumag I A", frule_tac c = b in subsetD[of "carr_dsumag I A" "carr_prodag I A"], assumption+, thin_tac "b \ carr_dsumag I A", thin_tac "carr_dsumag I A \ carr_prodag I A") apply (simp add:carr_prodag_def, (erule conjE)+, simp add:aGroup.ag_pOp_commute) apply (rule conjI) apply (rule allI, rule impI) apply (frule_tac X = a in prod_mOp_mem[of "I" "A"], simp add:subsetD) apply (frule_tac X = "prod_mOp I A a" and Y = a in prod_pOp_mem[of "I" "A"], simp add:subsetD, simp add:subsetD) apply (rule carr_prodag_mem_eq [of "I" "A"], assumption+, simp add:prod_zero_func) apply (rule ballI) apply (subst prod_pOp_mem_i, assumption+, simp add:subsetD, assumption) apply (subst prod_mOp_mem_i, assumption+, simp add:subsetD, assumption) apply (simp add:prod_zero_i) apply (frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)", thin_tac "prod_mOp I A a \ carr_prodag I A", thin_tac "prod_pOp I A (prod_mOp I A a) a \ carr_prodag I A", frule_tac c = a in subsetD[of "carr_dsumag I A" "carr_prodag I A"], assumption, thin_tac "carr_dsumag I A \ carr_prodag I A", simp add:carr_prodag_def, (erule conjE)+) apply (frule_tac x = l in bspec, assumption, thin_tac "\i\I. a i \ carrier (A i)") apply (rule aGroup.l_m, assumption+) apply (rule allI, rule impI) apply (frule prod_zero_func[of "I" "A"]) apply (frule_tac X = "prod_zero I A" and Y = a in prod_pOp_mem[of "I" "A"], assumption+, simp add:subsetD) apply (rule carr_prodag_mem_eq [of "I" "A"], assumption+, simp add:subsetD) apply (rule ballI) apply (subst prod_pOp_mem_i, assumption+) apply (simp add:subsetD, assumption) apply (simp add:prod_zero_i, frule_tac x = l in bspec, assumption, thin_tac "\k\I. aGroup (A k)", frule_tac c = a in subsetD[of "carr_dsumag I A" "carr_prodag I A"], assumption+, thin_tac "carr_dsumag I A \ carr_prodag I A", thin_tac "a \ carr_dsumag I A", thin_tac "prod_pOp I A (prod_zero I A) a \ carr_prodag I A") apply (simp add:carr_prodag_def, (erule conjE)+) apply (rule aGroup.l_zero, assumption) apply blast done lemma dsumag_pOp:"\k\I. aGroup (A k) \ pop (dsumag I A) = prod_pOp I A" apply (simp add:dsumag_def) done lemma dsumag_mOp:"\k\I. aGroup (A k) \ mop (dsumag I A) = prod_mOp I A" apply (simp add:dsumag_def) done lemma dsumag_zero:"\k\I. aGroup (A k) \ zero (dsumag I A) = prod_zero I A" apply (simp add:dsumag_def) done subsection "Characterization of a direct product" lemma direct_prod_mem_eq:"\\j\I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A); g \ carrier (a\\<^bsub>I\<^esub> A); \j\I. (PRoject I A j) f = (PRoject I A j) g\ \ f = g" apply (rule funcset_eq[of "f" "I" "g"]) apply (thin_tac "\j\I. aGroup (A j)", thin_tac "g \ carrier (a\\<^bsub>I\<^esub> A)", thin_tac "\j\I. (\\<^bsub>I,A,j\<^esub>) f = (\\<^bsub>I,A,j\<^esub>) g", simp add:prodag_def carr_prodag_def) apply (thin_tac "\j\I. aGroup (A j)", thin_tac "f \ carrier (a\\<^bsub>I\<^esub> A)", thin_tac "\j\I. (\\<^bsub>I,A,j\<^esub>) f = (\\<^bsub>I,A,j\<^esub>) g", simp add:prodag_def carr_prodag_def) apply (simp add:PRoject_def prodag_def) done lemma map_family_fun:"\\j\I. aGroup (A j); aGroup S; \j\I. ((g j) \ aHom S (A j)); x \ carrier S\ \ (\y \ carrier S. (\j\I. (g j) y)) x \ carrier (a\\<^bsub>I\<^esub> A)" apply (simp add:prodag_def carr_prodag_def) apply (simp add:aHom_mem) apply (rule Pi_I, simp add:Un_carrier_def) apply (frule_tac x = xa in bspec, assumption, thin_tac "\j\I. aGroup (A j)", frule_tac x = xa in bspec, assumption, thin_tac "\j\I. g j \ aHom S (A j)") apply (frule_tac G = "A xa" and f = "g xa" and a = x in aHom_mem[of "S"], assumption+, blast) done lemma map_family_aHom:"\\j\I. aGroup (A j); aGroup S; \j\I. ((g j) \ aHom S (A j))\ \ (\y \ carrier S. (\j\I. (g j) y)) \ aHom S (a\\<^bsub>I\<^esub> A)" apply (subst aHom_def, simp) apply (simp add:aGroup.ag_pOp_closed) apply (rule conjI) apply (rule Pi_I) apply (rule map_family_fun[of "I" "A" "S" "g"], assumption+) apply (rule ballI)+ apply (frule_tac x = a and y = b in aGroup.ag_pOp_closed[of "S"], assumption+) apply (frule_tac x = "a \\<^bsub>S\<^esub> b" in map_family_fun[of "I" "A" "S" "g"], assumption+, simp) apply (frule_tac x = a in map_family_fun[of "I" "A" "S" "g"], assumption+, simp, frule_tac x = b in map_family_fun[of "I" "A" "S" "g"], assumption+, simp) apply (frule prodag_aGroup[of "I" "A"]) apply (frule_tac x = "(\j\I. g j a)" and y = "(\j\I. g j b)" in aGroup.ag_pOp_closed[of "a\\<^bsub>I\<^esub> A"], assumption+) apply (simp only:prodag_carrier) apply (rule carr_prodag_mem_eq, assumption+) apply (rule ballI) apply (subst prodag_def, simp add:prod_pOp_def) apply (simp add:aHom_add) done lemma map_family_triangle:"\\j\I. aGroup (A j); aGroup S; \j\I. ((g j) \ aHom S (A j))\ \ \!f. f \ aHom S (a\\<^bsub>I\<^esub> A) \ (\j\I. compos S (PRoject I A j) f = (g j))" apply (rule ex_ex1I) apply (frule map_family_aHom[of "I" "A" "S" "g"], assumption+) apply (subgoal_tac "\j\I. compos S (\\<^bsub>I,A,j\<^esub>) (\y\carrier S. \j\I. g j y) = g j") apply blast apply (rule ballI) apply (simp add:compos_def) apply (rule funcset_eq[of _ "carrier S"]) apply (simp add:compose_def) apply (simp add:aHom_def) apply (rule ballI) apply (frule prodag_aGroup[of "I" "A"]) apply (frule prodag_carrier[of "I" "A"]) apply (frule_tac f = "\y\carrier S. \j\I. g j y" and a = x in aHom_mem[of "S" "a\\<^bsub>I\<^esub> A"], assumption+) apply (simp add:compose_def, simp add:PRoject_def) apply (rename_tac f f1) apply (erule conjE)+ apply (rule funcset_eq[of _ "carrier S"]) apply (simp add:aHom_def, simp add:aHom_def) apply (rule ballI) apply (frule prodag_aGroup[of "I" "A"]) apply (frule_tac f = f and a = x in aHom_mem[of "S" "a\\<^bsub>I\<^esub> A"], assumption+, frule_tac f = f1 and a = x in aHom_mem[of "S" "a\\<^bsub>I\<^esub> A"], assumption+) apply (rule_tac f = "f x" and g = "f1 x" in direct_prod_mem_eq[of "I" "A"], assumption+) apply (rule ballI) apply (rotate_tac 4, frule_tac x = j in bspec, assumption, thin_tac "\j\I. compos S (\\<^bsub>I,A,j\<^esub>) f = g j", frule_tac x = j in bspec, assumption, thin_tac "\j\I. compos S (\\<^bsub>I,A,j\<^esub>) f1 = g j", simp add:compos_def compose_def) apply (subgoal_tac "(\x\carrier S. (\\<^bsub>I,A,j\<^esub>) (f x)) x = g j x", subgoal_tac "(\x\carrier S. (\\<^bsub>I,A,j\<^esub>) (f1 x)) x = g j x", thin_tac "(\x\carrier S. (\\<^bsub>I,A,j\<^esub>) (f x)) = g j", thin_tac "(\x\carrier S. (\\<^bsub>I,A,j\<^esub>) (f1 x)) = g j", simp+) done lemma Ag_ind_triangle:"\\j\I. aGroup (A j); j \ I; f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) (B::'d set); j \ I\ \ compos (a\\<^bsub>I\<^esub> A) (compos (Ag_ind (a\\<^bsub>I\<^esub> A) f)(PRoject I A j) (ainvf\<^bsub>(a\\<^bsub>I\<^esub> A), (Ag_ind (a\\<^bsub>I\<^esub> A) f)\<^esub> (Agii (a\\<^bsub>I\<^esub> A) f))) (Agii (a\\<^bsub>I\<^esub> A) f) = PRoject I A j" apply (frule prodag_aGroup[of "I" "A"]) apply (frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (simp add:compos_def) apply (rule funcset_eq[of _ "carrier (a\\<^bsub>I\<^esub> A)"]) apply simp apply (simp add:PRoject_def prodag_carrier extensional_def) apply (rule ballI) apply (simp add:compose_def invfun_l) apply (simp add:aGroup.Agii_mem) apply (frule Ag_ind_bijec[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule_tac x = x in ainvf_l[of "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f" "Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply simp done (** Note f' a\\<^bsub>I\<^esub> A \ Ag_ind (a\\<^bsub>I\<^esub> A) f \ | \ | PRoject I A j \ | (PRoject I A j) o (f'\\<^sup>1) \ | A j , where f' = Agii (a\\<^bsub>I\<^esub> A) f **) definition ProjInd :: "['i set, 'i \ ('a, 'm) aGroup_scheme, ('i \ 'a) \ 'd, 'i] \ ('d \ 'a)" where "ProjInd I A f j = compos (Ag_ind (a\\<^bsub>I\<^esub> A) f)(PRoject I A j) (ainvf\<^bsub>(a\\<^bsub>I\<^esub> A), (Ag_ind (a\\<^bsub>I\<^esub> A) f)\<^esub> (Agii (a\\<^bsub>I\<^esub> A) f))" (** Note f' a\\<^bsub>I\<^esub> A \ Ag_ind (a\\<^bsub>I\<^esub> A) f \ | \ | PRoject I A j \ | PRojInd I A f j \ | A j **) lemma ProjInd_aHom:"\\j\ I. aGroup (A j); j \ I; f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) (B::'d set); j \ I\ \ (ProjInd I A f j) \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) f) (A j)" apply (frule prodag_aGroup[of "I" "A"]) apply (frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule_tac x = j in bspec, assumption) apply (frule aGroup.Ag_ind_aHom[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (simp add:ProjInd_def) apply (frule Ag_ind_bijec[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule ainvf_aHom[of "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f" "Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (frule project_aHom[of "I" "A" "j"], assumption) apply (simp add:aHom_compos) done lemma ProjInd_aHom1:"\\j\ I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) (B::'d set)\ \ \j\I. (ProjInd I A f j) \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) f) (A j)" apply (rule ballI) apply (simp add:ProjInd_aHom) done lemma ProjInd_mem_eq:"\\j\I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) B; aGroup S; x \ carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f); y \ carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f); \j\I. (ProjInd I A f j x = ProjInd I A f j y)\ \ x = y" apply (simp add:ProjInd_def) apply (simp add:compos_def compose_def) apply (frule prodag_aGroup[of "I" "A"]) apply (frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule aGroup.Ag_ind_aHom[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule Ag_ind_bijec[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule ainvf_aHom[of "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f" "Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (frule aHom_mem[of "Ag_ind (a\\<^bsub>I\<^esub> A) f" "a\\<^bsub>I\<^esub> A" "ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f" "x"], assumption+, frule aHom_mem[of "Ag_ind (a\\<^bsub>I\<^esub> A) f" "a\\<^bsub>I\<^esub> A" "ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f" "y"], assumption+) apply (frule direct_prod_mem_eq[of "I" "A" "(ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f) x" "(ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f) y"], assumption+) apply (thin_tac "ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) f) (a\\<^bsub>I\<^esub> A)") apply (frule ainvf_bijec[of "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f" "Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (thin_tac "bijec\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f") apply (unfold bijec_def, frule conjunct1, fold bijec_def) apply (frule injec_inj_on[of "Ag_ind (a\\<^bsub>I\<^esub> A) f" "a\\<^bsub>I\<^esub> A" "ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (simp add:injective_iff[THEN sym, of "ainvf\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) f\<^esub> Agii (a\\<^bsub>I\<^esub> A) f" "carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f)" "x" "y"]) done lemma ProjInd_mem_eq1:"\\j\I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) B; aGroup S; h \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) f) (Ag_ind (a\\<^bsub>I\<^esub> A) f); \j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) f) (ProjInd I A f j) h = ProjInd I A f j\ \ h = ag_idmap (Ag_ind (a\\<^bsub>I\<^esub> A) f)" apply (rule funcset_eq[of _ "carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f)"]) apply (simp add:aHom_def) apply (simp add:ag_idmap_def) apply (rule ballI) apply (simp add:ag_idmap_def) apply (frule prodag_aGroup[of "I" "A"], frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule_tac a = x in aHom_mem[of "Ag_ind (a\\<^bsub>I\<^esub> A) f" "Ag_ind (a\\<^bsub>I\<^esub> A) f" "h"], assumption+) apply (rule_tac x = "h x" and y = x in ProjInd_mem_eq[of "I" "A" "f" "B" "S"], assumption+) apply (rotate_tac 1, rule ballI, frule_tac x = j in bspec, assumption, thin_tac "\j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) f) (ProjInd I A f j) h = ProjInd I A f j") apply (simp add:compos_def compose_def) apply (subgoal_tac "(\x\carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f). ProjInd I A f j (h x)) x = ProjInd I A f j x", thin_tac "(\x\carrier (Ag_ind (a\\<^bsub>I\<^esub> A) f). ProjInd I A f j (h x)) = ProjInd I A f j") apply simp+ done lemma Ag_ind_triangle1:"\\j\I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) (B::'d set); j \ I\ \ compos (a\\<^bsub>I\<^esub> A) (ProjInd I A f j) (Agii (a\\<^bsub>I\<^esub> A) f) = PRoject I A j" apply (simp add:ProjInd_def) apply (simp add:Ag_ind_triangle) done lemma map_family_triangle1:"\\j\I. aGroup (A j); f \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to f (carrier (a\\<^bsub>I\<^esub> A)) (B::'d set); aGroup S; \j\I. ((g j) \ aHom S (A j))\ \ \!h. h \ aHom S (Ag_ind (a\\<^bsub>I\<^esub> A) f) \ (\j\I. compos S (ProjInd I A f j) h = (g j))" apply (frule prodag_aGroup[of "I" "A"]) apply (frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (frule Ag_ind_bijec[of "a\\<^bsub>I\<^esub> A" "f" "B"], assumption+) apply (rule ex_ex1I) apply (frule map_family_triangle[of "I" "A" "S" "g"], assumption+) apply (frule ex1_implies_ex) apply (erule exE) apply (erule conjE) apply (unfold bijec_def, frule conjunct2, fold bijec_def) apply (unfold surjec_def, frule conjunct1, fold surjec_def) apply (rename_tac fa, frule_tac f = fa in aHom_compos[of "S" "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f" _ "Agii (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (subgoal_tac "\j\I. compos S (ProjInd I A f j) (compos S (Agii (a\\<^bsub>I\<^esub> A) f) fa) = g j") apply blast apply (rule ballI) apply (frule_tac N = "A j" and f = fa and g = "Agii (a\\<^bsub>I\<^esub> A) f" and h = "ProjInd I A f j" in aHom_compos_assoc[of "S" "a\\<^bsub>I\<^esub> A" "Ag_ind (a\\<^bsub>I\<^esub> A) f"], assumption+) apply simp apply assumption+ apply (simp add:ProjInd_aHom) apply simp apply (thin_tac "compos S (ProjInd I A f j) (compos S (Agii (a\\<^bsub>I\<^esub> A) f) fa) = compos S (compos (a\\<^bsub>I\<^esub> A) (ProjInd I A f j) (Agii (a\\<^bsub>I\<^esub> A) f)) fa") apply (simp add:Ag_ind_triangle1) apply (rename_tac h h1) apply (erule conjE)+ apply (rule funcset_eq[of _ "carrier S"]) apply (simp add:aHom_def, simp add:aHom_def) apply (rule ballI) apply (simp add:compos_def) apply (frule_tac f = h and a = x in aHom_mem[of "S" "Ag_ind (a\\<^bsub>I\<^esub> A) f"], assumption+, frule_tac f = h1 and a = x in aHom_mem[of "S" "Ag_ind (a\\<^bsub>I\<^esub> A) f"], assumption+) apply (rule_tac x = "h x" and y = "h1 x" in ProjInd_mem_eq[of "I" "A" "f" "B" "S"], assumption+) apply (rule ballI) apply (rotate_tac 5, frule_tac x = j in bspec, assumption, thin_tac "\j\I. compose (carrier S) (ProjInd I A f j) h = g j", frule_tac x = j in bspec, assumption, thin_tac "\j\I. compose (carrier S) (ProjInd I A f j) h1 = g j") apply (simp add:compose_def, subgoal_tac "(\x\carrier S. ProjInd I A f j (h x)) x = g j x", thin_tac "(\x\carrier S. ProjInd I A f j (h x)) = g j", subgoal_tac "(\x\carrier S. ProjInd I A f j (h1 x)) x = g j x", thin_tac "(\x\carrier S. ProjInd I A f j (h1 x)) = g j", simp+) done lemma map_family_triangle2:"\I \ {}; \j\I. aGroup (A j); aGroup S; \j\I. g j \ aHom S (A j); ff \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to ff (carrier (a\\<^bsub>I\<^esub> A)) B; h1 \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) ff) S; \j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (g j) h1 = ProjInd I A ff j; h2 \ aHom S (Ag_ind (a\\<^bsub>I\<^esub> A) ff); \j\I. compos S (ProjInd I A ff j) h2 = g j\ \ \j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (ProjInd I A ff j) (compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) h2 h1) = ProjInd I A ff j" apply (rule ballI) apply (frule prodag_aGroup[of "I" "A"]) apply (frule_tac f = ff in aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" _ "B"], assumption+) apply (frule_tac N = "A j" and h = "ProjInd I A ff j" in aHom_compos_assoc[of "Ag_ind (a\\<^bsub>I\<^esub> A) ff" "S" "Ag_ind (a\\<^bsub>I\<^esub> A) ff" _ "h1" "h2"], assumption+) apply simp apply assumption+ apply (simp add:ProjInd_aHom) apply simp done lemma map_family_triangle3:"\\j\I. aGroup (A j); aGroup S; aGroup S1; \j\I. f j \ aHom S (A j); \j\I. g j \ aHom S1 (A j); h1 \ aHom S1 S; h2 \ aHom S S1; \j\I. compos S (g j) h2 = f j; \j\I. compos S1 (f j) h1 = g j\ \ \j\I. compos S (f j) (compos S h1 h2) = f j" apply (rule ballI) apply (frule_tac h = "f j" and N = "A j" in aHom_compos_assoc[of "S" "S1" "S" _ "h2" "h1"], assumption+) apply simp apply assumption+ apply simp apply simp done lemma map_family_triangle4:"\\j\I. aGroup (A j); aGroup S; \j\I. f j \ aHom S (A j)\ \ \j\I. compos S (f j) (ag_idmap S) = f j" apply (rule ballI) apply (frule_tac x = j in bspec, assumption, thin_tac "\j\I. aGroup (A j)", frule_tac x = j in bspec, assumption, thin_tac "\j\I. f j \ aHom S (A j)") apply (simp add:compos_aI_r) done lemma prod_triangle:"\I \ {}; \j\I. aGroup (A j); aGroup S; \j\I. g j \ aHom S (A j); ff \ carrier (a\\<^bsub>I\<^esub> A) \ B; bij_to ff (carrier (a\\<^bsub>I\<^esub> A)) B; h1 \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) ff) S; \j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (g j) h1 = ProjInd I A ff j; h2 \ aHom S (Ag_ind (a\\<^bsub>I\<^esub> A) ff); \j\I. compos S (ProjInd I A ff j) h2 = g j\ \ (compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) h2 h1) = ag_idmap (Ag_ind (a\\<^bsub>I\<^esub> A) ff)" apply (frule map_family_triangle2[of "I" "A" "S" "g" "ff" "B" "h1" "h2"], assumption+) apply (frule prodag_aGroup[of "I" "A"], frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" "ff" "B"], assumption+) apply (frule aHom_compos[of "Ag_ind (a\\<^bsub>I\<^esub> A) ff" "S" "Ag_ind (a\\<^bsub>I\<^esub> A) ff" "h1" "h2"], assumption+) apply (rule ProjInd_mem_eq1[of "I" "A" "ff" "B" "S" "compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) h2 h1"], assumption+) done lemma characterization_prodag:"\I \ {}; \j\(I::'i set). aGroup ((A j):: ('a, 'm) aGroup_scheme); aGroup (S::'d aGroup); \j\I. ((g j) \ aHom S (A j)); \ff. ff \ carrier (a\\<^bsub>I\<^esub> A) \ (B::'d set) \ bij_to ff (carrier (a\\<^bsub>I\<^esub> A)) B; \(S':: 'd aGroup). aGroup S' \ (\g'. (\j\I. (g' j) \ aHom S' (A j) \ (\! f. f \ aHom S' S \ (\j\I. compos S' (g j) f = (g' j)))))\ \ \h. bijec\<^bsub>(prodag I A),S\<^esub> h" apply (frule prodag_aGroup[of "I" "A"]) apply (erule exE) apply (frule_tac f = ff in aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" _ "B"], erule conjE, assumption, simp, erule conjE) apply (frule aGroup.Ag_ind_aGroup[of "a\\<^bsub>I\<^esub> A" _ "B"], assumption+, frule_tac a = S in forall_spec, assumption+) apply (rotate_tac -1, frule_tac x = g in spec, thin_tac "\g'. \j\I. g' j \ aHom S (A j) \ (\!f. f \ aHom S S \ (\j\I. compos S (g j) f = g' j))") apply (frule_tac a = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" in forall_spec, assumption+, thin_tac "\S'. aGroup S' \ (\g'. \j\I. g' j \ aHom S' (A j) \ (\!f. f \ aHom S' S \ (\j\I. compos S' (g j) f = g' j)))") apply (frule_tac x = "ProjInd I A ff" in spec, thin_tac "\g'. \j\I. g' j \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (A j) \ (\!f. f \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) ff) S \ (\j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (g j) f = g' j))") apply (frule_tac f = ff in ProjInd_aHom1[of "I" "A" _ "B"], assumption+) apply (simp add:nonempty_ex[of "I"], rotate_tac -2, frule ex1_implies_ex, thin_tac "\!f. f \ aHom (Ag_ind (a\\<^bsub>I\<^esub> A) ff) S \ (\j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (g j) f = ProjInd I A ff j)", rotate_tac -1, erule exE, erule conjE) apply (rename_tac ff h1, frule_tac f = ff in map_family_triangle1[of "I" "A" _ "B" "S" "g"], assumption+, rotate_tac -1, frule ex1_implies_ex, thin_tac "\!h. h \ aHom S (Ag_ind (a\\<^bsub>I\<^esub> A) ff) \ (\j\I. compos S (ProjInd I A ff j) h = g j)", rotate_tac -1, erule exE, erule conjE) apply (rename_tac ff h1 h2) apply (frule_tac ff = ff and ?h1.0 = h1 and ?h2.0 = h2 in prod_triangle[of "I" "A" "S" "g" _ "B"], assumption+, frule_tac ?S1.0 = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" in map_family_triangle3[of "I" "A" "S" _ "g"], assumption+, frule_tac f = h2 and g = h1 and M = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" in aHom_compos[of "S" _ "S" ], assumption+) apply (erule ex1E) apply (rotate_tac -1, frule_tac x = "compos S h1 h2" in spec, frule map_family_triangle4[of "I" "A" "S" "g"], assumption+, frule aGroup.aI_aHom[of "S"]) apply (frule_tac x = "aI\<^bsub>S\<^esub>" in spec, thin_tac "\y. y \ aHom S S \ (\j\I. compos S (g j) y = g j) \ y = f", simp, thin_tac "\j\I. compos S (ProjInd I A ff j) h2 = g j", thin_tac "\j\I. compos S (g j) f = g j", thin_tac "\j\I. compos (Ag_ind (a\\<^bsub>I\<^esub> A) ff) (g j) h1 = ProjInd I A ff j") apply (rotate_tac -1, frule sym, thin_tac "aI\<^bsub>S\<^esub> = f", simp, frule_tac A = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" and f = h1 and g = h2 in compos_aI_inj[of _ "S"], assumption+, frule_tac B = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" and f = h2 and g = h1 in compos_aI_surj[of "S"], assumption+) apply (frule_tac f = ff in Ag_ind_bijec[of "a\\<^bsub>I\<^esub> A" _ "B"], assumption+, frule_tac F = "Ag_ind (a\\<^bsub>I\<^esub> A) ff" and f = "Agii (a\\<^bsub>I\<^esub> A) ff" and g = h1 in compos_bijec[of "a\\<^bsub>I\<^esub> A" _ "S"], assumption+) apply (subst bijec_def, simp) apply (thin_tac "bijec\<^bsub>(a\\<^bsub>I\<^esub> A),Ag_ind (a\\<^bsub>I\<^esub> A) ff\<^esub> Agii (a\\<^bsub>I\<^esub> A) ff", thin_tac "injec\<^bsub>Ag_ind (a\\<^bsub>I\<^esub> A) ff,S\<^esub> h1", thin_tac "surjec\<^bsub>Ag_ind (a\\<^bsub>I\<^esub> A) ff,S\<^esub> h1") apply (rule exI, simp) done (*** Note. f S' \ S \ | g' j\ | g j \ | A j ***) chapter "Ring theory" section "Definition of a ring and an ideal" record 'a Ring = "'a aGroup" + tp :: "['a, 'a ] \ 'a" (infixl "\\<^sub>r\" 70) un :: "'a" ("1\<^sub>r\") locale Ring = fixes R (structure) assumes pop_closed: "pop R \ carrier R \ carrier R \ carrier R" and pop_aassoc : "\a \ carrier R; b \ carrier R; c \ carrier R\ \ (a \ b) \ c = a \ (b \ c)" and pop_commute:"\a \ carrier R; b \ carrier R\ \ a \ b = b \ a" and mop_closed:"mop R \ carrier R \ carrier R" and l_m :"a \ carrier R \ (-\<^sub>a a) \ a = \" and ex_zero: "\ \ carrier R" and l_zero:"a \ carrier R \ \ \ a = a" and tp_closed: "tp R \ carrier R \ carrier R \ carrier R" and tp_assoc : "\a \ carrier R; b \ carrier R; c \ carrier R\ \ (a \\<^sub>r b) \\<^sub>r c = a \\<^sub>r (b \\<^sub>r c)" and tp_commute: "\a \ carrier R; b \ carrier R\ \ a \\<^sub>r b = b \\<^sub>r a" and un_closed: "(1\<^sub>r) \ carrier R" and rg_distrib: "\a \ carrier R; b \ carrier R; c \ carrier R\ \ a \\<^sub>r (b \ c) = a \\<^sub>r b \ a \\<^sub>r c" and rg_l_unit: "a \ carrier R \ (1\<^sub>r) \\<^sub>r a = a" definition zeroring :: "('a, 'more) Ring_scheme \ bool" where "zeroring R \ Ring R \ carrier R = {\\<^bsub>R\<^esub>}" primrec nscal :: "('a, 'more) Ring_scheme => 'a => nat => 'a" where nscal_0: "nscal R x 0 = \\<^bsub>R\<^esub>" | nscal_suc: "nscal R x (Suc n) = (nscal R x n) \\<^bsub>R\<^esub> x" primrec npow :: "('a, 'more) Ring_scheme => 'a => nat => 'a" where npow_0: "npow R x 0 = 1\<^sub>r\<^bsub>R\<^esub>" | npow_suc: "npow R x (Suc n) = (npow R x n) \\<^sub>r\<^bsub>R\<^esub> x" primrec nprod :: "('a, 'more) Ring_scheme => (nat => 'a) => nat => 'a" where nprod_0: "nprod R f 0 = f 0" | nprod_suc: "nprod R f (Suc n) = (nprod R f n) \\<^sub>r\<^bsub>R\<^esub> (f (Suc n))" primrec nsum :: "('a, 'more) aGroup_scheme => (nat => 'a) => nat => 'a" where nsum_0: "nsum R f 0 = f 0" | nsum_suc: "nsum R f (Suc n) = (nsum R f n) \\<^bsub>R\<^esub> (f (Suc n))" abbreviation NSCAL :: "[nat, ('a, 'more) Ring_scheme, 'a] \ 'a" ("(3 _ \\<^bsub>_\<^esub> _)" [75,75,76]75) where "n \\<^bsub>R\<^esub> x == nscal R x n" abbreviation NPOW :: "['a, ('a, 'more) Ring_scheme, nat] \ 'a" ("(3_^\<^bsup>_ _\<^esup>)" [77,77,78]77) where "a^\<^bsup>R n\<^esup> == npow R a n" abbreviation SUM :: "('a, 'more) aGroup_scheme => (nat => 'a) => nat => 'a" ("(3\\<^sub>e _ _ _)" [85,85,86]85) where "\\<^sub>e G f n == nsum G f n" abbreviation NPROD :: "[('a, 'm) Ring_scheme, nat, nat \ 'a] \ 'a" ("(3e\\<^bsub>_,_\<^esub> _)" [98,98,99]98) where "e\\<^bsub>R,n\<^esub> f == nprod R f n" definition fSum :: "[_, (nat => 'a), nat, nat] \ 'a" where "fSum A f n m = (if n \ m then nsum A (cmp f (slide n))(m - n) else \\<^bsub>A\<^esub>)" abbreviation FSUM :: "[('a, 'more) aGroup_scheme, (nat \ 'a), nat, nat] \ 'a" ("(4\\<^sub>f _ _ _ _)" [85,85,85,86]85) where "\\<^sub>f G f n m == fSum G f n m" lemma (in aGroup) nsum_zeroGTr:"(\j \ n. f j = \) \ nsum A f n = \" apply (induct_tac n) apply (rule impI, simp) apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (cut_tac ex_zero) apply (simp add:l_zero[of \]) done lemma (in aGroup) nsum_zeroA:"\j \ n. f j = \ \ nsum A f n = \" apply (simp add:nsum_zeroGTr) done definition sr :: "[_ , 'a set] \ bool" where "sr R S == S \ carrier R \ 1\<^sub>r\<^bsub>R\<^esub> \ S \ (\x\S. \y \ S. x \\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> y) \ S \ x \\<^sub>r\<^bsub>R\<^esub> y \ S)" definition Sr :: "[_ , 'a set] \ _" where "Sr R S = R \carrier := S, pop := \x\S. \y\S. x \\<^bsub>R\<^esub> y, mop := \x\S. (-\<^sub>a\<^bsub>R\<^esub> x), zero := \\<^bsub>R\<^esub>, tp := \x\S. \y\S. x \\<^sub>r\<^bsub>R\<^esub> y, un := 1\<^sub>r\<^bsub>R\<^esub> \" (** sr is a subring without ring structure, Sr is a subring with Ring structure **) lemma (in Ring) Ring: "Ring R" .. lemma (in Ring) ring_is_ag:"aGroup R" apply (rule aGroup.intro, rule pop_closed, rule pop_aassoc, assumption+, rule pop_commute, assumption+, rule mop_closed, rule l_m, assumption, rule ex_zero, rule l_zero, assumption) done lemma (in Ring) ring_zero:"\ \ carrier R" by (simp add: ex_zero) lemma (in Ring) ring_one:"1\<^sub>r \ carrier R" by (simp add:un_closed) lemma (in Ring) ring_tOp_closed:"\ x \ carrier R; y \ carrier R\ \ x \\<^sub>r y \ carrier R" apply (cut_tac tp_closed) apply (frule funcset_mem[of "(\\<^sub>r)" "carrier R" "carrier R \ carrier R" "x"], assumption+, thin_tac "(\\<^sub>r) \ carrier R \ carrier R \ carrier R") apply (rule funcset_mem[of "(\\<^sub>r) x" "carrier R" "carrier R" "y"], assumption+) done lemma (in Ring) ring_tOp_commute:"\x \ carrier R; y \ carrier R\ \ x \\<^sub>r y = y \\<^sub>r x" by (simp add:tp_commute) lemma (in Ring) ring_distrib1:"\x \ carrier R; y \ carrier R; z \ carrier R \ \ x \\<^sub>r (y \ z) = x \\<^sub>r y \ x \\<^sub>r z" by (simp add:rg_distrib) lemma (in Ring) ring_distrib2:"\x \ carrier R; y \ carrier R; z \ carrier R \ \ (y \ z) \\<^sub>r x = y \\<^sub>r x \ z \\<^sub>r x" apply (subst tp_commute[of "y \ z" "x"]) apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_closed) apply assumption apply (subst ring_distrib1, assumption+) apply (simp add:tp_commute) done lemma (in Ring) ring_distrib3:"\a \ carrier R; b \ carrier R; x \ carrier R; y \ carrier R \ \ (a \ b) \\<^sub>r (x \ y) = a \\<^sub>r x \ a \\<^sub>r y \ b \\<^sub>r x \ b \\<^sub>r y" apply (subst ring_distrib2)+ apply (cut_tac ring_is_ag) apply (rule aGroup.ag_pOp_closed, assumption+) apply ((subst ring_distrib1)+, assumption+) apply (subst ring_distrib1, assumption+) apply (rule pop_aassoc [THEN sym, of "a \\<^sub>r x \ a \\<^sub>r y" "b \\<^sub>r x" "b \\<^sub>r y"]) apply (cut_tac ring_is_ag, rule aGroup.ag_pOp_closed, assumption) apply (simp add:ring_tOp_closed)+ done lemma (in Ring) rEQMulR: "\x \ carrier R; y \ carrier R; z \ carrier R; x = y \ \ x \\<^sub>r z = y \\<^sub>r z" by simp lemma (in Ring) ring_tOp_assoc:"\x \ carrier R; y \ carrier R; z \ carrier R \ \ (x \\<^sub>r y) \\<^sub>r z = x \\<^sub>r (y \\<^sub>r z)" by (simp add:tp_assoc) lemma (in Ring) ring_l_one:"x \ carrier R \ 1\<^sub>r \\<^sub>r x = x" by (simp add:rg_l_unit) lemma (in Ring) ring_r_one:"x \ carrier R \ x \\<^sub>r 1\<^sub>r = x" apply (subst ring_tOp_commute, assumption+) apply (simp add:un_closed) apply (simp add:ring_l_one) done lemma (in Ring) ring_times_0_x:"x \ carrier R \ \ \\<^sub>r x = \" apply (cut_tac ring_is_ag) apply (cut_tac ring_zero) apply (frule ring_distrib2 [of "x" "\" "\"], assumption+) apply (simp add:aGroup.ag_l_zero [of "R" "\"]) apply (frule ring_tOp_closed [of "\" "x"], assumption+) apply (frule sym, thin_tac "\ \\<^sub>r x = \ \\<^sub>r x \ \ \\<^sub>r x") apply (frule aGroup.ag_eq_sol2 [of "R" "\ \\<^sub>r x" "\ \\<^sub>r x" "\ \\<^sub>r x"], assumption+) apply (thin_tac "\ \\<^sub>r x \ \ \\<^sub>r x = \ \\<^sub>r x") apply (simp add:aGroup.ag_r_inv1) done lemma (in Ring) ring_times_x_0:"x \ carrier R \ x \\<^sub>r \ = \" apply (cut_tac ring_zero) apply (subst ring_tOp_commute, assumption+, simp add:ring_zero) apply (simp add:ring_times_0_x) done lemma (in Ring) rMulZeroDiv: "\ x \ carrier R; y \ carrier R; x = \ \ y = \ \ \ x \\<^sub>r y = \" apply (erule disjE, simp) apply (rule ring_times_0_x, assumption+) apply (simp, rule ring_times_x_0, assumption+) done lemma (in Ring) ring_inv1:"\ a \ carrier R; b \ carrier R \ \ -\<^sub>a (a \\<^sub>r b) = (-\<^sub>a a) \\<^sub>r b \ -\<^sub>a (a \\<^sub>r b) = a \\<^sub>r (-\<^sub>a b)" apply (cut_tac ring_is_ag) apply (rule conjI) apply (frule ring_distrib2 [THEN sym, of "b" "a" "-\<^sub>a a"], assumption+) apply (frule aGroup.ag_mOp_closed [of "R" "a"], assumption+) apply (simp add:aGroup.ag_r_inv1 [of "R" "a"]) apply (simp add:ring_times_0_x) apply (frule aGroup.ag_mOp_closed [of "R" "a"], assumption+) apply (frule ring_tOp_closed [of "a" "b"], assumption+) apply (frule ring_tOp_closed [of "-\<^sub>a a" "b"], assumption+) apply (frule aGroup.ag_eq_sol1 [of "R" "a \\<^sub>r b" "(-\<^sub>a a) \\<^sub>r b" "\"], assumption+) apply (rule ring_zero, assumption+) apply (thin_tac "a \\<^sub>r b \ (-\<^sub>a a) \\<^sub>r b = \") apply (frule sym) apply (thin_tac "(-\<^sub>a a) \\<^sub>r b = -\<^sub>a (a \\<^sub>r b) \ \") apply (frule aGroup.ag_mOp_closed [of "R" " a \\<^sub>r b"], assumption+) apply (simp add:aGroup.ag_r_zero) apply (frule ring_distrib1 [THEN sym, of "a" "b" "-\<^sub>a b"], assumption+) apply (simp add:aGroup.ag_mOp_closed) apply (simp add:aGroup.ag_r_inv1 [of "R" "b"]) apply (simp add:ring_times_x_0) apply (frule aGroup.ag_mOp_closed [of "R" "b"], assumption+) apply (frule ring_tOp_closed [of "a" "b"], assumption+) apply (frule ring_tOp_closed [of "a" "-\<^sub>a b"], assumption+) apply (frule aGroup.ag_eq_sol1 [THEN sym, of "R" "a \\<^sub>r b" "a \\<^sub>r (-\<^sub>a b)" "\"], assumption+) apply (simp add:ring_zero) apply assumption apply (frule aGroup.ag_mOp_closed [of "R" " a \\<^sub>r b"], assumption+) apply (simp add:aGroup.ag_r_zero) done lemma (in Ring) ring_inv1_1:"\a \ carrier R; b \ carrier R \ \ -\<^sub>a (a \\<^sub>r b) = (-\<^sub>a a) \\<^sub>r b" apply (simp add:ring_inv1) done lemma (in Ring) ring_inv1_2:"\ a \ carrier R; b \ carrier R \ \ -\<^sub>a (a \\<^sub>r b) = a \\<^sub>r (-\<^sub>a b)" apply (frule ring_inv1 [of "a" "b"], assumption+) apply (frule conjunct2) apply (thin_tac "-\<^sub>a a \\<^sub>r b = (-\<^sub>a a) \\<^sub>r b \ -\<^sub>a (a \\<^sub>r b) = a \\<^sub>r (-\<^sub>a b)") apply simp done lemma (in Ring) ring_times_minusl:"a \ carrier R \ -\<^sub>a a = (-\<^sub>a 1\<^sub>r) \\<^sub>r a" apply (cut_tac ring_one) apply (frule ring_inv1_1[of "1\<^sub>r" "a"], assumption+) apply (simp add:ring_l_one) done lemma (in Ring) ring_times_minusr:"a \ carrier R \ -\<^sub>a a = a \\<^sub>r (-\<^sub>a 1\<^sub>r)" apply (cut_tac ring_one) apply (frule ring_inv1_2[of "a" "1\<^sub>r"], assumption+) apply (simp add:ring_r_one) done lemma (in Ring) ring_inv1_3:"\a \ carrier R; b \ carrier R\ \ a \\<^sub>r b = (-\<^sub>a a) \\<^sub>r (-\<^sub>a b)" apply (cut_tac ring_is_ag) apply (subst aGroup.ag_inv_inv[THEN sym], assumption+) apply (frule aGroup.ag_mOp_closed[of "R" "a"], assumption+) apply (subst ring_inv1_1[THEN sym, of "-\<^sub>a a" "b"], assumption+) apply (subst ring_inv1_2[of "-\<^sub>a a" "b"], assumption+, simp) done lemma (in Ring) ring_distrib4:"\a \ carrier R; b \ carrier R; x \ carrier R; y \ carrier R \ \ a \\<^sub>r b \ (-\<^sub>a x \\<^sub>r y) = a \\<^sub>r (b \ (-\<^sub>a y)) \ (a \ (-\<^sub>a x)) \\<^sub>r y" apply (cut_tac ring_is_ag) apply (subst ring_distrib1, assumption+) apply (rule aGroup.ag_mOp_closed, assumption+) apply (subst ring_distrib2, assumption+) apply (rule aGroup.ag_mOp_closed, assumption+) apply (subst aGroup.pOp_assocTr43, assumption+) apply (rule ring_tOp_closed, assumption+)+ apply (rule aGroup.ag_mOp_closed, assumption+) apply (rule ring_tOp_closed, assumption+) apply (rule ring_tOp_closed) apply (simp add:aGroup.ag_mOp_closed)+ apply (subst ring_distrib1 [THEN sym, of "a" _], assumption+) apply (rule aGroup.ag_mOp_closed, assumption+) apply (simp add:aGroup.ag_l_inv1) apply (simp add:ring_times_x_0) apply (subst aGroup.ag_r_zero, assumption+) apply (simp add:ring_tOp_closed) apply (simp add: ring_inv1_1) done lemma (in Ring) rMulLC: "\x \ carrier R; y \ carrier R; z \ carrier R\ \ x \\<^sub>r (y \\<^sub>r z) = y \\<^sub>r (x \\<^sub>r z)" apply (subst ring_tOp_assoc [THEN sym], assumption+) apply (subst ring_tOp_commute [of "x" "y"], assumption+) apply (subst ring_tOp_assoc, assumption+) apply simp done lemma (in Ring) Zero_ring:"1\<^sub>r = \ \ zeroring R" apply (simp add:zeroring_def) apply (rule conjI) apply (rule Ring_axioms) apply (rule equalityI) apply (rule subsetI) apply (frule_tac x = x in ring_r_one, simp add:ring_times_x_0) apply (simp add:ring_zero) done lemma (in Ring) Zero_ring1:"\ (zeroring R) \ 1\<^sub>r \ \" apply (rule contrapos_pp, simp+, cut_tac Zero_ring, simp+) done lemma (in Ring) Sr_one:"sr R S \ 1\<^sub>r \ S" apply (simp add:sr_def) done lemma (in Ring) Sr_zero:"sr R S \ \ \ S" apply (cut_tac ring_is_ag, frule Sr_one[of "S"]) apply (simp add:sr_def) apply (erule conjE)+ apply (frule_tac x = "1\<^sub>r" in bspec, assumption, thin_tac "\x\S. \y\S. x \ -\<^sub>a y \ S \ x \\<^sub>r y \ S", frule_tac x = "1\<^sub>r" in bspec, assumption, thin_tac "\y\S. 1\<^sub>r \ -\<^sub>a y \ S \ 1\<^sub>r \\<^sub>r y \ S", erule conjE) apply (cut_tac ring_one, simp add:aGroup.ag_r_inv1[of "R" "1\<^sub>r"]) done lemma (in Ring) Sr_mOp_closed:"\sr R S; x \ S\ \ -\<^sub>a x \ S" apply (frule Sr_zero[of "S"]) apply (simp add:sr_def, (erule conjE)+) apply (cut_tac ring_is_ag) apply (frule_tac x = "\" in bspec, assumption, thin_tac "\x\S. \y\S. x \ -\<^sub>a y \ S \ x \\<^sub>r y \ S", frule_tac x = x in bspec, assumption, thin_tac "\y\S. \ \ -\<^sub>a y \ S \ \ \\<^sub>r y \ S", erule conjE) apply (frule subsetD[of "S" "carrier R" "\"], assumption+, frule subsetD[of "S" "carrier R" "x"], assumption+) apply (frule aGroup.ag_mOp_closed [of "R" "x"], assumption) apply (simp add:aGroup.ag_l_zero) done lemma (in Ring) Sr_pOp_closed:"\sr R S; x \ S; y \ S\ \ x \ y \ S" apply (frule Sr_mOp_closed[of "S" "y"], assumption+) apply (unfold sr_def, (erule conjE)+) apply (frule_tac x = x in bspec, assumption, thin_tac "\x\S. \y\S. x \ -\<^sub>a y \ S \ x \\<^sub>r y \ S", frule_tac x = "-\<^sub>a y" in bspec, assumption, thin_tac "\y\S. x \ -\<^sub>a y \ S \ x \\<^sub>r y \ S", erule conjE) apply (cut_tac ring_is_ag ) apply (frule subsetD[of "S" "carrier R" "y"], assumption+) apply (simp add:aGroup.ag_inv_inv) done lemma (in Ring) Sr_tOp_closed:"\sr R S; x \ S; y \ S\ \ x \\<^sub>r y \ S" by (simp add:sr_def) lemma (in Ring) Sr_ring:"sr R S \ Ring (Sr R S)" apply (simp add:Ring_def [of "Sr R S"], cut_tac ring_is_ag) apply (rule conjI) apply (simp add:Sr_def Sr_pOp_closed) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Sr_def, frule_tac x = a and y = b in Sr_pOp_closed, assumption+, frule_tac x = b and y = c in Sr_pOp_closed, assumption+, simp add:Sr_def sr_def, (erule conjE)+) apply (frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, frule_tac c = b in subsetD[of "S" "carrier R"], assumption+, frule_tac c = c in subsetD[of "S" "carrier R"], assumption+) apply (simp add:aGroup.ag_pOp_assoc) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Sr_def sr_def, (erule conjE)+, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, frule_tac c = b in subsetD[of "S" "carrier R"], assumption+) apply (simp add:aGroup.ag_pOp_commute) apply (rule conjI) apply ((subst Sr_def)+, simp) apply (simp add:Sr_mOp_closed) apply (rule conjI) apply (rule allI) apply ((subst Sr_def)+, simp add:Sr_mOp_closed, rule impI) apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, simp add:aGroup.ag_l_inv1) apply (rule conjI) apply (simp add:Sr_def Sr_zero) apply (rule conjI) apply (rule allI, simp add:Sr_def Sr_zero) apply (rule impI) apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, simp add:aGroup.ag_l_zero) apply (rule conjI) apply (simp add:Sr_def Sr_tOp_closed) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Sr_def Sr_tOp_closed) apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, frule_tac c = b in subsetD[of "S" "carrier R"], assumption+, frule_tac c = c in subsetD[of "S" "carrier R"], assumption+) apply (simp add:ring_tOp_assoc) apply (rule conjI) apply ((rule allI, rule impI)+, simp add:Sr_def) apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, frule_tac c = b in subsetD[of "S" "carrier R"], assumption+, simp add:ring_tOp_commute) apply (rule conjI) apply (simp add:Sr_def Sr_one) apply (rule conjI) apply (simp add:Sr_def Sr_pOp_closed Sr_tOp_closed) apply (rule allI, rule impI)+ apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+, frule_tac c = b in subsetD[of "S" "carrier R"], assumption+, frule_tac c = c in subsetD[of "S" "carrier R"], assumption+) apply (simp add:ring_distrib1) apply (simp add:Sr_def Sr_one) apply (rule allI, rule impI) apply (unfold sr_def, frule conjunct1, fold sr_def, frule_tac c = a in subsetD[of "S" "carrier R"], assumption+) apply (simp add:ring_l_one) done section "Calculation of elements" (** The author of this part is L. Chen, revised by H. Murao and Y. Santo **) subsection "nscale" lemma (in Ring) ring_tOp_rel:"\x\carrier R; xa\carrier R; y\carrier R; ya \ carrier R \ \ (x \\<^sub>r xa) \\<^sub>r (y \\<^sub>r ya) = (x \\<^sub>r y) \\<^sub>r (xa \\<^sub>r ya)" apply (frule ring_tOp_closed[of "y" "ya"], assumption+, simp add:ring_tOp_assoc[of "x" "xa"]) apply (simp add:ring_tOp_assoc[THEN sym, of "xa" "y" "ya"], simp add:ring_tOp_commute[of "xa" "y"], simp add:ring_tOp_assoc[of "y" "xa" "ya"]) apply (frule ring_tOp_closed[of "xa" "ya"], assumption+, simp add:ring_tOp_assoc[THEN sym, of "x" "y"]) done lemma (in Ring) nsClose: "\ n. \ x \ carrier R \ \ nscal R x n \ carrier R" apply (induct_tac n) apply (simp add:ring_zero) apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_closed) done lemma (in Ring) nsZero: "nscal R \ n = \" apply (cut_tac ring_is_ag) apply (induct_tac n) apply simp apply simp apply (cut_tac ring_zero, simp add:aGroup.ag_l_zero) done lemma (in Ring) nsZeroI: "\ n. x = \ \ nscal R x n = \" by (simp only:nsZero) lemma (in Ring) nsEqElm: "\ x \ carrier R; y \ carrier R; x = y \ \ (nscal R x n) = (nscal R y n)" by simp lemma (in Ring) nsDistr: "x \ carrier R \ (nscal R x n) \ (nscal R x m) = nscal R x (n + m)" apply (cut_tac ring_is_ag) apply (induct_tac m) apply simp apply (frule nsClose[of "x" "n"]) apply ( simp add:aGroup.ag_r_zero) apply simp apply (frule_tac x = x and n = n in nsClose, frule_tac x = x and n = na in nsClose) apply (subst aGroup.ag_pOp_assoc[THEN sym], assumption+, simp) done lemma (in Ring) nsDistrL: "\x \ carrier R; y \ carrier R \ \ (nscal R x n) \ (nscal R y n) = nscal R (x \ y) n" apply (cut_tac ring_is_ag) apply (induct_tac n) apply simp apply (cut_tac ring_zero, simp add:aGroup.ag_l_zero) apply simp apply (frule_tac x = x and n = n in nsClose, frule_tac x = y and n = n in nsClose) apply (subst aGroup.pOp_assocTr43[of R _ x _ y], assumption+) apply (frule_tac x = x and y = "n \\<^bsub>R\<^esub> y" in aGroup.ag_pOp_commute[of "R"], assumption+) apply simp apply (subst aGroup.pOp_assocTr43[THEN sym, of R _ _ x y], assumption+) apply simp done lemma (in Ring) nsMulDistrL:"\ x \ carrier R; y \ carrier R \ \ x \\<^sub>r (nscal R y n) = nscal R (x \\<^sub>r y) n" apply (induct_tac n) apply simp apply (simp add:ring_times_x_0) apply simp apply (subst ring_distrib1, assumption+) apply (rule nsClose, assumption+) apply simp done lemma (in Ring) nsMulDistrR:"\ x \ carrier R; y \ carrier R\ \ (nscal R y n) \\<^sub>r x = nscal R (y \\<^sub>r x) n" apply (frule_tac x = y and n = n in nsClose, simp add:ring_tOp_commute[of "n \\<^bsub>R\<^esub> y" "x"], simp add:nsMulDistrL, simp add:ring_tOp_commute[of "y" "x"]) done subsection "npow" lemma (in Ring) npClose:"x \ carrier R \ npow R x n \ carrier R" apply (induct_tac n) apply simp apply (simp add:ring_one) apply simp apply (rule ring_tOp_closed, assumption+) done lemma (in Ring) npMulDistr:"\ n m. x \ carrier R \ (npow R x n) \\<^sub>r (npow R x m) = npow R x (n + m)" apply (induct_tac m) apply simp apply (rule ring_r_one, simp add:npClose) apply simp apply (frule_tac x = x and n = n in npClose, frule_tac x = x and n = na in npClose) apply (simp add:ring_tOp_assoc[THEN sym]) done lemma (in Ring) npMulExp:"\n m. x \ carrier R \ npow R (npow R x n) m = npow R x (n * m)" apply (induct_tac m) apply simp apply simp apply (simp add:npMulDistr) apply (simp add:add.commute) done lemma (in Ring) npGTPowZero_sub: " \ n. \ x \ carrier R; npow R x m = \ \ \(m \ n) \ (npow R x n = \ )" apply (rule impI) apply (subgoal_tac "npow R x n = (npow R x (n-m)) \\<^sub>r (npow R x m)") apply simp apply (rule ring_times_x_0) apply (simp add:npClose) apply (thin_tac "x^\<^bsup>R m\<^esup> = \") apply (subst npMulDistr, assumption) apply simp done lemma (in Ring) npGTPowZero: "\ n. \ x \ carrier R; npow R x m = \; m \ n \ \ npow R x n = \" apply (cut_tac x = x and m = m and n = n in npGTPowZero_sub, assumption+) apply simp done lemma (in Ring) npOne: " npow R (1\<^sub>r) n = 1\<^sub>r" apply (induct_tac n) apply simp apply simp apply (rule ring_r_one, simp add:ring_one) done lemma (in Ring) npZero_sub: "0 < n \ npow R \ n = \" apply (induct_tac "n") apply simp apply simp apply (cut_tac ring_zero, frule_tac n = n in npClose[of "\"]) apply (simp add:ring_times_x_0) done lemma (in Ring) npZero: "0 < n \ npow R \ n = \" apply (simp add:npZero_sub) done lemma (in Ring) npMulElmL: "\ n. \ x \ carrier R; 0 \ n\ \ x \\<^sub>r (npow R x n) = npow R x (Suc n)" apply (simp only:npow_suc, frule_tac n = n and x = x in npClose, simp add:ring_tOp_commute) done lemma (in Ring) npMulEleL: "\ n. x \ carrier R \ (npow R x n) \\<^sub>r x = npow R x (Suc n)" by (simp add:npMulElmL[THEN sym]) lemma (in Ring) npMulElmR: "\ n. x \ carrier R \ (npow R x n) \\<^sub>r x = npow R x (Suc n)" apply ( frule_tac n = n in npClose[of "x"]) apply (simp only:ring_tOp_commute, subst npMulElmL, assumption, simp, simp) done lemma (in Ring) np_1:"a \ carrier R \ npow R a (Suc 0) = a" (* Y. Santo*) apply simp apply (simp add:ring_l_one) done subsection "nsum and fSum" lemma (in aGroup) nsum_memTr: "(\j \ n. f j \ carrier A) \ nsum A f n \ carrier A" apply (induct_tac "n") apply simp apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (frule_tac a = "Suc n" in forall_spec, simp, thin_tac "\j\Suc n. f j \ carrier A") apply (rule ag_pOp_closed, assumption+) done lemma (in aGroup) nsum_mem:"\j \ n. f j \ carrier A \ nsum A f n \ carrier A" apply (simp add:nsum_memTr) done lemma (in aGroup) nsum_eqTr:"(\j \ n. f j \ carrier A \ g j \ carrier A \ f j = g j) \ nsum A f n = nsum A g n" apply (induct_tac n) apply simp apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) done lemma (in aGroup) nsum_eq:"\\j \ n. f j \ carrier A; \j \ n. g j \ carrier A; \j \ n. f j = g j\ \ nsum A f n = nsum A g n" by (simp add:nsum_eqTr) lemma (in aGroup) nsum_cmp_assoc:"\\j \ n. f j \ carrier A; g \ {j. j \ n} \ {j. j \ n}; h \ {j. j \ n} \ {j. j \ n}\ \ nsum A (cmp (cmp f h) g) n = nsum A (cmp f (cmp h g)) n" apply (rule nsum_eq) apply (rule allI, rule impI, simp add:cmp_def) apply (frule_tac x = j in funcset_mem[of g "{j. j \ n}" "{j. j \ n}"], simp, frule_tac x = "g j" in funcset_mem[of h "{j. j \ n}" "{j. j \ n}"], assumption, simp) apply (rule allI, rule impI, simp add:cmp_def, frule_tac x = j in funcset_mem[of g "{j. j \ n}" "{j. j \ n}"], simp, frule_tac x = "g j" in funcset_mem[of h "{j. j \ n}" "{j. j \ n}"], assumption, simp) apply (rule allI, simp add:cmp_def) done lemma (in aGroup) fSum_Suc:"\j \ nset n (n + Suc m). f j \ carrier A \ fSum A f n (n + Suc m) = fSum A f n (n + m) \ f (n + Suc m)" by (simp add:fSum_def, simp add:cmp_def slide_def) lemma (in aGroup) fSum_eqTr:"(\j \ nset n (n + m). f j \ carrier A \ g j \ carrier A \ f j = g j) \ fSum A f n (n + m) = fSum A g n (n + m)" apply (induct_tac m) apply (simp add:fSum_def, simp add:cmp_def slide_def, simp add:nset_def) apply (rule impI) apply (subst fSum_Suc, rule ballI, simp, simp) apply (cut_tac n = n and m = na and f = g in fSum_Suc, rule ballI, simp, simp, thin_tac "\\<^sub>f A g n (Suc (n + na)) = \\<^sub>f A g n (n + na) \ g (Suc (n + na))") apply (cut_tac n = n and m = na in nsetnm_sub_mem, simp, thin_tac "\j. j \ nset n (n + na) \ j \ nset n (Suc (n + na))") apply (frule_tac x = "Suc (n + na)" in bspec, simp add:nset_def, simp) done lemma (in aGroup) fSum_eq:"\ \j \ nset n (n + m). f j \ carrier A; \j \ nset n (n + m). g j \ carrier A; (\j\ nset n (n + m). f j = g j)\ \ fSum A f n (n + m) = fSum A g n (n + m)" by (simp add:fSum_eqTr) lemma (in aGroup) fSum_eq1:"\n \ m; \j\nset n m. f j \ carrier A; \j\nset n m. g j \ carrier A; \j\nset n m. f j = g j\ \ fSum A f n m = fSum A g n m" apply (cut_tac fSum_eq[of n "m - n" f g]) apply simp+ done lemma (in aGroup) fSum_zeroTr:"(\j \ nset n (n + m). f j = \) \ fSum A f n (n + m) = \" apply (induct_tac m) apply (simp add:fSum_def cmp_def slide_def nset_def) apply (rule impI) apply (subst fSum_Suc) apply (rule ballI, simp add:ag_inc_zero) apply (frule_tac x = "n + Suc na" in bspec, simp add:nset_def, simp) apply (simp add:nset_def) apply (cut_tac ag_inc_zero, simp add:ag_l_zero) done lemma (in aGroup) fSum_zero:"\j \ nset n (n + m). f j = \ \ fSum A f n (n + m) = \" by (simp add:fSum_zeroTr) lemma (in aGroup) fSum_zero1:"\n < m; \j \ nset (Suc n) m. f j = \\ \ fSum A f (Suc n) m = \" apply (cut_tac fSum_zero[of "Suc n" "m - Suc n" f]) apply simp+ done lemma (in Ring) nsumMulEleL: "\ n. \ \ i. f i \ carrier R; x \ carrier R \ \ x \\<^sub>r (nsum R f n) = nsum R (\ i. x \\<^sub>r (f i)) n" apply (cut_tac ring_is_ag) apply (induct_tac "n") apply simp apply simp apply (subst ring_distrib1, assumption) apply (rule aGroup.nsum_mem, assumption) apply (rule allI, simp+) done lemma (in Ring) nsumMulElmL: "\ n. \ \ i. f i \ carrier R; x \ carrier R \ \ x \\<^sub>r (nsum R f n) = nsum R (\ i. x \\<^sub>r (f i)) n" apply (cut_tac ring_is_ag) apply (induct_tac "n") apply simp apply simp apply (subst ring_distrib1, assumption+) apply (simp add:aGroup.nsum_mem)+ done lemma (in aGroup) nsumTailTr: "(\j\(Suc n). f j \ carrier A) \ nsum A f (Suc n) = (nsum A (\ i. (f (Suc i))) n) \ (f 0)" apply (induct_tac "n") apply simp apply (rule impI, rule ag_pOp_commute) apply (cut_tac Nset_inc_0[of "Suc 0"], simp add:Pi_def, cut_tac n_in_Nsetn[of "Suc 0"], simp add:Pi_def) apply (rule impI) apply (cut_tac n = "Suc n" in Nsetn_sub_mem1, simp) apply (frule_tac a = 0 in forall_spec, simp, frule_tac a = "Suc (Suc n)" in forall_spec, simp) apply (cut_tac n = n in nsum_mem[of _ "\i. f (Suc i)"], rule allI, rule impI, frule_tac a = "Suc j" in forall_spec, simp, simp, thin_tac "\j\Suc (Suc n). f j \ carrier A") apply (subst ag_pOp_assoc, assumption+) apply (simp add:ag_pOp_commute[of "f 0"]) apply (subst ag_pOp_assoc[THEN sym], assumption+) apply simp done lemma (in aGroup) nsumTail: "\j \ (Suc n). f j \ carrier A \ nsum A f (Suc n) = (nsum A (\ i. (f (Suc i))) n) \ (f 0)" by (cut_tac nsumTailTr[of n f], simp) lemma (in aGroup) nsumElmTail: "\i. f i \ carrier A \ nsum A f (Suc n) = (nsum A (\ i. (f (Suc i))) n) \ (f 0)" apply (cut_tac n = n and f = f in nsumTail, rule allI, simp, simp) done lemma (in aGroup) nsum_addTr: "(\j \ n. f j \ carrier A \ g j \ carrier A) \ nsum A (\ i. (f i) \ (g i)) n = (nsum A f n) \ (nsum A g n)" apply (induct_tac "n") apply simp apply (simp, rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (thin_tac "\\<^sub>e A (\i. f i \ g i) n = \\<^sub>e A f n \ \\<^sub>e A g n") apply (rule aGroup.ag_add4_rel, rule aGroup_axioms) apply (rule aGroup.nsum_mem, rule aGroup_axioms, rule allI, simp) apply (rule aGroup.nsum_mem, rule aGroup_axioms, rule allI, simp) apply simp+ done lemma (in aGroup) nsum_add: "\ \j \ n. f j \ carrier A; \j \ n. g j \ carrier A\ \ nsum A (\ i. (f i) \ (g i)) n = (nsum A f n) \ (nsum A g n)" by (cut_tac nsum_addTr[of n f g], simp) lemma (in aGroup) nsumElmAdd: "\ \ i. f i \ carrier A; \ i. g i \ carrier A\ \ nsum A (\ i. (f i) \ (g i)) n = (nsum A f n) \ (nsum A g n)" apply (cut_tac nsum_add[of n f g]) apply simp apply (rule allI, simp)+ done lemma (in aGroup) nsum_add_nmTr: "(\j \ n. f j \ carrier A) \ (\j \ m. g j \ carrier A) \ nsum A (jointfun n f m g) (Suc (n + m)) = (nsum A f n) \ (nsum A g m)" apply (induct_tac m) apply (simp add:jointfun_def sliden_def) apply (rule impI) apply (rule ag_pOp_add_r) apply (rule nsum_mem, rule allI, erule conjE, rule impI, simp) apply (erule conjE, simp add:nsum_mem, simp) apply (rule nsum_eq[of n], simp+) apply (simp add:jointfun_def) apply (rule impI, simp) apply (erule conjE, simp add:sliden_def) apply (thin_tac "\\<^sub>e A (\i. if i \ n then f i else g (sliden (Suc n) i)) (n + na) \ g na = \\<^sub>e A f n \ \\<^sub>e A g na") apply (subst ag_pOp_assoc) apply (simp add:nsum_mem) apply (simp add:nsum_mem, simp) apply simp done lemma (in aGroup) nsum_add_nm: "\\j \ n. f j \ carrier A; \j \ m. g j \ carrier A\ \ nsum A (jointfun n f m g) (Suc (n + m)) = (nsum A f n) \ (nsum A g m)" apply (cut_tac nsum_add_nmTr[of n f m g]) apply simp done lemma (in Ring) npeSum2_sub_muly: "\ x \ carrier R; y \ carrier R \ \ y \\<^sub>r(nsum R (\i. nscal R ((npow R x (n-i)) \\<^sub>r (npow R y i)) (n choose i)) n) = nsum R (\i. nscal R ((npow R x (n-i)) \\<^sub>r (npow R y (i+1))) (n choose i)) n" apply (cut_tac ring_is_ag) apply (subst nsumMulElmL) apply (rule allI) apply (simp only:nsClose add:ring_tOp_closed add:npClose) apply assumption apply (simp only:nsMulDistrL add:nsClose add:ring_tOp_closed add:npClose) apply (simp only: rMulLC [of "y"] add:npClose) apply (simp del:npow_suc add:ring_tOp_commute[of y]) apply (rule aGroup.nsum_eq, assumption) apply (rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, rule ring_tOp_closed, assumption, simp add:npClose) apply (rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, rule npClose, assumption) apply (rule allI, rule impI) apply (frule_tac n = j in npClose[of y]) apply (simp add:ring_tOp_commute[of y]) done (********)(********)(********)(********) lemma binomial_n0: "(Suc n choose 0) = (n choose 0)" by simp lemma binomial_ngt_diff: "(n choose Suc n) = (Suc n choose Suc n) - (n choose n)" by (subst binomial_Suc_Suc, arith) lemma binomial_ngt_0: "(n choose Suc n) = 0" apply (subst binomial_ngt_diff, (subst binomial_n_n)+) apply simp done lemma diffLessSuc: "m \ n \ Suc (n-m) = Suc n - m" by arith lemma (in Ring) npow_suc_i: "\ x \ carrier R; i \ n \ \ npow R x (Suc n - i) = x \\<^sub>r (npow R x (n-i))" apply (subst diffLessSuc [THEN sym, of "i" "n"], assumption) apply (frule_tac n = "n - i" in npClose, simp add:ring_tOp_commute[of x]) done (** lemma (in Ring) nsumEqFunc_sub: "\ \ i. f i \ carrier R; \ i. g i \ carrier R \ \ ( \ i. i \ n \ f i = g i) \ (nsum0 R f n = nsum0 R g n)"; apply (induct_tac "n") apply simp+ done lemma (in Ring) nsumEqFunc: "\ \ i. f i \ carrier R; \ i. g i \ carrier R; \ i. i \ n \ f i = g i \ \ nsum0 R f n = nsum0 R g n" apply (cut_tac nsumEqFunc_sub [of "f" "g" "n"]) apply simp+ done nsumEqFunc \ nsum_eq **) (********)(********) lemma (in Ring) npeSum2_sub_mulx: "\ x \ carrier R; y \ carrier R \ \ x \\<^sub>r (nsum R (\ i. nscal R ((npow R x (n-i)) \\<^sub>r (npow R y i)) (n choose i)) n) = (nsum R (\i. nscal R ((npow R x (Suc n - Suc i)) \\<^sub>r (npow R y (Suc i))) (n choose Suc i)) n) \ (nscal R ((npow R x (Suc n - 0)) \\<^sub>r (npow R y 0)) (Suc n choose 0))" apply (cut_tac ring_is_ag) apply (simp only: binomial_n0) apply (subst aGroup.nsumElmTail [THEN sym, of R "\ i. nscal R ((npow R x (Suc n - i)) \\<^sub>r (npow R y i)) (n choose i)"], assumption+) apply (rule allI) apply (simp only:nsClose add:ring_tOp_closed add:npClose) apply (simp only:nsum_suc) apply (subst binomial_ngt_0) apply (simp only:nscal_0) apply (subst aGroup.ag_r_zero, assumption) apply (simp add:aGroup.nsum_mem nsClose ring_tOp_closed npClose) apply (subst nsumMulElmL [of _ "x"]) apply (rule allI, rule nsClose, rule ring_tOp_closed, simp add:npClose, simp add:npClose, assumption) apply (simp add: nsMulDistrL [of "x"] ring_tOp_closed npClose) apply (simp add:ring_tOp_assoc [THEN sym, of "x"] npClose) apply (rule aGroup.nsum_eq, assumption) apply (rule allI, rule impI, rule nsClose, (rule ring_tOp_closed)+, assumption, simp add:npClose, simp add:npClose) apply (rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, simp add:npClose) apply (rule allI, rule impI) apply (frule_tac n = "n - j" in npClose[of x], simp add:ring_tOp_commute[of x], subst npow_suc[THEN sym]) apply (simp add:Suc_diff_le) done lemma (in Ring) npeSum2_sub_mulx2: "\ x \ carrier R; y \ carrier R \ \ x \\<^sub>r (nsum R (\ i. nscal R ((npow R x (n-i)) \\<^sub>r (npow R y i)) (n choose i)) n) = (nsum R (\i. nscal R ((npow R x (n - i)) \\<^sub>r ((npow R y i) \\<^sub>r y )) (n choose Suc i)) n) \ (\ \ ((x \\<^sub>r (npow R x n)) \\<^sub>r (1\<^sub>r)))" apply (subst npeSum2_sub_mulx, assumption+, simp) apply (frule npClose[of x n]) apply (subst ring_tOp_commute[of x], assumption+) apply (cut_tac ring_is_ag) apply (cut_tac aGroup.nsum_eq[of R n "\i. (n choose Suc i) \\<^bsub>R\<^esub> (x^\<^bsup>R (n - i)\<^esup> \\<^sub>r y^\<^bsup>R (Suc i)\<^esup>)" "\i. (n choose Suc i) \\<^bsub>R\<^esub> (x^\<^bsup>R (n - i)\<^esup> \\<^sub>r (y^\<^bsup>R i\<^esup> \\<^sub>r y))"]) apply (simp del:npow_suc)+ apply (rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, simp only:npClose) apply (rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, rule ring_tOp_closed, simp add:npClose, assumption) apply (rule allI, rule impI) apply (frule_tac n = j in npClose[of y]) apply simp done lemma (in Ring) npeSum2: "\ n. \ x \ carrier R; y \ carrier R \ \ npow R (x \ y) n = nsum R (\ i. nscal R ((npow R x (n-i)) \\<^sub>r (npow R y i)) ( n choose i) ) n" apply (cut_tac ring_is_ag) apply (induct_tac "n") (*1*) apply simp apply (cut_tac ring_one, simp add:ring_r_one, simp add:aGroup.ag_l_zero) (*1:done*) apply (subst aGroup.nsumElmTail, assumption+) apply (rule allI) apply (simp add:nsClose ring_tOp_closed npClose) (** thm binomial_Suc_Suc **) apply (simp only:binomial_Suc_Suc) apply (simp only: nsDistr [THEN sym] add:npClose ring_tOp_closed) apply (subst aGroup.nsumElmAdd, assumption+) apply (rule allI, simp add:nsClose ring_tOp_closed npClose) apply (rule allI, simp add:nsClose add:ring_tOp_closed npClose) apply (subst aGroup.ag_pOp_assoc, assumption) apply (rule aGroup.nsum_mem, assumption, rule allI, rule impI, simp add:nsClose ring_tOp_closed npClose) apply (rule aGroup.nsum_mem, assumption, rule allI, rule impI, simp add:nsClose ring_tOp_closed npClose) apply (simp add:nsClose ring_tOp_closed npClose) apply (rule aGroup.ag_pOp_closed, assumption) apply (simp add:aGroup.ag_inc_zero) apply (rule ring_tOp_closed)+ apply (simp add:npClose, assumption, simp add:ring_one) apply (subst npMulElmL [THEN sym, of "x \ y"], simp add:aGroup.ag_pOp_closed, simp) apply simp apply (subst ring_distrib2 [of _ "x" "y"]) apply (rule aGroup.nsum_mem,assumption, rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, simp add:npClose, assumption+) apply (rule aGroup.gEQAddcross [THEN sym], assumption+, rule aGroup.nsum_mem, assumption, rule allI, rule impI, rule nsClose, (rule ring_tOp_closed)+, simp add:npClose, rule ring_tOp_closed, simp add:npClose, assumption) apply (rule aGroup.ag_pOp_closed, assumption) apply (rule aGroup.nsum_mem, assumption, rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, rule ring_tOp_closed, simp add:npClose, assumption) apply (rule aGroup.ag_pOp_closed, assumption, simp add:ring_zero) apply ((rule ring_tOp_closed)+, simp add:npClose,assumption, simp add:ring_one) apply (rule ring_tOp_closed, assumption, rule aGroup.nsum_mem, assumption, rule allI, rule impI, rule nsClose, rule ring_tOp_closed, (simp add:npClose)+) apply (rule ring_tOp_closed, assumption+, rule aGroup.nsum_mem, assumption, rule allI, rule impI, rule nsClose, rule ring_tOp_closed, simp add:npClose, simp add:npClose) apply (subst npeSum2_sub_muly [of "x" "y"], assumption+, simp) (* final part *) apply (subst npeSum2_sub_mulx2 [of x y], assumption+) apply (frule_tac n = na in npClose[of x], simp add:ring_tOp_commute[of _ x]) done lemma (in aGroup) nsum_zeroTr: "\ n. (\ i. i \ n \ f i = \) \ (nsum A f n = \)" apply (induct_tac "n") apply simp apply (rule impI) apply (cut_tac n = na in Nsetn_sub_mem1, simp) apply (subst aGroup.ag_l_zero, rule aGroup_axioms) apply (simp add:ag_inc_zero) apply simp done lemma (in Ring) npAdd: "\ x \ carrier R; y \ carrier R; npow R x m = \; npow R y n = \ \ \ npow R (x \ y) (m + n) = \" apply (subst npeSum2, assumption+) apply (rule aGroup.nsum_zeroTr [THEN mp]) apply (simp add:ring_is_ag) apply (rule allI, rule impI) apply (rule nsZeroI) apply (rule rMulZeroDiv, simp add:npClose, simp add:npClose) apply (case_tac "i \ n") apply (rule disjI1) apply (rule npGTPowZero [of "x" "m"], assumption+) apply arith apply (rule disjI2) apply (rule npGTPowZero [of "y" "n"], assumption+) apply (arith) done lemma (in Ring) npInverse: "\n. x \ carrier R \ npow R (-\<^sub>a x) n = npow R x n \ npow R (-\<^sub>a x) n = -\<^sub>a (npow R x n)" apply (induct_tac n) (* n=0 *) apply simp apply (erule disjE) apply simp apply (subst ring_inv1_2, simp add:npClose, assumption, simp) apply (cut_tac ring_is_ag) apply simp apply (subst ring_inv1_2[THEN sym, of _ x]) apply (rule aGroup.ag_mOp_closed, assumption+, simp add:npClose, assumption) apply (thin_tac "(-\<^sub>a x)^\<^bsup>R na\<^esup> = -\<^sub>a (x^\<^bsup>R na\<^esup>)", frule_tac n = na in npClose[of x], frule_tac x = "x^\<^bsup>R na\<^esup>" in aGroup.ag_mOp_closed[of R], simp add:npClose) apply (simp add: ring_inv1_1[of _ x]) apply (simp add:aGroup.ag_inv_inv[of R]) done lemma (in Ring) npMul: "\ n. \ x \ carrier R; y \ carrier R \ \ npow R (x \\<^sub>r y) n = (npow R x n) \\<^sub>r (npow R y n)" apply (induct_tac "n") (* n=0 *) apply simp apply (rule ring_r_one [THEN sym]) apply (simp add:ring_one) (* n>0 *) apply (simp only:npow_suc) apply (rule ring_tOp_rel[THEN sym]) apply (rule npClose, assumption+)+ done section "Ring homomorphisms" definition rHom :: "[('a, 'm) Ring_scheme, ('b, 'm1) Ring_scheme] \ ('a \ 'b) set" where "rHom A R = {f. f \ aHom A R \ (\x\carrier A. \y\carrier A. f ( x \\<^sub>r\<^bsub>A\<^esub> y) = (f x) \\<^sub>r\<^bsub>R\<^esub> (f y)) \ f (1\<^sub>r\<^bsub>A\<^esub>) = (1\<^sub>r\<^bsub>R\<^esub>)}" definition rInvim :: "[('a, 'm) Ring_scheme, ('b, 'm1) Ring_scheme, 'a \ 'b, 'b set] \ 'a set" where "rInvim A R f K = {a. a \ carrier A \ f a \ K}" definition rimg :: "[('a, 'm) Ring_scheme, ('b, 'm1) Ring_scheme, 'a \ 'b] \ 'b Ring" where "rimg A R f = \carrier= f `(carrier A), pop = pop R, mop = mop R, zero = zero R, tp = tp R, un = un R \" definition ridmap :: "('a, 'm) Ring_scheme \ ('a \ 'a)" where "ridmap R = (\x\carrier R. x)" definition r_isom :: "[('a, 'm) Ring_scheme, ('b, 'm1) Ring_scheme] \ bool" (infixr "\\<^sub>r" 100) where "r_isom R R' \ (\f\rHom R R'. bijec\<^bsub>R,R'\<^esub> f)" definition Subring :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme] \ bool" where "Subring R S == Ring S \ (carrier S \ carrier R) \ (ridmap S) \ rHom S R" lemma ridmap_surjec:"Ring A \ surjec\<^bsub>A,A\<^esub> (ridmap A)" by(simp add:surjec_def aHom_def ridmap_def Ring.ring_is_ag aGroup.ag_pOp_closed surj_to_def) lemma rHom_aHom:"f \ rHom A R \ f \ aHom A R" by (simp add:rHom_def) lemma rimg_carrier:"f \ rHom A R \ carrier (rimg A R f) = f ` (carrier A)" by (simp add:rimg_def) lemma rHom_mem:"\ f \ rHom A R; a \ carrier A \ \ f a \ carrier R" apply (simp add:rHom_def, frule conjunct1) apply (thin_tac "f \ aHom A R \ (\x\carrier A. \y\carrier A. f (x \\<^sub>r\<^bsub>A\<^esub> y) = f x \\<^sub>r\<^bsub>R\<^esub> f y) \ f 1\<^sub>r\<^bsub>A\<^esub> = 1\<^sub>r\<^bsub>R\<^esub>") apply (simp add:aHom_def, frule conjunct1) apply (thin_tac "f \ carrier A \ carrier R \ f \ extensional (carrier A) \ (\a\carrier A. \b\carrier A. f (a \\<^bsub>A\<^esub> b) = f a \\<^bsub>R\<^esub> f b)") apply (simp add:funcset_mem) done lemma rHom_func:"f \ rHom A R \ f \ carrier A \ carrier R" by (simp add:rHom_def aHom_def) lemma ringhom1:"\ Ring A; Ring R; x \ carrier A; y \ carrier A; f \ rHom A R \ \ f (x \\<^bsub>A\<^esub> y) = (f x) \\<^bsub>R\<^esub> (f y)" apply (simp add:rHom_def) apply (erule conjE) apply (frule Ring.ring_is_ag [of "A"]) apply (frule Ring.ring_is_ag [of "R"]) apply (rule aHom_add, assumption+) done lemma rHom_inv_inv:"\ Ring A; Ring R; x \ carrier A; f \ rHom A R \ \ f (-\<^sub>a\<^bsub>A\<^esub> x) = -\<^sub>a\<^bsub>R\<^esub> (f x)" apply (frule Ring.ring_is_ag [of "A"], frule Ring.ring_is_ag [of "R"]) apply (simp add:rHom_def, erule conjE) apply (simp add:aHom_inv_inv) done lemma rHom_0_0:"\ Ring A; Ring R; f \ rHom A R \ \ f (\\<^bsub>A\<^esub>) = \\<^bsub>R\<^esub>" apply (frule Ring.ring_is_ag [of "A"], frule Ring.ring_is_ag [of "R"]) apply (simp add:rHom_def, (erule conjE)+, simp add:aHom_0_0) done lemma rHom_tOp:"\ Ring A; Ring R; x \ carrier A; y \ carrier A; f \ rHom A R \ \ f (x \\<^sub>r\<^bsub>A\<^esub> y) = (f x) \\<^sub>r\<^bsub>R\<^esub> (f y)" by (simp add:rHom_def) lemma rHom_add:"\f \ rHom A R; x \ carrier A; y \ carrier A\ \ f (x \\<^bsub>A\<^esub> y) = (f x) \\<^bsub>R\<^esub> (f y)" by (simp add:rHom_def aHom_def) lemma rHom_one:"\ Ring A; Ring R;f \ rHom A R \ \ f (1\<^sub>r\<^bsub>A\<^esub>) = (1\<^sub>r\<^bsub>R\<^esub>)" by (simp add:rHom_def) lemma rHom_npow:"\ Ring A; Ring R; x \ carrier A; f \ rHom A R \ \ f (x^\<^bsup>A n\<^esup>) = (f x)^\<^bsup>R n\<^esup>" apply (induct_tac n) apply (simp add:rHom_one) apply (simp, frule_tac n = n in Ring.npClose[of "A" "x"], assumption+, subst rHom_tOp[of "A" "R" _ "x" "f"], assumption+, simp) done lemma rHom_compos:"\Ring A; Ring B; Ring C; f \ rHom A B; g \ rHom B C\ \ compos A g f \ rHom A C" apply (subst rHom_def, simp) apply (frule Ring.ring_is_ag[of "A"], frule Ring.ring_is_ag[of "B"], frule Ring.ring_is_ag[of "C"], frule rHom_aHom[of "f" "A" "B"], frule rHom_aHom[of "g" "B" "C"], simp add:aHom_compos) apply (rule conjI) apply ((rule ballI)+, simp add:compos_def compose_def, frule_tac x = x and y = y in Ring.ring_tOp_closed[of "A"], assumption+, simp) apply (simp add:rHom_tOp) apply (frule_tac a = x in rHom_mem[of "f" "A" "B"], assumption+, frule_tac a = y in rHom_mem[of "f" "A" "B"], assumption+, simp add:rHom_tOp) apply (frule Ring.ring_one[of "A"], frule Ring.ring_one[of "B"], simp add:compos_def compose_def, simp add:rHom_one) done lemma rimg_ag:"\Ring A; Ring R; f \ rHom A R\ \ aGroup (rimg A R f)" apply (frule Ring.ring_is_ag [of "A"], frule Ring.ring_is_ag [of "R"]) apply (simp add:rHom_def, (erule conjE)+) apply (subst aGroup_def) apply (simp add:rimg_def) apply (rule conjI) apply (rule Pi_I)+ apply (simp add:image_def) apply (erule bexE)+ apply simp apply (subst aHom_add [THEN sym, of "A" "R" "f"], assumption+) apply (blast dest: aGroup.ag_pOp_closed) apply (rule conjI) apply ((rule allI, rule impI)+, simp add:image_def, (erule bexE)+, simp) apply (frule_tac x = x and y = xa in aGroup.ag_pOp_closed, assumption+, frule_tac x = xa and y = xb in aGroup.ag_pOp_closed, assumption+) apply (simp add:aHom_add[of "A" "R" "f", THEN sym] aGroup.ag_pOp_assoc) apply (rule conjI) apply ((rule allI, rule impI)+, simp add:image_def, (erule bexE)+, simp) apply (simp add:aHom_add[of "A" "R" "f", THEN sym] aGroup.ag_pOp_commute) apply (rule conjI) apply (rule Pi_I) apply (simp add:image_def, erule bexE, simp) apply (simp add:aHom_inv_inv[THEN sym], frule_tac x = xa in aGroup.ag_mOp_closed[of "A"], assumption+, blast) apply (rule conjI) apply (rule allI, rule impI, simp add:image_def, (erule bexE)+, simp) apply (simp add:aHom_inv_inv[THEN sym], frule_tac x = x in aGroup.ag_mOp_closed[of "A"], assumption+, simp add:aHom_add[of "A" "R" "f", THEN sym]) apply (simp add:aGroup.ag_l_inv1 aHom_0_0) apply (rule conjI) apply (simp add:image_def) apply (frule aHom_0_0[THEN sym, of "A" "R" "f"], assumption+, frule Ring.ring_zero[of "A"], blast) apply (rule allI, rule impI, simp add:image_def, erule bexE, frule_tac a = x in aHom_mem[of "A" "R" "f"], assumption+, simp) apply (simp add:aGroup.ag_l_zero) done lemma rimg_ring:"\Ring A; Ring R; f \ rHom A R \ \ Ring (rimg A R f)" apply (unfold Ring_def [of "rimg A R f"]) apply (frule rimg_ag[of "A" "R" "f"], assumption+) apply (rule conjI, simp add:aGroup_def[of "rimg A R f"]) apply(rule conjI) apply (rule conjI, rule allI, rule impI) apply (frule aGroup.ag_inc_zero[of "rimg A R f"], subst aGroup.ag_pOp_commute, assumption+, simp add:aGroup.ag_r_zero[of "rimg A R f"]) apply (rule conjI) apply (rule Pi_I)+ apply (thin_tac "aGroup (rimg A R f)", simp add:rimg_def, simp add:image_def, (erule bexE)+, simp add:rHom_tOp[THEN sym]) apply (blast dest:Ring.ring_tOp_closed) apply ((rule allI)+, (rule impI)+) apply (thin_tac "aGroup (rimg A R f)", simp add:rimg_def, simp add:image_def, (erule bexE)+, simp) apply (frule_tac x = x and y = xa in Ring.ring_tOp_closed, assumption+, frule_tac x = xa and y = xb in Ring.ring_tOp_closed, assumption+, simp add:rHom_tOp[THEN sym], simp add:Ring.ring_tOp_assoc) apply (rule conjI, rule conjI, (rule allI)+, (rule impI)+) apply (thin_tac "aGroup (rimg A R f)", simp add:rimg_def, simp add:image_def, (erule bexE)+, simp, simp add:rHom_tOp[THEN sym], simp add:Ring.ring_tOp_commute) apply (thin_tac "aGroup (rimg A R f)", simp add:rimg_def, simp add:image_def) apply (subst rHom_one [THEN sym, of "A" "R" "f"], assumption+, frule Ring.ring_one[of "A"], blast) apply (rule conjI, (rule allI)+, (rule impI)+) apply (simp add:rimg_def, fold rimg_def, simp add:image_def, (erule bexE)+, simp) apply (frule rHom_aHom[of "f" "A" "R"], frule Ring.ring_is_ag [of "A"], frule Ring.ring_is_ag [of "R"], simp add:aHom_add[THEN sym], simp add:rHom_tOp[THEN sym]) apply (frule_tac x = xa and y = xb in aGroup.ag_pOp_closed[of "A"], assumption+, frule_tac x = x and y = xa in Ring.ring_tOp_closed[of "A"], assumption+, frule_tac x = x and y = xb in Ring.ring_tOp_closed[of "A"], assumption+, simp add:aHom_add[THEN sym], simp add:rHom_tOp[THEN sym], simp add:Ring.ring_distrib1) apply (rule allI, rule impI, thin_tac "aGroup (rimg A R f)") apply (simp add:rimg_def, simp add:image_def, erule bexE, simp add:rHom_tOp[THEN sym], frule_tac a = x in rHom_mem[of "f" "A" "R"], assumption+, simp add:Ring.ring_l_one) done definition ideal :: "[_ , 'a set] \ bool" where "ideal R I \ (R +> I) \ (\r\carrier R. \x\I. (r \\<^sub>r\<^bsub>R\<^esub> x \ I))" lemma (in Ring) ideal_asubg:"ideal R I \ R +> I" by (simp add:ideal_def) lemma (in Ring) ideal_pOp_closed:"\ideal R I; x \ I; y \ I \ \ x \ y \ I" apply (unfold ideal_def, frule conjunct1, fold ideal_def) apply (cut_tac ring_is_ag, simp add:aGroup.asubg_pOp_closed) done lemma (in Ring) ideal_nsum_closedTr:"ideal R I \ (\j \ n. f j \ I) \ nsum R f n \ I" apply (induct_tac n) apply (rule impI) apply simp apply (rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (rule ideal_pOp_closed, assumption+) apply simp done lemma (in Ring) ideal_nsum_closed:"\ideal R I; \j \ n. f j \ I\ \ nsum R f n \ I" by (simp add:ideal_nsum_closedTr) lemma (in Ring) ideal_subset1:"ideal R I \ I \ carrier R" apply (unfold ideal_def, frule conjunct1, fold ideal_def) apply (simp add:asubGroup_def sg_def, (erule conjE)+) apply (cut_tac ring_is_ag, simp add:aGroup.ag_carrier_carrier) done lemma (in Ring) ideal_subset:"\ideal R I; h \ I\ \ h \ carrier R" by (frule ideal_subset1[of "I"], simp add:subsetD) lemma (in Ring) ideal_ring_multiple:"\ideal R I; x \ I; r \ carrier R\ \ r \\<^sub>r x \ I" by (simp add:ideal_def) lemma (in Ring) ideal_ring_multiple1:"\ideal R I; x \ I; r \ carrier R \ \ x \\<^sub>r r \ I" apply (frule ideal_subset[of "I" "x"], assumption+) apply (simp add:ring_tOp_commute ideal_ring_multiple) done lemma (in Ring) ideal_npow_closedTr:"\ideal R I; x \ I\ \ 0 < n \ x^\<^bsup>R n\<^esup> \ I" apply (induct_tac n, simp) apply (rule impI) apply simp apply (case_tac "n = 0", simp) apply (frule ideal_subset[of "I" "x"], assumption+, simp add:ring_l_one) apply simp apply (frule ideal_subset[of "I" "x"], assumption+, rule ideal_ring_multiple, assumption+, simp add:ideal_subset) done lemma (in Ring) ideal_npow_closed:"\ideal R I; x \ I; 0 < n\ \ x^\<^bsup>R n\<^esup> \ I" by (simp add:ideal_npow_closedTr) lemma (in Ring) times_modTr:"\a \ carrier R; a' \ carrier R; b \ carrier R; b' \ carrier R; ideal R I; a \ (-\<^sub>a b) \ I; a' \ (-\<^sub>a b') \ I\ \ a \\<^sub>r a' \ (-\<^sub>a (b \\<^sub>r b')) \ I" apply (cut_tac ring_is_ag) apply (subgoal_tac "a \\<^sub>r a' \ (-\<^sub>a (b \\<^sub>r b')) = a \\<^sub>r a' \ (-\<^sub>a (a \\<^sub>r b')) \ (a \\<^sub>r b' \ (-\<^sub>a (b \\<^sub>r b')))") apply simp apply (simp add:ring_inv1_2[of "a" "b'"], simp add:ring_inv1_1[of "b" "b'"]) apply (frule aGroup.ag_mOp_closed[of "R" "b'"], assumption+) apply (simp add:ring_distrib1[THEN sym, of "a" "a'" "-\<^sub>a b'"]) apply (frule aGroup.ag_mOp_closed[of "R" "b"], assumption+) apply (frule ring_distrib2[THEN sym, of "b'" "a" "-\<^sub>a b" ], assumption+) apply simp apply (thin_tac "a \\<^sub>r a' \ (-\<^sub>a b) \\<^sub>r b' = a \\<^sub>r (a' \ -\<^sub>a b') \ (a \ -\<^sub>a b) \\<^sub>r b'", thin_tac "a \\<^sub>r b' \ (-\<^sub>a b) \\<^sub>r b' = (a \ -\<^sub>a b) \\<^sub>r b'") apply (frule ideal_ring_multiple[of "I" "a' \ (-\<^sub>a b')" "a"], assumption+, frule ideal_ring_multiple1[of "I" "a \ (-\<^sub>a b)" "b'"], assumption+) apply (simp add:ideal_pOp_closed) apply (frule ring_tOp_closed[of "a" "a'"], assumption+, frule ring_tOp_closed[of "a" "b'"], assumption+, frule ring_tOp_closed[of "b" "b'"], assumption+, frule aGroup.ag_mOp_closed[of "R" "b \\<^sub>r b'"], assumption+, frule aGroup.ag_mOp_closed[of "R" "a \\<^sub>r b'"], assumption+) apply (subst aGroup.ag_pOp_assoc[of "R"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+) apply (simp add:aGroup.ag_pOp_assoc[THEN sym, of "R" "-\<^sub>a (a \\<^sub>r b')" "a \\<^sub>r b'" "-\<^sub>a (b \\<^sub>r b')"], simp add:aGroup.ag_l_inv1 aGroup.ag_l_zero) done lemma (in Ring) ideal_inv1_closed:"\ ideal R I; x \ I \ \ -\<^sub>a x \ I" apply (cut_tac ring_is_ag) apply (unfold ideal_def, frule conjunct1, fold ideal_def) apply (simp add:aGroup.asubg_mOp_closed[of "R" "I"]) done lemma (in Ring) ideal_zero:"ideal R I \ \ \ I" apply (cut_tac ring_is_ag) apply (unfold ideal_def, frule conjunct1, fold ideal_def) apply (simp add:aGroup.asubg_inc_zero) done lemma (in Ring) ideal_zero_forall:"\I. ideal R I \ \ \ I" by (simp add:ideal_zero) lemma (in Ring) ideal_ele_sumTr1:"\ ideal R I; a \ carrier R; b \ carrier R; a \ b \ I; a \ I \ \ b \ I" apply (frule ideal_inv1_closed[of "I" "a"], assumption+) apply (frule ideal_pOp_closed[of "I" "-\<^sub>a a" "a \ b"], assumption+) apply (frule ideal_subset[of "I" "-\<^sub>a a"], assumption+) apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_assoc[THEN sym], simp add:aGroup.ag_l_inv1, simp add:aGroup.ag_l_zero) done lemma (in Ring) ideal_ele_sumTr2:"\ideal R I; a \ carrier R; b \ carrier R; a \ b \ I; b \ I\ \ a \ I" apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute[of "R" "a" "b"]) apply (simp add:ideal_ele_sumTr1[of "I" "b" "a"]) done lemma (in Ring) ideal_condition:"\I \ carrier R; I \ {}; \x\I. \y\I. x \ (-\<^sub>a y) \ I; \r\carrier R. \x\I. r \\<^sub>r x \ I \ \ ideal R I" apply (simp add:ideal_def) apply (cut_tac ring_is_ag) apply (rule aGroup.asubg_test[of "R" "I"], assumption+) done lemma (in Ring) ideal_condition1:"\I \ carrier R; I \ {}; \x\I. \y\I. x \ y \ I; \r\carrier R. \x\I. r \\<^sub>r x \ I \ \ ideal R I" apply (rule ideal_condition[of "I"], assumption+) apply (rule ballI)+ apply (cut_tac ring_is_ag, cut_tac ring_one, frule aGroup.ag_mOp_closed[of "R" "1\<^sub>r"], assumption+) apply (frule_tac x = "-\<^sub>a 1\<^sub>r " in bspec, assumption+, thin_tac "\r\carrier R. \x\I. r \\<^sub>r x \ I", rotate_tac -1, frule_tac x = y in bspec, assumption, thin_tac "\x\I. (-\<^sub>a 1\<^sub>r) \\<^sub>r x \ I") apply (frule_tac c = y in subsetD[of "I" "carrier R"], assumption+, simp add:ring_times_minusl[THEN sym], simp add:ideal_pOp_closed) done lemma (in Ring) zero_ideal:"ideal R {\}" apply (cut_tac ring_is_ag) apply (rule ideal_condition1) apply (simp add:ring_zero) apply simp apply simp apply (cut_tac ring_zero, simp add:aGroup.ag_l_zero) apply simp apply (rule ballI, simp add:ring_times_x_0) done lemma (in Ring) whole_ideal:"ideal R (carrier R)" apply (rule ideal_condition1) apply simp apply (cut_tac ring_zero, blast) apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_closed, simp add:ring_tOp_closed) done lemma (in Ring) ideal_inc_one:"\ideal R I; 1\<^sub>r \ I \ \ I = carrier R" apply (rule equalityI) apply (simp add:ideal_subset1) apply (rule subsetI, frule_tac r = x in ideal_ring_multiple[of "I" "1\<^sub>r"], assumption+, simp add:ring_r_one) done lemma (in Ring) ideal_inc_one1:"ideal R I \ (1\<^sub>r \ I) = (I = carrier R)" apply (rule iffI) apply (simp add:ideal_inc_one) apply (frule sym, thin_tac "I = carrier R", cut_tac ring_one, simp) done definition Unit :: "_ \ 'a \ bool" where "Unit R a \ a \ carrier R \ (\b\carrier R. a \\<^sub>r\<^bsub>R\<^esub> b = 1\<^sub>r\<^bsub>R\<^esub>)" lemma (in Ring) ideal_inc_unit:"\ideal R I; a \ I; Unit R a\ \ 1\<^sub>r \ I" by (simp add:Unit_def, erule conjE, erule bexE, frule_tac r = b in ideal_ring_multiple1[of "I" "a"], assumption+, simp) lemma (in Ring) proper_ideal:"\ideal R I; 1\<^sub>r \ I\ \ I \ carrier R" apply (rule contrapos_pp, simp+) apply (simp add: ring_one) done lemma (in Ring) ideal_inc_unit1:"\a \ carrier R; Unit R a; ideal R I; a \ I\ \ I = carrier R" apply (frule ideal_inc_unit[of "I" "a"], assumption+) apply (rule ideal_inc_one[of "I"], assumption+) done lemma (in Ring) int_ideal:"\ideal R I; ideal R J\ \ ideal R (I \ J)" apply (rule ideal_condition1) apply (frule ideal_subset1[of "I"], frule ideal_subset1[of "J"]) apply blast apply (frule ideal_zero[of "I"], frule ideal_zero[of "J"], blast) apply ((rule ballI)+, simp, (erule conjE)+, simp add:ideal_pOp_closed) apply ((rule ballI)+, simp, (erule conjE)+) apply (simp add:ideal_ring_multiple) done definition ideal_prod::"[_, 'a set, 'a set] \ 'a set" (infix "\\<^sub>r\" 90 ) where "ideal_prod R I J == \ {L. ideal R L \ {x.(\i\I. \j\J. x = i \\<^sub>r\<^bsub>R\<^esub> j)} \ L}" lemma (in Ring) set_sum_mem:"\a \ I; b \ J; I \ carrier R; J \ carrier R\ \ a \ b \ I \ J" apply (cut_tac ring_is_ag) apply (simp add:aGroup.set_sum, blast) done lemma (in Ring) sum_ideals:"\ideal R I1; ideal R I2\ \ ideal R (I1 \ I2)" apply (cut_tac ring_is_ag) apply (frule ideal_subset1[of "I1"], frule ideal_subset1[of "I2"]) apply (rule ideal_condition1) apply (rule subsetI, simp add:aGroup.set_sum, (erule bexE)+) apply (frule_tac h = h in ideal_subset[of "I1"], assumption+, frule_tac h = k in ideal_subset[of "I2"], assumption+, cut_tac ring_is_ag, simp add:aGroup.ag_pOp_closed) apply (frule ideal_zero[of "I1"], frule ideal_zero[of "I2"], frule set_sum_mem[of "\" "I1" "\" "I2"], assumption+, blast) apply (rule ballI)+ apply (simp add:aGroup.set_sum, (erule bexE)+, simp) apply (rename_tac x y i ia j ja) apply (frule_tac h = i in ideal_subset[of "I1"], assumption+, frule_tac h = ia in ideal_subset[of "I1"], assumption+, frule_tac h = j in ideal_subset[of "I2"], assumption+, frule_tac h = ja in ideal_subset[of "I2"], assumption+) apply (subst aGroup.pOp_assocTr43, assumption+) apply (frule_tac x = j and y = ia in aGroup.ag_pOp_commute[of "R"], assumption+, simp) apply (subst aGroup.pOp_assocTr43[THEN sym], assumption+) apply (frule_tac x = i and y = ia in ideal_pOp_closed[of "I1"], assumption+, frule_tac x = j and y = ja in ideal_pOp_closed[of "I2"], assumption+, blast) apply (rule ballI)+ apply (simp add:aGroup.set_sum, (erule bexE)+, simp) apply (rename_tac r x i j) apply (frule_tac h = i in ideal_subset[of "I1"], assumption+, frule_tac h = j in ideal_subset[of "I2"], assumption+) apply (simp add:ring_distrib1) apply (frule_tac x = i and r = r in ideal_ring_multiple[of "I1"], assumption+, frule_tac x = j and r = r in ideal_ring_multiple[of "I2"], assumption+, blast) done lemma (in Ring) sum_ideals_la1:"\ideal R I1; ideal R I2\ \ I1 \ (I1 \ I2)" apply (cut_tac ring_is_ag) apply (rule subsetI) apply (frule ideal_zero[of "I2"], frule_tac h = x in ideal_subset[of "I1"], assumption+, frule_tac x = x in aGroup.ag_r_zero[of "R"], assumption+) apply (subst aGroup.set_sum, assumption, simp add:ideal_subset1, simp add:ideal_subset1, simp, frule sym, thin_tac "x \ \ = x", blast) done lemma (in Ring) sum_ideals_la2:"\ideal R I1; ideal R I2 \ \ I2 \ (I1 \ I2)" apply (cut_tac ring_is_ag) apply (rule subsetI) apply (frule ideal_zero[of "I1"], frule_tac h = x in ideal_subset[of "I2"], assumption+, frule_tac x = x in aGroup.ag_l_zero[of "R"], assumption+) apply (subst aGroup.set_sum, assumption, simp add:ideal_subset1, simp add:ideal_subset1, simp, frule sym, thin_tac "\ \ x = x", blast) done lemma (in Ring) sum_ideals_cont:"\ideal R I; A \ I; B \ I \ \ A \ B \ I" apply (cut_tac ring_is_ag) apply (rule subsetI) apply (frule ideal_subset1[of I], frule subset_trans[of A I "carrier R"], assumption+, frule subset_trans[of B I "carrier R"], assumption+) apply (simp add:aGroup.set_sum[of R], (erule bexE)+, simp) apply (frule_tac c = h in subsetD[of "A" "I"], assumption+, frule_tac c = k in subsetD[of "B" "I"], assumption+) apply (simp add:ideal_pOp_closed) done lemma (in Ring) ideals_set_sum:"\ideal R A; ideal R B; x \ A \ B\ \ \h\A. \k\B. x = h \ k" apply (frule ideal_subset1[of A], frule ideal_subset1[of B]) apply (cut_tac ring_is_ag, simp add:aGroup.set_sum) done definition Rxa :: "[_, 'a ] \ 'a set" (infixl "\\<^sub>p" 200) where "Rxa R a = {x. \r\carrier R. x = (r \\<^sub>r\<^bsub>R\<^esub> a)}" lemma (in Ring) a_in_principal:"a \ carrier R \ a \ Rxa R a" apply (cut_tac ring_one, frule ring_l_one[THEN sym, of "a"]) apply (simp add:Rxa_def, blast) done lemma (in Ring) principal_ideal:"a \ carrier R \ ideal R (Rxa R a)" apply (rule ideal_condition1) apply (rule subsetI, simp add:Rxa_def, erule bexE, simp add:ring_tOp_closed) apply (frule a_in_principal[of "a"], blast) apply ((rule ballI)+, simp add:Rxa_def, (erule bexE)+, simp, subst ring_distrib2[THEN sym], assumption+, cut_tac ring_is_ag, frule_tac x = r and y = ra in aGroup.ag_pOp_closed, assumption+, blast) apply ((rule ballI)+, simp add:Rxa_def, (erule bexE)+, simp, simp add:ring_tOp_assoc[THEN sym]) apply (frule_tac x = r and y = ra in ring_tOp_closed, assumption, blast) done lemma (in Ring) rxa_in_Rxa:"\a \ carrier R; r \ carrier R\ \ r \\<^sub>r a \ Rxa R a" by (simp add:Rxa_def, blast) lemma (in Ring) Rxa_one:"Rxa R 1\<^sub>r = carrier R" apply (rule equalityI) apply (rule subsetI, simp add:Rxa_def, erule bexE) apply (simp add:ring_r_one) apply (rule subsetI, simp add:Rxa_def) apply (frule_tac t = x in ring_r_one[THEN sym], blast) done lemma (in Ring) Rxa_zero:"Rxa R \ = {\}" apply (rule equalityI) apply (rule subsetI) apply (simp add:Rxa_def, erule bexE, simp add:ring_times_x_0) apply (rule subsetI) apply (simp add:Rxa_def) apply (cut_tac ring_zero, frule ring_times_x_0[THEN sym, of "\"], blast) done lemma (in Ring) Rxa_nonzero:"\a \ carrier R; a \ \\ \ Rxa R a \ {\}" apply (rule contrapos_pp, simp+) apply (frule a_in_principal[of "a"]) apply simp done lemma (in Ring) ideal_cont_Rxa:"\ideal R I; a \ I\ \ Rxa R a \ I" apply (rule subsetI) apply (simp add:Rxa_def, erule bexE, simp) apply (simp add:ideal_ring_multiple) done lemma (in Ring) Rxa_mult_smaller:"\ a \ carrier R; b \ carrier R\ \ Rxa R (a \\<^sub>r b) \ Rxa R b" apply (frule rxa_in_Rxa[of b a], assumption, frule principal_ideal[of b]) apply (rule ideal_cont_Rxa[of "R \\<^sub>p b" "a \\<^sub>r b"], assumption+) done lemma (in Ring) id_ideal_psub_sum:"\ideal R I; a \ carrier R; a \ I\ \ I \ I \ Rxa R a" apply (cut_tac ring_is_ag) apply (simp add:psubset_eq) apply (frule principal_ideal) apply (rule conjI) apply (rule sum_ideals_la1, assumption+) apply (rule contrapos_pp) apply simp+ apply (frule sum_ideals_la2[of "I" "Rxa R a"], assumption+) apply (frule a_in_principal[of "a"], frule subsetD[of "Rxa R a" "I \ Rxa R a" "a"], assumption+) apply simp done lemma (in Ring) mul_two_principal_idealsTr:"\a \ carrier R; b \ carrier R; x \ Rxa R a; y \ Rxa R b\ \ \r\carrier R. x \\<^sub>r y = r \\<^sub>r (a \\<^sub>r b)" apply (simp add:Rxa_def, (erule bexE)+) apply simp apply (frule_tac x = ra and y = b in ring_tOp_closed, assumption+) apply (simp add:ring_tOp_assoc) apply (simp add:ring_tOp_assoc[THEN sym, of a _ b]) apply (simp add:ring_tOp_commute[of a], simp add:ring_tOp_assoc) apply (frule_tac x = a and y = b in ring_tOp_closed, assumption+, thin_tac "ra \\<^sub>r b \ carrier R", simp add:ring_tOp_assoc[THEN sym, of _ _ "a \\<^sub>r b"], frule_tac x = r and y = ra in ring_tOp_closed, assumption+) apply (simp add:ring_tOp_commute[of b a]) apply blast done primrec sum_pr_ideals::"[('a, 'm) Ring_scheme, nat \ 'a, nat] \ 'a set" where sum_pr0: "sum_pr_ideals R f 0 = Rxa R (f 0)" | sum_prn: "sum_pr_ideals R f (Suc n) = (Rxa R (f (Suc n))) \\<^bsub>R\<^esub> (sum_pr_ideals R f n)" lemma (in Ring) sum_of_prideals0: "\f. (\l \ n. f l \ carrier R) \ ideal R (sum_pr_ideals R f n)" apply (induct_tac n) apply (rule allI) apply (rule impI) apply simp apply (rule Ring.principal_ideal, rule Ring_axioms, assumption) (** case n **) apply (rule allI, rule impI) apply (frule_tac x = f in spec, thin_tac "\f. (\l \ n. f l \ carrier R) \ ideal R (sum_pr_ideals R f n)") apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (cut_tac a = "f (Suc n)" in principal_ideal, simp) apply (rule_tac ?I1.0 = "Rxa R (f (Suc n))" and ?I2.0 = "sum_pr_ideals R f n" in Ring.sum_ideals, rule Ring_axioms, assumption+) done lemma (in Ring) sum_of_prideals:"\\l \ n. f l \ carrier R\ \ ideal R (sum_pr_ideals R f n)" apply (simp add:sum_of_prideals0) done text \later, we show \sum_pr_ideals\ is the least ideal containing \{f 0, f 1,\, f n}\\ lemma (in Ring) sum_of_prideals1:"\f. (\l \ n. f l \ carrier R) \ f ` {i. i \ n} \ (sum_pr_ideals R f n)" apply (induct_tac n) apply (rule allI, rule impI) apply (simp, simp add:a_in_principal) apply (rule allI, rule impI) apply (frule_tac a = f in forall_spec, thin_tac "\f. (\l \ n. f l \ carrier R) \ f ` {i. i \ n} \ sum_pr_ideals R f n") apply (rule allI, cut_tac n = n in Nset_un, simp) apply (subst Nset_un) apply (cut_tac A = "{i. i \ (Suc n)}" and f = f and B = "carrier R" and ?A1.0 = "{i. i \ n}" and ?A2.0 = "{Suc n}" in im_set_un1, simp, rule Nset_un) apply (thin_tac "\f. (\l\n. f l \ carrier R) \ f ` {i. i \ n} \ sum_pr_ideals R f n", simp) apply (cut_tac n = n and f = f in sum_of_prideals, cut_tac n = n in Nsetn_sub_mem1, simp) apply (cut_tac a = "f (Suc n)" in principal_ideal, simp) apply (frule_tac ?I1.0 = "Rxa R (f (Suc n))" and ?I2.0 = "sum_pr_ideals R f n" in sum_ideals_la1, assumption+, cut_tac a = "f (Suc n)" in a_in_principal, simp, frule_tac A = "R \\<^sub>p f (Suc n)" and B = "R \\<^sub>p f (Suc n) \ sum_pr_ideals R f n" and c = "f (Suc n)" in subsetD, simp+) apply (frule_tac ?I1.0 = "Rxa R (f (Suc n))" and ?I2.0 = "sum_pr_ideals R f n" in sum_ideals_la2, assumption+) apply (rule_tac A = "f ` {j. j \ n}" and B = "sum_pr_ideals R f n" and C = "Rxa R (f (Suc n)) \ sum_pr_ideals R f n" in subset_trans, assumption+) done lemma (in Ring) sum_of_prideals2:"\l \ n. f l \ carrier R \ f ` {i. i \ n} \ (sum_pr_ideals R f n)" apply (simp add:sum_of_prideals1) done lemma (in Ring) sum_of_prideals3:"ideal R I \ \f. (\l \ n. f l \ carrier R) \ (f ` {i. i \ n} \ I) \ (sum_pr_ideals R f n \ I)" apply (induct_tac n) apply (rule allI, rule impI, erule conjE) apply simp apply (rule ideal_cont_Rxa[of I], assumption+) apply (rule allI, rule impI, erule conjE) apply (frule_tac a = f in forall_spec, thin_tac "\f. (\l \ n. f l \ carrier R) \ f `{i. i \ n} \ I \ sum_pr_ideals R f n \ I") apply (simp add:Nset_un) apply (thin_tac "\f. (\l \ n. f l \ carrier R) \ f ` {i. i \ n} \ I \ sum_pr_ideals R f n \ I") apply (frule_tac x = "Suc n" in spec, thin_tac "\l \ (Suc n). f l \ carrier R", simp) apply (cut_tac a = "Suc n" and A = "{i. i \ Suc n}" and f = f in mem_in_image2, simp) apply (frule_tac A = "f ` {i. i \ Suc n}" and B = I and c = "f (Suc n)" in subsetD, assumption+) apply (rule_tac A = "Rxa R (f (Suc n))" and B = "sum_pr_ideals R f n" in sum_ideals_cont[of I], assumption) apply (rule ideal_cont_Rxa[of I], assumption+) done lemma (in Ring) sum_of_prideals4:"\ideal R I; \l \ n. f l \ carrier R; (f ` {i. i \ n} \ I)\ \ sum_pr_ideals R f n \ I" apply (simp add:sum_of_prideals3) done lemma ker_ideal:"\Ring A; Ring R; f \ rHom A R\ \ ideal A (ker\<^bsub>A,R\<^esub> f)" apply (frule Ring.ring_is_ag[of "A"], frule Ring.ring_is_ag[of "R"]) apply (rule Ring.ideal_condition1, assumption+) apply (rule subsetI, simp add:ker_def) apply (simp add:rHom_def, frule conjunct1) apply (frule ker_inc_zero[of "A" "R" "f"], assumption+, blast) apply (rule ballI)+ apply (simp add:ker_def, (erule conjE)+) apply (simp add:aGroup.ag_pOp_closed) apply (simp add:rHom_def, frule conjunct1, simp add:aHom_add, frule Ring.ring_zero[of "R"], simp add:aGroup.ag_l_zero) apply (rule ballI)+ apply (simp add:ker_def, (erule conjE)+) apply (simp add:Ring.ring_tOp_closed) apply (simp add:rHom_tOp) apply (frule_tac a = r in rHom_mem[of "f" "A" "R"], assumption+, simp add:Ring.ring_times_x_0) done subsection "Ring of integers" definition Zr :: "int Ring" where "Zr = \ carrier = Zset, pop = \n\Zset. \m\Zset. (m + n), mop = \l\Zset. -l, zero = 0, tp = \m\Zset. \n\Zset. m * n, un = 1\" lemma ring_of_integers:"Ring Zr" apply (simp add:Ring_def) apply (rule conjI) apply (simp add:Zr_def Zset_def) apply (rule conjI) apply (simp add:Zr_def Zset_def) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:Zr_def Zset_def) apply (rule conjI) apply (simp add:Zr_def Zset_def) apply (rule conjI, rule allI, rule impI, simp add:Zr_def Zset_def) apply (rule conjI, simp add:Zr_def Zset_def) apply (rule conjI, rule allI, rule impI, simp add:Zr_def Zset_def) apply (rule conjI) apply (simp add:Zr_def Zset_def) apply (rule conjI, (rule allI, rule impI)+, simp add:Zr_def Zset_def) apply (rule conjI, (rule allI, rule impI)+, simp add:Zr_def Zset_def) apply (rule conjI) apply (simp add:Zr_def Zset_def) apply (rule conjI, (rule allI, rule impI)+, simp add:Zr_def Zset_def) apply (simp add: distrib_left) apply (rule allI, rule impI) apply (simp add:Zr_def Zset_def) done lemma Zr_zero:"\\<^bsub>Zr\<^esub> = 0" by (simp add:Zr_def) lemma Zr_one:"1\<^sub>r\<^bsub>Zr\<^esub> = 1" by (simp add:Zr_def) lemma Zr_minus:"-\<^sub>a\<^bsub>Zr\<^esub> n = - n" by (simp add:Zr_def Zset_def) lemma Zr_add:"n \\<^bsub>Zr\<^esub> m = n + m" by (simp add:Zr_def Zset_def) lemma Zr_times:"n \\<^sub>r\<^bsub>Zr\<^esub> m = n * m" by (simp add:Zr_def Zset_def) definition lev :: "int set \ int" where "lev I = Zleast {n. n \ I \ 0 < n}" lemma Zr_gen_Zleast:"\ideal Zr I; I \ {0::int}\ \ Rxa Zr (lev I) = I" apply (cut_tac ring_of_integers) apply (simp add:lev_def) apply (subgoal_tac "{n. n \ I \ 0 < n} \ {}") apply (subgoal_tac "{n. n \ I \ 0 < n} \ Zset") apply (subgoal_tac "LB {n. n \ I \ 0 < n} 0") apply (frule_tac A = "{n. n \ I \ 0 < n}" and n = 0 in Zleast, assumption+) apply (erule conjE)+ apply (fold lev_def) defer apply (simp add:LB_def) apply (simp add:Zset_def) apply (frule Ring.ideal_zero[of "Zr" "I"], assumption+, simp add:Zr_zero) apply (frule singleton_sub[of "0" "I"]) apply (frule sets_not_eq[of "I" "{0}"], assumption+, erule bexE, simp) apply (case_tac "0 < a", blast) apply (frule Ring.ring_one[of "Zr"]) apply (frule Ring.ring_is_ag[of "Zr"], frule aGroup.ag_mOp_closed[of "Zr" "1\<^sub>r\<^bsub>Zr\<^esub>"], assumption) apply (frule_tac x = a in Ring.ideal_ring_multiple[of "Zr" "I" _ "-\<^sub>a\<^bsub>Zr\<^esub> 1\<^sub>r\<^bsub>Zr\<^esub>"], assumption+) apply (simp add:Zr_one Zr_minus, thin_tac "ideal Zr I", thin_tac "Ring Zr", thin_tac "1 \ carrier Zr", thin_tac "-1 \ carrier Zr", thin_tac "aGroup Zr") apply (simp add:Zr_def Zset_def) apply (subgoal_tac "0 < - a", blast) apply arith apply (thin_tac "{n \ I. 0 < n} \ {}", thin_tac "{n \ I. 0 < n} \ Zset", thin_tac "LB {n \ I. 0 < n} 0") apply simp apply (erule conjE) apply (frule Ring.ideal_cont_Rxa[of "Zr" "I" "lev I"], assumption+) apply (rule equalityI, assumption, thin_tac "Rxa Zr (lev I) \ I") apply (rule subsetI) apply (simp add:Rxa_def, simp add:Zr_times) apply (cut_tac t = x and b = "lev I" in mult_div_mod_eq [symmetric]) apply (subgoal_tac "x = (x div lev I) * (lev I)", subgoal_tac "x div lev I \ carrier Zr", blast) apply (simp add:Zr_def Zset_def) apply (subgoal_tac "x mod lev I = 0", simp) apply (subst mult.commute, assumption) apply (subgoal_tac "x mod lev I \ I") - apply (thin_tac "x = lev I * (x div lev I) + x mod lev I") - apply (frule_tac a = x in pos_mod_conj[of "lev I"]) + apply (thin_tac "x = lev I * (x div lev I) + x mod lev I") + apply (frule_tac a = x in Divides.pos_mod_conj[of "lev I"]) apply (rule contrapos_pp, simp+) apply (erule conjE) apply (frule_tac a = "x mod (lev I)" in forall_spec) apply simp apply arith apply (frule_tac r = "x div (lev I)" in Ring.ideal_ring_multiple1[of "Zr" "I" "lev I"], assumption+, simp add:Zr_def Zset_def) apply (frule sym, thin_tac "x = lev I * (x div lev I) + x mod lev I") apply (rule_tac a = "lev I * (x div lev I)" and b = "x mod lev I " in Ring.ideal_ele_sumTr1[of "Zr" "I"], assumption+) apply (simp add:Zr_def Zset_def) apply (simp add:Zr_def Zset_def) apply (subst Zr_add) apply simp apply (simp add:Zr_times) done lemma Zr_pir:"ideal Zr I \ \n. Rxa Zr n = I" (** principal ideal ring *) apply (case_tac "I = {(0::int)}") apply (subgoal_tac "Rxa Zr 0 = I") apply blast apply (rule equalityI) apply (rule subsetI) apply (simp add:Rxa_def) apply (simp add:Zr_def Zset_def) apply (rule subsetI) apply (simp add:Rxa_def Zr_def Zset_def) apply (frule Zr_gen_Zleast [of "I"], assumption+) apply blast done section "Quotient rings" lemma (in Ring) mem_set_ar_cos:"\ideal R I; a \ carrier R\ \ a \\<^bsub>R\<^esub> I \ set_ar_cos R I" by (simp add:set_ar_cos_def, blast) lemma (in Ring) I_in_set_ar_cos:"ideal R I \ I \ set_ar_cos R I" apply (cut_tac ring_is_ag, frule ideal_asubg[of "I"], rule aGroup.unit_in_set_ar_cos, assumption+) done lemma (in Ring) ar_coset_same1:"\ideal R I; a \ carrier R; b \ carrier R; b \ (-\<^sub>a a) \ I \ \ a \\<^bsub>R\<^esub> I = b \\<^bsub>R\<^esub> I" apply (cut_tac ring_is_ag) apply (frule aGroup.b_ag_group[of "R"]) apply (simp add:ideal_def asubGroup_def) apply (erule conjE) apply (frule aGroup.ag_carrier_carrier[THEN sym, of "R"]) apply simp apply (frule Group.rcs_eq[of "b_ag R" "I" "a" "b"], assumption+) apply (frule aGroup.agop_gop [of "R"]) apply (frule aGroup.agiop_giop[of "R"]) apply simp apply (simp add:ar_coset_def rcs_def) done lemma (in Ring) ar_coset_same2:"\ideal R I; a \ carrier R; b \ carrier R; a \\<^bsub>R\<^esub> I = b \\<^bsub>R\<^esub> I\ \ b \ (-\<^sub>a a) \ I" apply (cut_tac ring_is_ag) apply (simp add:ar_coset_def) apply (frule aGroup.b_ag_group[of "R"]) apply (simp add:ideal_def asubGroup_def, frule conjunct1, fold asubGroup_def, fold ideal_def, simp add:asubGroup_def) apply (subgoal_tac "a \ carrier (b_ag R)", subgoal_tac "b \ carrier (b_ag R)") apply (simp add:Group.rcs_eq[THEN sym, of "b_ag R" "I" "a" "b"]) apply (frule aGroup.agop_gop [of "R"]) apply (frule aGroup.agiop_giop[of "R"]) apply simp apply (simp add:b_ag_def)+ done lemma (in Ring) ar_coset_same3:"\ideal R I; a \ carrier R; a \\<^bsub>R\<^esub> I = I\ \ a\I" apply (cut_tac ring_is_ag) apply (simp add:ar_coset_def) apply (rule Group.rcs_fixed [of "b_ag R" "I" "a" ]) apply (rule aGroup.b_ag_group, assumption) apply (simp add:ideal_def asubGroup_def) apply (simp add:b_ag_def) apply assumption done lemma (in Ring) ar_coset_same3_1:"\ideal R I; a \ carrier R; a \ I\ \ a \\<^bsub>R\<^esub> I \ I" apply (rule contrapos_pp, simp+) apply (simp add:ar_coset_same3) done lemma (in Ring) ar_coset_same4:"\ideal R I; a \ I\ \ a \\<^bsub>R\<^esub> I = I" apply (cut_tac ring_is_ag) apply (frule ideal_subset[of "I" "a"], assumption+) apply (simp add:ar_coset_def) apply (rule Group.rcs_Unit2 [of "b_ag R" "I""a"]) apply (rule aGroup.b_ag_group, assumption) apply (simp add:ideal_def asubGroup_def) apply assumption done lemma (in Ring) ar_coset_same4_1:"\ideal R I; a \\<^bsub>R\<^esub> I \ I\ \ a \ I" apply (rule contrapos_pp, simp+) apply (simp add:ar_coset_same4) done lemma (in Ring) belong_ar_coset1:"\ideal R I; a \ carrier R; x \ carrier R; x \ (-\<^sub>a a) \ I\ \ x \ a \\<^bsub>R\<^esub> I" apply (frule ar_coset_same1 [of "I" "a" "x"], assumption+) apply (subgoal_tac "x \ x \\<^bsub>R\<^esub> I") apply simp apply (cut_tac ring_is_ag) apply (subgoal_tac "carrier R = carrier (b_ag R)") apply (frule aGroup.agop_gop[THEN sym, of "R"]) apply (frule aGroup.agiop_giop [THEN sym, of "R"]) apply (simp add:ar_coset_def) apply (simp add:ideal_def asubGroup_def) apply (rule Group.a_in_rcs [of "b_ag R" "I" "x"]) apply (simp add: aGroup.b_ag_group) apply simp apply simp apply (simp add:b_ag_def) done lemma (in Ring) a_in_ar_coset:"\ideal R I; a \ carrier R\ \ a \ a \\<^bsub>R\<^esub> I" apply (rule belong_ar_coset1, assumption+) apply (cut_tac ring_is_ag) apply (simp add:aGroup.ag_r_inv1) apply (simp add:ideal_zero) done lemma (in Ring) ar_coset_subsetD:"\ideal R I; a \ carrier R; x \ a \\<^bsub>R\<^esub> I \ \ x \ carrier R" apply (subgoal_tac "carrier R = carrier (b_ag R)") apply (cut_tac ring_is_ag) apply (frule aGroup.agop_gop [THEN sym, of "R"]) apply (frule aGroup.agiop_giop [THEN sym, of "R"]) apply (simp add:ar_coset_def) apply (simp add:ideal_def asubGroup_def) apply (rule Group.rcs_subset_elem[of "b_ag R" "I" "a" "x"]) apply (simp add:aGroup.b_ag_group) apply simp apply assumption+ apply (simp add:b_ag_def) done lemma (in Ring) ar_cos_mem:"\ideal R I; a \ carrier R\ \ a \\<^bsub>R\<^esub> I \ set_rcs (b_ag R) I" apply (cut_tac ring_is_ag) apply (simp add:set_rcs_def ar_coset_def) apply (frule aGroup.ag_carrier_carrier[THEN sym, of "R"]) apply simp apply blast done lemma (in Ring) mem_ar_coset1:"\ideal R I; a \ carrier R; x \ a \\<^bsub>R\<^esub> I\ \ \h\I. h \ a = x" apply (cut_tac ring_is_ag) apply (frule aGroup.ag_carrier_carrier[THEN sym, of "R"]) apply (frule aGroup.agop_gop [THEN sym, of "R"]) apply (frule aGroup.agiop_giop [THEN sym, of "R"]) apply (simp add:ar_coset_def) apply (simp add:ideal_def asubGroup_def) apply (simp add:rcs_def) done lemma (in Ring) ar_coset_mem2:"\ideal R I; a \ carrier R; x \ a \\<^bsub>R\<^esub> I\ \ \h\I. x = a \ h" apply (cut_tac ring_is_ag) apply (frule mem_ar_coset1 [of "I" "a" "x"], assumption+) apply (erule bexE, frule_tac h = h in ideal_subset[of "I"], assumption+) apply (simp add:aGroup.ag_pOp_commute[of "R" _ "a"], frule sym, thin_tac "a \ h = x", blast) done lemma (in Ring) belong_ar_coset2:"\ideal R I; a \ carrier R; x \ a \\<^bsub>R\<^esub> I \ \ x \ (-\<^sub>a a) \ I" apply (cut_tac ring_is_ag) apply (frule mem_ar_coset1, assumption+, erule bexE) apply (frule sym, thin_tac "h \ a = x", simp) apply (frule_tac h = h in ideal_subset[of "I"], assumption) apply (frule aGroup.ag_mOp_closed[of "R" "a"], assumption) apply (subst aGroup.ag_pOp_assoc, assumption+, simp add:aGroup.ag_r_inv1, simp add:aGroup.ag_r_zero) done lemma (in Ring) ar_c_top: "\ideal R I; a \ carrier R; b \ carrier R\ \ (c_top (b_ag R) I (a \\<^bsub>R\<^esub> I) (b \\<^bsub>R\<^esub> I)) = (a \ b) \\<^bsub>R\<^esub> I" apply (cut_tac ring_is_ag, frule ideal_asubg, frule aGroup.asubg_nsubg[of "R" "I"], assumption, frule aGroup.b_ag_group[of "R"]) apply (simp add:ar_coset_def) apply (subst Group.c_top_welldef[THEN sym], assumption+) apply (simp add:aGroup.ag_carrier_carrier)+ apply (simp add:aGroup.agop_gop) done text\Following lemma is not necessary to define a quotient ring. But it makes clear that the binary operation2 of the quotient ring is well defined.\ lemma (in Ring) quotient_ring_tr1:"\ideal R I; a1 \ carrier R; a2 \ carrier R; b1 \ carrier R; b2 \ carrier R; a1 \\<^bsub>R\<^esub> I = a2 \\<^bsub>R\<^esub> I; b1 \\<^bsub>R\<^esub> I = b2 \\<^bsub>R\<^esub> I\ \ (a1 \\<^sub>r b1) \\<^bsub>R\<^esub> I = (a2 \\<^sub>r b2) \\<^bsub>R\<^esub> I" apply (rule ar_coset_same1, assumption+) apply (simp add: ring_tOp_closed)+ apply (frule ar_coset_same2 [of "I" "a1" "a2"], assumption+) apply (frule ar_coset_same2 [of "I" "b1" "b2"], assumption+) apply (frule ring_distrib4[of "a2" "b2" "a1" "b1"], assumption+) apply simp apply (rule ideal_pOp_closed[of "I"], assumption) apply (simp add:ideal_ring_multiple, simp add:ideal_ring_multiple1) done definition rcostOp :: "[_, 'a set] \ (['a set, 'a set] \ 'a set)" where "rcostOp R I = (\X\(set_rcs (b_ag R) I). \Y\(set_rcs (b_ag R) I). {z. \ x \ X. \ y \ Y. \h\I. (x \\<^sub>r\<^bsub>R\<^esub> y) \\<^bsub>R\<^esub> h = z})" lemma (in Ring) rcostOp:"\ideal R I; a \ carrier R; b \ carrier R\ \ rcostOp R I (a \\<^bsub>R\<^esub> I) (b \\<^bsub>R\<^esub> I) = (a \\<^sub>r b) \\<^bsub>R\<^esub> I" apply (cut_tac ring_is_ag) apply (frule ar_cos_mem[of "I" "a"], assumption+) apply (frule ar_cos_mem[of "I" "b"], assumption+) apply (simp add:rcostOp_def) apply (rule equalityI) apply (rule subsetI, simp) apply (erule bexE)+ apply (rule belong_ar_coset1, assumption+) apply (simp add:ring_tOp_closed) apply (frule sym, thin_tac "xa \\<^sub>r y \ h = x", simp) apply (rule aGroup.ag_pOp_closed, assumption) apply (frule_tac x = xa in ar_coset_mem2[of "I" "a"], assumption+, frule_tac x = y in ar_coset_mem2[of "I" "b"], assumption+, (erule bexE)+, simp) apply (rule ring_tOp_closed, rule aGroup.ag_pOp_closed, assumption+, simp add:ideal_subset) apply (rule aGroup.ag_pOp_closed, assumption+, simp add:ideal_subset, simp add:ideal_subset) apply (frule sym, thin_tac "xa \\<^sub>r y \ h = x", simp) apply (frule_tac x = xa in belong_ar_coset2[of "I" "a"], assumption+, frule_tac x = y in belong_ar_coset2[of "I" "b"], assumption+) apply (frule_tac x = xa in ar_coset_subsetD[of "I" "a"], assumption+, frule_tac x = y in ar_coset_subsetD[of "I" "b"], assumption+) apply (subst aGroup.ag_pOp_commute, assumption, simp add:ring_tOp_closed, simp add:ideal_subset) apply (subst aGroup.ag_pOp_assoc, assumption, simp add:ideal_subset, simp add:ring_tOp_closed, rule aGroup.ag_mOp_closed, simp add:ring_tOp_closed, simp add:ring_tOp_closed) apply (rule ideal_pOp_closed, assumption+) apply (rule_tac a = xa and a' = y and b = a and b' = b in times_modTr, assumption+) apply (rule subsetI, simp) apply (frule_tac x = x in ar_coset_mem2[of "I" "a \\<^sub>r b"], simp add:ring_tOp_closed, assumption) apply (erule bexE) apply simp apply (frule a_in_ar_coset[of "I" "a"], assumption+, frule a_in_ar_coset[of "I" "b"], assumption+) apply blast done definition qring :: "[('a, 'm) Ring_scheme, 'a set] \ \ carrier :: 'a set set, pop :: ['a set, 'a set] \ 'a set, mop :: 'a set \ 'a set, zero :: 'a set, tp :: ['a set, 'a set] \ 'a set, un :: 'a set \" where "qring R I = \ carrier = set_rcs (b_ag R) I, pop = c_top (b_ag R) I, mop = c_iop (b_ag R) I, zero = I, tp = rcostOp R I, un = 1\<^sub>r\<^bsub>R\<^esub> \\<^bsub>R\<^esub> I\" abbreviation QRING (infixl "'/\<^sub>r" 200) where "R /\<^sub>r I == qring R I" lemma (in Ring) carrier_qring:"ideal R I \ carrier (qring R I) = set_rcs (b_ag R) I" by (simp add:qring_def) lemma (in Ring) carrier_qring1:"ideal R I \ carrier (qring R I) = set_ar_cos R I" apply (cut_tac ring_is_ag) apply (simp add:carrier_qring set_rcs_def set_ar_cos_def) apply (simp add:ar_coset_def aGroup.ag_carrier_carrier) done lemma (in Ring) qring_ring:"ideal R I \ Ring (qring R I)" apply (cut_tac ring_is_ag) apply (frule ideal_asubg[of "I"], frule aGroup.asubg_nsubg[of "R" "I"], assumption, frule aGroup.b_ag_group[of "R"]) apply (subst Ring_def, simp) apply (rule conjI) apply (rule Pi_I)+ apply (simp add:carrier_qring, simp add:set_rcs_def, (erule bexE)+) apply (subst qring_def, simp) apply (subst Group.c_top_welldef[THEN sym, of "b_ag R" "I"], assumption+) apply (blast dest: Group.mult_closed[of "b_ag R"]) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:qring_def) apply (simp add:Group.Qg_tassoc[of "b_ag R" "I"]) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:qring_def) apply (simp add:set_rcs_def, (erule bexE)+, simp) apply (subst Group.c_top_welldef[THEN sym, of "b_ag R" "I"], assumption+)+ apply (simp add:aGroup.agop_gop) apply (simp add:aGroup.ag_carrier_carrier) apply (simp add:aGroup.ag_pOp_commute) apply (rule conjI) apply (simp add:qring_def Group.Qg_iop_closed) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:qring_def) apply (simp add:Group.Qg_i[of "b_ag R" "I"]) apply (rule conjI) apply (simp add:qring_def) apply (frule Group.nsg_sg[of "b_ag R" "I"], assumption) apply (simp add:Group.unit_rcs_in_set_rcs) apply (rule conjI) apply (rule allI, rule impI) apply (simp add:qring_def) apply (simp add:Group.Qg_unit[of "b_ag R" "I"]) apply (rule conjI) apply(rule Pi_I)+ apply (simp add:qring_def aGroup.aqgrp_carrier) apply (simp add:set_ar_cos_def, (erule bexE)+, simp add:rcostOp, blast dest: ring_tOp_closed) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:qring_def aGroup.aqgrp_carrier) apply (simp add:set_ar_cos_def, (erule bexE)+, simp add:rcostOp) apply (frule_tac x = aa and y = ab in ring_tOp_closed, assumption+, frule_tac x = ab and y = ac in ring_tOp_closed, assumption+, simp add:rcostOp, simp add:ring_tOp_assoc) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:qring_def aGroup.aqgrp_carrier) apply (simp add:set_ar_cos_def, (erule bexE)+, simp add:rcostOp, simp add:ring_tOp_commute) apply (rule conjI) apply (simp add:qring_def aGroup.aqgrp_carrier) apply (cut_tac ring_one, simp add:set_ar_cos_def, blast) apply (rule conjI) apply (rule allI, rule impI)+ apply (simp add:qring_def aGroup.aqgrp_carrier) apply (simp add:set_ar_cos_def, (erule bexE)+, simp) apply (simp add:ar_c_top rcostOp) apply (frule_tac x = ab and y = ac in aGroup.ag_pOp_closed, assumption+, frule_tac x = aa and y = ab in ring_tOp_closed, assumption+ , frule_tac x = aa and y = ac in ring_tOp_closed, assumption+) apply (simp add:ar_c_top rcostOp, simp add:ring_distrib1) apply (rule allI, rule impI) apply (simp add:qring_def aGroup.aqgrp_carrier) apply (simp add:set_ar_cos_def, erule bexE, simp) apply (cut_tac ring_one) apply (simp add:rcostOp, simp add:ring_l_one) done lemma (in Ring) qring_carrier:"ideal R I \ carrier (qring R I) = {X. \a\ carrier R. a \\<^bsub>R\<^esub> I = X}" apply (simp add:carrier_qring1 set_ar_cos_def) apply (rule equalityI) apply (rule subsetI, simp, erule bexE, frule sym, thin_tac "x = a \\<^bsub>R\<^esub> I", blast) apply (rule subsetI, simp, erule bexE, frule sym, thin_tac "a \\<^bsub>R\<^esub> I = x", blast) done lemma (in Ring) qring_mem:"\ideal R I; a \ carrier R\ \ a \\<^bsub>R\<^esub> I \ carrier (qring R I)" apply (simp add:qring_carrier) apply blast done lemma (in Ring) qring_pOp:"\ideal R I; a \ carrier R; b \ carrier R \ \ pop (qring R I) (a \\<^bsub>R\<^esub> I) (b \\<^bsub>R\<^esub> I) = (a \ b) \\<^bsub>R\<^esub> I" by (simp add:qring_def, simp add:ar_c_top) lemma (in Ring) qring_zero:"ideal R I \ zero (qring R I) = I" apply (simp add:qring_def) done lemma (in Ring) qring_zero_1:"\a \ carrier R; ideal R I; a \\<^bsub>R\<^esub> I = I\ \ a \ I" by (frule a_in_ar_coset [of "I" "a"], assumption+, simp) lemma (in Ring) Qring_fix1:"\a \ carrier R; ideal R I; a \ I\ \ a \\<^bsub>R\<^esub> I = I" apply (cut_tac ring_is_ag, frule aGroup.b_ag_group) apply (simp add:ar_coset_def) apply (frule ideal_asubg[of "I"], simp add:asubGroup_def) apply (simp add:Group.rcs_fixed2[of "b_ag R" "I"]) done lemma (in Ring) ar_cos_same:"\a \ carrier R; ideal R I; x \ a \\<^bsub>R\<^esub> I\ \ x \\<^bsub>R\<^esub> I = a \\<^bsub>R\<^esub> I" apply (cut_tac ring_is_ag) apply (rule ar_coset_same1[of "I" "x" "a"], assumption+) apply (rule ar_coset_subsetD[of "I"], assumption+) apply (frule ar_coset_mem2[of "I" "a" "x"], assumption+, erule bexE) apply (frule_tac h = h in ideal_subset[of "I"], assumption, simp add:aGroup.ag_p_inv) apply (frule_tac x = a in aGroup.ag_mOp_closed[of "R"], assumption+, frule_tac x = h in aGroup.ag_mOp_closed[of "R"], assumption+) apply (simp add:aGroup.ag_pOp_assoc[THEN sym], simp add:aGroup.ag_r_inv1 aGroup.ag_l_zero) apply (simp add:ideal_inv1_closed) done lemma (in Ring) qring_tOp:"\ideal R I; a \ carrier R; b \ carrier R\ \ tp (qring R I) (a \\<^bsub>R\<^esub> I) (b \\<^bsub>R\<^esub> I) = (a \\<^sub>r b) \\<^bsub>R\<^esub> I" by (simp add:qring_def, simp add:rcostOp) lemma rind_hom_well_def:"\Ring A; Ring R; f \ rHom A R; a \ carrier A \ \ f a = (f\\<^bsub>A,R\<^esub>) (a \\<^bsub>A\<^esub> (ker\<^bsub>A,R\<^esub> f))" apply (frule ker_ideal[of "A" "R" "f"], assumption+) apply (frule Ring.mem_set_ar_cos[of "A" "ker\<^bsub>A,R\<^esub> f" "a"], assumption+) apply (simp add:rind_hom_def) apply (rule someI2_ex) apply (frule Ring.a_in_ar_coset [of "A" "ker\<^bsub>A,R\<^esub> f" "a"], assumption+, blast) apply (frule_tac x = x in Ring.ar_coset_mem2[of "A" "ker\<^bsub>A,R\<^esub> f" "a"], assumption+, erule bexE, simp, frule_tac h = h in Ring.ideal_subset[of "A" "ker\<^bsub>A,R\<^esub> f"], assumption+) apply (frule_tac Ring.ring_is_ag[of "A"], frule_tac Ring.ring_is_ag[of "R"], simp add:rHom_def, frule conjunct1, simp add:aHom_add) apply (simp add:ker_def) apply (frule aHom_mem[of "A" "R" "f" "a"], assumption+, simp add:aGroup.ag_r_zero) done lemma (in Ring) set_r_ar_cos:"ideal R I \ set_rcs (b_ag R) I = set_ar_cos R I" apply (simp add:set_ar_cos_def set_rcs_def ar_coset_def) apply (cut_tac ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier) done lemma set_r_ar_cos_ker:"\Ring A; Ring R; f \ rHom A R \ \ set_rcs (b_ag A) (ker\<^bsub>A,R\<^esub> f) = set_ar_cos A (ker\<^bsub>A,R\<^esub> f)" apply (frule ker_ideal[of "A" "R" "f"], assumption+) apply (simp add:Ring.carrier_qring[THEN sym], simp add:Ring.carrier_qring1[THEN sym]) done lemma ind_hom_rhom:"\Ring A; Ring R; f \ rHom A R\ \ (f\\<^bsub>A,R\<^esub>) \ rHom (qring A (ker\<^bsub>A,R\<^esub> f)) R" apply (simp add:rHom_def [of "qring A (ker\<^bsub>A,R\<^esub> f)" "R"]) apply (rule conjI) apply (simp add:aHom_def) apply (rule conjI) apply (simp add:qring_def) apply (simp add:rind_hom_def extensional_def) apply (rule Pi_I) apply (frule Ring.ring_is_ag [of "A"], frule Ring.ring_is_ag [of "R"], frule aGroup.b_ag_group [of "R"]) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:set_ar_cos_def) apply (rule conjI) apply (rule impI) apply (erule bexE, simp) apply (frule ker_ideal [of "A" "R" "f"], assumption+) apply (frule_tac a = a in Ring.a_in_ar_coset [of "A" "ker\<^bsub>A,R\<^esub> f"], assumption+) apply (rule someI2_ex, blast) apply (frule_tac I = "ker\<^bsub>A,R\<^esub> f" and a = a and x = xa in Ring.ar_coset_subsetD[of "A"], assumption+) apply (simp add:aGroup.ag_carrier_carrier, simp add:rHom_mem) apply (simp add:set_r_ar_cos_ker, simp add:set_ar_cos_def, rule impI, blast) apply (rule conjI) apply (simp add:qring_def) apply (simp add:set_r_ar_cos_ker) apply (simp add:rind_hom_def extensional_def) apply (rule ballI)+ apply (simp add:qring_def) apply (simp add:set_r_ar_cos_ker) apply (simp add:set_ar_cos_def) apply ((erule bexE)+, simp) apply (frule ker_ideal[of "A" "R" "f"], assumption+) apply (simp add:Ring.ar_c_top) apply (frule Ring.ring_is_ag[of "A"], frule Ring.ring_is_ag[of "R"], frule_tac x = aa and y = ab in aGroup.ag_pOp_closed[of "A"], assumption+) apply (simp add:rind_hom_well_def[THEN sym]) apply (simp add:rHom_def, frule conjunct1, simp add:aHom_add) apply (rule conjI) apply (rule ballI)+ apply (frule ker_ideal[of "A" "R" "f"], assumption+, simp add:Ring.carrier_qring1, simp add:set_ar_cos_def, (erule bexE)+, simp add:qring_def Ring.rcostOp) apply (frule Ring.ring_is_ag[of "A"], frule_tac x = a and y = aa in Ring.ring_tOp_closed[of "A"], assumption+) apply (simp add:rind_hom_well_def[THEN sym], simp add:rHom_tOp) apply (simp add:qring_def) apply (frule Ring.ring_one[of "A"], simp add:rind_hom_well_def[THEN sym], simp add:rHom_one) done lemma ind_hom_injec:"\Ring A; Ring R; f \ rHom A R\ \ injec\<^bsub>(qring A (ker\<^bsub>A,R\<^esub> f)),R\<^esub> (f\\<^bsub>A,R\<^esub>)" apply (simp add:injec_def) apply (frule ind_hom_rhom [of "A" "R" "f"], assumption+) apply (frule rHom_aHom[of "f\\<^bsub>A,R\<^esub>" "A /\<^sub>r (ker\<^bsub>A,R\<^esub> f)" "R"], simp) apply (simp add:ker_def[of _ _ "f\\<^bsub>A,R\<^esub>"]) apply ((subst qring_def)+, simp) apply (simp add:set_r_ar_cos_ker) apply (frule Ring.ring_is_ag[of "A"], frule Ring.ring_is_ag[of "R"], frule ker_ideal[of "A" "R" "f"], assumption+) apply (rule equalityI) apply (rule subsetI) apply (simp, erule conjE) apply (simp add:set_ar_cos_def, erule bexE, simp) apply (simp add:rind_hom_well_def[THEN sym, of "A" "R" "f"], thin_tac "x = a \\<^bsub>A\<^esub> ker\<^bsub>A,R\<^esub> f") apply (rule_tac a = a in Ring.Qring_fix1[of "A" _ "ker\<^bsub>A,R\<^esub> f"], assumption+) apply (simp add:ker_def) apply (rule subsetI, simp) apply (simp add:Ring.I_in_set_ar_cos[of "A" "ker\<^bsub>A,R\<^esub> f"]) apply (frule Ring.ideal_zero[of "A" "ker\<^bsub>A,R\<^esub> f"], assumption+, frule Ring.ring_zero[of "A"]) apply (frule Ring.ar_coset_same4[of "A" "ker\<^bsub>A,R\<^esub> f" "\\<^bsub>A\<^esub>"], assumption+) apply (frule rind_hom_well_def[THEN sym, of "A" "R" "f" "\\<^bsub>A\<^esub>"], assumption+) apply simp apply (rule rHom_0_0, assumption+) done lemma rhom_to_rimg:"\Ring A; Ring R; f \ rHom A R\ \ f \ rHom A (rimg A R f)" apply (frule Ring.ring_is_ag[of "A"], frule Ring.ring_is_ag[of "R"]) apply (subst rHom_def, simp) apply (rule conjI) apply (subst aHom_def, simp) apply (rule conjI) apply (simp add:rimg_def) apply (rule conjI) apply (simp add:rHom_def aHom_def) apply ((rule ballI)+, simp add:rimg_def) apply (rule aHom_add, assumption+) apply (simp add:rHom_aHom, assumption+) apply (rule conjI) apply ((rule ballI)+, simp add:rimg_def, simp add:rHom_tOp) apply (simp add:rimg_def, simp add:rHom_one) done lemma ker_to_rimg:"\Ring A; Ring R; f \ rHom A R \ \ ker\<^bsub>A,R\<^esub> f = ker\<^bsub>A,(rimg A R f)\<^esub> f" apply (frule rhom_to_rimg [of "A" "R" "f"], assumption+) apply (simp add:ker_def) apply (simp add:rimg_def) done lemma indhom_eq:"\Ring A; Ring R; f \ rHom A R\ \ f\\<^bsub>A,(rimg A R f)\<^esub> = f\\<^bsub>A,R\<^esub>" apply (frule rimg_ring[of "A" "R" "f"], assumption+) apply (frule rhom_to_rimg[of "A" "R" "f"], assumption+, frule ind_hom_rhom[of "A" "rimg A R f"], assumption+, frule ind_hom_rhom[of "A" "R" "f"], assumption+) (** extensional **) apply (rule funcset_eq[of "f\\<^bsub>A,rimg A R f\<^esub> " "carrier (A /\<^sub>r (ker\<^bsub>A,R\<^esub> f))" "f\\<^bsub>A,R\<^esub>"]) apply (simp add:ker_to_rimg[THEN sym], simp add:rHom_def[of _ "rimg A R f"] aHom_def) apply (simp add:rHom_def[of _ "R"] aHom_def) apply (simp add:ker_to_rimg[THEN sym]) apply (rule ballI) apply (frule ker_ideal[of "A" "R" "f"], assumption+, simp add:Ring.carrier_qring1) apply (simp add:set_ar_cos_def, erule bexE, simp) apply (simp add:rind_hom_well_def[THEN sym]) apply (frule rind_hom_well_def[THEN sym, of "A" "rimg A R f" "f"], assumption+, simp add:ker_to_rimg[THEN sym]) done lemma indhom_bijec2_rimg:"\Ring A; Ring R; f \ rHom A R\ \ bijec\<^bsub>(qring A (ker\<^bsub>A,R\<^esub> f)),(rimg A R f)\<^esub> (f\\<^bsub>A,R\<^esub>)" apply (frule rimg_ring [of "A" "R" "f"], assumption+) apply (frule rhom_to_rimg[of "A" "R" "f"], assumption+) apply (frule ind_hom_rhom[of "A" "rimg A R f" "f"], assumption+) apply (frule ker_to_rimg[THEN sym, of "A" "R" "f"], assumption+) apply (frule indhom_eq[of "A" "R" "f"], assumption+) apply simp apply (simp add:bijec_def) apply (rule conjI) apply (simp add:injec_def) apply (rule conjI) apply (simp add:rHom_def) apply (frule ind_hom_injec [of "A" "R" "f"], assumption+) apply (simp add:injec_def) apply (simp add:ker_def [of _ _ "f\\<^bsub>A,R\<^esub>"]) apply (simp add:rimg_def) apply (simp add:surjec_def) apply (rule conjI) apply (simp add:rHom_def) apply (rule surj_to_test) apply (simp add:rHom_def aHom_def) apply (rule ballI) apply (simp add:rimg_carrier) apply (simp add:image_def) apply (erule bexE, simp) apply (frule_tac a1 = x in rind_hom_well_def[THEN sym, of "A" "R" "f"], assumption+) apply (frule ker_ideal[of "A" "R" "f"], assumption+, simp add:Ring.carrier_qring1, frule_tac a = x in Ring.mem_set_ar_cos[of "A" "ker\<^bsub>A,R\<^esub> f"], assumption+) apply blast done lemma surjec_ind_bijec:"\Ring A; Ring R; f \ rHom A R; surjec\<^bsub>A,R\<^esub> f\ \ bijec\<^bsub>(qring A (ker\<^bsub>A,R\<^esub> f)),R\<^esub> (f\\<^bsub>A,R\<^esub>)" apply (frule ind_hom_rhom[of "A" "R" "f"], assumption+) apply (simp add:surjec_def) apply (simp add:bijec_def) apply (simp add:ind_hom_injec) apply (simp add:surjec_def) apply (simp add:rHom_aHom) apply (rule surj_to_test) apply (simp add:rHom_def aHom_def) apply (rule ballI) apply (simp add:surj_to_def, frule sym, thin_tac "f ` carrier A = carrier R", simp, thin_tac "carrier R = f ` carrier A") apply (simp add:image_def, erule bexE) apply (frule_tac a1 = x in rind_hom_well_def[THEN sym, of "A" "R" "f"], assumption+) apply (frule ker_ideal[of "A" "R" "f"], assumption+, simp add:Ring.carrier_qring1, frule_tac a = x in Ring.mem_set_ar_cos[of "A" "ker\<^bsub>A,R\<^esub> f"], assumption+) apply blast done lemma ridmap_ind_bijec:"Ring A \ bijec\<^bsub>(qring A (ker\<^bsub>A,A\<^esub> (ridmap A))),A\<^esub> ((ridmap A)\\<^bsub>A,A\<^esub>)" apply (frule ridmap_surjec[of "A"]) apply (rule surjec_ind_bijec [of "A" "A" "ridmap A"], assumption+) apply (simp add:rHom_def, simp add:surjec_def) apply (rule conjI) apply (rule ballI)+ apply (frule_tac x = x and y = y in Ring.ring_tOp_closed[of "A"], assumption+, simp add:ridmap_def) apply (simp add:ridmap_def Ring.ring_one) apply assumption done lemma ker_of_idmap:"Ring A \ ker\<^bsub>A,A\<^esub> (ridmap A) = {\\<^bsub>A\<^esub>}" apply (simp add:ker_def) apply (simp add:ridmap_def) apply (rule equalityI) apply (rule subsetI) apply (simp add:CollectI) apply (rule subsetI) apply (simp add:CollectI) apply (simp add:Ring.ring_zero) done lemma ring_natural_isom:"Ring A \ bijec\<^bsub>(qring A {\\<^bsub>A\<^esub>}),A\<^esub> ((ridmap A)\\<^bsub>A,A\<^esub>)" apply (frule ridmap_ind_bijec) apply (simp add: ker_of_idmap) done (** A /\<^sub>r {0\<^sub>A}\<^sub> \ A **) definition pj :: "[('a, 'm) Ring_scheme, 'a set] \ ('a => 'a set)" where "pj R I = (\x. Pj (b_ag R) I x)" (* pj is projection homomorphism *) lemma pj_Hom:"\Ring R; ideal R I\ \ (pj R I) \ rHom R (qring R I)" apply (simp add:rHom_def) apply (rule conjI) apply (simp add:aHom_def) apply (rule conjI) apply (rule Pi_I) apply (simp add:qring_def) apply (frule Ring.ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:pj_def Pj_def) apply (simp add:set_rcs_def) apply blast apply (rule conjI) apply (simp add:pj_def Pj_def extensional_def) apply (frule Ring.ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier) apply (rule ballI)+ apply (frule Ring.ring_is_ag) apply (frule_tac x = a and y = b in aGroup.ag_pOp_closed, assumption+) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:pj_def Pj_def) apply (simp add:qring_def) apply (frule aGroup.b_ag_group) apply (simp add:aGroup.agop_gop [THEN sym]) apply (subst Group.c_top_welldef[of "b_ag R" "I"], assumption+) apply (frule Ring.ideal_asubg[of "R" "I"], assumption+) apply (simp add:aGroup.asubg_nsubg) apply assumption+ apply simp apply (rule conjI) apply (rule ballI)+ apply (simp add: qring_def) apply (frule_tac x = x and y = y in Ring.ring_tOp_closed, assumption+) apply (frule Ring.ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:pj_def Pj_def) apply (simp add:aGroup.ag_carrier_carrier) apply (frule_tac a1 = x and b1 = y in Ring.rcostOp [THEN sym, of "R" "I"], assumption+) apply (simp add:ar_coset_def) apply (simp add:qring_def) apply (frule Ring.ring_one) apply (frule Ring.ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:pj_def Pj_def) apply (simp add:ar_coset_def) done lemma pj_mem:"\Ring R; ideal R I; x \ carrier R\ \ pj R I x = x \\<^bsub>R\<^esub> I" apply (frule Ring.ring_is_ag) apply (simp add:aGroup.ag_carrier_carrier [THEN sym]) apply (simp add:pj_def Pj_def) apply (simp add:ar_coset_def) done lemma pj_zero:"\Ring R; ideal R I; x \ carrier R\ \ (pj R I x = \\<^bsub>(R /\<^sub>r I)\<^esub>) = (x \ I)" apply (rule iffI) apply (simp add:pj_mem Ring.qring_zero, simp add:Ring.qring_zero_1[of "R" "x" "I"]) apply (simp add:pj_mem Ring.qring_zero, rule Ring.Qring_fix1, assumption+) done lemma pj_surj_to:"\Ring R; ideal R J; X \ carrier (R /\<^sub>r J)\ \ \r\ carrier R. pj R J r = X" apply (simp add:qring_def set_rcs_def, fold ar_coset_def, simp add:b_ag_def, erule bexE, frule_tac x = a in pj_mem[of R J], assumption+, simp) apply blast done lemma invim_of_ideal:"\Ring R; ideal R I; ideal (qring R I) J \ \ ideal R (rInvim R (qring R I) (pj R I) J)" apply (rule Ring.ideal_condition, assumption) apply (simp add:rInvim_def) apply (subgoal_tac "\\<^bsub>R\<^esub> \ rInvim R (qring R I) (pj R I) J") apply (simp add:nonempty) apply (simp add:rInvim_def) apply (simp add: Ring.ring_zero) apply (frule Ring.ring_is_ag) apply (frule pj_Hom [of "R" "I"], assumption+) apply (frule Ring.qring_ring [of "R" "I"], assumption+) apply (frule rHom_0_0 [of "R" "R /\<^sub>r I" "pj R I"], assumption+) apply (simp add:Ring.ideal_zero) apply (rule ballI)+ apply (simp add:rInvim_def) apply (erule conjE)+ apply (rule conjI) apply (frule Ring.ring_is_ag) apply (rule aGroup.ag_pOp_closed, assumption+) apply (rule aGroup.ag_mOp_closed, assumption+) apply (frule pj_Hom [of "R" "I"], assumption+) apply (frule Ring.ring_is_ag) apply (frule_tac x = y in aGroup.ag_mOp_closed [of "R"], assumption+) apply (simp add:rHom_def) apply (erule conjE)+ apply (subst aHom_add [of "R" "R /\<^sub>r I" "pj R I"], assumption+) apply (simp add:Ring.qring_ring Ring.ring_is_ag) apply assumption+ apply (frule Ring.qring_ring [of "R" "I"], assumption+) apply (rule Ring.ideal_pOp_closed, assumption+) apply (subst aHom_inv_inv[of "R" "R /\<^sub>r I" "pj R I"], assumption+) apply (simp add:Ring.ring_is_ag) apply assumption+ apply (frule_tac x = "pj R I y" in Ring.ideal_inv1_closed [of "R /\<^sub>r I" "J"], assumption+) apply (rule ballI)+ apply (simp add:rInvim_def) apply (erule conjE) apply (simp add:Ring.ring_tOp_closed) apply (frule pj_Hom [of "R" "I"], assumption+) apply (subst rHom_tOp [of "R" "R /\<^sub>r I" _ _ "pj R I"], assumption+) apply (frule Ring.qring_ring[of "R" "I"], assumption+) apply (rule Ring.ideal_ring_multiple [of "R /\<^sub>r I" "J"]) apply (simp add:Ring.qring_ring) apply assumption+ apply (simp add:rHom_mem) done lemma pj_invim_cont_I:"\Ring R; ideal R I; ideal (qring R I) J\ \ I \ (rInvim R (qring R I) (pj R I) J)" apply (rule subsetI) apply (simp add:rInvim_def) apply (frule Ring.ideal_subset [of "R" "I"], assumption+) apply simp apply (frule pj_mem [of "R" "I" _], assumption+) apply (simp add:Ring.ar_coset_same4) apply (frule Ring.qring_ring[of "R" "I"], assumption+) apply (frule Ring.ideal_zero [of "qring R I" "J"], assumption+) apply (frule Ring.qring_zero[of "R" "I"], assumption) apply simp done lemma pj_invim_mono1:"\Ring R; ideal R I; ideal (qring R I) J1; ideal (qring R I) J2; J1 \ J2 \ \ (rInvim R (qring R I) (pj R I) J1) \ (rInvim R (qring R I) (pj R I) J2)" apply (rule subsetI) apply (simp add:rInvim_def) apply (simp add:subsetD) done lemma pj_img_ideal:"\Ring R; ideal R I; ideal R J; I \ J\ \ ideal (qring R I) ((pj R I)`J)" apply (rule Ring.ideal_condition [of "qring R I" "(pj R I) `J"]) apply (simp add:Ring.qring_ring) apply (rule subsetI, simp add:image_def) apply (erule bexE) apply (frule_tac h = xa in Ring.ideal_subset [of "R" "J"], assumption+) apply (frule pj_Hom [of "R" "I"], assumption+) apply (simp add:rHom_mem) apply (frule Ring.ideal_zero [of "R" "J"], assumption+) apply (simp add:image_def) apply blast apply (rule ballI)+ apply (simp add:image_def) apply (erule bexE)+ apply (frule pj_Hom [of "R" "I"], assumption+) apply (rename_tac x y s t) apply (frule_tac h = s in Ring.ideal_subset [of "R" "J"], assumption+) apply (frule_tac h = t in Ring.ideal_subset [of "R" "J"], assumption+) apply (simp add:rHom_def) apply (erule conjE)+ apply (frule Ring.ring_is_ag) apply (frule Ring.qring_ring [of "R" "I"], assumption+) apply (frule Ring.ring_is_ag [of "R /\<^sub>r I"]) apply (frule_tac x = t in aGroup.ag_mOp_closed [of "R"], assumption+) apply (frule_tac a1 = s and b1 = "-\<^sub>a\<^bsub>R\<^esub> t" in aHom_add [of "R" "R /\<^sub>r I" "pj R I", THEN sym], assumption+) apply (simp add:aHom_inv_inv) apply (frule_tac x = t in Ring.ideal_inv1_closed [of "R" "J"], assumption+) apply (frule_tac x = s and y = "-\<^sub>a\<^bsub>R\<^esub> t" in Ring.ideal_pOp_closed [of "R" "J"], assumption+) apply blast apply (rule ballI)+ apply (simp add:qring_def) apply (simp add:Ring.set_r_ar_cos) apply (simp add:set_ar_cos_def, erule bexE) apply simp apply (simp add:image_def) apply (erule bexE) apply (frule_tac x = xa in pj_mem [of "R" "I"], assumption+) apply (simp add:Ring.ideal_subset) apply simp apply (subst Ring.rcostOp, assumption+) apply (simp add:Ring.ideal_subset) apply (frule_tac x = xa and r = a in Ring.ideal_ring_multiple [of "R" "J"], assumption+) apply (frule_tac h = "a \\<^sub>r\<^bsub>R\<^esub> xa" in Ring.ideal_subset [of "R" "J"], assumption+) apply (frule_tac x1 = "a \\<^sub>r\<^bsub>R\<^esub> xa" in pj_mem [THEN sym, of "R" "I"], assumption+) apply simp apply blast done lemma npQring:"\Ring R; ideal R I; a \ carrier R\ \ npow (qring R I) (a \\<^bsub>R\<^esub> I) n = (npow R a n) \\<^bsub>R\<^esub> I" apply (induct_tac n) apply (simp add:qring_def) apply (simp add:qring_def) apply (rule Ring.rcostOp, assumption+) apply (rule Ring.npClose, assumption+) done section "Primary ideals, Prime ideals" definition maximal_set :: "['a set set, 'a set] \ bool" where "maximal_set S mx \ mx \ S \ (\s\S. mx \ s \ s = mx)" definition nilpotent :: "[_, 'a] \ bool" where "nilpotent R a \ (\(n::nat). a^\<^bsup>R n\<^esup> = \\<^bsub>R\<^esub>)" definition zero_divisor :: "[_, 'a] \ bool" where "zero_divisor R a \ (\x\ carrier R. x \ \\<^bsub>R\<^esub> \ x \\<^sub>r\<^bsub>R\<^esub> a = \\<^bsub>R\<^esub>)" definition primary_ideal :: "[_, 'a set] \ bool" where "primary_ideal R q \ ideal R q \ (1\<^sub>r\<^bsub>R\<^esub>) \ q \ (\x\ carrier R. \y\ carrier R. x \\<^sub>r\<^bsub>R\<^esub> y \ q \ (\n. (npow R x n) \ q \ y \ q))" definition prime_ideal :: "[_, 'a set] \ bool" where "prime_ideal R p \ ideal R p \ (1\<^sub>r\<^bsub>R\<^esub>) \ p \ (\x\ carrier R. \y\ carrier R. (x \\<^sub>r\<^bsub>R\<^esub> y \ p \ x \ p \ y \ p))" definition maximal_ideal :: "[_, 'a set] \ bool" where "maximal_ideal R mx \ ideal R mx \ 1\<^sub>r\<^bsub>R\<^esub> \ mx \ {J. (ideal R J \ mx \ J)} = {mx, carrier R}" lemma (in Ring) maximal_ideal_ideal:"\maximal_ideal R mx\ \ ideal R mx" by (simp add:maximal_ideal_def) lemma (in Ring) maximal_ideal_proper:"maximal_ideal R mx \ 1\<^sub>r \ mx" by (simp add:maximal_ideal_def) lemma (in Ring) prime_ideal_ideal:"prime_ideal R I \ ideal R I" by (simp add:prime_ideal_def) lemma (in Ring) prime_ideal_proper:"prime_ideal R I \ I \ carrier R" apply (simp add:prime_ideal_def, (erule conjE)+) apply (simp add:proper_ideal) done lemma (in Ring) prime_ideal_proper1:"prime_ideal R p \ 1\<^sub>r \ p" by (simp add:prime_ideal_def) lemma (in Ring) primary_ideal_ideal:"primary_ideal R q \ ideal R q" by (simp add:primary_ideal_def) lemma (in Ring) primary_ideal_proper1:"primary_ideal R q \ 1\<^sub>r \ q" by (simp add:primary_ideal_def) lemma (in Ring) prime_elems_mult_not:"\prime_ideal R P; x \ carrier R; y \ carrier R; x \ P; y \ P \ \ x \\<^sub>r y \ P" apply (simp add:prime_ideal_def, (erule conjE)+) apply (rule contrapos_pp, simp+) apply (frule_tac x = x in bspec, assumption, thin_tac "\x\carrier R. \y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P", frule_tac x = y in bspec, assumption, thin_tac "\y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P", simp) done lemma (in Ring) prime_is_primary:"prime_ideal R p \ primary_ideal R p" apply (unfold primary_ideal_def) apply (rule conjI, simp add:prime_ideal_def) apply (rule conjI, simp add:prime_ideal_def) apply ((rule ballI)+, rule impI) apply (simp add:prime_ideal_def, (erule conjE)+) apply (frule_tac x = x in bspec, assumption, thin_tac "\x\carrier R. \y\carrier R. x \\<^sub>r y \ p \ x \ p \ y \ p", frule_tac x = y in bspec, assumption, thin_tac "\y\carrier R. x \\<^sub>r y \ p \ x \ p \ y \ p", simp) apply (erule disjE) apply (frule_tac t = x in np_1[THEN sym]) apply (frule_tac a = x and A = p and b = "x^\<^bsup>R (Suc 0)\<^esup>" in eq_elem_in, assumption) apply blast apply simp done lemma (in Ring) maximal_prime_Tr0:"\maximal_ideal R mx; x \ carrier R; x \ mx\ \ mx \ (Rxa R x) = carrier R" apply (frule principal_ideal [of "x"]) apply (frule maximal_ideal_ideal[of "mx"]) apply (frule sum_ideals [of "mx" "Rxa R x"], assumption) apply (frule sum_ideals_la1 [of "mx" "Rxa R x"], assumption) apply (simp add:maximal_ideal_def) apply (erule conjE)+ apply (subgoal_tac "mx \ (Rxa R x) \ {J. ideal R J \ mx \ J}") apply simp apply (frule sum_ideals_la2 [of "mx" "Rxa R x"], assumption+) apply (frule a_in_principal [of "x"]) apply (frule subsetD [of "Rxa R x" "mx \ (Rxa R x)" "x"], assumption+) apply (thin_tac "{J. ideal R J \ mx \ J} = {mx, carrier R}") apply (erule disjE) apply simp apply simp apply (thin_tac "{J. ideal R J \ mx \ J} = {mx, carrier R}") apply simp done lemma (in Ring) maximal_prime:"maximal_ideal R mx \ prime_ideal R mx" apply (cut_tac ring_is_ag) apply (simp add:prime_ideal_def) apply (simp add:maximal_ideal_ideal) apply (simp add:maximal_ideal_proper) apply ((rule ballI)+, rule impI) apply (rule contrapos_pp, simp+, erule conjE) apply (frule_tac x = x in maximal_prime_Tr0[of "mx"], assumption+, frule_tac x = y in maximal_prime_Tr0[of "mx"], assumption+, frule maximal_ideal_ideal[of mx], frule ideal_subset1[of mx], frule_tac a = x in principal_ideal, frule_tac a = y in principal_ideal, frule_tac I = "R \\<^sub>p x" in ideal_subset1, frule_tac I = "R \\<^sub>p y" in ideal_subset1) apply (simp add:aGroup.set_sum) apply (cut_tac ring_one) apply (frule sym, thin_tac "{xa. \h\mx. \k\R \\<^sub>p x. xa = h \ k} = carrier R", frule sym, thin_tac "{x. \h\mx. \k\R \\<^sub>p y. x = h \ k} = carrier R") apply (frule_tac a = "1\<^sub>r" and B = "{xa. \i\mx. \j\(Rxa R x). xa = i \ j}" in eq_set_inc[of _ "carrier R"], assumption, frule_tac a = "1\<^sub>r" and B = "{xa. \i\mx. \j\(Rxa R y). xa = i \ j}" in eq_set_inc[of _ "carrier R"], assumption, thin_tac "carrier R = {xa. \i\mx. \j\(Rxa R x). xa = i \ j}", thin_tac "carrier R = {x. \i\mx. \j\(Rxa R y). x = i \ j}") apply (drule CollectD, (erule bexE)+, frule sym, thin_tac "1\<^sub>r = i \ j") apply (drule CollectD, (erule bexE)+, rotate_tac -1, frule sym, thin_tac "1\<^sub>r = ia \ ja") apply (frule_tac h = i in ideal_subset[of mx], assumption, frule_tac h = ia in ideal_subset[of mx], assumption, frule_tac h = j in ideal_subset, assumption+, frule_tac h = ja in ideal_subset, assumption+) apply (cut_tac ring_one) apply (frule_tac x = i and y = j in aGroup.ag_pOp_closed, assumption+) apply (frule_tac x = "i \ j" and y = ia and z = ja in ring_distrib1, assumption+) apply (frule_tac x = ia and y = i and z = j in ring_distrib2, assumption+, frule_tac x = ja and y = i and z = j in ring_distrib2, assumption+, simp) apply (thin_tac "1\<^sub>r \\<^sub>r ia = i \\<^sub>r ia \ j \\<^sub>r ia", thin_tac "1\<^sub>r \\<^sub>r ja = i \\<^sub>r ja \ j \\<^sub>r ja", simp add:ring_l_one[of "1\<^sub>r"]) apply (frule_tac x = ia and r = i in ideal_ring_multiple[of mx], assumption+, frule_tac x = i and r = j in ideal_ring_multiple1[of mx], assumption+, frule_tac x = i and r = ja in ideal_ring_multiple1[of mx], assumption+, frule_tac r = j and x = ia in ideal_ring_multiple[of mx], assumption+) apply (subgoal_tac "j \\<^sub>r ja \ mx") apply (frule_tac x = "i \\<^sub>r ia" and y = "j \\<^sub>r ia" in ideal_pOp_closed[of mx], assumption+) apply ( frule_tac x = "i \\<^sub>r ja" and y = "j \\<^sub>r ja" in ideal_pOp_closed[of mx], assumption+) apply (frule_tac x = "i \\<^sub>r ia \ j \\<^sub>r ia" and y = "i \\<^sub>r ja \ j \\<^sub>r ja" in ideal_pOp_closed[of mx], assumption+, thin_tac "i \ j = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)", thin_tac "ia \ ja = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)") apply (frule sym, thin_tac "1\<^sub>r = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)", simp) apply (simp add:maximal_ideal_def) apply (thin_tac "i \ j = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)", thin_tac "ia \ ja = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)", thin_tac "i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja) \ carrier R", thin_tac "1\<^sub>r = i \\<^sub>r ia \ j \\<^sub>r ia \ (i \\<^sub>r ja \ j \\<^sub>r ja)", thin_tac "i \\<^sub>r j \ mx", thin_tac "i \\<^sub>r ja \ mx", thin_tac "R \\<^sub>p y \ carrier R", thin_tac "R \\<^sub>p x \ carrier R", thin_tac "ideal R (R \\<^sub>p y)", thin_tac "ideal R (R \\<^sub>p x)") apply (simp add:Rxa_def, (erule bexE)+, simp) apply (simp add:ring_tOp_assoc) apply (simp add:ring_tOp_assoc[THEN sym]) apply (frule_tac x = x and y = ra in ring_tOp_commute, assumption+, simp) apply (simp add:ring_tOp_assoc, frule_tac x = x and y = y in ring_tOp_closed, assumption+) apply (frule_tac x1 = r and y1 = ra and z1 = "x \\<^sub>r y" in ring_tOp_assoc[THEN sym], assumption+, simp) apply (frule_tac x = r and y = ra in ring_tOp_closed, assumption+, rule ideal_ring_multiple[of mx], assumption+) done lemma (in Ring) chains_un:"\c \ chains {I. ideal R I \ I \ carrier R}; c \ {}\ \ ideal R (\c)" apply (rule ideal_condition1) apply (rule Union_least[of "c" "carrier R"]) apply (simp add:chains_def, erule conjE, frule_tac c = X in subsetD[of "c" "{I. ideal R I \ I \ carrier R}"], assumption+, simp add:psubset_imp_subset) apply (simp add:chains_def, erule conjE) apply (frule nonempty_ex[of "c"], erule exE) apply (frule_tac c = x in subsetD[of "c" "{I. ideal R I \ I \ carrier R}"], assumption+, simp, erule conjE) apply (frule_tac I = x in ideal_zero, blast) apply (rule ballI)+ apply simp apply (erule bexE)+ apply (simp add: chains_def chain_subset_def) apply (frule conjunct1) apply (frule conjunct2) apply (thin_tac "c \ {I. ideal R I \ I \ carrier R} \ (\x\c. \y\c. x \ y \ y \ x)") apply (frule_tac x = X in bspec, assumption, thin_tac "\x\c. \y\c. x \ y \ y \ x", frule_tac x = Xa in bspec, assumption, thin_tac "\y\c. X \ y \ y \ X") apply (frule_tac c = Xa in subsetD[of "c" "{I. ideal R I \ I \ carrier R}"], assumption+, frule_tac c = X in subsetD[of "c" "{I. ideal R I \ I \ carrier R}"], assumption+, simp) apply (erule conjE)+ apply (erule disjE, frule_tac c = x and A = X and B = Xa in subsetD, assumption+, frule_tac x = x and y = y and I = Xa in ideal_pOp_closed, assumption+, blast) apply (frule_tac c = y and A = Xa and B = X in subsetD, assumption+, frule_tac x = x and y = y and I = X in ideal_pOp_closed, assumption+, blast) apply (rule ballI)+ apply (simp, erule bexE) apply (simp add:chains_def, erule conjE) apply (frule_tac c = X in subsetD[of "c" "{I. ideal R I \ I \ carrier R}"], assumption+, simp, erule conjE) apply (frule_tac I = X and x = x and r = r in ideal_ring_multiple, assumption+, blast) done lemma (in Ring) zeroring_no_maximal:"zeroring R \ \ (\I. maximal_ideal R I)" apply (rule contrapos_pp, simp+, erule exE, frule_tac mx = x in maximal_ideal_ideal) apply (frule_tac I = x in ideal_zero) apply (simp add:zeroring_def, erule conjE, cut_tac ring_one, simp, thin_tac "carrier R = {\}", frule sym, thin_tac "1\<^sub>r = \", simp, thin_tac "\ = 1\<^sub>r") apply (simp add:maximal_ideal_def) done lemma (in Ring) id_maximal_Exist:"\(zeroring R) \ \I. maximal_ideal R I" apply (cut_tac A="{ I. ideal R I \ I \ carrier R }" in Zorn_Lemma2) apply (rule ballI) apply (case_tac "C={}", simp) apply (cut_tac zero_ideal) apply (simp add:zeroring_def) apply (cut_tac Ring, simp, frule not_sym, thin_tac "carrier R \ {\}") apply (cut_tac ring_zero, frule singleton_sub[of "\" "carrier R"], thin_tac "\ \ carrier R") apply (subst psubset_eq) apply blast apply (subgoal_tac "\C \ {I. ideal R I \ I \ carrier R}") apply (subgoal_tac "\x\C. x \ (\C)", blast) apply (rule ballI, rule Union_upper, assumption) apply (simp add:chains_un) apply (cut_tac A = C in Union_least[of _ "carrier R"]) apply (simp add:chains_def, erule conjE, frule_tac c = X and A = C in subsetD[of _ "{I. ideal R I \ I \ carrier R}"], assumption+, simp add:ideal_subset1, simp add:psubset_eq) apply (rule contrapos_pp, simp+, cut_tac ring_one, frule sym, thin_tac "\C = carrier R") apply (frule_tac B = "\C" in eq_set_inc[of "1\<^sub>r" "carrier R"], assumption, thin_tac "carrier R = \C") apply (simp, erule bexE) apply (simp add:chains_def, erule conjE) apply (frule_tac c = X and A = C in subsetD[of _ "{I. ideal R I \ I \ carrier R \ I \ carrier R}"], assumption+, simp, (erule conjE)+) apply (frule_tac I = X in ideal_inc_one, assumption+, simp) apply (erule bexE, simp, erule conjE) apply (subgoal_tac "maximal_ideal R M", blast) apply (simp add:maximal_ideal_def) apply (rule conjI, rule contrapos_pp, simp+, frule_tac I = M in ideal_inc_one, assumption+, simp) apply (rule equalityI) apply (rule subsetI, simp) apply (erule conjE) apply (frule_tac x = x in spec, thin_tac "\x. ideal R x \ x \ carrier R \ M \ x \ x = M", simp) apply (frule_tac I = x in ideal_subset1, simp add:psubset_eq) apply (case_tac "x = carrier R", simp) apply simp apply (rule subsetI, simp) apply (erule disjE) apply simp apply (simp add:whole_ideal) done definition ideal_Int :: "[_, 'a set set] \ 'a set" where "ideal_Int R S == \ S" lemma (in Ring) ideal_Int_ideal:"\S \ {I. ideal R I}; S\{}\ \ ideal R (\ S)" apply (rule ideal_condition1) apply (frule nonempty_ex[of "S"], erule exE) apply (frule_tac c = x in subsetD[of "S" "{I. ideal R I}"], assumption+) apply (simp, frule_tac I = x in ideal_subset1) apply (frule_tac B = x and A = S in Inter_lower) apply (rule_tac A = "\S" and B = x and C = "carrier R" in subset_trans, assumption+) apply (cut_tac ideal_zero_forall, blast) apply (simp, rule ballI) apply (rule ballI)+ apply simp apply (frule_tac x = X in bspec, assumption, thin_tac "\X\S. x \ X", frule_tac x = X in bspec, assumption, thin_tac "\X\S. y \ X") apply (frule_tac c = X in subsetD[of "S" "{I. ideal R I}"], assumption+, simp, rule_tac x = x and y = y in ideal_pOp_closed, assumption+) apply (rule ballI)+ apply (simp, rule ballI) apply (frule_tac x = X in bspec, assumption, thin_tac "\X\S. x \ X", frule_tac c = X in subsetD[of "S" "{I. ideal R I}"], assumption+, simp add:ideal_ring_multiple) done lemma (in Ring) sum_prideals_Int:"\\l \ n. f l \ carrier R; S = {I. ideal R I \ f ` {i. i \ n} \ I}\ \ (sum_pr_ideals R f n) = \ S" apply (rule equalityI) apply (subgoal_tac "\X\S. sum_pr_ideals R f n \ X") apply blast apply (rule ballI) apply (simp, erule conjE) apply (rule_tac I = X and n = n and f = f in sum_of_prideals4, assumption+) apply (subgoal_tac "(sum_pr_ideals R f n) \ S") apply blast apply (simp add:CollectI) apply (simp add: sum_of_prideals2) apply (simp add: sum_of_prideals) done text\This proves that \(sum_pr_ideals R f n)\ is the smallest ideal containing \f ` (Nset n)\\ primrec ideal_n_prod::"[('a, 'm) Ring_scheme, nat, nat \ 'a set] \ 'a set" where ideal_n_prod0: "ideal_n_prod R 0 J = J 0" | ideal_n_prodSn: "ideal_n_prod R (Suc n) J = (ideal_n_prod R n J) \\<^sub>r\<^bsub>R\<^esub> (J (Suc n))" abbreviation IDNPROD ("(3i\\<^bsub>_,_\<^esub> _)" [98,98,99]98) where "i\\<^bsub>R,n\<^esub> J == ideal_n_prod R n J" primrec ideal_pow :: "['a set, ('a, 'more) Ring_scheme, nat] \ 'a set" ("(3_/ \<^bsup>\_ _\<^esup>)" [120,120,121]120) where ip0: "I \<^bsup>\R 0\<^esup> = carrier R" | ipSuc: "I \<^bsup>\R (Suc n)\<^esup> = I \\<^sub>r\<^bsub>R\<^esub> (I \<^bsup>\R n\<^esup>)" lemma (in Ring) prod_mem_prod_ideals:"\ideal R I; ideal R J; i \ I; j \ J\ \ i \\<^sub>r j \ (I \\<^sub>r J)" apply (simp add:ideal_prod_def) apply (rule allI, rule impI, erule conjE, rename_tac X) apply (rule_tac A = "{x. \i\I. \j\J. x = Ring.tp R i j}" and B = X and c = "i \\<^sub>r j" in subsetD, assumption) apply simp apply blast done lemma (in Ring) ideal_prod_ideal:"\ideal R I; ideal R J \ \ ideal R (I \\<^sub>r J)" apply (rule ideal_condition1) apply (simp add:ideal_prod_def) apply (rule subsetI, simp) apply (cut_tac whole_ideal) apply (frule_tac x = "carrier R" in spec, thin_tac "\xa. ideal R xa \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ xa \ x \ xa") apply (subgoal_tac "{x. \i\I. \j\J. x = i \\<^sub>r j} \ carrier R", simp) apply (thin_tac "ideal R (carrier R) \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ carrier R \ x \ carrier R") apply (rule subsetI, simp, (erule bexE)+, simp) apply (frule_tac h = i in ideal_subset[of "I"], assumption+, frule_tac h = j in ideal_subset[of "J"], assumption+) apply (rule_tac x = i and y = j in ring_tOp_closed, assumption+) apply (frule ideal_zero[of "I"], frule ideal_zero[of "J"], subgoal_tac "\ \ I \\<^sub>r\<^bsub> R\<^esub> J", blast) apply (simp add:ideal_prod_def) apply (rule allI, rule impI, erule conjE) apply (rule ideal_zero, assumption) apply (rule ballI)+ apply (simp add:ideal_prod_def) apply (rule allI, rule impI) apply (frule_tac x = xa in spec, thin_tac "\xa. ideal R xa \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ xa \ x \ xa", frule_tac x = xa in spec, thin_tac "\x. ideal R x \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ x \ y \ x", erule conjE, simp, rule_tac x = x and y = y in ideal_pOp_closed, assumption+) apply (rule ballI)+ apply (simp add:ideal_prod_def) apply (rule allI, rule impI, erule conjE) apply (frule_tac x = xa in spec, thin_tac "\xa. ideal R xa \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ xa \ x \ xa", simp) apply (simp add:ideal_ring_multiple) done lemma (in Ring) ideal_prod_commute:"\ideal R I; ideal R J\ \ I \\<^sub>r J = J \\<^sub>r I" apply (simp add:ideal_prod_def) apply (subgoal_tac "{K. ideal R K \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ K} = {K. ideal R K \ {x. \i\J. \j\I. x = i \\<^sub>r j} \ K}") apply simp apply (rule equalityI) apply (rule subsetI, rename_tac X, simp, erule conjE) apply (rule subsetI, simp) apply ((erule bexE)+) apply (subgoal_tac "x \ {x. \i\I. \j\J. x = i \\<^sub>r j}", rule_tac c = x and A = "{x. \i\I. \j\J. x = i \\<^sub>r j}" and B = X in subsetD, assumption+, frule_tac h = i in ideal_subset[of "J"], assumption, frule_tac h = j in ideal_subset[of "I"], assumption, frule_tac x = i and y = j in ring_tOp_commute, assumption+, simp, blast) apply (rule subsetI, simp, erule conjE, rule subsetI, simp, (erule bexE)+, subgoal_tac "xa \ {x. \i\J. \j\I. x = i \\<^sub>r j}", rule_tac c = xa and A = "{x. \i\J. \j\I. x = i \\<^sub>r j}" and B = x in subsetD, assumption+, frule_tac h = i in ideal_subset[of "I"], assumption, frule_tac h = j in ideal_subset[of "J"], assumption, frule_tac x = i and y = j in ring_tOp_commute, assumption+, simp, blast) done lemma (in Ring) ideal_prod_subTr:"\ideal R I; ideal R J; ideal R C; \i\I. \j\J. i \\<^sub>r j \ C\ \ I \\<^sub>r J \ C" apply (simp add:ideal_prod_def) apply (rule_tac B = C and A = "{L. ideal R L \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ L}" in Inter_lower) apply simp apply (rule subsetI, simp, (erule bexE)+, simp) done lemma (in Ring) n_prod_idealTr: "(\k \ n. ideal R (J k)) \ ideal R (ideal_n_prod R n J)" apply (induct_tac n) apply (rule impI) apply simp apply (rule impI) apply (simp only:ideal_n_prodSn) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (rule ideal_prod_ideal, assumption) apply simp done lemma (in Ring) n_prod_ideal:"\\k \ n. ideal R (J k)\ \ ideal R (ideal_n_prod R n J)" apply (simp add:n_prod_idealTr) done lemma (in Ring) ideal_prod_la1:"\ideal R I; ideal R J\ \ (I \\<^sub>r J) \ I" apply (simp add:ideal_prod_def) apply (rule subsetI) apply (simp add:CollectI) apply (subgoal_tac "{x. \i\I. \j\J. x = i \\<^sub>r j} \ I") apply blast apply (thin_tac "\xa. ideal R xa \ {x. \i\I. \j\J. x = i \\<^sub>r j} \ xa \ x \ xa") apply (rule subsetI, simp add:CollectI, (erule bexE)+, frule_tac h = j in ideal_subset[of "J"], assumption+) apply (simp add:ideal_ring_multiple1) done lemma (in Ring) ideal_prod_el1:"\ideal R I; ideal R J; a \ (I \\<^sub>r J)\ \ a \ I" apply (frule ideal_prod_la1 [of "I" "J"], assumption+) apply (rule subsetD, assumption+) done lemma (in Ring) ideal_prod_la2:"\ideal R I; ideal R J \ \ (I \\<^sub>r J) \ J" apply (subst ideal_prod_commute, assumption+, rule ideal_prod_la1[of "J" "I"], assumption+) done lemma (in Ring) ideal_prod_sub_Int:"\ideal R I; ideal R J \ \ (I \\<^sub>r J) \ I \ J" by (simp add:ideal_prod_la1 ideal_prod_la2) lemma (in Ring) ideal_prod_el2:"\ideal R I; ideal R J; a \ (I \\<^sub>r J)\ \ a \ J" by (frule ideal_prod_la2 [of "I" "J"], assumption+, rule subsetD, assumption+) text\\i\\<^bsub>R,n\<^esub> J\ is the product of ideals\ lemma (in Ring) ele_n_prodTr0:"\\k \ (Suc n). ideal R (J k); a \ i\\<^bsub>R,(Suc n)\<^esub> J \ \ a \ (i\\<^bsub>R,n\<^esub> J) \ a \ (J (Suc n))" apply (simp add:Nset_Suc[of n]) apply (cut_tac n_prod_ideal[of n J]) apply (rule conjI) apply (rule ideal_prod_el1 [of "i\\<^bsub>R,n\<^esub> J" "J (Suc n)"], assumption, simp+) apply (rule ideal_prod_el2[of "i\\<^bsub>R,n\<^esub> J" "J (Suc n)"], assumption+, simp+) done lemma (in Ring) ele_n_prodTr1: "(\k \ n. ideal R (J k)) \ a \ ideal_n_prod R n J \ (\k \ n. a \ (J k))" apply (induct_tac n) (** n = 0 **) apply simp (** n **) apply (rule impI) apply (rule allI, rule impI) apply (cut_tac n = n in Nsetn_sub_mem1, simp) apply (erule conjE) apply (frule_tac n = n in ele_n_prodTr0[of _ J a]) apply simp apply (erule conjE, thin_tac "\k\Suc n. ideal R (J k)") apply simp apply (case_tac "k = Suc n", simp) apply (frule_tac m = k and n = "Suc n" in noteq_le_less, assumption+, thin_tac "k \ Suc n") apply (frule_tac x = k and n = "Suc n" in less_le_diff, simp) done lemma (in Ring) ele_n_prod:"\\k \ n. ideal R (J k); a \ ideal_n_prod R n J \ \ \k \ n. a \ (J k)" by (simp add: ele_n_prodTr1 [of "n" "J" "a"]) lemma (in Ring) idealprod_whole_l:"ideal R I \ (carrier R) \\<^sub>r\<^bsub>R\<^esub> I = I" apply (rule equalityI) apply (rule subsetI) apply (simp add:ideal_prod_def) apply (subgoal_tac "{x. \i\carrier R. \j\I. x = i \\<^sub>r j} \ I") apply blast apply (thin_tac "\xa. ideal R xa \ {x. \i\carrier R. \j\I. x = i \\<^sub>r j} \ xa \ x \ xa") apply (rule subsetI) apply simp apply ((erule bexE)+, simp) apply (thin_tac "xa = i \\<^sub>r j", simp add:ideal_ring_multiple) apply (rule subsetI) apply (simp add:ideal_prod_def) apply (rule allI, rule impI) apply (erule conjE) apply (rename_tac xa X) apply (cut_tac ring_one) apply (frule_tac h = xa in ideal_subset[of "I"], assumption, frule_tac x = xa in ring_l_one) apply (subgoal_tac "1\<^sub>r \\<^sub>r xa \ {x. \i\carrier R. \j\I. x = i \\<^sub>r j}") apply (rule_tac c = xa and A = "{x. \i\carrier R. \j\I. x = i \\<^sub>r j}" and B = X in subsetD, assumption+) apply simp apply simp apply (frule sym, thin_tac "1\<^sub>r \\<^sub>r xa = xa", blast) done lemma (in Ring) idealprod_whole_r:"ideal R I \ I \\<^sub>r (carrier R) = I" by (cut_tac whole_ideal, simp add:ideal_prod_commute[of "I" "carrier R"], simp add:idealprod_whole_l) lemma (in Ring) idealpow_1_self:"ideal R I \ I \<^bsup>\R (Suc 0)\<^esup> = I" apply simp apply (simp add:idealprod_whole_r) done lemma (in Ring) ideal_pow_ideal:"ideal R I \ ideal R (I \<^bsup>\R n\<^esup>)" apply (induct_tac n) apply (simp add:whole_ideal) apply simp apply (simp add:ideal_prod_ideal) done lemma (in Ring) ideal_prod_prime:"\ideal R I; ideal R J; prime_ideal R P; I \\<^sub>r J \ P \ \ I \ P \ J \ P" apply (rule contrapos_pp, simp+) apply (erule conjE, simp add:subset_eq, (erule bexE)+) apply (frule_tac i = x and j = xa in prod_mem_prod_ideals[of "I" "J"], assumption+) apply (frule_tac x = "x \\<^sub>r xa" in bspec, assumption, thin_tac "\x\I \\<^sub>r\<^bsub> R\<^esub> J. x \ P") apply (simp add: prime_ideal_def, (erule conjE)+) apply (frule_tac h = x in ideal_subset, assumption, frule_tac x = x in bspec, assumption, thin_tac "\x\carrier R. \y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P", frule_tac h = xa in ideal_subset, assumption, frule_tac x = xa in bspec, assumption, thin_tac "\y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P", simp) done lemma (in Ring) ideal_n_prod_primeTr:"prime_ideal R P \ (\k \ n. ideal R (J k)) \ (ideal_n_prod R n J \ P) \ (\i \ n. (J i) \ P)" apply (induct_tac n) apply simp apply (rule impI) apply (rule impI, simp) apply (cut_tac I = "i\\<^bsub>R,n\<^esub> J" and J = "J (Suc n)" in ideal_prod_prime[of _ _ "P"], rule_tac n = n and J = J in n_prod_ideal, rule allI, simp+) apply (erule disjE, simp) apply (cut_tac n = n in Nsetn_sub_mem1, blast) apply blast done lemma (in Ring) ideal_n_prod_prime:"\prime_ideal R P; \k \ n. ideal R (J k); ideal_n_prod R n J \ P\ \ \i \ n. (J i) \ P" apply (simp add:ideal_n_prod_primeTr) done definition ppa::"[_, nat \ 'a set, 'a set, nat] \ (nat \ 'a)" where "ppa R P A i l = (SOME x. x \ A \ x \ (P (skip i l)) \ x \ P i)" (** Note (ppa R P A) is used to prove prime_ideal_cont1, some element x of A such that x \ P j for (i \ j) and x \ P i **) lemma (in Ring) prod_primeTr:"\prime_ideal R P; ideal R A; \ A \ P; ideal R B; \ B \ P \ \ \x. x \ A \ x \ B \ x \ P" apply (simp add:subset_eq) apply (erule bexE)+ apply (subgoal_tac "x \\<^sub>r xa \ A \ x \\<^sub>r xa \ B \ x \\<^sub>r xa \ P") apply blast apply (rule conjI) apply (rule ideal_ring_multiple1, assumption+) apply (simp add:ideal_subset) apply (rule conjI) apply (rule ideal_ring_multiple, assumption+) apply (simp add:ideal_subset) apply (rule contrapos_pp, simp+) apply (simp add:prime_ideal_def, (erule conjE)+) apply (frule_tac h = x in ideal_subset[of "A"], assumption+, frule_tac h = xa in ideal_subset[of "B"], assumption+, frule_tac x = x in bspec, assumption, thin_tac "\x\carrier R. \y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P", frule_tac x = xa in bspec, assumption, thin_tac "\y\carrier R. x \\<^sub>r y \ P \ x \ P \ y \ P") apply simp done lemma (in Ring) prod_primeTr1:"\\k \ (Suc n). prime_ideal R (P k); ideal R A; \l \ (Suc n). \ (A \ P l); \k \ (Suc n). \l \ (Suc n). k = l \ \ (P k) \ (P l); i \ (Suc n)\ \ \l \ n. ppa R P A i l \ A \ ppa R P A i l \ (P (skip i l)) \ ppa R P A i l \ (P i)" apply (rule allI, rule impI) apply (cut_tac i = i and l = l in skip_il_neq_i) apply (rotate_tac 2) apply (frule_tac x = i in spec, thin_tac "\l \ (Suc n). \ A \ P l", simp) apply (cut_tac l = l in skip_mem[of _ "n" "i"], simp, frule_tac x = "skip i l" in spec, thin_tac "\k \ (Suc n). \l \ (Suc n). k = l \ \ P k \ P l", simp) apply (rotate_tac -1, frule_tac x = i in spec, thin_tac "\la \ (Suc n). skip i l = la \ \ P (skip i l) \ P la", simp) apply (cut_tac P = "P i" and A = A and B = "P (skip i l)" in prod_primeTr, simp, assumption+) apply (frule_tac x = "skip i l" in spec, thin_tac "\k\Suc n. prime_ideal R (P k)", simp, rule prime_ideal_ideal, assumption+) apply (simp add:ppa_def) apply (rule someI2_ex, assumption+) done lemma (in Ring) ppa_mem:"\\k \ (Suc n). prime_ideal R (P k); ideal R A; \l \ (Suc n). \ (A \ P l); \k \ (Suc n). \l \ (Suc n). k = l \ \ (P k) \ (P l); i \ (Suc n); l \ n\ \ ppa R P A i l \ carrier R" apply (frule_tac prod_primeTr1[of n P A], assumption+) apply (rotate_tac -1, frule_tac x = l in spec, thin_tac "\l\n. ppa R P A i l \ A \ ppa R P A i l \ P (skip i l) \ ppa R P A i l \ P i", simp) apply (simp add:ideal_subset) done lemma (in Ring) nsum_memrTr:"(\i \ n. f i \ carrier R) \ (\l \ n. nsum R f l \ carrier R)" apply (cut_tac ring_is_ag) apply (induct_tac n) (** n = 0 **) apply (rule impI, rule allI, rule impI) apply simp (** n **) apply (rule impI) apply (rule allI, rule impI) apply (rule aGroup.nsum_mem, assumption) apply (rule allI, simp) done lemma (in Ring) nsum_memr:"\i \ n. f i \ carrier R \ \l \ n. nsum R f l \ carrier R" by (simp add:nsum_memrTr) lemma (in Ring) nsum_ideal_incTr:"ideal R A \ (\i \ n. f i \ A) \ nsum R f n \ A" apply (induct_tac n) apply (rule impI) apply simp (** n **) apply (rule impI) apply simp apply (rule ideal_pOp_closed, assumption+) apply simp done lemma (in Ring) nsum_ideal_inc:"\ideal R A; \i \ n. f i \ A\ \ nsum R f n \ A" by (simp add:nsum_ideal_incTr) lemma (in Ring) nsum_ideal_excTr:"ideal R A \ (\i \ n. f i \ carrier R) \ (\j \ n. (\l \ {i. i \ n} -{j}. f l \ A) \ (f j \ A)) \ nsum R f n \ A" apply (induct_tac n) (** n = 0 **) apply simp (** n **) apply (rule impI) apply (erule conjE)+ apply (erule exE) apply (case_tac "j = Suc n", simp) apply ( thin_tac "(\j\n. f j \ A) \ \\<^sub>e R f n \ A") apply (erule conjE) apply (cut_tac n = n and f = f in nsum_ideal_inc[of A], assumption, rule allI, simp) apply (rule contrapos_pp, simp+) apply (frule_tac a = "\\<^sub>e R f n" and b = "f (Suc n)" in ideal_ele_sumTr1[of A], simp add:ideal_subset, simp, assumption+, simp) apply (erule conjE, frule_tac m = j and n = "Suc n" in noteq_le_less, assumption, frule_tac x = j and n = "Suc n" in less_le_diff, thin_tac "j \ Suc n", thin_tac "j < Suc n", simp, cut_tac n = n in Nsetn_sub_mem1, simp) apply (erule conjE, frule_tac x = "Suc n" in bspec, simp) apply (rule contrapos_pp, simp+) apply (frule_tac a = "\\<^sub>e R f n" and b = "f (Suc n)" in ideal_ele_sumTr2[of A]) apply (cut_tac ring_is_ag, rule_tac n = n in aGroup.nsum_mem[of R _ f], assumption+, rule allI, simp, simp, assumption+, simp) apply (subgoal_tac "\j\n. (\l\{i. i \ n} - {j}. f l \ A) \ f j \ A", simp, thin_tac "(\j\n. (\l\{i. i \ n} - {j}. f l \ A) \ f j \ A) \ \\<^sub>e R f n \ A") apply (subgoal_tac "\l\{i. i \ n} - {j}. f l \ A", blast, thin_tac "\\<^sub>e R f n \ f (Suc n) \ A", thin_tac "\\<^sub>e R f n \ A") apply (rule ballI) apply (frule_tac x = l in bspec, simp, assumption) done lemma (in Ring) nsum_ideal_exc:"\ideal R A; \i \ n. f i \ carrier R; \j \ n. (\l\{i. i \ n} -{j}. f l \ A) \ (f j \ A) \ \ nsum R f n \ A" by (simp add:nsum_ideal_excTr) lemma (in Ring) nprod_memTr:"(\i \ n. f i \ carrier R) \ (\l. l \ n \ nprod R f l \ carrier R)" apply (induct_tac n) apply (rule impI, rule allI, rule impI, simp) apply (rule impI, rule allI, rule impI) apply (case_tac "l \ n") apply (cut_tac n = n in Nset_Suc, blast) apply (cut_tac m = l and n = "Suc n" in Nat.le_antisym, assumption) apply (simp add: not_less) apply simp apply (rule ring_tOp_closed, simp) apply (cut_tac n = n in Nset_Suc, blast) done lemma (in Ring) nprod_mem:"\\i \ n. f i \ carrier R; l \ n\ \ nprod R f l \ carrier R" by (simp add:nprod_memTr) lemma (in Ring) ideal_nprod_incTr:"ideal R A \ (\i \ n. f i \ carrier R) \ (\l \ n. f l \ A) \ nprod R f n \ A" apply (induct_tac n) (** n = 0 **) apply simp (** n **) apply (rule impI) apply (erule conjE)+ apply simp apply (erule exE) apply (case_tac "l = Suc n", simp) apply (rule_tac x = "f (Suc n)" and r = "nprod R f n" in ideal_ring_multiple[of "A"], assumption+) apply (rule_tac n = "Suc n" and f = f and l = n in nprod_mem, assumption+, simp) apply (erule conjE) apply (frule_tac m = l and n = "Suc n" in noteq_le_less, assumption, frule_tac x = l and n = "Suc n" in less_le_diff, thin_tac "l \ Suc n", thin_tac "l < Suc n", simp) apply (rule_tac x = "nprod R f n" and r = "f (Suc n)" in ideal_ring_multiple1[of "A"], assumption+) apply blast apply simp done lemma (in Ring) ideal_nprod_inc:"\ideal R A; \i \ n. f i \ carrier R; \l \ n. f l \ A\ \ nprod R f n \ A" by (simp add:ideal_nprod_incTr) lemma (in Ring) nprod_excTr:"prime_ideal R P \ (\i \ n. f i \ carrier R) \ (\l \ n. f l \ P) \ nprod R f n \ P" apply (induct_tac n) (** n = 0 **) apply simp (* n = 0 done *) (** n **) apply (rule impI) apply (erule conjE)+ apply simp apply (rule_tac y = "f (Suc n)" and x = "nprod R f n" in prime_elems_mult_not[of "P"], assumption, rule_tac n = n in nprod_mem, rule allI, simp+) done lemma (in Ring) prime_nprod_exc:"\prime_ideal R P; \i \ n. f i \ carrier R; \l \ n. f l \ P\ \ nprod R f n \ P" by (simp add:nprod_excTr) definition nilrad :: "_ \ 'a set" where "nilrad R = {x. x \ carrier R \ nilpotent R x}" lemma (in Ring) id_nilrad_ideal:"ideal R (nilrad R)" apply (cut_tac ring_is_ag) apply (rule ideal_condition1[of "nilrad R"]) apply (rule subsetI) apply (simp add:nilrad_def CollectI) apply (simp add:nilrad_def) apply (cut_tac ring_zero) apply (subgoal_tac "nilpotent R \") apply blast apply (simp add:nilpotent_def) apply (frule np_1[of "\"], blast) apply (rule ballI)+ apply (simp add:nilrad_def nilpotent_def, (erule conjE)+) apply (erule exE)+ apply (simp add:aGroup.ag_pOp_closed[of "R"]) apply (frule_tac x = x and y = y and m = n and n = na in npAdd, assumption+, blast) apply (rule ballI)+ apply (simp add:nilrad_def nilpotent_def, erule conjE, erule exE) apply (simp add:ring_tOp_closed, frule_tac x = r and y = x and n = n in npMul, assumption+, simp, frule_tac x = r and n = n in npClose) apply (simp add:ring_times_x_0, blast) done definition rad_ideal :: "[_, 'a set ] \ 'a set" where "rad_ideal R I = {a. a \ carrier R \ nilpotent (qring R I) ((pj R I) a)}" lemma (in Ring) id_rad_invim:"ideal R I \ rad_ideal R I = (rInvim R (qring R I) (pj R I ) (nilrad (qring R I)))" apply (cut_tac ring_is_ag) apply (rule equalityI) apply (rule subsetI) apply (simp add:rad_ideal_def) apply (erule conjE)+ apply (simp add:rInvim_def) apply (simp add:nilrad_def) apply (subst pj_mem, rule Ring_axioms) apply assumption+ apply (simp add:qring_def ar_coset_def set_rcs_def) apply (simp add:aGroup.ag_carrier_carrier) apply blast apply (rule subsetI) apply (simp add:rInvim_def nilrad_def) apply (simp add: rad_ideal_def) done lemma (in Ring) id_rad_ideal:"ideal R I \ ideal R (rad_ideal R I)" (* thm invim_of_ideal *) apply (subst id_rad_invim [of "I"], assumption) apply (rule invim_of_ideal, rule Ring_axioms, assumption) apply (rule Ring.id_nilrad_ideal) apply (simp add:qring_ring) done lemma (in Ring) id_rad_cont_I:"ideal R I \ I \ (rad_ideal R I)" apply (simp add:rad_ideal_def) apply (rule subsetI, simp, simp add:ideal_subset) apply (simp add:nilpotent_def) apply (subst pj_mem, rule Ring_axioms, assumption+, simp add:ideal_subset) (* thm npQring *) apply (frule_tac h = x in ideal_subset[of "I"], assumption, frule_tac a = x in npQring[OF Ring, of "I" _ "Suc 0"], assumption, simp only:np_1, simp only:Qring_fix1, subst qring_zero[of "I"], assumption) apply blast done lemma (in Ring) id_rad_set:"ideal R I \ rad_ideal R I = {x. x \ carrier R \ (\n. npow R x n \ I)}" apply (simp add:rad_ideal_def) apply (rule equalityI) apply (rule subsetI) apply (simp add:nilpotent_def, erule conjE, erule exE) apply (simp add: pj_mem[OF Ring], simp add:npQring[OF Ring]) apply ( simp add:qring_zero) apply (frule_tac x = x and n = n in npClose) apply (frule_tac a = "x^\<^bsup>R n\<^esup>" in ar_coset_same3[of "I"], assumption+, blast) apply (rule subsetI, simp, erule conjE, erule exE) apply (simp add:nilpotent_def) apply (simp add: pj_mem[OF Ring], simp add:npQring[OF Ring], simp add:qring_zero) apply (frule_tac a = "x^\<^bsup>R n\<^esup>" in ar_coset_same4[of "I"], assumption+) apply blast done lemma (in Ring) rad_primary_prime:"primary_ideal R q \ prime_ideal R (rad_ideal R q)" apply (simp add:prime_ideal_def) apply (frule primary_ideal_ideal[of "q"]) apply (simp add:id_rad_ideal) apply (rule conjI) apply (rule contrapos_pp, simp+) apply (simp add:id_rad_set, erule conjE, erule exE) apply (simp add:npOne) apply (simp add:primary_ideal_proper1[of "q"]) apply ((rule ballI)+, rule impI) apply (rule contrapos_pp, simp+, erule conjE) apply (simp add:id_rad_set, erule conjE, erule exE) apply (simp add:npMul) apply (simp add:primary_ideal_def, (erule conjE)+) apply (frule_tac x = x and n = n in npClose, frule_tac x = y and n = n in npClose) apply (frule_tac x = "x^\<^bsup>R n\<^esup>" in bspec, assumption, thin_tac "\x\carrier R. \y\carrier R. x \\<^sub>r y \ q \ (\n. x^\<^bsup>R n\<^esup> \ q) \ y \ q", frule_tac x = "y^\<^bsup>R n\<^esup>" in bspec, assumption, thin_tac "\y\carrier R. x^\<^bsup>R n\<^esup> \\<^sub>r y \ q \ (\na. x^\<^bsup>R n\<^esup>^\<^bsup>R na\<^esup> \ q) \ y \ q", simp) apply (simp add:npMulExp) done lemma (in Ring) npow_notin_prime:"\prime_ideal R P; x \ carrier R; x \ P\ \ \n. npow R x n \ P" apply (rule allI) apply (induct_tac n) apply simp apply (simp add:prime_ideal_proper1) apply simp apply (frule_tac x = x and n = na in npClose) apply (simp add:prime_elems_mult_not) done lemma (in Ring) npow_in_prime:"\prime_ideal R P; x \ carrier R; \n. npow R x n \ P \ \ x \ P" apply (rule contrapos_pp, simp+) apply (frule npow_notin_prime, assumption+) apply blast done definition mul_closed_set::"[_, 'a set ] \ bool" where "mul_closed_set R S \ S \ carrier R \ (\s\S. \t\S. s \\<^sub>r\<^bsub>R\<^esub> t \ S)" locale Idomain = Ring + assumes idom: "\a \ carrier R; b \ carrier R; a \\<^sub>r b = \\ \ a = \ \ b = \" (* integral domain *) locale Corps = fixes K (structure) assumes f_is_ring: "Ring K" and f_inv: "\x\carrier K - {\}. \x' \ carrier K. x' \\<^sub>r x = 1\<^sub>r" (** integral domain **) lemma (in Ring) mul_closed_set_sub:"mul_closed_set R S \ S \ carrier R" by (simp add:mul_closed_set_def) lemma (in Ring) mul_closed_set_tOp_closed:"\mul_closed_set R S; s \ S; t \ S\ \ s \\<^sub>r t \ S" by (simp add:mul_closed_set_def) lemma (in Corps) f_inv_unique:"\ x \ carrier K - {\}; x' \ carrier K; x'' \ carrier K; x' \\<^sub>r x = 1\<^sub>r; x'' \\<^sub>r x = 1\<^sub>r \ \ x' = x''" apply (cut_tac f_is_ring) apply (cut_tac x = x' and y = x and z = x'' in Ring.ring_tOp_assoc[of K], assumption+, simp, assumption, simp) apply (simp add:Ring.ring_l_one[of K], simp add:Ring.ring_tOp_commute[of K x x''] Ring.ring_r_one[of K]) done definition invf :: "[_, 'a] \ 'a" where "invf K x = (THE y. y \ carrier K \ y \\<^sub>r\<^bsub>K\<^esub> x = 1\<^sub>r\<^bsub>K\<^esub>)" lemma (in Corps) invf_inv:"x \ carrier K - {\} \ (invf K x) \ carrier K \ (invf K x) \\<^sub>r x = 1\<^sub>r " apply (simp add:invf_def) apply (rule theI') apply (rule ex_ex1I) apply (cut_tac f_inv, blast) apply (rule_tac x' = xa and x'' = y in f_inv_unique[of x]) apply simp+ done definition npowf :: "_ \ 'a \ int \ 'a" where "npowf K x n = (if 0 \ n then npow K x (nat n) else npow K (invf K x) (nat (- n)))" abbreviation NPOWF :: "['a, _, int] \ 'a" ("(3_\<^bsub>_\<^esub>\<^bsup>_\<^esup>)" [77,77,78]77) where "a\<^bsub>K\<^esub>\<^bsup>n\<^esup> == npowf K a n" abbreviation IOP :: "['a, _] \ 'a" ("(_\<^bsup>\ _\<^esup>)" [87,88]87) where "a\<^bsup>\K\<^esup> == invf K a" lemma (in Idomain) idom_is_ring: "Ring R" .. lemma (in Idomain) idom_tOp_nonzeros:"\x \ carrier R; y \ carrier R; x \ \; y \ \\ \ x \\<^sub>r y \ \" apply (rule contrapos_pp, simp+) apply (cut_tac idom[of x y]) apply (erule disjE, simp+) done lemma (in Idomain) idom_potent_nonzero: "\x \ carrier R; x \ \\ \ npow R x n \ \ " apply (induct_tac n) apply simp (* case 0 *) apply (rule contrapos_pp, simp+) apply (frule ring_l_one[of "x", THEN sym]) apply simp apply (simp add:ring_times_0_x) (* case (Suc n) *) apply (rule contrapos_pp, simp+) apply (frule_tac n = n in npClose[of x], cut_tac a = "x^\<^bsup>R n\<^esup>" and b = x in idom, assumption+) apply (erule disjE, simp+) done lemma (in Idomain) idom_potent_unit:"\a \ carrier R; 0 < n\ \ (Unit R a) = (Unit R (npow R a n))" apply (rule iffI) apply (simp add:Unit_def, erule bexE) apply (simp add:npClose) apply (frule_tac x1 = a and y1 = b and n1 = n in npMul[THEN sym], assumption, simp add:npOne) apply (frule_tac x = b and n = n in npClose, blast) apply (case_tac "n = Suc 0", simp only: np_1) apply (simp add:Unit_def, erule conjE, erule bexE) apply (cut_tac x = a and n = "n - Suc 0" in npow_suc[of R], simp del:npow_suc, thin_tac "a^\<^bsup>R n\<^esup> = a^\<^bsup>R (n - Suc 0)\<^esup> \\<^sub>r a", frule_tac x = a and n = "n - Suc 0" in npClose, frule_tac x = "a^\<^bsup>R (n - Suc 0)\<^esup>" and y = a in ring_tOp_commute, assumption+, simp add:ring_tOp_assoc, frule_tac x = "a^\<^bsup>R (n - Suc 0)\<^esup>" and y = b in ring_tOp_closed, assumption+) apply blast done lemma (in Idomain) idom_mult_cancel_r:"\a \ carrier R; b \ carrier R; c \ carrier R; c \ \; a \\<^sub>r c = b \\<^sub>r c\ \ a = b" apply (cut_tac ring_is_ag) apply (frule ring_tOp_closed[of "a" "c"], assumption+, frule ring_tOp_closed[of "b" "c"], assumption+) apply (simp add:aGroup.ag_eq_diffzero[of "R" "a \\<^sub>r c" "b \\<^sub>r c"], simp add:ring_inv1_1, frule aGroup.ag_mOp_closed[of "R" "b"], assumption, simp add:ring_distrib2[THEN sym, of "c" "a" "-\<^sub>a b"]) apply (frule aGroup.ag_pOp_closed[of "R" "a" "-\<^sub>a b"], assumption+) apply (subst aGroup.ag_eq_diffzero[of R a b], assumption+) apply (rule contrapos_pp, simp+) apply (frule idom_tOp_nonzeros[of "a \ -\<^sub>a b" c], assumption+, simp) done lemma (in Idomain) idom_mult_cancel_l:"\a \ carrier R; b \ carrier R; c \ carrier R; c \ \; c \\<^sub>r a = c \\<^sub>r b\ \ a = b" apply (simp add:ring_tOp_commute) apply (simp add:idom_mult_cancel_r) done lemma (in Corps) invf_closed1:"x \ carrier K - {\} \ invf K x \ (carrier K) - {\}" apply (frule invf_inv[of x], erule conjE) apply (rule contrapos_pp, simp+) apply (cut_tac f_is_ring) apply ( simp add:Ring.ring_times_0_x[of K]) apply (frule sym, thin_tac "\ = 1\<^sub>r", simp, erule conjE) apply (frule Ring.ring_l_one[of K x], assumption) apply (rotate_tac -1, frule sym, thin_tac "1\<^sub>r \\<^sub>r x = x", simp add:Ring.ring_times_0_x) done lemma (in Corps) linvf:"x \ carrier K - {\} \ (invf K x) \\<^sub>r x = 1\<^sub>r" by (simp add:invf_inv) lemma (in Corps) field_is_ring:"Ring K" by (simp add:f_is_ring) lemma (in Corps) invf_one:"1\<^sub>r \ \ \ invf K (1\<^sub>r) = 1\<^sub>r" apply (cut_tac field_is_ring) apply (frule_tac Ring.ring_one) apply (cut_tac invf_closed1 [of "1\<^sub>r"]) apply (cut_tac linvf[of "1\<^sub>r"]) apply (simp add:Ring.ring_r_one[of "K"]) apply simp+ done lemma (in Corps) field_tOp_assoc:"\x \ carrier K; y \ carrier K; z \ carrier K\ \ x \\<^sub>r y \\<^sub>r z = x \\<^sub>r (y \\<^sub>r z)" apply (cut_tac field_is_ring) apply (simp add:Ring.ring_tOp_assoc) done lemma (in Corps) field_tOp_commute:"\x \ carrier K; y \ carrier K\ \ x \\<^sub>r y = y \\<^sub>r x" apply (cut_tac field_is_ring) apply (simp add:Ring.ring_tOp_commute) done lemma (in Corps) field_inv_inv:"\x \ carrier K; x \ \\ \ (x\<^bsup>\K\<^esup>)\<^bsup>\K\<^esup> = x" apply (cut_tac invf_closed1[of "x"]) apply (cut_tac invf_inv[of "x\<^bsup>\K\<^esup>"], erule conjE) apply (frule field_tOp_assoc[THEN sym, of "x\<^bsup>\ K\<^esup>\<^bsup>\ K\<^esup>" "x\<^bsup>\ K\<^esup>" "x"], simp, assumption, simp) apply (cut_tac field_is_ring, simp add:Ring.ring_l_one Ring.ring_r_one, erule conjE, cut_tac invf_inv[of x], erule conjE, simp add:Ring.ring_r_one) apply simp+ done lemma (in Corps) field_is_idom:"Idomain K" apply (rule Idomain.intro) apply (simp add:field_is_ring) apply (cut_tac field_is_ring) apply (rule Idomain_axioms.intro) apply (rule contrapos_pp, simp+, erule conjE) apply (cut_tac x = a in invf_closed1, simp, simp, erule conjE) apply (frule_tac x = "a\<^bsup>\ K\<^esup>" and y = a and z = b in field_tOp_assoc, assumption+) apply (simp add:linvf Ring.ring_times_x_0 Ring.ring_l_one) done lemma (in Corps) field_potent_nonzero:"\x \ carrier K; x \ \\ \ x^\<^bsup>K n\<^esup> \ \" apply (cut_tac field_is_idom) apply (cut_tac field_is_ring, simp add:Idomain.idom_potent_nonzero) done lemma (in Corps) field_potent_nonzero1:"\x \ carrier K; x \ \\ \ x\<^bsub>K\<^esub>\<^bsup>n\<^esup> \ \" apply (simp add:npowf_def) apply (case_tac "0 \ n") apply (simp add:field_potent_nonzero) apply simp apply (cut_tac invf_closed1[of "x"], simp+, (erule conjE)+) apply (simp add:field_potent_nonzero) apply simp done lemma (in Corps) field_nilp_zero:"\x \ carrier K; x^\<^bsup>K n\<^esup> = \\ \ x = \" by (rule contrapos_pp, simp+, simp add:field_potent_nonzero) lemma (in Corps) npowf_mem:"\a \ carrier K; a \ \\ \ npowf K a n \ carrier K" apply (simp add:npowf_def) apply (cut_tac field_is_ring) apply (case_tac "0 \ n", simp, simp add:Ring.npClose, simp) apply (cut_tac invf_closed1[of "a"], simp, erule conjE, simp add:Ring.npClose, simp) done lemma (in Corps) field_npowf_exp_zero:"\a \ carrier K; a \ \\ \ npowf K a 0 = 1\<^sub>r" by (cut_tac field_is_ring, simp add:npowf_def) lemma (in Corps) npow_exp_minusTr1:"\x \ carrier K; x \ \; 0 \ i\ \ 0 \ i - (int j) \ x\<^bsub>K\<^esub>\<^bsup>(i - (int j))\<^esup> = x^\<^bsup>K (nat i)\<^esup> \\<^sub>r (x\<^bsup>\K\<^esup>)^\<^bsup>K j\<^esup>" apply (cut_tac field_is_ring, cut_tac invf_closed1[of "x"], simp, simp add:npowf_def, erule conjE) apply (induct_tac "j", simp) apply (frule Ring.npClose[of "K" "x" "nat i"], assumption+, simp add:Ring.ring_r_one) apply (rule impI, simp) apply (subst zdiff) apply (simp add:add.commute[of "1"]) apply (cut_tac z = i and w = "int n + 1" in zdiff, simp only:minus_add_distrib, thin_tac "i - (int n + 1) = i + (- int n + - 1)") apply (cut_tac z = "i + - int n" in nat_diff_distrib[of "1"], simp, simp) apply (simp only:zdiff[of _ "1"], simp) apply (cut_tac field_is_idom) apply (frule_tac n = "nat i" in Ring.npClose[of "K" "x"], assumption+, frule_tac n = "nat i" in Ring.npClose[of "K" "x\<^bsup>\ K\<^esup>"], assumption+, frule_tac n = n in Ring.npClose[of "K" "x\<^bsup>\ K\<^esup>"], assumption+ ) apply (rule_tac a = "x^\<^bsup>K (nat (i + (- int n - 1)))\<^esup>" and b = "x^\<^bsup>K (nat i)\<^esup> \\<^sub>r (x\<^bsup>\ K\<^esup>^\<^bsup>K n\<^esup> \\<^sub>r x\<^bsup>\ K\<^esup>)" and c = x in Idomain.idom_mult_cancel_r[of "K"], assumption+) apply (simp add:Ring.npClose, rule Ring.ring_tOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+) apply (subgoal_tac "0 < nat (i - int n)") apply (subst Ring.npMulElmR, assumption+, simp, simp add:field_tOp_assoc[THEN sym, of "x^\<^bsup>K (nat i)\<^esup>" _ "x\<^bsup>\ K\<^esup>"]) apply (subst field_tOp_assoc[of _ _ x]) apply (rule Ring.ring_tOp_closed[of K], assumption+) apply (simp add: linvf) apply (subst Ring.ring_r_one[of K], assumption) apply auto apply (metis Ring.npClose) apply (simp only: uminus_add_conv_diff [symmetric] add.assoc [symmetric]) apply (simp add: algebra_simps nat_diff_distrib Suc_diff_Suc) apply (smt Ring.npMulElmR Suc_nat_eq_nat_zadd1 nat_diff_distrib' nat_int of_nat_0_le_iff) done lemma (in Corps) npow_exp_minusTr2:"\x \ carrier K; x \ \; 0 \ i; 0 \ j; 0 \ i - j\ \ x\<^bsub>K\<^esub>\<^bsup>(i - j)\<^esup> = x^\<^bsup>K (nat i)\<^esup> \\<^sub>r (x\<^bsup>\K\<^esup>)^\<^bsup>K (nat j)\<^esup>" apply (frule npow_exp_minusTr1[of "x" "i" "nat j"], assumption+) apply simp done lemma (in Corps) npowf_inv:"\x \ carrier K; x \ \; 0 \ j\ \ x\<^bsub>K\<^esub>\<^bsup>j\<^esup> = (x\<^bsup>\K\<^esup>)\<^bsub>K\<^esub>\<^bsup>(-j)\<^esup>" apply (simp add:npowf_def) apply (rule impI, simp add:zle) apply (simp add:field_inv_inv) done lemma (in Corps) npowf_inv1:"\x \ carrier K; x \ \; \ 0 \ j\ \ x\<^bsub>K\<^esub>\<^bsup>j\<^esup> = (x\<^bsup>\K\<^esup>)\<^bsub>K\<^esub>\<^bsup>(-j)\<^esup>" apply (simp add:npowf_def) done lemma (in Corps) npowf_inverse:"\x \ carrier K; x \ \\ \ x\<^bsub>K\<^esub>\<^bsup>j\<^esup> = (x\<^bsup>\K\<^esup>)\<^bsub>K\<^esub>\<^bsup>(-j)\<^esup>" apply (case_tac "0 \ j") apply (simp add:npowf_inv, simp add:npowf_inv1) done lemma (in Corps) npowf_expTr1:"\x \ carrier K; x \ \; 0 \ i; 0 \ j; 0 \ i - j\ \ x\<^bsub>K\<^esub>\<^bsup>(i - j)\<^esup> = x\<^bsub>K\<^esub>\<^bsup>i\<^esup> \\<^sub>r x\<^bsub>K\<^esub>\<^bsup>(- j)\<^esup>" apply (simp add:npow_exp_minusTr2) apply (simp add:npowf_def) done lemma (in Corps) npowf_expTr2:"\x \ carrier K; x \ \; 0 \ i + j\ \ x\<^bsub>K\<^esub>\<^bsup>(i + j)\<^esup> = x\<^bsub>K\<^esub>\<^bsup>i\<^esup> \\<^sub>r x\<^bsub>K\<^esub>\<^bsup>j\<^esup>" apply (cut_tac field_is_ring) apply (case_tac "0 \ i") apply (case_tac "0 \ j") apply (simp add:npowf_def, simp add:nat_add_distrib, rule Ring.npMulDistr[THEN sym], assumption+) apply (subst zminus_minus[THEN sym, of "i" "j"], subst npow_exp_minusTr2[of "x" "i" "-j"], assumption+) apply (simp add:zle, simp add:zless_imp_zle, simp add:npowf_def) apply (simp add:add.commute[of "i" "j"], subst zminus_minus[THEN sym, of "j" "i"], subst npow_exp_minusTr2[of "x" "j" "-i"], assumption+) apply (simp add:zle, simp add:zless_imp_zle, simp) apply (frule npowf_mem[of "x" "i"], assumption+, frule npowf_mem[of "x" "j"], assumption+, simp add:field_tOp_commute[of "x\<^bsub>K\<^esub>\<^bsup>i\<^esup>" "x\<^bsub>K\<^esub>\<^bsup>j\<^esup>"]) apply (simp add:npowf_def) done lemma (in Corps) npowf_exp_add:"\x \ carrier K; x \ \\ \ x\<^bsub>K\<^esub>\<^bsup>(i + j)\<^esup> = x\<^bsub>K\<^esub>\<^bsup>i\<^esup> \\<^sub>r x\<^bsub>K\<^esub>\<^bsup>j\<^esup>" apply (case_tac "0 \ i + j") apply (simp add:npowf_expTr2) apply (simp add:npowf_inv1[of "x" "i + j"]) apply (simp add:zle) apply (subgoal_tac "0 < -i + -j") prefer 2 apply simp apply (thin_tac "i + j < 0") apply (frule zless_imp_zle[of "0" "-i + -j"]) apply (thin_tac "0 < -i + -j") apply (cut_tac invf_closed1[of "x"]) apply (simp, erule conjE, frule npowf_expTr2[of "x\<^bsup>\K\<^esup>" "-i" "-j"], assumption+) apply (simp add:zdiff[THEN sym]) apply (simp add:npowf_inverse, simp) done lemma (in Corps) npowf_exp_1_add:"\x \ carrier K; x \ \\ \ x\<^bsub>K\<^esub>\<^bsup>(1 + j)\<^esup> = x \\<^sub>r x\<^bsub>K\<^esub>\<^bsup>j\<^esup>" apply (simp add:npowf_exp_add[of "x" "1" "j"]) apply (cut_tac field_is_ring) apply (simp add:npowf_def, simp add:Ring.ring_l_one) done lemma (in Corps) npowf_minus:"\x \ carrier K; x \ \\ \ (x\<^bsub>K\<^esub>\<^bsup>j\<^esup>)\<^bsup>\K\<^esup> = x\<^bsub>K\<^esub>\<^bsup>(- j)\<^esup>" apply (frule npowf_exp_add[of "x" "j" "-j"], assumption+) apply (simp add:field_npowf_exp_zero) apply (cut_tac field_is_ring) apply (frule npowf_mem[of "x" "j"], assumption+) apply (frule field_potent_nonzero1[of "x" "j"], assumption+) apply (cut_tac invf_closed1[of "x\<^bsub>K\<^esub>\<^bsup>j\<^esup>"], simp, erule conjE, frule Ring.ring_r_one[of "K" "(x\<^bsub>K\<^esub>\<^bsup>j\<^esup>)\<^bsup>\K\<^esup>"], assumption, simp, thin_tac "1\<^sub>r = x\<^bsub>K\<^esub>\<^bsup>j\<^esup> \\<^sub>r x\<^bsub>K\<^esub>\<^bsup>- j\<^esup>", frule npowf_mem[of "x" "-j"], assumption+) apply (simp add:field_tOp_assoc[THEN sym], simp add:linvf, simp add:Ring.ring_l_one, simp) done lemma (in Ring) residue_fieldTr:"\maximal_ideal R mx; x \ carrier(qring R mx); x \ \\<^bsub>(qring R mx)\<^esub>\ \\y\carrier (qring R mx). y \\<^sub>r\<^bsub>(qring R mx)\<^esub> x = 1\<^sub>r\<^bsub>(qring R mx)\<^esub>" apply (frule maximal_ideal_ideal[of "mx"]) apply (simp add:qring_carrier) apply (simp add:qring_zero) apply (simp add:qring_def) apply (erule bexE) apply (frule sym, thin_tac "a \\<^bsub>R\<^esub> mx = x", simp) apply (frule_tac a = a in ar_coset_same4_1[of "mx"], assumption+) apply (frule_tac x = a in maximal_prime_Tr0[of "mx"], assumption+) apply (cut_tac ring_one) apply (rotate_tac -2, frule sym, thin_tac "mx \ R \\<^sub>p a = carrier R") apply (frule_tac B = "mx \ R \\<^sub>p a" in eq_set_inc[of "1\<^sub>r" "carrier R"], assumption+, thin_tac "carrier R = mx \ R \\<^sub>p a") apply (frule ideal_subset1[of mx]) apply (frule_tac a = a in principal_ideal, frule_tac I = "R \\<^sub>p a" in ideal_subset1) apply (cut_tac ring_is_ag, simp add:aGroup.set_sum, (erule bexE)+) apply (thin_tac "ideal R (R \\<^sub>p a)", thin_tac "R \\<^sub>p a \ carrier R", simp add:Rxa_def, (erule bexE)+, simp, thin_tac "k = r \\<^sub>r a") apply (frule_tac a = r and b = a in rcostOp[of "mx"], assumption+) apply (frule_tac x = r and y = a in ring_tOp_closed, assumption+) apply (frule_tac a = "r \\<^sub>r a" and x = h and b = "1\<^sub>r" in aGroup.ag_eq_sol2[of "R"], assumption+) apply (simp add:ideal_subset) apply (simp add:ring_one, simp) apply (frule_tac a = h and b = "1\<^sub>r \ -\<^sub>a (r \\<^sub>r a)" and A = mx in eq_elem_in, assumption+) apply (frule_tac a = "r \\<^sub>r a" and b = "1\<^sub>r" in ar_coset_same1[of "mx"], rule ring_tOp_closed, assumption+, rule ring_one, assumption) apply (frule_tac a1 = "r \\<^sub>r a" and h1 = h in aGroup.arcos_fixed[THEN sym, of R mx], unfold ideal_def, erule conjE, assumption+, thin_tac "R +> mx \ (\r\carrier R. \x\mx. r \\<^sub>r x \ mx)", thin_tac "x = a \\<^bsub>R\<^esub> mx", thin_tac "1\<^sub>r = h \ r \\<^sub>r a", thin_tac "h = 1\<^sub>r \ -\<^sub>a (r \\<^sub>r a)", thin_tac "1\<^sub>r \ -\<^sub>a (r \\<^sub>r a) \ mx") apply (rename_tac b h k r) apply simp apply blast done (* constdefs (structure R) field_cd::"_ \ bool" "field_cd R == \x\(carrier R - {\}). \y\carrier R. y \\<^sub>r x = 1\<^sub>r" *) (* field condition *) (* constdefs (structure R) rIf :: "_ \ 'a \ 'a " *) (** rIf is ring_invf **) (* "rIf R == \x. (SOME y. y \ carrier R \ y \\<^sub>r x = 1\<^sub>r)" *) (* constdefs (structure R) Rf::"_ \ 'a field" "Rf R == \carrier = carrier R, pop = pop R, mop = mop R, zero = zero R, tp = tp R, un = un R, invf = rIf R\" *) (* constdefs (structure R) Rf :: "_ \ \ carrier :: 'a set, pOp :: ['a, 'a] \ 'a, mOp ::'a \ 'a, zero :: 'a, tOp :: ['a, 'a] \ 'a, one ::'a, iOp ::'a \ 'a\" "Rf R == \ carrier = carrier R, pOp = pOp R, mOp = mOp R, zero = zero R, tOp = tOp R, one = one R, iOp = ring_iOp R\" *) (* lemma (in Ring) rIf_mem:"\field_cd R; x \ carrier R - {\}\ \ rIf R x \ carrier R \ rIf R x \ \" apply (simp add:rIf_def) apply (rule someI2_ex) apply (simp add:field_cd_def, blast) apply (simp add:field_cd_def) apply (thin_tac "\x\carrier R - {\}. \y\carrier R. y \\<^sub>r x = 1\<^sub>r") apply (erule conjE)+ apply (rule contrapos_pp, simp+) apply (frule sym, thin_tac "\ \\<^sub>r x = 1\<^sub>r", simp add:ring_times_0_x) apply (frule ring_l_one[of "x"]) apply (simp add:ring_times_0_x) done lemma (in Ring) rIf:"\field_cd R; x \ carrier R - {\}\ \ (rIf R x) \\<^sub>r x = 1\<^sub>r" apply (simp add:rIf_def) apply (rule someI2_ex) apply (simp add:field_cd_def, blast) apply simp done lemma (in Ring) field_cd_integral:"field_cd R \ Idomain R" apply (rule Idomain.intro) apply assumption apply (rule Idomain_axioms.intro) apply (rule contrapos_pp, simp+, erule conjE) apply (cut_tac x = a in rIf_mem, assumption, simp, erule conjE) apply (frule_tac x = "rIf R a" and y = a and z = b in ring_tOp_assoc, assumption+, simp add:rIf) apply (simp add:ring_l_one ring_times_x_0) done lemma (in Ring) Rf_field:"field_cd R \ field (Rf R)" apply (rule field.intro) apply (simp add:Rf_def) apply (rule Ring.intro) apply (simp add:pop_closed) apply ( cut_tac ring_is_ag, simp add:aGroup.ag_pOp_assoc) apply (simp add:Rf_def, cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute) apply (simp add:mop_closed) apply (simp add: apply (rule conjI) prefer 2 apply (rule conjI) apply (rule univar_func_test, rule ballI) apply (simp, erule conjE, simp add:Rf_def) apply (rule rIf_mem, assumption+, simp) apply (rule allI, rule impI) apply (simp add:Rf_def) apply (frule_tac x = x in rIf, simp, assumption) apply (subst Rf_def, simp add:Ring_def) apply (cut_tac ring_is_ag) apply (rule conjI, simp add:aGroup_def) apply (rule conjI, (rule allI, rule impI)+, simp add:aGroup.ag_pOp_assoc) apply (rule conjI, (rule allI, rule impI)+, simp add:aGroup.ag_pOp_commute) apply (rule conjI, rule univar_func_test, rule ballI, simp add:aGroup.ag_mOp_closed) apply (rule conjI, rule allI, rule impI, simp add:aGroup.ag_l_inv1) apply (simp add:aGroup.ag_inc_zero) apply (rule conjI, rule allI, rule impI, simp add:aGroup.ag_l_zero) apply (rule conjI, rule bivar_func_test, (rule ballI)+, simp add:ring_tOp_closed) apply (rule conjI, (rule allI, rule impI)+, simp add:ring_tOp_assoc) apply (rule conjI, (rule allI, rule impI)+, simp add:ring_tOp_commute) apply (simp add:ring_one) apply (rule conjI, (rule allI, rule impI)+, simp add:ring_distrib1) apply (rule allI, rule impI, simp add:ring_l_one) done *) lemma (in Ring) residue_field_cd:"maximal_ideal R mx \ Corps (qring R mx)" apply (rule Corps.intro) apply (rule Ring.qring_ring, rule Ring_axioms) apply (simp add:maximal_ideal_ideal) apply (simp add:residue_fieldTr[of "mx"]) done (* lemma (in Ring) qRf_field:"maximal_ideal R mx \ field (Rf (qring R mx))" apply (frule maximal_ideal_ideal[of "mx"]) apply (frule qring_ring [of "mx"]) apply (frule residue_field_cd[of "mx"]) apply (rule Ring.Rf_field, assumption+) done lemma (in Ring) qRf_pj_rHom:"maximal_ideal R mx \ (pj R mx) \ rHom R (Rf (qring R mx))" apply (frule maximal_ideal_ideal[of "mx"]) apply (frule pj_Hom[OF Ring, of "mx"]) apply (simp add:rHom_def aHom_def Rf_def) done *) lemma (in Ring) maximal_set_idealTr: "maximal_set {I. ideal R I \ S \ I = {}} mx \ ideal R mx" by (simp add:maximal_set_def) lemma (in Ring) maximal_setTr:"\maximal_set {I. ideal R I \ S \ I = {}} mx; ideal R J; mx \ J \ \ S \ J \ {}" by (rule contrapos_pp, simp+, simp add:psubset_eq, erule conjE, simp add:maximal_set_def, blast) lemma (in Ring) mulDisj:"\mul_closed_set R S; 1\<^sub>r \ S; \ \ S; T = {I. ideal R I \ S \ I = {}}; maximal_set T mx \ \ prime_ideal R mx" apply (simp add:prime_ideal_def) apply (rule conjI, simp add:maximal_set_def, rule conjI, simp add:maximal_set_def) apply (rule contrapos_pp, simp+) apply ((erule conjE)+, blast) apply ((rule ballI)+, rule impI) apply (rule contrapos_pp, simp+, (erule conjE)+) apply (cut_tac a = x in id_ideal_psub_sum[of "mx"], simp add:maximal_set_def, assumption+, cut_tac a = y in id_ideal_psub_sum[of "mx"], simp add:maximal_set_def, assumption+) apply (frule_tac J = "mx \ R \\<^sub>p x" in maximal_setTr[of "S" "mx"], rule sum_ideals, simp add:maximal_set_def, simp add:principal_ideal, assumption, thin_tac "mx \ mx \ R \\<^sub>p x") apply (frule_tac J = "mx \ R \\<^sub>p y" in maximal_setTr[of "S" "mx"], rule sum_ideals, simp add:maximal_set_def, simp add:principal_ideal, assumption, thin_tac "mx \ mx \ R \\<^sub>p y") apply (frule_tac A = "S \ (mx \ R \\<^sub>p x)" in nonempty_ex, frule_tac A = "S \ (mx \ R \\<^sub>p y)" in nonempty_ex, (erule exE)+, simp, (erule conjE)+) apply (rename_tac x y s1 s2, thin_tac "S \ (mx \ R \\<^sub>p x) \ {}", thin_tac "S \ (mx \ R \\<^sub>p y) \ {}") apply (frule maximal_set_idealTr, frule_tac a = x in principal_ideal, frule_tac a = y in principal_ideal, frule ideal_subset1[of mx], frule_tac I = "R \\<^sub>p x" in ideal_subset1, frule_tac I = "R \\<^sub>p y" in ideal_subset1) apply (cut_tac ring_is_ag, simp add:aGroup.set_sum[of R mx], erule bexE, erule bexE, simp) apply (frule_tac s = s1 and t = s2 in mul_closed_set_tOp_closed, simp, assumption, simp, frule_tac c = h in subsetD[of mx "carrier R"], assumption+, frule_tac c = k and A = "R \\<^sub>p x" in subsetD[of _ "carrier R"], assumption+) apply ( cut_tac mul_closed_set_sub, frule_tac c = s2 in subsetD[of S "carrier R"], assumption+, simp add:ring_distrib2) apply ((erule bexE)+, simp, frule_tac c = ha in subsetD[of mx "carrier R"], assumption+, frule_tac c = ka and A = "R \\<^sub>p y" in subsetD[of _ "carrier R"], assumption+, simp add:ring_distrib1) apply (frule_tac x = h and r = ha in ideal_ring_multiple1[of mx], assumption+) apply (frule_tac x = h and r = ka in ideal_ring_multiple1[of mx], assumption+, frule_tac x = ha and r = k in ideal_ring_multiple[of mx], assumption+) apply (frule_tac a = x and b = y and x = k and y = ka in mul_two_principal_idealsTr, assumption+, erule bexE, frule_tac x = "x \\<^sub>r y" and r = r in ideal_ring_multiple[of mx], assumption+, rotate_tac -2, frule sym, thin_tac "k \\<^sub>r ka = r \\<^sub>r (x \\<^sub>r y)", simp) apply (frule_tac x = "h \\<^sub>r ha \ h \\<^sub>r ka" and y = "k \\<^sub>r ha \ k \\<^sub>r ka" in ideal_pOp_closed[of mx]) apply (rule ideal_pOp_closed, assumption+)+ apply (simp add:maximal_set_def) apply blast apply assumption done lemma (in Ring) ex_mulDisj_maximal:"\mul_closed_set R S; \ \ S; 1\<^sub>r \ S; T = {I. ideal R I \ S \ I = {}}\ \ \mx. maximal_set T mx" apply (cut_tac A="{ I. ideal R I \ S \ I = {}}" in Zorn_Lemma2) prefer 2 apply (simp add:maximal_set_def) apply (rule ballI) apply (case_tac "C = {}") apply (cut_tac zero_ideal, blast) apply (subgoal_tac "C \ chains {I. ideal R I \ I \ carrier R}") apply (frule chains_un, assumption) apply (subgoal_tac "S \ (\ C) = {}") apply (subgoal_tac "\x\C. x \ \ C", blast) apply (rule ballI, rule subsetI, simp add:CollectI) apply blast apply (rule contrapos_pp, simp+) apply (frule_tac A = S and B = "\ C" in nonempty_int) apply (erule exE) apply (simp, erule conjE, erule bexE) apply (simp add:chains_def, erule conjE) apply (frule_tac c = X and A = C and B = "{I. ideal R I \ S \ I = {}}" in subsetD, assumption+, thin_tac "C \ {I. ideal R I \ I \ carrier R}", thin_tac "C \ {I. ideal R I \ S \ I = {}}") apply (simp, blast) apply (simp add:chains_def chain_subset_def, erule conjE) apply (rule subsetI) apply (frule_tac c = x and A = C and B = "{I. ideal R I \ S \ I = {}}" in subsetD, assumption+, thin_tac "C \ {I. ideal R I \ S \ I = {}}", thin_tac "T = {I. ideal R I \ S \ I = {}}") apply (simp, thin_tac "\x\C. \y\C. x \ y \ y \ x", erule conjE) apply (simp add:psubset_eq ideal_subset1) apply (rule contrapos_pp, simp+) apply (rotate_tac -1, frule sym, thin_tac "x = carrier R", thin_tac "carrier R = x") apply (cut_tac ring_one, blast) done lemma (in Ring) ex_mulDisj_prime:"\mul_closed_set R S; \ \ S; 1\<^sub>r \ S\ \ \mx. prime_ideal R mx \ S \ mx = {}" apply (frule ex_mulDisj_maximal[of "S" "{I. ideal R I \ S \ I = {}}"], assumption+, simp, erule exE) apply (frule_tac mx = mx in mulDisj [of "S" "{I. ideal R I \ S \ I = {}}"], assumption+, simp, assumption) apply (simp add:maximal_set_def, (erule conjE)+, blast) done lemma (in Ring) nilradTr1:"\ zeroring R \ nilrad R = \ {p. prime_ideal R p}" apply (rule equalityI) (* nilrad R \ \Collect (prime_ideal R) *) apply (rule subsetI) apply (simp add:nilrad_def CollectI nilpotent_def) apply (erule conjE, erule exE) apply (rule allI, rule impI) apply (frule_tac prime_ideal_ideal) apply (frule sym, thin_tac "x^\<^bsup>R n\<^esup> = \", frule ideal_zero, simp) apply (case_tac "n = 0", simp) apply (frule Zero_ring1[THEN not_sym], simp) apply (rule_tac P = xa and x = x in npow_in_prime,assumption+, blast) apply (rule subsetI) apply (rule contrapos_pp, simp+) apply (frule id_maximal_Exist, erule exE, frule maximal_prime) apply (frule_tac a = I in forall_spec, assumption, frule_tac I = I in prime_ideal_ideal, frule_tac h = x and I = I in ideal_subset, assumption) apply (subgoal_tac "\ \ {s. \n. s = npow R x n} \ 1\<^sub>r \ {s. \n. s = npow R x n}") apply (subgoal_tac "mul_closed_set R {s. \n. s = npow R x n}") apply (erule conjE) apply (frule_tac S = "{s. \n. s = npow R x n}" in ex_mulDisj_prime, assumption+, erule exE, erule conjE) apply (subgoal_tac "x \ {s. \n. s = x^\<^bsup>R n\<^esup>}", blast) apply simp apply (cut_tac t = x in np_1[THEN sym], assumption, blast) apply (thin_tac "\ \ {s. \n. s = x^\<^bsup>R n\<^esup>} \ 1\<^sub>r \ {s. \n. s = x^\<^bsup>R n\<^esup>}", thin_tac "\xa. prime_ideal R xa \ x \ xa") apply (subst mul_closed_set_def) apply (rule conjI) apply (rule subsetI, simp, erule exE) apply (simp add:npClose) apply ((rule ballI)+, simp, (erule exE)+, simp) apply (simp add:npMulDistr, blast) apply (rule conjI) apply simp apply (rule contrapos_pp, simp+, erule exE) apply (frule sym, thin_tac "\ = x^\<^bsup>R n\<^esup>") apply (simp add:nilrad_def nilpotent_def) apply simp apply (cut_tac x1 = x in npow_0[THEN sym, of "R"], blast) done lemma (in Ring) nonilp_residue_nilrad:"\\ zeroring R; x \ carrier R; nilpotent (qring R (nilrad R)) (x \\<^bsub>R\<^esub> (nilrad R))\ \ x \\<^bsub>R\<^esub> (nilrad R) = \\<^bsub>(qring R (nilrad R))\<^esub>" apply (simp add:nilpotent_def) apply (erule exE) apply (cut_tac id_nilrad_ideal) apply (simp add:qring_zero) apply (cut_tac "Ring") apply (simp add:npQring) apply (frule_tac x = x and n = n in npClose) apply (frule_tac I = "nilrad R" and a = "x^\<^bsup>R n\<^esup>" in ar_coset_same3, assumption+) apply (rule_tac I = "nilrad R" and a = x in ar_coset_same4, assumption) apply (thin_tac "x^\<^bsup>R n\<^esup> \\<^bsub>R\<^esub> nilrad R = nilrad R", simp add:nilrad_def nilpotent_def, erule exE) apply (simp add:npMulExp, blast) done lemma (in Ring) ex_contid_maximal:"\ S = {1\<^sub>r}; \ \ S; ideal R I; I \ S = {}; T = {J. ideal R J \ S \ J = {} \ I \ J}\ \ \mx. maximal_set T mx" apply (cut_tac A="{J. ideal R J \ S \ J = {} \ I \ J}" in Zorn_Lemma2) apply (rule ballI) apply (case_tac "C = {}") (** case C = {} **) apply blast (** case C = {} done **) (** existence of sup in C **) apply (subgoal_tac "\C\{J. ideal R J \ S \ J = {} \ I \ J} \ (\x\C. x \ \C)") apply blast apply (rule conjI, simp add:CollectI) apply (subgoal_tac "C \ chains {I. ideal R I \ I \ carrier R}") apply (rule conjI, simp add:chains_un) apply (rule conjI) apply (rule contrapos_pp, simp+, erule bexE) apply (thin_tac " C \ chains {I. ideal R I \ I \ carrier R}") apply (simp add:chains_def, erule conjE) apply (frule_tac c = x and A = C and B = "{J. ideal R J \ 1\<^sub>r \ J \ I \ J}" in subsetD, assumption+, simp, thin_tac "C \ chains {I. ideal R I \ I \ carrier R}") apply (frule_tac A = C in nonempty_ex, erule exE, simp add:chains_def, erule conjE, frule_tac c = x and A = C and B = "{J. ideal R J \ 1\<^sub>r \ J \ I \ J}" in subsetD, assumption+, simp, (erule conjE)+) apply (rule_tac A = I and B = x and C = "\C" in subset_trans, assumption, rule_tac B = x and A = C in Union_upper, assumption+) apply (simp add:chains_def, erule conjE) apply (rule subsetI, simp) apply (frule_tac c = x and A = C and B = "{J. ideal R J \ 1\<^sub>r \ J \ I \ J}" in subsetD, assumption+, simp, (erule conjE)+) apply (subst psubset_eq, simp add:ideal_subset1) apply (rule contrapos_pp, simp+, simp add:ring_one) apply (rule ballI) apply (rule Union_upper, assumption) apply (erule bexE) apply (simp add:maximal_set_def) apply blast done lemma (in Ring) contid_maximal:"\S = {1\<^sub>r}; \ \ S; ideal R I; I \ S = {}; T = {J. ideal R J \ S \ J = {} \ I \ J}; maximal_set T mx\ \ maximal_ideal R mx" apply (simp add:maximal_set_def maximal_ideal_def) apply (erule conjE)+ apply (rule equalityI) (** {J. ideal R J \ mx \ J} \ {mx, carrier R} **) apply (rule subsetI, simp add:CollectI, erule conjE) apply (case_tac "x = mx", simp, simp) apply (subgoal_tac "1\<^sub>r \ x") apply (rule_tac I = x in ideal_inc_one, assumption+) apply (rule contrapos_pp, simp+) apply (drule spec[of _ mx]) apply (simp add:whole_ideal, rule subsetI, rule ideal_subset[of "mx"], assumption+) done lemma (in Ring) ideal_contained_maxid:"\\(zeroring R); ideal R I; 1\<^sub>r \ I\ \ \mx. maximal_ideal R mx \ I \ mx" apply (cut_tac ex_contid_maximal[of "{1\<^sub>r}" "I" "{J. ideal R J \ {1\<^sub>r} \ J = {} \ I \ J}"]) apply (erule exE, cut_tac mx = mx in contid_maximal[of "{1\<^sub>r}" "I" "{J. ideal R J \ {1\<^sub>r} \ J = {} \ I \ J}"]) apply simp apply (frule Zero_ring1, simp, assumption, simp, simp, simp, simp add:maximal_set_def, (erule conjE)+, blast, simp, frule Zero_ring1, simp) apply (assumption, simp, simp) done lemma (in Ring) nonunit_principal_id:"\a \ carrier R; \ (Unit R a)\ \ (R \\<^sub>p a) \ (carrier R)" apply (rule contrapos_pp, simp+) apply (frule sym, thin_tac "R \\<^sub>p a = carrier R") apply (cut_tac ring_one) apply (frule eq_set_inc[of "1\<^sub>r" "carrier R" "R \\<^sub>p a"], assumption, thin_tac "carrier R = R \\<^sub>p a", thin_tac "1\<^sub>r \ carrier R") apply (simp add:Rxa_def, erule bexE, simp add:ring_tOp_commute[of _ "a"], frule sym, thin_tac "1\<^sub>r = a \\<^sub>r r") apply (simp add:Unit_def) done lemma (in Ring) nonunit_contained_maxid:"\\(zeroring R); a \ carrier R; \ Unit R a \ \ \mx. maximal_ideal R mx \ a \ mx" apply (frule principal_ideal[of "a"], frule ideal_contained_maxid[of "R \\<^sub>p a"], assumption) apply (rule contrapos_pp, simp+, frule ideal_inc_one[of "R \\<^sub>p a"], assumption, simp add:nonunit_principal_id) apply (erule exE, erule conjE) apply (frule a_in_principal[of "a"]) apply (frule_tac B = mx in subsetD[of "R \\<^sub>p a" _ "a"], assumption, blast) done definition local_ring :: "_ \ bool" where "local_ring R == Ring R \ \ zeroring R \ card {mx. maximal_ideal R mx} = 1" lemma (in Ring) local_ring_diff:"\\ zeroring R; ideal R mx; mx \ carrier R; \a\ (carrier R - mx). Unit R a \ \ local_ring R \ maximal_ideal R mx" apply (subgoal_tac "{mx} = {m. maximal_ideal R m}") apply (cut_tac singletonI[of "mx"], simp) apply (frule sym, thin_tac "{mx} = {m. maximal_ideal R m}") apply (simp add:local_ring_def, simp add:Ring) apply (rule equalityI) apply (rule subsetI, simp) apply (simp add:maximal_ideal_def) apply (simp add:ideal_inc_one1[of "mx", THEN sym]) apply (thin_tac "x = mx", simp) apply (rule equalityI) apply (rule subsetI, simp, erule conjE) apply (case_tac "x \ mx") apply (frule_tac A = x and B = mx in sets_not_eq, assumption) apply (erule bexE) apply (frule_tac h = a and I = x in ideal_subset, assumption+) apply (frule_tac x = a in bspec, simp) apply (frule_tac I = x and a = a in ideal_inc_unit1, assumption+, simp) apply simp apply (rule subsetI, simp) apply (erule disjE) apply simp apply (simp add:whole_ideal ideal_subset1) apply (rule subsetI) apply simp apply (subgoal_tac "x \ mx", thin_tac "\a\carrier R - mx. Unit R a", simp add:maximal_ideal_def, (erule conjE)+) apply (subgoal_tac "mx \ {J. ideal R J \ x \ J}", simp) apply (thin_tac "{J. ideal R J \ x \ J} = {x, carrier R}") apply simp apply (rule contrapos_pp, simp+) apply (simp add:subset_eq, erule bexE) apply (frule_tac mx = x in maximal_ideal_ideal, frule_tac x = xa in bspec, thin_tac "\a\carrier R - mx. Unit R a", simp, simp add:ideal_subset) apply (frule_tac I = x and a = xa in ideal_inc_unit, assumption+, simp add:maximal_ideal_def) done lemma (in Ring) localring_unit:"\\ zeroring R; maximal_ideal R mx; \x. x \ mx \ Unit R (x \ 1\<^sub>r) \ \ local_ring R" apply (frule maximal_ideal_ideal[of "mx"]) apply (frule local_ring_diff[of "mx"], assumption) apply (simp add:maximal_ideal_def, erule conjE) apply (simp add:ideal_inc_one1[THEN sym, of "mx"]) apply (rule ballI, simp, erule conjE) apply (frule_tac x = a in maximal_prime_Tr0[of "mx"], assumption+) apply (frule sym, thin_tac "mx \ R \\<^sub>p a = carrier R", cut_tac ring_one, frule_tac a = "1\<^sub>r" and A = "carrier R" and B = "mx \ R \\<^sub>p a" in eq_set_inc, assumption+, thin_tac "carrier R = mx \ R \\<^sub>p a") apply (frule_tac a = a in principal_ideal, frule ideal_subset1[of mx], frule_tac I = "R \\<^sub>p a" in ideal_subset1) apply (cut_tac ring_is_ag, simp add:aGroup.set_sum, (erule bexE)+) apply (simp add:Rxa_def, erule bexE, simp) apply (frule sym, thin_tac "1\<^sub>r = h \ r \\<^sub>r a", frule_tac x = r and y = a in ring_tOp_closed, assumption+, frule_tac h = h in ideal_subset[of "mx"], assumption+) apply (frule_tac I = mx and x = h in ideal_inv1_closed, assumption) apply (frule_tac a = "-\<^sub>a h" in forall_spec, assumption, thin_tac "\x. x \ mx \ Unit R (x \ (h \ r \\<^sub>r a))", thin_tac "h \ r \\<^sub>r a = 1\<^sub>r") apply (frule_tac h = "-\<^sub>a h" in ideal_subset[of "mx"], assumption, frule_tac x1 = "-\<^sub>a h" and y1 = h and z1 = "r \\<^sub>r a" in aGroup.ag_pOp_assoc[THEN sym], assumption+, simp add:aGroup.ag_l_inv1 aGroup.ag_l_zero, thin_tac "k = r \\<^sub>r a", thin_tac "h \ r \\<^sub>r a \ carrier R", thin_tac "h \ carrier R", thin_tac "-\<^sub>a h \ mx", thin_tac "-\<^sub>a h \ (h \ r \\<^sub>r a) = r \\<^sub>r a") apply (simp add:ring_tOp_commute, simp add:Unit_def, erule bexE, simp add:ring_tOp_assoc, frule_tac x = r and y = b in ring_tOp_closed, assumption+, blast) apply simp done definition J_rad ::"_ \ 'a set" where "J_rad R = (if (zeroring R) then (carrier R) else \ {mx. maximal_ideal R mx})" (** if zeroring R then \ {mx. maximal_ideal R mx} is UNIV, hence we restrict UNIV to carrier R **) lemma (in Ring) zeroring_J_rad_empty:"zeroring R \ J_rad R = carrier R" by (simp add:J_rad_def) lemma (in Ring) J_rad_mem:"x \ J_rad R \ x \ carrier R" apply (simp add:J_rad_def) apply (case_tac "zeroring R", simp) apply simp apply (frule id_maximal_Exist, erule exE) apply (frule_tac a = I in forall_spec, assumption, thin_tac "\xa. maximal_ideal R xa \ x \ xa") apply (frule maximal_ideal_ideal, simp add:ideal_subset) done lemma (in Ring) J_rad_unit:"\\ zeroring R; x \ J_rad R\ \ \y. (y\ carrier R \ Unit R (1\<^sub>r \ (-\<^sub>a x) \\<^sub>r y))" apply (cut_tac ring_is_ag, rule allI, rule impI, rule contrapos_pp, simp+) apply (frule J_rad_mem[of "x"], frule_tac x = x and y = y in ring_tOp_closed, assumption, frule_tac x = "x \\<^sub>r y" in aGroup.ag_mOp_closed, assumption+) apply (cut_tac ring_one, frule_tac x = "1\<^sub>r" and y = "-\<^sub>a (x \\<^sub>r y)" in aGroup.ag_pOp_closed, assumption+) apply (frule_tac a = "1\<^sub>r \ -\<^sub>a (x \\<^sub>r y)" in nonunit_contained_maxid, assumption+, simp add:ring_inv1_1) apply (erule exE, erule conjE) apply (simp add:J_rad_def, frule_tac a = mx in forall_spec, assumption, thin_tac "\xa. maximal_ideal R xa \ x \ xa", frule_tac mx = mx in maximal_ideal_ideal, frule_tac I = mx and x = x and r = y in ideal_ring_multiple1, assumption+) apply (frule_tac I = mx and x = "x \\<^sub>r y" in ideal_inv1_closed, assumption+) apply (frule_tac I = mx and a = "1\<^sub>r" and b = "-\<^sub>a (x \\<^sub>r y)" in ideal_ele_sumTr2, assumption+) apply (simp add:maximal_ideal_def) done end diff --git a/thys/List-Infinite/CommonArith/Util_Div.thy b/thys/List-Infinite/CommonArith/Util_Div.thy --- a/thys/List-Infinite/CommonArith/Util_Div.thy +++ b/thys/List-Infinite/CommonArith/Util_Div.thy @@ -1,1239 +1,1239 @@ (* Title: Util_Div.thy Date: Oct 2006 Author: David Trachtenherz *) section \Results for division and modulo operators on integers\ theory Util_Div imports Util_Nat begin subsection \Additional (in-)equalities with \div\ and \mod\\ corollary Suc_mod_le_divisor: "0 < m \ Suc (n mod m) \ m" by (rule Suc_leI, rule mod_less_divisor) lemma mod_less_dividend: "\ 0 < m; m \ n \ \ n mod m < (n::nat)" by (rule less_le_trans[OF mod_less_divisor]) (*lemma mod_le_dividend: "n mod m \ (n::nat)"*) lemmas mod_le_dividend = mod_less_eq_dividend lemma diff_mod_le: "(t - r) mod m \ (t::nat)" by (rule le_trans[OF mod_le_dividend, OF diff_le_self]) (*corollary div_mult_cancel: "m div n * n = m - m mod (n::nat)"*) lemmas div_mult_cancel = minus_mod_eq_div_mult [symmetric] lemma mod_0_div_mult_cancel: "(n mod (m::nat) = 0) = (n div m * m = n)" apply (insert eq_diff_left_iff[OF mod_le_dividend le0, of n m]) apply (simp add: mult.commute minus_mod_eq_mult_div [symmetric]) done lemma div_mult_le: "(n::nat) div m * m \ n" by (simp add: mult.commute minus_mod_eq_mult_div [symmetric]) lemma less_div_Suc_mult: "0 < (m::nat) \ n < Suc (n div m) * m" apply (simp add: mult.commute minus_mod_eq_mult_div [symmetric]) apply (rule less_add_diff) by (rule mod_less_divisor) lemma nat_ge2_conv: "((2::nat) \ n) = (n \ 0 \ n \ 1)" by fastforce lemma Suc0_mod: "m \ Suc 0 \ Suc 0 mod m = Suc 0" by (case_tac m, simp_all) corollary Suc0_mod_subst: " \ m \ Suc 0; P (Suc 0) \ \ P (Suc 0 mod m)" by (blast intro: subst[OF Suc0_mod[symmetric]]) corollary Suc0_mod_cong: " m \ Suc 0 \ f (Suc 0 mod m) = f (Suc 0)" by (blast intro: arg_cong[OF Suc0_mod]) subsection \Additional results for addition and subtraction with \mod\\ lemma mod_Suc_conv: " ((Suc a) mod m = (Suc b) mod m) = (a mod m = b mod m)" by (simp add: mod_Suc) lemma mod_Suc': " 0 < n \ Suc m mod n = (if m mod n < n - Suc 0 then Suc (m mod n) else 0)" apply (simp add: mod_Suc) apply (intro conjI impI) apply simp apply (insert le_neq_trans[OF mod_less_divisor[THEN Suc_leI, of n m]], simp) done lemma mod_add:" ((a + k) mod m = (b + k) mod m) = ((a::nat) mod m = b mod m)" by (induct "k", simp_all add: mod_Suc_conv) corollary mod_sub_add: " k \ (a::nat) \ ((a - k) mod m = b mod m) = (a mod m = (b + k) mod m)" by (simp add: mod_add[where m=m and a="a-k" and b=b and k=k, symmetric]) lemma mod_sub_eq_mod_0_conv: " a + b \ (n::nat) \ ((n - a) mod m = b mod m) = ((n - (a + b)) mod m = 0)" by (insert mod_add[of "n-(a+b)" b m 0], simp) lemma mod_sub_eq_mod_swap: " \ a \ (n::nat); b \ n \ \ ((n - a) mod m = b mod m) = ((n - b) mod m = a mod m)" by (simp add: mod_sub_add add.commute) lemma le_mod_greater_imp_div_less: " \ a \ (b::nat); a mod m > b mod m \ \ a div m < b div m" apply (rule ccontr, simp add: linorder_not_less) apply (drule mult_le_mono1[of "b div m" _ m]) apply (drule add_less_le_mono[of "b mod m" "a mod m" "b div m * m" "a div m * m"]) apply simp_all done lemma less_mod_ge_imp_div_less: "\ a < (b::nat); a mod m \ b mod m \ \ a div m < b div m" apply (case_tac "m = 0", simp) apply (rule mult_less_cancel1[of m, THEN iffD1, THEN conjunct2]) apply (simp add: minus_mod_eq_mult_div [symmetric]) apply (rule order_less_le_trans[of _ "b - a mod m"]) apply (rule diff_less_mono) apply simp+ done corollary less_mod_0_imp_div_less: "\ a < (b::nat); b mod m = 0 \ \ a div m < b div m" by (simp add: less_mod_ge_imp_div_less) lemma mod_diff_right_eq: " (a::nat) \ b \ (b - a) mod m = (b - a mod m) mod m" proof - assume a_as:"a \ b" have "(b - a) mod m = (b - a + a div m * m) mod m" by simp also have "\ = (b + a div m * m - a) mod m" using a_as by simp also have "\ = (b + a div m * m - (a div m * m + a mod m)) mod m" by simp also have "\ = (b + a div m * m - a div m * m - a mod m) mod m" by (simp only: diff_diff_left[symmetric]) also have "\ = (b - a mod m) mod m" by simp finally show ?thesis . qed corollary mod_eq_imp_diff_mod_eq: " \ x mod m = y mod m; x \ (t::nat); y \ t \ \ (t - x) mod m = (t - y) mod m" by (simp only: mod_diff_right_eq) lemma mod_eq_imp_diff_mod_eq2: " \ x mod m = y mod m; (t::nat) \ x; t \ y \ \ (x - t) mod m = (y - t) mod m" apply (case_tac "m = 0", simp+) apply (subst mod_mult_self2[of "x - t" m t, symmetric]) apply (subst mod_mult_self2[of "y - t" m t, symmetric]) apply (simp only: add_diff_assoc2 diff_add_assoc gr0_imp_self_le_mult2) apply (simp only: mod_add) done lemma divisor_add_diff_mod_if: " (m + b mod m - a mod m) mod (m::nat)= ( if a mod m \ b mod m then (b mod m - a mod m) else (m + b mod m - a mod m))" apply (case_tac "m = 0", simp) apply clarsimp apply (subst diff_add_assoc, assumption) apply (simp only: mod_add_self1) apply (rule mod_less) apply (simp add: less_imp_diff_less) done corollary divisor_add_diff_mod_eq1: " a mod m \ b mod m \ (m + b mod m - a mod m) mod (m::nat) = b mod m - a mod m" by (simp add: divisor_add_diff_mod_if) corollary divisor_add_diff_mod_eq2: " b mod m < a mod m \ (m + b mod m - a mod m) mod (m::nat) = m + b mod m - a mod m" by (simp add: divisor_add_diff_mod_if) lemma mod_add_mod_if: " (a mod m + b mod m) mod (m::nat)= ( if a mod m + b mod m < m then a mod m + b mod m else a mod m + b mod m - m)" apply (case_tac "m = 0", simp_all) apply (clarsimp simp: linorder_not_less) apply (simp add: mod_if[of "a mod m + b mod m"]) apply (rule mod_less) apply (rule diff_less_conv[THEN iffD2], assumption) apply (simp add: add_less_mono) done corollary mod_add_mod_eq1: " a mod m + b mod m < m \ (a mod m + b mod m) mod (m::nat) = a mod m + b mod m" by (simp add: mod_add_mod_if) corollary mod_add_mod_eq2: " m \ a mod m + b mod m\ (a mod m + b mod m) mod (m::nat) = a mod m + b mod m - m" by (simp add: mod_add_mod_if) lemma mod_add1_eq_if: " (a + b) mod (m::nat) = ( if (a mod m + b mod m < m) then a mod m + b mod m else a mod m + b mod m - m)" by (simp add: mod_add_eq[symmetric, of a b] mod_add_mod_if) lemma mod_add_eq_mod_conv: "0 < (m::nat) \ ((x + a) mod m = b mod m ) = (x mod m = (m + b mod m - a mod m) mod m)" apply (simp only: mod_add_eq[symmetric, of x a]) apply (rule iffI) apply (drule sym) apply (simp add: mod_add_mod_if) apply (simp add: mod_add_left_eq le_add_diff_inverse2[OF trans_le_add1[OF mod_le_divisor]]) done lemma mod_diff1_eq: " (a::nat) \ b \ (b - a) mod m = (m + b mod m - a mod m) mod m" apply (case_tac "m = 0", simp) apply simp proof - assume a_as:"a \ b" and m_as: "0 < m" have a_mod_le_b_s: "a mod m \ b" by (rule le_trans[of _ a], simp only: mod_le_dividend, simp only: a_as) have "(b - a) mod m = (b - a mod m) mod m" using a_as by (simp only: mod_diff_right_eq) also have "\ = (b - a mod m + m) mod m" by simp also have "\ = (b + m - a mod m) mod m" using a_mod_le_b_s by simp also have "\ = (b div m * m + b mod m + m - a mod m) mod m" by simp also have "\ = (b div m * m + (b mod m + m - a mod m)) mod m" by (simp add: diff_add_assoc[OF mod_le_divisor, OF m_as]) also have "\ = ((b mod m + m - a mod m) + b div m * m) mod m" by simp also have "\ = (b mod m + m - a mod m) mod m" by simp also have "\ = (m + b mod m - a mod m) mod m" by (simp only: add.commute) finally show ?thesis . qed corollary mod_diff1_eq_if: " (a::nat) \ b \ (b - a) mod m = ( if a mod m \ b mod m then b mod m - a mod m else m + b mod m - a mod m)" by (simp only: mod_diff1_eq divisor_add_diff_mod_if) corollary mod_diff1_eq1: " \ (a::nat) \ b; a mod m \ b mod m \ \ (b - a) mod m = b mod m - a mod m" by (simp add: mod_diff1_eq_if) corollary mod_diff1_eq2: " \ (a::nat) \ b; b mod m < a mod m\ \ (b - a) mod m = m + b mod m - a mod m" by (simp add: mod_diff1_eq_if) subsubsection \Divisor subtraction with \div\ and \mod\\ lemma mod_diff_self1: " 0 < (n::nat) \ (m - n) mod m = m - n" by (case_tac "m = 0", simp_all) lemma mod_diff_self2: " m \ (n::nat) \ (n - m) mod m = n mod m" by (simp add: mod_diff_right_eq) lemma mod_diff_mult_self1: " k * m \ (n::nat) \ (n - k * m) mod m = n mod m" by (simp add: mod_diff_right_eq) lemma mod_diff_mult_self2: " m * k \ (n::nat) \ (n - m * k) mod m = n mod m" by (simp only: mult.commute[of m k] mod_diff_mult_self1) lemma div_diff_self1: "0 < (n::nat) \ (m - n) div m = 0" by (case_tac "m = 0", simp_all) lemma div_diff_self2: "(n - m) div m = n div m - Suc 0" apply (case_tac "m = 0", simp) apply (case_tac "n < m", simp) apply (case_tac "n = m", simp) apply (simp add: div_if) done lemma div_diff_mult_self1: " (n - k * m) div m = n div m - (k::nat)" apply (case_tac "m = 0", simp) apply (case_tac "n < k * m") apply simp apply (drule div_le_mono[OF less_imp_le, of n _ m]) apply simp apply (simp add: linorder_not_less) apply (rule iffD1[OF mult_cancel1_gr0[where k=m]], assumption) apply (subst diff_mult_distrib2) apply (simp only: minus_mod_eq_mult_div [symmetric]) apply (simp only: diff_commute[of _ "k*m"]) apply (simp only: mult.commute[of m]) apply (simp only: mod_diff_mult_self1) done lemma div_diff_mult_self2: " (n - m * k) div m = n div m - (k::nat)" by (simp only: mult.commute div_diff_mult_self1) subsubsection \Modulo equality and modulo of difference\ lemma mod_eq_imp_diff_mod_0:" (a::nat) mod m = b mod m \ (b - a) mod m = 0" (is "?P \ ?Q") proof - assume as1: ?P have "b - a = b div m * m + b mod m - (a div m * m + a mod m)" by simp also have "\ = b div m * m + b mod m - (a mod m + a div m * m)" by simp also have "\ = b div m * m + b mod m - a mod m - a div m * m" by simp also have "\ = b div m * m + b mod m - b mod m - a div m * m" using as1 by simp also have "\ = b div m * m - a div m * m" by (simp only: diff_add_inverse2) also have "\ = (b div m - a div m) * m" by (simp only: diff_mult_distrib) finally have "b - a = (b div m - a div m) * m" . hence "(b - a) mod m = (b div m - a div m) * m mod m" by (rule arg_cong) thus ?thesis by (simp only: mod_mult_self2_is_0) qed corollary mod_eq_imp_diff_dvd: " (a::nat) mod m = b mod m \ m dvd b - a" by (rule dvd_eq_mod_eq_0[THEN iffD2, OF mod_eq_imp_diff_mod_0]) lemma mod_neq_imp_diff_mod_neq0:" \ (a::nat) mod m \ b mod m; a \ b \ \ 0 < (b - a) mod m" apply (case_tac "m = 0", simp) apply (drule le_imp_less_or_eq, erule disjE) prefer 2 apply simp apply (drule neq_iff[THEN iffD1], erule disjE) apply (simp add: mod_diff1_eq1) apply (simp add: mod_diff1_eq2[OF less_imp_le] trans_less_add1[OF mod_less_divisor]) done corollary mod_neq_imp_diff_not_dvd:" \ (a::nat) mod m \ b mod m; a \ b \ \ \ m dvd b - a" by (simp add: dvd_eq_mod_eq_0 mod_neq_imp_diff_mod_neq0) lemma diff_mod_0_imp_mod_eq:" \ (b - a) mod m = 0; a \ b \ \ (a::nat) mod m = b mod m" apply (rule ccontr) apply (drule mod_neq_imp_diff_mod_neq0) apply simp_all done corollary diff_dvd_imp_mod_eq:" \ m dvd b - a; a \ b \ \ (a::nat) mod m = b mod m" by (rule dvd_eq_mod_eq_0[THEN iffD1, THEN diff_mod_0_imp_mod_eq]) lemma mod_eq_diff_mod_0_conv: " a \ (b::nat) \ (a mod m = b mod m) = ((b - a) mod m = 0)" apply (rule iffI) apply (rule mod_eq_imp_diff_mod_0, assumption) apply (rule diff_mod_0_imp_mod_eq, assumption+) done corollary mod_eq_diff_dvd_conv: " a \ (b::nat) \ (a mod m = b mod m) = (m dvd b - a)" by (rule dvd_eq_mod_eq_0[symmetric, THEN subst], rule mod_eq_diff_mod_0_conv) subsection \Some additional lemmata about integer \div\ and \mod\\ lemma zmod_eq_imp_diff_mod_0: "a mod m = b mod m \ (b - a) mod m = 0" for a b m :: int by (simp add: mod_diff_cong) (*lemma int_mod_distrib: "int (n mod m) = int n mod int m"*) lemmas int_mod_distrib = zmod_int lemma zdiff_mod_0_imp_mod_eq__pos:" \ (b - a) mod m = 0; 0 < (m::int) \ \ a mod m = b mod m" (is "\ ?P; ?Pm \ \ ?Q") proof - assume as1: ?P and as2: "0 < m" obtain r1 where a_r1:"r1 = a mod m" by blast obtain r2 where b_r2:"r2 = b mod m" by blast obtain q1 where a_q1: "q1 = a div m" by blast obtain q2 where b_q2: "q2 = b div m" by blast have a_r1_q1: "a = m * q1 + r1" using a_r1 a_q1 by simp have b_r2_q2: "b = m * q2 + r2" using b_r2 b_q2 by simp have "b - a = m * q2 + r2 - (m * q1 + r1)" using a_r1_q1 b_r2_q2 by simp also have "\ = m * q2 + r2 - m * q1 - r1" by simp also have "\ = m * q2 - m * q1 + r2 - r1" by simp finally have "b - a = m * (q2 - q1) + (r2 - r1)" by (simp add: right_diff_distrib) hence "(b - a) mod m = (r2 - r1) mod m" by (simp add: mod_add_eq) hence r2_r1_mod_m_0:"(r2 - r1) mod m = 0" (is "?R1") by (simp only: as1) have "r1 = r2" proof (rule notI[of "r1 \ r2", simplified]) assume as1': "r1 \ r2" have diff_le_s: "\a b (m::int). \ 0 \ a; b < m \ \ b - a < m" by simp - have s_r1:"0 \ r1 \ r1 < m" and s_r2:"0 \ r2 \ r2 < m" - by (simp add: as2 a_r1 b_r2 pos_mod_conj)+ + from as2 a_r1 b_r2 have s_r1:"0 \ r1 \ r1 < m" and s_r2:"0 \ r2 \ r2 < m" + by simp_all have mr2r1:"-m < r2 - r1" and r2r1m:"r2 - r1 < m" by (simp add: minus_less_iff[of m] s_r1 s_r2 diff_le_s)+ have "0 \ r2 - r1 \ (r2 - r1) mod m = (r2 - r1)" using r2r1m by (blast intro: mod_pos_pos_trivial) hence s1_pos: "0 \ r2 - r1 \ r2 - r1 = 0" using r2_r1_mod_m_0 by simp have "(r2-r1) mod -m = 0" by (simp add: zmod_zminus2_eq_if[of "r2-r1" m, simplified] r2_r1_mod_m_0) moreover have "r2 - r1 \ 0 \ (r2 - r1) mod -m = r2 - r1" using mr2r1 by (simp add: mod_neg_neg_trivial) ultimately have s1_neg:"r2 - r1 \ 0 \ r2 - r1 = 0" by simp have "r2 - r1 = 0" using s1_pos s1_neg linorder_linear by blast hence "r1 = r2" by simp thus False using as1' by blast qed thus ?thesis using a_r1 b_r2 by blast qed lemma zmod_zminus_eq_conv_pos: " 0 < (m::int) \ (a mod - m = b mod - m) = (a mod m = b mod m)" apply (simp only: mod_minus_right neg_equal_iff_equal) apply (simp only: zmod_zminus1_eq_if) apply (split if_split)+ apply (safe, simp_all) apply (insert pos_mod_bound[of m a] pos_mod_bound[of m b], simp_all) done lemma zmod_zminus_eq_conv: " ((a::int) mod - m = b mod - m) = (a mod m = b mod m)" apply (insert linorder_less_linear[of 0 m], elim disjE) apply (blast dest: zmod_zminus_eq_conv_pos) apply simp apply (simp add: zmod_zminus_eq_conv_pos[of "-m", symmetric]) done lemma zdiff_mod_0_imp_mod_eq:" (b - a) mod m = 0 \ (a::int) mod m = b mod m" by (metis dvd_eq_mod_eq_0 mod_eq_dvd_iff) lemma zmod_eq_diff_mod_0_conv: " ((a::int) mod m = b mod m) = ((b - a) mod m = 0)" apply (rule iffI) apply (rule zmod_eq_imp_diff_mod_0, assumption) apply (rule zdiff_mod_0_imp_mod_eq, assumption) done lemma "\(\(a::int) b m. (b - a) mod m = 0 \ a mod m \ b mod m)" by (simp add: zmod_eq_diff_mod_0_conv) lemma "\(a::nat) b m. (b - a) mod m = 0 \ a mod m \ b mod m" apply (rule_tac x=1 in exI) apply (rule_tac x=0 in exI) apply (rule_tac x=2 in exI) apply simp done lemma zmult_div_leq_mono:" \ (0::int) \ x; a \ b; 0 < d \ \ x * a div d \ x * b div d" by (metis mult_right_mono zdiv_mono1 mult.commute) lemma zmult_div_leq_mono_neg:" \ x \ (0::int); a \ b; 0 < d \ \ x * b div d \ x * a div d" by (metis mult_left_mono_neg zdiv_mono1) lemma zmult_div_pos_le:" \ (0::int) \ a; 0 \ b; b \ c \ \ a * b div c \ a" apply (case_tac "b = 0", simp) apply (subgoal_tac "b * a \ c * a") prefer 2 apply (simp only: mult_right_mono) apply (simp only: mult.commute) apply (subgoal_tac "a * b div c \ a * c div c") prefer 2 apply (simp only: zdiv_mono1) apply simp done lemma zmult_div_neg_le:" \ a \ (0::int); 0 < c; c \ b \ \ a * b div c \ a" apply (subgoal_tac "b * a \ c * a") prefer 2 apply (simp only: mult_right_mono_neg) apply (simp only: mult.commute) apply (subgoal_tac "a * b div c \ a * c div c") prefer 2 apply (simp only: zdiv_mono1) apply simp done lemma zmult_div_ge_0:"\ (0::int) \ x; 0 \ a; 0 < c \ \ 0 \ a * x div c" by (metis pos_imp_zdiv_nonneg_iff split_mult_pos_le) corollary zmult_div_plus_ge_0: " \ (0::int) \ x; 0 \ a; 0 \ b; 0 < c\ \ 0 \ a * x div c + b" by (insert zmult_div_ge_0[of x a c], simp) lemma zmult_div_abs_ge: " \ (0::int) \ b; b \ b'; 0 \ a; 0 < c\ \ \a * b div c\ \ \a * b' div c\" apply (insert zmult_div_ge_0[of b a c] zmult_div_ge_0[of "b'" a c], simp) by (metis zmult_div_leq_mono) lemma zmult_div_plus_abs_ge: " \ (0::int) \ b; b \ b'; 0 \ a; 0 < c \ \ \a * b div c + a\ \ \a * b' div c + a\" apply (insert zmult_div_plus_ge_0[of b a a c] zmult_div_plus_ge_0[of "b'" a a c], simp) by (metis zmult_div_leq_mono) subsection \Some further (in-)equality results for \div\ and \mod\\ lemma less_mod_eq_imp_add_divisor_le: " \ (x::nat) < y; x mod m = y mod m \ \ x + m \ y" apply (case_tac "m = 0") apply simp apply (rule contrapos_pp[of "x mod m = y mod m"]) apply blast apply (rule ccontr, simp only: not_not, clarify) proof - assume m_greater_0: "0 < m" assume x_less_y:"x < y" hence y_x_greater_0:"0 < y - x" by simp assume "x mod m = y mod m" hence y_x_mod_m: "(y - x) mod m = 0" by (simp only: mod_eq_imp_diff_mod_0) assume "\ x + m \ y" hence "y < x + m" by simp hence "y - x < x + m - x" by (simp add: diff_add_inverse diff_less_conv m_greater_0) hence y_x_less_m: "y - x < m" by simp have "(y - x) mod m = y - x" using y_x_less_m by simp hence "y - x = 0" using y_x_mod_m by simp thus False using y_x_greater_0 by simp qed lemma less_div_imp_mult_add_divisor_le: " (x::nat) < n div m \ x * m + m \ n" apply (case_tac "m = 0", simp) apply (case_tac "n < m", simp) apply (simp add: linorder_not_less) apply (subgoal_tac "m \ n - n mod m") prefer 2 apply (drule div_le_mono[of m _ m]) apply (simp only: div_self) apply (drule mult_le_mono2[of 1 _ m]) apply (simp only: mult_1_right minus_mod_eq_mult_div [symmetric]) apply (drule less_imp_le_pred[of x]) apply (drule mult_le_mono2[of x _ m]) apply (simp add: diff_mult_distrib2 minus_mod_eq_mult_div [symmetric] del: diff_diff_left) apply (simp only: le_diff_conv2[of m]) apply (drule le_diff_imp_le[of "m * x + m"]) apply (simp only: mult.commute[of _ m]) done lemma mod_add_eq_imp_mod_0: " ((n + k) mod (m::nat) = n mod m) = (k mod m = 0)" by (metis add_eq_if mod_add mod_add_self1 mod_self add.commute) lemma between_imp_mod_between: " \ b < (m::nat); m * k + a \ n; n \ m * k + b \ \ a \ n mod m \ n mod m \ b" apply (case_tac "m = 0", simp_all) apply (frule gr_implies_gr0) apply (subgoal_tac "k = n div m") prefer 2 apply (rule sym, rule div_nat_eqI) apply simp apply simp apply clarify apply (rule conjI) apply (rule add_le_imp_le_left[where c="m * (n div m)"], simp)+ done corollary between_imp_mod_le: " \ b < (m::nat); m * k \ n; n \ m * k + b \ \ n mod m \ b" by (insert between_imp_mod_between[of b m k 0 n], simp) corollary between_imp_mod_gr0: " \ (m::nat) * k < n; n < m * k + m \ \ 0 < n mod m" apply (case_tac "m = 0", simp_all) apply (rule Suc_le_lessD) apply (rule between_imp_mod_between[THEN conjunct1, of "m - Suc 0" m k "Suc 0" n]) apply simp_all done corollary le_less_div_conv: " 0 < m \ (k * m \ n \ n < Suc k * m) = (n div m = k)" by (auto simp add: ac_simps intro: div_nat_eqI dividend_less_times_div) lemma le_less_imp_div: " \ k * m \ n; n < Suc k * m \ \ n div m = k" by (auto simp add: ac_simps intro: div_nat_eqI) lemma div_imp_le_less: " \ n div m = k; 0 < m \ \ k * m \ n \ n < Suc k * m" by (auto simp add: ac_simps intro: dividend_less_times_div) lemma div_le_mod_le_imp_le: " \ (a::nat) div m \ b div m; a mod m \ b mod m \ \ a \ b" apply (rule subst[OF mult_div_mod_eq[of m a]]) apply (rule subst[OF mult_div_mod_eq[of m b]]) apply (rule add_le_mono) apply (rule mult_le_mono2) apply assumption+ done lemma le_mod_add_eq_imp_add_mod_le: " \ a \ b; (a + k) mod m = (b::nat) mod m \ \ a + k mod m \ b" by (metis add_le_mono2 diff_add_inverse le_add1 le_add_diff_inverse mod_diff1_eq mod_less_eq_dividend) corollary mult_divisor_le_mod_ge_imp_ge: " \ (m::nat) * k \ n; r \ n mod m \ \ m * k + r \ n" apply (insert le_mod_add_eq_imp_add_mod_le[of "m * k" n "n mod m" m]) apply (simp add: add.commute[of "m * k"]) done subsection \Additional multiplication results for \mod\ and \div\\ lemma mod_0_imp_mod_mult_right_0: " n mod m = (0::nat) \ n * k mod m = 0" by fastforce lemma mod_0_imp_mod_mult_left_0: " n mod m = (0::nat) \ k * n mod m = 0" by fastforce lemma mod_0_imp_div_mult_left_eq: " n mod m = (0::nat) \ k * n div m = k * (n div m)" by fastforce lemma mod_0_imp_div_mult_right_eq: " n mod m = (0::nat) \ n * k div m = k * (n div m)" by fastforce lemma mod_0_imp_mod_factor_0_left: " n mod (m * m') = (0::nat) \ n mod m = 0" by fastforce lemma mod_0_imp_mod_factor_0_right: " n mod (m * m') = (0::nat) \ n mod m' = 0" by fastforce subsection \Some factor distribution facts for \mod\\ lemma mod_eq_mult_distrib: " (a::nat) mod m = b mod m \ a * k mod (m * k) = b * k mod (m * k)" by simp lemma mod_mult_eq_imp_mod_eq: " (a::nat) mod (m * k) = b mod (m * k) \ a mod m = b mod m" apply (simp only: mod_mult2_eq) apply (drule_tac arg_cong[where f="\x. x mod m"]) apply (simp add: add.commute) done corollary mod_eq_mod_0_imp_mod_eq: " \ (a::nat) mod m' = b mod m'; m' mod m = 0 \ \ a mod m = b mod m" using mod_mod_cancel [of m m' a] mod_mod_cancel [of m m' b] by auto lemma mod_factor_imp_mod_0: " \(x::nat) mod (m * k) = y * k mod (m * k)\ \ x mod k = 0" (is "\ ?P1 \ \ ?Q") proof - assume as1: ?P1 have "y * k mod (m * k) = y mod m * k" by simp hence "x mod (m * k) = y mod m * k" using as1 by simp hence "y mod m * k = k * (x div k mod m) + x mod k" (is "?l1 = ?r1") by (simp only: ac_simps mod_mult2_eq) hence "(y mod m * k) mod k = ?r1 mod k" by simp hence "0 = ?r1 mod k" by simp thus "x mod k = 0" by (simp add: mod_add_eq) qed corollary mod_factor_div: " \(x::nat) mod (m * k) = y * k mod (m * k)\ \ x div k * k = x" by (blast intro: mod_factor_imp_mod_0[THEN mod_0_div_mult_cancel[THEN iffD1]]) lemma mod_factor_div_mod:" \ (x::nat) mod (m * k) = y * k mod (m * k); 0 < k \ \ x div k mod m = y mod m" (is "\ ?P1; ?P2 \ \ ?L = ?R") proof - assume as1: ?P1 assume as2: ?P2 have x_mod_k_0: "x mod k = 0" using as1 by (blast intro: mod_factor_imp_mod_0) have "?L * k + x mod k = x mod (k * m)" by (simp only: mod_mult2_eq mult.commute[of _ k]) hence "?L * k = x mod (k * m)" using x_mod_k_0 by simp hence "?L * k = y * k mod (m * k)" using as1 by (simp only: ac_simps) hence "?L * k = y mod m * k" by (simp only: mult_mod_left) thus ?thesis using as2 by simp qed subsection \More results about quotient \div\ with addition and subtraction\ lemma div_add1_eq_if: "0 < m \ (a + b) div (m::nat) = a div m + b div m + ( if a mod m + b mod m < m then 0 else Suc 0)" apply (simp only: div_add1_eq[of a b]) apply (rule arg_cong[of "(a mod m + b mod m) div m"]) apply (clarsimp simp: linorder_not_less) apply (rule le_less_imp_div[of "Suc 0" m "a mod m + b mod m"], simp) apply simp apply (simp only: add_less_mono[OF mod_less_divisor mod_less_divisor]) done corollary div_add1_eq1: " a mod m + b mod m < (m::nat) \ (a + b) div (m::nat) = a div m + b div m" apply (case_tac "m = 0", simp) apply (simp add: div_add1_eq_if) done corollary div_add1_eq1_mod_0_left: " a mod m = 0 \ (a + b) div (m::nat) = a div m + b div m" apply (case_tac "m = 0", simp) apply (simp add: div_add1_eq1) done corollary div_add1_eq1_mod_0_right: " b mod m = 0 \ (a + b) div (m::nat) = a div m + b div m" by (fastforce simp: div_add1_eq1_mod_0_left) corollary div_add1_eq2: " \ 0 < m; (m::nat) \ a mod m + b mod m \ \ (a + b) div (m::nat) = Suc (a div m + b div m)" by (simp add: div_add1_eq_if) lemma div_Suc: " 0 < n \ Suc m div n = (if Suc (m mod n) = n then Suc (m div n) else m div n)" apply (drule Suc_leI, drule le_imp_less_or_eq) apply (case_tac "n = Suc 0", simp) apply (split if_split, intro conjI impI) apply (rule_tac t="Suc m" and s="m + 1" in subst, simp) apply (subst div_add1_eq2, simp+) apply (insert le_neq_trans[OF mod_less_divisor[THEN Suc_leI, of n m]], simp) apply (rule_tac t="Suc m" and s="m + 1" in subst, simp) apply (subst div_add1_eq1, simp+) done lemma div_Suc': " 0 < n \ Suc m div n = (if m mod n < n - Suc 0 then m div n else Suc (m div n))" apply (simp add: div_Suc) apply (intro conjI impI) apply simp apply (insert le_neq_trans[OF mod_less_divisor[THEN Suc_leI, of n m]], simp) done lemma div_diff1_eq_if: " (b - a) div (m::nat) = b div m - a div m - (if a mod m \ b mod m then 0 else Suc 0)" apply (case_tac "m = 0", simp) apply (case_tac "b < a") apply (frule less_imp_le[of b]) apply (frule div_le_mono[of _ _ m]) apply simp apply (simp only: linorder_not_less neq0_conv) proof - assume le_as: "a \ b" and m_as: "0 < m" have div_le:"a div m \ b div m" using le_as by (simp only: div_le_mono) have "b - a = b div m * m + b mod m - (a div m * m + a mod m)" by simp also have "\ = b div m * m + b mod m - a div m * m - a mod m" by simp also have "\ = b div m * m - a div m * m + b mod m - a mod m" by (simp only: diff_add_assoc2[OF mult_le_mono1[OF div_le]]) finally have b_a_s1: "b - a = (b div m - a div m) * m + b mod m - a mod m" (is "?b_a = ?b_a1") by (simp only: diff_mult_distrib) hence b_a_div_s: "(b - a) div m = ((b div m - a div m) * m + b mod m - a mod m) div m" by (rule arg_cong) show ?thesis proof (cases "a mod m \ b mod m") case True hence as': "a mod m \ b mod m" . have "(b - a) div m = ?b_a1 div m" using b_a_div_s . also have "\ = ((b div m - a div m) * m + (b mod m - a mod m)) div m" using as' by simp also have "\ = b div m - a div m + (b mod m - a mod m) div m" apply (simp only: add.commute) by (simp only: div_mult_self1[OF less_imp_neq[OF m_as, THEN not_sym]]) finally have b_a_div_s': "(b - a) div m = \" . have "(b mod m - a mod m) div m = 0" by (rule div_less, rule less_imp_diff_less, rule mod_less_divisor, rule m_as) thus ?thesis using b_a_div_s' as' by simp next case False hence as1': "\ a mod m \ b mod m" . hence as': "b mod m < a mod m" by simp have a_div_less: "a div m < b div m" using le_as as' by (blast intro: le_mod_greater_imp_div_less) have "b div m - a div m = b div m - a div m - (Suc 0 - Suc 0)" by simp also have "\ = b div m - a div m + Suc 0 - Suc 0" by simp also have "\ = b div m - a div m - Suc 0 + Suc 0" by (simp only: diff_add_assoc2 a_div_less[THEN zero_less_diff[THEN iffD2], THEN Suc_le_eq[THEN iffD2]]) finally have b_a_div_s': "b div m - a div m = \" . have "(b - a) div m = ?b_a1 div m" using b_a_div_s . also have "\ = ((b div m - a div m - Suc 0 + Suc 0) * m + b mod m - a mod m ) div m" using b_a_div_s' by (rule arg_cong) also have "\ = ((b div m - a div m - Suc 0) * m + Suc 0 * m + b mod m - a mod m ) div m" by (simp only: add_mult_distrib) also have "\ = ((b div m - a div m - Suc 0) * m + m + b mod m - a mod m ) div m" by simp also have "\ = ((b div m - a div m - Suc 0) * m + (m + b mod m - a mod m) ) div m" by (simp only: add.assoc m_as diff_add_assoc[of "a mod m" "m + b mod m"] trans_le_add1[of "a mod m" m, OF mod_le_divisor]) also have "\ = b div m - a div m - Suc 0 + (m + b mod m - a mod m) div m" by (simp only: add.commute div_mult_self1[OF less_imp_neq[OF m_as, THEN not_sym]]) finally have b_a_div_s': "(b - a) div m = \" . have div_0_s: "(m + b mod m - a mod m) div m = 0" by (rule div_less, simp only: add_diff_less m_as as') show ?thesis by (simp add: as1' b_a_div_s' div_0_s) qed qed corollary div_diff1_eq: " (b - a) div (m::nat) = b div m - a div m - (m + a mod m - Suc (b mod m)) div m" apply (case_tac "m = 0", simp) apply (simp only: neq0_conv) apply (rule subst[of "if a mod m \ b mod m then 0 else Suc 0" "(m + a mod m - Suc(b mod m)) div m"]) prefer 2 apply (rule div_diff1_eq_if) apply (split if_split, rule conjI) apply simp apply (clarsimp simp: linorder_not_le) apply (rule sym) apply (drule Suc_le_eq[of "b mod m", THEN iffD2]) apply (simp only: diff_add_assoc) apply (simp only: div_add_self1) apply (simp add: less_imp_diff_less) done corollary div_diff1_eq1: " a mod m \ b mod m \ (b - a) div (m::nat) = b div m - a div m" by (simp add: div_diff1_eq_if) corollary div_diff1_eq1_mod_0: " a mod m = 0 \ (b - a) div (m::nat) = b div m - a div m" by (simp add: div_diff1_eq1) corollary div_diff1_eq2: " b mod m < a mod m \ (b - a) div (m::nat) = b div m - Suc (a div m)" by (simp add: div_diff1_eq_if) subsection \Further results about \div\ and \mod\\ subsubsection \Some auxiliary facts about \mod\\ lemma diff_less_divisor_imp_sub_mod_eq: " \ (x::nat) \ y; y - x < m \ \ x = y - (y - x) mod m" by simp lemma diff_ge_divisor_imp_sub_mod_less: " \ (x::nat) \ y; m \ y - x; 0 < m \ \ x < y - (y - x) mod m" apply (simp only: less_diff_conv) apply (simp only: le_diff_conv2 add.commute[of m]) apply (rule less_le_trans[of _ "x + m"]) apply simp_all done lemma le_imp_sub_mod_le: " (x::nat) \ y \ x \ y - (y - x) mod m" apply (case_tac "m = 0", simp_all) apply (case_tac "m \ y - x") apply (drule diff_ge_divisor_imp_sub_mod_less[of x y m]) apply simp_all done lemma mod_less_diff_mod: " \ n mod m < r; r \ m; r \ (n::nat) \ \ (n - r) mod m = m + n mod m - r" apply (case_tac "r = m") apply (simp add: mod_diff_self2) apply (simp add: mod_diff1_eq[of r n m]) done lemma mod_0_imp_mod_pred: " \ 0 < (n::nat); n mod m = 0 \ \ (n - Suc 0) mod m = m - Suc 0" apply (case_tac "m = 0", simp_all) apply (simp only: Suc_le_eq[symmetric]) apply (simp only: mod_diff1_eq) apply (case_tac "m = Suc 0") apply simp_all done lemma mod_pred: " 0 < n \ (n - Suc 0) mod m = ( if n mod m = 0 then m - Suc 0 else n mod m - Suc 0)" apply (split if_split, rule conjI) apply (simp add: mod_0_imp_mod_pred) apply clarsimp apply (case_tac "m = Suc 0", simp) apply (frule subst[OF Suc0_mod[symmetric], where P="\x. x \ n mod m"], simp) apply (simp only: mod_diff1_eq1) apply (simp add: Suc0_mod) done corollary mod_pred_Suc_mod: " 0 < n \ Suc ((n - Suc 0) mod m) mod m = n mod m" apply (case_tac "m = 0", simp) apply (simp add: mod_pred) done corollary diff_mod_pred: " a < b \ (b - Suc a) mod m = ( if a mod m = b mod m then m - Suc 0 else (b - a) mod m - Suc 0)" apply (rule_tac t="b - Suc a" and s="b - a - Suc 0" in subst, simp) apply (subst mod_pred, simp) apply (simp add: mod_eq_diff_mod_0_conv) done corollary diff_mod_pred_Suc_mod: " a < b \ Suc ((b - Suc a) mod m) mod m = (b - a) mod m" apply (case_tac "m = 0", simp) apply (simp add: diff_mod_pred mod_eq_diff_mod_0_conv) done lemma mod_eq_imp_diff_mod_eq_divisor: " \ a < b; 0 < m; a mod m = b mod m \ \ Suc ((b - Suc a) mod m) = m" apply (drule mod_eq_imp_diff_mod_0[of a]) apply (frule iffD2[OF zero_less_diff]) apply (drule mod_0_imp_mod_pred[of "b-a" m], assumption) apply simp done lemma sub_diff_mod_eq: " r \ t \ (t - (t - r) mod m) mod (m::nat) = r mod m" by (metis mod_diff_right_eq diff_diff_cancel diff_le_self) lemma sub_diff_mod_eq': " r \ t \ (k * m + t - (t - r) mod m) mod (m::nat) = r mod m" apply (simp only: diff_mod_le[of t r m, THEN add_diff_assoc, symmetric]) apply (simp add: sub_diff_mod_eq) done lemma mod_eq_Suc_0_conv: "Suc 0 < k \ ((x + k - Suc 0) mod k = 0) = (x mod k = Suc 0)" apply (simp only: mod_pred) apply (case_tac "x mod k = Suc 0") apply simp_all done lemma mod_eq_divisor_minus_Suc_0_conv: "Suc 0 < k \ (x mod k = k - Suc 0) = (Suc x mod k = 0)" by (simp only: mod_Suc, split if_split, fastforce) subsubsection \Some auxiliary facts about \div\\ lemma sub_mod_div_eq_div: "((n::nat) - n mod m) div m = n div m" apply (case_tac "m = 0", simp) apply (simp add: minus_mod_eq_mult_div) done lemma mod_less_imp_diff_div_conv: " \ n mod m < r; r \ m + n mod m\ \ (n - r) div m = n div m - Suc 0" apply (case_tac "m = 0", simp) apply (simp only: neq0_conv) apply (case_tac "n < m", simp) apply (simp only: linorder_not_less) apply (rule div_nat_eqI) apply (simp_all add: algebra_simps minus_mod_eq_mult_div [symmetric]) done corollary mod_0_le_imp_diff_div_conv: " \ n mod m = 0; 0 < r; r \ m \ \ (n - r) div m = n div m - Suc 0" by (simp add: mod_less_imp_diff_div_conv) corollary mod_0_less_imp_diff_Suc_div_conv: " \ n mod m = 0; r < m \ \ (n - Suc r) div m = n div m - Suc 0" by (drule mod_0_le_imp_diff_div_conv[where r="Suc r"], simp_all) corollary mod_0_imp_diff_Suc_div_conv: " (n - r) mod m = 0 \ (n - Suc r) div m = (n - r) div m - Suc 0" apply (case_tac "m = 0", simp) apply (rule_tac t="n - Suc r" and s="n - r - Suc 0" in subst, simp) apply (rule mod_0_le_imp_diff_div_conv, simp+) done corollary mod_0_imp_sub_1_div_conv: " n mod m = 0 \ (n - Suc 0) div m = n div m - Suc 0" apply (case_tac "m = 0", simp) apply (simp add: mod_0_less_imp_diff_Suc_div_conv) done corollary sub_Suc_mod_div_conv: " (n - Suc (n mod m)) div m = n div m - Suc 0" apply (case_tac "m = 0", simp) apply (simp add: mod_less_imp_diff_div_conv) done lemma div_le_conv: "0 < m \ n div m \ k = (n \ Suc k * m - Suc 0)" apply (rule iffI) apply (drule mult_le_mono1[of _ _ m]) apply (simp only: mult.commute[of _ m] minus_mod_eq_mult_div [symmetric]) apply (drule le_diff_conv[THEN iffD1]) apply (rule le_trans[of _ "m * k + n mod m"], assumption) apply (simp add: add.commute[of m]) apply (simp only: diff_add_assoc[OF Suc_leI]) apply (rule add_le_mono[OF le_refl]) apply (rule less_imp_le_pred) apply (rule mod_less_divisor, assumption) apply (drule div_le_mono[of _ _ m]) apply (simp add: mod_0_imp_sub_1_div_conv) done lemma le_div_conv: "0 < (m::nat) \ (n \ k div m) = (n * m \ k)" apply (rule iffI) apply (drule mult_le_mono1[of _ _ m]) apply (simp add: div_mult_cancel) apply (drule div_le_mono[of _ _ m]) apply simp done lemma less_mult_imp_div_less: "n < k * m \ n div m < (k::nat)" apply (case_tac "k = 0", simp) apply (case_tac "m = 0", simp) apply simp apply (drule less_imp_le_pred[of n]) apply (drule div_le_mono[of _ _ m]) apply (simp add: mod_0_imp_sub_1_div_conv) done lemma div_less_imp_less_mult: "\ 0 < (m::nat); n div m < k \ \ n < k * m" apply (rule ccontr, simp only: linorder_not_less) apply (drule div_le_mono[of _ _ m]) apply simp done lemma div_less_conv: "0 < (m::nat) \ (n div m < k) = (n < k * m)" apply (rule iffI) apply (rule div_less_imp_less_mult, assumption+) apply (rule less_mult_imp_div_less, assumption) done lemma div_eq_0_conv: "(n div (m::nat) = 0) = (m = 0 \ n < m)" apply (rule iffI) apply (case_tac "m = 0", simp) apply (rule ccontr) apply (simp add: linorder_not_less) apply (drule div_le_mono[of _ _ m]) apply simp apply fastforce done lemma div_eq_0_conv': "0 < m \ (n div (m::nat) = 0) = (n < m)" by (simp add: div_eq_0_conv) corollary div_gr_imp_gr_divisor: "x < n div (m::nat) \ m \ n" apply (drule gr_implies_gr0, drule neq0_conv[THEN iffD2]) apply (simp add: div_eq_0_conv) done lemma mod_0_less_div_conv: " n mod (m::nat) = 0 \ (k * m < n) = (k < n div m)" apply (case_tac "m = 0", simp) apply fastforce done lemma add_le_divisor_imp_le_Suc_div: " \ x div m \ n; y \ m \ \ (x + y) div m \ Suc n" apply (case_tac "m = 0", simp) apply (simp only: div_add1_eq_if[of _ x]) apply (drule order_le_less[of y, THEN iffD1], fastforce) done text \List of definitions and lemmas\ thm minus_mod_eq_mult_div [symmetric] mod_0_div_mult_cancel div_mult_le less_div_Suc_mult thm Suc0_mod Suc0_mod_subst Suc0_mod_cong thm mod_Suc_conv thm mod_add mod_sub_add thm mod_sub_eq_mod_0_conv mod_sub_eq_mod_swap thm le_mod_greater_imp_div_less thm mod_diff_right_eq mod_eq_imp_diff_mod_eq thm divisor_add_diff_mod_if divisor_add_diff_mod_eq1 divisor_add_diff_mod_eq2 thm mod_add_eq mod_add1_eq_if thm mod_diff1_eq_if mod_diff1_eq mod_diff1_eq1 mod_diff1_eq2 thm nat_mod_distrib int_mod_distrib thm zmod_zminus_eq_conv thm mod_eq_imp_diff_mod_0 zmod_eq_imp_diff_mod_0 thm mod_neq_imp_diff_mod_neq0 diff_mod_0_imp_mod_eq zdiff_mod_0_imp_mod_eq thm zmod_eq_diff_mod_0_conv mod_eq_diff_mod_0_conv thm less_mod_eq_imp_add_divisor_le thm mod_add_eq_imp_mod_0 thm mod_eq_mult_distrib mod_factor_imp_mod_0 mod_factor_div mod_factor_div_mod thm mod_diff_self1 mod_diff_self2 mod_diff_mult_self1 mod_diff_mult_self2 thm div_diff_self1 div_diff_self2 div_diff_mult_self1 div_diff_mult_self2 thm le_less_imp_div div_imp_le_less thm le_less_div_conv thm diff_less_divisor_imp_sub_mod_eq diff_ge_divisor_imp_sub_mod_less le_imp_sub_mod_le thm sub_mod_div_eq_div thm mod_less_imp_diff_div_conv mod_0_le_imp_diff_div_conv mod_0_less_imp_diff_Suc_div_conv mod_0_imp_sub_1_div_conv thm sub_Suc_mod_div_conv thm mod_less_diff_mod mod_0_imp_mod_pred thm mod_pred mod_pred_Suc_mod thm mod_eq_imp_diff_mod_eq_divisor thm diff_mod_le sub_diff_mod_eq sub_diff_mod_eq' thm div_diff1_eq_if div_diff1_eq div_diff1_eq1 div_diff1_eq2 thm div_le_conv end diff --git a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy --- a/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy +++ b/thys/Ordinary_Differential_Equations/Numerics/Example_Utilities.thy @@ -1,2708 +1,2708 @@ theory Example_Utilities imports Init_ODE_Solver begin definition "true_form = Less (floatarith.Num 0) (floatarith.Num 1)" lemma open_true_form[intro, simp]: "open_form true_form" by (auto simp: true_form_def) lemma max_Var_form_true_form[simp]: "max_Var_form true_form = 0" by (auto simp: true_form_def) lemma interpret_form_true_form[simp]: "interpret_form true_form = (\_. True)" by (auto simp: true_form_def) lemmas [simp] = length_aforms_of_ivls declare INF_cong_simp [cong] SUP_cong_simp [cong] image_cong_simp [cong del] declare [[ cd_patterns "_ = interpret_floatariths ?fas _" "_ = interpret_floatarith ?fa _"]] concrete_definition reify_example for i j k uses reify_example hide_const (open) Print.file_output definition "file_output s f = (if s = STR '''' then f (\_. ()) else if s = STR ''-'' then f print else Print.file_output s f)" definition "aforms_of_point xs = aforms_of_ivls xs xs" definition "unit_matrix_list D = concat (map (\i. map (\j. if i = j then 1 else 0::real) [0..) l x \ list_all2 (\) x u}" context includes lifting_syntax begin lemma list_interval_transfer[transfer_rule]: "((list_all2 A) ===> (list_all2 A) ===> rel_set (list_all2 A)) list_interval list_interval" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding list_interval_def by transfer_prover end lemma in_list_interval_lengthD: "x \ list_interval a b \ length x = length a" by (auto simp: list_interval_def list_all2_lengthD) context includes floatarith_notation begin definition "varvec_fas' D C = ((map Var [0..b. (map (\i. (Num (C i)) + Var (D + D * D) * (mvmult_fa D D (map Var [D..i. (map (\j. (Num (C i)) + Var (D + D * D) * Var (D + D * i + j)) [0.. \for illustration\ assumes[simp]: "D=3" "rf = real_of_float" shows "interpret_floatariths (varvec_fas D (\i. [a, b, c] ! i)) [a, b, c, d11, d12, d13, d21, d22, d23, d31, d32, d33, 2] = [rf a, rf b, rf c, rf a + 2 * rf d11, rf a + 2 * rf d12, rf a + 2 * rf d13, rf b + 2 * rf d21, rf b + 2 * rf d22, rf b + 2 * rf d23, rf c + 2 * rf d31, rf c + 2 * rf d32, rf c + 2 * rf d33]" by (simp add: varvec_fas_def mvmult_fa_def eval_nat_numeral) definition "vareq_projections n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ cs \ \(color) coding for partial derivatives\ = [(i + n * (x + 1)::nat, i + n * (y + 1), c). (i, c) \ zip ds cs, (x, y) \ ps]" definition "varvec_aforms_line D X line = approx_floatariths 30 (varvec_fas D (\i. float_of (fst (X ! i)))) (take (D + D*D) X @ line)" definition "varvec_aforms_head D X s = varvec_aforms_line D X (aforms_of_point [s])" definition "varvec_aforms_vec D X s = varvec_aforms_line D (map (\x. (fst x, zero_pdevs)) X) [aform_of_ivl 0 s]" definition "shows_aforms_vareq n \ \dimension\ ps \ \pairs of coordinates to project onto\ ds \ \partial derivatives w.r.t. which variables\ csl \ \color coding for partial derivatives ('arrow' heads)\ csh \ \color coding for partial derivatives (lines)\ s \ \scale vectors for partial derivatives\ (no_str::string) \ \default string if no C1 info is present\ X \ \affine form with C1 info\ = (case (varvec_aforms_head n X s, varvec_aforms_vec n X s) of (Some X, Some Y) \ shows_sep (\(x, y, c). shows_segments_of_aform x y X c) shows_nl (vareq_projections n ps ds csl) o shows_nl o shows_sep (\(x, y, c). shows_segments_of_aform x y Y c) shows_nl (vareq_projections n ps ds csh) o shows_nl | _ \ shows_string no_str o shows_nl)" abbreviation "print_string s \ print (String.implode s)" abbreviation "print_show s \ print_string (s '''')" value [code] "print_show (shows_aforms_vareq 3 [(x, y). x \ [0..<3], y \ [0..<3], x < y] [0..<3] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] [''0x0000ff'', ''0x00ff00'', ''0xff0000''] (FloatR 1 (-1)) ''# no C1 info'' ((((\(a, b). aforms_of_ivls a b) (with_unit_matrix 3 ([10, 20, 30], [12, 22, 32]))))))" method_setup guess_rhs = \ let fun compute ctxt var lhs = let val lhs' = Code_Evaluation.dynamic_value_strict ctxt lhs; val clhs' = Thm.cterm_of ctxt lhs'; val inst = Thm.instantiate (TVars.empty, Vars.make [(var, clhs')]); in PRIMITIVE inst end; fun eval_schematic_rhs ctxt t = (case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) t of SOME (lhs, Var var) => compute ctxt var lhs | _ => no_tac); in Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (SUBGOAL (fn (t, _) => eval_schematic_rhs ctxt t)))) end \ lemma length_aforms_of_point[simp]: "length (aforms_of_point xs) = length xs" by (auto simp: aforms_of_point_def) definition "aform2d_plot_segments x y a = shows_segments_of_aform x y a ''0x000000''" lemma list_of_eucl_prod[simp]: "list_of_eucl (x, y) = list_of_eucl x @ list_of_eucl y" by (auto simp: list_of_eucl_def Basis_list_prod_def intro!: nth_equalityI) lemma list_of_eucl_real[simp]: "list_of_eucl (x::real) = [x]" by (auto simp: list_of_eucl_def Basis_list_real_def) lemma Joints_aforms_of_ivls_self[simp]: "xs \ Joints (aforms_of_ivls xs xs)" by (auto intro!: aforms_of_ivls) lemma Joints_aforms_of_ivls_self_eq[simp]: "Joints (aforms_of_ivls xs xs) = {xs}" apply (auto ) by (auto simp: aforms_of_ivls_def Joints_def valuate_def aform_val_def intro!: nth_equalityI) lemma local_lipschitz_c1_euclideanI: fixes T::"real set" and X::"'a::euclidean_space set" and f::"real \ 'a \ 'a" assumes f': "\t x. t \ T \ x \ X \ (f t has_derivative f' t x) (at x)" assumes cont_f': "\i. i \ Basis \ continuous_on (T \ X) (\(t, x). f' t x i)" assumes "open T" assumes "open X" shows "local_lipschitz T X f" using assms apply (intro c1_implies_local_lipschitz[where f'="\(t, x). Blinfun (f' t x)"]) apply (auto simp: bounded_linear_Blinfun_apply has_derivative_bounded_linear split_beta' intro!: has_derivative_Blinfun continuous_on_blinfun_componentwise) apply (subst continuous_on_cong[OF refl]) defer apply assumption apply auto apply (subst bounded_linear_Blinfun_apply) apply (rule has_derivative_bounded_linear) by auto definition "list_aform_of_aform (x::real aform) = (fst x, list_of_pdevs (snd x))" primrec split_aforms_list:: "(real aform) list list \ nat \ nat \ (real aform) list list" where "split_aforms_list Xs i 0 = Xs" | "split_aforms_list Xs i (Suc n) = split_aforms_list (concat (map (\x. (let (a, b) = split_aforms x i in [a, b])) Xs)) i n" definition "shows_aforms x y c X = shows_lines (map (\b. (shows_segments_of_aform x y b c ''\'')) X)" end definition "the_odo ode_fas safe_form = the(mk_ode_ops ode_fas safe_form)" locale ode_interpretation = fixes safe_form::form and safe_set::"'a::executable_euclidean_space set" and ode_fas::"floatarith list" and ode::"'a \ 'a" and finite::"'n::enum" assumes dims: "DIM('a) = CARD('n)" assumes len: "length ode_fas = CARD('n)" assumes safe_set_form: "safe_set = {x. interpret_form safe_form (list_of_eucl x)}" assumes interpret_fas: "\x. x \ safe_set \ einterpret ode_fas (list_of_eucl x) = ode x" assumes odo: "mk_ode_ops ode_fas safe_form \ None" assumes isFDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length ode_fas) [0.. the_odo ode_fas safe_form" lemmas odo_def = the_odo_def lemma odo_simps[simp]: "ode_expression odo = ode_fas" "safe_form_expr odo = safe_form" using odo by (auto simp: odo_def ode_expression_mk_ode_ops safe_form_expr_mk_ode_ops) lemma safe_set: "safe_set = aform.Csafe odo" using odo dims safe_set_form isFDERIV unfolding aform.Csafe_def aform.safe_def aform.safe_form_def aform.ode_e_def by (auto simp: mk_ode_ops_def safe_set_form len split: if_splits) lemma ode: "\x. x \ safe_set \ ode x = aform.ode odo x" by (auto simp: aform.ode_def aform.ode_e_def interpret_fas) sublocale auto_ll_on_open ode safe_set by (rule aform.auto_ll_on_open_congI[OF safe_set[symmetric] ode[symmetric]]) lemma ode_has_derivative_ode_d1: "(ode has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" if "x \ safe_set" for x proof - from aform.fderiv[OF that[unfolded safe_set]] have "(aform.ode odo has_derivative blinfun_apply (aform.ode_d1 odo x)) (at x)" by simp moreover from topological_tendstoD[OF tendsto_ident_at open_domain(2) that] have "\\<^sub>F x' in at x. x' \ safe_set" . then have "\\<^sub>F x' in at x. aform.ode odo x' = ode x'" by eventually_elim (auto simp: ode) ultimately show ?thesis by (rule has_derivative_transform_eventually) (auto simp: ode that) qed sublocale c1_on_open_euclidean ode "aform.ode_d1 odo" safe_set apply unfold_locales subgoal by simp subgoal by (simp add: ode_has_derivative_ode_d1) subgoal by (rule aform.cont_fderiv') (auto intro!: continuous_intros simp: safe_set) done sublocale transfer_eucl_vec for a::'a and n::'n by unfold_locales (simp add: dims) lemma flow_eq: "t \ existence_ivl0 x \ aform.flow0 odo x t = flow0 x t" and Dflow_eq: "t \ existence_ivl0 x \ aform.Dflow odo x t = Dflow x t" and ex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = existence_ivl0 x" and poincare_mapsto_eq: "closed a \ aform.poincare_mapsto odo a b c d e = poincare_mapsto a b c d e" and flowsto_eq: "aform.flowsto odo = flowsto" apply - subgoal by (rule flow0_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule Dflow_cong[symmetric]) (auto simp: safe_set ode) subgoal by (rule existence_ivl0_cong[symmetric]) (auto simp: safe_set ode) subgoal apply (subst aform.poincare_mapsto_cong[OF safe_set[symmetric]]) by (auto simp: ode) subgoal apply (intro ext) apply (subst flowsto_congI[OF safe_set ode]) by (auto simp: safe_set) done definition "avf \ \x::'n rvec. cast (aform.ode odo (cast x)::'a)::'n rvec" context includes lifting_syntax begin lemma aform_ode_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_ve) aform.ode aform.ode" unfolding aform.ode_def by transfer_prover lemma cast_aform_ode: "cast (aform.ode odo (cast (x::'n rvec))::'a) = aform.ode odo x" by transfer simp lemma aform_safe_transfer[transfer_rule]: "((=) ===> rel_ve ===> (=)) aform.safe aform.safe" unfolding aform.safe_def by transfer_prover lemma aform_Csafe_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.Csafe aform.Csafe" unfolding aform.Csafe_def by transfer_prover lemma cast_safe_set: "(cast ` safe_set::'n rvec set) = aform.Csafe odo" unfolding safe_set by transfer simp lemma aform_ode_d_raw_transfer[transfer_rule]: "((=) ===> (=) ===> rel_ve ===> rel_ve ===> rel_ve ===> rel_ve) aform.ode_d_raw aform.ode_d_raw" unfolding aform.ode_d_raw_def by transfer_prover lemma aform_ode_d_raw_aux_transfer: "((=) ===> rel_ve ===> rel_ve ===> rel_ve) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0) (\x xb xa. if xb \ aform.Csafe x then aform.ode_d_raw x 0 xb 0 xa else 0)" by transfer_prover lemma aform_ode_d1_transfer[transfer_rule]: "((=) ===> rel_ve ===> rel_blinfun rel_ve rel_ve) aform.ode_d1 aform.ode_d1" apply (auto simp: rel_blinfun_def aform.ode_d1_def intro!: rel_funI) unfolding aform.ode_d.rep_eq using aform_ode_d_raw_aux_transfer apply - apply (drule rel_funD, rule refl) apply (drule rel_funD, assumption) apply (drule rel_funD; assumption) done lemma cast_bl_transfer[transfer_rule]: "(rel_blinfun (=) (=) ===> rel_blinfun rel_ve rel_ve) id_blinfun cast_bl" by (auto simp: rel_ve_cast rel_blinfun_def intro!: rel_funI dest!: rel_funD) lemma cast_bl_transfer'[transfer_rule]: "(rel_blinfun rel_ve rel_ve ===> rel_blinfun (=) (=)) id_blinfun cast_bl" apply (auto simp: rel_ve_cast rel_blinfun_def cast_cast intro!: rel_funI dest!: rel_funD) by (subst cast_cast) auto lemma rel_blinfun_eq[relator_eq]: "rel_blinfun (=) (=) = (=)" by (auto simp: Rel_def rel_blinfun_def blinfun_ext rel_fun_eq intro!: rel_funI ext) lemma cast_aform_ode_D1: "cast_bl (aform.ode_d1 odo (cast (x::'n rvec))::'a\\<^sub>L'a) = (aform.ode_d1 odo x::'n rvec \\<^sub>L 'n rvec)" by transfer simp end definition "vf \ \x. cast (ode (cast x))" definition "vf' \ \x::'n rvec. cast_bl (aform.ode_d1 odo (cast x::'a)) ::'n rvec \\<^sub>L 'n rvec" definition "vX \ cast ` safe_set" sublocale a?: transfer_c1_on_open_euclidean a n ode "aform.ode_d1 odo" safe_set vf vf' vX for a::'a and n::'n by unfold_locales (simp_all add: dims vf_def vf'_def vX_def) sublocale av: transfer_c1_on_open_euclidean a n "aform.ode odo" "aform.ode_d1 odo" "(aform.Csafe odo)" avf vf' vX for a::'a and n::'n apply unfold_locales unfolding vX_def by (simp_all add: dims avf_def safe_set) lemma vflow_eq: "t \ v.existence_ivl0 x \ aform.flow0 odo x t = v.flow0 x t" thm flow_eq[of t "cast x"] flow_eq[of t "cast x", untransferred] apply (subst flow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set .. lemma vf'_eq: "vf' = aform.ode_d1 odo" unfolding vf'_def cast_aform_ode_D1 .. lemma vDflow_eq: "t \ v.existence_ivl0 x \ aform.Dflow odo x t = v.Dflow x t" apply (subst Dflow_eq[of t "cast x", untransferred, symmetric]) apply simp unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq .. lemma vex_ivl_eq: "t \ aform.existence_ivl0 odo x \ aform.existence_ivl0 odo x = v.existence_ivl0 x" apply (subst ex_ivl_eq[of t "cast x", untransferred, symmetric]) unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto context includes lifting_syntax begin lemma id_cast_eucl1_transfer_eq: "(\x. x) = (\x. (fst x, 1\<^sub>L o\<^sub>L snd x o\<^sub>L 1\<^sub>L))" by auto lemma cast_eucl1_transfer[transfer_rule]: "(rel_prod (=) (rel_blinfun (=) (=)) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) (\x. x) cast_eucl1" unfolding cast_eucl1_def id_cast_eucl1_transfer_eq apply transfer_prover_start apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply (transfer_step) apply simp done end lemma avpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = av.v.poincare_mapsto a b c d e" if "closed a" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto lemma vpoincare_mapsto_eq: "aform.poincare_mapsto odo a (b::'n eucl1 set) c d e = v.poincare_mapsto a b c d e" if "closed a" proof - have "closed (cast ` a::'a set)" using that by transfer auto from poincare_mapsto_eq[of "cast ` a::'a set" "cast_eucl1 ` b::('a \ 'a \\<^sub>L 'a) set" "cast ` c::'a set" "cast ` d::'a set" "cast_eucl1 ` e::('a \ 'a \\<^sub>L 'a) set", OF this, untransferred] have "v.poincare_mapsto a b c d e = av.v.poincare_mapsto a b c d e" by auto also have "\ = aform.poincare_mapsto odo a (b::'n eucl1 set) c d e" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?thesis by simp qed lemma avflowsto_eq: "aform.flowsto odo = (av.v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "av.v.flowsto a b c d = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto then show ?case by simp qed lemma vflowsto_eq: "aform.flowsto odo = (v.flowsto::'n eucl1 set \ _)" proof (intro ext, goal_cases) case (1 a b c d) have "aform.flowsto odo (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d) = flowsto (cast_eucl1 ` a::'a c1_info set) b (cast_eucl1 ` c) (cast_eucl1 ` d)" by (subst flowsto_eq) auto from this[untransferred] have "v.flowsto a b c d = av.v.flowsto a b c d" by auto also have "\ = aform.flowsto odo a b c d" unfolding avf_def vX_def cast_aform_ode cast_safe_set vf'_eq by auto finally show ?case by simp qed context includes lifting_syntax begin lemma flow1_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_prod rel_ve (rel_blinfun rel_ve rel_ve)) flow1_of_list flow1_of_list" unfolding flow1_of_list_def blinfun_of_list_def o_def flow1_of_vec1_def by transfer_prover lemma c1_info_of_appr_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (rel_option (list_all2 (=))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr aform.c1_info_of_appr" unfolding aform.c1_info_of_appr_def by transfer_prover lemma c0_info_of_appr_transfer[transfer_rule]: "((list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr aform.c0_info_of_appr" unfolding aform.c0_info_of_appr_def by transfer_prover lemma aform_scaleR2_transfer[transfer_rule]: "((=) ===> (=) ===> rel_set (rel_prod A B) ===> rel_set (rel_prod A B)) scaleR2 scaleR2" if [unfolded Rel_def, transfer_rule]: "((=) ===> B ===> B) (*\<^sub>R) (*\<^sub>R)" unfolding scaleR2_def by transfer_prover lemma scaleR_rel_blinfun_transfer[transfer_rule]: "((=) ===> rel_blinfun rel_ve rel_ve ===> rel_blinfun rel_ve rel_ve) (*\<^sub>R) (*\<^sub>R)" apply (auto intro!: rel_funI simp: rel_blinfun_def blinfun.bilinear_simps) apply (drule rel_funD) apply assumption apply (rule scaleR_transfer[THEN rel_funD, THEN rel_funD]) apply auto done lemma c1_info_of_appre_transfer[transfer_rule]: "(rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appre aform.c1_info_of_appre" unfolding aform.c1_info_of_appre_def by transfer_prover lemma c1_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprs aform.c1_info_of_apprs" unfolding aform.c1_info_of_apprs_def by transfer_prover lemma c1_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_appr' aform.c1_info_of_appr'" unfolding aform.c1_info_of_appr'_def by transfer_prover lemma c0_info_of_apprs_transfer[transfer_rule]: "((=) ===> rel_set rel_ve) aform.c0_info_of_apprs aform.c0_info_of_apprs" unfolding aform.c0_info_of_apprs_def by transfer_prover lemma c0_info_of_appr'_transfer[transfer_rule]: "(rel_option (list_all2 (=)) ===> rel_set rel_ve) aform.c0_info_of_appr' aform.c0_info_of_appr'" unfolding aform.c0_info_of_appr'_def by transfer_prover lemma aform_Csafe_vX[simp]: "aform.Csafe odo = (vX::'n rvec set)" by (simp add: vX_def cast_safe_set) definition blinfuns_of_lvivl::"real list \ real list \ ('b \\<^sub>L 'b::executable_euclidean_space) set" where "blinfuns_of_lvivl x = blinfun_of_list ` list_interval (fst x) (snd x)" lemma blinfun_of_list_transfer[transfer_rule]: "(list_all2 (=) ===> rel_blinfun rel_ve rel_ve) blinfun_of_list blinfun_of_list" unfolding blinfun_of_list_def by transfer_prover lemma blinfuns_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl blinfuns_of_lvivl" unfolding blinfuns_of_lvivl_def by transfer_prover definition "blinfuns_of_lvivl' x = (case x of None \ UNIV | Some x \ blinfuns_of_lvivl x)" lemma blinfuns_of_lvivl'_transfer[transfer_rule]: "(rel_option (rel_prod (list_all2 (=)) (list_all2 (=))) ===> rel_set (rel_blinfun rel_ve rel_ve)) blinfuns_of_lvivl' blinfuns_of_lvivl'" unfolding blinfuns_of_lvivl'_def by transfer_prover lemma atLeastAtMost_transfer[transfer_rule]: "(A ===> A ===> rel_set A) atLeastAtMost atLeastAtMost" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding atLeastAtMost_def atLeast_def atMost_def by transfer_prover lemma set_of_ivl_transfer[transfer_rule]: "(rel_prod A A ===> rel_set A) set_of_ivl set_of_ivl" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" "bi_unique A" unfolding set_of_ivl_def by transfer_prover lemma set_of_lvivl_transfer[transfer_rule]: "(rel_prod (list_all2 (=)) (list_all2 (=)) ===> rel_set rel_ve) set_of_lvivl set_of_lvivl" unfolding set_of_lvivl_def by transfer_prover lemma set_of_lvivl_eq: "set_of_lvivl I = (eucl_of_list ` list_interval (fst I) (snd I)::'b::executable_euclidean_space set)" if [simp]: "length (fst I) = DIM('b)" "length (snd I) = DIM('b)" proof (auto simp: set_of_lvivl_def list_interval_def set_of_ivl_def, goal_cases) case (1 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "list_of_eucl x"] lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "list_of_eucl x" "snd I"] show ?case by force next case (2 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "fst I" "x"] show ?case by (auto simp: list_all2_lengthD) next case (3 x) with lv_rel_le[where 'a='b, param_fo, OF lv_relI lv_relI, of "x" "snd I"] show ?case by (auto simp: list_all2_lengthD) qed lemma bounded_linear_matrix_vector_mul[THEN bounded_linear_compose, bounded_linear_intros]: "bounded_linear ((*v) x)" for x::"real^'x^'y" unfolding linear_linear by (rule matrix_vector_mul_linear) lemma blinfun_of_list_eq: "blinfun_of_list x = blinfun_of_vmatrix (eucl_of_list x::((real, 'b) vec, 'b) vec)" if "length x = CARD('b::enum)*CARD('b)" unfolding blinfun_of_list_def apply (transfer fixing: x) apply (rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear apply (auto intro!: bounded_linear_intros) proof goal_cases case (1 b) have "(eucl_of_list x::((real, 'b) vec, 'b) vec) *v b = (eucl_of_list x::((real, 'b) vec, 'b) vec) *v eucl_of_list (list_of_eucl b)" by simp also have "\ = (\ij Basis_list ! j)) *\<^sub>R Basis_list ! i)" by (subst eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list) (auto simp: that) also have "\ = (\i\Basis. \j\Basis. (b \ j * x ! (index Basis_list i * CARD('b) + index Basis_list j)) *\<^sub>R i)" apply (subst sum_list_Basis_list[symmetric])+ apply (subst sum_list_sum_nth)+ by (auto simp add: atLeast0LessThan scaleR_sum_left intro!: sum.cong) finally show ?case by simp qed lemma blinfuns_of_lvivl_eq: "blinfuns_of_lvivl x = (blinfun_of_vmatrix ` set_of_lvivl x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "length (fst x) = CARD('b::enum)*CARD('b)" "length (snd x) = CARD('b)*CARD('b)" apply (subst set_of_lvivl_eq) subgoal by (simp add: that) subgoal by (simp add: that) unfolding blinfuns_of_lvivl_def image_image by (auto simp: that blinfun_of_list_eq[symmetric] in_list_interval_lengthD cong: image_cong) lemma range_blinfun_of_vmatrix[simp]: "range blinfun_of_vmatrix = UNIV" apply auto apply transfer subgoal for x by (rule image_eqI[where x="matrix x"]) auto done lemma blinfun_of_vmatrix_image: "blinfun_of_vmatrix ` aform.set_of_lvivl' x = (blinfuns_of_lvivl' x::((real, 'b) vec \\<^sub>L (real, 'b) vec) set)" if "aform.lvivl'_invar (CARD('b)*CARD('b::enum)) x" using that by (auto simp: aform.set_of_lvivl'_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_eq aform.lvivl'_invar_def split: option.splits) lemma one_step_result123: "solves_one_step_until_time_aform optns odo X0i t1 t2 E dE \ (x0, d0) \ aform.c1_info_of_appre X0i \ t \ {t1 .. t2} \ set_of_lvivl E \ S \ blinfuns_of_lvivl' dE \ dS \ length (fst E) = CARD('n) \ length (snd E) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dE \ aform.c1_info_invare DIM('a) X0i \ aform.D odo = DIM('a) \ (t \ existence_ivl0 (x0::'a) \ flow0 x0 t \ S) \ Dflow x0 t o\<^sub>L d0 \ dS" apply (transfer fixing: optns X0i t1 t2 t E dE) subgoal premises prems for x0 d0 S dS proof - have "t \ aform.existence_ivl0 odo x0 \ aform.flow0 odo x0 t \ S \ aform.Dflow odo x0 t o\<^sub>L d0 \ dS" apply (rule one_step_in_ivl[of t t1 t2 x0 d0 X0i "fst E" "snd E" S dE dS odo optns]) using prems by (auto simp: eucl_of_list_prod set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image aform.D_def solves_one_step_until_time_aform_def) with vflow_eq[of t x0] vDflow_eq[of t x0] vex_ivl_eq[symmetric, of t x0] show ?thesis by simp qed done lemmas one_step_result12 = one_step_result123[THEN conjunct1] and one_step_result3 = one_step_result123[THEN conjunct2] lemmas one_step_result1 = one_step_result12[THEN conjunct1] and one_step_result2 = one_step_result12[THEN conjunct2] lemma plane_of_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) plane_of plane_of" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding plane_of_def by transfer_prover lemma below_halfspace_transfer[transfer_rule]: "(rel_sctn A ===> rel_set A) below_halfspace below_halfspace" if [transfer_rule]: "(A ===> A ===> (=)) (\) (\)" "bi_total A" unfolding below_halfspace_def le_halfspace_def by transfer_prover definition "rel_nres A a b \ (a, b) \ \{(a, b). A a b}\nres_rel" lemma FAILi_transfer[transfer_rule]: "(rel_nres B) FAILi FAILi" by (auto simp: rel_nres_def nres_rel_def) lemma RES_transfer[transfer_rule]: "(rel_set B ===> rel_nres B) RES RES" by (auto simp: rel_nres_def nres_rel_def rel_set_def intro!: rel_funI RES_refine) context includes autoref_syntax begin lemma RETURN_dres_nres_relI: "(fi, f) \ A \ B \ (\x. dRETURN (fi x), (\x. RETURN (f x))) \ A \ \B\dres_nres_rel" by (auto simp: dres_nres_rel_def dest: fun_relD) end lemma br_transfer[transfer_rule]: "((B ===> C) ===> (B ===> (=)) ===> rel_set (rel_prod B C)) br br" if [transfer_rule]: "bi_total B" "bi_unique C" "bi_total C" unfolding br_def by transfer_prover lemma aform_appr_rel_transfer[transfer_rule]: "(rel_set (rel_prod (list_all2 (=)) (rel_set rel_ve))) aform.appr_rel aform.appr_rel" unfolding aform.appr_rel_br by (transfer_prover) lemma appr1_rel_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (list_all2 (=)) (rel_option (list_all2 (=)))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1_rel aform.appr1_rel" unfolding aform.appr1_rel_internal by transfer_prover lemma relAPP_transfer[transfer_rule]: "((rel_set (rel_prod B C) ===> D) ===> rel_set (rel_prod B C) ===> D) Relators.relAPP Relators.relAPP" unfolding relAPP_def by transfer_prover lemma prod_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (rel_prod B D) (rel_prod C E))) prod_rel prod_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding prod_rel_def_internal by transfer_prover lemma Domain_transfer[transfer_rule]: "(rel_set (rel_prod A B) ===> rel_set A) Domain Domain" if [transfer_rule]: "bi_total A" "bi_unique A" "bi_total B" "bi_unique B" unfolding Domain_unfold by transfer_prover lemma set_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod (rel_set B) (rel_set C))) set_rel set_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" unfolding set_rel_def_internal by transfer_prover lemma relcomp_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod C D) ===> rel_set (rel_prod B D)) relcomp relcomp" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding relcomp_unfold by transfer_prover lemma Union_rel_transfer[transfer_rule]: "(rel_set (rel_prod B (rel_set C)) ===> rel_set (rel_prod C (rel_set D)) ===> rel_set (rel_prod B (rel_set D))) Union_rel Union_rel" if [transfer_rule]: "bi_total B" "bi_unique B" "bi_unique C" "bi_total C" "bi_unique D" "bi_total D" unfolding Union_rel_internal top_fun_def top_bool_def by transfer_prover lemma fun_rel_transfer[transfer_rule]: "(rel_set (rel_prod B C) ===> rel_set (rel_prod D E) ===> rel_set (rel_prod (B ===> D) (C ===> E))) Relators.fun_rel Relators.fun_rel" if [transfer_rule]: "bi_unique B" "bi_unique C" "bi_unique D" "bi_total D" "bi_unique E" "bi_total E" unfolding fun_rel_def_internal by transfer_prover lemma c1_info_of_apprse_transfer[transfer_rule]: "(list_all2 (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) ===> rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))) aform.c1_info_of_apprse aform.c1_info_of_apprse" unfolding aform.c1_info_of_apprse_def by transfer_prover term scaleR2_rel (* "scaleR2_rel" :: "('b \ ('c \ 'd) set) set \ (((ereal \ ereal) \ 'b) \ ('c \ 'd) set) set" *) lemma scaleR2_rel_transfer[transfer_rule]: "(rel_set (rel_prod (=) (rel_set (rel_prod (=) (=)))) ===> rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (=)) (rel_set (rel_prod (=) (=))))) scaleR2_rel scaleR2_rel" unfolding scaleR2_rel_internal by transfer_prover lemma appr1_rele_transfer[transfer_rule]: "(rel_set (rel_prod (rel_prod (rel_prod (=) (=)) (rel_prod (list_all2 (=)) (rel_option (list_all2 (=))))) (rel_set (rel_prod rel_ve (rel_blinfun rel_ve rel_ve))))) aform.appr1e_rel aform.appr1e_rel" unfolding scaleR2_rel_internal by transfer_prover lemma flow1_of_vec1_times: "flow1_of_vec1 ` (A \ B) = A \ blinfun_of_vmatrix ` B" by (auto simp: flow1_of_vec1_def vec1_of_flow1_def) lemma stable_on_transfer[transfer_rule]: "(rel_set rel_ve ===> rel_set rel_ve ===> (=)) v.stable_on stable_on" unfolding stable_on_def v.stable_on_def by transfer_prover theorem solves_poincare_map_aform: "solves_poincare_map_aform optns odo (\x. dRETURN (symstart x)) [S] guards ivl sctn roi XS RET dRET \ (symstart, symstarta) \ fun_rel (aform.appr1e_rel) (clw_rel aform.appr_rel \\<^sub>r clw_rel aform.appr1e_rel) \ (\X0. (\(CX, X). flowsto (X0 - trap \ UNIV) {0..} (CX \ UNIV) X) (symstarta X0)) \ stable_on (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) trap \ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns symstart S guards ivl sctn roi XS RET dRET) subgoal premises prems for symstarta trap proof - have "aform.poincare_mapsto odo (set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS - trap \ UNIV) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (flow1_of_vec1 ` ({eucl_of_list (fst RET)..eucl_of_list (snd RET)} \ aform.set_of_lvivl' dRET))" apply (rule solves_poincare_map[OF _ RETURN_dres_nres_relI RETURN_rule, of optns odo symstart S guards ivl sctn roi XS "fst RET" "snd RET" dRET symstarta trap]) subgoal using prems(1) by (simp add: solves_poincare_map_aform_def) subgoal using prems(2) by (auto simp: fun_rel_def_internal) subgoal for X0 using prems(3)[of X0] vflowsto_eq by auto subgoal unfolding aform.stable_on_def proof (safe, goal_cases) case (1 t x0) from 1 have a: "t \ v.existence_ivl0 x0" using vex_ivl_eq by blast with 1 have b: "v.flow0 x0 t \ trap" using vflow_eq by simp have c: "v.flow0 x0 s \ vX - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)" if s: "s \ {0<..t}" for s using 1(4)[rule_format, OF s] apply (subst (asm) vflow_eq) unfolding aform_Csafe_vX[symmetric] using s a by (auto dest!: a.v.ivl_subset_existence_ivl) from a b c show ?case using prems(4)[unfolded v.stable_on_def, rule_format, OF b a 1(3) c] by simp qed subgoal using prems by auto subgoal using prems by (auto simp: aform.D_def) subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto subgoal using prems by auto done then show ?thesis using vflow_eq vex_ivl_eq vflowsto_eq prems apply (subst vpoincare_mapsto_eq[symmetric]) by (auto simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times) qed done theorem solves_poincare_map_aform': "solves_poincare_map_aform' optns odo S guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ length (normal S) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) (below_halfspace (map_sctn eucl_of_list S)) (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns S guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map'[of optns odo S guards ivl sctn roi XS "fst RET" "snd RET" dRET] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_aform'_def) done theorem solves_poincare_map_onto_aform: "solves_poincare_map_onto_aform optns odo guards ivl sctn roi XS RET dRET\ (\X. X \ set XS \ aform.c1_info_invare DIM('a) X) \ aform.D odo = DIM('a) \ length (normal sctn) = DIM('a) \ length (fst ivl) = DIM('a) \ length (snd ivl) = DIM('a) \ (\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)) \ length (fst RET) = CARD('n) \ length (snd RET) = CARD('n) \ aform.lvivl'_invar (CARD('n) * CARD('n)) dRET \ poincare_mapsto ((set_of_lvivl ivl::('a set)) \ plane_of (map_sctn eucl_of_list sctn)) (aform.c1_info_of_apprse XS) UNIV (aform.Csafe odo - set_of_lvivl ivl \ plane_of (map_sctn eucl_of_list sctn)) (set_of_lvivl RET \ blinfuns_of_lvivl' dRET)" apply (transfer fixing: optns guards ivl sctn roi XS RET dRET) subgoal using solves_poincare_map_onto[of optns odo guards ivl sctn roi XS "fst RET" "snd RET" dRET, where 'n='n, unfolded aform.poincare_maps_onto_def] using vflow_eq vex_ivl_eq vflowsto_eq apply (subst vpoincare_mapsto_eq[symmetric]) by (auto intro!: closed_Int simp: set_of_lvivl_def set_of_ivl_def blinfun_of_vmatrix_image flow1_of_vec1_times aform.D_def solves_poincare_map_onto_aform_def) done end end subsection \Example Utilities!\ hide_const floatarith.Max floatarith.Min lemma degree_sum_pdevs_scaleR_Basis: "degree (sum_pdevs (\i. pdevs_scaleR (a i) i) (Basis::'b::euclidean_space set)) = Max ((\i. degree (a i)) ` Basis)" apply (rule antisym) subgoal apply (rule degree_le) by (auto ) subgoal apply (rule Max.boundedI) apply simp apply simp apply (auto simp: intro!: degree_leI) by (auto simp: euclidean_eq_iff[where 'a='b]) done lemma Inf_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Inf_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Inf_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Inf_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_left) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma Sup_aform_eucl_of_list_aform: assumes "length a = DIM('b::executable_euclidean_space)" shows "Sup_aform (eucl_of_list_aform a::'b aform) = eucl_of_list (map Sup_aform a)" using assms apply (auto simp: eucl_of_list_aform_def Sup_aform_def[abs_def] algebra_simps eucl_of_list_inner inner_sum_left intro!: euclidean_eqI[where 'a='b]) apply (auto simp: tdev_def inner_sum_left abs_inner inner_Basis if_distrib cong: if_cong) apply (rule sum.mono_neutral_cong_right) apply simp by (auto simp: degree_sum_pdevs_scaleR_Basis) lemma eucl_of_list_map_Inf_aform_leI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "eucl_of_list (map Inf_aform a) \ x" using Inf_aform_le_Affine[OF assms(1)] assms(2) by (auto simp: Inf_aform_eucl_of_list_aform) lemma eucl_of_list_map_Sup_aform_geI: assumes "x \ Affine (eucl_of_list_aform a::'b::executable_euclidean_space aform)" assumes "length a = DIM('b)" shows "x \ eucl_of_list (map Sup_aform a)" using Sup_aform_ge_Affine[OF assms(1)] assms(2) by (auto simp: Sup_aform_eucl_of_list_aform) lemma mem_Joints_appendE: assumes "x \ Joints (xs @ ys)" obtains x1 x2 where "x = x1 @ x2" "x1 \ Joints xs" "x2 \ Joints ys" using assms by (auto simp: Joints_def valuate_def) lemma c1_info_of_appr_subsetI1: fixes X1::"'b::executable_euclidean_space set" assumes subset: "{eucl_of_list (map Inf_aform (fst R)) .. eucl_of_list (map Sup_aform (fst R))} \ X1" assumes len: "length (fst R) = DIM('b)" shows "aform.c1_info_of_appr R \ X1 \ UNIV" using len apply (auto simp: aform.c1_info_of_appr_def flow1_of_list_def split: option.splits intro!: subsetD[OF subset] elim!: mem_Joints_appendE) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal by (auto intro!: eucl_of_list_mem_eucl_of_list_aform eucl_of_list_map_Inf_aform_leI eucl_of_list_map_Sup_aform_geI) subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) apply (simp add: Joints_imp_length_eq eucl_of_list_map_Inf_aform_leI eucl_of_list_mem_eucl_of_list_aform) done subgoal apply (subst (2) eucl_of_list_take_DIM[symmetric, OF refl]) apply (auto simp: min_def) by (simp add: Joints_imp_length_eq eucl_of_list_map_Sup_aform_geI eucl_of_list_mem_eucl_of_list_aform) done lemmas [simp] = compute_tdev syntax product_aforms::"(real aform) list \ (real aform) list \ (real aform) list" (infixr "\\<^sub>a" 70) lemma matrix_inner_Basis_list: includes vec_syntax assumes "k < CARD('n) * CARD('m)" shows "(f::(('n::enum rvec, 'm::enum) vec)) \ Basis_list ! k = vec_nth (vec_nth f (enum_class.enum ! (k div CARD('n)))) (enum_class.enum ! (k mod CARD('n)))" proof - have "f \ Basis_list ! k = (\x\UNIV. \xa\UNIV. if enum_class.enum ! (k mod CARD('n)) = xa \ enum_class.enum ! (k div CARD('n)) = x then f $ x $ xa else 0)" using assms unfolding inner_vec_def apply (auto simp: Basis_list_vec_def concat_map_map_index) apply (subst (2) sum.cong[OF refl]) apply (subst sum.cong[OF refl]) apply (subst (2) vec_nth_Basis2) apply (force simp add: Basis_vec_def Basis_list_real_def) apply (rule refl) apply (rule refl) apply (auto simp: if_distribR if_distrib axis_eq_axis Basis_list_real_def cong: if_cong) done also have "\ = f $ enum_class.enum ! (k div CARD('n)) $ enum_class.enum ! (k mod CARD('n))" apply (subst if_conn) apply (subst sum.delta') apply simp by (simp add: sum.delta') finally show ?thesis by simp qed lemma list_of_eucl_matrix: includes vec_syntax shows "(list_of_eucl (M::(('n::enum rvec, 'm::enum) vec))) = concat (map (\i. map (\j. M $ (enum_class.enum ! i)$ (enum_class.enum ! j) ) [0.. real_of_float (lapprox_rat 20 i j))" definition "udec x = (case quotient_of x of (i, j) \ real_of_float (rapprox_rat 20 i j))" lemma ldec: "ldec x \ real_of_rat x" and udec: "real_of_rat x \ udec x" apply (auto simp: ldec_def udec_def split: prod.splits intro!: lapprox_rat[le] rapprox_rat[ge]) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) apply (metis Fract_of_int_quotient less_eq_real_def less_int_code(1) of_rat_rat quotient_of_denom_pos quotient_of_div) done context includes floatarith_notation begin definition "matrix_of_degrees\<^sub>e = (let ur = Rad_of (Var 0); vr = Rad_of (Var 1) in [Cos ur, Cos vr, 0, Sin ur, Sin vr, 0, 0, 0, 0])" definition "matrix_of_degrees u v = approx_floatariths 30 matrix_of_degrees\<^sub>e (aforms_of_point ([u, v, 0]))" lemma interpret_floatariths_matrix_of_degrees: "interpret_floatariths matrix_of_degrees\<^sub>e (([u::real, v::real, 0])) = [cos (rad_of u), cos (rad_of v), 0, sin (rad_of u), sin (rad_of v), 0, 0, 0, 0]" by (auto simp: matrix_of_degrees\<^sub>e_def Let_def inverse_eq_divide) definition "num_options p sstep m N a projs print_fun = \ precision = p, adaptive_atol = FloatR 1 (- a), adaptive_rtol = FloatR 1 (- a), method_id = 2, start_stepsize = FloatR 1 (- sstep), iterations = 40, halve_stepsizes = 40, widening_mod = 10, rk2_param = FloatR 1 0, default_reduce = correct_girard (p) (m) (N), printing_fun = (\a b. let _ = fold (\(x, y, c) _. print_fun (String.implode (shows_segments_of_aform (x) (y) b c ''\''))) projs (); _ = print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\'')) in () ), tracing_fun = (\a b. let _ = print_fun (String.implode (''# '' @ a @ ''\'')) in case b of Some b \ (let _ = () in print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\''))) | None \ ()) \" definition "num_options_c1 p sstep m N a projs dcolors print_fun = (let no = num_options p sstep m N a (map (\(x, y, c, ds). (x, y, c)) projs) print_fun; D = length dcolors in no \printing_fun:= (\a b. let _ = printing_fun no a b in if a then () else fold (\(x, y, c, ds) _. print_fun (String.implode (shows_aforms_vareq D [(x, y)] ds dcolors dcolors (FloatR 1 (-1)) ''# no C1 info'' b ''''))) projs () ) \)" definition "num_options_code p sstep m N a projs print_fun = num_options (nat_of_integer p) (int_of_integer sstep) (nat_of_integer m) (nat_of_integer N) (int_of_integer a) (map (\(i, j, k). (nat_of_integer i, nat_of_integer j, k)) projs) print_fun" definition "ro s n M g0 g1 inter_step = \max_tdev_thres = FloatR 1 s, pre_split_reduce = correct_girard 30 n M, pre_inter_granularity = FloatR 1 g0, post_inter_granularity = (FloatR 1 g1), pre_collect_granularity = FloatR 1 g0, max_intersection_step = FloatR 1 inter_step\" definition "code_ro s n m g0 g1 inter_step = ro (int_of_integer s) (nat_of_integer n) (nat_of_integer m) (int_of_integer g0) (int_of_integer g1) (int_of_integer inter_step)" fun xsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [1,0,0] x)" fun xsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec' x (y0, y1) (z0, z1) = (([x, y0, z0], [x, y1, z1]), Sctn [-1,0,0] (-x))" fun ysec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, 1,0] y)" fun ysec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec' (x0, x1) y (z0, z1) = (([x0, y, z0], [x1, y, z1]), Sctn [0, -1,0] (-y))" fun zsec:: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, 1] z)" fun zsec':: "real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "zsec' (x0, x1) (y0, y1) z = (([x0, y0, z], [x1, y1, z]), Sctn [0, 0, -1] (-z))" fun xsec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "xsec2 x (y0, y1) = (([x, y0], [x, y1]), Sctn [1,0] x)" fun xsec2':: "real \ real \ real \(real list \ real list) \ real list sctn" where "xsec2' x (y0, y1) = (([x, y0], [x, y1]), Sctn [-1,0] (-x))" fun ysec2:: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2 (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, 1] y)" fun ysec2':: "real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec2' (x0, x1) y = (([x0, y], [x1, y]), Sctn [0, -1] (-y))" fun ysec4:: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4 (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, 1,0, 0] (y))" fun ysec4':: "real \ real \ real \ real \ real \ real \ real \ (real list \ real list) \ real list sctn" where "ysec4' (x0, x1) y (z0, z1) (m0, m1) = (([x0, y, z0, m0], [x1, y, z1, m1]), Sctn [0, -1,0, 0] (-y))" definition "code_sctn N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := 1]) c" definition "code_sctn' N n c = Sctn ((replicate (nat_of_integer N) (0::real))[nat_of_integer n := -1]) (-c)" definition "lrat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition "urat i j = real_of_float (lapprox_rat 30 (int_of_integer i) (int_of_integer j))" definition [simp]: "TAG_optns (optns::string \ ((String.literal \ unit) \ (real aform) numeric_options)) = True" lemma TAG_optns: "P \ (TAG_optns optns \ P)" by auto definition [simp]: "TAG_reach_optns (roi::real aform reach_options) = True" lemma TAG_reach_optns: "P \ (TAG_reach_optns optns \ P)" by auto definition [simp]: "TAG_sctn (b::bool) = True" lemma TAG_sctn: "P \ (TAG_sctn optns \ P)" by auto subsection \Integrals and Computation\ lemma has_vderiv_on_PairD: assumes "((\t. (f t, g t)) has_vderiv_on fg') T" shows "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" proof - from assms have "((\x. (f x, g x)) has_derivative (\xa. xa *\<^sub>R fg' t)) (at t within T)" if "t \ T" for t by (auto simp: has_vderiv_on_def has_vector_derivative_def that intro!: derivative_eq_intros) from diff_chain_within[OF this has_derivative_fst[OF has_derivative_ident]] diff_chain_within[OF this has_derivative_snd[OF has_derivative_ident]] show "(f has_vderiv_on (\t. fst (fg' t))) T" "(g has_vderiv_on (\t. snd (fg' t))) T" by (auto simp: has_vderiv_on_def has_vector_derivative_def o_def) qed lemma solves_autonomous_odeI: assumes "((\t. (t, phi t)) solves_ode (\t x. (1, f (fst x) (snd x)))) S (T \ X)" shows "(phi solves_ode f) S X" proof (rule solves_odeI) from solves_odeD[OF assms] have "((\t. (t, phi t)) has_vderiv_on (\t. (1, f (fst (t, phi t)) (snd (t, phi t))))) S" "\t. t \ S \ (phi t) \ X" by auto from has_vderiv_on_PairD(2)[OF this(1)] this(2) show "(phi has_vderiv_on (\t. f t (phi t))) S" "\t. t \ S \ phi t \ X" by auto qed lemma integral_solves_autonomous_odeI: fixes f::"real \ 'a::executable_euclidean_space" assumes "(phi solves_ode (\t _. f t)) {a .. b} X" "phi a = 0" assumes "a \ b" shows "(f has_integral phi b) {a .. b}" proof - have "(f has_integral phi b - phi a) {a..b}" apply (rule fundamental_theorem_of_calculus[of a b phi f]) unfolding has_vderiv_on_def[symmetric] apply fact using solves_odeD[OF assms(1)] by (simp add: has_vderiv_on_def) then show ?thesis by (simp add: assms) qed lemma zero_eq_eucl_of_list_rep_DIM: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate (DIM('a)) 0)" by (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner) lemma zero_eq_eucl_of_list_rep: "(0::'a::executable_euclidean_space) = eucl_of_list (replicate D 0)" if "D \ DIM('a)" proof - from that have "replicate D (0::real) = replicate (DIM('a)) 0 @ replicate (D - DIM('a)) 0" by (auto simp: replicate_add[symmetric]) also have "eucl_of_list \ = (eucl_of_list (replicate DIM('a) 0)::'a)" by (rule eucl_of_list_append_zeroes) also have "\ = 0" by (rule zero_eq_eucl_of_list_rep_DIM[symmetric]) finally show ?thesis by simp qed lemma one_has_ivl_integral: "((\x. 1::real) has_ivl_integral b - a) a b" using has_integral_const_real[of "1::real" a b] has_integral_const_real[of "1::real" b a] by (auto simp: has_ivl_integral_def split: if_splits) lemma Joints_aforms_of_point_self[simp]: "xs \ Joints (aforms_of_point xs)" by (simp add: aforms_of_point_def) lemma bind_eq_dRETURN_conv: "(f \ g = dRETURN S) \ (\R. f = dRETURN R \ g R = dRETURN S)" by (cases f) auto end lemma list_of_eucl_memI: "list_of_eucl (x::'x::executable_euclidean_space) \ S" if "x \ eucl_of_list ` S" "\x. x \ S \ length x = DIM('x)" using that by auto lemma Joints_aforms_of_ivls_append_point: "Joints (xs @ aforms_of_ivls p p) = (\x. x @ p) ` Joints xs" using aform.set_of_appr_of_ivl_append_point[unfolded aform_ops_def approximate_set_ops.simps] . context ode_interpretation begin theorem solves_one_step_ivl: assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, None) t1 t2 (ls, us) None" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done from t T have t: "t \ {t1 .. t2}" by auto show "t \ existence_ivl0 x0 \ flow0 x0 t \ S" by (rule one_step_result12[OF r aform.c1_info_of_appre_c0_I[OF x0] t S subset_UNIV lens]) (auto simp: aform.c1_info_invare_None lens X) qed theorem solves_one_step_ivl': assumes T: "T \ {t1 .. t2}" assumes X: "X \ {eucl_of_list lx .. eucl_of_list ux}" "length lx = DIM('a)" "length ux = DIM('a)" assumes DS: "list_interval lds uds \ list_interval ld ud" and lends: "length lds = DIM('a)*DIM('a)" "length uds = DIM('a)*DIM('a)" assumes S: "{eucl_of_list ls::'a .. eucl_of_list us} \ S" assumes lens0: "length ls = DIM('a)" "length us = DIM('a)" \ \TODO: this could be verified\ "length dx0s = DIM('a)*DIM('a)" assumes [simp]: "aform.D odo = DIM('a)" assumes r: "solves_one_step_until_time_aform optns odo ((1,1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s)) t1 t2 (ls, us) (Some (lds, uds))" shows "t \ T \ x0 \ X \ t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" proof (intro impI) assume t: "t \ T" and x0: "x0 \ X" from S have S: "set_of_lvivl (ls, us) \ S" by (auto simp: set_of_lvivl_def set_of_ivl_def) have lens: "length (fst (ls, us)) = CARD('n)" "length (snd (ls, us)) = CARD('n)" by (auto simp: lens0) have x0: "list_of_eucl x0 \ Joints (aforms_of_ivls lx ux)" apply (rule aforms_of_ivls) subgoal by (simp add: X) subgoal by (simp add: X) using subsetD[OF X(1) x0] apply (auto simp: eucl_le[where 'a='a] X) apply (metis assms(3) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) apply (metis assms(4) dim length_Basis_list list_of_eucl_eucl_of_list list_of_eucl_nth nth_Basis_list_in_Basis) done have x0dx0: "(x0, blinfun_of_list dx0s) \ aform.c1_info_of_appre ((1, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI[where x="list_of_eucl x0@dx0s"]) using lens0 apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) apply (rule x0) done from t T have t: "t \ {t1 .. t2}" by auto have DS: "blinfuns_of_lvivl' (Some (lds, uds)) \ blinfun_of_list ` list_interval ld ud" using DS by (auto simp: blinfuns_of_lvivl'_def blinfuns_of_lvivl_def) have inv: "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lds, uds))" "aform.c1_info_invare DIM('a) ((1::ereal, 1), aforms_of_ivls lx ux, Some (aforms_of_point dx0s))" by (auto simp: aform.lvivl'_invar_def lends aform.c1_info_invare_def X lens0 power2_eq_square aform.c1_info_invar_def) from one_step_result123[OF r x0dx0 t S DS lens inv \aform.D _ = _\] show "t \ existence_ivl0 x0 \ flow0 x0 t \ S \ Dflow x0 t o\<^sub>L blinfun_of_list dx0s \ blinfuns_of_lvivl (ld, ud)" by (auto simp: blinfuns_of_lvivl_def) qed end definition "zero_aforms D = map (\_. (0, zero_pdevs)) [0..pf. solves_one_step_until_time_aform (snd soptns pf) a b c d e f)" definition "solves_poincare_map_aform'_fo soptns a b c d e f g h i = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_aform' (snd soptns pf) a b c d e f g h i)" definition "solves_poincare_map_onto_aform_fo soptns a b c d e f g h = file_output (String.implode (fst soptns)) (\pf. solves_poincare_map_onto_aform (snd soptns pf) a b c d e f g h)" lemma solves_one_step_until_time_aform_foI: "solves_one_step_until_time_aform (snd optns (\_. ())) a b c d e f" if "solves_one_step_until_time_aform_fo optns a b c d e f" using that by (auto simp: solves_one_step_until_time_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_aform'_foI: "solves_poincare_map_aform' (snd optns (\_. ())) a b c d e f g h i" if "solves_poincare_map_aform'_fo optns a b c d e f g h i" using that by (auto simp: solves_poincare_map_aform'_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) lemma solves_poincare_map_onto_aform_foI: "solves_poincare_map_onto_aform (snd optns (\_. ())) a b c d e f g h" if "solves_poincare_map_onto_aform_fo optns a b c d e f g h" using that by (auto simp: solves_poincare_map_onto_aform_fo_def file_output_def Print.file_output_def print_def[abs_def] split: if_splits) definition "can_mk_ode_ops fas safe_form \ mk_ode_ops fas safe_form \ None" theorem solve_one_step_until_time_aform_integral_bounds: fixes f::"real \ 'a::executable_euclidean_space" assumes "a \ b" assumes ba: "b - a \ {t1 .. t2}" assumes a: "a \ {a1 .. a2}" assumes ls_us_subset: "{eucl_of_list ls .. eucl_of_list us} \ {l .. u}" assumes fas: "\xs::real list. length xs > 0 \ interpret_form safe_form xs \ (1::real, f (xs ! 0)) = einterpret fas xs" assumes D: "D = DIM('a) + 1" "D = CARD('i::enum)" assumes lenlu: "length ls + 1 = D" "length us + 1 = D" assumes lfas: "length fas = D" assumes mv: "can_mk_ode_ops fas safe_form" assumes FDERIV: "\xs. interpret_form safe_form xs \ isFDERIV (length fas) [0.. {l .. u}" proof - have lens0: "length ((x::real) # replicate (D - 1) 0) = DIM(real \ 'a)" for x using assms by auto have a0: "(a, 0) \ {eucl_of_list (a1 # replicate (D - 1) 0)..eucl_of_list (a2 # replicate (D - 1) 0)}" using assms by (auto simp: eucl_of_list_prod) let ?U = "{x::real\'a. interpret_form safe_form (list_of_eucl x)}" interpret ode_interpretation safe_form ?U fas "\x. (1::real, f (fst x))" "undefined::'i" apply unfold_locales subgoal using assms by simp subgoal using assms by simp subgoal using mv by (simp add: D lfas) subgoal for x apply (cases x) by (rule HOL.trans[OF fas[symmetric]]) (auto simp: fas) subgoal using mv by (simp add: can_mk_ode_ops_def) subgoal by (rule FDERIV) done have lens: "length (0 # ls) = DIM(real \ 'a)" "length (t2 # us) = DIM(real \ 'a)" "aform.D odo = DIM(real \ 'a)" using lenlu by (simp_all add: lfas aform.D_def D aform.ode_e_def ) have D_odo: "aform.D odo = DIM(real \ 'a)" by (auto simp: aform.D_def aform.ode_e_def lfas D) from solves_one_step_ivl[rule_format, OF order_refl order_refl lens0 lens0 order_refl lens(1,2) D_odo, unfolded odo_def, OF sos ba a0] have lsus: "flow0 (a, 0) (b - a) \ {eucl_of_list (0#ls)..eucl_of_list (t2#us)}" and exivl: "b - a \ existence_ivl0 (a, 0)" by auto have flow: "flow0 (a, 0) (b - a) \ UNIV \ {l..u}" using lsus apply (rule rev_subsetD) using ls_us_subset by (auto simp: eucl_of_list_prod) from ivl_subset_existence_ivl[OF exivl] \a \ b\ exivl have "0 \ existence_ivl0 (a, 0)" by (auto simp del: existence_ivl_initial_time_iff) from mem_existence_ivl_iv_defined(2)[OF this] have safe: "(a, 0::'a) \ ?U" by simp from flow_solves_ode[OF UNIV_I this] have fs: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) solves_ode (\_ x. (1, f (fst x)))) (existence_ivl0 (a, 0)) ?U" by simp with solves_odeD[OF fs] have vdp: "((\t. (fst (flow0 (a, 0) t), snd (flow0 (a, 0) t))) has_vderiv_on (\t. (1, f (fst (flow0 (a, 0) t))))) (existence_ivl0 (a, 0))" by simp have "fst (flow0 (a, 0) t) = a + t" if "t \ existence_ivl0 (a, 0)" for t proof - have "fst (flow0 (a, 0) 0) = a" using safe by auto have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) (existence_ivl0 (a, 0))" using has_vderiv_on_PairD[OF vdp] by auto then have "((\t. fst (flow0 (a, 0) t)) has_vderiv_on (\x. 1)) {0--t}" apply (rule has_vderiv_on_subset) using closed_segment_subset_existence_ivl[OF that] by auto from fundamental_theorem_of_calculus_ivl_integral[OF this, THEN ivl_integral_unique] one_has_ivl_integral[of t 0, THEN ivl_integral_unique] safe show ?thesis by auto qed with vdp have "((\t. (t, snd (flow0 (a, 0) t))) solves_ode (\t x. (1, f (a + fst x)))) (existence_ivl0 (a, 0)) ((existence_ivl0 (a, 0)) \ UNIV)" apply (intro solves_odeI) apply auto apply (auto simp: has_vderiv_on_def has_vector_derivative_def) proof goal_cases case (1 x) have r: "((\(x, y). (x - a, y::'a)) has_derivative (\x. x)) (at x within t)" for x t by (auto intro!: derivative_eq_intros) from 1 have "((\x. (a + x, snd (flow0 (a, 0) x))) has_derivative (\xa. (xa, xa *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto from has_derivative_in_compose2[OF r subset_UNIV _ this, simplified] 1 have "((\x. (x, snd (flow0 (a, 0) x))) has_derivative (\y. (y, y *\<^sub>R f (a + x)))) (at x within existence_ivl0 (a, 0))" by auto then show ?case by simp qed from solves_autonomous_odeI[OF this] have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) (existence_ivl0 (a, 0)) UNIV" by simp \ \TODO: do non-autonomous -- autonomous conversion automatically!\ then have "((\t. snd (flow0 (a, 0) t)) solves_ode (\b c. f (a + b))) {0 .. b - a} UNIV" apply (rule solves_ode_on_subset) using exivl by (rule ivl_subset_existence_ivl) (rule order_refl) from integral_solves_autonomous_odeI[OF this] have "((\b. f (a + b)) has_integral snd (flow0 (a, 0) (b - a))) (cbox 0 (b - a))" using \a \ b\ safe by auto from has_integral_affinity[OF this, where m=1 and c="-a"] have "(f has_integral snd (flow0 (a, 0) (b - a))) {a..b}" by auto then have "integral {a..b} f = snd (flow0 (a, 0) (b - a))" by blast also have "\ \ {l .. u}" using flow by auto finally show ?thesis by simp qed lemma [code_computation_unfold]: "numeral x = real_of_int (Code_Target_Int.positive x)" by simp lemma [code_computation_unfold]: "numeral x \ Float (Code_Target_Int.positive x) 0" by (simp add: Float_def float_of_numeral) definition no_print::"String.literal \ unit" where "no_print x = ()" lemmas [approximation_preproc] = list_of_eucl_real list_of_eucl_prod append.simps named_theorems DIM_simps lemmas [DIM_simps] = DIM_real DIM_prod length_nth_simps add_numeral_special add_numeral_special card_sum card_prod card_bit0 card_bit1 card_num0 card_num1 numeral_times_numeral numeral_mult mult_1_right mult_1 aform.D_def lemma numeral_refl: "numeral x = numeral x" by simp lemma plain_floatarith_approx_eq_SomeD: "approx prec fa [] = Some (the (approx prec fa []))" if "plain_floatarith 0 fa" using that by (auto dest!: plain_floatarith_approx_not_None[where p=prec and XS=Nil]) definition [simp]: "approx1 p f xs = real_of_float (lower (the (approx p f xs)))" definition [simp]: "approx2 p f xs = real_of_float (upper (the (approx p f xs)))" definition [simp]: "approx_defined p f xs \ approx p f xs \ None" definition "approxs p fs xs = those (map (\f. approx p f xs) fs)" definition [simp]: "approxs1 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o lower) y) | None \ replicate (length f) 0)" definition [simp]: "approxs2 p f xs = (case approxs p f xs of Some y \ (map (real_of_float o upper) y) | None \ replicate (length f) 0)" definition "approxs_defined p fs xs \ (those (map (\f. approx p f xs) fs) \ None)" lemma length_approxs: "length (approxs1 p f xs) = length f" "length (approxs2 p f xs) = length f" by (auto simp: approxs_def dest!: those_eq_SomeD split: option.splits) lemma real_in_approxI: "x \ {(approx1 prec fa []) .. (approx2 prec fa [])}" if "x = interpret_floatarith fa []" "approx_defined prec fa []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma real_subset_approxI: "{a .. b} \ {(approx1 prec fa []) .. (approx2 prec fb [])}" if "a = interpret_floatarith fa []" "b = interpret_floatarith fb []" "approx_defined prec fa []" "approx_defined prec fb []" using that by (force dest: approx_emptyD simp: set_of_eq) lemma approxs_eq_Some_lengthD: "length ys = length fas" if "approxs prec fas XS = Some ys" using that by (auto simp: approxs_def dest!: those_eq_SomeD) lemma approxs_pointwise: "interpret_floatarith (fas ! ia) xs \ {real_of_float (lower (ys ! ia)) .. (upper (ys ! ia))}" if "approxs prec fas XS = Some ys" "bounded_by xs XS" "ia < length fas" proof - from those_eq_SomeD[OF that(1)[unfolded approxs_def]] have ys: "ys = map (the \ (\f. approx prec f XS)) fas" and ex: "\y. i < length fas \ approx prec (fas ! i) XS = Some y" for i by auto from ex[of ia] that obtain ivl where ivl: "approx prec (fas ! ia) XS = Some ivl" by auto from approx[OF that(2) this] have "interpret_floatarith (fas ! ia) xs \\<^sub>r ivl" by auto moreover have "ys ! ia = ivl" unfolding ys apply (auto simp: o_def) apply (subst nth_map) apply (simp add: that) using ivl by simp ultimately show ?thesis using that by (auto simp: approxs_eq_Some_lengthD set_of_eq split: prod.splits) qed lemmas approxs_pointwise_le = approxs_pointwise[simplified, THEN conjunct1] and approxs_pointwise_ge = approxs_pointwise[simplified, THEN conjunct2] lemma approxs_eucl: "eucl_of_list (interpret_floatariths fas xs) \ {eucl_of_list (map lower ys) .. eucl_of_list (map upper ys)::'a::executable_euclidean_space}" if "approxs prec fas XS = Some ys" "length fas = DIM('a)" "bounded_by xs XS" using that by (auto simp: eucl_le[where 'a='a] eucl_of_list_inner o_def approxs_eq_Some_lengthD intro!: approxs_pointwise_le approxs_pointwise_ge) lemma plain_floatariths_approx_eq_SomeD: "approxs prec fas [] = Some (the (approxs prec fas []))" if "list_all (plain_floatarith 0) fas" using that apply (induction fas) apply (auto simp: approxs_def split: option.splits dest!: plain_floatarith_approx_eq_SomeD) subgoal for a fas aa apply (cases "those (map (\f. approx prec f []) fas)") by auto done lemma approxs_definedD: "approxs prec fas xs = Some (the (approxs prec fas xs))" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approxs_defined_ne_None[simp]: "approxs prec fas xs \ None" if "approxs_defined prec fas xs" using that by (auto simp: approxs_defined_def approxs_def) lemma approx_subset_euclI: "{eucl_of_list (approxs2 prec fals [])::'a .. eucl_of_list (approxs1 prec faus [])} \ {l .. u}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma eucl_subset_approxI: "{l .. u} \ {eucl_of_list (approxs1 prec fals [])::'a .. eucl_of_list (approxs2 prec faus [])}" if "list_of_eucl l = interpret_floatariths fals []" and "list_of_eucl u = interpret_floatariths faus []" and "length fals = DIM('a::executable_euclidean_space)" and "length faus = DIM('a::executable_euclidean_space)" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that by (auto intro!: bounded_by_Nil dest!: approxs_eucl[where 'a='a] list_of_eucl_eqD plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) lemma approx_subset_listI: "list_interval (approxs2 prec fals []) (approxs1 prec faus []) \ list_interval l u" if "l = interpret_floatariths fals []" and "u = interpret_floatariths faus []" and "length fals = length l" and "length faus = length u" and "approxs_defined prec fals []" and "approxs_defined prec faus []" using that apply (auto simp: list_interval_def list_all2_conv_all_nth dest: approxs_eq_Some_lengthD intro!: bounded_by_Nil dest!: plain_floatariths_approx_eq_SomeD[where prec=prec] split: option.splits) apply (rule order_trans) apply (rule approxs_pointwise_ge) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) apply (subst interpret_floatariths_nth) apply (auto dest: approxs_eq_Some_lengthD) apply (rule approxs_pointwise_le[ge]) apply assumption apply (rule bounded_by_Nil) apply (auto dest: approxs_eq_Some_lengthD) done definition "unit_list D n = (replicate D 0)[n:=1]" definition "mirror_sctn (sctn::real list sctn) = Sctn (map uminus (normal sctn)) (- pstn sctn)" definition "mirrored_sctn b (sctn::real list sctn) = (if b then mirror_sctn sctn else sctn)" lemma mirror_sctn_simps[simp]: "pstn (mirror_sctn sctn) = - pstn sctn" "normal (mirror_sctn sctn) = map uminus (normal sctn)" by (cases sctn) (auto simp: mirror_sctn_def) lemma length_unit_list[simp]: "length (unit_list D n) = D" by (auto simp: unit_list_def) lemma eucl_of_list_unit_list[simp]: "(eucl_of_list (unit_list D n)::'a::executable_euclidean_space) = Basis_list ! n" if "D = DIM('a)" "n < D" using that by (auto simp: unit_list_def eucl_of_list_inner inner_Basis nth_list_update' intro!: euclidean_eqI[where 'a='a]) lemma le_eucl_of_list_iff: "(t::'a::executable_euclidean_space) \ eucl_of_list uX0 \ (\i. i < DIM('a) \ t \ Basis_list ! i \ uX0 ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma eucl_of_list_le_iff: "eucl_of_list uX0 \ (t::'a::executable_euclidean_space) \ (\i. i < DIM('a) \ uX0 ! i \ t \ Basis_list ! i)" if "length uX0 = DIM('a)" using that apply (auto simp: eucl_le[where 'a='a] eucl_of_list_inner) using nth_Basis_list_in_Basis apply force by (metis Basis_list in_Basis_index_Basis_list index_less_size_conv length_Basis_list) lemma Joints_aforms_of_ivls: "Joints (aforms_of_ivls lX0 uX0) = list_interval lX0 uX0" if "list_all2 (\) lX0 uX0" using that apply (auto simp: list_interval_def dest: Joints_aforms_of_ivlsD1[OF _ that] Joints_aforms_of_ivlsD2[OF _ that] list_all2_lengthD intro!: aforms_of_ivls) by (auto simp: list_all2_conv_all_nth) lemma list_of_eucl_in_list_interval_iff: "list_of_eucl x0 \ list_interval lX0 uX0 \ x0 \ {eucl_of_list lX0 .. eucl_of_list uX0::'a}" if "length lX0 = DIM('a::executable_euclidean_space)" "length uX0 = DIM('a::executable_euclidean_space)" using that by (auto simp: list_interval_def eucl_of_list_le_iff le_eucl_of_list_iff list_all2_conv_all_nth) text \TODO: make a tactic out of this?!\ lemma file_output_iff: "file_output s f = f (\_. ())" by (auto simp: file_output_def print_def[abs_def] Print.file_output_def) context ode_interpretation begin lemma poincare_mapsto_subset: "poincare_mapsto P X0 SS CX R" if "poincare_mapsto P' Y0 RR CZ S" "X0 \ Y0" "CZ \ CX" "S \ R" "RR = SS" "P = P'" using that by (force simp: poincare_mapsto_def) theorem solves_poincare_map_aform'_derivI: assumes solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (Sctn (unit_list D n) (lP ! n)) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and dim: "aform.D odo = DIM('a)" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and guards: "(\a xs b ba ro. (xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a))" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and SS: "SS = {x::'a. x \ Basis_list ! n \ lP ! n}" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (Sctn (unit_list D n) (lP ! n))) = DIM('a)" by (auto simp: D) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) auto then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed definition guards_invar::"nat \ (((real list \ real list) \ real list sctn) list \ (real \ real pdevs) reach_options) list \ bool" where "guards_invar D guards = (\(xs, ro) \ set guards. \((a, b), ba) \ set xs. length a = D \ length b = D \ length (normal ba) = D)" theorem solves_poincare_map_aform'I: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_aform'_fo optns odo (Sctn (unit_list D n) (lP ! n)) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map P x \ R" proof - note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (((Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) (below_halfspace (map_sctn eucl_of_list (Sctn (unit_list D n) (lP ! n)))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_aform'[OF solves, OF 1 dim 2 3 4 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) (below_halfspace (map_sctn eucl_of_list (((Sctn (unit_list D n) (lP ! n)))))) UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal using assms by (simp add: below_halfspace_def le_halfspace_def[abs_def]) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map P x \ R" by (auto simp: poincare_mapsto_def) qed definition "poincare_map_from_outside = poincare_map" theorem poincare_maps_onto_aformI: assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" assumes D: "D = DIM('a)" assumes guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and dim: "aform.D odo = DIM('a)" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" and solves: "solves_poincare_map_onto_aform_fo optns odo guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, None)] (lR, uR) None" shows "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" proof - note solves = solves[unfolded solves_poincare_map_onto_aform_fo_def file_output_iff] have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, None)] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens) have 2: "length (normal ((mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) from guards have guards: "(xs, ro) \ set guards \ ((a, b), ba) \ set xs \ length a = DIM('a) \ length b = DIM('a) \ length (normal ba) = DIM('a)" for xs ro a b ba by (auto simp: guards_invar_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) None" by (auto simp: lens) have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, None)]) UNIV (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' None)" by (rule solves_poincare_map_onto_aform[OF solves, OF 1 dim 2 3 guards 5]) then have "poincare_mapsto P (X0 \ UNIV::('a \ ('a \\<^sub>L 'a)) set) UNIV UNIV (R \ UNIV)" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_apprse_def aform.c1_info_of_appre_def aform.c1_info_of_appr_def) apply (rule image_eqI) apply (rule eucl_of_list_list_of_eucl[symmetric]) apply (rule aforms_of_ivls) by (auto simp add: lens subset_iff le_eucl_of_list_iff eucl_of_list_le_iff) subgoal by simp subgoal using R by (auto simp: set_of_lvivl_def set_of_ivl_def) subgoal by simp subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "\x\X0. returns_to P x \ poincare_map_from_outside P x \ R" by (auto simp: poincare_mapsto_def poincare_map_from_outside_def) qed end lemmas [simp] = length_approxs context includes ode_ops.lifting begin lift_definition empty_ode_ops::"ode_ops" is "([], true_form)" by auto end ML \val ode_numerics_conv = @{computation_check terms: Trueprop Not solves_one_step_until_time_aform_fo solves_poincare_map_aform'_fo solves_poincare_map_onto_aform_fo num_options num_options_c1 ro (* nat *) Suc "0::nat" "1::nat" "(+)::nat \ nat \ nat" "(-) ::nat \ nat \ nat" "(=)::nat\nat\bool" "(^)::nat\nat\nat" (* int / integer*) "(=)::int\int\bool" "(+)::int\int\int" "uminus::_\int" "uminus::_\integer" int_of_integer integer_of_int "0::int" "1::int" "(^)::int\nat\int" (* real *) "(=)::real\real\bool" "real_of_float" "(/)::real\real\real" "(^)::real\nat\real" "uminus::real\_" "(+)::real\real\real" "(-)::real\real\real" "(*)::real\real\real" real_divl real_divr real_of_int "0::real" "1::real" (* rat *) Fract "0::rat" "1::rat" "(+)::rat\rat\rat" "(-)::rat\rat\rat" "(*)::rat\rat\rat" "uminus::rat\_" "(/)::rat\rat\rat" "(^)::rat\nat\rat" (* ereal *) "1::ereal" (* lists: *) "replicate::_\real\_" "unit_list::_\_\real list" "Nil::(nat \ nat \ string) list" "Cons::_\_\(nat \ nat \ string) list" "Nil::(nat \ nat \ string \ nat list) list" "Cons::_\_\(nat \ nat \ string \ nat list) list" "Nil::real list" "Cons::_\_\real list" "Nil::nat list" "Cons::_\_\nat list" "Nil::string list" "Cons::_\_\string list" "Nil::real aform list" "Cons::_\_\real aform list" "Nil::(float interval) option list" "Cons::_\_\(float interval) option list" "nth::_\_\real" "upt" (* products: *) "Pair::_\_\(nat \ string)" "Pair::_\_\(nat \ nat \ string)" "Pair::_\_\char list \ nat list" "Pair::_\_\nat \ char list \ nat list" "Pair::_\_\nat \ nat \ char list \ nat list" "Pair::_\_\char list \ ((String.literal \ unit) \ (real \ real pdevs) numeric_options)" "Pair::_\_\ereal\ereal" "Pair::_\_\real aform list \ real aform list option" "Pair::_\_\(ereal \ ereal) \ real aform list \ real aform list option" "Pair::_\_\real aform" "Pair::_\_\real list \ real list" "Nil::(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Cons::_\_\(((real list \ real list) \ real list sctn) list \ (real aform) reach_options) list" "Nil::((real list \ real list) \ real list sctn) list" "Cons::_\_\((real list \ real list) \ real list sctn) list" "Pair::_\_\((real list \ real list) \ real list sctn) list \ real aform reach_options" "Nil::((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" "Cons::_\_\((ereal \ ereal) \ (real aform) list \ (real aform) list option) list" (* option *) "None::(real aform) list option" "Some::_\(real aform) list option" "None::(real list \ real list) option" "Some::_\(real list \ real list) option" (* aforms *) "aform_of_ivl::_\_\real aform" aforms_of_ivls aforms_of_point (* pdevs*) "zero_pdevs::real pdevs" "zero_aforms::_ \ real aform list" (* ode_ops *) mk_ode_ops init_ode_ops empty_ode_ops can_mk_ode_ops "the::ode_ops option \ ode_ops" the_odo (* Characters/Strings *) String.Char String.implode "Nil::string" "Cons::_\_\string" (* float *) "(=)::float\float\bool" "(+)::float\float\float" "uminus::_\float" "(-)::_\_\float" Float float_of_int float_of_nat (* approximation... *) approx approx1 approx2 approxs1 approxs2 approx_defined approxs_defined (* floatarith *) "0::floatarith" "1::floatarith" "(+)::_\_\floatarith" "(-)::_\_\floatarith" "(*)::_\_\floatarith" "(/)::_\_\floatarith" "inverse::_\floatarith" "uminus::_\floatarith" "Sum\<^sub>e::_\nat list\floatarith" Sin Half Tan R\<^sub>e Norm (* form *) true_form Half (* Printing *) print no_print (* sections *) xsec xsec' ysec ysec' zsec zsec' xsec2 xsec2' ysec2 ysec2' ysec4 ysec4' mirrored_sctn Code_Target_Nat.natural Code_Target_Int.positive Code_Target_Int.negative Code_Numeral.positive Code_Numeral.negative datatypes: bool num floatarith "floatarith list" form "real list sctn" "real \ real" } \ ML \fun ode_numerics_tac ctxt = CONVERSION (ode_numerics_conv ctxt) THEN' (resolve_tac ctxt [TrueI])\ lemma eq_einterpretI: assumes "list_of_eucl (VS::'a::executable_euclidean_space) = interpret_floatariths fas xs" assumes "length fas = DIM('a)" shows "VS = eucl_of_list (interpret_floatariths fas xs)" using assms apply (subst (asm) list_of_eucl_eucl_of_list[symmetric]) apply (auto intro!: ) by (metis eucl_of_list_list_of_eucl) lemma one_add_square_ne_zero[simp]: "1 + t * t \ 0" for t::real by (metis semiring_normalization_rules(12) sum_squares_eq_zero_iff zero_neq_one) lemma abs_rat_bound: "abs (x - y) \ e / f" if "y \ {yl .. yu}" "x \ {yu - real_divl p e f.. yl + real_divl p e f}" for x y e::real proof - note \x \ _\ also have "{yu - real_divl p e f.. yl + real_divl p e f} \ {yu - e / f .. yl + e / f}" by (auto intro!: diff_mono real_divl) finally show ?thesis using that unfolding abs_diff_le_iff by auto qed lemma in_ivl_selfI: "a \ {a .. a}" for a::real by auto lemma pi4_bnds: "pi / 4 \ {real_divl 80 (lb_pi 80) 4 .. real_divr 80 (ub_pi 80) 4}" using pi_boundaries[of 80] unfolding atLeastAtMost_iff by (intro conjI real_divl[le] real_divr[ge] divide_right_mono) auto lemma abs_minus_leI: "\x - x'\ \ e" if "x \ {x' - e .. x' + e}" for x e::real using that by auto lemmas [DIM_simps] = Suc_numeral One_nat_def[symmetric] TrueI Suc_1 length_approxs arith_simps lemma (in ode_interpretation) length_ode_e[DIM_simps]: "length (ode_expression odo) = DIM('a)" by (auto simp: len) named_theorems solves_one_step_ivl_thms context ode_interpretation begin lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl[OF _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = TAG_optns[OF solves_one_step_ivl'[OF _ _ _ _ _ _ _ _ _ _ _ _ solves_one_step_until_time_aform_foI], rotated -1, of optns _ _ _ _ _ _ _ _ _ _ _ _ _ _ optns for optns] lemmas [solves_one_step_ivl_thms] = solves_poincare_map_aform'I poincare_maps_onto_aformI end lemma TAG_optnsI: "TAG_optns optns" by simp named_theorems poincare_tac_theorems lemmas [DIM_simps] = one_less_numeral_iff rel_simps abbreviation "point_ivl a \ {a .. a}" lemma isFDERIV_compute: "isFDERIV D vs fas xs \ (list_all (\i. list_all (\j. isDERIV (vs ! i) (fas ! j) xs) [0.. length fas = D \ length vs = D" unfolding isFDERIV_def by (auto simp: list.pred_set) theorem (in ode_interpretation) solves_poincare_map_aform'_derivI'[solves_one_step_ivl_thms]: \ \TODO: replace @{thm solves_poincare_map_aform'_derivI}\ assumes "TAG_optns optns" assumes "TAG_reach_optns roi" assumes "TAG_sctn mirrored" and D: "D = DIM('a)" assumes DS: "list_interval lDR uDR \ list_interval lDS uDS" and ode_fas: "aform.D odo = DIM('a)" and guards: "guards_invar DIM('a) guards" and P: "P = {eucl_of_list lP .. eucl_of_list uP}" and plane: "uP ! n = lP ! n" and X0: "X0 \ {eucl_of_list lX0 .. eucl_of_list uX0}" and nD: "n < DIM('a)" and R: "{eucl_of_list lR .. eucl_of_list uR} \ R" and lens: "length (lP) = DIM('a)" "length (uP) = DIM('a)" "length (lX0) = DIM('a)" "length (uX0) = DIM('a)" "length (lR) = DIM('a)" "length (uR) = DIM('a)" "length DX0 = DIM('a)*DIM('a)" "length lDR = CARD('n) * CARD('n)" "length uDR = CARD('n) * CARD('n)" and SS: "SS = {x::'a. if mirrored then x \ Basis_list ! n \ lP ! n else x \ Basis_list ! n \ lP ! n}" assumes solves: "solves_poincare_map_aform'_fo optns odo (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))) guards (lP, uP) (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))) roi [((1,1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] (lR, uR) (Some (lDR, uDR))" shows "\x\X0. returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" proof (rule ballI) fix x assume "x \ X0" then have la2: "list_all2 (\) lX0 uX0" using X0 by (force simp: subset_iff eucl_of_list_le_iff le_eucl_of_list_iff lens list_all2_conv_all_nth) have 1: "\X. X \ set [((1::ereal, 1::ereal), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))] \ aform.c1_info_invare DIM('a) X" for X by (auto simp: aform.c1_info_invare_def aform.c1_info_invar_def lens power2_eq_square) have 2: "length (normal (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 3: "length (fst (lP, uP)) = DIM('a)" "length (snd (lP, uP)) = DIM('a)" by (auto simp: lens) have 4: "length (normal (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n)))) = DIM('a)" by (auto simp: D mirrored_sctn_def) have 5: "length (fst (lR, uR)) = CARD('n)" "length (snd (lR, uR)) = CARD('n)" "aform.lvivl'_invar (CARD('n) * CARD('n)) (Some (lDR, uDR))" by (auto simp: lens aform.lvivl'_invar_def) note solves = solves[unfolded solves_poincare_map_aform'_fo_def file_output_iff] have "poincare_mapsto (set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (aform.c1_info_of_apprse [((1, 1), aforms_of_ivls lX0 uX0, Some (aforms_of_point DX0))]) (below_halfspace (map_sctn eucl_of_list (mirrored_sctn (\mirrored) (Sctn (unit_list D n) (lP ! n))))) (aform.Csafe odo - set_of_lvivl (lP, uP) \ plane_of (map_sctn eucl_of_list (mirrored_sctn mirrored (Sctn (unit_list D n) (lP ! n))))) (set_of_lvivl (lR, uR) \ blinfuns_of_lvivl' (Some (lDR, uDR)))" by (rule solves_poincare_map_aform'[OF solves, OF 1 ode_fas 4 3 2 _ 5]) (use guards in \auto simp: guards_invar_def\) then have "poincare_mapsto P (X0 \ {blinfun_of_list DX0}::('a \ ('a \\<^sub>L 'a)) set) SS UNIV (R \ blinfuns_of_lvivl (lDS, uDS))" apply (rule poincare_mapsto_subset) subgoal using X0 apply (auto simp: aform.c1_info_of_appre_def aform.c1_info_of_appr_def aform.c1_info_of_apprse_def) subgoal for x0 apply (rule image_eqI[where x="list_of_eucl x0@DX0"]) using lens apply (auto simp: flow1_of_list_def aforms_of_point_def Joints_aforms_of_ivls_append_point) apply (rule imageI) using X0 by (auto simp: Joints_aforms_of_ivls la2 list_of_eucl_in_list_interval_iff) done subgoal by simp subgoal using R DS by (auto simp: set_of_lvivl_def set_of_ivl_def blinfuns_of_lvivl'_def blinfuns_of_lvivl_def lens) subgoal using assms by (auto simp: below_halfspace_def le_halfspace_def[abs_def] mirrored_sctn_def mirror_sctn_def) subgoal using assms by (fastforce simp add: P set_of_lvivl_def set_of_ivl_def plane_of_def le_eucl_of_list_iff eucl_of_list_le_iff mirrored_sctn_def mirror_sctn_def) done then show "returns_to P x \ return_time P differentiable at x within SS \ (\D. (poincare_map P has_derivative blinfun_apply D) (at x within SS) \ poincare_map P x \ R \ D o\<^sub>L blinfun_of_list DX0 \ blinfuns_of_lvivl (lDS, uDS))" using \x \ X0\ by (auto simp: poincare_mapsto_def) qed lemmas [DIM_simps] = aform.ode_e_def ML \ structure ODE_Numerics_Tac = struct fun mk_nat n = HOLogic.mk_number @{typ nat} n fun mk_int n = HOLogic.mk_number @{typ int} n fun mk_integer n = @{term integer_of_int} $ (HOLogic.mk_number @{typ int} n) fun mk_bool b = if b then @{term True} else @{term False} fun mk_numeralT n = let fun mk_bit 0 ty = Type (@{type_name bit0}, [ty]) | mk_bit 1 ty = Type (@{type_name bit1}, [ty]); fun bin_of n = if n = 1 then @{typ num1} else if n = 0 then @{typ num0} else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r (bin_of q) end; in bin_of n end; fun print_tac' ctxt s = K (print_tac ctxt s) val using_master_directory = File.full_path o Resources.master_directory o Proof_Context.theory_of; fun using_master_directory_term ctxt s = (if s = "-" orelse s = "" then s else Path.explode s |> using_master_directory ctxt |> Path.implode) |> HOLogic.mk_string fun real_in_approx_tac ctxt p = let val inst_approx = (TVars.empty, Vars.make [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_in_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt end fun real_subset_approx_tac ctxt p = let val inst_approx = (TVars.empty, Vars.make [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) val approx_thm = Thm.instantiate inst_approx @{thm real_subset_approxI} in resolve_tac ctxt [approx_thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' ode_numerics_tac ctxt THEN' ode_numerics_tac ctxt end fun basic_nt_ss ctxt nt = put_simpset HOL_basic_ss ctxt addsimps Named_Theorems.get ctxt nt fun DIM_tac defs ctxt = (Simplifier.simp_tac (basic_nt_ss ctxt @{named_theorems DIM_simps} addsimps defs)) fun subset_approx_preconds_tac ctxt p thm = let val inst_approx = (TVars.empty, Vars.make [((("prec", 0), @{typ nat}), mk_nat p |> Thm.cterm_of ctxt)]) in resolve_tac ctxt [Thm.instantiate inst_approx thm] THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (reify_floatariths_tac ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (DIM_tac [] ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) THEN' SOLVED' (ode_numerics_tac ctxt) end val cfg_trace = Attrib.setup_config_bool @{binding ode_numerics_trace} (K false) fun tracing_tac ctxt = if Config.get ctxt cfg_trace then print_tac ctxt else K all_tac fun tracing_tac' ctxt = fn s => K (tracing_tac ctxt s) fun eucl_subset_approx_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm eucl_subset_approxI} fun approx_subset_eucl_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_euclI} fun approx_subset_list_tac ctxt p = subset_approx_preconds_tac ctxt p @{thm approx_subset_listI} val static_simpset = Simplifier.simpset_of @{context} fun nth_tac ctxt = Simplifier.simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms nth_Cons_0 nth_Cons_Suc numeral_nat}) fun nth_list_eq_tac ctxt n = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => case try (Thm.term_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) concl of SOME (@{const List.nth(real)} $ xs $ Var _, @{const List.nth(real)} $ ys $ Var _) => let val i = find_index (op=) (HOLogic.dest_list xs ~~ HOLogic.dest_list ys) val thm = Goal.prove context [] [] (HOLogic.mk_eq (@{const List.nth(real)} $ xs $ HOLogic.mk_number @{typ nat} i, @{const List.nth(real)} $ ys $ HOLogic.mk_number @{typ nat} i) |> HOLogic.mk_Trueprop) (fn {context, ...} => HEADGOAL (nth_tac context)) in SOLVE (HEADGOAL (resolve_tac context [thm])) end | _ => no_tac ) ctxt n fun numeric_precond_step_tac defs thms p = Subgoal.FOCUS_PARAMS (fn {context, concl, ...} => let val prems = Logic.strip_imp_prems (Thm.term_of concl) val conclusion = Logic.strip_imp_concl (Thm.term_of concl) in (case conclusion |> HOLogic.dest_Trueprop of @{const Set.member(real)} $ _ $ _ => tracing_tac context "numeric_precond_step: real in approx" THEN HEADGOAL (real_in_approx_tac context p) | Const(@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ _ $ _) $ (Const (@{const_name atLeastAtMost}, _) $ Var _ $ Var _) => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (real_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) $ _ => tracing_tac context "numeric_precond_step: approx subset eucl" THEN HEADGOAL (approx_subset_eucl_tac context p) | Const (@{const_name less_eq}, _) $ _ $ (Const (@{const_name atLeastAtMost}, _) $ (Const (@{const_name eucl_of_list}, _) $ Var _) $ _) => tracing_tac context "numeric_precond_step: eucl subset approx" THEN HEADGOAL (eucl_subset_approx_tac context p) | Const (@{const_name less_eq}, _) $ (@{const list_interval(real)} $ _ $ _) $ (@{const list_interval(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: approx subset list" THEN HEADGOAL (approx_subset_list_tac context p) | @{const HOL.eq(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const less(nat)} $ _ $ _ => tracing_tac context "numeric_precond_step: DIM_tac" THEN HEADGOAL (SOLVED' (DIM_tac [] context)) | @{const HOL.eq(real)} $ (@{const nth(real)} $ _ $ _) $ (@{const nth(real)} $ _ $ _) => tracing_tac context "numeric_precond_step: nth_list_eq_tac" THEN HEADGOAL (SOLVED' (nth_list_eq_tac context)) | Const (@{const_name "HOL.eq"}, _) $ _ $ (Const (@{const_name eucl_of_list}, _) $ (@{const interpret_floatariths} $ _ $ _)) => tracing_tac context "numeric_precond_step: reify floatariths" THEN HEADGOAL (resolve_tac context @{thms eq_einterpretI} THEN' reify_floatariths_tac context) | t as _ $ _ => let val (c, args) = strip_comb t in if member (op=) [@{const "solves_one_step_until_time_aform_fo"}, @{const "solves_poincare_map_aform'_fo"}, @{const "solves_poincare_map_onto_aform_fo"}, @{const "can_mk_ode_ops"} ] c then tracing_tac context "numeric_precond_step: ode_numerics_tac" THEN HEADGOAL ( CONVERSION (Simplifier.rewrite (put_simpset HOL_basic_ss context addsimps defs)) THEN' tracing_tac' context "numeric_precond_step: ode_numerics_tac (unfolded)" THEN' ode_numerics_tac context) else if member (op=) [@{const "isFDERIV"}] c then tracing_tac context "numeric_precond_step: isFDERIV" THEN HEADGOAL (SOLVED'(Simplifier.asm_full_simp_tac (put_simpset static_simpset context addsimps (@{thms isFDERIV_def less_Suc_eq_0_disj isDERIV_Power_iff} @ thms @ defs)) THEN' tracing_tac' context "numeric_precond_step: simplified isFDERIV" )) else tracing_tac context "numeric_precond_step: boolean, try thms" THEN HEADGOAL (SOLVED' (resolve_tac context thms)) end | _ => tracing_tac context "numeric_precond_step: boolean constant" THEN no_tac ) end) fun integral_bnds_tac_gen_start sstep d p m N atol filename ctxt i = let val insts = (TVars.make [((("'i", 0), @{sort "{enum}"}), mk_numeralT (d + 1) |> Thm.ctyp_of ctxt)], Vars.make [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, (@{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ @{term "[(0::nat, 1::nat, ''0x000000'')]"})) |> Thm.cterm_of ctxt), ((("safe_form", 0), @{typ form}), @{cterm true_form}) ]) in resolve_tac ctxt [Thm.instantiate insts @{thm solve_one_step_until_time_aform_integral_bounds}] i THEN (Lin_Arith.tac ctxt i ORELSE Simplifier.simp_tac ctxt i) end fun integral_bnds_tac_gen sstep d p m N atol filename thms ctxt = integral_bnds_tac_gen_start sstep d p m N atol filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac [] thms p ctxt) val integral_bnds_tac = integral_bnds_tac_gen 5 fun mk_proj (m, n, s) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s] fun mk_projs projs = HOLogic.mk_list @{typ "nat \ nat \ string"} (map mk_proj projs) fun mk_string_list ds = HOLogic.mk_list @{typ "string"} (map HOLogic.mk_string ds) fun mk_nat_list ds = HOLogic.mk_list @{typ "nat"} (map mk_nat ds) fun mk_proj_c1 (m, n, s, ds) = HOLogic.mk_tuple [mk_nat m, mk_nat n, HOLogic.mk_string s, mk_nat_list ds] fun mk_projs_c1 projs = HOLogic.mk_list @{typ "nat \ nat \ string \ nat list"} (map mk_proj_c1 projs) fun TAG_optns_thm p sstep m N atol projs filename ctxt = Thm.instantiate (TVars.empty, Vars.make [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs projs) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt = Thm.instantiate (TVars.empty, Vars.make [((("optns", 0), @{typ "string \ ((String.literal \ unit) \(real aform) numeric_options)"}), HOLogic.mk_prod (using_master_directory_term ctxt filename, @{term num_options_c1} $ mk_nat p $ mk_int sstep $ mk_nat m $ mk_nat N $ mk_int atol $ mk_projs_c1 projs $ mk_string_list ds) |> Thm.cterm_of ctxt)]) @{thm TAG_optnsI} fun ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun ode_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = ode_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode_bnds_tac = ode_bnds_tac_gen 5 fun ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_c1_thm p sstep m N atol projs ds filename ctxt] fun ode'_bnds_tac_gen sstep ode_defs p m N atol projs ds filename ctxt = ode'_bnds_tac_gen_start sstep p m N atol projs ds filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD (numeric_precond_step_tac ode_defs [] p ctxt) val ode'_bnds_tac = ode'_bnds_tac_gen 5 fun poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt = tracing_tac' ctxt "solves_one_step_ivl_thms" THEN' resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' tracing_tac' ctxt "resolved solves_one_step_ivl_thms" THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare_bnds_tac = poincare_bnds_tac_gen 5 fun poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt = resolve_tac ctxt (Named_Theorems.get ctxt @{named_theorems solves_one_step_ivl_thms}) THEN' resolve_tac ctxt [TAG_optns_thm p sstep m N atol projs filename ctxt] fun poincare'_bnds_tac_gen sstep ode_defs p m N atol projs filename ctxt = poincare'_bnds_tac_gen_start sstep p m N atol projs filename ctxt THEN_ALL_NEW_FWD REPEAT_ALL_NEW_FWD ( numeric_precond_step_tac ode_defs (Named_Theorems.get ctxt @{named_theorems poincare_tac_theorems}) p ctxt) val poincare'_bnds_tac = poincare'_bnds_tac_gen 5 end \ lemma (in auto_ll_on_open) Poincare_Banach_fixed_pointI: assumes "convex S" and c: "complete S" "S \ {}" and "S \ T" assumes derivative_bounded: "\x\S. poincare_map \ x \ S \ (\D. (poincare_map \ has_derivative D) (at x within T) \ onorm D \ B)" assumes B: "B < 1" shows "\!x. x \ S \ poincare_map \ x = x" using c _ B proof (rule banach_fix) from derivative_bounded c show "0 \ B" by (auto dest!: has_derivative_bounded_linear onorm_pos_le) from derivative_bounded show "poincare_map \ ` S \ S" by auto obtain D where D: "\x \ S. (poincare_map \ has_derivative D x) (at x within T) \ onorm (D x) \ B" apply atomize_elim apply (rule bchoice) using derivative_bounded by auto with \S \ T\ have "(\x. x \ S \ (poincare_map \ has_derivative D x) (at x within S))" by (auto intro: has_derivative_subset) from bounded_derivative_imp_lipschitz[of S "poincare_map \" D B, OF this] \convex S\ D c \0 \ B\ have "B-lipschitz_on S (poincare_map \)" by auto then show "\x\S. \y\S. dist (poincare_map \ x) (poincare_map \ y) \ B * dist x y" by (auto simp: lipschitz_on_def) qed ML \open ODE_Numerics_Tac\ lemma isFDERIV_product: "isFDERIV n xs fas vs \ length fas = n \ length xs = n \ list_all (\(x, f). isDERIV x f vs) (List.product xs fas)" apply (auto simp: isFDERIV_def list_all2_iff in_set_zip list_all_length product_nth) apply auto - apply (metis gr_implies_not_zero gr_zeroI less_mult_imp_div_less pos_mod_bound) + apply (metis add_lessD1 less_mult_imp_div_less mod_less_divisor add_0) done end diff --git a/thys/Padic_Field/Padic_Field_Powers.thy b/thys/Padic_Field/Padic_Field_Powers.thy --- a/thys/Padic_Field/Padic_Field_Powers.thy +++ b/thys/Padic_Field/Padic_Field_Powers.thy @@ -1,11439 +1,11441 @@ theory Padic_Field_Powers imports Ring_Powers Padic_Field_Polynomials Generated_Boolean_Algebra Padic_Field_Topology begin text\This theory is intended to develop the necessary background on subsets of powers of a $p$-adic field to prove Macintyre's quantifier elimination theorem. In particular, we define semi-algebraic subsets of $\mathbb{Q}_p^n$, semi-algebraic functions $\mathbb{Q}_p^n \to \mathbb{Q}_p$, and semi- algebraic mappings $\mathbb{Q}_p^n \to \mathbb{Q}_p^m$ for arbitrary $n, m \in \mathbb{N}$. In addition we prove that many common sets and functions are semi-algebraic. We are closely following the paper \cite{denef1986} by Denef, where an algebraic proof of Mactinyre's theorem is developed.\ (**************************************************************************************************) (**************************************************************************************************) section\Cartesian Powers of $p$-adic Fields\ (**************************************************************************************************) (**************************************************************************************************) lemma list_tl: "tl (t#x) = x" using List.list.sel(3) by auto lemma list_hd: "hd (t#x) = t" unfolding List.list.sel(1) by auto sublocale padic_fields < cring_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding cring_coord_rings_axioms_def cring_coord_rings_def using Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.is_UP_cring) by auto sublocale padic_fields < Qp: domain_coord_rings Q\<^sub>p "UP Q\<^sub>p" unfolding domain_coord_rings_def cring_coord_rings_axioms_def cring_coord_rings_def using Qp.domain_axioms Qp.zero_not_one UPQ.R_cring apply (simp add: UPQ.UP_cring_axioms) by auto context padic_fields begin no_notation Zp.to_fun (infixl\\\ 70) no_notation ideal_prod (infixl "\\" 80) notation evimage (infixr "\\" 90) and euminus_set ("_ \<^sup>c\" 70) type_synonym padic_tuple = "padic_number list" type_synonym padic_function = "padic_number \ padic_number" type_synonym padic_nary_function = "padic_tuple \ padic_number" type_synonym padic_function_tuple = "padic_nary_function list" type_synonym padic_nary_function_poly = "nat \ padic_nary_function" (**************************************************************************************************) (**************************************************************************************************) subsection\Polynomials over $\mathbb{Q}_p$ and Polynomial Maps\ (**************************************************************************************************) (**************************************************************************************************) lemma last_closed': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "t \ carrier Q\<^sub>p" using assms last_closed[of n "x@[t]" Q\<^sub>p] by (metis (full_types) cartesian_power_car_memE gr0I last_snoc length_append_singleton less_not_refl zero_less_Suc) lemma segment_in_car': assumes "x@[t] \ carrier (Q\<^sub>p\<^bsup>Suc n\<^esup>)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" proof- have 0: "length x = n" by (metis Suc_inject assms cartesian_power_car_memE length_append_singleton) have "set x \ set (x@[t])" by (metis rotate1.simps(2) set_rotate1 set_subset_Cons) then have 1: "set x \ carrier Q\<^sub>p" using assms cartesian_power_car_memE''[of "x@[t]" Q\<^sub>p "Suc n"] by blast show ?thesis using 0 1 assms cartesian_power_car_memI[of x n Q\<^sub>p] by blast qed lemma Qp_zero: "Q\<^sub>p\<^bsup>0\<^esup> = nil_ring" unfolding cartesian_power_def by simp lemma Qp_zero_carrier: "carrier (Q\<^sub>p\<^bsup>0\<^esup>) = {[]}" by (simp add: Qp_zero) text\Abbreviation for constant polynomials\ abbreviation(input) Qp_to_IP where "Qp_to_IP k \ Qp.indexed_const k" lemma Qp_to_IP_car: assumes "k \ carrier Q\<^sub>p" shows "Qp_to_IP k \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Qp.indexed_const_closed by blast lemma(in cring_coord_rings) smult_closed: assumes "a \ carrier R" assumes "q \ carrier (R[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>R[\\<^bsub>n\<^esub>]\<^esub> q \ carrier (R[\\<^bsub>n\<^esub>])" using assms unfolding coord_ring_def using Pring_smult_closed by (simp add: R.Pring_smult_closed) lemma Qp_poly_smult_cfs: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) m = a \ (P m)" using assms unfolding coord_ring_def using Qp.Pring_smult_cfs by blast lemma Qp_smult_r_distr: assumes "a \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q) = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> q)" using assms unfolding coord_ring_def using Qp.Pring_smult_r_distr by blast lemma Qp_smult_l_distr: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(a \ b) \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P = (a \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P) \\<^bsub> Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (b \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> P)" using assms unfolding coord_ring_def using Qp.Pring_smult_l_distr by blast abbreviation(input) Qp_funs where "Qp_funs n \ Fun\<^bsub>n\<^esub> Q\<^sub>p" (**************************************************************************************************) (**************************************************************************************************) subsection\Evaluation of Polynomials in $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) Qp_ev where "Qp_ev P q \ (eval_at_point Q\<^sub>p q P)" lemma Qp_ev_one: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \" unfolding coord_ring_def by (metis Qp.Pring_one eval_at_point_const Qp.one_closed assms) lemma Qp_ev_zero: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> a = \"unfolding coord_ring_def by (metis Qp.Pring_zero eval_at_point_const Qp.zero_closed assms) lemma Qp_eval_pvar_pow: assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "k < n" assumes "(m::nat) \ 0" shows "Qp_ev ((pvar Q\<^sub>p k)[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> m) a = ((a!k)[^]m)" by (metis eval_at_point_nat_pow eval_pvar assms(1) assms(2) pvar_closed) text\composition of polynomials over $\mathbb{Q}_p$\ definition Qp_poly_comp where "Qp_poly_comp m fs = poly_compose (length fs) m fs" text\lemmas about polynomial maps\ lemma Qp_is_poly_tupleI: assumes "\i. i < length fs\ fs!i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_poly_tuple m fs" unfolding is_poly_tuple_def using assms using cartesian_power_car_memE'' cartesian_power_car_memI' by blast lemma Qp_is_poly_tuple_append: assumes "is_poly_tuple m fs" assumes "is_poly_tuple m gs" shows "is_poly_tuple m (fs@gs)" proof(rule Qp_is_poly_tupleI) show "\i. i < length (fs @ gs) \ (fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" proof- fix i assume A: "i < length (fs @ gs)" show "(fs @ gs) ! i \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" apply(cases "i < length fs") using assms is_poly_tupleE[of m fs i] nth_append[of fs gs i] apply presburger proof- assume B: "\ i < length fs" then have C: "length fs \ i \ i < length (fs @ gs)" using A not_le by blast then have "i - length fs < length gs" using length_append[of fs gs] by linarith then show ?thesis using A assms is_poly_tupleE[of m gs "i - length fs"] nth_append[of fs gs i] B by presburger qed qed qed lemma Qp_poly_mapE: assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j \ carrier Q\<^sub>p" using assms poly_map_closed cartesian_power_car_memE' by blast lemma Qp_poly_mapE': assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = length fs" unfolding poly_map_def using Qp.cring_axioms poly_tuple_evalE' by (metis assms restrict_def) lemma Qp_poly_mapE'': assumes "is_poly_tuple n fs" assumes "length fs = m" assumes "n \ 0" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "j < m" shows "(poly_map n fs as)!j = (Qp_ev (fs!j) as)" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) nth_map restrict_apply') lemma poly_map_apply: assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "poly_map n fs as = poly_tuple_eval fs as" unfolding poly_map_def restrict_def by (simp add: assms) lemma poly_map_pullbackI: assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ poly_map n fs \\<^bsub>n\<^esub> S" using assms poly_map_apply by blast lemma poly_map_pullbackI': assumes "is_poly_tuple n fs" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "poly_map n fs as \ S" shows "as \ ((poly_map n fs) -` S)" by (simp add: assms(3)) text\lemmas about polynomial composition\ lemma poly_compose_ring_hom: assumes "is_poly_tuple m fs" assumes "length fs = n" shows "(ring_hom_ring (Q\<^sub>p[\\<^bsub>n\<^esub>]) (Q\<^sub>p[\\<^bsub>m\<^esub>]) (Qp_poly_comp m fs))" unfolding Qp_poly_comp_def by (simp add: assms(1) assms(2) poly_compose_ring_hom) lemma poly_compose_closed: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "(Qp_poly_comp m fs f) \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" using Qp.cring_axioms assms unfolding Qp_poly_comp_def using poly_compose_closed by blast lemma poly_compose_add: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_add unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_mult: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "Qp_poly_comp m fs (f \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> g) = (Qp_poly_comp m fs f) \\<^bsub>Q\<^sub>p[\\<^bsub>m\<^esub>]\<^esub> (Qp_poly_comp m fs g)" using Qp.cring_axioms assms poly_compose_mult unfolding is_poly_tuple_def Qp_poly_comp_def by blast lemma poly_compose_const: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "a \ carrier Q\<^sub>p" shows "Qp_poly_comp m fs (Qp_to_IP a) = Qp_to_IP a" using Qp.cring_axioms assms poly_compose_const unfolding is_poly_tuple_def Qp_poly_comp_def by metis lemma Qp_poly_comp_eval: assumes "is_poly_tuple m fs" assumes "length fs = n" assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" shows "Qp_ev (Qp_poly_comp m fs f) as = Qp_ev f (poly_map m fs as)" proof- have "(restrict (poly_tuple_eval fs) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) as) = poly_tuple_eval fs as" unfolding restrict_def by (simp add: assms) thus ?thesis using assms Qp.cring_axioms poly_compose_eval unfolding Qp_poly_comp_def poly_map_def by presburger qed (**************************************************************************************************) (**************************************************************************************************) subsection\Mapping Univariate Polynomials to Multivariable Polynomials in One Variable\ (**************************************************************************************************) (**************************************************************************************************) abbreviation(input) to_Qp_x where "to_Qp_x \ (IP_to_UP (0::nat) :: (nat multiset \ padic_number) \ nat \ padic_number)" abbreviation(input) from_Qp_x where "from_Qp_x \ UP_to_IP Q\<^sub>p (0::nat)" lemma from_Qp_x_closed: assumes "q \ carrier Q\<^sub>p_x" shows "from_Qp_x q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" using assms UP_to_IP_closed unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_closed: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "to_Qp_x q \ carrier Q\<^sub>p_x" using assms Qp.IP_to_UP_closed[of q "0::nat"] Qp.cring_axioms unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma to_Qp_x_from_Qp_x: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" shows "from_Qp_x (to_Qp_x q) = q" using assms UP_to_IP_inv[of q "0::nat"] Qp.Pring_car unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_to_Qp_x: assumes "q \ carrier Q\<^sub>p_x" shows "to_Qp_x (from_Qp_x q) = q" by (meson UPQ.IP_to_UP_inv assms) text\ring hom properties of these maps\ lemma to_Qp_x_ring_hom: "to_Qp_x \ ring_hom (Q\<^sub>p[\\<^bsub>1\<^esub>]) Q\<^sub>p_x" using IP_to_UP_ring_hom[of "0::nat"] ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_ring_hom: "from_Qp_x \ ring_hom Q\<^sub>p_x (Q\<^sub>p[\\<^bsub>1\<^esub>])" using UP_to_IP_ring_hom ring_hom_ring.homh unfolding coord_ring_def by (metis One_nat_def lessThan_0 lessThan_Suc) lemma from_Qp_x_add: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis (mono_tags, lifting) assms(1) assms(2) from_Qp_x_ring_hom ring_hom_add) lemma from_Qp_x_mult: assumes "a \ carrier Q\<^sub>p_x" assumes "b \ carrier Q\<^sub>p_x" shows "from_Qp_x (a \\<^bsub>Q\<^sub>p_x\<^esub> b) = from_Qp_x a \\<^bsub>Q\<^sub>p[\\<^bsub>1\<^esub>]\<^esub> from_Qp_x b" by (metis assms(1) assms(2) from_Qp_x_ring_hom ring_hom_mult) text\equivalence of evaluation maps\ lemma Qp_poly_Qp_x_eval: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" shows "Qp_ev P a = (to_Qp_x P)\(Qp.to_R a)" proof- have 0: "(IP_to_UP 0 P) \ (a ! 0) = ((IP_to_UP 0 P) \ (if 0 < length a then a ! 0 else \))" using assms by (metis (mono_tags, lifting) cartesian_power_car_memE gr0I zero_neq_one) have 1: "closed_fun Q\<^sub>p (\n. if n < length a then a ! n else \)" proof fix n show "(if n < length a then a ! n else \) \ carrier Q\<^sub>p" apply(cases "n < length a") using assms apply (metis cartesian_power_car_memE cartesian_power_car_memE') by (meson Qp.cring_axioms cring.cring_simprules(2)) qed have 2: " P \ Pring_set Q\<^sub>p {0::nat}" using assms unfolding coord_ring_def by (metis Qp.Pring_car UPQ.UP_to_IP_closed assms(1) to_Qp_x_closed to_Qp_x_from_Qp_x) have 3: "total_eval Q\<^sub>p (\i. if i < length a then a ! i else \) P = IP_to_UP 0 P \ (if 0 < length a then a ! 0 else \)" using 1 2 assms IP_to_UP_poly_eval[of P "0::nat" "(\i. if i < length a then a ! i else \)" ] UPQ.to_fun_def by presburger then show ?thesis using 0 unfolding eval_at_point_def by blast qed lemma Qp_x_Qp_poly_eval: assumes "P \ carrier Q\<^sub>p_x" assumes "a \ carrier Q\<^sub>p" shows "P \ a = Qp_ev (from_Qp_x P) (to_R1 a)" proof- have "Qp_ev (from_Qp_x P) (to_R1 a) = (to_Qp_x (from_Qp_x P))\(Qp.to_R (Qp.to_R1 a))" using Qp_poly_Qp_x_eval assms(1) assms(2) from_Qp_x_closed Qp.to_R1_closed by blast then show ?thesis using assms by (metis UPQ.IP_to_UP_inv Qp.to_R_to_R1) qed (**************************************************************************************************) (**************************************************************************************************) subsection\$n^{th}$-Power Sets over $\mathbb{Q}_p$\ (**************************************************************************************************) (**************************************************************************************************) definition P_set where "P_set (n::nat) = {a \ nonzero Q\<^sub>p. (\y \ carrier Q\<^sub>p . (y[^] n) = a)}" lemma P_set_carrier: "P_set n \ carrier Q\<^sub>p" unfolding P_set_def nonzero_def by blast lemma P_set_memI: assumes "a \ carrier Q\<^sub>p" assumes "a \ \" assumes "b \ carrier Q\<^sub>p" assumes "b[^](n::nat) = a" shows "a \ P_set n" unfolding P_set_def using assms by (metis (mono_tags, lifting) mem_Collect_eq not_nonzero_Qp) lemma P_set_nonzero: "P_set n \ nonzero Q\<^sub>p" unfolding P_set_def by blast lemma P_set_nonzero': assumes "a \ P_set n" shows "a \ nonzero Q\<^sub>p" "a \ carrier Q\<^sub>p" using assms P_set_nonzero P_set_carrier apply blast using assms P_set_carrier by blast lemma P_set_one: assumes "n \ 0" shows "\ \ P_set (n::nat)" proof- have 0: "\[^]n = \" using Qp.nat_pow_one by blast have 1: "\ \ carrier Q\<^sub>p" by blast then show ?thesis using one_nonzero unfolding P_set_def using 0 by blast qed lemma zeroth_P_set: "P_set 0 = {\}" proof show "P_set 0 \ {\}" unfolding P_set_def proof fix x assume "x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^](0::nat)) = a}" then have "\y\carrier Q\<^sub>p. (y[^](0::nat)) = x" by blast then obtain a where a_def: "a \ carrier Q\<^sub>p \ (a[^](0::nat)) = x" by blast then show "x \ {\}" using Qp.nat_pow_0 by blast qed show "{\} \ P_set 0" using P_set_memI[of \ \ 0] Qp.nat_pow_one Qp.one_closed local.one_neq_zero by blast qed lemma P_set_mult_closed: assumes "n \ 0" assumes "a \ P_set n" assumes "b \ P_set n" shows "a \ b \ P_set n" proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ (a0 [^] n = a)" using assms(2) unfolding P_set_def by blast obtain b0 where b0_def: "b0 \ carrier Q\<^sub>p \ (b0 [^] n = b)" using assms(3) unfolding P_set_def by blast have "(a0 \ b0) [^] n = a0 [^] n \ b0 [^] n" using a0_def b0_def Qp.nat_pow_distrib by blast then have 0: "a \ b = (a0 \ b0) [^] n" using a0_def b0_def by blast have 1: "a0 \ b0 \ carrier Q\<^sub>p" by (meson Qp.cring_axioms a0_def b0_def cring.cring_simprules(5)) have 2: "a \ b \ nonzero Q\<^sub>p" using assms nonzero_is_submonoid unfolding P_set_def by (metis (no_types, lifting) "0" "1" Qp.integral Qp_nat_pow_nonzero a0_def b0_def mem_Collect_eq not_nonzero_Qp) then show ?thesis using 0 1 assms unfolding P_set_def by blast qed lemma P_set_inv_closed: assumes "a \ P_set n" shows "inv a \ P_set n" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis Qp.inv_one singletonD) next case False then show ?thesis proof- obtain a0 where a0_def: "a0 \ carrier Q\<^sub>p \ a0[^]n = a" using assms P_set_def[of n] by blast have "a0 \ nonzero Q\<^sub>p" apply(rule ccontr) proof- assume "a0 \ nonzero Q\<^sub>p " then have "a0 = \" using a0_def by (meson not_nonzero_Qp) then show False using a0_def assms by (metis (mono_tags, lifting) False P_set_def Qp.cring_axioms \a0 \ nonzero Q\<^sub>p\ cring_def mem_Collect_eq neq0_conv ring.pow_zero) qed then have "(inv a0)[^]n = inv a" using a0_def \a0 \ carrier Q\<^sub>p \ (a0[^]n) = a\ \a0 \ nonzero Q\<^sub>p\ Units_nonzero monoid.nat_pow_of_inv[of Q\<^sub>p a n] Qp.nat_pow_of_inv Units_eq_nonzero by presburger then show ?thesis by (metis P_set_memI Qp.nat_pow_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero \a0 \ nonzero Q\<^sub>p\ a0_def inv_in_frac(1) inv_in_frac(2)) qed qed lemma P_set_val: assumes "a \ P_set (n::nat)" shows "(ord a) mod n = 0" proof(cases "n = 0") case True then show ?thesis using assms zeroth_P_set by (metis mod_by_0 of_nat_0 ord_one singletonD) next case False then show ?thesis proof- obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^] n) = a" using assms P_set_def by blast have an: "a \ nonzero Q\<^sub>p" using P_set_def assms by blast have bn: "b \ nonzero Q\<^sub>p" proof(rule ccontr) assume "b \ nonzero Q\<^sub>p" then have "b = \\<^bsub> Q\<^sub>p\<^esub>" using b_def not_nonzero_Qp by metis then have "(b[^] n) = \" using False Qp.cring_axioms cring_def ring.pow_zero by blast then show False using b_def an Qp.not_nonzero_memI by blast qed then have "ord a = n * (ord b)" using b_def an nonzero_nat_pow_ord by blast then show ?thesis by (metis mod_mult_self1_is_0) qed qed lemma P_set_pow: assumes "n > 0" assumes "s \ P_set n" shows "s[^]k \ P_set (n*k)" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ y[^]n = s" using assms unfolding P_set_def by blast then have 0: "y \ nonzero Q\<^sub>p" using assms P_set_nonzero'(1) Qp_nonzero_nat_pow by blast have 1: "y[^](n*k) = s[^] k" using 0 y_def assms Qp.nat_pow_pow by blast hence 2: "s[^]k \ nonzero Q\<^sub>p" using 0 by (metis Qp_nat_pow_nonzero) thus ?thesis unfolding P_set_def using 1 y_def by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) text\ In this section we introduce the notion of a $p$-adic semialgebraic set. Intuitively, these are the subsets of $\mathbb{Q}_p^n$ which are definable by first order quantifier-free formulas in the standard first-order language of rings, with an additional relation symbol included for the relation $\text{ val}(x) \leq \text{ val}(y)$, interpreted according to the definiton of the $p$-adic valuation on $\mathbb{Q}_p$. In fact, by Macintyre's quantifier elimination theorem for the first-order theory of $\mathbb{Q}_p$ in this language, one can equivalently remove the ``quantifier-free" clause from the latter definition. The definition we give here is also equivalent, and due to Denef in \cite{denef1986}. The given definition here is desirable mainly for its utility in producing a proof of Macintyre's theorem, which is our overarching goal. \ (********************************************************************************************) (********************************************************************************************) subsubsection\Defining Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition basic_semialg_set where "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). \y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)}" lemma basic_semialg_set_zero_set: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "Qp_ev P q = \" assumes "n \ 0" shows "q \ basic_semialg_set (m::nat) (n::nat) P" proof- have "\ = (\[^]n)" using assms(4) Qp.nat_pow_zero by blast then show ?thesis unfolding basic_semialg_set_def using assms Qp.cring_axioms cring.cring_simprules(2) by blast qed lemma basic_semialg_set_def': assumes "n \ 0" assumes "P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "basic_semialg_set (m::nat) (n::nat) P = {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ (P_set n)}" proof show "basic_semialg_set m n P \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" proof fix x assume A: "x \ basic_semialg_set m n P" show "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" apply(cases "Qp_ev P x = \") using A basic_semialg_set_def apply blast unfolding basic_semialg_set_def P_set_def proof assume A0: "Qp_ev P x \ \" have A1: " \y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n)" using A basic_semialg_set_def by blast have A2: "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A basic_semialg_set_def by blast show " x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (Qp_ev P x = \ \ Qp_ev P x \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a})" by (metis (mono_tags, lifting) A1 A2 Qp.nonzero_memI assms(2) eval_at_point_closed mem_Collect_eq) qed qed show "{q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} \ basic_semialg_set m n P" proof fix x assume A: " x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n}" then have A':"x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by blast show "x \ basic_semialg_set m n P" using A A' apply(cases "Qp_ev P x = \") using assms basic_semialg_set_zero_set[of P m x n] apply blast proof- assume B: "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ P_set n} " assume B': "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assume B'': "Qp_ev P x \ \ " show "x \ basic_semialg_set m n P" unfolding basic_semialg_set_def P_set_def proof have "\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n) " using A nonzero_def [of Q\<^sub>p] unfolding P_set_def proof - assume "x \ {q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>). Qp_ev P q = \ \ Qp_ev P q \ {a \ nonzero Q\<^sub>p. \y\carrier Q\<^sub>p. (y[^]n) = a}}" then have "Qp_ev P x \ nonzero Q\<^sub>p \ (\r. r \ carrier Q\<^sub>p \ (r[^]n) = Qp_ev P x)" using B'' by blast then show ?thesis by blast qed then show "x \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ (\y\carrier Q\<^sub>p. Qp_ev P x = (y[^]n))" using B' by blast qed qed qed qed lemma basic_semialg_set_memI: assumes "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" assumes "y \ carrier Q\<^sub>p" assumes "Qp_ev P q = (y[^]n)" shows "q \ basic_semialg_set m n P" using assms(1) assms(2) assms(3) basic_semialg_set_def by blast lemma basic_semialg_set_memE: assumes "q \ basic_semialg_set m n P" shows "q \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" "\y \ carrier Q\<^sub>p. Qp_ev P q = (y[^]n)" using assms basic_semialg_set_def apply blast using assms basic_semialg_set_def by blast definition is_basic_semialg :: "nat \ ((nat \ int) \ (nat \ int)) set list set \ bool" where "is_basic_semialg m S \ (\ (n::nat) \ 0. (\ P \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]). S = basic_semialg_set m n P))" abbreviation(input) basic_semialgs where "basic_semialgs m \ {S. (is_basic_semialg m S)}" definition semialg_sets where "semialg_sets n = gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" lemma carrier_is_semialg: "(carrier (Q\<^sub>p\<^bsup>n\<^esup>)) \ semialg_sets n " unfolding semialg_sets_def using gen_boolean_algebra.universe by blast lemma empty_set_is_semialg: " {} \ semialg_sets n" using carrier_is_semialg[of n] unfolding semialg_sets_def using gen_boolean_algebra.complement by (metis Diff_cancel) lemma semialg_intersect: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms(1) assms(2) gen_boolean_algebra_intersect semialg_sets_def by blast lemma semialg_union: assumes "A \ semialg_sets n" assumes "B \ semialg_sets n" shows "(A \ B) \ semialg_sets n " using assms gen_boolean_algebra.union semialg_sets_def by blast lemma semialg_complement: assumes "A \ semialg_sets n" shows "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) \ semialg_sets n " using assms gen_boolean_algebra.complement semialg_sets_def by blast lemma semialg_zero: assumes "A \ semialg_sets 0" shows "A = {[]} \ A = {}" using assms unfolding semialg_sets_def cartesian_power_def proof- assume A0: " A \ gen_boolean_algebra (carrier (RDirProd_list (R_list 0 Q\<^sub>p))) (basic_semialgs 0)" show " A = {[]} \ A = {}" proof- have "A \ {[]} \ A = {}" proof assume A1: "A \ {[]}" show "A = {}" proof- have "(R_list 0 Q\<^sub>p) = []" by simp then have "(carrier (RDirProd_list (R_list 0 Q\<^sub>p))) = {[]}" using RDirProd_list_nil by simp then show ?thesis using A0 A1 by (metis gen_boolean_algebra_subset subset_singletonD) qed qed then show ?thesis by linarith qed qed lemma basic_semialg_is_semialg: assumes "is_basic_semialg n A" shows "A \ semialg_sets n" by (metis (no_types, lifting) assms gen_boolean_algebra.simps inf_absorb1 is_basic_semialg_def mem_Collect_eq basic_semialg_set_def semialg_sets_def subsetI) lemma basic_semialg_is_semialg': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "A \ semialg_sets n" using assms basic_semialg_is_semialg is_basic_semialg_def by blast definition is_semialgebraic where "is_semialgebraic n S = (S \ semialg_sets n)" lemma is_semialgebraicE: assumes "is_semialgebraic n S" shows "S \ semialg_sets n" using assms is_semialgebraic_def by blast lemma is_semialgebraic_closed: assumes "is_semialgebraic n S" shows "S \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using is_semialgebraicE[of n S] unfolding semialg_sets_def using assms gen_boolean_algebra_subset is_semialgebraicE semialg_sets_def by blast lemma is_semialgebraicI: assumes "S \ semialg_sets n" shows "is_semialgebraic n S" by (simp add: assms is_semialgebraic_def) lemma basic_semialg_is_semialgebraic: assumes "is_basic_semialg n A" shows "is_semialgebraic n A" using assms basic_semialg_is_semialg is_semialgebraicI by blast lemma basic_semialg_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "m \0" assumes "A = basic_semialg_set n m f" shows "is_semialgebraic n A" using assms(1) assms(2) assms(3) basic_semialg_is_semialg' is_semialgebraicI by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Algebraic Sets over $p$-adic Fields\ (********************************************************************************************) (********************************************************************************************) lemma p_times_square_not_square: assumes "a \ nonzero Q\<^sub>p" shows "\ \ (a [^] (2::nat)) \ P_set (2::nat)" proof assume A: "\ \ (a[^](2::nat)) \ P_set (2::nat)" then have "\ \ (a[^](2::nat)) \ nonzero Q\<^sub>p" unfolding P_set_def by blast then obtain b where b_def: "b \ carrier Q\<^sub>p \ b[^](2::nat) = \ \ (a[^](2::nat))" using A P_set_def by blast have "b \ nonzero Q\<^sub>p" apply(rule ccontr) using b_def assms by (metis A P_set_nonzero'(1) Qp.nat_pow_zero not_nonzero_Qp zero_neq_numeral) then have LHS: "ord (b[^](2::nat)) = 2* (ord b)" using nonzero_nat_pow_ord by presburger have "ord( \ \ (a[^](2::nat))) = 1 + 2* ord a" using assms nonzero_nat_pow_ord Qp_nat_pow_nonzero ord_mult ord_p p_nonzero by presburger then show False using b_def LHS by presburger qed lemma p_times_square_not_square': assumes "a \ carrier Q\<^sub>p" shows "\ \ (a [^] (2::nat)) = \ \ a = \" by (metis Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero assms p_nonzero) lemma zero_set_semialg_set: assumes "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "Qp_ev q a = \ \( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev q a) [^] (2::nat)) = y[^](2::nat)) " proof show "Qp_ev q a = \ \ \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" proof- assume "Qp_ev q a = \" then have "\ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" by (metis Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null zero_neq_numeral) then have "\ \ carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (\[^](2::nat))" using Qp.cring_axioms cring.cring_simprules(2) by blast then show "\y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^] (2::nat)) = (y[^] (2::nat))" by blast qed show " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat)) \ Qp_ev q a = \" proof- assume A: " \y\carrier Q\<^sub>p. \ \ (Qp_ev q a[^](2::nat)) = (y[^](2::nat))" then obtain b where b_def: "b\carrier Q\<^sub>p \ \ \ (Qp_ev q a[^](2::nat)) = (b[^](2::nat))" by blast show "Qp_ev q a = \" proof(rule ccontr) assume " Qp_ev q a \ \" then have " Qp_ev q a \ nonzero Q\<^sub>p" using assms eval_at_point_closed[of a n q] nonzero_def proof - have "Qp_ev q a \ carrier Q\<^sub>p" using \\a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>); q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ \ Qp_ev q a \ carrier Q\<^sub>p\ \a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ by fastforce then have "Qp_ev q a \ {r \ carrier Q\<^sub>p. r \ \}" using \Qp_ev q a \ \\ by force then show ?thesis by (metis nonzero_def ) qed then have "\ \ (Qp_ev q a[^](2::nat)) \ nonzero Q\<^sub>p" by (metis Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero not_nonzero_Qp p_nonzero p_times_square_not_square') then have "\ \ (Qp_ev q a[^](2::nat)) \ P_set (2::nat)" using b_def unfolding P_set_def by blast then show False using \Qp_ev q a \ nonzero Q\<^sub>p\ p_times_square_not_square by blast qed qed qed lemma alg_as_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" shows "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" proof have 00: "\x. x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ Qp_ev q x = \ \ (Qp_ev P x) [^] (2::nat)" using assms eval_at_point_smult MP.nat_pow_closed Qp.int_inc_closed eval_at_point_nat_pow by presburger show "V\<^bsub>Q\<^sub>p\<^esub> n P \ basic_semialg_set n 2 q" proof fix x assume A: "x \ V\<^bsub>Q\<^sub>p\<^esub> n P " show "x \ basic_semialg_set n (2::nat) q " proof- have P: "Qp_ev P x = \" using A zero_setE(2) by blast have "Qp_ev q x = \" proof- have "Qp_ev q x = \ \ (Qp_ev (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) x)" using assms eval_at_point_smult[of x n "(P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" \] basic_semialg_set_def by (meson A MP.nat_pow_closed Qp.int_inc_closed zero_setE(1)) then show ?thesis by (metis A P Qp.int_inc_closed Qp.integral_iff Qp.nat_pow_zero Qp.zero_closed assms(1) eval_at_point_nat_pow neq0_conv zero_less_numeral zero_setE(1)) qed then have 0: "Qp_ev q x = \ \ Qp_ev q x \ P_set (2::nat)" by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A zero_setE(1) by blast then show ?thesis using 0 basic_semialg_set_def' by (metis (no_types, opaque_lifting) Qp.nat_pow_zero Qp.zero_closed \eval_at_point Q\<^sub>p x q = \\ basic_semialg_set_memI zero_neq_numeral) qed qed show "basic_semialg_set n 2 q \ V\<^bsub>Q\<^sub>p\<^esub> n P" proof fix x assume A: "x \ basic_semialg_set n 2 q" have 0: "\ Qp_ev q x \ P_set 2" proof assume "Qp_ev q x \ P_set 2" then have 0: "Qp_ev q x \ nonzero Q\<^sub>p \ (\y \ carrier Q\<^sub>p . (y[^] (2::nat)) = Qp_ev q x)" using P_set_def by blast have "( \y \ carrier Q\<^sub>p. \ \ ((Qp_ev P x) [^] (2::nat)) = y[^](2::nat))" proof- obtain y where y_def: "y \ carrier Q\<^sub>p \ (y[^] (2::nat)) = Qp_ev q x" using 0 by blast have "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_memE(1) by blast then have "Qp_ev q x = \ \ ((Qp_ev P x) [^] (2::nat))" using assms eval_at_point_scalar_mult 00 by blast then have "(y[^] (2::nat)) = \ \ ((Qp_ev P x) [^] (2::nat))" using y_def by blast then show ?thesis using y_def by blast qed then have "Qp_ev P x = \" by (metis (no_types, lifting) A assms(1) basic_semialg_set_def mem_Collect_eq zero_set_semialg_set) then have "Qp_ev q x = \" using assms eval_at_point_smult by (metis "00" A Qp.int_inc_closed Qp.nat_pow_zero Qp.r_null basic_semialg_set_memE(1) zero_neq_numeral) then show False using 0 Qp.not_nonzero_memI by blast qed show " x \ V\<^bsub>Q\<^sub>p\<^esub> n P" apply(rule zero_setI) using A basic_semialg_set_memE(1) apply blast using A 0 00[of x] by (metis assms(1) basic_semialg_set_memE(1) basic_semialg_set_memE(2) zero_set_semialg_set) qed qed lemma is_zero_set_imp_basic_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_basic_semialg n S" unfolding is_basic_semialg_def proof- obtain q where q_def: "q = \ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (P[^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat))" by blast have 0: "zero_set Q\<^sub>p n P = basic_semialg_set n (2::nat) q" using alg_as_semialg[of P n q] q_def assms(1) by linarith have "(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms(1) by blast then have "\ \\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub>(P [^]\<^bsub>Q\<^sub>p[\\<^bsub>n\<^esub>]\<^esub> (2::nat)) \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using assms q_def Qp.int_inc_closed local.smult_closed by blast then have 1: "q \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" by (metis q_def ) then show "\m. m \ 0 \ (\P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). S = basic_semialg_set n m P)" using 0 assms by (metis zero_neq_numeral) qed lemma is_zero_set_imp_semialg: assumes "P \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "S = zero_set Q\<^sub>p n P" shows "is_semialgebraic n S" using assms(1) assms(2) basic_semialg_is_semialg is_semialgebraicI is_zero_set_imp_basic_semialg by blast text\Algebraic sets are semialgebraic\ lemma is_algebraic_imp_is_semialg: assumes "is_algebraic Q\<^sub>p n S" shows "is_semialgebraic n S" proof(rule is_semialgebraicI) obtain ps where ps_def: "finite ps \ ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ S = affine_alg_set Q\<^sub>p n ps" using is_algebraicE by (metis assms) have "ps \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n ps \ semialg_sets n" apply(rule finite.induct[of ps]) apply (simp add: ps_def) using affine_alg_set_empty[of n] apply (simp add: carrier_is_semialg) proof fix A a assume IH: "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ affine_alg_set Q\<^sub>p n A \ semialg_sets n" assume P: "insert a A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" have "A \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" using P by blast then show "affine_alg_set Q\<^sub>p n (insert a A) \ semialg_sets n" using IH P semialg_intersect[of "affine_alg_set Q\<^sub>p n A" n "affine_alg_set Q\<^sub>p n {a}" ] is_zero_set_imp_semialg affine_alg_set_insert[of n a A] by (metis Int_commute affine_alg_set_singleton insert_subset is_semialgebraicE) qed then show "S \ semialg_sets n" using ps_def by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\Basic Lemmas about the Semialgebraic Predicate\ (********************************************************************************************) (********************************************************************************************) text\Finite and cofinite sets are semialgebraic\ lemma finite_is_semialg: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "finite F" shows "is_semialgebraic n F" using Qp.finite_sets_are_algebraic is_algebraic_imp_is_semialg[of n F] assms(1) assms(2) by blast definition is_cofinite where "is_cofinite n F = finite (ring_pow_comp Q\<^sub>p n F)" lemma is_cofiniteE: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "finite (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)" using assms(2) is_cofinite_def by (simp add: ring_pow_comp_def) lemma complement_is_semialg: assumes "is_semialgebraic n F" shows "is_semialgebraic n ((carrier (Q\<^sub>p\<^bsup>n\<^esup>)) - F)" using assms is_semialgebraic_def semialg_complement by blast lemma cofinite_is_semialgebraic: assumes "F \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "is_cofinite n F" shows "is_semialgebraic n F" using assms ring_pow_comp_inv[of F Q\<^sub>p n] complement_is_semialg[of n "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] finite_is_semialg[of "(carrier (Q\<^sub>p\<^bsup>n\<^esup>) - F)"] is_cofiniteE[of F] by (simp add: ring_pow_comp_def) lemma diff_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A - B)" apply(rule is_semialgebraicI) using assms unfolding semialg_sets_def using gen_boolean_algebra_diff is_semialgebraicE semialg_sets_def by blast lemma intersection_is_semialg: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_intersect by blast lemma union_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic n B" shows "is_semialgebraic n (A \ B)" using assms(1) assms(2) is_semialgebraicE is_semialgebraicI semialg_union by blast lemma carrier_is_semialgebraic: "is_semialgebraic n (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using carrier_is_semialg by (simp add: carrier_is_semialg is_semialgebraic_def) lemma empty_is_semialgebraic: "is_semialgebraic n {}" by (simp add: empty_set_is_semialg is_semialgebraic_def) (********************************************************************************************) (********************************************************************************************) subsubsection\One-Dimensional Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition one_var_semialg where "one_var_semialg S = ((to_R1 ` S) \ (semialg_sets 1))" definition univ_basic_semialg_set where "univ_basic_semialg_set (m::nat) P = {a \ carrier Q\<^sub>p. (\y \ carrier Q\<^sub>p. (P \ a = (y[^]m)))}" text\Equivalence of univ\_basic\_semialg\_sets and semialgebraic subsets of $\mathbb{Q}^1$ \ lemma univ_basic_semialg_set_to_semialg_set: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "to_R1 ` (univ_basic_semialg_set m P) = basic_semialg_set 1 m (from_Qp_x P)" proof show "(\a. [a]) ` univ_basic_semialg_set m P \ basic_semialg_set 1 m (from_Qp_x P)" proof fix x assume A: "x \ (\a. [a]) ` univ_basic_semialg_set m P" then obtain b y where by_def:"b \ carrier Q\<^sub>p \ y \ carrier Q\<^sub>p \ (P \ b) = (y[^]m) \ x = [b]" unfolding univ_basic_semialg_set_def by blast then have "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A Qp.to_R1_closed[of b] unfolding univ_basic_semialg_set_def by blast then show "x \ basic_semialg_set 1 m (from_Qp_x P)" using by_def Qp_x_Qp_poly_eval assms unfolding basic_semialg_set_def by blast qed show "basic_semialg_set 1 m (from_Qp_x P) \ (\a. [a]) ` univ_basic_semialg_set m P" proof fix x assume A: "x \ basic_semialg_set 1 m (from_Qp_x P)" then obtain b where b_def: "b \ carrier Q\<^sub>p \ x = [b]" unfolding basic_semialg_set_def by (metis (mono_tags, lifting) mem_Collect_eq Qp.to_R1_to_R Qp.to_R_pow_closed) obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev (from_Qp_x P) [b] = (y[^]m))" using A b_def unfolding basic_semialg_set_def by blast have " P \ b = (y[^]m)" using assms y_def b_def Qp_x_Qp_poly_eval by blast then show " x \ (\a. [a]) ` univ_basic_semialg_set m P" using y_def b_def unfolding basic_semialg_set_def univ_basic_semialg_set_def by blast qed qed definition is_univ_semialgebraic where "is_univ_semialgebraic S = (S \ carrier Q\<^sub>p \ is_semialgebraic 1 (to_R1 ` S))" lemma is_univ_semialgebraicE: assumes "is_univ_semialgebraic S" shows "is_semialgebraic 1 (to_R1 ` S)" using assms is_univ_semialgebraic_def by blast lemma is_univ_semialgebraicI: assumes "is_semialgebraic 1 (to_R1 ` S)" shows "is_univ_semialgebraic S" proof- have "S \ carrier Q\<^sub>p" proof fix x assume "x \ S" then have "(to_R1 x) \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms by (smt Collect_mono_iff gen_boolean_algebra_subset image_def is_semialgebraicE mem_Collect_eq semialg_sets_def Qp.to_R1_carrier) then show "x \ carrier Q\<^sub>p" using assms by (metis nth_Cons_0 Qp.to_R_pow_closed) qed then show ?thesis using assms unfolding is_univ_semialgebraic_def by blast qed lemma univ_basic_semialg_set_is_univ_semialgebraic: assumes "P \ carrier Q\<^sub>p_x" assumes "m \ 0" shows "is_univ_semialgebraic (univ_basic_semialg_set m P)" using assms by (metis (mono_tags, lifting) basic_semialg_is_semialgebraic' from_Qp_x_closed is_univ_semialgebraic_def mem_Collect_eq subsetI univ_basic_semialg_set_def univ_basic_semialg_set_to_semialg_set) lemma intersection_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms intersection_is_semialg[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis le_infI1 Qp.to_R1_intersection) lemma union_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A \ B)" using assms union_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (metis Un_subset_iff image_Un) lemma diff_is_univ_semialgebraic: assumes "is_univ_semialgebraic A" assumes "is_univ_semialgebraic B" shows "is_univ_semialgebraic (A - B)" using assms diff_is_semialgebraic[of 1 "((\a. [a]) ` A)" "((\a. [a]) ` B)"] unfolding is_univ_semialgebraic_def by (smt Diff_subset subset_trans Qp.to_R1_diff) lemma finite_is_univ_semialgebraic: assumes "A \ carrier Q\<^sub>p" assumes "finite A" shows "is_univ_semialgebraic A" using assms finite_is_semialg[of "((\a. [a]) ` A)" ] to_R1_finite[of A] unfolding is_univ_semialgebraic_def by (metis Qp.to_R1_carrier Qp.to_R1_subset) (********************************************************************************************) (********************************************************************************************) subsubsection\Defining the $p$-adic Valuation Semialgebraically\ (********************************************************************************************) (********************************************************************************************) lemma Qp_square_root_criterion0: assumes "p \ 2" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "val a \ val b" assumes "a \ \" assumes "b \ \" assumes "val a \ 0" shows "\y \ carrier Q\<^sub>p. a[^](2::nat) \\<^bsub>Q\<^sub>p\<^esub> \\b[^](2::nat) = (y [^] (2::nat))" proof- have 0: "(to_Zp a) \ carrier Z\<^sub>p" using assms(2) to_Zp_closed by blast have 1: "(to_Zp b) \ carrier Z\<^sub>p" using assms(3) to_Zp_closed by blast have 2: "a \ \\<^sub>p" using val_ring_val_criterion assms(2) assms(5) assms(7) by blast have 3: "b \ \\<^sub>p" using assms val_ring_val_criterion[of b] dual_order.trans by blast have 4: "val_Zp (to_Zp b) = val b" using 3 Zp_def \_def padic_fields.to_Zp_val padic_fields_axioms by blast have 5: "val_Zp (to_Zp a) = val a" using Q\<^sub>p_def Zp_def assms(2) assms(7) padic_fields.Qp_val_ringI padic_fields.to_Zp_val padic_fields_axioms by blast have "\y \ carrier Z\<^sub>p. (to_Zp a)[^]\<^bsub>Z\<^sub>p\<^esub>(2::nat) \\<^bsub>Z\<^sub>p\<^esub> \

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

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

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

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

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

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

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

b) \ nonzero Q\<^sub>p" using B by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) inv_in_frac(3)) then have "val a \ val b" using F0 F1 A by blast then have "val (a \
b) \ 0" using F0 F1 val_fract assms(1) local.eint_minus_ineq' by presburger obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat)) = (y[^](2::nat))" using Qp_val_ring_alt_def0 by (meson B False Qp.integral Qp.nonzero_closed \(a \
b) \ nonzero Q\<^sub>p\ \0 \ val (a \
b)\ assms(1) assms(2) inv_in_frac(1) inv_in_frac(2) ord_nonneg val_ringI) then have "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = (b[^](4::nat)) \ (y[^](2::nat))" by presburger then have F2: "(b[^](4::nat)) \ (\ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis Qp.nat_pow_pow assms(2) mult_2_right numeral_Bit0) have F3: "((b[^](4::nat)) \ \) \\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat))[^] (2::nat)) \ (y[^](2::nat))" proof- have 0: "(\[^](3::nat)) \ (a \
b[^](4::nat)) \ carrier Q\<^sub>p " proof- have "(a \
b[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed Qp_nat_pow_nonzero by presburger then show ?thesis by (meson Qp.cring_axioms cring.cring_simprules(5) p_natpow_closed(1)) qed have 1: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then show ?thesis using 0 F2 ring.ring_simprules(23)[of Q\<^sub>p "\" "(\[^](3::nat))\ ((a \
b)[^](4::nat))" "(b[^](4::nat))"] Qp.cring_axioms Qp.nonzero_mult_closed Qp.ring_axioms Qp_nat_pow_nonzero \(a \
b) \ nonzero Q\<^sub>p\ p_nonzero by blast qed have F4: "(b[^](4::nat)) \ carrier Q\<^sub>p" using Qp.nat_pow_closed assms(2) by blast then have "((b[^](4::nat)) \ \) = (b[^](4::nat))" using Qp.r_one by blast then have F5: "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> ((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F3 by presburger have "((b[^](4::nat)) \((\[^](3::nat))\ ((a \
b)[^](4::nat)))) = (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat)))" proof- have 0: "(b[^](4::nat)) \ carrier Q\<^sub>p" using F4 by blast have 1: "(\[^](3::nat)) \ carrier Q\<^sub>p" by blast have 2: "((a \
b)[^](4::nat)) \ carrier Q\<^sub>p" using F0 Qp.nat_pow_closed assms(1) fract_closed by blast show ?thesis using 0 1 2 monoid.m_assoc[of Q\<^sub>p] comm_monoid.m_comm[of Q\<^sub>p] using Qp.m_lcomm by presburger qed then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b[^](4::nat)) \ ((a \
b)[^](4::nat))) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F5 by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\((b \(a \
b))[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" using F0 Qp.nat_pow_distrib assms(1) assms(2) fract_closed by presburger then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = ((b[^](2::nat)) [^] (2::nat)) \ (y[^](2::nat))" by (metis F0 assms(1) local.fract_cancel_right) then have "(b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" using Qp.nat_pow_closed Qp.nat_pow_distrib assms(2) y_def by blast then have "((b[^](2::nat))\ y) \ carrier Q\<^sub>p \ (b[^](4::nat))\\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\(a[^](4::nat)) = (((b[^](2::nat))\ y)[^](2::nat))" by (meson Qp.cring_axioms Qp.nat_pow_closed assms(2) cring.cring_simprules(5) y_def) then show ?thesis by blast qed qed qed show "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat)) \ val a \ val b" proof- assume A: "\y \ carrier Q\<^sub>p. (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" show "val a \ val b" proof(cases "a = \") case True then show ?thesis using eint_ord_code(3) local.val_zero by presburger next case False have "b \ \" proof(rule ccontr) assume "\ b \ \" then have "\y \ carrier Q\<^sub>p. (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by (metis (no_types, lifting) Qp.add.r_cancel_one' Qp.nat_pow_closed Qp.nonzero_memE(2) Qp_nonzero_nat_pow assms(1) assms(2) local.monom_term_car not_nonzero_Qp p_natpow_closed(1) zero_less_numeral) then obtain y where y_def: "y \ carrier Q\<^sub>p \ (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" by blast have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" proof- have 00: "(\[^](3::nat)) \ nonzero Q\<^sub>p" using Qp_nat_pow_nonzero p_nonzero by blast have 01: "(a[^](4::nat)) \ nonzero Q\<^sub>p" using False Qp_nat_pow_nonzero assms(1) not_nonzero_Qp Qp.nonzero_memI by presburger then show ?thesis using ord_mult[of "\[^](3::nat)" "a[^](4::nat)"] by (metis (no_types, lifting) "00" False assms(1) nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral ord_p_pow_nat) qed have 1: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 2* (ord y)" proof- have "y \ \" proof(rule ccontr) assume " \ y \ \" then have "(\[^](3::nat))\ (a[^](4::nat)) = \" using y_def Qp.cring_axioms cring_def pos2 ring.pow_zero by blast then show False by (metis False Qp.integral Qp.nat_pow_closed Qp.nonzero_pow_nonzero Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) p_natpow_closed(1) p_nonzero) qed then show ?thesis using y_def by (metis nonzero_nat_pow_ord not_nonzero_Qp of_nat_numeral) qed then show False using 0 by presburger qed then have F0: "b \ nonzero Q\<^sub>p" using assms(2) not_nonzero_Qp by metis have F1: "a \ nonzero Q\<^sub>p" using False assms(1) not_nonzero_Qp by metis obtain y where y_def: "y \ carrier Q\<^sub>p \ (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat)) = (y[^](2::nat))" using A by blast show ?thesis proof(rule ccontr) assume " \ val a \ val b " then have F2: "ord a < ord b" using F0 F1 assms by (metis False \b \ \\ eint_ord_simps(1) leI val_def) have 0: "ord ((\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using F0 ord_mult F1 Qp_nat_pow_nonzero nonzero_nat_pow_ord ord_p_pow_nat p_natpow_closed(2) by presburger have 1: " ord (b[^](4::nat)) = 4* ord b" using F0 nonzero_nat_pow_ord by presburger have 2: "(4 * (ord b)) > 4 * (ord a)" using F2 by linarith have 3: "(4 * (ord b)) \ 3 + 4* ord a" proof(rule ccontr) assume "\ (4 * (ord b)) \ 3 + 4* ord a" then have "(4 * (ord b)) > 3 + 4* ord a" by linarith then have 30: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 3 + 4* ord a" using "0" "1" F0 F1 Qp_nat_pow_nonzero Qp.nat_pow_closed assms(1) monom_term_car not_nonzero_Qp ord_ultrametric_noteq p_natpow_closed(1) p_nonzero by (metis Qp.integral) have "y \ nonzero Q\<^sub>p" proof(rule ccontr) assume A: "y \ nonzero Q\<^sub>p" then have "y = \" using y_def Qp.nonzero_memI by blast then have "b [^] 4 \ \ [^] 3 \ a [^] 4 = \" by (smt "0" "1" A F0 False Qp.integral Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_mult_closed Qp.nonzero_pow_nonzero Qp.pow_zero assms(1) diff_ord_nonzero not_nonzero_Qp p_nonzero pos2 y_def) then show False by (smt "0" "1" A F0 F1 Qp.integral Qp.nat_pow_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero assms(1) diff_ord_nonzero not_nonzero_Qp p_natpow_closed(1) p_nonzero y_def) qed then have 31: "ord ((b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))) = 2* ord y" using nonzero_nat_pow_ord y_def by presburger then show False using 30 by presburger qed show False using 2 3 by presburger qed qed qed qed text\The polynomial in two variables which semialgebraically defines the valuation relation\ definition Qp_val_poly where "Qp_val_poly = (pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" lemma Qp_val_poly_closed: "Qp_val_poly \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" proof- have "(pvar Q\<^sub>p 1) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed one_less_numeral_iff semiring_norm(76) by blast then have 0: "(pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using ring.Pring_is_ring[of Q\<^sub>p "{0::nat..2-1}"] monoid.nat_pow_closed[of "coord_ring Q\<^sub>p 2"] Qp.cring_axioms cring.axioms(1) ring.Pring_is_monoid by blast have 1: "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using local.pvar_closed pos2 by blast have 2: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" using 1 local.smult_closed p_natpow_closed(1) by blast then show ?thesis unfolding Qp_val_poly_def using 0 by blast qed lemma Qp_val_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "Qp_ev Qp_val_poly [a, b] = (b[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (a[^](4::nat))" proof- have 0: "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" proof(rule cartesian_power_car_memI) show "length [a, b] = 2" by simp have "set [a,b] = {a,b}" by auto then show "set [a, b] \ carrier Q\<^sub>p" using assms by (simp add: \a \ carrier Q\<^sub>p\ \b \ carrier Q\<^sub>p\) qed obtain f where f_def: "f = ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))" by blast obtain g where g_def: "g = (\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)))" by blast have 1: "Qp_val_poly = f \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> g" unfolding Qp_val_poly_def using f_def g_def by blast have 1: "Qp_ev (pvar Q\<^sub>p (0::nat)) [a,b] = a" using eval_pvar by (metis \[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)\ nth_Cons_0 pos2) have 2: "Qp_ev (pvar Q\<^sub>p (1::nat)) [a,b] = b" using eval_pvar by (metis (no_types, lifting) "0" One_nat_def add_diff_cancel_right' assms(2) cartesian_power_car_memE gr_zeroI less_numeral_extra(1) less_numeral_extra(4) list.size(4) nth_Cons_pos Qp.to_R1_closed Qp.to_R_to_R1 zero_less_diff) have 3: "Qp_ev ((pvar Q\<^sub>p 1)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (b[^](4::nat))" by (metis "0" "2" eval_at_point_nat_pow local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 4: "Qp_ev ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)) [a,b] = (a[^](4::nat))" using "0" "1" eval_at_point_nat_pow local.pvar_closed pos2 by presburger then have 5: "Qp_ev (poly_scalar_mult Q\<^sub>p (\[^](3::nat)) ((pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat))) [a,b] = (\[^](3::nat))\ (a[^](4::nat))" using eval_at_point_smult[of "[a,b]" 2 "(pvar Q\<^sub>p 0)[^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub>(4::nat)" "\[^](3::nat)" ] 2 by (metis "0" MP.nat_pow_closed eval_at_point_scalar_mult local.pvar_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis proof- have 00: "[a, b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" by (simp add: "0") have 01: " pvar Q\<^sub>p 1 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed one_less_numeral_iff semiring_norm(76)) have 02: "\[^](3::nat) \\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (pvar Q\<^sub>p 0 [^]\<^bsub>Q\<^sub>p[\\<^bsub>2\<^esub>]\<^esub> (4::nat)) \ carrier (Q\<^sub>p[\\<^bsub>2\<^esub>])" by (meson MP.nat_pow_closed local.pvar_closed local.smult_closed p_natpow_closed(1) zero_less_numeral) then show ?thesis unfolding Qp_val_poly_def using 00 01 02 by (metis (no_types, lifting) "3" "4" MP.nat_pow_closed eval_at_point_add eval_at_point_smult local.pvar_closed p_natpow_closed(1) zero_less_numeral) qed qed lemma Qp_2I: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "[a,b] \ carrier (Q\<^sub>p\<^bsup>2\<^esup>)" apply(rule cartesian_power_car_memI) using assms apply (simp add: assms(1) assms(2)) using assms by (simp add: assms(1) assms(2)) lemma pair_id: assumes "length as = 2" shows "as = [as!0, as!1]" using assms by (smt One_nat_def diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons' nth_equalityI numeral_2_eq_2) lemma Qp_val_semialg: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "val b \ val a \ [a,b] \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val a \ val b \ [a, b] \ basic_semialg_set 2 2 Qp_val_poly" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def by (metis (mono_tags, lifting) assms(1) assms(2) mem_Collect_eq) show "[a, b] \ basic_semialg_set 2 2 Qp_val_poly \ val a \ val b" using Qp_val_alt_def[of a b] Qp_2I[of a b] Qp_val_poly_eval[of a b] unfolding basic_semialg_set_def using assms(1) assms(2) by blast qed definition val_relation_set where "val_relation_set = {as \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (as!1) \ val (as!0)}" lemma val_relation_setE: assumes "as \ val_relation_set" shows "as!0 \ carrier Q\<^sub>p \ as!1 \ carrier Q\<^sub>p \ as = [as!0,as!1] \ val (as!1) \ val (as!0)" using assms unfolding val_relation_set_def by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff pair_id pos2 semiring_norm(76)) lemma val_relation_setI: assumes "as!0 \ carrier Q\<^sub>p" assumes "as!1 \ carrier Q\<^sub>p" assumes "length as = 2" assumes "val (as!1) \ val(as!0)" shows "as \ val_relation_set" unfolding val_relation_set_def using assms Qp_2I[of "as!0" "as!1"] by (metis (no_types, lifting) mem_Collect_eq pair_id) lemma val_relation_semialg: "val_relation_set = basic_semialg_set 2 (2::nat) Qp_val_poly" proof show "val_relation_set \ basic_semialg_set 2 (2::nat) Qp_val_poly" proof fix as assume A: "as \ val_relation_set" have 0: "length as = 2" unfolding val_relation_set_def by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq val_relation_set_def) have 1: "as = [as ! 0, as ! 1]" by (metis (no_types, lifting) A cartesian_power_car_memE mem_Collect_eq pair_id val_relation_set_def) show "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" using A 1 val_relation_setE[of as] Qp_val_semialg[of "as!0" "as!1"] by presburger qed show "basic_semialg_set 2 (2::nat) Qp_val_poly \ val_relation_set" proof fix as assume "as \ basic_semialg_set 2 (2::nat) Qp_val_poly" then show "as \ val_relation_set" using val_relation_setI[of as] by (smt cartesian_power_car_memE cartesian_power_car_memE' mem_Collect_eq one_less_numeral_iff Qp_val_semialg basic_semialg_set_def val_relation_set_def padic_fields_axioms pair_id pos2 semiring_norm(76)) qed qed lemma val_relation_is_semialgebraic: "is_semialgebraic 2 val_relation_set" proof - have "{rs \ carrier (Q\<^sub>p\<^bsup>2\<^esup>). val (rs ! 0) \ val (rs ! 1)} = basic_semialg_set (Suc 1) (Suc 1) Qp_val_poly" using Suc_1 val_relation_semialg val_relation_set_def by presburger then show ?thesis by (metis (no_types) Qp_val_poly_closed Suc_1 basic_semialg_is_semialgebraic' val_relation_set_def zero_neq_numeral) qed lemma Qp_val_ring_is_semialg: obtains P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" proof- obtain P where P_def: "P = (\[^](3::nat)) \\<^bsub>Q\<^sub>p_x \<^esub>(X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat) \\<^bsub>Q\<^sub>p_x\<^esub> \\<^bsub>Q\<^sub>p_x\<^esub>" by blast have 0: "P \ carrier Q\<^sub>p_x" proof- have 0: "(X_poly Q\<^sub>p) \ carrier Q\<^sub>p_x" using UPQ.X_closed by blast then show ?thesis using P_def UPQ.P.nat_pow_closed p_natpow_closed(1) by blast qed have 1: "\\<^sub>p = univ_basic_semialg_set 2 P" proof show "\\<^sub>p \ univ_basic_semialg_set 2 P" proof fix x assume A: "x \ \\<^sub>p" show "x \ univ_basic_semialg_set 2 P" proof- have x_car: "x \ carrier Q\<^sub>p" using A val_ring_memE by blast then have "(\y \ carrier Q\<^sub>p. \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat)))" using A Qp_val_ring_alt_def[of x] by blast then obtain y where y_def: "y \ carrier Q\<^sub>p \ \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat)) = (y[^](2::nat))" by blast have "y \ carrier Q\<^sub>p \ P \ x = (y[^](2::nat))" proof- have "P \ x = \ \\<^bsub>Q\<^sub>p\<^esub> (\[^](3::nat))\ (x[^](4::nat))" proof- have "((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat)) \ carrier Q\<^sub>p_x" using UPQ.monom_closed p_natpow_closed(1) by blast then have "P \ x = (((\[^](3::nat)) \\<^bsub>Q\<^sub>p_x\<^esub> (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using P_def x_car UPQ.to_fun_plus by blast then have 0: "P \ x = (\[^](3::nat)) \(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) \\<^bsub>Q\<^sub>p\<^esub> (\\<^bsub>Q\<^sub>p_x\<^esub> \ x)" using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_smult p_natpow_closed(1) x_car by presburger have "(( (X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> (4::nat))\ x) = (x[^](4::nat))" using UPQ.to_fun_X_pow x_car by blast then have "P \ x = (\[^](3::nat)) \(x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using "0" UPQ.to_fun_one x_car by presburger then show ?thesis using y_def Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1) x_car by presburger qed then show ?thesis using y_def by blast qed then show ?thesis unfolding univ_basic_semialg_set_def using x_car by blast qed qed show "univ_basic_semialg_set 2 P \ \\<^sub>p" proof fix x assume A: "x \ univ_basic_semialg_set (2::nat) P" then obtain y where y_def: "y \ carrier Q\<^sub>p \ (P \ x) = (y[^](2::nat))" unfolding univ_basic_semialg_set_def by blast have x_car: "x \ carrier Q\<^sub>p" using A by (metis (no_types, lifting) mem_Collect_eq univ_basic_semialg_set_def) have 0: "(P \ x) = (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \" using P_def x_car UPQ.UP_one_closed UPQ.monom_closed UPQ.monom_rep_X_pow UPQ.to_fun_monom UPQ.to_fun_one UPQ.to_fun_plus p_natpow_closed(1) by presburger have 1: "y \ carrier Q\<^sub>p \ (\[^](3::nat)) \ (x[^](4::nat)) \\<^bsub>Q\<^sub>p\<^esub> \ = (y[^](2::nat))" using "0" y_def by blast then show "x \ \\<^sub>p" using x_car Qp_val_ring_alt_def[of x] y_def by (metis Qp.add.m_comm Qp.one_closed local.monom_term_car p_natpow_closed(1)) qed qed show ?thesis using 0 1 that by blast qed lemma Qp_val_ring_is_univ_semialgebraic: "is_univ_semialgebraic \\<^sub>p" proof- obtain P where "P \ carrier Q\<^sub>p_x \ \\<^sub>p = univ_basic_semialg_set 2 P" using Qp_val_ring_is_semialg by blast then show ?thesis by (metis univ_basic_semialg_set_is_univ_semialgebraic zero_neq_numeral) qed lemma Qp_val_ring_is_semialgebraic: "is_semialgebraic 1 (to_R1` \\<^sub>p)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def by blast (********************************************************************************************) (********************************************************************************************) subsubsection\Inverse Images of Semialgebraic Sets by Polynomial Maps\ (********************************************************************************************) (********************************************************************************************) lemma basic_semialg_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>])" assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S = basic_semialg_set k m f" assumes "m \0" shows "poly_map n fs \\<^bsub>n\<^esub> S = basic_semialg_set n m (Qp_poly_comp n fs f)" proof show "poly_map n fs \\<^bsub>n\<^esub> S \ basic_semialg_set n m (Qp_poly_comp n fs f)" proof fix x assume A: "x \ poly_map n fs \\<^bsub>n\<^esub> S" then have 0: "poly_map n fs x \ S" proof - have "\n f. {rs. rs \ S} \ {rs \ carrier (Q\<^sub>p\<^bsup>k\<^esup>). \r. r \ carrier Q\<^sub>p \ Qp_ev f rs = (r[^](n::nat))}" by (metis (no_types) Collect_mem_eq \S = basic_semialg_set k m f\ basic_semialg_set_def eq_iff) then show ?thesis using A by blast qed have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A assms by (meson evimage_eq) have "\y \ (carrier Q\<^sub>p). Qp_ev f (poly_map n fs x) = (y[^]m)" using A 0 assms basic_semialg_set_def by blast then have "\y \ (carrier Q\<^sub>p). Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using 1 assms Qp_poly_comp_eval by blast then show "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" using "1" basic_semialg_set_def by blast qed show "basic_semialg_set n m (Qp_poly_comp n fs f) \ poly_map n fs \\<^bsub>n\<^esub> S" proof fix x assume A: "x \ basic_semialg_set n m (Qp_poly_comp n fs f)" have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using A basic_semialg_set_def by blast have 1: "(poly_map n fs x) \ carrier (Q\<^sub>p\<^bsup>k\<^esup>)" using "0" poly_map_closed assms(2) assms(3) by blast show "x \ poly_map n fs \\<^bsub>n\<^esub> S" proof- have "\y \ carrier Q\<^sub>p. Qp_ev (Qp_poly_comp n fs f) x = (y[^]m)" using A basic_semialg_set_def by blast then have 2: "\y \ carrier Q\<^sub>p. Qp_ev f (poly_map n fs x) = (y[^]m)" using assms Qp_poly_comp_eval by (metis (no_types, lifting) A basic_semialg_set_def mem_Collect_eq) have 3: "poly_map n fs x \ S" using assms 0 1 basic_semialg_set_def[of k m f] "2" by blast show ?thesis using "0" "3" by blast qed qed qed lemma basic_semialg_pullback': assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "A \ basic_semialgs k" shows "poly_map n fs \\<^bsub>n\<^esub> A \ (basic_semialgs n)" proof- obtain f m where fm_def: "m \0 \f \ carrier (Q\<^sub>p[\\<^bsub>k\<^esub>]) \ A = basic_semialg_set k m f" using assms by (metis is_basic_semialg_def mem_Collect_eq) then have "poly_map n fs \\<^bsub>n\<^esub> A = basic_semialg_set n m (Qp_poly_comp n fs f)" using assms basic_semialg_pullback[of f k n fs A m] by linarith then show ?thesis unfolding is_basic_semialg_def by (metis (mono_tags, lifting) assms(1) assms(2) fm_def mem_Collect_eq poly_compose_closed) qed lemma semialg_pullback: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "poly_map n fs \\<^bsub>n\<^esub> S \ semialg_sets n" unfolding semialg_sets_def apply(rule gen_boolean_algebra.induct[of S "(carrier (Q\<^sub>p\<^bsup>k\<^esup>))" "basic_semialgs k"]) using assms semialg_sets_def apply blast apply (metis assms(1) assms(2) carrier_is_semialgebraic evimageI2 extensional_vimage_closed is_semialgebraicE poly_map_closed semialg_sets_def subsetI subset_antisym) apply (metis Int_absorb2 assms(1) assms(2) basic_semialg_is_semialg basic_semialg_is_semialgebraic basic_semialg_pullback' is_semialgebraic_closed mem_Collect_eq semialg_sets_def) apply (metis evimage_Un semialg_sets_def semialg_union) by (metis assms(1) assms(2) carrier_is_semialgebraic diff_is_semialgebraic evimage_Diff extensional_vimage_closed is_semialgebraicE is_semialgebraicI poly_map_closed poly_map_pullbackI semialg_sets_def subsetI subset_antisym) lemma pullback_is_semialg: assumes "is_poly_tuple n fs" assumes "length fs = k" assumes "S \ semialg_sets k" shows "is_semialgebraic n (poly_map n fs \\<^bsub>n\<^esub> S)" using assms(1) assms(2) assms(3) is_semialgebraicI padic_fields_axioms semialg_pullback by blast text\Equality and inequality sets for a pair of polynomials\ definition val_ineq_set where "val_ineq_set n f g = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" lemma poly_map_length : assumes "length fs = m" assumes "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" shows "length (poly_map n fs as) = m" using assms unfolding poly_map_def poly_tuple_eval_def by (metis (no_types, lifting) length_map restrict_apply') lemma val_ineq_set_pullback: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g = poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set " proof show "val_ineq_set n f g \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" proof fix x assume "x \ val_ineq_set n f g" then have 0: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ val (Qp_ev f x) \ val (Qp_ev g x)" by (metis (mono_tags, lifting) mem_Collect_eq val_ineq_set_def) have 1: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" unfolding poly_map_def poly_tuple_eval_def using 0 by (metis (no_types, lifting) Cons_eq_map_conv list.simps(8) restrict_apply') have 2: "poly_map n [g,f] x \ val_relation_set" apply(rule val_relation_setI) using 1 0 assms apply (metis eval_at_point_closed nth_Cons_0) using 1 0 assms apply (metis One_nat_def eval_at_point_closed diff_Suc_1 less_numeral_extra(1) nth_Cons_pos Qp.to_R_to_R1) using poly_map_length assms 0 apply (metis "1" Qp_2I cartesian_power_car_memE eval_at_point_closed) by (metis "0" "1" One_nat_def nth_Cons_0 nth_Cons_Suc) have 3: "is_poly_tuple n [g, f]" using assms by (smt One_nat_def diff_Suc_1 Qp_is_poly_tupleI length_Suc_conv less_SucE less_one list.size(3) nth_Cons') then show "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" using 0 1 2 by blast qed show "poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set \ val_ineq_set n f g" proof fix x have 0: "is_poly_tuple n [g, f]" using Qp_is_poly_tupleI assms by (metis (no_types, lifting) diff_Suc_1 length_Cons less_Suc0 less_SucE list.size(3) nth_Cons') assume A: "x \ poly_map n [g,f] \\<^bsub>n\<^esub> val_relation_set" then have 1: "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ poly_map n [g,f] x \ val_relation_set" using 0 by (meson evimageD extensional_vimage_closed subsetD) have 2: "poly_map n [g,f] x = [Qp_ev g x, Qp_ev f x]" by (metis "1" Qp_poly_mapE' length_0_conv poly_map_cons) show "x \ val_ineq_set n f g" using 0 1 2 unfolding val_ineq_set_def val_relation_set_def by (metis (no_types, lifting) "1" list.inject mem_Collect_eq nth_Cons_0 poly_map_apply val_relation_setE) qed qed lemma val_ineq_set_is_semialg: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "val_ineq_set n f g \ semialg_sets n" proof- have 0: "val_relation_set \ semialg_sets 2" using val_relation_semialg basic_semialg_is_semialg' by (metis Qp_val_poly_closed zero_neq_numeral) show ?thesis using val_ineq_set_pullback semialg_pullback[of n "[g,f]" 2 "val_relation_set" ] by (metis (no_types, lifting) "0" assms(1) assms(2) diff_Suc_1 Qp_is_poly_tupleI length_Cons less_Suc0 less_SucE list.size(3) nth_Cons_0 nth_Cons_pos numeral_2_eq_2 zero_neq_numeral) qed lemma val_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n (val_ineq_set n f g)" using assms(1) assms(2) is_semialgebraicI val_ineq_set_is_semialg by blast lemma val_ineq_setI: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ (val_ineq_set n f g)" shows "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" "val (Qp_ev f x) \ val (Qp_ev g x)" using assms unfolding val_ineq_set_def apply blast using assms unfolding val_ineq_set_def by blast lemma val_ineq_setE: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" assumes "val (Qp_ev f x) \ val (Qp_ev g x)" shows "x \ (val_ineq_set n f g)" using assms unfolding val_ineq_set_def by blast lemma val_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast lemma val_eq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" proof- have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 1: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" using assms val_ineq_set_is_semialgebraic unfolding val_ineq_set_def by blast have 2: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} \ {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule equalityI, rule subsetI , rule IntI) unfolding mem_Collect_eq using le_less apply blast apply (metis order_refl) apply(rule subsetI, erule IntE) unfolding mem_Collect_eq by (meson less_le_trans not_less_iff_gr_or_eq) show ?thesis unfolding 2 apply(rule intersection_is_semialg) using 0 apply blast using 1 by blast qed lemma equalityI'': assumes "\x. A x \ B x" assumes "\x. B x \ A x" shows "{x. A x} = {x. B x}" using assms by blast lemma val_strict_ineq_set_is_semialgebraic: assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" proof- have 0: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} - {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule equalityI', rule DiffI) unfolding le_less mem_Collect_eq apply blast unfolding mem_Collect_eq using neq_iff apply blast apply(erule DiffE) unfolding mem_Collect_eq by blast show ?thesis unfolding 0 apply(rule diff_is_semialgebraic) using assms val_ineq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def apply blast using assms val_eq_set_is_semialgebraic[of f n g] unfolding val_ineq_set_def by blast qed lemma constant_poly_val_exists: shows "\g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]). (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" proof- obtain a where a_def: "a \ carrier Q\<^sub>p \ val a = c" by (meson Qp.minus_closed Qp.nonzero_closed dist_nonempty' p_nonzero) obtain g where g_def: "g = coord_const a" by blast show ?thesis using a_def g_def Qp_to_IP_car by (metis (no_types, opaque_lifting) Qp_to_IP_car a_def eval_at_point_const g_def le_less subset_iff) qed lemma val_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)}" apply(rule val_ineq_set_is_semialgebraic') using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) \ c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_ineq_set_is_semialgebraic''': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)}" apply(rule val_ineq_set_is_semialgebraic') using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) \ val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c \ val (Qp_ev f x)}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 unfolding 1 by blast qed lemma val_eq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)}" apply(rule val_eq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) = c}" apply(rule equalityI'') using g_def apply fastforce using g_def by metis show ?thesis using 0 unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using assms apply blast using g_def by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < val (Qp_ev g x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev f x) < c}" apply(rule equalityI'') using g_def apply fastforce using g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma val_strict_ineq_set_is_semialgebraic'': assumes "f \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" proof- obtain g where g_def: "g \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ (\ x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) = c)" using constant_poly_val_exists by blast have 0: "is_semialgebraic n {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)}" apply(rule val_strict_ineq_set_is_semialgebraic) using g_def apply blast using assms by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). val (Qp_ev g x) < val (Qp_ev f x)} = {x \ carrier (Q\<^sub>p\<^bsup>n\<^esup>). c < val (Qp_ev f x)}" apply(rule equalityI'') using assms g_def apply fastforce using assms g_def by fastforce show ?thesis using 0 g_def unfolding 1 by blast qed lemma(in cring) R1_memE: assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "x = [(hd x)]" using assms cartesian_power_car_memE by (metis diff_is_0_eq' hd_conv_nth le_eq_less_or_eq length_0_conv length_tl list.exhaust list.sel(3) normalize.cases nth_Cons_0 zero_neq_one) lemma(in cring) R1_memE': assumes "x \ carrier (R\<^bsup>1\<^esup>)" shows "hd x \ carrier R" using R1_memE assms cartesian_power_car_memE[of x R 1] cartesian_power_car_memE'[of x R 1 0] by (metis hd_conv_nth less_numeral_extra(1) list.size(3) zero_neq_one) lemma univ_val_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x \ c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c}" apply(rule val_ineq_set_is_semialgebraic'') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) \ c} = to_R1 ` {x \ carrier Q\<^sub>p. val x \ c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x \ c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a \ c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) \ c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_strict_ineq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x < c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) < c} = to_R1 ` {x \ carrier Q\<^sub>p. val x < c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) then show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c}" using A 0 unfolding mem_Collect_eq by (metis (no_types, lifting) image_iff mem_Collect_eq) qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x < c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a < c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) < c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed lemma univ_val_eq_set_is_univ_semialgebraic: "is_univ_semialgebraic {x \ carrier Q\<^sub>p. val x = c}" proof- have 0: "is_semialgebraic 1 {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c}" apply(rule val_eq_set_is_semialgebraic') using pvar_closed by blast have 1: "{x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (Qp_ev (pvar Q\<^sub>p 0) x) = c} = to_R1 ` {x \ carrier Q\<^sub>p. val x = c}" proof(rule equalityI') show " \x. x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c} \ x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" proof- fix x assume A: "x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" then have 0: "x = [(hd x)] \ hd x \ carrier Q\<^sub>p" using Qp.R1_memE Qp.R1_memE' by blast have 1: "eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0) = hd x" using A 0 by (metis (no_types, lifting) One_nat_def eval_pvar lessI nth_Cons_0 Qp.to_R1_closed) show "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c}" using A 0 unfolding mem_Collect_eq 1 by blast qed show "\x. x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} \ x \ {x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>). val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c}" proof fix x assume A: "x \ (\a. [a]) ` {x \ carrier Q\<^sub>p. val x = c} " then obtain a where a_def: "x = [a] \ a \ carrier Q\<^sub>p \ val a = c" by blast then have 0: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using cartesian_power_car_memI Qp.to_R1_closed by presburger then have 1: "(eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = a" using a_def by (metis eval_pvar less_one Qp.to_R_to_R1) show "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>) \ val (eval_at_point Q\<^sub>p x (pvar Q\<^sub>p 0)) = c" unfolding 1 using a_def 0 by blast qed qed show ?thesis using 0 unfolding 1 using is_univ_semialgebraicI by blast qed (********************************************************************************************) (********************************************************************************************) subsubsection\One Dimensional $p$-adic Balls are Semialgebraic\ (********************************************************************************************) (********************************************************************************************) lemma coord_ring_one_def: "Pring Q\<^sub>p {(0::nat)} = (Q\<^sub>p[\\<^bsub>1\<^esub>])" proof- have "{(0::nat)} = {..<1}" by auto thus ?thesis unfolding coord_ring_def by auto qed lemma times_p_pow_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]n \ a" shows "val b = val a + n" using val_mult[of "\[^]n" a] assms unfolding assms(2) val_p_int_pow by (metis add.commute p_intpow_closed(1)) lemma times_p_pow_neg_val: assumes "a \ carrier Q\<^sub>p" assumes "b = \[^]-n \ a" shows "val b = val a - n" by (metis Qp.m_comm Qp_int_pow_nonzero assms(1) assms(2) p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) lemma eint_minus_int_pos: assumes "a - eint n \ 0" shows "a \ n" using assms apply(induction a) apply (metis diff_ge_0_iff_ge eint_ord_simps(1) idiff_eint_eint zero_eint_def) by simp text\\p\-adic balls as pullbacks of polynomial maps\ lemma balls_as_pullbacks: assumes "c \ carrier Q\<^sub>p" shows "\P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]). to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p)" proof- obtain P0 where P0_def: "P0 = (to_poly (\[^](-n))) \\<^bsub>Q\<^sub>p_x\<^esub>((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c)" by blast have 0: "P0 \ carrier Q\<^sub>p_x" proof- have P0: "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have P1: "(to_poly (\[^](-n))) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast then show ?thesis using P0_def P0 P1 by blast qed have 1: "\x. x \ carrier Q\<^sub>p \ P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" proof- fix x assume A: "x \ carrier Q\<^sub>p" have P0: "(to_poly (\[^](-n))) \ x = (\[^](-n))" using A UPQ.to_fun_to_poly p_intpow_closed(1) by blast have P1: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (x \\<^bsub>Q\<^sub>p\<^esub> c)" by (metis A UPQ.to_fun_X_minus X_poly_minus_def assms) have P2: "to_poly (\[^](-n)) \ carrier Q\<^sub>p_x" using UPQ.to_poly_closed p_intpow_closed(1) by blast have P3: "((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ carrier Q\<^sub>p_x" using UPQ.X_closed UPQ.to_poly_closed assms by blast have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = to_poly (\[^]- n) \ x \ (((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x)" using A P0_def P0 P1 P2 P3 to_fun_mult[of "to_poly (\[^](-n))" "(X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c" x] UPQ.to_fun_mult by blast then have "to_poly (\[^]- n) \\<^bsub>Q\<^sub>p_x\<^esub> ((X_poly Q\<^sub>p) \\<^bsub>Q\<^sub>p_x\<^esub> to_poly c) \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c) " by (metis P0 P1) then show "P0 \ x = (\[^](-n)) \ (x \\<^bsub>Q\<^sub>p\<^esub> c)" using P0_def by metis qed have 2: " (\a. [a]) ` B\<^bsub>n\<^esub>[c] = poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof show "(\a. [a]) ` B\<^bsub>n\<^esub>[c] \ poly_map 1 [from_Qp_x P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" proof fix x assume A: "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" then obtain a where a_def: "x = [a] \ a \ B\<^bsub>n\<^esub>[c]" by blast have P0: "P0 \ a \ \\<^sub>p" proof- have "B\<^bsub>n\<^esub>[c] \ carrier Q\<^sub>p" using c_ball_in_Qp by blast hence a_closed: "a \ carrier Q\<^sub>p" using a_def by blast have P0: "P0 \ a = (\[^](-n)) \ (a \ c)" using 1 a_def c_ballE(1) by blast then have P1: "val (P0 \ a) = val (\[^](-n)) + val (a \ c)" using val_mult[of "\[^]-n" "a \ c"] a_closed assms Qp.minus_closed p_intpow_closed(1) by presburger then have P2: "val (P0 \ a) = val (a \\<^bsub>Q\<^sub>p\<^esub> c) - n" by (metis P0 Qp.m_comm Qp.minus_closed Qp_int_pow_nonzero assms local.a_closed p_intpow_closed(1) p_intpow_inv'' p_nonzero val_fract val_p_int_pow) have P3: "val (a \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using a_def c_ballE(2) by blast then have "val (P0 \ a) \ -n + n" using P2 by (metis add.commute diff_conv_add_uminus diff_self local.eint_minus_ineq' zero_eint_def) then have P4: "val (P0 \ a) \ 0" by (metis add.commute add.right_inverse zero_eint_def) have P5: "P0 \ a \ carrier Q\<^sub>p" using "0" UPQ.to_fun_closed local.a_closed by blast then show ?thesis using P4 using val_ring_val_criterion by blast qed have "poly_map 1 [from_Qp_x P0] x = [Qp_ev (from_Qp_x P0) [a]] " using a_def poly_map_def[of 1 "[from_Qp_x P0]"] poly_tuple_eval_def[of ] by (metis Qp_poly_mapE' c_ballE(1) length_0_conv poly_map_cons Qp.to_R1_closed) then have "poly_map 1 [from_Qp_x P0] x = [P0 \ a] " using Qp_x_Qp_poly_eval[of P0 a] by (metis "0" a_def c_ballE(1)) then have P1: "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using P0 by blast have P2: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A c_ballE(1) Qp.to_R1_closed by blast have P3: "is_poly_tuple 1 [from_Qp_x P0]" apply(rule Qp_is_poly_tupleI) by (metis "0" Qp_is_poly_tupleI from_Qp_x_closed gr_implies_not0 is_poly_tupleE is_poly_tuple_Cons list.size(3) zero_neq_one) show "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p" using P3 P2 P1 unfolding evimage_def poly_map_def by blast qed have 20: "is_poly_tuple 1 [from_Qp_x P0]" using 0 UP_to_IP_closed[of P0 "0::nat"] unfolding is_poly_tuple_def by (metis (no_types, lifting) empty_set from_Qp_x_closed list.simps(15) singletonD subset_code(1)) show "poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> (\a. [a]) ` \\<^sub>p \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" proof fix x assume A: "x \ poly_map 1 [UP_to_IP Q\<^sub>p 0 P0] \\<^bsub>1\<^esub> ((\a. [a]) ` \\<^sub>p)" have A0: "(\a. [a]) ` \\<^sub>p \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using Qp_val_ring_is_univ_semialgebraic is_univ_semialgebraic_def Qp.to_R1_car_subset Qp_val_ring_is_semialgebraic is_semialgebraic_closed by presburger have "poly_map 1 [from_Qp_x P0] x \ ((\a. [a]) ` \\<^sub>p)" using A0 A 20 by blast then obtain a where a_def: "a \ \\<^sub>p \ (poly_map 1 [from_Qp_x P0] x) = [a]" by blast have x_closed: "x \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using A by (meson evimage_eq) then obtain y where y_def: "x = [y] \ y \ carrier Q\<^sub>p" using A by (metis Qp.to_R1_to_R Qp.to_R_pow_closed) have "(poly_map 1 [from_Qp_x P0] x) = [(Qp_ev (from_Qp_x P0) [y])]" unfolding poly_map_def poly_tuple_eval_def using x_closed by (smt "20" One_nat_def length_Suc_conv list.size(3) nth_Cons_0 nth_map poly_tuple_eval_closed poly_tuple_eval_def restrict_apply' Qp.to_R1_to_R y_def zero_less_Suc) then have "(poly_map 1 [from_Qp_x P0] x) = [P0 \ y]" by (metis "0" Qp_x_Qp_poly_eval y_def) then have "[a] = [P0 \ y]" using a_def by presburger then have A1: "a = (\[^](-n)) \ (y \\<^bsub>Q\<^sub>p\<^esub> c)" using 1[of y] y_def by blast have "y \ B\<^bsub>n\<^esub>[c]" proof- have B0: "val a = val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n" using A1 y_def Qp.minus_closed assms times_p_pow_neg_val by blast have B1: "val a \ 0" using a_def val_ring_memE by blast then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) - n \ 0" using B0 by metis then have "val (y \\<^bsub>Q\<^sub>p\<^esub> c) \ n" using eint_minus_int_pos by blast then show "y \ B\<^bsub>n\<^esub>[c]" using c_ballI y_def by blast qed then show "x \ (\a. [a]) ` B\<^bsub>n\<^esub>[c]" using y_def by blast qed qed then show ?thesis by (meson "0" from_Qp_x_closed) qed lemma ball_is_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1` B\<^bsub>n\<^esub>[c])" proof- obtain P where P_def: "P \ carrier (Q\<^sub>p[\\<^bsub>1\<^esub>]) \ to_R1` B\<^bsub>n\<^esub>[c] = poly_map 1 [P] \\<^bsub>1\<^esub> (to_R1 ` \\<^sub>p) " using assms balls_as_pullbacks[of c n] by meson have "is_poly_tuple 1 [P]" using P_def unfolding is_poly_tuple_def by (metis (no_types, opaque_lifting) list.inject list.set_cases neq_Nil_conv subset_code(1)) then show ?thesis using assms P_def pullback_is_semialg[of 1 "[P]" 1 "((\a. [a]) ` \\<^sub>p) "] by (metis (mono_tags, lifting) One_nat_def Qp_val_ring_is_semialgebraic is_semialgebraic_def length_Suc_conv list.distinct(1) list.size(3)) qed lemma ball_is_univ_semialgebraic: assumes "c \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (B\<^bsub>n\<^esub>[c])" using assms ball_is_semialgebraic c_ball_in_Qp is_univ_semialgebraic_def by presburger abbreviation Qp_to_R1_set where "Qp_to_R1_set S \ to_R1 ` S" (********************************************************************************************) (********************************************************************************************) subsubsection\Finite Unions and Intersections of Semialgebraic Sets\ (********************************************************************************************) (********************************************************************************************) definition are_semialgebraic where "are_semialgebraic n Xs = (\ x. x \ Xs \ is_semialgebraic n x)" lemma are_semialgebraicI: assumes "\x. x \ Xs \ is_semialgebraic n x " shows "are_semialgebraic n Xs" using are_semialgebraic_def assms by blast lemma are_semialgebraicE: assumes "are_semialgebraic n Xs" assumes "x \ Xs" shows "is_semialgebraic n x" using are_semialgebraic_def assms(1) assms(2) by blast definition are_univ_semialgebraic where "are_univ_semialgebraic Xs = (\ x. x \ Xs \ is_univ_semialgebraic x)" lemma are_univ_semialgebraicI: assumes "\x. x \ Xs \ is_univ_semialgebraic x " shows "are_univ_semialgebraic Xs" using are_univ_semialgebraic_def assms by blast lemma are_univ_semialgebraicE: assumes "are_univ_semialgebraic Xs" assumes "x \ Xs" shows "is_univ_semialgebraic x" using are_univ_semialgebraic_def assms(1) assms(2) by blast lemma are_univ_semialgebraic_semialgebraic: assumes "are_univ_semialgebraic Xs" shows "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" apply(rule are_semialgebraicI) using are_univ_semialgebraicE assms image_iff is_univ_semialgebraicE by (metis (no_types, lifting)) lemma to_R1_set_union: "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" using image_iff by blast lemma to_R1_inter: assumes "Xs \ {}" shows "to_R1 ` (\ Xs) = \ (Qp_to_R1_set ` Xs)" proof show "to_R1 ` (\ Xs) \ \ (Qp_to_R1_set ` Xs)" by blast show "\ (Qp_to_R1_set ` Xs) \ to_R1 ` (\ Xs)" proof fix x assume A: "x \ \ (Qp_to_R1_set ` Xs)" then have 0: "\S. S \ Xs \ x \ (Qp_to_R1_set S)" by blast obtain S where "S \ Xs \ x \ (Qp_to_R1_set S)" using assms 0 by blast then obtain b where b_def: "b \ S \ x = [b]" by blast have "b \ (\ Xs)" using "0" b_def by blast then show "x \ to_R1 ` (\ Xs)" using b_def by blast qed qed lemma finite_union_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply (simp add: empty_is_semialgebraic) by (metis Sup_insert insert_subset is_semialgebraicI union_is_semialgebraic) lemma finite_union_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n " shows "is_semialgebraic n (\ Xs)" using assms(1) assms(2) finite_union_is_semialgebraic by blast lemma(in padic_fields) finite_union_is_semialgebraic'': assumes "finite S" assumes "\x. x \ S \ is_semialgebraic m (F x)" shows "is_semialgebraic m (\ x \ S. F x)" using assms finite_union_is_semialgebraic[of "F`S" m] unfolding is_semialgebraic_def by blast lemma finite_union_is_univ_semialgebraic': assumes "finite Xs" assumes "are_univ_semialgebraic Xs" shows "is_univ_semialgebraic (\ Xs)" proof- have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms finite_union_is_semialgebraic'[of "((`) (\a. [a]) ` Xs)"] to_R1_set_union[of Xs] by (metis (no_types, lifting) are_semialgebraicE are_univ_semialgebraic_semialgebraic finite_imageI is_semialgebraicE subsetI) then show ?thesis using is_univ_semialgebraicI by blast qed lemma finite_intersection_is_semialgebraic: assumes "finite Xs" shows "Xs \ semialg_sets n \ Xs \{} \ is_semialgebraic n (\ Xs)" apply(rule finite.induct[of Xs]) apply (simp add: assms) apply auto[1] proof fix A::"((nat \ int) \ (nat \ int)) set list set set" fix a assume 0: "finite A" assume 1: "A \ semialg_sets n \ A \ {} \ is_semialgebraic n (\ A) " assume 2: "insert a A \ semialg_sets n \ insert a A \ {}" show "is_semialgebraic n (\ (insert a A))" proof(cases "A = {}") case True then have "insert a A = {a}" by simp then show ?thesis by (metis "2" cInf_singleton insert_subset is_semialgebraicI) next case False then have "A \ semialg_sets n \ A \ {}" using "2" by blast then have "is_semialgebraic n (\ A) " using "1" by linarith then show ?thesis using 0 1 2 intersection_is_semialg by (metis Inf_insert insert_subset is_semialgebraicI) qed qed lemma finite_intersection_is_semialgebraic': assumes "finite Xs" assumes "Xs \ semialg_sets n \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (simp add: assms(1) assms(2) finite_intersection_is_semialgebraic) lemma finite_intersection_is_semialgebraic'': assumes "finite Xs" assumes "are_semialgebraic n Xs \ Xs \{}" shows " is_semialgebraic n (\ Xs)" by (meson are_semialgebraicE assms(1) assms(2) finite_intersection_is_semialgebraic' is_semialgebraicE subsetI) lemma finite_intersection_is_univ_semialgebraic: assumes "finite Xs" assumes "are_univ_semialgebraic Xs" assumes "Xs \ {}" shows "is_univ_semialgebraic (\ Xs)" proof- have "are_semialgebraic 1 (Qp_to_R1_set ` Xs)" using are_univ_semialgebraic_semialgebraic assms(2) by blast then have "is_semialgebraic 1 (\ (Qp_to_R1_set ` Xs))" using assms finite_intersection_is_semialgebraic''[of "Qp_to_R1_set ` Xs" 1] by blast then have "is_semialgebraic 1 (Qp_to_R1_set (\ Xs))" using assms to_R1_inter[of Xs] by presburger then show ?thesis using is_univ_semialgebraicI by blast qed (**************************************************************************************************) (**************************************************************************************************) subsection\Cartesian Products of Semialgebraic Sets\ (**************************************************************************************************) (**************************************************************************************************) lemma Qp_times_basic_semialg_right: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" proof show "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ basic_semialg_set (n + m) k a" proof fix x assume "x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" then obtain as bs where as_bs_def: "as \ (basic_semialg_set n k a) \ bs \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ x = as@bs" using cartesian_product_memE[of x "basic_semialg_set n k a" "carrier (Q\<^sub>p\<^bsup>m\<^esup>)" Q\<^sub>p n] basic_semialg_set_def by (metis (no_types, lifting) append_take_drop_id basic_semialg_set_memE(1) subsetI) have 0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using as_bs_def basic_semialg_set_memE(1) cartesian_product_closed' by blast have 1: "(Qp_ev a x = Qp_ev a as)" using as_bs_def poly_eval_cartesian_prod[of as n bs m a] assms basic_semialg_set_memE(1) by blast obtain y where y_def: "y \ carrier Q\<^sub>p \ (Qp_ev a as = (y[^]k))" using as_bs_def using basic_semialg_set_memE(2)[of as n k a] by blast show "x \ basic_semialg_set (n + m) k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "0") apply (simp add: y_def) using "1" y_def by blast qed show "basic_semialg_set (n + m) k a \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" proof fix x assume A: "x \ basic_semialg_set (n + m) k a" have A0: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A basic_semialg_set_memE(1) by blast have A1: "set x \ carrier Q\<^sub>p" using A0 by (metis (no_types, lifting) cartesian_power_car_memE cartesian_power_car_memE' in_set_conv_nth subsetI) have A2: "length x = n + m" using A0 cartesian_power_car_memE by blast obtain as where as_def: "as = take n x" by blast obtain bs where bs_def: "bs = drop n x" by blast have 0: "x = as@bs" using A as_def bs_def by (metis append_take_drop_id) have 1: "as \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" apply(rule cartesian_power_car_memI) using as_def A2 apply (simp add: A2 min.absorb2) by (metis (no_types, lifting) A1 as_def dual_order.trans set_take_subset) have 2: "bs \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" apply(rule cartesian_power_car_memI) using bs_def A2 apply (simp add: A2) by (metis A1 bs_def order.trans set_drop_subset) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)" using basic_semialg_set_memE A by meson have 3: "as \ basic_semialg_set n k a" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using \y \ carrier Q\<^sub>p \ Qp_ev a x = (y[^]k)\ apply blast using y_def A 1 0 2 assms(1) poly_eval_cartesian_prod by blast show " x \ cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>))" using 3 2 "0" by (metis (mono_tags, lifting) as_def basic_semialg_set_memE(1) bs_def cartesian_product_memI subsetI) qed qed lemma Qp_times_basic_semialg_right_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: " a \ carrier (Q\<^sub>p[\\<^bsub>n+m\<^esub>])" using assms poly_ring_car_mono'(2) by blast have 2: "is_semialgebraic (n + m) (basic_semialg_set (n + m) k a)" using assms basic_semialg_is_semialgebraic'[of a "n+m" k "basic_semialg_set (n + m) k a"] "1" by blast show ?thesis using 0 2 by metis qed lemma Qp_times_basic_semialg_right_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_right_is_semialgebraic[of k P] using assms(1) by blast qed lemma cartesian_product_memE': assumes "x \ cartesian_product A B" obtains a b where "a \ A \ b \ B \ x = a@b" using assms unfolding cartesian_product_def by blast lemma Qp_times_basic_semialg_left: assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) = basic_semialg_set (n+m) k (shift_vars n m a)" proof show "cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) \ basic_semialg_set (n + m) k (shift_vars n m a)" proof fix x assume A: "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" then obtain as bs where as_bs_def: "as \ (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ bs \ (basic_semialg_set n k a) \ x = as@bs " using cartesian_product_memE' by blast have 0: "Qp_ev (shift_vars n m a) x = Qp_ev a bs" using A as_bs_def assms shift_vars_eval[of a n as m bs ] by (metis (no_types, lifting) basic_semialg_set_memE(1)) obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev a bs = (y[^]k)" using as_bs_def basic_semialg_set_memE(2) by blast have 1: "x \ carrier (Q\<^sub>p\<^bsup>n+m\<^esup>)" using A as_bs_def by (metis (no_types, lifting) add.commute basic_semialg_set_memE(1) cartesian_product_closed') show "x \ basic_semialg_set (n + m) k (shift_vars n m a)" apply(rule basic_semialg_set_memI[of _ _ y]) apply (simp add: "1") using y_def apply blast using "0" y_def by blast qed show "basic_semialg_set (n + m) k (shift_vars n m a) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a) " proof fix x assume A: "x \ basic_semialg_set (n + m) k (shift_vars n m a)" then obtain y where y_def: "y \ carrier Q\<^sub>p \ Qp_ev (shift_vars n m a) x = (y[^]k)" using assms basic_semialg_set_memE[of x "n + m" k "shift_vars n m a"] shift_vars_closed[of a m] Qp.cring_axioms by blast have "x \ carrier (Q\<^sub>p\<^bsup>m+n\<^esup>)" using A basic_semialg_set_memE(1) by (metis add.commute) then have "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (carrier (Q\<^sub>p\<^bsup>n\<^esup>))" using cartesian_product_carrier by blast then obtain as bs where as_bs_def: "x = as@bs \ as \ carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ bs \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson cartesian_product_memE') have "bs \ (basic_semialg_set n k a)" apply(rule basic_semialg_set_memI[of _ _ y]) using as_bs_def apply blast apply (simp add: y_def) using y_def shift_vars_eval[of a n as m bs ] as_bs_def assms(1) by metis then show "x \ cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a)" using as_bs_def unfolding cartesian_product_def by blast qed qed lemma Qp_times_basic_semialg_left_is_semialgebraic: assumes "k > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialg_set n k a))" using basic_semialg_is_semialgebraic'[of a "n+m" k] Qp_times_basic_semialg_left by (metis assms(1) assms(2) basic_semialg_is_semialgebraic is_basic_semialg_def neq0_conv shift_vars_closed) lemma Qp_times_basic_semialg_left_is_semialgebraic': assumes "A \ basic_semialgs n" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) A)" proof- obtain k P where "k \ 0 \ P\carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])\ A = basic_semialg_set n k P" using assms is_basic_semialg_def by (metis mem_Collect_eq) then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic[of k P] using assms(1) by blast qed lemma product_of_basic_semialgs_is_semialg: assumes "k > 0" assumes "l > 0" assumes "a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>])" assumes "b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>])" shows "is_semialgebraic (n + m) (cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b))" proof- have 0: "cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = basic_semialg_set (n+ m) k a" using Qp_times_basic_semialg_right assms by presburger have 1: "cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b) = basic_semialg_set (m + n) l (shift_vars m n b)" using Qp_times_basic_semialg_left assms by blast have 2: "(cartesian_product (basic_semialg_set n k a) (basic_semialg_set m l b)) = cartesian_product (basic_semialg_set n k a) (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialg_set m l b)" proof- have 0: "basic_semialg_set n k a \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" using basic_semialg_set_memE(1) by blast have 1: "carrier (Q\<^sub>p\<^bsup>m\<^esup>) \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" by simp have 2: "carrier (Q\<^sub>p\<^bsup>n\<^esup>) \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by simp have 3: "basic_semialg_set m l b \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using basic_semialg_set_memE(1) by blast show ?thesis using 0 1 2 3 cartesian_product_intersection[of "(basic_semialg_set n k a)" Q\<^sub>p n "(carrier (Q\<^sub>p\<^bsup>m\<^esup>))" m "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))" "(basic_semialg_set m l b)"] by (smt Collect_cong inf.absorb_iff1 inf.absorb_iff2) qed then show ?thesis using Qp_times_basic_semialg_left_is_semialgebraic Qp_times_basic_semialg_right_is_semialgebraic assms by (metis (no_types, lifting) add.commute intersection_is_semialg) qed lemma product_of_basic_semialgs_is_semialg': assumes "A \ (basic_semialgs n)" assumes "B \ (basic_semialgs m)" shows "is_semialgebraic (n + m) (cartesian_product A B)" proof- obtain k a where ka_def: "k > 0 \ a \ carrier (Q\<^sub>p[\\<^bsub>n\<^esub>]) \ A = (basic_semialg_set n k a)" using assms by (metis is_basic_semialg_def mem_Collect_eq neq0_conv) obtain l b where lb_def: "l > 0 \ b \ carrier (Q\<^sub>p[\\<^bsub>m\<^esub>]) \ B = (basic_semialg_set m l b)" by (metis assms(2) gr_zeroI is_basic_semialg_def mem_Collect_eq) show ?thesis using ka_def lb_def assms product_of_basic_semialgs_is_semialg by blast qed lemma car_times_semialg_is_semialg: assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms is_semialgebraic_def semialg_sets_def apply blast apply (simp add: carrier_is_semialgebraic cartesian_product_carrier) proof- show "\A. A \ basic_semialgs m \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" proof- fix A assume A: "A \ basic_semialgs m " then have " (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)) = A" by (metis basic_semialg_set_memE(1) inf_absorb1 is_basic_semialg_def mem_Collect_eq subsetI) then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)))" using add.commute[of n m] assms A Qp_times_basic_semialg_left_is_semialgebraic' by (simp add: \n + m = m + n\) qed show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- fix A C assume A: " A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" " is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" then have B: "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A \ cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) C)" using union_is_semialgebraic by blast show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (A \ C))" proof- have 0: "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(1) gen_boolean_algebra_subset by blast have 1: " C \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using A(3) gen_boolean_algebra_subset by blast then show ?thesis using 0 A B using cartesian_product_binary_union_right[of A Q\<^sub>p m C "(carrier (Q\<^sub>p\<^bsup>n\<^esup>))"] unfolding is_semialgebraic_def semialg_sets_def by presburger qed qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) A)" then have "A \ carrier (Q\<^sub>p\<^bsup>m\<^esup>)" using gen_boolean_algebra_subset by blast then show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - A))" using A cartesian_product_car_complement_right[of A Q\<^sub>p m n] unfolding is_semialgebraic_def semialg_sets_def by (metis (mono_tags, lifting) semialg_complement semialg_sets_def) qed qed lemma basic_semialg_times_semialg_is_semialg: assumes "A \ basic_semialgs n" assumes "is_semialgebraic m B" shows " is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of B "carrier (Q\<^sub>p\<^bsup>m\<^esup>)""basic_semialgs m"]) using assms(2) is_semialgebraic_def semialg_sets_def apply blast using Qp_times_basic_semialg_right_is_semialgebraic' assms(1) apply blast apply (metis assms(1) basic_semialg_is_semialgebraic inf.absorb1 is_semialgebraic_closed mem_Collect_eq product_of_basic_semialgs_is_semialg') apply (metis (no_types, lifting) cartesian_product_binary_union_right is_semialgebraicI is_semialgebraic_closed semialg_sets_def union_is_semialgebraic) proof- show "\Aa. Aa \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m) \ is_semialgebraic (n + m) (cartesian_product A Aa) \ is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - Aa))" proof- fix B assume A: "B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) (basic_semialgs m)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B))" using A assms cartesian_product_complement_right[of B Q\<^sub>p m A n] add.commute[of n m] proof - have f1: "\n B. \ is_semialgebraic n B \ B \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)" by (meson is_semialgebraic_closed) have "is_basic_semialg n A" using \A \ {S. is_basic_semialg n S}\ by blast then have f2: "is_semialgebraic n A" using padic_fields.basic_semialg_is_semialgebraic padic_fields_axioms by blast have "B \ semialg_sets m" using \B \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) {S. is_basic_semialg m S}\ semialg_sets_def by blast then have "is_semialgebraic m B" by (meson padic_fields.is_semialgebraicI padic_fields_axioms) then show ?thesis using f2 f1 by (metis (no_types) Qp_times_basic_semialg_right_is_semialgebraic' \A \ {S. is_basic_semialg n S}\ \\B \ carrier (Q\<^sub>p\<^bsup>m\<^esup>); A \ carrier (Q\<^sub>p\<^bsup>n\<^esup>)\ \ cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>)) - cartesian_product A B = cartesian_product A (carrier (Q\<^sub>p\<^bsup>m\<^esup>) - B)\ \is_semialgebraic (n + m) (cartesian_product A B)\ diff_is_semialgebraic) qed qed qed text\Semialgebraic sets are closed under cartesian products\ lemma cartesian_product_is_semialgebraic: assumes "is_semialgebraic n A" assumes "is_semialgebraic m B" shows "is_semialgebraic (n + m) (cartesian_product A B)" apply(rule gen_boolean_algebra.induct[of A "carrier (Q\<^sub>p\<^bsup>n\<^esup>)""basic_semialgs n"]) using assms is_semialgebraicE semialg_sets_def apply blast using assms car_times_semialg_is_semialg apply blast using assms basic_semialg_times_semialg_is_semialg apply (simp add: Int_absorb2 basic_semialg_is_semialgebraic is_semialgebraic_closed) proof- show "\A C. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product C B) \ is_semialgebraic (n + m) (cartesian_product (A \ C) B)" proof- fix A C assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" "C \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product C B)" show "is_semialgebraic (n + m) (cartesian_product (A \ C) B)" using A cartesian_product_binary_union_left[of A Q\<^sub>p n C B] by (metis (no_types, lifting) gen_boolean_algebra_subset union_is_semialgebraic) qed show "\A. A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n) \ is_semialgebraic (n + m) (cartesian_product A B) \ is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" proof- fix A assume A: "A \ gen_boolean_algebra (carrier (Q\<^sub>p\<^bsup>n\<^esup>)) (basic_semialgs n)" "is_semialgebraic (n + m) (cartesian_product A B)" show "is_semialgebraic (n + m) (cartesian_product (carrier (Q\<^sub>p\<^bsup>n\<^esup>) - A) B)" using assms A cartesian_product_complement_left[of A Q\<^sub>p n B m] unfolding is_semialgebraic_def semialg_sets_def by (metis car_times_semialg_is_semialg diff_is_semialgebraic is_semialgebraicE is_semialgebraicI is_semialgebraic_closed semialg_sets_def) qed qed (**************************************************************************************************) (**************************************************************************************************) subsection\$N^{th}$ Power Residues\ (**************************************************************************************************) (**************************************************************************************************) definition nth_root_poly where "nth_root_poly (n::nat) a = ((X_poly Q\<^sub>p) [^]\<^bsub>Q\<^sub>p_x\<^esub> n) \\<^bsub>Q\<^sub>p_x\<^esub> (to_poly a)" lemma nth_root_poly_closed: assumes "a \ carrier Q\<^sub>p" shows "nth_root_poly n a \ carrier Q\<^sub>p_x" using assms unfolding nth_root_poly_def by (meson UPQ.P.minus_closed UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_poly_closed) lemma nth_root_poly_eval: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" shows "(nth_root_poly n a) \ b = (b[^]n) \\<^bsub>Q\<^sub>p\<^esub> a" using assms unfolding nth_root_poly_def using UPQ.P.nat_pow_closed UPQ.X_closed UPQ.to_fun_X_pow UPQ.to_fun_diff UPQ.to_fun_to_poly UPQ.to_poly_closed by presburger text\Hensel's lemma gives us this criterion for the existence of \n\-th roots\ lemma nth_root_poly_root: assumes "(n::nat) > 1" assumes "a \ \\<^sub>p" assumes "a \ \" assumes "val (\ \\<^bsub>Q\<^sub>p\<^esub> a) > 2* val ([n]\\)" shows "(\ b \ \\<^sub>p. ((b[^]n) = a))" proof- obtain \ where alpha_def: "\ \ carrier Z\<^sub>p \ \ \ = a" using assms(2) by blast have 0: "\ \ carrier Z\<^sub>p" by (simp add: alpha_def) have 1: "\ \ \\<^bsub>Z\<^sub>p\<^esub>" using assms alpha_def inc_of_one by blast obtain N where N_def: "N = [n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" by blast have N_closed: "N \ carrier Z\<^sub>p" using N_def Zp_nat_mult_closed by blast have 2: "\ ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = ([n]\ \)" proof(induction n) case 0 have 00: "[(0::nat)] \\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = \\<^bsub>Z\<^sub>p\<^esub>" using Zp_nat_inc_zero by blast have 01: "[(0::nat)] \\<^bsub>Q\<^sub>p\<^esub> \ = \" using Qp.nat_inc_zero by blast then show ?case using 00 inc_of_nat by blast next case (Suc n) then show ?case using inc_of_nat by blast qed have 3: "val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \) = val (\ \\<^bsub>Q\<^sub>p\<^esub> a)" using alpha_def Zp.one_closed ring_hom_one[of \ Z\<^sub>p Q\<^sub>p] inc_is_hom Zp.ring_hom_minus[of Q\<^sub>p \ "\\<^bsub>Z\<^sub>p\<^esub>" \ ] Qp.ring_axioms unfolding \_def by (metis Q\<^sub>p_def Zp.minus_closed Zp_def padic_fields.val_of_inc padic_fields_axioms) have 4: "([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) \ nonzero Z\<^sub>p" proof- have 40: "int n \ 0" using of_nat_0_le_iff by blast have "nat (int n) = n" using nat_int by blast hence "[n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> = [int n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>" using 40 unfolding add_pow_def int_pow_def nat_pow_def proof - have "(if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) 0 else rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n) = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n" by (meson of_nat_less_0_iff) then show "rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) n = (let f = rec_nat \\<^bsub>add_monoid Z\<^sub>p\<^esub> (\n f. f \\<^bsub>add_monoid Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) in if int n < 0 then inv\<^bsub>add_monoid Z\<^sub>p\<^esub> f (nat (- int n)) else f (nat (int n)))" using \nat (int n) = n\ by presburger qed thus ?thesis using Zp_char_0[of n] Zp.not_nonzero_memE Zp_char_0' assms(1) gr_implies_not_zero by blast qed then have 5: "val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>) = val ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))" using 2 ord_of_inc by (metis N_closed N_def val_of_inc) then have 6: "(val_Zp (\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub> \)) > 2*(val_Zp ([n]\\<^bsub>Z\<^sub>p\<^esub> \\<^bsub>Z\<^sub>p\<^esub>))" using assms 3 by presburger have "\ b \ carrier Z\<^sub>p. (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" using Zp_nth_root_lemma[of \ n] assms "0" "1" "6" by blast then obtain b where b_def: "b \ carrier Z\<^sub>p \ (b[^]\<^bsub>Z\<^sub>p\<^esub>n= \)" by blast then have "\ (b [^]\<^bsub>Z\<^sub>p\<^esub>n) = a" using alpha_def by blast then have "(\ b) [^] n = a" by (metis Qp.nat_inc_zero Q\<^sub>p_def Qp.nat_pow_zero Zp.nat_pow_0 Zp.nat_pow_zero Zp_nat_inc_zero \_def alpha_def assms(3) b_def frac_inc_of_nat inc_of_one inc_pow not_nonzero_Qp) then show ?thesis using b_def by blast qed text\All points sufficiently close to 1 have nth roots\ lemma eint_nat_times_2: "2*(n::nat) = 2*eint n" using times_eint_simps(1) by (metis mult.commute mult_2_right of_nat_add) lemma P_set_of_one: "P_set 1 = nonzero Q\<^sub>p" apply(rule equalityI) apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq apply blast apply(rule subsetI) unfolding P_set_def nonzero_def mem_Collect_eq using Qp.nat_pow_eone by blast lemma nth_power_fact: assumes "(n::nat) \ 1" shows "\ (m::nat) > 0. \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof(cases "n = 1") case True have "\ u \ carrier Q\<^sub>p. ac 1 u = 1 \ val u = 0 \ u \ P_set n" unfolding True P_set_of_one by (metis iless_Suc_eq padic_fields.val_ring_memE padic_fields.zero_in_val_ring padic_fields_axioms val_nonzero zero_eint_def) then show ?thesis by blast next case F: False obtain m where m_def: "m = 1 + nat ( 2*(ord ([n]\\<^bsub>Q\<^sub>p\<^esub> (\))))" by blast then have m_pos: "m > 0" by linarith have "\ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n" proof fix u assume A: "u \ carrier Q\<^sub>p" show " ac m u = 1 \ val u = 0 \ u \ P_set n" proof assume B: "ac m u = 1 \ val u = 0" then have 0: "val u = val \" by (smt A ac_def not_nonzero_Qp val_one val_ord zero_eint_def) have 1: "ac m u = ac m \" by (metis B Qp.one_nonzero ac_p ac_p_int_pow_factor angular_component_factors_x angular_component_p inc_of_one m_pos p_nonzero) have 2: "u \ nonzero Q\<^sub>p" proof- have "ac m \ = 0" by (meson ac_def) then have "u \ \" by (metis B zero_neq_one) then show ?thesis using A not_nonzero_Qp Qp.nonzero_memI by presburger qed then have 3: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using m_pos 0 1 ac_ord_prop[of "\" u "0::int" m] by (metis B Qp.one_nonzero add.right_neutral eint.inject val_ord zero_eint_def) show "u \ P_set n" proof(cases "u = \") case True then show ?thesis by (metis P_set_one insert_iff zeroth_P_set) next case False have F0: "u \ \\<^sub>p" apply(rule val_ring_memI, rule A) unfolding 0 val_one by auto have F1: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) \ m" using False 3 by blast have "ord (\ \ u) \ m" by (metis A F1 False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) val_ord) hence F2: "ord (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using m_def F1 A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(1) int_nat_eq of_nat_1 of_nat_add val_ord[of "\ \ u"] eint_nat_times_2 by linarith have "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(val ([n]\ \))" proof- have 0: "val (\ \\<^bsub>Q\<^sub>p\<^esub> u) > 2*(ord ([n]\ \))" using F2 val_ord[of "\ \ u"] A False Qp.not_eq_diff_nonzero Qp.one_closed eint_ord_simps(2) by presburger have "n > 0" using assms by linarith hence "eint (ord ([n] \ \)) = val ([n] \ \)" using val_ord_nat_inc[of n] by blast hence "2*ord ([n]\ \) = 2*val ([n]\ \)" by (metis inc_of_nat times_eint_simps(1)) thus ?thesis using 0 val_ord[of "\ \ u"] assms by presburger qed then have "(\ b \ \\<^sub>p. ((b[^]n) = u))" using m_def False nth_root_poly_root[of n u] F0 assms F by linarith then have "(\ b \ carrier Q\<^sub>p. ((b[^]n) = u))" using val_ring_memE by blast then show "u \ P_set n" using P_set_def[of n] 2 by blast qed qed qed then show ?thesis using m_pos by blast qed definition pow_res where "pow_res (n::nat) x = {a. a \ carrier Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. (a = x \ (y[^]n)))}" lemma nonzero_pow_res: assumes "x \ nonzero Q\<^sub>p" shows "pow_res (n::nat) x \ nonzero Q\<^sub>p" proof fix a assume "a \ pow_res n x" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (a = x \ (y[^]n))" using pow_res_def by blast then show "a \ nonzero Q\<^sub>p" using assms Qp.Units_m_closed Qp_nat_pow_nonzero Units_eq_nonzero by blast qed lemma pow_res_of_zero: shows "pow_res n \ = {\}" unfolding pow_res_def apply(rule equalityI) apply(rule subsetI) unfolding mem_Collect_eq apply (metis Qp.integral_iff Qp.nat_pow_closed Qp.nonzero_closed Qp.zero_closed insertCI) apply(rule subsetI) unfolding mem_Collect_eq by (metis Qp.nat_pow_one Qp.one_nonzero Qp.r_one Qp.zero_closed equals0D insertE) lemma equal_pow_resI: assumes "x \ carrier Q\<^sub>p" assumes "y \ pow_res n x" shows "pow_res n x = pow_res n y" proof have y_closed: "y \ carrier Q\<^sub>p" using assms unfolding pow_res_def by blast obtain c where c_def: "c \ nonzero Q\<^sub>p \ y = x \ (c[^]n)" using assms pow_res_def by blast have "((inv c)[^]n) = inv (c[^]n)" using c_def Qp.field_axioms Qp.nat_pow_of_inv Units_eq_nonzero by blast then have "y \ ((inv c)[^]n) = x" using y_closed c_def assms Qp.inv_cancelL(2) Qp.nonzero_closed Qp_nat_pow_nonzero Units_eq_nonzero by presburger then have P0: "(inv c) \ nonzero Q\<^sub>p \ x =y \ ((inv c)[^]n) " using c_def nonzero_inverse_Qp by blast show "pow_res n x \ pow_res n y" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = x \ (b[^]n)" using A pow_res_def by blast then have 0: "b \ nonzero Q\<^sub>p \ a = y \ ((inv c)[^]n) \ (b[^]n)" using \y \ inv c [^] n = x\ by blast have "b \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" proof- have "(inv c)[^]n \ (b[^]n) = ((inv c) \ b)[^]n" using c_def b_def assms P0 Qp.nat_pow_distrib Qp.nonzero_closed by presburger then have " y \ (((inv c)[^]n) \ (b[^]n)) = y \ (((inv c) \ b)[^]n)" by presburger then show ?thesis using y_closed 0 P0 Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) by presburger qed then have "((inv c) \ b) \ nonzero Q\<^sub>p \ a = y \ (((inv c) \ b)[^]n)" by (metis P0 Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed not_nonzero_Qp) then show "a \ pow_res n y" using pow_res_def \a \ carrier Q\<^sub>p\ by blast qed show "pow_res n y \ pow_res n x" proof fix a assume A: "a \ pow_res n y" then have 0: "a \ carrier Q\<^sub>p" by (metis (no_types, lifting) mem_Collect_eq pow_res_def) obtain b where b_def: "b \ nonzero Q\<^sub>p \ a = y \ (b[^]n)" using A pow_res_def by blast then have "a = (x \ (c[^]n)) \ (b[^]n)" using c_def by blast then have "a = x \ ((c[^]n) \ (b[^]n))" by (meson Qp.m_assoc Qp.nat_pow_closed Qp.nonzero_closed assms(1) b_def c_def) then have "a = x \ ((c \ b)[^]n)" using Qp.nat_pow_distrib Qp.nonzero_closed b_def c_def by presburger then have "(c \ b) \ nonzero Q\<^sub>p \ a = x \ ((c \ b)[^]n)" by (metis Qp.integral_iff Qp.nonzero_closed Qp.nonzero_mult_closed b_def c_def not_nonzero_Qp) then show "a \ pow_res n x" using pow_res_def 0 by blast qed qed lemma zeroth_pow_res: assumes "x \ carrier Q\<^sub>p" shows "pow_res 0 x = {x}" apply(rule equalityI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms apply (metis Qp.nat_pow_0 Qp.r_one singletonI) apply(rule subsetI) unfolding pow_res_def mem_Collect_eq using assms by (metis Qp.nat_pow_0 Qp.one_nonzero Qp.r_one equals0D insertE) lemma Zp_car_zero_res: assumes "x \ carrier Z\<^sub>p" shows "x 0 = 0" using assms unfolding Zp_def using Zp_def Zp_defs(3) padic_set_zero_res prime by blast lemma zeroth_ac: assumes "x \ carrier Q\<^sub>p" shows "ac 0 x = 0" apply(cases "x = \ ") unfolding ac_def apply presburger using assms angular_component_closed[of x] Zp_car_zero_res unfolding nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_imp_nonzero: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "x \ nonzero Q\<^sub>p" using assms unfolding ac_def nonzero_def mem_Collect_eq by presburger lemma nonzero_ac_val_ord: assumes "x \ carrier Q\<^sub>p" assumes "ac m x \ 0" shows "val x = ord x" using nonzero_ac_imp_nonzero assms val_ord by blast lemma pow_res_equal_ord: assumes "n > 0" shows "\m > 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof- obtain m where m_def_0: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ val u = 0 \ u \ P_set n)" using assms nth_power_fact[of n] by (metis less_imp_le_nat less_one linorder_neqE_nat nat_le_linear zero_less_iff_neq_zero) then have m_def: "m > 0 \ ( \ u \ carrier Q\<^sub>p. ac m u = 1 \ ord u = 0 \ u \ P_set n)" by (smt nonzero_ac_val_ord zero_eint_def) have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y " then have 0: "ac m (x \
y) = 1" using ac_inv[of y m] ac_mult by (metis ac_inv'''(1) ac_mult' m_def nonzero_inverse_Qp) have 1: "ord (x \
y) = 0" using A ord_fract by presburger have 2: "(x \
y) \ nonzero Q\<^sub>p" using A by (metis Qp.nonzero_closed Qp.nonzero_mult_closed local.fract_cancel_right nonzero_inverse_Qp not_nonzero_Qp zero_fract) have 3: "(x \
y) \ P_set n" using m_def 0 1 2 nonzero_def by (smt Qp.nonzero_closed) then obtain b where b_def: "b \ carrier Q\<^sub>p \ (b[^]n) = (x \
y)" using P_set_def by blast then have "(x \
y) \ y = (b[^]n) \ y" by presburger then have "x = (b[^]n) \ y" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_left) then have "x = y \(b[^]n)" using A b_def by (metis Qp.nonzero_closed local.fract_cancel_right) then have "x \ pow_res n y" unfolding pow_res_def using A b_def by (metis (mono_tags, lifting) "2" Qp.nat_pow_0 Qp.nonzero_closed Qp_nonzero_nat_pow mem_Collect_eq not_gr_zero) then show "pow_res n x = pow_res n y" using A equal_pow_resI[of x y n] unfolding nonzero_def by (metis (mono_tags, lifting) A Qp.nonzero_closed equal_pow_resI) qed qed qed then show ?thesis using m_def by blast qed lemma pow_res_equal: assumes "n > 0" shows "\m> 0. \x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = (ord y mod n) \ pow_res n x = pow_res n y" proof- obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms pow_res_equal_ord by meson have "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix x show "\y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof fix y show "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n \ pow_res n x = pow_res n y" proof assume A: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y mod int n" show "pow_res n x = pow_res n y" proof- have A0: "x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p" using A by blast have A1: "ac m x = ac m y" using A by blast have A2: "ord x = ord y mod int n" using A by blast obtain k where k_def: "k = ord x" by blast obtain l where l_def: "ord y = ord x + (l:: int)*(int n)" using assms A2 by (metis A k_def mod_eqE mod_mod_trivial mult_of_nat_commute) have m_def': "\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y" using m_def by blast have 0: "ord (y \ (\[^](- l*n))) = ord x" proof- have 0: "ord (y \ (\[^](- l*n))) = ord y + (ord (\[^](- l*n)))" using ord_mult p_nonzero A0 Qp_int_pow_nonzero by blast have 1: "ord (\[^](- l*n)) = - l*n" using ord_p_pow_int[of "-l*n"] by blast then have "ord (y \ (\[^](- l*n))) = ord y - l*n" using 0 by linarith then show ?thesis using k_def l_def by linarith qed have 1: "ac m (y \ (\[^](- l*n))) = ac m y" using assms ac_p_int_pow_factor_right[of ] m_def A Qp.nonzero_closed by presburger have 2: "y \ (\[^](- l*n)) \ nonzero Q\<^sub>p" using A0 Qp_int_pow_nonzero[of \ "- l*n"] Qp.cring_axioms nonzero_def cring.cring_simprules(5) fract_cancel_left not_nonzero_Qp p_intpow_inv'' p_nonzero zero_fract Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_memI Qp.nonzero_mult_closed minus_mult_commute mult_minus_right p_intpow_closed(1) p_intpow_closed(2) by presburger then have 3: "pow_res n (y \ (\[^](- l*n))) = pow_res n x" using 2 A0 m_def'[of "y \ (\[^](- l*n))" x] "0" "1" A1 by linarith have 4: "(y \ (\[^](- l*n))) = (y \ ((\[^]- l)[^]n))" using Qp_int_nat_pow_pow[of \ "-l" n] p_nonzero by presburger have "y \ (\[^](- l*n))\ pow_res n y " using "2" "4" Qp_int_pow_nonzero nonzero_def p_nonzero unfolding pow_res_def nonzero_def proof - assume a1: "\x n. x \ {a \ carrier Q\<^sub>p. a \ \} \ x [^] (n::int) \ {a \ carrier Q\<^sub>p. a \ \}" assume a2: "\ \ {a \ carrier Q\<^sub>p. a \ \}" assume a3: "y \ \ [^] (- l * int n) \ {a \ carrier Q\<^sub>p. a \ \}" have f4: "\ [^] (- 1 * l) \ {r \ carrier Q\<^sub>p. r \ \}" using a2 a1 by presburger have f5: "- l = - 1 * l" by linarith then have f6: "y \ \ [^] (- 1 * l * int n) = y \ (\ [^] (- 1 * l)) [^] n" using \y \ \ [^] (- l * int n) = y \ (\ [^] - l) [^] n\ by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. r \ \}" using f5 a3 by presburger then have "y \ (\ [^] (- 1 * l)) [^] n \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f4 by blast then have "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra. ra \ {r \ carrier Q\<^sub>p. r \ \} \ r = y \ ra [^] n}" using f6 f5 by presburger then show "y \ \ [^] (- l * int n) \ {r \ carrier Q\<^sub>p. \ra\{r \ carrier Q\<^sub>p. r \ \}. r = y \ ra [^] n}" by meson qed then have "pow_res n (y \ (\[^](- l*n))) = pow_res n y" using equal_pow_resI[of "(y \ (\[^](- l*n)))" y n] "2" A0 assms Qp.nonzero_mult_closed p_intpow_closed(2) by (metis (mono_tags, opaque_lifting) "3" A Qp.nonzero_closed equal_pow_resI) then show ?thesis using 3 by blast qed qed qed qed then show ?thesis using m_def by blast qed definition pow_res_classes where "pow_res_classes n = {S. \x \ nonzero Q\<^sub>p. S = pow_res n x }" lemma pow_res_semialg_def: assumes "x \ nonzero Q\<^sub>p" assumes "n \ 1" shows "\P \ carrier Q\<^sub>p_x. pow_res n x = (univ_basic_semialg_set n P) - {\}" proof- have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (inv x) \ a = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" proof fix a assume A: "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))" unfolding pow_res_def by blast then obtain y where y_def: "y \ nonzero Q\<^sub>p \a = x \ (y[^]n)" by blast then have "y \ nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" proof - show ?thesis by (metis (no_types, opaque_lifting) Qp.m_assoc Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ y [^] n)\ assms(1) local.fract_cancel_right nonzero_inverse_Qp y_def) qed then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" using assms \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. a = x \ (y[^]n))\ by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)} \ pow_res n x" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. inv x \ a = (y[^]n)}" show "a \ pow_res n x" proof- have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using A by blast then obtain y where y_def: "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ a = x \(y[^]n)" by (metis Qp.l_one Qp.m_assoc Qp.nonzero_closed Qp.not_nonzero_memI \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = y [^] n)\ assms(1) field_inv(2) inv_in_frac(1)) then show ?thesis by (metis (mono_tags, lifting) \a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))\ mem_Collect_eq pow_res_def) qed qed qed obtain P where P_def: "P = up_ring.monom Q\<^sub>p_x (inv x) 1" by blast have P_closed: "P \ carrier Q\<^sub>p_x" using P_def Qp.nonzero_closed Qp.nonzero_memE(2) UPQ.is_UP_monomE(1) UPQ.is_UP_monomI assms(1) inv_in_frac(1) by presburger have P_eval: "\a. a \ carrier Q\<^sub>p \ (P \ a) = (inv x) \ a" using P_def to_fun_monom[of ] by (metis Qp.nat_pow_eone Qp.nonzero_closed Qp.not_nonzero_memI assms(1) inv_in_frac(1)) have 0: "pow_res n x = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "pow_res n x \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume "a \ pow_res n x" then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using 0 by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using P_eval by (metis (mono_tags, lifting) mem_Collect_eq) qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ pow_res n x" proof fix a assume "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y\nonzero Q\<^sub>p \ inv x \ a = (y[^]n)" using P_eval \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then have "a \ carrier Q\<^sub>p \ (\y\nonzero Q\<^sub>p. inv x \ a = (y[^]n))" using \a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}\ by blast then show "a \ pow_res n x" using 0 by blast qed qed have 1: "univ_basic_semialg_set n P - {\} = {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. (P \ a) = (y[^]n)}" proof show "univ_basic_semialg_set n P - {\} \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" proof fix a assume A: "a \ univ_basic_semialg_set n P - {\}" then have A0: "a \ carrier Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" unfolding univ_basic_semialg_set_def by blast then have A0': "a \ nonzero Q\<^sub>p \ (\y\carrier Q\<^sub>p. P \ a = (y[^]n))" using A by (metis DiffD2 not_nonzero_Qp singletonI) then obtain y where y_def: "y\carrier Q\<^sub>p \ P \ a = (y[^]n)" by blast have A1: "(P \ a) \ \" using P_eval A0' Qp.integral_iff Qp.nonzero_closed Qp.nonzero_memE(2) assms(1) inv_in_frac(1) inv_in_frac(2) by presburger have A2: "y \ nonzero Q\<^sub>p" proof- have A20: "(y[^]n) \\" using A1 y_def by blast have "y \ \" apply(rule ccontr) using A20 assms by (metis Qp.nat_pow_eone Qp.semiring_axioms Qp.zero_closed le_zero_eq semiring.nat_pow_zero) then show ?thesis using y_def A1 not_nonzero_Qp Qp.not_nonzero_memE by blast qed then have "y \ nonzero Q\<^sub>p \ P \ a = (y[^]n)" using y_def by blast then show "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" using A0 nonzero_def by blast qed show "{a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)} \ univ_basic_semialg_set n P - {\}" proof fix a assume A: "a \ {a \ carrier Q\<^sub>p. \y\nonzero Q\<^sub>p. P \ a = (y[^]n)}" then obtain y where y_def: "y\nonzero Q\<^sub>p \ P \ a = (y[^]n)" by blast then have "y \\ \ y\ carrier Q\<^sub>p \ P \ a = (y[^]n)" by (metis (mono_tags, opaque_lifting) Qp.nonzero_closed Qp.not_nonzero_memI) then have "a \\" using P_eval by (metis Qp.m_comm Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.zero_closed assms(1) inv_in_frac(1) zero_fract) then show "a \ univ_basic_semialg_set n P - {\}" unfolding univ_basic_semialg_set_def using A \y \ \ \ y \ carrier Q\<^sub>p \ P \ a = (y[^]n)\ by blast qed qed show ?thesis using 0 1 by (metis (no_types, lifting) P_closed) qed lemma pow_res_is_univ_semialgebraic: assumes "x \ carrier Q\<^sub>p" shows "is_univ_semialgebraic (pow_res n x)" proof(cases "n = 0") case True have T0: "pow_res n x = {x}" unfolding True using assms by (simp add: assms zeroth_pow_res) have "[x] \ carrier (Q\<^sub>p\<^bsup>1\<^esup>)" using assms Qp.to_R1_closed by blast hence "is_semialgebraic 1 {[x]}" using is_algebraic_imp_is_semialg singleton_is_algebraic by blast thus ?thesis unfolding T0 using assms by (simp add: \x \ carrier Q\<^sub>p\ finite_is_univ_semialgebraic) next case False show ?thesis proof(cases "x = \") case True then show ?thesis using finite_is_univ_semialgebraic False pow_res_of_zero by (metis Qp.zero_closed empty_subsetI finite.emptyI finite.insertI insert_subset) next case F: False then show ?thesis using False pow_res_semialg_def[of x n] diff_is_univ_semialgebraic[of _ "{\}"] finite_is_univ_semialgebraic[of "{\}"] by (metis Qp.zero_closed assms empty_subsetI finite.emptyI finite.insertI insert_subset less_one less_or_eq_imp_le linorder_neqE_nat not_nonzero_Qp univ_basic_semialg_set_is_univ_semialgebraic) qed qed lemma pow_res_is_semialg: assumes "x \ carrier Q\<^sub>p" shows "is_semialgebraic 1 (to_R1 ` (pow_res n x))" using assms pow_res_is_univ_semialgebraic is_univ_semialgebraicE by blast lemma pow_res_refl: assumes "x \ carrier Q\<^sub>p" shows "x \ pow_res n x" proof- have "x = x \ (\ [^]n)" using assms Qp.nat_pow_one Qp.r_one by presburger thus ?thesis using assms unfolding pow_res_def mem_Collect_eq using Qp.one_nonzero by blast qed lemma equal_pow_resE: assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "n > 0" assumes "pow_res n a = pow_res n b" shows "\ s \ P_set n. a = b \ s" proof- have "a \ pow_res n b" using assms pow_res_refl by blast then obtain y where y_def: " y \ nonzero Q\<^sub>p \ a = b \ y[^]n" unfolding pow_res_def by blast thus ?thesis unfolding P_set_def using Qp.nonzero_closed Qp_nat_pow_nonzero by blast qed lemma pow_res_one: assumes "x \ nonzero Q\<^sub>p" shows "pow_res 1 x = nonzero Q\<^sub>p" proof show "pow_res 1 x \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of x 1] by blast show "nonzero Q\<^sub>p \ pow_res 1 x" proof fix y assume A: "y \ nonzero Q\<^sub>p" then have 0: "\ \ nonzero Q\<^sub>p \ y = x \ ((inv x)\ y)[^](1::nat)" using assms Qp.m_comm Qp.nat_pow_eone Qp.nonzero_closed Qp.nonzero_mult_closed Qp.one_nonzero local.fract_cancel_right nonzero_inverse_Qp by presburger have 1: "(inv x)\ y \ nonzero Q\<^sub>p" using A assms by (metis Qp.Units_m_closed Units_eq_nonzero nonzero_inverse_Qp) then show "y \ pow_res 1 x" unfolding pow_res_def using 0 1 A Qp.nonzero_closed by blast qed qed lemma pow_res_zero: assumes "n > 0" shows "pow_res n \ = {\}" proof show "pow_res n \ \ {\}" unfolding pow_res_def using Qp.l_null Qp.nat_pow_closed Qp.nonzero_closed by blast show "{\} \ pow_res n \" using assms unfolding pow_res_def using Qp.l_null Qp.one_closed Qp.one_nonzero empty_subsetI insert_subset by blast qed lemma equal_pow_resI': assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ P_set n" assumes "a = b \ c" assumes "n > 0" shows "pow_res n a = pow_res n b" proof- obtain y where y_def: "c = y[^]n \ y \ carrier Q\<^sub>p" using assms unfolding P_set_def by blast have c_nonzero: "c \ nonzero Q\<^sub>p" using P_set_nonzero'(1) assms(3) by blast have y_nonzero: "y \ nonzero Q\<^sub>p" using y_def c_nonzero Qp_nonzero_nat_pow assms(5) by blast have 0: "a \ pow_res n b" using assms y_nonzero y_def unfolding pow_res_def by blast show ?thesis apply(cases "b = \") using pow_res_zero Qp.l_null Qp.nonzero_closed assms(4) c_nonzero apply presburger by (metis "0" assms(1) assms(2) assms(5) not_nonzero_Qp equal_pow_resI) qed lemma equal_pow_resI'': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "a \ inv b \ P_set n" shows "pow_res n a = pow_res n b" using assms equal_pow_resI'[of a b "a \ inv b" n] Qp.nonzero_closed local.fract_cancel_right by blast lemma equal_pow_resI''': assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ nonzero Q\<^sub>p" assumes "c \ nonzero Q\<^sub>p" assumes "pow_res n (c \ a) = pow_res n (c \ b)" shows "pow_res n a = pow_res n b" proof- have 0: "c \a \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(2) assms(4)) have 1: "c \b \ nonzero Q\<^sub>p" by (meson Localization.submonoid.m_closed Qp.nonzero_is_submonoid assms(3) assms(4)) have "c\a \ pow_res n (c\b)" proof(cases "n = 1") case True then show ?thesis using assms 0 1 pow_res_one[of "c\b"] by blast next case False then have "n \ 2" using assms(1) by linarith then show ?thesis using assms 0 1 pow_res_refl[of "c\a" n] unfolding nonzero_def by blast qed then obtain y where y_def: "y \ nonzero Q\<^sub>p \ (c \ a) = (c \ b)\y[^]n" using assms unfolding pow_res_def by blast then have "a = b\y[^]n" using assms by (metis Qp.m_assoc Qp.m_lcancel Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero) then show ?thesis by (metis P_set_memI Qp.nonzero_closed Qp.nonzero_mult_closed Qp.not_nonzero_memI Qp_nat_pow_nonzero assms(1) assms(3) equal_pow_resI' y_def) qed lemma equal_pow_resI'''': assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "a = b \ u" assumes "u \ P_set n" shows "pow_res n a = pow_res n b" proof(cases "a = \") case True then have "b = \" using assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral Qp.nonzero_closed Qp.not_nonzero_memI mem_Collect_eq) then show ?thesis using pow_res_zero using True by blast next case False then have 0: "a \ nonzero Q\<^sub>p" using Qp.not_nonzero_memE assms(2) by blast have 1: "b \ nonzero Q\<^sub>p" using 0 assms unfolding P_set_def by (metis (no_types, lifting) Qp.integral_iff Qp.nonzero_closed mem_Collect_eq not_nonzero_Qp) have 2: "a \ (inv b)\ P_set n" using assms 0 1 by (metis P_set_nonzero'(2) Qp.inv_cancelR(1) Qp.m_comm Qp.nonzero_memE(2) Units_eq_nonzero inv_in_frac(1)) then show ?thesis using equal_pow_resI'' by (meson "0" "1" assms(1) equal_pow_resI) qed lemma Zp_Units_ord_zero: assumes "a \ Units Z\<^sub>p" shows "ord_Zp a = 0" proof- have "inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p" apply(rule Zp.nonzero_memI, rule Zp.Units_inv_closed, rule assms) using assms Zp.Units_inverse in_Units_imp_not_zero by blast then have "ord_Zp (a \\<^bsub>Z\<^sub>p\<^esub> inv \<^bsub>Z\<^sub>p\<^esub> a) = ord_Zp a + ord_Zp (inv \<^bsub>Z\<^sub>p\<^esub> a)" using assms ord_Zp_mult Zp.Units_nonzero zero_not_one by (metis Zp.zero_not_one) then show ?thesis by (smt Zp.Units_closed Zp.Units_r_inv Zp.integral_iff Zp.nonzero_closed \inv\<^bsub>Z\<^sub>p\<^esub> a \ nonzero Z\<^sub>p\ assms ord_Zp_one ord_pos) qed lemma pow_res_nth_pow: assumes "a \ nonzero Q\<^sub>p" assumes "n > 0" shows "pow_res n (a[^]n) = pow_res n \" proof show "pow_res n (a [^] n) \ pow_res n \" proof fix x assume A: "x \ pow_res n (a [^] n)" then show "x \ pow_res n \" by (metis P_set_memI Qp.l_one Qp.nat_pow_closed Qp.nonzero_closed Qp.nonzero_memE(2) Qp.nonzero_pow_nonzero Qp.one_closed assms(1) assms(2) equal_pow_resI') qed show "pow_res n \ \ pow_res n (a [^] n)" proof fix x assume A: "x \ pow_res n \" then obtain y where y_def: "y \ nonzero Q\<^sub>p \ x = \\y[^]n" unfolding pow_res_def by blast then have 0: "x = y[^]n" using Qp.l_one Qp.nonzero_closed by blast have "y[^]n = a[^]n \ (inv a \ y)[^]n" proof- have "a[^]n \ (inv a \ y)[^]n = a[^]n \ (inv a)[^]n \ y[^]n" using Qp.Units_inv_closed Qp.m_assoc Qp.nat_pow_closed Qp.nat_pow_distrib Qp.nonzero_closed Units_eq_nonzero assms(1) y_def by presburger then show ?thesis by (metis Qp.Units_inv_inv Qp.inv_cancelR(1) Qp.nat_pow_distrib Qp.nonzero_closed Qp.nonzero_mult_closed Units_eq_nonzero assms(1) nonzero_inverse_Qp y_def) qed then show "x \ pow_res n (a [^] n)" using y_def A assms unfolding pow_res_def mem_Collect_eq by (metis "0" Qp.integral Qp.m_closed Qp.nonzero_closed Qp.not_nonzero_memI inv_in_frac(1) inv_in_frac(2) not_nonzero_Qp) qed qed lemma pow_res_of_p_pow: assumes "n > 0" shows "pow_res n (\[^]((l::int)*n)) = pow_res n \" proof- have 0: "\[^]((l::int)*n) = (\[^]l)[^]n" using Qp_p_int_nat_pow_pow by blast have "\[^]((l::int)*n) \ P_set n" using P_set_memI[of _ "\[^]l"] by (metis "0" Qp.not_nonzero_memI Qp_int_pow_nonzero p_intpow_closed(1) p_nonzero) thus ?thesis using "0" assms p_intpow_closed(2) pow_res_nth_pow by presburger qed lemma pow_res_nonzero: assumes "n > 0" assumes "a \ nonzero Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n b" shows "b \ nonzero Q\<^sub>p" using assms nonzero_pow_res[of a n] pow_res_zero[of n] by (metis insert_subset not_nonzero_Qp) lemma pow_res_mult: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" assumes "b \ carrier Q\<^sub>p" assumes "c \ carrier Q\<^sub>p" assumes "d \ carrier Q\<^sub>p" assumes "pow_res n a = pow_res n c" assumes "pow_res n b = pow_res n d" shows "pow_res n (a \ b) = pow_res n (c \ d)" proof(cases "a \ nonzero Q\<^sub>p") case True then have "c \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where alpha_def: "\ \ nonzero Q\<^sub>p \ a = c \ \[^]n" using assms True pow_res_refl[of a n] unfolding assms unfolding pow_res_def by blast show ?thesis proof(cases "b \ nonzero Q\<^sub>p") case T: True then have "d \ nonzero Q\<^sub>p" using assms pow_res_nonzero by blast then obtain \ where beta_def: "\ \ nonzero Q\<^sub>p \ b = d \ \[^]n" using T pow_res_refl[of b n] unfolding assms unfolding pow_res_def using assms by blast then have "a \ b = (c \ d) \ (\[^]n \ \[^] n)" using Qp.m_assoc Qp.m_lcomm Qp.nonzero_closed Qp.nonzero_mult_closed Qp_nat_pow_nonzero alpha_def assms(3) assms(4) assms(5) by presburger then have 0: "a \ b = (c \ d) \ ((\ \ \)[^] n)" by (metis Qp.nat_pow_distrib Qp.nonzero_closed alpha_def beta_def) show ?thesis apply(intro equal_pow_resI'[of _ _ "(\ \ \)[^] n"] Qp.ring_simprules assms P_set_memI[of _ "\ \ \"] Qp.nat_pow_closed nonzero_memE 0 Qp_nat_pow_nonzero ) using alpha_def beta_def apply auto apply(intro nonzero_memI Qp.nonzero_mult_closed) using alpha_def beta_def nonzero_memE apply auto by (meson Qp.integral_iff) next case False then have "b = \" by (meson assms not_nonzero_Qp) then have "d = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \b = \\ assms by presburger qed next case False then have "a = \" by (meson assms not_nonzero_Qp) then have "c = \" using assms by (metis False not_nonzero_Qp pow_res_nonzero) then show ?thesis using Qp.r_null \a = \\ assms Qp.l_null by presburger qed lemma pow_res_p_pow_factor: assumes "n > 0" assumes "a \ carrier Q\<^sub>p" shows "pow_res n a = pow_res n (\[^]((l::int)*n) \ a)" proof(cases "a = \") case True then show ?thesis using Qp.r_null p_intpow_closed(1) by presburger next case False then show ?thesis using assms pow_res_of_p_pow using Qp.m_comm Qp.one_closed Qp.r_one p_intpow_closed(1) pow_res_mult by presburger qed lemma pow_res_classes_finite: assumes "n \ 1" shows "finite (pow_res_classes n)" proof(cases "n = 1") case True have "pow_res_classes n = {(nonzero Q\<^sub>p)}" using True pow_res_one unfolding pow_res_classes_def using Qp.one_nonzero by blast then show ?thesis by auto next case False then have n_bound: "n \ 2" using assms by linarith obtain m where m_def: "m > 0 \ (\x y. x \ nonzero Q\<^sub>p \ y \ nonzero Q\<^sub>p \ ac m x = ac m y \ ord x = ord y \ pow_res n x = pow_res n y)" using assms False pow_res_equal_ord n_bound by (metis gr_zeroI le_numeral_extra(2)) obtain f where f_def: "f = (\ \ \. (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)))" by blast have 0: "\x. x \ nonzero Q\<^sub>p \ pow_res n x = f (ac m x) (ord x)" proof- fix x assume A: "x \ nonzero Q\<^sub>p" obtain \ where eta_def: "\ = ac m x" by blast obtain \ where nu_def: "\ = ord x" by blast have "\y \pow_res n x. ac m y = ac m x \ ord y = ord x" using pow_res_refl A assms neq0_conv Qp.nonzero_closed by blast hence "pow_res n x \ (pow_res_classes n) \ (\ y \ (pow_res n x). ac m y = \ \ ord y = \)" unfolding nu_def eta_def using assms unfolding pow_res_classes_def using A by blast then have 0: "(\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" by blast have "f \ \ = (SOME y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \))" using f_def by blast then have 1: "f \ \ \ (pow_res_classes n) \ ((\ y \ (f \ \). ac m y = \ \ ord y = \))" using 0 someI_ex[of "\ y. y \ (pow_res_classes n) \ (\ x \ y. ac m x = \ \ ord x = \)"] unfolding f_def by blast then obtain y where y_def: "y \ (f \ \) \ ac m y = ac m x \ ord y = ord x" unfolding nu_def eta_def by blast obtain a where a_def: "a \ nonzero Q\<^sub>p \ (f \ \) = pow_res n a" using 1 unfolding pow_res_classes_def by blast then have 2: "y \ pow_res n a" using y_def by blast have 3: "y \ nonzero Q\<^sub>p" using y_def nonzero_pow_res[of a n] a_def by blast then have 4: "pow_res n y = pow_res n a" using 3 y_def a_def equal_pow_resI[of y a n] n_bound Qp.nonzero_closed by (metis equal_pow_resI) have 5: "pow_res n y = f \ \" using 4 a_def by blast then show "pow_res n x = f (ac m x) (ord x)" unfolding eta_def nu_def using "3" A m_def y_def by blast qed obtain N where N_def: "N > 0 \ (\ u \ carrier Q\<^sub>p. ac N u = 1 \ val u = 0 \ u \ P_set n)" using n_bound nth_power_fact assms by blast have 1: "\x. x \ nonzero Q\<^sub>p \ (\y \ nonzero Q\<^sub>p. ord y \ 0 \ ord y < n \ pow_res n x = pow_res n y)" proof- fix x assume x_def: "x \ nonzero Q\<^sub>p" then obtain k where k_def: "k = ord x mod n" by blast then obtain l where l_def: "ord x = (int n)*l + k" using cancel_div_mod_rules(2)[of n "ord x"0] unfolding k_def by (metis group_add_class.add.right_cancel) have "x = (\[^](ord x)) \ \ (angular_component x)" using x_def angular_component_factors_x by blast then have "x = (\[^](n*l + k)) \ \ (angular_component x)" unfolding l_def by blast hence "x = \[^](int n*l) \ (\[^] k) \ \ (angular_component x)" by (metis p_intpow_add) hence 0: "x = (\[^]l)[^]n \ (\[^] k) \ \ (angular_component x)" using p_pow_factor[of n l k] \x = \ [^] (int n * l + k) \ \ (angular_component x)\ by presburger have 1: "\ (angular_component x) \ carrier Q\<^sub>p" using x_def angular_component_closed inc_closed by blast hence 2: "x = (\[^]l)[^]n \ ((\[^] k) \ \ (angular_component x))" using 0 by (metis Qp.m_assoc Qp.nat_pow_closed p_intpow_closed(1)) obtain a where a_def: "a = (\[^] k) \ \ (angular_component x)" by blast have 30: "angular_component x \ Units Z\<^sub>p" using angular_component_unit x_def by blast then have 3: "\ (angular_component x) \ Units Q\<^sub>p" by (metis Units_eq_nonzero Zp.Units_closed in_Units_imp_not_zero inc_of_nonzero not_nonzero_Qp) have 4: "\ (angular_component x) \ nonzero Q\<^sub>p" using 3 Units_nonzero_Qp by blast have a_nonzero: "a \ nonzero Q\<^sub>p" unfolding a_def 4 by (meson "3" Qp.UnitsI(1) Qp.Units_m_closed Units_nonzero_Qp p_intpow_closed(1) p_intpow_inv) have 5: "x = a \(\[^]l)[^]n" using 2 a_nonzero unfolding a_def using Qp.m_comm Qp.nat_pow_closed Qp.nonzero_closed p_intpow_closed(1) by presburger hence "x \ pow_res n a" unfolding pow_res_def using Qp.nonzero_closed Qp_int_pow_nonzero p_nonzero x_def by blast hence 6:"pow_res n a = pow_res n x" using x_def a_def equal_pow_resI[of x a n] a_nonzero n_bound Qp.nonzero_closed equal_pow_resI by blast have 7: "ord (\ (angular_component x)) = 0" proof- have "ord_Zp (angular_component x) = 0" using 30 Zp_Units_ord_zero by blast then have "val_Zp (angular_component x) = 0" using "30" unit_imp_val_Zp0 by blast then have "val (\ (angular_component x)) = 0" by (metis angular_component_closed val_of_inc x_def) then show ?thesis using angular_component_closed x_def by (metis "30" Zp.Units_closed \ord_Zp (angular_component x) = 0\ in_Units_imp_not_zero not_nonzero_Qp ord_of_inc) qed have 8: "ord a = k" unfolding a_def using 3 4 7 ord_mult[of "\ [^] k" "\ (angular_component x)"] ord_p_pow_int[of k] p_pow_nonzero using Qp_int_pow_nonzero p_nonzero by presburger have 9: "k < n" unfolding k_def using assms by auto - show " \y\nonzero Q\<^sub>p. 0 \ ord y \ ord y < int n \ pow_res n x = pow_res n y" - by (metis "6" "8" a_nonzero assms k_def less_imp_of_nat_less of_nat_eq_0_iff pos_mod_conj rel_simps(45) zero_less_iff_neq_zero) + 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 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 " 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 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 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 ) 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/Probabilistic_Prime_Tests/Jacobi_Symbol.thy b/thys/Probabilistic_Prime_Tests/Jacobi_Symbol.thy --- a/thys/Probabilistic_Prime_Tests/Jacobi_Symbol.thy +++ b/thys/Probabilistic_Prime_Tests/Jacobi_Symbol.thy @@ -1,545 +1,545 @@ (* File: Jacobi_Symbol.thy Authors: Daniel Stüwe, Manuel Eberl The Jacobi symbol, a generalisation of the Legendre symbol. This is used in the Solovay--Strassen test. *) section \The Jacobi Symbol\ theory Jacobi_Symbol imports Legendre_Symbol Algebraic_Auxiliaries begin text \ The Jacobi symbol is a generalisation of the Legendre symbol to non-primes \cite{Legendre_Symbol, Jacobi_Symbol}. It is defined as \[\left(\frac{a}{n}\right) = \left(\frac{a}{p_1}\right)^{k_1} \ldots \left(\frac{a}{p_l}\right)^{k_l}\] where $(\frac{a}{p})$ denotes the Legendre symbol, \a\ is an integer, \n\ is an odd natural number and $p_1^{k_1}\ldots p_l^{k_l}$ is its prime factorisation. There is, however, a fairly natural generalisation to all non-zero integers for \n\. It is less clear what a good choice for \n = 0\ is; Mathematica and Maxima adopt the convention that $(\frac{\pm 1}{0}) = 1$ and $(\frac{a}{0}) = 0$ otherwise. However, we chose the slightly different convention $(\frac{a}{0}) = 0$ for \<^emph>\all\ \a\ because then the Jacobi symbol is completely multiplicative in both arguments without any restrictions. \ definition Jacobi :: "int \ int \ int" where "Jacobi a n = (if n = 0 then 0 else (\p\#prime_factorization n. Legendre a p))" lemma Jacobi_0_right [simp]: "Jacobi a 0 = 0" by (simp add: Jacobi_def) lemma Jacobi_mult_left [simp]: "Jacobi (a * b) n = Jacobi a n * Jacobi b n" proof (cases "n = 0") case False have *: "{# Legendre (a * b) p . p \# prime_factorization n #} = {# Legendre a p * Legendre b p . p \# prime_factorization n #}" by (meson Legendre_mult in_prime_factors_imp_prime image_mset_cong) show ?thesis using False unfolding Jacobi_def * prod_mset.distrib by auto qed auto lemma Jacobi_mult_right [simp]: "Jacobi a (n * m) = Jacobi a n * Jacobi a m" by (cases "m = 0"; cases "n = 0") (auto simp: Jacobi_def prime_factorization_mult) lemma prime_p_Jacobi_eq_Legendre[intro!]: "prime p \ Jacobi a p = Legendre a p" unfolding Jacobi_def prime_factorization_prime by simp lemma Jacobi_mod [simp]: "Jacobi (a mod m) n = Jacobi a n" if "n dvd m" proof - have *: "{# Legendre (a mod m) p . p \# prime_factorization n #} = {# Legendre a p . p \# prime_factorization n #}" using that by (intro image_mset_cong, subst Legendre_mod) (auto intro: dvd_trans[OF in_prime_factors_imp_dvd]) thus ?thesis by (simp add: Jacobi_def) qed lemma Jacobi_mod_cong: "[a = b] (mod n) \ Jacobi a n = Jacobi b n" by (metis Jacobi_mod cong_def dvd_refl) lemma Jacobi_1_eq_1 [simp]: "p \ 0 \ Jacobi 1 p = 1" by (simp add: Jacobi_def in_prime_factors_imp_prime cong: image_mset_cong) lemma Jacobi_eq_0_not_coprime: assumes "n \ 0" "\coprime a n" shows "Jacobi a n = 0" proof - from assms have "\p. p dvd gcd a n \ prime p" by (intro prime_divisor_exists) auto then obtain p where p: "p dvd a" "p dvd n" "prime p" by auto hence "Legendre a p = 0" using assms by (auto simp: prime_int_iff) thus ?thesis using p assms unfolding Jacobi_def by (auto simp: image_iff prime_factors_dvd) qed lemma Jacobi_p_eq_2'[simp]: "n > 0 \ Jacobi a (2^n) = a mod 2" by (auto simp add: Jacobi_def prime_factorization_prime_power) lemma Jacobi_prod_mset[simp]: "n \ 0 \ Jacobi (prod_mset M) n = (\q\#M. Jacobi q n)" by (induction M) simp_all lemma non_trivial_coprime_neq: "1 < a \ 1 < b \ coprime a b \ a \ b" for a b :: int by auto lemma odd_odd_even: fixes a b :: int assumes "odd a" "odd b" shows "even ((a*b-1) div 2) = even ((a-1) div 2 + (b-1) div 2)" using assms by (auto elim!: oddE simp: algebra_simps) lemma prime_nonprime_wlog [case_names primes nonprime sym]: assumes "\p q. prime p \ prime q \ P p q" assumes "\p q. \prime p \ P p q" assumes "\p q. P p q \ P q p" shows "P p q" by (cases "prime p"; cases "prime q") (auto intro: assms) lemma Quadratic_Reciprocity_Jacobi: fixes p q :: int assumes "coprime p q" and "2 < p" "2 < q" and "odd p" "odd q" shows "Jacobi p q * Jacobi q p = (- 1) ^ (nat ((p - 1) div 2 * ((q - 1) div 2)))" using assms proof (induction "nat p" "nat q" arbitrary: p q rule: measure_induct_rule[where f = "\(a, b). a + b", split_format(complete), simplified]) case (1 p q) thus ?case proof (induction p q rule: prime_nonprime_wlog) case (sym p q) thus ?case by (simp only: add_ac coprime_commute mult_ac) blast next case (primes p q) from \prime p\ \prime q\ have "prime (nat p)" "prime (nat q)" "p \ q" using prime_int_nat_transfer primes(4) non_trivial_coprime_neq prime_gt_1_int by blast+ with Quadratic_Reciprocity_int and prime_p_Jacobi_eq_Legendre show ?case using \prime p\ \prime q\ primes(5-) by presburger next case (nonprime p q) from \\prime p\ obtain a b where *: "p = a * b" "1 < b" "1 < a" using \2 < p\ prime_divisor_exists_strong[of p] by auto hence odd_ab: "odd a" "odd b" using \odd p\ by simp_all moreover have "2 < b" and "2 < a" using odd_ab and * by presburger+ moreover have "coprime a q" and "coprime b q" using \coprime p q\ unfolding * by simp_all ultimately have IH: "Jacobi a q * Jacobi q a = (- 1) ^ nat ((a - 1) div 2 * ((q - 1) div 2))" "Jacobi b q * Jacobi q b = (- 1) ^ nat ((b - 1) div 2 * ((q - 1) div 2))" by (auto simp: * nonprime) have pos: "0 < q" "0 < p" "0 < a" "0 < b" using * \2 < q\ by simp_all have "Jacobi p q * Jacobi q p = (Jacobi a q * Jacobi q a) * (Jacobi b q * Jacobi q b)" using * by simp also have "... = (- 1) ^ nat ((a - 1) div 2 * ((q - 1) div 2)) * (- 1) ^ nat ((b - 1) div 2 * ((q - 1) div 2))" using IH by presburger also from odd_odd_even[OF odd_ab] have "... = (- 1) ^ nat ((p - 1) div 2 * ((q - 1) div 2))" unfolding * minus_one_power_iff using \2 < q\ * by (auto simp add: even_nat_iff pos_imp_zdiv_nonneg_iff) finally show ?case . qed qed lemma Jacobi_values: "Jacobi p q \ {1, -1, 0}" proof (cases "q = 0") case False hence "\Legendre p x\ = 1" if "x \# prime_factorization q" "Jacobi p q \ 0" for x using that prod_mset_zero_iff Legendre_values[of p x] unfolding Jacobi_def is_unit_prod_mset_iff set_image_mset by fastforce then have "is_unit (prod_mset (image_mset (Legendre p) (prime_factorization q)))" if "Jacobi p q \ 0" using that False unfolding Jacobi_def is_unit_prod_mset_iff by auto thus ?thesis by (auto simp: Jacobi_def) qed auto lemma Quadratic_Reciprocity_Jacobi': fixes p q :: int assumes "coprime p q" and "2 < p" "2 < q" and "odd p" "odd q" shows "Jacobi q p = (if p mod 4 = 3 \ q mod 4 = 3 then -1 else 1) * Jacobi p q" proof - have aux: "a \ {1, -1, 0} \ c \ 0 \ a*b = c \ b = c * a" for b c a :: int by auto from Quadratic_Reciprocity_Jacobi[OF assms] have "Jacobi q p = (-1) ^ nat ((p - 1) div 2 * ((q - 1) div 2)) * Jacobi p q" using Jacobi_values by (fastforce intro!: aux) also have "(-1 :: int) ^ nat ((p - 1) div 2 * ((q - 1) div 2)) = (if even ((p - 1) div 2) \ even ((q - 1) div 2) then 1 else - 1)" unfolding minus_one_power_iff using \2 < p\ \2 < q\ by (auto simp: even_nat_iff) also have "... = (if p mod 4 = 3 \ q mod 4 = 3 then -1 else 1)" using \odd p\ \odd q\ by presburger finally show ?thesis . qed lemma dvd_odd_square: "8 dvd a\<^sup>2 - 1" if "odd a" for a :: int proof - obtain x where "a = 2*x + 1" using \odd a\ by (auto elim: oddE) thus ?thesis by(cases "odd x") (auto elim: oddE simp: power2_eq_square algebra_simps) qed lemma odd_odd_even': fixes a b :: int assumes "odd a" "odd b" shows "even (((a * b)\<^sup>2 - 1) div 8) \ even (((a\<^sup>2 - 1) div 8) + ((b\<^sup>2 - 1) div 8))" proof - obtain x where [simp]: "a = 2*x + 1" using \odd a\ by (auto elim: oddE) obtain y where [simp]: "b = 2*y + 1" using \odd b\ by (auto elim: oddE) show ?thesis by (cases "even x"; cases "even y"; elim oddE evenE) (auto simp: power2_eq_square algebra_simps) qed lemma odd_odd_even_nat': fixes a b :: nat assumes "odd a" "odd b" shows "even (((a * b)\<^sup>2 - 1) div 8) \ even (((a\<^sup>2 - 1) div 8) + ((b\<^sup>2 - 1) div 8))" proof - obtain x where [simp]: "a = 2*x + 1" using \odd a\ by (auto elim: oddE) obtain y where [simp]: "b = 2*y + 1" using \odd b\ by (auto elim: oddE) show ?thesis by (cases "even x"; cases "even y"; elim oddE evenE) (auto simp: power2_eq_square algebra_simps) qed lemma supplement2_Jacobi: "odd p \ p > 1 \ Jacobi 2 p = (- 1) ^ (((nat p)\<^sup>2 - 1) div 8)" proof (induction p rule: prime_divisors_induct) case (factor p x) then have "odd x" by force have "2 < p" using \odd (p * x)\ prime_gt_1_int[OF \prime p\] by (cases "p = 2") auto have "odd p" using prime_odd_int[OF \prime p\ \2 < p\] . have "0 < x" using \1 < (p * x)\ prime_gt_0_int[OF \prime p\] and less_trans less_numeral_extra(1) zero_less_mult_pos by blast have base_case : "Jacobi 2 p = (- 1) ^ (((nat p)\<^sup>2 - 1) div 8)" using \2 < p\ \prime p\ supplement2_Legendre and prime_p_Jacobi_eq_Legendre by presburger show ?case proof (cases "x = 1") case True thus ?thesis using base_case by force next case False have "Jacobi 2 (p * x) = Jacobi 2 p * Jacobi 2 x" using \2 < p\ \0 < x\ by simp also have "Jacobi 2 x = (- 1) ^ (((nat x)\<^sup>2 - 1) div 8)" using \odd x\ \0 < x\ \x \ 1\ by (intro factor.IH) auto also note base_case also have "(-1) ^ (((nat p)\<^sup>2 - 1) div 8) * (-1) ^ (((nat x)\<^sup>2 - 1) div 8) = (-1 :: int) ^ (((nat (p * x))\<^sup>2 - 1) div 8)" unfolding minus_one_power_iff using \2 < p\ \0 < x\ \odd x\ \odd p\ and odd_odd_even_nat' using [[linarith_split_limit = 0]] by (force simp add: nat_mult_distrib even_nat_iff) finally show ?thesis . qed qed simp_all lemma mod_nat_wlog [consumes 1, case_names modulo]: fixes P :: "nat \ bool" assumes "b > 0" assumes "\k. k \ {0.. n mod b = k \ P n" shows "P n" using assms and mod_less_divisor by fastforce lemma mod_int_wlog [consumes 1, case_names modulo]: fixes P :: "int \ bool" assumes "b > 0" assumes "\k. 0 \ k \ k < b \ n mod b = k \ P n" shows "P n" - using assms and pos_mod_conj by blast + using \b > 0\ assms(2) [of \n mod b\] by simp lemma supplement2_Jacobi': assumes "odd p" and "p > 1" shows "Jacobi 2 p = (if p mod 8 = 1 \ p mod 8 = 7 then 1 else -1)" proof - have "0 < (4 :: nat)" by simp then have *: "even ((p\<^sup>2 - 1) div 8) = (p mod 8 = 1 \ p mod 8 = 7)" if "odd p" for p :: nat proof(induction p rule: mod_nat_wlog) case (modulo k) then consider "p mod 4 = 1" | "p mod 4 = 3" using \odd p\ by (metis dvd_0_right even_even_mod_4_iff even_numeral mod_exhaust_less_4) then show ?case proof (cases) case 1 then obtain l where l: "p = 4 * l + 1" using mod_natE by blast have "even l = ((4 * l + 1) mod 8 = 1 \ (4 * l + 1) mod 8 = 7)" by presburger thus ?thesis by (simp add: l power2_eq_square algebra_simps) next case 2 then obtain l where l: "p = 4 * l + 3" using mod_natE by blast have "odd l = ((3 + l * 4) mod 8 = Suc 0 \ (3 + l * 4) mod 8 = 7)" by presburger thus ?thesis by (simp add: l power2_eq_square algebra_simps) qed qed have [simp]: "nat p mod 8 = nat (p mod 8)" using \p > 1\ using nat_mod_distrib[of p 8] by simp from assms have "odd (nat p)" by (simp add: even_nat_iff) show ?thesis unfolding supplement2_Jacobi[OF assms] minus_one_power_iff *[OF \odd (nat p)\] by (simp add: nat_eq_iff) qed theorem supplement1_Jacobi: "odd p \ 1 < p \ Jacobi (-1) p = (-1) ^ (nat ((p - 1) div 2))" proof (induction p rule: prime_divisors_induct) case (factor p x) then have "odd x" by force have "2 < p" using \odd (p * x)\ prime_gt_1_int[OF \prime p\] by (cases "p = 2") auto have "prime (nat p)" using \prime p\ prime_int_nat_transfer by blast have "Jacobi (-1) p = Legendre (-1) p" using prime_p_Jacobi_eq_Legendre[OF \prime p\] . also have "... = (-1) ^ ((nat p - 1) div 2)" using \prime p\ \2 < p\ and supplement1_Legendre[of "nat p"] by (metis int_nat_eq nat_mono_iff nat_numeral_as_int prime_gt_0_int prime_int_nat_transfer) also have "((nat p - 1) div 2) = nat ((p - 1) div 2)" by force finally have base_case: "Jacobi (-1) p = (-1) ^ nat ((p - 1) div 2)" . show ?case proof (cases "x = 1") case True then show ?thesis using base_case by simp next case False have "0 < x" using \1 < (p * x)\ prime_gt_0_int[OF \prime p\] by (meson int_one_le_iff_zero_less not_less not_less_iff_gr_or_eq zero_less_mult_iff) have "odd p" using \prime p\ \2 < p\ by (simp add: prime_odd_int) have "Jacobi (-1) (p * x) = Jacobi (-1) p * Jacobi (-1) x" using \2 < p\ \0 < x\ by simp also note base_case also have "Jacobi (-1) x = (-1) ^ nat ((x - 1) div 2)" using \0 < x\ False \odd x\ factor.IH by fastforce also have "(- 1) ^ nat ((p - 1) div 2) * (- 1) ^ nat ((x - 1) div 2) = (- 1 :: int) ^ nat ((p*x - 1) div 2)" unfolding minus_one_power_iff using \2 < p\ \0 < x\ and \odd x\ \odd p\ by (fastforce elim!: oddE simp: even_nat_iff algebra_simps) finally show ?thesis . qed qed simp_all theorem supplement1_Jacobi': "odd n \ 1 < n \ Jacobi (-1) n = (if n mod 4 = 1 then 1 else -1)" by (simp add: even_nat_iff minus_one_power_iff supplement1_Jacobi) presburger? lemma Jacobi_0_eq_0: "\is_unit n \ Jacobi 0 n = 0" by (cases "prime_factorization n = {#}") (auto simp: Jacobi_def prime_factorization_empty_iff image_iff intro: Nat.gr0I) lemma is_unit_Jacobi_aux: "is_unit x \ Jacobi a x = 1" unfolding Jacobi_def using prime_factorization_empty_iff[of x] by auto lemma is_unit_Jacobi[simp]: "Jacobi a 1 = 1" "Jacobi a (-1) = 1" using is_unit_Jacobi_aux by simp_all lemma Jacobi_neg_right [simp]: "Jacobi a (-n) = Jacobi a n" proof - have * : "-n = (-1) * n" by simp show ?thesis unfolding * by (subst Jacobi_mult_right) auto qed lemma Jacobi_neg_left: assumes "odd n" "1 < n" shows "Jacobi (-a) n = (if n mod 4 = 1 then 1 else -1) * Jacobi a n" proof - have * : "-a = (-1) * a" by simp show ?thesis unfolding * Jacobi_mult_left supplement1_Jacobi'[OF assms] .. qed function jacobi_code :: "int \ int \ int" where "jacobi_code a n = ( if n = 0 then 0 else if n = 1 then 1 else if a = 1 then 1 else if n < 0 then jacobi_code a (-n) else if even n then if even a then 0 else jacobi_code a (n div 2) else if a < 0 then (if n mod 4 = 1 then 1 else -1) * jacobi_code (-a) n else if a = 0 then 0 else if a \ n then jacobi_code (a mod n) n else if even a then (if n mod 8 \ {1, 7} then 1 else -1) * jacobi_code (a div 2) n else if coprime a n then (if n mod 4 = 3 \ a mod 4 = 3 then -1 else 1) * jacobi_code n a else 0)" by auto termination proof (relation "measure (\(a, n). nat(abs(a) + abs(n)*2) + (if n < 0 then 1 else 0) + (if a < 0 then 1 else 0))", goal_cases) case (5 a n) thus ?case by (fastforce intro!: less_le_trans[OF pos_mod_bound]) qed auto lemmas [simp del] = jacobi_code.simps lemma Jacobi_code [code]: "Jacobi a n = jacobi_code a n" proof (induction a n rule: jacobi_code.induct) case (1 a n) show ?case proof (cases "n = 0") case 2: False then show ?thesis proof (cases "n = 1") case 3: False then show ?thesis proof (cases "a = 1") case 4: False then show ?thesis proof (cases "n < 0") case True then show ?thesis using 2 3 4 1(1) by (subst jacobi_code.simps) simp next case 5: False then show ?thesis proof (cases "even n") case True then show ?thesis using 2 3 4 5 1(2) by (elim evenE, subst jacobi_code.simps) (auto simp: prime_p_Jacobi_eq_Legendre) next case 6: False then show ?thesis proof (cases "a < 0") case True then show ?thesis using 2 3 4 5 6 by(subst jacobi_code.simps, subst 1(3)[symmetric]) (simp_all add: Jacobi_neg_left) next case 7: False then show ?thesis proof (cases "a = 0") case True have *: "\ is_unit n" using 3 5 by simp then show ?thesis using Jacobi_0_eq_0[OF *] 2 3 4 5 7 True by (subst jacobi_code.simps) simp next case 8: False then show ?thesis proof (cases "a \ n") case True then show ?thesis using 2 3 4 5 6 7 8 1(4) by (subst jacobi_code.simps) simp next case 9: False then show ?thesis proof (cases "even a") case True hence "a = 2 * (a div 2)" by simp also have "Jacobi \ n = Jacobi 2 n * Jacobi (a div 2) n" by simp also have "Jacobi (a div 2) n = jacobi_code (a div 2) n" using 2 3 4 5 6 7 8 9 True by (intro 1(5)) also have "Jacobi 2 n = (if n mod 8 \ {1, 7} then 1 else - 1)" using 2 3 5 supplement2_Jacobi'[OF 6] by simp also have "\ * jacobi_code (a div 2) n = jacobi_code a n" using 2 3 4 5 6 7 8 9 True by (subst (2) jacobi_code.simps) (simp only: if_False if_True HOL.simp_thms) finally show ?thesis . next case 10: False note foo = 1 2 3 then show ?thesis proof (cases "coprime a n") case True note this_case = 2 3 4 5 6 7 8 9 10 True have "2 < a" using 10 4 7 by presburger moreover have "2 < n" using 3 5 6 by presburger ultimately have "jacobi_code a n = (if n mod 4 = 3 \ a mod 4 = 3 then - 1 else 1) * jacobi_code n a" using this_case by (subst jacobi_code.simps) simp also have "jacobi_code n a = Jacobi n a" using this_case by (intro 1(6) [symmetric]) auto also have "(if n mod 4 = 3 \ a mod 4 = 3 then -1 else 1) * \ = Jacobi a n" using this_case and \2 < a\ by (intro Quadratic_Reciprocity_Jacobi' [symmetric]) (auto simp: coprime_commute) finally show ?thesis .. next case False have *: "0 < a" "0 < n" using 5 7 8 9 by linarith+ show ?thesis using 1 2 3 4 5 6 7 8 9 10 False * by (subst jacobi_code.simps) (auto simp: Jacobi_eq_0_not_coprime) qed qed qed qed qed qed qed qed (subst jacobi_code.simps, simp) qed (subst jacobi_code.simps, simp) qed (subst jacobi_code.simps, simp) qed lemma Jacobi_eq_0_imp_not_coprime: assumes "p \ 0" "p \ 1" shows "Jacobi n p = 0 \ \coprime n p" using assms Jacobi_mod_cong coprime_iff_invertible_int by force lemma Jacobi_eq_0_iff_not_coprime: assumes "p \ 0" "p \ 1" shows "Jacobi n p = 0 \ \coprime n p" proof - from assms and Jacobi_eq_0_imp_not_coprime show ?thesis using Jacobi_eq_0_not_coprime by auto qed end \ No newline at end of file diff --git a/thys/Rank_Nullity_Theorem/Mod_Type.thy b/thys/Rank_Nullity_Theorem/Mod_Type.thy --- a/thys/Rank_Nullity_Theorem/Mod_Type.thy +++ b/thys/Rank_Nullity_Theorem/Mod_Type.thy @@ -1,557 +1,555 @@ (* Title: Mod_Type.thy Author: Jose Divasón Author: Jesús Aransay *) section\Class for modular arithmetic\ theory Mod_Type imports "HOL-Library.Numeral_Type" "HOL-Analysis.Cartesian_Euclidean_Space" Dual_Order begin subsection\Definition and properties\ text\Class for modular arithmetic. It is inspired by the locale mod\_type.\ class mod_type = times + wellorder + neg_numeral + fixes Rep :: "'a => int" and Abs :: "int => 'a" assumes type: "type_definition Rep Abs {0.. int CARD ('a)" by (rule Rep_less_n [THEN order_less_imp_le]) lemma Rep_inject_sym: "x = y \ Rep x = Rep y" by (rule type_definition.Rep_inject [OF type, symmetric]) lemma Rep_inverse: "Abs (Rep x) = x" by (rule type_definition.Rep_inverse [OF type]) lemma Abs_inverse: "m \ {0.. Rep (Abs m) = m" by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod int CARD ('a))) = m mod int CARD ('a)" - by (simp add: Abs_inverse pos_mod_conj [OF size0]) + using size0 by (auto simp add: Abs_inverse) lemma Rep_Abs_0: "Rep (Abs 0) = 0" apply (rule Abs_inverse [of 0]) using size0 by simp lemma Rep_0: "Rep 0 = 0" by (simp add: zero_def Rep_Abs_0) lemma Rep_Abs_1: "Rep (Abs 1) = 1" by (simp add: Abs_inverse size1) lemma Rep_1: "Rep 1 = 1" by (simp add: one_def Rep_Abs_1) lemma Rep_mod: "Rep x mod int CARD ('a) = Rep x" apply (rule_tac x=x in type_definition.Abs_cases [OF type]) apply (simp add: type_definition.Abs_inverse [OF type]) done lemmas Rep_simps = Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1 subsection\Conversion between a modular class and the subset of natural numbers associated.\ text\Definitions to make transformations among elements of a modular class and naturals\ definition to_nat :: "'a => nat" where "to_nat = nat \ Rep" definition Abs' :: "int => 'a" where "Abs' x = Abs(x mod int CARD ('a))" definition from_nat :: "nat \ 'a" where "from_nat = (Abs' \ int)" lemma bij_Rep: "bij_betw (Rep) (UNIV::'a set) {0.. Rep x" using bij_Rep unfolding bij_betw_def by auto lemma bij_Abs: "bij_betw (Abs) {0..xa\{0.. x" and x2: "x < int CARD('a)" show "x \ int ` {0::nat.. {0::nat..'a) {0::nat..'a) (int ` {0::nat..'a)` {0::nat..'a) = (to_nat::'a=>nat)" proof (unfold the_inv_into_def fun_eq_iff from_nat_def to_nat_def o_def, clarify) fix x::"'a" show "(THE y::nat. y \ {0::nat.. Abs' (int y) = x) = nat (Rep x)" proof (rule the_equality, auto) show " Abs' (Rep x) = x" by (metis Abs'_def Rep_inverse Rep_mod) show "nat (Rep x) < CARD('a)" by (metis (full_types) Rep_less_n nat_int size0 zless_nat_conj) assume x: "\ (0::int) \ Rep x" show "(0::nat) < CARD('a)" and "Abs' (0::int) = x" using Rep_ge_0 x by auto next fix y::nat assume y: "y < CARD('a)" have "(Rep(Abs'(int y)::'a)) = (Rep((Abs(int y mod int CARD('a)))::'a))" unfolding Abs'_def .. also have "... = (Rep (Abs (int y)::'a))" using zmod_int[of y "CARD('a)"] using y mod_less by auto also have "... = (int y)" proof (rule Abs_inverse) show "int y \ {0::int.. (UNIV::'a set)) = 0" proof (rule Least_equality, auto) fix y::"'a" have "(0::'a) \ Abs (Rep y mod int CARD('a))" using strict_mono_Rep unfolding strict_mono_def by (metis (opaque_lifting, mono_tags) Rep_0 Rep_ge_0 strict_mono_Rep strict_mono_less_eq) also have "... = y" by (metis Rep_inverse Rep_mod) finally show "(0::'a) \ y" . qed lemma add_to_nat_def: "x + y = from_nat (to_nat x + to_nat y)" unfolding from_nat_def to_nat_def o_def using Rep_ge_0[of x] using Rep_ge_0[of y] using Rep_less_n[of x] Rep_less_n[of y] unfolding Abs'_def unfolding add_def[of x y] by auto lemma to_nat_1: "to_nat 1 = 1" by (simp add: to_nat_def Rep_1) lemma add_def': shows "x + y = Abs' (Rep x + Rep y)" unfolding Abs'_def using add_def by simp lemma Abs'_0: shows "Abs' (CARD('a))=(0::'a)" by (metis (opaque_lifting, mono_tags) Abs'_def mod_self zero_def) lemma Rep_plus_one_le_card: assumes a: "a + 1 \ 0" shows "(Rep a) + 1 < CARD ('a)" proof (rule ccontr) assume "\ Rep a + 1 < CARD('a)" hence to_nat_eq_card: "Rep a + 1 = CARD('a)" using Rep_less_n by (simp add: add1_zle_eq order_class.less_le) have "a+1 = Abs' (Rep a + Rep (1::'a))" using add_def' by auto also have "... = Abs' ((Rep a) + 1)" using Rep_1 by simp also have "... = Abs' (CARD('a))" unfolding to_nat_eq_card .. also have "... = 0" using Abs'_0 by auto finally show False using a by contradiction qed lemma to_nat_plus_one_less_card: "\a. a+1 \ 0 --> to_nat a + 1 < CARD('a)" proof (clarify) fix a assume a: "a + 1 \ 0" have "Rep a + 1 < int CARD('a)" using Rep_plus_one_le_card[OF a] by auto hence "nat (Rep a + 1) < nat (int CARD('a))" unfolding zless_nat_conj using size0 by fast thus "to_nat a + 1 < CARD('a)" unfolding to_nat_def o_def using nat_add_distrib[OF Rep_ge_0] by simp qed corollary to_nat_plus_one_less_card': assumes "a+1 \ 0" shows "to_nat a + 1 < CARD('a)" using to_nat_plus_one_less_card assms by simp lemma strict_mono_to_nat: "strict_mono to_nat" using strict_mono_Rep unfolding strict_mono_def to_nat_def using Rep_ge_0 by (metis comp_apply nat_less_eq_zless) lemma to_nat_eq [simp]: "to_nat x = to_nat y \ x = y" using injD [OF bij_betw_imp_inj_on[OF bij_to_nat]] by blast lemma mod_type_forall_eq [simp]: "(\j::'a. (to_nat j) P j) = (\a. P a)" proof (auto) fix a assume a: "\j. (to_nat::'a=>nat) j < CARD('a) \ P j" have "(to_nat::'a=>nat) a < CARD('a)" using bij_to_nat unfolding bij_betw_def by auto thus "P a" using a by auto qed lemma to_nat_from_nat: assumes t:"to_nat j = k" shows "from_nat k = j" proof - have "from_nat k = from_nat (to_nat j)" unfolding t .. also have "... = from_nat (the_inv_into {0.. from_nat ` {0.. b" shows "to_nat a \ to_nat b" proof (cases "a=b") case True thus ?thesis by auto next case False hence "a (n::'a)" using least_0 by (metis (full_types) Least_le UNIV_I) lemma to_nat_from_nat_id: assumes x: "x'a) {0.. {0.. j" and "j from_nat j" proof (cases "i=j") case True have "(from_nat i::'a) = from_nat j" using True by simp thus ?thesis by simp next case False hence "i b" proof - have "a + 1 = (from_nat (to_nat (a + 1))::'a)" using from_nat_to_nat_id [of "a+1",symmetric] . also have "... \ (from_nat (to_nat (b::'a))::'a)" proof (rule from_nat_mono') have "to_nat a < to_nat b" using ab by (metis to_nat_mono) hence "to_nat a + 1 \ to_nat b" by simp thus "to_nat b < CARD ('a)" using bij_to_nat unfolding bij_betw_def by auto hence "to_nat a + 1 < CARD ('a)" by (metis \to_nat a + 1 \ to_nat b\ preorder_class.le_less_trans) thus "to_nat (a + 1) \ to_nat b" by (metis \to_nat a + 1 \ to_nat b\ to_nat_suc) qed also have "... = b" by (metis from_nat_to_nat_id) finally show "a + (1::'a) \ b" . qed lemma le_Suc': assumes ab: "a + 1 \ b" and less_card: "(to_nat a) + 1 < CARD ('a)" shows "a < b" proof - have "a = (from_nat (to_nat a)::'a)" using from_nat_to_nat_id [of "a",symmetric] . also have "... < (from_nat (to_nat b)::'a)" proof (rule from_nat_mono) show "to_nat b < CARD('a)" using bij_to_nat unfolding bij_betw_def by auto have "to_nat (a + 1) \ to_nat b" using ab by (metis to_nat_mono') hence "to_nat (a) + 1 \ to_nat b" using to_nat_suc[OF less_card] by auto thus "to_nat a < to_nat b" by simp qed finally show "a < b" by (metis to_nat_from_nat) qed lemma Suc_le: assumes less_card: "(to_nat a) + 1 < CARD ('a)" shows "a < a + 1" proof - have "(to_nat a) < (to_nat a) + 1" by simp hence "(to_nat a) < to_nat (a + 1)" by (metis less_card to_nat_suc) hence "(from_nat (to_nat a)::'a) < from_nat (to_nat (a + 1))" by (rule from_nat_mono, metis less_card to_nat_suc) thus "a < a + 1" by (metis to_nat_from_nat) qed lemma Suc_le': fixes a::'a assumes "a + 1 \ 0" shows "a < a + 1" using Suc_le to_nat_plus_one_less_card assms by blast lemma from_nat_not_eq: assumes a_eq_to_nat: "a \ to_nat b" and a_less_card: "a b" proof (rule ccontr) assume "\ from_nat a \ b" hence "from_nat a = b" by simp hence "to_nat ((from_nat a)::'a) = to_nat b" by auto thus False by (metis a_eq_to_nat a_less_card to_nat_from_nat_id) qed lemma Suc_less: fixes i::'a assumes "i j" shows "i+1a::'a. a \ -1" proof (clarify) fix a::'a have zero_ge_card_1: "0 \ int CARD('a) - 1" using size1 by auto have card_less: "int CARD('a) - 1 < int CARD('a)" by auto have not_zero: "1 mod int CARD('a) \ 0" by (metis (opaque_lifting, mono_tags) Rep_Abs_1 Rep_mod zero_neq_one) have int_card: "int (CARD('a) - 1) = int CARD('a) - 1" using of_nat_diff[of 1 "CARD ('a)"] using size1 by simp have "a = Abs' (Rep a)" by (metis (opaque_lifting, mono_tags) Rep_0 add_0_right add_def' monoid_add_class.add.right_neutral) also have "... = Abs' (int (nat (Rep a)))" by (metis Rep_ge_0 int_nat_eq) also have "... \ Abs' (int (CARD('a) - 1))" proof (rule from_nat_mono'[unfolded from_nat_def o_def, of "nat (Rep a)" "CARD('a) - 1"]) show "nat (Rep a) \ CARD('a) - 1" using Rep_less_n using int_card nat_le_iff by auto show "CARD('a) - 1 < CARD('a)" using finite_UNIV_card_ge_0 finite_mod_type by fastforce qed also have "... = - 1" unfolding Abs'_def unfolding minus_def zmod_zminus1_eq_if unfolding Rep_1 apply (rule cong [of Abs], rule refl) unfolding if_not_P [OF not_zero] unfolding int_card unfolding mod_pos_pos_trivial[OF zero_ge_card_1 card_less] using mod_pos_pos_trivial[OF _ size1] by presburger finally show "a \ -1" by fastforce qed lemma a_eq_minus_1: "\a::'a. a+1 = 0 \ a = -1" by (metis eq_neg_iff_add_eq_0) lemma forall_from_nat_rw: shows "(\x\{0..x. P (from_nat x))" proof (auto) fix y assume *: "\x\{0.. (UNIV::'a set)" by auto from this obtain x where x1: "from_nat y = (from_nat x::'a)" and x2: "x\{0.. CARD('a)" shows "a+1 \ 0" proof (rule ccontr, simp) assume a_plus_one_zero: "a + 1 = 0" hence rep_eq_card: "Rep a + 1 = CARD('a)" using assms to_nat_0 Suc_eq_plus1 Suc_lessI Zero_not_Suc to_nat_less_card to_nat_suc by (metis (opaque_lifting, mono_tags)) moreover have "Rep a + 1 < CARD('a)" using Abs'_0 Rep_1 Suc_eq_plus1 Suc_lessI Suc_neq_Zero add_def' assms rep_eq_card to_nat_0 to_nat_less_card to_nat_suc by (metis (opaque_lifting, mono_tags)) ultimately show False by fastforce qed lemma from_nat_suc: shows "from_nat (j + 1) = from_nat j + 1" unfolding from_nat_def o_def Abs'_def add_def' Rep_1 Rep_Abs_mod unfolding of_nat_add apply (subst mod_add_left_eq) unfolding of_nat_1 .. lemma to_nat_plus_1_set: shows "to_nat a + 1 \ {1..Instantiations\ instantiation bit0 and bit1:: (finite) mod_type begin definition "(Rep::'a bit0 => int) x = Rep_bit0 x" definition "(Abs::int => 'a bit0) x = Abs_bit0' x" definition "(Rep::'a bit1 => int) x = Rep_bit1 x" definition "(Abs::int => 'a bit1) x = Abs_bit1' x" instance proof show "(0::'a bit0) = Abs (0::int)" unfolding Abs_bit0_def Abs_bit0'_def zero_bit0_def by auto show "(1::int) < int CARD('a bit0)" by (metis bit0.size1) show "type_definition (Rep::'a bit0 => int) (Abs:: int => 'a bit0) {0::int..x::'a bit0. Rep_bit0 x \ {0::int..x::'a bit0. Abs_bit0 (Rep_bit0 x mod int CARD('a bit0)) = x" by (metis Rep_bit0_inverse bit0.Rep_mod) show "\y::int. y \ {0::int.. Rep_bit0 ((Abs_bit0::int => 'a bit0) (y mod int CARD('a bit0))) = y" by (metis bit0.Abs_inverse bit0.Rep_mod) qed show "(1::'a bit0) = Abs (1::int)" unfolding Abs_bit0_def Abs_bit0'_def one_bit0_def by (metis bit0.of_nat_eq of_nat_1 one_bit0_def) fix x y :: "'a bit0" show "x + y = Abs ((Rep x + Rep y) mod int CARD('a bit0))" unfolding Abs_bit0_def Rep_bit0_def plus_bit0_def Abs_bit0'_def by fastforce show "x * y = Abs (Rep x * Rep y mod int CARD('a bit0))" unfolding Abs_bit0_def Rep_bit0_def times_bit0_def Abs_bit0'_def by fastforce show "x - y = Abs ((Rep x - Rep y) mod int CARD('a bit0))" unfolding Abs_bit0_def Rep_bit0_def minus_bit0_def Abs_bit0'_def by fastforce show "- x = Abs (- Rep x mod int CARD('a bit0))" unfolding Abs_bit0_def Rep_bit0_def uminus_bit0_def Abs_bit0'_def by fastforce show "(0::'a bit1) = Abs (0::int)" unfolding Abs_bit1_def Abs_bit1'_def zero_bit1_def by auto show "(1::int) < int CARD('a bit1)" by (metis bit1.size1) show "(1::'a bit1) = Abs (1::int)" unfolding Abs_bit1_def Abs_bit1'_def one_bit1_def by (metis bit1.of_nat_eq of_nat_1 one_bit1_def) fix x y :: "'a bit1" show "x + y = Abs ((Rep x + Rep y) mod int CARD('a bit1))" unfolding Abs_bit1_def Abs_bit1'_def Rep_bit1_def plus_bit1_def by fastforce show "x * y = Abs (Rep x * Rep y mod int CARD('a bit1))" unfolding Abs_bit1_def Rep_bit1_def times_bit1_def Abs_bit1'_def by fastforce show "x - y = Abs ((Rep x - Rep y) mod int CARD('a bit1))" unfolding Abs_bit1_def Rep_bit1_def minus_bit1_def Abs_bit1'_def by fastforce show "- x = Abs (- Rep x mod int CARD('a bit1))" unfolding Abs_bit1_def Rep_bit1_def uminus_bit1_def Abs_bit1'_def by fastforce show "type_definition (Rep::'a bit1 => int) (Abs:: int => 'a bit1) {0::int..x::'a bit1. Rep_bit1 x \ {0::int..x::'a bit1. Abs_bit1 (Rep_bit1 x mod int CARD('a bit1)) = x" by (metis Rep_bit1_inverse bit1.Rep_mod) show "\y::int. y \ {0::int.. Rep_bit1 ((Abs_bit1::int => 'a bit1) (y mod int CARD('a bit1))) = y" by (metis bit1.Abs_inverse bit1.Rep_mod) qed show "strict_mono (Rep::'a bit0 => int)" unfolding strict_mono_def by (metis Rep_bit0_def less_bit0_def) show "strict_mono (Rep::'a bit1 => int)" unfolding strict_mono_def by (metis Rep_bit1_def less_bit1_def) qed end end - - diff --git a/thys/SumSquares/FourSquares.thy b/thys/SumSquares/FourSquares.thy --- a/thys/SumSquares/FourSquares.thy +++ b/thys/SumSquares/FourSquares.thy @@ -1,461 +1,461 @@ (* Title: FourSquares.thy Author: Roelof Oosterhuis, 2007 Rijksuniversiteit Groningen *) section \Lagrange's four-square theorem\ theory FourSquares imports "HOL-Number_Theory.Number_Theory" begin context fixes sum4sq_nat :: "nat \ nat \ nat \ nat \ nat" defines "sum4sq_nat a b c d \ a^2+b^2+c^2+d^2" fixes is_sum4sq_nat :: "nat \ bool" defines "is_sum4sq_nat n \ (\ a b c d. n = sum4sq_nat a b c d)" begin (* TODO: move this lemma to the distribution (?). (It is used also in QuadForm and TwoSquares) *) private lemma best_division_abs: "(n::int) > 0 \ \ k. 2 * \a - k*n\ \ n" proof - assume a: "n > 0" define k where "k = a div n" have h: "a - k * n = a mod n" by (simp add: div_mult_mod_eq algebra_simps k_def) thus ?thesis proof (cases "2 * (a mod n) \ n") case True hence "2 * \a - k*n\ \ n" using h pos_mod_sign a by auto thus ?thesis by blast next case False hence "2 * (n - a mod n) \ n" by auto have "a - (k+1)*n = a mod n - n" using h by (simp add: algebra_simps) hence "2 * \a - (k+1)*n\ \ n" using h pos_mod_bound[of n a] a False by fastforce thus ?thesis by blast qed qed text \Shows that all nonnegative integers can be written as the sum of four squares. The proof consists of the following steps: \begin{itemize}\item For every prime $p=2n+1$ the two sets of residue classes $$\{x^2 \bmod p\mid 0\le x\le n\} ~{\rm and}~ \{-1-y^2 \bmod p \mid 0 \le y \le n\}$$ both contain $n+1$ different elements and therefore they must have at least one element in common. \item Hence there exist $x,y$ such that $x^2+y^2+1^2+0^2$ is a multiple of $p$. \item The next step is to show, by an infinite descent, that $p$ itself can be written as the sum of four squares. \item Finally, using the multiplicity of this form, the same holds for all positive numbers. \end{itemize}\ private definition sum4sq_int :: "int \ int \ int \ int \ int" where "sum4sq_int = (\(a,b,c,d). a^2+b^2+c^2+d^2)" private definition is_sum4sq_int :: "int \ bool" where "is_sum4sq_int n \ (\ a b c d. n = sum4sq_int(a,b,c,d))" private lemma mult_sum4sq_int: "sum4sq_int(a,b,c,d) * sum4sq_int(p,q,r,s) = sum4sq_int(a*p+b*q+c*r+d*s, a*q-b*p-c*s+d*r, a*r+b*s-c*p-d*q, a*s-b*r+c*q-d*p)" by (unfold sum4sq_int_def, simp add: eval_nat_numeral field_simps) private lemma sum4sq_int_nat_eq: "sum4sq_nat a b c d = sum4sq_int (a, b, c, d)" unfolding sum4sq_nat_def sum4sq_int_def by simp private lemma is_sum4sq_int_nat_eq: "is_sum4sq_nat n = is_sum4sq_int (int n)" unfolding is_sum4sq_nat_def is_sum4sq_int_def proof assume "\a b c d. n = sum4sq_nat a b c d" thus "\a b c d. int n = sum4sq_int (a, b, c, d)" using sum4sq_int_nat_eq by force next assume "\a b c d. int n = sum4sq_int (a, b, c, d)" then obtain a b c d where "int n = sum4sq_int (a, b, c, d)" by blast hence "int n = sum4sq_int (int (nat \a\), int (nat \b\), int (nat \c\), int (nat \d\))" unfolding sum4sq_int_def by simp hence "int n = int (sum4sq_nat (nat \a\) (nat \b\) (nat \c\) (nat \d\))" using sum4sq_int_nat_eq by presburger thus "\a b c d. n = sum4sq_nat a b c d" by auto qed private lemma is_mult_sum4sq_int: "is_sum4sq_int x \ is_sum4sq_int y \ is_sum4sq_int (x*y)" by (unfold is_sum4sq_int_def, auto simp only: mult_sum4sq_int, blast) private lemma is_mult_sum4sq_nat: "is_sum4sq_nat x \ is_sum4sq_nat y \ is_sum4sq_nat (x*y)" using is_mult_sum4sq_int is_sum4sq_int_nat_eq by simp private lemma mult_oddprime_is_sum4sq: "\ prime (nat p); odd p \ \ \ t. 0 t

is_sum4sq_int (p*t)" proof - assume p1: "prime (nat p)" then have p0: "p > 1" and "prime p" by (simp_all add: prime_int_nat_transfer prime_nat_iff) assume p2: "odd p" then obtain n where n: "p = 2*n+1" using oddE by blast with p1 have n0: "n > 0" by (auto simp add: prime_nat_iff) let ?C = "{0 ..< p}" let ?D = "{0 .. n}" let ?f = "%x. x^2 mod p" let ?g = "%x. (-1-x^2) mod p" let ?A = "?f ` ?D" let ?B = "?g ` ?D" have finC: "finite ?C" by simp have finD: "finite ?D" by simp from p0 have AsubC: "?A \ ?C" and BsubC: "?B \ ?C" - by (auto simp add: pos_mod_conj) + by auto with finC have finA: "finite ?A" and finB: "finite ?B" by (auto simp add: finite_subset) from AsubC BsubC have AunBsubC: "?A \ ?B \ ?C" by (rule Un_least) from p0 have cardC: "card ?C = nat p" using card_atLeastZeroLessThan_int by blast from n0 have cardD: "card ?D = 1+ nat n" by simp have cardA: "card ?A = card ?D" proof - have "inj_on ?f ?D" proof (unfold inj_on_def, auto) fix x fix y assume x0: "0 \ x" and xn: "x \ n" and y0: "0 \ y" and yn: " y \ n" and xyp: "x^2 mod p = y^2 mod p" with p0 have "[x^2 = y^2] (mod p)" using cong_def by blast hence "p dvd x^2-y^2" using cong_iff_dvd_diff by blast hence "p dvd (x+y)*(x-y)" by (simp add: power2_eq_square algebra_simps) hence "p dvd x+y \ p dvd x-y" using \prime p\ p0 by (auto dest: prime_dvd_multD) moreover { assume "p dvd x+y" moreover from xn yn n have "x+y < p" by auto ultimately have "\ x+y > 0" by (auto simp add: zdvd_not_zless) with x0 y0 have "x = y" by auto } \ \both are zero\ moreover { assume ass: "p dvd x-y" have "x = y" proof (rule ccontr, case_tac "x-y \ 0") assume "x-y \ 0" and "x \ y" hence "x-y > 0" by auto with ass have "\ x-y < p" by (auto simp add: zdvd_not_zless) with xn y0 n p0 show "False" by auto next assume "\ 0 \ x-y" hence "y-x > 0" by auto moreover from x0 yn n p0 have "y-x < p" by auto ultimately have "\ p dvd y-x" by (auto simp add: zdvd_not_zless) moreover from ass have "p dvd -(x-y)" by (simp only: dvd_minus_iff) ultimately show "False" by auto qed } ultimately show "x=y" by auto qed with finD show ?thesis by (simp only: inj_on_iff_eq_card) qed have cardB: "card ?B = card ?D" proof - have "inj_on ?g ?D" proof (unfold inj_on_def, auto) fix x fix y assume x0: "0 \ x" and xn: "x \ n" and y0: "0 \ y" and yn: " y \ n" and xyp: "(-1-x^2) mod p = (-1-y^2) mod p" with p0 have "[-1-y^2 = -1-x^2] (mod p)" by (simp only: cong_def) hence "p dvd (-1-y^2) - (-1-x^2)" by (simp only: cong_iff_dvd_diff) moreover have "-1-y^2 - (-1-x^2) = x^2 - y^2" by arith ultimately have "p dvd x^2-y^2" by simp hence "p dvd (x+y)*(x-y)" by (simp add: power2_eq_square algebra_simps) with p1 have "p dvd x+y \ p dvd x-y" using \prime p\ p0 by (auto dest: prime_dvd_multD) moreover { assume "p dvd x+y" moreover from xn yn n have "x+y < p" by auto ultimately have "\ x+y > 0" by (auto simp add: zdvd_not_zless) with x0 y0 have "x = y" by auto } \ \both are zero\ moreover { assume ass: "p dvd x-y" have "x = y" proof (rule ccontr, case_tac "x-y \ 0") assume "x-y \ 0" and "x \ y" hence "x-y > 0" by auto with ass have "\ x-y < p" by (auto simp add: zdvd_not_zless) with xn y0 n p0 show "False" by auto next assume "\ 0 \ x-y" hence "y-x > 0" by auto moreover from x0 yn n p0 have "y-x < p" by auto ultimately have "\ p dvd y-x" by (auto simp add: zdvd_not_zless) moreover from ass have "p dvd -(x-y)" by (simp only: dvd_minus_iff) ultimately show "False" by auto qed } ultimately show "x=y" by auto qed with finD show ?thesis by (simp only: inj_on_iff_eq_card) qed have "?A \ ?B \ {}" proof (rule ccontr, auto) assume ABdisj: "?A \ ?B = {}" from cardA cardB cardD have "2 + 2*(nat n) = card ?A + card ?B" by auto also with finA finB ABdisj have "\ = card (?A \ ?B)" by (simp only: card_Un_disjoint) also with finC AunBsubC have "\ \ card ?C" by (simp only: card_mono) also with cardC have "\ = nat p" by simp finally have "2 + 2*(nat n) \ nat p" by simp with n show "False" by arith qed then obtain z where "z \ ?A \ z \ ?B" by auto then obtain x y where xy: "x \ ?D \ y \ ?D \ z = x^2 mod p \ z = (-1-y^2) mod p" by blast with p0 have "[x^2=-1-y^2](mod p)" by (simp add: cong_def) hence "p dvd x^2-(-1-y^2)" by (simp only: cong_iff_dvd_diff) moreover have "x^2-(-1-y^2)=x^2+y^2+1" by arith ultimately have "p dvd sum4sq_int(x,y,1,0)" by (auto simp add: sum4sq_int_def) then obtain t where t: "p * t = sum4sq_int(x,y,1,0)" by (auto simp only: dvd_def eq_refl) hence "is_sum4sq_int (p*t)" by (unfold is_sum4sq_int_def, auto) moreover have "t > 0 \ t < p" proof have "x^2 \ 0 \ y^2 \ 0" by simp hence "x^2+y^2+1 > 0" by arith with t have "p*t > 0" by (unfold sum4sq_int_def, auto) moreover { assume "t < 0" with p0 have "p*t < p*0" by (simp only: zmult_zless_mono2) hence "p*t < 0" by simp } moreover { assume "t = 0" hence "p*t = 0" by simp } ultimately have "\ t < 0 \ t \ 0" by auto thus "t > 0" by simp from xy have "x^2 \ n^2 \ y^2 \ n^2" by (auto simp add: power_mono) hence "x^2+y^2+1 \ 2*n^2 + 1" by auto with t have contr: "p*t \ 2*n^2+1" by (simp add: sum4sq_int_def) moreover { assume "t > n+1" with p0 have "p*(n+1) < p*t" by (simp only: zmult_zless_mono2) with n have "p*t > (2*n+1)*n + (2*n+1)*1" by (simp only: distrib_left) hence "p*t > 2*n*n + n + 2*n + 1" by (simp only: distrib_right mult_1_left) with n0 have "p*t > 2*n^2 + 1" by (simp add: power2_eq_square) } ultimately have "\ t > n+1" by auto with n0 n show "t < p" by auto qed ultimately show ?thesis by blast qed private lemma zprime_is_sum4sq: "prime (nat p) \ is_sum4sq_int p" proof (cases) assume p2: "p=2" hence "p = sum4sq_int(1,1,0,0)" by (auto simp add: sum4sq_int_def) thus ?thesis by (auto simp add: is_sum4sq_int_def) next assume "\ p =2" and prp: "prime (nat p)" hence "\ nat p = 2" by simp with prp have "2 < nat p" using prime_nat_iff by force moreover with prp have "odd (nat p)" using prime_odd_nat[of "nat p"] by blast ultimately have "odd p" by (simp add: even_nat_iff) with prp have "\ t. 0 t

is_sum4sq_int (p*t)" by (rule mult_oddprime_is_sum4sq) then obtain a b c d t where pt_sol: "0 t

p*t = sum4sq_int(a,b,c,d)" by (unfold is_sum4sq_int_def, blast) hence Qt: "0 t

(\ a1 a2 a3 a4. p*t = sum4sq_int(a1,a2,a3,a4))" (is "?Q t") by blast have "?Q 1" proof (rule ccontr) assume nQ1: "\ ?Q 1" have "\ ?Q t" proof (induct t rule: infinite_descent0_measure[where V="\x. (nat x)- 1"], clarify) fix x a b c d assume "nat x - 1 = 0" and "x > 0" and s: "p*x=sum4sq_int(a,b,c,d)" and "x < p" moreover hence "x = 1" by arith ultimately have "?Q 1" by auto with nQ1 show False by auto next fix x assume "0 < nat x - 1" and "\ \ ?Q x" then obtain a1 a2 a3 a4 where ass: "1 x

p*x = sum4sq_int(a1,a2,a3,a4)" by auto have "\ y. nat y - 1 < nat x - 1 \ ?Q y" proof (cases) assume evx: "even x" hence "even (x*p)" by simp with ass have ev1234: "even (a1^2+a2^2 + a3^2+a4^2)" by (auto simp add: sum4sq_int_def ac_simps) have "\ b1 b2 b3 b4. p*x=sum4sq_int(b1,b2,b3,b4) \ even (b1+b2) \ even (b3+b4)" proof (cases) assume ev12: "even (a1^2+a2^2)" with ev1234 ass show ?thesis by auto next assume "\ even (a1^2+a2^2)" hence odd12: "odd (a1^2+a2^2)" by simp with ev1234 have odd34: "odd (a3^2+a4^2)" by auto show ?thesis proof (cases) assume ev1: "even (a1^2)" with odd12 have odd2: "odd (a2^2)" by simp show ?thesis proof (cases) assume "even (a3^2)" moreover from ass have "p*x = sum4sq_int(a1,a3,a2,a4)" by (auto simp add: sum4sq_int_def) ultimately show ?thesis using odd2 odd34 ev1 by auto next assume "\ even (a3^2)" moreover from ass have "p*x = sum4sq_int(a1,a4,a2,a3)" by (auto simp add: sum4sq_int_def) ultimately show ?thesis using odd34 odd2 ev1 by auto qed next assume odd1: "\ even (a1^2)" with odd12 have ev2: "even (a2^2)" by simp show ?thesis proof (cases) assume "even (a3^2)" moreover from ass have "sum4sq_int(a1,a4,a2,a3)=p*x" by (auto simp add: sum4sq_int_def) ultimately show ?thesis using odd34 odd1 ev2 by force next assume "\ even (a3^2)" moreover from ass have "sum4sq_int(a1,a3,a2,a4)=p*x" by (auto simp add: sum4sq_int_def) ultimately show ?thesis using odd34 odd1 ev2 by force qed qed qed then obtain b1 b2 b3 b4 where b: "p*x = sum4sq_int(b1,b2,b3,b4) \ even (b1+b2) \ even (b3+b4)" by auto then obtain c1 c3 where c13: "b1+b2 = 2*c1 \ b3+b4 = 2*c3" using evenE[of "b1+b2"] evenE[of "b3+b4"] by meson from b have "even (b1-b2) \ even (b3-b4)" by simp then obtain c2 c4 where c24: "b1-b2 = 2*c2 \ b3-b4 = 2*c4" using evenE[of "b1-b2"] evenE[of "b3-b4"] by meson from evx obtain y where y: "x = 2*y" using evenE by blast hence "4*(p*y) = 2*(p*x)" by (simp add: ac_simps) also from b have "\ = 2*b1^2 + 2*b2^2 + 2*b3^2 + 2*b4^2" by (auto simp: sum4sq_int_def) also have "\ = (b1 + b2)^2 + (b1 - b2)^2 + (b3 + b4)^2 + (b3 - b4)^2" by (auto simp add: power2_eq_square algebra_simps) also with c13 c24 have "\ = 4*(c1^2 + c2^2 + c3^2 + c4^2)" by (auto simp add: power_mult_distrib) finally have "p * y = sum4sq_int(c1,c2,c3,c4)" by (auto simp add: sum4sq_int_def) moreover from y ass have "0 < y \ y < p \ (nat y) - 1 < (nat x) - 1" by arith ultimately show ?thesis by blast next assume xodd: "\ even x" with ass have "\ c1 c2 c3 c4. 2*\a1-c1*x\\x \ 2*\a2-c2*x\\x \ 2*\a3-c3*x\\x \ 2*\a4-c4*x\\x" by (simp add: best_division_abs) then obtain b1 c1 b2 c2 b3 c3 b4 c4 where bc_def: "b1 = a1-c1*x \ b2 = a2-c2*x \ b3 = a3-c3*x \ b4 = a4-c4*x" and "2*\b1\\x \ 2*\b2\\x \ 2*\b3\\x \ 2*\b4\\x" by blast moreover have "2*\b1\\x \ 2*\b2\\x \ 2*\b3\\x \ 2*\b4\\x" using xodd by fastforce ultimately have bc_abs: "2*\b1\ 2*\b2\ 2*\b3\ 2*\b4\ x dvd ?A2 \ x dvd ?A3 \ x dvd ?A4" proof (safe) from bc_def have "?A1 = (b1+c1*x)*b1 + (b2+c2*x)*b2 + (b3+c3*x)*b3 + (b4+c4*x)*b4" by simp also with y have "\ = x*(y + c1*b1 + c2*b2 + c3*b3 + c4*b4)" by (auto simp add: distrib_left power2_eq_square ac_simps) finally show "x dvd ?A1" by auto from bc_def have "?A2 = (b1+c1*x)*b2 - (b2+c2*x)*b1 - (b3+c3*x)*b4 + (b4+c4*x)*b3" by simp also have "\ = x*(c1*b2 - c2*b1 - c3*b4 + c4*b3)" by (auto simp add: distrib_left right_diff_distrib ac_simps) finally show "x dvd ?A2" by auto from bc_def have "?A3 = (b1+c1*x)*b3 + (b2+c2*x)*b4 - (b3+c3*x)*b1 - (b4+c4*x)*b2" by simp also have "\ = x*(c1*b3 + c2*b4 - c3*b1 - c4*b2)" by (auto simp add: distrib_left right_diff_distrib ac_simps) finally show "x dvd ?A3" by auto from bc_def have "?A4 = (b1+c1*x)*b4 - (b2+c2*x)*b3 + (b3+c3*x)*b2 - (b4+c4*x)*b1" by simp also have "\ = x*(c1*b4 - c2*b3 + c3*b2 - c4*b1)" by (auto simp add: distrib_left right_diff_distrib ac_simps) finally show "x dvd ?A4" by auto qed then obtain d1 d2 d3 d4 where d: "?A1=x*d1 \ ?A2=x*d2 \ ?A3=x*d3 \ ?A4=x*d4" by (auto simp add: dvd_def) let ?D = "sum4sq_int(d1,d2,d3,d4)" from d have "x^2*?D = ?A" by (auto simp only: sum4sq_int_def power_mult_distrib distrib_left) also have "\ = sum4sq_int(a1,a2,a3,a4)*sum4sq_int(b1,b2,b3,b4)" by (simp only: mult_sum4sq_int) also with y ass have "\ = (p*x)*(x*y)" by (auto simp add: sum4sq_int_def) also have "\ = x^2*(p*y)" by (simp only: power2_eq_square ac_simps) finally have "x^2*(?D - p*y) = 0" by (auto simp add: right_diff_distrib) with ass have "p*y = ?D" by auto moreover have y_l_x: "y < x" proof - have "4*b1^2 = (2*\b1\)^2 \ 4*b2^2 = (2*\b2\)^2 \ 4*b3^2 = (2*\b3\)^2 \ 4*b4^2 = (2*\b4\)^2" by simp with bc_abs have "4*b1^2 4*b2^2 4*b3^2 4*b4^2b\" x 2 for b] by auto hence "?B < x^2" by auto with y have "x*(x-y) > 0" by (auto simp add: power2_eq_square right_diff_distrib) moreover from ass have "x > 0" by simp ultimately show ?thesis using zero_less_mult_pos by fastforce qed moreover have "y > 0" proof - have b2pos: "b1^2 \ 0 \ b2^2 \ 0 \ b3^2 \ 0 \ b4^2 \ 0" by simp hence "?B = 0 \ ?B > 0" by arith moreover { assume "?B = 0" moreover from b2pos have "?B-b1^2 \ 0 \ ?B-b2^2 \ 0 \ ?B-b3^2 \ 0 \ ?B-b4^2 \ 0" by arith ultimately have "b1^2 \ 0 \ b2^2 \ 0 \ b3^2 \ 0 \ b4^2 \ 0" by auto with b2pos have "b1^2 = 0 \ b2^2 = 0 \ b3^2 = 0 \ b4^2 = 0" by arith hence "b1 = 0 \ b2 = 0 \ b3 = 0 \ b4 = 0" by auto with bc_def have "x dvd a1 \ x dvd a2 \ x dvd a3 \ x dvd a4" by auto hence "x^2 dvd a1^2 \ x^2 dvd a2^2 \ x^2 dvd a3^2 \ x^2 dvd a4^2" by simp hence "x^2 dvd a1^2+a2^2+a3^2+a4^2" by (simp only: dvd_add) with ass have "x^2 dvd p*x" by (auto simp only: sum4sq_int_def) hence "x*x dvd x*p" by (simp only: power2_eq_square ac_simps) with ass have "nat x dvd nat p" by (simp add: nat_dvd_iff) moreover from ass prp have "x \ 0 \ x \ 1 \ x \ p \ prime (nat p)" by simp ultimately have "False" unfolding prime_nat_iff by auto } moreover { assume "?B > 0" with y have "x*y > 0" by simp moreover from ass have "x > 0" by simp ultimately have ?thesis using zero_less_mult_pos by blast } ultimately show ?thesis by auto qed moreover with y_l_x have "(nat y) - 1 < (nat x) - 1" by arith moreover from y_l_x ass have "y < p" by auto ultimately show ?thesis by blast qed thus "\ y. nat y - 1 < nat x - 1 \ \ \ ?Q y" by blast qed with Qt show "False" by simp qed thus "is_sum4sq_int p" by (auto simp add: is_sum4sq_int_def) qed private lemma prime_is_sum4sq: "prime p \ is_sum4sq_nat p" using zprime_is_sum4sq is_sum4sq_int_nat_eq by simp theorem sum_of_four_squares: "is_sum4sq_nat n" proof (induction n rule: nat_less_induct) case (1 n) show ?case proof (cases "n>1") case False hence "n = 0 \ n = 1" by auto moreover have "0 = sum4sq_nat 0 0 0 0" "1 = sum4sq_nat 1 0 0 0" unfolding sum4sq_nat_def by auto ultimately show ?thesis unfolding is_sum4sq_nat_def by blast next case True then obtain p m where dec: "prime p \ n = p * m" using prime_factor_nat[of n] by (auto elim: dvdE) moreover hence "m int a < int b" apply simp done lemma zless_trans:"\(i::int) < j; j < k\ \ i < k" apply simp done lemma zmult_pos_bignumTr0:"\L. \m. L < m \ z < x + int m" by (subgoal_tac "\m. (nat((abs z) + (abs x))) < m \ z < x + int m", blast, rule allI, rule impI, arith) lemma zle_less_trans:"\(i::int) \ j; j < k\ \ i < k" apply (simp add:less_le) done lemma zless_le_trans:"\(i::int) < j; j \ k\ \ i < k" apply (simp add:less_le) done lemma zmult_pos_bignumTr:"0 < (a::int) \ \l. \m. l < m \ z < x + (int m) * a" apply (cut_tac zmult_pos_bignumTr0[of "z" "x"]) apply (erule exE) apply (subgoal_tac "\m. L < m \ z < x + int m * a", blast) apply (rule allI, rule impI) apply (drule_tac a = m in forall_spec, assumption) apply (subgoal_tac "0 \ int m") apply (frule_tac a = "int m" and b = a in pos_zmult_pos, assumption) apply (cut_tac order_refl[of "x"]) apply (frule_tac z' = "int m" and z = "int m * a" in zadd_zle_mono[of "x" "x"], assumption+) apply (rule_tac y = "x + int m" and z = "x + (int m)* a" in less_le_trans[of "z"], assumption+) apply simp done lemma ale_shift:"\(x::ant)\ y; y = z\ \ x \ z" by simp lemma aneg_na_0[simp]:"a < 0 \ na a = 0" by (simp add:na_def) lemma amult_an_an:"an (m * n) = (an m) * (an n)" apply (simp add:an_def) apply (simp add: of_nat_mult a_z_z) done definition adiv :: "[ant, ant] \ ant" (infixl "adiv" 200) where "x adiv y = ant ((tna x) div (tna y))" definition amod :: "[ant, ant] \ ant" (infixl "amod" 200) where "x amod y = ant ((tna x) mod (tna y))" lemma apos_amod_conj:"0 < ant b \ 0 \ (ant a) amod (ant b) \ (ant a) amod (ant b) < (ant b)" by (simp add:amod_def tna_ant, simp only:ant_0[THEN sym], simp add:aless_zless) lemma amod_adiv_equality: "(ant a) = (a div b) *\<^sub>a (ant b) + ant (a mod b)" apply (simp add:adiv_def tna_ant a_z_z a_zpz asprod_mult) done lemma asp_z_Z:"z *\<^sub>a ant x \ Z\<^sub>\" by (simp add:asprod_mult z_in_aug_inf) lemma apos_in_aug_inf:"0 \ a \ a \ Z\<^sub>\" by (simp add:aug_inf_def, rule contrapos_pp, simp+, cut_tac minf_le_any[of "0"], frule ale_antisym[of "0" "-\"], assumption+, simp) lemma amult_1_both:"\0 < (w::ant); x * w = 1\ \ x = 1 \ w = 1" apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "w"], (erule disjE)+, simp, (frule sym, thin_tac "\ = 1", simp only:ant_1[THEN sym], simp del:ant_1)) apply (erule disjE, erule exE, simp, (frule sym, thin_tac "-\ = 1", simp only:ant_1[THEN sym], simp del:ant_1), simp) apply (frule sym, thin_tac "-\ = 1", simp only:ant_1[THEN sym], simp del:ant_1) apply ((erule disjE)+, erule exE, simp, frule_tac aless_imp_le[of "0" "-\"], cut_tac minf_le_any[of "0"], frule ale_antisym[of "0" "-\"], assumption+, simp only:ant_0[THEN sym], simp, frule sym, thin_tac "-\ = 1", simp only:ant_1[THEN sym], simp del:ant_1) apply ((erule disjE)+, (erule exE)+, simp only:ant_1[THEN sym], simp del:ant_1 add:a_z_z, (cut_tac a = z and b = za in mult.commute, simp, cut_tac z = za and z' = z in times_1_both, assumption+), simp) apply (erule exE, simp, cut_tac x = z and y = 0 in less_linear, erule disjE, simp, frule sym, thin_tac "-\ = 1", simp only:ant_1[THEN sym], simp del:ant_1, erule disjE, simp add:ant_0, simp, frule sym, thin_tac "\ = 1", simp only:ant_1[THEN sym], simp del:ant_1, erule disjE, erule exE, simp, frule sym, thin_tac "\ = 1", simp only:ant_1[THEN sym], simp del:ant_1, simp) done lemma poss_int_neq_0:"0 < (z::int) \ z \ 0" by simp lemma aadd_neg_negg[simp]:"\a \ (0::ant); b < 0\ \ a + b < 0" apply (frule ale_minus[of "a" "0"], simp, frule aless_minus[of "b" "0"], simp) apply (frule aadd_pos_poss[of "-a" "-b"], assumption+, simp add:aminus_add_distrib[THEN sym, of "a" "b"], frule aless_minus[of "0" "-(a + b)"], simp add:a_minus_minus) done lemma aadd_two_negg[simp]:"\a < (0::ant); b < 0\ \ a + b < 0" by auto lemma amin_aminTr:"(z::ant) \ z' \ amin z w \ amin z' w" by (simp add:amin_def) lemma amin_le1:"(z::ant) \ z' \ (amin z w) \ z'" by (simp add:amin_def, simp add:aneg_le, rule impI, frule aless_le_trans[of "w" "z" "z'"], assumption+, simp add:aless_imp_le) lemma amin_le2:"(z::ant) \ z' \ (amin w z) \ z'" by (simp add:amin_def, rule impI, frule ale_trans[of "w" "z" "z'"], assumption+) lemma Amin_geTr:"(\j \ n. f j \ Z\<^sub>\) \ (\j \ n. z \ (f j)) \ z \ (Amin n f)" apply (induct_tac n) apply (rule impI, erule conjE, simp) apply (rule impI, (erule conjE)+, cut_tac Nsetn_sub_mem1[of n], simp, drule_tac x = "Suc n" in spec, simp, rule_tac z = z and x = "Amin n f" and y = "f(Suc n)" in amin_ge1, simp+) done lemma Amin_ge:"\\j \ n. f j \ Z\<^sub>\; \j \ n. z \ (f j)\ \ z \ (Amin n f)" by (simp add:Amin_geTr) definition Abs :: "ant \ ant" where "Abs z = (if z < 0 then -z else z)" lemma Abs_pos:"0 \ Abs z" by (simp add:Abs_def, rule conjI, rule impI, cut_tac aless_minus[of "z" "0"], simp, assumption, rule impI, simp add:aneg_less[of "z" "0"]) lemma Abs_x_plus_x_pos:"0 \ (Abs x) + x" apply (case_tac "x < 0", simp add:Abs_def, simp add:aadd_minus_inv) apply (simp add:aneg_less, simp add:Abs_def, simp add:aneg_less[THEN sym, of "0" "x"], simp add:aneg_less[of "x" "0"], simp add:aadd_two_pos) done lemma Abs_ge_self:"x \ Abs x" apply (simp add:Abs_def, rule impI, cut_tac ale_minus[of "x" "0"], simp add:aminus_0, simp add:aless_imp_le) done lemma na_1:"na 1 = Suc 0" apply (simp only:ant_1[THEN sym], simp only:na_def, simp only:ant_0[THEN sym], simp only:aless_zless[of "1" "0"], simp, subgoal_tac "\ \ 1", simp) apply (simp only:ant_1[THEN sym], simp only:tna_ant, rule not_sym, simp only:ant_1[THEN sym], simp del:ant_1) done lemma ant_int:"ant (int n) = an n" by (simp add:an_def) lemma int_nat:"0 < z \ int (nat z) = z" by arith lemma int_ex_nat:"0 < z \ \n. int n = z" by (cut_tac int_nat[of z], blast, assumption) lemma eq_nat_pos_ints: "\nat (z::int) = nat (z'::int); 0 \ z; 0 \ z'\ \ z = z'" by simp lemma a_p1_gt[simp]:"\a \ \; a \ -\\ \ a < a + 1" apply (cut_tac aadd_poss_less[of a 1], simp add:aadd_commute, assumption+) apply (cut_tac zposs_aposss[of 1], simp) done lemma gt_na_poss:"(na a) < m \ 0 < m" apply (simp add:na_def) done lemma azmult_less:"\a \ \; na a < m; 0 < x\ \ a < int m *\<^sub>a x" apply (cut_tac mem_ant[of "a"]) apply (erule disjE) apply (case_tac "x = \") apply simp apply (subst less_le[of "-\" "\"]) apply simp apply (frule aless_imp_le[of "0" "x"], frule apos_neq_minf[of "x"]) apply (cut_tac mem_ant[of "x"], simp, erule exE, simp) apply (simp add:asprod_amult a_z_z) apply (simp, erule exE, simp) apply (frule_tac a = "ant z" in gt_na_poss[of _ "m"]) apply (case_tac "x = \", simp) apply (frule aless_imp_le[of "0" "x"]) apply (frule apos_neq_minf[of "x"]) apply (cut_tac mem_ant[of "x"], simp, erule exE, simp add:asprod_amult a_z_z) apply (subst aless_zless) apply (cut_tac a = "ant z" in gt_na_poss[of _ "m"], assumption) apply (smt a0_less_int_conv aposs_na_poss int_less_mono int_nat na_def of_nat_0_le_iff pos_zmult_pos tna_ant z_neq_inf) done lemma zmult_gt_one:"\2 \ m; 0 < xa\ \ 1 < int m * xa" by (metis ge2_zmult_pos mult.commute) lemma zmult_pos:"\ 0 < m; 0 < (a::int)\ \ 0 < (int m) * a" by (frule zmult_zless_mono2[of "0" "a" "int m"], simp, simp) lemma ant_int_na:"\0 \ a; a \ \ \ \ ant (int (na a)) = a" by (frule an_na[of "a"], assumption, simp add:an_def) lemma zpos_nat:"0 \ (z::int) \ \n. z = int n" apply (subgoal_tac "z = int (nat z)") apply blast apply simp done section "nsets" lemma nsetTr1:"\j \ nset a b; j \ a\ \ j \ nset (Suc a) b" apply (simp add:nset_def) done lemma nsetTr2:"j \ nset (Suc a) (Suc b) \ j - Suc 0 \ nset a b" apply (simp add:nset_def, erule conjE, simp add:skip_im_Tr4[of "j" "b"]) done lemma nsetTr3:"\j \ Suc (Suc 0); j - Suc 0 \ nset (Suc 0) (Suc n)\ \ Suc 0 < j - Suc 0" apply (simp add:nset_def, erule conjE, subgoal_tac "j \ 0", rule contrapos_pp, simp+) done lemma Suc_leD1:"Suc m \ n \ m < n" apply (insert lessI[of "m"], rule less_le_trans[of "m" "Suc m" "n"], assumption+) done lemma leI1:"n < m \ \ ((m::nat) \ n)" apply (rule contrapos_pp, simp+) done lemma neg_zle:"\ (z::int) \ z' \ z' < z" apply (simp add: not_le) done lemma nset_m_m:"nset m m = {m}" by (simp add:nset_def, rule equalityI, rule subsetI, simp, rule subsetI, simp) lemma nset_Tr51:"\j \ nset (Suc 0) (Suc (Suc n)); j \ Suc 0\ \ j - Suc 0 \ nset (Suc 0) (Suc n)" apply (simp add:nset_def, (erule conjE)+, frule_tac m = j and n = "Suc (Suc n)" and l = "Suc 0" in diff_le_mono, simp) done lemma nset_Tr52:"\j \ Suc (Suc 0); Suc 0 \ j - Suc 0\ \ \ j - Suc 0 \ Suc 0" by auto lemma nset_Suc:"nset (Suc 0) (Suc (Suc n)) = nset (Suc 0) (Suc n) \ {Suc (Suc n)}" by (auto simp add:nset_def) lemma AinequalityTr0:"x \ -\ \ \L. (\N. L < N \ (an m) < (x + an N))" apply (case_tac "x = \", simp add:an_def) apply (cut_tac mem_ant[of "x"], simp, erule exE, simp add:an_def a_zpz, simp add:aless_zless, cut_tac x = z in zmult_pos_bignumTr0[of "int m"], simp) done lemma AinequalityTr:"\0 < b \ b \ \; x \ -\\ \ \L. (\N. L < N \ (an m) < (x + (int N) *\<^sub>a b))" apply (frule_tac AinequalityTr0[of "x" "m"], erule exE, subgoal_tac "\N. L < N \ an m < x + (int N) *\<^sub>a b", blast, rule allI, rule impI) apply (drule_tac a = N in forall_spec, assumption, erule conjE, cut_tac N = N in asprod_ge[of "b"], assumption, thin_tac "x \ - \", thin_tac "b \ \", thin_tac "an m < x + an N", simp) apply (frule_tac x = "an N" and y = "int N *\<^sub>a b" and z = x in aadd_le_mono, simp only:aadd_commute[of _ "x"]) done lemma two_inequalities:"\\(n::nat). x < n \ P n; \(n::nat). y < n \ Q n\ \ \n. (max x y) < n \ (P n) \ (Q n)" by auto lemma multi_inequalityTr0:"(\j \ (n::nat). (x j) \ -\ ) \ (\L. (\N. L < N \ (\l \ n. (an m) < (x l) + (an N))))" apply (induct_tac n) apply (rule impI, simp) apply (rule AinequalityTr0[of "x 0" "m"], assumption) (** n **) apply (rule impI) apply (subgoal_tac "\l. l \ n \ l \ (Suc n)", simp) apply (erule exE) apply (frule_tac a = "Suc n" in forall_spec, simp) apply (frule_tac x = "x (Suc n)" in AinequalityTr0[of _ "m"]) apply (erule exE) apply (subgoal_tac "\N. (max L La) < N \ (\l \ (Suc n). an m < x l + an N)", blast) apply (rule allI, rule impI, rule allI, rule impI) apply (rotate_tac 1) apply (case_tac "l = Suc n", simp, drule_tac m = l and n = "Suc n" in noteq_le_less, assumption+, drule_tac x = l and n = "Suc n" in less_le_diff, simp, simp) done lemma multi_inequalityTr1:"\\j \ (n::nat). (x j) \ - \\ \ \L. (\N. L < N \ (\l \ n. (an m) < (x l) + (an N)))" by (simp add:multi_inequalityTr0) lemma gcoeff_multi_inequality:"\\N. 0 < N \ (\j \ (n::nat). (x j) \ -\ \ 0 < (b N j) \ (b N j) \ \)\ \ \L. (\N. L < N \ (\l \ n.(an m) < (x l) + (int N) *\<^sub>a (b N l)))" apply (subgoal_tac "\j \ n. x j \ - \") apply (frule multi_inequalityTr1[of "n" "x" "m"]) apply (erule exE) apply (subgoal_tac "\N. L < N \ (\l \ n. an m < x l + (int N) *\<^sub>a (b N l))") apply blast apply (rule allI, rule impI, rule allI, rule impI, drule_tac a = N in forall_spec, simp, drule_tac a = l in forall_spec, assumption, drule_tac a = N in forall_spec, assumption, drule_tac a = l in forall_spec, assumption, drule_tac a = l in forall_spec, assumption) apply (cut_tac b = "b N l" and N = N in asprod_ge, simp, simp, (erule conjE)+, simp, thin_tac "x l \ - \", thin_tac "b N l \ \") apply (frule_tac x = "an N" and y = "int N *\<^sub>a b N l" and z = "x l" in aadd_le_mono, simp add:aadd_commute, rule allI, rule impI, cut_tac lessI[of "(0::nat)"], drule_tac a = "Suc 0" in forall_spec, assumption) apply simp done primrec m_max :: "[nat, nat \ nat] \ nat" where m_max_0: "m_max 0 f = f 0" | m_max_Suc: "m_max (Suc n) f = max (m_max n f) (f (Suc n))" (** maximum value of f **) lemma m_maxTr:"\l \ n. (f l) \ m_max n f" apply (induct_tac n) apply simp apply (rule allI, rule impI) apply simp apply (case_tac "l = Suc n", simp) apply (cut_tac m = l and n = "Suc n" in noteq_le_less, assumption+, thin_tac "l \ Suc n", thin_tac "l \ Suc n", frule_tac x = l and n = "Suc n" in less_le_diff, thin_tac "l < Suc n", simp) apply (drule_tac a = l in forall_spec, assumption) apply simp done lemma m_max_gt:"l \ n \ (f l) \ m_max n f" apply (simp add:m_maxTr) done lemma ASum_zero:" (\j \ n. f j \ Z\<^sub>\) \ (\l \ n. f l = 0) \ ASum f n = 0" apply (induct_tac n) apply (rule impI, erule conjE, simp) apply (rule impI) apply (subgoal_tac "(\j\n. f j \ Z\<^sub>\) \ (\l\n. f l = 0)", simp) apply (simp add:aadd_0_l, erule conjE, thin_tac "(\j\n. f j \ Z\<^sub>\) \ (\l\n. f l = 0) \ ASum f n = 0") apply (rule conjI) apply (rule allI, rule impI, drule_tac a = j in forall_spec, simp, assumption+) apply (thin_tac "\j\Suc n. f j \ Z\<^sub>\") apply (rule allI, rule impI, drule_tac a = l in forall_spec, simp+) done lemma eSum_singleTr:"(\j \ n. f j \ Z\<^sub>\) \ (j \ n \ (\l \{h. h \ n} - {j}. f l = 0)) \ ASum f n = f j" apply (induct_tac n) apply (simp, rule impI, (erule conjE)+) apply (case_tac "j \ n") apply simp apply (simp add:aadd_0_r) apply simp apply (simp add:nat_not_le_less[of j]) apply (frule_tac m = n and n = j in Suc_leI) apply (frule_tac m = j and n = "Suc n" in le_antisym, assumption+, simp) apply (cut_tac n = n in ASum_zero [of _ "f"]) apply (subgoal_tac "(\j\n. f j \ Z\<^sub>\) \ (\l\n. f l = 0)") apply (thin_tac "\j\Suc n. f j \ Z\<^sub>\", thin_tac "\l\{h. h \ Suc n} - {Suc n}. f l = 0", simp only:mp) apply (simp add:aadd_0_l) apply (thin_tac "(\j\n. f j \ Z\<^sub>\) \ (\l\n. f l = 0) \ ASum f n = 0") apply (rule conjI, thin_tac "\l\{h. h \ Suc n} - {Suc n}. f l = 0", simp) apply (thin_tac "\j\Suc n. f j \ Z\<^sub>\", simp) done lemma eSum_single:"\\j \ n. f j \ Z\<^sub>\ ; j \ n; \l \ {h. h \ n} - {j}. f l = 0\ \ ASum f n = f j" apply (simp add:eSum_singleTr) done lemma ASum_eqTr:"(\j \ n. f j \ Z\<^sub>\) \ (\j \ n. g j \ Z\<^sub>\) \ (\j \ n. f j = g j) \ ASum f n = ASum g n" apply (induct_tac n) apply (rule impI, simp) apply (rule impI, (erule conjE)+) apply simp done lemma ASum_eq:"\\j \ n. f j \ Z\<^sub>\; \j \ n. g j \ Z\<^sub>\; \j \ n. f j = g j\ \ ASum f n = ASum g n" by (cut_tac ASum_eqTr[of n f g], simp) definition Kronecker_delta :: "[nat, nat] \ ant" ("(\\<^bsub>_ _\<^esub>)" [70,71]70) where "\\<^bsub>i j\<^esub> = (if i = j then 1 else 0)" definition K_gamma :: "[nat, nat] \ int" ("(\\<^bsub>_ _\<^esub>)" [70,71]70) where "\\<^bsub>i j\<^esub> = (if i = j then 0 else 1)" abbreviation TRANSPOS ("(\\<^bsub>_ _\<^esub>)" [90,91]90) where "\\<^bsub>i j\<^esub> == transpos i j" lemma Kdelta_in_Zinf:"\j \ (Suc n); k \ (Suc n)\ \ z *\<^sub>a (\\<^bsub>j k\<^esub>) \ Z\<^sub>\" apply (simp add:Kronecker_delta_def) apply (simp add:z_in_aug_inf Zero_in_aug_inf) apply (simp add:asprod_n_0 Zero_in_aug_inf) done lemma Kdelta_in_Zinf1:"\j \ n; k \ n\ \ \\<^bsub>j k\<^esub> \ Z\<^sub>\" apply (simp add:Kronecker_delta_def) apply (simp add:z_in_aug_inf Zero_in_aug_inf) apply (rule impI) apply (simp only:ant_1[THEN sym], simp del:ant_1 add:z_in_aug_inf) done primrec m_zmax :: "[nat, nat \ int] \ int" where m_zmax_0: "m_zmax 0 f = f 0" | m_zmax_Suc: "m_zmax (Suc n) f = zmax (m_zmax n f) (f (Suc n))" lemma m_zmax_gt_eachTr: "(\j \ n. f j \ Zset) \ (\j \ n. (f j) \ m_zmax n f)" apply (induct_tac n) apply (rule impI, rule allI, rule impI, simp) apply (rule impI) apply simp apply (rule allI, rule impI) apply (case_tac "j = Suc n", simp) apply (simp add:zmax_def) apply (drule_tac m = j and n = "Suc n" in noteq_le_less, assumption, drule_tac x = j and n = "Suc n" in less_le_diff, simp) apply (drule_tac a = j in forall_spec, assumption) apply (simp add:zmax_def) done lemma m_zmax_gt_each:"(\j \ n. f j \ Zset) \ (\j \ n. (f j) \ m_zmax n f)" apply (simp add:m_zmax_gt_eachTr) done lemma n_notin_Nset_pred:" 0 < n \ \ n \ (n - Suc 0)" apply simp done lemma Nset_preTr:"\0 < n; j \ (n - Suc 0)\ \ j \ n" apply simp done lemma Nset_preTr1:"\0 < n; j \ (n - Suc 0)\ \ j \ n" apply simp done lemma transpos_noteqTr:"\0 < n; k \ (n - Suc 0); j \ n; j \ n\ \ j \ (\\<^bsub>j n\<^esub>) k" apply (rule contrapos_pp, simp+) apply (simp add:transpos_def) apply (case_tac "k = j", simp, simp) apply (case_tac "k = n", simp) apply (simp add:n_notin_Nset_pred) done chapter "Elementary properties of a valuation" section "Definition of a valuation" definition valuation :: "[('b, 'm) Ring_scheme, 'b \ ant] \ bool" where "valuation K v \ v \ extensional (carrier K) \ v \ carrier K \ Z\<^sub>\ \ v (\\<^bsub>K\<^esub>) = \ \ (\x\((carrier K) - {\\<^bsub>K\<^esub>}). v x \ \) \ (\x\(carrier K). \y\(carrier K). v (x \\<^sub>r\<^bsub>K\<^esub> y) = (v x) + (v y)) \ (\x\(carrier K). 0 \ (v x) \ 0 \ (v (1\<^sub>r\<^bsub>K\<^esub> \\<^bsub>K\<^esub> x))) \ (\x. x \ carrier K \ (v x) \ \ \ (v x) \ 0)" lemma (in Corps) invf_closed:"x \ carrier K - {\} \ x\<^bsup>\ K\<^esup> \ carrier K" by (cut_tac invf_closed1[of x], simp, assumption) lemma (in Corps) valuation_map:"valuation K v \ v \ carrier K \ Z\<^sub>\" by (simp add:valuation_def) lemma (in Corps) value_in_aug_inf:"\valuation K v; x \ carrier K\ \ v x \ Z\<^sub>\" by (simp add:valuation_def, (erule conjE)+, simp add:funcset_mem) lemma (in Corps) value_of_zero:"valuation K v \ v (\) = \" by (simp add:valuation_def) lemma (in Corps) val_nonzero_noninf:"\valuation K v; x \ carrier K; x \ \\ \ (v x) \ \" by (simp add:valuation_def) lemma (in Corps) value_inf_zero:"\valuation K v; x \ carrier K; v x = \\ \ x = \" by (rule contrapos_pp, simp+, frule val_nonzero_noninf[of v x], assumption+, simp) lemma (in Corps) val_nonzero_z:"\valuation K v; x \ carrier K; x \ \\ \ \z. (v x) = ant z" by (frule value_in_aug_inf[of v x], assumption+, frule val_nonzero_noninf[of v x], assumption+, cut_tac mem_ant[of "v x"], simp add:aug_inf_def) lemma (in Corps) val_nonzero_z_unique:"\valuation K v; x \ carrier K; x \ \\ \ \!z. (v x) = ant z" by (rule ex_ex1I, simp add:val_nonzero_z, simp) lemma (in Corps) value_noninf_nonzero:"\valuation K v; x \ carrier K; v x \ \\ \ x \ \" by (rule contrapos_pp, simp+, simp add:value_of_zero) lemma (in Corps) val1_neq_0:"\valuation K v; x \ carrier K; v x = 1\ \ x \ \" apply (rule contrapos_pp, simp+, simp add:value_of_zero) apply (simp only:ant_1[THEN sym], cut_tac z_neq_inf[THEN not_sym, of 1], simp) done lemma (in Corps) val_Zmin_sym:"\valuation K v; x \ carrier K; y \ carrier K\ \ amin (v x) (v y) = amin (v y ) (v x)" by (simp add:amin_commute) lemma (in Corps) val_t2p:"\valuation K v; x \ carrier K; y \ carrier K\ \ v (x \\<^sub>r y ) = v x + v y" by (simp add:valuation_def) lemma (in Corps) val_axiom4:"\valuation K v; x \ carrier K; 0 \ v x\ \ 0 \ v (1\<^sub>r \ x)" by (simp add:valuation_def) lemma (in Corps) val_axiom5:"valuation K v \ \x. x \ carrier K \ v x \ \ \ v x \ 0" by (simp add:valuation_def) lemma (in Corps) val_field_nonzero:"valuation K v \ carrier K \ {\}" by (rule contrapos_pp, simp+, frule val_axiom5[of v], erule exE, (erule conjE)+, simp add:value_of_zero) lemma (in Corps) val_field_1_neq_0:"valuation K v \ 1\<^sub>r \ \" apply (rule contrapos_pp, simp+) apply (frule val_axiom5[of v]) apply (erule exE, (erule conjE)+) apply (cut_tac field_is_ring, frule_tac t = x in Ring.ring_l_one[THEN sym, of "K"], assumption+, simp add:Ring.ring_times_0_x, simp add:value_of_zero) done lemma (in Corps) value_of_one:"valuation K v \ v (1\<^sub>r) = 0" apply (cut_tac field_is_ring, frule Ring.ring_one[of "K"]) apply (frule val_t2p[of v "1\<^sub>r" "1\<^sub>r"], assumption+, simp add:Ring.ring_l_one, frule val_field_1_neq_0[of v], frule val_nonzero_z[of v "1\<^sub>r"], assumption+, erule exE, simp add:a_zpz) done lemma (in Corps) has_val_one_neq_zero:"valuation K v \ 1\<^sub>r \ \" by (frule value_of_one[of "v"], rule contrapos_pp, simp+, simp add:value_of_zero) lemma (in Corps) val_minus_one:"valuation K v \ v (-\<^sub>a 1\<^sub>r) = 0" apply (cut_tac field_is_ring, frule Ring.ring_one[of "K"], frule Ring.ring_is_ag[of "K"], frule val_field_1_neq_0[of v], frule aGroup.ag_inv_inj[of "K" "1\<^sub>r" "\"], assumption+, simp add:Ring.ring_zero, assumption) apply (frule val_nonzero_z[of v "-\<^sub>a 1\<^sub>r"], rule aGroup.ag_mOp_closed, assumption+, simp add:aGroup.ag_inv_zero, erule exE, frule val_t2p [THEN sym, of v "-\<^sub>a 1\<^sub>r" "-\<^sub>a 1\<^sub>r"]) apply (simp add:aGroup.ag_mOp_closed[of "K" "1\<^sub>r"], simp add:aGroup.ag_mOp_closed[of "K" "1\<^sub>r"], frule Ring.ring_inv1_3[THEN sym, of "K" "1\<^sub>r" "1\<^sub>r"], assumption+, simp add:Ring.ring_l_one, simp add:value_of_one a_zpz) done lemma (in Corps) val_minus_eq:"\valuation K v; x \ carrier K\ \ v (-\<^sub>a x) = v x" apply (cut_tac field_is_ring, simp add:Ring.ring_times_minusl[of K x], subst val_t2p[of v], assumption+, frule Ring.ring_is_ag[of "K"], rule aGroup.ag_mOp_closed, assumption+, simp add:Ring.ring_one, assumption, simp add:val_minus_one, simp add:aadd_0_l) done lemma (in Corps) value_of_inv:"\valuation K v; x \ carrier K; x \ \\ \ v (x\<^bsup>\K\<^esup>) = - (v x)" apply (cut_tac invf_inv[of x], erule conjE, frule val_t2p[of v "x\<^bsup>\K\<^esup>" x], assumption+, simp+, simp add:value_of_one, simp add:a_inv) apply simp done lemma (in Corps) val_exp_ring:"\ valuation K v; x \ carrier K; x \ \\ \ (int n) *\<^sub>a (v x) = v (x^\<^bsup>K n\<^esup>)" apply (cut_tac field_is_ring, induct_tac n, simp add:Ring.ring_r_one, simp add:value_of_one) apply (drule sym, simp) apply (subst val_t2p[of v _ x], assumption+, rule Ring.npClose, assumption+, frule val_nonzero_z[of v x], assumption+, erule exE, simp add:asprod_mult a_zpz, simp add: distrib_right) done text\exponent in a field\ lemma (in Corps) val_exp:"\ valuation K v; x \ carrier K; x \ \\ \ z *\<^sub>a (v x) = v (x\<^bsub>K\<^esub>\<^bsup>z\<^esup>)" apply (simp add:npowf_def) apply (case_tac "0 \ z", simp, frule val_exp_ring [of v x "nat z"], assumption+, simp, simp) apply (simp add:zle, cut_tac invf_closed1[of x], simp, cut_tac val_exp_ring [THEN sym, of v "x\<^bsup>\ K\<^esup>" "nat (- z)"], simp, thin_tac "v (x\<^bsup>\ K\<^esup>^\<^bsup>K (nat (- z))\<^esup>) = (- z) *\<^sub>a v (x\<^bsup>\ K\<^esup>)", erule conjE) apply (subst value_of_inv[of v x], assumption+) apply (frule val_nonzero_z[of v x], assumption+, erule exE, simp, simp add:asprod_mult aminus, simp+) done lemma (in Corps) value_zero_nonzero:"\valuation K v; x \ carrier K; v x = 0\ \ x \ \" by (frule value_noninf_nonzero[of v x], assumption+, simp, assumption) lemma (in Corps) v_ale_diff:"\valuation K v; x \ carrier K; y \ carrier K; x \ \; v x \ v y \ \ 0 \ v(y \\<^sub>r x\<^bsup>\ K\<^esup>)" apply (frule value_in_aug_inf[of v x], simp+, frule value_in_aug_inf[of v y], simp+, frule val_nonzero_z[of v x], assumption+, erule exE) apply (cut_tac invf_closed[of x], simp+, simp add:val_t2p, simp add:value_of_inv[of v "x"], frule_tac x = "ant z" in ale_diff_pos[of _ "v y"], simp add:diff_ant_def) apply simp done lemma (in Corps) amin_le_plusTr:"\valuation K v; x \ carrier K; y \ carrier K; v x \ \; v y \ \; v x \ v y\ \ amin (v x) (v y) \ v ( x \ y)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag, frule value_noninf_nonzero[of v x], assumption+, frule v_ale_diff[of v x y], assumption+, cut_tac invf_closed1[of x], frule Ring.ring_tOp_closed[of K y "x\<^bsup>\ K\<^esup>"], assumption+, simp, frule Ring.ring_one[of "K"], frule aGroup.ag_pOp_closed[of "K" "1\<^sub>r" "y \\<^sub>r x\<^bsup>\ K\<^esup>"], assumption+, frule val_axiom4[of v "y \\<^sub>r ( x\<^bsup>\ K\<^esup>)"], assumption+) apply (frule aadd_le_mono[of "0" "v (1\<^sub>r \ y \\<^sub>r x\<^bsup>\ K\<^esup>)" "v x"], simp add:aadd_0_l, simp add:aadd_commute[of _ "v x"], simp add:val_t2p[THEN sym, of v x], simp add:Ring.ring_distrib1 Ring.ring_r_one, simp add:Ring.ring_tOp_commute[of "K" "x"], simp add:Ring.ring_tOp_assoc, simp add:linvf, simp add:Ring.ring_r_one, cut_tac amin_le_l[of "v x" "v y"], rule ale_trans[of "amin (v x) (v y)" "v x" "v (x \ y)"], assumption+) apply simp done lemma (in Corps) amin_le_plus:"\valuation K v; x \ carrier K; y \ carrier K\ \ (amin (v x) (v y)) \ (v (x \ y))" apply (cut_tac field_is_ring, frule Ring.ring_is_ag) apply (case_tac "v x = \ \ v y = \") apply (erule disjE, simp, frule value_inf_zero[of v x], assumption+, simp add:aGroup.ag_l_zero amin_def, frule value_inf_zero[of v y], assumption+, simp add:aGroup.ag_r_zero amin_def, simp, erule conjE) apply (cut_tac z = "v x" and w = "v y" in ale_linear, erule disjE, simp add:amin_le_plusTr, frule_tac amin_le_plusTr[of v y x], assumption+, simp add:aGroup.ag_pOp_commute amin_commute) done lemma (in Corps) value_less_eq:"\ valuation K v; x \ carrier K; y \ carrier K; (v x) < (v y)\ \ (v x) = (v (x \ y))" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule amin_le_plus[of v x y], assumption+, frule aless_imp_le[of "v x" "v y"], simp add: amin_def) apply (frule amin_le_plus[of v "x \ y" "-\<^sub>a y"], rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, simp add:val_minus_eq, frule aGroup.ag_mOp_closed[of "K" "y"], assumption+, simp add:aGroup.ag_pOp_assoc[of "K" "x" "y"], simp add:aGroup.ag_r_inv1, simp add:aGroup.ag_r_zero, simp add:amin_def) apply (case_tac "\ (v (x \\<^bsub>K\<^esub> y) \ (v y))", simp+) done lemma (in Corps) value_less_eq1:"\valuation K v; x \ carrier K; y \ carrier K; (v x) < (v y)\ \ v x = v (y \ x)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule value_less_eq[of v x y], assumption+) apply (subst aGroup.ag_pOp_commute, assumption+) done lemma (in Corps) val_1px:"\valuation K v; x \ carrier K; 0 \ (v (1\<^sub>r \ x))\ \ 0 \ (v x)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Ring.ring_one[of "K"]) apply (rule contrapos_pp, simp+, case_tac "x = \\<^bsub>K\<^esub>", simp add:aGroup.ag_r_zero, simp add:value_of_zero, simp add: aneg_le[of "0" "v x"], frule value_less_eq[of v x "1\<^sub>r"], assumption+, simp add:value_of_one) apply (drule sym, simp add:aGroup.ag_pOp_commute[of "K" "x"]) done lemma (in Corps) val_1mx:"\valuation K v; x \ carrier K; 0 \ (v (1\<^sub>r \ (-\<^sub>a x)))\ \ 0 \ (v x)" by (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule val_1px[of v "-\<^sub>a x"], simp add:aGroup.ag_mOp_closed, assumption, simp add:val_minus_eq) section "The normal valuation of v" definition Lv :: "[('r, 'm) Ring_scheme , 'r \ ant] \ ant" (* Least nonnegative value *) where "Lv K v = AMin {x. x \ v ` carrier K \ 0 < x}" definition n_val :: "[('r, 'm) Ring_scheme, 'r \ ant] \ ('r \ ant)" where "n_val K v = (\x\ carrier K. (THE l. (l * (Lv K v)) = v x))" (* normal valuation *) definition Pg :: "[('r, 'm) Ring_scheme, 'r \ ant] \ 'r" (* Positive generator *) where "Pg K v = (SOME x. x \ carrier K - {\\<^bsub>K\<^esub>} \ v x = Lv K v)" lemma (in Corps) vals_pos_nonempty:"valuation K v \ {x. x \ v ` carrier K \ 0 < x} \ {}" using val_axiom5[of v] value_noninf_nonzero[of v] value_of_inv[THEN sym, of v] by (auto simp: ex_image_cong_iff) (metis Ring.ring_is_ag aGroup.ag_mOp_closed aGroup.ag_pOp_closed aGroup.ag_r_inv1 f_is_ring zero_lt_inf) lemma (in Corps) vals_pos_LBset:"valuation K v \ {x. x \ v ` carrier K \ 0 < x} \ LBset 1" by (rule subsetI, simp add:LBset_def, erule conjE, rule_tac x = x in gt_a0_ge_1, assumption) lemma (in Corps) Lv_pos:"valuation K v \ 0 < Lv K v" apply (simp add:Lv_def, frule vals_pos_nonempty[of v], frule vals_pos_LBset[of v], simp only:ant_1[THEN sym], frule AMin[of "{x. x \ v ` carrier K \ 0 < x}" "1"], assumption+, erule conjE) apply simp done lemma (in Corps) AMin_z:"valuation K v \ \a. AMin {x. x \ v ` carrier K \ 0 < x} = ant a" apply (frule vals_pos_nonempty[of v], frule vals_pos_LBset[of v], simp only:ant_1[THEN sym], frule AMin[of "{x. x \ v ` carrier K \ 0 < x}" "1"], assumption+, erule conjE) apply (frule val_axiom5[of v], erule exE, (erule conjE)+, cut_tac x = "v x" in aless_linear[of _ "0"], simp, erule disjE, frule_tac x = x in value_noninf_nonzero[of v], assumption+, frule_tac x1 = x in value_of_inv[THEN sym, of v], assumption+) apply (frule_tac x = "v x" in aless_minus[of _ "0"], simp, cut_tac x = x in invf_closed1, simp, erule conjE, frule valuation_map[of v], frule_tac a = "x\<^bsup>\ K\<^esup>" in mem_in_image[of "v" "carrier K" "Z\<^sub>\"], simp) apply (drule_tac a = "v (x\<^bsup>\ K\<^esup>)" in forall_spec, simp, frule_tac x = "x\<^bsup>\ K\<^esup>" in val_nonzero_noninf[of v], thin_tac "v (x\<^bsup>\ K\<^esup>) \ v ` carrier K", thin_tac "{x \ v ` carrier K. 0 < x} \ LBset 1", thin_tac "AMin {x \ v ` carrier K. 0 < x} \ v ` carrier K", thin_tac "0 < AMin {x \ v ` carrier K. 0 < x}", simp, thin_tac "v (x\<^bsup>\ K\<^esup>) \ v ` carrier K", thin_tac "{x \ v ` carrier K. 0 < x} \ LBset 1", thin_tac "AMin {x \ v ` carrier K. 0 < x} \ v ` carrier K", thin_tac "0 < AMin {x \ v ` carrier K. 0 < x}", simp) apply (rule noninf_mem_Z[of "AMin {x \ v ` carrier K. 0 < x}"], frule image_sub[of v "carrier K" "Z\<^sub>\" "carrier K"], rule subset_refl) apply (rule subsetD[of "v ` carrier K" "Z\<^sub>\" "AMin {x \ v ` carrier K. 0 < x}"], assumption+) apply auto by (metis (no_types, lifting) aneg_le aug_inf_noninf_is_z image_eqI value_in_aug_inf z_less_i) lemma (in Corps) Lv_z:"valuation K v \ \z. Lv K v = ant z" by (simp add:Lv_def, rule AMin_z, assumption+) lemma (in Corps) AMin_k:"valuation K v \ \k\ carrier K - {\}. AMin {x. x \ v ` carrier K \ 0 < x} = v k" apply (frule vals_pos_nonempty[of v], frule vals_pos_LBset[of v], simp only:ant_1[THEN sym], frule AMin[of "{x. x \ v ` carrier K \ 0 < x}" "1"], assumption+, erule conjE) apply (thin_tac "\x\{x. x \ v ` carrier K \ 0 < x}. AMin {x. x \ v ` carrier K \ 0 < x} \ x") apply (simp add:image_def, erule conjE, erule bexE, thin_tac "{x. (\xa\carrier K. x = v xa) \ 0 < x} \ LBset 1", thin_tac "\x. (\xa\carrier K. x = v xa) \ 0 < x", subgoal_tac "x \ carrier K - {\}", blast, frule AMin_z[of v], erule exE, simp) apply (simp add:image_def, thin_tac "AMin {x. (\xa\carrier K. x = v xa) \ 0 < x} = ant a", rule contrapos_pp, simp+, frule sym, thin_tac "v (\) = ant a", simp add:value_of_zero) done lemma (in Corps) val_Pg:" valuation K v \ Pg K v \ carrier K - {\} \ v (Pg K v) = Lv K v" apply (frule AMin_k[of v], unfold Lv_def, unfold Pg_def) apply (rule someI2_ex) apply (erule bexE, drule sym, unfold Lv_def, blast) apply simp done lemma (in Corps) amin_generateTr:"valuation K v \ \w\carrier K - {\}. \z. v w = z *\<^sub>a AMin {x. x \ v ` carrier K \ 0 < x}" apply (frule vals_pos_nonempty[of v], frule vals_pos_LBset[of v], simp only:ant_1[THEN sym], frule AMin[of "{x. x \ v ` carrier K \ 0 < x}" "1"], assumption+, frule AMin_z[of v], erule exE, simp, thin_tac "\x. x \ v ` carrier K \ 0 < x", (erule conjE)+, rule ballI, simp, erule conjE, frule_tac x = w in val_nonzero_noninf[of v], assumption+, frule_tac x = w in value_in_aug_inf[of v], assumption+, simp add:aug_inf_def, cut_tac a = "v w" in mem_ant, simp, erule exE, cut_tac a = z and b = a in amod_adiv_equality) apply (case_tac "z mod a = 0", simp add:ant_0 aadd_0_r, blast, thin_tac "{x. x \ v ` carrier K \ 0 < x} \ LBset 1", thin_tac "v w \ \", thin_tac "v w \ - \") apply (frule AMin_k[of v], erule bexE, drule sym, drule sym, drule sym, rotate_tac -1, drule sym) apply (cut_tac z = z in z_in_aug_inf, cut_tac z = "(z div a)" and x = a in asp_z_Z, cut_tac z = "z mod a" in z_in_aug_inf, frule_tac a = "ant z" and b = "(z div a) *\<^sub>a ant a" and c = "ant (z mod a)" in ant_sol, assumption+, subst asprod_mult, simp, assumption, simp, frule_tac x = k and z = "z div a" in val_exp[of v], (erule conjE)+, assumption, simp, simp, thin_tac "(z div a) *\<^sub>a v k = v (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)", erule conjE) apply (frule_tac x = k and n = "z div a" in field_potent_nonzero1, assumption+, frule_tac a = k and n = "z div a" in npowf_mem, assumption, frule_tac x1 = "k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>" in value_of_inv[THEN sym, of v], assumption+, simp add:diff_ant_def, thin_tac "- v (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>) = v ((k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>)", cut_tac x = "k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>" in invf_closed1, simp, simp, erule conjE, frule_tac x1 = w and y1 = "(k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>" in val_t2p[THEN sym, of v], assumption+, simp, cut_tac field_is_ring, thin_tac "v w + v ((k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>) = ant (z mod a)", thin_tac "v (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>) + ant (z mod a) = v w", frule_tac x = w and y = "(k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>" in Ring.ring_tOp_closed[of "K"], assumption+) apply (frule valuation_map[of v], frule_tac a = "w \\<^sub>r (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>" in mem_in_image[of "v" "carrier K" "Z\<^sub>\"], assumption+, simp) apply (thin_tac "AMin {x. x \ v ` carrier K \ 0 < x} = v k", thin_tac "v \ carrier K \ Z\<^sub>\", subgoal_tac "0 < v (w \\<^sub>r (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup> )", drule_tac a = "v (w \\<^sub>r (k\<^bsub>K\<^esub>\<^bsup>(z div a)\<^esup>)\<^bsup>\ K\<^esup>)" in forall_spec, simp add:image_def) -apply (drule sym, simp) -apply (frule_tac b = a and a = z in pos_mod_conj, erule conjE, + apply (drule sym, simp) + +apply (frule_tac b = a and a = z in Divides.pos_mod_conj, erule conjE, simp, simp, - frule_tac b = a and a = z in pos_mod_conj, erule conjE, simp) + frule_tac b = a and a = z in Divides.pos_mod_conj, erule conjE, simp) done lemma (in Corps) val_principalTr1:"\ valuation K v\ \ Lv K v \ v ` (carrier K - {\}) \ (\w\v ` carrier K. \a. w = a * Lv K v) \ 0 < Lv K v" apply (rule conjI, frule val_Pg[of v], erule conjE, simp add:image_def, frule sym, thin_tac "v (Pg K v) = Lv K v", erule conjE, blast) apply (rule conjI, rule ballI, simp add:image_def, erule bexE) apply ( frule_tac x = x in value_in_aug_inf[of v], assumption, frule sym, thin_tac "w = v x", simp add:aug_inf_def, cut_tac a = w in mem_ant, simp, erule disjE, erule exE, frule_tac x = x in value_noninf_nonzero[of v], assumption+, simp, frule amin_generateTr[of v]) apply (drule_tac x = x in bspec, simp, erule exE, frule AMin_z[of v], erule exE, simp add:Lv_def, simp add:asprod_mult, frule sym, thin_tac "za * a = z", simp, subst a_z_z[THEN sym], blast) apply (simp add:Lv_def, frule AMin_z[of v], erule exE, simp, frule Lv_pos[of v], simp add:Lv_def, frule_tac m1 = a in a_i_pos[THEN sym], blast, simp add:Lv_pos) done lemma (in Corps) val_principalTr2:"\valuation K v; c \ v ` (carrier K - {\}) \ (\w\v ` carrier K. \a. w = a * c) \ 0 < c; d \ v ` (carrier K - {\}) \ (\w\v ` carrier K. \a. w = a * d) \ 0 < d\ \ c = d" apply ((erule conjE)+, drule_tac x = d in bspec, simp add:image_def, erule bexE, blast, drule_tac x = c in bspec, simp add:image_def, erule bexE, blast) apply ((erule exE)+, drule sym, simp, simp add:image_def, (erule bexE)+, simp, (erule conjE)+, frule_tac x = x in val_nonzero_z[of v], assumption+, erule exE, frule_tac x = xa in val_nonzero_z[of v], assumption+, erule exE, simp) apply ( subgoal_tac "a \ \ \ a \ -\", subgoal_tac "aa \ \ \ aa \ -\", cut_tac a = a in mem_ant, cut_tac a = aa in mem_ant, simp, (erule exE)+, simp add:a_z_z, thin_tac "c = ant z", frule sym, thin_tac "zb * z = za", simp) apply (subgoal_tac "0 < zb", cut_tac a = zc and b = zb in mult.commute, simp, simp add:pos_zmult_eq_1_iff, rule contrapos_pp, simp+, cut_tac x = 0 and y = zb in less_linear, simp, thin_tac "\ 0 < zb", erule disjE, simp, frule_tac i = 0 and j = z and k = zb in zmult_zless_mono_neg, assumption+, simp add:mult.commute) apply (rule contrapos_pp, simp+, thin_tac "a \ \ \ a \ - \", erule disjE, simp, rotate_tac 5, drule sym, simp, simp, rotate_tac 5, drule sym, simp) apply (rule contrapos_pp, simp+, erule disjE, simp, rotate_tac 4, drule sym, simp, simp, rotate_tac 4, drule sym, simp) done lemma (in Corps) val_principal:"valuation K v \ \!x0. x0 \ v ` (carrier K - {\}) \ (\w \ v ` (carrier K). \(a::ant). w = a * x0) \ 0 < x0" by (rule ex_ex1I, frule val_principalTr1[of v], blast, rule_tac c = x0 and d = y in val_principalTr2[of v], assumption+) lemma (in Corps) n_val_defTr:"\valuation K v; w \ carrier K\ \ \!a. a * Lv K v = v w" apply (rule ex_ex1I, frule AMin_k[of v], frule Lv_pos[of v], simp add:Lv_def, erule bexE, frule_tac x = k in val_nonzero_z[of v], simp, simp, erule exE, simp, (erule conjE)+) apply (case_tac "w = \\<^bsub>K\<^esub>", simp add:value_of_zero, frule_tac m = z in a_i_pos, blast) apply (frule amin_generateTr[of v], drule_tac x = w in bspec, simp, simp) apply ( erule exE, simp add:asprod_mult, subst a_z_z[THEN sym], blast) apply (frule AMin_k[of v]) apply (erule bexE, frule Lv_pos[of v], simp add:Lv_def) apply ( erule conjE, frule_tac x = k in val_nonzero_z[of v], assumption+, erule exE, simp) apply ( case_tac "w = \\<^bsub>K\<^esub>", simp del:a_i_pos add:value_of_zero, subgoal_tac "y = \", simp, rule contrapos_pp, simp+, cut_tac a = a in mem_ant, simp, erule disjE, simp, erule exE, simp add:a_z_z) apply (rule contrapos_pp, simp+, cut_tac a = y in mem_ant, simp, erule disjE, simp, erule exE, simp add:a_z_z, frule_tac x = w in val_nonzero_z[of v], assumption+, erule exE, simp, cut_tac a = a in mem_ant, erule disjE, simp, frule sym, thin_tac "- \ = ant za", simp, erule disjE, erule exE, simp add:a_z_z) apply (cut_tac a = y in mem_ant, erule disjE, simp, rotate_tac 3, drule sym, simp, erule disjE, erule exE, simp add:a_z_z, frule sym, thin_tac "zb * z = za", simp, simp, rotate_tac 3, drule sym, simp, simp, frule sym, thin_tac "\ = ant za", simp) done lemma (in Corps) n_valTr:"\ valuation K v; x \ carrier K\ \ (THE l. (l * (Lv K v)) = v x)*(Lv K v) = v x" by (rule theI', rule n_val_defTr, assumption+) lemma (in Corps) n_val:"\valuation K v; x \ carrier K\ \ (n_val K v x)*(Lv K v) = v x" by (frule n_valTr[of v x], assumption+, simp add:n_val_def) lemma (in Corps) val_pos_n_val_pos:"\valuation K v; x \ carrier K\ \ (0 \ v x) = (0 \ n_val K v x)" apply (frule n_val[of v x], assumption+, drule sym, frule Lv_pos[of v], frule Lv_z[of v], erule exE, simp) apply (frule_tac w = z and x = 0 and y = "n_val K v x" in amult_pos_mono_r, simp add:amult_0_l) done lemma (in Corps) n_val_in_aug_inf:"\valuation K v; x \ carrier K\ \ n_val K v x \ Z\<^sub>\" apply (cut_tac field_is_ring, frule Ring.ring_zero[of "K"], frule Lv_pos[of v], frule Lv_z[of v], erule exE, simp add:aug_inf_def) apply (rule contrapos_pp, simp+) apply (case_tac "x = \\<^bsub>K\<^esub>", simp, frule n_val[of v "\"], simp add:value_of_zero, simp add:value_of_zero) apply (frule n_val[of v x], simp, frule val_nonzero_z[of v x], assumption+, erule exE, simp, rotate_tac -2, drule sym, simp) done lemma (in Corps) n_val_0:"\valuation K v; x \ carrier K; v x = 0\ \ n_val K v x = 0" by (frule Lv_z[of v], erule exE, frule Lv_pos[of v], frule n_val[of v x], simp, simp, rule_tac z = z and a = "n_val K v x" in a_a_z_0, assumption+) lemma (in Corps) value_n0_n_val_n0:"\valuation K v; x \ carrier K; v x \ 0\ \ n_val K v x \ 0" apply (frule n_val[of v x], rule contrapos_pp, simp+, frule Lv_z[of v], erule exE, simp, simp only:ant_0[THEN sym]) apply (rule contrapos_pp, simp+, simp add:a_z_z) done lemma (in Corps) val_0_n_val_0:"\valuation K v; x \ carrier K\ \ (v x = 0) = (n_val K v x = 0)" apply (rule iffI, simp add:n_val_0) apply (rule contrapos_pp, simp+, frule value_n0_n_val_n0[of v x], assumption+) apply simp done lemma (in Corps) val_noninf_n_val_noninf:"\valuation K v; x \ carrier K\ \ (v x \ \) = (n_val K v x \ \)" by (frule Lv_z[of v], erule exE, frule Lv_pos[of v], simp, frule n_val[THEN sym, of v x],simp, simp, thin_tac "v x = n_val K v x * ant z", rule iffI, rule contrapos_pp, simp+, cut_tac mem_ant[of "n_val K v x"], erule disjE, simp, erule disjE, erule exE, simp add:a_z_z, simp, simp) lemma (in Corps) val_inf_n_val_inf:"\valuation K v; x \ carrier K\ \ (v x = \) = (n_val K v x = \)" by (cut_tac val_noninf_n_val_noninf[of v x], simp, assumption+) lemma (in Corps) val_eq_n_val_eq:"\valuation K v; x \ carrier K; y \ carrier K\ \ (v x = v y) = (n_val K v x = n_val K v y)" apply (subst n_val[THEN sym, of v x], assumption+, subst n_val[THEN sym, of v y], assumption+, frule Lv_pos[of v], frule Lv_z[of v], erule exE, simp, frule_tac s = z in zless_neq[THEN not_sym, of "0"]) apply (rule iffI) apply (rule_tac z = z in amult_eq_eq_r[of _ "n_val K v x" "n_val K v y"], assumption+) apply simp done lemma (in Corps) val_poss_n_val_poss:"\valuation K v; x \ carrier K\ \ (0 < v x) = (0 < n_val K v x)" apply (simp add:less_le, frule val_pos_n_val_pos[of v x], assumption+, rule iffI, erule conjE, simp, simp add:value_n0_n_val_n0[of v x]) apply (drule sym, erule conjE, simp, frule_tac val_0_n_val_0[THEN sym, of v x], assumption+, simp) done lemma (in Corps) n_val_Pg:"valuation K v \ n_val K v (Pg K v) = 1" apply (frule val_Pg[of v], simp, (erule conjE)+, frule n_val[of v "Pg K v"], simp, frule Lv_z[of v], erule exE, simp, frule Lv_pos[of v], simp, frule_tac i = 0 and j = z in zless_neq) apply (rotate_tac -1, frule not_sym, thin_tac "0 \ z", subgoal_tac "n_val K v (Pg K v) * ant z = 1 * ant z", rule_tac z = z in adiv_eq[of _ "n_val K v (Pg K v)" "1"], assumption+, simp add:amult_one_l) done lemma (in Corps) n_val_valuationTr1:"valuation K v \ \x\carrier K. n_val K v x \ Z\<^sub>\" by (rule ballI, frule n_val[of v], assumption, frule_tac x = x in value_in_aug_inf[of v], assumption, frule Lv_pos[of v], simp add:aug_inf_def, frule Lv_z[of v], erule exE, simp, rule contrapos_pp, simp+) lemma (in Corps) n_val_t2p:"\valuation K v; x \ carrier K; y \ carrier K\ \ n_val K v (x \\<^sub>r y) = n_val K v x + (n_val K v y)" apply (cut_tac field_is_ring, frule Ring.ring_tOp_closed[of K x y], assumption+, frule n_val[of v "x \\<^sub>r y"], assumption+, frule Lv_pos[of "v"], simp add:val_t2p, frule n_val[THEN sym, of v x], assumption+, frule n_val[THEN sym, of v y], assumption+, simp, frule Lv_z[of v], erule exE, simp) apply (subgoal_tac "ant z \ 0") apply (frule_tac z1 = z in amult_distrib1[THEN sym, of _ "n_val K v x" "n_val K v y"], simp, thin_tac "n_val K v x * ant z + n_val K v y * ant z = (n_val K v x + n_val K v y) * ant z", rule_tac z = z and a = "n_val K v (x \\<^sub>r y)" and b = "n_val K v x + n_val K v y" in adiv_eq, simp, assumption+, simp) done lemma (in Corps) n_val_valuationTr2:"\ valuation K v; x \ carrier K; y \ carrier K\ \ amin (n_val K v x) (n_val K v y) \ (n_val K v ( x \ y))" apply (frule n_val[THEN sym, of v x], assumption+, frule n_val[THEN sym, of v y], assumption+, frule n_val[THEN sym, of v "x \ y"], cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], rule aGroup.ag_pOp_closed, assumption+) apply (frule amin_le_plus[of v x y], assumption+, simp, simp add:amult_commute[of _ "Lv K v"], frule Lv_z[of v], erule exE, simp, frule Lv_pos[of v], simp, simp add:amin_amult_pos, simp add:amult_pos_mono_l) done lemma (in Corps) n_val_valuation:"valuation K v \ valuation K (n_val K v)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag) apply (frule Lv_z[of v], erule exE, frule Lv_pos[of v], simp, subst valuation_def) apply (rule conjI, simp add:n_val_def restrict_def extensional_def) apply (rule conjI, simp add:n_val_valuationTr1) apply (rule conjI, frule n_val[of v \], simp add:Ring.ring_zero, frule Lv_z[of v], erule exE, frule Lv_pos[of v], cut_tac mem_ant[of "n_val K v (\)"], erule disjE, simp add:value_of_zero, erule disjE, erule exE, simp add:a_z_z value_of_zero, assumption+) apply (rule conjI, rule ballI, frule_tac x = x in val_nonzero_noninf[of v], simp+, simp add:val_noninf_n_val_noninf) apply (rule conjI, (rule ballI)+, simp add:n_val_t2p, rule conjI, rule ballI, rule impI, frule Lv_z[of v], erule exE, frule Lv_pos[of v], simp, frule_tac x = x in n_val[of v], simp, frule_tac w1 = z and x1 = 0 and y1 = "n_val K v x" in amult_pos_mono_r[THEN sym], simp add:amult_0_l, frule_tac x = x in val_axiom4[of v], assumption+, frule_tac x1 = "1\<^sub>r \ x" in n_val[THEN sym, of v], frule Ring.ring_is_ag[of "K"], rule aGroup.ag_pOp_closed, assumption+, simp add:Ring.ring_one, assumption, frule_tac w = z and x = 0 and y = "n_val K v (1\<^sub>r \ x)" in amult_pos_mono_r, simp add:amult_0_l) apply (frule val_axiom5[of v], erule exE, (erule conjE)+, frule_tac x = x in value_n0_n_val_n0[of v], assumption+, frule_tac x = x in val_noninf_n_val_noninf, simp, blast) done lemma (in Corps) n_val_le_val:"\valuation K v; x \ carrier K; 0 \ (v x)\ \ (n_val K v x) \(v x)" by (subst n_val[THEN sym, of v x], assumption+, frule Lv_pos[of v], simp add:val_pos_n_val_pos[of v x], frule Lv_z[of v], erule exE, cut_tac b = z and x = "n_val K v x" in amult_pos, simp+, simp add:asprod_amult, simp add:amult_commute) lemma (in Corps) n_val_surj:"valuation K v \ \x\ carrier K. n_val K v x = 1" apply (frule Lv_z[of v], erule exE, frule Lv_pos[of v], frule AMin_k[of v], erule bexE, frule_tac x = k in n_val[of v], simp, simp add:Lv_def) apply (subgoal_tac "n_val K v k * ant z = 1 * ant z", subgoal_tac "z \ 0", frule_tac z = z and a = "n_val K v k" and b = 1 in amult_eq_eq_r, assumption, blast, simp, simp add:amult_one_l) done lemma (in Corps) n_value_in_aug_inf:"\valuation K v; x \ carrier K\ \ n_val K v x \ Z\<^sub>\" by (frule n_val[of v x], assumption, simp add:aug_inf_def, rule contrapos_pp, simp+, frule Lv_pos[of v], frule Lv_z[of v], erule exE, simp, frule value_in_aug_inf[of v x], assumption+, simp add:aug_inf_def) lemma (in Corps) val_surj_n_valTr:"\valuation K v; \x \ carrier K. v x = 1\ \ Lv K v = 1" apply (erule bexE, frule_tac x = x in n_val[of v], simp, frule Lv_pos[of v]) apply (frule_tac w = "Lv K v" and x = "n_val K v x" in amult_1_both) apply simp+ done lemma (in Corps) val_surj_n_val:"\valuation K v; \x \ carrier K. v x = 1\ \ (n_val K v) = v" apply (rule funcset_eq[of _ "carrier K"], simp add:n_val_def restrict_def extensional_def, simp add:valuation_def) apply (rule ballI, frule val_surj_n_valTr[of v], assumption+, frule_tac x = x in n_val[of v], assumption+, simp add:amult_one_r) done lemma (in Corps) n_val_n_val:"valuation K v \ n_val K (n_val K v) = n_val K v" by (frule n_val_valuation[of v], frule n_val_surj[of v], simp add:val_surj_n_val) lemma nnonzero_annonzero:"0 < N \ an N \ 0" apply (simp only:an_0[THEN sym]) apply (subst aneq_natneq, simp) done section "Valuation ring" definition Vr :: "[('r, 'm) Ring_scheme, 'r \ ant] \ ('r, 'm) Ring_scheme" where "Vr K v = Sr K ({x. x \ carrier K \ 0 \ (v x)})" definition vp :: "[('r, 'm) Ring_scheme, 'r \ ant] \ 'r set" where "vp K v = {x. x \ carrier (Vr K v) \ 0 < (v x)}" definition r_apow :: "[('r, 'm) Ring_scheme, 'r set, ant] \ 'r set" where "r_apow R I a = (if a = \ then {\\<^bsub>R\<^esub>} else (if a = 0 then carrier R else I\<^bsup>\R (na a)\<^esup>))" (** 0 \ a and a \ -\ **) abbreviation RAPOW ("(3_\<^bsup> _ _\<^esup>)" [62,62,63]62) where "I\<^bsup>R a\<^esup> == r_apow R I a" lemma (in Ring) ring_pow_apow:"ideal R I \ I\<^bsup>\R n\<^esup> = I\<^bsup>R (an n)\<^esup>" apply (simp add:r_apow_def) apply (case_tac "n = 0", simp) apply (simp add:nnonzero_annonzero) apply (simp add:an_neq_inf na_an) done lemma (in Ring) r_apow_Suc:"ideal R I \ I\<^bsup>R (an (Suc 0))\<^esup> = I" apply (cut_tac an_1, simp add:r_apow_def) apply (simp add:a0_neq_1[THEN not_sym]) apply (simp only:ant_1[THEN sym]) apply (simp del:ant_1 add:z_neq_inf[of 1, THEN not_sym]) apply (simp add:na_1) apply (simp add:idealprod_whole_r) done lemma (in Ring) apow_ring_pow:"ideal R I \ I\<^bsup>\R n\<^esup> = I\<^bsup>R (an n)\<^esup>" apply (simp add:r_apow_def) apply (case_tac "n = 0", simp add:an_0) apply (simp add: aless_nat_less[THEN sym], cut_tac an_neq_inf[of n], simp add: less_le[of 0 "an n"] na_an) done lemma (in Corps) Vr_ring:"valuation K v \ Ring (Vr K v)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], simp add:Vr_def, rule Ring.Sr_ring, assumption+) apply (simp add:sr_def) apply (intro conjI subsetI) apply (simp_all add: value_of_one Ring.ring_one[of "K"]) apply ((rule allI, rule impI)+, (erule conjE)+, rule conjI, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (frule_tac x = x and y = "-\<^sub>a y" in amin_le_plus[of v], assumption+, rule aGroup.ag_mOp_closed, assumption+, simp add:val_minus_eq[of v]) apply ( frule_tac z = 0 and x = "v x" and y = "v y" in amin_ge1, assumption+, frule_tac i = 0 and j = "amin (v x) (v y)" and k = "v (x \ -\<^sub>a y)" in ale_trans, assumption+, simp) by (simp add: Ring.ring_tOp_closed aadd_two_pos val_t2p) lemma (in Corps) val_pos_mem_Vr:"\valuation K v; x \ carrier K\ \ (0 \ (v x)) = (x \ carrier (Vr K v))" by (rule iffI, (simp add:Vr_def Sr_def)+) lemma (in Corps) val_poss_mem_Vr:"\valuation K v; x \ carrier K; 0 < (v x)\ \ x \ carrier (Vr K v)" by (frule aless_imp_le[of "0" "v x"], simp add:val_pos_mem_Vr) lemma (in Corps) Vr_one:"valuation K v \ 1\<^sub>r\<^bsub>K\<^esub> \ carrier (Vr K v)" by (cut_tac field_is_ring, frule Ring.ring_one[of "K"], frule val_pos_mem_Vr[of v "1\<^sub>r"], assumption+, simp add:value_of_one) lemma (in Corps) Vr_mem_f_mem:"\valuation K v; x \ carrier (Vr K v)\ \ x \ carrier K" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_0_f_0:"valuation K v \ \\<^bsub>Vr K v\<^esub> = \" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_1_f_1:"valuation K v \ 1\<^sub>r\<^bsub>(Vr K v)\<^esub> = 1\<^sub>r" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_pOp_f_pOp:"\valuation K v; x \ carrier (Vr K v); y \ carrier (Vr K v)\ \ x \\<^bsub>Vr K v\<^esub> y = x \ y" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_mOp_f_mOp:"\valuation K v; x \ carrier (Vr K v)\ \ -\<^sub>a\<^bsub>(Vr K v)\<^esub> x = -\<^sub>a x" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_tOp_f_tOp:"\valuation K v; x \ carrier (Vr K v); y \ carrier(Vr K v)\ \ x \\<^sub>r\<^bsub>(Vr K v)\<^esub> y = x \\<^sub>r y" by (simp add:Vr_def Sr_def) lemma (in Corps) Vr_pOp_le:"\valuation K v; x \ carrier K; y \ carrier (Vr K v)\ \ v x \ (v x + (v y))" apply (frule val_pos_mem_Vr[THEN sym, of v y], simp add:Vr_mem_f_mem, simp, frule aadd_pos_le[of "v y" "v x"], simp add:aadd_commute) done lemma (in Corps) Vr_integral:"valuation K v \ Idomain (Vr K v)" apply (simp add:Idomain_def, simp add:Vr_ring, simp add:Idomain_axioms_def, rule allI, rule impI, rule allI, (rule impI)+, simp add:Vr_tOp_f_tOp, simp add:Vr_0_f_0) apply (rule contrapos_pp, simp+, erule conjE, cut_tac field_is_idom, frule_tac x = a in Vr_mem_f_mem[of v], assumption, frule_tac x = b in Vr_mem_f_mem[of v], assumption, frule_tac x = a and y = b in Idomain.idom_tOp_nonzeros[of "K"], assumption+, simp) done lemma (in Corps) Vr_exp_mem:"\valuation K v; x \ carrier (Vr K v)\ \ x^\<^bsup>K n\<^esup> \ carrier (Vr K v)" by (frule Vr_ring[of v], induct_tac n, simp add:Vr_one, simp add:Vr_tOp_f_tOp[THEN sym, of v], simp add:Ring.ring_tOp_closed) lemma (in Corps) Vr_exp_f_exp:"\valuation K v; x \ carrier (Vr K v)\ \ x^\<^bsup>(Vr K v) n\<^esup> = x^\<^bsup>K n\<^esup>" apply (induct_tac n, simp, simp add:Vr_1_f_1, simp, thin_tac "x^\<^bsup>(Vr K v) n\<^esup> = x^\<^bsup>K n\<^esup>") apply (rule Vr_tOp_f_tOp, assumption+, simp add:Vr_exp_mem, assumption) done lemma (in Corps) Vr_potent_nonzero:"\valuation K v; x \ carrier (Vr K v) - {\\<^bsub>Vr K v\<^esub>}\ \ x^\<^bsup>K n\<^esup> \ \\<^bsub>Vr K v\<^esub>" apply (frule Vr_mem_f_mem[of v x], simp, simp add:Vr_0_f_0, erule conjE) apply (frule Vr_mem_f_mem[of v x], assumption+, simp add:field_potent_nonzero) done lemma (in Corps) elem_0_val_if:"\valuation K v; x \ carrier K; v x = 0\ \ x \ carrier (Vr K v) \ x\<^bsup>\ K\<^esup> \ carrier (Vr K v)" apply (frule val_pos_mem_Vr[of v x], assumption, simp) apply (frule value_zero_nonzero[of "v" "x"], simp add:Vr_mem_f_mem, simp) apply (frule value_of_inv[of v x], assumption+, simp, subst val_pos_mem_Vr[THEN sym, of v "x\<^bsup>\K\<^esup>"], assumption+, cut_tac invf_closed[of x], simp+) done lemma (in Corps) elem0val:"\valuation K v; x \ carrier K; x \ \\ \ (v x = 0) = ( x \ carrier (Vr K v) \ x\<^bsup>\ K\<^esup> \ carrier (Vr K v))" apply (rule iffI, rule elem_0_val_if[of v], assumption+, erule conjE) apply (simp add:val_pos_mem_Vr[THEN sym, of v x], frule Vr_mem_f_mem[of v "x\<^bsup>\K\<^esup>"], assumption+, simp add:val_pos_mem_Vr[THEN sym, of v "x\<^bsup>\K\<^esup>"], simp add:value_of_inv, frule ale_minus[of "0" "- v x"], simp add:a_minus_minus) done lemma (in Corps) ideal_inc_elem0val_whole:"\ valuation K v; x \ carrier K; v x = 0; ideal (Vr K v) I; x \ I\ \ I = carrier (Vr K v)" apply (frule elem_0_val_if[of v x], assumption+, erule conjE, frule value_zero_nonzero[of v x], assumption+, frule Vr_ring[of v], frule_tac I = I and x = x and r = "x\<^bsup>\K\<^esup>" in Ring.ideal_ring_multiple[of "Vr K v"], assumption+, cut_tac invf_closed1[of x], simp+, (erule conjE)+) apply (simp add:Vr_tOp_f_tOp, cut_tac invf_inv[of x], simp+, simp add: Vr_1_f_1[THEN sym, of v], simp add:Ring.ideal_inc_one, simp+) done lemma (in Corps) vp_mem_Vr_mem:"\valuation K v; x \ (vp K v)\ \ x \ carrier (Vr K v)" by (rule val_poss_mem_Vr[of v x], assumption+, (simp add:vp_def Vr_def Sr_def)+) lemma (in Corps) vp_mem_val_poss:"\ valuation K v; x \ carrier K\ \ (x \ vp K v) = (0 < (v x))" by (simp add:vp_def, simp add:Vr_def Sr_def less_ant_def) lemma (in Corps) Pg_in_Vr:"valuation K v \ Pg K v \ carrier (Vr K v)" by (frule val_Pg[of v], erule conjE, frule Lv_pos[of v], drule sym, simp, erule conjE, simp add:val_poss_mem_Vr) lemma (in Corps) vp_ideal:"valuation K v \ ideal (Vr K v) (vp K v)" apply (cut_tac field_is_ring, frule Vr_ring[of v], rule Ring.ideal_condition1, assumption+, rule subsetI, simp add:vp_mem_Vr_mem, simp add:vp_def) apply (frule val_Pg[of v], frule Lv_pos[of v], simp, (erule conjE)+, drule sym, simp, frule val_poss_mem_Vr[of v "Pg K v"], assumption+, blast) apply ((rule ballI)+, frule_tac x = x in vp_mem_Vr_mem[of v], assumption) apply ( frule_tac x = y in vp_mem_Vr_mem[of v], assumption, simp add:vp_def, frule Ring.ring_is_ag[of "Vr K v"], frule_tac x = x and y = y in aGroup.ag_pOp_closed, assumption+, simp) apply (simp add:Vr_pOp_f_pOp, cut_tac x = "v x" and y = "v y" in amin_le_l, frule_tac x = x and y = y in amin_le_plus, (simp add:Vr_mem_f_mem)+, (frule_tac z = 0 and x = "v x" and y = "v y" in amin_gt, assumption+), rule_tac x = 0 and y = "amin (v x) (v y)" and z = "v (x \ y)" in less_le_trans, assumption+) apply ((rule ballI)+, frule_tac x1 = r in val_pos_mem_Vr[THEN sym, of v], simp add:Vr_mem_f_mem, simp, frule_tac x = x in vp_mem_Vr_mem[of v], simp add:Vr_pOp_f_pOp, simp add:vp_def, simp add:Ring.ring_tOp_closed, simp add:Vr_tOp_f_tOp) apply (frule_tac x = r in Vr_mem_f_mem[of v], assumption+, frule_tac x = x in Vr_mem_f_mem[of v], assumption+, simp add:val_t2p, simp add:aadd_pos_poss) done lemma (in Corps) vp_not_whole:"valuation K v \ (vp K v) \ carrier (Vr K v)" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Vr_ring[of v]) apply (rule contrapos_pp, simp+, drule sym, frule Ring.ring_one[of "Vr K v"], simp, simp add:Vr_1_f_1, frule Ring.ring_one[of "K"]) apply (simp only:vp_mem_val_poss[of v "1\<^sub>r"], simp add:value_of_one) done lemma (in Ring) elem_out_ideal_nonzero:"\ideal R I; x \ carrier R; x \ I\ \ x \ \\<^bsub>R\<^esub>" by (rule contrapos_pp, simp+, frule ideal_zero[of I], simp) lemma (in Corps) vp_prime:"valuation K v \ prime_ideal (Vr K v) (vp K v)" apply (simp add:prime_ideal_def, simp add:vp_ideal) apply (rule conjI) (** if the unit is contained in (vp K v), then (vp K v) is the whole ideal, this contradicts vp_not_whole **) apply (rule contrapos_pp, simp+, frule Vr_ring[of v], frule vp_ideal[of v], frule Ring.ideal_inc_one[of "Vr K v" "vp K v"], assumption+, simp add:vp_not_whole[of v]) (* done*) (** if x \\<^bsub>(Vr K v)\<^esub> y is in (vp K v), then 0 < v (x \\<^sub>K y). We have 0 \ (v x) and 0 \ (v y), because x and y are elements of Vr K v. Since v (x \\<^sub>K y) = (v x) + (v y), we have 0 < (v x) or 0 < (v y). To obtain the final conclusion, we suppose \ (x \ vp K v \ y \ vp K v) then, we have (v x) = 0 and (v y) = 0. Frome this, we have v (x \\<^sub>K y) = 0. Contradiction. *) apply ((rule ballI)+, rule impI, rule contrapos_pp, simp+, (erule conjE)+, frule Vr_ring[of v]) apply ( frule_tac x = x in Vr_mem_f_mem[of v], assumption) apply ( frule_tac x = y in Vr_mem_f_mem[of v], assumption) apply ( frule vp_ideal[of v], frule_tac x = x in Ring.elem_out_ideal_nonzero[of "Vr K v" "vp K v"], assumption+) apply ( frule_tac x = y in Ring.elem_out_ideal_nonzero[of "Vr K v" "vp K v"], assumption+, simp add:Vr_0_f_0, simp add:Vr_tOp_f_tOp) apply ( frule_tac x = "x \\<^sub>r y" in vp_mem_val_poss[of v], cut_tac field_is_ring, simp add:Ring.ring_tOp_closed, simp) apply (cut_tac field_is_ring, frule_tac x = x and y = y in Ring.ring_tOp_closed, assumption+, simp add:Ring.ring_tOp_closed[of "Vr K v"], simp add:vp_def, simp add:aneg_less, frule_tac x1 = x in val_pos_mem_Vr[THEN sym, of v], assumption+, frule_tac x1 = y in val_pos_mem_Vr[THEN sym, of v], assumption+, frule_tac P = "x \ carrier (Vr K v)" and Q = "0 \ v x" in eq_prop, assumption, frule_tac P = "y \ carrier (Vr K v)" and Q = "0 \ v y" in eq_prop, assumption, frule_tac x = "v x" and y = 0 in ale_antisym, assumption+, frule_tac x = "v y" and y = 0 in ale_antisym, assumption+, simp add:val_t2p aadd_0_l) done lemma (in Corps) vp_pow_ideal:"valuation K v \ ideal (Vr K v) ((vp K v)\<^bsup>\(Vr K v) n\<^esup>)" by (frule Vr_ring[of v], frule vp_ideal[of v], simp add:Ring.ideal_pow_ideal) lemma (in Corps) vp_apow_ideal:"\valuation K v; 0 \ n\ \ ideal (Vr K v) ((vp K v)\<^bsup>(Vr K v) n\<^esup>)" apply (frule Vr_ring[of v]) apply (case_tac "n = 0", simp add:r_apow_def, simp add:Ring.whole_ideal[of "Vr K v"]) apply (case_tac "n = \", simp add:r_apow_def, simp add:Ring.zero_ideal) apply (simp add:r_apow_def, simp add:vp_pow_ideal) done lemma (in Corps) mem_vp_apow_mem_Vr:"\valuation K v; 0 \ N; x \ vp K v \<^bsup>(Vr K v) N\<^esup>\ \ x \ carrier (Vr K v)" by (frule Vr_ring[of v], frule vp_apow_ideal[of v N], assumption, simp add:Ring.ideal_subset) lemma (in Corps) elem_out_vp_unit:"\valuation K v; x \ carrier (Vr K v); x \ vp K v\ \ v x = 0" by (metis Vr_mem_f_mem ale_antisym aneg_le val_pos_mem_Vr vp_mem_val_poss) lemma (in Corps) vp_maximal:"valuation K v \ maximal_ideal (Vr K v) (vp K v)" apply (frule Vr_ring[of v], simp add:maximal_ideal_def, simp add:vp_ideal) (** we know that vp is not a whole ideal, and so vp does not include 1 **) apply (frule vp_not_whole[of v], rule conjI, rule contrapos_pp, simp+, frule vp_ideal[of v], frule Ring.ideal_inc_one[of "Vr K v" "vp K v"], assumption+) apply simp (** onemore condition of maximal ideal **) apply (rule equalityI, rule subsetI, simp, erule conjE, case_tac "x = vp K v", simp, simp, rename_tac X) (** show exists a unit contained in X **) apply (frule_tac A = X in sets_not_eq[of _ "vp K v"], assumption+, erule bexE, frule_tac I = X and h = a in Ring.ideal_subset[of "Vr K v"], assumption+, frule_tac x = a in elem_out_vp_unit[of v], assumption+) (** since v a = 0, we see a is a unit **) apply (frule_tac x = a and I = X in ideal_inc_elem0val_whole [of v], simp add:Vr_mem_f_mem, assumption+) apply (rule subsetI, simp, erule disjE, simp add:prime_ideal_def, simp add:vp_ideal, simp add:Ring.whole_ideal, rule subsetI, simp add:vp_mem_Vr_mem) done lemma (in Corps) ideal_sub_vp:"\ valuation K v; ideal (Vr K v) I; I \ carrier (Vr K v)\ \ I \ (vp K v)" apply (frule Vr_ring[of v], rule contrapos_pp, simp+) apply (simp add:subset_eq, erule bexE) apply (frule_tac h = x in Ring.ideal_subset[of "Vr K v" I], assumption+, frule_tac x = x in elem_out_vp_unit[of v], assumption+, frule_tac x = x in ideal_inc_elem0val_whole[of v _ I], simp add:Vr_mem_f_mem, assumption+, simp) done lemma (in Corps) Vr_local:"\valuation K v; maximal_ideal (Vr K v) I\ \ (vp K v) = I" apply (frule Vr_ring[of v], frule ideal_sub_vp[of v I], simp add:Ring.maximal_ideal_ideal) apply (simp add:maximal_ideal_def, frule conjunct2, fold maximal_ideal_def, frule conjunct1, rule Ring.proper_ideal, assumption+,simp add:maximal_ideal_def, assumption) apply (rule equalityI) prefer 2 apply assumption apply (rule contrapos_pp, simp+, frule sets_not_eq[of "vp K v" I], assumption+, erule bexE) apply (frule_tac x = a in vp_mem_Vr_mem[of v], frule Ring.maximal_ideal_ideal[of "Vr K v" "I"], assumption, frule_tac x = a in Ring.elem_out_ideal_nonzero[of "Vr K v" "I"], assumption+, frule vp_ideal[of v], rule Ring.ideal_subset[of "Vr K v" "vp K v"], assumption+) apply (frule_tac a = a in Ring.principal_ideal[of "Vr K v"], assumption+, frule Ring.maximal_ideal_ideal[of "Vr K v" I], assumption+, frule_tac ?I2.0 = "Vr K v \\<^sub>p a"in Ring.sum_ideals[of "Vr K v" "I"], simp add:Ring.maximal_ideal_ideal, assumption, frule_tac ?I2.0 = "Vr K v \\<^sub>p a"in Ring.sum_ideals_la1[of "Vr K v" "I"], assumption+, frule_tac ?I2.0 = "Vr K v \\<^sub>p a"in Ring.sum_ideals_la2[of "Vr K v" "I"], assumption+, frule_tac a = a in Ring.a_in_principal[of "Vr K v"], assumption+, frule_tac A = "Vr K v \\<^sub>p a" and B = "I \\<^bsub>(Vr K v)\<^esub> (Vr K v \\<^sub>p a)" and c = a in subsetD, assumption+) thm Ring.sum_ideals_cont[of "Vr K v" "vp K v" I ] apply (frule_tac B = "Vr K v \\<^sub>p a" in Ring.sum_ideals_cont[of "Vr K v" "vp K v" I], simp add:vp_ideal, assumption) apply (frule_tac a = a in Ring.ideal_cont_Rxa[of "Vr K v" "vp K v"], simp add:vp_ideal, assumption+) apply (simp add:maximal_ideal_def, (erule conjE)+, subgoal_tac "I \\<^bsub>(Vr K v)\<^esub> (Vr K v \\<^sub>p a) \ {J. ideal (Vr K v) J \ I \ J}", simp, thin_tac "{J. ideal (Vr K v) J \ I \ J} = {I, carrier (Vr K v)}") apply (erule disjE, simp) apply (cut_tac A = "carrier (Vr K v)" and B = "I \\<^bsub>Vr K v\<^esub> Vr K v \\<^sub>p a" and C = "vp K v" in subset_trans, simp, assumption, frule Ring.ideal_subset1[of "Vr K v" "vp K v"], simp add:vp_ideal, frule equalityI[of "vp K v" "carrier (Vr K v)"], assumption+, frule vp_not_whole[of v], simp) apply blast done lemma (in Corps) v_residue_field:"valuation K v \ Corps ((Vr K v) /\<^sub>r (vp K v))" by (frule Vr_ring[of v], rule Ring.residue_field_cd [of "Vr K v" "vp K v"], assumption+, simp add:vp_maximal) lemma (in Corps) Vr_n_val_Vr:"valuation K v \ carrier (Vr K v) = carrier (Vr K (n_val K v))" by (simp add:Vr_def Sr_def, rule equalityI, (rule subsetI, simp, erule conjE, simp add:val_pos_n_val_pos), (rule subsetI, simp, erule conjE, simp add:val_pos_n_val_pos[THEN sym])) section "Ideals in a valuation ring" lemma (in Corps) Vr_has_poss_elem:"valuation K v \ \x\carrier (Vr K v) - {\\<^bsub>Vr K v\<^esub>}. 0 < v x" apply (frule val_Pg[of v], erule conjE, frule Lv_pos[of v], drule sym, subst Vr_0_f_0, assumption+) apply (frule aeq_ale[of "Lv K v" "v (Pg K v)"], frule aless_le_trans[of "0" "Lv K v" "v (Pg K v)"], assumption+, frule val_poss_mem_Vr[of v "Pg K v"], simp, assumption, blast) done lemma (in Corps) vp_nonzero:"valuation K v \ vp K v \ {\\<^bsub>Vr K v\<^esub>}" apply (frule Vr_has_poss_elem[of v], erule bexE, simp, erule conjE, frule_tac x1 = x in vp_mem_val_poss[THEN sym, of v], simp add:Vr_mem_f_mem, simp, rule contrapos_pp, simp+) done lemma (in Corps) field_frac_mul:"\x \ carrier K; y \ carrier K; y \ \\ \ x = (x \\<^sub>r (y\<^bsup>\K\<^esup>)) \\<^sub>r y" apply (cut_tac invf_closed[of y], cut_tac field_is_ring, simp add:Ring.ring_tOp_assoc, subst linvf[of y], simp, simp add:Ring.ring_r_one[of K], simp) done lemma (in Corps) elems_le_val:"\valuation K v; x \ carrier K; y \ carrier K; x \ \; v x \ (v y)\ \ \r\carrier (Vr K v). y = r \\<^sub>r x" apply (frule ale_diff_pos[of "v x" "v y"], simp add:diff_ant_def, simp add:value_of_inv[THEN sym, of v x], cut_tac invf_closed[of "x"], simp only:val_t2p[THEN sym, of v y "x\<^bsup>\K\<^esup>"]) apply (cut_tac field_is_ring, frule_tac x = y and y = "x\<^bsup>\K\<^esup>" in Ring.ring_tOp_closed[of "K"], assumption+, simp add:val_pos_mem_Vr[of v "y \\<^sub>r (x\<^bsup>\K\<^esup>)"], frule field_frac_mul[of y x], assumption+, blast) apply simp done lemma (in Corps) val_Rxa_gt_a:"\valuation K v; x \ carrier (Vr K v) - {\}; y \ carrier (Vr K v); y \ Rxa (Vr K v) x\ \ v x \ (v y)" apply (simp add:Rxa_def, erule bexE, simp add:Vr_tOp_f_tOp, (erule conjE)+, frule_tac x = r in Vr_mem_f_mem[of v], assumption+, frule_tac x = x in Vr_mem_f_mem[of v], assumption+) apply (subst val_t2p, assumption+, simp add:val_pos_mem_Vr[THEN sym, of v], frule_tac y = "v r" in aadd_le_mono[of "0" _ "v x"], simp add:aadd_0_l) done lemma (in Corps) val_Rxa_gt_a_1:"\valuation K v; x \ carrier (Vr K v); y \ carrier (Vr K v); x \ \; v x \ (v y)\ \ y \ Rxa (Vr K v) x" apply (frule_tac x = x in Vr_mem_f_mem[of v], assumption+, frule_tac x = y in Vr_mem_f_mem[of v], assumption+, frule v_ale_diff[of v x y], assumption+, cut_tac invf_closed[of x], cut_tac field_is_ring, frule Ring.ring_tOp_closed[of K y "x\<^bsup>\K\<^esup>"], assumption+) apply (simp add:val_pos_mem_Vr[of "v" "y \\<^sub>r (x\<^bsup>\K\<^esup>)"], frule field_frac_mul[of "y" "x"], assumption+, simp add:Rxa_def, simp add:Vr_tOp_f_tOp, blast, simp) done lemma (in Corps) eqval_inv:"\valuation K v; x \ carrier K; y \ carrier K; y \ \; v x = v y\ \ 0 = v (x \\<^sub>r (y\<^bsup>\K\<^esup>))" by (cut_tac invf_closed[of y], simp add:val_t2p value_of_inv, simp add:aadd_minus_r, simp) lemma (in Corps) eq_val_eq_idealTr:"\valuation K v; x \ carrier (Vr K v) - {\}; y \ carrier (Vr K v); v x \ (v y)\ \ Rxa (Vr K v) y \ Rxa (Vr K v) x" apply (frule val_Rxa_gt_a_1[of v x y], simp+, erule conjE) apply (frule_tac x = x in Vr_mem_f_mem[of v], assumption+, frule Vr_ring[of v], frule Ring.principal_ideal[of "Vr K v" "x"], assumption, frule Ring.ideal_cont_Rxa[of "Vr K v" "(Vr K v) \\<^sub>p x" "y"], assumption+) done lemma (in Corps) eq_val_eq_ideal:"\valuation K v; x \ carrier (Vr K v); y \ carrier (Vr K v); v x = v y\ \ Rxa (Vr K v) x = Rxa (Vr K v) y" apply (case_tac "x = \\<^bsub>K\<^esub>", simp add:value_of_zero, frule value_inf_zero[of v y], simp add:Vr_mem_f_mem, rule sym, assumption, simp) apply (rule equalityI, rule eq_val_eq_idealTr[of v y x], assumption+, drule sym, simp, rule contrapos_pp, simp+, simp add:value_of_zero, frule Vr_mem_f_mem[of v x], assumption+, frule value_inf_zero[of v x], assumption+, rule sym, assumption, simp, simp, simp) apply (rule eq_val_eq_idealTr[of v x y], assumption+, simp, assumption, rule aeq_ale, assumption+) done lemma (in Corps) eq_ideal_eq_val:"\valuation K v; x \ carrier (Vr K v); y \ carrier (Vr K v); Rxa (Vr K v) x = Rxa (Vr K v) y\ \ v x = v y" apply (case_tac "x = \\<^bsub>K\<^esub>", simp, drule sym, frule Vr_ring[of v], frule Ring.a_in_principal[of "Vr K v" y], assumption+, simp, thin_tac "Vr K v \\<^sub>p y = Vr K v \\<^sub>p (\)", simp add:Rxa_def, erule bexE, simp add:Vr_0_f_0[of v, THEN sym]) apply (simp add:Vr_tOp_f_tOp, simp add:Vr_0_f_0, frule_tac x = r in Vr_mem_f_mem[of v], assumption+, cut_tac field_is_ring, simp add:Ring.ring_times_x_0) apply (frule Vr_ring[of v], frule val_Rxa_gt_a[of v x y], simp, simp) apply (drule sym, frule Ring.a_in_principal[of "Vr K v" "y"], simp, simp) apply (frule val_Rxa_gt_a[of v y x], simp, rule contrapos_pp, simp+, frule Ring.a_in_principal[of "Vr K v" "x"], assumption+, simp add:Rxa_def, erule bexE, simp add:Vr_tOp_f_tOp, cut_tac field_is_ring, frule_tac x = r in Vr_mem_f_mem[of v], assumption+, simp add:Ring.ring_times_x_0, simp, frule Ring.a_in_principal[of "Vr K v" "x"], assumption+, simp, rule ale_antisym, assumption+) done lemma (in Corps) zero_val_gen_whole: "\valuation K v; x \ carrier (Vr K v)\ \ (v x = 0) = (Rxa (Vr K v) x = carrier (Vr K v))" apply (frule Vr_mem_f_mem[of v x], assumption, frule Vr_ring[of v]) apply (rule iffI, frule Ring.principal_ideal[of "Vr K v" "x"], assumption+, frule Ring.a_in_principal[of "Vr K v" "x"], assumption+, rule ideal_inc_elem0val_whole[of v x "Vr K v \\<^sub>p x"], assumption+, frule Ring.ring_one[of "Vr K v"], frule eq_set_inc[of "1\<^sub>r\<^bsub>(Vr K v)\<^esub>" "carrier (Vr K v)" "Vr K v \\<^sub>p x"], drule sym, assumption, thin_tac "1\<^sub>r\<^bsub>(Vr K v)\<^esub> \ carrier (Vr K v)", thin_tac "Vr K v \\<^sub>p x = carrier (Vr K v)") apply (simp add:Rxa_def, erule bexE, simp add:Vr_1_f_1, simp add:Vr_tOp_f_tOp, frule value_of_one[of v], simp, frule_tac x = r in Vr_mem_f_mem[of v], assumption+, cut_tac field_is_ring, simp add:val_t2p, simp add:val_pos_mem_Vr[THEN sym, of v], rule contrapos_pp, simp+, cut_tac less_le[THEN sym, of "0" "v x"], drule not_sym, simp, frule_tac x = "v r" and y = "v x" in aadd_pos_poss, assumption+, simp) done lemma (in Corps) elem_nonzeroval_gen_proper:"\ valuation K v; x \ carrier (Vr K v); v x \ 0\ \ Rxa (Vr K v) x \ carrier (Vr K v)" apply (rule contrapos_pp, simp+) apply (simp add: zero_val_gen_whole[THEN sym]) done text\We prove that Vr K v is a principal ideal ring\ definition LI :: "[('r, 'm) Ring_scheme, 'r \ ant, 'r set] \ ant" where (** The least nonzero value of I **) "LI K v I = AMin (v ` I)" definition Ig :: "[('r, 'm) Ring_scheme, 'r \ ant, 'r set] \ 'r" where (** Generator of I **) "Ig K v I = (SOME x. x \ I \ v x = LI K v I)" lemma (in Corps) val_in_image:"\valuation K v; ideal (Vr K v) I; x \ I\ \ v x \ v ` I" by (simp add:image_def, blast) lemma (in Corps) I_vals_nonempty:"\valuation K v; ideal (Vr K v) I\ \ v ` I \ {}" by (frule Vr_ring[of v], frule Ring.ideal_zero[of "Vr K v" "I"], assumption+, rule contrapos_pp, simp+) lemma (in Corps) I_vals_LBset:"\ valuation K v; ideal (Vr K v) I\ \ v ` I \ LBset 0" apply (frule Vr_ring[of v], rule subsetI, simp add:LBset_def, simp add:image_def) apply (erule bexE, frule_tac h = xa in Ring.ideal_subset[of "Vr K v" "I"], assumption+) apply (frule_tac x1 = xa in val_pos_mem_Vr[THEN sym, of v], simp add:Vr_mem_f_mem, simp) done lemma (in Corps) LI_pos:"\valuation K v; ideal (Vr K v) I\ \ 0 \ LI K v I" apply (simp add:LI_def, frule I_vals_LBset[of v], simp add:ant_0[THEN sym], frule I_vals_nonempty[of v], simp only:ant_0) apply (simp only:ant_0[THEN sym], frule AMin[of "v ` I" "0"], assumption, erule conjE, frule subsetD[of "v ` I" "LBset (ant 0)" "AMin (v ` I)"], assumption+, simp add:LBset_def) done lemma (in Corps) LI_poss:"\valuation K v; ideal (Vr K v) I; I \ carrier (Vr K v)\ \ 0 < LI K v I" apply (subst less_le) apply (simp add:LI_pos) apply (rule contrapos_pp, simp+) apply (simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp add:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, simp only:ant_0) apply (simp only:ant_0[THEN sym], frule AMin[of "v ` I" "0"], assumption, erule conjE, frule subsetD[of "v ` I" "LBset (ant 0)" "AMin (v ` I)"], assumption+, simp add:LBset_def) apply (thin_tac "\x\I. ant 0 \ v x", thin_tac "v ` I \ {x. ant 0 \ x}", simp add:image_def, erule bexE, simp add:ant_0) apply (frule Vr_ring[of v], frule_tac h = x in Ring.ideal_subset[of "Vr K v" "I"], assumption+, frule_tac x = x in zero_val_gen_whole[of v], assumption+, simp, frule_tac a = x in Ring.ideal_cont_Rxa[of "Vr K v" "I"], assumption+, simp, frule Ring.ideal_subset1[of "Vr K v" "I"], assumption+) apply (frule equalityI[of "I" "carrier (Vr K v)"], assumption+, simp) done lemma (in Corps) LI_z:"\valuation K v; ideal (Vr K v) I; I \ {\\<^bsub>Vr K v\<^esub>}\ \ \z. LI K v I = ant z" apply (frule Vr_ring[of v], frule Ring.ideal_zero[of "Vr K v" "I"], assumption+, cut_tac mem_ant[of "LI K v I"], frule LI_pos[of v I], assumption, erule disjE, simp, cut_tac minf_le_any[of "0"], frule ale_antisym[of "0" "-\"], assumption+, simp) apply (erule disjE, simp, frule singleton_sub[of "\\<^bsub>Vr K v\<^esub>" "I"], frule sets_not_eq[of "I" "{\\<^bsub>Vr K v\<^esub>}"], assumption+, erule bexE, simp) apply (simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp only:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE) apply (frule_tac x = a in val_in_image[of v I], assumption+, drule_tac x = "v a" in bspec, simp, simp add:Vr_0_f_0, frule_tac x = a in val_nonzero_z[of v], simp add:Ring.ideal_subset Vr_mem_f_mem, assumption+, erule exE, simp, cut_tac x = "ant z" in inf_ge_any, frule_tac x = "ant z" in ale_antisym[of _ "\"], assumption+, simp) done lemma (in Corps) LI_k:"\valuation K v; ideal (Vr K v) I\ \ \k\ I. LI K v I = v k" by (simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp only:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE, thin_tac "\x\v ` I. AMin (v ` I) \ x", simp add:image_def) lemma (in Corps) LI_infinity:"\valuation K v; ideal (Vr K v) I\ \ (LI K v I = \) = (I = {\\<^bsub>Vr K v\<^esub>})" apply (frule Vr_ring[of v]) apply (rule iffI) apply (rule contrapos_pp, simp+, frule Ring.ideal_zero[of "Vr K v" "I"], assumption+, frule singleton_sub[of "\\<^bsub>Vr K v\<^esub>" "I"], frule sets_not_eq[of "I" "{\\<^bsub>Vr K v\<^esub>}"], assumption+, erule bexE, frule_tac h = a in Ring.ideal_subset[of "Vr K v" "I"], assumption+, simp add:Vr_0_f_0, frule_tac x = a in Vr_mem_f_mem[of v], assumption+, frule_tac x = a in val_nonzero_z[of v], assumption+, erule exE, simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp only:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE) apply (frule_tac h = a in Ring.ideal_subset[of "Vr K v" "I"], assumption+, frule_tac x = a in val_in_image[of v I], assumption+, drule_tac x = "v a" in bspec, simp) apply (frule_tac x = a in val_nonzero_z[of v], assumption+, erule exE, simp, cut_tac x = "ant z" in inf_ge_any, frule_tac x = "ant z" in ale_antisym[of _ "\"], assumption+, simp) apply (frule sym, thin_tac "I = {\\<^bsub>Vr K v\<^esub>}", simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp only:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE, drule sym, simp, simp add:Vr_0_f_0 value_of_zero) done lemma (in Corps) val_Ig:"\valuation K v; ideal (Vr K v) I\ \ (Ig K v I) \ I \ v (Ig K v I) = LI K v I" by (simp add:Ig_def, rule someI2_ex, frule LI_k[of v I], assumption+, erule bexE, drule sym, blast, assumption) lemma (in Corps) Ig_nonzero:"\valuation K v; ideal (Vr K v) I; I \ {\\<^bsub>Vr K v\<^esub>}\ \ (Ig K v I) \ \" by (rule contrapos_pp, simp+, frule LI_infinity[of v I], assumption+, frule val_Ig[of v I], assumption+, erule conjE, simp add:value_of_zero) lemma (in Corps) Vr_ideal_npowf_closed:"\valuation K v; ideal (Vr K v) I; x \ I; 0 < n\ \ x\<^bsub>K\<^esub>\<^bsup>n\<^esup> \ I" by (simp add:npowf_def, frule Vr_ring[of v], frule Ring.ideal_npow_closed[of "Vr K v" "I" "x" "nat n"], assumption+, simp, frule Ring.ideal_subset[of "Vr K v" "I" "x"], assumption+, simp add:Vr_exp_f_exp) lemma (in Corps) Ig_generate_I:"\valuation K v; ideal (Vr K v) I\ \ (Vr K v) \\<^sub>p (Ig K v I) = I" apply (frule Vr_ring[of v]) apply (case_tac "I = carrier (Vr K v)", frule sym, thin_tac "I = carrier (Vr K v)", frule Ring.ring_one[of "Vr K v"], simp, simp add: Vr_1_f_1, frule val_Ig[of v I], assumption+, erule conjE, frule LI_pos[of v I], assumption+, simp add: LI_def cong del: image_cong_simp, frule I_vals_LBset[of v], assumption+, simp only: ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE, frule val_in_image[of v I "1\<^sub>r"], assumption+, drule_tac x = "v (1\<^sub>r)" in bspec, assumption+, simp add: value_of_one ant_0 cong del: image_cong_simp, simp add: zero_val_gen_whole[of v "Ig K v I"]) apply (frule val_Ig[of v I], assumption+, (erule conjE)+, frule Ring.ideal_cont_Rxa[of "Vr K v" "I" "Ig K v I"], assumption+, rule equalityI, assumption+) apply (case_tac "LI K v I = \", frule LI_infinity[of v I], simp, simp add:Rxa_def, simp add:Ring.ring_times_x_0, frule Ring.ring_zero, blast) apply (rule subsetI, case_tac "v x = 0", frule_tac x = x in Vr_mem_f_mem[of v], simp add:Ring.ideal_subset, frule_tac x = x in zero_val_gen_whole[of v], simp add:Ring.ideal_subset, simp, frule_tac a = x in Ring.ideal_cont_Rxa[of "Vr K v" "I"], assumption+, simp, frule Ring.ideal_subset1[of "Vr K v" "I"], assumption, frule equalityI[of "I" "carrier (Vr K v)"], assumption+, simp) apply (simp add:LI_def, frule I_vals_LBset[of v], assumption+, simp only:ant_0[THEN sym], frule I_vals_nonempty[of v], assumption+, frule AMin[of "v ` I" "0"], assumption, erule conjE, frule_tac x = "v x" in bspec, frule_tac x = x in val_in_image[of v I], assumption+, simp) apply (drule_tac x = x in bspec, assumption, frule_tac y = x in eq_val_eq_idealTr[of v "Ig K v I"], simp add:Ring.ideal_subset, rule contrapos_pp, simp+, simp add:value_of_zero, simp add:Ring.ideal_subset, simp) apply (frule_tac a = x in Ring.a_in_principal[of "Vr K v"], simp add:Ring.ideal_subset, rule subsetD, assumption+) done lemma (in Corps) Pg_gen_vp:"valuation K v \ (Vr K v) \\<^sub>p (Pg K v) = vp K v" apply (frule vp_ideal[of v], frule Ig_generate_I[of v "vp K v"], assumption+, frule vp_not_whole[of v], frule eq_val_eq_ideal[of v "Ig K v (vp K v)" "Pg K v"], frule val_Ig [of v "vp K v"], assumption+, erule conjE, simp add:vp_mem_Vr_mem) apply (frule val_Pg[of v], erule conjE, frule Lv_pos[of v], rotate_tac -2, drule sym, simp, simp add:val_poss_mem_Vr) apply (thin_tac "Vr K v \\<^sub>p Ig K v (vp K v) = vp K v", frule val_Pg[of v], erule conjE, simp, frule val_Ig[of v "vp K v"], assumption+, erule conjE, simp, thin_tac "v (Pg K v) = Lv K v", thin_tac "Ig K v (vp K v) \ vp K v \ v (Ig K v (vp K v)) = LI K v (vp K v)", simp add:LI_def Lv_def, subgoal_tac "v ` vp K v = {x. x \ v ` carrier K \ 0 < x}", simp) apply (thin_tac "ideal (Vr K v) (vp K v)", thin_tac "Pg K v \ carrier K", thin_tac "Pg K v \ \", rule equalityI, rule subsetI, simp add:image_def vp_def, erule exE, erule conjE, (erule conjE)+, frule_tac x = xa in Vr_mem_f_mem[of v], assumption+, simp, blast) apply (rule subsetI, simp add:image_def vp_def, erule conjE, erule bexE, simp, frule_tac x = xa in val_poss_mem_Vr[of v], assumption+, cut_tac y = "v xa" in less_le[of "0"], simp, blast, simp) done lemma (in Corps) vp_gen_t:"valuation K v \ \t\carrier (Vr K v). vp K v = (Vr K v) \\<^sub>p t" by (frule Pg_gen_vp[of v], frule Pg_in_Vr[of v], blast) lemma (in Corps) vp_gen_nonzero:"\valuation K v; vp K v = (Vr K v) \\<^sub>p t\ \ t \ \\<^bsub>Vr K v\<^esub>" apply (rule contrapos_pp, simp+, cut_tac Ring.Rxa_zero[of "Vr K v"], drule sym, simp, simp add:vp_nonzero) apply (simp add:Vr_ring) done lemma (in Corps) n_value_idealTr:"\valuation K v; 0 \ n\ \ (vp K v) \<^bsup>\(Vr K v) n\<^esup> = Vr K v \\<^sub>p ((Pg K v)^\<^bsup>(Vr K v) n\<^esup>)" apply (frule Vr_ring[of v], frule Pg_gen_vp[THEN sym, of v], simp add:vp_ideal, frule val_Pg[of v], simp, (erule conjE)+) apply (subst Ring.principal_ideal_n_pow[of "Vr K v" "Pg K v" "Vr K v \\<^sub>p Pg K v"], assumption+, frule Lv_pos[of v], rotate_tac -2, frule sym, thin_tac "v (Pg K v) = Lv K v", simp, simp add:val_poss_mem_Vr, simp+) done lemma (in Corps) ideal_pow_vp:"\valuation K v; ideal (Vr K v) I; I \ carrier (Vr K v); I \ {\\<^bsub>Vr K v\<^esub>}\ \ I = (vp K v)\<^bsup>\ (Vr K v) (na (n_val K v (Ig K v I)))\<^esup>" apply (frule Vr_ring[of v], frule Ig_generate_I[of v I], assumption+) apply (frule n_val[of v "Ig K v I"], frule val_Ig[of v I], assumption+, erule conjE, simp add:Ring.ideal_subset[of "Vr K v" "I" "Ig K v I"] Vr_mem_f_mem) apply (frule val_Pg[of v], erule conjE, rotate_tac -1, drule sym, simp, frule Ig_nonzero[of v I], assumption+, frule LI_pos[of v I], assumption+, frule Lv_pos[of v], frule val_Ig[of v I], assumption+, (erule conjE)+, rotate_tac -1, drule sym, simp, frule val_pos_n_val_pos[of v "Ig K v I"], simp add:Ring.ideal_subset Vr_mem_f_mem, simp) apply (frule zero_val_gen_whole[THEN sym, of v "Ig K v I"], simp add:Ring.ideal_subset, simp, rotate_tac -1, drule not_sym, cut_tac less_le[THEN sym, of "0" "v (Ig K v I)"], simp, thin_tac "0 \ v (Ig K v I)", frule Ring.ideal_subset[of "Vr K v" I "Ig K v I"], assumption+, frule Vr_mem_f_mem[of v "Ig K v I"], assumption+, frule val_poss_n_val_poss[of v "Ig K v I"], assumption+, simp) apply (frule Ig_nonzero[of v I], frule val_nonzero_noninf[of v "Ig K v I"], assumption+, simp add:val_noninf_n_val_noninf[of v "Ig K v I"], frule val_poss_mem_Vr[of v "Pg K v"], assumption+, subst n_value_idealTr[of v "na (n_val K v (Ig K v I))"], assumption+, simp add:na_def) apply (frule eq_val_eq_ideal[of v "Ig K v I" "(Pg K v)^\<^bsup>(Vr K v) (na (n_val K v (Ig K v I)))\<^esup>"], assumption+, rule Ring.npClose, assumption+, simp add:Vr_exp_f_exp[of v "Pg K v"], subst val_exp_ring[THEN sym, of v "Pg K v" "na (n_val K v (Ig K v I))"], assumption+) apply (frule Lv_z[of v], erule exE, simp, rotate_tac 6, drule sym, simp, subst asprod_amult, simp add:val_poss_n_val_poss[of v "Ig K v I"], frule val_nonzero_noninf[of v "Ig K v I"], assumption+, frule val_noninf_n_val_noninf[of v "Ig K v I"], assumption+, simp, rule aposs_na_poss[of "n_val K v (Ig K v I)"], assumption+) apply (fold an_def) apply (subst an_na[THEN sym, of "n_val K v (Ig K v I)"], frule val_nonzero_noninf[of v "Ig K v I"], assumption+, frule val_noninf_n_val_noninf[of v "Ig K v I"], assumption+, simp, simp add:aless_imp_le, simp) apply simp done lemma (in Corps) ideal_apow_vp:"\valuation K v; ideal (Vr K v) I\ \ I = (vp K v)\<^bsup> (Vr K v) (n_val K v (Ig K v I))\<^esup>" apply (frule Vr_ring[of v]) apply (case_tac "v (Ig K v I) = \", frule val_Ig[of v I], assumption, frule val_inf_n_val_inf[of v "Ig K v I"], simp add:Ring.ideal_subset Vr_mem_f_mem, simp, simp add:r_apow_def, simp add:LI_infinity[of v I]) apply (case_tac "v (Ig K v I) = 0", frule val_0_n_val_0[of v "Ig K v I"], frule val_Ig[of v I], assumption+, erule conjE, simp add:Ring.ideal_subset Vr_mem_f_mem, simp, frule val_Ig[of v I], assumption, frule zero_val_gen_whole[of v "Ig K v I"], simp add:Ring.ideal_subset, (erule conjE)+, simp, frule Ring.ideal_cont_Rxa[of "Vr K v" "I" "Ig K v I"], assumption+) apply (simp, frule Ring.ideal_subset1[of "Vr K v" "I"], assumption+, frule equalityI[of "I" "carrier (Vr K v)"], assumption+, simp add:r_apow_def) apply (frule val_noninf_n_val_noninf[of v "Ig K v I"], frule val_Ig[of v I], assumption, simp add:Ring.ideal_subset Vr_mem_f_mem, simp, frule value_n0_n_val_n0[of v "Ig K v I"], frule val_Ig[of v I], assumption, simp add:Ring.ideal_subset Vr_mem_f_mem, assumption) apply (simp add:r_apow_def, rule ideal_pow_vp, assumption+, frule elem_nonzeroval_gen_proper[of v "Ig K v I"], frule val_Ig[of v I], assumption+, erule conjE, simp add:Ring.ideal_subset, assumption, simp add:Ig_generate_I) apply (frule val_Ig[of v I], assumption+, erule conjE, simp, simp add:LI_infinity[of v I]) done (* A note to the above lemma (in Corps). Let K be a field and v be a valuation. Let R be the valuaiton ring of v, and let P be the maximal ideal of R. If I is an ideal of R such that I \ 0 and I \ R, then I = P^n. Here n = nat znt n_valuation K G a i v (I_gen K v I)) which is nat of the integer part of the normal value of (I_gen K v I). Let b be a generator of I, then n = v (b) / v (p), where p is a generator of P in R: I = P \<^bsup>\R n\<^esup> Here P = vp K v, R = Vr K v, b = Ig K v I,, n = nat n_val K v (Ig K v I). It is easy to see that n = v\<^sup>* b. Here v\<^sup>* is the normal valuation derived from v. *) lemma (in Corps) ideal_apow_n_val:"\valuation K v; x \ carrier (Vr K v)\ \ (Vr K v) \\<^sub>p x = (vp K v)\<^bsup>(Vr K v) (n_val K v x)\<^esup>" apply (frule Vr_ring[of v], frule Ring.principal_ideal[of "Vr K v" "x"], assumption+, frule ideal_apow_vp[of v "Vr K v \\<^sub>p x"], assumption+) apply (frule val_Ig[of v "Vr K v \\<^sub>p x"], assumption+, erule conjE, frule Ring.ideal_subset[of "Vr K v" "Vr K v \\<^sub>p x" "Ig K v (Vr K v \\<^sub>p x)"], assumption+, frule Ig_generate_I[of v "Vr K v \\<^sub>p x"], assumption+) apply (frule eq_ideal_eq_val[of v "Ig K v (Vr K v \\<^sub>p x)" x], assumption+, thin_tac "Vr K v \\<^sub>p Ig K v (Vr K v \\<^sub>p x) = Vr K v \\<^sub>p x", thin_tac "v (Ig K v (Vr K v \\<^sub>p x)) = LI K v (Vr K v \\<^sub>p x)", frule n_val[THEN sym, of v x], simp add:Vr_mem_f_mem, simp, thin_tac "v x = n_val K v x * Lv K v", frule n_val[THEN sym, of v "Ig K v (Vr K v \\<^sub>p x)"], simp add:Vr_mem_f_mem, simp, thin_tac "v (Ig K v (Vr K v \\<^sub>p x)) = n_val K v x * Lv K v") apply (frule Lv_pos[of v], frule Lv_z[of v], erule exE, simp, frule_tac s = z in zless_neq[THEN not_sym, of "0"], frule_tac z = z in adiv_eq[of _ "n_val K v (Ig K v (Vr K v \\<^sub>p x))" "n_val K v x"], assumption+, simp) done lemma (in Corps) t_gen_vp:"\valuation K v; t \ carrier K; v t = 1\ \ (Vr K v) \\<^sub>p t = vp K v" (* apply (frule val_surj_n_val[of v], blast) apply (frule ideal_apow_n_val[of v t]) apply (cut_tac a0_less_1) apply (rule val_poss_mem_Vr[of v t], assumption+, simp) apply (simp add:r_apow_def) apply (simp only:ant_1[THEN sym], simp only:ant_0[THEN sym]) apply (simp only:aeq_zeq, simp) apply (cut_tac z_neq_inf[THEN not_sym, of "1"], simp) apply (simp only:an_1[THEN sym]) apply (simp add:na_an) apply (rule Ring.idealprod_whole_r[of "Vr K v" "vp K v"]) apply (simp add:Vr_ring) apply (simp add:vp_ideal) done *) proof - assume a1:"valuation K v" and a2:"t \ carrier K" and a3:"v t = 1" from a1 and a2 and a3 have h1:"t \ carrier (Vr K v)" apply (cut_tac a0_less_1) apply (rule val_poss_mem_Vr[of v t], assumption+, simp) done from a1 and a2 and a3 have h2:"n_val K v = v" apply (subst val_surj_n_val[of v]) apply assumption apply blast apply simp done from a1 and h1 have h3:"Vr K v \\<^sub>p t = vp K v\<^bsup> (Vr K v) (n_val K v t)\<^esup>" apply (simp add:ideal_apow_n_val[of v t]) done from a1 and a3 and h2 and h3 show ?thesis apply (simp add:r_apow_def) apply (simp only:ant_1[THEN sym], simp only:ant_0[THEN sym]) apply (simp only:aeq_zeq, simp) apply (cut_tac z_neq_inf[THEN not_sym, of "1"], simp) apply (simp only:an_1[THEN sym]) apply (simp add:na_an) apply (rule Ring.idealprod_whole_r[of "Vr K v" "vp K v"]) apply (simp add:Vr_ring) apply (simp add:vp_ideal) done qed lemma (in Corps) t_vp_apow:"\valuation K v; t \ carrier K; v t = 1\ \ (Vr K v) \\<^sub>p (t^\<^bsup>(Vr K v) n\<^esup>) = (vp K v)\<^bsup>(Vr K v) (an n)\<^esup>" (* apply (frule Vr_ring[of v], subst Ring.principal_ideal_n_pow[THEN sym, of "Vr K v" t "vp K v" n], assumption+) apply (cut_tac a0_less_1, rule val_poss_mem_Vr[of v], assumption+) apply (simp, simp add:t_gen_vp, simp add:r_apow_def) apply (rule conjI, rule impI, simp only:an_0[THEN sym], frule an_inj[of n 0], simp) apply (rule impI) apply (rule conjI, rule impI) apply (simp add:an_def) apply (rule impI, cut_tac an_nat_pos[of n], simp add:na_an) done *) proof - assume a1:"valuation K v" and a2:"t \ carrier K" and a3:"v t = 1" from a1 have h1:"Ring (Vr K v)" by (simp add:Vr_ring[of v]) from a1 and a2 and a3 have h2:"t \ carrier (Vr K v)" apply (cut_tac a0_less_1) apply (rule val_poss_mem_Vr) apply assumption+ apply simp done from a1 and a2 and a3 and h1 and h2 show ?thesis apply (subst Ring.principal_ideal_n_pow[THEN sym, of "Vr K v" t "vp K v" n]) apply assumption+ apply (simp add:t_gen_vp) apply (simp add:r_apow_def) apply (rule conjI, rule impI, simp only:an_0[THEN sym], frule an_inj[of n 0], simp) apply (rule impI) apply (rule conjI, rule impI) apply (simp add:an_def) apply (rule impI, cut_tac an_nat_pos[of n], simp add:na_an) done qed lemma (in Corps) nonzeroelem_gen_nonzero:"\valuation K v; x \ \; x \ carrier (Vr K v)\ \ Vr K v \\<^sub>p x \ {\\<^bsub>Vr K v\<^esub>}" by (frule Vr_ring[of v], frule_tac a = x in Ring.a_in_principal[of "Vr K v"], assumption+, rule contrapos_pp, simp+, simp add:Vr_0_f_0) subsection "Amin lemma (in Corps)s " lemma (in Corps) Amin_le_addTr:"valuation K v \ (\j \ n. f j \ carrier K) \ Amin n (v \ f) \ (v (nsum K f n))" apply (induct_tac n) apply (rule impI, simp) apply (rule impI, simp, frule_tac x = "\\<^sub>e K f n" and y = "f (Suc n)" in amin_le_plus[of v], cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], cut_tac n = n in aGroup.nsum_mem[of K _ f], assumption, rule allI, simp add:funcset_mem, assumption, simp) apply (frule_tac z = "Amin n (\u. v (f u))" and z' = "v (\\<^sub>e K f n)" and w = "v (f (Suc n))" in amin_aminTr, rule_tac i = "amin (Amin n (\u. v (f u))) (v (f (Suc n)))" and j = "amin (v (\\<^sub>e K f n)) (v (f (Suc n)))" and k = "v (\\<^sub>e K f n \ (f (Suc n)))" in ale_trans, assumption+) done lemma (in Corps) Amin_le_add:"\valuation K v; \j \ n. f j \ carrier K\ \ Amin n (v \ f) \ (v (nsum K f n))" by (frule Amin_le_addTr[of v n f], simp) lemma (in Corps) value_ge_add:"\valuation K v; \j \ n. f j \ carrier K; \j \ n. z \ ((v \ f) j)\ \ z \ (v (\\<^sub>e K f n))" apply (frule Amin_le_add[of v n f], assumption+, cut_tac Amin_ge[of n "v \ f" z], rule ale_trans, assumption+) apply (rule allI, rule impI, simp add:comp_def Zset_def, rule value_in_aug_inf[of v], assumption+, simp+) done lemma (in Corps) Vr_ideal_powTr1:"\valuation K v; ideal (Vr K v) I; I \ carrier (Vr K v); b \ I\ \ b \ (vp K v)" by (frule ideal_sub_vp[of v I], assumption+, simp add:subsetD) section \pow of vp and \n_value\ -- convergence --\ lemma (in Corps) n_value_x_1:"\valuation K v; 0 \ n; x \ (vp K v) \<^bsup>(Vr K v) n\<^esup>\ \ n \ (n_val K v x)" (* 1. prove that x \ carrier (Vr K v) and that x \ carrier K *) apply ((case_tac "n = \", simp add:r_apow_def, simp add:Vr_0_f_0, cut_tac field_is_ring, frule Ring.ring_zero[of "K"], frule val_inf_n_val_inf[of v \], assumption+, simp add:value_of_zero), (case_tac "n = 0", simp add:r_apow_def, subst val_pos_n_val_pos[THEN sym, of v x], assumption+, simp add:Vr_mem_f_mem, subst val_pos_mem_Vr[of v x], assumption+, simp add:Vr_mem_f_mem, assumption, simp add:r_apow_def, frule Vr_ring[of v], frule vp_pow_ideal[of v "na n"], frule Ring.ideal_subset[of "Vr K v" "(vp K v) \<^bsup>\(Vr K v) (na n)\<^esup>" x], assumption+, frule Vr_mem_f_mem[of v x], assumption+)) (* 1. done *) (** 2. Show that v (I_gen K v (vpr K v)^\<^sup>Vr K v\<^sup> \<^sup>nat n) \ v x. the key lemma (in Corps) is "val_Rxa_gt_a" **) apply (case_tac "x = \\<^bsub>K\<^esub>", simp, frule value_of_zero[of v], simp add:val_inf_n_val_inf, simp add:n_value_idealTr[of v "na n"], frule val_Pg[of v], erule conjE, simp, erule conjE, frule Lv_pos[of v], rotate_tac -4, frule sym, thin_tac "v (Pg K v) = Lv K v", simp, frule val_poss_mem_Vr[of v "Pg K v"], assumption+, frule val_Rxa_gt_a[of v "Pg K v^\<^bsup>(Vr K v) (na n)\<^esup>" x], frule Vr_integral[of v], simp only:Vr_0_f_0[of v, THEN sym], frule Idomain.idom_potent_nonzero[of "Vr K v" "Pg K v" "na n"], assumption+, simp, simp add:Ring.npClose, assumption+) apply (thin_tac "x \ Vr K v \\<^sub>p (Pg K v^\<^bsup>(Vr K v) (na n)\<^esup>)", thin_tac "ideal (Vr K v) (Vr K v \\<^sub>p (Pg K v^\<^bsup>(Vr K v) (na n)\<^esup>))") apply (simp add:Vr_exp_f_exp[of v "Pg K v"], simp add:val_exp_ring[THEN sym, of v "Pg K v"], simp add:n_val[THEN sym, of v x], frule val_nonzero_z[of v "Pg K v"], assumption+, erule exE, simp, frule aposs_na_poss[of "n"], simp add: less_le, simp add:asprod_amult, frule_tac w = z in amult_pos_mono_r[of _ "ant (int (na n))" "n_val K v x"], simp, cut_tac an_na[of "n"], simp add:an_def, assumption+) done lemma (in Corps) n_value_x_1_nat:"\valuation K v; x \ (vp K v)\<^bsup>\(Vr K v) n\<^esup> \ \ (an n) \ (n_val K v x)" apply (cut_tac an_nat_pos[of "n"]) apply( frule n_value_x_1[of "v" "an n" "x"], assumption+) apply (simp add:r_apow_def) apply (case_tac "n = 0", simp, simp) apply (cut_tac aless_nat_less[THEN sym, of "0" "n"]) apply simp unfolding less_le apply simp apply (cut_tac an_neq_inf [of "n"]) apply simp apply (simp add: na_an) apply assumption done lemma (in Corps) n_value_x_2:"\valuation K v; x \ carrier (Vr K v); n \ (n_val K v x); 0 \ n\ \ x \ (vp K v) \<^bsup>(Vr K v) n\<^esup>" apply (frule Vr_ring[of v], frule val_Pg[of v], erule conjE, simp, erule conjE, drule sym, frule Lv_pos[of v], simp, frule val_poss_mem_Vr[of v "Pg K v"], assumption+) apply (case_tac "n = \", simp add:r_apow_def, cut_tac inf_ge_any[of "n_val K v x"], frule ale_antisym[of "n_val K v x" "\"], assumption+, frule val_inf_n_val_inf[THEN sym, of v "x"], simp add:Vr_mem_f_mem, simp, frule value_inf_zero[of v x], simp add:Vr_mem_f_mem, simp+, simp add:Vr_0_f_0) apply (case_tac "n = 0", simp add:r_apow_def, simp add:r_apow_def, subst n_value_idealTr[of v "na n"], assumption+, simp add:apos_na_pos) apply (rule val_Rxa_gt_a_1[of v "Pg K v^\<^bsup>(Vr K v) (na n)\<^esup>" x], assumption+, rule Ring.npClose, assumption+, simp add:Vr_0_f_0[THEN sym, of v], frule Vr_integral[of v], frule val_poss_mem_Vr[of v "Pg K v"], assumption+, simp add:Idomain.idom_potent_nonzero) apply (simp add:Vr_exp_f_exp, simp add:val_exp_ring[THEN sym, of v], rotate_tac -5, drule sym, frule Lv_z[of v], erule exE, simp, frule aposs_na_poss[of "n"], simp add: less_le, simp add:asprod_amult, subst n_val[THEN sym, of v x], assumption+, simp add:Vr_mem_f_mem, simp, subst amult_pos_mono_r[of _ "ant (int (na n))" "n_val K v x"], assumption, cut_tac an_na[of "n"], simp add:an_def, assumption+) done lemma (in Corps) n_value_x_2_nat:"\valuation K v; x \ carrier (Vr K v); (an n) \ ((n_val K v) x)\ \ x \ (vp K v)\<^bsup>\(Vr K v) n\<^esup>" by (frule n_value_x_2[of v x "an n"], assumption+, simp, simp add:r_apow_def, case_tac "an n = \", simp add:an_def, simp, case_tac "n = 0", simp, subgoal_tac "an n \ 0", simp, simp add:na_an, rule contrapos_pp, simp+, simp add:an_def) lemma (in Corps) n_val_n_pow:"\valuation K v; x \ carrier (Vr K v); 0 \ n\ \ (n \ (n_val K v x)) = (x \ (vp K v) \<^bsup>(Vr K v) n\<^esup>)" by (rule iffI, simp add:n_value_x_2, simp add:n_value_x_1) lemma (in Corps) eqval_in_vpr_apow:"\valuation K v; x \ carrier K; 0 \ n; y \ carrier K; n_val K v x = n_val K v y; x \ (vp K v)\<^bsup>(Vr K v) n\<^esup>\ \ y \ (vp K v) \<^bsup>(Vr K v) n\<^esup>" apply (frule n_value_x_1[of v n x], assumption+, simp, rule n_value_x_2[of v y n], assumption+, frule mem_vp_apow_mem_Vr[of v n x], assumption+) apply (frule val_pos_mem_Vr[THEN sym, of v x], assumption+, simp, simp add:val_pos_n_val_pos[of v x], simp add:val_pos_n_val_pos[THEN sym, of v y], simp add:val_pos_mem_Vr, assumption+) done lemma (in Corps) convergenceTr:"\valuation K v; x \ carrier K; b \ carrier K; b \ (vp K v)\<^bsup>(Vr K v) n\<^esup>; (Abs (n_val K v x)) \ n\ \ x \\<^sub>r b \ (vp K v)\<^bsup>(Vr K v) (n + (n_val K v x))\<^esup>" (** Valuation ring is a ring **) apply (cut_tac Abs_pos[of "n_val K v x"], frule ale_trans[of "0" "Abs (n_val K v x)" "n"], assumption+, thin_tac "0 \ Abs (n_val K v x)") apply (frule Vr_ring[of v], frule_tac aadd_le_mono[of "Abs (n_val K v x)" "n" "n_val K v x"], cut_tac Abs_x_plus_x_pos[of "n_val K v x"], frule ale_trans[of "0" "Abs (n_val K v x) + n_val K v x" "n + n_val K v x"], assumption+, thin_tac "0 \ Abs (n_val K v x) + n_val K v x", thin_tac "Abs (n_val K v x) + n_val K v x \ n + n_val K v x", rule n_value_x_2[of v "x \\<^sub>r b" "n + n_val K v x"], assumption+) apply (frule n_value_x_1[of v n b], assumption+) apply (frule aadd_le_mono[of "n" "n_val K v b" "n_val K v x"], frule ale_trans[of "0" "n + n_val K v x" "n_val K v b + n_val K v x"], assumption) apply (thin_tac "0 \ n + n_val K v x", thin_tac "n \ n_val K v b", thin_tac "n + n_val K v x \ n_val K v b + n_val K v x", simp add:aadd_commute[of "n_val K v b" "n_val K v x"]) apply (frule n_val_valuation[of v], simp add:val_t2p[THEN sym, of "n_val K v" x b], cut_tac field_is_ring, frule Ring.ring_tOp_closed[of "K" "x" "b"], assumption+, simp add:val_pos_n_val_pos[THEN sym, of v "x \\<^sub>r b"], simp add:val_pos_mem_Vr, frule n_val_valuation[of v], subst val_t2p[of "n_val K v"], assumption+, frule n_value_x_1[of v n b], assumption+, simp add:aadd_commute[of "n_val K v x" "n_val K v b"], rule aadd_le_mono[of n "n_val K v b" "n_val K v x"], assumption+) done lemma (in Corps) convergenceTr1:"\valuation K v; x \ carrier K; b \ (vp K v)\<^bsup>(Vr K v) (n + Abs (n_val K v x))\<^esup>; 0 \ n\ \ x \\<^sub>r b \ (vp K v) \<^bsup>(Vr K v) n\<^esup>" apply (cut_tac field_is_ring, frule Vr_ring[of v], frule vp_apow_ideal[of v "n + Abs (n_val K v x)"], cut_tac Abs_pos[of "n_val K v x"], rule aadd_two_pos[of "n" "Abs (n_val K v x)"], assumption+) apply (frule Ring.ideal_subset[of "Vr K v" "vp K v\<^bsup> (Vr K v) (n + Abs (n_val K v x))\<^esup>" "b"], assumption+, frule Vr_mem_f_mem[of v b], assumption, frule convergenceTr[of v x b "n + Abs (n_val K v x)"], assumption+, rule aadd_pos_le[of "n" "Abs (n_val K v x)"], assumption) apply (frule apos_in_aug_inf[of "n"], cut_tac Abs_pos[of "n_val K v x"], frule apos_in_aug_inf[of "Abs (n_val K v x)"], frule n_value_in_aug_inf[of v x], assumption+, frule aadd_assoc_i[of "n" "Abs (n_val K v x)" "n_val K v x"], assumption+, cut_tac Abs_x_plus_x_pos[of "n_val K v x"]) apply (frule_tac Ring.ring_tOp_closed[of K x b], assumption+, rule n_value_x_2[of v "x \\<^sub>r b" n], assumption+) apply (subst val_pos_mem_Vr[THEN sym, of v "x \\<^sub>r b"], assumption+, subst val_pos_n_val_pos[of v "x \\<^sub>r b"], assumption+) apply (frule n_value_x_1[of "v" "n + Abs(n_val K v x) + n_val K v x" "x \\<^sub>r b"], subst aadd_assoc_i, assumption+, rule aadd_two_pos[of "n"], assumption+, rule ale_trans[of "0" "n + Abs (n_val K v x) + n_val K v x" "n_val K v (x \\<^sub>r b)"], simp, simp add:aadd_two_pos, assumption, frule n_value_x_1[of "v" "n + Abs (n_val K v x)" " b"], cut_tac Abs_pos[of "n_val K v x"], rule aadd_two_pos[of "n" "Abs (n_val K v x)"], assumption+) apply (frule n_val_valuation[of v], subst val_t2p[of "n_val K v"], assumption+) apply (frule aadd_le_mono[of "n + Abs (n_val K v x)" "n_val K v b" "n_val K v x"], simp add:aadd_commute[of "n_val K v b" "n_val K v x"], rule ale_trans[of "n" "n + (Abs (n_val K v x) + n_val K v x)" "n_val K v x + n_val K v b"], frule aadd_pos_le[of "Abs (n_val K v x) + n_val K v x" "n"], simp add:aadd_commute[of "n"], assumption+) done lemma (in Corps) vp_potent_zero:"\valuation K v; 0 \ n\ \ (n = \) = (vp K v \<^bsup>(Vr K v) n\<^esup> = {\\<^bsub>Vr K v\<^esub>})" apply (rule iffI) apply (simp add:r_apow_def, rule contrapos_pp, simp+, frule apos_neq_minf[of "n"], cut_tac mem_ant[of "n"], simp, erule exE, simp, simp add:ant_0[THEN sym], thin_tac "n = ant z") apply (case_tac "z = 0", simp add:ant_0, simp add:r_apow_def, frule Vr_ring[of v], frule Ring.ring_one[of "Vr K v"], simp, simp add:Vr_0_f_0, simp add:Vr_1_f_1, frule value_of_one[of v], simp, simp add:value_of_zero, cut_tac n = z in zneq_aneq[of _ "0"], simp only:ant_0) apply (simp add:r_apow_def, frule_tac n = "na (ant z)" in n_value_idealTr[of v], simp add:na_def, simp, thin_tac "vp K v \<^bsup>\(Vr K v) (na (ant z))\<^esup> = {\\<^bsub>Vr K v\<^esub>}", frule Vr_ring[of v], frule Pg_in_Vr[of v], frule_tac n = "na (ant z)" in Ring.npClose[of "Vr K v" "Pg K v"], assumption) apply (frule_tac a = "(Pg K v)^\<^bsup>(Vr K v) (na (ant z))\<^esup>" in Ring.a_in_principal[of "Vr K v"], assumption, simp, frule Vr_integral[of "v"], frule val_Pg[of v], simp, (erule conjE)+, frule_tac n = "na (ant z)" in Idomain.idom_potent_nonzero[of "Vr K v" "Pg K v"], assumption+, simp add:Vr_0_f_0, simp) done lemma (in Corps) Vr_potent_eqTr1:"\valuation K v; 0 \ n; 0 \ m; (vp K v) \<^bsup>(Vr K v) n\<^esup> = (vp K v) \<^bsup>(Vr K v) m\<^esup>; m = 0\ \ n = m" (*** compare the value of the generator of each ideal ***) (** express each ideal as a principal ideal **) apply (frule Vr_ring[of v], simp add:r_apow_def, case_tac "n = 0", simp, case_tac "n = \", simp, frule val_Pg[of v], erule conjE, simp, erule conjE, rotate_tac -3, drule sym, frule Lv_pos[of v], simp, frule val_poss_mem_Vr[of v "Pg K v"], assumption+, drule sym, simp, simp add:Vr_0_f_0) apply (simp, drule sym, frule Ring.ring_one[of "Vr K v"], simp, frule n_value_x_1_nat[of v "1\<^sub>r\<^bsub>(Vr K v)\<^esub>" "na n"], assumption, simp add:an_na, simp add:Vr_1_f_1, frule n_val_valuation[of v], simp add:value_of_one[of "n_val K v"]) done lemma (in Corps) Vr_potent_eqTr2:"\valuation K v; (vp K v) \<^bsup>\(Vr K v) n\<^esup> = (vp K v) \<^bsup>\(Vr K v) m\<^esup>\ \ n = m" (** 1. express each ideal as a principal ideal **) apply (frule Vr_ring[of v], frule val_Pg[of v], simp, (erule conjE)+, rotate_tac -1, frule sym, thin_tac "v (Pg K v) = Lv K v", frule Lv_pos[of v], simp) apply (subgoal_tac "0 \ int n", subgoal_tac "0 \ int m", frule n_value_idealTr[of "v" "m"]) apply simp apply simp apply( thin_tac "vp K v \<^bsup>\(Vr K v) m\<^esup> = Vr K v \\<^sub>p (Pg K v^\<^bsup>(Vr K v) m\<^esup>)", frule n_value_idealTr[of "v" "n"], simp, simp, thin_tac "vp K v \<^bsup>\(Vr K v) n\<^esup> = Vr K v \\<^sub>p (Pg K v^\<^bsup>(Vr K v) m\<^esup>)", frule val_poss_mem_Vr[of "v" "Pg K v"], assumption+) (** 2. the value of generators should coincide **) apply (frule Lv_z[of v], erule exE, rotate_tac -4, drule sym, simp, frule eq_ideal_eq_val[of "v" "Pg K v^\<^bsup>(Vr K v) n\<^esup>" "Pg K v^\<^bsup>(Vr K v) m\<^esup>"]) apply (rule Ring.npClose, assumption+, rule Ring.npClose, assumption+) apply (simp only:Vr_exp_f_exp, simp add:val_exp_ring[THEN sym, of v "Pg K v"], thin_tac "Vr K v \\<^sub>p (Pg K v^\<^bsup>K n\<^esup>) = Vr K v \\<^sub>p (Pg K v^\<^bsup>K m\<^esup>)") apply (case_tac "n = 0", simp, case_tac "m = 0", simp, simp only:of_nat_0_less_iff[THEN sym, of "m"], simp only:asprod_amult a_z_z, simp only:ant_0[THEN sym], simp only:aeq_zeq, simp) apply (auto simp add: asprod_mult) done lemma (in Corps) Vr_potent_eq:"\valuation K v; 0 \ n; 0 \ m; (vp K v) \<^bsup>(Vr K v) n\<^esup> = (vp K v) \<^bsup>(Vr K v) m\<^esup>\ \ n = m" apply (frule n_val_valuation[of v], case_tac "m = 0", simp add:Vr_potent_eqTr1) apply (case_tac "n = 0", frule sym, thin_tac "vp K v\<^bsup> (Vr K v) n\<^esup> = vp K v\<^bsup> (Vr K v) m\<^esup>", frule Vr_potent_eqTr1[of v m n], assumption+, rule sym, assumption, frule vp_potent_zero[of "v" "n"], assumption+) apply (case_tac "n = \", simp, thin_tac "vp K v\<^bsup> (Vr K v) \\<^esup> = {\\<^bsub>Vr K v\<^esub>}", frule vp_potent_zero[THEN sym, of v m], assumption+, simp, simp, frule vp_potent_zero[THEN sym, of v "m"], assumption+, simp, thin_tac "vp K v\<^bsup> (Vr K v) m\<^esup> \ {\\<^bsub>Vr K v\<^esub>}") apply (frule aposs_na_poss[of "n"], subst less_le, simp, frule aposs_na_poss[of "m"], subst less_le, simp, simp add:r_apow_def, frule Vr_potent_eqTr2[of "v" "na n" "na m"], assumption+, thin_tac "vp K v \<^bsup>\(Vr K v) (na n)\<^esup> = vp K v \<^bsup>\(Vr K v) (na m)\<^esup>", simp add:aeq_nat_eq[THEN sym]) done text\the following two lemma (in Corps) s are used in completion of K\ lemma (in Corps) Vr_prime_maximalTr1:"\valuation K v; x \ carrier (Vr K v); Suc 0 < n\ \ x \\<^sub>r\<^bsub>(Vr K v)\<^esub> (x^\<^bsup>K (n - Suc 0)\<^esup>) \ (Vr K v) \\<^sub>p (x^\<^bsup>K n\<^esup>)" apply (frule Vr_ring[of v], subgoal_tac "x^\<^bsup>K n\<^esup> = x^\<^bsup>K (Suc (n - Suc 0))\<^esup>", simp del:Suc_pred, rotate_tac -1, drule sym) apply (subst Vr_tOp_f_tOp, assumption+, subst Vr_exp_f_exp[of v, THEN sym], assumption+, simp only:Ring.npClose, simp del:Suc_pred) apply (cut_tac field_is_ring, frule Ring.npClose[of K x "n - Suc 0"], frule Vr_mem_f_mem[of v x], assumption+, frule Vr_mem_f_mem[of v x], assumption+) apply (simp add:Ring.ring_tOp_commute[of K x "x^\<^bsup>K (n - Suc 0)\<^esup>"]) apply (rule Ring.a_in_principal, assumption) apply (frule Ring.npClose[of "Vr K v" x n], assumption, simp add:Vr_exp_f_exp) apply (simp only:Suc_pred) done lemma (in Corps) Vr_prime_maximalTr2:"\ valuation K v; x \ vp K v; x \ \; Suc 0 < n\ \ x \ Vr K v \\<^sub>p (x^\<^bsup>K n\<^esup>) \ x^\<^bsup>K (n - Suc 0)\<^esup> \ (Vr K v) \\<^sub>p (x^\<^bsup>K n\<^esup>)" apply (frule Vr_ring[of v]) apply (frule vp_mem_Vr_mem[of v x], assumption, frule Ring.npClose[of "Vr K v" x n], simp only:Vr_exp_f_exp) apply (cut_tac field_is_ring, cut_tac field_is_idom, frule Vr_mem_f_mem[of v x], assumption+, frule Idomain.idom_potent_nonzero[of K x n], assumption+) apply (rule conjI) apply (rule contrapos_pp, simp+) apply (frule val_Rxa_gt_a[of v "x^\<^bsup>K n\<^esup>" x], simp, simp add:Vr_exp_f_exp, assumption+) apply (simp add:val_exp_ring[THEN sym, of v x n]) apply (frule val_nonzero_z[of v x], assumption+, erule exE, simp add:asprod_amult a_z_z) apply (simp add:vp_mem_val_poss[of v x]) apply (rule contrapos_pp, simp+) apply (frule val_Rxa_gt_a[of v "x^\<^bsup>K n\<^esup>" "x^\<^bsup>K (n - Suc 0)\<^esup>"]) apply (simp, frule Ring.npClose[of "Vr K v" "x" "n - Suc 0"], assumption+) apply (simp add:Vr_exp_f_exp) apply (frule Ring.npClose[of "Vr K v" "x" "n - Suc 0"], assumption+, simp add:Vr_exp_f_exp, assumption) apply (simp add:val_exp_ring[THEN sym, of v x]) apply (simp add:vp_mem_val_poss[of "v" "x"]) apply (frule val_nonzero_z[of "v" "x"], assumption+, erule exE, simp add:asprod_amult a_z_z) done lemma (in Corps) Vring_prime_maximal:"\valuation K v; prime_ideal (Vr K v) I; I \ {\\<^bsub>Vr K v\<^esub>}\ \ maximal_ideal (Vr K v) I" apply (frule Vr_ring[of v], frule Ring.prime_ideal_proper[of "Vr K v" "I"], assumption+, frule Ring.prime_ideal_ideal[of "Vr K v" "I"], assumption+, frule ideal_pow_vp[of v I], frule n_value_idealTr[of "v" "na (n_val K v (Ig K v I))"], simp, simp, assumption+) apply (case_tac "na (n_val K v (Ig K v I)) = 0", simp, frule Ring.Rxa_one[of "Vr K v"], simp, frule Suc_leI[of "0" "na (n_val K v (Ig K v I))"], thin_tac "0 < na (n_val K v (Ig K v I))") apply (case_tac "na (n_val K v (Ig K v I)) = Suc 0", simp, frule Pg_in_Vr[of v]) apply (frule vp_maximal[of v], frule Ring.maximal_ideal_ideal[of "Vr K v" "vp K v"], assumption+, subst Ring.idealprod_whole_r[of "Vr K v" "vp K v"], assumption+) apply (rotate_tac -1, drule not_sym, frule le_neq_implies_less[of "Suc 0" "na (n_val K v (Ig K v I))"], assumption+, thin_tac "Suc 0 \ na (n_val K v (Ig K v I))", thin_tac "Suc 0 \ na (n_val K v (Ig K v I))", thin_tac "Vr K v \\<^sub>p 1\<^sub>r\<^bsub>Vr K v\<^esub> = carrier (Vr K v)") apply (frule val_Pg[of v], simp, (erule conjE)+, frule Lv_pos[of v], rotate_tac -2, drule sym) apply (frule val_poss_mem_Vr[of "v" "Pg K v"], frule vp_mem_val_poss[THEN sym, of "v" "Pg K v"], assumption+, simp) apply (frule Vr_prime_maximalTr2[of v "Pg K v" "na (n_val K v (Ig K v I))"], simp add:vp_mem_val_poss[of v "Pg K v"], assumption+, erule conjE) apply (frule Ring.npMulDistr[of "Vr K v" "Pg K v" "na 1" "na (n_val K v (Ig K v I)) - Suc 0"], assumption+, simp add:na_1) apply (rotate_tac 8, drule sym) apply (frule Ring.a_in_principal[of "Vr K v" "Pg K v^\<^bsup>(Vr K v) (na (n_val K v (Ig K v I)))\<^esup>"], simp add:Ring.npClose) apply (simp add:Vr_exp_f_exp[of "v"]) apply (simp add:Ring.ring_l_one[of "Vr K v"]) apply (frule n_value_idealTr[THEN sym, of v "na (n_val K v (Ig K v I))"], simp) apply (simp add:Vr_exp_f_exp) apply (rotate_tac 6, drule sym, simp) apply (thin_tac "I \ carrier (Vr K v)", thin_tac "I = vp K v \<^bsup>\(Vr K v) (na (n_val K v (Ig K v I)))\<^esup>", thin_tac "v (Pg K v) = Lv K v", thin_tac "(Vr K v) \\<^sub>p ((Pg K v) \\<^sub>r\<^bsub>(Vr K v)\<^esub> ((Pg K v)^\<^bsup>K (na ((n_val K v) (Ig K v I)) - (Suc 0))\<^esup>)) = I", thin_tac "Pg K v \ carrier K", thin_tac "Pg K v \ \", thin_tac "Pg K v^\<^bsup>K (na ((n_val K v) (Ig K v I)))\<^esup> = Pg K v \\<^sub>r\<^bsub>Vr K v\<^esub> Pg K v^\<^bsup>K ((na ((n_val K v) (Ig K v I))) - Suc 0)\<^esup>") apply (simp add:prime_ideal_def, erule conjE, drule_tac x = "Pg K v" in bspec, assumption, drule_tac x = "Pg K v^\<^bsup>K (na (n_val K v (Ig K v I)) - Suc 0)\<^esup> " in bspec) apply (simp add:Vr_exp_f_exp[THEN sym, of v]) apply (rule Ring.npClose[of "Vr K v" "Pg K v"], assumption+) apply simp done text\From the above lemma (in Corps) , we see that a valuation ring is of dimension one.\ lemma (in Corps) field_frac1:"\1\<^sub>r \ \; x \ carrier K\ \ x = x \\<^sub>r ((1\<^sub>r)\<^bsup>\K\<^esup>)" by (simp add:invf_one, cut_tac field_is_ring, simp add:Ring.ring_r_one[THEN sym]) lemma (in Corps) field_frac2:"\x \ carrier K; x \ \\ \ x = (1\<^sub>r) \\<^sub>r ((x\<^bsup>\K\<^esup>)\<^bsup>\K\<^esup>)" by (cut_tac field_is_ring, simp add:field_inv_inv, simp add:Ring.ring_l_one[THEN sym]) lemma (in Corps) val_nonpos_inv_pos:"\valuation K v; x \ carrier K; \ 0 \ (v x)\ \ 0 < (v (x\<^bsup>\K\<^esup>))" by (case_tac "x = \\<^bsub>K\<^esub>", simp add:value_of_zero, frule Vr_ring[of v], simp add:aneg_le[of "0" "v x"], frule value_of_inv[THEN sym, of v x], assumption+, frule aless_minus[of "v x" "0"], simp) lemma (in Corps) frac_Vr_is_K:"\valuation K v; x \ carrier K\ \ \s\carrier (Vr K v). \t\carrier (Vr K v) - {\}. x = s \\<^sub>r (t\<^bsup>\K\<^esup>)" apply (frule Vr_ring[of v], frule has_val_one_neq_zero[of v]) apply (case_tac "x = \\<^bsub>K\<^esub>", frule Ring.ring_one[of "Vr K v"], frule field_frac1[of x], simp only:Vr_1_f_1, frule Ring.ring_zero[of "Vr K v"], simp add:Vr_0_f_0 Vr_1_f_1, blast) apply (case_tac "0 \ (v x)", frule val_pos_mem_Vr[THEN sym, of v x], assumption+, simp, frule field_frac1[of x], assumption+, frule has_val_one_neq_zero[of v], frule Ring.ring_one[of "Vr K v"], simp only:Vr_1_f_1, blast) apply (frule val_nonpos_inv_pos[of v x], assumption+, cut_tac invf_inv[of x], erule conjE, frule val_poss_mem_Vr[of v "x\<^bsup>\K\<^esup>"], assumption+) apply (frule Ring.ring_one[of "Vr K v"], simp only:Vr_1_f_1, frule field_frac2[of x], assumption+) apply (cut_tac invf_closed1[of x], blast, simp+) done lemma (in Corps) valuations_eqTr1:"\valuation K v; valuation K v'; Vr K v = Vr K v'; \x\carrier (Vr K v). v x = v' x\ \ v = v'" apply (rule funcset_eq [of _ "carrier K"], simp add:valuation_def, simp add:valuation_def, rule ballI, frule_tac x = x in frac_Vr_is_K[of v], assumption+, (erule bexE)+, simp, erule conjE) apply (frule_tac x = t in Vr_mem_f_mem[of v'], assumption, cut_tac x = t in invf_closed1, simp, simp, erule conjE) apply (frule_tac x = s in Vr_mem_f_mem[of "v'"], assumption+, simp add:val_t2p, simp add:value_of_inv) done lemma (in Corps) ridmap_rhom:"\ valuation K v; valuation K v'; carrier (Vr K v) \ carrier (Vr K v')\ \ ridmap (Vr K v) \ rHom (Vr K v) (Vr K v')" apply (frule Vr_ring[of "v"], frule Vr_ring[of "v'"], subst rHom_def, simp, rule conjI) apply (simp add:aHom_def, rule conjI, rule Pi_I, simp add:ridmap_def subsetD, simp add:ridmap_def restrict_def extensional_def, (rule ballI)+, frule Ring.ring_is_ag[of "Vr K v"], simp add:aGroup.ag_pOp_closed, simp add:Vr_pOp_f_pOp subsetD) apply (rule conjI, (rule ballI)+, simp add:ridmap_def, simp add:Ring.ring_tOp_closed, simp add:Vr_tOp_f_tOp subsetD, frule Ring.ring_one[of "Vr K v"], frule Ring.ring_one[of "Vr K v'"], simp add:Vr_1_f_1, simp add:ridmap_def ) done lemma (in Corps) contract_ideal:"\valuation K v; valuation K v'; carrier (Vr K v) \ carrier (Vr K v')\ \ ideal (Vr K v) (carrier (Vr K v) \ vp K v')" apply (frule_tac ridmap_rhom[of "v" "v'"], assumption+, frule Vr_ring[of "v"], frule Vr_ring[of "v'"]) apply (cut_tac TwoRings.i_contract_ideal[of "Vr K v" "Vr K v'" "ridmap (Vr K v)" "vp K v'"], subgoal_tac "(i_contract (ridmap (Vr K v)) (Vr K v) (Vr K v') (vp K v')) = (carrier (Vr K v) \ vp K v')") apply simp apply(thin_tac "ideal (Vr K v) (i_contract (ridmap (Vr K v)) (Vr K v) (Vr K v') (vp K v'))", simp add:i_contract_def invim_def ridmap_def, blast) apply (simp add:TwoRings_def TwoRings_axioms_def, simp) apply (simp add:vp_ideal) done lemma (in Corps) contract_prime:"\valuation K v; valuation K v'; carrier (Vr K v) \ carrier (Vr K v')\ \ prime_ideal (Vr K v) (carrier (Vr K v) \ vp K v')" apply (frule_tac ridmap_rhom[of "v" "v'"], assumption+, frule Vr_ring[of "v"], frule Vr_ring[of "v'"], cut_tac TwoRings.i_contract_prime[of "Vr K v" "Vr K v'" "ridmap (Vr K v)" "vp K v'"]) apply (subgoal_tac "(i_contract (ridmap (Vr K v)) (Vr K v) (Vr K v') (vp K v')) = (carrier (Vr K v) \ vp K v')", simp, thin_tac "prime_ideal (Vr K v) (i_contract (ridmap (Vr K v)) (Vr K v) (Vr K v') (vp K v'))", simp add:i_contract_def invim_def ridmap_def, blast) apply (simp add:TwoRings_def TwoRings_axioms_def, simp) apply (simp add:vp_prime) done (* \x\carrier K. 0 \ (v x) \ 0 \ (v' x) *) lemma (in Corps) valuation_equivTr:"\valuation K v; valuation K v'; x \ carrier K; 0 < (v' x); carrier (Vr K v) \ carrier (Vr K v')\ \ 0 \ (v x)" apply (rule contrapos_pp, simp+, frule val_nonpos_inv_pos[of "v" "x"], assumption+, case_tac "x = \\<^bsub>K\<^esub>", simp add:value_of_zero[of "v"]) apply ( cut_tac invf_closed1[of "x"], simp, erule conjE, frule aless_imp_le[of "0" "v (x\<^bsup>\K\<^esup>)"]) apply (simp add:val_pos_mem_Vr[of v "x\<^bsup>\K\<^esup>"], frule subsetD[of "carrier (Vr K v)" "carrier (Vr K v')" "x\<^bsup>\K\<^esup>"], assumption+, frule val_pos_mem_Vr[THEN sym, of "v'" "x\<^bsup>\K\<^esup>"], assumption+) apply (simp, simp add:value_of_inv[of "v'" "x"], cut_tac ale_minus[of "0" "- v' x"], thin_tac "0 \ - v' x", simp only:a_minus_minus, cut_tac aneg_less[THEN sym, of "v' x" "- 0"], simp, assumption, simp) done lemma (in Corps) contract_maximal:"\valuation K v; valuation K v'; carrier (Vr K v) \ carrier (Vr K v')\ \ maximal_ideal (Vr K v) (carrier (Vr K v) \ vp K v')" apply (frule Vr_ring[of "v"], frule Vr_ring[of "v'"], rule Vring_prime_maximal, assumption+, simp add:contract_prime) apply (frule vp_nonzero[of "v'"], frule vp_ideal[of "v'"], frule Ring.ideal_zero[of "Vr K v'" "vp K v'"], assumption+, frule sets_not_eq[of "vp K v'" "{\\<^bsub>(Vr K v')\<^esub>}"], simp add: singleton_sub[of "\\<^bsub>(Vr K v')\<^esub>" "carrier (Vr K v')"], erule bexE, simp add:Vr_0_f_0) apply (case_tac "a \ carrier (Vr K v)", blast, frule_tac x = a in vp_mem_Vr_mem[of "v'"], assumption+, frule_tac x = a in Vr_mem_f_mem[of "v'"], assumption+, subgoal_tac "a \ carrier (Vr K v)", blast, frule_tac x1 = a in val_pos_mem_Vr[THEN sym, of "v"], assumption+, simp, frule val_nonpos_inv_pos[of "v"], assumption+) apply (frule_tac y = "v (a\<^bsup>\K\<^esup>)" in aless_imp_le[of "0"], cut_tac x = a in invf_closed1, simp, frule_tac x = "a\<^bsup>\K\<^esup>" in val_poss_mem_Vr[of v], simp, assumption+) apply (frule_tac c = "a\<^bsup>\K\<^esup>" in subsetD[of "carrier (Vr K v)" "carrier (Vr K v')"], assumption+) apply ( frule_tac x = "a\<^bsup>\K\<^esup>" in val_pos_mem_Vr[of "v'"], simp, simp only:value_of_inv[of "v'"], simp, simp add:value_of_inv[of "v'"]) apply (frule_tac y = "- v' a" in ale_minus[of "0"], simp add:a_minus_minus, frule_tac x = a in vp_mem_val_poss[of "v'"], assumption+, simp) done section "Equivalent valuations" definition v_equiv :: "[_ , 'r \ ant, 'r \ ant] \ bool" where "v_equiv K v1 v2 \ n_val K v1 = n_val K v2" lemma (in Corps) valuation_equivTr1:"\valuation K v; valuation K v'; \x\carrier K. 0 \ (v x) \ 0 \ (v' x)\ \ carrier (Vr K v) \ carrier (Vr K v')" apply (frule Vr_ring[of "v"], frule Vr_ring[of "v'"]) apply (rule subsetI, case_tac "x = \\<^bsub>K\<^esub>", simp, simp add:Vr_def Sr_def, frule_tac x1 = x in val_pos_mem_Vr[THEN sym, of "v"], frule_tac x = x in Vr_mem_f_mem[of "v"], simp, frule_tac x = x in Vr_mem_f_mem[of "v"], assumption+) apply (drule_tac x = x in bspec, simp add:Vr_mem_f_mem) apply simp apply (subst val_pos_mem_Vr[THEN sym, of v'], assumption+, simp add:Vr_mem_f_mem, assumption+) done lemma (in Corps) valuation_equivTr2:"\valuation K v; valuation K v'; carrier (Vr K v) \ carrier (Vr K v'); vp K v = carrier (Vr K v) \ vp K v'\ \ carrier (Vr K v') \ carrier (Vr K v)" apply (frule Vr_ring[of "v"], frule Vr_ring[of "v'"]) apply (rule subsetI) apply (case_tac "x = \\<^bsub>(Vr K v')\<^esub>", simp, subst Vr_0_f_0[of "v'"], assumption+, subst Vr_0_f_0[of "v", THEN sym], assumption, simp add:Ring.ring_zero) apply (rule contrapos_pp, simp+) apply (frule_tac x = x in Vr_mem_f_mem[of "v'"], assumption+) apply (simp add:val_pos_mem_Vr[THEN sym, of "v"]) apply (cut_tac y = "v x" in aneg_le[of "0"], simp) apply (simp add:Vr_0_f_0[of "v'"]) apply (frule_tac x = "v x" in aless_minus[of _ "0"], simp, thin_tac "v x < 0", thin_tac "\ 0 \ v x") apply (simp add:value_of_inv[THEN sym, of "v"]) apply (cut_tac x = x in invf_closed1, simp, simp, erule conjE) apply (frule_tac x1 = "x\<^bsup>\K\<^esup>" in vp_mem_val_poss[THEN sym, of "v"], assumption, simp, erule conjE) apply (frule vp_ideal [of "v'"]) apply (frule_tac x = "x\<^bsup>\K\<^esup>" and r = x in Ring.ideal_ring_multiple[of "Vr K v'" "vp K v'"], assumption+) apply (frule_tac x = "x\<^bsup>\K\<^esup>" in vp_mem_Vr_mem[of "v'"], assumption+) apply (frule_tac x = x and y = "x\<^bsup>\K\<^esup>" in Ring.ring_tOp_commute[of "Vr K v'"], assumption+, simp, thin_tac "x \\<^sub>r\<^bsub>Vr K v'\<^esub> x\<^bsup>\ K\<^esup> = x\<^bsup>\ K\<^esup> \\<^sub>r\<^bsub>Vr K v'\<^esub> x") apply (simp add:Vr_tOp_f_tOp) apply (cut_tac x = x in linvf, simp, simp) apply (cut_tac field_is_ring, frule Ring.ring_one[of "K"]) apply (frule ideal_inc_elem0val_whole[of "v'" "1\<^sub>r" "vp K v'"], assumption+, simp add:value_of_one, assumption+) apply (frule vp_not_whole[of "v'"], simp) done lemma (in Corps) eq_carr_eq_Vring:" \valuation K v; valuation K v'; carrier (Vr K v) = carrier (Vr K v')\ \ Vr K v = Vr K v'" apply (simp add:Vr_def Sr_def) done lemma (in Corps) valuations_equiv:"\valuation K v; valuation K v'; \x\carrier K. 0 \ (v x) \ 0 \ (v' x)\ \ v_equiv K v v'" (** step0. preliminaries. **) apply (frule Vr_ring[of "v"], frule Vr_ring[of "v'"]) (** step1. show carrier (Vr K v) \ carrier (Vr K v') **) apply (frule valuation_equivTr1[of "v" "v'"], assumption+) (** step2. maximal_ideal (Vr K v) (carrier (Vr K v) \ (vp K v')). contract of the maximal ideal is prime, and a prime is maximal **) apply (frule contract_maximal [of "v" "v'"], assumption+) (** step3. Vring is a local ring, we have (vp K v) = (carrier (Vr K v) \ (vp K v')) **) apply (frule Vr_local[of "v" "(carrier (Vr K v) \ vp K v')"], assumption+) (** step4. show carrier (Vr K v') \ carrier (Vr K v) **) apply (frule valuation_equivTr2[of "v" "v'"], assumption+, frule equalityI[of "carrier (Vr K v)" "carrier (Vr K v')"], assumption+, thin_tac "carrier (Vr K v) \ carrier (Vr K v')", thin_tac "carrier (Vr K v') \ carrier (Vr K v)") (** step5. vp K v' = vp K v **) apply (frule vp_ideal[of "v'"], frule Ring.ideal_subset1[of "Vr K v'" "vp K v'"], assumption, simp add:Int_absorb1, thin_tac "\x\carrier K. 0 \ v x \ 0 \ v' x", thin_tac "vp K v' \ carrier (Vr K v')", thin_tac "ideal (Vr K v') (vp K v')", thin_tac "maximal_ideal (Vr K v) (vp K v')") (** step6. to show v_equiv K v v', we check whether the normal valuations derived from the valuations have the same value or not. if (Vr K (n_valuation K v)) = (Vr K (n_valuation K v')), then we have only to check the values of the elements in this valuation ring. We see (Vr K v) = (Vr K (n_valuation K G a i v)). **) apply (simp add:v_equiv_def, rule valuations_eqTr1[of "n_val K v" "n_val K v'"], (simp add:n_val_valuation)+, rule eq_carr_eq_Vring[of "n_val K v" "n_val K v'"], (simp add:n_val_valuation)+, subst Vr_n_val_Vr[THEN sym, of "v"], assumption+, subst Vr_n_val_Vr[THEN sym, of "v'"], assumption+) apply (rule ballI, frule n_val_valuation[of "v"], frule n_val_valuation[of "v'"], frule_tac x1 = x in val_pos_mem_Vr[THEN sym, of "n_val K v"], simp add:Vr_mem_f_mem, simp, frule Vr_n_val_Vr[THEN sym, of "v"], simp, thin_tac "carrier (Vr K (n_val K v)) = carrier (Vr K v')", frule_tac x1 = x in val_pos_mem_Vr[THEN sym, of "v'"], simp add:Vr_mem_f_mem, simp, frule_tac x = x in val_pos_n_val_pos[of "v'"], simp add:Vr_mem_f_mem, simp, frule_tac x = x in ideal_apow_n_val[of "v"], simp add:Vr_n_val_Vr[THEN sym, of "v"], simp) apply (frule eq_carr_eq_Vring[of "v" "v'"], assumption+, frule_tac x = x in ideal_apow_n_val[of "v'"], assumption, simp add:Vr_n_val_Vr[THEN sym, of "v"], thin_tac "Vr K v' \\<^sub>p x = vp K v'\<^bsup> (Vr K v') (n_val K v x)\<^esup>", frule_tac n = "n_val K v' x" and m = "n_val K v x" in Vr_potent_eq[of "v'"], assumption+, frule sym, assumption+) done lemma (in Corps) val_equiv_axiom1:"valuation K v \ v_equiv K v v" apply (simp add:v_equiv_def) done lemma (in Corps) val_equiv_axiom2:"\ valuation K v; valuation K v'; v_equiv K v v'\ \ v_equiv K v' v" apply (simp add:v_equiv_def) done lemma (in Corps) val_equiv_axiom3:"\ valuation K v; valuation K v'; valuation K v'; v_equiv K v v'; v_equiv K v' v''\ \ v_equiv K v v''" apply (simp add:v_equiv_def) done lemma (in Corps) n_val_equiv_val:"\ valuation K v\ \ v_equiv K v (n_val K v)" apply (frule valuations_equiv[of "v" "n_val K v"], simp add:n_val_valuation) apply (rule ballI, rule impI, simp add:val_pos_n_val_pos, assumption) done section "Prime divisors" definition prime_divisor :: "[_, 'b \ ant] \ ('b \ ant) set" ("(2P\<^bsub> _ _\<^esub>)" [96,97]96) where "P\<^bsub>K v\<^esub> = {v'. valuation K v' \ v_equiv K v v'}" definition prime_divisors :: "_ \ ('b \ ant) set set" ("Pds\" 96) where "Pds\<^bsub>K\<^esub> = {P. \v. valuation K v \ P = P\<^bsub> K v\<^esub> }" definition normal_valuation_belonging_to_prime_divisor :: "[_ , ('b \ ant) set] \ ('b \ ant)" ("(\\<^bsub>_ _\<^esub>)" [96,97]96) where "\\<^bsub>K P\<^esub> = n_val K (SOME v. v \ P)" lemma (in Corps) val_in_P_valuation:"\valuation K v; v' \ P\<^bsub>K v\<^esub>\ \ valuation K v'" apply (simp add:prime_divisor_def) done lemma (in Corps) vals_in_P_equiv:"\ valuation K v; v' \ P\<^bsub>K v\<^esub>\ \ v_equiv K v v'" apply (simp add:prime_divisor_def) done lemma (in Corps) v_in_prime_v:"valuation K v \ v \ P\<^bsub>K v\<^esub>" apply (simp add:prime_divisor_def, frule val_equiv_axiom1[of "v"], assumption+) done lemma (in Corps) some_in_prime_divisor:"valuation K v \ (SOME w. w \ P\<^bsub>K v\<^esub>) \ P\<^bsub>K v\<^esub>" apply (subgoal_tac "P\<^bsub> K v\<^esub> \ {}", rule nonempty_some[of "P\<^bsub> K v\<^esub>"], assumption+, frule v_in_prime_v[of "v"]) apply blast done lemma (in Corps) valuation_some_in_prime_divisor:"valuation K v \ valuation K (SOME w. w \ P\<^bsub>K v\<^esub>)" apply (frule some_in_prime_divisor[of "v"], simp add:prime_divisor_def) done lemma (in Corps) valuation_some_in_prime_divisor1:"P \ Pds \ valuation K (SOME w. w \ P)" apply (simp add:prime_divisors_def, erule exE) apply (simp add:valuation_some_in_prime_divisor) done lemma (in Corps) representative_of_pd_valuation: "P \ Pds \ valuation K (\\<^bsub>K P\<^esub>)" apply (simp add:prime_divisors_def, erule exE, erule conjE, simp add:normal_valuation_belonging_to_prime_divisor_def, frule_tac v = v in valuation_some_in_prime_divisor) apply (rule n_val_valuation, assumption+) done lemma (in Corps) some_in_P_equiv:"valuation K v \ v_equiv K v (SOME w. w \ P\<^bsub>K v\<^esub>)" apply (frule some_in_prime_divisor[of v]) apply (rule vals_in_P_equiv, assumption+) done lemma (in Corps) n_val_n_val1:"P \ Pds \ n_val K (\\<^bsub>K P\<^esub>) = (\\<^bsub>K P\<^esub>)" apply (simp add: normal_valuation_belonging_to_prime_divisor_def, frule valuation_some_in_prime_divisor1[of P]) apply (rule n_val_n_val[of "SOME v. v \ P"], assumption+) done lemma (in Corps) P_eq_val_equiv:"\valuation K v; valuation K v'\ \ (v_equiv K v v') = (P\<^bsub>K v\<^esub> = P\<^bsub>K v'\<^esub>)" apply (rule iffI, rule equalityI, rule subsetI, simp add:prime_divisor_def, erule conjE, frule val_equiv_axiom2[of "v" "v'"], assumption+, rule val_equiv_axiom3[of "v'" "v"], assumption+, rule subsetI, simp add:prime_divisor_def, erule conjE) apply (rule val_equiv_axiom3[of "v" "v'"], assumption+, frule v_in_prime_v[of "v"], simp, thin_tac "P\<^bsub>K v\<^esub> = P\<^bsub>K v'\<^esub>", simp add:prime_divisor_def, rule val_equiv_axiom2[of "v'" "v"], assumption+) done lemma (in Corps) unique_n_valuation:"\ P \ Pds\<^bsub>K\<^esub>; P' \ Pds\ \ (P = P') = (\\<^bsub>K P\<^esub> = \\<^bsub>K P'\<^esub>)" apply (rule iffI, simp) apply (simp add:prime_divisors_def, (erule exE)+, (erule conjE)+) apply (simp add:normal_valuation_belonging_to_prime_divisor_def, frule_tac v = v in some_in_P_equiv, frule_tac v = va in some_in_P_equiv, subgoal_tac "v_equiv K (SOME w. w \ P\<^bsub>K v\<^esub>) (SOME w. w \ P\<^bsub>K va\<^esub>)") apply (frule_tac v = v in some_in_prime_divisor, frule_tac v = va in some_in_prime_divisor, frule_tac v = v and v' = "SOME w. w \ P\<^bsub>K v\<^esub>" and v'' = "SOME w. w \ P\<^bsub>K va\<^esub>" in val_equiv_axiom3) apply (simp add:prime_divisor_def, simp add:prime_divisor_def, assumption+, frule_tac v = va and v' = "SOME w. w \ P\<^bsub>K va\<^esub>" in val_equiv_axiom2, simp add:prime_divisor_def, assumption+) apply (frule_tac v = v and v' = "SOME w. w \ P\<^bsub>K va\<^esub>" and v'' = va in val_equiv_axiom3, simp add:prime_divisor_def, simp add:prime_divisor_def, assumption+, frule_tac v = v and v' = va in P_eq_val_equiv, assumption+) apply simp apply (simp add:v_equiv_def) done lemma (in Corps) n_val_representative:"P \ Pds \ (\\<^bsub>K P\<^esub>) \ P" apply (simp add:prime_divisors_def, erule exE, erule conjE, simp add:normal_valuation_belonging_to_prime_divisor_def, frule_tac v = v in valuation_some_in_prime_divisor, frule_tac v = "SOME w. w \ P\<^bsub>K v\<^esub>" in n_val_equiv_val, frule_tac v = v in some_in_P_equiv, frule_tac v = v and v' = "SOME w. w \ P\<^bsub> K v\<^esub>" and v'' = "n_val K (SOME w. w \ P\<^bsub>K v\<^esub>)" in val_equiv_axiom3, assumption+, frule_tac v = v in n_val_valuation, simp add:prime_divisor_def, simp add:n_val_valuation) done lemma (in Corps) val_equiv_eq_pdiv:"\ P \ Pds\<^bsub>K\<^esub>; P'\ Pds\<^bsub>K\<^esub>; valuation K v; valuation K v'; v_equiv K v v'; v \ P; v' \ P' \ \ P = P'" apply (simp add:prime_divisors_def, (erule exE)+, (erule conjE)+) apply (rename_tac w w', frule_tac v = w in vals_in_P_equiv[of _ "v"], simp, frule_tac v = w' in vals_in_P_equiv[of _ "v'"], simp, frule_tac v = w and v' = v and v'' = v' in val_equiv_axiom3, assumption+, frule_tac v = w' in val_equiv_axiom2[of _ "v'"], assumption+, frule_tac v = w and v' = v' and v'' = w' in val_equiv_axiom3, assumption+) apply simp+ apply (simp add:P_eq_val_equiv) done lemma (in Corps) distinct_p_divisors:"\ P \ Pds\<^bsub>K\<^esub>; P' \ Pds\<^bsub>K\<^esub>\ \ (\ P = P') = (\ v_equiv K (\\<^bsub>K P\<^esub>) (\\<^bsub>K P'\<^esub>))" apply (rule iffI, rule contrapos_pp, simp+, frule val_equiv_eq_pdiv[of "P" "P'" "\\<^bsub>K P\<^esub>" "\\<^bsub>K P'\<^esub>"], assumption+, simp add: representative_of_pd_valuation, simp add: representative_of_pd_valuation, assumption) apply (rule n_val_representative[of "P"], assumption, rule n_val_representative[of "P'"], assumption, simp, rule contrapos_pp, simp+, frule sym, thin_tac "P = P'", simp, frule representative_of_pd_valuation[of P], frule val_equiv_axiom1[of "\\<^bsub>K P\<^esub>"], simp) done section "Approximation" definition valuations :: "[_ , nat, nat \ ('r \ ant)] \ bool" where "valuations K n vv \ (\j \ n. valuation K (vv j))" definition vals_nonequiv :: "[_, nat, nat \ ('r \ ant)] \ bool" where "vals_nonequiv K n vv \ valuations K n vv \ (\j\n. \l \ n. j \ l \ \ (v_equiv K (vv j) (vv l)))" definition Ostrowski_elem :: "[_, nat, nat \ ('b \ ant), 'b] \ bool" where "Ostrowski_elem K n vv x \ (0 < (vv 0 (1\<^sub>r\<^bsub>K\<^esub> \\<^bsub>K\<^esub> (-\<^sub>a\<^bsub>K\<^esub> x)))) \ (\j\nset (Suc 0) n. 0 < (vv j x))" (** vv 0, vv 1, vv 2,\, vv n are valuations **) lemma (in Corps) Ostrowski_elem_0:"\vals_nonequiv K n vv; x \ carrier K; Ostrowski_elem K n vv x\ \ 0 < (vv 0 (1\<^sub>r \ (-\<^sub>a x)))" apply (simp add:Ostrowski_elem_def) done lemma (in Corps) Ostrowski_elem_Suc:"\vals_nonequiv K n vv; x \ carrier K; Ostrowski_elem K n vv x; j \ nset (Suc 0) n\ \ 0 < (vv j x)" apply (simp add:Ostrowski_elem_def) done lemma (in Corps) vals_nonequiv_valuation:"\vals_nonequiv K n vv; m \ n\ \ valuation K (vv m)" apply (simp add:vals_nonequiv_def, erule conjE) apply (thin_tac "\j\n. \l\ n. j \ l \ \ v_equiv K (vv j) (vv l)") apply (simp add:valuations_def) done lemma (in Corps) vals_nonequiv:"\ vals_nonequiv K (Suc (Suc n)) vv; i \ (Suc (Suc n)); j \ (Suc (Suc n)); i \ j\ \ \ (v_equiv K (vv i) (vv j))" apply (simp add:vals_nonequiv_def) done lemma (in Corps) skip_vals_nonequiv:"vals_nonequiv K (Suc (Suc n)) vv \ vals_nonequiv K (Suc n) (compose {l. l \ (Suc n)} vv (skip j))" apply (subst vals_nonequiv_def) apply (rule conjI) apply (subst valuations_def, rule allI, rule impI, simp add:compose_def) apply (cut_tac l = ja and n = "Suc n" in skip_mem[of _ _ "j"], simp, frule_tac m = "skip j ja" in vals_nonequiv_valuation[of "Suc (Suc n)" "vv"], simp, assumption) apply ((rule allI, rule impI)+, rule impI, cut_tac l = ja and n = "Suc n" in skip_mem[of _ _ "j"], simp, cut_tac l = l and n = "Suc n" in skip_mem[of _ _ "j"], simp+) apply (cut_tac i = ja and j = l in skip_inj[of _ "Suc n" _ "j"], simp+, simp add:compose_def, rule_tac i = "skip j ja" and j = "skip j l" in vals_nonequiv[of "n"], assumption+) done lemma (in Corps) not_v_equiv_reflex:"\valuation K v; valuation K v'; \ v_equiv K v v'\ \ \ v_equiv K v' v " apply (simp add:v_equiv_def) done lemma (in Corps) nonequiv_ex_Ostrowski_elem:"\valuation K v; valuation K v'; \ v_equiv K v v'\ \ \x\carrier K. 0 \ (v x) \ (v' x) < 0" apply (subgoal_tac "\ (\x\carrier K. 0 \ (v x) \ 0 \ (v' x))") prefer 2 apply (rule contrapos_pp, simp+, frule valuations_equiv[of "v" "v'"], assumption+, simp add:val_equiv_axiom2[of v v']) apply (simp, erule bexE, erule conjE, simp add:aneg_le) apply blast done lemma (in Corps) field_op_minus:"\a \ carrier K; b \ carrier K; b \ \\ \ -\<^sub>a (a \\<^sub>r (b\<^bsup>\K\<^esup>)) = (-\<^sub>a a) \\<^sub>r (b\<^bsup>\K\<^esup>)" apply (cut_tac invf_closed1[of "b"], simp, erule conjE, cut_tac field_is_ring, simp add:Ring.ring_inv1[of "K" "a" "b\<^bsup>\K\<^esup>"], simp) done lemma (in Corps) field_one_plus_frac1:"\a \ carrier K; b \ carrier K; b \ \\ \ 1\<^sub>r \ (a \\<^sub>r (b\<^bsup>\K\<^esup>)) = (b \ a) \\<^sub>r (b\<^bsup>\K\<^esup>)" apply (cut_tac field_is_ring, cut_tac invf_closed1[of b], simp+, erule conjE, cut_tac field_is_idom) apply (rule Idomain.idom_mult_cancel_r[of K "1\<^sub>r \ (a \\<^sub>r (b\<^bsup>\K\<^esup>))" "(b \ a) \\<^sub>r (b\<^bsup>\K\<^esup>)" "b"], assumption+, frule Idomain.idom_is_ring[of "K"], frule Ring.ring_is_ag[of "K"], rule aGroup.ag_pOp_closed [of "K"], assumption+, simp add:Ring.ring_one,rule Ring.ring_tOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+, frule Ring.ring_is_ag[of "K"], rule aGroup.ag_pOp_closed, assumption+, subst Ring.ring_distrib2[of "K" "b"], assumption+, simp add:Ring.ring_one, simp add:Ring.ring_tOp_closed, simp add:Ring.ring_l_one) thm Ring.ring_distrib2[of K "b\<^bsup>\K\<^esup>"] apply (subst Ring.ring_distrib2[of K "b\<^bsup>\K\<^esup>"], assumption+, simp add:Ring.ring_tOp_commute[of "K" "b" "b\<^bsup>\K\<^esup>"], subst linvf[of "b"], simp, subst Ring.ring_distrib2[of "K" "b"], assumption+, simp add:Ring.ring_one, simp add:Ring.ring_tOp_closed, simp add:Ring.ring_l_one, simp) done lemma (in Corps) field_one_plus_frac2:"\a \ carrier K; b \ carrier K; a \ b \ \\ \ 1\<^sub>r \ (-\<^sub>a (a \\<^sub>r (a \ b)\<^bsup>\K\<^esup>)) = b \\<^sub>r ((a \ b)\<^bsup>\K\<^esup>)" apply (frule field_op_minus[of "a" "a \ b"], cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], simp add:aGroup.ag_pOp_closed, assumption, simp, thin_tac "-\<^sub>a (a \\<^sub>r (a \ b)\<^bsup>\ K\<^esup>) = (-\<^sub>a a) \\<^sub>r (a \ b)\<^bsup>\ K\<^esup>") apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule aGroup.ag_mOp_closed[of "K" "a"], assumption, frule field_one_plus_frac1[of "-\<^sub>a a" "a \ b"], simp add:aGroup.ag_pOp_closed, simp, simp, thin_tac "1\<^sub>r \ (-\<^sub>a a) \\<^sub>r (a \ b)\<^bsup>\ K\<^esup> = (a \ b \ -\<^sub>a a) \\<^sub>r (a \ b)\<^bsup>\ K\<^esup>", simp add:aGroup.ag_pOp_assoc[of "K" "a" "b" "-\<^sub>a a"], simp add:aGroup.ag_pOp_commute[of "K" "b" "-\<^sub>a a"], simp add:aGroup.ag_pOp_assoc[THEN sym], simp add:aGroup.ag_r_inv1, simp add:aGroup.ag_l_zero) done lemma (in Corps) field_one_plus_frac3:"\x \ carrier K; x \ 1\<^sub>r; 1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x) \ \ \ \ 1\<^sub>r \ -\<^sub>a x \\<^sub>r (1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\ K\<^esup> = (1\<^sub>r \ -\<^sub>a x^\<^bsup>K (Suc (Suc 0))\<^esup>) \\<^sub>r (1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\ K\<^esup>" apply (cut_tac field_is_ring, frule Ring.ring_is_ag, frule Ring.ring_one, cut_tac invf_closed1[of "1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)"], simp, erule conjE) apply (subst Ring.ring_inv1_1, assumption+, subst field_one_plus_frac1[of "-\<^sub>a x" "1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)"]) apply (rule aGroup.ag_mOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, subst Ring.ring_distrib1, assumption+, rule aGroup.ag_mOp_closed, assumption+) apply (simp add:Ring.ring_r_one) apply (simp add:Ring.ring_inv1_2[THEN sym, of K x x]) apply (subgoal_tac "1\<^sub>r \ (x \ -\<^sub>a x \\<^sub>r x) \ -\<^sub>a x = 1\<^sub>r \ -\<^sub>a x^\<^bsup>K (Suc (Suc 0))\<^esup>", simp, frule Ring.ring_tOp_closed[of K x x], assumption+) apply (frule Ring.ring_tOp_closed[of K x x], assumption+, frule aGroup.ag_mOp_closed[of K "x \\<^sub>r x"], assumption+, frule aGroup.ag_mOp_closed[of K x], assumption+) apply (subst aGroup.ag_pOp_assoc, assumption+, rule aGroup.ag_pOp_closed, assumption+) apply (rule aGroup.ag_pOp_add_l[of K "x \ -\<^sub>a x \\<^sub>r x \ -\<^sub>a x" "-\<^sub>a x^\<^bsup>K (Suc (Suc 0))\<^esup>" "1\<^sub>r"], assumption+, (rule aGroup.ag_pOp_closed, assumption+)+, rule aGroup.ag_mOp_closed, assumption+, rule Ring.npClose, assumption+, subst aGroup.ag_pOp_commute, assumption+, simp add:aGroup.ag_pOp_assoc aGroup.ag_r_inv1 aGroup.ag_r_zero) apply (simp add:Ring.ring_l_one) apply simp apply (rule aGroup.ag_pOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+, rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed[of K x], assumption+) done lemma (in Corps) OstrowskiTr1:"\ valuation K v; s \ carrier K; t \ carrier K; 0 \ (v s); v t < 0\ \ s \ t \ \" apply (rule contrapos_pp, simp+, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], simp only:aGroup.ag_plus_zero[THEN sym, of "K" "s" "t"]) apply (simp add:val_minus_eq[of "v" "t"]) done lemma (in Corps) OstrowskiTr2:"\valuation K v; s \ carrier K; t \ carrier K; 0 \ (v s); v t < 0\ \ 0 < (v (1\<^sub>r \ (-\<^sub>a ((t \\<^sub>r ((s \ t)\<^bsup>\K\<^esup>))))))" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule OstrowskiTr1[of "v" "s" "t"], assumption+, frule field_one_plus_frac2[of "t" "s"], assumption+, simp add:aGroup.ag_pOp_commute) apply (subst aGroup.ag_pOp_commute[of "K" "s" "t"], assumption+, simp, simp add:aGroup.ag_pOp_commute[of "K" "t" "s"], thin_tac "1\<^sub>r \ -\<^sub>a (t \\<^sub>r (s \ t)\<^bsup>\ K\<^esup>) = s \\<^sub>r (s \ t)\<^bsup>\ K\<^esup>", frule aGroup.ag_pOp_closed[of "K" "s" "t"], assumption+, cut_tac invf_closed1[of "s \ t"], simp, erule conjE) apply (simp add:val_t2p, simp add:value_of_inv, frule aless_le_trans[of "v t" "0" "v s"], assumption+, frule value_less_eq[THEN sym, of v t s], assumption+, simp add:aGroup.ag_pOp_commute, frule aless_diff_poss[of "v t" "v s"], simp add:diff_ant_def, simp) done lemma (in Corps) OstrowskiTr3:"\valuation K v; s \ carrier K; t \ carrier K; 0 \ (v t); v s < 0\ \ 0 < (v (t \\<^sub>r (( s \ t)\<^bsup>\K\<^esup>)))" apply (frule aless_le_trans[of "v s" "0" "v t"], assumption+, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule aGroup.ag_pOp_closed[of "K" "s" "t"], assumption+, frule OstrowskiTr1[of v t s], assumption+, frule value_less_eq[THEN sym, of v s t], assumption+) apply (simp add:aGroup.ag_pOp_commute[of K t s], cut_tac invf_closed1[of "s \ t"], simp) apply ( erule conjE, simp add:val_t2p[of v], simp add:value_of_inv) apply (cut_tac aless_diff_poss[of "v s" "v t"], simp add:diff_ant_def, simp+) done lemma (in Corps) restrict_Ostrowski_elem:"\ x \ carrier K; Ostrowski_elem K (Suc (Suc n)) vv x\ \ Ostrowski_elem K (Suc n) vv x" apply (simp add:Ostrowski_elem_def, erule conjE, rule ballI, simp add:nset_def, insert lessI [of "Suc n"]) done lemma (in Corps) restrict_vals_nonequiv:"vals_nonequiv K (Suc (Suc n)) vv \ vals_nonequiv K (Suc n) vv" apply (simp add:vals_nonequiv_def, erule conjE, simp add:valuations_def) done lemma (in Corps) restrict_vals_nonequiv1:"vals_nonequiv K (Suc (Suc n)) vv \ vals_nonequiv K (Suc n) (compose {h. h \ (Suc n)} vv (skip 1))" apply (simp add:vals_nonequiv_def, (erule conjE)+, rule conjI, thin_tac "\j\Suc (Suc n). \l\Suc (Suc n). j \ l \ \ v_equiv K (vv j) (vv l)", simp add:valuations_def, rule allI, rule impI, simp add:compose_def skip_def nset_def) apply ((rule allI, rule impI)+, rule impI) apply (simp add:compose_def skip_def nset_def) done lemma (in Corps) restrict_vals_nonequiv2:"\vals_nonequiv K (Suc (Suc n)) vv\ \ vals_nonequiv K (Suc n) (compose {j. j \ (Suc n)} vv (skip 2))" apply (simp add:vals_nonequiv_def, (erule conjE)+, rule conjI, thin_tac "\j\Suc (Suc n). \l\Suc (Suc n). j \ l \ \ v_equiv K (vv j) (vv l)", simp add:valuations_def, rule allI, rule impI) apply (simp add:compose_def skip_def nset_def, (rule allI, rule impI)+, rule impI, simp add:compose_def skip_def nset_def) done lemma (in Corps) OstrowskiTr31:"\valuation K v; s \ carrier K; 0 < (v (1\<^sub>r \ (-\<^sub>a s)))\ \ s \ \" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"]) apply (rule contrapos_pp, simp+) apply (simp add:aGroup.ag_inv_zero, frule Ring.ring_one[of "K"], simp add:aGroup.ag_r_zero) apply (simp add:value_of_one) done lemma (in Corps) OstrowskiTr32:"\valuation K v; s \ carrier K; 0 < (v (1\<^sub>r \ (-\<^sub>a s)))\ \ 0 \ (v s)" apply (rule contrapos_pp, simp+, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], simp add:aneg_le, frule has_val_one_neq_zero[of "v"]) apply (frule OstrowskiTr31[of v s], assumption+, frule not_sym, frule Ring.ring_one[of "K"]) apply (frule value_less_eq[THEN sym, of v "-\<^sub>a s" "1\<^sub>r"], simp add:aGroup.ag_mOp_closed, assumption+, simp add:val_minus_eq) apply (simp add:value_of_one, frule aGroup.ag_mOp_closed[of "K" "s"], assumption+, simp add:aGroup.ag_pOp_commute[of "K" "-\<^sub>a s" "1\<^sub>r"], simp add:val_minus_eq) done lemma (in Corps) OstrowskiTr4:"\valuation K v; s \ carrier K; t \ carrier K; 0 < (v (1\<^sub>r \ (-\<^sub>a s))); 0 < (v (1\<^sub>r \ (-\<^sub>a t)))\ \ 0 < (v (1\<^sub>r \ (-\<^sub>a (s \\<^sub>r t))))" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Ring.ring_one[of "K"]) apply (subgoal_tac "1\<^sub>r \ (-\<^sub>a (s \\<^sub>r t)) = 1\<^sub>r \ (-\<^sub>a s) \ (s \\<^sub>r (1\<^sub>r \ (-\<^sub>a t)))", simp, thin_tac "1\<^sub>r \ -\<^sub>a (s \\<^sub>r t) = 1\<^sub>r \ -\<^sub>a s \ s \\<^sub>r (1\<^sub>r \ -\<^sub>a t)") apply (frule aGroup.ag_mOp_closed[of K s], assumption+, frule aGroup.ag_pOp_closed[of K "1\<^sub>r" "-\<^sub>a s"], assumption+, frule aGroup.ag_mOp_closed[of "K" "t"], assumption+, frule aGroup.ag_pOp_closed[of "K" "1\<^sub>r" "-\<^sub>a t"], assumption+, frule Ring.ring_tOp_closed[of "K" "s" "1\<^sub>r \ (-\<^sub>a t)"], assumption+, frule amin_le_plus[of v "1\<^sub>r \ (-\<^sub>a s)" "s \\<^sub>r (1\<^sub>r \ (-\<^sub>a t))"], assumption+) apply (frule amin_gt[of "0" "v (1\<^sub>r \ -\<^sub>a s)" "v (s \\<^sub>r (1\<^sub>r \ -\<^sub>a t))"]) apply (simp add:val_t2p, frule OstrowskiTr32[of v s], assumption+, rule aadd_pos_poss[of "v s" "v (1\<^sub>r \ -\<^sub>a t)"], assumption+, simp add:Ring.ring_distrib1) apply (frule aGroup.ag_mOp_closed[of K t], assumption, simp add:Ring.ring_distrib1 Ring.ring_r_one, frule aGroup.ag_mOp_closed[of K s], assumption+, subst aGroup.pOp_assocTr43, assumption+, simp add:Ring.ring_tOp_closed, simp add:aGroup.ag_l_inv1 aGroup.ag_r_zero, simp add:Ring.ring_inv1_2) done lemma (in Corps) OstrowskiTr5:"\ vals_nonequiv K (Suc (Suc n)) vv; s \ carrier K; t \ carrier K; 0 \ (vv (Suc 0)) s \ 0 \ (vv (Suc (Suc 0))) t; Ostrowski_elem K (Suc n) (compose {j. j \ (Suc n)} vv (skip 1)) s; Ostrowski_elem K (Suc n) (compose {j. j \ (Suc n)} vv (skip 2)) t\ \ Ostrowski_elem K (Suc (Suc n)) vv (s \\<^sub>r t)" apply (erule conjE, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule_tac x = s and y = t in Ring.ring_tOp_closed[of "K"], assumption+, frule skip_vals_nonequiv[of n "vv" "1"], frule skip_vals_nonequiv[of n "vv" "2"], subst Ostrowski_elem_def, rule conjI) apply (rule OstrowskiTr4, simp add:vals_nonequiv_valuation[of "Suc (Suc n)" "vv" "0"], assumption+, frule Ostrowski_elem_0[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 1)" "s"], assumption+, simp add:skip_def compose_def, frule Ostrowski_elem_0[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 2)" "t"], assumption+, simp add:skip_def compose_def) apply (rule ballI, case_tac "j = Suc 0", frule_tac j = " Suc 0" in Ostrowski_elem_Suc[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 2)" "t"], assumption+, simp add:nset_def) apply ( thin_tac "Ostrowski_elem K (Suc n) (compose {j. j \ Suc n} vv (skip 1)) s", thin_tac "Ostrowski_elem K (Suc n) (compose {j. j \ Suc n} vv (skip 2)) t", thin_tac "vals_nonequiv K (Suc n) (compose {l. l \ Suc n} vv (skip 1))", frule vals_nonequiv_valuation[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 2)" "Suc 0"]) apply simp+ apply (simp add:skip_def compose_def, simp add:val_t2p, simp add:aadd_pos_poss) (** Ostrowski_elem_Suc case j = Suc (Suc 0) **) apply (case_tac "j = Suc (Suc 0)", frule vals_nonequiv_valuation[of "Suc n" "compose {j. j \ Suc n} vv (skip 1)" "Suc 0"], simp, frule_tac j = " Suc 0" in Ostrowski_elem_Suc[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 1)" "s"], assumption+, simp add:nset_def, simp add:skip_def compose_def, simp add:val_t2p, rule aadd_poss_pos, assumption+) apply (frule_tac j = j in nsetTr1[of _ "Suc 0" "Suc (Suc n)"], assumption, frule_tac j = j in nsetTr2[of _ "Suc 0" "Suc n"], thin_tac "j \ nset (Suc (Suc 0)) (Suc (Suc n))", frule_tac j = "j - Suc 0" in Ostrowski_elem_Suc[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 1)" "s"], assumption+) apply (frule_tac j = "j - Suc 0" in Ostrowski_elem_Suc[of "Suc n" "compose {j. j \ (Suc n)} vv (skip 2)" "t"], assumption+, thin_tac "Ostrowski_elem K (Suc n) (compose {j. j \ (Suc n)} vv (skip 1)) s", thin_tac "Ostrowski_elem K (Suc n) (compose {j. j \ (Suc n)} vv (skip 2)) t", thin_tac "vals_nonequiv K (Suc n) (compose {j. j \ (Suc n)} vv (skip 1))", thin_tac "vals_nonequiv K (Suc n) (compose {j. j \ (Suc n)} vv (skip 2))") apply (simp add:compose_def skip_def nset_def, (erule conjE)+, simp, subgoal_tac "\ (j - Suc 0 \ Suc 0)", simp) apply (frule_tac m = j in vals_nonequiv_valuation[of "Suc (Suc n)"], assumption+, simp add:val_t2p, rule_tac x = "vv j s" and y = "vv j t" in aadd_pos_poss, simp add:aless_imp_le, assumption) apply simp done lemma (in Corps) one_plus_x_nonzero:"\valuation K v; x \ carrier K; v x < 0\ \ 1\<^sub>r \ x \ carrier K \ v (1\<^sub>r \ x) < 0" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Ring.ring_one[of "K"], frule aGroup.ag_pOp_closed[of "K" "1\<^sub>r" "x"], assumption+, simp) apply (frule value_less_eq[of "v" "x" "1\<^sub>r"], assumption+, simp add:value_of_one, simp add:aGroup.ag_pOp_commute) done lemma (in Corps) val_neg_nonzero:"\valuation K v; x \ carrier K; v x < 0\ \ x \ \" apply (rule contrapos_pp, simp+, simp add:value_of_zero) apply (frule aless_imp_le[of "\" "0"], cut_tac inf_ge_any[of "0"], frule ale_antisym[of "0" "\"], assumption+, simp) done lemma (in Corps) OstrowskiTr6:"\valuation K v; x \ carrier K; \ 0 \ (v x)\ \ (1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)) \ carrier K - {\}" apply (simp add:aneg_le, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule aGroup.ag_mOp_closed[of "K" "x"], assumption+, frule one_plus_x_nonzero[of "v" "-\<^sub>a x"], assumption+, simp add:val_minus_eq, erule conjE) apply (rule conjI, rule aGroup.ag_pOp_closed[of "K"], assumption+, simp add:Ring.ring_one, rule Ring.ring_tOp_closed[of "K"], assumption+) apply (frule val_t2p[of v x "1\<^sub>r \ (-\<^sub>a x)"], assumption+, frule val_neg_nonzero[of v x], assumption+, frule val_nonzero_z[of v x], assumption+, erule exE, frule_tac z = z in aadd_less_mono_z[of "v (1\<^sub>r \ (-\<^sub>a x))" "0"], simp add:aadd_0_l, simp only:aadd_commute[of "v (1\<^sub>r \ -\<^sub>a x)"], frule_tac x = "ant z + v (1\<^sub>r \ -\<^sub>a x)" and y ="ant z" in aless_trans[of _ _ "0"], assumption, drule sym, simp) apply (frule_tac x = x and y = "1\<^sub>r \ -\<^sub>a x" in Ring.ring_tOp_closed[of "K"], assumption+, frule one_plus_x_nonzero[of v "x \\<^sub>r (1\<^sub>r \ (-\<^sub>a x))"], assumption+, erule conjE, rule val_neg_nonzero[of v], assumption+) done lemma (in Corps) OstrowskiTr7:"\valuation K v; x \ carrier K; \ 0 \ (v x)\ \ 1\<^sub>r \ -\<^sub>a (x \\<^sub>r ((1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\K\<^esup>)) = (1\<^sub>r \ -\<^sub>a x \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)) \\<^sub>r ((1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\K\<^esup>)" apply (cut_tac field_is_ring, frule OstrowskiTr6[of v x], assumption+, simp, erule conjE, cut_tac field_is_idom, cut_tac invf_closed1[of "1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)"], simp, frule Ring.ring_is_ag[of "K"], frule aGroup.ag_mOp_closed[of "K" "x"], assumption+, frule Ring.ring_one[of "K"], frule aGroup.ag_pOp_closed[of "K" "1\<^sub>r" "-\<^sub>a x"], assumption+, rule Idomain.idom_mult_cancel_r[of K "1\<^sub>r \ -\<^sub>a (x \\<^sub>r ((1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\K\<^esup>))" "(1\<^sub>r \ -\<^sub>a x \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)) \\<^sub>r ((1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\K\<^esup>)" "(1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))"], assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule aGroup.ag_mOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+, simp, rule Ring.ring_tOp_closed, assumption+, (rule aGroup.ag_pOp_closed, assumption+)+, rule Ring.ring_tOp_closed, assumption+, simp, assumption+, subst Ring.ring_tOp_assoc, assumption+, rule aGroup.ag_pOp_closed, assumption+, simp add:Ring.ring_tOp_closed, simp, simp) apply (subst linvf[of "1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x)"], simp, (subst Ring.ring_distrib2, assumption+)+, erule conjE) apply (rule aGroup.ag_mOp_closed, assumption, rule Ring.ring_tOp_closed, assumption+, subst Ring.ring_r_one, assumption+) apply (rule aGroup.ag_pOp_closed, assumption+, rule Ring.ring_tOp_closed, assumption+, erule conjE, simp add:Ring.ring_inv1_1, simp add:Ring.ring_tOp_assoc[of K "-\<^sub>a x" "(1\<^sub>r \ x \\<^sub>r (1\<^sub>r \ -\<^sub>a x))\<^bsup>\ K\<^esup>"], simp add:linvf, simp add:Ring.ring_r_one Ring.ring_l_one, frule Ring.ring_tOp_closed[of K x "1\<^sub>r \ -\<^sub>a x"], assumption+, simp add:aGroup.ag_pOp_assoc, simp add:aGroup.ag_pOp_commute) apply simp done lemma (in Corps) Ostrowski_elem_nonzero:"\vals_nonequiv K (Suc n) vv; x \ carrier K; Ostrowski_elem K (Suc n) vv x\ \ x \ \" apply (simp add:Ostrowski_elem_def, frule conjunct1, fold Ostrowski_elem_def, frule vals_nonequiv_valuation[of "Suc n" "vv" "0"], simp) apply (rule contrapos_pp, simp+, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], simp add:aGroup.ag_inv_zero, frule Ring.ring_one[of "K"], simp add:aGroup.ag_r_zero, simp add:value_of_one) done lemma (in Corps) Ostrowski_elem_not_one:"\vals_nonequiv K (Suc n) vv; x \ carrier K; Ostrowski_elem K (Suc n) vv x\ \ 1\<^sub>r \ -\<^sub>a x \ \" apply (frule vals_nonequiv_valuation [of "Suc n" "vv" "Suc 0"], simp, simp add:Ostrowski_elem_def, frule conjunct2, fold Ostrowski_elem_def) apply (subgoal_tac "0 < (vv (Suc 0) x)", rule contrapos_pp, simp+, cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Ring.ring_one[of "K"], simp only:aGroup.ag_eq_diffzero[THEN sym, of "K" "1\<^sub>r" "x"], drule sym, simp, simp add:value_of_one, subgoal_tac "Suc 0 \ nset (Suc 0) (Suc n)", simp, simp add:nset_def) done lemma (in Corps) val_unit_cond:"\ valuation K v; x \ carrier K; 0 < (v (1\<^sub>r \ -\<^sub>a x))\ \ v x = 0" apply (cut_tac field_is_ring, frule Ring.ring_is_ag[of "K"], frule Ring.ring_one[of "K"]) apply (frule aGroup.ag_mOp_closed[of "K" "1\<^sub>r"], assumption+, frule has_val_one_neq_zero[of v]) apply (frule aGroup.ag_pOp_assoc[of "K" "-\<^sub>a 1\<^sub>r" "1\<^sub>r" "-\<^sub>a x"], assumption+, simp add:aGroup.ag_mOp_closed, simp add:aGroup.ag_l_inv1, frule aGroup.ag_mOp_closed[of "K" "x"], assumption+, simp add:aGroup.ag_l_zero) apply (subgoal_tac "v (-\<^sub>a x) = v ( -\<^sub>a 1\<^sub>r \ (1\<^sub>r \ -\<^sub>a x))") prefer 2 apply simp apply (thin_tac "-\<^sub>a x = -\<^sub>a 1\<^sub>r \ (1\<^sub>r \ -\<^sub>a x)", frule value_less_eq[of "v" "-\<^sub>a 1\<^sub>r" "1\<^sub>r \ -\<^sub>a x"], assumption+, rule aGroup.ag_pOp_closed, assumption+, simp add:val_minus_eq value_of_one, simp add:val_minus_eq) apply (simp add: value_of_one) done end diff --git a/thys/Virtual_Substitution/OptimizationProofs.thy b/thys/Virtual_Substitution/OptimizationProofs.thy --- a/thys/Virtual_Substitution/OptimizationProofs.thy +++ b/thys/Virtual_Substitution/OptimizationProofs.thy @@ -1,702 +1,704 @@ subsection "Optimization Proofs" theory OptimizationProofs imports Optimizations begin lemma neg_nnf : "\\. (\ eval (nnf (Neg \)) \) = eval (nnf \) \" apply(induction \) apply(simp_all) using aNeg_aEval apply blast using aNeg_aEval by blast theorem eval_nnf : "\\. eval \ \ = eval (nnf \) \" apply(induction \)apply(simp_all) using neg_nnf by blast theorem negation_free_nnf : "negation_free (nnf \)" proof(induction "depth \" arbitrary : \ rule: nat_less_induct ) case 1 then show ?case proof(induction \) case (And \1 \2) then show ?case apply simp by (metis less_Suc_eq_le max.cobounded1 max.cobounded2) next case (Or \1 \2) then show ?case apply simp by (metis less_Suc_eq_le max.cobounded1 max.cobounded2) next case (Neg \) then show ?case proof (induction \) case (And \1 \2) then show ?case apply simp by (metis less_Suc_eq max_less_iff_conj not_less_eq) next case (Or \1 \2) then show ?case apply simp by (metis less_Suc_eq max_less_iff_conj not_less_eq) next case (Neg \) then show ?case by (metis Suc_eq_plus1 add_lessD1 depth.simps(6) lessI nnf.simps(12)) qed auto qed auto qed lemma groupQuantifiers_eval : "eval F L = eval (groupQuantifiers F) L" apply(induction F arbitrary: L rule:groupQuantifiers.induct) unfolding doubleExist unwrapExist unwrapExist' unwrapExist'' doubleForall unwrapForall unwrapForall' unwrapForall'' apply (auto) using doubleExist doubleExist unwrapExist unwrapExist' unwrapExist'' doubleForall unwrapForall unwrapForall' unwrapForall'' apply auto by metis+ theorem simp_atom_eval : "aEval a xs = eval (simp_atom a) xs" proof(cases a) case (Less p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Eq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Leq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) next case (Neq p) then show ?thesis by(cases "get_if_const p")(simp_all add:get_if_const_insertion) qed lemma simpfm_eval : "\L. eval \ L = eval (simpfm \) L" apply(induction \) apply(simp_all add: simp_atom_eval eval_and eval_or) using eval_neg by blast lemma exQ_clearQuantifiers: assumes ExQ : "\xs. eval (clearQuantifiers \) xs = eval \ xs" shows "eval (clearQuantifiers (ExQ \)) xs = eval (ExQ \) xs" proof- define \' where "\' = clearQuantifiers \" have h : "freeIn 0 \' \ (eval (lowerFm 0 1 \') xs = eval (ExQ \') xs)" using eval_lowerFm by simp have "eval (clearQuantifiers (ExQ \)) xs = eval (if freeIn 0 \' then lowerFm 0 1 \' else ExQ \') xs" using \'_def by simp also have "... = eval (ExQ \) xs" apply(cases "freeIn 0 \'") using h ExQ \'_def by(simp_all) finally show ?thesis by simp qed lemma allQ_clearQuantifiers : assumes AllQ : "\xs. eval (clearQuantifiers \) xs = eval \ xs" shows "eval (clearQuantifiers (AllQ \)) xs = eval (AllQ \) xs" proof- define \' where "\' = clearQuantifiers \" have "freeIn 0 \' \ (eval (ExQ \') xs) = eval (AllQ \') xs" by (simp add: var_not_in_eval2) then have h : "freeIn 0 \' \ (eval (lowerFm 0 1 \') xs = eval (AllQ \') xs)" using eval_lowerFm by simp have "eval (clearQuantifiers (AllQ \)) xs = eval (if freeIn 0 \' then lowerFm 0 1 \' else AllQ \') xs" using \'_def by simp also have "... = eval (AllQ \) xs" apply(cases "freeIn 0 \'") using h AllQ \'_def by(simp_all) finally show ?thesis by simp qed lemma clearQuantifiers_eval : "eval (clearQuantifiers \) xs = eval \ xs" proof(induction \ arbitrary : xs) case (Atom x) then show ?case using simp_atom_eval by simp next case (And \1 \2) then show ?case using eval_and by simp next case (Or \1 \2) then show ?case using eval_or by simp next case (Neg \) then show ?case using eval_neg by auto next case (ExQ \) then show ?case using exQ_clearQuantifiers by simp next case (AllQ \) then show ?case using allQ_clearQuantifiers by simp next case (ExN x1 \) then show ?case proof(induction x1 arbitrary:\) case 0 then show ?case by auto next case (Suc x1) show ?case using Suc(1)[of "ExQ \", OF exQ_clearQuantifiers[OF Suc(2)]] apply simp using Suc_eq_plus1 \eval (clearQuantifiers (ExN x1 (ExQ \))) xs = eval (ExN x1 (ExQ \)) xs\ eval.simps(10) unwrapExist' by presburger qed next case (AllN x1 \) then show ?case proof(induction x1 arbitrary:\) case 0 then show ?case by auto next case (Suc x1) show ?case using Suc(1)[of "AllQ \", OF allQ_clearQuantifiers[OF Suc(2)]] apply simp using unwrapForall' by force qed qed auto lemma push_forall_eval_AllQ : "\xs. eval (AllQ \) xs = eval (push_forall (AllQ \)) xs" proof(induction \) case TrueF then show ?case by simp next case FalseF then show ?case by simp next case (Atom x) then show ?case using aEval_lowerAtom eval.simps(1) eval.simps(8) push_forall.simps(11) by presburger next case (And \1 \2) {fix xs have "eval (AllQ (And \1 \2)) xs = (\x. eval \1 (x#xs) \ eval \2 (x#xs))" by simp also have "... = ((\x. eval \1 (x#xs)) \ (\x. eval \2 (x#xs)))" by blast also have "... = eval (push_forall (AllQ (And \1 \2))) xs" using And eval_and by(simp) finally have "eval (AllQ (And \1 \2)) xs = eval (push_forall (AllQ (And \1 \2))) xs" by simp } then show ?case by simp next case (Or \1 \2) then show ?case proof(cases "freeIn 0 \1") case True then have h : "freeIn 0 \1" by simp then show ?thesis proof(cases "freeIn 0 \2") case True {fix xs have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using eval_lowerFm h eval.simps(7) by blast then have h1 : "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using h var_not_in_eval2 by blast have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using eval_lowerFm True eval.simps(7) by blast then have h2 : "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using True var_not_in_eval2 by blast have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h h1 h2 True eval_or) } then show ?thesis by simp next case False {fix xs have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using eval_lowerFm h eval.simps(7) by blast then have "\x. eval \1 (x # xs) = eval (lowerFm 0 1 \1) xs" using True var_not_in_eval2 by blast then have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h False eval_or) } then show ?thesis by simp qed next case False then have h : "\freeIn 0 \1" by simp then show ?thesis proof(cases "freeIn 0 \2") case True {fix xs have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using eval_lowerFm True eval.simps(7) by blast then have "\x. eval \2 (x # xs) = eval (lowerFm 0 1 \2) xs" using True var_not_in_eval2 by blast then have "eval (AllQ (Or \1 \2)) xs = eval (push_forall (AllQ (Or \1 \2))) xs" by(simp add:h True eval_or) } then show ?thesis by simp next case False then show ?thesis by(simp add:h False eval_or) qed qed next case (Neg \) {fix xs have "freeIn 0 (Neg \) \ (eval (ExQ (Neg \)) xs) = eval (AllQ (Neg \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (Neg \) \ (eval (lowerFm 0 1 (Neg \)) xs = eval (AllQ (Neg \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (Neg \))) xs = eval (if freeIn 0 (Neg \) then lowerFm 0 1 (Neg \) else AllQ (Neg \)) xs" by simp also have "... = eval (AllQ (Neg \)) xs" apply(cases "freeIn 0 (Neg \)") using h by(simp_all) finally have "eval (push_forall (AllQ (Neg \))) xs = eval (AllQ (Neg \)) xs" by simp } then show ?case by simp next case (ExQ \) {fix xs have "freeIn 0 (ExQ \) \ (eval (ExQ (ExQ \)) xs) = eval (AllQ (ExQ \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (ExQ \) \ (eval (lowerFm 0 1 (ExQ \)) xs = eval (AllQ (ExQ \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (ExQ \))) xs = eval (if freeIn 0 (ExQ \) then lowerFm 0 1 (ExQ \) else AllQ (ExQ \)) xs" by simp also have "... = eval (AllQ (ExQ \)) xs" apply(cases "freeIn 0 (ExQ \)") using h by(simp_all) finally have "eval (push_forall (AllQ (ExQ \))) xs = eval (AllQ (ExQ \)) xs" by simp } then show ?case by simp next case (AllQ \) {fix xs have "freeIn 0 (AllQ \) \ (eval (ExQ (AllQ \)) xs) = eval (AllQ (AllQ \)) xs" using var_not_in_eval2 eval.simps(7) eval.simps(8) by blast then have h : "freeIn 0 (AllQ \) \ (eval (lowerFm 0 1 (AllQ \)) xs = eval (AllQ (AllQ \)) xs)" using eval_lowerFm by blast have "eval (push_forall (AllQ (AllQ \))) xs = eval (if freeIn 0 (AllQ \) then lowerFm 0 1 (AllQ \) else AllQ (AllQ \)) xs" by simp also have "... = eval (AllQ (AllQ \)) xs" apply(cases "freeIn 0 (AllQ \)") using h AllQ by(simp_all) finally have "eval (push_forall (AllQ (AllQ \))) xs = eval (AllQ (AllQ \)) xs" by simp } then show ?case by simp next case (ExN x1 \) then show ?case using eval.simps(7) eval.simps(8) eval_lowerFm push_forall.simps(17) var_not_in_eval2 by presburger next case (AllN x1 \) then show ?case using eval.simps(7) eval.simps(8) eval_lowerFm push_forall.simps(18) var_not_in_eval2 by presburger qed lemma push_forall_eval : "\xs. eval \ xs = eval (push_forall \) xs" proof(induction \) case (Atom x) then show ?case using simp_atom_eval by simp next case (And \1 \2) then show ?case using eval_and by auto next case (Or \1 \2) then show ?case using eval_or by auto next case (Neg \) then show ?case using eval_neg by auto next case (AllQ \) then show ?case using push_forall_eval_AllQ by blast next case (ExN x1 \) then show ?case using eval.simps(10) push_forall.simps(7) by presburger qed auto lemma map_fm_binders_negation_free : assumes "negation_free \" shows "negation_free (map_fm_binders f \ n)" using assms apply(induction \ arbitrary : n) by auto lemma negation_free_and : assumes "negation_free \" assumes "negation_free \" shows "negation_free (and \ \)" using assms unfolding and_def by simp lemma negation_free_or : assumes "negation_free \" assumes "negation_free \" shows "negation_free (or \ \)" using assms unfolding or_def by simp lemma push_forall_negation_free_all : assumes "negation_free \" shows "negation_free (push_forall (AllQ \))" using assms proof(induction \) case (And \1 \2) show ?case apply auto apply(rule negation_free_and) using And by auto next case (Or \1 \2) show ?case apply auto apply(rule negation_free_or) using Or map_fm_binders_negation_free negation_free_or by auto next case (ExQ \) then show ?case using map_fm_binders_negation_free by auto next case (AllQ \) then show ?case using map_fm_binders_negation_free by auto next case (ExN x1 \) then show ?case using map_fm_binders_negation_free by auto next case (AllN x1 \) then show ?case using map_fm_binders_negation_free by auto qed auto lemma push_forall_negation_free : assumes "negation_free \" shows "negation_free(push_forall \)" using assms proof(induction \) case (Atom A) then show ?case apply(cases A) by auto next case (And \1 \2) then show ?case by (auto simp add: and_def) next case (Or \1 \2) then show ?case by (auto simp add: or_def) next case (AllQ \) then show ?case using push_forall_negation_free_all by auto qed auto lemma to_list_insertion: "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(to_list v p)]" proof- have h1 : "insertion f p = insertion f (\i\MPoly_Type.degree p v. isolate_variable_sparse p v i * Var v ^ i)" using sum_over_zero by auto have h2 : "insertion f (Var v) = f v" by (auto simp: monomials_Var coeff_Var insertion_code) define d where "d = MPoly_Type.degree p v" define g where "g = (\x. insertion f (isolate_variable_sparse p v x) * f v ^ x)" have h3 : "insertion f (isolate_variable_sparse p v d) * f v ^ d = g d" using g_def by auto show ?thesis unfolding h1 insertion_sum' insertion_mult insertion_pow h2 apply auto unfolding d_def[symmetric] g_def[symmetric] h3 proof(induction d) case 0 then show ?case by auto next case (Suc d) show ?case apply (auto simp add: Suc ) unfolding g_def by auto qed qed lemma to_list_p: "p = sum_list [term * (Var v) ^ i. (term,i)\(to_list v p)]" proof- define d where "d = MPoly_Type.degree p v" have "(\i\MPoly_Type.degree p v. isolate_variable_sparse p v i * Var v ^ i) = (\(term, i)\to_list v p. term * Var v ^ i)" unfolding to_list.simps d_def[symmetric] apply(induction d) by auto then show ?thesis using sum_over_zero[of p v] by auto qed fun chophelper :: "(real mpoly * nat) list \ (real mpoly * nat) list \ (real mpoly * nat) list * (real mpoly * nat) list" where "chophelper [] L = (L,[])"| "chophelper ((p,i)#L) R = (if p=0 then chophelper L (R @ [(p,i)]) else (R,(p,i)#L))" lemma preserve : assumes "(a,b)=chophelper L L'" shows "a@b=L'@L" using assms proof(induction L arbitrary : a b L') case Nil then show ?case using assms by auto next case (Cons A L) then show ?case proof(cases A) case (Pair p i) show ?thesis using Cons unfolding Pair apply(cases "p=0") by auto qed qed lemma compare : assumes "(a,b)=chophelper L L'" shows "chop L = b" using assms proof(induction L arbitrary : a b L') case Nil then show ?case by auto next case (Cons A L) then show ?case proof(cases A) case (Pair p i) show ?thesis using Cons unfolding Pair apply(cases "p=0") by auto qed qed lemma allzero: assumes "\(p,i)\set(L'). p=0" assumes "(a,b)=chophelper L L'" shows "\(p,i)\set(a). p=0" using assms proof(induction L arbitrary : a b L') case Nil then show ?case by auto next case (Cons t L) then show ?case proof(cases t) case (Pair p i) show ?thesis proof(cases "p=0") case True have h1: "\x\set (L' @ [(0, i)]). case x of (p, i) \ p = 0" using Cons(2) by auto show ?thesis using Cons(1)[OF h1] Cons(3) True unfolding Pair by auto next case False then show ?thesis using Cons unfolding Pair by auto qed qed qed lemma separate: assumes "(a,b)=chophelper (to_list v p) []" shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\a] + sum_list [insertion f term * (f v) ^ i. (term,i)\b]" using to_list_insertion[of f p v] preserve[OF assms, symmetric] unfolding List.append.left_neutral by (simp del: to_list.simps) lemma chopped : assumes "(a,b)=chophelper (to_list v p) []" shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\b]" proof- have h1 : "\(p, i)\set []. p = 0" by auto have "(\(term, i)\a. insertion f term * f v ^ i) = 0" using allzero[OF h1 assms] proof(induction a) case Nil then show ?case by auto next case (Cons a1 a2) then show ?case apply(cases a1) by simp qed then show ?thesis using separate[OF assms, of f] by auto qed lemma insertion_chop : shows "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(chop (to_list v p))]" proof(cases "chophelper (to_list v p) []") case (Pair a b) show ?thesis using chopped[OF Pair[symmetric], of f] unfolding compare[OF Pair[symmetric], symmetric] . qed lemma sorted : "sorted_wrt (\(_,i).\(_,i'). ix. (isolate_variable_sparse p v x, x)) [0..x. (isolate_variable_sparse p v x, x)) [0..(chop (to_list v p))] * (f v)^i" proof- have h : "sorted_wrt (\(_, i) (_, y). i < y) (chop (to_list v p))" proof- define L where "L = to_list v p" show ?thesis using sublist[of "to_list v p"] sorted[of v p] unfolding L_def[symmetric] by (metis sorted_wrt_append sublist_def) qed then have "\(term,d)\set(chop (to_list v p)). d\i" unfolding assms[symmetric] by fastforce then have simp : "\(term,d)\set(chop(to_list v p)). f v ^ (d - i) * f v ^ i = f v ^ d" unfolding semiring_normalization_rules(26) by(auto simp del: to_list.simps) have "insertion f p = sum_list [insertion f term * (f v) ^ i. (term,i)\(chop (to_list v p))]" using insertion_chop[of f p v] . also have "...= (\(term, d)\chop (to_list v p). insertion f term * f v ^ (d-i) * f v ^ i)" using simp by (smt Pair_inject case_prodE map_eq_conv mult.assoc split_cong) also have "... = (\(term, d)\chop (to_list v p). insertion f term * f v ^ (d - i)) * f v ^ i" proof- define d where "d = chop(to_list v p)" define a where "a = f v ^ i" define b where "b = (\(term, d). insertion f term * f v ^ (d - i))" have h : "(\(term, d)\d. insertion f term * f v ^ (d - i) * a) = (\(term, d)\d. b (term,d) * a)" using b_def by auto show ?thesis unfolding d_def[symmetric] a_def[symmetric] b_def[symmetric] h apply(induction d) apply simp apply auto by (simp add: ring_class.ring_distribs(2)) qed finally show ?thesis by auto qed lemma insert_Var_Zero : "insertion f (Var v) = f v" unfolding insertion_code monomials_Var apply auto unfolding coeff_Var by simp lemma decreasePower_insertion : assumes "decreasePower v p = (p',i)" shows "insertion f p = insertion f p'* (f v)^i" proof(cases "chop (to_list v p)") case Nil then show ?thesis using assms by auto next case (Cons a list) then show ?thesis proof(cases a) case (Pair coef i') have i'_def : "i'=i" using Cons assms Pair by auto have chop: "chop (to_list v p) = (coef, i) # list" using Cons assms unfolding i'_def Pair by auto have p'_def : "p' = (\(term, x)\chop (to_list v p). term * Var v ^ (x - i))" using assms Cons Pair by auto have p'_insertion : "insertion f p' = (\(term, x)\chop (to_list v p). insertion f term * f v ^ (x - i))" proof- define d where "d = chop (to_list v p)" have "insertion f p' = insertion f (\(term, x)\chop (to_list v p). term * Var v ^ (x - i))" using p'_def by auto also have "... = (\(term, x)\chop (to_list v p). insertion f (term * Var v ^ (x - i)))" unfolding d_def[symmetric] apply(induction d) apply simp apply(simp add:insertion_add) by auto also have "... = (\(term, x)\chop (to_list v p). insertion f term * f v ^ (x - i))" unfolding insertion_mult insertion_pow insert_Var_Zero by auto finally show ?thesis by auto qed have h : "(coef, i') # list = chop (to_list v p)" using Cons unfolding Pair by auto show ?thesis unfolding p'_insertion using move_exp[OF h, of f] unfolding i'_def . qed qed lemma unpower_eval: "eval (unpower v \) L = eval \ L" proof(induction \ arbitrary: v L) case TrueF then show ?case by auto next case FalseF then show ?case by auto next case (Atom At) then show ?case proof(cases At) case (Less p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis proof(cases "i=0") case True then show ?thesis unfolding Less unpower.simps h by auto next case False obtain x where x_def : "Suc x = i" using False using not0_implies_Suc by auto have h1 : "i mod 2 = 0 \ (insertion (nth_default 0 L) q < 0 \ insertion (nth_default 0 L) (Var v) \ 0) = (insertion (nth_default 0 L) q * nth_default 0 L v ^ i < 0)" proof - assume "i mod 2 = 0" then have "\r. \ (r::real) ^ i < 0" by presburger then show ?thesis by (metis \\thesis. (\x. Suc x = i \ thesis) \ thesis\ insert_Var_Zero linorder_neqE_linordered_idom mult_less_0_iff power_0_Suc power_eq_0_iff) qed - show ?thesis unfolding Less unpower.simps h x_def[symmetric] apply simp - unfolding x_def p apply(cases "i mod 2 = 0") using h1 apply simp_all - by (smt insert_Var_Zero insertion_neg mod_Suc mod_eq_0D mult_less_0_iff nat.inject odd_power_less_zero power_0 power_Suc0_right power_eq_0_iff x_def zero_less_Suc zero_less_power) + show ?thesis + unfolding Less unpower.simps h x_def [symmetric] apply simp + unfolding x_def p apply (cases \even i\) + using h1 apply (auto simp add: insertion_neg insert_Var_Zero mult_less_0_iff not_less zero_less_mult_iff elim: oddE) + done qed next case (Eq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis unfolding Eq unpower.simps h apply simp apply(cases i) apply simp apply simp unfolding p apply simp by (metis insert_Var_Zero) next case (Leq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis proof(cases "i=0") case True then show ?thesis unfolding Leq unpower.simps h by auto next case False obtain x where x_def : "Suc x = i" using False using not0_implies_Suc by auto define a where "a = insertion (nth_default 0 L) q" define x' where "x' = nth_default 0 L v" show ?thesis unfolding Leq unpower.simps h x_def[symmetric] apply simp unfolding x_def p apply(cases "i mod 2 = 0") unfolding insert_Var_Zero insertion_mult insertion_pow insertion_neg apply simp_all unfolding a_def[symmetric] x'_def[symmetric] proof- assume "i mod 2 = 0" then have "x' ^ i \0" by (simp add: \i mod 2 = 0\ even_iff_mod_2_eq_zero zero_le_even_power) then show "(a \ 0 \ x' = 0) = (a * x' ^ i \ 0)" using Rings.ordered_semiring_0_class.mult_nonpos_nonneg[of a "x'^i"] apply auto unfolding Rings.linordered_ring_strict_class.mult_le_0_iff apply auto by (simp add: False power_0_left) next assume h: "i mod 2 = Suc 0" show "(a = 0 \ a < 0 \ 0 \ x' \ 0 < a \ x' \ 0) = (a * x' ^ i \ 0)" using h by (smt even_iff_mod_2_eq_zero mult_less_cancel_right mult_neg_neg mult_nonneg_nonpos mult_pos_pos not_mod2_eq_Suc_0_eq_0 power_0_Suc x_def zero_le_power_eq zero_less_mult_pos2 zero_less_power) qed qed next case (Neq p) obtain q i where h: "decreasePower v p = (q, i)" using prod.exhaust_sel by blast have p : "\f. insertion f p = insertion f q* (f v)^i" using decreasePower_insertion[OF h] by auto show ?thesis unfolding Neq unpower.simps h apply simp apply(cases i) apply simp apply simp unfolding p apply simp by (metis insert_Var_Zero) qed qed auto lemma to_list_filter: "p = sum_list [term * (Var v) ^ i. (term,i)\((filter (\(x,_). x\0) (to_list v p)))]" proof- define L where "L = to_list v p" have "(\(term, i)\to_list v p. term * Var v ^ i) = (\(term, i)\filter (\(x, _). x \ 0) (to_list v p). term * Var v ^ i)" unfolding L_def[symmetric] apply(induction L) by auto then show ?thesis using to_list_p[of p v] by auto qed end \ No newline at end of file diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1558 +1,1558 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" begin subsection \Implicit bit representation of \<^typ>\int\\ lemma bin_last_def: "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: "(\k::int. k div 2) 0 = 0" "(\k::int. k div 2) 1 = 0" "(\k::int. k div 2) (- 1) = - 1" "(\k::int. k div 2) Numeral1 = 0" "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that by (rule bit_eqI) lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (metis bit_eq_iff) lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by simp lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc del: take_bit_minus_one_eq_mask) lemma sbintrunc_0_numeral [simp]: "(signed_take_bit :: nat \ int \ int) 0 1 = -1" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by simp lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by simp lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq mask_eq_exp_minus_1) lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (fact signed_take_bit_int_greater_eq_minus_exp) lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (fact signed_take_bit_int_less_exp) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) lemma bin_cat_eq_push_bit_add_take_bit: \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ abbreviation (input) bin_sc :: \nat \ bool \ int \ int\ where \bin_sc n b i \ set_bit i n b\ lemma bin_sc_0 [simp]: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" by (simp add: set_bit_int_def) lemma bin_sc_Suc [simp]: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" by (simp add: set_bit_int_def set_bit_Suc unset_bit_Suc bin_last_def) lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" by (simp add: bit_simps) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" by (simp add: bit_simps) lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ apply (simp_all add: fun_eq_iff bit_eq_iff) apply (simp_all add: bit_simps bin_nth_sc_gen) done lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (rule bit_eqI) apply (cases x) apply (auto simp add: bit_simps bin_sc_eq) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: set_bit_int_def unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: set_bit_int_def set_bit_greater_eq) lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc_0 bin_sc_Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc_Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = of_bool b + 2 * (x div 2)" by (fact bin_sc_0) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" by (fact bin_sc_Suc) lemma bin_last_set_bit: "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" by (cases n) (simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" by (cases n) (simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" by (fact bin_sc_numeral) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by (simp add: msb_int_def set_bit_int_def) lemma word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ apply (rule bit_word_eqI) apply (cases x) apply (simp_all add: bit_simps bin_sc_eq) done lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI) (simp add: word_size bin_nth_sc_gen nth_bintr bit_simps) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemmas shiftl_int_def = shiftl_eq_mult[of x for x::int] lemmas shiftr_int_def = shiftr_eq_div[of x for x::int] subsubsection \Basic simplification rules\ context includes bit_operations_syntax begin lemmas int_not_def = not_int_def lemma int_not_simps: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact and.left_neutral) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact or.left_neutral) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact xor.left_neutral) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0: "x << 0 = x" and int_shiftl_Suc: "x << Suc n = 2 * x << n" by (auto simp add: shiftl_int_def) lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" by (fact push_bit_of_0) lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" by simp lemma bin_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" by (cases n) (simp_all add: push_bit_eq_mult) lemma bin_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (fact bit_push_bit_iff_int) lemma bin_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int by (simp add: bit_iff_odd_drop_bit) lemma bin_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" by (simp add: drop_bit_Suc drop_bit_half) lemma bin_nth_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: bit_simps) lemma bin_nth_conv_AND: fixes x :: int shows "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" by (fact bit_iff_and_push_bit_not_eq_0) lemma int_shiftl_numeral [simp]: "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 w))" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" using int_shiftl_numeral [of Num.One w] by (simp only: numeral_eq_Suc push_bit_Suc) simp lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" by (fact push_bit_nonnegative_int_iff) lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" by (fact push_bit_negative_int_iff) lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" by (fact bit_push_bit_iff_int) lemma int_0shiftr: "drop_bit x (0 :: int) = 0" by (fact drop_bit_of_0) lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" by (fact drop_bit_minus_one) lemma int_shiftr_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" by (fact drop_bit_nonnegative_int_iff) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" by (fact drop_bit_negative_int_iff) lemma int_shiftr_numeral [simp]: "drop_bit (numeral w') (1 :: int) = 0" "drop_bit (numeral w') (numeral num.One :: int) = 0" "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" by (simp_all add: numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "drop_bit (Suc 0) (1 :: int) = 0" "drop_bit (Suc 0) (numeral num.One :: int) = 0" "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" shows "bit (x - y) m = bit x m" proof - from \bin_sign x = 0\ have \x \ 0\ by (simp add: sign_Pls_ge_0) moreover from x y have \x < 2 ^ n\ by simp ultimately have \q < n\ if \bit x q\ for q using that by (metis bit_take_bit_iff take_bit_int_eq_self) then have \bit (x + NOT (mask n)) m = bit x m\ using \m < n\ by (simp add: disjunctive_add bit_simps) also have \x + NOT (mask n) = x - y\ using y by (simp flip: minus_exp_eq_not_mask) finally show ?thesis . qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: "bin_sc n True i = i OR (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) end subsection \More lemmas on words\ lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by (simp add: bin_sign_def not_le msb_int_def) lemma msb_bin_sc: "msb (bin_sc n b x) \ msb x" by (simp add: msb_conv_bin_sign) lemma msb_word_def: \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ for a :: \'a::len word\ by (simp add: bin_sign_def bit_simps msb_word_iff_bit) lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no: "Bit_Operations.set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma clearBit_no: "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int - by auto (metis pos_mod_conj)+ + using pos_mod_sign [of n b] pos_mod_bound [of n b] by (safe, auto) lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int by (fact take_bit_push_bit) lemma uint_shiftl: "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) context includes bit_operations_syntax begin lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) end lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed lemma bintrunc_id: "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: take_bit_int_eq_self_iff le_less_trans) lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" if "n = m" "a = c" "take_bit m b = take_bit m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma bin_sc_pos: "0 \ i \ 0 \ bin_sc n b i" by (metis bin_sign_sc sign_Pls_ge_0) code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy --- a/thys/Word_Lib/Reversed_Bit_Lists.thy +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -1,2232 +1,2229 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" Aligned Singleton_Bit_Shifts Legacy_Aliases begin context includes bit_operations_syntax begin lemma horner_sum_of_bool_2_concat: \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. for ws :: \'a::len word list\ proof (induction ws) case Nil then show ?case by simp next case (Cons w ws) moreover have \horner_sum of_bool 2 (map (bit w) [0.. proof transfer fix k :: int have \map (\n. n < LENGTH('a) \ bit k n) [0.. by simp then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. by (simp only: horner_sum_bit_eq_take_bit) qed ultimately show ?case by (simp add: horner_sum_append) qed subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) subsection \Explicit bit representation of \<^typ>\int\\ primrec bl_to_bin_aux :: "bool list \ int \ int" where Nil: "bl_to_bin_aux [] w = w" | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" definition bl_to_bin :: "bool list \ int" where "bl_to_bin bs = bl_to_bin_aux bs 0" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (w div 2) (odd w # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" lemma bin_to_bl_aux_zero_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" by (cases n) auto lemma bin_to_bl_aux_minus1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" by (cases n) auto lemma bin_to_bl_aux_one_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) simp_all lemma bin_to_bl_aux_Bit1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) simp_all lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" by (induct n arbitrary: w bs) auto lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" by (simp add: bin_to_bl_def bin_to_bl_aux_append) lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" by (auto simp: bin_to_bl_def) lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (induct n arbitrary: w bs) auto lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" by (simp add: bin_to_bl_def size_bin_to_bl_aux) lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" apply (induct bs arbitrary: w n) apply auto apply (simp_all only: add_Suc [symmetric]) apply (auto simp add: bin_to_bl_def) done lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" unfolding bl_to_bin_def apply (rule box_equals) apply (rule bl_bin_bl') prefer 2 apply (rule bin_to_bl_aux.Z) apply simp done lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" apply (rule_tac box_equals) defer apply (rule bl_bin_bl) apply (rule bl_bin_bl) apply simp done lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" by (auto simp: bl_to_bin_def) lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" by (auto simp: bl_to_bin_def) lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" by (simp add: bin_to_bl_def bin_to_bl_zero_aux) lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (take_bit n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = take_bit n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (take_bit m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: "bin_to_bl_aux n (take_bit m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp apply clarsimp apply (case_tac "m") apply (clarsimp simp: bin_to_bl_zero_aux) apply (erule thin_rl) apply (induct_tac n) apply (auto simp add: take_bit_Suc) done lemma bin_to_bl_bintr: "bin_to_bl n (take_bit m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" by (induct n) auto lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (fact size_bin_to_bl_aux) lemma len_bin_to_bl: "length (bin_to_bl n w) = n" by (fact size_bin_to_bl) (* FIXME: duplicate *) lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" by (induction bs arbitrary: w) (simp_all add: bin_sign_def) lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (signed_take_bit n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (signed_take_bit n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: "bit (bl_to_bin_aux bl w) n = (n < size bl \ rev bl ! n \ n \ length bl \ bit w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done lemma bin_nth_of_bl: "bit (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) lemma bin_nth_bl: "n < m \ bit w n = nth (rev (bin_to_bl m w)) n" by (metis bin_bl_bin bin_nth_of_bl nth_bintr size_bin_to_bl) lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bit w (m - 1 - n) else bl ! (n - m))" apply (induction bl arbitrary: w) apply simp_all apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) apply (metis One_nat_def Suc_pred add_diff_cancel_left' add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bit w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" apply (rule nth_equalityI) apply simp apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) done lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" by (simp add: takefill_bintrunc) lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (subst mult_2 [of \2 ^ length bs\]) apply (simp only: add.assoc) apply (rule pos_add_strict) apply simp_all done qed lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" proof (induct bs) case Nil then show ?case by simp next case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] show ?case by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" by (metis bin_bl_bin bintr_lt2p bl_bin_bl) lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (rule add_le_imp_le_left [of \2 ^ length bs\]) apply (rule add_increasing) apply simp_all done qed lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply (unfold bl_to_bin_def) apply (rule xtrans(4)) apply (rule bl_to_bin_ge2p_aux) apply simp done lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (w div 2)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bl_to_bin bl div 2)" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: "bl \ [] \ bl_to_bin_aux (butlast bl) w = bl_to_bin_aux bl w div 2" by (induct bl arbitrary: w) auto lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bl_to_bin bl div 2" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: "take_bit m (bl_to_bin_aux bl w) = bl_to_bin_aux (drop (length bl - m) bl) (take_bit (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp next case (Cons b bl) show ?case proof (cases "m - length bl") case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp with Cons show ?thesis by simp next case (Suc n) then have "m - Suc (length bl) = n" by simp with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) qed qed lemma trunc_bl2bin: "take_bit m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) lemma trunc_bl2bin_len [simp]: "take_bit (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) lemma bl2bin_drop: "bl_to_bin (drop k bl) = take_bit (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) apply (cases "k \ length bl") apply auto done lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m (((\w. w div 2) ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done lemma last_bin_last': "size xs > 0 \ last xs \ odd (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto lemma last_bin_last: "size xs > 0 \ last xs \ odd (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') lemma bin_last_last: "odd w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: "drop m (bin_to_bl_aux n bin bs) = bin_to_bl_aux (n - m) bin (drop (m - n) bs)" apply (induction n arbitrary: m bin bs) apply auto apply (case_tac "m \ n") apply (auto simp add: not_le Suc_diff_le) apply (case_tac "m - n") apply auto apply (use Suc_diff_Suc in fastforce) done lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" by (simp add: bin_to_bl_def drop_bin2bl_aux) lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" apply (induct m arbitrary: w bs) apply clarsimp apply clarsimp apply (simp add: bin_to_bl_aux_alt) apply (simp add: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (simp add: bin_to_bl_def) apply (simp add: take_bin2bl_lem drop_bit_Suc) done lemma bin_to_bl_drop_bit: "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" using bin_split_take by simp lemma bin_split_take1: "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" using bin_split_take by simp lemma bl_bin_bl_rep_drop: "bin_to_bl n (bl_to_bin bl) = replicate (n - length bl) False @ drop (length bl - n) bl" by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: "bl_to_bin_aux bs (concat_bit nv v w) = concat_bit (nv + length bs) (bl_to_bin_aux bs v) w" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: "bin_to_bl_aux (nv + nw) (concat_bit nw w v) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = concat_bit (length bs) (bl_to_bin bs) w" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: "bin_to_bl (nv + nw) (concat_bit nw w v) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) lemmas bl_to_bin_aux_app_cat = trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: "bl_to_bin (bsa @ bs) = concat_bit (length bs) (bl_to_bin bs) (bl_to_bin bsa)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: "bin_to_bl (n + nw) (concat_bit nw wa w) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ lemma bl_to_bin_app_cat_alt: "concat_bit n w (bl_to_bin cs) = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" apply (unfold bl_to_bin_def) apply (induct n) apply simp apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) apply simp done lemma bin_exhaust: "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int apply (cases \even bin\) apply (auto elim!: evenE oddE) apply fastforce apply fastforce done primrec rbl_succ :: "bool list \ bool list" where Nil: "rbl_succ Nil = Nil" | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" primrec rbl_pred :: "bool list \ bool list" where Nil: "rbl_pred Nil = Nil" | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" primrec rbl_add :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_add Nil x = Nil" | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" primrec rbl_mult :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_mult Nil x = Nil" | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in if y then rbl_add ws x else ws)" lemma size_rbl_pred: "length (rbl_pred bl) = length bl" by (induct bl) auto lemma size_rbl_succ: "length (rbl_succ bl) = length bl" by (induct bl) auto lemma size_rbl_add: "length (rbl_add bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) lemmas rbl_sizes [simp] = size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult lemmas rbl_Nils = rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_add_take2: "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def rbl_add_app2) done lemma rbl_mult_take2: "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" apply (rule trans) apply (rule rbl_mult_app2 [symmetric]) apply simp apply (rule_tac f = "rbl_mult bla" in arg_cong) apply (rule append_take_drop_id) done lemma rbl_add_split: "P (rbl_add (y # ys) (x # xs)) = (\ws. length ws = length ys \ ws = rbl_add ys xs \ (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ (\ y \ P (x # ws)))" by (cases y) (auto simp: Let_def) lemma rbl_mult_split: "P (rbl_mult (y # ys) xs) = (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" by (auto simp: Let_def) lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" proof (unfold bin_to_bl_def, induction n arbitrary: bin) case 0 then show ?case by simp next case (Suc n) obtain b k where \bin = of_bool b + 2 * k\ using bin_exhaust by blast moreover have \(2 * k - 1) div 2 = k - 1\ using even_succ_div_2 [of \2 * (k - 1)\] by simp ultimately show ?case using Suc [of \bin div 2\] by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt) apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" by (simp add: length_concat comp_def sum_list_triv) lemma bin_cat_foldl_lem: "foldl (\u k. concat_bit n k u) x xs = concat_bit (size xs * n) (foldl (\u k. concat_bit n k u) y xs) x" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) apply (drule_tac x = "concat_bit n a y" in meta_spec) apply (simp add: bin_cat_assoc_sym) done lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (unfold bin_rcat_eq_foldl) apply (rule sym) apply (induct wl) apply (auto simp add: bl_to_bin_append) apply (simp add: bl_to_bin_aux_alt sclem) apply (simp add: bin_cat_foldl_lem [symmetric]) done lemma bin_last_bl_to_bin: "odd (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) lemma bin_rest_bl_to_bin: "bl_to_bin bs div 2 = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" apply (induction n arbitrary: v w bs cs) apply auto apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done lemma bl_or_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" by (induct n arbitrary: v w bs cs) simp_all lemma bl_and_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" by (induction n arbitrary: v w bs cs) simp_all lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" by (induct n arbitrary: w cs) auto lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" by (simp add: bin_to_bl_def bl_not_aux_bin) lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" by (simp add: bin_to_bl_def bl_and_aux_bin) lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" by (simp add: bin_to_bl_def bl_or_aux_bin) lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" using bl_xor_aux_bin by (simp add: bin_to_bl_def) subsection \Type \<^typ>\'a word\\ lift_definition of_bl :: \bool list \ 'a::len word\ is bl_to_bin . lift_definition to_bl :: \'a::len word \ bool list\ is \bin_to_bl LENGTH('a)\ by (simp add: bl_to_bin_inj) lemma to_bl_eq: \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ for w :: \'a::len word\ by transfer simp lemma bit_of_bl_iff [bit_simps]: \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ by transfer (simp add: bin_nth_of_bl ac_simps) lemma rev_to_bl_eq: \rev (to_bl w) = map (bit w) [0.. for w :: \'a::len word\ apply (rule nth_equalityI) apply (simp add: to_bl.rep_eq) apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) done lemma to_bl_eq_rev: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ using rev_to_bl_eq [of w] apply (subst rev_is_rev_conv [symmetric]) apply (simp add: rev_map) done lemma of_bl_rev_eq: \of_bl (rev bs) = horner_sum of_bool 2 bs\ apply (rule bit_word_eqI) apply (simp add: bit_of_bl_iff) apply transfer apply (simp add: bit_horner_sum_bit_iff ac_simps) done lemma of_bl_eq: \of_bl bs = horner_sum of_bool 2 (rev bs)\ using of_bl_rev_eq [of \rev bs\] by simp lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ apply (rule bit_word_eqI) apply (auto simp add: bit_simps to_bl_eq_rev nth_append rev_nth nth_butlast not_less simp flip: bit_Suc) apply (metis Suc_pred len_gt_0 less_eq_decr_length_iff not_bit_length verit_la_disequality) done lemma length_to_bl_eq: \length (to_bl w) = LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma word_rotr_eq: \word_rotr n w = of_bl (rotater n (to_bl w))\ apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) done done lemma word_rotl_eq: \word_rotl n w = of_bl (rotate n (to_bl w))\ proof - have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ by (simp add: rotater_rev') then show ?thesis apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) done done qed lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by transfer (simp add: fun_eq_iff) \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (standard; transfer) apply (auto dest: sym) done global_interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl "{bl. length bl = LENGTH('a::len)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_to_bl_iff: \hd (to_bl w) \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ by (simp add: to_bl_eq_rev hd_map hd_rev) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" by (simp add: hd_to_bl_iff bit_last_iff bin_sign_def) lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: "bit (of_bl bl::'a::len word) n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by transfer simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by transfer simp lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by transfer simp lemma ucast_down_bl: \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ if \is_down (ucast :: 'a::len word \ 'b::len word)\ using that by transfer simp lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by transfer (simp add: bl_to_bin_app_cat) lemma ucast_of_bl_up: \ucast (of_bl bl :: 'a::len word) = of_bl bl\ if \size bl \ size (of_bl bl :: 'a::len word)\ using that apply transfer apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff) apply (subst (asm) trunc_bl2bin_len [symmetric]) apply (auto simp only: bit_take_bit_iff) done lemma word_rev_tf: "to_bl (of_bl bl::'a::len word) = rev (takefill False (LENGTH('a)) (rev bl))" by transfer (simp add: bl_bin_bl_rtf) lemma word_rep_drop: "to_bl (of_bl bl::'a::len word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len word) ::'a::len word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app: \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ for w :: \'a::len word\ using that by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma word_0_bl [simp]: "of_bl [] = 0" by transfer simp lemma word_1_bl: "of_bl [True] = 1" by transfer (simp add: bl_to_bin_def) lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by transfer (simp add: bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" by transfer (simp add: rbl_succ) lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" by transfer (simp add: rbl_pred) lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) lemma of_bl_length_less: \(of_bl x :: 'a::len word) < 2 ^ k\ if \length x = k\ \k < LENGTH('a)\ proof - from that have \length x < LENGTH('a)\ by simp then have \(of_bl x :: 'a::len word) < 2 ^ length x\ apply (simp add: of_bl_eq) apply transfer apply (simp add: take_bit_horner_sum_bit_eq) apply (subst length_rev [symmetric]) apply (simp only: horner_sum_of_bool_2_less) done with that show ?thesis by simp qed lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" by transfer (simp add: bl_not_bin) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_xor_bin) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_or_bin) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) lemma bin_nth_uint': "bit (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "bit w n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) lemma to_bl_nth: "n < size w \ to_bl w ! n = bit w (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: \map (bit w) [0.. for w :: \'a::len word\ proof (rule nth_equalityI) show \length (map (bit w) [0.. by simp fix m assume \m < length (map (bit w) [0.. then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed lemma to_bl_unfold: \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) lemma nth_rev_to_bl: \rev (to_bl w) ! n \ bit w n\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold) lemma nth_to_bl: \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold rev_nth) lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma [code abstract]: \Word.the_int (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (fact to_bl_eq_rev) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth nth_append nth_butlast nth_bin_to_bl simp flip: bit_Suc) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" apply (rule bit_word_eqI) apply (simp add: bit_simps) subgoal for n apply (cases n) apply simp_all done done lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" apply (rule bit_word_eqI) apply (simp add: bit_simps) subgoal for n apply (cases n) apply (simp_all add: nth_rev_to_bl) done done lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) lemma to_bl_double_eq: \to_bl (2 * w) = tl (to_bl w) @ [False]\ using bl_shiftl1 [of w] by (simp add: shiftl1_def ac_simps) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: \shiftr1 w = of_bl (butlast (to_bl w))\ proof (rule bit_word_eqI) fix n assume \n < LENGTH('a)\ show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ proof (cases \n = LENGTH('a) - 1\) case True then show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff) next case False with \n < LENGTH('a)\ have \n < LENGTH('a) - 1\ by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast word_size to_bl_nth) qed qed lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" proof (rule nth_equalityI) fix n assume \n < length (to_bl (sshiftr1 w))\ then have \n < LENGTH('a)\ by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) apply (simp_all add: to_bl_nth word_size hd_conv_nth bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (rule nth_equalityI) apply (simp_all add: word_size to_bl_nth bit_simps) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (rule nth_equalityI) apply (simp_all add: word_size nth_to_bl bit_simps) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" apply (rule nth_equalityI) apply (auto simp add: word_size to_bl_nth bit_simps dest: bit_imp_le_length) done lemma take_sshiftr': "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (cases n) apply (auto simp add: hd_to_bl_iff bit_simps not_less word_size) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_simps nth_Cons split: nat.split) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps nth_append) done lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" by (simp flip: shiftl_of_bl) lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" apply (rule bit_word_eqI) apply (simp add: bit_simps) apply (cases bl rule: rev_cases) apply auto done lemma shiftr_bl_of: "length bl \ LENGTH('a) \ (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" by (rule bit_word_eqI) (auto simp add: bit_simps rev_nth) lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add: bit_simps intro!: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size bit_simps) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done lemma slice1_eq_of_bl: \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ for w :: \'a::len word\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('b)\ show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ by (cases \m \ n\; cases \LENGTH('a) \ n\) (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) qed lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice_def) (* TODO: neg_numeral *) lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop) using drop_takefill apply force done lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) apply (metis diff_add_inverse) done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_eq_of_bl by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma revcast_eq_of_bl: \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ for w :: \'a::len word\ by (simp add: revcast_def slice1_eq_of_bl) lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" apply (rule nth_equalityI) apply simp apply (cases \LENGTH('a) \ LENGTH('b)\) apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" apply (rule bit_word_eqI) apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) done lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply transfer apply (simp add: bl_to_bin_app_cat bin_cat_num) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len word \ 'd::len word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done lemma word_rcat_bl: \word_rcat wl = of_bl (concat (map to_bl wl))\ proof - define ws where \ws = rev wl\ moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) apply transfer apply simp done ultimately show ?thesis by simp qed lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_rotl_eq to_bl_use_of_bl) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_rotr_eq to_bl_use_of_bl) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] end locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map end lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_eq_word_rotr_word_rotl) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe - apply (simp add: nat_mult_distrib) - apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj - [THEN conjunct2, THEN order_less_imp_le]] - nat_mod_distrib) - apply (simp add: nat_mod_distrib) + apply (auto simp add: nat_mult_distrib nat_mod_distrib) + using nat_0_le nat_minus_as_int zmod_int apply presburger done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemma max_word_bl: "to_bl (- 1::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto context includes bit_operations_syntax begin lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) declare bin_to_bl_def [simp] lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [bit x 0]" by (simp add: word_and_1) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [bit x 0]" using word_and_1_bl [of x] by (simp add: ac_simps) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" apply (rule bit_word_eqI) apply (auto simp: rev_nth bit_simps cong: rev_conj_cong) done lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" apply (auto simp add: to_bl_unfold) apply (rule bit_word_eqI) apply auto done lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: last_conv_nth bit_0 lsb_odd) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" by (simp add: is_aligned_mask eq_zero_set_bl bl_and_mask word_size) lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" apply (rule nth_equalityI) using assms apply (simp_all add: nth_append not_less word_size to_bl_nth is_aligned_imp_not_bit) done lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed lemma is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (rule nth_equalityI) apply (auto simp add: nth_append to_bl_nth word_size bit_simps not_less nth_Cons le_diff_conv) subgoal for i apply (cases \Suc (i + n) - LENGTH('a)\) apply simp_all done done lemma xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x XOR 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" apply (auto simp add: to_bl_eq_rev take_map drop_map take_rev drop_rev bit_simps) apply (rule nth_equalityI) apply (auto simp add: bit_simps rev_nth nth_append Suc_diff_Suc) done lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" apply (simp flip: push_bit_eq_mult subtract_mask(1) take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (simp add: take_bit_push_bit) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (simp flip: push_bit_eq_mult take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (rule nth_equalityI) apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma map_bits_rev_to_bl: "map (bit x) [0.. of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" proof - define ys where \ys = rev xs\ have \take_bit (length ys) (horner_sum of_bool 2 ys :: 'a word) = horner_sum of_bool 2 ys\ by transfer (simp add: take_bit_horner_sum_bit_eq min_def) then have \(of_bl (rev ys) :: 'a word) \ mask (length ys)\ by (simp only: of_bl_rev_eq less_eq_mask_iff_take_bit_eq_self) with ys_def show ?thesis by simp qed text\Some auxiliaries for sign-shifting by the entire word length or more\ lemma sshiftr_clamp_pos: assumes "LENGTH('a) \ n" "0 \ sint x" shows "(x::'a::len word) >>> n = 0" apply (rule bit_word_eqI) using assms apply (auto simp add: bit_simps bit_last_iff) done lemma sshiftr_clamp_neg: assumes "LENGTH('a) \ n" "sint x < 0" shows "(x::'a::len word) >>> n = -1" apply (rule bit_word_eqI) using assms apply (auto simp add: bit_simps bit_last_iff) done lemma sshiftr_clamp: assumes "LENGTH('a) \ n" shows "(x::'a::len word) >>> n = x >>> LENGTH('a)" apply (rule bit_word_eqI) using assms apply (auto simp add: bit_simps bit_last_iff) done text\ Like @{thm shiftr1_bl_of}, but the precondition is stronger because we need to pick the msb out of the list. \ lemma sshiftr1_bl_of: "length bl = LENGTH('a) \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd bl # butlast bl)" apply (rule word_bl.Rep_eqD) apply (subst bl_sshiftr1[of "of_bl bl :: 'a word"]) by (simp add: word_bl.Abs_inverse) text\ Like @{thm sshiftr1_bl_of}, with a weaker precondition. We still get a direct equation for @{term \sshiftr1 (of_bl bl)\}, it's just uglier. \ lemma sshiftr1_bl_of': "LENGTH('a) \ length bl \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd (drop (length bl - LENGTH('a)) bl) # butlast (drop (length bl - LENGTH('a)) bl))" apply (subst of_bl_drop'[symmetric, of "length bl - LENGTH('a)"]) using sshiftr1_bl_of[of "drop (length bl - LENGTH('a)) bl"] by auto text\ Like @{thm shiftr_bl_of}. \ lemma sshiftr_bl_of: assumes "length bl = LENGTH('a)" shows "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - from assms obtain b bs where \bl = b # bs\ by (cases bl) simp_all then have *: \bl ! 0 \ b\ \hd bl \ b\ by simp_all show ?thesis apply (rule bit_word_eqI) using assms * by (auto simp add: bit_simps nth_append rev_nth not_less) qed text\Like @{thm shiftr_bl}\ lemma sshiftr_bl: "x >>> n \ of_bl (replicate n (msb x) @ take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" unfolding word_msb_alt by (smt (z3) length_to_bl_eq sshiftr_bl_of word_bl.Rep_inverse) end lemma of_bl_drop_eq_take_bit: \of_bl (drop n xs) = take_bit (length xs - n) (of_bl xs)\ by (simp add: of_bl_drop take_bit_eq_mask) lemma of_bl_take_to_bl_eq_drop_bit: \of_bl (take n (to_bl w)) = drop_bit (LENGTH('a) - n) w\ if \n \ LENGTH('a)\ for w :: \'a::len word\ using that shiftr_bl [of w \LENGTH('a) - n\] by (simp add: shiftr_def) end