diff --git a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy --- a/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Berlekamp_Type_Based.thy @@ -1,3145 +1,3145 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) section \The Berlekamp Algorithm\ theory Berlekamp_Type_Based imports Jordan_Normal_Form.Matrix_Kernel Jordan_Normal_Form.Gauss_Jordan_Elimination Jordan_Normal_Form.Missing_VectorSpace Polynomial_Factorization.Square_Free_Factorization Polynomial_Factorization.Missing_Multiset Finite_Field Chinese_Remainder_Poly Poly_Mod_Finite_Field "HOL-Computational_Algebra.Field_as_Ring" begin hide_const (open) up_ring.coeff up_ring.monom Modules.module subspace Modules.module_hom subsection \Auxiliary lemmas\ context fixes g :: "'b \ 'a :: comm_monoid_mult" begin lemma prod_list_map_filter: "prod_list (map g (filter f xs)) * prod_list (map g (filter (\ x. \ f x) xs)) = prod_list (map g xs)" by (induct xs, auto simp: ac_simps) lemma prod_list_map_partition: assumes "List.partition f xs = (ys, zs)" shows "prod_list (map g xs) = prod_list (map g ys) * prod_list (map g zs)" using assms by (subst prod_list_map_filter[symmetric, of _ f], auto simp: o_def) end lemma coprime_id_is_unit: fixes a::"'b::semiring_gcd" shows "coprime a a \ is_unit a" using dvd_unit_imp_unit by auto lemma dim_vec_of_list[simp]: "dim_vec (vec_of_list x) = length x" by (transfer, auto) lemma length_list_of_vec[simp]: "length (list_of_vec A) = dim_vec A" by (transfer', auto) lemma list_of_vec_vec_of_list[simp]: "list_of_vec (vec_of_list a) = a" proof - { fix aa :: "'a list" have "map (\n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0..n. if n < length aa then aa ! n else undef_vec (n - length aa)) [0.. list_of_vec) {v. dim_vec v = n}" proof (rule comp_inj_on) show "inj_on list_of_vec {v. dim_vec v = n}" by (auto simp add: inj_on_def, transfer, auto simp add: mk_vec_def) show "inj_on Poly (list_of_vec ` {v. dim_vec v = n})" proof (auto simp add: inj_on_def) fix x y::"'c vec" assume "n = dim_vec x" and dim_xy: "dim_vec y = dim_vec x" and Poly_eq: "Poly (list_of_vec x) = Poly (list_of_vec y)" note [simp del] = nth_list_of_vec show "list_of_vec x = list_of_vec y" proof (rule nth_equalityI, auto simp: dim_xy) have length_eq: "length (list_of_vec x ) = length (list_of_vec y)" using dim_xy by (transfer, auto) fix i assume "i < dim_vec x" thus "list_of_vec x ! i = list_of_vec y ! i" using Poly_eq unfolding poly_eq_iff coeff_Poly_eq using dim_xy unfolding nth_default_def by (auto, presburger) qed qed qed corollary inj_Poly_list_of_vec: "inj_on (Poly \ list_of_vec) (carrier_vec n)" using inj_Poly_list_of_vec' unfolding carrier_vec_def . lemma list_of_vec_rw_map: "list_of_vec m = map (\n. m $ n) [0.. []" shows "degree (Poly xs) < length xs" using xs by (induct xs, auto intro: Poly.simps(1)) lemma vec_of_list_list_of_vec[simp]: "vec_of_list (list_of_vec a) = a" by (transfer, auto simp add: mk_vec_def) lemma row_mat_of_rows_list: assumes b: "b < length A" and nc: "\i. i < length A \ length (A ! i) = nc" shows "(row (mat_of_rows_list nc A) b) = vec_of_list (A ! b)" proof (auto simp add: vec_eq_iff) show "dim_col (mat_of_rows_list nc A) = length (A ! b)" unfolding mat_of_rows_list_def using b nc by auto fix i assume i: "i < length (A ! b)" show "row (mat_of_rows_list nc A) b $ i = vec_of_list (A ! b) $ i" using i b nc unfolding mat_of_rows_list_def row_def by (transfer, auto simp add: mk_vec_def mk_mat_def) qed lemma degree_Poly_list_of_vec: assumes n: "x \ carrier_vec n" and n0: "n > 0" shows "degree (Poly (list_of_vec x)) < n" proof - have x_dim: "dim_vec x = n" using n by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] n0 n x_dim) have "degree (Poly (list_of_vec x)) < length (list_of_vec x)" by (rule degree_Poly'[OF l]) also have "... = n" using x_dim by auto finally show ?thesis . qed lemma list_of_vec_nth: assumes i: "i < dim_vec x" shows "list_of_vec x ! i = x $ i" using i by (transfer, auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth': assumes i: "i < dim_vec x" shows "coeff (Poly (list_of_vec x)) i = x $ i" using i by (auto simp add: list_of_vec_nth nth_default_def) lemma list_of_vec_row_nth: assumes x: "x < dim_col A" shows "list_of_vec (row A i) ! x = A $$ (i, x)" using x unfolding row_def by (transfer', auto simp add: mk_vec_def) lemma coeff_Poly_list_of_vec_nth: assumes x: "x < dim_col A" shows "coeff (Poly (list_of_vec (row A i))) x = A $$ (i, x)" proof - have "coeff (Poly (list_of_vec (row A i))) x = nth_default 0 (list_of_vec (row A i)) x" unfolding coeff_Poly_eq by simp also have "... = A $$ (i, x)" using x list_of_vec_row_nth unfolding nth_default_def by (auto simp del: nth_list_of_vec) finally show ?thesis . qed lemma inj_on_list_of_vec: "inj_on list_of_vec (carrier_vec n)" unfolding inj_on_def unfolding list_of_vec_rw_map by auto lemma vec_of_list_carrier[simp]: "vec_of_list x \ carrier_vec (length x)" unfolding carrier_vec_def by simp lemma card_carrier_vec: "card (carrier_vec n:: 'b::finite vec set) = CARD('b) ^ n" proof - let ?A = "UNIV::'b set" let ?B = "{xs. set xs \ ?A \ length xs = n}" let ?C = "(carrier_vec n:: 'b::finite vec set)" have "card ?C = card ?B" proof - have "bij_betw (list_of_vec) ?C ?B" proof (unfold bij_betw_def, auto) show "inj_on list_of_vec (carrier_vec n)" by (rule inj_on_list_of_vec) fix x::"'b list" assume n: "n = length x" thus "x \ list_of_vec ` carrier_vec (length x)" unfolding image_def by auto (rule bexI[of _ "vec_of_list x"], auto) qed thus ?thesis using bij_betw_same_card by blast qed also have "... = card ?A ^ n" by (rule card_lists_length_eq, simp) finally show ?thesis . qed lemma finite_carrier_vec[simp]: "finite (carrier_vec n:: 'b::finite vec set)" by (rule card_ge_0_finite, unfold card_carrier_vec, auto) lemma row_echelon_form_dim0_row: assumes "A \ carrier_mat 0 n" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_dim0_col: assumes "A \ carrier_mat n 0" shows "row_echelon_form A" using assms unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma row_echelon_form_one_dim0[simp]: "row_echelon_form (1\<^sub>m 0)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto lemma Poly_list_of_vec_0[simp]: "Poly (list_of_vec (0\<^sub>v 0)) = [:0:]" by (simp add: poly_eq_iff nth_default_def) lemma monic_normalize: assumes "(p :: 'b :: {field,euclidean_ring_gcd} poly) \ 0" shows "monic (normalize p)" by (simp add: assms normalize_poly_old_def) lemma exists_factorization_prod_list: fixes P::"'b::field poly list" assumes "degree (prod_list P) > 0" and "\ u. u \ set P \ degree u > 0 \ monic u" and "square_free (prod_list P)" shows "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" using assms proof (induct P) case Nil thus ?case by auto next case (Cons x P) have sf_P: "square_free (prod_list P)" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons mult.commute square_free_factor) have deg_x: "degree x > 0" using Cons.prems by auto have distinct_P: "distinct P" by (meson Cons.prems(2) Cons.prems(3) distinct.simps(2) square_free_prod_list_distinct) have "\A. finite A \ x = \A \ A \ {q. irreducible q \ monic q}" proof (rule monic_square_free_irreducible_factorization) show "monic x" by (simp add: Cons.prems(2)) show "square_free x" by (metis Cons.prems(3) dvd_triv_left prod_list.Cons square_free_factor) qed from this obtain A where fin_A: "finite A" and xA: "x = \A" and A: "A \ {q. irreducible\<^sub>d q \ monic q}" by auto obtain A' where s: "set A' = A" and length_A': "length A' = card A" using \finite A\ distinct_card finite_distinct_list by force have A_not_empty: "A \ {}" using xA deg_x by auto have x_prod_list_A': "x = prod_list A'" proof - have "x = \A" using xA by simp also have "... = prod id A" by simp also have "... = prod id (set A')" unfolding s by simp also have "... = prod_list (map id A')" by (rule prod.distinct_set_conv_list, simp add: card_distinct length_A' s) also have "... = prod_list A'" by auto finally show ?thesis . qed show ?case proof (cases "P = []") case True show ?thesis proof (rule exI[of _ "A'"], auto simp add: True) show "prod_list A' = x" using x_prod_list_A' by simp show "Suc 0 \ length A'" using A_not_empty using s length_A' by (simp add: Suc_leI card_gt_0_iff fin_A) show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed next case False have hyp: "\Q. prod_list Q = prod_list P \ length P \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule Cons.hyps[OF _ _ sf_P]) have set_P: "set P \ {}" using False by auto have "prod_list P = prod_list (map id P)" by simp also have "... = prod id (set P)" using prod.distinct_set_conv_list[OF distinct_P, of id] by simp also have "... = \(set P)" by simp finally have "prod_list P = \(set P)" . hence "degree (prod_list P) = degree (\(set P))" by simp also have "... = degree (prod id (set P))" by simp also have "... = (\i\(set P). degree (id i))" proof (rule degree_prod_eq_sum_degree) show "\i\set P. id i \ 0" using Cons.prems(2) by force qed also have "... > 0" by (metis Cons.prems(2) List.finite_set set_P gr0I id_apply insert_iff list.set(2) sum_pos) finally show "degree (prod_list P) > 0" by simp show "\u. u \ set P \ degree u > 0 \ monic u" using Cons.prems by auto qed from this obtain Q where QP: "prod_list Q = prod_list P" and length_PQ: "length P \ length Q" and monic_irr_Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast show ?thesis proof (rule exI[of _ "A' @ Q"], auto simp add: monic_irr_Q) show "prod_list A' * prod_list Q = x * prod_list P" unfolding QP x_prod_list_A' by auto have "length A' \ 0" using A_not_empty using s length_A' by auto thus "Suc (length P) \ length A' + length Q" using QP length_PQ by linarith show "\u. u \ set A' \ irreducible u" using s A by auto show "\u. u \ set A' \ monic u" using s A by auto qed qed qed lemma normalize_eq_imp_smult: fixes p :: "'b :: {euclidean_ring_gcd} poly" assumes n: "normalize p = normalize q" shows "\ c. c \ 0 \ q = smult c p" proof(cases "p = 0") case True with n show ?thesis by (auto intro:exI[of _ 1]) next case p0: False have degree_eq: "degree p = degree q" using n degree_normalize by metis hence q0: "q \ 0" using p0 n by auto have p_dvd_q: "p dvd q" using n by (simp add: associatedD1) from p_dvd_q obtain k where q: "q = k * p" unfolding dvd_def by (auto simp: ac_simps) with q0 have "k \ 0" by auto then have "degree k = 0" using degree_eq degree_mult_eq p0 q by fastforce then obtain c where k: "k = [: c :]" by (metis degree_0_id) with \k \ 0\ have "c \ 0" by auto have "q = smult c p" unfolding q k by simp with \c \ 0\ show ?thesis by auto qed lemma prod_list_normalize: fixes P :: "'b :: {idom_divide,normalization_semidom_multiplicative} poly list" shows "normalize (prod_list P) = prod_list (map normalize P)" proof (induct P) case Nil show ?case by auto next case (Cons p P) have "normalize (prod_list (p # P)) = normalize p * normalize (prod_list P)" using normalize_mult by auto also have "... = normalize p * prod_list (map normalize P)" using Cons.hyps by auto also have "... = prod_list (normalize p # (map normalize P))" by auto also have "... = prod_list (map normalize (p # P))" by auto finally show ?case . qed lemma prod_list_dvd_prod_list_subset: fixes A::"'b::comm_monoid_mult list" assumes dA: "distinct A" and dB: "distinct B" (*Maybe this condition could be avoided*) and s: "set A \ set B" shows "prod_list A dvd prod_list B" proof - have "prod_list A = prod_list (map id A)" by auto also have "... = prod id (set A)" by (rule prod.distinct_set_conv_list[symmetric, OF dA]) also have "... dvd prod id (set B)" by (rule prod_dvd_prod_subset[OF _ s], auto) also have "... = prod_list (map id B)" by (rule prod.distinct_set_conv_list[OF dB]) also have "... = prod_list B" by simp finally show ?thesis . qed end lemma gcd_monic_constant: "gcd f g \ {1, f}" if "monic f" and "degree g = 0" for f g :: "'a :: {field_gcd} poly" proof (cases "g = 0") case True moreover from \monic f\ have "normalize f = f" by (rule normalize_monic) ultimately show ?thesis by simp next case False with \degree g = 0\ have "is_unit g" by simp then have "Rings.coprime f g" by (rule is_unit_right_imp_coprime) then show ?thesis by simp qed lemma distinct_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "distinct (find_base_vectors A)" proof - note non_pivot_base = non_pivot_base[OF ref A] let ?pp = "set (pivot_positions A)" from A have dim: "dim_row A = nr" "dim_col A = nc" by auto { fix j j' assume j: "j < nc" "j \ snd ` ?pp" and j': "j' < nc" "j' \ snd ` ?pp" and neq: "j' \ j" from non_pivot_base(2)[OF j] non_pivot_base(4)[OF j' j neq] have "non_pivot_base A (pivot_positions A) j \ non_pivot_base A (pivot_positions A) j'" by auto } hence inj: "inj_on (non_pivot_base A (pivot_positions A)) (set [j\[0.. snd ` ?pp])" unfolding inj_on_def by auto thus ?thesis unfolding find_base_vectors_def Let_def unfolding distinct_map dim by auto qed lemma length_find_base_vectors: fixes A::"'a::field mat" assumes ref: "row_echelon_form A" and A: "A \ carrier_mat nr nc" shows "length (find_base_vectors A) = card (set (find_base_vectors A))" using distinct_card[OF distinct_find_base_vectors[OF ref A]] by auto subsection \Previous Results\ definition power_poly_f_mod :: "'a::field poly \ 'a poly \ nat \ 'a poly" where "power_poly_f_mod modulus = (\a n. a ^ n mod modulus)" lemma power_poly_f_mod_binary: "power_poly_f_mod m a n = (if n = 0 then 1 mod m - else let (d, r) = Divides.divmod_nat n 2; + else let (d, r) = Euclidean_Division.divmod_nat n 2; rec = power_poly_f_mod m ((a * a) mod m) d in if r = 0 then rec else (rec * a) mod m)" for m a :: "'a :: {field_gcd} poly" proof - note d = power_poly_f_mod_def show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding d by simp next case False - obtain q r where div: "Divides.divmod_nat n 2 = (q,r)" by force - hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding divmod_nat_def by auto + obtain q r where div: "Euclidean_Division.divmod_nat n 2 = (q,r)" by force + hence n: "n = 2 * q + r" and r: "r = 0 \ r = 1" unfolding Euclidean_Division.divmod_nat_def by auto have id: "a ^ (2 * q) = (a * a) ^ q" by (simp add: power_mult_distrib semiring_normalization_rules) show ?thesis proof (cases "r = 0") case True show ?thesis using power_mod [of "a * a" m q] - by (auto simp add: divmod_nat_def Let_def True n d div id) + by (auto simp add: Euclidean_Division.divmod_nat_def Let_def True n d div id) next case False with r have r: "r = 1" by simp show ?thesis by (auto simp add: d r div Let_def mod_simps) (simp add: n r mod_simps ac_simps power_mult_distrib power_mult power2_eq_square) qed qed qed fun power_polys where "power_polys mul_p u curr_p (Suc i) = curr_p # power_polys mul_p u ((curr_p * mul_p) mod u) i" | "power_polys mul_p u curr_p 0 = []" context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma fermat_theorem_mod_ring [simp]: fixes a::"'a mod_ring" shows "a ^ CARD('a) = a" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis proof transfer fix a assume "a \ {0.. 0" then have a: "1 \ a" "a < int CARD('a)" by simp_all then have [simp]: "a mod int CARD('a) = a" by simp from a have "\ int CARD('a) dvd a" by (auto simp add: zdvd_not_zless) then have "\ CARD('a) dvd nat \a\" by simp with a have "\ CARD('a) dvd nat a" by simp with prime_card have "[nat a ^ (CARD('a) - 1) = 1] (mod CARD('a))" by (rule fermat_theorem) with a have "int (nat a ^ (CARD('a) - 1) mod CARD('a)) = 1" by (simp add: cong_def) with a have "a ^ (CARD('a) - 1) mod CARD('a) = 1" by (simp add: of_nat_mod) then have "a * (a ^ (CARD('a) - 1) mod int CARD('a)) = a" by simp then have "(a * (a ^ (CARD('a) - 1) mod int CARD('a))) mod int CARD('a) = a mod int CARD('a)" by (simp only:) then show "a ^ CARD('a) mod int CARD('a) = a" by (simp add: mod_simps semiring_normalization_rules(27)) qed qed lemma mod_eq_dvd_iff_poly: "((x::'a mod_ring poly) mod n = y mod n) = (n dvd x - y)" proof assume H: "x mod n = y mod n" hence "x mod n - y mod n = 0" by simp hence "(x mod n - y mod n) mod n = 0" by simp hence "(x - y) mod n = 0" by (simp add: mod_diff_eq) thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) next assume H: "n dvd x - y" then obtain k where k: "x-y = n*k" unfolding dvd_def by blast hence "x = n*k + y" using diff_eq_eq by blast hence "x mod n = (n*k + y) mod n" by simp thus "x mod n = y mod n" by (simp add: mod_add_left_eq) qed lemma cong_gcd_eq_poly: "gcd a m = gcd b m" if "[(a::'a mod_ring poly) = b] (mod m)" using that by (simp add: cong_def) (metis gcd_mod_left mod_by_0) lemma coprime_h_c_poly: fixes h::"'a mod_ring poly" assumes "c1 \ c2" shows "coprime (h - [:c1:]) (h - [:c2:])" proof (intro coprimeI) fix d assume "d dvd h - [:c1:]" and "d dvd h - [:c2:]" hence "h mod d = [:c1:] mod d" and "h mod d = [:c2:] mod d" using mod_eq_dvd_iff_poly by simp+ hence "[:c1:] mod d = [:c2:] mod d" by simp hence "d dvd [:c2 - c1:]" by (metis (no_types) mod_eq_dvd_iff_poly diff_pCons right_minus_eq) thus "is_unit d" by (metis (no_types) assms dvd_trans is_unit_monom_0 monom_0 right_minus_eq) qed lemma coprime_h_c_poly2: fixes h::"'a mod_ring poly" assumes "coprime (h - [:c1:]) (h - [:c2:])" and "\ is_unit (h - [:c1:])" shows "c1 \ c2" using assms coprime_id_is_unit by blast lemma degree_minus_eq_right: fixes p::"'b::ab_group_add poly" shows "degree q < degree p \ degree (p - q) = degree p" using degree_add_eq_left[of "-q" p] degree_minus by auto lemma coprime_prod: fixes A::"'a mod_ring set" and g::"'a mod_ring \ 'a mod_ring poly" assumes "\x\A. coprime (g a) (g x)" shows "coprime (g a) (prod (\x. g x) A)" proof - have f: "finite A" by simp show ?thesis using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = (g x) * (\c\A. g c)" by (simp add: insert.hyps(2)) with insert.prems show ?case by (auto simp: insert.hyps(3) prod_coprime_right) qed auto qed lemma coprime_prod2: fixes A::"'b::semiring_gcd set" assumes "\x\A. coprime (a) (x)" and f: "finite A" shows "coprime (a) (prod (\x. x) A)" using f using assms proof (induct A) case (insert x A) have "(\c\insert x A. c) = (x) * (\c\A. c)" by (simp add: insert.hyps) with insert.prems show ?case by (simp add: insert.hyps prod_coprime_right) qed auto lemma divides_prod: fixes g::"'a mod_ring \ 'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2 \ coprime (g c1) (g c2)" assumes "\c\ A. g c dvd f" shows "(\c\A. g c) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. g c) = g x * (\c\ A. g c)" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "g x dvd f" using insert.prems by auto show "prod g A dvd f" using insert.hyps(3) insert.prems by auto from insert show "Rings.coprime (g x) (prod g A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto qed (* Proof of equation 9 in the book by Knuth x^p - x = (x-0)(x-1)...(x-(p-1)) (mod p) *) lemma poly_monom_identity_mod_p: "monom (1::'a mod_ring) (CARD('a)) - monom 1 1 = prod (\x. [:0,1:] - [:x:]) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof - let ?f="(\x::'a mod_ring. [:0,1:] - [:x:])" have "?rhs dvd ?lhs" proof (rule divides_prod) { fix a::"'a mod_ring" have "poly ?lhs a = 0" by (simp add: poly_monom) hence "([:0,1:] - [:a:]) dvd ?lhs" using poly_eq_0_iff_dvd by fastforce } thus "\x\UNIV::'a mod_ring set. [:0, 1:] - [:x:] dvd monom 1 CARD('a) - monom 1 1" by fast show "\c1 c2. c1 \ UNIV \ c2 \ UNIV \ c1 \ (c2 :: 'a mod_ring) \ coprime ([:0, 1:] - [:c1:]) ([:0, 1:] - [:c2:])" by (auto dest!: coprime_h_c_poly[of _ _ "[:0,1:]"]) qed from this obtain g where g: "?lhs = ?rhs * g" using dvdE by blast have degree_lhs_card: "degree ?lhs = CARD('a)" proof - have "degree (monom (1::'a mod_ring) 1) = 1" by (simp add: degree_monom_eq) moreover have d_c: "degree (monom (1::'a mod_ring) CARD('a)) = CARD('a)" by (simp add: degree_monom_eq) ultimately have "degree (monom (1::'a mod_ring) 1) < degree (monom (1::'a mod_ring) CARD('a))" using prime_card unfolding prime_nat_iff by auto hence "degree ?lhs = degree (monom (1::'a mod_ring) CARD('a))" by (rule degree_minus_eq_right) thus ?thesis unfolding d_c . qed have degree_rhs_card: "degree ?rhs = CARD('a)" proof - have "degree (prod ?f UNIV) = sum (degree \ ?f) UNIV \ coeff (prod ?f UNIV) (sum (degree \ ?f) UNIV) = 1" by (rule degree_prod_sum_monic, auto) moreover have "sum (degree \ ?f) UNIV = CARD('a)" by auto ultimately show ?thesis by presburger qed have monic_lhs: "monic ?lhs" using degree_lhs_card by auto have monic_rhs: "monic ?rhs" by (rule monic_prod, simp) have degree_eq: "degree ?rhs = degree ?lhs" unfolding degree_lhs_card degree_rhs_card .. have g_not_0: "g \ 0" using g monic_lhs by auto have degree_g0: "degree g = 0" proof - have "degree (?rhs * g) = degree ?rhs + degree g" by (rule degree_monic_mult[OF monic_rhs g_not_0]) thus ?thesis using degree_eq g by simp qed have monic_g: "monic g" using monic_factor g monic_lhs monic_rhs by auto have "g = 1" using monic_degree_0[OF monic_g] degree_g0 by simp thus ?thesis using g by auto qed (* Proof of equation 10 in the book by Knuth v(x)^p - v(x) = (v(x)-0)(v(x)-1)...(v(x)-(p-1)) (mod p) *) lemma poly_identity_mod_p: "v^(CARD('a)) - v = prod (\x. v - [:x:]) (UNIV::'a mod_ring set)" proof - have id: "monom 1 1 \\<^sub>p v = v" "[:0, 1:] \\<^sub>p v = v" unfolding pcompose_def apply (auto) by (simp add: fold_coeffs_def) have id2: "monom 1 (CARD('a)) \\<^sub>p v = v ^ (CARD('a))" by (metis id(1) pcompose_hom.hom_power x_pow_n) show ?thesis using arg_cong[OF poly_monom_identity_mod_p, of "\ f. f \\<^sub>p v"] unfolding pcompose_hom.hom_minus pcompose_hom.hom_prod id pcompose_const id2 . qed lemma coprime_gcd: fixes h::"'a mod_ring poly" assumes "Rings.coprime (h-[:c1:]) (h-[:c2:])" shows "Rings.coprime (gcd f(h-[:c1:])) (gcd f (h-[:c2:]))" using assms coprime_divisors by blast lemma divides_prod_gcd: fixes h::"'a mod_ring poly" assumes "\c1 c2. c1 \ A \ c2 \ A \ c1 \ c2\ coprime (h-[:c1:]) (h-[:c2:])" shows "(\c\A. gcd f (h - [:c:])) dvd f" proof - have finite_A: "finite A" using finite[of A] . thus ?thesis using assms proof (induct A) case (insert x A) have "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps(2)) also have "... dvd f" proof (rule divides_mult) show "gcd f (h - [:x:]) dvd f" by simp show "(\c\A. gcd f (h - [:c:])) dvd f" using insert.hyps(3) insert.prems by auto show "Rings.coprime (gcd f (h - [:x:])) (\c\A. gcd f (h - [:c:]))" by (rule prod_coprime_right) (metis Berlekamp_Type_Based.coprime_h_c_poly coprime_gcd coprime_iff_coprime insert.hyps(2)) qed finally show ?case . qed auto qed lemma monic_prod_gcd: assumes f: "finite A" and f0: "(f :: 'b :: {field_gcd} poly) \ 0" shows "monic (\c\A. gcd f (h - [:c:]))" using f proof (induct A) case (insert x A) have rw: "(\c\insert x A. gcd f (h - [:c:])) = (gcd f (h - [:x:])) * (\c\ A. gcd f (h - [:c:]))" by (simp add: insert.hyps) show ?case proof (unfold rw, rule monic_mult) show "monic (gcd f (h - [:x:]))" using poly_gcd_monic[of f] f0 using insert.prems insert_iff by blast show "monic (\c\A. gcd f (h - [:c:]))" using insert.hyps(3) insert.prems by blast qed qed auto lemma coprime_not_unit_not_dvd: fixes a::"'b::semiring_gcd" assumes "a dvd b" and "coprime b c" and "\ is_unit a" shows "\ a dvd c" using assms coprime_divisors coprime_id_is_unit by fastforce lemma divides_prod2: fixes A::"'b::semiring_gcd set" assumes f: "finite A" and "\a\A. a dvd c" and "\a1 a2. a1 \ A \ a2 \ A \ a1 \ a2 \ coprime a1 a2" shows "\A dvd c" using assms proof (induct A) case (insert x A) have "\(insert x A) = x * \A" by (simp add: insert.hyps(1) insert.hyps(2)) also have "... dvd c" proof (rule divides_mult) show "x dvd c" by (simp add: insert.prems) show "\A dvd c" using insert by auto from insert show "Rings.coprime x (\A)" by (auto intro: prod_coprime_right) qed finally show ?case . qed auto lemma coprime_polynomial_factorization: fixes a1 :: "'b :: {field_gcd} poly" assumes irr: "as \ {q. irreducible q \ monic q}" and "finite as" and a1: "a1 \ as" and a2: "a2 \ as" and a1_not_a2: "a1 \ a2" shows "coprime a1 a2" proof (rule ccontr) assume not_coprime: "\ coprime a1 a2" let ?b= "gcd a1 a2" have b_dvd_a1: "?b dvd a1" and b_dvd_a2: "?b dvd a2" by simp+ have irr_a1: "irreducible a1" using a1 irr by blast have irr_a2: "irreducible a2" using a2 irr by blast have a2_not0: "a2 \ 0" using a2 irr by auto have degree_a1: "degree a1 \ 0" using irr_a1 by auto have degree_a2: "degree a2 \ 0" using irr_a2 by auto have not_a2_dvd_a1: "\ a2 dvd a1" proof (rule ccontr, simp) assume a2_dvd_a1: "a2 dvd a1" from this obtain k where k: "a1 = a2 * k" unfolding dvd_def by auto have k_not0: "k \ 0" using degree_a1 k by auto show False proof (cases "degree a2 = degree a1") case False have "degree a2 < degree a1" using False a2_dvd_a1 degree_a1 divides_degree by fastforce hence "\ irreducible a1" using degree_a2 a2_dvd_a1 degree_a2 by (metis degree_a1 irreducible\<^sub>dD(2) irreducible\<^sub>d_multD irreducible_connect_field k neq0_conv) thus False using irr_a1 by contradiction next case True have "degree a1 = degree a2 + degree k" unfolding k using degree_mult_eq[OF a2_not0 k_not0] by simp hence "degree k = 0" using True by simp hence "k = 1" using monic_factor a1 a2 irr k monic_degree_0 by auto hence "a1 = a2" using k by simp thus False using a1_not_a2 by contradiction qed qed have b_not0: "?b \ 0" by (simp add: a2_not0) have degree_b: "degree ?b > 0" using not_coprime[simplified] b_not0 is_unit_gcd is_unit_iff_degree by blast have "degree ?b < degree a2" by (meson b_dvd_a1 b_dvd_a2 irreducibleD' dvd_trans gcd_dvd_1 irr_a2 not_a2_dvd_a1 not_coprime) hence "\ irreducible\<^sub>d a2" using degree_a2 b_dvd_a2 degree_b by (metis degree_smult_eq irreducible\<^sub>d_dvd_smult less_not_refl3) thus False using irr_a2 by auto qed (* Proof of equation 14 in the book by Knuth *) theorem Berlekamp_gcd_step: fixes f::"'a mod_ring poly" and h::"'a mod_ring poly" assumes hq_mod_f: "[h^(CARD('a)) = h] (mod f)" and monic_f: "monic f" and sf_f: "square_free f" shows "f = prod (\c. gcd f (h - [:c:])) (UNIV::'a mod_ring set)" (is "?lhs = ?rhs") proof (cases "f=0") case True thus ?thesis using coeff_0 monic_f zero_neq_one by auto next case False note f_not_0 = False show ?thesis proof (rule poly_dvd_antisym) show "?rhs dvd f" using coprime_h_c_poly by (intro divides_prod_gcd, auto) have "monic ?rhs" by (rule monic_prod_gcd[OF _ f_not_0], simp) thus "coeff f (degree f) = coeff ?rhs (degree ?rhs)" using monic_f by auto next show "f dvd ?rhs" proof - let ?p = "CARD('a)" obtain P where finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" using monic_square_free_irreducible_factorization[OF monic_f sf_f] by auto have f_dvd_hqh: "f dvd (h^?p - h)" using hq_mod_f unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "f = \P" using f_desc_square_free by simp also have "... dvd ?rhs" proof (rule divides_prod2[OF finite_P]) show "\a1 a2. a1 \ P \ a2 \ P \ a1 \ a2 \ coprime a1 a2" using coprime_polynomial_factorization[OF P finite_P] by simp show "\a\P. a dvd (\c\UNIV. gcd f (h - [:c:]))" proof fix fi assume fi_P: "fi \ P" show "fi dvd ?rhs" proof (rule dvd_prod, auto) show "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible (fi)" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) have fi_dvd_hc: "\c\UNIV::'a mod_ring set. fi dvd (h-[:c:])" by (rule irreducible_dvd_prod[OF _ fi_dvd_prod_hc], simp add: irr_fi) thus "\c. fi dvd h - [:c:]" by simp qed qed qed finally show "f dvd ?rhs" . qed qed qed (******* Implementation of Berlekamp's algorithm (type-based version) *******) subsection \Definitions\ definition berlekamp_mat :: "'a mod_ring poly \ 'a mod_ring mat" where "berlekamp_mat u = (let n = degree u; mul_p = power_poly_f_mod u [:0,1:] (CARD('a)); xks = power_polys mul_p u 1 n in mat_of_rows_list n (map (\ cs. let coeffs_cs = (coeffs cs); k = n - length (coeffs cs) in (coeffs cs) @ replicate k 0) xks))" definition berlekamp_resulting_mat :: "('a mod_ring) poly \ 'a mod_ring mat" where "berlekamp_resulting_mat u = (let Q = berlekamp_mat u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then Q $$ (i,j) - 1 else Q $$ (i,j)) in (gauss_jordan_single (transpose_mat QI)))" definition berlekamp_basis :: "'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_basis u = (map (Poly o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" lemma berlekamp_basis_code[code]: "berlekamp_basis u = (map (poly_of_list o list_of_vec) (find_base_vectors (berlekamp_resulting_mat u)))" unfolding berlekamp_basis_def poly_of_list_def .. primrec berlekamp_factorization_main :: "nat \ 'a mod_ring poly list \ 'a mod_ring poly list \ nat \ 'a mod_ring poly list" where "berlekamp_factorization_main i divs (v # vs) n = (if v = 1 then berlekamp_factorization_main i divs vs n else if length divs = n then divs else let facts = [ w . u \ divs, s \ [0 ..< CARD('a)], w \ [gcd u (v - [:of_int s:])], w \ 1]; (lin,nonlin) = List.partition (\ q. degree q = i) facts in lin @ berlekamp_factorization_main i nonlin vs (n - length lin))" | "berlekamp_factorization_main i divs [] n = divs" definition berlekamp_monic_factorization :: "nat \ 'a mod_ring poly \ 'a mod_ring poly list" where "berlekamp_monic_factorization d f = (let vs = berlekamp_basis f; n = length vs; fs = berlekamp_factorization_main d [f] vs n in fs)" subsection \Properties\ lemma power_polys_works: fixes u::"'b::unique_euclidean_semiring" assumes i: "i < n" and c: "curr_p = curr_p mod u" (*Equivalent to degree curr_p < degree u*) shows "power_polys mult_p u curr_p n ! i = curr_p * mult_p ^ i mod u" using i c proof (induct n arbitrary: curr_p i) case 0 thus ?case by simp next case (Suc n) have p_rw: "power_polys mult_p u curr_p (Suc n) ! i = (curr_p # power_polys mult_p u (curr_p * mult_p mod u) n) ! i" by simp show ?case proof (cases "i=0") case True show ?thesis using Suc.prems unfolding p_rw True by auto next case False note i_not_0 = False show ?thesis proof (cases "i < n") case True note i_less_n = True have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (i - 1)" unfolding p_rw using nth_Cons_pos False by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (i-1) mod u" by (rule Suc.hyps) (auto simp add: i_less_n less_imp_diff_less) also have "... = curr_p * mult_p ^ i mod u" using False by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . next case False hence i_n: "i = n" using Suc.prems by auto have "power_polys mult_p u curr_p (Suc n) ! i = power_polys mult_p u (curr_p * mult_p mod u) n ! (n - 1)" unfolding p_rw using nth_Cons_pos i_n i_not_0 by auto also have "... = (curr_p * mult_p mod u) * mult_p ^ (n-1) mod u" proof (rule Suc.hyps) show "n - 1 < n" using i_n i_not_0 by linarith show "curr_p * mult_p mod u = curr_p * mult_p mod u mod u" by simp qed also have "... = curr_p * mult_p ^ i mod u" using i_n [symmetric] i_not_0 by (cases i) (simp_all add: algebra_simps mod_simps) finally show ?thesis . qed qed qed lemma length_power_polys[simp]: "length (power_polys mult_p u curr_p n) = n" by (induct n arbitrary: curr_p, auto) (* Equation 12 *) lemma Poly_berlekamp_mat: assumes k: "k < degree u" shows "Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k) mod u" proof - let ?map ="(map (\cs. coeffs cs @ replicate (degree u - length (coeffs cs)) 0) (power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)))" have "row (berlekamp_mat u) k = row (mat_of_rows_list (degree u) ?map) k" by (simp add: berlekamp_mat_def Let_def) also have "... = vec_of_list (?map ! k)" proof- { fix i assume i: "i < degree u" let ?c= "power_polys (power_poly_f_mod u [:0, 1:] CARD('a)) u 1 (degree u) ! i" let ?coeffs_c="(coeffs ?c)" have "?c = 1*([:0, 1:] ^ CARD('a) mod u)^i mod u" proof (unfold power_poly_f_mod_def, rule power_polys_works[OF i]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = [:0, 1:] ^ (CARD('a) * i) mod u" by (simp add: power_mod power_mult) finally have c_rw: "?c = [:0, 1:] ^ (CARD('a) * i) mod u" . have "length ?coeffs_c \ degree u" proof - show ?thesis proof (cases "?c = 0") case True thus ?thesis by auto next case False have "length ?coeffs_c = degree (?c) + 1" by (rule length_coeffs[OF False]) also have "... = degree ([:0, 1:] ^ (CARD('a) * i) mod u) + 1" using c_rw by simp also have "... \ degree u" by (metis One_nat_def add.right_neutral add_Suc_right c_rw calculation coeffs_def degree_0 degree_mod_less discrete gr_implies_not0 k list.size(3) one_neq_zero) finally show ?thesis . qed qed then have "length ?coeffs_c + (degree u - length ?coeffs_c) = degree u" by auto } with k show ?thesis by (intro row_mat_of_rows_list, auto) qed finally have row_rw: "row (berlekamp_mat u) k = vec_of_list (?map ! k)" . have "Poly (list_of_vec (row (berlekamp_mat u) k)) = Poly (list_of_vec (vec_of_list (?map ! k)))" unfolding row_rw .. also have "... = Poly (?map ! k)" by simp also have "... = [:0,1:]^(CARD('a) * k) mod u" proof - let ?cs = "(power_polys (power_poly_f_mod u [:0, 1:] (nat (int CARD('a)))) u 1 (degree u)) ! k" let ?c = "coeffs ?cs @ replicate (degree u - length (coeffs ?cs)) 0" have map_k_c: "?map ! k = ?c" by (rule nth_map, simp add: k) have "(Poly (?map ! k)) = Poly (coeffs ?cs)" unfolding map_k_c Poly_append_replicate_0 .. also have "... = ?cs" by simp also have "... = power_polys ([:0, 1:] ^ CARD('a) mod u) u 1 (degree u) ! k" by (simp add: power_poly_f_mod_def) also have "... = 1* ([:0,1:]^(CARD('a)) mod u) ^ k mod u" proof (rule power_polys_works[OF k]) show "1 = 1 mod u" using k mod_poly_less by force qed also have "... = ([:0,1:]^(CARD('a)) mod u) ^ k mod u" by auto also have "... = [:0,1:]^(CARD('a) * k) mod u" by (simp add: power_mod power_mult) finally show ?thesis . qed finally show ?thesis . qed corollary Poly_berlekamp_cong_mat: assumes k: "k < degree u" shows "[Poly (list_of_vec (row (berlekamp_mat u) k)) = [:0,1:]^(CARD('a) * k)] (mod u)" using Poly_berlekamp_mat[OF k] unfolding cong_def by auto lemma mat_of_rows_list_dim[simp]: "mat_of_rows_list n vs \ carrier_mat (length vs) n" "dim_row (mat_of_rows_list n vs) = length vs" "dim_col (mat_of_rows_list n vs) = n" unfolding mat_of_rows_list_def by auto lemma berlekamp_mat_closed[simp]: "berlekamp_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_mat u) = degree u" "dim_col (berlekamp_mat u) = degree u" unfolding carrier_mat_def berlekamp_mat_def Let_def by auto lemma vec_of_list_coeffs_nth: assumes i: "i \ {..degree h}" and h_not0: "h \ 0" shows "vec_of_list (coeffs h) $ i = coeff h i" proof - have "vec_of_list (map (coeff h) [0..i. f i mod z) A" using f by (induct, auto simp add: poly_mod_add_left) lemma prime_not_dvd_fact: assumes kn: "k < n" and prime_n: "prime n" shows "\ n dvd fact k" using kn proof (induct k) case 0 thus ?case using prime_n unfolding prime_nat_iff by auto next case (Suc k) show ?case proof (rule ccontr, unfold not_not) assume "n dvd fact (Suc k)" also have "... = Suc k * \{1..k}" unfolding fact_Suc unfolding fact_prod by simp finally have "n dvd Suc k * \{1..k}" . hence "n dvd Suc k \ n dvd \{1..k}" using prime_dvd_mult_eq_nat[OF prime_n] by blast moreover have "\ n dvd Suc k" by (simp add: Suc.prems(1) nat_dvd_not_less) moreover hence "\ n dvd \{1..k}" using Suc.hyps Suc.prems using Suc_lessD fact_prod[of k] by (metis of_nat_id) ultimately show False by simp qed qed lemma dvd_choose_prime: assumes kn: "k < n" and k: "k \ 0" and n: "n \ 0" and prime_n: "prime n" shows "n dvd (n choose k)" proof - have "n dvd (fact n)" by (simp add: fact_num_eq_if n) moreover have "\ n dvd (fact k * fact (n-k))" proof (rule ccontr, simp) assume "n dvd fact k * fact (n - k)" hence "n dvd fact k \ n dvd fact (n - k)" using prime_dvd_mult_eq_nat[OF prime_n] by simp moreover have "\ n dvd (fact k)" by (rule prime_not_dvd_fact[OF kn prime_n]) moreover have "\ n dvd fact (n - k)" using prime_not_dvd_fact[OF _ prime_n] kn k by simp ultimately show False by simp qed moreover have "(fact n::nat) = fact k * fact (n-k) * (n choose k)" using binomial_fact_lemma kn by auto ultimately show ?thesis using prime_n by (auto simp add: prime_dvd_mult_iff) qed lemma add_power_poly_mod_ring: fixes x :: "'a mod_ring poly" shows "(x + y) ^ CARD('a) = x ^ CARD('a) + y ^ CARD('a)" proof - let ?A="{0..CARD('a)}" let ?f="\k. of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k)" have A_rw: "?A = insert CARD('a) (insert 0 (?A - {0} - {CARD('a)}))" by fastforce have sum0: "sum ?f (?A - {0} - {CARD('a)}) = 0" proof (rule sum.neutral, rule) fix xa assume xa: "xa \ {0..CARD('a)} - {0} - {CARD('a)}" have card_dvd_choose: "CARD('a) dvd (CARD('a) choose xa)" proof (rule dvd_choose_prime) show "xa < CARD('a)" using xa by simp show "xa \ 0" using xa by simp show "CARD('a) \ 0" by simp show "prime CARD('a)" by (rule prime_card) qed hence rw0: "of_int (CARD('a) choose xa) = (0 :: 'a mod_ring)" by transfer simp have "of_nat (CARD('a) choose xa) = [:of_int (CARD('a) choose xa) :: 'a mod_ring:]" by (simp add: of_nat_poly) also have "... = [:0:]" using rw0 by simp finally show "of_nat (CARD('a) choose xa) * x ^ xa * y ^ (CARD('a) - xa) = 0" by auto qed have "(x + y)^CARD('a) = (\k = 0..CARD('a). of_nat (CARD('a) choose k) * x ^ k * y ^ (CARD('a) - k))" unfolding binomial_ring by (rule sum.cong, auto) also have "... = sum ?f (insert CARD('a) (insert 0 (?A - {0} - {CARD('a)})))" using A_rw by simp also have "... = ?f 0 + ?f CARD('a) + sum ?f (?A - {0} - {CARD('a)})" by auto also have "... = x^CARD('a) + y^CARD('a)" unfolding sum0 by auto finally show ?thesis . qed lemma power_poly_sum_mod_ring: fixes f :: "'b \ 'a mod_ring poly" assumes f: "finite A" shows "(sum f A) ^ CARD('a) = sum (\i. (f i) ^ CARD('a)) A" using f by (induct, auto simp add: add_power_poly_mod_ring) lemma poly_power_card_as_sum_of_monoms: fixes h :: "'a mod_ring poly" shows "h ^ CARD('a) = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof - have "h ^ CARD('a) = (\i\degree h. monom (coeff h i) i) ^ CARD('a)" by (simp add: poly_as_sum_of_monoms) also have "... = (\i\degree h. (monom (coeff h i) i) ^ CARD('a))" by (simp add: power_poly_sum_mod_ring) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" show "monom (coeff h x) x ^ CARD('a) = monom (coeff h x) (CARD('a) * x)" by (unfold poly_eq_iff, auto simp add: monom_power) qed finally show ?thesis . qed lemma degree_Poly_berlekamp_le: assumes i: "i < degree u" shows "degree (Poly (list_of_vec (row (berlekamp_mat u) i))) < degree u" by (metis Poly_berlekamp_mat degree_0 degree_mod_less gr_implies_not0 i linorder_neqE_nat) (* Equation 12: alternative statement. *) lemma monom_card_pow_mod_sum_berlekamp: assumes i: "i < degree u" shows "monom 1 (CARD('a) * i) mod u = (\j 0" using i by simp hence set_rw: "{..degree u - 1} = {.. degree u - 1" by auto have "monom 1 (CARD('a) * i) mod u = [:0, 1:] ^ (CARD('a) * i) mod u" using x_as_monom x_pow_n by metis also have "... = ?p" unfolding Poly_berlekamp_mat[OF i] by simp also have "... = (\i\degree u - 1. monom (coeff ?p i) i)" using degree_le2 poly_as_sum_of_monoms' by fastforce also have "... = (\ij {.. v = (\i = 0.. v = (\i = 0.. v = col A j \ v" using j row_transpose by auto also have "... = (\i = 0..iiiiiiiiiii degree u" by (metis Suc_leI assms coeffs_0_eq_Nil degree_0 length_coeffs_degree list.size(3) not_le_imp_less order.asym) thus ?thesis by simp qed lemma vec_of_list_coeffs_nth': assumes i: "i \ {..degree h}" and h_not0: "h \ 0" assumes "degree h < degree u" shows "vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) $ i = coeff h i" using assms by (transfer', auto simp add: mk_vec_def coeffs_nth length_coeffs_degree nth_append) lemma vec_of_list_coeffs_replicate_nth_0: assumes i: "i \ {.. {..{..degree h}") case True thus ?thesis using assms vec_of_list_coeffs_nth' h_not0 by simp next case False have c0: "coeff h i = 0" using False le_degree by auto thus ?thesis using assms False h_not0 by (transfer', auto simp add: mk_vec_def length_coeffs_degree nth_append c0) qed qed (* Equation 13 *) lemma equation_13: fixes u h defines H: "H \ vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0)" assumes deg_le: "degree h < degree u" (*Mandatory from equation 8*) shows "[h^CARD('a) = h] (mod u) \ (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H" (is "?lhs = ?rhs") proof - have f: "finite {..degree u}" by auto have [simp]: "dim_vec H = degree u" unfolding H using dim_vec_of_list_h deg_le by simp let ?B = "(berlekamp_mat u)" let ?f = "\i. (transpose_mat ?B *\<^sub>v H) $ i" show ?thesis proof assume rhs: ?rhs have dimv_h_dimr_B: "dim_vec H = dim_row ?B" by (metis berlekamp_mat_closed(2) berlekamp_mat_closed(3) dim_mult_mat_vec index_transpose_mat(2) rhs) have degree_h_less_dim_H: "degree h < dim_vec H" by (auto simp add: deg_le) have set_rw: "{..degree u - 1} = {.. degree u - 1" using deg_le by simp hence "h = (\j\degree u - 1. monom (coeff h j) j)" using poly_as_sum_of_monoms' by fastforce also have "... = (\jj {..j H) j)" by (rule sum.cong, auto) also have "... = (\ji = 0.. {.. H) x = monom (\i = 0..ji = 0..i = 0..ji = 0..ji = 0..ji = 0.. {0..jji = 0..i = 0..i = 0..x. x mod u"], rule sum.cong, rule) fix x assume x: "x \ {0.. {..i\degree h. monom (coeff h i) (CARD('a) * i)) mod u" proof (rule arg_cong[of _ _ "\x. x mod u"]) let ?f="\i. monom (coeff h i) (CARD('a) * i)" have ss0: "(\i = degree h + 1 ..< dim_vec H. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0..< dim_vec H} = {0..degree h} \ {degree h + 1 ..< dim_vec H}" using degree_h_less_dim_H by auto have "(\i = 0..i = 0..degree h. ?f i) + (\i = degree h + 1 ..< dim_vec H. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i = 0..degree h. ?f i)" unfolding ss0 by auto finally show "(\i = 0..i\degree h. ?f i)" by (simp add: atLeast0AtMost) qed also have "... = h^CARD('a) mod u" using poly_power_card_as_sum_of_monoms by auto finally show ?lhs unfolding cong_def using deg_le by (simp add: mod_poly_less) next assume lhs: ?lhs have deg_le': "degree h \ degree u - 1" using deg_le by auto have set_rw: "{..ii \ degree u - 1. monom (coeff h i) i)" by simp also have "... = (\i\degree h. monom (coeff h i) i)" unfolding poly_as_sum_of_monoms using poly_as_sum_of_monoms' deg_le' by auto also have "... = (\i\degree h. monom (coeff h i) i) mod u" by (simp add: deg_le mod_poly_less poly_as_sum_of_monoms) also have "... = (\i\degree h. monom (coeff h i) (CARD('a)*i)) mod u" using lhs unfolding cong_def poly_as_sum_of_monoms poly_power_card_as_sum_of_monoms by auto also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i)) mod u" by (rule arg_cong[of _ _ "\x. x mod u"], rule sum.cong, simp_all add: mult_monom) also have "... = (\i\degree h. monom (coeff h i) 0 * monom 1 (CARD('a)*i) mod u)" by (simp add: poly_mod_sum) also have "... = (\i\degree h. monom (coeff h i) 0 * (monom 1 (CARD('a)*i) mod u))" proof (rule sum.cong, rule) fix x assume x: "x \ {..degree h}" have h_rw: "monom (coeff h x) 0 mod u = monom (coeff h x) 0" by (metis deg_le degree_pCons_eq_if gr_implies_not_zero linorder_neqE_nat mod_poly_less monom_0) have "monom (coeff h x) 0 * monom 1 (CARD('a) * x) = smult (coeff h x) (monom 1 (CARD('a) * x))" by (simp add: monom_0) also have "... mod u = Polynomial.smult (coeff h x) (monom 1 (CARD('a) * x) mod u)" using mod_smult_left by auto also have "... = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" by (simp add: monom_0) finally show "monom (coeff h x) 0 * monom 1 (CARD('a) * x) mod u = monom (coeff h x) 0 * (monom 1 (CARD('a) * x) mod u)" . qed also have "... = (\i\degree h. monom (coeff h i) 0 * (\j {..degree h}" have "(monom 1 (CARD('a) * x) mod u) = (\jjiji=degree h+1 ..< degree u. ?f i) = 0" by (rule sum.neutral, simp add: coeff_eq_0) have set_rw: "{0.. {degree h+1..i=0..i=0..degree h. ?f i) + (\i=degree h+1 ..< degree u. ?f i)" unfolding set_rw by (rule sum.union_disjoint, auto) also have "... = (\i=0..degree h. ?f i)" using ss0 by simp finally show ?thesis by (simp add: atLeast0AtMost atLeast0LessThan) qed also have "... = (\ijijjijiijii. i < degree u \ coeff h i = (\ji. i < degree u \ H $ i = (\jjjjv H = H" proof (rule eq_vecI) fix i show "dim_vec (transpose_mat ?B *\<^sub>v H) = dim_vec (H)" by auto assume i: "i < dim_vec (H)" have "(transpose_mat ?B *\<^sub>v H) $ i = row (transpose_mat ?B) i \ H" using i by simp also have "... = (\j = 0..jv H) $ i = H $ i" . qed qed qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma exists_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\s. fi dvd (h - [:s:])" proof - let ?p = "CARD('a)" have f_dvd_hqh: "f dvd (h^?p - h)" using h unfolding cong_def using mod_eq_dvd_iff_poly by blast also have hq_h_rw: "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by (rule poly_identity_mod_p) finally have f_dvd_hc: "f dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" by simp have "fi dvd f" using f_desc_square_free fi_P using dvd_prod_eqI finite_P by blast hence "fi dvd (h^?p - h)" using dvd_trans f_dvd_hqh by auto also have "... = prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" unfolding hq_h_rw by simp finally have fi_dvd_prod_hc: "fi dvd prod (\c. h - [:c:]) (UNIV::'a mod_ring set)" . have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (simp add: irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis using irreducible_dvd_prod[OF _ fi_dvd_prod_hc] irr_fi by auto qed corollary exists_unique_s_factor_dvd_h_s: fixes fi::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "f = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and fi_P: "fi \ P" and h: "h \ {v. [v^(CARD('a)) = v] (mod f)}" shows "\!s. fi dvd (h - [:s:])" proof - obtain c where fi_dvd: "fi dvd (h - [:c:])" using assms exists_s_factor_dvd_h_s by blast have irr_fi: "irreducible fi" using fi_P P by blast have fi_not_unit: "\ is_unit fi" by (simp add: irr_fi irreducible\<^sub>dD(1) poly_dvd_1) show ?thesis proof (rule ex1I[of _ c], auto simp add: fi_dvd) fix c2 assume fi_dvd_hc2: "fi dvd h - [:c2:]" have *: "fi dvd (h - [:c:]) * (h - [:c2:])" using fi_dvd by auto hence "fi dvd (h - [:c:]) \ fi dvd (h - [:c2:])" using irr_fi by auto thus "c2 = c" using coprime_h_c_poly coprime_not_unit_not_dvd fi_dvd fi_dvd_hc2 fi_not_unit by blast qed qed lemma exists_two_distint: "\a b::'a mod_ring. a \ b" by (rule exI[of _ 0], rule exI[of _ 1], auto) lemma coprime_cong_mult_factorization_poly: fixes f::"'b::{field} poly" and a b p :: "'c :: {field_gcd} poly" assumes finite_P: "finite P" and P: "P \ {q. irreducible q}" and p: "\p\P. [a=b] (mod p)" and coprime_P: "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" shows "[a = b] (mod (\a\P. a))" using finite_P P p coprime_P proof (induct P) case empty thus ?case by simp next case (insert p P) have ab_mod_pP: "[a=b] (mod (p * \P))" proof (rule coprime_cong_mult_poly) show "[a = b] (mod p)" using insert.prems by auto show "[a = b] (mod \P)" using insert.prems insert.hyps by auto from insert show "Rings.coprime p (\P)" by (auto intro: prod_coprime_right) qed thus ?case by (simp add: insert.hyps(1) insert.hyps(2)) qed end context assumes "SORT_CONSTRAINT('a::prime_card)" begin lemma W_eq_berlekamp_mat: fixes u::"'a mod_ring poly" shows "{v. [v^CARD('a) = v] (mod u) \ degree v < degree u} = {h. let H = vec_of_list ((coeffs h) @ replicate (degree u - length (coeffs h)) 0) in (transpose_mat (berlekamp_mat u)) *\<^sub>v H = H \ degree h < degree u}" using equation_13 by (auto simp add: Let_def) lemma transpose_minus_1: assumes "dim_row Q = dim_col Q" shows "transpose_mat (Q - (1\<^sub>m (dim_row Q))) = (transpose_mat Q - (1\<^sub>m (dim_row Q)))" using assms unfolding mat_eq_iff by auto lemma system_iff: fixes v::"'b::comm_ring_1 vec" assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" proof - have t1:"transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v) \ (transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v)" by (subst minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have t2:"(transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v) \ transpose_mat Q *\<^sub>v v - v = 0\<^sub>v (dim_vec v)" by (subst (asm) minus_mult_distrib_mat_vec, insert sq_Q[symmetric] v, auto) have "transpose_mat Q *\<^sub>v v - v = v - v \ transpose_mat Q *\<^sub>v v = v" proof - assume a1: "transpose_mat Q *\<^sub>v v - v = v - v" have f2: "transpose_mat Q *\<^sub>v v \ carrier_vec (dim_vec v)" by (metis dim_mult_mat_vec index_transpose_mat(2) sq_Q v carrier_vec_dim_vec) then have f3: "0\<^sub>v (dim_vec v) + transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v" by (meson left_zero_vec) have f4: "0\<^sub>v (dim_vec v) = transpose_mat Q *\<^sub>v v - v" using a1 by auto have f5: "- v \ carrier_vec (dim_vec v)" by simp then have f6: "- v + transpose_mat Q *\<^sub>v v = v - v" using f2 a1 using comm_add_vec minus_add_uminus_vec by fastforce have "v - v = - v + v" by auto then have "transpose_mat Q *\<^sub>v v = transpose_mat Q *\<^sub>v v - v + v" using f6 f4 f3 f2 by (metis (no_types, lifting) a1 assoc_add_vec comm_add_vec f5 carrier_vec_dim_vec) then show ?thesis using a1 by auto qed hence "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q *\<^sub>v v) - v = v - v)" by auto also have "... = ((transpose_mat Q *\<^sub>v v) - v = 0\<^sub>v (dim_vec v))" by auto also have "... = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using t1 t2 by auto finally show ?thesis. qed lemma system_if_mat_kernel: assumes sq_Q: "dim_row Q = dim_col Q" and v: "dim_row Q = dim_vec v" shows "(transpose_mat Q *\<^sub>v v = v) \ v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q))))" proof - have "(transpose_mat Q *\<^sub>v v = v) = ((transpose_mat Q - 1\<^sub>m (dim_row Q)) *\<^sub>v v = 0\<^sub>v (dim_vec v))" using assms system_iff by blast also have "... = (v \ mat_kernel (transpose_mat (Q - (1\<^sub>m (dim_row Q)))))" unfolding mat_kernel_def unfolding transpose_minus_1[OF sq_Q] unfolding v by auto finally show ?thesis . qed lemma degree_u_mod_irreducible\<^sub>d_factor_0: fixes v and u::"'a mod_ring poly" defines W: "W \ {v. [v ^ CARD('a) = v] (mod u)}" assumes v: "v \ W" and finite_U: "finite U" and u_U: "u = \U" and U_irr_monic: "U \ {q. irreducible q \ monic q}" and fi_U: "fi \ U" shows "degree (v mod fi) = 0" proof - have deg_fi: "degree fi > 0" using U_irr_monic using fi_U irreducible\<^sub>dD[of fi] by auto have "fi dvd u" using u_U U_irr_monic finite_U dvd_prod_eqI fi_U by blast moreover have "u dvd (v^CARD('a) - v)" using v unfolding W cong_def by (simp add: mod_eq_dvd_iff_poly) ultimately have "fi dvd (v^CARD('a) - v)" by (rule dvd_trans) then have fi_dvd_prod_vc: "fi dvd prod (\c. v - [:c:]) (UNIV::'a mod_ring set)" by (simp add: poly_identity_mod_p) have irr_fi: "irreducible fi" using fi_U U_irr_monic by blast have fi_not_unit: "\ is_unit fi" using irr_fi by (auto simp: poly_dvd_1) have fi_dvd_vc: "\c. fi dvd v - [:c:]" using irreducible_dvd_prod[OF _ fi_dvd_prod_vc] irr_fi by auto from this obtain a where "fi dvd v - [:a:]" by blast hence "v mod fi = [:a:] mod fi" using mod_eq_dvd_iff_poly by blast also have "... = [:a:]" by (simp add: deg_fi mod_poly_less) finally show ?thesis by simp qed (* Also polynomials over a field as a vector space in HOL-Algebra.*) definition "poly_abelian_monoid = \carrier = UNIV::'a mod_ring poly set, monoid.mult = ((*)), one = 1, zero = 0, add = (+), module.smult = smult\" interpretation vector_space_poly: vectorspace class_ring poly_abelian_monoid rewrites [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 0" and [simp]: "\\<^bsub>poly_abelian_monoid\<^esub> = 1" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (+)" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = (*)" and [simp]: "carrier poly_abelian_monoid = UNIV" and [simp]: "(\\<^bsub>poly_abelian_monoid\<^esub>) = smult" apply unfold_locales apply (auto simp: poly_abelian_monoid_def class_field_def smult_add_left smult_add_right Units_def) by (metis add.commute add.right_inverse) lemma subspace_Berlekamp: assumes f: "degree f \ 0" shows "subspace (class_ring :: 'a mod_ring ring) {v. [v^(CARD('a)) = v] (mod f) \ (degree v < degree f)} poly_abelian_monoid" proof - { fix v :: "'a mod_ring poly" and w :: "'a mod_ring poly" assume a1: "v ^ card (UNIV::'a set) mod f = v mod f" assume "w ^ card (UNIV::'a set) mod f = w mod f" then have "(v ^ card (UNIV::'a set) + w ^ card (UNIV::'a set)) mod f = (v + w) mod f" using a1 by (meson mod_add_cong) then have "(v + w) ^ card (UNIV::'a set) mod f = (v + w) mod f" by (simp add: add_power_poly_mod_ring) } note r=this thus ?thesis using f by (unfold_locales, auto simp: zero_power mod_smult_left smult_power cong_def degree_add_less) qed lemma berlekamp_resulting_mat_closed[simp]: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" proof - let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" let ?G="(gauss_jordan_single ?A)" have "?G \carrier_mat (degree u) (degree u)" by (rule gauss_jordan_single(2)[of ?A], auto) thus "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" "dim_row (berlekamp_resulting_mat u) = degree u" "dim_col (berlekamp_resulting_mat u) = degree u" unfolding berlekamp_resulting_mat_def Let_def by auto qed (*find_base vectors returns a basis:*) lemma berlekamp_resulting_mat_basis: "kernel.basis (degree u) (berlekamp_resulting_mat u) (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule find_base_vectors(3)) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp let ?A="(transpose_mat (mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j))))" have "row_echelon_form (gauss_jordan_single ?A)" by (rule gauss_jordan_single(3)[of ?A], auto) thus "row_echelon_form (berlekamp_resulting_mat u)" unfolding berlekamp_resulting_mat_def Let_def by auto qed lemma set_berlekamp_basis_eq: "(set (berlekamp_basis u)) = (Poly \ list_of_vec)` (set (find_base_vectors (berlekamp_resulting_mat u)))" by (auto simp add: image_def o_def berlekamp_basis_def) lemma berlekamp_resulting_mat_constant: assumes deg_u: "degree u = 0" shows "berlekamp_resulting_mat u = 1\<^sub>m 0" by (unfold mat_eq_iff, auto simp add: deg_u) context fixes u::"'a::prime_card mod_ring poly" begin lemma set_berlekamp_basis_constant: assumes deg_u: "degree u = 0" shows "set (berlekamp_basis u) = {}" proof - have one_carrier: "1\<^sub>m 0 \ carrier_mat 0 0" by auto have m: "mat_kernel (1\<^sub>m 0) = {(0\<^sub>v 0) :: 'a mod_ring vec}" unfolding mat_kernel_def by auto have r: "row_echelon_form (1\<^sub>m 0 :: 'a mod_ring mat)" unfolding row_echelon_form_def pivot_fun_def Let_def by auto have "set (find_base_vectors (1\<^sub>m 0)) \ {0\<^sub>v 0 :: 'a mod_ring vec}" using find_base_vectors(1)[OF r one_carrier] unfolding m . hence "set (find_base_vectors (1\<^sub>m 0) :: 'a mod_ring vec list) = {}" using find_base_vectors(2)[OF r one_carrier] using subset_singletonD by fastforce thus ?thesis unfolding set_berlekamp_basis_eq unfolding berlekamp_resulting_mat_constant[OF deg_u] by auto qed (*Maybe [simp]*) lemma row_echelon_form_berlekamp_resulting_mat: "row_echelon_form (berlekamp_resulting_mat u)" by (rule gauss_jordan_single(3), auto simp add: berlekamp_resulting_mat_def Let_def) lemma mat_kernel_berlekamp_resulting_mat_degree_0: assumes d: "degree u = 0" shows "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (auto simp add: mat_kernel_def mult_mat_vec_def d) lemma in_mat_kernel_berlekamp_resulting_mat: assumes x: "transpose_mat (berlekamp_mat u) *\<^sub>v x = x" and x_dim: "x \ carrier_vec (degree u)" shows "x \ mat_kernel (berlekamp_resulting_mat u)" proof - let ?QI="(mat(dim_row (berlekamp_mat u)) (dim_row (berlekamp_mat u)) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" have *: "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (degree u)) = transpose_mat ?QI" by auto have "(transpose_mat (berlekamp_mat u) - 1\<^sub>m (dim_row (berlekamp_mat u))) *\<^sub>v x = 0\<^sub>v (dim_vec x)" using system_iff[of "berlekamp_mat u" x] x_dim x by auto hence "transpose_mat ?QI *\<^sub>v x = 0\<^sub>v (degree u)" using x_dim * by auto hence "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" unfolding berlekamp_resulting_mat_def Let_def using gauss_jordan_single(1)[of "transpose_mat ?QI" "degree u" "degree u" _ x] x_dim by auto thus ?thesis by (auto simp add: mat_kernel_def x_dim) qed private abbreviation "V \ kernel.VK (degree u) (berlekamp_resulting_mat u)" private abbreviation "W \ vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u) \ (degree v < degree u)}" interpretation V: vectorspace class_ring V proof - interpret k: kernel "(degree u)" "(degree u)" "(berlekamp_resulting_mat u)" by (unfold_locales; auto) show "vectorspace class_ring V" by intro_locales qed lemma linear_Poly_list_of_vec: shows "(Poly \ list_of_vec) \ module_hom class_ring V (vector_space_poly.vs {v. [v^(CARD('a)) = v] (mod u)})" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed qed lemma linear_Poly_list_of_vec': assumes "degree u > 0" shows "(Poly \ list_of_vec) \ module_hom R V W" proof (auto simp add: LinearCombinations.module_hom_def Matrix.module_vec_def) fix m1 m2::" 'a mod_ring vec" assume m1: "m1 \ mat_kernel (berlekamp_resulting_mat u)" and m2: "m2 \ mat_kernel (berlekamp_resulting_mat u)" have m1_rw: "list_of_vec m1 = map (\n. m1 $ n) [0..n. m2 $ n) [0.. carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m1], auto) moreover have "m2 \ carrier_vec (degree u)" by (rule mat_kernelD(1)[OF _ m2], auto) ultimately have dim_eq: "dim_vec m1 = dim_vec m2" by auto show "Poly (list_of_vec (m1 + m2)) = Poly (list_of_vec m1) + Poly (list_of_vec m2)" unfolding poly_eq_iff m1_rw m2_rw plus_vec_def using dim_eq by (transfer', auto simp add: mk_vec_def nth_default_def) next fix r m assume m: "m \ mat_kernel (berlekamp_resulting_mat u)" show "Poly (list_of_vec (r \\<^sub>v m)) = smult r (Poly (list_of_vec m))" unfolding poly_eq_iff list_of_vec_rw_map[of m] smult_vec_def by (transfer', auto simp add: mk_vec_def nth_default_def) next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (cases "degree u = 0") case True have "mat_kernel (berlekamp_resulting_mat u) = {0\<^sub>v 0}" by (rule mat_kernel_berlekamp_resulting_mat_degree_0[OF True]) hence x_0: "x = 0\<^sub>v 0" using x by blast show ?thesis by (auto simp add: zero_power x_0 cong_def) next case False note deg_u = False show ?thesis proof - let ?QI="(mat (degree u) (degree u) (\(i, j). if i = j then berlekamp_mat u $$ (i, j) - 1 else berlekamp_mat u $$ (i, j)))" let ?H="vec_of_list (coeffs (Poly (list_of_vec x)) @ replicate (degree u - length (coeffs (Poly (list_of_vec x)))) 0)" have x_dim: "dim_vec x = degree u" using x unfolding mat_kernel_def by auto hence x_carrier[simp]: "x \ carrier_vec (degree u)" by (metis carrier_vec_dim_vec) have x_kernel: "berlekamp_resulting_mat u *\<^sub>v x = 0\<^sub>v (degree u)" using x unfolding mat_kernel_def by auto have t_QI_x_0: "(transpose_mat ?QI) *\<^sub>v x = 0\<^sub>v (degree u)" using gauss_jordan_single(1)[of "(transpose_mat ?QI)" "degree u" "degree u" "gauss_jordan_single (transpose_mat ?QI)" x] using x_kernel unfolding berlekamp_resulting_mat_def Let_def by auto have l: "(list_of_vec x) \ []" by (auto simp add: list_of_vec_rw_map vec_of_dim_0[symmetric] deg_u x_dim) have deg_le: "degree (Poly (list_of_vec x)) < degree u" using degree_Poly_list_of_vec using x_carrier deg_u by blast show "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" proof (unfold equation_13[OF deg_le]) have QR_rw: "?QI = berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))" by auto have "dim_row (berlekamp_mat u) = dim_vec ?H" by (auto, metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) moreover have "?H \ mat_kernel (transpose_mat (berlekamp_mat u - 1\<^sub>m (dim_row (berlekamp_mat u))))" proof - have Hx: "?H = x" proof (unfold vec_eq_iff, auto) let ?H'="vec_of_list (strip_while ((=) 0) (list_of_vec x) @ replicate (degree u - length (strip_while ((=) 0) (list_of_vec x))) 0)" show "length (strip_while ((=) 0) (list_of_vec x)) + (degree u - length (strip_while ((=) 0) (list_of_vec x))) = dim_vec x" by (metis le_add_diff_inverse length_list_of_vec length_strip_while_le x_dim) fix i assume i: "i < dim_vec x" have "?H $ i = coeff (Poly (list_of_vec x)) i" proof (rule vec_of_list_coeffs_replicate_nth[OF _ deg_le]) show "i \ {.. carrier_vec (degree u)" using deg_le dim_vec_of_list_h Hx by auto moreover have "transpose_mat (berlekamp_mat u - 1\<^sub>m (degree u)) *\<^sub>v ?H = 0\<^sub>v (degree u)" using t_QI_x_0 Hx QR_rw by auto ultimately show ?thesis by (auto simp add: mat_kernel_def) qed ultimately show "transpose_mat (berlekamp_mat u) *\<^sub>v ?H = ?H" using system_if_mat_kernel[of "berlekamp_mat u" ?H] by auto qed qed qed next fix x assume x: "x \ mat_kernel (berlekamp_resulting_mat u)" show "degree (Poly (list_of_vec x)) < degree u" by (rule degree_Poly_list_of_vec, insert assms x, auto simp: mat_kernel_def) qed lemma berlekamp_basis_eq_8: assumes v: "v \ set (berlekamp_basis u)" shows "[v ^ CARD('a) = v] (mod u)" proof - { fix x assume x: "x \ set (find_base_vectors (berlekamp_resulting_mat u))" have "set (find_base_vectors (berlekamp_resulting_mat u)) \ mat_kernel (berlekamp_resulting_mat u)" proof (rule find_base_vectors(1)) show "row_echelon_form (berlekamp_resulting_mat u)" by (rule row_echelon_form_berlekamp_resulting_mat) show "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by simp qed hence "x \ mat_kernel (berlekamp_resulting_mat u)" using x by auto hence "[Poly (list_of_vec x) ^ CARD('a) = Poly (list_of_vec x)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto } thus "[v ^ CARD('a) = v] (mod u)" using v unfolding set_berlekamp_basis_eq by auto qed lemma surj_Poly_list_of_vec: assumes deg_u: "degree u > 0" shows "(Poly \ list_of_vec)` (carrier V) = carrier W" proof (auto simp add: image_def) fix xa assume xa: "xa \ mat_kernel (berlekamp_resulting_mat u)" thus "[Poly (list_of_vec xa) ^ CARD('a) = Poly (list_of_vec xa)] (mod u)" using linear_Poly_list_of_vec unfolding LinearCombinations.module_hom_def Matrix.module_vec_def by auto show "degree (Poly (list_of_vec xa)) < degree u" proof (rule degree_Poly_list_of_vec[OF _ deg_u]) show "xa \ carrier_vec (degree u)" using xa unfolding mat_kernel_def by simp qed next fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" show "\xa \ mat_kernel (berlekamp_resulting_mat u). x = Poly (list_of_vec xa)" proof (rule bexI[of _ "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)"]) let ?X = "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)" show "x = Poly (list_of_vec (vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0)))" by auto have X: "?X \ carrier_vec (degree u)" unfolding carrier_vec_def by (auto, metis Suc_leI coeffs_0_eq_Nil deg_x degree_0 le_add_diff_inverse length_coeffs_degree linordered_semidom_class.add_diff_inverse list.size(3) order.asym) have t: "transpose_mat (berlekamp_mat u) *\<^sub>v ?X = ?X" using equation_13[OF deg_x] x by auto show "vec_of_list (coeffs x @ replicate (degree u - length (coeffs x)) 0) \ mat_kernel (berlekamp_resulting_mat u)" by (rule in_mat_kernel_berlekamp_resulting_mat[OF t X]) qed qed lemma card_set_berlekamp_basis: "card (set (berlekamp_basis u)) = length (berlekamp_basis u)" proof - have b: "berlekamp_resulting_mat u \ carrier_mat (degree u) (degree u)" by auto have "(set (berlekamp_basis u)) = (Poly \ list_of_vec) ` set (find_base_vectors (berlekamp_resulting_mat u))" unfolding set_berlekamp_basis_eq .. also have " card ... = card (set (find_base_vectors (berlekamp_resulting_mat u)))" proof (rule card_image, rule subset_inj_on[OF inj_Poly_list_of_vec]) show "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier_vec (degree u)" using find_base_vectors(1)[OF row_echelon_form_berlekamp_resulting_mat b] unfolding carrier_vec_def mat_kernel_def by auto qed also have "... = length (find_base_vectors (berlekamp_resulting_mat u))" by (rule length_find_base_vectors[symmetric, OF row_echelon_form_berlekamp_resulting_mat b]) finally show ?thesis unfolding berlekamp_basis_def by auto qed context assumes deg_u0[simp]: "degree u > 0" begin interpretation Berlekamp_subspace: vectorspace class_ring W by (rule vector_space_poly.subspace_is_vs[OF subspace_Berlekamp], simp) lemma linear_map_Poly_list_of_vec': "linear_map class_ring V W (Poly \ list_of_vec)" proof (auto simp add: linear_map_def) show "vectorspace class_ring V" by intro_locales show "vectorspace class_ring W" by (rule Berlekamp_subspace.vectorspace_axioms) show "mod_hom class_ring V W (Poly \ list_of_vec)" proof (rule mod_hom.intro, unfold mod_hom_axioms_def) show "module class_ring V" by intro_locales show "module class_ring W" using Berlekamp_subspace.vectorspace_axioms by intro_locales show "Poly \ list_of_vec \ module_hom class_ring V W" by (rule linear_Poly_list_of_vec'[OF deg_u0]) qed qed lemma berlekamp_basis_basis: "Berlekamp_subspace.basis (set (berlekamp_basis u))" proof (unfold set_berlekamp_basis_eq, rule linear_map.linear_inj_image_is_basis) show "linear_map class_ring V W (Poly \ list_of_vec)" by (rule linear_map_Poly_list_of_vec') show "inj_on (Poly \ list_of_vec) (carrier V)" proof (rule subset_inj_on[OF inj_Poly_list_of_vec]) show "carrier V \ carrier_vec (degree u)" by (auto simp add: mat_kernel_def) qed show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show b: "V.basis (set (find_base_vectors (berlekamp_resulting_mat u)))" by (rule berlekamp_resulting_mat_basis) show "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using b unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed qed lemma finsum_sum: fixes f::"'a mod_ring poly" assumes f: "finite B" and a_Pi: "a \ B \ carrier R" and V: "B \ carrier W" shows "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" using f a_Pi V proof (induct B) case empty thus ?case unfolding Berlekamp_subspace.module.M.finsum_empty by auto next case (insert x V) have hyp: "(\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) V" proof (rule insert.hyps) show "a \ V \ carrier R" using insert.prems unfolding class_field_def by auto show "V \ carrier W" using insert.prems by simp qed have "(\\<^bsub>W\<^esub>v\insert x V. a v \\<^bsub>W\<^esub> v) = (a x \\<^bsub>W\<^esub> x) \\<^bsub>W\<^esub> (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" proof (rule abelian_monoid.finsum_insert) show "abelian_monoid W" by (unfold_locales) show "finite V" by fact show "x \ V" by fact show "(\v. a v \\<^bsub>W\<^esub> v) \ V \ carrier W" proof (unfold Pi_def, rule, rule allI, rule impI) fix v assume v: "v\V" show "a v \\<^bsub>W\<^esub> v \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a v \ carrier class_ring" using insert.prems v unfolding Pi_def by (simp add: class_field_def) show "v \ carrier W" using v insert.prems by auto qed qed show "a x \\<^bsub>W\<^esub> x \ carrier W" proof (rule Berlekamp_subspace.smult_closed) show "a x \ carrier class_ring" using insert.prems unfolding Pi_def by (simp add: class_field_def) show "x \ carrier W" using insert.prems by auto qed qed also have "... = (a x \\<^bsub>W\<^esub> x) + (\\<^bsub>W\<^esub>v \ V. a v \\<^bsub>W\<^esub> v)" by auto also have "... = (a x \\<^bsub>W\<^esub> x) + sum (\v. smult (a v) v) V" unfolding hyp by simp also have "... = (smult (a x) x) + sum (\v. smult (a v) v) V" by simp also have "... = sum (\v. smult (a v) v) (insert x V)" by (simp add: insert.hyps(1) insert.hyps(2)) finally show ?case . qed lemma exists_vector_in_Berlekamp_subspace_dvd: fixes p_i::"'a mod_ring poly" assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v. v \ {h. [h^(CARD('a)) = h] (mod u) \ degree h < degree u} \ v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1)" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by auto have irr_pj: "irreducible p_j" using pj P by auto obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m f_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m f_desc_square_free f_not_0 by auto obtain i where mi: "m i = p_i" and i: "i < n" using P_m pi by blast obtain j where mj: "m j = p_j" and j: "j < n" using P_m pj by blast have ij: "i \ j" using mi mj pi_pj by auto obtain s_i and s_j::"'a mod_ring" where si_sj: "s_i \ s_j" using exists_two_distint by blast let ?u="\x. if x = i then [:s_i:] else if x = j then [:s_j:] else [:0:]" have degree_si: "degree [:s_i:] = 0" by auto have degree_sj: "degree [:s_j:] = 0" by auto have "\!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\a\{i. i < n}. [v = ?u a] (mod m a))" proof (rule chinese_remainder_unique_poly) show "\a\{i. i < n}. \b\{i. i < n}. a \ b \ Rings.coprime (m a) (m b)" proof (rule+) fix a b assume "a \ {i. i < n}" and "b \ {i. i < n}" and "a \ b" thus "Rings.coprime (m a) (m b)" using coprime_polynomial_factorization[OF P finite_P, simplified] P_m by (metis image_eqI inj_onD inj_on_m) qed show "\i\{i. i < n}. m i \ 0" by (rule not_zero) show "0 < degree (prod m {i. i < n})" unfolding degree_prod using deg_u0 by blast qed from this obtain v where v: "\a\{i. i < n}. [v = ?u a] (mod m a)" and degree_v: "degree v < (\i\{i. i < n}. degree (m i))" by blast show ?thesis proof (rule exI[of _ v], auto) show vp_v_mod: "[v ^ CARD('a) = v] (mod u)" proof (unfold f_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P]) show "P \ {q. irreducible q}" using P by blast show "\p\P. [v ^ CARD('a) = v] (mod p)" proof (rule ballI) fix p assume p: "p \ P" hence irr_p: "irreducible\<^sub>d p" using P by auto obtain k where mk: "m k = p" and k: "k < n" using P_m p by blast have "[v = ?u k] (mod p)" using v mk k by auto moreover have "?u k mod p = ?u k" apply (rule mod_poly_less) using irreducible\<^sub>dD(1)[OF irr_p] by auto ultimately obtain s where v_mod_p: "v mod p = [:s:]" unfolding cong_def by force hence deg_v_p: "degree (v mod p) = 0" by auto have "v mod p = [:s:]" by (rule v_mod_p) also have "... = [:s:]^CARD('a)" unfolding poly_const_pow by auto also have "... = (v mod p) ^ CARD('a)" using v_mod_p by auto also have "... = (v mod p) ^ CARD('a) mod p" using calculation by auto also have "... = v^CARD('a) mod p" using power_mod by blast finally show "[v ^ CARD('a) = v] (mod p)" unfolding cong_def .. qed show "\p1 p2. p1 \ P \ p2 \ P \ p1 \ p2 \ coprime p1 p2" using P coprime_polynomial_factorization finite_P by auto qed have "[v = ?u i] (mod m i)" using v i by auto hence v_pi_si_mod: "v mod p_i = [:s_i:] mod p_i" unfolding cong_def mi by auto also have "... = [:s_i:]" apply (rule mod_poly_less) using irr_pi by auto finally have v_pi_si: "v mod p_i = [:s_i:]" . have "[v = ?u j] (mod m j)" using v j by auto hence v_pj_sj_mod: "v mod p_j = [:s_j:] mod p_j" unfolding cong_def mj using ij by auto also have "... = [:s_j:]" apply (rule mod_poly_less) using irr_pj by auto finally have v_pj_sj: "v mod p_j = [:s_j:]" . show "v mod p_i = v mod p_j \ False" using si_sj v_pi_si v_pj_sj by auto show "degree (v mod p_i) = 0" unfolding v_pi_si by simp show "degree (v mod p_j) = 0" unfolding v_pj_sj by simp show "\s. gcd w (v - [:s:]) \ w \ gcd w (v - [:s:]) \ 1" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" using v_pi_si_mod mod_eq_dvd_iff_poly by blast have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" using v_pj_sj_mod mod_eq_dvd_iff_poly by blast have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using vp_v_mod cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" proof (rule ccontr, simp) assume gcd_w: "gcd w (v - [:s_i:]) = w" show False apply (rule \v mod p_i = v mod p_j \ False\) by (metis irreducibleE \degree (v mod p_i) = 0\ gcd_greatest_iff gcd_w irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_si) qed show "gcd w (v - [:s_i:]) \ 1" by (metis irreducibleE gcd_greatest_iff irr_pi pi_dvd_v_si pi_dvd_w) qed show "degree v < degree u" proof - have "(\i | i < n. degree (m i)) = degree (prod m {i. i < n})" by (rule degree_prod_eq_sum_degree[symmetric, OF not_zero]) thus ?thesis using degree_v unfolding degree_prod by auto qed qed qed lemma exists_vector_in_Berlekamp_basis_dvd_aux: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j" proof (rule ccontr, auto) have V_in_carrier: "B \ carrier W" using basis_V unfolding Berlekamp_subspace.basis_def by auto assume all_eq: "\v\B. v mod p_i = v mod p_j" obtain x where x: "x \ {h. [h ^ CARD('a) = h] (mod u) \ degree h < degree u}" and x_pi_pj: "x mod p_i \ x mod p_j" and "degree (x mod p_i) = 0" and "degree (x mod p_j) = 0" "(\s. gcd w (x - [:s:]) \ w \ gcd w (x - [:s:]) \ 1)" using exists_vector_in_Berlekamp_subspace_dvd[OF _ _ _ pi pj _ _ _ _ w_dvd_f monic_w pi_dvd_w] assms by meson have x_in: "x \ carrier W" using x by auto hence "(\!a. a \ B \\<^sub>E carrier class_ring \ Berlekamp_subspace.lincomb a B = x)" using Berlekamp_subspace.basis_criterion[OF finite_V V_in_carrier] using basis_V by (simp add: class_field_def) from this obtain a where a_Pi: "a \ B \\<^sub>E carrier class_ring" and lincomb_x: "Berlekamp_subspace.lincomb a B = x" by blast have fs_ss: "(\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) = sum (\v. smult (a v) v) B" proof (rule finsum_sum) show "finite B" by fact show "a \ B \ carrier class_ring" using a_Pi by auto show "B \ carrier W" by (rule V_in_carrier) qed have "x mod p_i = Berlekamp_subspace.lincomb a B mod p_i" using lincomb_x by simp also have " ... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_i" unfolding Berlekamp_subspace.lincomb_def .. also have "... = (sum (\v. smult (a v) v) B) mod p_i" unfolding fs_ss .. also have "... = sum (\v. smult (a v) v mod p_i) B" using finite_V poly_mod_sum by blast also have "... = sum (\v. smult (a v) (v mod p_i)) B" by (meson mod_smult_left) also have "... = sum (\v. smult (a v) (v mod p_j)) B" using all_eq by auto also have "... = sum (\v. smult (a v) v mod p_j) B" by (metis mod_smult_left) also have "... = (sum (\v. smult (a v) v) B) mod p_j" by (metis (mono_tags, lifting) finite_V poly_mod_sum sum.cong) also have "... = (\\<^bsub>W\<^esub>v\B. a v \\<^bsub>W\<^esub> v) mod p_j" unfolding fs_ss .. also have "... = Berlekamp_subspace.lincomb a B mod p_j" unfolding Berlekamp_subspace.lincomb_def .. also have "... = x mod p_j" using lincomb_x by simp finally have "x mod p_i = x mod p_j" . thus False using x_pi_pj by contradiction qed lemma exists_vector_in_Berlekamp_basis_dvd: assumes basis_V: "Berlekamp_subspace.basis B" and finite_V: "finite B" (*This should be removed, since the Berlekamp subspace is a finite set*) assumes finite_P: "finite P" and f_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" and pi: "p_i \ P" and pj: "p_j \ P" and pi_pj: "p_i \ p_j" and monic_f: "monic u" and sf_f: "square_free u" and not_irr_w: "\ irreducible w" and w_dvd_f: "w dvd u" and monic_w: "monic w" and pi_dvd_w: "p_i dvd w" and pj_dvd_w: "p_j dvd w" shows "\v \ B. v mod p_i \ v mod p_j \ degree (v mod p_i) = 0 \ degree (v mod p_j) = 0 \ \This implies that the algorithm decreases the degree of the reducible polynomials in each step:\ \ (\s. gcd w (v - [:s:]) \ w \ \ coprime w (v - [:s:]))" proof - have f_not_0: "u \ 0" using monic_f by auto have irr_pi: "irreducible p_i" using pi P by fast have irr_pj: "irreducible p_j" using pj P by fast obtain v where vV: "v \ B" and v_pi_pj: "v mod p_i \ v mod p_j" using assms exists_vector_in_Berlekamp_basis_dvd_aux by blast have v: "v \ {v. [v ^ CARD('a) = v] (mod u)}" using basis_V vV unfolding Berlekamp_subspace.basis_def by auto have deg_v_pi: "degree (v mod p_i) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pi]) from this obtain s_i where v_pi_si: "v mod p_i = [:s_i:]" using degree_eq_zeroE by blast have deg_v_pj: "degree (v mod p_j) = 0" by (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF v finite_P f_desc_square_free P pj]) from this obtain s_j where v_pj_sj: "v mod p_j = [:s_j:]" using degree_eq_zeroE by blast have si_sj: "s_i \ s_j" using v_pi_si v_pj_sj v_pi_pj by auto have "(\s. gcd w (v - [:s:]) \ w \ \ Rings.coprime w (v - [:s:]))" proof (rule exI[of _ s_i], rule conjI) have pi_dvd_v_si: "p_i dvd v - [:s_i:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pi_si) have pj_dvd_v_sj: "p_j dvd v - [:s_j:]" by (metis mod_eq_dvd_iff_poly mod_mod_trivial v_pj_sj) have w_eq: "w = prod (\c. gcd w (v - [:c:])) (UNIV::'a mod_ring set)" proof (rule Berlekamp_gcd_step) show "[v ^ CARD('a) = v] (mod w)" using v cong_dvd_modulus_poly w_dvd_f by blast show "square_free w" by (rule square_free_factor[OF w_dvd_f sf_f]) show "monic w" by (rule monic_w) qed show "gcd w (v - [:s_i:]) \ w" by (metis irreducibleE deg_v_pi gcd_greatest_iff irr_pj is_unit_field_poly mod_eq_dvd_iff_poly mod_poly_less neq0_conv pj_dvd_w v_pi_pj v_pi_si) show "\ Rings.coprime w (v - [:s_i:])" using irr_pi pi_dvd_v_si pi_dvd_w by (simp add: irreducible\<^sub>dD(1) not_coprimeI) qed thus ?thesis using v_pi_pj vV deg_v_pi deg_v_pj by auto qed lemma exists_bijective_linear_map_W_vec: assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" shows "\f. linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f \ bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" proof - let ?B="carrier_vec (card P)::'a mod_ring vec set" have u_not_0: "u \ 0" using deg_u0 degree_0 by force obtain m and n::nat where P_m: "P = m ` {i. i < n}" and inj_on_m: "inj_on m {i. i < n}" using finite_imp_nat_seg_image_inj_on[OF finite_P] by blast hence n: "n = card P" by (simp add: card_image) have degree_prod: "degree (prod m {i. i < n}) = degree u" by (metis P_m u_desc_square_free inj_on_m prod.reindex_cong) have not_zero: "\i\{i. i < n}. m i \ 0" using P_m u_desc_square_free u_not_0 by auto have deg_sum_eq: "(\i\{i. i < n}. degree (m i)) = degree u" by (metis degree_prod degree_prod_eq_sum_degree not_zero) have coprime_mi_mj:"\i\{i. i < n}. \j\{i. i < n}. i \ j \ coprime (m i) (m j)" proof (rule+) fix i j assume i: "i \ {i. i < n}" and j: "j \ {i. i < n}" and ij: "i \ j" show "coprime (m i) (m j)" proof (rule coprime_polynomial_factorization[OF P finite_P]) show "m i \ P" using i P_m by auto show "m j \ P" using j P_m by auto show "m i \ m j" using inj_on_m i ij j unfolding inj_on_def by blast qed qed let ?f = "\v. vec n (\i. coeff (v mod (m i)) 0)" interpret vec_VS: vectorspace class_ring "(module_vec TYPE('a mod_ring) n)" by (rule VS_Connect.vec_vs) interpret linear_map class_ring W "(module_vec TYPE('a mod_ring) n)" ?f by (intro_locales, unfold mod_hom_axioms_def LinearCombinations.module_hom_def, auto simp add: vec_eq_iff module_vec_def mod_smult_left poly_mod_add_left) have "linear_map class_ring W (module_vec TYPE('a mod_ring) n) ?f" by (intro_locales) moreover have inj_f: "inj_on ?f (carrier W)" proof (rule Ke0_imp_inj, auto simp add: mod_hom.ker_def) show "[0 ^ CARD('a) = 0] (mod u)" by (simp add: cong_def zero_power) show "vec n (\i. 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" by (auto simp add: module_vec_def) fix x assume x: "[x ^ CARD('a) = x] (mod u)" and deg_x: "degree x < degree u" and v: "vec n (\i. coeff (x mod m i) 0) = \\<^bsub>module_vec TYPE('a mod_ring) n\<^esub>" have cong_0: "\i\{i. i < n}. [x = (\i. 0) i] (mod m i)" proof (rule, unfold cong_def) fix i assume i: "i \ {i. i < n}" have deg_x_mod_mi: "degree (x mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "x \ {v. [v ^ CARD('a) = v] (mod u)}" using x by auto show "m i \ P" using P_m i by auto qed thus "x mod m i = 0 mod m i" using v unfolding module_vec_def by (auto, metis i leading_coeff_neq_0 mem_Collect_eq index_vec index_zero_vec(1)) qed moreover have deg_x2: "degree x < (\i\{i. i < n}. degree (m i))" using deg_sum_eq deg_x by simp moreover have "\i\{i. i < n}. [0 = (\i. 0) i] (mod m i)" by (auto simp add: cong_def) moreover have "degree 0 < (\i\{i. i < n}. degree (m i))" using degree_prod deg_sum_eq deg_u0 by force moreover have "\!x. degree x < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [x = (\i. 0) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) ultimately show "x = 0" by blast qed moreover have "?f ` (carrier W) = ?B" proof (auto simp add: image_def) fix xa show "n = card P" by (auto simp add: n) next fix x::"'a mod_ring vec" assume x: "x \ carrier_vec (card P)" have " \!v. degree v < (\i\{i. i < n}. degree (m i)) \ (\i\{i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" proof (rule chinese_remainder_unique_poly[OF not_zero]) show "0 < degree (prod m {i. i < n})" using deg_u0 degree_prod by linarith qed (insert coprime_mi_mj, auto) from this obtain v where deg_v: "degree v < (\i\{i. i < n}. degree (m i))" and v_x_cong: "(\i \ {i. i < n}. [v = (\i. [:x $ i:]) i] (mod m i))" by auto show "\xa. [xa ^ CARD('a) = xa] (mod u) \ degree xa < degree u \ x = vec n (\i. coeff (xa mod m i) 0)" proof (rule exI[of _ v], auto) show v: "[v ^ CARD('a) = v] (mod u)" proof (unfold u_desc_square_free, rule coprime_cong_mult_factorization_poly[OF finite_P], auto) fix y assume y: "y \ P" thus "irreducible y" using P by blast obtain i where i: "i \ {i. i < n}" and mi: "y = m i" using P_m y by blast have "irreducible (m i)" using i P_m P by auto moreover have "[v = [:x $ i:]] (mod m i)" using v_x_cong i by auto ultimately have v_mi_eq_xi: "v mod m i = [:x $ i:]" by (auto simp: cong_def intro!: mod_poly_less) have xi_pow_xi: "[:x $ i:]^CARD('a) = [:x $ i:]" by (simp add: poly_const_pow) hence "(v mod m i)^CARD('a) = v mod m i" using v_mi_eq_xi by auto hence "(v mod m i)^CARD('a) = (v^CARD('a) mod m i)" by (metis mod_mod_trivial power_mod) thus "[v ^ CARD('a) = v] (mod y)" unfolding mi cong_def v_mi_eq_xi xi_pow_xi by simp next fix p1 p2 assume "p1 \ P" and "p2 \ P" and "p1 \ p2" then show "Rings.coprime p1 p2" using coprime_polynomial_factorization[OF P finite_P] by auto qed show "degree v < degree u" using deg_v deg_sum_eq degree_prod by presburger show "x = vec n (\i. coeff (v mod m i) 0)" proof (unfold vec_eq_iff, rule conjI) show "dim_vec x = dim_vec (vec n (\i. coeff (v mod m i) 0))" using x n by simp show "\ii. coeff (v mod m i) 0)). x $ i = vec n (\i. coeff (v mod m i) 0) $ i" proof (auto) fix i assume i: "i < n" have deg_mi: "irreducible (m i)" using i P_m P by auto have deg_v_mi: "degree (v mod m i) = 0" proof (rule degree_u_mod_irreducible\<^sub>d_factor_0[OF _ finite_P u_desc_square_free P]) show "v \ {v. [v ^ CARD('a) = v] (mod u)}" using v by fast show "m i \ P" using P_m i by auto qed have "v mod m i = [:x $ i:] mod m i" using v_x_cong i unfolding cong_def by auto also have "... = [:x $ i:]" using deg_mi by (auto intro!: mod_poly_less) finally show "x $ i = coeff (v mod m i) 0" by simp qed qed qed qed ultimately show ?thesis unfolding bij_betw_def n by auto qed lemma fin_dim_kernel_berlekamp: "V.fin_dim" proof - have "finite (set (find_base_vectors (berlekamp_resulting_mat u)))" by auto moreover have "set (find_base_vectors (berlekamp_resulting_mat u)) \ carrier V" and "V.gen_set (set (find_base_vectors (berlekamp_resulting_mat u)))" using berlekamp_resulting_mat_basis[of u] unfolding V.basis_def by auto ultimately show ?thesis unfolding V.fin_dim_def by auto qed lemma Berlekamp_subspace_fin_dim: "Berlekamp_subspace.fin_dim" proof (rule linear_map.surj_fin_dim[OF linear_map_Poly_list_of_vec']) show "(Poly \ list_of_vec) ` carrier V = carrier W" using surj_Poly_list_of_vec[OF deg_u0] by auto show "V.fin_dim" by (rule fin_dim_kernel_berlekamp) qed context fixes P assumes finite_P: "finite P" and u_desc_square_free: "u = (\a\P. a)" and P: "P \ {q. irreducible q \ monic q}" begin interpretation RV: vec_space "TYPE('a mod_ring)" "card P" . lemma Berlekamp_subspace_eq_dim_vec: "Berlekamp_subspace.dim = RV.dim" proof - obtain f where lm_f: "linear_map class_ring W (module_vec TYPE('a mod_ring) (card P)) f" and bij_f: "bij_betw f (carrier W) (carrier_vec (card P)::'a mod_ring vec set)" using exists_bijective_linear_map_W_vec[OF finite_P u_desc_square_free P] by blast show ?thesis proof (rule linear_map.dim_eq[OF lm_f Berlekamp_subspace_fin_dim]) show "inj_on f (carrier W)" by (rule bij_betw_imp_inj_on[OF bij_f]) show " f ` carrier W = carrier RV.V" using bij_f unfolding bij_betw_def by auto qed qed lemma Berlekamp_subspace_dim: "Berlekamp_subspace.dim = card P" using Berlekamp_subspace_eq_dim_vec RV.dim_is_n by simp corollary card_berlekamp_basis_number_factors: "card (set (berlekamp_basis u)) = card P" unfolding Berlekamp_subspace_dim[symmetric] by (rule Berlekamp_subspace.dim_basis[symmetric], auto simp add: berlekamp_basis_basis) lemma length_berlekamp_basis_numbers_factors: "length (berlekamp_basis u) = card P" using card_set_berlekamp_basis card_berlekamp_basis_number_factors by auto end end end end context assumes "SORT_CONSTRAINT('a :: prime_card)" begin context fixes f :: "'a mod_ring poly" and n assumes sf: "square_free f" and n: "n = length (berlekamp_basis f)" and monic_f: "monic f" begin lemma berlekamp_basis_length_factorization: assumes f: "f = prod_list us" and d: "\ u. u \ set us \ degree u > 0" shows "length us \ n" proof (cases "degree f = 0") case True have "us = []" proof (rule ccontr) assume "us \ []" from this obtain u where u: "u \ set us" by fastforce hence deg_u: "degree u > 0" using d by auto have "degree f = degree (prod_list us)" unfolding f .. also have "... = sum_list (map degree us)" proof (rule degree_prod_list_eq) fix p assume p: "p \ set us" show "p \ 0" using d[OF p] degree_0 by auto qed also have " ... \ degree u" by (simp add: member_le_sum_list u) finally have "degree f > 0" using deg_u by auto thus False using True by auto qed thus ?thesis by simp next case False hence f_not_0: "f \ 0" using degree_0 by fastforce obtain P where fin_P: "finite P" and f_P: "f = \P" and P: "P \ {p. irreducible p \ monic p}" using monic_square_free_irreducible_factorization[OF monic_f sf] by auto have n_card_P: "n = card P" using P False f_P fin_P length_berlekamp_basis_numbers_factors n by blast have distinct_us: "distinct us" using d f sf square_free_prod_list_distinct by blast let ?us'="(map normalize us)" have distinct_us': "distinct ?us'" proof (auto simp add: distinct_map) show "distinct us" by (rule distinct_us) show "inj_on normalize (set us)" proof (auto simp add: inj_on_def, rule ccontr) fix x y assume x: "x \ set us" and y: "y \ set us" and n: "normalize x = normalize y" and x_not_y: "x \ y" from normalize_eq_imp_smult[OF n] obtain c where c0: "c \ 0" and y_smult: "y = smult c x" by blast have sf_xy: "square_free (x*y)" proof (rule square_free_factor[OF _ sf]) have "x*y = prod_list [x,y]" by simp also have "... dvd prod_list us" by (rule prod_list_dvd_prod_list_subset, auto simp add: x y x_not_y distinct_us) also have "... = f" unfolding f .. finally show "x * y dvd f" . qed have "x * y = smult c (x*x)" using y_smult mult_smult_right by auto hence sf_smult: "square_free (smult c (x*x))" using sf_xy by auto have "x*x dvd (smult c (x*x))" by (simp add: dvd_smult) hence "\ square_free (smult c (x*x))" by (metis d square_free_def x) thus "False" using sf_smult by contradiction qed qed have length_us_us': "length us = length ?us'" by simp have f_us': "f = prod_list ?us'" proof - have "f = normalize f" using monic_f f_not_0 by (simp add: normalize_monic) also have "... = prod_list ?us'" by (unfold f, rule prod_list_normalize[of us]) finally show ?thesis . qed have "\Q. prod_list Q = prod_list ?us' \ length ?us' \ length Q \ (\u. u \ set Q \ irreducible u \ monic u)" proof (rule exists_factorization_prod_list) show "degree (prod_list ?us') > 0" using False f_us' by auto show "square_free (prod_list ?us')" using f_us' sf by auto fix u assume u: "u \ set ?us'" have u_not0: "u \ 0" using d u degree_0 by fastforce have "degree u > 0" using d u by auto moreover have "monic u" using u monic_normalize[OF u_not0] by auto ultimately show "degree u > 0 \ monic u" by simp qed from this obtain Q where Q_us': "prod_list Q = prod_list ?us'" and length_us'_Q: "length ?us' \ length Q" and Q: "(\u. u \ set Q \ irreducible u \ monic u)" by blast have distinct_Q: "distinct Q" proof (rule square_free_prod_list_distinct) show "square_free (prod_list Q)" using Q_us' f_us' sf by auto show "\u. u \ set Q \ degree u > 0" using Q irreducible_degree_field by auto qed have set_Q_P: "set Q = P" proof (rule monic_factorization_uniqueness) show "\(set Q) = \P" using Q_us' by (metis distinct_Q f_P f_us' list.map_ident prod.distinct_set_conv_list) qed (insert P Q fin_P, auto) hence "length Q = card P" using distinct_Q distinct_card by fastforce have "length us = length ?us'" by (rule length_us_us') also have "... \ length Q" using length_us'_Q by auto also have "... = card (set Q)" using distinct_card[OF distinct_Q] by simp also have "... = card P" using set_Q_P by simp finally show ?thesis using n_card_P by simp qed lemma berlekamp_basis_irreducible: assumes f: "f = prod_list us" and n_us: "length us = n" and us: "\ u. u \ set us \ degree u > 0" and u: "u \ set us" shows "irreducible u" proof (fold irreducible_connect_field, intro irreducible\<^sub>dI[OF us[OF u]]) fix q r :: "'a mod_ring poly" assume dq: "degree q > 0" and qu: "degree q < degree u" and dr: "degree r > 0" and uqr: "u = q * r" with us[OF u] have q: "q \ 0" and r: "r \ 0" by auto from split_list[OF u] obtain xs ys where id: "us = xs @ u # ys" by auto let ?us = "xs @ q # r # ys" have f: "f = prod_list ?us" unfolding f id uqr by simp { fix x assume "x \ set ?us" with us[unfolded id] dr dq have "degree x > 0" by auto } from berlekamp_basis_length_factorization[OF f this] have "length ?us \ n" by simp also have "\ = length us" unfolding n_us by simp also have "\ < length ?us" unfolding id by simp finally show False by simp qed end lemma not_irreducible_factor_yields_prime_factors: assumes uf: "u dvd (f :: 'b :: {field_gcd} poly)" and fin: "finite P" and fP: "f = \P" and P: "P \ {q. irreducible q \ monic q}" and u: "degree u > 0" "\ irreducible u" shows "\ pi pj. pi \ P \ pj \ P \ pi \ pj \ pi dvd u \ pj dvd u" proof - from finite_distinct_list[OF fin] obtain ps where Pps: "P = set ps" and dist: "distinct ps" by auto have fP: "f = prod_list ps" unfolding fP Pps using dist by (simp add: prod.distinct_set_conv_list) note P = P[unfolded Pps] have "set ps \ P" unfolding Pps by auto from uf[unfolded fP] P dist this show ?thesis proof (induct ps) case Nil with u show ?case using divides_degree[of u 1] by auto next case (Cons p ps) from Cons(3) have ps: "set ps \ {q. irreducible q \ monic q}" by auto from Cons(2) have dvd: "u dvd p * prod_list ps" by simp obtain k where gcd: "u = gcd p u * k" by (meson dvd_def gcd_dvd2) from Cons(3) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of u] *(2) have "gcd p u = 1 \ gcd p u = p" by auto thus ?case proof assume "gcd p u = 1" then have "Rings.coprime p u" by (rule gcd_eq_1_imp_coprime) with dvd have "u dvd prod_list ps" using coprime_dvd_mult_right_iff coprime_imp_coprime by blast from Cons(1)[OF this ps] Cons(4-5) show ?thesis by auto next assume "gcd p u = p" with gcd have upk: "u = p * k" by auto hence p: "p dvd u" by auto from dvd[unfolded upk] *(3) have kps: "k dvd prod_list ps" by auto from dvd u * have dk: "degree k > 0" by (metis gr0I irreducible_mult_unit_right is_unit_iff_degree mult_zero_right upk) from ps kps have "\ q \ set ps. q dvd k" proof (induct ps) case Nil with dk show ?case using divides_degree[of k 1] by auto next case (Cons p ps) from Cons(3) have dvd: "k dvd p * prod_list ps" by simp obtain l where gcd: "k = gcd p k * l" by (meson dvd_def gcd_dvd2) from Cons(2) have *: "monic p" "irreducible p" "p \ 0" by auto from monic_irreducible_gcd[OF *(1), of k] *(2) have "gcd p k = 1 \ gcd p k = p" by auto thus ?case proof assume "gcd p k = 1" with dvd have "k dvd prod_list ps" by (metis dvd_triv_left gcd_greatest_mult mult.left_neutral) from Cons(1)[OF _ this] Cons(2) show ?thesis by auto next assume "gcd p k = p" with gcd have upk: "k = p * l" by auto hence p: "p dvd k" by auto thus ?thesis by auto qed qed then obtain q where q: "q \ set ps" and dvd: "q dvd k" by auto from dvd upk have qu: "q dvd u" by auto from Cons(4) q have "p \ q" by auto thus ?thesis using q p qu Cons(5) by auto qed qed qed lemma berlekamp_factorization_main: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and vs: "vs = vs1 @ vs2" and vsf: "vs = berlekamp_basis f" and n_bb: "n = length (berlekamp_basis f)" and n: "n = length us1 + n2" and us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" and us1: "\ u. u \ set us1 \ monic u \ irreducible u" and divs: "\ d. d \ set divs \ monic d \ degree d > 0" and vs1: "\ u v i. v \ set vs1 \ u \ set us1 \ set divs \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1,u}" and f: "f = prod_list (us1 @ divs)" and deg_f: "degree f > 0" and d: "\ g. g dvd f \ degree g = d \ irreducible g" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - have mon_f: "monic f" unfolding f by (rule monic_prod_list, insert divs us1, auto) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \ P" "P \ {q. irreducible q \ monic q}" by auto hence f0: "f \ 0" by auto show ?thesis using vs n us divs f us1 vs1 proof (induct vs2 arbitrary: divs n2 us1 vs1) case (Cons v vs2) show ?case proof (cases "v = 1") case False from Cons(2) vsf have v: "v \ set (berlekamp_basis f)" by auto from berlekamp_basis_eq_8[OF this] have vf: "[v ^ CARD('a) = v] (mod f)" . let ?gcd = "\ u i. gcd u (v - [:of_int i:])" let ?gcdn = "\ u i. gcd u (v - [:of_nat i:])" let ?map = "\ u. (map (\ i. ?gcd u i) [0 ..< CARD('a)])" define udivs where "udivs \ \ u. filter (\ w. w \ 1) (?map u)" { obtain xs where xs: "[0.. u. [w. i \ [0 ..< CARD('a)], w \ [?gcd u i], w \ 1])" unfolding udivs_def xs by (intro ext, auto simp: o_def, induct xs, auto) } note udivs_def' = this define facts where "facts \ [ w . u \ divs, w \ udivs u]" { fix u assume u: "u \ set divs" then obtain bef aft where divs: "divs = bef @ u # aft" by (meson split_list) from Cons(5)[OF u] have mon_u: "monic u" by simp have uf: "u dvd f" unfolding Cons(6) divs by auto from vf uf have vu: "[v ^ CARD('a) = v] (mod u)" by (rule cong_dvd_modulus_poly) from square_free_factor[OF uf sf_f] have sf_u: "square_free u" . let ?g = "?gcd u" from mon_u have u0: "u \ 0" by auto have "u = (\c\UNIV. gcd u (v - [:c:]))" using Berlekamp_gcd_step[OF vu mon_u sf_u] . also have "\ = (\i \ {0..< int CARD('a)}. ?g i)" by (rule sym, rule prod.reindex_cong[OF to_int_mod_ring_hom.inj_f range_to_int_mod_ring[symmetric]], simp add: of_int_of_int_mod_ring) finally have u_prod: "u = (\i \ {0..< int CARD('a)}. ?g i)" . let ?S = "{0.. ?S" hence "?g i \ 1" by auto moreover have mgi: "monic (?g i)" by (rule poly_gcd_monic, insert u0, auto) ultimately have "degree (?g i) > 0" using monic_degree_0 by blast note this mgi } note gS = this have int_set: "int ` set [0.. ?S" and j: "j \ ?S" and gij: "?g i = ?g j" show "i = j" proof (rule ccontr) define S where "S = {0.. S" "j \ S" "finite S" using i j unfolding S_def by auto assume ij: "i \ j" have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = ?g i * ?g j * (\i \ S. ?g i)" unfolding id using S ij by auto also have "\ = ?g i * ?g i * (\i \ S. ?g i)" unfolding gij by simp finally have dvd: "?g i * ?g i dvd u" unfolding dvd_def by auto with sf_u[unfolded square_free_def, THEN conjunct2, rule_format, OF gS(1)[OF i]] show False by simp qed qed have "u = (\i \ {0..< int CARD('a)}. ?g i)" by fact also have "\ = (\i \ ?S. ?g i)" by (rule sym, rule prod.setdiff_irrelevant, auto) also have "\ = \ (set (udivs u))" unfolding udivs_def set_filter set_map by (rule sym, rule prod.reindex_cong[of ?g, OF inj _ refl], auto simp: int_set[symmetric]) finally have u_udivs: "u = \(set (udivs u))" . { fix w assume mem: "w \ set (udivs u)" then obtain i where w: "w = ?g i" and i: "i \ ?S" unfolding udivs_def set_filter set_map int_set by auto have wu: "w dvd u" by (simp add: w) let ?v = "\ j. v - [:of_nat j:]" define j where "j = nat i" from i have j: "of_int i = (of_nat j :: 'a mod_ring)" "j < CARD('a)" unfolding j_def by auto from gS[OF i, folded w] have *: "degree w > 0" "monic w" "w \ 0" by auto from w have "w dvd ?v j" using j by simp hence gcdj: "?gcdn w j = w" by (metis gcd.commute gcd_left_idem j(1) w) { fix j' assume j': "j' < CARD('a)" have "?gcdn w j' \ {1,w}" proof (rule ccontr) assume not: "?gcdn w j' \ {1,w}" with gcdj have neq: "int j' \ int j" by auto (* next step will yield contradiction to square_free u *) let ?h = "?gcdn w j'" from *(3) not have deg: "degree ?h > 0" using monic_degree_0 poly_gcd_monic by auto have hw: "?h dvd w" by auto have "?h dvd ?gcdn u j'" using wu using dvd_trans by auto also have "?gcdn u j' = ?g j'" by simp finally have hj': "?h dvd ?g j'" by auto from divides_degree[OF this] deg u0 have degj': "degree (?g j') > 0" by auto hence j'1: "?g j' \ 1" by auto with j' have mem': "?g j' \ set (udivs u)" unfolding udivs_def by auto from degj' j' have j'S: "int j' \ ?S" by auto from i j have jS: "int j \ ?S" by auto from inj_on_contraD[OF inj neq j'S jS] have neq: "w \ ?g j'" using w j by auto have cop: "\ coprime w (?g j')" using hj' hw deg by (metis coprime_not_unit_not_dvd poly_dvd_1 Nat.neq0_conv) obtain w' where w': "?g j' = w'" by auto from u_udivs sf_u have "square_free (\(set (udivs u)))" by simp from square_free_prodD[OF this finite_set mem mem'] cop neq show False by simp qed } from gS[OF i, folded w] i this have "degree w > 0" "monic w" "\ j. j < CARD('a) \ ?gcdn w j \ {1,w}" by auto } note udivs = this let ?is = "filter (\ i. ?g i \ 1) (map int [0 ..< CARD('a)])" have id: "udivs u = map ?g ?is" unfolding udivs_def filter_map o_def .. have dist: "distinct (udivs u)" unfolding id distinct_map proof (rule conjI[OF distinct_filter], unfold distinct_map) have "?S = set ?is" unfolding int_set[symmetric] by auto thus "inj_on ?g (set ?is)" using inj by auto qed (auto simp: inj_on_def) from u_udivs prod.distinct_set_conv_list[OF dist, of id] have "prod_list (udivs u) = u" by auto note udivs this dist } note udivs = this have facts: "facts = concat (map udivs divs)" unfolding facts_def by auto obtain lin nonlin where part: "List.partition (\ q. degree q = d) facts = (lin,nonlin)" by force from Cons(6) have "f = prod_list us1 * prod_list divs" by auto also have "prod_list divs = prod_list facts" unfolding facts using udivs(4) by (induct divs, auto) finally have f: "f = prod_list us1 * prod_list facts" . note facts' = facts { fix u assume u: "u \ set facts" from u[unfolded facts] obtain u' where u': "u' \ set divs" and u: "u \ set (udivs u')" by auto from u' udivs(1-2)[OF u' u] prod_list_dvd[OF u, unfolded udivs(4)[OF u']] have "degree u > 0" "monic u" "\ u' \ set divs. u dvd u'" by auto } note facts = this have not1: "(v = 1) = False" using False by auto have "us = us1 @ (if length divs = n2 then divs else let (lin, nonlin) = List.partition (\q. degree q = d) facts in lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding Cons(4) facts_def udivs_def' berlekamp_factorization_main.simps Let_def not1 if_False by (rule arg_cong[where f = "\ x. us1 @ x"], rule if_cong, simp_all) (* takes time *) hence res: "us = us1 @ (if length divs = n2 then divs else lin @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin))" unfolding part by auto show ?thesis proof (cases "length divs = n2") case False with res have us: "us = (us1 @ lin) @ berlekamp_factorization_main d nonlin vs2 (n2 - length lin)" by auto from Cons(2) have vs: "vs = (vs1 @ [v]) @ vs2" by auto have f: "f = prod_list ((us1 @ lin) @ nonlin)" unfolding f using prod_list_partition[OF part] by simp { fix u assume "u \ set ((us1 @ lin) @ nonlin)" with part have "u \ set facts \ set us1" by auto with facts Cons(7) have "degree u > 0" by (auto simp: irreducible_degree_field) } note deg = this from berlekamp_basis_length_factorization[OF sf_f n_bb mon_f f deg, unfolded Cons(3)] have "n2 \ length lin" by auto hence n: "n = length (us1 @ lin) + (n2 - length lin)" unfolding Cons(3) by auto show ?thesis proof (rule Cons(1)[OF vs n us _ f]) fix u assume "u \ set nonlin" with part have "u \ set facts" by auto from facts[OF this] show "monic u \ degree u > 0" by auto next fix u assume u: "u \ set (us1 @ lin)" { assume *: "\ (monic u \ irreducible\<^sub>d u)" with Cons(7) u have "u \ set lin" by auto with part have uf: "u \ set facts" and deg: "degree u = d" by auto from facts[OF uf] obtain u' where "u' \ set divs" and uu': "u dvd u'" by auto from this(1) have "u' dvd f" unfolding Cons(6) using prod_list_dvd[of u'] by auto with uu' have "u dvd f" by (rule dvd_trans) from facts[OF uf] d[OF this deg] * have False by auto } thus "monic u \ irreducible u" by auto next fix w u i assume w: "w \ set (vs1 @ [v])" and u: "u \ set (us1 @ lin) \ set nonlin" and i: "i < CARD('a)" from u part have u: "u \ set us1 \ set facts" by auto show "gcd u (w - [:of_nat i:]) \ {1, u}" proof (cases "u \ set us1") case True from Cons(7)[OF this] have "monic u" "irreducible u" by auto thus ?thesis by (rule monic_irreducible_gcd) next case False with u have u: "u \ set facts" by auto show ?thesis proof (cases "w = v") case True from u[unfolded facts'] obtain u' where u: "u \ set (udivs u')" and u': "u' \ set divs" by auto from udivs(3)[OF u' u i] show ?thesis unfolding True . next case False with w have w: "w \ set vs1" by auto from u obtain u' where u': "u' \ set divs" and dvd: "u dvd u'" using facts(3)[of u] dvd_refl[of u] by blast from w have "w \ set vs1 \ w = v" by auto from facts(1-2)[OF u] have u: "monic u" by auto from Cons(8)[OF w _ i] u' have "gcd u' (w - [:of_nat i:]) \ {1, u'}" by auto with dvd u show ?thesis by (rule monic_gcd_dvd) qed qed qed next case True with res have us: "us = us1 @ divs" by auto from Cons(3) True have n: "n = length us" unfolding us by auto show ?thesis unfolding us[symmetric] proof (intro conjI ballI) show f: "f = prod_list us" unfolding us using Cons(6) by simp { fix u assume "u \ set us" hence "degree u > 0" using Cons(5) Cons(7)[unfolded irreducible\<^sub>d_def] unfolding us by (auto simp: irreducible_degree_field) } note deg = this fix u assume u: "u \ set us" thus "monic u" unfolding us using Cons(5) Cons(7) by auto show "irreducible u" by (rule berlekamp_basis_irreducible[OF sf_f n_bb mon_f f n[symmetric] deg u]) qed qed next case True (* v = 1 *) with Cons(4) have us: "us = us1 @ berlekamp_factorization_main d divs vs2 n2" by simp from Cons(2) True have vs: "vs = (vs1 @ [1]) @ vs2" by auto show ?thesis proof (rule Cons(1)[OF vs Cons(3) us Cons(5-7)], goal_cases) case (3 v u i) show ?case proof (cases "v = 1") case False with 3 Cons(8)[of v u i] show ?thesis by auto next case True hence deg: "degree (v - [: of_nat i :]) = 0" by (metis (no_types, opaque_lifting) degree_pCons_0 diff_pCons diff_zero pCons_one) from 3(2) Cons(5,7)[of u] have "monic u" by auto from gcd_monic_constant[OF this deg] show ?thesis . qed qed qed next case Nil with vsf have vs1: "vs1 = berlekamp_basis f" by auto from Nil(3) have us: "us = us1 @ divs" by auto from Nil(4,6) have md: "\ u. u \ set us \ monic u \ degree u > 0" unfolding us by (auto simp: irreducible_degree_field) from Nil(7)[unfolded vs1] us have no_further_splitting_possible: "\ u v i. v \ set (berlekamp_basis f) \ u \ set us \ i < CARD('a) \ gcd u (v - [:of_nat i:]) \ {1, u}" by auto from Nil(5) us have prod: "f = prod_list us" by simp show ?case proof (intro conjI ballI) fix u assume u: "u \ set us" from md[OF this] have mon_u: "monic u" and deg_u: "degree u > 0" by auto from prod u have uf: "u dvd f" by (simp add: prod_list_dvd) from monic_square_free_irreducible_factorization[OF mon_f sf_f] obtain P where P: "finite P" "f = \P" "P \ {q. irreducible q \ monic q}" by auto show "irreducible u" proof (rule ccontr) assume irr_u: "\ irreducible u" from not_irreducible_factor_yields_prime_factors[OF uf P deg_u this] obtain pi pj where pij: "pi \ P" "pj \ P" "pi \ pj" "pi dvd u" "pj dvd u" by blast from exists_vector_in_Berlekamp_basis_dvd[OF deg_f berlekamp_basis_basis[OF deg_f, folded vs1] finite_set P pij(1-3) mon_f sf_f irr_u uf mon_u pij(4-5), unfolded vs1] obtain v s where v: "v \ set (berlekamp_basis f)" and gcd: "gcd u (v - [:s:]) \ {1,u}" using is_unit_gcd by auto from surj_of_nat_mod_ring[of s] obtain i where i: "i < CARD('a)" and s: "s = of_nat i" by auto from no_further_splitting_possible[OF v u i] gcd[unfolded s] show False by auto qed qed (insert prod md, auto) qed qed lemma berlekamp_monic_factorization: fixes f::"'a mod_ring poly" assumes sf_f: "square_free f" and us: "berlekamp_monic_factorization d f = us" and d: "\ g. g dvd f \ degree g = d \ irreducible g" and deg: "degree f > 0" and mon: "monic f" shows "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" proof - from us[unfolded berlekamp_monic_factorization_def Let_def] deg have us: "us = [] @ berlekamp_factorization_main d [f] (berlekamp_basis f) (length (berlekamp_basis f))" by (auto) have id: "berlekamp_basis f = [] @ berlekamp_basis f" "length (berlekamp_basis f) = length [] + length (berlekamp_basis f)" "f = prod_list ([] @ [f])" by auto show "f = prod_list us \ (\ u \ set us. monic u \ irreducible u)" by (rule berlekamp_factorization_main[OF sf_f id(1) refl refl id(2) us _ _ _ id(3)], insert mon deg d, auto) qed end end diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field_Factorization_Record_Based.thy @@ -1,606 +1,606 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Record Based Version\ theory Finite_Field_Factorization_Record_Based imports Finite_Field_Factorization Matrix_Record_Based Poly_Mod_Finite_Field_Record_Based "HOL-Types_To_Sets.Types_To_Sets" Jordan_Normal_Form.Matrix_IArray_Impl Jordan_Normal_Form.Gauss_Jordan_IArray_Impl Polynomial_Interpolation.Improved_Code_Equations Polynomial_Factorization.Missing_List begin hide_const(open) monom coeff text \Whereas @{thm finite_field_factorization} provides a result for a polynomials over GF(p), we now develop a theorem which speaks about integer polynomials modulo p.\ lemma (in poly_mod_prime_type) finite_field_factorization_modulo_ring: assumes g: "(g :: 'a mod_ring poly) = of_int_poly f" and sf: "square_free_m f" and fact: "finite_field_factorization g = (d,gs)" and c: "c = to_int_mod_ring d" and fs: "fs = map to_int_poly gs" shows "unique_factorization_m f (c, mset fs)" proof - have [transfer_rule]: "MP_Rel f g" unfolding g MP_Rel_def by (simp add: Mp_f_representative) have sg: "square_free g" by (transfer, rule sf) have [transfer_rule]: "M_Rel c d" unfolding M_Rel_def c by (rule M_to_int_mod_ring) have fs_gs[transfer_rule]: "list_all2 MP_Rel fs gs" unfolding fs list_all2_map1 MP_Rel_def[abs_def] Mp_to_int_poly by (simp add: list.rel_refl) have [transfer_rule]: "rel_mset MP_Rel (mset fs) (mset gs)" using fs_gs using rel_mset_def by blast have [transfer_rule]: "MF_Rel (c,mset fs) (d,mset gs)" unfolding MF_Rel_def by transfer_prover from finite_field_factorization[OF sg fact] have uf: "unique_factorization Irr_Mon g (d,mset gs)" by auto from uf[untransferred] show "unique_factorization_m f (c, mset fs)" . qed text \We now have to implement @{const finite_field_factorization}.\ context fixes p :: int and ff_ops :: "'i arith_ops_record" (* finite-fields *) begin fun power_poly_f_mod_i :: "('i list \ 'i list) \ 'i list \ nat \ 'i list" where "power_poly_f_mod_i modulus a n = (if n = 0 then modulus (one_poly_i ff_ops) - else let (d,r) = Divides.divmod_nat n 2; + else let (d,r) = Euclidean_Division.divmod_nat n 2; rec = power_poly_f_mod_i modulus (modulus (times_poly_i ff_ops a a)) d in if r = 0 then rec else modulus (times_poly_i ff_ops rec a))" declare power_poly_f_mod_i.simps[simp del] fun power_polys_i :: "'i list \ 'i list \ 'i list \ nat \ 'i list list" where "power_polys_i mul_p u curr_p (Suc i) = curr_p # power_polys_i mul_p u (mod_field_poly_i ff_ops (times_poly_i ff_ops curr_p mul_p) u) i" | "power_polys_i mul_p u curr_p 0 = []" lemma length_power_polys_i[simp]: "length (power_polys_i x y z n) = n" by (induct n arbitrary: x y z, auto) definition berlekamp_mat_i :: "'i list \ 'i mat" where "berlekamp_mat_i u = (let n = degree_i u; ze = arith_ops_record.zero ff_ops; on = arith_ops_record.one ff_ops; mul_p = power_poly_f_mod_i (\ v. mod_field_poly_i ff_ops v u) [ze, on] (nat p); xks = power_polys_i mul_p u [on] n in mat_of_rows_list n (map (\ cs. cs @ replicate (n - length cs) ze) xks))" definition berlekamp_resulting_mat_i :: "'i list \ 'i mat" where "berlekamp_resulting_mat_i u = (let Q = berlekamp_mat_i u; n = dim_row Q; QI = mat n n (\ (i,j). if i = j then arith_ops_record.minus ff_ops (Q $$ (i,j)) (arith_ops_record.one ff_ops) else Q $$ (i,j)) in (gauss_jordan_single_i ff_ops (transpose_mat QI)))" definition berlekamp_basis_i :: "'i list \ 'i list list" where "berlekamp_basis_i u = (map (poly_of_list_i ff_ops o list_of_vec) (find_base_vectors_i ff_ops (berlekamp_resulting_mat_i u)))" primrec berlekamp_factorization_main_i :: "'i \ 'i \ nat \ 'i list list \ 'i list list \ nat \ 'i list list" where "berlekamp_factorization_main_i ze on d divs (v # vs) n = ( if v = [on] then berlekamp_factorization_main_i ze on d divs vs n else if length divs = n then divs else let of_int = arith_ops_record.of_int ff_ops; facts = filter (\ w. w \ [on]) [ gcd_poly_i ff_ops u (minus_poly_i ff_ops v (if s = 0 then [] else [of_int (int s)])) . u \ divs, s \ [0 ..< nat p]]; (lin,nonlin) = List.partition (\ q. degree_i q = d) facts in lin @ berlekamp_factorization_main_i ze on d nonlin vs (n - length lin))" | "berlekamp_factorization_main_i ze on d divs [] n = divs" definition berlekamp_monic_factorization_i :: "nat \ 'i list \ 'i list list" where "berlekamp_monic_factorization_i d f = (let vs = berlekamp_basis_i f in berlekamp_factorization_main_i (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops) d [f] vs (length vs))" partial_function (tailrec) dist_degree_factorize_main_i :: "'i \ 'i \ nat \ 'i list \ 'i list \ nat \ (nat \ 'i list) list \ (nat \ 'i list) list" where [code]: "dist_degree_factorize_main_i ze on dv v w d res = (if v = [on] then res else if d + d > dv then (dv, v) # res else let w = power_poly_f_mod_i (\ f. mod_field_poly_i ff_ops f v) w (nat p); d = Suc d; gd = gcd_poly_i ff_ops (minus_poly_i ff_ops w [ze,on]) v in if gd = [on] then dist_degree_factorize_main_i ze on dv v w d res else let v' = div_field_poly_i ff_ops v gd in dist_degree_factorize_main_i ze on (degree_i v') v' (mod_field_poly_i ff_ops w v') d ((d,gd) # res))" definition distinct_degree_factorization_i :: "'i list \ (nat \ 'i list) list" where "distinct_degree_factorization_i f = (let ze = arith_ops_record.zero ff_ops; on = arith_ops_record.one ff_ops in if degree_i f = 1 then [(1,f)] else dist_degree_factorize_main_i ze on (degree_i f) f [ze,on] 0 [])" definition finite_field_factorization_i :: "'i list \ 'i \ 'i list list" where "finite_field_factorization_i f = (if degree_i f = 0 then (lead_coeff_i ff_ops f,[]) else let a = lead_coeff_i ff_ops f; u = smult_i ff_ops (arith_ops_record.inverse ff_ops a) f; gs = (if use_distinct_degree_factorization then distinct_degree_factorization_i u else [(1,u)]); (irr,hs) = List.partition (\ (i,f). degree_i f = i) gs in (a,map snd irr @ concat (map (\ (i,g). berlekamp_monic_factorization_i i g) hs)))" end context prime_field_gen begin lemma power_polys_i: assumes i: "i < n" and [transfer_rule]: "poly_rel f f'" "poly_rel g g'" and h: "poly_rel h h'" shows "poly_rel (power_polys_i ff_ops g f h n ! i) (power_polys g' f' h' n ! i)" using i h proof (induct n arbitrary: h h' i) case (Suc n h h' i) note * = this note [transfer_rule] = *(3) show ?case proof (cases i) case 0 with Suc show ?thesis by auto next case (Suc j) with *(2-) have "j < n" by auto note IH = *(1)[OF this] show ?thesis unfolding Suc by (simp, rule IH, transfer_prover) qed qed simp lemma power_poly_f_mod_i: assumes m: "(poly_rel ===> poly_rel) m (\ x'. x' mod m')" shows "poly_rel f f' \ poly_rel (power_poly_f_mod_i ff_ops m f n) (power_poly_f_mod m' f' n)" proof - from m have m: "\ x x'. poly_rel x x' \ poly_rel (m x) (x' mod m')" unfolding rel_fun_def by auto show "poly_rel f f' \ poly_rel (power_poly_f_mod_i ff_ops m f n) (power_poly_f_mod m' f' n)" proof (induct n arbitrary: f f' rule: less_induct) case (less n f f') note f[transfer_rule] = less(2) show ?case proof (cases "n = 0") case True show ?thesis by (simp add: True power_poly_f_mod_i.simps power_poly_f_mod_binary, rule m[OF poly_rel_one]) next case False hence n: "(n = 0) = False" by simp - obtain q r where div: "Divides.divmod_nat n 2 = (q,r)" by force - from this[unfolded divmod_nat_def] n have "q < n" by auto + obtain q r where div: "Euclidean_Division.divmod_nat n 2 = (q,r)" by force + from this[unfolded Euclidean_Division.divmod_nat_def] n have "q < n" by auto note IH = less(1)[OF this] have rec: "poly_rel (power_poly_f_mod_i ff_ops m (m (times_poly_i ff_ops f f)) q) (power_poly_f_mod m' (f' * f' mod m') q)" by (rule IH, rule m, transfer_prover) have other: "poly_rel (m (times_poly_i ff_ops (power_poly_f_mod_i ff_ops m (m (times_poly_i ff_ops f f)) q) f)) (power_poly_f_mod m' (f' * f' mod m') q * f' mod m')" by (rule m, rule poly_rel_times[unfolded rel_fun_def, rule_format, OF rec f]) show ?thesis unfolding power_poly_f_mod_i.simps[of _ _ _ n] Let_def power_poly_f_mod_binary[of _ _ n] div split n if_False using rec other by auto qed qed qed lemma berlekamp_mat_i[transfer_rule]: "(poly_rel ===> mat_rel R) (berlekamp_mat_i p ff_ops) berlekamp_mat" proof (intro rel_funI) fix f f' let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" assume f[transfer_rule]: "poly_rel f f'" have deg: "degree_i f = degree f'" by transfer_prover { fix i j assume i: "i < degree f'" and j: "j < degree f'" define cs where "cs = (\cs :: 'i list. cs @ replicate (degree f' - length cs) ?ze)" define cs' where "cs' = (\cs :: 'a mod_ring poly. coeffs cs @ replicate (degree f' - length (coeffs cs)) 0)" define poly where "poly = power_polys_i ff_ops (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) f [?on] (degree f')" define poly' where "poly' = (power_polys (power_poly_f_mod f' [:0, 1:] (nat p)) f' 1 (degree f'))" have *: "poly_rel (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) (power_poly_f_mod f' [:0, 1:] (nat p))" by (rule power_poly_f_mod_i, transfer_prover, simp add: poly_rel_def one zero) have [transfer_rule]: "poly_rel (poly ! i) (poly' ! i)" unfolding poly_def poly'_def by (rule power_polys_i[OF i f *], simp add: poly_rel_def one) have *: "list_all2 R (cs (poly ! i)) (cs' (poly' ! i))" unfolding cs_def cs'_def by transfer_prover from list_all2_nthD[OF *[unfolded poly_rel_def], of j] j have "R (cs (poly ! i) ! j) (cs' (poly' ! i) ! j)" unfolding cs_def by auto hence "R (mat_of_rows_list (degree f') (map (\cs. cs @ replicate (degree f' - length cs) ?ze) (power_polys_i ff_ops (power_poly_f_mod_i ff_ops (\v. mod_field_poly_i ff_ops v f) [?ze, ?on] (nat p)) f [?on] (degree f'))) $$ (i, j)) (mat_of_rows_list (degree f') (map (\cs. coeffs cs @ replicate (degree f' - length (coeffs cs)) 0) (power_polys (power_poly_f_mod f' [:0, 1:] (nat p)) f' 1 (degree f'))) $$ (i, j))" unfolding mat_of_rows_list_def length_map length_power_polys_i power_polys_works length_power_polys index_mat[OF i j] split unfolding poly_def cs_def poly'_def cs'_def using i by auto } note main = this show "mat_rel R (berlekamp_mat_i p ff_ops f) (berlekamp_mat f')" unfolding berlekamp_mat_i_def berlekamp_mat_def Let_def nat_p[symmetric] deg unfolding mat_rel_def by (intro conjI allI impI, insert main, auto) qed lemma berlekamp_resulting_mat_i[transfer_rule]: "(poly_rel ===> mat_rel R) (berlekamp_resulting_mat_i p ff_ops) berlekamp_resulting_mat" proof (intro rel_funI) fix f f' assume "poly_rel f f'" from berlekamp_mat_i[unfolded rel_fun_def, rule_format, OF this] have bmi: "mat_rel R (berlekamp_mat_i p ff_ops f) (berlekamp_mat f')" . show "mat_rel R (berlekamp_resulting_mat_i p ff_ops f) (berlekamp_resulting_mat f')" unfolding berlekamp_resulting_mat_def Let_def berlekamp_resulting_mat_i_def by (rule gauss_jordan_i[unfolded rel_fun_def, rule_format], insert bmi, auto simp: mat_rel_def one intro!: minus[unfolded rel_fun_def, rule_format]) qed lemma berlekamp_basis_i[transfer_rule]: "(poly_rel ===> list_all2 poly_rel) (berlekamp_basis_i p ff_ops) berlekamp_basis" unfolding berlekamp_basis_i_def[abs_def] berlekamp_basis_code[abs_def] o_def by transfer_prover lemma berlekamp_factorization_main_i[transfer_rule]: "((=) ===> list_all2 poly_rel ===> list_all2 poly_rel ===> (=) ===> list_all2 poly_rel) (berlekamp_factorization_main_i p ff_ops (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops)) berlekamp_factorization_main" proof (intro rel_funI, clarify, goal_cases) case (1 _ d xs xs' ys ys' _ n) let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" let ?of_int = "arith_ops_record.of_int ff_ops" from 1(2) 1(1) show ?case proof (induct ys ys' arbitrary: xs xs' n rule: list_all2_induct) case (Cons y ys y' ys' xs xs' n) note trans[transfer_rule] = Cons(1,2,4) obtain clar0 clar1 clar2 where clarify: "\ s u. gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)])) = clar0 s u" "[0..u. concat (map (\s. if gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)])) \ [?on] then [gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)]))] else []) [0..concat (map (\u. map (\s. gcd_poly_i ff_ops u (minus_poly_i ff_ops y (if s = 0 then [] else [?of_int (int s)]))) [0.. [?on]]" have Facts: "Facts = facts" unfolding Facts_def facts_def clarify proof (induct xs) case (Cons x xs) show ?case by (simp add: Cons, induct clar1, auto) qed simp define facts' where "facts' = concat (map (\u. concat (map (\x. if gcd u (y' - [:of_nat x:]) \ 1 then [gcd u (y' - [:of_int (int x):])] else []) [0.. x. of_int (int x) = of_nat x" "[?on] = one_poly_i ff_ops" by (auto simp: one_poly_i_def) have facts[transfer_rule]: "list_all2 poly_rel facts facts'" unfolding facts_def facts'_def apply (rule concat_transfer[unfolded rel_fun_def, rule_format]) apply (rule list.map_transfer[unfolded rel_fun_def, rule_format, OF _ trans(3)]) apply (rule concat_transfer[unfolded rel_fun_def, rule_format]) apply (rule list_all2_map_map) proof (unfold id) fix f f' x assume [transfer_rule]: "poly_rel f f'" and x: "x \ set [0.. int x" "int x < p" by auto from of_int[OF this] have rel[transfer_rule]: "R (?of_int (int x)) (of_nat x)" by auto { assume "0 < x" with * have *: "0 < int x" "int x < p" by auto have "(of_nat x :: 'a mod_ring) = of_int (int x)" by simp also have "\ \ 0" unfolding of_int_of_int_mod_ring using * unfolding p by (transfer', auto) } with rel have [transfer_rule]: "poly_rel (if x = 0 then [] else [?of_int (int x)]) [:of_nat x:]" unfolding poly_rel_def by (auto simp add: cCons_def p) show "list_all2 poly_rel (if gcd_poly_i ff_ops f (minus_poly_i ff_ops y (if x = 0 then [] else [?of_int (int x)])) \ one_poly_i ff_ops then [gcd_poly_i ff_ops f (minus_poly_i ff_ops y (if x = 0 then [] else [?of_int (int x)]))] else []) (if gcd f' (y' - [:of_nat x:]) \ 1 then [gcd f' (y' - [:of_nat x:])] else [])" by transfer_prover qed have id1: "berlekamp_factorization_main_i p ff_ops ?ze ?on d xs (y # ys) n = ( if y = [?on] then berlekamp_factorization_main_i p ff_ops ?ze ?on d xs ys n else if length xs = n then xs else (let fac = facts; (lin, nonlin) = List.partition (\q. degree_i q = d) fac in lin @ berlekamp_factorization_main_i p ff_ops ?ze ?on d nonlin ys (n - length lin)))" unfolding berlekamp_factorization_main_i.simps Facts[symmetric] by (simp add: o_def Facts_def Let_def) have id2: "berlekamp_factorization_main d xs' (y' # ys') n = ( if y' = 1 then berlekamp_factorization_main d xs' ys' n else if length xs' = n then xs' else (let fac = facts'; (lin, nonlin) = List.partition (\q. degree q = d) fac in lin @ berlekamp_factorization_main d nonlin ys' (n - length lin)))" by (simp add: o_def facts'_def nat_p) have len: "length xs = length xs'" by transfer_prover have id3: "(y = [?on]) = (y' = 1)" by (transfer_prover_start, transfer_step+, simp add: one_poly_i_def finite_field_ops_int_def) show ?case proof (cases "y' = 1") case True hence id4: "(y' = 1) = True" by simp show ?thesis unfolding id1 id2 id3 id4 if_True by (rule Cons(3), transfer_prover) next case False hence id4: "(y' = 1) = False" by simp note id1 = id1[unfolded id3 id4 if_False] note id2 = id2[unfolded id4 if_False] show ?thesis proof (cases "length xs' = n") case True thus ?thesis unfolding id1 id2 Let_def len using trans by simp next case False hence id: "(length xs' = n) = False" by simp have id': "length [q\facts . degree_i q = d] = length [q\facts'. degree q = d]" by transfer_prover have [transfer_rule]: "list_all2 poly_rel (berlekamp_factorization_main_i p ff_ops ?ze ?on d [x\facts . degree_i x \ d] ys (n - length [q\facts . degree_i q = d])) (berlekamp_factorization_main d [x\facts' . degree x \ d] ys' (n - length [q\facts' . degree q = d]))" unfolding id' by (rule Cons(3), transfer_prover) show ?thesis unfolding id1 id2 Let_def len id if_False unfolding partition_filter_conv o_def split by transfer_prover qed qed qed simp qed lemma berlekamp_monic_factorization_i[transfer_rule]: "((=) ===> poly_rel ===> list_all2 poly_rel) (berlekamp_monic_factorization_i p ff_ops) berlekamp_monic_factorization" unfolding berlekamp_monic_factorization_i_def[abs_def] berlekamp_monic_factorization_def[abs_def] Let_def by transfer_prover lemma dist_degree_factorize_main_i: "poly_rel F f \ poly_rel G g \ list_all2 (rel_prod (=) poly_rel) Res res \ list_all2 (rel_prod (=) poly_rel) (dist_degree_factorize_main_i p ff_ops (arith_ops_record.zero ff_ops) (arith_ops_record.one ff_ops) (degree_i F) F G d Res) (dist_degree_factorize_main f g d res)" proof (induct f g d res arbitrary: F G Res rule: dist_degree_factorize_main.induct) case (1 v w d res V W Res) let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" note simp = dist_degree_factorize_main.simps[of v w d] dist_degree_factorize_main_i.simps[of p ff_ops ?ze ?on "degree_i V" V W d] have v[transfer_rule]: "poly_rel V v" by (rule 1) have w[transfer_rule]: "poly_rel W w" by (rule 1) have res[transfer_rule]: "list_all2 (rel_prod (=) poly_rel) Res res" by (rule 1) have [transfer_rule]: "poly_rel [?on] 1" by (simp add: one poly_rel_def) have id1: "(V = [?on]) = (v = 1)" unfolding finite_field_ops_int_def by transfer_prover have id2: "degree_i V = degree v" by transfer_prover note simp = simp[unfolded id1 id2] note IH = 1(1,2) show ?case proof (cases "v = 1") case True with res show ?thesis unfolding id2 simp by simp next case False with id1 have "(v = 1) = False" by auto note simp = simp[unfolded this if_False] note IH = IH[OF False] show ?thesis proof (cases "degree v < d + d") case True thus ?thesis unfolding id2 simp using res v by auto next case False hence "(degree v < d + d) = False" by auto note simp = simp[unfolded this if_False] let ?P = "power_poly_f_mod_i ff_ops (\f. mod_field_poly_i ff_ops f V) W (nat p)" let ?G = "gcd_poly_i ff_ops (minus_poly_i ff_ops ?P [?ze, ?on]) V" let ?g = "gcd (w ^ CARD('a) mod v - monom 1 1) v" define G where "G = ?G" define g where "g = ?g" note simp = simp[unfolded Let_def, folded G_def g_def] note IH = IH[OF False refl refl refl] have [transfer_rule]: "poly_rel [?ze,?on] (monom 1 1)" unfolding poly_rel_def by (auto simp: coeffs_monom one zero) have id: "w ^ CARD('a) mod v = power_poly_f_mod v w (nat p)" unfolding power_poly_f_mod_def by (simp add: p) have P[transfer_rule]: "poly_rel ?P (w ^ CARD('a) mod v)" unfolding id by (rule power_poly_f_mod_i[OF _ w], transfer_prover) have g[transfer_rule]: "poly_rel G g" unfolding G_def g_def by transfer_prover have id3: "(G = [?on]) = (g = 1)" by transfer_prover note simp = simp[unfolded id3] show ?thesis proof (cases "g = 1") case True from IH(1)[OF this[unfolded g_def] v P res] True show ?thesis unfolding id2 simp by simp next case False have vg: "poly_rel (div_field_poly_i ff_ops V G) (v div g)" by transfer_prover have "poly_rel (mod_field_poly_i ff_ops ?P (div_field_poly_i ff_ops V G)) (w ^ CARD('a) mod v mod (v div g))" by transfer_prover note IH = IH(2)[OF False[unfolded g_def] refl vg[unfolded G_def g_def] this[unfolded G_def g_def], folded g_def G_def] have "list_all2 (rel_prod (=) poly_rel) ((Suc d, G) # Res) ((Suc d, g) # res)" using g res by auto note IH = IH[OF this] from False have "(g = 1) = False" by simp note simp = simp[unfolded this if_False] show ?thesis unfolding id2 simp using IH by simp qed qed qed qed lemma distinct_degree_factorization_i[transfer_rule]: "(poly_rel ===> list_all2 (rel_prod (=) poly_rel)) (distinct_degree_factorization_i p ff_ops) distinct_degree_factorization" proof fix F f assume f[transfer_rule]: "poly_rel F f" have id: "(degree_i F = 1) = (degree f = 1)" by transfer_prover note d = distinct_degree_factorization_i_def distinct_degree_factorization_def let ?ze = "arith_ops_record.zero ff_ops" let ?on = "arith_ops_record.one ff_ops" show "list_all2 (rel_prod (=) poly_rel) (distinct_degree_factorization_i p ff_ops F) (distinct_degree_factorization f)" proof (cases "degree f = 1") case True with id f show ?thesis unfolding d by auto next case False from False id have "?thesis = (list_all2 (rel_prod (=) poly_rel) (dist_degree_factorize_main_i p ff_ops ?ze ?on (degree_i F) F [?ze, ?on] 0 []) (dist_degree_factorize_main f (monom 1 1) 0 []))" unfolding d Let_def by simp also have \ by (rule dist_degree_factorize_main_i[OF f], auto simp: poly_rel_def coeffs_monom one zero) finally show ?thesis . qed qed lemma finite_field_factorization_i[transfer_rule]: "(poly_rel ===> rel_prod R (list_all2 poly_rel)) (finite_field_factorization_i p ff_ops) finite_field_factorization" unfolding finite_field_factorization_i_def finite_field_factorization_def Let_def lead_coeff_i_def' by transfer_prover text \Since the implementation is sound, we can now combine it with the soundness result of the finite field factorization.\ lemma finite_field_i_sound: assumes f': "f' = of_int_poly_i ff_ops (Mp f)" and berl_i: "finite_field_factorization_i p ff_ops f' = (c',fs')" and sq: "square_free_m f" and fs: "fs = map (to_int_poly_i ff_ops) fs'" and c: "c = arith_ops_record.to_int ff_ops c'" shows "unique_factorization_m f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" proof - define f'' :: "'a mod_ring poly" where "f'' = of_int_poly (Mp f)" have rel_f[transfer_rule]: "poly_rel f' f''" by (rule poly_rel_of_int_poly[OF f'], simp add: f''_def) interpret pff: idom_ops "poly_ops ff_ops" poly_rel by (rule idom_ops_poly) obtain c'' fs'' where berl: "finite_field_factorization f'' = (c'',fs'')" by force from rel_funD[OF finite_field_factorization_i rel_f, unfolded rel_prod_conv assms(2) split berl] have rel[transfer_rule]: "R c' c''" "list_all2 poly_rel fs' fs''" by auto from to_int[OF rel(1)] have cc': "c = to_int_mod_ring c''" unfolding c by simp have c: "c \ {0 ..< p}" unfolding cc' by (metis Divides.pos_mod_bound Divides.pos_mod_sign M_to_int_mod_ring atLeastLessThan_iff gr_implies_not_zero nat_le_0 nat_p not_le poly_mod.M_def zero_less_card_finite) { fix f assume "f \ set fs'" with rel(2) obtain f' where "poly_rel f f'" unfolding list_all2_conv_all_nth set_conv_nth by auto hence "is_poly ff_ops f" using fun_cong[OF Domainp_is_poly, of f] unfolding Domainp_iff[abs_def] by auto } hence fs': "Ball (set fs') (is_poly ff_ops)" by auto define mon :: "'a mod_ring poly \ bool" where "mon = monic" have [transfer_rule]: "(poly_rel ===> (=)) (monic_i ff_ops) mon" unfolding mon_def by (rule poly_rel_monic) have len: "length fs' = length fs''" by transfer_prover have fs': "fs = map to_int_poly fs''" unfolding fs proof (rule nth_map_conv[OF len], intro allI impI) fix i assume i: "i < length fs'" obtain f g where id: "fs' ! i = f" "fs'' ! i = g" by auto from i rel(2)[unfolded list_all2_conv_all_nth[of _ fs' fs'']] id have "poly_rel f g" by auto from to_int_poly_i[OF this] have "to_int_poly_i ff_ops f = to_int_poly g" . thus "to_int_poly_i ff_ops (fs' ! i) = to_int_poly (fs'' ! i)" unfolding id . qed have f: "f'' = of_int_poly f" unfolding poly_eq_iff f''_def by (simp add: to_int_mod_ring_hom.injectivity to_int_mod_ring_of_int_M Mp_coeff) have *: "unique_factorization_m f (c, mset fs)" using finite_field_factorization_modulo_ring[OF f sq berl cc' fs'] by auto have fs': "(\fi\set fs. set (coeffs fi) \ {0.. int poly \ int \ int poly list" where "finite_field_factorization_main p f_ops f \ let (c',fs') = finite_field_factorization_i p f_ops (of_int_poly_i f_ops (poly_mod.Mp p f)) in (arith_ops_record.to_int f_ops c', map (to_int_poly_i f_ops) fs')" lemma(in prime_field_gen) finite_field_factorization_main: assumes res: "finite_field_factorization_main p ff_ops f = (c,fs)" and sq: "square_free_m f" shows "unique_factorization_m f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" proof - obtain c' fs' where res': "finite_field_factorization_i p ff_ops (of_int_poly_i ff_ops (Mp f)) = (c', fs')" by force show ?thesis by (rule finite_field_i_sound[OF refl res' sq], insert res[unfolded finite_field_factorization_main_def res'], auto) qed definition finite_field_factorization_int :: "int \ int poly \ int \ int poly list" where "finite_field_factorization_int p = ( if p \ 65535 then finite_field_factorization_main p (finite_field_ops32 (uint32_of_int p)) else if p \ 4294967295 then finite_field_factorization_main p (finite_field_ops64 (uint64_of_int p)) else finite_field_factorization_main p (finite_field_ops_integer (integer_of_int p)))" context poly_mod_prime begin lemmas finite_field_factorization_main_integer = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops_integer, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas finite_field_factorization_main_uint32 = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops32, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemmas finite_field_factorization_main_uint64 = prime_field_gen.finite_field_factorization_main [OF prime_field.prime_field_finite_field_ops64, unfolded prime_field_def mod_ring_locale_def, unfolded poly_mod_type_simps, internalize_sort "'a :: prime_card", OF type_to_set, unfolded remove_duplicate_premise, cancel_type_definition, OF non_empty] lemma finite_field_factorization_int: assumes sq: "poly_mod.square_free_m p f" and result: "finite_field_factorization_int p f = (c,fs)" shows "poly_mod.unique_factorization_m p f (c, mset fs) \ c \ {0 ..< p} \ (\ fi \ set fs. set (coeffs fi) \ {0 ..< p})" using finite_field_factorization_main_integer[OF _ sq, of c fs] finite_field_factorization_main_uint32[OF _ _ sq, of c fs] finite_field_factorization_main_uint64[OF _ _ sq, of c fs] result[unfolded finite_field_factorization_int_def] by (auto split: if_splits) end end diff --git a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy --- a/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy +++ b/thys/Berlekamp_Zassenhaus/Finite_Field_Record_Based.thy @@ -1,1657 +1,1657 @@ (* Authors: Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada *) subsection \Finite Fields\ text \We provide four implementations for $GF(p)$ -- the field with $p$ elements for some prime $p$ -- one by int, one by integers, one by 32-bit numbers and one 64-bit implementation. Correctness of the implementations is proven by transfer rules to the type-based version of $GF(p)$.\ theory Finite_Field_Record_Based imports Finite_Field Arithmetic_Record_Based Native_Word.Uint32 Native_Word.Uint64 "HOL-Library.Code_Target_Numeral" Native_Word.Code_Target_Int_Bit begin (* mod on standard case which can immediately be mapped to target languages without considering special cases *) definition mod_nonneg_pos :: "integer \ integer \ integer" where "x \ 0 \ y > 0 \ mod_nonneg_pos x y = (x mod y)" code_printing \ \FIXME illusion of partiality\ constant mod_nonneg_pos \ (SML) "IntInf.mod/ ( _,/ _ )" and (Eval) "IntInf.mod/ ( _,/ _ )" and (OCaml) "Z.rem" and (Haskell) "Prelude.mod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ (k '% l))" definition mod_nonneg_pos_int :: "int \ int \ int" where "mod_nonneg_pos_int x y = int_of_integer (mod_nonneg_pos (integer_of_int x) (integer_of_int y))" lemma mod_nonneg_pos_int[simp]: "x \ 0 \ y > 0 \ mod_nonneg_pos_int x y = (x mod y)" unfolding mod_nonneg_pos_int_def using mod_nonneg_pos_def by simp context fixes p :: int begin definition plus_p :: "int \ int \ int" where "plus_p x y \ let z = x + y in if z \ p then z - p else z" definition minus_p :: "int \ int \ int" where "minus_p x y \ if y \ x then x - y else x + p - y" definition uminus_p :: "int \ int" where "uminus_p x = (if x = 0 then 0 else p - x)" definition mult_p :: "int \ int \ int" where "mult_p x y = (mod_nonneg_pos_int (x * y) p)" fun power_p :: "int \ nat \ int" where "power_p x n = (if n = 0 then 1 else - let (d,r) = Divides.divmod_nat n 2; + let (d,r) = Euclidean_Division.divmod_nat n 2; rec = power_p (mult_p x x) d in if r = 0 then rec else mult_p rec x)" text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p :: "int \ int" where "inverse_p x = (if x = 0 then 0 else power_p x (nat (p - 2)))" definition divide_p :: "int \ int \ int" where "divide_p x y = mult_p x (inverse_p y)" definition finite_field_ops_int :: "int arith_ops_record" where "finite_field_ops_int \ Arith_Ops_Record 0 1 plus_p mult_p minus_p uminus_p divide_p inverse_p (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) (\ x . x) (\ x . x) (\ x. 0 \ x \ x < p)" end context fixes p :: uint32 begin definition plus_p32 :: "uint32 \ uint32 \ uint32" where "plus_p32 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p32 :: "uint32 \ uint32 \ uint32" where "minus_p32 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p32 :: "uint32 \ uint32" where "uminus_p32 x = (if x = 0 then 0 else p - x)" definition mult_p32 :: "uint32 \ uint32 \ uint32" where "mult_p32 x y = (x * y mod p)" lemma int_of_uint32_shift: "int_of_uint32 (drop_bit k n) = (int_of_uint32 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint32_0_iff: "int_of_uint32 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint32_0: "int_of_uint32 0 = 0" unfolding int_of_uint32_0_iff by simp lemma int_of_uint32_ge_0: "int_of_uint32 n \ 0" by (transfer, auto) lemma two_32: "2 ^ LENGTH(32) = (4294967296 :: int)" by simp lemma int_of_uint32_plus: "int_of_uint32 (x + y) = (int_of_uint32 x + int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_minus: "int_of_uint32 (x - y) = (int_of_uint32 x - int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mult: "int_of_uint32 (x * y) = (int_of_uint32 x * int_of_uint32 y) mod 4294967296" by (transfer, unfold uint_word_ariths two_32, rule refl) lemma int_of_uint32_mod: "int_of_uint32 (x mod y) = (int_of_uint32 x mod int_of_uint32 y)" by (transfer, unfold uint_mod two_32, rule refl) lemma int_of_uint32_inv: "0 \ x \ x < 4294967296 \ int_of_uint32 (uint32_of_int x) = x" by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p32 :: "uint32 \ uint32 \ uint32" where "power_p32 x n = (if n = 0 then 1 else let rec = power_p32 (mult_p32 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p32 rec x)" by pat_completeness auto termination proof - { fix n :: uint32 assume "n \ 0" with int_of_uint32_ge_0[of n] int_of_uint32_0_iff[of n] have "int_of_uint32 n > 0" by auto hence "0 < int_of_uint32 n" "int_of_uint32 n div 2 < int_of_uint32 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint32 n))", auto simp: int_of_uint32_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p32 :: "uint32 \ uint32" where "inverse_p32 x = (if x = 0 then 0 else power_p32 x (p - 2))" definition divide_p32 :: "uint32 \ uint32 \ uint32" where "divide_p32 x y = mult_p32 x (inverse_p32 y)" definition finite_field_ops32 :: "uint32 arith_ops_record" where "finite_field_ops32 \ Arith_Ops_Record 0 1 plus_p32 mult_p32 minus_p32 uminus_p32 divide_p32 inverse_p32 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint32_of_int int_of_uint32 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint32_code [code_unfold]: "drop_bit 1 x = (uint32_shiftr x 1)" by (simp add: uint32_shiftr_def) (* ******************************************************************************** *) subsubsection \Transfer Relation\ locale mod_ring_locale = fixes p :: int and ty :: "'a :: nontriv itself" assumes p: "p = int CARD('a)" begin lemma nat_p: "nat p = CARD('a)" unfolding p by simp lemma p2: "p \ 2" unfolding p using nontriv[where 'a = 'a] by auto lemma p2_ident: "int (CARD('a) - 2) = p - 2" using p2 unfolding p by simp definition mod_ring_rel :: "int \ 'a mod_ring \ bool" where "mod_ring_rel x x' = (x = to_int_mod_ring x')" (* domain transfer rules *) lemma Domainp_mod_ring_rel [transfer_domain_rule]: "Domainp (mod_ring_rel) = (\ v. v \ {0 ..< p})" proof - { fix v :: int assume *: "0 \ v" "v < p" have "Domainp mod_ring_rel v" proof show "mod_ring_rel v (of_int_mod_ring v)" unfolding mod_ring_rel_def using * p by auto qed } note * = this show ?thesis by (intro ext iffI, insert range_to_int_mod_ring[where 'a = 'a] *, auto simp: mod_ring_rel_def p) qed (* left/right/bi-unique *) lemma bi_unique_mod_ring_rel [transfer_rule]: "bi_unique mod_ring_rel" "left_unique mod_ring_rel" "right_unique mod_ring_rel" unfolding mod_ring_rel_def bi_unique_def left_unique_def right_unique_def by auto (* left/right-total *) lemma right_total_mod_ring_rel [transfer_rule]: "right_total mod_ring_rel" unfolding mod_ring_rel_def right_total_def by simp (* ************************************************************************************ *) subsubsection \Transfer Rules\ (* 0 / 1 *) lemma mod_ring_0[transfer_rule]: "mod_ring_rel 0 0" unfolding mod_ring_rel_def by simp lemma mod_ring_1[transfer_rule]: "mod_ring_rel 1 1" unfolding mod_ring_rel_def by simp (* addition *) lemma plus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "plus_p p x y = ((x + y) mod p)" proof (cases "p \ x + y") case False thus ?thesis using x y unfolding plus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x + y - p" "x + y - p < p" by auto from True have id: "plus_p p x y = x + y - p" unfolding plus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_plus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (plus_p p) (+)" proof - { fix x y :: "'a mod_ring" have "plus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x + y)" by (transfer, subst plus_p_mod_def, auto, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* subtraction *) lemma minus_p_mod_def: assumes x: "x \ {0 ..< p}" and y: "y \ {0 ..< p}" shows "minus_p p x y = ((x - y) mod p)" proof (cases "x - y < 0") case False thus ?thesis using x y unfolding minus_p_def Let_def by auto next case True from True x y have *: "p > 0" "0 \ x - y + p" "x - y + p < p" by auto from True have id: "minus_p p x y = x - y + p" unfolding minus_p_def by auto show ?thesis unfolding id using * using mod_pos_pos_trivial by fastforce qed lemma mod_ring_minus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (minus_p p) (-)" proof - { fix x y :: "'a mod_ring" have "minus_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x - y)" by (transfer, subst minus_p_mod_def, auto simp: p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* unary minus *) lemma mod_ring_uminus[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (uminus_p p) uminus" proof - { fix x :: "'a mod_ring" have "uminus_p p (to_int_mod_ring x) = to_int_mod_ring (uminus x)" by (transfer, auto simp: uminus_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* multiplication *) lemma mod_ring_mult[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (mult_p p) ((*))" proof - { fix x y :: "'a mod_ring" have "mult_p p (to_int_mod_ring x) (to_int_mod_ring y) = to_int_mod_ring (x * y)" by (transfer, auto simp: mult_p_def p) } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *) qed (* equality *) lemma mod_ring_eq[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> (=)) (=) (=)" by (intro rel_funI, auto simp: mod_ring_rel_def) (* power *) lemma mod_ring_power[transfer_rule]: "(mod_ring_rel ===> (=) ===> mod_ring_rel) (power_p p) (^)" proof (intro rel_funI, clarify, unfold binary_power[symmetric], goal_cases) fix x y n assume xy: "mod_ring_rel x y" from xy show "mod_ring_rel (power_p p x n) (binary_power y n)" proof (induct y n arbitrary: x rule: binary_power.induct) case (1 x n y) note 1(2)[transfer_rule] show ?case proof (cases "n = 0") case True thus ?thesis by (simp add: mod_ring_1) next case False - obtain d r where id: "Divides.divmod_nat n 2 = (d,r)" by force + obtain d r where id: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force let ?int = "power_p p (mult_p p y y) d" let ?gfp = "binary_power (x * x) d" from False have id': "?thesis = (mod_ring_rel (if r = 0 then ?int else mult_p p ?int y) (if r = 0 then ?gfp else ?gfp * x))" unfolding power_p.simps[of _ _ n] binary_power.simps[of _ n] Let_def id split by simp have [transfer_rule]: "mod_ring_rel ?int ?gfp" by (rule 1(1)[OF False refl id[symmetric]], transfer_prover) show ?thesis unfolding id' by transfer_prover qed qed qed declare power_p.simps[simp del] lemma ring_finite_field_ops_int: "ring_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end locale prime_field = mod_ring_locale p ty for p and ty :: "'a :: prime_card itself" begin lemma prime: "prime p" unfolding p using prime_card[where 'a = 'a] by simp (* mod *) lemma mod_ring_mod[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) ((\ x y. if y = 0 then x else 0)) (mod)" proof - { fix x y :: "'a mod_ring" have "(if to_int_mod_ring y = 0 then to_int_mod_ring x else 0) = to_int_mod_ring (x mod y)" unfolding modulo_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* normalize *) lemma mod_ring_normalize[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) ((\ x. if x = 0 then 0 else 1)) normalize" proof - { fix x :: "'a mod_ring" have "(if to_int_mod_ring x = 0 then 0 else 1) = to_int_mod_ring (normalize x)" unfolding normalize_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* unit_factor *) lemma mod_ring_unit_factor[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (\ x. x) unit_factor" proof - { fix x :: "'a mod_ring" have "to_int_mod_ring x = to_int_mod_ring (unit_factor x)" unfolding unit_factor_mod_ring_def by auto } note * = this show ?thesis by (intro rel_funI, auto simp: mod_ring_rel_def *[symmetric]) qed (* inverse *) lemma mod_ring_inverse[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel) (inverse_p p) inverse" proof (intro rel_funI) fix x y assume [transfer_rule]: "mod_ring_rel x y" show "mod_ring_rel (inverse_p p x) (inverse y)" unfolding inverse_p_def inverse_mod_ring_def apply (transfer_prover_start) apply (transfer_step)+ apply (unfold p2_ident) apply (rule refl) done qed (* division *) lemma mod_ring_divide[transfer_rule]: "(mod_ring_rel ===> mod_ring_rel ===> mod_ring_rel) (divide_p p) (/)" unfolding divide_p_def[abs_def] divide_mod_ring_def[abs_def] inverse_mod_ring_def[symmetric] by transfer_prover lemma mod_ring_rel_unsafe: assumes "x < CARD('a)" shows "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" proof - have id: "of_nat x = (of_int (int x) :: 'a mod_ring)" by simp show "mod_ring_rel (int x) (of_nat x)" "0 < x \ of_nat x \ (0 :: 'a mod_ring)" unfolding id unfolding mod_ring_rel_def proof (auto simp add: assms of_int_of_int_mod_ring) assume "0 < x" with assms have "of_int_mod_ring (int x) \ (0 :: 'a mod_ring)" by (metis (no_types) less_imp_of_nat_less less_irrefl of_nat_0_le_iff of_nat_0_less_iff to_int_mod_ring_hom.hom_zero to_int_mod_ring_of_int_mod_ring) thus "of_int_mod_ring (int x) = (0 :: 'a mod_ring) \ False" by blast qed qed lemma finite_field_ops_int: "field_ops (finite_field_ops_int p) mod_ring_rel" by (unfold_locales, auto simp: finite_field_ops_int_def bi_unique_mod_ring_rel right_total_mod_ring_rel mod_ring_divide mod_ring_plus mod_ring_minus mod_ring_uminus mod_ring_inverse mod_ring_mod mod_ring_unit_factor mod_ring_normalize mod_ring_mult mod_ring_eq mod_ring_0 mod_ring_1 Domainp_mod_ring_rel) end text \Once we have proven the soundness of the implementation, we do not care any longer that @{typ "'a mod_ring"} has been defined internally via lifting. Disabling the transfer-rules will hide the internal definition in further applications of transfer.\ lifting_forget mod_ring.lifting text \For soundness of the 32-bit implementation, we mainly prove that this implementation implements the int-based implementation of the mod-ring.\ context mod_ring_locale begin context fixes pp :: "uint32" assumes ppp: "p = int_of_uint32 pp" and small: "p \ 65535" begin lemmas uint32_simps = int_of_uint32_0 int_of_uint32_plus int_of_uint32_minus int_of_uint32_mult definition urel32 :: "uint32 \ int \ bool" where "urel32 x y = (y = int_of_uint32 x \ y < p)" definition mod_ring_rel32 :: "uint32 \ 'a mod_ring \ bool" where "mod_ring_rel32 x y = (\ z. urel32 x z \ mod_ring_rel z y)" lemma urel32_0: "urel32 0 0" unfolding urel32_def using p2 by (simp, transfer, simp) lemma urel32_1: "urel32 1 1" unfolding urel32_def using p2 by (simp, transfer, simp) lemma le_int_of_uint32: "(x \ y) = (int_of_uint32 x \ int_of_uint32 y)" by (transfer, simp add: word_le_def) lemma urel32_plus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (plus_p32 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" let ?p = "int_of_uint32 pp" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p32_def plus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_minus: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (minus_p32 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel32_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint32 using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p32_def minus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_uminus: assumes "urel32 x y" shows "urel32 (uminus_p32 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint32 x" from assms int_of_uint32_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel32_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint32_0_iff using rel small by (auto simp: uint32_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le True if_True using True by (auto simp: uint32_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p32_def uminus_p_def Let_def urel32_def unfolding ppp le False if_False using False by (auto simp: uint32_simps) qed qed lemma urel32_mult: assumes "urel32 x y" "urel32 x' y'" shows "urel32 (mult_p32 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel32_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 65536 * 65536" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 4294967296" by simp show ?thesis unfolding id using small rel unfolding mult_p32_def mult_p_def Let_def urel32_def unfolding ppp by (auto simp: uint32_simps, unfold int_of_uint32_mod int_of_uint32_mult, subst mod_pos_pos_trivial[of _ 4294967296], insert le, auto) qed lemma urel32_eq: assumes "urel32 x y" "urel32 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint32 x" let ?x' = "int_of_uint32 x'" from assms int_of_uint32_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel32_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel32_normalize: assumes x: "urel32 x y" shows "urel32 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel32_eq[OF x urel32_0] using urel32_0 urel32_1 by auto lemma urel32_mod: assumes x: "urel32 x x'" and y: "urel32 y y'" shows "urel32 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel32_eq[OF y urel32_0] using urel32_0 x by auto lemma urel32_power: "urel32 x x' \ urel32 y (int y') \ urel32 (power_p32 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel32_eq[OF y urel32_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel32_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel32_eq[OF y urel32_0] by auto from y have \int y' = int_of_uint32 y\ \int y' < p\ by (simp_all add: urel32_def) - obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force - from divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel32 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel32_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint32.rep_eq modulo_uint32.rep_eq uint_mod \int y' = int_of_uint32 y\) done from urel32_eq[OF this urel32_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel32 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel32_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel32_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p32.simps[of _ _ y] dr' id if_False rem using IH urel32_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel32_inverse: assumes x: "urel32 x x'" shows "urel32 (inverse_p32 pp x) (inverse_p p x')" proof - have p: "urel32 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel32_def unfolding ppp by (simp add: int_of_uint32.rep_eq minus_uint32.rep_eq uint_sub_if') show ?thesis unfolding inverse_p32_def inverse_p_def urel32_eq[OF x urel32_0] using urel32_0 urel32_power[OF x p] by auto qed lemma mod_ring_0_32: "mod_ring_rel32 0 0" using urel32_0 mod_ring_0 unfolding mod_ring_rel32_def by blast lemma mod_ring_1_32: "mod_ring_rel32 1 1" using urel32_1 mod_ring_1 unfolding mod_ring_rel32_def by blast lemma mod_ring_uminus32: "(mod_ring_rel32 ===> mod_ring_rel32) (uminus_p32 pp) uminus" using urel32_uminus mod_ring_uminus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_plus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (plus_p32 pp) (+)" using urel32_plus mod_ring_plus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_minus32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (minus_p32 pp) (-)" using urel32_minus mod_ring_minus unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_mult32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (mult_p32 pp) ((*))" using urel32_mult mod_ring_mult unfolding mod_ring_rel32_def rel_fun_def by blast lemma mod_ring_eq32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> (=)) (=) (=)" using urel32_eq mod_ring_eq unfolding mod_ring_rel32_def rel_fun_def by blast lemma urel32_inj: "urel32 x y \ urel32 x z \ y = z" using urel32_eq[of x y x z] by auto lemma urel32_inj': "urel32 x z \ urel32 y z \ x = y" using urel32_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel32: "bi_unique mod_ring_rel32" "left_unique mod_ring_rel32" "right_unique mod_ring_rel32" using bi_unique_mod_ring_rel urel32_inj' unfolding mod_ring_rel32_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel32_def) lemma right_total_mod_ring_rel32: "right_total mod_ring_rel32" unfolding mod_ring_rel32_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel32 (uint32_of_int z) z" unfolding urel32_def using small unfolding ppp by (auto simp: int_of_uint32_inv) with zy show "\ x z. urel32 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel32: "Domainp mod_ring_rel32 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel32 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel32_def proof let ?i = "int_of_uint32" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel32 x (?i x)" unfolding urel32_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel32 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel32 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel32_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops32: "ring_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, auto simp: finite_field_ops32_def bi_unique_mod_ring_rel32 right_total_mod_ring_rel32 mod_ring_plus32 mod_ring_minus32 mod_ring_uminus32 mod_ring_mult32 mod_ring_eq32 mod_ring_0_32 mod_ring_1_32 Domainp_mod_ring_rel32) end end context prime_field begin context fixes pp :: "uint32" assumes *: "p = int_of_uint32 pp" "p \ 65535" begin lemma mod_ring_normalize32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. if x = 0 then 0 else 1) normalize" using urel32_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_mod32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (\x y. if y = 0 then x else 0) (mod)" using urel32_mod[OF *] mod_ring_mod unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor32: "(mod_ring_rel32 ===> mod_ring_rel32) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_inverse32: "(mod_ring_rel32 ===> mod_ring_rel32) (inverse_p32 pp) inverse" using urel32_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel32_def[OF *] rel_fun_def by blast lemma mod_ring_divide32: "(mod_ring_rel32 ===> mod_ring_rel32 ===> mod_ring_rel32) (divide_p32 pp) (/)" using mod_ring_inverse32 mod_ring_mult32[OF *] unfolding divide_p32_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops32: "field_ops (finite_field_ops32 pp) mod_ring_rel32" by (unfold_locales, insert ring_finite_field_ops32[OF *], auto simp: ring_ops_def finite_field_ops32_def mod_ring_divide32 mod_ring_inverse32 mod_ring_mod32 mod_ring_normalize32) end end (* now there is 64-bit time *) context fixes p :: uint64 begin definition plus_p64 :: "uint64 \ uint64 \ uint64" where "plus_p64 x y \ let z = x + y in if z \ p then z - p else z" definition minus_p64 :: "uint64 \ uint64 \ uint64" where "minus_p64 x y \ if y \ x then x - y else (x + p) - y" definition uminus_p64 :: "uint64 \ uint64" where "uminus_p64 x = (if x = 0 then 0 else p - x)" definition mult_p64 :: "uint64 \ uint64 \ uint64" where "mult_p64 x y = (x * y mod p)" lemma int_of_uint64_shift: "int_of_uint64 (drop_bit k n) = (int_of_uint64 n) div (2 ^ k)" apply transfer apply transfer apply (simp add: take_bit_drop_bit min_def) apply (simp add: drop_bit_eq_div) done lemma int_of_uint64_0_iff: "int_of_uint64 n = 0 \ n = 0" by (transfer, rule uint_0_iff) lemma int_of_uint64_0: "int_of_uint64 0 = 0" unfolding int_of_uint64_0_iff by simp lemma int_of_uint64_ge_0: "int_of_uint64 n \ 0" by (transfer, auto) lemma two_64: "2 ^ LENGTH(64) = (18446744073709551616 :: int)" by simp lemma int_of_uint64_plus: "int_of_uint64 (x + y) = (int_of_uint64 x + int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_minus: "int_of_uint64 (x - y) = (int_of_uint64 x - int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mult: "int_of_uint64 (x * y) = (int_of_uint64 x * int_of_uint64 y) mod 18446744073709551616" by (transfer, unfold uint_word_ariths two_64, rule refl) lemma int_of_uint64_mod: "int_of_uint64 (x mod y) = (int_of_uint64 x mod int_of_uint64 y)" by (transfer, unfold uint_mod two_64, rule refl) lemma int_of_uint64_inv: "0 \ x \ x < 18446744073709551616 \ int_of_uint64 (uint64_of_int x) = x" by transfer (simp add: take_bit_int_eq_self unsigned_of_int) context includes bit_operations_syntax begin function power_p64 :: "uint64 \ uint64 \ uint64" where "power_p64 x n = (if n = 0 then 1 else let rec = power_p64 (mult_p64 x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p64 rec x)" by pat_completeness auto termination proof - { fix n :: uint64 assume "n \ 0" with int_of_uint64_ge_0[of n] int_of_uint64_0_iff[of n] have "int_of_uint64 n > 0" by auto hence "0 < int_of_uint64 n" "int_of_uint64 n div 2 < int_of_uint64 n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_uint64 n))", auto simp: int_of_uint64_shift *) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p64 :: "uint64 \ uint64" where "inverse_p64 x = (if x = 0 then 0 else power_p64 x (p - 2))" definition divide_p64 :: "uint64 \ uint64 \ uint64" where "divide_p64 x y = mult_p64 x (inverse_p64 y)" definition finite_field_ops64 :: "uint64 arith_ops_record" where "finite_field_ops64 \ Arith_Ops_Record 0 1 plus_p64 mult_p64 minus_p64 uminus_p64 divide_p64 inverse_p64 (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) uint64_of_int int_of_uint64 (\ x. 0 \ x \ x < p)" end lemma shiftr_uint64_code [code_unfold]: "drop_bit 1 x = (uint64_shiftr x 1)" by (simp add: uint64_shiftr_def) text \For soundness of the 64-bit implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "uint64" assumes ppp: "p = int_of_uint64 pp" and small: "p \ 4294967295" begin lemmas uint64_simps = int_of_uint64_0 int_of_uint64_plus int_of_uint64_minus int_of_uint64_mult definition urel64 :: "uint64 \ int \ bool" where "urel64 x y = (y = int_of_uint64 x \ y < p)" definition mod_ring_rel64 :: "uint64 \ 'a mod_ring \ bool" where "mod_ring_rel64 x y = (\ z. urel64 x z \ mod_ring_rel z y)" lemma urel64_0: "urel64 0 0" unfolding urel64_def using p2 by (simp, transfer, simp) lemma urel64_1: "urel64 1 1" unfolding urel64_def using p2 by (simp, transfer, simp) lemma le_int_of_uint64: "(x \ y) = (int_of_uint64 x \ int_of_uint64 y)" by (transfer, simp add: word_le_def) lemma urel64_plus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (plus_p64 pp x x') (plus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" let ?p = "int_of_uint64 pp" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using small rel unfolding plus_p64_def plus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_minus: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (minus_p64 pp x x') (minus_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel64_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_uint64 using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using small rel unfolding minus_p64_def minus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_uminus: assumes "urel64 x y" shows "urel64 (uminus_p64 pp x) (uminus_p p y)" proof - let ?x = "int_of_uint64 x" from assms int_of_uint64_ge_0 have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel64_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_uint64_0_iff using rel small by (auto simp: uint64_simps) show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le True if_True using True by (auto simp: uint64_simps) next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using small rel unfolding uminus_p64_def uminus_p_def Let_def urel64_def unfolding ppp le False if_False using False by (auto simp: uint64_simps) qed qed lemma urel64_mult: assumes "urel64 x y" "urel64 x' y'" shows "urel64 (mult_p64 pp x x') (mult_p p y y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel64_def by auto from rel have "?x * ?x' < p * p" by (metis mult_strict_mono') also have "\ \ 4294967296 * 4294967296" by (rule mult_mono, insert p2 small, auto) finally have le: "?x * ?x' < 18446744073709551616" by simp show ?thesis unfolding id using small rel unfolding mult_p64_def mult_p_def Let_def urel64_def unfolding ppp by (auto simp: uint64_simps, unfold int_of_uint64_mod int_of_uint64_mult, subst mod_pos_pos_trivial[of _ 18446744073709551616], insert le, auto) qed lemma urel64_eq: assumes "urel64 x y" "urel64 x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_uint64 x" let ?x' = "int_of_uint64 x'" from assms int_of_uint64_ge_0 have id: "y = ?x" "y' = ?x'" unfolding urel64_def by auto show ?thesis unfolding id by (transfer, transfer) rule qed lemma urel64_normalize: assumes x: "urel64 x y" shows "urel64 (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel64_eq[OF x urel64_0] using urel64_0 urel64_1 by auto lemma urel64_mod: assumes x: "urel64 x x'" and y: "urel64 y y'" shows "urel64 (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel64_eq[OF y urel64_0] using urel64_0 x by auto lemma urel64_power: "urel64 x x' \ urel64 y (int y') \ urel64 (power_p64 pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' = 0") case True hence y: "y = 0" using urel64_eq[OF y urel64_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel64_1) next case False hence id: "(y = 0) = False" "(y' = 0) = False" using urel64_eq[OF y urel64_0] by auto from y have \int y' = int_of_uint64 y\ \int y' < p\ by (simp_all add: urel64_def) - obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force - from divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have "urel64 (y AND 1) r'" using \int y' < p\ small apply (simp add: urel64_def and_one_eq r') apply (auto simp add: ppp and_one_eq) apply (simp add: of_nat_mod int_of_uint64.rep_eq modulo_uint64.rep_eq uint_mod \int y' = int_of_uint64 y\) done from urel64_eq[OF this urel64_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel64 (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel64_def using small unfolding ppp apply transfer apply transfer apply (auto simp add: drop_bit_Suc take_bit_int_eq_self) done note IH = 1(1)[OF False refl dr'[symmetric] urel64_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p64.simps[of _ _ y] dr' id if_False rem using IH urel64_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel64_inverse: assumes x: "urel64 x x'" shows "urel64 (inverse_p64 pp x) (inverse_p p x')" proof - have p: "urel64 (pp - 2) (int (nat (p - 2)))" using p2 small unfolding urel64_def unfolding ppp by (simp add: int_of_uint64.rep_eq minus_uint64.rep_eq uint_sub_if') show ?thesis unfolding inverse_p64_def inverse_p_def urel64_eq[OF x urel64_0] using urel64_0 urel64_power[OF x p] by auto qed lemma mod_ring_0_64: "mod_ring_rel64 0 0" using urel64_0 mod_ring_0 unfolding mod_ring_rel64_def by blast lemma mod_ring_1_64: "mod_ring_rel64 1 1" using urel64_1 mod_ring_1 unfolding mod_ring_rel64_def by blast lemma mod_ring_uminus64: "(mod_ring_rel64 ===> mod_ring_rel64) (uminus_p64 pp) uminus" using urel64_uminus mod_ring_uminus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_plus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (plus_p64 pp) (+)" using urel64_plus mod_ring_plus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_minus64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (minus_p64 pp) (-)" using urel64_minus mod_ring_minus unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_mult64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (mult_p64 pp) ((*))" using urel64_mult mod_ring_mult unfolding mod_ring_rel64_def rel_fun_def by blast lemma mod_ring_eq64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> (=)) (=) (=)" using urel64_eq mod_ring_eq unfolding mod_ring_rel64_def rel_fun_def by blast lemma urel64_inj: "urel64 x y \ urel64 x z \ y = z" using urel64_eq[of x y x z] by auto lemma urel64_inj': "urel64 x z \ urel64 y z \ x = y" using urel64_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel64: "bi_unique mod_ring_rel64" "left_unique mod_ring_rel64" "right_unique mod_ring_rel64" using bi_unique_mod_ring_rel urel64_inj' unfolding mod_ring_rel64_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel64_def) lemma right_total_mod_ring_rel64: "right_total mod_ring_rel64" unfolding mod_ring_rel64_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel64 (uint64_of_int z) z" unfolding urel64_def using small unfolding ppp by (auto simp: int_of_uint64_inv) with zy show "\ x z. urel64 x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel64: "Domainp mod_ring_rel64 = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel64 x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel64_def proof let ?i = "int_of_uint64" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" using small unfolding ppp by (transfer, auto simp: word_less_def) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel64 x (?i x)" unfolding urel64_def using small * unfolding ppp by (transfer, auto simp: word_less_def) qed next assume "\a b. x = a \ (\z. urel64 a z \ mod_ring_rel z b)" then obtain b z where xz: "urel64 x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel64_def using small unfolding ppp by (transfer, auto simp: word_less_def) qed qed lemma ring_finite_field_ops64: "ring_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, auto simp: finite_field_ops64_def bi_unique_mod_ring_rel64 right_total_mod_ring_rel64 mod_ring_plus64 mod_ring_minus64 mod_ring_uminus64 mod_ring_mult64 mod_ring_eq64 mod_ring_0_64 mod_ring_1_64 Domainp_mod_ring_rel64) end end context prime_field begin context fixes pp :: "uint64" assumes *: "p = int_of_uint64 pp" "p \ 4294967295" begin lemma mod_ring_normalize64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. if x = 0 then 0 else 1) normalize" using urel64_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_mod64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (\x y. if y = 0 then x else 0) (mod)" using urel64_mod[OF *] mod_ring_mod unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor64: "(mod_ring_rel64 ===> mod_ring_rel64) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_inverse64: "(mod_ring_rel64 ===> mod_ring_rel64) (inverse_p64 pp) inverse" using urel64_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel64_def[OF *] rel_fun_def by blast lemma mod_ring_divide64: "(mod_ring_rel64 ===> mod_ring_rel64 ===> mod_ring_rel64) (divide_p64 pp) (/)" using mod_ring_inverse64 mod_ring_mult64[OF *] unfolding divide_p64_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops64: "field_ops (finite_field_ops64 pp) mod_ring_rel64" by (unfold_locales, insert ring_finite_field_ops64[OF *], auto simp: ring_ops_def finite_field_ops64_def mod_ring_divide64 mod_ring_inverse64 mod_ring_mod64 mod_ring_normalize64) end end (* and a final implementation via integer *) context fixes p :: integer begin definition plus_p_integer :: "integer \ integer \ integer" where "plus_p_integer x y \ let z = x + y in if z \ p then z - p else z" definition minus_p_integer :: "integer \ integer \ integer" where "minus_p_integer x y \ if y \ x then x - y else (x + p) - y" definition uminus_p_integer :: "integer \ integer" where "uminus_p_integer x = (if x = 0 then 0 else p - x)" definition mult_p_integer :: "integer \ integer \ integer" where "mult_p_integer x y = (x * y mod p)" lemma int_of_integer_0_iff: "int_of_integer n = 0 \ n = 0" using integer_eqI by auto lemma int_of_integer_0: "int_of_integer 0 = 0" unfolding int_of_integer_0_iff by simp lemma int_of_integer_plus: "int_of_integer (x + y) = (int_of_integer x + int_of_integer y)" by simp lemma int_of_integer_minus: "int_of_integer (x - y) = (int_of_integer x - int_of_integer y)" by simp lemma int_of_integer_mult: "int_of_integer (x * y) = (int_of_integer x * int_of_integer y)" by simp lemma int_of_integer_mod: "int_of_integer (x mod y) = (int_of_integer x mod int_of_integer y)" by simp lemma int_of_integer_inv: "int_of_integer (integer_of_int x) = x" by simp lemma int_of_integer_shift: "int_of_integer (drop_bit k n) = (int_of_integer n) div (2 ^ k)" by transfer (simp add: int_of_integer_pow shiftr_integer_conv_div_pow2) context includes bit_operations_syntax begin function power_p_integer :: "integer \ integer \ integer" where "power_p_integer x n = (if n \ 0 then 1 else let rec = power_p_integer (mult_p_integer x x) (drop_bit 1 n) in if n AND 1 = 0 then rec else mult_p_integer rec x)" by pat_completeness auto termination proof - { fix n :: integer assume "\ (n \ 0)" hence "n > 0" by auto hence "int_of_integer n > 0" by (simp add: less_integer.rep_eq) hence "0 < int_of_integer n" "int_of_integer n div 2 < int_of_integer n" by auto } note * = this show ?thesis by (relation "measure (\ (x,n). nat (int_of_integer n))", auto simp: * int_of_integer_shift) qed end text \In experiments with Berlekamp-factorization (where the prime $p$ is usually small), it turned out that taking the below implementation of inverse via exponentiation is faster than the one based on the extended Euclidean algorithm.\ definition inverse_p_integer :: "integer \ integer" where "inverse_p_integer x = (if x = 0 then 0 else power_p_integer x (p - 2))" definition divide_p_integer :: "integer \ integer \ integer" where "divide_p_integer x y = mult_p_integer x (inverse_p_integer y)" definition finite_field_ops_integer :: "integer arith_ops_record" where "finite_field_ops_integer \ Arith_Ops_Record 0 1 plus_p_integer mult_p_integer minus_p_integer uminus_p_integer divide_p_integer inverse_p_integer (\ x y . if y = 0 then x else 0) (\ x . if x = 0 then 0 else 1) (\ x . x) integer_of_int int_of_integer (\ x. 0 \ x \ x < p)" end lemma shiftr_integer_code [code_unfold]: "drop_bit 1 x = (integer_shiftr x 1)" unfolding shiftr_integer_code using integer_of_nat_1 by auto text \For soundness of the integer implementation, we mainly prove that this implementation implements the int-based implementation of GF(p).\ context mod_ring_locale begin context fixes pp :: "integer" assumes ppp: "p = int_of_integer pp" begin lemmas integer_simps = int_of_integer_0 int_of_integer_plus int_of_integer_minus int_of_integer_mult definition urel_integer :: "integer \ int \ bool" where "urel_integer x y = (y = int_of_integer x \ y \ 0 \ y < p)" definition mod_ring_rel_integer :: "integer \ 'a mod_ring \ bool" where "mod_ring_rel_integer x y = (\ z. urel_integer x z \ mod_ring_rel z y)" lemma urel_integer_0: "urel_integer 0 0" unfolding urel_integer_def using p2 by simp lemma urel_integer_1: "urel_integer 1 1" unfolding urel_integer_def using p2 by simp lemma le_int_of_integer: "(x \ y) = (int_of_integer x \ int_of_integer y)" by (rule less_eq_integer.rep_eq) lemma urel_integer_plus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (plus_p_integer pp x x') (plus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" let ?p = "int_of_integer pp" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(pp \ x + x') = (?p \ ?x + ?x')" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?p \ ?x + ?x'") case True hence True: "(?p \ ?x + ?x') = True" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?p \ ?x + ?x') = False" by simp show ?thesis unfolding id using rel unfolding plus_p_integer_def plus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_minus: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (minus_p_integer pp x x') (minus_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' \ p" unfolding urel_integer_def by auto have le: "(x' \ x) = (?x' \ ?x)" unfolding le_int_of_integer using rel by auto show ?thesis proof (cases "?x' \ ?x") case True hence True: "(?x' \ ?x) = True" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x' \ ?x) = False" by simp show ?thesis unfolding id using rel unfolding minus_p_integer_def minus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma urel_integer_uminus: assumes "urel_integer x y" shows "urel_integer (uminus_p_integer pp x) (uminus_p p y)" proof - let ?x = "int_of_integer x" from assms have id: "y = ?x" and rel: "0 \ ?x" "?x < p" unfolding urel_integer_def by auto have le: "(x = 0) = (?x = 0)" unfolding int_of_integer_0_iff using rel by auto show ?thesis proof (cases "?x = 0") case True hence True: "(?x = 0) = True" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le True if_True using True by auto next case False hence False: "(?x = 0) = False" by simp show ?thesis unfolding id using rel unfolding uminus_p_integer_def uminus_p_def Let_def urel_integer_def unfolding ppp le False if_False using False by auto qed qed lemma pp_pos: "int_of_integer pp > 0" using ppp nontriv[where 'a = 'a] unfolding p by (simp add: less_integer.rep_eq) lemma urel_integer_mult: assumes "urel_integer x y" "urel_integer x' y'" shows "urel_integer (mult_p_integer pp x x') (mult_p p y y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" and rel: "0 \ ?x" "?x < p" "0 \ ?x'" "?x' < p" unfolding urel_integer_def by auto from rel(1,3) have xx: "0 \ ?x * ?x'" by simp show ?thesis unfolding id using rel unfolding mult_p_integer_def mult_p_def Let_def urel_integer_def unfolding ppp mod_nonneg_pos_int[OF xx pp_pos] using xx pp_pos by simp qed lemma urel_integer_eq: assumes "urel_integer x y" "urel_integer x' y'" shows "(x = x') = (y = y')" proof - let ?x = "int_of_integer x" let ?x' = "int_of_integer x'" from assms have id: "y = ?x" "y' = ?x'" unfolding urel_integer_def by auto show ?thesis unfolding id integer_eq_iff .. qed lemma urel_integer_normalize: assumes x: "urel_integer x y" shows "urel_integer (if x = 0 then 0 else 1) (if y = 0 then 0 else 1)" unfolding urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_1 by auto lemma urel_integer_mod: assumes x: "urel_integer x x'" and y: "urel_integer y y'" shows "urel_integer (if y = 0 then x else 0) (if y' = 0 then x' else 0)" unfolding urel_integer_eq[OF y urel_integer_0] using urel_integer_0 x by auto lemma urel_integer_power: "urel_integer x x' \ urel_integer y (int y') \ urel_integer (power_p_integer pp x y) (power_p p x' y')" including bit_operations_syntax proof (induct x' y' arbitrary: x y rule: power_p.induct[of _ p]) case (1 x' y' x y) note x = 1(2) note y = 1(3) show ?case proof (cases "y' \ 0") case True hence y: "y = 0" "y' = 0" using urel_integer_eq[OF y urel_integer_0] by auto show ?thesis unfolding y True by (simp add: power_p.simps urel_integer_1) next case False hence id: "(y \ 0) = False" "(y' = 0) = False" using False y by (auto simp add: urel_integer_def not_le) (metis of_int_integer_of of_int_of_nat_eq of_nat_0_less_iff) - obtain d' r' where dr': "Divides.divmod_nat y' 2 = (d',r')" by force - from divmod_nat_def[of y' 2, unfolded dr'] + obtain d' r' where dr': "Euclidean_Division.divmod_nat y' 2 = (d',r')" by force + from Euclidean_Division.divmod_nat_def[of y' 2, unfolded dr'] have r': "r' = y' mod 2" and d': "d' = y' div 2" by auto have aux: "\ y'. int (y' mod 2) = int y' mod 2" by presburger have "urel_integer (y AND 1) r'" unfolding r' using y unfolding urel_integer_def unfolding ppp apply (auto simp add: and_one_eq) apply (simp add: of_nat_mod) done from urel_integer_eq[OF this urel_integer_0] have rem: "(y AND 1 = 0) = (r' = 0)" by simp have div: "urel_integer (drop_bit 1 y) (int d')" unfolding d' using y unfolding urel_integer_def unfolding ppp shiftr_integer_conv_div_pow2 by auto from id have "y' \ 0" by auto note IH = 1(1)[OF this refl dr'[symmetric] urel_integer_mult[OF x x] div] show ?thesis unfolding power_p.simps[of _ _ "y'"] power_p_integer.simps[of _ _ y] dr' id if_False rem using IH urel_integer_mult[OF IH x] by (auto simp: Let_def) qed qed lemma urel_integer_inverse: assumes x: "urel_integer x x'" shows "urel_integer (inverse_p_integer pp x) (inverse_p p x')" proof - have p: "urel_integer (pp - 2) (int (nat (p - 2)))" using p2 unfolding urel_integer_def unfolding ppp by auto show ?thesis unfolding inverse_p_integer_def inverse_p_def urel_integer_eq[OF x urel_integer_0] using urel_integer_0 urel_integer_power[OF x p] by auto qed lemma mod_ring_0__integer: "mod_ring_rel_integer 0 0" using urel_integer_0 mod_ring_0 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_1__integer: "mod_ring_rel_integer 1 1" using urel_integer_1 mod_ring_1 unfolding mod_ring_rel_integer_def by blast lemma mod_ring_uminus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (uminus_p_integer pp) uminus" using urel_integer_uminus mod_ring_uminus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_plus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (plus_p_integer pp) (+)" using urel_integer_plus mod_ring_plus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_minus_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (minus_p_integer pp) (-)" using urel_integer_minus mod_ring_minus unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_mult_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (mult_p_integer pp) ((*))" using urel_integer_mult mod_ring_mult unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma mod_ring_eq_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> (=)) (=) (=)" using urel_integer_eq mod_ring_eq unfolding mod_ring_rel_integer_def rel_fun_def by blast lemma urel_integer_inj: "urel_integer x y \ urel_integer x z \ y = z" using urel_integer_eq[of x y x z] by auto lemma urel_integer_inj': "urel_integer x z \ urel_integer y z \ x = y" using urel_integer_eq[of x z y z] by auto lemma bi_unique_mod_ring_rel_integer: "bi_unique mod_ring_rel_integer" "left_unique mod_ring_rel_integer" "right_unique mod_ring_rel_integer" using bi_unique_mod_ring_rel urel_integer_inj' unfolding mod_ring_rel_integer_def bi_unique_def left_unique_def right_unique_def by (auto simp: urel_integer_def) lemma right_total_mod_ring_rel_integer: "right_total mod_ring_rel_integer" unfolding mod_ring_rel_integer_def right_total_def proof fix y :: "'a mod_ring" from right_total_mod_ring_rel[unfolded right_total_def, rule_format, of y] obtain z where zy: "mod_ring_rel z y" by auto hence zp: "0 \ z" "z < p" unfolding mod_ring_rel_def p using range_to_int_mod_ring[where 'a = 'a] by auto hence "urel_integer (integer_of_int z) z" unfolding urel_integer_def unfolding ppp by auto with zy show "\ x z. urel_integer x z \ mod_ring_rel z y" by blast qed lemma Domainp_mod_ring_rel_integer: "Domainp mod_ring_rel_integer = (\x. 0 \ x \ x < pp)" proof fix x show "Domainp mod_ring_rel_integer x = (0 \ x \ x < pp)" unfolding Domainp.simps unfolding mod_ring_rel_integer_def proof let ?i = "int_of_integer" assume *: "0 \ x \ x < pp" hence "0 \ ?i x \ ?i x < p" unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) hence "?i x \ {0 ..< p}" by auto with Domainp_mod_ring_rel have "Domainp mod_ring_rel (?i x)" by auto from this[unfolded Domainp.simps] obtain b where b: "mod_ring_rel (?i x) b" by auto show "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" proof (intro exI, rule conjI[OF refl], rule exI, rule conjI[OF _ b]) show "urel_integer x (?i x)" unfolding urel_integer_def using * unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed next assume "\a b. x = a \ (\z. urel_integer a z \ mod_ring_rel z b)" then obtain b z where xz: "urel_integer x z" and zb: "mod_ring_rel z b" by auto hence "Domainp mod_ring_rel z" by auto with Domainp_mod_ring_rel have "0 \ z" "z < p" by auto with xz show "0 \ x \ x < pp" unfolding urel_integer_def unfolding ppp by (simp add: le_int_of_integer less_integer.rep_eq) qed qed lemma ring_finite_field_ops_integer: "ring_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, auto simp: finite_field_ops_integer_def bi_unique_mod_ring_rel_integer right_total_mod_ring_rel_integer mod_ring_plus_integer mod_ring_minus_integer mod_ring_uminus_integer mod_ring_mult_integer mod_ring_eq_integer mod_ring_0__integer mod_ring_1__integer Domainp_mod_ring_rel_integer) end end context prime_field begin context fixes pp :: "integer" assumes *: "p = int_of_integer pp" begin lemma mod_ring_normalize_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. if x = 0 then 0 else 1) normalize" using urel_integer_normalize[OF *] mod_ring_normalize unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_mod_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (\x y. if y = 0 then x else 0) (mod)" using urel_integer_mod[OF *] mod_ring_mod unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_unit_factor_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (\x. x) unit_factor" using mod_ring_unit_factor unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_inverse_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer) (inverse_p_integer pp) inverse" using urel_integer_inverse[OF *] mod_ring_inverse unfolding mod_ring_rel_integer_def[OF *] rel_fun_def by blast lemma mod_ring_divide_integer: "(mod_ring_rel_integer ===> mod_ring_rel_integer ===> mod_ring_rel_integer) (divide_p_integer pp) (/)" using mod_ring_inverse_integer mod_ring_mult_integer[OF *] unfolding divide_p_integer_def divide_mod_ring_def inverse_mod_ring_def[symmetric] rel_fun_def by blast lemma finite_field_ops_integer: "field_ops (finite_field_ops_integer pp) mod_ring_rel_integer" by (unfold_locales, insert ring_finite_field_ops_integer[OF *], auto simp: ring_ops_def finite_field_ops_integer_def mod_ring_divide_integer mod_ring_inverse_integer mod_ring_mod_integer mod_ring_normalize_integer) end end context prime_field begin (* four implementations of modular integer arithmetic for finite fields *) thm finite_field_ops64 finite_field_ops32 finite_field_ops_integer finite_field_ops_int end context mod_ring_locale begin (* four implementations of modular integer arithmetic for finite rings *) thm ring_finite_field_ops64 ring_finite_field_ops32 ring_finite_field_ops_integer ring_finite_field_ops_int end end diff --git a/thys/Dict_Construction/Test_Lazy_Case.thy b/thys/Dict_Construction/Test_Lazy_Case.thy --- a/thys/Dict_Construction/Test_Lazy_Case.thy +++ b/thys/Dict_Construction/Test_Lazy_Case.thy @@ -1,47 +1,47 @@ subsection \Interaction with \Lazy_Case\\ theory Test_Lazy_Case imports Dict_Construction Lazy_Case.Lazy_Case Show.Show_Instances begin datatype 'a tree = Node | Fork 'a "'a tree list" lemma map_tree[code]: "map_tree f t = (case t of Node \ Node | Fork x ts \ Fork (f x) (map (map_tree f) ts))" by (induction t) auto experiment begin (* FIXME proper qualified path *) text \ Dictionary construction of @{const map_tree} requires the [@{attribute fundef_cong}] rule of @{const Test_Lazy_Case.tree.case_lazy}. \ declassify valid: map_tree thm valid lemma "Test__Lazy__Case_tree_map__tree = map_tree" by (fact valid) end definition i :: "(unit \ (bool list \ string \ nat option) list) option \ string" where "i = show" experiment begin -text \This currently requires @{theory Lazy_Case.Lazy_Case} because of @{const divmod_nat}.\ +text \This currently requires @{theory Lazy_Case.Lazy_Case} because of @{const Euclidean_Division.divmod_nat}.\ (* FIXME get rid of Lazy_Case dependency *) declassify valid: i thm valid lemma "Test__Lazy__Case_i = i" by (fact valid) end end \ No newline at end of file diff --git a/thys/Polynomial_Factorization/Prime_Factorization.thy b/thys/Polynomial_Factorization/Prime_Factorization.thy --- a/thys/Polynomial_Factorization/Prime_Factorization.thy +++ b/thys/Polynomial_Factorization/Prime_Factorization.thy @@ -1,801 +1,801 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Prime Factorization\ text \This theory contains not-completely naive algorithms to test primality and to perform prime factorization. More precisely, it corresponds to prime factorization algorithm A in Knuth's textbook \cite{Knuth}.\ theory Prime_Factorization imports "HOL-Computational_Algebra.Primes" Missing_List Missing_Multiset begin subsection \Definitions\ definition primes_1000 :: "nat list" where "primes_1000 = [2, 3, 5, 7, 11, 13, 17, 19, 23, 29, 31, 37, 41, 43, 47, 53, 59, 61, 67, 71, 73, 79, 83, 89, 97, 101, 103, 107, 109, 113, 127, 131, 137, 139, 149, 151, 157, 163, 167, 173, 179, 181, 191, 193, 197, 199, 211, 223, 227, 229, 233, 239, 241, 251, 257, 263, 269, 271, 277, 281, 283, 293, 307, 311, 313, 317, 331, 337, 347, 349, 353, 359, 367, 373, 379, 383, 389, 397, 401, 409, 419, 421, 431, 433, 439, 443, 449, 457, 461, 463, 467, 479, 487, 491, 499, 503, 509, 521, 523, 541, 547, 557, 563, 569, 571, 577, 587, 593, 599, 601, 607, 613, 617, 619, 631, 641, 643, 647, 653, 659, 661, 673, 677, 683, 691, 701, 709, 719, 727, 733, 739, 743, 751, 757, 761, 769, 773, 787, 797, 809, 811, 821, 823, 827, 829, 839, 853, 857, 859, 863, 877, 881, 883, 887, 907, 911, 919, 929, 937, 941, 947, 953, 967, 971, 977, 983, 991, 997]" lemma primes_1000: "primes_1000 = filter prime [0..<1001]" by eval definition next_candidates :: "nat \ nat \ nat list" where "next_candidates n = (if n = 0 then (1001,primes_1000) else (n + 30, [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]))" definition "candidate_invariant n = (n = 0 \ n mod 30 = (11 :: nat))" partial_function (tailrec) remove_prime_factor :: "nat \ nat \ nat list \ nat \ nat list " where - [code]: "remove_prime_factor p n ps = (case Divides.divmod_nat n p of (n',m) \ + [code]: "remove_prime_factor p n ps = (case Euclidean_Division.divmod_nat n p of (n',m) \ if m = 0 then remove_prime_factor p n' (p # ps) else (n,ps))" partial_function (tailrec) prime_factorization_nat_main :: "nat \ nat \ nat list \ nat list \ nat list" where [code]: "prime_factorization_nat_main n j is ps = (case is of [] \ (case next_candidates j of (j,is) \ prime_factorization_nat_main n j is ps) - | (i # is) \ (case Divides.divmod_nat n i of (n',m) \ + | (i # is) \ (case Euclidean_Division.divmod_nat n i of (n',m) \ if m = 0 then case remove_prime_factor i n' (i # ps) of (n',ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' j is ps' else if i * i \ n then prime_factorization_nat_main n j is ps else (n # ps)))" partial_function (tailrec) prime_nat_main :: "nat \ nat \ nat list \ bool" where [code]: "prime_nat_main n j is = (case is of [] \ (case next_candidates j of (j,is) \ prime_nat_main n j is) | (i # is) \ (if i dvd n then i \ n else if i * i \ n then prime_nat_main n j is else True))" definition prime_nat :: "nat \ bool" where "prime_nat n \ if n < 2 then False else \ \TODO: integrate precomputed map\ case next_candidates 0 of (j,is) \ prime_nat_main n j is" definition prime_factorization_nat :: "nat \ nat list" where "prime_factorization_nat n \ rev (if n < 2 then [] else case next_candidates 0 of (j,is) \ prime_factorization_nat_main n j is [])" definition divisors_nat :: "nat \ nat list" where "divisors_nat n \ if n = 0 then [] else remdups_adj (sort (map prod_list (subseqs (prime_factorization_nat n))))" definition divisors_int_pos :: "int \ int list" where "divisors_int_pos x \ map int (divisors_nat (nat (abs x)))" definition divisors_int :: "int \ int list" where "divisors_int x \ let xs = divisors_int_pos x in xs @ (map uminus xs)" subsection \Proofs\ lemma remove_prime_factor: assumes res: "remove_prime_factor i n ps = (m,qs)" and i: "i > 1" and n: "n \ 0" shows "\ rs. qs = rs @ ps \ n = m * prod_list rs \ \ i dvd m \ set rs \ {i}" using res n proof (induct n arbitrary: ps rule: less_induct) case (less n ps) - obtain n' mo where dm: "Divides.divmod_nat n i = (n',mo)" by force - hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: divmod_nat_def) + obtain n' mo where dm: "Euclidean_Division.divmod_nat n i = (n',mo)" by force + hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: Euclidean_Division.divmod_nat_def) from less(2)[unfolded remove_prime_factor.simps[of i n] dm] have res: "(if mo = 0 then remove_prime_factor i n' (i # ps) else (n, ps)) = (m, qs)" by auto from less(3) have n: "n \ 0" by auto with n' i have "n' < n" by auto note IH = less(1)[OF this] show ?case proof (cases "mo = 0") case True with mo n' have n: "n = n' * i" by auto with \n \ 0\ have n': "n' \ 0" by auto from True res have "remove_prime_factor i n' (i # ps) = (m,qs)" by auto from IH[OF this n'] obtain rs where "qs = rs @ i # ps" and "n' = m * prod_list rs \ \ i dvd m \ set rs \ {i}" by auto thus ?thesis by (intro exI[of _ "rs @ [i]"], unfold n, auto) next case False with mo have i_n: "\ i dvd n" by auto from False res have id: "m = n" "qs = ps" by auto show ?thesis unfolding id using i_n by auto qed qed lemma prime_sqrtI: assumes n: "n \ 2" and small: "\ j. 2 \ j \ j < i \ \ j dvd n" and i: "\ i * i \ n" shows "prime (n::nat)" unfolding prime_nat_iff proof (intro conjI impI allI) show "1 < n" using n by auto fix j assume jn: "j dvd n" from jn obtain k where njk: "n = j * k" unfolding dvd_def by auto with \1 < n\ have jn: "j \ n" by (metis dvd_imp_le jn neq0_conv not_less0) show "j = 1 \ j = n" proof (rule ccontr) assume "\ ?thesis" with njk n have j1: "j > 1 \ j \ n" by simp have "\ j k. 1 < j \ j \ k \ n = j * k" proof (cases "j \ k") case True thus ?thesis unfolding njk using j1 by blast next case False show ?thesis by (rule exI[of _ k], rule exI[of _ j], insert \1 < n\ j1 njk False, auto) (metis Suc_lessI mult_0_right neq0_conv) qed then obtain j k where j1: "1 < j" and jk: "j \ k" and njk: "n = j * k" by auto with small[of j] have ji: "j \ i" unfolding dvd_def by force from mult_mono[OF ji ji] have "i * i \ j * j" by auto with i have "j * j > n" by auto from this[unfolded njk] have "k < j" by auto with jk show False by auto qed qed lemma candidate_invariant_0: "candidate_invariant 0" unfolding candidate_invariant_def by auto lemma next_candidates: assumes res: "next_candidates n = (m,ps)" and n: "candidate_invariant n" shows "candidate_invariant m" "sorted ps" "{i. prime i \ n \ i \ i < m} \ set ps" "set ps \ {2..} \ {n.. []" "n < m" unfolding candidate_invariant_def proof - note res = res[unfolded next_candidates_def] note n = n[unfolded candidate_invariant_def] show "m = 0 \ m mod 30 = 11" using res n by (auto split: if_splits) show "sorted ps" using res n by (auto split: if_splits simp: primes_1000_def sorted2_simps simp del: sorted_wrt.simps(2)) show "set ps \ {2..} \ {n.. []" using res n by (auto split: if_splits simp: primes_1000_def) show "n < m" using res by (auto split: if_splits) show "{i. prime i \ n \ i \ i < m} \ set ps" proof (cases "n = 0") case True hence *: "m = 1001" "ps = primes_1000" using res by auto show ?thesis unfolding * True primes_1000 by auto next case False hence n: "n mod 30 = 11" and m: "m = n + 30" and ps: "ps = [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]" using res n by auto { fix i assume *: "prime i" "n \ i" "i < n + 30" "i \ set ps" from n * have i11: "i \ 11" by auto define j where "j = i - n" have i: "i = n + j" using \n \ i\ j_def by auto have "i mod 30 = (j + n) mod 30" using \n \ i\ unfolding j_def by simp also have "\ = (j mod 30 + n mod 30) mod 30" by (simp add: mod_simps) also have "\ = (j mod 30 + 11) mod 30" unfolding n by simp finally have i30: "i mod 30 = (j mod 30 + 11) mod 30" by simp have 2: "2 dvd (30 :: nat)" and 112: "11 mod (2 :: nat) = 1" by simp_all have "(j + 11) mod 2 = (j + 1) mod 2" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\j. j mod 2"] have 2: "i mod 2 = (j mod 2 + 1) mod 2" by (simp add: mod_simps mod_mod_cancel [OF 2]) have 3: "3 dvd (30 :: nat)" and 113: "11 mod (3 :: nat) = 2" by simp_all have "(j + 11) mod 3 = (j + 2) mod 3" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 3"] have 3: "i mod 3 = (j mod 3 + 2) mod 3" by (simp add: mod_simps mod_mod_cancel [OF 3]) have 5: "5 dvd (30 :: nat)" and 115: "11 mod (5 :: nat) = 1" by simp_all have "(j + 11) mod 5 = (j + 1) mod 5" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 5"] have 5: "i mod 5 = (j mod 5 + 1) mod 5" by (simp add: mod_simps mod_mod_cancel [OF 5]) from n *(2-)[unfolded ps i, simplified] have "j \ {1,3,5,7,9,11,13,15,17,19,21,23,25,27,29} \ j \ {4,10,16,22,28} \ j \ {14,24}" (is "j \ ?j2 \ j \ ?j3 \ j \ ?j5") by simp presburger moreover { assume "j \ ?j2" hence "j mod 2 = 1" by auto with 2 have "i mod 2 = 0" by auto with i11 have "2 dvd i" "i \ 2" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j3" hence "j mod 3 = 1" by auto with 3 have "i mod 3 = 0" by auto with i11 have "3 dvd i" "i \ 3" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j5" hence "j mod 5 = 4" by auto with 5 have "i mod 5 = 0" by auto with i11 have "5 dvd i" "i \ 5" by auto with *(1) have False unfolding prime_nat_iff by auto } ultimately have False by blast } thus ?thesis unfolding m ps by auto qed qed lemma prime_test_iterate2: assumes small: "\ j. 2 \ j \ j < (i :: nat) \ \ j dvd n" and odd: "odd n" and n: "n \ 3" and i: "i \ 3" "odd i" and mod: "\ i dvd n" and j: "2 \ j" "j < i + 2" shows "\ j dvd n" proof assume dvd: "j dvd n" with small[OF j(1)] have "j \ i" by linarith with dvd mod have "j > i" by (cases "i = j", auto) with j have "j = Suc i" by simp with i have "even j" by auto with dvd j(1) have "2 dvd n" by (metis dvd_trans) with odd show False by auto qed lemma prime_divisor: assumes "j \ 2" and "j dvd n" shows "\ p :: nat. prime p \ p dvd j \ p dvd n" proof - let ?pf = "prime_factors j" from assms have "j > 0" by auto from prime_factorization_nat[OF this] have "j = (\p\?pf. p ^ multiplicity p j)" by auto with \j \ 2\ have "?pf \ {}" by auto then obtain p where p: "p \ ?pf" by auto hence pr: "prime p" by auto define rem where "rem = (\p\?pf - {p}. p ^ multiplicity p j)" from p have mult: "multiplicity p j \ 0" by (auto simp: prime_factors_multiplicity) have "finite ?pf" by simp have "j = (\p\?pf. p ^ multiplicity p j)" by fact also have "?pf = (insert p (?pf - {p}))" using p by auto also have "(\p\insert p (?pf - {p}). p ^ multiplicity p j) = p ^ multiplicity p j * rem" unfolding rem_def by (subst prod.insert, auto) also have "\ = p * (p ^ (multiplicity p j - 1) * rem)" using mult by (cases "multiplicity p j", auto) finally have pj: "p dvd j" unfolding dvd_def by blast with \j dvd n\ have "p dvd n" by (metis dvd_trans) with pj pr show ?thesis by blast qed lemma prime_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_nat_main n jj is \ res = prime n" proof (induct ni arbitrary: n i "is" jj res rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_nat_main n jjj iis" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) with res[unfolded simps] have res: "res = (if i' dvd n then n \ i' else if i' * i' \ n then prime_nat_main n jj iis else True)" by simp from 1(11) Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) have "set iis \ {i'..}" by(auto simp: Cons) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by(auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 n _ _ i'(3) sd_iis 1(10) iis] show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_nat_main n jj iis" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis by (rule IH[OF _ dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = True" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True have "i' \ 2" using i i' by auto from \i' dvd n\ obtain k where "n = i' * k" .. with n have "k \ 0" by (cases "k = 0", auto) with \n = i' * k\ have *: "i' < n \ i' = n" by auto with True res have "res \ i' = n" by auto also have "\ = prime n" using * proof assume "i' < n" with \i' \ 2\ \i' dvd n\ have "\ prime n" by (auto simp add: prime_nat_iff) with \i' < n\ show ?thesis by auto next assume "i' = n" with dvd n have "prime n" by (simp add: prime_nat_iff') with \i' = n\ show ?thesis by auto qed finally show ?thesis . qed qed qed lemma prime_factorization_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_factorization_nat_main n jj is ps \ \ qs. res = qs @ ps \ Ball (set qs) prime \ n = prod_list qs" proof (induct ni arbitrary: n i "is" jj res ps rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res ps) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_factorization_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_factorization_nat_main n jjj iis ps" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) - obtain n' m where dm: "Divides.divmod_nat n i' = (n',m)" by force - hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: divmod_nat_def) + obtain n' m where dm: "Euclidean_Division.divmod_nat n i' = (n',m)" by force + hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: Euclidean_Division.divmod_nat_def) have m: "(m = 0) = (i' dvd n)" unfolding m by auto from Cons res[unfolded simps] dm m n' have res: "res = ( if i' dvd n then case remove_prime_factor i' (n div i') (i' # ps) of (n', ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' jj iis ps' else if i' * i' \ n then prime_factorization_nat_main n jj iis ps else n # ps)" by simp from 1(11) i Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" "i' > 1" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) Cons have "set iis \ {i'..}" by(auto) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by (auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 _ _ _ i'(3) sd_iis 1(10) iis] { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_factorization_nat_main n jj iis ps" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast show ?thesis by (rule IH[OF _ n dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = n # ps" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True note i_n = this obtain n'' qs where rp: "remove_prime_factor i' (n div i') (i' # ps) = (n'',qs)" by force with res True have res: "res = (if n'' = 1 then qs else prime_factorization_nat_main n'' jj iis qs)" by auto have pi: "prime i'" unfolding prime_nat_iff proof (intro conjI allI impI) show "1 < i'" using i' i by auto fix j assume ji: "j dvd i'" with i' i have j0: "j \ 0" by (cases "j = 0", auto) from ji i_n have jn: "j dvd n" by (metis dvd_trans) with dvd[of j] have j: "2 > j \ j \ i'" by linarith from ji \1 < i'\ have "j \ i'" unfolding dvd_def by (simp add: dvd_imp_le ji) with j j0 show "j = 1 \ j = i'" by linarith qed from True n' have id: "n = n' * i'" by auto from n id have "n' \ 0" by (cases "n = 0", auto) with id have "i' \ n" by auto from remove_prime_factor[OF rp[folded n'] \1 < i'\ \n' \ 0\] obtain rs where qs: "qs = rs @ i' # ps" and n': "n' = n'' * prod_list rs" and i_n'': "\ i' dvd n''" and rs: "set rs \ {i'}" by auto { fix j assume "j dvd n''" hence "j dvd n" unfolding id n' by auto } note dvd' = this show ?thesis proof (cases "n'' = 1") case False with res have res: "res = prime_factorization_nat_main n'' jj iis qs" by simp from i i' have "i' \ 2" by simp from False n' \n' \ 0\ have n2: "n'' \ 2" by (cases "n'' = 0"; auto) have lrs: "prod_list rs \ 0" using n' \n' \ 0\ by (cases "prod_list rs = 0", auto) with \i' \ 2\ have "prod_list rs * i' \ 2" by (cases "prod_list rs", auto) hence nn'': "n > n''" unfolding id n' using n2 by simp have "i' \ n" unfolding id n' using pi False by fastforce with \i' \ n\ i' have "n > i" by auto with nn'' i i' have less: "n - i > n'' - Suc i'" by simp { fix j assume 2: "2 \ j" and ji: "j < Suc i'" have "\ j dvd n''" proof (cases "j = i'") case False with ji have "j < i'" by auto from dvd' dvd[OF 2 this] show ?thesis by blast qed (insert i_n'', auto) } from IH[OF _ n2 this iis res] less obtain ss where res: "res = ss @ qs \ Ball (set ss) prime \ n'' = prod_list ss" by auto thus ?thesis unfolding id n' qs using pi rs by auto next case True with res have res: "res = qs" by auto show ?thesis unfolding id n' res qs True using rs \prime i'\ by (intro exI[of _ "rs @ [i']"], auto) qed qed qed qed lemma prime_nat[simp]: "prime_nat n = prime n" proof (cases "n < 2") case True thus ?thesis unfolding prime_nat_def prime_nat_iff by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) from n can have res: "prime_nat n = prime_nat_main n jj is" unfolding prime_nat_def by auto show ?thesis using prime_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub res] cann(4) by auto qed lemma prime_factorization_nat: fixes n :: nat defines "pf \ prime_factorization_nat n" shows "Ball (set pf) prime" and "n \ 0 \ prod_list pf = n" and "n = 0 \ pf = []" proof - note pf = pf_def[unfolded prime_factorization_nat_def] have "Ball (set pf) prime \ (n \ 0 \ prod_list pf = n) \ (n = 0 \ pf = [])" proof (cases "n < 2") case True thus ?thesis using pf by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) let ?pfm = "prime_factorization_nat_main n jj is []" from pf[unfolded can] False have res: "pf = rev ?pfm" by simp from prime_factorization_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub refl, of Nil] cann(4) have "Ball (set ?pfm) prime" "n = prod_list ?pfm" by auto thus ?thesis unfolding res using n by auto qed thus "Ball (set pf) prime" "n \ 0 \ prod_list pf = n" "n = 0 \ pf = []" by auto qed lemma prod_mset_multiset_prime_factorization_nat [simp]: "(x::nat) \ 0 \ prod_mset (prime_factorization x) = x" by simp (* TODO Move *) lemma prime_factorization_unique'': fixes A :: "'a :: {factorial_semiring_multiplicative} multiset" assumes "\p. p \# A \ prime p" assumes "prod_mset A = normalize x" shows "prime_factorization x = A" proof - have "prod_mset A \ 0" by (auto dest: assms(1)) with assms(2) have "x \ 0" by simp hence "prod_mset (prime_factorization x) = prod_mset A" by (simp add: assms prod_mset_prime_factorization) with assms show ?thesis by (intro prime_factorization_unique') auto qed lemma multiset_prime_factorization_nat_correct: "prime_factorization n = mset (prime_factorization_nat n)" proof - note pf = prime_factorization_nat[of n] show ?thesis proof (cases "n = 0") case True thus ?thesis using pf(3) by simp next case False note pf = pf(1) pf(2)[OF False] show ?thesis proof (rule prime_factorization_unique'') show "prime p" if "p \# mset (prime_factorization_nat n)" for p using pf(1) that by simp let ?l = "\i\#prime_factorization n. i" let ?r = "\i\#mset (prime_factorization_nat n). i" show "prod_mset (mset (prime_factorization_nat n)) = normalize n" by (simp add: pf(2) prod_mset_prod_list) qed qed qed lemma multiset_prime_factorization_code[code_unfold]: "prime_factorization = (\n. mset (prime_factorization_nat n))" by (intro ext multiset_prime_factorization_nat_correct) lemma divisors_nat: "n \ 0 \ set (divisors_nat n) = {p. p dvd n}" "distinct (divisors_nat n)" "divisors_nat 0 = []" proof - show "distinct (divisors_nat n)" "divisors_nat 0 = []" unfolding divisors_nat_def by auto assume n: "n \ 0" from n have "n > 0" by auto { fix x have "(x dvd n) = (x \ 0 \ (\p. multiplicity p x \ multiplicity p n))" proof (cases "x = 0") case False with \n > 0\ show ?thesis by (auto simp: dvd_multiplicity_eq) next case True with n show ?thesis by auto qed } note dvd = this let ?dn = "set (divisors_nat n)" let ?mf = "\ (n :: nat). prime_factorization n" have "?dn = prod_list ` set (subseqs (prime_factorization_nat n))" unfolding divisors_nat_def using n by auto also have "\ = prod_mset ` mset ` set (subseqs (prime_factorization_nat n))" by (force simp: prod_mset_prod_list) also have "mset ` set (subseqs (prime_factorization_nat n)) = { ps. ps \# mset (prime_factorization_nat n)}" unfolding multiset_of_subseqs by simp also have "\ = { ps. ps \# ?mf n}" thm multiset_prime_factorization_code[symmetric] unfolding multiset_prime_factorization_nat_correct[symmetric] by auto also have "prod_mset ` \ = {p. p dvd n}" (is "?l = ?r") proof - { fix x assume "x dvd n" from this[unfolded dvd] have x: "x \ 0" by auto from \x dvd n\ \x \ 0\ \n \ 0\ have sub: "?mf x \# ?mf n" by (subst prime_factorization_subset_iff_dvd) auto have "prod_mset (?mf x) = x" using x by (simp add: prime_factorization_nat) hence "x \ ?l" using sub by force } moreover { fix x assume "x \ ?l" then obtain ps where x: "x = prod_mset ps" and sub: "ps \# ?mf n" by auto have "x dvd n" using prod_mset_subset_imp_dvd[OF sub] n x by simp } ultimately show ?thesis by blast qed finally show "set (divisors_nat n) = {p. p dvd n}" . qed lemma divisors_int_pos: "x \ 0 \ set (divisors_int_pos x) = {i. i dvd x \ i > 0}" "distinct (divisors_int_pos x)" "divisors_int_pos 0 = []" proof - show "divisors_int_pos 0 = []" by code_simp show "distinct (divisors_int_pos x)" unfolding divisors_int_pos_def using divisors_nat(2)[of "nat (abs x)"] by (simp add: distinct_map inj_on_def) assume x: "x \ 0" let ?x = "nat (abs x)" from x have xx: "?x \ 0" by auto from x have 0: "\ y. y dvd x \ y \ 0" by auto have id: "int ` {p. int p dvd x} = {i. i dvd x \ 0 < i}" (is "?l = ?r") proof - { fix y assume "y \ ?l" then obtain p where y: "y = int p" and dvd: "int p dvd x" by auto have "y \ ?r" unfolding y using dvd 0[OF dvd] by auto } moreover { fix y assume "y \ ?r" hence dvd: "y dvd x" and y0: "y > 0" by auto define n where "n = nat y" from y0 have y: "y = int n" unfolding n_def by auto with dvd have "y \ ?l" by auto } ultimately show ?thesis by blast qed from xx show "set (divisors_int_pos x) = {i. i dvd x \ i > 0}" by (simp add: divisors_int_pos_def divisors_nat id) qed lemma divisors_int: "x \ 0 \ set (divisors_int x) = {i. i dvd x}" "distinct (divisors_int x)" "divisors_int 0 = []" proof - show "divisors_int 0 = []" by code_simp show "distinct (divisors_int x)" proof (cases "x = 0") case True show ?thesis unfolding True by code_simp next case False from divisors_int_pos(1)[OF False] divisors_int_pos(2) show ?thesis unfolding divisors_int_def Let_def distinct_append distinct_map inj_on_def by auto qed assume x: "x \ 0" show "set (divisors_int x) = {i. i dvd x}" unfolding divisors_int_def Let_def set_append set_map divisors_int_pos(1)[OF x] using x by auto (metis (no_types, lifting) dvd_mult_div_cancel image_eqI linorder_neqE_linordered_idom mem_Collect_eq minus_dvd_iff minus_minus mult_zero_left neg_less_0_iff_less) qed definition divisors_fun :: "('a \ ('a :: {comm_monoid_mult,zero}) list) \ bool" where "divisors_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x }) \ (\ x. distinct (df x))" lemma divisors_funD: "divisors_fun df \ x \ 0 \ d dvd x \ d \ set (df x)" unfolding divisors_fun_def by auto definition divisors_pos_fun :: "('a \ ('a :: {comm_monoid_mult,zero,ord}) list) \ bool" where "divisors_pos_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x \ d > 0}) \ (\ x. distinct (df x))" lemma divisors_pos_funD: "divisors_pos_fun df \ x \ 0 \ d dvd x \ d > 0 \ d \ set (df x)" unfolding divisors_pos_fun_def by auto lemma divisors_fun_nat: "divisors_fun divisors_nat" unfolding divisors_fun_def using divisors_nat by auto lemma divisors_fun_int: "divisors_fun divisors_int" unfolding divisors_fun_def using divisors_int by auto lemma divisors_pos_fun_int: "divisors_pos_fun divisors_int_pos" unfolding divisors_pos_fun_def using divisors_int_pos by auto end diff --git a/thys/Polynomial_Interpolation/Improved_Code_Equations.thy b/thys/Polynomial_Interpolation/Improved_Code_Equations.thy --- a/thys/Polynomial_Interpolation/Improved_Code_Equations.thy +++ b/thys/Polynomial_Interpolation/Improved_Code_Equations.thy @@ -1,73 +1,73 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Improved Code Equations\ text \This theory contains improved code equations for certain algorithms.\ theory Improved_Code_Equations imports "HOL-Computational_Algebra.Polynomial" "HOL-Library.Code_Target_Nat" begin subsection \@{const divmod_integer}.\ text \We improve @{thm divmod_integer_code} by deleting @{const sgn}-expressions.\ text \We guard the application of divmod-abs' with the condition @{term "x \ 0 \ y \ 0"}, so that application can be ensured on non-negative values. Hence, one can drop "abs" in target language setup.\ definition divmod_abs' where "x \ 0 \ y \ 0 \ divmod_abs' x y = Code_Numeral.divmod_abs x y" (* led to an another 10 % improvement on factorization example *) lemma divmod_integer_code''[code]: "divmod_integer k l = (if k = 0 then (0, 0) else if l > 0 then (if k > 0 then divmod_abs' k l else case divmod_abs' (- k) l of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, l - s)) else if l = 0 then (0, k) else apsnd uminus (if k < 0 then divmod_abs' (-k) (-l) else case divmod_abs' k (-l) of (r, s) \ if s = 0 then (- r, 0) else (- r - 1, - l - s)))" unfolding divmod_integer_code by (cases "l = 0"; cases "l < 0"; cases "l > 0"; auto split: prod.splits simp: divmod_abs'_def divmod_abs_def) code_printing \ \FIXME illusion of partiality\ constant divmod_abs' \ (SML) "IntInf.divMod/ ( _,/ _ )" and (Eval) "Integer.div'_mod/ ( _ )/ ( _ )" and (OCaml) "Z.div'_rem" and (Haskell) "divMod/ ( _ )/ ( _ )" and (Scala) "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k '/% l))" -subsection \@{const Divides.divmod_nat}.\ -text \We implement @{const Divides.divmod_nat} via @{const divmod_integer} +subsection \@{const Euclidean_Division.divmod_nat}.\ +text \We implement @{const Euclidean_Division.divmod_nat} via @{const divmod_integer} instead of invoking both division and modulo separately, and we further simplify the case-analysis which is performed in @{thm divmod_integer_code''}.\ -lemma divmod_nat_code'[code]: "Divides.divmod_nat m n = ( +lemma divmod_nat_code'[code]: "Euclidean_Division.divmod_nat m n = ( let k = integer_of_nat m; l = integer_of_nat n in map_prod nat_of_integer nat_of_integer (if k = 0 then (0, 0) else if l = 0 then (0,k) else divmod_abs' k l))" using divmod_nat_code [of m n] by (simp add: divmod_abs'_def integer_of_nat_eq_of_nat Let_def) subsection \@{const binomial}\ lemma binomial_code[code]: "n choose k = (if k \ n then fact n div (fact k * fact (n - k)) else 0)" using binomial_eq_0[of n k] binomial_altdef_nat[of k n] by simp end diff --git a/thys/Promela/Promela.thy b/thys/Promela/Promela.thy --- a/thys/Promela/Promela.thy +++ b/thys/Promela/Promela.thy @@ -1,2645 +1,2657 @@ section "Formalization of Promela semantics" theory Promela imports PromelaDatastructures PromelaInvariants PromelaStatistics begin text \Auxiliary\ lemma mod_integer_le: - "a \ b \ 0 < a \ x mod (a + 1) \ b" for a b x :: integer - by (metis add_pos_nonneg discrete not_less order.strict_trans2 - unique_euclidean_semiring_numeral_class.pos_mod_bound zero_le_one) + \x mod (a + 1) \ b\ if \a \ b\ \0 < a\ for a b x :: integer +using that including integer.lifting proof transfer + fix a b x :: int + assume \0 < a\ \a \ b\ + have \x mod (a + 1) < a + 1\ + by (rule pos_mod_bound) (use \0 < a\ in simp) + with \a \ b\ show \x mod (a + 1) \ b\ + by simp +qed lemma mod_integer_ge: - "b \ 0 \ 0 < a \ b \ x mod (a+1)" for a b x :: integer - by (metis dual_order.trans less_add_one order.strict_trans - unique_euclidean_semiring_numeral_class.pos_mod_sign) + \b \ x mod (a + 1)\ if \b \ 0\ \0 < a\ for a b x :: integer +using that including integer.lifting proof transfer + fix a b x :: int + assume \b \ 0\ \0 < a\ + then have \0 \ x mod (a + 1)\ + by simp + with \b \ 0\ show \b \ x mod (a + 1)\ + by simp +qed text \ After having defined the datastructures, we present in this theory how to construct the transition system and how to generate the successors of a state, \ie the real semantics of a Promela program. For the first task, we take the enriched AST as input, the second one operates on the transition system. \ subsection \Misc Helpers\ definition add_label :: "String.literal \ labels \ nat \ labels" where "add_label l lbls pos = ( case lm.lookup l lbls of None \ lm.update l pos lbls | Some _ \ abortv STR ''Label given twice: '' l (\_. lbls))" definition min_prio :: "edge list \ integer \ integer" where "min_prio es start = Min ((prio ` set es) \ {start})" lemma min_prio_code [code]: "min_prio es start = fold (\e pri. if prio e < pri then prio e else pri) es start" proof - from Min.set_eq_fold have "Min (set (start # map prio es)) = fold min (map prio es) start" by metis also have "... = fold (min \ prio) es start" by (simp add: fold_map) also have "... = fold (\e pri. if prio e < pri then prio e else pri) es start" by (auto intro!: fold_cong) finally show ?thesis by (simp add: min_prio_def) qed definition for_all :: "('a \ bool) \ 'a list \ bool" where "for_all f xs \ (\x \ set xs. f x)" lemma for_all_code[code]: "for_all f xs \ foldli xs id (\kv \. f kv) True" by (simp add: for_all_def foldli_conj) definition find_remove :: "('a \ bool) \ 'a list \ 'a option \ 'a list" where "find_remove P xs = (case List.find P xs of None \ (None, xs) | Some x \ (Some x, List.remove1 x xs))" lemma find_remove_code [code]: "find_remove P [] = (None, [])" "find_remove P (x#xs) = (if P x then (Some x, xs) else apsnd (Cons x) (find_remove P xs))" by (induct xs) (auto simp add: find_remove_def dest: find_SomeD split: option.split) lemma find_remove_subset: "find_remove P xs = (res, xs') \ set xs' \ set xs" unfolding find_remove_def using set_remove1_subset by (force split: option.splits) lemma find_remove_length: "find_remove P xs = (res, xs') \ length xs' \ length xs" unfolding find_remove_def by (induct xs arbitrary: res xs') (auto split: if_splits option.splits) subsection \Variable handling\ text \ Handling variables, with their different scopes (global vs. local), and their different types (array vs channel vs bounded) is one of the main challenges of the implementation. \ fun lookupVar :: "variable \ integer option \ integer" where "lookupVar (Var _ val) None = val" | "lookupVar (Var _ _) (Some _) = abort STR ''Array used on var'' (\_.0)" | "lookupVar (VArray _ _ vals) None = vals !! 0" (* sic! *) | "lookupVar (VArray _ siz vals) (Some idx) = vals !! nat_of_integer idx" primrec checkVarValue :: "varType \ integer \ integer" where "checkVarValue (VTBounded lRange hRange) val = ( if val \ hRange \ val \ lRange then val else \ \overflowing is well-defined and may actually be used (e.g. bool)\ if lRange = 0 \ val > 0 then val mod (hRange + 1) else \ \we do not want to implement C-semantics (ie type casts)\ abort STR ''Value overflow'' (\_. lRange))" | "checkVarValue VTChan val = ( if val < min_var_value \ val > max_var_value then abort STR ''Value overflow'' (\_. 0) else val)" lemma [simp]: "variable_inv (Var VTChan 0)" by simp context fixes type :: varType assumes "varType_inv type" begin lemma checkVarValue_bounded: "checkVarValue type val \ {min_var_value..max_var_value}" using \varType_inv type\ by (cases type) (auto intro: mod_integer_le mod_integer_ge) lemma checkVarValue_bounds: "min_var_value \ checkVarValue type val" "checkVarValue type val \ max_var_value" using checkVarValue_bounded [of val] by simp_all lemma checkVarValue_Var: "variable_inv (Var type (checkVarValue type val))" using \varType_inv type\ by (simp add: checkVarValue_bounds) end fun editVar :: "variable \ integer option \ integer \ variable" where "editVar (Var type _ ) None val = Var type (checkVarValue type val)" | "editVar (Var _ _) (Some _) _ = abort STR ''Array used on var'' (\_. Var VTChan 0)" | "editVar (VArray type siz vals) None val = ( let lv = IArray.list_of vals in let v' = lv[0:=checkVarValue type val] in VArray type siz (IArray v'))" | "editVar (VArray type siz vals) (Some idx) val = ( let lv = IArray.list_of vals in let v' = lv[(nat_of_integer idx):=checkVarValue type val] in VArray type siz (IArray v'))" lemma editVar_variable_inv: assumes "variable_inv v" shows "variable_inv (editVar v idx val)" proof (cases v) case (Var type val) with assms have "varType_inv type" by simp with Var show ?thesis by (cases idx) (auto intro!: checkVarValue_Var simp del: checkVarValue.simps variable_inv.simps) next case (VArray type siz vals) with assms have [simp, intro!]: "varType_inv type" by simp show ?thesis proof (cases idx) case None with assms VArray show ?thesis by (cases "IArray.list_of vals") (auto intro!: checkVarValue_bounds) next case (Some i) note upd_cases = in_set_upd_cases[where l="IArray.list_of vals" and i="nat_of_integer i"] from Some VArray assms show ?thesis by (cases type) (auto elim!: upd_cases intro!: mod_integer_le mod_integer_ge simp add: min_var_value_def) qed qed definition getVar' :: "bool \ String.literal \ integer option \ 'a gState_scheme \ pState \ integer option" where "getVar' gl v idx g p = ( let vars = if gl then gState.vars g else pState.vars p in map_option (\x. lookupVar x idx) (lm.lookup v vars))" definition setVar' :: "bool \ String.literal \ integer option \ integer \ 'a gState_scheme \ pState \ 'a gState_scheme * pState" where "setVar' gl v idx val g p = ( if gl then if v = STR ''_'' then (g,p) \ \\''_''\ is a write-only scratch variable\ else case lm.lookup v (gState.vars g) of None \ abortv STR ''Unknown global variable: '' v (\_. (g,p)) | Some x \ (g\gState.vars := lm.update v (editVar x idx val) (gState.vars g)\ , p) else case lm.lookup v (pState.vars p) of None \ abortv STR ''Unknown proc variable: '' v (\_. (g,p)) | Some x \ (g, p\pState.vars := lm.update v (editVar x idx val) (pState.vars p)\))" lemma setVar'_gState_inv: assumes "gState_inv prog g" shows "gState_inv prog (fst (setVar' gl v idx val g p))" unfolding setVar'_def using assms by (auto simp add: gState_inv_def lm.correct intro: editVar_variable_inv split: option.splits) lemma setVar'_gState_progress_rel: assumes "gState_inv prog g" shows "(g, fst (setVar' gl v idx val g p)) \ gState_progress_rel prog" apply (intro gState_progress_relI) apply (fact assms) apply (fact setVar'_gState_inv[OF assms]) apply (auto simp: setVar'_def lm.correct split: option.splits) done lemma vardict_inv_process_names: assumes "vardict_inv ss proc v" and "lm.lookup k v = Some x" shows "k \ process_names ss proc" using assms by (auto simp add: lm.correct vardict_inv_def) lemma vardict_inv_variable_inv: assumes "vardict_inv ss proc v" and "lm.lookup k v = Some x" shows "variable_inv x" using assms by (auto simp add: lm.correct vardict_inv_def) lemma vardict_inv_updateI: assumes "vardict_inv ss proc vs" and "x \ process_names ss proc" and "variable_inv v" shows "vardict_inv ss proc (lm.update x v vs)" using assms by (auto simp add: lm.correct vardict_inv_def) lemma update_vardict_inv: assumes "vardict_inv ss proc v" and "lm.lookup k v = Some x" and "variable_inv x'" shows "vardict_inv ss proc (lm.update k x' v)" using assms by (auto intro!: vardict_inv_updateI vardict_inv_process_names) lemma setVar'_pState_inv: assumes "pState_inv prog p" shows "pState_inv prog (snd (setVar' gl v idx val g p))" unfolding setVar'_def using assms by (auto split: if_splits option.splits simp add: pState_inv_def intro: update_vardict_inv editVar_variable_inv vardict_inv_variable_inv) lemma setVar'_cl_inv: assumes "cl_inv (g,p)" shows "cl_inv (setVar' gl v idx val g p)" unfolding setVar'_def using assms by (auto split: if_splits option.splits) definition withVar' :: "bool \ String.literal \ integer option \ (integer \ 'x) \ 'a gState_scheme \ pState \ 'x" where "withVar' gl v idx f g p = f (the (getVar' gl v idx g p))" definition withChannel' :: "bool \ String.literal \ integer option \ (nat \ channel \ 'x) \ 'a gState_scheme \ pState \ 'x" where "withChannel' gl v idx f g p = ( let error = \_. abortv STR ''Variable is not a channel: '' v (\_. f 0 InvChannel) in let abort = \_. abortv STR ''Channel already closed / invalid: '' v (\_. f 0 InvChannel) in withVar' gl v idx (\i. let i = nat_of_integer i in if i \ length (channels g) then error () else let c = channels g ! i in case c of InvChannel \ abort () | _ \ f i c) g p)" subsection \Expressions\ text \Expressions are free of side-effects. This is in difference to SPIN, where @{term run} is an expression with side-effect. We treat @{term run} as a statement.\ abbreviation "trivCond x \ if x then 1 else 0" fun exprArith :: "'a gState_scheme \ pState \ expr \ integer" and pollCheck :: "'a gState_scheme \ pState \ channel \ recvArg list \ bool \ bool" and recvArgsCheck :: "'a gState_scheme \ pState \ recvArg list \ integer list \ bool" where "exprArith g p (ExprConst x) = x" | "exprArith g p (ExprMConst x _) = x" | "exprArith g p ExprTimeOut = trivCond (timeout g)" | "exprArith g p (ExprLen (ChanRef (VarRef gl name None))) = withChannel' gl name None ( \_ c. case c of Channel _ _ q \ integer_of_nat (length q) | HSChannel _ \ 0) g p" | "exprArith g p (ExprLen (ChanRef (VarRef gl name (Some idx)))) = withChannel' gl name (Some (exprArith g p idx)) ( \_ c. case c of Channel _ _ q \ integer_of_nat (length q) | HSChannel _ \ 0) g p" | "exprArith g p (ExprEmpty (ChanRef (VarRef gl name None))) = trivCond (withChannel' gl name None ( \_ c. case c of Channel _ _ q \ (q = []) | HSChannel _ \ True) g p)" | "exprArith g p (ExprEmpty (ChanRef (VarRef gl name (Some idx)))) = trivCond (withChannel' gl name (Some (exprArith g p idx)) ( \_ c. case c of Channel _ _ q \ (q = []) | HSChannel _ \ True) g p)" | "exprArith g p (ExprFull (ChanRef(VarRef gl name None))) = trivCond (withChannel' gl name None ( \_ c. case c of Channel cap _ q \ integer_of_nat (length q) \ cap | HSChannel _ \ False) g p)" | "exprArith g p (ExprFull (ChanRef(VarRef gl name (Some idx)))) = trivCond (withChannel' gl name (Some (exprArith g p idx)) ( \_ c. case c of Channel cap _ q \ integer_of_nat (length q) \ cap | HSChannel _ \ False) g p)" | "exprArith g p (ExprVarRef (VarRef gl name None)) = withVar' gl name None id g p" | "exprArith g p (ExprVarRef (VarRef gl name (Some idx))) = withVar' gl name (Some (exprArith g p idx)) id g p" | "exprArith g p (ExprUnOp UnOpMinus expr) = 0 - exprArith g p expr" | "exprArith g p (ExprUnOp UnOpNeg expr) = ((exprArith g p expr) + 1) mod 2" | "exprArith g p (ExprBinOp BinOpAdd lexpr rexpr) = (exprArith g p lexpr) + (exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpSub lexpr rexpr) = (exprArith g p lexpr) - (exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpMul lexpr rexpr) = (exprArith g p lexpr) * (exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpDiv lexpr rexpr) = (exprArith g p lexpr) div (exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpMod lexpr rexpr) = (exprArith g p lexpr) mod (exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpGr lexpr rexpr) = trivCond (exprArith g p lexpr > exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpLe lexpr rexpr) = trivCond (exprArith g p lexpr < exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpGEq lexpr rexpr) = trivCond (exprArith g p lexpr \ exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpLEq lexpr rexpr) = trivCond (exprArith g p lexpr \ exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpEq lexpr rexpr) = trivCond (exprArith g p lexpr = exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpNEq lexpr rexpr) = trivCond (exprArith g p lexpr \ exprArith g p rexpr)" | "exprArith g p (ExprBinOp BinOpAnd lexpr rexpr) = trivCond (exprArith g p lexpr \ 0 \ exprArith g p rexpr \ 0)" | "exprArith g p (ExprBinOp BinOpOr lexpr rexpr) = trivCond (exprArith g p lexpr \ 0 \ exprArith g p rexpr \ 0)" | "exprArith g p (ExprCond cexpr texpr fexpr) = (if exprArith g p cexpr \ 0 then exprArith g p texpr else exprArith g p fexpr)" | "exprArith g p (ExprPoll (ChanRef (VarRef gl name None)) rs srt) = trivCond (withChannel' gl name None ( \_ c. pollCheck g p c rs srt) g p)" | "exprArith g p (ExprPoll (ChanRef (VarRef gl name (Some idx))) rs srt) = trivCond (withChannel' gl name (Some (exprArith g p idx)) ( \_ c. pollCheck g p c rs srt) g p)" | "pollCheck g p InvChannel _ _ = abort STR ''Channel already closed / invalid.'' (\_. False)" | "pollCheck g p (HSChannel _) _ _ = False" | "pollCheck g p (Channel _ _ q) rs srt = ( if q = [] then False else if \ srt then recvArgsCheck g p rs (hd q) else List.find (recvArgsCheck g p rs) q \ None)" | "recvArgsCheck _ _ [] [] = True" | "recvArgsCheck _ _ _ [] = abort STR ''Length mismatch on receiving.'' (\_. False)" | "recvArgsCheck _ _ [] _ = abort STR ''Length mismatch on receiving.'' (\_. False)" | "recvArgsCheck g p (r#rs) (v#vs) = (( case r of RecvArgConst c \ c = v | RecvArgMConst c _ \ c = v | RecvArgVar var \ True | RecvArgEval e \ exprArith g p e = v ) \ recvArgsCheck g p rs vs)" text \@{const getVar'} etc.\ do operate on name, index, \ldots directly. Lift them to use @{const VarRef} instead.\ fun liftVar where "liftVar f (VarRef gl v idx) argm g p = f gl v (map_option (exprArith g p) idx) argm g p" definition "getVar v = liftVar (\gl v idx arg. getVar' gl v idx) v ()" definition "setVar = liftVar setVar'" definition "withVar = liftVar withVar'" primrec withChannel where "withChannel (ChanRef v) = liftVar withChannel' v" lemma setVar_gState_progress_rel: assumes "gState_inv prog g" shows "(g, fst (setVar v val g p)) \ gState_progress_rel prog" unfolding setVar_def by (cases v) (simp add: setVar'_gState_progress_rel[OF assms]) lemmas setVar_gState_inv = setVar_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma setVar_pState_inv: assumes "pState_inv prog p" shows "pState_inv prog (snd (setVar v val g p))" unfolding setVar_def by (cases v) (auto simp add: setVar'_pState_inv assms) lemma setVar_cl_inv: assumes "cl_inv (g,p)" shows "cl_inv (setVar v val g p)" unfolding setVar_def by (cases v) (auto simp add: setVar'_cl_inv assms) subsection \Variable declaration\ lemma channel_inv_code [code]: "channel_inv (Channel cap ts q) \ cap \ max_array_size \ 0 \ cap \ for_all varType_inv ts \ length ts \ max_array_size \ length q \ max_array_size \ for_all (\x. length x = length ts \ for_all (\y. y \ min_var_value \ y \ max_var_value) x) q" "channel_inv (HSChannel ts) \ for_all varType_inv ts \ length ts \ max_array_size" by (auto simp add: for_all_def) force+ primrec toVariable :: "'a gState_scheme \ pState \ varDecl \ String.literal * variable * channels" where "toVariable g p (VarDeclNum lb hb name siz init) = ( let type = VTBounded lb hb in if \ varType_inv type then abortv STR ''Invalid var def (varType_inv failed): '' name (\_. (name, Var VTChan 0, [])) else let init = checkVarValue type (case init of None \ 0 | Some e \ exprArith g p e); v = (case siz of None \ Var type init | Some s \ if nat_of_integer s \ max_array_size then VArray type (nat_of_integer s) (IArray.tabulate (s, \_. init)) else abortv STR ''Invalid var def (array too large): '' name (\_. Var VTChan 0)) in (name, v, []))" | "toVariable g p (VarDeclChan name siz types) = ( let size = (case siz of None \ 1 | Some s \ nat_of_integer s); chans = (case types of None \ [] | Some (cap, tys) \ let C = (if cap = 0 then HSChannel tys else Channel cap tys []) in if \ channel_inv C then abortv STR ''Invalid var def (channel_inv failed): '' name (\_. []) else replicate size C); cidx = (case types of None \ 0 | Some _ \ integer_of_nat (length (channels g))); v = (case siz of None \ Var VTChan cidx | Some s \ if nat_of_integer s \ max_array_size then VArray VTChan (nat_of_integer s) (IArray.tabulate (s, \i. if cidx = 0 then 0 else i + cidx)) else abortv STR ''Invalid var def (array too large): '' name (\_. Var VTChan 0)) in (name, v, chans))" lemma toVariable_variable_inv: assumes "gState_inv prog g" shows "variable_inv (fst (snd (toVariable g p v)))" using assms apply (cases v) apply (auto intro!: checkVarValue_Var simp del: variable_inv.simps checkVarValue.simps varType_inv.simps split: if_splits option.splits) apply (auto intro!: mod_integer_ge mod_integer_le simp add: min_var_value_def) apply (simp_all add: assms gState_inv_def max_channels_def max_var_value_def min_var_value_def max_array_size_def) including integer.lifting apply (transfer', simp)+ done lemma toVariable_channels_inv: shows "\x \ set (snd (snd (toVariable g p v))). channel_inv x" by (cases v) auto lemma toVariable_channels_inv': shows "toVariable g p v = (a,b,c) \ \x \ set c. channel_inv x" using toVariable_channels_inv by (metis snd_conv) lemma toVariable_variable_inv': shows "gState_inv prog g \ toVariable g p v = (a,b,c) \ variable_inv b" by (metis snd_conv fst_conv toVariable_variable_inv) definition mkChannels :: "'a gState_scheme \ pState \ channels \ 'a gState_scheme * pState" where "mkChannels g p cs = ( if cs = [] then (g,p) else let l = length (channels g) in if l + length cs > max_channels then abort STR ''Too much channels'' (\_. (g,p)) else let cs\<^sub>p = map integer_of_nat [l..channels := channels g @ cs\; p' = p\pState.channels := pState.channels p @ cs\<^sub>p\ in (g', p'))" lemma mkChannels_gState_progress_rel: "gState_inv prog g \ set cs \ Collect channel_inv \ (g, fst (mkChannels g p cs)) \ gState_progress_rel prog" unfolding mkChannels_def by (intro gState_progress_relI) (auto simp add: gState_inv_def gState.defs cl_inv_def) lemmas mkChannels_gState_inv = mkChannels_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma mkChannels_pState_inv: "pState_inv prog p \ cl_inv (g,p) \ pState_inv prog (snd (mkChannels g p cs))" unfolding mkChannels_def including integer.lifting apply (auto simp add: pState_inv_def pState.defs gState_inv_def dest!: cl_inv_lengthD) apply (transfer', simp)+ done lemma mkChannels_cl_inv: "cl_inv (g,p) \ cl_inv (mkChannels g p cs)" unfolding mkChannels_def by (auto simp add: pState.defs dest: cl_inv_lengthD intro!: cl_invI) definition mkVarChannel :: "varDecl \ ((var_dict \ var_dict) \ 'a gState_scheme * pState \ 'a gState_scheme * pState) \ 'a gState_scheme \ pState \ 'a gState_scheme * pState" where "mkVarChannel v upd g p = ( let (k,v,cs) = toVariable g p v; (g',p') = upd (lm.update k v) (g,p) in mkChannels g' p' cs)" lemma mkVarChannel_gState_inv: assumes "gState_inv prog g" and "\k v' cs. toVariable g p v = (k,v',cs) \ gState_inv prog (fst (upd (lm.update k v') (g,p)))" shows "gState_inv prog (fst (mkVarChannel v upd g p))" using assms unfolding mkVarChannel_def by (force split: varDecl.split prod.split intro!: mkChannels_gState_inv dest: toVariable_channels_inv') lemma mkVarChannel_gState_progress_rel: assumes "gState_inv prog g" and "\k v' cs. toVariable g p v = (k,v',cs) \ (g, fst (upd (lm.update k v') (g,p))) \ gState_progress_rel prog" shows "(g, fst (mkVarChannel v upd g p)) \ gState_progress_rel prog" proof - obtain k v' cs where 1: "toVariable g p v = (k,v',cs)" by (metis prod.exhaust) obtain g' p' where 2: "(g',p') = upd (lm.update k v') (g,p)" by (metis prod.exhaust) with 1 assms have *: "(g, g') \ gState_progress_rel prog" by (metis fst_conv) also from 1 2 have "(g', fst (mkChannels g' p' cs)) \ gState_progress_rel prog" by (force intro!: mkChannels_gState_progress_rel gState_progress_rel_gState_invI2[OF *] dest: toVariable_channels_inv') finally have "(g, fst (mkChannels g' p' cs)) \ gState_progress_rel prog" . thus ?thesis using 1 2 by (auto simp add: mkVarChannel_def split: prod.split) qed lemma mkVarChannel_pState_inv: assumes "pState_inv prog p" and "cl_inv (g,p)" and "\k v' cs. toVariable g p v = (k,v',cs) \ cl_inv (upd (lm.update k v') (g,p))" and "\k v' cs. toVariable g p v = (k,v',cs) \ pState_inv prog (snd (upd (lm.update k v') (g,p)))" shows "pState_inv prog (snd (mkVarChannel v upd g p))" using assms unfolding mkVarChannel_def by (force split: varDecl.split prod.split intro!: mkChannels_pState_inv) lemma mkVarChannel_cl_inv: assumes "cl_inv (g,p)" and "\k v' cs. toVariable g p v = (k,v',cs) \ cl_inv (upd (lm.update k v') (g,p))" shows "cl_inv (mkVarChannel v upd g p)" using assms unfolding mkVarChannel_def by (force split: varDecl.split prod.splits intro!: mkChannels_cl_inv) definition mkVarChannelProc :: "procVarDecl \ 'a gState_scheme \ pState \ 'a gState_scheme * pState" where "mkVarChannelProc v g p = ( let v' = case v of ProcVarDeclNum lb hb name siz init \ VarDeclNum lb hb name siz init | ProcVarDeclChan name siz \ VarDeclChan name siz None; (k,v,cs) = toVariable g p v' in mkVarChannel v' (apsnd \ pState.vars_update) g p)" lemma mkVarChannelProc_gState_progress_rel: assumes "gState_inv prog g" shows "(g, fst (mkVarChannelProc v g p)) \ gState_progress_rel prog" unfolding mkVarChannelProc_def using assms by (auto intro!: mkVarChannel_gState_progress_rel) lemmas mkVarChannelProc_gState_inv = mkVarChannelProc_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma toVariable_name: "toVariable g p (VarDeclNum lb hb name sz init) = (x,a,b) \ x = name" "toVariable g p (VarDeclChan name sz t) = (x, a, b) \ x = name" by (auto split: if_splits) declare toVariable.simps[simp del] lemma statesDecls_process_names: assumes "v \ statesDecls (states prog !! (pState.idx p))" shows "procVarDeclName v \ process_names (states prog !! (pState.idx p)) (processes prog !! (pState.idx p))" using assms by (cases "processes prog !! (pState.idx p)") (auto simp add: statesNames_def) lemma mkVarChannelProc_pState_inv: assumes "pState_inv prog p" and "gState_inv prog g" and "cl_inv (g, p)" and decl: "v \ statesDecls (states prog !! (pState.idx p))" shows "pState_inv prog (snd (mkVarChannelProc v g p))" unfolding mkVarChannelProc_def using assms statesDecls_process_names[OF decl] by (auto intro!: mkVarChannel_pState_inv) (auto dest: toVariable_name split: procVarDecl.splits intro: toVariable_variable_inv' vardict_inv_updateI simp add: pState_inv_def) lemma mkVarChannelProc_cl_inv: assumes "cl_inv (g,p)" shows "cl_inv (mkVarChannelProc v g p)" unfolding mkVarChannelProc_def using assms by (auto intro!: mkVarChannel_cl_inv) subsection \Folding\ text \ Fold over lists (and lists of lists) of @{typ step}/@{typ stmnt}. The folding functions are doing a bit more than that, e.g.\ ensuring the offset into the program array is correct. \ definition step_fold' where "step_fold' g steps (lbls :: labels) pri pos (nxt :: edgeIndex) (onxt :: edgeIndex option) iB = foldr (\step (pos, nxt, lbls, es). let (e,enxt,lbls) = g step (lbls, pri, pos, nxt, onxt, iB) in (pos + length e, enxt, lbls, es@e) ) steps (pos, nxt, lbls, [])" definition step_fold where "step_fold g steps lbls pri pos nxt onxt iB = ( let (_,nxt,lbls,s) = step_fold' g steps lbls pri pos nxt onxt iB in (s,nxt,lbls))" lemma step_fold'_cong: assumes "lbls = lbls'" and "pri = pri'" and "pos = pos'" and "steps = steps'" and "nxt = nxt'" and "onxt = onxt'" and "iB = iB'" and "\x d. x \ set steps \ g x d = g' x d" shows "step_fold' g steps lbls pri pos nxt onxt iB = step_fold' g' steps' lbls' pri' pos' nxt' onxt' iB'" unfolding step_fold'_def by (auto intro: foldr_cong simp add: assms) lemma step_fold_cong[fundef_cong]: assumes "lbls = lbls'" and "pri = pri'" and "pos = pos'" and "steps = steps'" and "nxt = nxt'" and "onxt = onxt'" and "iB = iB'" and "\x d. x \ set steps \ g x d = g' x d" shows "step_fold g steps lbls pri pos nxt onxt iB = step_fold g' steps' lbls' pri' pos' nxt' onxt' iB'" unfolding step_fold_def by (auto simp: assms cong: step_fold'_cong) fun step_foldL_step where "step_foldL_step _ _ _ [] (pos, nxt, lbls, es, is) = (pos, nxt, lbls, es, is)" | "step_foldL_step g pri onxt (s#steps) (pos, nxt, lbls, es, is) = ( let (pos', nxt', lbls', ss') = step_fold' g steps lbls pri pos nxt onxt False in let (s', nxt'', lbls'') = g s (lbls',pri,pos',nxt',onxt,True) in let rs = butlast s'; s'' = last s' in (pos' + length rs, nxt, lbls'', es@ss'@rs, s''#is))" definition step_foldL where "step_foldL g stepss lbls pri pos nxt onxt = foldr (step_foldL_step g pri onxt) stepss (pos,nxt,lbls,[],[])" lemma step_foldL_step_cong: assumes "pri = pri'" and "onxt = onxt'" and "s = s'" and "d = d'" and "\x d. x \ set s \ g x d = g' x d" shows "step_foldL_step g pri onxt s d = step_foldL_step g' pri' onxt' s' d'" using assms by (cases d', cases s') (simp_all cong: step_fold'_cong) lemma step_foldL_cong[fundef_cong]: assumes "lbls = lbls'" and "pri = pri'" and "pos = pos'" and "stepss = stepss'" and "nxt = nxt'" and "onxt = onxt'" and "\x x' d. x \ set stepss \ x' \ set x \ g x' d = g' x' d" shows "step_foldL g stepss lbls pri pos nxt onxt = step_foldL g' stepss' lbls' pri' pos' nxt' onxt'" unfolding step_foldL_def using assms apply (cases stepss') apply simp apply (force intro!: foldr_cong step_foldL_step_cong) done subsection \Starting processes\ definition modProcArg :: "(procArg * integer) \ String.literal * variable" where "modProcArg x = ( case x of (ProcArg ty name, val) \ if varType_inv ty then let init = checkVarValue ty val in (name, Var ty init) else abortv STR ''Invalid proc arg (varType_inv failed)'' name (\_. (name, Var VTChan 0)))" definition emptyProc :: "pState" \ \The empty process.\ where "emptyProc = \pid = 0, vars = lm.empty (), pc = 0, channels = [], idx = 0 \" lemma vardict_inv_empty: "vardict_inv ss proc (lm.empty())" unfolding vardict_inv_def by (simp add: lm.correct) lemma emptyProc_cl_inv[simp]: "cl_inv (g, emptyProc)" by (simp add: cl_inv_def emptyProc_def) lemma emptyProc_pState_inv: assumes "program_inv prog" shows "pState_inv prog emptyProc" proof - from assms have "IArray.length (states prog !! 0) > 0" by (intro program_inv_length_states) (auto simp add: program_inv_def) with assms show ?thesis unfolding pState_inv_def program_inv_def emptyProc_def by (auto simp add: vardict_inv_empty) qed fun mkProc :: "'a gState_scheme \ pState \ String.literal \ expr list \ process \ nat \ 'a gState_scheme * pState" where "mkProc g p name args (sidx, start, argDecls, decls) pidN = ( let start = case start of Index x \ x | _ \ abortv STR ''Process start is not index: '' name (\_. 0) in \ \sanity check\ if length args \ length argDecls then abortv STR ''Signature mismatch: '' name (\_. (g,emptyProc)) else let \ \evaluate args (in the context of the calling process)\ eArgs = map (exprArith g p) args; \ \replace the init part of \argDecls\\ argVars = map modProcArg (zip argDecls eArgs); \ \add \_pid\ to vars\ pidI = integer_of_nat pidN; argVars = (STR ''_pid'', Var (VTBounded 0 pidI) pidI)#argVars; argVars = lm.to_map argVars; \ \our new process\ p = \ pid = pidN, vars = argVars, pc = start, channels = [], idx = sidx \ in \ \apply the declarations\ foldl (\(g,p) d. mkVarChannel d (apsnd \ pState.vars_update) g p) (g,p) decls)" lemma mkProc_gState_progress_rel: assumes "gState_inv prog g" shows "(g, fst (mkProc g p name args (processes prog !! sidx) pidN)) \ gState_progress_rel prog" proof - obtain sidx' start argDecls decls where p: "processes prog !! sidx = (sidx', start, argDecls, decls)" by (metis prod.exhaust) from assms have "\p. (g, fst (foldl (\(g,p) d. mkVarChannel d (apsnd \ pState.vars_update) g p) (g,p) decls)) \ gState_progress_rel prog" proof (induction decls arbitrary: g p) case (Cons d decls) obtain g' p' where new: "(g',p') = (mkVarChannel d (apsnd \ pState.vars_update) g p)" by (metis prod.exhaust) hence "g' = fst ..." by (metis fst_conv) with Cons.prems have g_g': "(g,g') \ gState_progress_rel prog" by (auto intro: mkVarChannel_gState_progress_rel) also note Cons.IH[OF g_g'[THEN gState_progress_rel_gState_invI2], of p'] finally show ?case by (auto simp add: o_def new) qed simp thus ?thesis using assms p by auto qed lemmas mkProc_gState_inv = mkProc_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma mkProc_pState_inv: assumes "program_inv prog" and "gState_inv prog g" and "pidN \ max_procs" and "pidN > 0" and "sidx < IArray.length (processes prog)" and "fst (processes prog !! sidx) = sidx" shows "pState_inv prog (snd (mkProc g p name args (processes prog !! sidx) pidN))" proof - obtain sidx' start argDecls decls where "processes prog !! sidx = (sidx', start, argDecls, decls)" by (metis prod.exhaust) with assms have p_def: "processes prog !! sidx = (sidx, start, argDecls, decls)" "IArray.list_of (processes prog) ! sidx = (sidx, start, argDecls, decls)" by simp_all with assms have "(sidx,start,argDecls,decls) \ set (IArray.list_of (processes prog))" by (force dest: nth_mem) with assms obtain s where s: "start = Index s" "s < IArray.length (states prog !! sidx)" unfolding program_inv_def by auto hence P_inv: "pState_inv prog \ pid = pidN, vars = lm.to_map ((STR ''_pid'', Var (VTBounded 0 (integer_of_nat pidN)) (integer_of_nat pidN)) # map modProcArg (zip argDecls (map (exprArith g p) args))), pc = s, channels = [], idx = sidx\" unfolding pState_inv_def using assms[unfolded program_inv_def] including integer.lifting apply (simp add: p_def) apply (intro lm_to_map_vardict_inv) apply auto apply (simp add: max_procs_def max_var_value_def) apply transfer' apply simp apply transfer' apply simp apply (simp add: min_var_value_def) apply transfer' apply simp apply (simp add: max_var_value_def max_procs_def) apply transfer' apply simp apply (drule set_zip_leftD) apply (force simp add: modProcArg_def split: procArg.splits if_splits intro!: image_eqI) apply (clarsimp simp add: modProcArg_def split: procArg.splits if_splits simp del: variable_inv.simps) apply (intro checkVarValue_Var) apply assumption done from p_def have "varDeclName ` set decls \ process_names (states prog !! sidx) (processes prog !! sidx)" by auto with \gState_inv prog g\ have F_inv: "\p. \ pState_inv prog p; sidx = pState.idx p; cl_inv (g,p) \ \ pState_inv prog (snd (foldl (\(g,p) d. mkVarChannel d (apsnd \ pState.vars_update) g p) (g,p) decls))" proof (induction decls arbitrary: g p) case (Cons d ds) hence decl: "varDeclName d \ process_names (states prog !! pState.idx p) (processes prog !! pState.idx p)" by simp obtain g' p' where new: "(g',p') = (mkVarChannel d (apsnd \ pState.vars_update) g p)" by (metis prod.exhaust) hence p': "p' = snd ..." and g': "g' = fst ..." by (metis snd_conv fst_conv)+ with Cons.prems have "pState_inv prog p'" apply (auto intro!: mkVarChannel_pState_inv) apply (simp add: pState_inv_def) apply (intro vardict_inv_updateI) apply simp apply (cases d) apply (force dest!: toVariable_name) apply (force dest!: toVariable_name) apply (intro toVariable_variable_inv') apply assumption+ done moreover from p' Cons.prems have "pState.idx p' = sidx" by (auto simp add: mkVarChannel_def mkChannels_def split: prod.split) moreover from new Cons.prems have "cl_inv (g',p')" by (auto intro!: mkVarChannel_cl_inv) moreover from g' Cons.prems have "gState_inv prog g'" by (auto intro!: mkVarChannel_gState_inv) moreover from Cons.prems have "varDeclName ` set ds \ process_names (states prog !! sidx) (processes prog !! sidx)" by simp ultimately have "pState_inv prog (snd (foldl (\(g,p) d. mkVarChannel d (apsnd \ pState.vars_update) g p) (g',p') ds))" using Cons.IH[of p' g'] by (simp add: o_def) with new show ?case by (simp add: o_def) qed simp show ?thesis by (auto simp add: p_def s cl_inv_def intro: F_inv[OF P_inv]) (blast intro: emptyProc_pState_inv assms) qed lemma mkProc_cl_inv: assumes "cl_inv (g,p)" shows "cl_inv (mkProc g p name args (processes prog !! sidx) pidN)" proof - note IArray.sub_def [simp del] obtain sidx' start argDecls decls where [simp]: "processes prog !! sidx = (sidx', start, argDecls, decls)" by (metis prod.exhaust) have P_inv: "\s v. cl_inv (g, \pid = pidN, vars = v, pc = s, channels = [], idx = sidx' \)" by (simp add: cl_inv_def) have "\p. cl_inv(g,p) \ cl_inv (foldl (\(g,p) d. mkVarChannel d (apsnd \ pState.vars_update) g p) (g,p) decls)" proof (induction decls arbitrary: g p) case (Cons d ds) obtain g' p' where new: "(g',p') = (mkVarChannel d (apsnd \ pState.vars_update) g p)" by (metis prod.exhaust) with Cons.prems have "cl_inv (g',p')" by (auto intro!: mkVarChannel_cl_inv) from Cons.IH[OF this] new show ?case by (simp add: o_def) qed simp from this[OF P_inv] show ?thesis by auto qed declare mkProc.simps[simp del] definition runProc :: "String.literal \ expr list \ program \ 'a gState_scheme \ pState \ 'a gState_scheme * pState" where "runProc name args prog g p = ( if length (procs g) \ max_procs then abort STR ''Too many processes'' (\_. (g,p)) else let pid = length (procs g) + 1 in case lm.lookup name (proc_data prog) of None \ abortv STR ''No such process: '' name (\_. (g,p)) | Some proc_idx \ let (g', proc) = mkProc g p name args (processes prog !! proc_idx) pid in (g'\procs := procs g @ [proc]\, p))" lemma runProc_gState_progress_rel: assumes "program_inv prog" and "gState_inv prog g" and "pState_inv prog p" and "cl_inv (g,p)" shows "(g, fst (runProc name args prog g p)) \ gState_progress_rel prog" proof (cases "length (procs g) < max_procs") note IArray.sub_def [simp del] case True thus ?thesis proof (cases "lm.lookup name (proc_data prog)") case (Some proc_idx) hence *: "proc_idx < IArray.length (processes prog)" "fst (processes prog !! proc_idx) = proc_idx" using assms by (simp_all add: lm.correct program_inv_def) obtain g' p' where new: "(g',p') = mkProc g p name args (processes prog !! proc_idx) (length (procs g) + 1)" by (metis prod.exhaust) hence g': "g' = fst ..." and p': "p' = snd ..." by (metis snd_conv fst_conv)+ from assms g' have "(g, g') \ gState_progress_rel prog " by (auto intro!: mkProc_gState_progress_rel) moreover from * assms True p' have "pState_inv prog p'" by (auto intro!: mkProc_pState_inv) moreover from assms new have "cl_inv (g',p')" by (auto intro!: mkProc_cl_inv) ultimately show ?thesis using True Some new assms unfolding runProc_def gState_progress_rel_def by (clarsimp split: prod.split) (auto simp add: gState_inv_def cl_inv_def) next case None with assms show ?thesis by (auto simp add: runProc_def) qed next case False with assms show ?thesis by (auto simp add: runProc_def) qed lemmas runProc_gState_inv = runProc_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma runProc_pState_id: "snd (runProc name args prog g p) = p" unfolding runProc_def by (auto split: if_splits split: option.split prod.split) lemma runProc_pState_inv: assumes "pState_inv prog p" shows "pState_inv prog (snd (runProc name args prog g p))" by (simp add: assms runProc_pState_id) lemma runProc_cl_inv: assumes "program_inv prog" and "gState_inv prog g" and "pState_inv prog p" and "cl_inv (g,p)" shows "cl_inv (runProc name args prog g p)" proof - obtain g' p' where *: "runProc name args prog g p = (g',p')" by (metis prod.exhaust) with runProc_gState_progress_rel[OF assms, of name args] have "length (channels g) \ length (channels g')" by (simp add: gState_progress_rel_def) moreover from * runProc_pState_id have "p' = p" by (metis snd_conv) ultimately show ?thesis by (metis \cl_inv (g,p)\ * cl_inv_trans) qed subsection \AST to edges\ type_synonym ast = "AST.module list" text \In this section, the AST is translated into the transition system.\ text \ Handling atomic blocks is non-trivial. Therefore, we do this in an extra pass: @{term lp} and @{term hp} are the positions of the start and the end of the atomic block. Every edge pointing into this range is therefore marked as @{term Atomic}. If they are pointing somewhere else, they are set to @{term InAtomic}, meaning: they start \emph{in} an atomic block, but leave it afterwards. \ definition atomize :: "nat \ nat \ edge list \ edge list" where "atomize lp hp es = fold (\e es. let e' = case target e of LabelJump _ None \ \ \Labels are checked again later on, when they\ \ \are going to be resolved. Hence it is safe to say\ \ \\atomic\ here, especially as the later algorithm\ \ \relies on targets in atomic blocks to be marked as such.\ e\ atomic := InAtomic \ | LabelJump _ (Some via) \ if lp \ via \ hp \ via then e\ atomic := Atomic \ else e\ atomic := InAtomic \ | Index p' \ if lp \ p' \ hp \ p' then e\ atomic := Atomic \ else e\ atomic := InAtomic \ in e'#es) es []" fun skip \ \No-(edge)\ where "skip (lbls, pri, pos, nxt, _) = ([[\cond = ECExpr (ExprConst 1), effect = EEId, target = nxt, prio = pri, atomic = NonAtomic\]], Index pos, lbls)" text \ The AST is walked backwards. This allows to know the next state directly. Parameters used: \begin{description} \item[lbls] Map of Labels \item[pri] Current priority \item[pos] Current position in the array \item[nxt] Next state \item[onxt] Previous 'next state' (where to jump after a 'do') \item[inBlock] Needed for certain constructs to calculate the layout of the array \end{description} \ fun stepToState :: "step \ (labels * integer * nat * edgeIndex * edgeIndex option * bool) \ edge list list * edgeIndex * labels" and stmntToState :: "stmnt \ (labels * integer * nat * edgeIndex * edgeIndex option * bool) \ edge list list * edgeIndex * labels" where "stepToState (StepStmnt s None) data = stmntToState s data" | "stepToState (StepStmnt s (Some u)) (lbls, pri, pos, nxt, onxt, _) = ( let \ \the \unless\ part\ (ues,_,lbls') = stmntToState u (lbls, pri, pos, nxt, onxt, True); u = last ues; ues = butlast ues; pos' = pos + length ues; \ \find minimal current priority\ pri = min_prio u pri; \ \the guarded part --\ \ \priority is decreased, because there is now a new unless part with\ \ \higher prio\ (ses,spos,lbls'') = stmntToState s (lbls', pri - 1, pos', nxt, onxt, False); \ \add an edge to the unless part for each generated state\ ses = map (List.append u) ses in (ues@ses,spos,lbls''))" | "stepToState (StepDecl decls) (lbls, pri, pos, nxt, onxt, _) = ( let edgeF = \d (lbls,pri,pos,nxt,_). ([[\cond = ECTrue, effect = EEDecl d, target = nxt, prio = pri, atomic = NonAtomic\]], Index pos, lbls) in step_fold edgeF decls lbls pri pos nxt onxt False)" | "stepToState StepSkip (lbls,_,_,nxt,_) = ([],nxt,lbls)" | "stmntToState (StmntAtomic steps) (lbls, pri, pos, nxt, onxt, inBlock) = ( let (es,pos',lbls') = step_fold stepToState steps lbls pri pos nxt onxt inBlock in let es' = map (atomize pos (pos + length es)) es in (es', pos', lbls'))" | "stmntToState (StmntLabeled l s) (lbls, pri, pos, d) = ( let (es, pos', lbls) = stmntToState s (lbls, pri, pos, d); \ \We don't resolve goto-chains. If the labeled stmnt returns only a jump,\ \ \use this goto state.\ lpos = case pos' of Index p \ p | _ \ pos; lbls' = add_label l lbls lpos in (es, pos', lbls'))" | "stmntToState (StmntDo stepss) (lbls, pri, pos, nxt, onxt, inBlock) = ( let \ \construct the different branches\ \ \\nxt\ in those branches points current pos (it is a loop after all)\ \ \\onxt\ then is the current \nxt\ (needed for break, f.ex.)\ (_,_,lbls,es,is) = step_foldL stepToState stepss lbls pri (pos+1) (Index pos) (Some nxt); \ \put the branch starting points (\is\) into the array\ es' = concat is # es in if inBlock then \ \inside another DO or IF or UNLESS\ \ \\\\ append branches again, so they can be consumed\ (es' @ [concat is], Index pos, lbls) else (es', Index pos, lbls) )" | "stmntToState (StmntIf stepss) (lbls, pri, pos, nxt, onxt, _) = ( let (pos,_,lbls,es,is) = step_foldL stepToState stepss lbls pri pos nxt onxt in (es @ [concat is], Index pos, lbls))" | "stmntToState (StmntSeq steps) (lbls, pri, pos, nxt, onxt, inBlock) = step_fold stepToState steps lbls pri pos nxt onxt inBlock" | "stmntToState (StmntAssign v e) (lbls, pri, pos, nxt, _) = ([[\cond = ECTrue, effect = EEAssign v e, target = nxt, prio = pri, atomic = NonAtomic\]], Index pos, lbls)" | "stmntToState (StmntAssert e) (lbls, pri, pos, nxt, _) = ([[\cond = ECTrue, effect = EEAssert e, target = nxt, prio = pri, atomic = NonAtomic\]], Index pos, lbls)" | "stmntToState (StmntCond e) (lbls, pri, pos, nxt, _) = ([[\cond = ECExpr e, effect = EEId, target = nxt, prio = pri, atomic = NonAtomic\]], Index pos, lbls)" | "stmntToState StmntElse (lbls, pri, pos, nxt, _) = ([[\cond = ECElse, effect = EEId, target = nxt, prio = pri, atomic = NonAtomic \]], Index pos, lbls)" | "stmntToState StmntBreak (lbls,pri,_,_,Some onxt,_) = ([[\cond = ECTrue, effect = EEGoto, target = onxt, prio = pri, atomic = NonAtomic \]], onxt, lbls)" | "stmntToState StmntBreak (_,_,_,_,None,_) = abort STR ''Misplaced break'' (\_. ([],Index 0,lm.empty()))" | "stmntToState (StmntRun n args) (lbls, pri, pos, nxt, onxt, _) = ([[\cond = ECRun n, effect = EERun n args, target = nxt, prio = pri, atomic = NonAtomic \]], Index pos,lbls)" | "stmntToState (StmntGoTo l) (lbls, pri, pos, _) = ([[\cond = ECTrue, effect = EEGoto, target = LabelJump l None, prio = pri, atomic = NonAtomic \]], LabelJump l (Some pos), lbls)" | "stmntToState (StmntSend v e srt) (lbls, pri, pos, nxt, _) = ([[\cond = ECSend v, effect = EESend v e srt, target = nxt, prio = pri, atomic = NonAtomic \]], Index pos, lbls)" | "stmntToState (StmntRecv v r srt rem) (lbls, pri, pos, nxt, _) = ([[\cond = ECRecv v r srt, effect = EERecv v r srt rem, target = nxt, prio = pri, atomic = NonAtomic \]], Index pos, lbls)" | "stmntToState StmntSkip d = skip d" subsubsection \Setup\ definition endState :: "edge list" where \ \An extra state added to each process marking its end.\ "endState = [\ cond = ECFalse, effect = EEEnd, target = Index 0, prio = 0, atomic = NonAtomic\]" definition resolveLabel :: "String.literal \ labels \ nat" where "resolveLabel l lbls = ( case lm.lookup l lbls of None \ abortv STR ''Unresolved label: '' l (\_. 0) | Some pos \ pos)" primrec resolveLabels :: "edge list list \ labels \ edge list \ edge list" where "resolveLabels _ _ [] = []" | "resolveLabels edges lbls (e#es) = ( let check_atomic = \pos. fold (\e a. a \ inAtomic e) (edges ! pos) True in case target e of Index _ \ e | LabelJump l None \ let pos = resolveLabel l lbls in e\target := Index pos, atomic := if inAtomic e then if check_atomic pos then Atomic else InAtomic else NonAtomic \ | LabelJump l (Some via) \ let pos = resolveLabel l lbls in e\target := Index pos, \ \NB: \isAtomic\ instead of \inAtomic\, cf \atomize()\\ atomic := if isAtomic e then if check_atomic pos \ check_atomic via then Atomic else InAtomic else atomic e \ ) # (resolveLabels edges lbls es)" definition calculatePrios :: "edge list list \ (integer * edge list) list" where "calculatePrios ess = map (\es. (min_prio es 0, es)) ess" definition toStates :: "step list \ states * edgeIndex * labels" where "toStates steps = ( let (states,pos,lbls) = step_fold stepToState steps (lm.empty()) 0 1 (Index 0) None False; pos = (case pos of Index _ \ pos | LabelJump l _ \ Index (resolveLabel l lbls)); states = endState # states; states = map (resolveLabels states lbls) states; states = calculatePrios states in case pos of Index s \ if s < length states then (IArray states, pos, lbls) else abort STR ''Start index out of bounds'' (\_. (IArray states, Index 0, lbls)))" lemma toStates_inv: assumes "toStates steps = (ss,start,lbls)" shows "\s. start = Index s \ s < IArray.length ss" and "IArray.length ss > 0" using assms unfolding toStates_def calculatePrios_def by (auto split: prod.splits edgeIndex.splits if_splits) (* returns: states * is_active * name * labels * process *) primrec toProcess :: "nat \ proc \ states * nat * String.literal * (labels * process)" where "toProcess sidx (ProcType act name args decls steps) = ( let (states, start, lbls) = toStates steps; act = (case act of None \ 0 | Some None \ 1 | Some (Some x) \ nat_of_integer x) in (states, act, name, lbls, sidx, start, args, decls))" | "toProcess sidx (Init decls steps) = ( let (states, start, lbls) = toStates steps in (states, 1, STR '':init:'', lbls, sidx, start, [], decls))" lemma toProcess_sidx: "toProcess sidx p = (ss,a,n,l,idx,r) \ idx = sidx" by (cases p) (simp_all split: prod.splits) lemma toProcess_states_nonempty: "toProcess sidx p = (ss,a,n,l,idx,r) \ IArray.length ss > 0" by (cases p) (force split: prod.splits dest: toStates_inv(2))+ lemma toProcess_start: "toProcess sidx p = (ss,a,n,l,idx,start,r) \ \s. start = Index s \ s < IArray.length ss" by (cases p) (force split: prod.splits dest: toStates_inv(1))+ lemma toProcess_startE: assumes "toProcess sidx p = (ss,a,n,l,idx,start,r)" obtains s where "start = Index s" "s < IArray.length ss" using toProcess_start[OF assms] by blast text \ The main construction function. Takes an AST and returns an initial state, and the program (= transition system). \ definition setUp :: "ast \ program \ gState" where "setUp ast = ( let (decls, procs, _) = preprocess ast; assertVar = Var (VTBounded 0 1) 0; pre_procs = map (case_prod toProcess) (List.enumerate 1 procs); procs = IArray ((0, Index 0, [], []) # map (\(_,_,_,_,p). p) pre_procs); labels = IArray (lm.empty() # map (\(_,_,_,l,_). l) pre_procs); states = IArray (IArray [(0,[])] # map (\(s,_). s) pre_procs); names = IArray (STR ''invalid'' # map (\(_,_,n,_). n) pre_procs); proc_data = lm.to_map (map (\(_,_,n,_,idx,_). (n,idx)) pre_procs); prog = \ processes = procs, labels = labels, states = states, proc_names = names, proc_data = proc_data \; g = \ vars = lm.sng (STR ''__assert__'') assertVar, channels = [InvChannel], timeout = False, procs = [] \; g' = foldl (\g d. fst (mkVarChannel d (apfst \ gState.vars_update) g emptyProc) ) g decls; g'' = foldl (\g (_,a,name,_). foldl (\g name. fst (runProc name [] prog g emptyProc) ) g (replicate a name) ) g' pre_procs in (prog, g''))" lemma setUp_program_inv': "program_inv (fst (setUp ast))" proof (rule program_invI, goal_cases) case 1 show ?case by (simp add: setUp_def split: prod.split) next case 2 show ?case by (simp add: setUp_def split: prod.split) next case 3 thus ?case by (auto simp add: setUp_def o_def split: prod.splits dest!: toProcess_states_nonempty) next case 4 thus ?case unfolding setUp_def by (auto simp add: lm.correct o_def in_set_enumerate_eq nth_enumerate_eq dest!: subsetD[OF Misc.ran_map_of] toProcess_sidx split: prod.splits) (* TODO: Change name Misc.ran_map_of \ ran_map_of_ss, as it collides with AList.ran_map_of *) next case 5 thus ?case apply (auto simp add: setUp_def o_def split: prod.splits) apply (frule toProcess_sidx) apply (frule toProcess_start) apply (auto simp: in_set_enumerate_eq nth_enumerate_eq) done qed lemma setUp_program_inv: assumes "setUp ast = (prog,g)" shows "program_inv prog" using assms setUp_program_inv' by (metis fst_conv) lemma setUp_gState_inv: assumes "setUp ast = (prog, g)" shows "gState_inv prog g" proof - from assms have p_INV: "program_inv prog" by (fact setUp_program_inv) { fix prog :: program assume *: "program_inv prog" let ?g = "\ vars = lm.sng (STR ''__assert__'') (Var (VTBounded 0 1) 0), channels = [InvChannel], timeout = False, procs = [] \" have g1: "gState_inv prog ?g" by (simp add: gState_inv_def max_channels_def lm_correct max_var_value_def) { fix g decls assume "gState_inv prog g" hence "gState_inv prog (foldl (\g d. fst (mkVarChannel d (apfst \ gState.vars_update) g emptyProc) ) g decls)" apply (rule foldl_rule) apply (intro mkVarChannel_gState_inv) apply simp apply (frule_tac g=\ in toVariable_variable_inv') apply assumption apply (auto simp add: gState_inv_def lm.correct) done } note g2 = this[OF g1] { fix g :: "'a gState_scheme" and pre_procs assume "gState_inv prog g" hence "gState_inv prog (foldl (\g (_,a,name,_). foldl (\g name. fst (runProc name [] prog g emptyProc) ) g (replicate a name) ) g pre_procs)" apply (rule foldl_rule) apply (clarsimp split: prod.splits) apply (rule foldl_rule) apply (auto intro!: runProc_gState_inv emptyProc_pState_inv *) done } note this[OF g2] } note g_INV = this from assms p_INV show ?thesis unfolding setUp_def by (auto split: prod.splits intro!: g_INV) qed subsection \Semantic Engine\ text \ After constructing the transition system, we are missing the final part: The successor function on this system. We use SPIN-nomenclature and call it \emph{semantic engine}. \ definition "assertVar \ VarRef True (STR ''__assert__'') None" subsubsection \Evaluation of Edges\ fun evalRecvArgs :: "recvArg list \ integer list \ gState\<^sub>I \ pState \ gState\<^sub>I * pState" where "evalRecvArgs [] [] g l = (g,l)" | "evalRecvArgs _ [] g l = abort STR ''Length mismatch on receiving.'' (\_. (g,l))" | "evalRecvArgs [] _ g l = abort STR ''Length mismatch on receiving.'' (\_. (g,l))" | "evalRecvArgs (r#rs) (v#vs) g l = ( let (g,l) = case r of RecvArgVar var \ setVar var v g l | _ \ (g,l) in evalRecvArgs rs vs g l)" primrec evalCond :: "edgeCond \ gState\<^sub>I \ pState \ bool" where "evalCond ECTrue _ _ \ True" | "evalCond ECFalse _ _ \ False" | "evalCond (ECExpr e) g l \ exprArith g l e \ 0" | "evalCond (ECRun _) g l \ length (procs g) < 255" | "evalCond ECElse g l \ gState\<^sub>I.else g" | "evalCond (ECSend v) g l \ withChannel v (\_ c. case c of Channel cap _ q \ integer_of_nat (length q) < cap | HSChannel _ \ True) g l" | "evalCond (ECRecv v rs srt) g l \ withChannel v (\i c. case c of HSChannel _ \ handshake g \ 0 \ recvArgsCheck g l rs (hsdata g) | _ \ pollCheck g l c rs srt) g l" fun evalHandshake :: "edgeCond \ nat \ gState\<^sub>I \ pState \ bool" where "evalHandshake (ECRecv v _ _) h g l \ h = 0 \ withChannel v (\i c. case c of HSChannel _ \ i = h | Channel _ _ _ \ False) g l" | "evalHandshake _ h _ _ \ h = 0" primrec evalEffect :: "edgeEffect \ program \ gState\<^sub>I \ pState \ gState\<^sub>I * pState" where "evalEffect EEEnd _ g l = (g,l)" | "evalEffect EEId _ g l = (g,l)" | "evalEffect EEGoto _ g l = (g,l)" | "evalEffect (EEAssign v e) _ g l = setVar v (exprArith g l e) g l" | "evalEffect (EEDecl d) _ g l = mkVarChannelProc d g l" | "evalEffect (EERun name args) prog g l = runProc name args prog g l" | "evalEffect (EEAssert e) _ g l = ( if exprArith g l e = 0 then setVar assertVar 1 g l else (g,l))" | "evalEffect (EESend v es srt) _ g l = withChannel v (\i c. let ab = \_. abort STR ''Length mismatch on sending.'' (\_. (g,l)); es = map (exprArith g l) es in if \ for_all (\x. x \ min_var_value \ x \ max_var_value) es then abort STR ''Invalid Channel'' (\_. (g,l)) else case c of Channel cap ts q \ if length ts \ length es \ \ (length q < max_array_size) then ab() else let q' = if \ srt then q@[es] else let q = map lexlist q; q' = insort (lexlist es) q in map unlex q'; g = gState.channels_update (\cs. cs[ i := Channel cap ts q' ]) g in (g,l) | HSChannel ts \ if length ts \ length es then ab() else (g\hsdata := es, handshake := i\, l) | InvChannel \ abort STR ''Trying to send on invalid channel'' (\_. (g,l)) ) g l" | "evalEffect (EERecv v rs srt rem) _ g l = withChannel v (\i c. case c of Channel cap ts qs \ if qs = [] then abort STR ''Recv from empty channel'' (\_. (g,l)) else let (q', qs') = if \ srt then (hd qs, tl qs) else apfst the (find_remove (recvArgsCheck g l rs) qs); (g,l) = evalRecvArgs rs q' g l; g = if rem then gState.channels_update (\cs. cs[ i := Channel cap ts qs']) g else g \ \messages are not removed -- so no need to update anything\ in (g,l) | HSChannel _ \ let (g,l) = evalRecvArgs rs (hsdata g) g l in let g = g\ handshake := 0, hsdata := [] \ in (g,l) | InvChannel \ abort STR ''Receiving on invalid channel'' (\_. (g,l)) ) g l" lemma statesDecls_effect: assumes "ef \ effect ` edgeSet ss" and "ef = EEDecl d" shows "d \ statesDecls ss" proof - from assms obtain e where "e \ edgeSet ss" "ef = effect e" by auto thus ?thesis using assms unfolding statesDecls_def by (auto simp add: edgeDecls_def intro!: bexI[where x = e] split: edgeEffect.split) qed lemma evalRecvArgs_pState_inv: assumes "pState_inv prog p" shows "pState_inv prog (snd (evalRecvArgs rargs xs g p))" using assms proof (induction rargs xs arbitrary: p g rule: list_induct2') case (4 r rs x xs) thus ?case proof (cases r) case (RecvArgVar v) obtain g' p' where new: "setVar v x g p = (g',p')" by (metis prod.exhaust) hence "p' = snd (setVar v x g p)" by simp with "4" have "pState_inv prog p'" by (auto intro!: setVar_pState_inv) from "4.IH"[OF this] RecvArgVar new show ?thesis by simp qed simp_all qed simp_all lemma evalRecvArgs_pState_inv': assumes "evalRecvArgs rargs xs g p = (g', p')" and "pState_inv prog p" shows "pState_inv prog p'" using assms evalRecvArgs_pState_inv by (metis snd_conv) lemma evalRecvArgs_gState_progress_rel: assumes "gState_inv prog g" shows "(g, fst (evalRecvArgs rargs xs g p)) \ gState_progress_rel prog" using assms proof (induction rargs xs arbitrary: p g rule: list_induct2') case (4 r rs x xs) thus ?case proof (cases r) case (RecvArgVar v) obtain g' p' where new: "setVar v x g p = (g',p')" by (metis prod.exhaust) hence "g' = fst (setVar v x g p)" by simp with "4" have "(g, g') \ gState_progress_rel prog" by (auto intro!: setVar_gState_progress_rel) also hence "gState_inv prog g'" by (rule gState_progress_rel_gState_invI2) note "4.IH"[OF this, of p'] finally show ?thesis using RecvArgVar new by simp qed simp_all qed simp_all lemmas evalRecvArgs_gState_inv = evalRecvArgs_gState_progress_rel[THEN gState_progress_rel_gState_invI2] lemma evalRecvArgs_cl_inv: assumes "cl_inv (g,p)" shows "cl_inv (evalRecvArgs rargs xs g p)" using assms proof (induction rargs xs arbitrary: p g rule: list_induct2') case (4 r rs x xs) thus ?case proof (cases r) case (RecvArgVar v) with 4 have "cl_inv (setVar v x g p)" by (auto intro!: setVar_cl_inv) with "4.IH" RecvArgVar show ?thesis by (auto split: prod.splits) qed simp_all qed simp_all lemma evalEffect_pState_inv: assumes "pState_inv prog p" and "gState_inv prog g" and "cl_inv (g,p)" and "e \ effect ` edgeSet (states prog !! pState.idx p)" shows "pState_inv prog (snd (evalEffect e prog g p))" using assms proof (cases e) case (EEDecl d) with assms have "d \ statesDecls (states prog !! pState.idx p)" using statesDecls_effect by simp with assms EEDecl show ?thesis by (auto simp del: IArray.sub_def intro!: mkVarChannelProc_pState_inv) next case (EESend c es srt) then obtain v where "ChanRef v = c" by (cases c) simp with EESend assms show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def split: channel.split) next case (EERecv c es srt) then obtain v where "ChanRef v = c" by (cases c) simp with EERecv assms show ?thesis by (cases v) (auto simp: withChannel'_def withVar'_def split: prod.splits channel.split dest: evalRecvArgs_pState_inv') qed (clarsimp_all intro!: setVar_pState_inv runProc_pState_inv) lemma evalEffect_gState_progress_rel: assumes "program_inv prog" and "gState_inv prog g" and "pState_inv prog p" and "cl_inv (g,p)" shows "(g, fst (evalEffect e prog g p)) \ gState_progress_rel prog" using assms proof (cases e) case EEAssert with assms show ?thesis by (auto intro: setVar_gState_progress_rel) next case EEAssign with assms show ?thesis by (auto intro: setVar_gState_progress_rel) next case EEDecl with assms show ?thesis by (auto intro: mkVarChannelProc_gState_progress_rel) next case EERun with assms show ?thesis by (auto intro: runProc_gState_progress_rel) next case (EESend c es srt) then obtain v where v: "c = ChanRef v" by (metis chanRef.exhaust) obtain idx where idx: "nat_of_integer (the (getVar v g p)) = idx" by blast note idx' = idx[symmetric, unfolded getVar_def] show ?thesis proof (cases "idx < length (gState.channels g)") case True note DEF = True EESend v idx' assms show ?thesis proof (cases "gState.channels g ! idx") case (Channel cap ts q) with True have "Channel cap ts q \ set (gState.channels g)" by (metis nth_mem) with assms have "channel_inv (Channel cap ts q)" by (auto simp add: gState_inv_def simp del: channel_inv.simps) with Channel DEF show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def for_all_def split: channel.split intro!: gState_progress_rel_channels_update) next case HSChannel with DEF show ?thesis by (cases v) (auto simp: withChannel'_def withVar'_def gState_progress_rel_def gState_inv_def split: channel.split) next case InvChannel with DEF show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def) qed next case False with EESend idx' v assms show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def) qed next case (EERecv c rs srt rem) then obtain v where v: "c = ChanRef v" by (metis chanRef.exhaust) obtain idx where idx: "nat_of_integer (the (getVar v g p)) = idx" by blast note idx' = idx[symmetric, unfolded getVar_def] show ?thesis proof (cases "idx < length (gState.channels g)") case True note DEF = True EERecv v idx' assms show ?thesis proof (cases "gState.channels g ! idx") note channel_inv.simps[simp del] case (Channel cap ts q) with True have "Channel cap ts q \ set (gState.channels g)" by (metis nth_mem) with assms have c_inv: "channel_inv (Channel cap ts q)" by (auto simp add: gState_inv_def simp del: channel_inv.simps) moreover obtain res q' where "apfst the (find_remove (recvArgsCheck g p rs) q) = (res, q')" by (metis prod.exhaust) moreover hence "q' = snd (find_remove (recvArgsCheck g p rs) q)" by (simp add: apfst_def map_prod_def split: prod.splits) with find_remove_subset find_remove_length have "set q' \ set q" "length q' \ length q" by (metis surjective_pairing)+ with c_inv have "channel_inv (Channel cap ts q')" by (auto simp add: channel_inv.simps) moreover { assume "q \ []" hence "set (tl q) \ set q" using tl_subset by auto with c_inv have "channel_inv (Channel cap ts (tl q))" by (auto simp add: channel_inv.simps) } moreover { fix res g' p' assume "evalRecvArgs rs res g p = (g',p')" with evalRecvArgs_gState_progress_rel assms have "(g,g') \ gState_progress_rel prog" by (metis fst_conv) hence "length (channels g) \ length (channels g')" by (simp add: gState_progress_rel_def) } ultimately show ?thesis using Channel DEF apply (cases v) apply (auto simp add: withChannel'_def withVar'_def for_all_def split: channel.split prod.split elim: fstE intro!: evalRecvArgs_gState_progress_rel gState_progress_rel_channels_update_step) apply force+ done next case HSChannel obtain g' p' where *: "evalRecvArgs rs (hsdata g) g p = (g',p')" by (metis prod.exhaust) with assms have "(g,g') \ gState_progress_rel prog" by (auto elim!: fstE intro!: evalRecvArgs_gState_progress_rel) also hence "gState_inv prog g'" by blast hence "(g',g'\handshake := 0, hsdata := []\) \ gState_progress_rel prog" by (auto simp add: gState_progress_rel_def gState_inv_def) finally have "(g,g'\handshake := 0, hsdata := []\) \ gState_progress_rel prog" . with DEF HSChannel * show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def for_all_def split: channel.split prod.split) next case InvChannel with DEF show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def) qed next case False with EERecv idx' v assms show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def) qed qed simp_all lemma evalEffect_cl_inv: assumes "cl_inv (g,p)" and "program_inv prog" and "gState_inv prog g" and "pState_inv prog p" shows "cl_inv (evalEffect e prog g p)" using assms proof (cases e) case EERun with assms show ?thesis by (force intro!: runProc_cl_inv) next case (EESend c es srt) then obtain v where "ChanRef v = c" by (cases c) simp with EESend assms show ?thesis by (cases v) (auto simp add: withChannel'_def withVar'_def split: channel.split intro!: cl_inv_channels_update) next case (EERecv c es srt) then obtain v where "ChanRef v = c" by (cases c) simp with EERecv assms show ?thesis apply (cases v) apply (auto simp add: withChannel'_def withVar'_def split: channel.split prod.split intro!: cl_inv_channels_update) apply (metis evalRecvArgs_cl_inv)+ done qed (simp_all add: setVar_cl_inv mkVarChannelProc_cl_inv) subsubsection \Executable edges\ text \ To find a successor global state, we first need to find all those edges which are executable (\ie the condition evaluates to true). \ type_synonym choices = "(edge * pState) list" \ \A choice is an executable edge and the process it belongs to.\ definition getChoices :: "gState\<^sub>I \ pState \ edge list \ choices" where "getChoices g p = foldl (\E e. if evalHandshake (cond e) (handshake g) g p \ evalCond (cond e) g p then (e,p)#E else E) []" lemma getChoices_sub_edges_fst: "fst ` set (getChoices g p es) \ set es" unfolding getChoices_def by (rule foldl_rule_aux) auto lemma getChoices_sub_edges: "(a,b) \ set (getChoices g p es) \ a \ set es" using getChoices_sub_edges_fst by force lemma getChoices_p_snd: "snd ` set (getChoices g p es) \ {p}" unfolding getChoices_def by (rule foldl_rule_aux) auto lemma getChoices_p: "(a,b) \ set (getChoices g p es) \ b = p" using getChoices_p_snd by force definition sort_by_pri where "sort_by_pri min_pri edges = foldl (\es e. let idx = nat_of_integer (abs (prio e)) in if idx > min_pri then abort STR ''Invalid priority'' (\_. es) else let ep = e # (es ! idx) in es[idx := ep] ) (replicate (min_pri + 1) []) edges" lemma sort_by_pri_edges': assumes "set edges \ A" shows "set (sort_by_pri min_pri edges) \ {xs. set xs \ A}" using assms unfolding sort_by_pri_def apply (rule_tac I="\\ _. (\x \ set \. set x \ A) \ length \ = min_pri + 1" in foldl_rule_aux_P) apply simp apply (force dest!: subsetD[OF set_update_subset_insert] split: if_splits) apply force done lemma sort_by_pri_edges: assumes "set edges \ A" and "es \ set (sort_by_pri min_pri edges)" shows "set es \ A" using sort_by_pri_edges'[OF assms(1)] assms by auto lemma sort_by_pri_length: "length (sort_by_pri min_pri edges) = min_pri + 1" unfolding sort_by_pri_def by (rule foldl_rule_aux_P [where I="\\ _. length \ = min_pri + 1"]) simp_all definition executable :: "states iarray \ gState\<^sub>I \ choices nres" \ \Find all executable edges\ where "executable ss g = ( let procs = procs g in nfoldli procs (\_. True) (\p E. if (exclusive g = 0 \ exclusive g = pid p) then do { let (min_pri, edges) = (ss !! pState.idx p) !! pc p; ASSERT(set edges \ edgeSet (ss !! pState.idx p)); (E',_,_) \ if min_pri = 0 then do { WHILE\<^sub>T (\(E,brk,_). E = [] \ brk = 0) (\ (_, _, ELSE). do { let g = g\gState\<^sub>I.else := ELSE\; E = getChoices g p edges in if E = [] then ( if \ ELSE then RETURN (E, 0::nat, True) else RETURN (E, 1, False)) else RETURN (E, 1, ELSE) }) ([], 0::nat, False) } else do { let min_pri = nat_of_integer (abs min_pri); let pri_edges = sort_by_pri min_pri edges; ASSERT (\es \ set pri_edges. set es \ edgeSet (ss !! pState.idx p)); let pri_edges = IArray pri_edges; WHILE\<^sub>T (\(E,pri,_). E = [] \ pri \ min_pri) (\(_, pri, ELSE). do { let es = pri_edges !! pri; let g = g\gState\<^sub>I.else := ELSE\; let E = getChoices g p es; if E = [] then ( if \ ELSE then RETURN (E,pri,True) else RETURN (E, pri + 1, False)) else RETURN (E, pri, ELSE) }) ([], 0, False) }; RETURN (E'@E) } else RETURN E ) [] )" definition "while_rel1 = measure (\x. if x = [] then 1 else 0) <*lex*> measure (\x. if x = 0 then 1 else 0) <*lex*> measure (\x. if \ x then 1 else 0)" lemma wf_while_rel1: "wf while_rel1" unfolding while_rel1_def by auto definition "while_rel2 mp = measure (\x. if x = [] then 1 else 0) <*lex*> measure (\x. (mp + 1) - x) <*lex*> measure (\x. if \ x then 1 else 0)" lemma wf_while_rel2: "wf (while_rel2 mp)" unfolding while_rel2_def by auto lemma executable_edgeSet: assumes "gState_inv prog g" and "program_inv prog" and "ss = states prog" shows "executable ss g \ SPEC (\cs. \(e,p) \ set cs. e \ edgeSet (ss !! pState.idx p) \ pState_inv prog p \ cl_inv (g,p))" unfolding executable_def using assms apply (refine_rcg refine_vcg nfoldli_rule[where I="\_ _ cs. \(e,p) \ set cs. e \ edgeSet (ss !! pState.idx p) \ pState_inv prog p \ cl_inv (g,p)"]) apply simp apply (rename_tac p l1 l2 \ e p') apply (subgoal_tac "pState_inv prog p") apply (clarsimp simp add: edgeSet_def pState_inv_def)[] apply (rule_tac x="IArray.list_of (IArray.list_of (states prog) ! pState.idx p) ! pc p" in bexI) apply simp apply (force intro!: nth_mem) apply (force simp add: gState_inv_def) apply (refine_rcg refine_vcg wf_while_rel1 WHILET_rule[where I="\(cs,_,_). \(e,p) \ set cs. e \ edgeSet (ss !! pState.idx p) \ pState_inv prog p \ cl_inv (g,p)" and R=while_rel1]) apply (vc_solve solve: asm_rl simp add: while_rel1_def) apply (frule getChoices_p) using getChoices_sub_edges apply (force simp add: gState_inv_def) using sort_by_pri_edges apply fastforce apply (rename_tac min_pri pri_edges) apply (rule_tac I="\(cs,_,_). \(e,p) \ set cs. e \ edgeSet (ss !! pState.idx p) \ pState_inv prog p \ cl_inv (g,p)" and R="while_rel2 (nat_of_integer (abs min_pri))" in WHILET_rule) apply (simp add: wf_while_rel2) apply simp apply (refine_rcg refine_vcg) apply (clarsimp_all split: prod.split simp add: while_rel2_def) apply (metis diff_less_mono lessI) apply (rename_tac idx' else e p') apply (frule getChoices_p) apply (frule getChoices_sub_edges) apply (subgoal_tac "sort_by_pri (nat_of_integer \min_pri\) pri_edges ! idx' \ set (sort_by_pri (nat_of_integer \min_pri\) pri_edges)") apply (auto simp add: assms gState_inv_def)[] apply (force simp add: sort_by_pri_length) apply (auto simp add: assms) done lemma executable_edgeSet': assumes "gState_inv prog g" and "program_inv prog" shows "executable (states prog) g \ SPEC (\cs. \(e,p) \ set cs. e \ edgeSet ((states prog) !! pState.idx p) \ pState_inv prog p \ cl_inv(g,p))" using executable_edgeSet[where ss = "states prog"] assms by simp schematic_goal executable_refine: "RETURN (?ex s g) \ executable s g" unfolding executable_def by (refine_transfer) concrete_definition executable_impl for s g uses executable_refine subsubsection \Successor calculation\ function to\<^sub>I where "to\<^sub>I \ gState.vars = v, channels = ch, timeout = t, procs = p \ = \ gState.vars = v, channels = ch, timeout = False, procs = p, handshake = 0, hsdata = [], exclusive = 0, gState\<^sub>I.else = False \" by (metis gState.cases) (metis gState.ext_inject) termination by lexicographic_order function "from\<^sub>I" where "from\<^sub>I \ gState.vars = v, channels = ch, timeout = t, procs = p, \ = m \ = \ gState.vars = v, channels = ch, timeout = t, procs = p \" by (metis gState.surjective) (metis gState.ext_inject) termination by lexicographic_order function reset\<^sub>I where "reset\<^sub>I \ gState.vars = v, channels = ch, timeout = t, procs = p, handshake = hs, hsdata = hsd, exclusive = _, gState\<^sub>I.else = _ \ = \ gState.vars = v, channels = ch, timeout = False, procs = p, handshake = 0, hsdata = if hs \ 0 then hsd else [], exclusive = 0, gState\<^sub>I.else = False \" by (metis (full_types) gState\<^sub>I.surjective unit.exhaust) (metis gState.select_convs gState\<^sub>I.select_convs) termination by lexicographic_order lemma gState_inv_to\<^sub>I: "gState_inv prog g = gState_inv prog (to\<^sub>I g)" by (cases g, simp add: gState_inv_def cl_inv_def) lemma gState_inv_from\<^sub>I: "gState_inv prog g = gState_inv prog (from\<^sub>I g)" by (cases g, simp add: gState_inv_def cl_inv_def) lemma gState_inv_reset\<^sub>I: "gState_inv prog g = gState_inv prog (reset\<^sub>I g)" by (cases g, simp add: gState_inv_def cl_inv_def) lemmas gState_inv_I_simps = gState_inv_to\<^sub>I gState_inv_from\<^sub>I gState_inv_reset\<^sub>I definition removeProcs \ \Remove ended processes, if there is no running one with a higher pid.\ where "removeProcs ps = foldr (\p (dead,sd,ps,dcs). if dead \ pc p = 0 then (True, True, ps, pState.channels p @ dcs) else (False, sd, p#ps, dcs)) ps (True, False, [], [])" lemma removeProcs_subset': "set (fst (snd (snd (removeProcs ps)))) \ set ps" unfolding removeProcs_def apply (subst foldr_conv_foldl) apply (rule foldl_rule_aux_P[where I="\(_,_,ps',_) _. set ps' \ set ps"]) apply simp apply (clarsimp split: if_splits) apply force apply (rename_tac p) apply (case_tac "p = x") apply (subst set_rev[symmetric]) apply simp apply force apply force done lemma removeProcs_length': "length (fst (snd (snd (removeProcs ps)))) \ length ps" unfolding removeProcs_def apply (subst foldr_conv_foldl) apply (rule foldl_rule_aux_P[where I="\(_,_,ps',_) ps''. length ps' + length ps'' \ length ps"]) apply (auto split: if_splits) done lemma removeProcs_subset: "removeProcs ps = (dead,sd,ps',dcs) \ set ps' \ set ps" using removeProcs_subset' by (metis fst_conv snd_conv) lemma removeProcs_length: "removeProcs ps = (dead,sd,ps',dcs) \ length ps' \ length ps" using removeProcs_length' by (metis fst_conv snd_conv) definition cleanChans :: "integer list \ channels \ channels" \ \Mark channels of closed processes as invalid.\ where "cleanChans dchans cs = snd (foldl (\(i,cs) c. if List.member dchans i then (i + 1, cs@[InvChannel]) else (i + 1, cs@[c])) (0, []) cs)" lemma cleanChans_channel_inv: assumes "set cs \ Collect channel_inv" shows "set (cleanChans dchans cs) \ Collect channel_inv" using assms unfolding cleanChans_def apply (rule_tac foldl_rule_aux) apply (force split: if_splits)+ done lemma cleanChans_length: "length (cleanChans dchans cs) = length cs" unfolding cleanChans_def by (rule foldl_rule_aux_P[where I="\(_,cs') cs''. length cs' + length cs'' = length cs"]) (force split: if_splits)+ definition checkDeadProcs :: "'a gState_scheme \ 'a gState_scheme" where "checkDeadProcs g = ( let (_, soDied, procs, dchans) = removeProcs (procs g) in if soDied then g\ procs := procs, channels := cleanChans dchans (channels g)\ else g)" lemma checkDeadProcs_gState_progress_rel: assumes "gState_inv prog g" shows "(g, checkDeadProcs g) \ gState_progress_rel prog" using assms cleanChans_length[where cs="channels g"] cleanChans_channel_inv[where cs="channels g"] unfolding checkDeadProcs_def apply (intro gState_progress_relI) apply assumption apply (clarsimp split: if_splits prod.splits) apply (frule removeProcs_length) apply (frule removeProcs_subset) apply (auto simp add: gState_inv_def cl_inv_def)[] apply (clarsimp split: if_splits prod.splits)+ done lemma gState_progress_rel_exclusive: "(g, g') \ gState_progress_rel prog \ (g, g'\exclusive := p\) \ gState_progress_rel prog" by (simp add: gState_progress_rel_def gState_inv_def cl_inv_def) definition applyEdge :: "program \ edge \ pState \ gState\<^sub>I \ gState\<^sub>I nres" where "applyEdge prog e p g = do { let (g',p') = evalEffect (effect e) prog g p; ASSERT ((g,g') \ gState_progress_rel prog); ASSERT (pState_inv prog p'); ASSERT (cl_inv (g',p')); let p'' = (case target e of Index t \ if t < IArray.length (states prog !! pState.idx p') then p'\pc := t\ else abort STR ''Edge target out of bounds'' (\_. p') | _ \ abort STR ''Edge target not Index'' (\_. p')); ASSERT (pState_inv prog p''); let g'' = g'\procs := list_update (procs g') (pid p'' - 1) p''\; ASSERT ((g',g'') \ gState_progress_rel prog); let g''' = (if isAtomic e \ handshake g'' = 0 then g''\ exclusive := pid p'' \ else g''); ASSERT ((g',g''') \ gState_progress_rel prog); let g\<^sub>f = (if pc p'' = 0 then checkDeadProcs g''' else g'''); ASSERT ((g''',g\<^sub>f) \ gState_progress_rel prog); RETURN g\<^sub>f }" lemma applyEdge_gState_progress_rel: assumes "program_inv prog" and "gState_inv prog g" and "pState_inv prog p" and "cl_inv (g,p)" and "e \ edgeSet (states prog !! pState.idx p)" shows "applyEdge prog e p g \ SPEC (\g'. (g,g') \ gState_progress_rel prog)" using assms unfolding applyEdge_def apply (intro refine_vcg) apply (force elim: fstE intro!: evalEffect_gState_progress_rel) apply (force elim: sndE intro!: evalEffect_pState_inv) apply (drule subst) apply (rule evalEffect_cl_inv) apply assumption+ apply (cases "target e") apply (clarsimp simp add: pState_inv_def) apply simp apply (frule gState_progress_rel_gState_invI2) apply (cases "target e") apply (clarsimp split: if_splits) apply (intro gState_progress_relI) apply assumption apply (auto simp add: gState_inv_def cl_inv_def dest!: subsetD[OF set_update_subset_insert])[] apply (simp add: cl_inv_def) apply simp apply (intro gState_progress_relI) apply assumption apply (auto simp add: gState_inv_def cl_inv_def dest!: subsetD[OF set_update_subset_insert])[] apply (simp add: cl_inv_def) apply simp apply (clarsimp split: if_splits) apply (intro gState_progress_relI) apply assumption apply (auto simp add: gState_inv_def cl_inv_def dest!: subsetD[OF set_update_subset_insert])[] apply (simp add: cl_inv_def) apply simp apply (auto split: if_splits intro!: gState_progress_rel_exclusive)[] apply (force intro!: checkDeadProcs_gState_progress_rel) apply (blast intro!: gState_progress_rel_trans) done schematic_goal applyEdge_refine: "RETURN (?ae prog e p g) \ applyEdge prog e p g" unfolding applyEdge_def by (refine_transfer) concrete_definition applyEdge_impl for e p g uses applyEdge_refine definition nexts :: "program \ gState \ gState ls nres" \ \The successor function\ where "nexts prog g = ( let f = from\<^sub>I; g = to\<^sub>I g in REC (\D g. do { E \ executable (states prog) g; if E = [] then if handshake g \ 0 then \ \HS not possible -- remove current step\ RETURN (ls.empty()) else if exclusive g \ 0 then \ \Atomic blocks -- just return current state\ RETURN (ls.sng (f g)) else if \ timeout g then \ \Set timeout\ D (g\timeout := True\) else \ \If all else fails: stutter\ RETURN (ls.sng (f (reset\<^sub>I g))) else \ \Setting the internal variables (exclusive, handshake, ...) to 0\ \ \is safe -- they are either set by the edges, or not thought\ \ \to be used outside executable.\ let g = reset\<^sub>I g in nfoldli E (\_. True) (\(e,p) G. applyEdge prog e p g \ (\ g'. if handshake g' \ 0 \ isAtomic e then do { G\<^sub>R \ D g'; if ls.isEmpty G\<^sub>R \ handshake g' = 0 then \ \this only happens if the next step is a handshake, which fails\ \ \hence we stay at the current state\ RETURN (ls.ins (f g') G) else RETURN (ls.union G\<^sub>R G) } else RETURN (ls.ins (f g') G))) (ls.empty()) }) g \ (\G. if ls.isEmpty G then RETURN (ls.sng (f g)) else RETURN G) )" lemma gState_progress_rel_intros: "(to\<^sub>I g, gI') \ gState_progress_rel prog \ (g, from\<^sub>I gI') \ gState_progress_rel prog" "(gI, gI') \ gState_progress_rel prog \ (gI, reset\<^sub>I gI') \ gState_progress_rel prog" "(to\<^sub>I g, gI') \ gState_progress_rel prog \ (to\<^sub>I g, gI'\timeout := t\) \ gState_progress_rel prog" unfolding gState_progress_rel_def by (cases g, cases gI', cases gI, force simp add: gState_inv_def cl_inv_def)+ lemma gState_progress_rel_step_intros: "(to\<^sub>I g, g') \ gState_progress_rel prog \ (reset\<^sub>I g', g'') \ gState_progress_rel prog \ (g, from\<^sub>I g'') \ gState_progress_rel prog" "(to\<^sub>I g, g') \ gState_progress_rel prog \ (reset\<^sub>I g', g'') \ gState_progress_rel prog \ (to\<^sub>I g, g'') \ gState_progress_rel prog" unfolding gState_progress_rel_def by (cases g, cases g', cases g'', force simp add: gState_inv_def cl_inv_def)+ lemma cl_inv_reset\<^sub>I: "cl_inv(g,p) \ cl_inv(reset\<^sub>I g, p)" by (cases g) (simp add: cl_inv_def) lemmas refine_helpers = gState_progress_rel_intros gState_progress_rel_step_intros cl_inv_reset\<^sub>I lemma nexts_SPEC: assumes "gState_inv prog g" and "program_inv prog" shows "nexts prog g \ SPEC (\gs. \g' \ ls.\ gs. (g,g') \ gState_progress_rel prog)" using assms unfolding nexts_def apply (refine_rcg refine_vcg REC_rule[where pre="\g'. (to\<^sub>I g,g') \ gState_progress_rel prog"]) apply (simp add: gState_inv_to\<^sub>I) apply (rule order_trans[OF executable_edgeSet']) apply (drule gState_progress_rel_gState_invI2) apply assumption apply assumption apply (refine_rcg refine_vcg nfoldli_rule[where I="\_ _ \. \g' \ ls.\ \. (g,g') \ gState_progress_rel prog"] order_trans[OF applyEdge_gState_progress_rel]) apply (vc_solve intro: refine_helpers solve: asm_rl simp add: ls.correct) apply (rule order_trans) apply (rprems) apply (vc_solve intro: refine_helpers solve: asm_rl simp add: ls.correct) apply (rule order_trans) apply rprems apply (vc_solve intro: refine_helpers solve: asm_rl simp add: ls.correct) done lemma RETURN_dRETURN: "RETURN f \ f' \ nres_of (dRETURN f) \ f'" unfolding nres_of_def by simp lemma executable_dRETURN: "nres_of (dRETURN (executable_impl prog g)) \ executable prog g" using executable_impl.refine by (simp add: RETURN_dRETURN) lemma applyEdge_dRETURN: "nres_of (dRETURN (applyEdge_impl prog e p g)) \ applyEdge prog e p g" using applyEdge_impl.refine by (simp add: RETURN_dRETURN) schematic_goal nexts_code_aux: "nres_of (?nexts prog g) \ nexts prog g" unfolding nexts_def by (refine_transfer the_resI executable_dRETURN applyEdge_dRETURN) concrete_definition nexts_code_aux for prog g uses nexts_code_aux prepare_code_thms nexts_code_aux_def subsubsection \Handle non-termination\ text \ A Promela model may include non-terminating parts. Therefore we cannot guarantee, that @{const nexts} will actually terminate. To avoid having to deal with this in the model checker, we fail in case of non-termination. \ (* TODO: Integrate such a concept into refine_transfer! *) definition SUCCEED_abort where "SUCCEED_abort msg dm m = ( case m of RES X \ if X={} then Code.abort msg (\_. dm) else RES X | _ \ m)" definition dSUCCEED_abort where "dSUCCEED_abort msg dm m = ( case m of dSUCCEEDi \ Code.abort msg (\_. dm) | _ \ m)" definition ref_succeed where "ref_succeed m m' \ m \ m' \ (m=SUCCEED \ m'=SUCCEED)" lemma dSUCCEED_abort_SUCCEED_abort: "\ RETURN dm' \ dm; ref_succeed (nres_of m') m \ \ nres_of (dSUCCEED_abort msg (dRETURN dm') (m')) \ SUCCEED_abort msg dm m" unfolding dSUCCEED_abort_def SUCCEED_abort_def ref_succeed_def by (auto split: dres.splits nres.splits) text \The final successor function now incorporates: \begin{enumerate} \item @{const Promela.nexts} \item handling of non-termination \end{enumerate}\ definition nexts_code where "nexts_code prog g = the_res (dSUCCEED_abort (STR ''The Universe is broken!'') (dRETURN (ls.sng g)) (nexts_code_aux prog g))" lemma nexts_code_SPEC: assumes "gState_inv prog g" and "program_inv prog" shows "g' \ ls.\ (nexts_code prog g) \ (g,g') \ gState_progress_rel prog" unfolding nexts_code_def unfolding dSUCCEED_abort_def using assms using order_trans[OF nexts_code_aux.refine nexts_SPEC[OF assms(1,2)]] by (auto split: dres.splits simp: ls.correct) subsection \Finiteness of the state space\ inductive_set reachable_states for P :: program and g\<^sub>s :: gState \ \start state\ where "g\<^sub>s \ reachable_states P g\<^sub>s" | "g \ reachable_states P g\<^sub>s \ x \ ls.\ (nexts_code P g) \ x \ reachable_states P g\<^sub>s" lemmas reachable_states_induct[case_names init step] = reachable_states.induct[split_format (complete)] lemma reachable_states_finite: assumes "program_inv prog" and "gState_inv prog g" shows "finite (reachable_states prog g)" proof (rule finite_subset[OF _ gStates_finite[of _ g]]) define INV where "INV g' \ g' \ (gState_progress_rel prog)\<^sup>* `` {g} \ gState_inv prog g'" for g' { fix g' have "g' \ reachable_states prog g \ INV g'" proof (induct rule: reachable_states_induct) case init with assms show ?case by (simp add: INV_def) next case (step g g') from step(2,3) have "(g, g') \ gState_progress_rel prog" using nexts_code_SPEC[OF _ \program_inv prog\] unfolding INV_def by auto thus ?case using step(2) unfolding INV_def by auto qed } thus "reachable_states prog g \ (gState_progress_rel prog)\<^sup>* `` {g}" unfolding INV_def by auto qed subsection \Traces\ text \When trying to generate a lasso, we have a problem: We only have a list of global states. But what are the transitions to come from one to the other? This problem shall be tackled by @{term replay}: Given two states, it generates a list of transitions that was taken.\ (* Give a list of edges that lead from g\<^sub>1 to g\<^sub>2 *) definition replay :: "program \ gState \ gState \ choices nres" where "replay prog g\<^sub>1 g\<^sub>2 = ( let g\<^sub>1 = to\<^sub>I g\<^sub>1; check = \g. from\<^sub>I g = g\<^sub>2 in REC\<^sub>T (\D g. do { E \ executable (states prog) g; if E = [] then if check g then RETURN [] else if \ timeout g then D (g\timeout := True\) else abort STR ''Stuttering should not occur on replay'' (\_. RETURN []) else let g = reset\<^sub>I g in nfoldli E (\E. E = []) (\(e,p) _. applyEdge prog e p g \ (\g'. if handshake g' \ 0 \ isAtomic e then do { E\<^sub>R \ D g'; if E\<^sub>R = [] then if check g' then RETURN [(e,p)] else RETURN [] else RETURN ((e,p) # E\<^sub>R) } else if check g' then RETURN [(e,p)] else RETURN []) ) [] }) g\<^sub>1 )" lemma abort_refine[refine_transfer]: "nres_of (f ()) \ F () \ nres_of (abort s f) \ abort s F" "f() \ dSUCCEED \ abort s f \ dSUCCEED" by auto schematic_goal replay_code_aux: "RETURN (?replay prog g\<^sub>1 g\<^sub>2) \ replay prog g\<^sub>1 g\<^sub>2" unfolding replay_def applyEdge_def by (refine_transfer the_resI executable_dRETURN) concrete_definition replay_code for prog g\<^sub>1 g\<^sub>2 uses replay_code_aux prepare_code_thms replay_code_def subsubsection \Printing of traces\ (* Might go to another theory *) definition procDescr :: "(integer \ string) \ program \ pState \ string" where "procDescr f prog p = ( let name = String.explode (proc_names prog !! pState.idx p); id = f (integer_of_nat (pid p)) in name @ '' ('' @ id @ '')'')" definition printInitial :: "(integer \ string) \ program \ gState \ string" where "printInitial f prog g\<^sub>0 = ( let psS = printList (procDescr f prog) (procs g\<^sub>0) [] [] '' '' in ''Initially running: '' @ psS)" abbreviation "lf \ CHR 0x0A" fun printConfig :: "(integer \ string) \ program \ gState option \ gState \ string" where "printConfig f prog None g\<^sub>0 = printInitial f prog g\<^sub>0" | "printConfig f prog (Some g\<^sub>0) g\<^sub>1 = ( let eps = replay_code prog g\<^sub>0 g\<^sub>1 in let print = (\(e,p). procDescr f prog p @ '': '' @ printEdge f (pc p) e) in if eps = [] \ g\<^sub>1 = g\<^sub>0 then '' -- stutter --'' else printList print eps [] [] (lf#'' ''))" definition "printConfigFromAST f \ printConfig f o fst o setUp" subsection \Code export\ code_identifier code_module "PromelaInvariants" \ (SML) Promela | code_module "PromelaDatastructures" \ (SML) Promela definition "executable_triv prog g = executable_impl (snd prog) g" definition "apply_triv prog g ep = applyEdge_impl prog (fst ep) (snd ep) (reset\<^sub>I g)" export_code setUp printProcesses printConfigFromAST nexts_code executable_triv apply_triv extractLTLs lookupLTL checking SML export_code setUp printProcesses printConfigFromAST nexts_code executable_triv apply_triv extractLTLs lookupLTL in SML file \Promela.sml\ end diff --git a/thys/Subresultants/Binary_Exponentiation.thy b/thys/Subresultants/Binary_Exponentiation.thy --- a/thys/Subresultants/Binary_Exponentiation.thy +++ b/thys/Subresultants/Binary_Exponentiation.thy @@ -1,59 +1,59 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Binary Exponentiation\ text \This theory defines the standard algorithm for binary exponentiation, or exponentiation by squaring.\ theory Binary_Exponentiation imports Main begin -declare divmod_nat_def[termination_simp] +declare Euclidean_Division.divmod_nat_def[termination_simp] context monoid_mult begin fun binary_power :: "'a \ nat \ 'a" where "binary_power x n = (if n = 0 then 1 else - let (d,r) = Divides.divmod_nat n 2; + let (d,r) = Euclidean_Division.divmod_nat n 2; rec = binary_power (x * x) d in if r = 0 then rec else rec * x)" lemma binary_power[simp]: "binary_power = (^)" proof (intro ext) fix x n show "binary_power x n = x ^ n" proof (induct x n rule: binary_power.induct) case (1 x n) show ?case proof (cases "n = 0") case False note IH = 1[OF False] - obtain d r where n2: "Divides.divmod_nat n 2 = (d,r)" by force - from divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto + obtain d r where n2: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force + from Euclidean_Division.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto hence r: "r = 0 \ r = 1" by auto let ?rec = "binary_power (x * x) d" have "binary_power x n = (if r = 0 then ?rec else ?rec * x)" unfolding binary_power.simps[of x n] n2 using False by auto also have "\ = ?rec * x ^ r" using r by (cases "r = 0", auto) also have "?rec = (x * x) ^ d" by (rule IH[OF _ refl], simp add: n2) also have "\ = x ^ (d + d)" unfolding power_add using power2_eq_square power_even_eq power_mult by auto also have "\ * x ^ r = x ^ (d + d + r)" by (simp add: power_add) also have "d + d + r = n" unfolding dr by presburger finally show ?thesis . qed auto qed qed lemma binary_power_code_unfold[code_unfold]: "(^) = binary_power" by simp declare binary_power.simps[simp del] end end diff --git a/thys/Subresultants/Dichotomous_Lazard.thy b/thys/Subresultants/Dichotomous_Lazard.thy --- a/thys/Subresultants/Dichotomous_Lazard.thy +++ b/thys/Subresultants/Dichotomous_Lazard.thy @@ -1,173 +1,173 @@ section \Dichotomous Lazard\ text \This theory contains Lazard's optimization in the computation of the subresultant PRS as described by Ducos \cite[Section 2]{Ducos}.\ theory Dichotomous_Lazard imports "HOL-Computational_Algebra.Polynomial_Factorial" (* to_fract *) begin lemma power_fract[simp]: "(Fract a b)^n = Fract (a^n) (b^n)" by (induct n, auto simp: fract_collapse) lemma range_to_fract_dvd_iff: assumes b: "b \ 0" shows "Fract a b \ range to_fract \ b dvd a" proof assume "b dvd a" then obtain c where a: "a = b * c" unfolding dvd_def by auto have "Fract a b = Fract c 1" using b unfolding a by (simp add: eq_fract) thus "Fract a b \ range to_fract" unfolding to_fract_def by auto next assume "Fract a b \ range to_fract" then obtain c where "Fract a b = Fract c 1" unfolding to_fract_def by auto hence "a = b * c" using b by (simp add: eq_fract) thus "b dvd a" .. qed lemma Fract_cases_coprime [cases type: fract]: fixes q :: "'a :: factorial_ring_gcd fract" obtains (Fract) a b where "q = Fract a b" "b \ 0" "coprime a b" proof - obtain a b where q: "q = Fract a b" and b0: "b \ 0" by (cases q, auto) define g where g: "g = gcd a b" define A where A: "A = a div g" define B where B: "B = b div g" have a: "a = A * g" unfolding A g by simp have b: "b = B * g" unfolding B g by simp from b0 b have 0: "B \ 0" by auto have q: "q = Fract A B" unfolding q a b by (subst eq_fract, auto simp: b0 0 g) have cop: "coprime A B" unfolding A B g using b0 by (simp add: div_gcd_coprime) assume "\a b. q = Fract a b \ b \ 0 \ coprime a b \ thesis" from this[OF q 0 cop] show ?thesis . qed lemma to_fract_power_le: fixes a :: "'a :: factorial_ring_gcd fract" assumes no_fract: "a * b ^ e \ range to_fract" and a: "a \ range to_fract" and le: "f \ e" shows "a * b ^ f \ range to_fract" proof - obtain bn bd where b: "b = Fract bn bd" and bd: "bd \ 0" and copb: "coprime bn bd" by (cases b, auto) obtain an where a: "a = Fract an 1" using a unfolding to_fract_def by auto have id: "a * b ^ e = Fract (an * bn^e) (bd^e)" unfolding a b power_fract mult_fract by simp have 0: "bd^e \ 0" for e using bd by auto from no_fract[unfolded id range_to_fract_dvd_iff[OF 0]] have dvd: "bd ^ e dvd an * bn ^ e" . from copb have copb: "coprime (bd ^ e) (bn ^ e)" for e by (simp add: ac_simps) from dvd copb [of e] bd have "bd ^ e dvd an" by (simp add: coprime_dvd_mult_left_iff) hence "bd ^ f dvd an" using le by (rule power_le_dvd) hence dvd: "bd ^ f dvd an * bn ^ f" by simp from le obtain g where e: "e = f + g" using le_Suc_ex by blast have id': "a * b ^ f = Fract (an * bn^f) (bd^f)" unfolding a b power_fract mult_fract by simp show ?thesis unfolding id' range_to_fract_dvd_iff[OF 0] by (rule dvd) qed lemma div_divide_to_fract: assumes "x \ range to_fract" and "x = (y :: 'a :: idom_divide fract) / z" and "x' = y' div z'" and "y = to_fract y'" "z = to_fract z'" shows "x = to_fract x'" proof (cases "z' = 0") case True thus ?thesis using assms by auto next case False from assms obtain r where "to_fract y' / to_fract z' = to_fract r" by auto thus ?thesis using False assms by (simp add: eq_fract(1) to_fract_def) qed -declare divmod_nat_def[termination_simp] +declare Euclidean_Division.divmod_nat_def [termination_simp] fun dichotomous_Lazard :: "'a :: idom_divide \ 'a \ nat \ 'a" where "dichotomous_Lazard x y n = (if n \ 1 then if n = 1 then x else 1 else - let (d,r) = Divides.divmod_nat n 2; + let (d,r) = Euclidean_Division.divmod_nat n 2; rec = dichotomous_Lazard x y d; recsq = rec * rec div y in if r = 0 then recsq else recsq * x div y)" lemma dichotomous_Lazard_main: fixes x :: "'a :: idom_divide" assumes "\ i. i \ n \ (to_fract x)^i / (to_fract y)^(i - 1) \ range to_fract" shows "to_fract (dichotomous_Lazard x y n) = (to_fract x)^n / (to_fract y)^(n-1)" using assms proof (induct x y n rule: dichotomous_Lazard.induct) case (1 x y n) let ?f = to_fract consider (0) "n = 0" | (1) "n = 1" | (n) "\ n \ 1" by linarith thus ?case proof cases case n - obtain d r where n2: "Divides.divmod_nat n 2 = (d,r)" by force - from divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto + obtain d r where n2: "Euclidean_Division.divmod_nat n 2 = (d,r)" by force + from Euclidean_Division.divmod_nat_def[of n 2] n2 have dr: "d = n div 2" "r = n mod 2" by auto hence r: "r = 0 \ r = 1" by auto define rec where "rec = dichotomous_Lazard x y d" let ?sq = "rec * rec div y" have res: "dichotomous_Lazard x y n = (if r = 0 then ?sq else ?sq * x div y)" unfolding dichotomous_Lazard.simps[of x y n] n2 Let_def rec_def using n by auto have ndr: "n = d + d + r" unfolding dr by presburger from ndr r n have d0: "d \ 0" by auto have IH: "?f rec = ?f x ^ d / ?f y ^ (d - 1)" using 1(1)[OF n refl n2[symmetric] 1(2), folded rec_def] ndr by auto have "?f (rec * rec) = ?f x ^ d / ?f y ^ (d - 1) * ?f x ^ d / ?f y ^ (d - 1)" using IH by simp also have "\ = ?f x ^ (d + d) / ?f y ^ (d - 1 + (d - 1))" unfolding power_add by simp also have "d - 1 + (d - 1) = d + d - 2" using d0 by simp finally have id: "?f (rec * rec) = ?f x ^ (d + d) / ?f y ^ (d + d - 2)" . let ?dd = "(?f x ^ (d + d) / ?f y ^ (d + d - 2)) / ?f y" let ?d = "?f x ^ (d + d) / ?f y ^ (d + d - 1)" have dd: "?dd = ?d" using d0 by (cases d, auto) have sq: "?f ?sq = ?d" unfolding dd[symmetric] proof (rule sym, rule div_divide_to_fract[OF _ refl refl id[symmetric] refl], unfold dd) show "?d \ range ?f" by (rule 1(2), insert ndr, auto) qed show ?thesis proof (cases "r = 0") case True with res sq show ?thesis unfolding ndr by auto next case False with r have r: "r = 1" by auto let ?sq' = "?sq * x div y" from False res have res: "dichotomous_Lazard x y n = ?sq'" by simp from sq have id: "?f (?sq * x) = ?f x ^ (d + d + r) / ?f y ^ (d + d - 1)" unfolding r by simp let ?dd = "(?f x ^ (d + d + r) / ?f y ^ (d + d - 1)) / ?f y" let ?d = "?f x ^ (d + d + r) / ?f y ^ (d + d + r - 1)" have dd: "?dd = ?d" using d0 unfolding r by (cases d, auto) have sq': "?f ?sq' = ?d" unfolding dd[symmetric] proof (rule sym, rule div_divide_to_fract[OF _ refl refl id[symmetric] refl], unfold dd) show "?d \ range ?f" by (rule 1(2), unfold ndr, auto) qed show ?thesis unfolding res sq' unfolding ndr by simp qed qed auto qed lemma dichotomous_Lazard: fixes x :: "'a :: factorial_ring_gcd" assumes "(to_fract x)^n / (to_fract y)^(n-1) \ range to_fract" shows "to_fract (dichotomous_Lazard x y n) = (to_fract x)^n / (to_fract y)^(n-1)" proof (rule dichotomous_Lazard_main) fix i assume i: "i \ n" show "to_fract x ^ i / to_fract y ^ (i - 1) \ range to_fract" proof (cases i) case (Suc j) have id: "to_fract x ^ i / to_fract y ^ (i - 1) = to_fract x * (to_fract x / to_fract y) ^ j" unfolding Suc by (simp add: power_divide) from Suc i have "n \ 0" and j: "j \ n - 1" by auto hence idd: "to_fract x * (to_fract x / to_fract y) ^ (n - 1) = (to_fract x)^n / (to_fract y)^(n-1)" by (cases n, auto simp: power_divide) show ?thesis unfolding id by (rule to_fract_power_le[OF _ _ j], unfold idd, insert assms, auto) next case 0 have "1 = to_fract 1" by simp hence "1 \ range to_fract" by blast thus ?thesis using 0 by auto qed qed declare dichotomous_Lazard.simps[simp del] end